123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118(************************************************************************)(* * 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) *)(************************************************************************)openUtilopenNamesopenConstropenContextopenCErrorsopenEvar_refineropenTacexpropenLocusopenContext.Named.DeclarationopenLtac_pretypemoduleNamedDecl=Context.Named.Declaration(* The instantiate tactic *)letinstantiate_evarevk(ist,rawc)=letopenProofview.NotationsinProofview.tclENV>>=funenv->Proofview.Goal.enterbeginfungl->letsigma=Proofview.Goal.sigmaglinletevi=Evd.findsigmaevkinletfiltered=Evd.evar_filtered_envenveviinletconstrvars=Tacinterp.extract_ltac_constr_valuesistfilteredinletlvar={ltac_constrs=constrvars;ltac_uconstrs=Names.Id.Map.empty;ltac_idents=Names.Id.Map.empty;ltac_genargs=ist.Geninterp.lfun;}inletsigma'=w_refine(evk,evi)(lvar,rawc)envsigmainProofview.Unsafe.tclEVARSsigma'endletevar_listsigmac=letrecevrecaccc=matchEConstr.kindsigmacwith|Evar(evk,_asev)->ev::acc|_->EConstr.foldsigmaevrecacccinevrec[]cletinstantiate_tacncido=Proofview.Goal.enterbeginfungl->letsigma=Proofview.Goal.sigmaglinletenv=Proofview.Goal.envglinletconcl=Proofview.Goal.conclglinletevl=matchidowithConclLocation()->evar_listsigmaconcl|HypLocation(id,hloc)->letdecl=Environ.lookup_namedidenvinmatchhlocwithInHyp->(matchdeclwith|LocalAssum(_,typ)->evar_listsigma(EConstr.of_constrtyp)|_->user_errPp.(str"Please be more specific: in type or value?"))|InHypTypeOnly->evar_listsigma(EConstr.of_constr(NamedDecl.get_typedecl))|InHypValueOnly->(matchdeclwith|LocalDef(_,body,_)->evar_listsigma(EConstr.of_constrbody)|_->user_errPp.(str"Not a defined hypothesis."))inifList.lengthevl<nthenuser_errPp.(str"Not enough uninstantiated existential variables.");ifn<=0thenuser_errPp.(str"Incorrect existential variable index.");letevk,_=List.nthevl(n-1)ininstantiate_evarevkcendletinstantiate_tac_by_nameidc=Proofview.Goal.enterbeginfungl->letsigma=Proofview.Goal.sigmaglinletevk=tryEvd.evar_keyidsigmawithNot_found->user_errPp.(str"Unknown existential variable.")ininstantiate_evarevkcendletlet_evarnametyp=letsrc=(Loc.tagEvar_kinds.GoalEvar)inProofview.Goal.enterbeginfungl->letsigma=Tacmach.New.projectglinletenv=Proofview.Goal.envglinletsigma,_=Typing.sort_ofenvsigmatypinletid=matchnamewith|Name.Anonymous->letid=Namegen.id_of_name_using_hdcharenvsigmatypnameinNamegen.next_ident_away_in_goalid(Termops.vars_of_envenv)|Name.Nameid->idinlet(sigma,evar)=Evarutil.new_evarenvsigma~src~naming:(Namegen.IntroFreshid)typinTacticals.New.tclTHEN(Proofview.Unsafe.tclEVARSsigma)(Tactics.pose_tac(Name.Nameid)evar)endlethget_evarn=letopenEConstrinProofview.Goal.enterbeginfungl->letsigma=Tacmach.New.projectglinletconcl=Proofview.Goal.conclglinletevl=evar_listsigmaconclinifList.lengthevl<nthenuser_errPp.(str"Not enough uninstantiated existential variables.");ifn<=0thenuser_errPp.(str"Incorrect existential variable index.");letev=List.nthevl(n-1)inletev_type=EConstr.existential_typesigmaevinletr=Retyping.relevance_of_type(Proofview.Goal.envgl)sigmaev_typeinTactics.change_concl(mkLetIn(make_annotName.Anonymousr,mkEvarev,ev_type,concl))end