123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691692693694695696697698699700701702703704705706707708709710711712713714715716717718719720721722723724725726727728729730731732733734735736737738739740741742743744745746747748749750751752753754755756757758759760761762763764765766767768769770771772773774775776777778779780781782783784785786787788789790791792793794795796797798799800801802803804805806807808809810811812813814815816817818819820821822823824825826827828829830831832833834835836837838839840841842843844845846847848849850851852853854855856857858859860861862863864865866867868869870871872873874875876877878879880881882883884885886887888889890891892893894895896897898899900901(************************************************************************)(* * 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) *)(************************************************************************)openPpopenCErrorsopenUtilopenNamesopenConstropenEConstropenDeclarationsopenTacmachopenTactypesmoduleRelDecl=Context.Rel.DeclarationmoduleNamedDecl=Context.Named.Declaration(*********************)(* Tacticals *)(*********************)typetactic=Proofview.V82.tacopenEvdexceptionFailErrorofint*Pp.tLazy.tletcatch_failerror(e,info)=matchewith|FailError(lvl,s)whenlvl>0->Exninfo.iraise(FailError(lvl-1,s),info)|e->Control.check_for_interrupt()letunpackageglsig=(ref(glsig.sigma)),glsig.itletrepackagerv={it=v;sigma=!r;}letapply_sig_tacrtacg=Control.check_for_interrupt();(* Breakpoint *)letglsigma=tac(repackagerg)inr:=glsigma.sigma;glsigma.it(* [goal_goal_list : goal sigma -> goal list sigma] *)letgoal_goal_listgls={it=[gls.it];sigma=gls.sigma;}(* identity tactic without any message *)lettclIDTACgls=goal_goal_listgls(* the message printing identity tactic *)lettclIDTAC_MESSAGEsgls=Feedback.msg_info(hov0s);tclIDTACgls(* General failure tactic *)lettclFAIL_ssgls=user_err~hdr:"Refiner.tclFAIL_s"(strs)(* The Fail tactic *)lettclFAILlvlsg=raise(FailError(lvl,lazys))lettclFAIL_lazylvlsg=raise(FailError(lvl,s))letstart_tacgls=letsigr,g=unpackageglsin(sigr,[g])letfinish_tac(sigr,gl)=repackagesigrgl(* Apply [tacfi.(i)] on the first n subgoals, [tacli.(i)] on the last
m subgoals, and [tac] on the others *)letthens3parts_tactacfitactacli(sigr,gs)=letnf=Array.lengthtacfiinletnl=Array.lengthtacliinletng=List.lengthgsinifng<nf+nlthenuser_err~hdr:"Refiner.thensn_tac"(str"Not enough subgoals.");letgll=(List.map_i(funi->apply_sig_tacsigr(ifi<nfthentacfi.(i)elseifi>=ng-nlthentacli.(nl-ng+i)elsetac))0gs)in(sigr,List.flattengll)(* Apply [taci.(i)] on the first n subgoals and [tac] on the others *)letthensf_tactacitac=thens3parts_tactacitac[||](* Apply [tac i] on the ith subgoal (no subgoals number check) *)letthensi_tactac(sigr,gs)=letgll=List.map_i(funi->apply_sig_tacsigr(taci))1gsin(sigr,List.flattengll)letthen_tactac=thensf_tac[||]tac(* [tclTHENS3PARTS tac1 [|t1 ; ... ; tn|] tac2 [|t'1 ; ... ; t'm|] gls]
applies the tactic [tac1] to [gls] then, applies [t1], ..., [tn] to
the first [n] resulting subgoals, [t'1], ..., [t'm] to the last [m]
subgoals and [tac2] to the rest of the subgoals in the middle. Raises an
error if the number of resulting subgoals is strictly less than [n+m] *)lettclTHENS3PARTStac1tacfitactacligls=finish_tac(thens3parts_tactacfitactacli(then_tactac1(start_tacgls)))(* [tclTHENSFIRSTn tac1 [|t1 ; ... ; tn|] tac2 gls] applies the tactic [tac1]
to [gls] and applies [t1], ..., [tn] to the first [n] resulting
subgoals, and [tac2] to the others subgoals. Raises an error if
the number of resulting subgoals is strictly less than [n] *)lettclTHENSFIRSTntac1tacitac=tclTHENS3PARTStac1tacitac[||](* [tclTHENSLASTn tac1 tac2 [|t1 ;...; tn|] gls] applies the tactic [tac1]
to [gls] and applies [t1], ..., [tn] to the last [n] resulting
subgoals, and [tac2] to the other subgoals. Raises an error if the
number of resulting subgoals is strictly less than [n] *)lettclTHENSLASTntac1tactaci=tclTHENS3PARTStac1[||]tactaci(* [tclTHEN_i tac taci gls] applies the tactic [tac] to [gls] and applies
[(taci i)] to the i_th resulting subgoal (starting from 1), whatever the
number of subgoals is *)lettclTHEN_itactacigls=finish_tac(thensi_tactaci(then_tactac(start_tacgls)))(* [tclTHEN tac1 tac2 gls] applies the tactic [tac1] to [gls] and applies
[tac2] to every resulting subgoals *)lettclTHENtac1tac2=tclTHENS3PARTStac1[||]tac2[||](* [tclTHENSV tac1 [t1 ; ... ; tn] gls] applies the tactic [tac1] to
[gls] and applies [t1],..., [tn] to the [n] resulting subgoals. Raises
an error if the number of resulting subgoals is not [n] *)lettclTHENSVtac1tac2v=tclTHENS3PARTStac1tac2v(tclFAIL_s"Wrong number of tactics.")[||]lettclTHENStac1tac2l=tclTHENSVtac1(Array.of_listtac2l)(* [tclTHENLAST tac1 tac2 gls] applies the tactic [tac1] to [gls] and [tac2]
to the last resulting subgoal *)lettclTHENLASTtac1tac2=tclTHENSLASTntac1tclIDTAC[|tac2|](* [tclTHENFIRST tac1 tac2 gls] applies the tactic [tac1] to [gls] and [tac2]
to the first resulting subgoal *)lettclTHENFIRSTtac1tac2=tclTHENSFIRSTntac1[|tac2|]tclIDTAC(* [tclTHENLIST [t1;..;tn]] applies [t1] then [t2] ... then [tn]. More
convenient than [tclTHEN] when [n] is large. *)letrectclTHENLIST=function[]->tclIDTAC|t1::tacl->tclTHENt1(tclTHENLISTtacl)(* [tclMAP f [x1..xn]] builds [(f x1);(f x2);...(f xn)] *)lettclMAPtacfunl=List.fold_right(funx->(tclTHEN(tacfunx)))ltclIDTAC(* PROGRESS tac ptree applies tac to the goal ptree and fails if tac leaves
the goal unchanged *)lettclPROGRESStacptree=letrslt=tacptreeinifGoal.V82.progressrsltptreethenrsltelseuser_err~hdr:"Refiner.PROGRESS"(str"Failed to progress.")(* Execute tac, show the names of new hypothesis names created by tac
in the "as" format and then forget everything. From the logical
point of view [tclSHOWHYPS tac] is therefore equivalent to idtac,
except that it takes the time and memory of tac and prints "as"
information). The resulting (unchanged) goals are printed *after*
the as-expression, which forces pg to some gymnastic.
TODO: Have something similar (better?) in the xml protocol.
NOTE: some tactics delete hypothesis and reuse names (induction,
destruct), this is not detected by this tactical. *)lettclSHOWHYPS(tac:tactic)(goal:Goal.goalEvd.sigma):Goal.goallistEvd.sigma=letoldhyps=pf_hypsgoalinletrslt:Goal.goallistEvd.sigma=tacgoalinlet{it=gls;sigma=sigma;}=rsltinlethyps=List.map(fungl->pf_hyps{it=gl;sigma=sigma;})glsinletcmpd1d2=Names.Id.equal(NamedDecl.get_idd1)(NamedDecl.get_idd2)inletnewhyps=List.map(funhypl->List.subtractcmphyploldhyps)hypsinlets=letfrst=reftrueinList.fold_left(funacclh->acc^(if!frstthen(frst:=false;"")else" | ")^(List.fold_left(funaccd->(Names.Id.to_string(NamedDecl.get_idd))^" "^acc)""lh))""newhypsinFeedback.msg_notice(str"<infoH>"++(hov0(strs))++(str"</infoH>"));tclIDTACgoal;;(* ORELSE0 t1 t2 tries to apply t1 and if it fails, applies t2 *)lettclORELSE0t1t2g=tryt1gwith(* Breakpoint *)|ewhenCErrors.noncriticale->lete=Exninfo.captureeincatch_failerrore;t2g(* ORELSE t1 t2 tries to apply t1 and if it fails or does not progress,
then applies t2 *)lettclORELSEt1t2=tclORELSE0(tclPROGRESSt1)t2(* applies t1;t2then if t1 succeeds or t2else if t1 fails
t2* are called in terminal position (unless t1 produces more than
1 subgoal!) *)lettclORELSE_THENt1t2thent2elsegls=matchtrySome(tclPROGRESSt1gls)withewhenCErrors.noncriticale->lete=Exninfo.captureeincatch_failerrore;Nonewith|None->t2elsegls|Somesgl->letsigr,gl=unpackagesglinfinish_tac(then_tact2then(sigr,gl))(* TRY f tries to apply f, and if it fails, leave the goal unchanged *)lettclTRYf=(tclORELSE0ftclIDTAC)lettclTHENTRYfg=(tclTHENf(tclTRYg))(* Try the first tactic that does not fail in a list of tactics *)letrectclFIRST=function|[]->tclFAIL_s"No applicable tactic."|t::rest->tclORELSE0t(tclFIRSTrest)(* Fails if a tactic did not solve the goal *)lettclCOMPLETEtac=tclTHENtac(tclFAIL_s"Proof is not complete.")(* Iteration tacticals *)lettclDOnt=letrecdoreck=ifk<0thenuser_err~hdr:"Refiner.tclDO"(str"Wrong argument : Do needs a positive integer.");ifInt.equalk0thentclIDTACelseifInt.equalk1thentelse(tclTHENt(dorec(k-1)))indorecn(* Beware: call by need of CAML, g is needed *)letrectclREPEATtg=tclORELSE_THENt(tclREPEATt)tclIDTACglettclAT_LEAST_ONCEt=(tclTHENt(tclREPEATt))(************************************************************************)(* Tacticals applying on hypotheses *)(************************************************************************)letnthDeclmgl=tryList.nth(pf_hypsgl)(m-1)withFailure_->user_errPp.(str"No such assumption.")letnthHypIdmgl=nthDeclmgl|>NamedDecl.get_idletnthHypmgl=mkVar(nthHypIdmgl)letlastDeclgl=nthDecl1glletlastHypIdgl=nthHypId1glletlastHypgl=nthHyp1glletnLastDeclsngl=tryList.firstnn(pf_hypsgl)withFailure_->user_errPp.(str"Not enough hypotheses in the goal.")letnLastHypsIdngl=List.map(NamedDecl.get_id)(nLastDeclsngl)letnLastHypsngl=List.mapmkVar(nLastHypsIdngl)letonNthDeclmtacgl=tac(nthDeclmgl)glletonNthHypIdmtacgl=tac(nthHypIdmgl)glletonNthHypmtacgl=tac(nthHypmgl)glletonLastDecl=onNthDecl1letonLastHypId=onNthHypId1letonLastHyp=onNthHyp1letonHypsfindtacgl=tac(findgl)glletonNLastDeclsntac=onHyps(nLastDeclsn)tacletonNLastHypsIdntac=onHyps(nLastHypsIdn)tacletonNLastHypsntac=onHyps(nLastHypsn)tacletafterHypidgl=fst(List.split_when(NamedDecl.get_id%>Id.equalid)(pf_hypsgl))(***************************************)(* Clause Tacticals *)(***************************************)(* The following functions introduce several tactic combinators and
functions useful for working with clauses. A clause is either None
or (Some id), where id is an identifier. This type is useful for
defining tactics that may be used either to transform the
conclusion (None) or to transform a hypothesis id (Some id). --
--Eduardo (8/8/97)
*)letfullGoalgl=None::List.mapOption.make(pf_ids_of_hypsgl)letonAllHypstacgl=tclMAPtac(pf_ids_of_hypsgl)glletonAllHypsAndConcltacgl=tclMAPtac(fullGoalgl)glletonClausetacclgls=lethyps()=pf_ids_of_hypsglsintclMAPtac(Locusops.simple_clause_ofhypscl)glsletonClauseLRtacclgls=lethyps()=pf_ids_of_hypsglsintclMAPtac(List.rev(Locusops.simple_clause_ofhypscl))glsletifOnHyppredtac1tac2idgl=ifpred(id,pf_get_hyp_typglid)thentac1idglelsetac2idgl(************************************************************************)(* Elimination Tacticals *)(************************************************************************)(* The following tacticals allow to apply a tactic to the
branches generated by the application of an elimination
tactic.
Two auxiliary types --branch_args and branch_assumptions-- are
used to keep track of some information about the ``branches'' of
the elimination. *)letfix_empty_or_and_patternnvl=(* 1- The syntax does not distinguish between "[ ]" for one clause with no
names and "[ ]" for no clause at all *)(* 2- More generally, we admit "[ ]" for any disjunctive pattern of
arbitrary length *)matchlwith|IntroOrPattern[[]]->IntroOrPattern(List.makenv[])|_->lletcheck_or_and_pattern_size?loccheck_andnamesbranchsigns=letn=Array.lengthbranchsignsinletmsgp1p2=strbrk"a conjunctive pattern made of "++intp1++(ifp1==p2thenmt()elsestr" or "++intp2)++str" patterns"inleterr1p1p2=user_err?loc(str"Expects "++msgp1p2++str".")inleterrnn=user_err?loc(str"Expects a disjunctive pattern with "++intn++str" branches.")inleterr1'p1p2=user_err?loc(strbrk"Expects a disjunctive pattern with 1 branch or "++msgp1p2++str".")inleterrforthcoming?loc=user_err?loc(strbrk"Unexpected non atomic pattern.")inmatchnameswith|IntroAndPatternl->ifnot(Int.equaln1)thenerrnn;letl'=List.filterCAst.(function{v=IntroForthcoming_}->true|{v=IntroNaming_}|{v=IntroAction_}->false)linifl'!=[]thenerrforthcoming?loc:(List.hdl').CAst.loc;ifcheck_andthenletp1=List.count(funx->x)branchsigns.(0)inletp2=List.lengthbranchsigns.(0)inletp=List.lengthlinifnot(Int.equalpp1||Int.equalpp2)thenerr1p1p2;ifInt.equalpp1thenIntroAndPattern(List.extendbranchsigns.(0)(CAst.make@@IntroNamingNamegen.IntroAnonymous)l)elsenameselsenames|IntroOrPatternll->ifnot(Int.equaln(List.lengthll))thenifInt.equaln1thenletp1=List.count(funx->x)branchsigns.(0)inletp2=List.lengthbranchsigns.(0)inerr1'p1p2elseerrnn;namesletget_and_check_or_and_pattern_gen?loccheck_andnamesbranchsigns=letnames=check_or_and_pattern_size?loccheck_andnamesbranchsignsinmatchnameswith|IntroAndPatternl->[|l|]|IntroOrPatternl->Array.of_listlletget_and_check_or_and_pattern?loc=get_and_check_or_and_pattern_gen?loctrueletcompute_induction_namescheck_andbranchletsigns=function|None->Array.make(Array.lengthbranchletsigns)[]|Some{CAst.loc;v=names}->letnames=fix_empty_or_and_pattern(Array.lengthbranchletsigns)namesinget_and_check_or_and_pattern_gencheck_and?locnamesbranchletsigns(* Compute the let-in signature of case analysis or standard induction scheme *)letcompute_constructor_signatures~rec_flag((_,kasity),u)=letrecanalreccrecargs=matchc,recargswith|RelDecl.LocalAssum_::c,recarg::rest->letrest=analreccrestinbeginmatchDeclareops.dest_recargrecargwith|Norec|Nested_->true::rest|Mrec(_,j)->ifrec_flag&&Int.equaljkthentrue::true::restelsetrue::restend|RelDecl.LocalDef_::c,rest->false::analreccrest|[],[]->[]|_->anomaly(Pp.str"compute_constructor_signatures.")inlet(mib,mip)=Global.lookup_inductiveityinletmap(ctx,_)=List.skipn(Context.Rel.lengthmib.mind_params_ctxt)(List.revctx)inletlc=Array.mapmapmip.mind_nf_lcinletlrecargs=Declareops.dest_subtermsmip.mind_recargsinArray.map2analreclclrecargsletelimination_sort_of_goalgl=pf_applyRetyping.get_sort_family_ofgl(pf_conclgl)letelimination_sort_of_hypidgl=pf_applyRetyping.get_sort_family_ofgl(pf_get_hyp_typglid)letelimination_sort_of_clause=function|None->elimination_sort_of_goal|Someid->elimination_sort_of_hypid(* Change evars *)lettclEVARSsigmagls=tclIDTAC{glswithEvd.sigma=sigma}letpf_with_evarsglsevkgls=letevd,a=glsevglsintclTHEN(tclEVARSevd)(ka)glsletpf_constr_of_globalgrk=pf_with_evars(fungls->pf_applyEvd.fresh_globalglsgr)k(** Tacticals of Ltac defined directly in term of Proofview *)moduleNew=structopenProofviewopenProofview.NotationsopenTacmach.NewlettclIDTAC=tclUNIT()lettclTHENt1t2=t1<*>t2lettclFAIL?infolvlmsg=letinfo=matchinfowith(* If the backtrace points here it means the caller didn't save
the backtrace correctly *)|None->Exninfo.reify()|Someinfo->infointclZERO~info(FailError(lvl,lazymsg))lettclZEROMSG?info?locmsg=letinfo=matchinfowith(* If the backtrace points here it means the caller didn't save
the backtrace correctly *)|None->Exninfo.reify()|Someinfo->infoinletinfo=matchlocwith|None->info|Someloc->Loc.add_locinfolocinleterr=UserError(None,msg)intclZERO~infoerrletcatch_failerrore=trycatch_failerrore;tclUNIT()withewhenCErrors.noncriticale->let_,info=Exninfo.captureeintclZERO~infoe(* spiwack: I chose to give the Ltac + the same semantics as
[Proofview.tclOR], however, for consistency with the or-else
tactical, we may consider wrapping the first argument with
[tclPROGRESS]. It strikes me as a bad idea, but consistency can be
considered valuable. *)lettclORt1t2=tclINDEPENDENTbeginProofview.tclORt1beginfune->catch_failerrore<*>t2endendlettclORDt1t2=tclINDEPENDENTbeginProofview.tclORt1beginfune->catch_failerrore<*>t2()endendlettclONCE=Proofview.tclONCElettclEXACTLY_ONCEt=Proofview.tclEXACTLY_ONCE(FailError(0,lazy(assertfalse)))tlettclIFCATCHtttte=tclINDEPENDENTbeginProofview.tclIFCATCHttt(fune->catch_failerrore<*>te())endlettclORELSE0t1t2=tclINDEPENDENTbegintclORELSEt1beginfune->catch_failerrore<*>t2endendlettclORELSE0Lt1t2=tclINDEPENDENTLbegintclORELSEt1beginfune->catch_failerrore<*>t2endendlettclORELSEt1t2=tclORELSE0(tclPROGRESSt1)t2lettclTHENS3PARTSt1l1repeatl2=tclINDEPENDENTbegint1<*>Proofview.tclORELSE(* converts the [SizeMismatch] error into an ltac error *)begintclEXTEND(Array.to_listl1)repeat(Array.to_listl2)endbeginfunction(e,info)->matchewith|SizeMismatch(i,_)->leterrmsg=str"Incorrect number of goals"++spc()++str"(expected "++inti++str(String.plurali" tactic")++str")"intclFAIL0errmsg|reraise->tclZERO~inforeraiseendendlettclTHENSFIRSTnt1lrepeat=tclTHENS3PARTSt1lrepeat[||]lettclTHENFIRSTnt1l=tclTHENSFIRSTnt1l(tclUNIT())lettclTHENFIRSTt1t2=tclTHENFIRSTnt1[|t2|]lettclBINDFIRSTt1t2=t1>>=funans->Proofview.Unsafe.tclGETGOALS>>=fungls->matchglswith|[]->tclFAIL0(str"Expect at least one goal.")|hd::tl->Proofview.Unsafe.tclSETGOALS[hd]<*>t2ans>>=funans->Proofview.Unsafe.tclNEWGOALStl<*>Proofview.tclUNITanslettclTHENSLASTnt1repeatl=tclTHENS3PARTSt1[||]repeatllettclTHENLASTnt1l=tclTHENS3PARTSt1[||](tclUNIT())llettclTHENLASTt1t2=tclTHENLASTnt1[|t2|]letoption_of_failurefx=trySome(fx)withFailure_->NonelettclBINDLASTt1t2=t1>>=funans->Proofview.Unsafe.tclGETGOALS>>=fungls->matchoption_of_failureList.sep_lastglswith|None->tclFAIL0(str"Expect at least one goal.")|Some(last,firstn)->Proofview.Unsafe.tclSETGOALS[last]<*>t2ans>>=funans->Proofview.Unsafe.tclGETGOALS>>=funnewgls->tclEVARMAP>>=funsigma->letfirstn=Proofview.Unsafe.undefinedsigmafirstninProofview.Unsafe.tclSETGOALS(firstn@newgls)<*>Proofview.tclUNITanslettclTHENStl=tclINDEPENDENTbegint<*>Proofview.tclORELSE(* converts the [SizeMismatch] error into an ltac error *)begintclDISPATCHlendbeginfunction(e,info)->matchewith|SizeMismatch(i,_)->leterrmsg=str"Incorrect number of goals"++spc()++str"(expected "++inti++str(String.plurali" tactic")++str")"intclFAIL0errmsg|reraise->tclZERO~inforeraiseendendlettclTHENLISTl=List.fold_lefttclTHEN(tclUNIT())l(* [tclMAP f [x1..xn]] builds [(f x1);(f x2);...(f xn)] *)lettclMAPtacfunl=List.fold_right(funx->(tclTHEN(tacfunx)))l(tclUNIT())lettclTRYt=tclORELSE0t(tclUNIT())lettclTRYbt=tclORELSE0L(t<*>tclUNITtrue)(tclUNITfalse)lettclIFTHENELSEt1t2t3=tclINDEPENDENTbeginProofview.tclIFCATCHt1(fun()->t2)(fun(e,info)->Proofview.tclORELSEt3(fune'->tclZERO~infoe))endlettclIFTHENSVELSEt1at3=Proofview.tclIFCATCHt1(fun()->tclDISPATCH(Array.to_lista))(fun_->t3)lettclIFTHENFIRSTELSEt1t2t3=Proofview.tclIFCATCHt1(fun()->tclEXTEND[t2](tclUNIT())[])(fun_->t3)lettclIFTHENTRYELSEMUSTt1t2=tclIFTHENELSEt1(tclTRYt2)t2lettclIFTHENFIRSTTRYELSEMUSTt1t2=tclIFTHENFIRSTELSEt1(tclTRYt2)t2(* Try the first tactic that does not fail in a list of tactics *)letrectclFIRST=function|[]->letinfo=Exninfo.reify()intclZEROMSG~info(str"No applicable tactic.")|t::rest->tclORELSE0t(tclFIRSTrest)letrectclFIRST_PROGRESS_ONtac=function|[]->tclFAIL0(str"No applicable tactic")|[a]->taca(* so that returned failure is the one from last item *)|a::tl->tclORELSE(taca)(tclFIRST_PROGRESS_ONtactl)letrectclDOnt=ifn<0thenletinfo=Exninfo.reify()intclZEROMSG~info(str"Wrong argument : Do needs a positive integer.")elseifn=0thentclUNIT()elseifn=1thentelsetclTHENt(tclDO(n-1)t)letrectclREPEAT0t=tclINDEPENDENTbeginProofview.tclIFCATCHt(fun()->tclCHECKINTERRUPT<*>tclREPEAT0t)(fune->catch_failerrore<*>tclUNIT())endlettclREPEATt=tclREPEAT0(tclPROGRESSt)letrectclREPEAT_MAIN0t=Proofview.tclIFCATCHt(fun()->tclTRYFOCUS11(tclREPEAT_MAIN0t))(fune->catch_failerrore<*>tclUNIT())lettclREPEAT_MAINt=tclREPEAT_MAIN0(tclPROGRESSt)lettclCOMPLETEt=t>>=funres->(tclINDEPENDENT(letinfo=Exninfo.reify()intclZEROMSG~info(str"Proof is not complete.")))<*>tclUNITres(* Try the first that solves the current goal *)lettclSOLVEtacl=tclFIRST(List.maptclCOMPLETEtacl)lettclPROGRESSt=Proofview.tclINDEPENDENT(Proofview.tclPROGRESSt)(* Check that holes in arguments have been resolved *)letcheck_evarsenvsigmaextsigmaorigsigma=letreachable=lazy(Evarutil.reachable_from_evarssigma(Evar.Map.domain(Evd.undefined_maporigsigma)))inletrecis_undefined_up_to_restrictionsigmaevk=ifEvd.memorigsigmaevkthenNoneelseletevi=Evd.findsigmaevkinmatchEvd.evar_bodyeviwith|Evd.Evar_empty->Some(evk,evi)|Evd.Evar_definedc->matchConstr.kind(EConstr.Unsafe.to_constrc)with|Evar(evk,l)->is_undefined_up_to_restrictionsigmaevk|_->(* We make the assumption that there is no way to refine an
evar remaining after typing from the initial term given to
apply/elim and co tactics, is it correct? *)Noneinletrest=Evd.fold_undefined(funevkeviacc->matchis_undefined_up_to_restrictionsigmaevkwith|Some(evk',evi)->(* If [evk'] descends from [evk] which descends itself from
an originally undefined evar in [origsigma], it is a not
a fresh undefined hole from [sigma]. *)ifEvar.Set.memevk(Lazy.forcereachable)thenaccelse(evk',evi)::acc|_->acc)extsigma[]inmatchrestwith|[]->()|(evk,evi)::_->let(loc,_)=evi.Evd.evar_sourceinPretype_errors.error_unsolvable_implicit?locenvsigmaevkNonelettclMAPDELAYEDWITHHOLESaccept_unresolved_holesltac=letrecaux=function|[]->tclUNIT()|x::l->Proofview.Goal.enterbeginfungl->letenv=Proofview.Goal.envglinletsigma_initial=Proofview.Goal.sigmaglinlet(sigma,x)=xenvsigma_initialinProofview.Unsafe.tclEVARSsigma<*>tacx>>=fun()->auxl>>=fun()->ifaccept_unresolved_holesthentclUNIT()elsetclEVARMAP>>=funsigma_final->trylet()=check_evarsenvsigma_finalsigmasigma_initialintclUNIT()withewhenCErrors.noncriticale->lete,info=Exninfo.captureeintclZERO~infoeendinauxl(* The following is basically
tclMAPDELAYEDWITHHOLES accept_unresolved_holes [fun _ _ -> (sigma,())] (fun () -> tac)
but with value not necessarily in unit *)lettclWITHHOLESaccept_unresolved_holestacsigma=tclEVARMAP>>=funsigma_initial->ifsigma==sigma_initialthentacelseletcheck_evars_ifx=ifnotaccept_unresolved_holesthentclEVARMAP>>=funsigma_final->tclENV>>=funenv->trylet()=check_evarsenvsigma_finalsigmasigma_initialintclUNITxwithewhenCErrors.noncriticale->lete,info=Exninfo.captureeintclZERO~infoeelsetclUNITxinProofview.Unsafe.tclEVARSsigma<*>tac>>=check_evars_iflettclDELAYEDWITHHOLEScheckxtac=Proofview.Goal.enterbeginfungl->letenv=Proofview.Goal.envglinletsigma=Proofview.Goal.sigmaglinlet(sigma,x)=xenvsigmaintclWITHHOLEScheck(tacx)sigmaendlettclTIMEOUTnt=Proofview.tclOR(Proofview.tclTIMEOUTnt)beginfunction(e,info)->matchewith|Logic_monad.Tac_Timeoutase->letinfo=Exninfo.reify()inProofview.tclZERO~info(FailError(0,lazy(CErrors.printe)))|e->Proofview.tclZERO~infoeendlettclTIMEst=Proofview.tclTIMEstletnthDeclmgl=lethyps=Proofview.Goal.hypsglintryList.nthhyps(m-1)withFailure_->CErrors.user_errPp.(str"No such assumption.")letnLastDeclsgln=tryList.firstnn(Proofview.Goal.hypsgl)withFailure_->CErrors.user_errPp.(str"Not enough hypotheses in the goal.")letnthHypIdmgl=(* We only use [id] *)nthDeclmgl|>NamedDecl.get_idletnthHypmgl=mkVar(nthHypIdmgl)letonNthHypIdmtac=Proofview.Goal.enterbeginfungl->tac(nthHypIdmgl)endletonNthHypmtac=Proofview.Goal.enterbeginfungl->tac(nthHypmgl)endletonLastHypId=onNthHypId1letonLastHyp=onNthHyp1letonNthDeclmtac=Proofview.Goal.enterbeginfungl->Proofview.tclUNIT(nthDeclmgl)>>=tacendletonLastDecl=onNthDecl1letnLastHypsIdgln=List.map(NamedDecl.get_id)(nLastDeclsgln)letnLastHypsgln=List.mapmkVar(nLastHypsIdgln)letifOnHyppredtac1tac2id=Proofview.Goal.enterbeginfungl->lettyp=Tacmach.New.pf_get_hyp_typidglinifpf_applypredgl(id,typ)thentac1idelsetac2idendletonHypsfindtac=Proofview.Goal.enterbeginfungl->tac(findgl)endletonNLastDeclsntac=onHyps(fungl->nLastDeclsgln)tacletonNLastHypsIdntac=onHyps(fungl->nLastHypsIdgln)tacletonNLastHypsntac=onHyps(fungl->nLastHypsgln)tacletafterHypidtac=Proofview.Goal.enterbeginfungl->lethyps=Proofview.Goal.hypsglinletrem,_=List.split_when(NamedDecl.get_id%>Id.equalid)hypsintacremendletfullGoalgl=lethyps=Tacmach.New.pf_ids_of_hypsglinNone::List.mapOption.makehypslettryAllHypstac=Proofview.Goal.enterbeginfungl->lethyps=Tacmach.New.pf_ids_of_hypsglintclFIRST_PROGRESS_ONtachypsendlettryAllHypsAndConcltac=Proofview.Goal.enterbeginfungl->tclFIRST_PROGRESS_ONtac(fullGoalgl)endletonClausetaccl=Proofview.Goal.enterbeginfungl->lethyps=Tacmach.New.pf_ids_of_hypsglintclMAPtac(Locusops.simple_clause_of(fun()->hyps)cl)endletfullGoalgl=None::List.mapOption.make(Tacmach.New.pf_ids_of_hypsgl)letonAllHypstac=Proofview.Goal.enterbeginfungl->tclMAPtac(Tacmach.New.pf_ids_of_hypsgl)endletonAllHypsAndConcltac=Proofview.Goal.enterbeginfungl->tclMAPtac(fullGoalgl)endletelimination_sort_of_goalgl=(* Retyping will expand evars anyway. *)letc=Proofview.Goal.conclglinpf_applyRetyping.get_sort_family_ofglcletelimination_sort_of_hypidgl=(* Retyping will expand evars anyway. *)letc=pf_get_hyp_typidglinpf_applyRetyping.get_sort_family_ofglcletelimination_sort_of_clauseidgl=matchidwith|None->elimination_sort_of_goalgl|Someid->elimination_sort_of_hypidglletpf_constr_of_globalref=Proofview.tclEVARMAP>>=funsigma->Proofview.tclENV>>=funenv->let(sigma,c)=Evd.fresh_globalenvsigmarefinProofview.Unsafe.tclEVARSsigma<*>Proofview.tclUNITclettclTYPEOFTHEN?refreshctac=Proofview.Goal.enter(fungl->letenv=Proofview.Goal.envglinletsigma=Proofview.Goal.sigmaglinlet(sigma,t)=Typing.type_of?refreshenvsigmacinProofview.Unsafe.tclEVARSsigma<*>tacsigmat)lettclSELECT=Goal_select.tclSELECTend