12345678910111213141516171819202122232425262728293031323334353637383940414243444546474849505152535455565758596061626364656667686970717273747576777879808182838485868788899091929394959697989910010110210310410510610710810911011111211311411511611711811912012112212312412512612712812913013113213313413513613713813914014114214314414514614714814915015115215315415515615715815916016116216316416516616716816917017117217317417517617717817918018118218318418518618718818919019119219319419519619719819920020120220320420520620720820921021121221321421521621721821922022122222322422522622722822923023123223323423523623723823924024124224324424524624724824925025125225325425525625725825926026126226326426526626726826927027127227327427527627727827928028128228328428528628728828929029129229329429529629729829930030130230330430530630730830931031131231331431531631731831932032132232332432532632732832933033133233333433533633733833934034134234334434534634734834935035135235335435535635735835936036136236336436536636736836937037137237337437537637737837938038138238338438538638738838939039139239339439539639739839940040140240340440540640740840941041141241341441541641741841942042142242342442542642742842943043143243343443543643743843944044144244344444544644744844945045145245345445545645745845946046146246346446546646746846947047147247347447547647747847948048148248348448548648748848949049149249349449549649749849950050150250350450550650750850951051151251351451551651751851952052152252352452552652752852953053153253353453553653753853954054154254354454554654754854955055155255355455555655755855956056156256356456556656756856957057157257357457557657757857958058158258358458558658758858959059159259359459559659759859960060160260360460560660760860961061161261361461561661761861962062162262362462562662762862963063163263363463563663763863964064164264364464564664764864965065165265365465565665765865966066166266366466566666766866967067167267367467567667767867968068168268368468568668768868969069169269369469569669769869970070170270370470570670770870971071171271371471571671771871972072172272372472572672772872973073173273373473573673773873974074174274374474574674774874975075175275375475575675775875976076176276376476576676776876977077177277377477577677777877978078178278378478578678778878979079179279379479579679779879980080180280380480580680780880981081181281381481581681781881982082182282382482582682782882983083183283383483583683783883984084184284384484584684784884985085185285385485585685785885986086186286386486586686786886987087187287387487587687787887988088188288388488588688788888989089189289389489589689789889990090190290390490590690790890991091191291391491591691791891992092192292392492592692792892993093193293393493593693793893994094194294394494594694794894995095195295395495595695795895996096196296396496596696796896997097197297397497597697797897998098198298398498598698798898999099199299399499599699799899910001001100210031004100510061007100810091010101110121013101410151016101710181019102010211022102310241025102610271028102910301031103210331034103510361037103810391040104110421043104410451046104710481049105010511052105310541055105610571058105910601061106210631064106510661067106810691070107110721073107410751076107710781079108010811082108310841085108610871088108910901091109210931094109510961097109810991100110111021103110411051106110711081109111011111112111311141115111611171118111911201121112211231124112511261127112811291130113111321133113411351136113711381139114011411142114311441145114611471148114911501151115211531154115511561157115811591160116111621163116411651166116711681169117011711172117311741175117611771178117911801181118211831184118511861187118811891190119111921193119411951196119711981199120012011202120312041205120612071208120912101211121212131214121512161217121812191220122112221223122412251226122712281229123012311232123312341235123612371238123912401241124212431244124512461247124812491250125112521253125412551256125712581259126012611262126312641265126612671268126912701271127212731274127512761277127812791280128112821283128412851286128712881289129012911292129312941295129612971298129913001301130213031304130513061307130813091310131113121313131413151316131713181319132013211322132313241325132613271328132913301331133213331334133513361337133813391340134113421343134413451346134713481349135013511352135313541355135613571358135913601361136213631364136513661367136813691370(************************************************************************)(* * 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) *)(************************************************************************)(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)openLtac_pluginopenNamesopenPpopenGenargopenTermopenContextmoduleCoqConstr=ConstropenCoqConstropenVarsopenLibnamesopenTacticsopenTermopsopenTacmach.OldopenGlob_termopenUtilopenEvdopenTacexpropenTacinterpopenPretypingopenPpconstropenPrinteropenNamegenopenEvar_kindsopenConstrexpropenConstrexpr_opstypessrtermkind=|InParens|WithAt|NoFlag|Cpatternleterrorstrm=CErrors.user_errletloc_errorlocmsg=CErrors.user_err?loc(strmsg)letppnl=Feedback.msg_info(* 0 cost pp function. Active only if env variable SSRDEBUG is set *)(* or if SsrDebug is Set *)letpp_ref=ref(fun_->())letssr_pps=Feedback.msg_debug(str"SSR: "++Lazy.forces)let_=tryignore(Sys.getenv"SSRMATCHINGDEBUG");pp_ref:=ssr_ppwithNot_found->()letdebugb=ifbthenpp_ref:=ssr_ppelsepp_ref:=fun_->()let_=Goptions.declare_bool_option{Goptions.optkey=["Debug";"SsrMatching"];Goptions.optdepr=false;Goptions.optread=(fun_->!pp_ref==ssr_pp);Goptions.optwrite=debug}letpps=!pp_refs(** Utils *)(* {{{ *****************************************************************)letenv_sizeenv=List.length(Environ.named_contextenv)letsafeDestAppc=matchkindcwithApp(f,a)->f,a|_->c,[||](* Toplevel constr must be globalized twice ! *)letglob_constristgenvsigmat=matcht,istwith|(_,Somece),Someist->letvars=Id.Map.fold(funx_accu->Id.Set.addxaccu)ist.lfunId.Set.emptyinletltacvars={Constrintern.empty_ltac_signwithConstrintern.ltac_vars=vars}inConstrintern.intern_genWithoutTypeConstraint~ltacvars:ltacvarsgenvsigmace|(rc,None),_->rc|(_,Some_),None->CErrors.anomalyPp.(str"glob_constr: term with no ist")(* Term printing utilities functions for deciding bracketing. *)letpr_parenprxx=hov1(str"("++prxx++str")")(* String lexing utilities *)letskip_wscharss=letrecloopi=matchs.[i]with'\n'..' '->loop(i+1)|_->iinloop(* We also guard characters that might interfere with the ssreflect *)(* tactic syntax. *)letguard_termkindsi=matchs.[i]with|'('->false|'{'|'/'|'='->true|_->kind=InParens(* The call 'guard s i' should return true if the contents of s *)(* starting at i need bracketing to avoid ambiguities. *)letpr_guardedguardprcc=lets=Pp.string_of_ppcmds(prcc)^"$"inifguards(skip_wscharss0)thenpr_parenprccelseprcc(* More sensible names for constr printers *)letwith_global_env_evmfx=letenv=Global.env()inletsigma=Evd.from_envenvinfenvsigmaxletprl_glob_constr=with_global_env_evmpr_lglob_constr_envletpr_glob_constr=with_global_env_evmpr_glob_constr_envletprl_constr_expr=pr_lconstr_exprletpr_constr_expr=pr_constr_exprletprl_glob_constr_and_exprenvsigma=function|_,Somec->prl_constr_exprenvsigmac|c,None->prl_glob_constrcletpr_glob_constr_and_exprenvsigma=function|_,Somec->pr_constr_exprenvsigmac|c,None->pr_glob_constrc(** Adding a new uninterpreted generic argument type *)letadd_genargtagpr=letwit=Genarg.make0taginlettag=Geninterp.Val.createtaginletglobistx=(ist,x)inletsubst_x=xinletinterpistx=Ftactic.return(Geninterp.Val.Dyn(tag,x))inletgen_prenvsigma___=prenvsigmainlet()=Genintern.register_intern0witglobinlet()=Genintern.register_subst0witsubstinlet()=Geninterp.register_interp0witinterpinlet()=Geninterp.register_val0wit(Some(Geninterp.Val.Basetag))inPptactic.declare_extra_genarg_pprulewitgen_prgen_prgen_pr;wit(** Constructors for constr_expr *)letisCVar=function{CAst.v=CRef(qid,_)}->qualid_is_identqid|_->falseletdestCVar=function|{CAst.v=CRef(qid,_)}whenqualid_is_identqid->qualid_basenameqid|_->CErrors.anomaly(str"not a CRef.")letisGLambdac=matchDAst.getcwithGLambda(Name_,_,_,_)->true|_->falseletdestGLambdac=matchDAst.getcwithGLambda(Nameid,_,_,c)->(id,c)|_->CErrors.anomaly(str"not a GLambda")letisGHolec=matchDAst.getcwithGHole_->true|_->falseletmkCHole~loc=CAst.make?loc@@CHole(None,IntroAnonymous,None)letmkCLambda?locnametyt=CAst.make?loc@@CLambdaN([CLocalAssum([CAst.make?locname],DefaultExplicit,ty)],t)letmkCLetIn?locnamebot=CAst.make?loc@@CLetIn((CAst.make?locname),bo,None,t)letmkCCast?loctty=CAst.make?loc@@CCast(t,DEFAULTcast,ty)(** Constructors for rawconstr *)letmkRHole=DAst.make@@GHole(InternalHole,IntroAnonymous,None)letmkRAppfargs=ifargs=[]thenfelseDAst.make@@GApp(f,args)letmkRCastrcrt=DAst.make@@GCast(rc,DEFAULTcast,rt)letmkRLambdanst=DAst.make@@GLambda(n,Explicit,s,t)letnf_evarsigmac=EConstr.Unsafe.to_constr(Evarutil.nf_evarsigma(EConstr.of_constrc))(* }}} *)exceptionNoProgress(** Unification procedures. *)(* To enforce the rigidity of the rooted match we always split *)(* top applications, so the unification procedures operate on *)(* arrays of patterns and terms. *)(* We perform three kinds of unification: *)(* EQ: exact conversion check *)(* FO: first-order unification of evars, without conversion *)(* HO: higher-order unification with conversion *)(* The subterm unification strategy is to find the first FO *)(* match, if possible, and the first HO match otherwise, then *)(* compute all the occurrences that are EQ matches for the *)(* relevant subterm. *)(* Additional twists: *)(* - If FO/HO fails then we attempt to fill evars using *)(* typeclasses before raising an outright error. We also *)(* fill typeclasses even after a successful match, since *)(* beta-reduction and canonical instances may leave *)(* undefined evars. *)(* - We do postchecks to rule out matches that are not *)(* closed or that assign to a global evar; these can be *)(* disabled for rewrite or dependent family matches. *)(* - We do a full FO scan before turning to HO, as the FO *)(* comparison can be much faster than the HO one. *)letunif_EQenvsigmapc=letenv=Environ.set_universes(Evd.universessigma)envinletevars=existential_opt_value0sigmaintrylet_=Reduction.convenvp~evarscintruewith_->falseletunif_EQ_argsenvsigmapaa=letn=Array.lengthpainletrecloopi=(i=n)||unif_EQenvsigmapa.(i)a.(i)&&loop(i+1)inloop0letunif_HOenvisepc=tryEvarconv.unify_delayenvisepcwithEvarconv.UnableToUnify(ise,err)->raisePretype_errors.(PretypeError(env,ise,CannotUnify(p,c,Someerr)))letunif_HO_argsenvise0paica=letn=Array.lengthpainletrecloopisej=ifj=ntheniseelseloop(unif_HOenvise(EConstr.of_constrpa.(j))(EConstr.of_constrca.(i+j)))(j+1)inloopise00(* FO unification should boil down to calling w_unify with no_delta, but *)(* alas things are not so simple: w_unify does partial type-checking, *)(* which breaks down when the no-delta flag is on (as the Coq type system *)(* requires full convertibility. The workaround here is to convert all *)(* evars into metas, since 8.2 does not TC metas. This means some lossage *)(* for HO evars, though hopefully Miller patterns can pick up some of *)(* those cases, and HO matching will mop up the rest. *)letflags_FOenv=letoracle=Environ.oracleenvinletts=Conv_oracle.get_transp_stateoracleinletflags={(Unification.default_no_delta_unify_flagsts).Unification.core_unify_flagswithUnification.modulo_conv_on_closed_terms=None;Unification.modulo_eta=true;Unification.modulo_betaiota=true;Unification.modulo_delta_types=ts}in{Unification.core_unify_flags=flags;Unification.merge_unify_flags=flags;Unification.subterm_unify_flags=flags;Unification.allow_K_in_toplevel_higher_order_unification=false;Unification.resolve_evars=(Unification.default_no_delta_unify_flagsts).Unification.resolve_evars}letunif_FOenvisepc=Unification.w_unifyenviseReduction.CONV~flags:(flags_FOenv)(EConstr.of_constrp)(EConstr.of_constrc)(* Perform evar substitution in main term and prune substitution. *)letnf_open_termsigma0isec=letc=EConstr.Unsafe.to_constrcinlets=iseands'=refsigma0inletrecnfc'=matchkindc'with|Evarex->begintrynf(existential_value0sex)with_->letk,a=exinleta'=List.mapnfainifnot(Evd.mem!s'k)thens':=Evd.add!s'k(Evarutil.nf_evar_infos(Evd.findsk));mkEvar(k,a')end|_->mapnfc'inletcopy_defkevi()=ifevar_bodyevi!=Evd.Evar_emptythen()elsematchEvd.evar_body(Evd.findsk)with|Evar_definedc'->letc'=EConstr.of_constr(nf(EConstr.Unsafe.to_constrc'))ins':=Evd.definekc'!s'|_->()inletc'=nfcinlet_=Evd.foldcopy_defsigma0()in!s',Evd.evar_universe_contexts,EConstr.of_constrc'letunif_end?(solve_TC=true)envsigma0ise0ptok=letise=Evarconv.solve_unif_constraints_with_heuristicsenvise0inlettcs=Evd.get_typeclass_evarsiseinlets,uc,t=nf_open_termsigma0iseptinletise1=create_evar_defssinletise1=Evd.set_typeclass_evarsise1(Evar.Set.filter(funev->Evd.is_undefinedise1ev)tcs)inletise1=Evd.set_universe_contextise1ucinletise2=ifsolve_TCthenTypeclasses.resolve_typeclasses~fail:trueenvise1elseise1inifnot(okise)thenraiseNoProgresselseifise2==ise1then(s,uc,t)elselets,uc',t=nf_open_termsigma0ise2tins,UState.unionucuc',tletunify_HOenvsigma0t1t2=letsigma=unif_HOenvsigma0t1t2inletsigma,uc,_=unif_end~solve_TC:falseenvsigma0sigmat2(fun_->true)inEvd.set_universe_contextsigmaucletpf_unify_HOglt1t2=letenv,sigma0,si=pf_envgl,projectgl,sig_itglinletsigma=unify_HOenvsigma0t1t2inre_sigsisigma(* This is what the definition of iter_constr should be... *)letiter_constr_LRfc=matchkindcwith|Evar(k,a)->List.iterfa|Cast(cc,_,t)->fcc;ft|Prod(_,t,b)|Lambda(_,t,b)->ft;fb|LetIn(_,v,t,b)->fv;ft;fb|App(cf,a)->fcf;Array.iterfa|Case(_,_,pms,(_,p),iv,v,b)->fv;Array.iterfpms;fp;iter_invertfiv;Array.iter(fun(_,c)->fc)b|Fix(_,(_,t,b))|CoFix(_,(_,t,b))->fori=0toArray.lengtht-1doft.(i);fb.(i)done|Proj(_,a)->fa|Array(_u,t,def,ty)->Array.iterft;fdef;fty|(Rel_|Meta_|Var_|Sort_|Const_|Ind_|Construct_|Int_|Float_)->()(* The comparison used to determine which subterms matches is KEYED *)(* CONVERSION. This looks for convertible terms that either have the same *)(* same head constant as pat if pat is an application (after beta-iota), *)(* or start with the same constr constructor (esp. for LetIn); this is *)(* disregarded if the head term is let x := ... in x, and casts are always *)(* ignored and removed). *)(* Record projections get special treatment: in addition to the projection *)(* constant itself, ssreflect also recognizes head constants of canonical *)(* projections. *)exceptionNoMatchtypessrdir=L2R|R2Lletpr_dir_side=functionL2R->str"LHS"|R2L->str"RHS"letinv_dir=functionL2R->R2L|R2L->L2Rtypepattern_class=|KpatFixed|KpatConst|KpatEvarofEvar.t|KpatLet|KpatLam|KpatRigid|KpatFlex|KpatProjofConstant.ttypetpattern={up_k:pattern_class;up_FO:constr;up_f:constr;up_a:constrarray;up_t:constr;(* equation proof term or matched term *)up_dir:ssrdir;(* direction of the rule *)up_ok:constr->evar_map->bool;(* progress test for rewrite *)}letall_ok__=trueletproj_nparamsc=try1+Structures.Structure.projection_nparamscwithNot_found->0letisRigidc=matchkindcwith|(Prod_|Sort_|Lambda_|Case_|Fix_|CoFix_|Int_|Float_|Array_)->true|(Rel_|Var_|Meta_|Evar(_,_)|Cast(_,_,_)|LetIn(_,_,_,_)|App(_,_)|Const(_,_)|Ind((_,_),_)|Construct(((_,_),_),_)|Proj(_,_))->falselethole_var=mkVar(Id.of_string"_")letpr_constr_patenvsigmac0=letrecwipe_evarc=ifisEvarcthenhole_varelsemapwipe_evarcinpr_constr_envenvsigma(wipe_evarc0)letehole_var=EConstr.mkVar(Id.of_string"_")letpr_econstr_patenvsigmac0=letrecwipe_evarc=letopenEConstrinifisEvarsigmacthenehole_varelsemapsigmawipe_evarcinletdummy_decl=letdummy_prod=mkProd(make_annotAnonymousSorts.Relevant,mkProp,mkProp)inletna=make_annot(EConstr.destVarsigmaehole_var)Sorts.RelevantinContext.Named.Declaration.(LocalAssum(na,dummy_prod))inletenv=Environ.push_nameddummy_declenvinpr_econstr_envenvsigma(wipe_evarc0)(* Turn (new) evars into metas *)letevars_for_FO~hackenvsigma0(ise0:evar_map)c0=letise=refise0inletsigma=refise0inletnenv=env_sizeenv+ifhackthen1else0inletrecputc=matchkindcwith|Evar(k,aasex)->begintryput(existential_value0!sigmaex)withNotInstantiatedEvar->ifEvd.memsigma0kthenmapputcelseletevi=Evd.find!sigmakinletdc=List.firstn(max0(List.lengtha-nenv))(evar_filtered_contextevi)inletabs_dc(d,c)=function|Context.Named.Declaration.LocalDef(x,b,t)->d,mkNamedLetInx(putb)(putt)c|Context.Named.Declaration.LocalAssum(x,t)->mkVarx.binder_name::d,mkNamedProdx(putt)cinleta,t=Context.Named.fold_insideabs_dc~init:([],(put@@EConstr.Unsafe.to_constrevi.evar_concl))(EConstr.Unsafe.to_named_contextdc)inletm=Evarutil.new_meta()inise:=meta_declarem(EConstr.of_constrt)!ise;sigma:=Evd.definek(EConstr.of_constr(applistc(mkMetam)a))!sigma;put(existential_value0!sigmaex)end|_->mapputcinletc1=putc0in!ise,c1(* Compile a match pattern from a term; t is the term to fill. *)(* p_origin can be passed to obtain a better error message *)letmk_tpattern?p_origin?(hack=false)envsigma0(ise,t)okdirp=letk,f,a=letf,a=Reductionops.whd_betaiota_stackenvise(EConstr.of_constrp)inletf=EConstr.Unsafe.to_constrfinleta=List.mapEConstr.Unsafe.to_constrainmatchkindfwith|Const(p,_)->letnp=proj_nparamspinifnp=0||np>List.lengthathenKpatConst,f,aelseleta1,a2=List.chopnpainKpatProjp,(applistcfa1),a2|Proj(p,arg)->KpatProj(Projection.constantp),f,a|Var_|Ind_|Construct_->KpatFixed,f,a|Evar(k,_)->ifEvd.memsigma0kthenKpatEvark,f,aelseifa<>[]thenKpatFlex,f,aelse(matchp_originwithNone->CErrors.user_errPp.(str"indeterminate pattern")|Some(dir,rule)->errorstrm(str"indeterminate "++pr_dir_sidedir++str" in "++pr_constr_patenviserule))|LetIn(_,v,_,b)->ifb<>mkRel1thenKpatLet,f,aelseKpatFlex,v,a|Lambda_->KpatLam,f,a|_->KpatRigid,f,ainletaa=Array.of_listainletise',p'=evars_for_FO~hackenvsigma0ise(mkApp(f,aa))inise',{up_k=k;up_FO=p';up_f=f;up_a=aa;up_ok=ok;up_dir=dir;up_t=t}(* Specialize a pattern after a successful match: assign a precise head *)(* kind and arity for Proj and Flex patterns. *)letungen_upatlhs(sigma,uc,t)u=letf,a=safeDestApplhsinletk=matchkindfwith|Var_|Ind_|Construct_->KpatFixed|Const_->KpatConst|Evar(k,_)->ifis_definedsigmakthenraiseNoMatchelseKpatEvark|LetIn_->KpatLet|Lambda_->KpatLam|_->KpatRigidinsigma,uc,{uwithup_k=k;up_FO=lhs;up_f=f;up_a=a;up_t=t}letnb_cs_proj_argsisepcfu=letopenStructuresinletopenValuePatterninletnak=letopenCanonicalSolutioninlet_,{cvalue_arguments}=find(Global.env())ise(GlobRef.ConstRefpc,k)inList.lengthcvalue_argumentsinletnargs_of_projt=matchkindtwith|App(_,args)->Array.lengthargs|Proj_->0(* if splay_app calls expand_projection, this has to be
the number of arguments including the projected *)|_->assertfalseintrymatchkindfwith|Prod_->naProd_cs|Sorts->na(Sort_cs(Sorts.familys))|Const(c',_)whenConstant.CanOrd.equalc'pc->nargs_of_proju.up_f|Proj(c',_)whenConstant.CanOrd.equal(Names.Projection.constantc')pc->nargs_of_proju.up_f|Var_|Ind_|Construct_|Const_->na(Const_cs(fst@@destReff))|_->-1withNot_found->-1letisEvar_kkf=matchkindfwithEvar(k',_)->k=k'|_->falseletnb_argsc=matchkindcwithApp(_,a)->Array.lengtha|_->0letmkSubArgia=ifi=Array.lengthathenaelseArray.suba0iletmkSubAppfia=ifi=0thenfelsemkApp(f,mkSubArgia)letsplay_appise=letrecloopca=matchkindcwith|App(f,a')->loopf(Array.appenda'a)|Cast(c',_,_)->loopc'a|Evarex->(tryloop(existential_value0iseex)awith_->c,a)|_->c,ainfunc->matchkindcwith|App(f,a)->loopfa|Cast_|Evar_->loopc[||]|_->c,[||]letfilter_upatisei0fnufpats=letna=Array.lengthu.up_ainifn<nathenfpatselseletnp=matchu.up_kwith|KpatConstwheneq_constr_nounivsu.up_ff->na|KpatFixedwheneq_constr_nounivsu.up_ff->na|KpatEvarkwhenisEvar_kkf->na|KpatLetwhenisLetInf->na|KpatLamwhenisLambdaf->na|KpatRigidwhenisRigidf->na|KpatFlex->na|KpatProjpc->letnp=na+nb_cs_proj_argsisepcfuinifn<npthen-1elsenp|_->-1inifnp<nathenfpatselselet()=if!i0<nptheni0:=nin(u,np)::fpatsleteq_prim_projct=matchkindtwith|Proj(p,_)->Constant.CanOrd.equal(Projection.constantp)c|_->falseletfilter_upat_FOi0fnufpats=letnp=nb_argsu.up_FOinifn<npthenfpatselseletok=matchu.up_kwith|KpatConst->eq_constr_nounivsu.up_ff|KpatFixed->eq_constr_nounivsu.up_ff|KpatEvark->isEvar_kkf|KpatLet->isLetInf|KpatLam->isLambdaf|KpatRigid->isRigidf|KpatProjpc->equalf(mkConstpc)||eq_prim_projpcf|KpatFlex->i0:=n;trueinifokthenbeginif!i0<nptheni0:=np;(u,np)::fpatsendelsefpatsexceptionFoundUnifof(evar_map*UState.t*tpattern)(* Note: we don't update env as we descend into the term, as the primitive *)(* unification procedure always rejects subterms with bound variables. *)letdont_impact_evars_insigma0cl=letevs_in_cl=Evd.evars_of_termsigma0clinfunsigma->Evar.Set.for_all(funk->trylet_=Evd.find_undefinedsigmakintruewithNot_found->false)evs_in_cl(* We are forced to duplicate code between the FO/HO matching because we *)(* have to work around several kludges in unify.ml: *)(* - w_unify drops into second-order unification when the pattern is an *)(* application whose head is a meta. *)(* - w_unify tries to unify types without subsumption when the pattern *)(* head is an evar or meta (e.g., it fails on ?1 = nat when ?1 : Type). *)(* - w_unify expands let-in (zeta conversion) eagerly, whereas we want to *)(* match a head let rigidly. *)letmatch_upats_FOupatsenvsigma0iseorig_c=letdont_impact_evars=dont_impact_evars_insigma0(EConstr.of_constrorig_c)inletrecloopc=letf,a=splay_appisecinleti0=ref(-1)inletfpats=List.fold_right(filter_upat_FOi0f(Array.lengtha))upats[]inwhile!i0>=0doleti=!i0ini0:=-1;letc'=mkSubAppfiainletone_match(u,np)=letskip=ifi<=nptheni<npelseifu.up_k==KpatFlexthenbegini0:=i-1;falseendelsebeginif!i0<nptheni0:=np;trueendinifskip||not(closed0c')then()elsetrylet_=matchu.up_kwith|KpatFlex->letkludgev=mkLambda(make_annotAnonymousSorts.Relevant,mkProp,v)inunif_FOenvise(kludgeu.up_FO)(kludgec')|KpatLet->letkludgevla=letvl,a=safeDestAppvlainletx,v,t,b=destLetInvlinmkApp(mkLambda(x,t,b),Array.consva)inunif_FOenvise(kludgeu.up_FO)(kludgec')|_->unif_FOenviseu.up_FOc'inletise'=(* Unify again using HO to assign evars *)letp=mkApp(u.up_f,u.up_a)intryunif_HOenvise(EConstr.of_constrp)(EConstr.of_constrc')withewhenCErrors.noncriticale->raiseNoMatchinletlhs=mkSubAppfiainletpt'=unif_endenvsigma0ise'(EConstr.of_constru.up_t)(u.up_oklhs)inletpt'=pi1pt',pi2pt',EConstr.Unsafe.to_constr(pi3pt')inraise(FoundUnif(ungen_upatlhspt'u))withFoundUnif(s,_,_)assig_uwhendont_impact_evarss->raisesig_u|Not_found->CErrors.anomaly(str"incomplete ise in match_upats_FO.")|ewhenCErrors.noncriticale->()inList.iterone_matchfpatsdone;iter_constr_LRloopf;Array.iterloopaintrylooporig_cwithInvalid_argument_->CErrors.anomaly(str"IN FO.")letmatch_upats_HO~on_instanceupatsenvsigma0isec=letdont_impact_evars=dont_impact_evars_insigma0(EConstr.of_constrc)inletit_did_match=reffalseinletfailed_because_of_TC=reffalseinletrecauxupatsenvsigma0isec=letf,a=splay_appisecinleti0=ref(-1)inletfpats=List.fold_right(filter_upatisei0f(Array.lengtha))upats[]inwhile!i0>=0doleti=!i0ini0:=-1;letone_match(u,np)=letskip=ifi<=nptheni<npelseifu.up_k==KpatFlexthenbegini0:=i-1;falseendelsebeginif!i0<nptheni0:=np;trueendinifskipthen()elsetryletise'=matchu.up_kwith|KpatFixed|KpatConst->ise|KpatEvar_->let_,pka=destEvaru.up_fand_,ka=destEvarfinletfoldisepkk=unif_HOenvise(EConstr.of_constrpk)(EConstr.of_constrk)inList.fold_left2foldisepkaka|KpatLet->letx,v,t,b=destLetInfinlet_,pv,_,pb=destLetInu.up_finletise'=unif_HOenvise(EConstr.of_constrpv)(EConstr.of_constrv)inunif_HO(Environ.push_rel(Context.Rel.Declaration.LocalAssum(x,t))env)ise'(EConstr.of_constrpb)(EConstr.of_constrb)|KpatFlex|KpatProj_->unif_HOenvise(EConstr.of_constru.up_f)(EConstr.of_constr(mkSubAppf(i-Array.lengthu.up_a)a))|_->unif_HOenvise(EConstr.of_constru.up_f)(EConstr.of_constrf)inletise''=unif_HO_argsenvise'u.up_a(i-Array.lengthu.up_a)ainletlhs=mkSubAppfiainletpt'=unif_endenvsigma0ise''(EConstr.of_constru.up_t)(u.up_oklhs)inletpt'=pi1pt',pi2pt',EConstr.Unsafe.to_constr(pi3pt')inon_instance(ungen_upatlhspt'u)withFoundUnif(s,_,_)assig_uwhendont_impact_evarss->raisesig_u|NoProgress->it_did_match:=true|Pretype_errors.PretypeError(_,_,Pretype_errors.UnsatisfiableConstraints_)->failed_because_of_TC:=true|ewhenCErrors.noncriticale->()inList.iterone_matchfpatsdone;iter_constr_LR(auxupatsenvsigma0ise)f;Array.iter(auxupatsenvsigma0ise)ainauxupatsenvsigma0isec;if!it_did_matchthenraiseNoProgress;!failed_because_of_TCletfixed_upatevd=function|{up_k=KpatFlex|KpatEvar_|KpatProj_}->false|{up_t=t}->not(occur_existentialevd(EConstr.of_constrt))(** FIXME *)letdo_oncerf=match!rwithSome_->()|None->r:=Some(f())letassert_doner=match!rwithSomex->x|None->CErrors.anomaly(str"do_once never called.")letassert_done_multiresr=match!rwith|None->CErrors.anomaly(str"do_once never called.")|Some(e,n,xs)->r:=Some(e,n+1,xs);tryList.nthxsnwithFailure_->raiseNoMatchtypesubst=Environ.env->constr->constr->int->constrtypefind_P=Environ.env->constr->int->k:subst->constrtypeconclude=unit->constr*ssrdir*(Evd.evar_map*UState.t*constr)(* upats_origin makes a better error message only *)letmk_tpattern_matcher?(all_instances=false)?(raise_NoMatch=false)?upats_originsigma0occ(ise,upats)=letnocc=ref0andskip_occ=reffalseinletuse_occ,occ_list=matchoccwith|Some(true,ol)->ol=[],ol|Some(false,ol)->ol<>[],ol|None->false,[]inletmax_occ=List.fold_rightmaxocc_list0inletsubst_occ=letocc_set=Array.makemax_occ(notuse_occ)inlet_=List.iter(funi->occ_set.(i-1)<-use_occ)occ_listinlet_=ifmax_occ=0thenskip_occ:=use_occinfun()->incrnocc;if!nocc=max_occthenskip_occ:=use_occ;if!nocc<=max_occthenocc_set.(!nocc-1)elsenotuse_occinletupat_that_matched=refNoneinletmatch_EQenvsigmau=matchu.up_kwith|KpatLet->letx,pv,t,pb=destLetInu.up_finletenv'=Environ.push_rel(Context.Rel.Declaration.LocalAssum(x,t))envinletmatch_letf=matchkindfwith|LetIn(_,v,_,b)->unif_EQenvsigmapvv&&unif_EQenv'sigmapbb|_->falseinmatch_let|KpatFixed->eq_constr_nounivsu.up_f|KpatConst->eq_constr_nounivsu.up_f|KpatLam->func->(matchkindcwith|Lambda_->unif_EQenvsigmau.up_fc|_->false)|_->unif_EQenvsigmau.up_finletp2tp=mkApp(p.up_f,p.up_a)inletsourceenv=matchupats_origin,upatswith|None,[p]->(iffixed_upatisepthenstr"term "elsestr"partial term ")++pr_constr_patenvise(p2tp)++spc()|Some(dir,rule),[p]->str"The "++pr_dir_sidedir++str" of "++pr_constr_patenviserule++fnl()++ws4++pr_constr_patenvise(p2tp)++fnl()|Some(dir,rule),_->str"The "++pr_dir_sidedir++str" of "++pr_constr_patenviserule++spc()|_,[]|None,_::_::_->CErrors.anomaly(str"mk_tpattern_matcher with no upats_origin.")inleton_instance,instances=letinstances=ref[]in(funx->ifall_instancestheninstances:=!instances@[x]elseraise(FoundUnifx)),(fun()->!instances)inletrecuniquize=function|[]->[]|(sigma,_,{up_f=f;up_a=a;up_t=t}asx)::xs->lett=nf_evarsigmatinletf=nf_evarsigmafinleta=Array.map(nf_evarsigma)ainletneq(sigma1,_,{up_f=f1;up_a=a1;up_t=t1})=lett1=nf_evarsigma1t1inletf1=nf_evarsigma1f1inleta1=Array.map(nf_evarsigma1)a1innot(equaltt1&&equalff1&&CArray.for_all2equalaa1)inx::uniquize(List.filterneqxs)in((funenvch~k->do_onceupat_that_matched(fun()->letfailed_because_of_TC=reffalseintryifnotall_instancesthenmatch_upats_FOupatsenvsigma0isec;failed_because_of_TC:=match_upats_HO~on_instanceupatsenvsigma0isec;raiseNoMatchwithFoundUnifsigma_u->env,0,[sigma_u]|(NoMatch|NoProgress)whenall_instances&&instances()<>[]->env,0,uniquize(instances())|NoMatchwhen(notraise_NoMatch)->if!failed_because_of_TCthenerrorstrm(sourceenv++strbrk"matches but type classes inference fails")elseerrorstrm(sourceenv++str"does not match any subterm of the goal")|NoProgresswhen(notraise_NoMatch)->letdir=matchupats_originwithSome(d,_)->d|_->CErrors.anomaly(str"mk_tpattern_matcher with no upats_origin.")inerrorstrm(str"all matches of "++sourceenv++str"are equal to the "++pr_dir_side(inv_dirdir))|NoProgress->raiseNoMatch);letsigma,_,({up_f=pf;up_a=pa}asu)=ifall_instancesthenassert_done_multiresupat_that_matchedelseList.hd(pi3(assert_doneupat_that_matched))in(* pp(lazy(str"sigma@tmatch=" ++ pr_evar_map None sigma)); *)if!skip_occthen((*ignore(k env u.up_t 0);*)c)elseletmatch_EQ=match_EQenvsigmauinletpn=Array.lengthpainletrecsubst_loop(env,hasacc)c'=if!skip_occthenc'elseletf,a=splay_appsigmac'inifArray.lengtha>=pn&&match_EQf&&unif_EQ_argsenvsigmapaathenleta1,a2=Array.chop(Array.lengthpa)ainletfa1=mkApp(f,a1)inletf'=ifsubst_occ()thenkenvu.up_tfa1helsefa1inmkApp(f',Array.map_left(subst_loopacc)a2)else(* TASSI: clear letin values to avoid unfolding *)letinc_hrd(env,h')=letctx_item=matchrdwith|Context.Rel.Declaration.LocalAssum_asx->x|Context.Rel.Declaration.LocalDef(x,_,y)->Context.Rel.Declaration.LocalAssum(x,y)inEConstr.push_relctx_itemenv,h'+1inletselfaccc=EConstr.of_constr(subst_loopacc(EConstr.Unsafe.to_constrc))inletf=EConstr.of_constrfinletf'=map_constr_with_binders_left_to_rightenvsigmainc_hselfaccfinletf'=EConstr.Unsafe.to_constrf'inmkApp(f',Array.map_left(subst_loopacc)a)insubst_loop(env,h)c):find_P),((fun()->letenv,(sigma,uc,({up_f=pf;up_a=pa}asu))=match!upat_that_matchedwith|Some(env,_,x)->env,List.hdx|Nonewhenraise_NoMatch->raiseNoMatch|None->CErrors.anomaly(str"companion function never called.")inletp'=mkApp(pf,pa)inifmax_occ<=!noccthenp',u.up_dir,(sigma,uc,u.up_t)elseerrorstrm(str"Only "++int!nocc++str" < "++intmax_occ++str(String.plural!nocc" occurrence")++matchupats_originwith|None->str" of"++spc()++pr_constr_patenvsigmap'|Some(dir,rule)->str" of the "++pr_dir_sidedir++fnl()++ws4++pr_constr_patenvsigmap'++fnl()++str"of "++pr_constr_patenvsigmarule)):conclude)type('ident,'term)ssrpattern=|Tof'term|In_Tof'term|X_In_Tof'ident*'term|In_X_In_Tof'ident*'term|E_In_X_In_Tof'term*'ident*'term|E_As_X_In_Tof'term*'ident*'termletpr_patternpr_identpr_term=function|Tt->pr_termt|In_Tt->str"in "++pr_termt|X_In_T(x,t)->pr_identx++str" in "++pr_termt|In_X_In_T(x,t)->str"in "++pr_identx++str" in "++pr_termt|E_In_X_In_T(e,x,t)->pr_terme++str" in "++pr_identx++str" in "++pr_termt|E_As_X_In_T(e,x,t)->pr_terme++str" as "++pr_identx++str" in "++pr_termtletpr_pattern_auxpr_constr=function|Tt->pr_constrt|In_Tt->str"in "++pr_constrt|X_In_T(x,t)->pr_constrx++str" in "++pr_constrt|In_X_In_T(x,t)->str"in "++pr_constrx++str" in "++pr_constrt|E_In_X_In_T(e,x,t)->pr_constre++str" in "++pr_constrx++str" in "++pr_constrt|E_As_X_In_T(e,x,t)->pr_constre++str" as "++pr_constrx++str" in "++pr_constrtletpp_patternenv(sigma,p)=pr_pattern_aux(funt->pr_econstr_patenvsigma(pi3(nf_open_termsigmasigma(EConstr.of_constrt))))ptypecpattern={kind:ssrtermkind;pattern:Genintern.glob_constr_and_expr;interpretation:Geninterp.interp_signoption}letpr_term{kind;pattern;_}=letenv=Global.env()inletsigma=Evd.from_envenvinpr_guarded(guard_termkind)(pr_glob_constr_and_exprenvsigma)patternletprl_term{kind;pattern;_}=letenv=Global.env()inletsigma=Evd.from_envenvinpr_guarded(guard_termkind)(prl_glob_constr_and_exprenvsigma)patternletpr_cpattern=pr_termletpr_pattern_w_ids=pr_patternpr_idprl_termletmk_termkcist={kind=k;pattern=(mkRHole,Somec);interpretation=ist}letmk_lterm=mk_termNoFlagletglob_ssrtermgs=function|{kind;pattern=(_,Somec);interpretation=None}->letx=Tacintern.intern_constrgscin{kind;pattern=(fstx,Somec);interpretation=None}|ct->ct(* ssrterm conbinators *)letcombineCGt1t2fg=letmk_isti1i2=matchi1,i2with|None,Somei->Somei|Somei,None->Somei|None,None->None|Somei,Somejwheni==j->Somei|_->CErrors.anomaly(Pp.str"combineCG: different ist")inmatcht1,t2with|{kind=x;pattern=(t1,None);interpretation=i1},{pattern=(t2,None);interpretation=i2}->{kind=x;pattern=(gt1t2,None);interpretation=mk_isti1i2}|{kind=x;pattern=(_,Somet1);interpretation=i1},{pattern=(_,Somet2);interpretation=i2}->{kind=x;pattern=(mkRHole,Some(ft1t2));interpretation=mk_isti1i2}|_,{pattern=(_,None);_}->CErrors.anomaly(str"have: mixed C-G constr.")|_->CErrors.anomaly(str"have: mixed G-C constr.")letloc_ofCG=function|{pattern=(s,None);_}->Glob_ops.loc_of_glob_constrs|{pattern=(_,Somes);_}->Constrexpr_ops.constr_locs(* This piece of code asserts the following notations are reserved *)(* Reserved Notation "( a 'in' b )" (at level 0). *)(* Reserved Notation "( a 'as' b )" (at level 0). *)(* Reserved Notation "( a 'in' b 'in' c )" (at level 0). *)(* Reserved Notation "( a 'as' b 'in' c )" (at level 0). *)letglob_cpatterngsp=pp(lazy(str"globbing pattern: "++pr_termp));letglobx=(glob_ssrtermgs(mk_ltermxNone)).patterninletencodeksl=letname=Name(Id.of_string("_ssrpat_"^s))in{kind=k;pattern=(mkRCastmkRHole(mkRLambdanamemkRHole(mkRAppmkRHolel)),None);interpretation=None}inletbind_int1t2=letmkCHole=mkCHole~loc:Noneinletn=Name(destCVart1)infst(glob(mkCCastmkCHole(mkCLambdanmkCHolet2)))inletcheck_vart2=ifnot(isCVart2)thenloc_error(constr_loct2)"Only identifiers are allowed here"inmatchpwith|{pattern=(_,None);_}asx->x|{kind=k;pattern=(v,Somet);_}asorig->ifk=Cpatternthenglob_ssrtermgs{kind=InParens;pattern=(v,Somet);interpretation=None}elsematcht.CAst.vwith|CNotation(_,(InConstrEntry,"( _ in _ )"),([t1;t2],[],[],[]))->(trymatchglobt1,globt2with|(r1,None),(r2,None)->encodek"In"[r1;r2]|(r1,Some_),(r2,Some_)whenisCVart1->encodek"In"[r1;r2;bind_int1t2]|(r1,Some_),(r2,Some_)->encodek"In"[r1;r2]|_->CErrors.anomaly(str"where are we?.")with_whenisCVart1->encodek"In"[bind_int1t2])|CNotation(_,(InConstrEntry,"( _ in _ in _ )"),([t1;t2;t3],[],[],[]))->check_vart2;encodek"In"[fst(globt1);bind_int2t3]|CNotation(_,(InConstrEntry,"( _ as _ )"),([t1;t2],[],[],[]))->encodek"As"[fst(globt1);fst(globt2)]|CNotation(_,(InConstrEntry,"( _ as _ in _ )"),([t1;t2;t3],[],[],[]))->check_vart2;encodek"As"[fst(globt1);bind_int2t3]|_->glob_ssrtermgsorig;;letglob_rpatternsp=matchpwith|Tt->T(glob_cpatternst)|In_Tt->In_T(glob_ssrtermst)|X_In_T(x,t)->X_In_T(x,glob_ssrtermst)|In_X_In_T(x,t)->In_X_In_T(x,glob_ssrtermst)|E_In_X_In_T(e,x,t)->E_In_X_In_T(glob_ssrtermse,x,glob_ssrtermst)|E_As_X_In_T(e,x,t)->E_As_X_In_T(glob_ssrtermse,x,glob_ssrtermst)letsubst_ssrterms{kind;pattern;interpretation}={kind;pattern=Tacsubst.subst_glob_constr_and_exprspattern;interpretation}letsubst_rpatterns=function|Tt->T(subst_ssrtermst)|In_Tt->In_T(subst_ssrtermst)|X_In_T(x,t)->X_In_T(x,subst_ssrtermst)|In_X_In_T(x,t)->In_X_In_T(x,subst_ssrtermst)|E_In_X_In_T(e,x,t)->E_In_X_In_T(subst_ssrtermse,x,subst_ssrtermst)|E_As_X_In_T(e,x,t)->E_As_X_In_T(subst_ssrtermse,x,subst_ssrtermst)letinterp_ssrtermist{kind;pattern;_}={kind;pattern;interpretation=Someist}letinterp_rpatterns=function|Tt->T(interp_ssrtermst)|In_Tt->In_T(interp_ssrtermst)|X_In_T(x,t)->X_In_T(interp_ssrtermsx,interp_ssrtermst)|In_X_In_T(x,t)->In_X_In_T(interp_ssrtermsx,interp_ssrtermst)|E_In_X_In_T(e,x,t)->E_In_X_In_T(interp_ssrtermse,interp_ssrtermsx,interp_ssrtermst)|E_As_X_In_T(e,x,t)->E_As_X_In_T(interp_ssrtermse,interp_ssrtermsx,interp_ssrtermst)letinterp_rpattern0ist__t=interp_rpatternisttlettag_of_cpatternp=p.kindletloc_of_cpattern=loc_ofCGtypeocc=(bool*intlist)optiontyperpattern=(cpattern,cpattern)ssrpatternletpr_rpattern=pr_patternpr_cpatternpr_cpatternletwit_rpatternty=add_genarg"rpatternty"(funenvsigma->pr_patternpr_cpatternpr_cpattern)typepattern=Evd.evar_map*(constr,constr)ssrpatternletid_of_cpattern{pattern=(c1,c2);_}=letopenCAstinmatchDAst.getc1,c2with|_,Some{v=CRef(qid,_)}whenqualid_is_identqid->Some(qualid_basenameqid)|_,Some{v=CAppExpl((qid,_),[])}whenqualid_is_identqid->Some(qualid_basenameqid)|GRef(GlobRef.VarRefx,_),None->Somex|_->Noneletid_of_Ctermt=matchid_of_cpatterntwith|Somex->x|None->loc_error(loc_of_cpatternt)"Only identifiers are allowed here"letinterp_open_constristenvsigmagc=Tacinterp.interp_open_constristenvsigmagcletpf_intern_termenvsigma{pattern=c;interpretation=ist;_}=glob_constristenvsigmacletinterp_ssrtermistenvsigmat=interp_ssrtermisttletinterp_termenvsigma=function|{pattern=c;interpretation=Someist;_}->on_sndEConstr.Unsafe.to_constr(interp_open_constristenvsigmac)|_->errorstrm(str"interpreting a term with no ist")letthinidsigmagoal=letids=Id.Set.singletonidinletevi=Evd.findsigmagoalinletenv=Evd.evar_filtered_env(Global.env())eviinletcl=Evd.evar_concleviinletsigma=Evd.clear_metassigmainletans=trySome(Evarutil.clear_hyps_in_evienvsigma(Environ.named_context_valenv)clids)withEvarutil.ClearDependencyError_->Noneinmatchanswith|None->sigma|Some(sigma,hyps,concl)->let(sigma,evk)=Evarutil.new_pure_evar~src:(Loc.tagEvar_kinds.GoalEvar)~typeclass_candidate:falsehypssigmaconclinletsigma=Evd.remove_future_goalsigmaevkinletid=Evd.evar_identgoalsigmainletproof=EConstr.mkEvar(evk,Evd.evar_identity_subst@@Evd.findsigmaevk)inletsigma=Evd.definegoalproofsigmainmatchidwith|None->sigma|Someid->Evd.renameevkidsigma(*
let pr_ist { lfun= lfun } =
prlist_with_sep spc
(fun (id, Geninterp.Val.Dyn(ty,_)) ->
pr_id id ++ str":" ++ Geninterp.Val.pr ty) (Id.Map.bindings lfun)
*)letinterp_pattern?wit_ssrpatternargenvsigma0redredty=pp(lazy(str"interpreting: "++pr_rpatternred));letxInTxy=X_In_T(x,y)andinXInTxy=In_X_In_T(x,y)inletinTx=In_TxandeInXInText=E_In_X_In_T(e,x,t)inleteAsXInText=E_As_X_In_T(e,x,t)inletmkG?(k=NoFlag)xist={kind=k;pattern=(x,None);interpretation=ist}inletist_ofx=x.interpretationinletdecode({interpretation=ist;_}ast)?reccallfg=trymatchDAst.get(pf_intern_termenvsigma0t)with|GCast(t,DEFAULTcast,c)whenisGHolet&&isGLambdac->let(x,c)=destGLambdacinfx{kind=NoFlag;pattern=(c,None);interpretation=ist}|GVaridwhenOption.has_someist&&letist=Option.getistinId.Map.memidist.lfun&¬(Option.is_emptyreccall)&¬(Option.is_emptywit_ssrpatternarg)->letv=Id.Map.findid(Option.getist).lfuninOption.getreccall(Value.cast(topwit(Option.getwit_ssrpatternarg))v)|it->gtwithewhenCErrors.noncriticale->gtinletdecodeGisttfg=decode(mkGtist)fginletbad_encid_=CErrors.anomaly(str"bad encoding for pattern "++strid++str".")inletcleanup_XinEhxrpsigma=leth_k=matchkindhwithEvar(k,_)->k|_->assertfalseinletto_clean,update=(* handle rename if x is already used *)letctx=Environ.named_contextenvinletlen=Context.Named.lengthctxinletname=refNoneintryignore(Context.Named.lookupxctx);(name,funk->if!name=Nonethenletnctx=Evd.evar_context(Evd.findsigmak)inletnlen=Context.Named.lengthnctxinifnlen>lenthenbeginname:=Some(Context.Named.Declaration.get_id(List.nthnctx(nlen-len-1)))end)withNot_found->ref(Somex),fun_->()inletnew_evars=letrecauxacct=matchkindtwith|Evar(k,_)->ifk=h_k||List.memkacc||Evd.memsigma0kthenaccelse(updatek;k::acc)|_->CoqConstr.foldauxacctinaux[](nf_evarsigmarp)inletsigma=List.fold_left(funsigmae->ifEvd.is_definedsigmaethensigmaelse(* clear may be recursive *)ifOption.is_empty!to_cleanthensigmaelseletname=Option.get!to_cleaninpp(lazy(pr_idname));thinnamesigmae)sigmanew_evarsinsigmainletred=letrecdecode_red=function|T{kind=k;pattern=(t,None);interpretation=ist}->beginmatchDAst.gettwith|GCast(c,DEFAULTcast,t)whenisGHolec&&let(id,t)=destGLambdatinletid=Id.to_stringidinletlen=String.lengthidin(len>8&&String.subid08="_ssrpat_")->let(id,t)=destGLambdatinletid=Id.to_stringidinletlen=String.lengthidin(matchString.subid8(len-8),DAst.gettwith|"In",GApp(_,[t])->decodeGisttxInT(funx->Tx)|"In",GApp(_,[e;t])->decodeGistt(eInXInT(mkGeist))(bad_encid)|"In",GApp(_,[e;t;e_in_t])->decodeGistt(eInXInT(mkGeist))(fun_->decodeGiste_in_txInT(fun_->assertfalse))|"As",GApp(_,[e;t])->decodeGistt(eAsXInT(mkGeist))(bad_encid)|_->bad_encid())|_->decode~reccall:decode_red(mkG~ktist)xInT(funx->Tx)end|Tt->decode~reccall:decode_redtxInT(funx->Tx)|In_Tt->decodetinXInTinT|X_In_T(e,t)->decodet(eInXInTe)(funx->xInT(id_of_Cterme)x)|In_X_In_T(e,t)->inXInT(id_of_Cterme)t|E_In_X_In_T(e,x,rp)->eInXInTe(id_of_Ctermx)rp|E_As_X_In_T(e,x,rp)->eAsXInTe(id_of_Ctermx)rpindecode_redredinpp(lazy(str"decoded as: "++pr_pattern_w_idsred));letred=matchredtywith|None->red|Some(ty,ist)->letty={kind=NoFlag;pattern=ty;interpretation=Someist}inmatchredwith|Tt->T(combineCGtty(mkCCast?loc:(loc_ofCGt))mkRCast)|X_In_T(x,t)->letgty=pf_intern_termenvsigma0tyinE_As_X_In_T(mkG(mkRCastmkRHolegty)(ist_ofty),x,t)|E_In_X_In_T(e,x,t)->letty=mkG(pf_intern_termenvsigma0ty)(ist_ofty)inE_In_X_In_T(combineCGety(mkCCast?loc:(loc_ofCGt))mkRCast,x,t)|E_As_X_In_T(e,x,t)->letty=mkG(pf_intern_termenvsigma0ty)(ist_ofty)inE_As_X_In_T(combineCGety(mkCCast?loc:(loc_ofCGt))mkRCast,x,t)|red->redinpp(lazy(str"typed as: "++pr_pattern_w_idsred));letmkXLetIn?locx{kind;pattern=(g,c);interpretation}=matchcwith|Someb->{kind;pattern=(g,Some(mkCLetIn?locx(mkCHole~loc)b));interpretation}|None->{kind;pattern=DAst.make?loc@@GLetIn(x,DAst.make?loc@@GHole(BinderTypex,IntroAnonymous,None),None,g),None;interpretation}inmatchredwith|Tt->letsigma,t=interp_termenvsigma0tinsigma,Tt|In_Tt->letsigma,t=interp_termenvsigma0tinsigma,In_Tt|X_In_T(x,rp)|In_X_In_T(x,rp)->letmkxp=matchredwithX_In_T_->X_In_T(x,p)|_->In_X_In_T(x,p)inletrp=mkXLetIn(Namex)rpinletsigma,rp=interp_termenvsigma0rpinlet_,h,_,rp=destLetInrpinletsigma=cleanup_XinEhxrpsigmainletrp=subst1h(nf_evarsigmarp)insigma,mkhrp|E_In_X_In_T(e,x,rp)|E_As_X_In_T(e,x,rp)->letmkexp=matchredwithE_In_X_In_T_->E_In_X_In_T(e,x,p)|_->E_As_X_In_T(e,x,p)inletrp=mkXLetIn(Namex)rpinletsigma,rp=interp_termenvsigma0rpinlet_,h,_,rp=destLetInrpinletsigma=cleanup_XinEhxrpsigmainletrp=subst1h(nf_evarsigmarp)inletsigma,e=interp_termenvsigmaeinsigma,mkehrp;;letinterp_cpatternenvsigmaredredty=interp_patternenvsigma(Tred)redty;;letinterp_rpattern~wit_ssrpatternargenvsigmared=interp_pattern~wit_ssrpatternargenvsigmaredNone;;letid_of_pattern=function|_,Tt->(matchkindtwithVarid->Someid|_->None)|_->None(* The full occurrence set *)letnoindex=Some(false,[])(* calls do_subst on every sub-term identified by (pattern,occ) *)leteval_pattern?raise_NoMatchenv0sigma0concl0patternoccdo_subst=letfssigmax=nf_evarsigmaxinletpop_evarsigmaep=let{Evd.evar_body=e_body}ase_def=Evd.findsigmaeinlete_body=matche_bodywithEvar_definedc->EConstr.Unsafe.to_constrc|_->errorstrm(str"Matching the pattern "++pr_constr_envenv0sigma0p++str" did not instantiate ?"++int(Evar.repre)++spc()++str"Does the variable bound by the \"in\" construct occur "++str"in the pattern?")inletsigma=Evd.add(Evd.removesigmae)e{e_defwithEvd.evar_body=Evar_empty}insigma,e_bodyinletex_valuehole=matchkindholewithEvar(e,_)->e|_->assertfalseinletmk_upat_for?hackenvsigma0(sigma,t)?(p=t)ok=letsigma,pat=mk_tpattern?hackenvsigma0(sigma,p)okL2R(fssigmat)insigma,[pat]inmatchpatternwith|None->do_substenv0concl0concl01,UState.empty|Some(sigma,(Trp|In_Trp))->letrp=fssigmarpinletise=create_evar_defssigmainletocc=matchpatternwithSome(_,T_)->occ|_->noindexinletrp=mk_upat_forenv0sigma0(ise,rp)all_okinletfind_T,end_T=mk_tpattern_matcher?raise_NoMatchsigma0occrpinletconcl=find_Tenv0concl01~k:do_substinlet_,_,(_,us,_)=end_T()inconcl,us|Some(sigma,(X_In_T(hole,p)|In_X_In_T(hole,p)))->letp=fssigmapinletocc=matchpatternwithSome(_,X_In_T_)->occ|_->noindexinletex=ex_valueholeinletrp=mk_upat_for~hack:trueenv0sigma0(sigma,p)all_okinletfind_T,end_T=mk_tpattern_matchersigma0noindexrpin(* we start from sigma, so hole is considered a rigid head *)letholep=mk_upat_forenv0sigma(sigma,hole)all_okinletfind_X,end_X=mk_tpattern_matcher?raise_NoMatchsigmaoccholepinletconcl=find_Tenv0concl01~k:(funenvc_h->letp_sigma=unify_HOenv(create_evar_defssigma)(EConstr.of_constrc)(EConstr.of_constrp)inletsigma,e_body=pop_evarp_sigmaexpinfsp_sigma(find_Xenv(fssigmap)h~k:(funenv_->do_substenve_body)))inlet_=end_X()inlet_,_,(_,us,_)=end_T()inconcl,us|Some(sigma,E_In_X_In_T(e,hole,p))->letp,e=fssigmap,fssigmaeinletex=ex_valueholeinletrp=mk_upat_for~hack:trueenv0sigma0(sigma,p)all_okinletfind_T,end_T=mk_tpattern_matchersigma0noindexrpinletholep=mk_upat_forenv0sigma(sigma,hole)all_okinletfind_X,end_X=mk_tpattern_matchersigmanoindexholepinletre=mk_upat_forenv0sigma0(sigma,e)all_okinletfind_E,end_E=mk_tpattern_matcher?raise_NoMatchsigma0occreinletconcl=find_Tenv0concl01~k:(funenvc_h->letp_sigma=unify_HOenv(create_evar_defssigma)(EConstr.of_constrc)(EConstr.of_constrp)inletsigma,e_body=pop_evarp_sigmaexpinfsp_sigma(find_Xenv(fssigmap)h~k:(funenvc_h->find_Eenve_bodyh~k:do_subst)))inlet_,_,(_,us,_)=end_E()inlet_=end_X()inlet_=end_T()inconcl,us|Some(sigma,E_As_X_In_T(e,hole,p))->letp,e=fssigmap,fssigmaeinletex=ex_valueholeinletrp=lete_sigma=unify_HOenv0sigma(EConstr.of_constrhole)(EConstr.of_constre)ine_sigma,fse_sigmapinletrp=mk_upat_for~hack:trueenv0sigma0rpall_okinletfind_TE,end_TE=mk_tpattern_matchersigma0noindexrpinletholep=mk_upat_forenv0sigma(sigma,hole)all_okinletfind_X,end_X=mk_tpattern_matchersigmaoccholepinletconcl=find_TEenv0concl01~k:(funenvc_h->letp_sigma=unify_HOenv(create_evar_defssigma)(EConstr.of_constrc)(EConstr.of_constrp)inletsigma,e_body=pop_evarp_sigmaexpinfsp_sigma(find_Xenv(fssigmap)h~k:(funenvc_h->lete_sigma=unify_HOenvsigma(EConstr.of_constre_body)(EConstr.of_constre)inlete_body=fse_sigmaeindo_substenve_bodye_bodyh)))inlet_=end_X()inlet_,_,(_,us,_)=end_TE()inconcl,us;;letredex_of_pattern?(resolve_typeclasses=false)env(sigma,p)=lete=matchpwith|In_T_|In_X_In_T_->CErrors.anomaly(str"pattern without redex.")|Te|X_In_T(e,_)|E_As_X_In_T(e,_,_)|E_In_X_In_T(e,_,_)->einletsigma=ifnotresolve_typeclassesthensigmaelseTypeclasses.resolve_typeclasses~fail:falseenvsigmainnf_evarsigmae,Evd.evar_universe_contextsigmaletfill_occ_pattern?raise_NoMatchenvsigmaclpatocch=letdo_make_rel,occ=ifocc=Some(true,[])thenfalse,Some(false,[1])elsetrue,occinletfind_R,conclude=letr=refNonein(funenvc_h'->do_oncer(fun()->c);ifdo_make_relthenmkRel(h'+h-1)elsec),(fun_->if!r=Nonethenfst(redex_of_patternenvpat)elseassert_doner)inletcl,us=eval_pattern?raise_NoMatchenvsigmacl(Somepat)occfind_Rinlete=concludeclin(e,us),cl;;(* clenup interface for external use *)letmk_tpattern?p_originenvsigma0sigma_tfdirc=mk_tpattern?p_originenvsigma0sigma_tfdirc;;leteval_pattern?raise_NoMatchenv0sigma0concl0patternoccdo_subst=fst(eval_pattern?raise_NoMatchenv0sigma0concl0patternoccdo_subst);;letpf_fill_occenvconcloccsigma0p(sigma,t)okh=letp=EConstr.Unsafe.to_constrpinletconcl=EConstr.Unsafe.to_constrconclinletise=create_evar_defssigmainletise,u=mk_tpatternenvsigma0(ise,EConstr.Unsafe.to_constrt)okL2Rpinletfind_U,end_U=mk_tpattern_matcher~raise_NoMatch:truesigma0occ(ise,[u])inletconcl=find_Uenvconclh~k:(fun___->mkRel)inletrdx,_,(sigma,uc,p)=end_U()insigma,uc,EConstr.of_constrp,EConstr.of_constrconcl,EConstr.of_constrrdxletfill_occ_termenvsigma0clocc(sigma,t)=tryletsigma',uc,t',cl,_=pf_fill_occenvcloccsigma0t(sigma,t)all_ok1inifsigma'!=sigma0thenCErrors.user_errPp.(str"matching impacts evars")elsecl,t'withNoMatch->tryletsigma',uc,t'=unif_endenvsigma0(create_evar_defssigma)t(fun_->true)inifsigma'!=sigma0thenraiseNoMatchelsecl,t'with_->errorstrm(str"partial term "++pr_econstr_patenvsigmat++str" does not match any subterm of the goal")letpf_fill_occ_termglocct=letsigma0=projectglandenv=pf_envglandconcl=pf_conclglinletcl,t=fill_occ_termenvsigma0conclocctincl,tletcpattern_of_idid={kind=NoFlag;pattern=DAst.make@@GRef(GlobRef.VarRefid,None),None;interpretation=SomeGeninterp.({lfun=Id.Map.empty;poly=false;extra=Tacinterp.TacStore.empty})}letis_wildcard({pattern=(l,r);_}:cpattern):bool=matchDAst.getl,rwith|_,Some{CAst.v=CHole_}|GHole_,None->true|_->false(* "ssrpattern" *)letpf_merge_ucucgl=re_sig(sig_itgl)(Evd.merge_universe_context(projectgl)uc)letpf_unsafe_merge_ucucgl=re_sig(sig_itgl)(Evd.set_universe_context(projectgl)uc)(** All the pattern types reuse the same dynamic toplevel tag *)letwit_ssrpatternarg=wit_rpatterntyletinterp_rpattern=interp_rpattern~wit_ssrpatternargletssrpatterntac_istarg=letopenProofview.NotationsinProofview.Goal.enterbeginfungl->letsigma0=Proofview.Goal.sigmaglinletconcl0=Proofview.Goal.conclglinletenv=Proofview.Goal.envglinletpat=interp_rpatternenvsigma0arginletconcl0=EConstr.Unsafe.to_constrconcl0inlet(t,uc),concl_x=fill_occ_patternenvsigma0concl0patnoindex1inlett=EConstr.of_constrtinletconcl_x=EConstr.of_constrconcl_xinletsigma,tty=Typing.type_ofenvsigma0tinletconcl=EConstr.mkLetIn(make_annot(Name(Id.of_string"selected"))Sorts.Relevant,t,tty,concl_x)inProofview.Unsafe.tclEVARSsigma<*>convert_concl~cast:false~check:trueconclDEFAULTcastend(* Register "ssrpattern" tactic *)let()=letmltac_ist=letarg=letv=Id.Map.find(Names.Id.of_string"pattern")ist.lfuninValue.cast(topwitwit_ssrpatternarg)vinssrpatterntacistarginletname={mltac_plugin="ssrmatching_plugin";mltac_tactic="ssrpattern";}inlet()=Tacenv.register_ml_tacticname[|mltac|]inlettac=CAst.make(TacFun([Name(Id.of_string"pattern")],CAst.make(TacML({mltac_name=name;mltac_index=0},[]))))inletobj()=Tacenv.register_ltactruefalse(Id.of_string"ssrpattern")tacinMltop.declare_cache_objobj"ssrmatching_plugin"letssrinstancesofarg=Proofview.Goal.enterbeginfungl->letokrhslhsise=truein(* not (equal lhs (Evarutil.nf_evar ise rhs)) in *)letenv=Proofview.Goal.envglinletsigma=Proofview.Goal.sigmaglinletconcl=Proofview.Goal.conclglinletconcl=EConstr.to_constr~abort_on_undefined_evars:falsesigmaconclinletsigma0,cpat=interp_cpatternenvsigmaargNoneinletpat=matchcpatwithTx->x|_->errorstrm(str"Not supported")inletetpat,tpat=mk_tpatternenvsigma(sigma0,pat)(okpat)L2Rpatinletfind,conclude=mk_tpattern_matcher~all_instances:true~raise_NoMatch:truesigmaNone(etpat,[tpat])inletprintenvpc_=ppnl(hov1(str"instance:"++spc()++pr_constr_envenv(Proofview.Goal.sigmagl)p++spc()++str"matches:"++spc()++pr_constr_envenv(Proofview.Goal.sigmagl)c));cinppnl(str"BEGIN INSTANCES");trywhiletruedoignore(findenvconcl1~k:print)done;raiseNoMatchwithNoMatch->ppnl(str"END INSTANCES");Tacticals.tclIDTACendmoduleInternal=structletwit_rpatternty=wit_rpatterntyletglob_rpattern=glob_rpatternletsubst_rpattern=subst_rpatternletinterp_rpattern=interp_rpattern0letpr_rpattern=pr_rpatternletmk_rpatternx=xletmk_lterm=mk_ltermletmk_term=mk_termletglob_cpattern=glob_cpatternletsubst_ssrterm=subst_ssrtermletinterp_ssrterm=interp_ssrtermletpr_ssrterm=pr_termend(* vim: set filetype=ocaml foldmethod=marker: *)