123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691692693694695696697698699700701702703704705706707708709710711712713714715716717718719720721722723724725726727728729730731732733734735736737738739740741742743744745746747748749750751752753754755756757758759760761762763764765766767768769770771772773774775776777778779780781782783784785786787788789790791792793794795796797798799800801802803804805806807808809810811812813814815816817818819820821822823824825826827828829830831832833834835836837838839840841842843844845846847848849850851852853854855856857858859860861862863864865866867868869870871872873874875876877878879880881882883884885886887888889890891892893894895896897898899900901902903904905906907908909910911912913914915916917918919920921922923924925926927928929930931932933934935936937938939940941942943944945946947948949950951952953954955956957958959960961962963964965966967968969970971972973974975976977978979980981982983984985986987988989990991992993994995996997998999100010011002100310041005100610071008100910101011101210131014101510161017101810191020102110221023102410251026102710281029103010311032103310341035103610371038103910401041104210431044104510461047104810491050105110521053105410551056105710581059106010611062106310641065106610671068106910701071107210731074107510761077107810791080108110821083108410851086108710881089109010911092109310941095109610971098109911001101110211031104110511061107110811091110111111121113111411151116111711181119112011211122112311241125112611271128112911301131113211331134113511361137113811391140114111421143114411451146114711481149115011511152115311541155115611571158115911601161116211631164116511661167116811691170117111721173117411751176117711781179118011811182118311841185118611871188118911901191119211931194119511961197119811991200120112021203120412051206120712081209121012111212121312141215121612171218121912201221122212231224122512261227122812291230123112321233123412351236123712381239124012411242124312441245124612471248124912501251125212531254125512561257125812591260126112621263126412651266126712681269127012711272127312741275127612771278127912801281128212831284128512861287128812891290129112921293129412951296129712981299130013011302130313041305130613071308130913101311131213131314131513161317131813191320132113221323132413251326132713281329133013311332133313341335133613371338133913401341134213431344134513461347134813491350135113521353135413551356135713581359136013611362136313641365136613671368136913701371137213731374137513761377137813791380138113821383138413851386138713881389139013911392139313941395139613971398139914001401140214031404140514061407140814091410141114121413141414151416141714181419142014211422142314241425142614271428142914301431143214331434143514361437143814391440144114421443144414451446144714481449145014511452145314541455145614571458145914601461146214631464146514661467146814691470147114721473147414751476147714781479148014811482148314841485148614871488148914901491149214931494149514961497149814991500150115021503150415051506150715081509151015111512151315141515151615171518151915201521152215231524152515261527152815291530153115321533153415351536153715381539154015411542154315441545154615471548154915501551155215531554155515561557155815591560156115621563156415651566156715681569157015711572157315741575157615771578157915801581158215831584158515861587158815891590159115921593159415951596159715981599(************************************************************************)(* * 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. *)openUtilopenNamesopenEvdopenTermopenConstropenContextopenTermopsopenPrinteropenLocusopsopenLtac_pluginopenTacmachopenTacticalsopenLibnamesopenSsrmatching_pluginopenSsrmatchingopenSsrastopenSsrprintersmoduleRelDecl=Context.Rel.DeclarationmoduleNamedDecl=Context.Named.Declaration(* Defining grammar rules with "xx" in it automatically declares keywords too,
* we thus save the lexer to restore it at the end of the file *)letfrozen_lexer=CLexer.get_keyword_state();;leterrorstrmx=CErrors.user_err~hdr:"ssreflect"xletallocc=Some(false,[])(** Bound assumption argument *)(* The Ltac API does have a type for assumptions but it is level-dependent *)(* and therefore impractical to use for complex arguments, so we substitute *)(* our own to have a uniform representation. Also, we refuse to intern *)(* idents that match global/section constants, since this would lead to *)(* fragile Ltac scripts. *)lethyp_id(SsrHyp(_,id))=idlethyp_err?locmsgid=CErrors.user_err?loc~hdr:"ssrhyp"Pp.(strmsg++Id.printid)letnot_section_idid=not(Termops.is_section_variableid)lethyps_ids=List.maphyp_idletreccheck_hyps_uniqids=function|SsrHyp(loc,id)::_whenList.memidids->hyp_err?loc"Duplicate assumption "id|SsrHyp(_,id)::hyps->check_hyps_uniq(id::ids)hyps|[]->()letcheck_hyp_existshyps(SsrHyp(_,id))=tryignore(Context.Named.lookupidhyps)withNot_found->errorstrmPp.(str"No assumption is named "++Id.printid)lettest_hyp_existshyps(SsrHyp(_,id))=tryignore(Context.Named.lookupidhyps);truewithNot_found->falselethoikf=functionHypx->fx|Idx->fxlethoi_id=hoikhyp_idletmk_hinttac=false,[Sometac]letmk_orhinttacs=true,tacsletnullhint=true,[]letnohint=false,[]type'atac_a=(goal*'a)sigma->(goal*'a)listsigmaletprojectgl=gl.Evd.sigmaletre_sigitsigma={Evd.it=it;Evd.sigma=sigma}letpush_ctxagl=re_sig(sig_itgl,a)(projectgl)letpush_ctxsagl=re_sig(List.map(funx->x,a)(sig_itgl))(projectgl)letpull_ctxgl=letg,a=sig_itglinre_sigg(projectgl),aletpull_ctxsgl=letg,a=List.split(sig_itgl)inre_sigg(projectgl),aletwith_ctxfgl=letgl,ctx=pull_ctxglinletrc,ctx=fctxinrc,push_ctxctxglletwithout_ctxfgl=letgl,_ctx=pull_ctxglinfgllettac_ctxtgl=letgl,a=pull_ctxglinletgl=tglinpush_ctxsagllettclTHEN_iat1t2gl=letgal=t1glinletgoals,sigma=sig_itgal,projectgalinlet_,opened,sigma=List.fold_left(fun(i,opened,sigma)g->letgl=t2i(re_siggsigma)ini+1,sig_itgl::opened,projectgl)(1,[],sigma)goalsinre_sig(List.flatten(List.revopened))sigmalettclTHEN_at1t2gl=tclTHEN_iat1(fun_->t2)gllettclTHENS_at1tlgl=tclTHEN_iat1(funi->List.nthtl(i-1))glletrectclTHENLIST_a=function|[]->tac_ctxtclIDTAC|t1::tacl->tclTHEN_at1(tclTHENLIST_atacl)(* like tclTHEN_i but passes to the tac "i of n" and not just i *)lettclTHEN_i_maxtactacigl=letmaxi=ref0intclTHEN_ia(tclTHEN_iatac(funi->maxi:=maxi!maxi;tac_ctxtclIDTAC))(funigl->tacii!maxigl)gllettac_on_allgltac=letgoals=sig_itglinletopened,sigma=List.fold_left(fun(opened,sigma)g->letgl=tac(re_siggsigma)insig_itgl::opened,projectgl)([],projectgl)goalsinre_sig(List.flatten(List.revopened))sigma(* Used to thread data between intro patterns at run time *)typetac_ctx={tmp_ids:(Id.t*Name.tref)list;wild_ids:Id.tlist;delayed_clears:Id.tlist;}letnew_ctx()={tmp_ids=[];wild_ids=[];delayed_clears=[]}letwith_fresh_ctxtgl=letgl=push_ctx(new_ctx())glinletgl=tglinfst(pull_ctxsgl)openGenargopenStdargopenPpleterrorstrmx=CErrors.user_err~hdr:"ssreflect"xletanomalys=CErrors.anomaly(strs)(* Tentative patch from util.ml *)letarray_fold_right_fromnfva=letrecfoldn=ifn>=Array.lengthvthenaelsefv.(n)(fold(succn))infoldnletarray_app_tlvl=ifArray.lengthv=0theninvalid_arg"array_app_tl";array_fold_right_from1(funel->e::l)vlletarray_list_of_tlv=ifArray.lengthv=0theninvalid_arg"array_list_of_tl";array_fold_right_from1(funel->e::l)v[](* end patch *)letoption_assert_getomsg=matchowith|None->CErrors.anomalymsg|Somex->x(** Constructors for rawconstr *)openGlob_termletmkRHole=DAst.make@@GHole(Evar_kinds.InternalHole,Namegen.IntroAnonymous,None)letrecmkRHolesn=ifn>0thenmkRHole::mkRHoles(n-1)else[]letrecisRHolescl=matchclwith|[]->true|c::l->matchDAst.getcwithGHole_->isRHolesl|_->falseletmkRAppfargs=ifargs=[]thenfelseDAst.make@@GApp(f,args)letmkRVarid=DAst.make@@GRef(GlobRef.VarRefid,None)letmkRltacVarid=DAst.make@@GVar(id)letmkRCastrcrt=DAst.make@@GCast(rc,CastConvrt)letmkRType=DAst.make@@GSort(UAnonymous{rigid=true})letmkRProp=DAst.make@@GSort(UNamed[GProp,0])letmkRArrowrt1rt2=DAst.make@@GProd(Anonymous,Explicit,rt1,rt2)letmkRConstructc=DAst.make@@GRef(GlobRef.ConstructRefc,None)letmkRIndmind=DAst.make@@GRef(GlobRef.IndRefmind,None)letmkRLambdanst=DAst.make@@GLambda(n,Explicit,s,t)letrecmkRnatn=ifn<=0thenDAst.make@@GRef(Coqlib.lib_ref"num.nat.O",None)elsemkRApp(DAst.make@@GRef(Coqlib.lib_ref"num.nat.S",None))[mkRnat(n-1)]letglob_constristgenv=function|_,Somece->letvars=Id.Map.fold(funx_accu->Id.Set.addxaccu)ist.Tacinterp.lfunId.Set.emptyinletltacvars={Constrintern.empty_ltac_signwithConstrintern.ltac_vars=vars}inConstrintern.intern_genPretyping.WithoutTypeConstraint~ltacvarsgenvEvd.(from_envgenv)ce|rc,None->rcletpf_intern_termistgl(_,c)=glob_constrist(pf_envgl)cletintern_termistenv(_,c)=glob_constristenvc(* Estimate a bound on the number of arguments of a raw constr. *)(* This is not perfect, because the unifier may fail to *)(* typecheck the partial application, so we use a minimum of 5. *)(* Also, we don't handle delayed or iterated coercions to *)(* FUNCLASS, which is probably just as well since these can *)(* lead to infinite arities. *)letsplay_open_constrenv(sigma,c)=lett=Retyping.get_type_ofenvsigmacinReductionops.splay_prodenvsigmatletisAppIndenvsigmac=letc=Reductionops.clos_whd_flagsCClosure.allenvsigmacinletc,_=decompose_app_vectsigmacinEConstr.isIndsigmac(** Generic argument-based globbing/typing utilities *)letinterp_refineistglrc=letconstrvars=Tacinterp.extract_ltac_constr_valuesist(pf_envgl)inletvars={Glob_ops.empty_lvarwithLtac_pretype.ltac_constrs=constrvars;ltac_genargs=ist.Tacinterp.lfun}inletkind=Pretyping.OfType(pf_conclgl)inletflags={Pretyping.use_typeclasses=Pretyping.UseTC;solve_unification_constraints=true;fail_evar=false;expand_evars=true;program_mode=false;polymorphic=false;}inletsigma,c=Pretyping.understand_ltacflags(pf_envgl)(projectgl)varskindrcin(* ppdebug(lazy(str"sigma@interp_refine=" ++ pr_evar_map None sigma)); *)debug_ssr(fun()->str"c@interp_refine="++Printer.pr_econstr_env(pf_envgl)sigmac);(sigma,(sigma,c))letinterp_open_constrenvsigma0istgc=let(sigma,(c,_))=Tacinterp.interp_open_constr_with_bindingsistenvsigma0(gc,Tactypes.NoBindings)in(sigma0,(sigma,c))letinterp_termenvsigmaist(_,c)=snd(interp_open_constrenvsigmaistc)letof_ftacticftacgl=letr=refNoneinlettac=Ftactic.runftac(funans->r:=Someans;Proofview.tclUNIT())inlettac=Proofview.V82.of_tactictacinlet{sigma=sigma}=tacglinletans=match!rwith|None->assertfalse(* If the tactic failed we should not reach this point *)|Someans->ansin(sigma,ans)letinterp_witwitistglx=letglobarg=in_gen(glbwitwit)xinletarg=Tacinterp.interp_genargistglobarginlet(sigma,arg)=of_ftacticargglinsigma,Tacinterp.Value.cast(topwitwit)argletinterp_hypistgl(SsrHyp(loc,id))=lets,id'=interp_witwit_hypistglCAst.(make?locid)inifnot_section_idid'thens,SsrHyp(loc,id')elsehyp_err?loc"Can't clear section hypothesis "id'letinterp_hypsistglghyps=lethyps=List.mapsnd(List.map(interp_hypistgl)ghyps)incheck_hyps_uniq[]hyps;Tacmach.projectgl,hyps(* Old terms *)letmk_termkc=k,(mkRHole,Somec)letmk_ltermc=mk_termNoFlagc(* New terms *)letmk_ast_closure_termat={annotation=a;body=t;interp_env=None;glob_env=None;}letglob_ast_closure_term(ist:Genintern.glob_sign)t=letist={ast_ltacvars=ist.Genintern.ltacvars;ast_intern_sign=ist.Genintern.intern_sign;ast_extra=ist.Genintern.extra;}in{twithglob_env=Someist}letsubst_ast_closure_term(_s:Mod_subst.substitution)t=(* _s makes sense only for glob constr *)tletinterp_ast_closure_term(ist:Geninterp.interp_sign)(gl:'goalEvd.sigma)t=(* gl is only useful if we want to interp *now*, later we have
* a potentially different gl.sigma *)Tacmach.projectgl,{twithinterp_env=Someist}letssrterm_of_ast_closure_term{body;annotation}=letc=matchannotationwith|`Parens->InParens|`At->WithAt|_->NoFlaginmk_termcbodyletssrdgens_of_parsed_dgens=function|[],clr->{dgens=[];gens=[];clr}|[gens],clr->{dgens=[];gens;clr}|[dgens;gens],clr->{dgens;gens;clr}|_->assertfalseletnbargs_open_constrenvoc=letpl,_=splay_open_constrenvocinList.lengthplletpf_nbargsenvsigmac=nbargs_open_constrenv(sigma,c)letinternal_names=ref[]letadd_internal_namept=internal_names:=pt::!internal_namesletis_internal_names=List.exists(funp->ps)!internal_nameslettmp_tag="_the_"lettmp_post="_tmp_"letmk_tmp_idi=Id.of_string(Printf.sprintf"%s%s%s"tmp_tag(CString.ordinali)tmp_post)letnew_tmp_idctx=letid=mk_tmp_id(1+List.lengthctx.tmp_ids)inletorig=refAnonymousin(id,orig),{ctxwithtmp_ids=(id,orig)::ctx.tmp_ids};;letmk_internal_ids=lets'=Printf.sprintf"_%s_"sinlets'=String.map(func->ifc=' 'then'_'elsec)s'inadd_internal_name((=)s');Id.of_strings'letsame_prefixstn=letrecloopi=i=n||s.[i]=t.[i]&&loop(i+1)inloop0letskip_digitss=letn=String.lengthsinletrecloopi=ifi<n&&is_digits.[i]thenloop(i+1)elseiinloopletmk_tagged_idti=Id.of_string(Printf.sprintf"%s%d_"ti)letis_taggedts=letn=String.lengths-1andm=String.lengthtinm<n&&s.[n]='_'&&same_prefixstm&&skip_digitssm=nletevar_tag="_evar_"let_=add_internal_name(is_taggedevar_tag)letmk_evar_namen=Name(mk_tagged_idevar_tagn)letssr_anon_hyp="Hyp"letwildcard_tag="_the_"letwildcard_post="_wildcard_"letmk_wildcard_idi=Id.of_string(Printf.sprintf"%s%s%s"wildcard_tag(CString.ordinali)wildcard_post)lethas_wildcard_tags=letn=String.lengthsinletm=String.lengthwildcard_taginletm'=String.lengthwildcard_postinn<m+m'+2&&same_prefixswildcard_tagm&&String.subs(n-m')m'=wildcard_post&&skip_digitssm=n-m'-2let_=add_internal_namehas_wildcard_tagletnew_wild_idctx=leti=1+List.lengthctx.wild_idsinletid=mk_wildcard_idiinid,{ctxwithwild_ids=id::ctx.wild_ids}letdischarged_tag="_discharged_"letmk_discharged_idid=Id.of_string(Printf.sprintf"%s%s_"discharged_tag(Id.to_stringid))lethas_discharged_tags=letm=String.lengthdischarged_tagandn=String.lengths-1inm<n&&s.[n]='_'&&same_prefixsdischarged_tagmlet_=add_internal_namehas_discharged_tagletis_discharged_idid=has_discharged_tag(Id.to_stringid)letmax_suffixm(t,j0astj0)id=lets=Id.to_stringidinletn=String.lengths-1inletdn=String.lengtht-1-ninleti0=j0-dninifnot(i0>=m&&s.[n]='_'&&same_prefixstm)thentj0elseletrecloopi=ifi<i0&&s.[i]='0'thenloop(i+1)elseif(ifi<i0thenskip_digitssi=nelsele_s_ti)thens,ielsetj0andle_s_ti=letds=s.[i]anddt=t.[i+dn]inifds=dttheni=n||le_s_t(i+1)elsedt<ds&&skip_digitssi=ninloopm(** creates a fresh (w.r.t. `gl_ids` and internal names) inaccessible name of the form _tXX_ *)letmk_anon_idtgl_ids=letm,si0,id0=lets=ref(Printf.sprintf"_%s_"t)inifis_internal_name!sthens:="_"^!s;letn=String.length!s-1inletrecloopij=letd=!s.[i]inifnot(is_digitd)theni+1,jelseloop(i-1)(ifd='0'thenjelsei)inletm,j=loop(n-1)ninm,(!s,j),Id.of_string_soft!sinifnot(List.memid0gl_ids)thenid0elselets,i=List.fold_left(max_suffixm)si0gl_idsinletopenBytesinlets=of_stringsinletn=lengths-1inletrecloopi=ifgetsi='9'then(setsi'0';loop(i-1))elseifi<mthen(setsn'0';setsm'1';cats(of_string"_"))else(setsi(Char.chr(Char.code(getsi)+1));s)inId.of_string_soft(Bytes.to_string(loop(n-1)))letconvert_concl_no_checkt=Tactics.convert_concl~check:falsetDEFAULTcastletconvert_concl~checkt=Tactics.convert_concl~checktDEFAULTcastletrename_hd_prodorig_name_refgl=matchEConstr.kind(projectgl)(pf_conclgl)with|Prod(x,src,tgt)->letx={xwithbinder_name=!orig_name_ref}inProofview.V82.of_tactic(convert_concl_no_check(EConstr.mkProd(x,src,tgt)))gl|_->CErrors.anomaly(str"gentac creates no product")(* Reduction that preserves the Prod/Let spine of the "in" tactical. *)letinc_safen=ifn=0thennelsen+1letrecsafe_depthsc=matchEConstr.kindscwith|LetIn({binder_name=Namex},_,_,c')whenis_discharged_idx->safe_depthsc'+1|LetIn(_,_,_,c')|Prod(_,_,c')->inc_safe(safe_depthsc')|_->0letred_safe(r:Reductionops.reduction_function)esc0=letrecred_toecn=matchEConstr.kindscwith|Prod(x,t,c')whenn>0->lett'=restinlete'=EConstr.push_rel(RelDecl.LocalAssum(x,t'))einEConstr.mkProd(x,t',red_toe'c'(n-1))|LetIn(x,b,t,c')whenn>0->lett'=restinlete'=EConstr.push_rel(RelDecl.LocalAssum(x,t'))einEConstr.mkLetIn(x,resb,t',red_toe'c'(n-1))|_->rescinred_toec0(safe_depthsc0)letis_id_constrsigmac=matchEConstr.kindsigmacwith|Lambda(_,_,c)whenEConstr.isRelsigmac->1=EConstr.destRelsigmac|_->falseletred_product_skip_idenvsigmac=matchEConstr.kindsigmacwith|App(hd,args)whenArray.lengthargs=1&&is_id_constrsigmahd->args.(0)|_->tryTacred.red_productenvsigmacwith_->cletssrevaltacistgtac=Tacinterp.tactic_of_valueistgtac(** Open term to lambda-term coercion *)(* {{{ ************************************)(* This operation takes a goal gl and an open term (sigma, t), and *)(* returns a term t' where all the new evars in sigma are abstracted *)(* with the mkAbs argument, i.e., for mkAbs = mkLambda then there is *)(* some duplicate-free array args of evars of sigma such that the *)(* term mkApp (t', args) is convertible to t. *)(* This makes a useful shorthand for local definitions in proofs, *)(* i.e., pose succ := _ + 1 means pose succ := fun n : nat => n + 1, *)(* and, in context of the 4CT library, pose mid := maps id means *)(* pose mid := fun d : detaSet => @maps d d (@id (datum d)) *)(* Note that this facility does not extend to set, which tries *)(* instead to fill holes by matching a goal subterm. *)(* The argument to "have" et al. uses product abstraction, e.g. *)(* have Hmid: forall s, (maps id s) = s. *)(* stands for *)(* have Hmid: forall (d : dataSet) (s : seq d), (maps id s) = s. *)(* We also use this feature for rewrite rules, so that, e.g., *)(* rewrite: (plus_assoc _ 3). *)(* will execute as *)(* rewrite (fun n => plus_assoc n 3) *)(* i.e., it will rewrite some subterm .. + (3 + ..) to .. + 3 + ... *)(* The convention is also used for the argument of the congr tactic, *)(* e.g., congr (x + _ * 1). *)(* Replace new evars with lambda variables, retaining local dependencies *)(* but stripping global ones. We use the variable names to encode the *)(* the number of dependencies, so that the transformation is reversible. *)letenv_sizeenv=List.length(Environ.named_contextenv)letpf_conclgl=EConstr.Unsafe.to_constr(pf_conclgl)letpf_get_hypglx=EConstr.Unsafe.to_named_decl(pf_get_hypglx)letpf_e_type_ofglt=letsigma,env,it=projectgl,pf_envgl,sig_itglinletsigma,ty=Typing.type_ofenvsigmatinre_sigitsigma,tyletpf_resolve_typeclasses~where~failgl=letsigma,env,it=projectgl,pf_envgl,sig_itglinletfilter=letevset=Evarutil.undefined_evars_of_termsigmawhereinfunk_->Evar.Set.memkevsetinletsigma=Typeclasses.resolve_typeclasses~filter~failenvsigmainre_sigitsigmaletresolve_typeclasses~where~failenvsigma=letfilter=letevset=Evarutil.undefined_evars_of_termsigmawhereinfunk_->Evar.Set.memkevsetinletsigma=Typeclasses.resolve_typeclasses~filter~failenvsigmainsigmaletnf_evarsigmat=EConstr.Unsafe.to_constr(Evarutil.nf_evarsigma(EConstr.of_constrt))letabs_evars2envsigma0rigid(sigma,c0)=letc0=EConstr.to_constr~abort_on_undefined_evars:falsesigmac0inletsigma0,ucst=sigma0,Evd.evar_universe_contextsigmainletnenv=env_sizeenvinletabs_evarnk=letevi=Evd.findsigmakinletconcl=EConstr.Unsafe.to_constrevi.evar_conclinletdc=EConstr.Unsafe.to_named_context(CList.firstnn(evar_filtered_contextevi))inletabs_dcc=function|NamedDecl.LocalDef(x,b,t)->mkNamedLetInxbt(mkArrowtx.binder_relevancec)|NamedDecl.LocalAssum(x,t)->mkNamedProdxtcinlett=Context.Named.fold_insideabs_dc~init:concldcinnf_evarsigmatinletrecputevlistc=matchConstr.kindcwith|Evar(k,a)->ifList.mem_assockevlist||Evd.memsigma0k||List.memkrigidthenevlistelseletn=max0(List.lengtha-nenv)inlett=abs_evarnkin(k,(n,t))::putevlistt|_->Constr.foldputevlistcinletevlist=put[]c0inifevlist=[]then0,EConstr.of_constrc0,[],ucstelseletreclookupki=function|[]->0,0|(k',(n,_))::evl->ifk=k'theni,nelselookupk(i+1)evlinletrecgetic=matchConstr.kindcwith|Evar(ev,a)->letj,n=lookupevievlistinifj=0thenConstr.map(geti)celseifn=0thenmkReljelseleta=Array.of_listainmkApp(mkRelj,Array.initn(funk->getia.(n-1-k)))|_->Constr.map_with_binders((+)1)geticinletrecloopci=function|(_,(n,t))::evl->loop(mkLambda(make_annot(mk_evar_namen)Sorts.Relevant,get(i-1)t,c))(i-1)evl|[]->cinList.lengthevlist,EConstr.of_constr(loop(get1c0)1evlist),List.mapfstevlist,ucstletpf_abs_evars2glrigidc=abs_evars2(pf_envgl)(projectgl)rigidcletabs_evarsenvsigmat=abs_evars2envsigma[]tletpf_abs_evarsglt=pf_abs_evars2gl[]t(* As before but if (?i : T(?j)) and (?j : P : Prop), then the lambda for i
* looks like (fun evar_i : (forall pi : P. T(pi))) thanks to "loopP" and all
* occurrences of evar_i are replaced by (evar_i evar_j) thanks to "app".
*
* If P can be solved by ssrautoprop (that defaults to trivial), then
* the corresponding lambda looks like (fun evar_i : T(c)) where c is
* the solution found by ssrautoprop.
*)letssrautoprop_tac=ref(Proofview.Goal.enter(fungl->assertfalse))(* Thanks to Arnaud Spiwack for this snippet *)letcall_on_evartaces=let{it=gs;sigma=s}=tac{it=e;sigma=s;}ings,sopenPpletpp_=()(* FIXME *)moduleIntset=Evar.Setletabs_evars_pirrelenvsigma0(sigma,c0)=pp(lazy(str"==PF_ABS_EVARS_PIRREL=="));pp(lazy(str"c0= "++Printer.pr_constr_envenvsigmac0));letc0=nf_evarsigma0(nf_evarsigmac0)inletnenv=env_sizeenvinletabs_evarnk=letevi=Evd.findsigmakinletconcl=EConstr.Unsafe.to_constrevi.evar_conclinletdc=EConstr.Unsafe.to_named_context(CList.firstnn(evar_filtered_contextevi))inletabs_dcc=function|NamedDecl.LocalDef(x,b,t)->mkNamedLetInxbt(mkArrowtx.binder_relevancec)|NamedDecl.LocalAssum(x,t)->mkNamedProdxtcinlett=Context.Named.fold_insideabs_dc~init:concldcinnf_evarsigma0(nf_evarsigmat)inletrecputevlistc=matchConstr.kindcwith|Evar(k,a)->ifList.mem_assockevlist||Evd.memsigma0kthenevlistelseletn=max0(List.lengtha-nenv)inletk_ty=Retyping.get_sort_family_ofenvsigma(Evd.evar_concl(Evd.findsigmak))inletis_prop=k_ty=InPropinlett=abs_evarnkin(k,(n,t,is_prop))::putevlistt|_->Constr.foldputevlistcinletevlist=put[]c0inifevlist=[]then0,c0elseletpr_constrt=Printer.pr_econstr_envenvsigma(Reductionops.nf_betaenvsigma0(EConstr.of_constrt))inpp(lazy(str"evlist="++pr_list(fun()->str";")(fun(k,_)->Evar.printk)evlist));letevplist=letdepev=List.fold_left(funevs(_,(_,t,_))->lett=EConstr.of_constrtinIntset.unionevs(Evarutil.undefined_evars_of_termsigmat))Intset.emptyevlistinList.filter(fun(i,(_,_,b))->b&&Intset.memidepev)evlistinletevlist,evplist,sigma=ifevplist=[]thenevlist,[],sigmaelseList.fold_left(fun(ev,evp,sigma)(i,(_,t,_)asp)->tryletng,sigma=call_on_evar(Proofview.V82.of_tactic!ssrautoprop_tac)isigmainif(ng<>[])thenerrorstrm(str"Should we tell the user?");List.filter(fun(j,_)->j<>i)ev,evp,sigmawith_->ev,p::evp,sigma)(evlist,[],sigma)(List.revevplist)inletc0=nf_evarsigmac0inletevlist=List.map(fun(x,(y,t,z))->x,(y,nf_evarsigmat,z))evlistinletevplist=List.map(fun(x,(y,t,z))->x,(y,nf_evarsigmat,z))evplistinpp(lazy(str"c0= "++pr_constrc0));letreclookupki=function|[]->0,0|(k',(n,_,_))::evl->ifk=k'theni,nelselookupk(i+1)evlinletrecgetevlistic=matchConstr.kindcwith|Evar(ev,a)->letj,n=lookupevievlistinifj=0thenConstr.map(getevlisti)celseifn=0thenmkReljelseleta=Array.of_listainmkApp(mkRelj,Array.initn(funk->getevlistia.(n-1-k)))|_->Constr.map_with_binders((+)1)(getevlist)icinletrecappextra_argsic=matchdecompose_appcwith|hd,argswhenisRelhd&&destRelhd=i->letj=destRelhdinmkApp(mkRelj,Array.of_list(List.map(Vars.lift(i-1))extra_args@args))|_->Constr.map_with_binders((+)1)(appextra_args)icinletrecloopPevlistci=function|(_,(n,t,_))::evl->lett=getevlist(i-1)tinletn=Name(Id.of_string(ssr_anon_hyp^string_of_intn))inloopPevlist(mkProd(make_annotnSorts.Relevant,t,c))(i-1)evl|[]->cinletrecloopci=function|(_,(n,t,_))::evl->letevs=Evarutil.undefined_evars_of_termsigma(EConstr.of_constrt)inlett_evplist=List.filter(fun(k,_)->Intset.memkevs)evplistinlett=loopPt_evplist(gett_evplist1t)1t_evplistinlett=getevlist(i-1)tinletextra_args=List.map(fun(k,_)->mkRel(fst(lookupkievlist)))(List.revt_evplist)inletc=ifextra_args=[]thencelseappextra_args1cinloop(mkLambda(make_annot(mk_evar_namen)Sorts.Relevant,t,c))(i-1)evl|[]->cinletres=loop(getevlist1c0)1evlistinpp(lazy(str"res= "++pr_constrres));List.lengthevlist,resletpf_abs_evars_pirrelglc=abs_evars_pirrel(pf_envgl)(projectgl)c(* Strip all non-essential dependencies from an abstracted term, generating *)(* standard names for the abstracted holes. *)letnb_evar_deps=function|Nameid->lets=Id.to_stringidinifnot(is_taggedevar_tags)then0elseletm=String.lengthevar_tagin(tryint_of_string(String.subsm(String.lengths-1-m))with_->0)|_->0lettype_idenvsigmat=Id.of_string(Namegen.hdcharenvsigmat)letpf_type_idglt=type_id(pf_envgl)(projectgl)tletpfe_type_ofglt=letsigma,ty=pf_type_ofgltinre_sig(sig_itgl)sigma,tyletpfe_new_typegl=letsigma,env,it=projectgl,pf_envgl,sig_itglinletsigma,t=Evarutil.new_Typesigmainre_sigitsigma,tletpfe_type_relevance_ofglt=letgl,ty=pfe_type_ofgltingl,ty,pf_applyRetyping.relevance_of_termgltletpf_type_ofglt=letsigma,ty=pf_type_ofgl(EConstr.of_constrt)inre_sig(sig_itgl)sigma,EConstr.Unsafe.to_constrtyletabs_ctermenvsigmanc0=ifn<=0thenc0elseletc0=EConstr.Unsafe.to_constrc0inletnoargs=[|0|]inleteva=Array.makennoargsinletrecstripic=matchConstr.kindcwith|App(f,a)whenisRelf->letj=i-destRelfinifj>=n||eva.(j)=noargsthenmkApp(f,Array.map(stripi)a)elseletdp=eva.(j)inletnd=Array.lengthdp-1inletmkargk=stripia.(ifk<ndthendp.(k+1)-jelsek+dp.(0))inmkApp(f,Array.init(Array.lengtha-dp.(0))mkarg)|_->Constr.map_with_binders((+)1)stripicinletrecstrip_ndepsjic=matchConstr.kindcwith|Prod(x,t,c1)wheni<j->letdl,c2=strip_ndepsj(i+1)c1inifVars.noccurn1c2thendl,Vars.lift(-1)c2elsei::dl,mkProd(x,stripit,c2)|LetIn(x,b,t,c1)wheni<j->let_,_,c1'=destProdc1inletdl,c2=strip_ndepsj(i+1)c1'inifVars.noccurn1c2thendl,Vars.lift(-1)c2elsei::dl,mkLetIn(x,stripib,stripit,c2)|_->[],stripicinletrecstrip_evarsic=matchConstr.kindcwith|Lambda(x,t1,c1)wheni<n->letna=nb_evar_depsx.binder_nameinletdl,t2=strip_ndeps(i+na)it1inletna'=List.lengthdlineva.(i)<-Array.of_list(na-na'::dl);letx'=ifna'=0thenName(type_idenvsigma(EConstr.of_constrt2))elsemk_evar_namena'inmkLambda({xwithbinder_name=x'},t2,strip_evars(i+1)c1)(* if noccurn 1 c2 then lift (-1) c2 else
mkLambda (Name (pf_type_id gl t2), t2, c2) *)|_->stripicinEConstr.of_constr(strip_evars0c0)letpf_abs_ctermglnc0=abs_cterm(pf_envgl)(projectgl)nc0(* }}} *)letpf_merge_ucucgl=re_sig(sig_itgl)(Evd.merge_universe_contextgl.Evd.sigmauc)letpf_merge_uc_ofsigmagl=letucst=Evd.evar_universe_contextsigmainpf_merge_ucucstglletrecconstr_namesigmac=matchEConstr.kindsigmacwith|Varid->Nameid|Cast(c',_,_)->constr_namesigmac'|Const(cn,_)->Name(Label.to_id(Constant.labelcn))|App(c',_)->constr_namesigmac'|_->Anonymousletpf_mkprodglc?(name=constr_name(projectgl)c)cl=letgl,t,r=pfe_type_relevance_ofglcinifname<>Anonymous||EConstr.Vars.noccurn(projectgl)1clthengl,EConstr.mkProd(make_annotnamer,t,cl)elsegl,EConstr.mkProd(make_annot(Name(pf_type_idglt))r,t,cl)letpf_abs_prodnameglccl=pf_mkprodglc~name(Termops.subst_term(projectgl)ccl)(** look up a name in the ssreflect internals module *)letssrdirpath=DirPath.make[Id.of_string"ssreflect"]letssrqidname=Libnames.make_qualidssrdirpath(Id.of_stringname)letmkSsrRefname=letqn=Format.sprintf"plugins.ssreflect.%s"nameinifCoqlib.has_refqnthenCoqlib.lib_refqnelseCErrors.user_errPp.(str"Small scale reflection library not loaded ("++strname++str")")letmkSsrRRefname=(DAst.make@@GRef(mkSsrRefname,None)),NoneletmkSsrConstnameenvsigma=EConstr.fresh_globalenvsigma(mkSsrRefname)letpf_mkSsrConstnamegl=letsigma,env,it=projectgl,pf_envgl,sig_itglinlet(sigma,t)=mkSsrConstnameenvsigmaint,re_sigitsigmaletpf_fresh_globalnamegl=letsigma,env,it=projectgl,pf_envgl,sig_itglinletsigma,t=Evd.fresh_globalenvsigmanameinEConstr.Unsafe.to_constrt,re_sigitsigmaletmkProttcgl=letprot,gl=pf_mkSsrConst"protect_term"glinEConstr.mkApp(prot,[|t;c|]),glletmkEtaAppcnimin=letopenEConstrinifn=0thencelseletnargs,mkarg=ifn<0then-n,(funi->mkRel(imin+i))elseletimax=imin+n-1inn,(funi->mkRel(imax-i))inmkApp(c,Array.initnargsmkarg)letmkRefltcgl=letsigma=projectglinlet(sigma,refl)=EConstr.fresh_global(pf_envgl)sigmaCoqlib.(lib_ref"core.eq.refl")inEConstr.mkApp(refl,[|t;c|]),{glwithsigma}letdischarge_hyp(id',(id,mode))gl=letcl'=Vars.subst_varid(pf_conclgl)inletdecl=pf_get_hypglidinmatchdecl,modewith|NamedDecl.LocalAssum_,_|NamedDecl.LocalDef_,"("->letid'={(NamedDecl.get_annotdecl)withbinder_name=Nameid'}inProofview.V82.of_tactic(Tactics.apply_type~typecheck:true(EConstr.of_constr(mkProd(id',NamedDecl.get_typedecl,cl')))[EConstr.of_constr(mkVarid)])gl|NamedDecl.LocalDef(_,v,t),_->letid'={(NamedDecl.get_annotdecl)withbinder_name=Nameid'}inProofview.V82.of_tactic(convert_concl~check:true(EConstr.of_constr(mkLetIn(id',v,t,cl'))))gl(* wildcard names *)letclear_wildswildsgl=Proofview.V82.of_tactic(Tactics.clear(List.filter(funid->List.memidwilds)(pf_ids_of_hypsgl)))glletclear_with_wildswildsclr0gl=letextend_clrclrnd=letid=NamedDecl.get_idndinifList.memidclr||not(List.memidwilds)thenclrelseletvars=Termops.global_vars_set_of_decl(pf_envgl)(projectgl)ndinletoccursid'=Id.Set.memid'varsinifList.existsoccursclrthenid::clrelseclrinProofview.V82.of_tactic(Tactics.clear(Context.Named.fold_insideextend_clr~init:clr0(Tacmach.pf_hypsgl)))glletclear_wilds_and_tmp_and_delayed_idsgl=let_,ctx=pull_ctxglintac_ctx(tclTHEN(clear_with_wildsctx.wild_idsctx.delayed_clears)(clear_wilds(List.mapfstctx.tmp_ids@ctx.wild_ids)))glletview_errorsgv=errorstrm(str("Cannot "^s^" view ")++pr_termgv)openLocus(****************************** tactics ***********************************)letrewritetac?(under=false)dirc=(* Due to the new optional arg ?tac, application shouldn't be too partial *)letopenProofview.NotationsinProofview.Goal.enterbeginfun_->Equality.general_rewrite~where:None~l2r:(dir=L2R)AllOccurrences~freeze:true~dep:false~with_evars:false(c,Tactypes.NoBindings)<*>ifunderthenProofview.cycle1elseProofview.tclUNIT()end(**********************`:********* hooks ************************************)typename_hint=(int*EConstr.typesarray)optionrefletpf_abs_ssrterm?(resolve_typeclasses=false)istglt=letsigma,ctast=interp_term(pf_envgl)(projectgl)isttinletsigma,_ast=letenv=pf_envglinifnotresolve_typeclassesthentelseletsigma=Typeclasses.resolve_typeclasses~fail:falseenvsigmainsigma,Evarutil.nf_evarsigmactinletn,c,abstracted_away,ucst=pf_abs_evarsgltinList.fold_leftEvd.removesigmaabstracted_away,pf_abs_ctermglnc,ucst,nlettop_id=mk_internal_id"top assumption"letssr_n_tacseedn=Proofview.Goal.enterbeginfungl->letname=ifn=-1thenseedelse("ssr"^seed^string_of_intn)inletfailmsg=CErrors.user_err(Pp.strmsg)inlettacname=tryTacenv.locate_tactic(Libnames.qualid_of_ident(Id.of_stringname))withNot_found->tryTacenv.locate_tactic(ssrqidname)withNot_found->ifn=-1thenfail"The ssreflect library was not loaded"elsefail("The tactic "^name^" was not found")inlettacexpr=Tacexpr.Reference(ArgArg(Loc.tag@@tacname))inTacinterp.eval_tactic@@CAst.make(Tacexpr.TacArgtacexpr)endletdonetacn=ssr_n_tac"done"nopenConstrexpropenUtil(** Constructors for constr_expr *)letmkCProploc=CAst.make?loc@@CSort(UNamed[CProp,0])letmkCTypeloc=CAst.make?loc@@CSort(UAnonymous{rigid=true})letmkCVar?locid=CAst.make?loc@@CRef(qualid_of_ident?locid,None)letrecmkCHoles?locn=ifn<=0then[]else(CAst.make?loc@@CHole(None,Namegen.IntroAnonymous,None))::mkCHoles?loc(n-1)letmkCHoleloc=CAst.make?loc@@CHole(None,Namegen.IntroAnonymous,None)letmkCLambda?locnametyt=CAst.make?loc@@CLambdaN([CLocalAssum([CAst.make?locname],DefaultExplicit,ty)],t)letmkCArrow?loctyt=CAst.make?loc@@CProdN([CLocalAssum([CAst.makeAnonymous],DefaultExplicit,ty)],t)letmkCCast?loctty=CAst.make?loc@@CCast(t,CastConvty)letrecisCHoles=function{CAst.v=CHole_}::cl->isCHolescl|cl->cl=[]letrecisCxHoles=function({CAst.v=CHole_},None)::ch->isCxHolesch|_->falseletpf_interp_ty?(resolve_typeclasses=false)envsigma0istty=letn_binders=ref0inletty=matchtywith|a,(t,None)->letrecforce_typety=DAst.(map(function|GProd(x,k,s,t)->incrn_binders;GProd(x,k,s,force_typet)|GLetIn(x,v,oty,t)->incrn_binders;GLetIn(x,v,oty,force_typet)|_->DAst.get(mkRCasttymkRType)))tyina,(force_typet,None)|_,(_,Somety)->letrecforce_typety=CAst.(map(function|CProdN(abs,t)->n_binders:=!n_binders+List.length(List.flatten(List.map(functionCLocalAssum(nal,_,_)->nal|CLocalDef(na,_,_)->[na]|CLocalPattern_->(* We count a 'pat for 1; TO BE CHECKED *)[CAst.makeName.Anonymous])abs));CProdN(abs,force_typet)|CLetIn(n,v,oty,t)->incrn_binders;CLetIn(n,v,oty,force_typet)|_->(mkCCastty(mkCTypeNone)).v))tyinmk_termNoFlag(force_typety)inletstrip_cast(sigma,t)=letopenEConstrinletrecauxt=matchkind_of_typesigmatwith|CastType(t,ty)when!n_binders=0&&isSortsigmaty->t|ProdType(n,s,t)->decrn_binders;mkProd(n,s,auxt)|LetInType(n,v,ty,t)->decrn_binders;mkLetIn(n,v,ty,auxt)|_->anomaly"pf_interp_ty: ssr Type cast deleted by typecheck"insigma,auxtinletsigma,ctyasty=strip_cast(interp_termenvsigma0istty)inletty=ifnotresolve_typeclassesthentyelseletsigma=Typeclasses.resolve_typeclasses~fail:falseenvsigmainsigma,Evarutil.nf_evarsigmactyinletn,c,_,ucst=abs_evarsenvsigma0tyinletlam_c=abs_ctermenvsigma0ncinletctx,c=EConstr.decompose_lam_n_assumsigmanlam_cinn,EConstr.it_mkProd_or_LetIncctx,lam_c,ucst;;(* TASSI: given (c : ty), generates (c ??? : ty[???/...]) with m evars *)exceptionNotEnoughProductsletsaturate?(beta=false)?(bi_types=false)envsigmac?(ty=Retyping.get_type_ofenvsigmac)m=letreclooptyargssigman=letopenEConstrinifn=0thenletargs=List.revargsin(ifbetathenReductionops.whd_betaenvsigmaelsefunx->x)(EConstr.mkApp(c,Array.of_list(List.mapsndargs))),ty,args,sigmaelsematchkind_of_typesigmatywith|ProdType(_,src,tgt)->letsigma=create_evar_defssigmainlet(sigma,x)=Evarutil.new_evarenvsigma(ifbi_typesthenReductionops.nf_betaiotaenvsigmasrcelsesrc)inloop(EConstr.Vars.subst1xtgt)((m-n,x)::args)sigma(n-1)|CastType(t,_)->looptargssigman|LetInType(_,v,_,t)->loop(EConstr.Vars.subst1vt)argssigman|SortType_->assertfalse|AtomicType_->letty=(* FIXME *)(Reductionops.whd_allenvsigma)tyinmatchkind_of_typesigmatywith|ProdType_->looptyargssigman|_->raiseNotEnoughProductsinloopty[]sigmamletpf_saturate?beta?bi_typesglc?tym=letenv,sigma,si=pf_envgl,projectgl,sig_itglinlett,ty,args,sigma=saturate?beta?bi_typesenvsigmac?tymint,ty,args,re_sigsisigmaletpf_partial_solutiongltevl=letsigma,g=projectgl,sig_itglinletsigma=Goal.V82.partial_solution(pf_envgl)sigmagtinre_sig(List.map(funx->(fst(EConstr.destEvarsigmax)))evl)sigmaletdependent_apply_error=tryCErrors.user_err(Pp.str"Could not fill dependent hole in \"apply\"")witherr->err(* TASSI: Sometimes Coq's apply fails. According to my experience it may be
* related to goals that are products and with beta redexes. In that case it
* guesses the wrong number of implicit arguments for your lemma. What follows
* is just like apply, but with a user-provided number n of implicits.
*
* Refine.refine function that handles type classes and evars but fails to
* handle "dependently typed higher order evars".
*
* Refiner.refiner that does not handle metas with a non ground type but works
* with dependently typed higher order metas. *)letapplyn~with_evars?beta?(with_shelve=false)?(first_goes_last=false)nt=Proofview.V82.tacticbeginfungl->ifwith_evarsthenletrefinegl=lett,ty,args,gl=pf_saturate?beta~bi_types:truegltnin(* pp(lazy(str"sigma@saturate=" ++ pr_evar_map None (project gl))); *)letgl=pf_unify_HOglty(Tacmach.pf_conclgl)inletgs=CList.map_filter(fun(_,e)->ifEConstr.isEvar(projectgl)ethenSomeeelseNone)argsinpf_partial_solutiongltgsinProofview.(V82.of_tactic(Tacticals.New.tclTHENLIST[V82.tacticrefine;(ifwith_shelvethenshelve_unifiableelsetclUNIT());(iffirst_goes_lastthencycle1elsetclUNIT())]))glelselett,gl=ifn=0thent,glelseletsigma,si=projectgl,sig_itglinletrecloopsigmaboargs=function(* saturate with metas *)|0->EConstr.mkApp(t,Array.of_list(List.revargs)),re_sigsisigma|n->matchEConstr.kindsigmabowith|Lambda(_,ty,bo)->ifnot(EConstr.Vars.closed0sigmaty)thenraisedependent_apply_error;letm=Evarutil.new_meta()inloop(meta_declaremtysigma)bo((EConstr.mkMetam)::args)(n-1)|_->assertfalseinloopsigmat[]ninpp(lazy(str"Refiner.refiner "++Printer.pr_econstr_env(pf_envgl)(projectgl)t));Proofview.(V82.of_tactic(Tacticals.New.tclTHENLIST[Logic.refiner~check:falseEConstr.Unsafe.(to_constrt);(iffirst_goes_lastthencycle1elsetclUNIT())]))glendletrefine_with?(first_goes_last=false)?beta?(with_evars=true)oc=letopenProofview.NotationsinProofview.Goal.enterbeginfungl->letenv=Proofview.Goal.envglinletsigma=Proofview.Goal.sigmaglinletuct=Evd.evar_universe_context(fstoc)inletn,oc=abs_evars_pirrelenvsigma(fstoc,EConstr.to_constr~abort_on_undefined_evars:false(fstoc)(sndoc))inProofview.Unsafe.tclEVARS(Evd.set_universe_contextsigmauct)<*>Proofview.tclORELSE(applyn~with_evars~first_goes_last~with_shelve:true?betan(EConstr.of_constroc))(fun_->Proofview.tclZEROdependent_apply_error)end(* We wipe out all the keywords generated by the grammar rules we defined. *)(* The user is supposed to Require Import ssreflect or Require ssreflect *)(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *)(* consequence the extended ssreflect grammar. *)let()=CLexer.set_keyword_statefrozen_lexer;;(** Basic tactics *)letrecfst_prodredtac=Proofview.Goal.enterbeginfungl->letconcl=Proofview.Goal.conclglinmatchEConstr.kind(Proofview.Goal.sigmagl)conclwith|Prod(id,_,tgt)|LetIn(id,_,_,tgt)->tacid.binder_name|_->ifredthenTacticals.New.tclZEROMSG(str"No product even after head-reduction.")elseTacticals.New.tclTHENTactics.hnf_in_concl(fst_prodtruetac)endletintroid?(orig=refAnonymous)name=letopenProofview.NotationsinProofview.Goal.enterbeginfungl->letenv=Proofview.Goal.envglinletsigma=Proofview.Goal.sigmaglinletg=Proofview.Goal.conclglinmatchEConstr.kindsigmagwith|App(hd,_)whenEConstr.isLambdasigmahd->convert_concl_no_check(Reductionops.whd_betaenvsigmag)|_->Tacticals.New.tclIDTACend<*>(fst_prodfalse(funid->orig:=id;Tactics.intro_mustbe_forcename))letanontacdeclgl=letid=matchRelDecl.get_namedeclwith|Nameid->ifis_discharged_ididthenidelsemk_anon_id(Id.to_stringid)(Tacmach.pf_ids_of_hypsgl)|_->mk_anon_idssr_anon_hyp(Tacmach.pf_ids_of_hypsgl)inProofview.V82.of_tactic(introidid)glletrecintro_anongl=tryanontac(List.hd(fst(EConstr.decompose_prod_n_assum(projectgl)1(Tacmach.pf_conclgl))))glwitherr0->trytclTHEN(Proofview.V82.of_tacticTactics.red_in_concl)intro_anonglwithewhenCErrors.noncriticale->raiseerr0(* with _ -> CErrors.error "No product even after reduction" *)letis_pf_varsigmac=EConstr.isVarsigmac&¬_section_id(EConstr.destVarsigmac)lethyp_of_varsigmav=SsrHyp(Loc.tag@@EConstr.destVarsigmav)letinterp_clrsigma=function|Someclr,(k,c)when(k=NoFlag||k=WithAt)&&is_pf_varsigmac->hyp_of_varsigmac::clr|Someclr,_->clr|None,_->[](** Basic tacticals *)(** Multipliers *)(* {{{ ***********************************************************)(* tactical *)lettclIDtac=taclettclDOTRYntac=letopenTacticals.Newinifn<=0thentclIDTACelseletrecloopi=ifi=nthentclTRYtacelsetclTRY(tclTHENtac(loop(i+1)))inloop1lettclDOntac=letprefixi=str"At iteration "++inti++str": "inlettac_err_atigl=tryProofview.V82.of_tactictacglwith|CErrors.UserError(l,s)ase->let_,info=Exninfo.captureeinlete'=CErrors.UserError(l,prefixi++s)inExninfo.iraise(e',info)inletrecloopigl=ifi=nthentac_err_atiglelse(tclTHEN(tac_err_ati)(loop(i+1)))glinProofview.V82.tactic~nf_evars:false(loop1)lettclAT_LEAST_ONCEt=letopenTacticals.NewintclTHENt(tclREPEATt)lettclMULT=function|0,May->Tacticals.New.tclREPEAT|1,May->Tacticals.New.tclTRY|n,May->tclDOTRYn|0,Must->tclAT_LEAST_ONCE|n,Mustwhenn>1->tclDOn|_->tclIDletold_cleartacclr=check_hyps_uniq[]clr;Proofview.V82.of_tactic(Tactics.clear(hyps_idsclr))letcleartacclr=check_hyps_uniq[]clr;Tactics.clear(hyps_idsclr)(* }}} *)(** Generalize tactic *)(* XXX the k of the redex should percolate out *)letpf_interp_gen_auxglto_ind((oclr,occ),t)=letpat=interp_cpattern(pf_envgl)(projectgl)tNonein(* UGLY API *)letgl=pf_merge_uc_of(fstpat)glinletcl,env,sigma=Tacmach.pf_conclgl,pf_envgl,projectglinlet(c,ucst),cl=tryfill_occ_pattern~raise_NoMatch:trueenvsigma(EConstr.Unsafe.to_constrcl)patocc1withNoMatch->redex_of_patternenvpat,(EConstr.Unsafe.to_constrcl)inletgl=pf_merge_ucucstglinletc=EConstr.of_constrcinletcl=EConstr.of_constrclinletclr=interp_clrsigma(oclr,(tag_of_cpatternt,c))inifnot(occur_existentialsigmac)theniftag_of_cpatternt=WithAtthenifnot(EConstr.isVarsigmac)thenerrorstrm(str"@ can be used with variables only")elsematchTacmach.pf_get_hypgl(EConstr.destVarsigmac)with|NamedDecl.LocalAssum_->errorstrm(str"@ can be used with let-ins only")|NamedDecl.LocalDef(name,b,ty)->true,pat,EConstr.mkLetIn(map_annotName.mk_namename,b,ty,cl),c,clr,ucst,glelseletgl,ccl=pf_mkprodglcclinfalse,pat,ccl,c,clr,ucst,glelseifto_ind&&occ=Nonethenletnv,p,_,ucst'=pf_abs_evarsgl(fstpat,c)inletucst=UState.unionucstucst'inifnv=0thenanomaly"occur_existential but no evars"elseletgl,pty,rp=pfe_type_relevance_ofglpinfalse,pat,EConstr.mkProd(make_annot(constr_name(projectgl)c)rp,pty,Tacmach.pf_conclgl),p,clr,ucst,glelseCErrors.user_err?loc:(loc_of_cpatternt)(str"generalized term didn't match")letapply_typexxs=Proofview.V82.of_tactic(Tactics.apply_type~typecheck:truexxs)letgenclrtacclcsclr=lettclmyORELSEtac1tac2gl=trytac1glwithewhenCErrors.noncriticale->tac2eglin(* apply_type may give a type error, but the useful message is
* the one of clear. You type "move: x" and you get
* "x is used in hyp H" instead of
* "The term H has type T x but is expected to have type T x0". *)tclTHEN(tclmyORELSE(apply_typeclcs)(funtype_errgl->tclTHEN(tclTHEN(Proofview.V82.of_tactic(Tactics.elim_type(EConstr.of_constr(UnivGen.constr_of_monomorphic_global@@Coqlib.(lib_ref"core.False.type")))))(old_cleartacclr))(fungl->raisetype_err)gl))(old_cleartacclr)letgentacgen=Proofview.V82.tacticbeginfungl->(* ppdebug(lazy(str"sigma@gentac=" ++ pr_evar_map None (project gl))); *)letconv,_,cl,c,clr,ucst,gl=pf_interp_gen_auxglfalsegenindebug_ssr(fun()->str"c@gentac="++pr_econstr_env(pf_envgl)(projectgl)c);letgl=pf_merge_ucucstglinifconvthentclTHEN(Proofview.V82.of_tactic(convert_concl~check:truecl))(old_cleartacclr)glelsegenclrtaccl[c]clrglendletgenstac(gens,clr)=Tacticals.New.tclTHENLIST(cleartacclr::List.rev_mapgentacgens)letgen_tmp_ids?(ist=Geninterp.({lfun=Id.Map.empty;poly=false;extra=Tacinterp.TacStore.empty}))gl=letgl,ctx=pull_ctxglinpush_ctxsctx(tclTHENLIST(List.map(fun(id,orig_ref)->tclTHEN(Proofview.V82.of_tactic(gentac((None,Some(false,[])),cpattern_of_idid)))(rename_hd_prodorig_ref))ctx.tmp_ids)gl);;letpf_interp_gento_indgengl=let_,_,a,b,c,ucst,gl=pf_interp_gen_auxglto_indgenin(a,b,c),pf_merge_ucucstglletpfLIFTf=letopenProofview.Notationsinlethack=refNoneinProofview.V82.tactic(fungl->letg=sig_itglinletx,gl=fglinhack:=Some(x,projectgl);re_sig[g](projectgl))>>=fun()->letx,sigma=option_assert_get!hack(Pp.str"pfLIFT")inProofview.Unsafe.tclEVARSsigma<*>Proofview.tclUNITx;;letis_protecthdenvsigma=let_,protectC=mkSsrConst"protect_term"envsigmainEConstr.eq_constr_nounivssigmahdprotectCletabs_wgenkeep_letfgen(gl,args,c)=letsigma,env=projectgl,pf_envglinletevar_closedtp=ifoccur_existentialsigmatthenCErrors.user_err?loc:(loc_of_cpatternp)~hdr:"ssreflect"(pr_econstr_patenvsigmat++str" contains holes and matches no subterm of the goal")inmatchgenwith|_,Some((x,mode),None)whenmode="@"||(mode=" "&&keep_let)->letx=hoi_idxinletdecl=Tacmach.pf_get_hypglxingl,(ifNamedDecl.is_local_defdeclthenargselseEConstr.mkVarx::args),EConstr.mkProd_or_LetIn(decl|>NamedDecl.to_rel_decl|>RelDecl.set_name(Name(fx)))(EConstr.Vars.subst_varxc)|_,Some((x,_),None)->letx=hoi_idxinlethyp=Tacmach.pf_get_hypglxinletx'=make_annot(Name(fx))(NamedDecl.get_relevancehyp)inletprod=EConstr.mkProd(x',NamedDecl.get_typehyp,EConstr.Vars.subst_varxc)ingl,EConstr.mkVarx::args,prod|_,Some((x,"@"),Somep)->letx=hoi_idxinletcp=interp_cpattern(pf_envgl)(projectgl)pNoneinletgl=pf_merge_uc_of(fstcp)glinlet(t,ucst),c=tryfill_occ_pattern~raise_NoMatch:trueenvsigma(EConstr.Unsafe.to_constrc)cpNone1withNoMatch->redex_of_patternenvcp,(EConstr.Unsafe.to_constrc)inletc=EConstr.of_constrcinlett=EConstr.of_constrtinevar_closedtp;letut=red_product_skip_idenvsigmatinletgl,ty,r=pfe_type_relevance_ofgltinpf_merge_ucucstgl,args,EConstr.mkLetIn(make_annot(Name(fx))r,ut,ty,c)|_,Some((x,_),Somep)->letx=hoi_idxinletcp=interp_cpattern(pf_envgl)(projectgl)pNoneinletgl=pf_merge_uc_of(fstcp)glinlet(t,ucst),c=tryfill_occ_pattern~raise_NoMatch:trueenvsigma(EConstr.Unsafe.to_constrc)cpNone1withNoMatch->redex_of_patternenvcp,(EConstr.Unsafe.to_constrc)inletc=EConstr.of_constrcinlett=EConstr.of_constrtinevar_closedtp;letgl,ty,r=pfe_type_relevance_ofgltinpf_merge_ucucstgl,t::args,EConstr.mkProd(make_annot(Name(fx))r,ty,c)|_->gl,args,cletclr_of_wgengenclrs=matchgenwith|clr,Some((x,_),None)->letx=hoi_idxincleartacclr::cleartac[SsrHyp(Loc.tagx)]::clrs|clr,_->cleartacclr::clrsletreduct_in_concl~checkt=Tactics.reduct_in_concl~check(t,DEFAULTcast)letunfoldcl=letmoduleR=ReductionopsinletmoduleF=CClosure.RedFlagsinreduct_in_concl~check:false(R.clos_norm_flags(F.mkflags(List.map(func->F.fCONST(fst(destConst(EConstr.Unsafe.to_constrc))))cl@[F.fBETA;F.fMATCH;F.fFIX;F.fCOFIX])))openProofviewopenNotationslettacSIGMA=Goal.enter_one~__LOC__beginfung->letk=Goal.goalginletsigma=Goal.sigmagintclUNIT(Tacmach.re_sigksigma)endlettclINTERP_AST_CLOSURE_TERM_AS_CONSTRc=tclINDEPENDENTLbegintacSIGMA>>=fungl->letold_ssrterm=mkRHole,Somec.Ssrast.bodyinletist=option_assert_getc.Ssrast.interp_envPp.(str"tclINTERP_AST_CLOSURE_TERM_AS_CONSTR: term with no ist")inletsigma,t=interp_witStdarg.wit_constristglold_ssrterminUnsafe.tclEVARSsigma<*>tclUNITtendlettacREDUCE_TO_QUANTIFIED_INDty=tacSIGMA>>=fungl->trytclUNIT(Tacmach.pf_reduce_to_quantified_indglty)withe->tclZEROelettacTYPEOFc=Goal.enter_one~__LOC__(fung->letsigma,env=Goal.sigmag,Goal.envginletsigma,ty=Typing.type_ofenvsigmacinUnsafe.tclEVARSsigma<*>tclUNITty)(** This tactic creates a partial proof realizing the introduction rule, but
does not check anything. *)letunsafe_introenvdeclb=letopenContext.Named.DeclarationinRefine.refine~typecheck:falsebeginfunsigma->letctx=Environ.named_context_valenvinletnctx=EConstr.push_named_context_valdeclctxinletinst=EConstr.identity_subst_val(Environ.named_context_valenv)inletninst=EConstr.mkRel1::instinletnb=EConstr.Vars.subst1(EConstr.mkVar(get_iddecl))binletsigma,ev=Evarutil.new_pure_evar~principal:truenctxsigmanbinsigma,EConstr.mkNamedLambda_or_LetIndecl(EConstr.mkEvar(ev,ninst))endletset_decl_idid=letopenContextinfunction|Rel.Declaration.LocalAssum(name,ty)->Named.Declaration.LocalAssum({namewithbinder_name=id},ty)|Rel.Declaration.LocalDef(name,ty,t)->Named.Declaration.LocalDef({namewithbinder_name=id},ty,t)letrecdecompose_assumenvsigmaorig_goal=letopenContextinmatchEConstr.kindsigmaorig_goalwith|Prod(name,ty,t)->Rel.Declaration.LocalAssum(name,ty),t,true|LetIn(name,ty,t1,t2)->Rel.Declaration.LocalDef(name,ty,t1),t2,true|_->letgoal=Reductionops.whd_allnoletenvsigmaorig_goalinmatchEConstr.kindsigmagoalwith|Prod(name,ty,t)->Rel.Declaration.LocalAssum(name,ty),t,false|LetIn(name,ty,t1,t2)->Rel.Declaration.LocalDef(name,ty,t1),t2,false|App(hd,args)whenEConstr.isLetInsigmahd->(* hack *)let_,v,_,b=EConstr.destLetInsigmahdinletctx,t,_=decompose_assumenvsigma(EConstr.mkApp(EConstr.Vars.subst1vb,args))inctx,t,false|_->CErrors.user_errPp.(str"No assumption in "++Printer.pr_econstr_envenvsigmagoal)lettclFULL_BETAIOTA=Goal.enterbeginfungl->letr,_=Redexpr.reduction_of_red_expr(Goal.envgl)Genredexpr.(Lazy{rBeta=true;rMatch=true;rFix=true;rCofix=true;rZeta=false;rDelta=false;rConst=[]})inTactics.e_reduct_in_concl~check:false(r,Constr.DEFAULTcast)endtypeintro_id=|Anon|IdofId.t|Seedofstring(** [intro id k] introduces the first premise (product or let-in) of the goal
under the name [id], reducing the head of the goal (using beta, iota, delta
but not zeta) if necessary. If [id] is None, a name is generated, that will
not be user accessible. If the goal does not start with a product or a
let-in even after reduction, it fails. In case of success, the original name
and final id are passed to the continuation [k] which gets evaluated. *)lettclINTRO~id~conclusion:k=Goal.enterbeginfungl->letopenContextinletenv,sigma,g=Goal.(envgl,sigmagl,conclgl)inletdecl,t,no_red=decompose_assumenvsigmaginletoriginal_name=Rel.Declaration.get_namedeclinletalready_used=Tacmach.New.pf_ids_of_hypsglinletid=matchid,original_namewith|Idid,_->id|Seedid,_->mk_anon_ididalready_used|Anon,Nameid->ifis_discharged_ididthenidelsemk_anon_id(Id.to_stringid)already_used|Anon,Anonymous->letids=Tacmach.New.pf_ids_of_hypsglinmk_anon_idssr_anon_hypidsinifList.memidalready_usedthenerrorstrmPp.(Id.printid++str" already used");unsafe_introenv(set_decl_ididdecl)t<*>(ifno_redthentclUNIT()elsetclFULL_BETAIOTA)<*>k~orig_name:original_name~new_name:idendletreturn~orig_name:_~new_name:_=tclUNIT()lettclINTRO_IDid=tclINTRO~id:(Idid)~conclusion:returnlettclINTRO_ANON?seed()=matchseedwith|None->tclINTRO~id:Anon~conclusion:return|Someseed->tclINTRO~id:(Seedseed)~conclusion:returnlettclRENAME_HD_PRODname=Goal.enterbeginfungl->letconcl=Goal.conclglinletsigma=Goal.sigmaglinmatchEConstr.kindsigmaconclwith|Prod(x,src,tgt)->convert_concl_no_checkEConstr.(mkProd({xwithbinder_name=name},src,tgt))|_->CErrors.anomaly(Pp.str"rename_hd_prod: no head product")endlettcl0G~defaulttac=numgoals>>=funng->ifng=0thentclUNITdefaultelsetacletrectclFIRSTa=function|[]->Tacticals.New.tclZEROMSGPp.(str"No applicable tactic.")|tac::rest->tclORELSEtac(fun_->tclFIRSTarest)letrectclFIRSTitacn=ifn<0thenTacticals.New.tclZEROMSGPp.(str"tclFIRSTi")elsetclORELSE(tclFIRSTitac(n-1))(fun_->tacn)lettacCONSTR_NAME?namec=matchnamewith|Somen->tclUNITn|None->Goal.enter_one~__LOC__(fung->letsigma=Goal.sigmagintclUNIT(constr_namesigmac))lettacMKPRODc?namecl=tacTYPEOFc>>=funt->tacCONSTR_NAME?namec>>=funname->Goal.enter_one~__LOC__beginfung->letsigma,env=Goal.sigmag,Goal.envginletr=Retyping.relevance_of_termenvsigmacinifname<>Names.Name.Anonymous||EConstr.Vars.noccurnsigma1clthentclUNIT(EConstr.mkProd(make_annotnamer,t,cl))elseletname=Names.Id.of_string(Namegen.hdcharenvsigmat)intclUNIT(EConstr.mkProd(make_annot(Name.Namename)r,t,cl))endlettacINTERP_CPATTERNcp=tacSIGMA>>=beginfungl->tclUNIT(Ssrmatching.interp_cpattern(pf_envgl)(projectgl)cpNone)endlettacUNIFYab=tacSIGMA>>=beginfungl->letgl=Ssrmatching.pf_unify_HOglabinUnsafe.tclEVARS(Tacmach.projectgl)endlettclOPTIONod=matchowith|None->d>>=tclUNIT|Somex->tclUNITxlettacIS_INJECTION_CASE?tyt=begintclOPTIONty(tacTYPEOFt)>>=funty->tacREDUCE_TO_QUANTIFIED_INDty>>=fun((mind,_),_)->tclUNIT(Coqlib.check_ind_ref"core.eq.type"mind)endlettclWITHTOPtac=Goal.enterbeginfungl->lettop=mk_anon_id"top_assumption"(Tacmach.New.pf_ids_of_hypsgl)intclINTRO_IDtop<*>tac(EConstr.mkVartop)<*>Tactics.clear[top]endlettacMK_SSR_CONSTname=Proofview.tclENV>>=funenv->Proofview.tclEVARMAP>>=funsigma->matchmkSsrConstnameenvsigmawith|sigma,c->Unsafe.tclEVARSsigma<*>tclUNITc|exceptionewhenCErrors.noncriticale->tclLIFT(Proofview.NonLogical.raise(e,Exninfo.null))lettacDEST_CONSTc=Proofview.tclEVARMAP>>=funsigma->matchEConstr.destConstsigmacwith|c,_->tclUNITc|exceptionewhenCErrors.noncriticale->tclLIFT(Proofview.NonLogical.raise(e,Exninfo.null))(* TASSI: This version of unprotects inlines the unfold tactic definition,
* since we don't want to wipe out let-ins, and it seems there is no flag
* to change that behaviour in the standard unfold code *)letunprotecttac=tacMK_SSR_CONST"protect_term">>=tacDEST_CONST>>=funprot->Tacticals.New.onClause(funidopt->lethyploc=Option.map(funid->id,InHyp)idoptinTactics.reduct_option~check:false(Reductionops.clos_norm_flags(CClosure.RedFlags.mkflags[CClosure.RedFlags.fBETA;CClosure.RedFlags.fCONSTprot;CClosure.RedFlags.fMATCH;CClosure.RedFlags.fFIX;CClosure.RedFlags.fCOFIX]),DEFAULTcast)hyploc)allHypsAndConclmoduletypeStateType=sigtypestatevalinit:stateendmoduleMakeState(S:StateType)=structletstate_field:S.stateProofview_monad.StateStore.field=Proofview_monad.StateStore.field()(* FIXME: should not inject fresh_state, but initialize it at the beginning *)letlift_upd_stateupds=letopenProofview_monad.StateStoreinletold_state=Option.defaultS.init(getsstate_field)inupdold_state>>=funnew_state->tclUNIT(setsstate_fieldnew_state)lettacUPDATEupd=Goal.enterbeginfungl->lets0=Goal.stateglinGoal.enter_one~__LOC__(fun_->lift_upd_stateupds0)>>=funs->Unsafe.tclGETGOALS>>=fungls->letgls=List.map(fungs->letg=Proofview_monad.drop_stategsinProofview_monad.goal_with_stategs)glsinUnsafe.tclSETGOALSglsendlettclGETk=Goal.enterbeginfungl->letopenProofview_monad.StateStoreink(Option.defaultS.init(get(Goal.stategl)state_field))endlettclGET1k=Goal.enter_onebeginfungl->letopenProofview_monad.StateStoreink(Option.defaultS.init(get(Goal.stategl)state_field))endlettclSETnew_s=letopenProofview_monad.StateStoreinUnsafe.tclGETGOALS>>=fungls->letgls=List.map(fungs->letg=Proofview_monad.drop_stategsinlets=Proofview_monad.get_stategsinProofview_monad.goal_with_stateg(setsstate_fieldnew_s))glsinUnsafe.tclSETGOALSglsletgetg=Option.defaultS.init(Proofview_monad.StateStore.get(Goal.stateg)state_field)endletis_construct_refsigmacr=EConstr.isConstructsigmac&&GlobRef.equal(GlobRef.ConstructRef(fst(EConstr.destConstructsigmac)))rletis_ind_refsigmacr=EConstr.isIndsigmac&&GlobRef.equal(GlobRef.IndRef(fst(EConstr.destIndsigmac)))rletis_const_refsigmacr=EConstr.isConstsigmac&&GlobRef.equal(GlobRef.ConstRef(fst(EConstr.destConstsigmac)))r(* vim: set filetype=ocaml foldmethod=marker: *)