123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691692693694695696697698699700701702703704705706707708709710711712713714715716717718719720721722723724725726727728729730731732733734735736737738739740741742743744745746747748749750751752753754755756757758759760761762763764765766767768769770771772773774775776777778779780781782783784785786787788789790791792793794795796797798799800801802803804805806807808809810811812813814815816817818819820821822823824825826827828829830831832833834835836837838839840841842843844845846847848849850851852853854855856857858859860861862863864865866867868869870871872873874875876877878879880881882883884885886887888889890891892893894895896897898899900901902903904905906907908909910911912913914915916917918919920921922923924925926927928929930931932933934935936937938939940941942943944945946947948949950951952953954955956957958959960961962963964965966967968969970971972973974975976977978979980981982983984985986987988989990991992993994995996997998999100010011002100310041005100610071008100910101011101210131014101510161017101810191020102110221023102410251026102710281029103010311032103310341035103610371038103910401041104210431044104510461047104810491050105110521053105410551056105710581059106010611062106310641065106610671068106910701071107210731074107510761077107810791080108110821083108410851086108710881089109010911092109310941095109610971098109911001101110211031104110511061107110811091110111111121113111411151116111711181119112011211122112311241125112611271128112911301131113211331134113511361137113811391140114111421143114411451146114711481149115011511152115311541155115611571158115911601161116211631164116511661167116811691170117111721173117411751176117711781179118011811182118311841185118611871188118911901191119211931194119511961197119811991200120112021203120412051206120712081209121012111212121312141215121612171218121912201221122212231224122512261227122812291230123112321233123412351236123712381239124012411242124312441245124612471248124912501251125212531254125512561257125812591260126112621263126412651266126712681269127012711272127312741275127612771278127912801281128212831284128512861287128812891290129112921293129412951296129712981299130013011302130313041305130613071308130913101311131213131314131513161317131813191320132113221323132413251326132713281329133013311332133313341335133613371338133913401341134213431344134513461347134813491350135113521353135413551356135713581359136013611362136313641365136613671368136913701371137213731374137513761377137813791380138113821383138413851386138713881389139013911392139313941395139613971398139914001401140214031404140514061407140814091410141114121413141414151416141714181419142014211422142314241425142614271428142914301431143214331434143514361437143814391440144114421443144414451446144714481449145014511452145314541455145614571458145914601461146214631464146514661467146814691470147114721473147414751476147714781479148014811482148314841485148614871488148914901491149214931494149514961497149814991500150115021503150415051506150715081509151015111512151315141515151615171518151915201521152215231524152515261527152815291530153115321533153415351536153715381539154015411542154315441545154615471548154915501551155215531554155515561557155815591560156115621563156415651566156715681569157015711572157315741575157615771578157915801581158215831584158515861587158815891590159115921593159415951596159715981599160016011602160316041605160616071608160916101611161216131614161516161617161816191620162116221623162416251626162716281629163016311632163316341635163616371638163916401641164216431644164516461647164816491650165116521653165416551656165716581659166016611662166316641665166616671668166916701671167216731674167516761677167816791680168116821683168416851686168716881689169016911692169316941695169616971698169917001701170217031704170517061707170817091710171117121713171417151716171717181719172017211722172317241725172617271728172917301731173217331734173517361737173817391740174117421743174417451746174717481749175017511752175317541755175617571758175917601761176217631764176517661767176817691770177117721773177417751776177717781779178017811782178317841785178617871788178917901791179217931794179517961797179817991800180118021803180418051806180718081809181018111812181318141815181618171818181918201821182218231824182518261827182818291830183118321833183418351836183718381839184018411842184318441845184618471848184918501851185218531854185518561857185818591860186118621863186418651866186718681869187018711872187318741875187618771878187918801881188218831884188518861887188818891890189118921893189418951896189718981899190019011902190319041905190619071908190919101911191219131914191519161917191819191920192119221923192419251926192719281929193019311932193319341935193619371938193919401941194219431944194519461947(************************************************************************)(* * 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) *)(************************************************************************)moduleCVars=VarsopenPpopenCErrorsopenUtilopenNamesopenNameopsopenTermopenConstropenContextopenTermopsopenEConstropenVarsopenNamegenopenInductiveopenInductiveopsopenLibnamesopenGlobnamesopenReductionopsopenTypingopenRetypingopenTacmach.NewopenLogicopenHipatternopenTacticals.NewopenTacticsopenTacredopenCoqlibopenDeclarationsopenIndrecopenClenvopenInd_tablesopenEqschemesopenLocusopenLocusopsopenTactypesopenProofview.NotationsopenUnificationopenContext.Named.DeclarationmoduleNamedDecl=Context.Named.Declaration(* Options *)typeinj_flags={keep_proof_equalities:bool;injection_in_context:bool;injection_pattern_l2r_order:bool;}openGoptionsletuse_injection_pattern_l2r_order=function|None->true|Someflags->flags.injection_pattern_l2r_orderletinjection_in_context=reffalseletuse_injection_in_context=function|None->!injection_in_context|Someflags->flags.injection_in_contextlet()=declare_bool_option{optdepr=false;optkey=["Structural";"Injection"];optread=(fun()->!injection_in_context);optwrite=(funb->injection_in_context:=b)}(* Rewriting tactics *)typedep_proof_flag=bool(* true = support rewriting dependent proofs *)typefreeze_evars_flag=bool(* true = don't instantiate existing evars *)typeorientation=booltypeconditions=|Naive(* Only try the first occurrence of the lemma (default) *)|FirstSolved(* Use the first match whose side-conditions are solved *)|AllMatches(* Rewrite all matches whose side-conditions are solved *)(* Warning : rewriting from left to right only works
if there exists in the context a theorem named <eqname>_<suffsort>_r
with type (A:<sort>)(x:A)(P:A->Prop)(P x)->(y:A)(eqname A y x)->(P y).
If another equality myeq is introduced, then corresponding theorems
myeq_ind_r, myeq_rec_r and myeq_rect_r have to be proven. See below.
-- Eduardo (19/8/97)
*)letrewrite_core_unif_flags={modulo_conv_on_closed_terms=None;use_metas_eagerly_in_conv_on_closed_terms=true;use_evars_eagerly_in_conv_on_closed_terms=false;modulo_delta=TransparentState.empty;modulo_delta_types=TransparentState.empty;check_applied_meta_types=true;use_pattern_unification=true;use_meta_bound_pattern_unification=true;allowed_evars=Evarsolve.AllowedEvars.all;restrict_conv_on_strict_subterms=false;modulo_betaiota=false;modulo_eta=true;}letrewrite_unif_flags={core_unify_flags=rewrite_core_unif_flags;merge_unify_flags=rewrite_core_unif_flags;subterm_unify_flags=rewrite_core_unif_flags;allow_K_in_toplevel_higher_order_unification=false;(* allow_K does not matter in practice because calls w_typed_unify *)resolve_evars=true}letfreeze_initial_evarssigmaflagsclause=(* We take evars of the type: this may include old evars! For excluding *)(* all old evars, including the ones occurring in the rewriting lemma, *)(* we would have to take the clenv_value *)letnewevars=lazy(Evarutil.undefined_evars_of_termsigma(clenv_typeclause))inletinitial=Evd.undefined_mapsigmainletallowedevk=ifEvar.Map.memevkinitialthenfalseelseEvar.Set.memevk(Lazy.forcenewevars)inletallowed_evars=Evarsolve.AllowedEvars.from_predallowedin{flagswithcore_unify_flags={flags.core_unify_flagswithallowed_evars};merge_unify_flags={flags.merge_unify_flagswithallowed_evars};subterm_unify_flags={flags.subterm_unify_flagswithallowed_evars}}letmake_flagsfrzevarssigmaflagsclause=iffrzevarsthenfreeze_initial_evarssigmaflagsclauseelseflagsletside_tactacsidetac=matchsidetacwith|None->tac|Somesidetac->tclTHENSFIRSTntac[|Proofview.tclUNIT()|]sidetacletinstantiate_lemma_allfrzevarsglctyll2rconcl=letenv=Proofview.Goal.envglinletsigma=projectglinleteqclause=pf_applyClenv.make_clenv_bindinggl(c,ty)linlet(equiv,args)=decompose_app_vectsigma(Clenv.clenv_typeeqclause)inletarglen=Array.lengthargsinlet()=ifarglen<2thenuser_errPp.(str"The term provided is not an applied relation.")inletc1=args.(arglen-2)inletc2=args.(arglen-1)inlettry_occ(evd',c')=letclenv=Clenv.update_clenv_evdeqclauseevd'inClenv.clenv_pose_dependent_evars~with_evars:trueclenvinletflags=make_flagsfrzevars(Tacmach.New.projectgl)rewrite_unif_flagseqclauseinletoccs=w_unify_to_subterm_all~flagsenveqclause.evd((ifl2rthenc1elsec2),concl)inList.maptry_occoccsletinstantiate_lemmaglcctll2rconcl=letsigma=Proofview.Goal.sigmaglinlett=trysnd(reduce_to_quantified_ind(pf_envgl)sigmact)withUserError_->ctinleteqclause=Clenv.make_clenv_binding(pf_envgl)sigma(c,t)lineqclauseletrewrite_conv_closed_core_unif_flags={modulo_conv_on_closed_terms=SomeTransparentState.full;(* We have this flag for historical reasons, it has e.g. the consequence *)(* to rewrite "?x+2" in "y+(1+1)=0" or to rewrite "?x+?x" in "2+(1+1)=0" *)use_metas_eagerly_in_conv_on_closed_terms=true;use_evars_eagerly_in_conv_on_closed_terms=false;(* Combined with modulo_conv_on_closed_terms, this flag allows since 8.2 *)(* to rewrite e.g. "?x+(2+?x)" in "1+(1+2)=0" *)modulo_delta=TransparentState.empty;modulo_delta_types=TransparentState.full;check_applied_meta_types=true;use_pattern_unification=true;(* To rewrite "?n x y" in "y+x=0" when ?n is *)(* a preexisting evar of the goal*)use_meta_bound_pattern_unification=true;allowed_evars=Evarsolve.AllowedEvars.all;restrict_conv_on_strict_subterms=false;modulo_betaiota=false;modulo_eta=true;}letrewrite_conv_closed_unif_flags={core_unify_flags=rewrite_conv_closed_core_unif_flags;merge_unify_flags=rewrite_conv_closed_core_unif_flags;subterm_unify_flags=rewrite_conv_closed_core_unif_flags;allow_K_in_toplevel_higher_order_unification=false;resolve_evars=false}letrewrite_keyed_core_unif_flags={modulo_conv_on_closed_terms=SomeTransparentState.full;(* We have this flag for historical reasons, it has e.g. the consequence *)(* to rewrite "?x+2" in "y+(1+1)=0" or to rewrite "?x+?x" in "2+(1+1)=0" *)use_metas_eagerly_in_conv_on_closed_terms=true;use_evars_eagerly_in_conv_on_closed_terms=false;(* Combined with modulo_conv_on_closed_terms, this flag allows since 8.2 *)(* to rewrite e.g. "?x+(2+?x)" in "1+(1+2)=0" *)modulo_delta=TransparentState.full;modulo_delta_types=TransparentState.full;check_applied_meta_types=true;use_pattern_unification=true;(* To rewrite "?n x y" in "y+x=0" when ?n is *)(* a preexisting evar of the goal*)use_meta_bound_pattern_unification=true;allowed_evars=Evarsolve.AllowedEvars.all;restrict_conv_on_strict_subterms=false;modulo_betaiota=true;modulo_eta=true;}letrewrite_keyed_unif_flags={core_unify_flags=rewrite_keyed_core_unif_flags;merge_unify_flags=rewrite_keyed_core_unif_flags;subterm_unify_flags=rewrite_keyed_core_unif_flags;allow_K_in_toplevel_higher_order_unification=false;resolve_evars=false}letrewrite_elimwith_evarsfrzevarsclsce=Proofview.Goal.enterbeginfungl->letflags=ifUnification.is_keyed_unification()thenrewrite_keyed_unif_flagselserewrite_conv_closed_unif_flagsinletflags=make_flagsfrzevars(Tacmach.New.projectgl)flagscingeneral_elim_clausewith_evarsflagsclsceendlettclNOTSAMEGOALtac=letgoalgl=Proofview.Goal.goalglinProofview.Goal.enterbeginfungl->letsigma=projectglinletev=goalglintac>>=fun()->Proofview.Goal.goals>>=fungls->letcheckaccugl'=gl'>>=fungl'->letaccu=accu||Proofview.Progress.goal_equal~evd:sigma~extended_evd:(projectgl')ev(goalgl')inProofview.tclUNITaccuinProofview.Monad.List.fold_leftcheckfalsegls>>=funhas_same->ifhas_samethentclZEROMSG(str"Tactic generated a subgoal identical to the original goal.")elseProofview.tclUNIT()end(* Ad hoc asymmetric general_elim_clause *)letgeneral_elim_clausewith_evarsfrzevarsclsrewelim=letopenPretype_errorsinProofview.tclORELSEbeginmatchclswith|None->(* was tclWEAK_PROGRESS which only fails for tactics generating one
subgoal and did not fail for useless conditional rewritings generating
an extra condition *)tclNOTSAMEGOAL(rewrite_elimwith_evarsfrzevarsclsrewelim)|Some_->rewrite_elimwith_evarsfrzevarsclsrewelimendbeginfunction(e,info)->matchewith|PretypeError(env,evd,NoOccurrenceFound(c',_))->Proofview.tclZERO~info(PretypeError(env,evd,NoOccurrenceFound(c',cls)))|e->Proofview.tclZERO~infoeendletgeneral_elim_clausewith_evarsfrzevarstacclsctll2relim=letstrat,tac=matchtacwith|None->Naive,None|Some(tac,Naive)->Naive,Sometac|Some(tac,FirstSolved)->FirstSolved,Some(tclCOMPLETEtac)|Some(tac,AllMatches)->AllMatches,Some(tclCOMPLETEtac)inlettry_clausec=side_tac(tclTHEN(Proofview.Unsafe.tclEVARSc.evd)(general_elim_clausewith_evarsfrzevarsclscelim))tacinProofview.Goal.enterbeginfungl->lettyp=matchclswith|None->pf_conclgl|Someid->pf_get_hyp_typidglinmatchstratwith|Naive->letcs=instantiate_lemmaglctll2rtypintry_clausecs|FirstSolved->letcs=instantiate_lemma_allfrzevarsglctll2rtypintclFIRST(List.maptry_clausecs)|AllMatches->letcs=instantiate_lemma_allfrzevarsglctll2rtypintclMAPtry_clausecsend(* The next function decides in particular whether to try a regular
rewrite or a generalized rewrite.
Approach is to break everything, if [eq] appears in head position
then regular rewrite else try general rewrite.
If occurrences are set, use general rewrite.
*)let(forward_general_setoid_rewrite_clause,general_setoid_rewrite_clause)=Hook.make()(* Do we have a JMeq instance on twice the same domains ? *)letjmeq_same_domenvsigma=function|None->true(* already checked in Hipattern.find_eq_data_decompose *)|Somet->letrels,t=decompose_prod_assumsigmatinletenv=push_rel_contextrelsenvinmatchdecompose_appsigmatwith|_,[dom1;_;dom2;_]->is_convenvsigmadom1dom2|_->falseleteq_elimination_refl2rsort=letname=ifl2rthenmatchsortwith|InProp->"core.eq.ind_r"|InSProp->"core.eq.sind_r"|InSet|InType->"core.eq.rect_r"elsematchsortwith|InProp->"core.eq.ind"|InSProp->"core.eq.sind"|InSet|InType->"core.eq.rect"inifCoqlib.has_refnamethenSome(Coqlib.lib_refname)elseNone(* find_elim determines which elimination principle is necessary to
eliminate lbeq on sort_of_gl. *)letfind_elimhdcncllft2rgtdepclsot=Proofview.Goal.enter_onebeginfungl->letsigma=projectglinletis_global_existsgrc=Coqlib.has_refgr&&isRefXsigma(Coqlib.lib_refgr)cinletinccl=Option.is_emptyclsinletenv=Proofview.Goal.envglinletis_eq=is_global_exists"core.eq.type"hdcnclinletis_jmeq=is_global_exists"core.JMeq.type"hdcncl&&jmeq_same_domenvsigmaotinif(is_eq||is_jmeq)&¬depthenletsort=elimination_sort_of_clauseclsglinletc=matchEConstr.kindsigmahdcnclwith|Ind(ind_sp,u)->beginmatchlft2rgt,clswith|Sometrue,None|Somefalse,Some_->beginmatchifis_eqtheneq_elimination_reftruesortelseNonewith|Somer->destConstRefr|None->letc1=destConstRef(lookup_eliminatorenvind_spsort)inletmp,l=Constant.repr2(Constant.make1(Constant.canonicalc1))inletl'=Label.of_id(add_suffix(Label.to_idl)"_r")inletc1'=Global.constant_of_delta_kn(KerName.makempl')inifnot(Environ.mem_constantc1'(Global.env()))thenuser_err~hdr:"Equality.find_elim"(str"Cannot find rewrite principle "++Label.printl'++str".");c1'end|_->beginmatchifis_eqtheneq_elimination_reffalsesortelseNonewith|Somer->destConstRefr|None->destConstRef(lookup_eliminatorenvind_spsort)endend|_->(* cannot occur since we checked that we are in presence of
Logic.eq or Jmeq just before *)assertfalseinpf_constr_of_global(GlobRef.ConstRefc)elseletscheme_name=matchdep,lft2rgt,incclwith(* Non dependent case *)|false,Sometrue,true->rew_l2r_scheme_kind|false,Sometrue,false->rew_r2l_scheme_kind|false,_,false->rew_l2r_scheme_kind|false,_,true->rew_r2l_scheme_kind(* Dependent case *)|true,Sometrue,true->rew_l2r_dep_scheme_kind|true,Sometrue,false->rew_l2r_forward_dep_scheme_kind|true,_,true->rew_r2l_dep_scheme_kind|true,_,false->rew_r2l_forward_dep_scheme_kindinmatchEConstr.kindsigmahdcnclwith|Ind(ind,u)->find_schemescheme_nameind>>=func->pf_constr_of_global(GlobRef.ConstRefc)|_->assertfalseendlettype_of_clauseclsgl=matchclswith|None->Proofview.Goal.conclgl|Someid->pf_get_hyp_typidglletleibniz_rewrite_ebindings_clauseclslft2rgttacctlwith_evarsfrzevarsdep_proof_okhdcncl=Proofview.Goal.enterbeginfungl->letevd=Proofview.Goal.sigmaglinletenv=Proofview.Goal.envglinletisatomic=isProdevd(whd_zetaenvevdhdcncl)inletdep_fun=ifisatomicthendependentelsedependent_no_evarinlettype_of_cls=type_of_clauseclsglinletdep=dep_proof_ok&&dep_funevdctype_of_clsinfind_elimhdcncllft2rgtdepcls(Somet)>>=funelim->general_elim_clausewith_evarsfrzevarstacclsctl(matchlft2rgtwithNone->false|Someb->b)(ElimTermelim)endletadjust_rewriting_directionargslft2rgt=matchargswith|[_]->(* equality to a constant, like in eq_true *)(* more natural to see -> as the rewriting to the constant *)ifnotlft2rgtthenuser_errPp.(str"Rewriting non-symmetric equality not allowed from right-to-left.");None|_->(* other equality *)Somelft2rgtletrewrite_side_tactacsidetac=side_tactac(Option.mapfstsidetac)(* Main function for dispatching which kind of rewriting it is about *)letgeneral_rewrite~where:cls~l2r:lft2rgtoccs~freeze:frzevars~dep:dep_proof_ok~with_evars?tac((c,l):constrwith_bindings)=ifnot(Locusops.is_all_occurrencesoccs)then(rewrite_side_tac(Hook.getforward_general_setoid_rewrite_clauseclslft2rgtoccs(c,l)~new_goals:[])tac)elseProofview.Goal.enterbeginfungl->letsigma=Tacmach.New.projectglinletenv=Proofview.Goal.envglinletctype=get_type_ofenvsigmacinletrels,t=decompose_prod_assumsigma(whd_betaiotazetaenvsigmactype)inmatchmatch_with_equality_typeenvsigmatwith|Some(hdcncl,args)->(* Fast path: direct leibniz-like rewrite *)letlft2rgt=adjust_rewriting_directionargslft2rgtinleibniz_rewrite_ebindings_clauseclslft2rgttacc(it_mkProd_or_LetIntrels)lwith_evarsfrzevarsdep_proof_okhdcncl|None->Proofview.tclORELSEbeginrewrite_side_tac(Hook.getforward_general_setoid_rewrite_clauseclslft2rgtoccs(c,l)~new_goals:[])tacendbeginfunction|(e,info)->Proofview.tclEVARMAP>>=funsigma->letenv'=push_rel_contextrelsenvinletrels',t'=splay_prod_assumenv'sigmatin(* Search for underlying eq *)matchmatch_with_equality_typeenv'sigmat'with|Some(hdcncl,args)->letlft2rgt=adjust_rewriting_directionargslft2rgtinleibniz_rewrite_ebindings_clauseclslft2rgttacc(it_mkProd_or_LetInt'(rels'@rels))lwith_evarsfrzevarsdep_proof_okhdcncl|None->Proofview.tclZERO~infoe(* error "The provided term does not end with an equality or a declared rewrite relation." *)endendletgeneral_rewrite_clausel2rwith_evars?tacccl=letoccs_of=occurrences_map(List.fold_left(funacc->functionArgArgx->x::acc|ArgVar_->acc)[])inmatchcl.onhypswith|Somel->(* If a precise list of locations is given, success is mandatory for
each of these locations. *)letrecdo_hyps=function|[]->Proofview.tclUNIT()|((occs,id),_)::l->tclTHENFIRST(general_rewrite~where:(Someid)~l2r(occs_ofoccs)~freeze:false~dep:true~with_evars?tacc)(do_hypsl)inifcl.concl_occs==NoOccurrencesthendo_hypslelsetclTHENFIRST(general_rewrite~where:None~l2r(occs_ofcl.concl_occs)~freeze:false~dep:true~with_evars?tacc)(do_hypsl)|None->(* Otherwise, if we are told to rewrite in all hypothesis via the
syntax "* |-", we fail iff all the different rewrites fail *)letrecdo_hyps_atleastonce=function|[]->tclZEROMSG(Pp.str"Nothing to rewrite.")|id::l->tclIFTHENFIRSTTRYELSEMUST(general_rewrite~where:(Someid)~l2rAllOccurrences~freeze:false~dep:true~with_evars?tacc)(do_hyps_atleastoncel)inletdo_hyps=(* If the term to rewrite uses an hypothesis H, don't rewrite in H *)letidsgl=letids_in_c=Termops.global_vars_set(Proofview.Goal.envgl)(projectgl)(fstc)inletids_of_hyps=pf_ids_of_hypsglinId.Set.fold(funidl->List.removeId.equalidl)ids_in_cids_of_hypsinProofview.Goal.enterbeginfungl->do_hyps_atleastonce(idsgl)endinifcl.concl_occs==NoOccurrencesthendo_hypselsetclIFTHENFIRSTTRYELSEMUST(general_rewrite~where:None~l2r(occs_ofcl.concl_occs)~freeze:false~dep:true~with_evars?tacc)do_hypsletapply_special_clear_requestclear_flagf=Proofview.Goal.enterbeginfungl->letsigma=Tacmach.New.projectglinletenv=Proofview.Goal.envglintrylet(sigma,(c,bl))=fenvsigmainapply_clear_requestclear_flag(use_clear_hyp_by_default())cwithewhennoncriticale->tclIDTACendtypemulti=|Preciselyofint|UpToofint|RepeatStar|RepeatPlusletgeneral_multi_rewritewith_evarslcltac=letdo1l2rf=Proofview.Goal.enterbeginfungl->letsigma=Tacmach.New.projectglinletenv=Proofview.Goal.envglinlet(sigma,c)=fenvsigmaintclWITHHOLESwith_evars(general_rewrite_clausel2rwith_evars?tacccl)sigmaendinletrecdoNl2rc=function|Preciselynwhenn<=0->Proofview.tclUNIT()|Precisely1->do1l2rc|Preciselyn->tclTHENFIRST(do1l2rc)(doNl2rc(Precisely(n-1)))|RepeatStar->tclREPEAT_MAIN(do1l2rc)|RepeatPlus->tclTHENFIRST(do1l2rc)(doNl2rcRepeatStar)|UpTonwhenn<=0->Proofview.tclUNIT()|UpTon->tclTHENFIRST(tclTRY(do1l2rc))(doNl2rc(UpTo(n-1)))inletrecloop=function|[]->Proofview.tclUNIT()|(l2r,m,clear_flag,c)::l->tclTHENFIRST(tclTHEN(doNl2rcm)(apply_special_clear_requestclear_flagc))(loopl)inlooplletrewriteLRc=general_rewrite~where:None~l2r:trueAllOccurrences~freeze:true~dep:true~with_evars:false(c,NoBindings)letrewriteRLc=general_rewrite~where:None~l2r:falseAllOccurrences~freeze:true~dep:true~with_evars:false(c,NoBindings)(* Replacing tactics *)letclasses_dirpath=DirPath.make(List.mapId.of_string["Classes";"Coq"])letinit_setoid()=ifis_dirpath_prefix_ofclasses_dirpath(Lib.cwd())then()elsecheck_required_library["Coq";"Setoids";"Setoid"]letcheck_setoidcl=letconcloccs=Locusops.occurrences_map(funx->x)cl.concl_occsinOption.fold_left(List.fold_left(funb((occ,_),_)->b||(not(Locusops.is_all_occurrences(Locusops.occurrences_map(funx->x)occ)))))(not(Locusops.is_all_occurrencesconcloccs)&&(concloccs<>NoOccurrences))cl.onhypsletreplace_coreclausel2req=ifcheck_setoidclausetheninit_setoid();tclTHENFIRST(assert_asfalseNoneNoneeq)(onLastHypId(funid->tclTHEN(tclTRY(general_rewrite_clausel2rfalse(mkVarid,NoBindings)clause))(clear[id])))(* eq,sym_eq : equality on Type and its symmetry theorem
c1 c2 : c1 is to be replaced by c2
unsafe : If true, do not check that c1 and c2 are convertible
tac : Used to prove the equality c1 = c2
gl : goal *)letreplace_using_leibnizclausec1c2l2runsafetry_prove_eq_opt=lettry_prove_eq=matchtry_prove_eq_optwith|None->Proofview.tclUNIT()|Sometac->tclCOMPLETEtacinProofview.Goal.enterbeginfungl->letget_type_of=pf_applyget_type_ofglinlett1=get_type_ofc1andt2=get_type_ofc2inletevd=ifunsafethenSome(Tacmach.New.projectgl)elsetrySome(Evarconv.unify_delay(Proofview.Goal.envgl)(Tacmach.New.projectgl)t1t2)withEvarconv.UnableToUnify_->Noneinmatchevdwith|None->tclFAIL0(str"Terms do not have convertible types")|Someevd->lete,sym=trylib_ref"core.eq.type",lib_ref"core.eq.sym"withUserError_->trylib_ref"core.identity.type",lib_ref"core.identity.sym"withUserError_->user_err(strbrk"Need a registration for either core.eq.type and core.eq.sym or core.identity.type and core.identity.sym.")inTacticals.New.pf_constr_of_globalsym>>=funsym->Tacticals.New.pf_constr_of_globale>>=fune->leteq=applist(e,[t1;c1;c2])intclTHENLAST(replace_coreclausel2req)(tclFIRST[assumption;tclTHEN(applysym)assumption;try_prove_eq])endletreplacec1c2=replace_using_leibnizonConclc2c1falsefalseNoneletreplace_byc1c2tac=replace_using_leibnizonConclc2c1falsefalse(Sometac)letreplace_in_clause_maybe_byc1c2cltac_opt=replace_using_leibnizclc2c1falsefalsetac_opt(* End of Eduardo's code. The rest of this file could be improved
using the functions match_with_equation, etc that I defined
in Pattern.ml.
-- Eduardo (19/8/97)
*)(* Tactics for equality reasoning with the "eq" relation. This code
will work with any equivalence relation which is substitutive *)(* [find_positions t1 t2]
will find the positions in the two terms which are suitable for
discrimination, or for injection. Obviously, if there is a
position which is suitable for discrimination, then we want to
exploit it, and not bother with injection. So when we find a
position which is suitable for discrimination, we will just raise
an exception with that position.
So the algorithm goes like this:
if [t1] and [t2] start with the same constructor, then we can
continue to try to find positions in the arguments of [t1] and
[t2].
if [t1] and [t2] do not start with the same constructor, then we
have found a discrimination position
if one [t1] or [t2] do not start with a constructor and the two
terms are not already convertible, then we have found an injection
position.
A discriminating position consists of a constructor-path and a pair
of operators. The constructor-path tells us how to get down to the
place where the two operators, which must differ, can be found.
An injecting position has two terms instead of the two operators,
since these terms are different, but not manifestly so.
A constructor-path is a list of pairs of (operator * int), where
the int (based at 0) tells us which argument of the operator we
descended into.
*)exceptionDiscrFoundof(constructor*int)list*constructor*constructorletkeep_proof_equalities_for_injection=reffalselet()=declare_bool_option{optdepr=false;optkey=["Keep";"Proof";"Equalities"];optread=(fun()->!keep_proof_equalities_for_injection);optwrite=(funb->keep_proof_equalities_for_injection:=b)}letkeep_proof_equalities=function|None->!keep_proof_equalities_for_injection|Someflags->flags.keep_proof_equalities(* [keep_proofs] is relevant for types in Prop with elimination in Type *)(* In particular, it is relevant for injection but not for discriminate *)letfind_positionsenvsigma~keep_proofs~no_discrt1t2=letprojectenvsortsposnt1t2=letty1=get_type_ofenvsigmat1inlets=get_sort_family_of~truncation_style:trueenvsigmaty1inifList.mem_fSorts.family_equalssortsthen[(List.revposn,t1,t2)]else[]inletrecfindrecsortsposnt1t2=lethd1,args1=whd_all_stackenvsigmat1inlethd2,args2=whd_all_stackenvsigmat2inmatch(EConstr.kindsigmahd1,EConstr.kindsigmahd2)with|Construct((ind1,i1assp1),u1),Construct(sp2,_)whenInt.equal(List.lengthargs1)(constructor_nallargsenvsp1)->letsorts'=CList.intersectSorts.family_equalsorts(sorts_below(top_allowed_sortenv(fstsp1)))in(* both sides are fully applied constructors, so either we descend,
or we can discriminate here. *)ifConstruct.CanOrd.equalsp1sp2thenletnparams=inductive_nparamsenvind1inletparams1,rargs1=List.chopnparamsargs1inlet_,rargs2=List.chopnparamsargs2inlet(mib,mip)=lookup_mind_specifenvind1inletparams1=List.mapEConstr.Unsafe.to_constrparams1inletu1=EInstance.kindsigmau1inletctxt=(get_constructor((ind1,u1),mib,mip,params1)i1).cs_argsinletadjusti=CVars.adjust_rel_to_rel_contextctxt(i+1)-1inList.flatten(List.map2_i(funi->findrecsorts'((sp1,adjusti)::posn))0rargs1rargs2)elseifList.mem_fSorts.family_equalInTypesorts'&¬no_discrthen(* see build_discriminator *)raise(DiscrFound(List.revposn,sp1,sp2))else(* if we cannot eliminate to Type, we cannot discriminate but we
may still try to project *)projectenvsortsposn(applist(hd1,args1))(applist(hd2,args2))|_->lett1_0=applist(hd1,args1)andt2_0=applist(hd2,args2)inifis_convenvsigmat1_0t2_0then[]elseprojectenvsortsposnt1_0t2_0intryletsorts=ifkeep_proofsthen[InSet;InType;InProp]else[InSet;InType]inInr(findrecsorts[]t1t2)withDiscrFound(path,c1,c2)->Inl(path,c1,c2)letuse_keep_proofs=function|None->!keep_proof_equalities_for_injection|Someb->bletdiscriminableenvsigmat1t2=matchfind_positionsenvsigma~keep_proofs:false~no_discr:falset1t2with|Inl_->true|_->falseletinjectableenvsigma~keep_proofst1t2=matchfind_positionsenvsigma~keep_proofs:(use_keep_proofskeep_proofs)~no_discr:truet1t2with|Inl_->assertfalse|Inr[]|Inr[([],_,_)]->false|Inr_->true(* Once we have found a position, we need to project down to it. If
we are discriminating, then we need to produce False on one of the
branches of the discriminator, and True on the other one. So the
result type of the case-expressions is always Prop.
If we are injecting, then we need to discover the result-type.
This can be difficult, since the type of the two terms at the
injection-position can be different, and we need to find a
dependent sigma-type which generalizes them both.
We can get an approximation to the right type to choose by:
(0) Before beginning, we reserve a patvar for the default
value of the match, to be used in all the bogus branches.
(1) perform the case-splits, down to the site of the injection. At
each step, we have a term which is the "head" of the next
case-split. At the point when we actually reach the end of our
path, the "head" is the term to return. We compute its type, and
then, backwards, make a sigma-type with every free debruijn
reference in that type. We can be finer, and first do a S(TRONG)NF
on the type, so that we get the fewest number of references
possible.
(2) This gives us a closed type for the head, which we use for the
types of all the case-splits.
(3) Now, we can compute the type of one of T1, T2, and then unify
it with the type of the last component of the result-type, and this
will give us the bindings for the other arguments of the tuple.
*)(* The algorithm, then is to perform successive case-splits. We have
the result-type of the case-split, and also the type of that
result-type. We have a "direction" we want to follow, i.e. a
constructor-number, and in all other "directions", we want to juse
use the default-value.
After doing the case-split, we call the afterfun, with the updated
environment, to produce the term for the desired "direction".
The assumption is made here that the result-type is not manifestly
functional, so we can just use the length of the branch-type to
know how many lambda's to stick in.
*)(* [descend_then env sigma head dirn]
returns the number of products introduced, and the environment
which is active, in the body of the case-branch given by [dirn],
along with a continuation, which expects to be fed:
(1) the value of the body of the branch given by [dirn]
(2) the default-value
(3) the type of the default-value, which must also be the type of
the body of the [dirn] branch
the continuation then constructs the case-split.
*)letdescend_thenenvsigmaheaddirn=letIndType(indf,_)asindt=tryfind_rectypeenvsigma(get_type_ofenvsigmahead)withNot_found->user_errPp.(str"Cannot project on an inductive type derived from a dependency.")inletindp,_=(dest_ind_familyindf)inletind,_=check_privacyenvindpinlet(mib,mip)=lookup_mind_specifenvindinletcstr=get_constructorsenvindfinletdirn_nlams=cstr.(dirn-1).cs_nargsinletdirn_env=Environ.push_rel_contextcstr.(dirn-1).cs_argsenvin(dirn_nlams,dirn_env,(funsigmadirnval(dfltval,resty)->letdeparsign=make_arity_signatureenvsigmatrueindfinletp=it_mkLambda_or_LetIn(lift(mip.mind_nrealargs+1)resty)deparsigninletbuild_branchi=letresult=ifInt.equalidirnthendirnvalelsedfltvalinletcs_args=List.map(fund->map_rel_declEConstr.of_constrd)cstr.(i-1).cs_argsinletargs=name_contextenvsigmacs_argsinit_mkLambda_or_LetInresultargsinletbrl=List.mapbuild_branch(List.interval1(Array.lengthmip.mind_consnames))inletrci=Sorts.Relevantin(* TODO relevance *)letci=make_case_infoenvindrciRegularStyleinInductiveops.make_case_or_projectenvsigmaindtciphead(Array.of_listbrl)))(* Now we need to construct the discriminator, given a discriminable
position. This boils down to:
(1) If the position is directly beneath us, then we need to do a
case-split, with result-type Prop, and stick True and False into
the branches, as is convenient.
(2) If the position is not directly beneath us, then we need to
call descend_then, to descend one step, and then recursively
construct the discriminator.
*)(* [construct_discriminator env sigma dirn c ind special default]]
constructs a case-split on [c] of type [ind], with the [dirn]-th
branch giving [special], and all the rest giving [default]. *)letbuild_selectorenvsigmadirncindspecialdefault=letIndType(indf,_)asindt=tryfind_rectypeenvsigmaindwithNot_found->(* one can find Rel(k) in case of dependent constructors
like T := c : (A:Set)A->T and a discrimination
on (c bool true) = (c bool false)
CP : changed assert false in a more informative error
*)user_err~hdr:"Equality.construct_discriminator"(str"Cannot discriminate on inductive constructors with \
dependent types.")inlet(indp,_)=dest_ind_familyindfinletind,_=check_privacyenvindpinlettyp=Retyping.get_type_ofenvsigmadefaultinlet(mib,mip)=lookup_mind_specifenvindinletdeparsign=make_arity_signatureenvsigmatrueindfinletp=it_mkLambda_or_LetIntypdeparsigninletcstrs=get_constructorsenvindfinletbuild_branchi=letendpt=ifInt.equalidirnthenspecialelsedefaultinletargs=List.map(fund->map_rel_declEConstr.of_constrd)cstrs.(i-1).cs_argsinit_mkLambda_or_LetInendptargsinletbrl=List.mapbuild_branch(List.interval1(Array.lengthmip.mind_consnames))inletrci=Sorts.Relevantin(* TODO relevance *)letci=make_case_infoenvindrciRegularStyleinletans=Inductiveops.make_case_or_projectenvsigmaindtcipc(Array.of_listbrl)inansletbuild_coq_False()=pf_constr_of_global(lib_ref"core.False.type")letbuild_coq_True()=pf_constr_of_global(lib_ref"core.True.type")letbuild_coq_I()=pf_constr_of_global(lib_ref"core.True.I")letrecbuild_discriminatorenvsigmatrue_0false_0dirnc=function|[]->letind=get_type_ofenvsigmacinbuild_selectorenvsigmadirncindtrue_0(fstfalse_0)|((sp,cnum),argnum)::l->let(cnum_nlams,cnum_env,kont)=descend_thenenvsigmaccnuminletnewc=mkRel(cnum_nlams-argnum)inletsubval=build_discriminatorcnum_envsigmatrue_0false_0dirnnewclinkontsigmasubvalfalse_0(* Note: discrimination could be more clever: if some elimination is
not allowed because of a large impredicative constructor in the
path (see allowed_sorts in find_positions), the positions could
still be discrimated by projecting first instead of putting the
discrimination combinator inside the projecting combinator. Example
of relevant situation:
Inductive t:Set := c : forall A:Set, A -> nat -> t.
Goal ~ c _ 0 0 = c _ 0 1. intro. discriminate H.
*)letgen_absurdityid=Proofview.Goal.enterbeginfungl->letenv=pf_envglinletsigma=projectglinlethyp_typ=pf_get_hyp_typidglinifis_empty_typeenvsigmahyp_typthensimplest_elim(mkVarid)elsetclZEROMSG(str"Not the negation of an equality.")end(* Precondition: eq is leibniz equality
returns ((eq_elim t t1 P i t2), absurd_term)
where P=[e:t]discriminator
absurd_term=False
*)letind_scheme_of_eqlbeqto_kind=let(mib,mip)=Global.lookup_inductive(destIndReflbeq.eq)inletfrom_kind=inductive_sort_familymipin(* use ind rather than case by compatibility *)letkind=Elimschemes.nondep_elim_schemefrom_kindto_kindinfind_schemekind(destIndReflbeq.eq)>>=func->Proofview.tclUNIT(GlobRef.ConstRefc)letdiscrimination_pfe(t,t1,t2)discriminatorlbeqto_kind=build_coq_I()>>=funi->ind_scheme_of_eqlbeqto_kind>>=funeq_elim->pf_constr_of_globaleq_elim>>=funeq_elim->Proofview.tclUNIT(applist(eq_elim,[t;t1;mkNamedLambda(make_annoteSorts.Relevant)tdiscriminator;i;t2]))typeequality={eq_data:(coq_eq_data*(EConstr.t*EConstr.t*EConstr.t));(* equality data + A : Type, t1 : A, t2 : A *)eq_clenv:clausenv;(* clause [M : R A t1 t2] where [R] is the equality from above *)}leteq_baseid=Id.of_string"e"letdiscr_positionsenvsigma{eq_data=(lbeq,(t,t1,t2));eq_clenv=eq_clause}cpathdirn=build_coq_True()>>=funtrue_0->build_coq_False()>>=funfalse_0->letfalse_ty=Retyping.get_type_ofenvsigmafalse_0inletfalse_kind=Retyping.get_sort_family_ofenvsigmafalse_0inlete=next_ident_awayeq_baseid(vars_of_envenv)inlete_env=push_named(Context.Named.Declaration.LocalAssum(make_annoteSorts.Relevant,t))envinletdiscriminator=tryProofview.tclUNIT(build_discriminatore_envsigmatrue_0(false_0,false_ty)dirn(mkVare)cpath)withUserError_asex->let_,info=Exninfo.captureexinProofview.tclZERO~infoexindiscriminator>>=fundiscriminator->discrimination_pfe(t,t1,t2)discriminatorlbeqfalse_kind>>=funpf->(* pf : eq t t1 t2 -> False *)letpf=EConstr.mkApp(pf,[|clenv_valueeq_clause|])intclTHENS(assert_afterAnonymousfalse_0)[onLastHypIdgen_absurdity;(Logic.refiner~check:trueEConstr.Unsafe.(to_constrpf))]letdiscrEqeq=let{eq_data=(_,(_,t1,t2));eq_clenv=eq_clause}=eqinletsigma=eq_clause.evdinProofview.Goal.enterbeginfungl->letenv=Proofview.Goal.envglinmatchfind_positionsenvsigma~keep_proofs:false~no_discr:falset1t2with|Inr_->letinfo=Exninfo.reify()intclZEROMSG~info(str"Not a discriminable equality.")|Inl(cpath,(_,dirn),_)->discr_positionsenvsigmaeqcpathdirnendletonEqualitywith_evarstac(c,lbindc)=Proofview.Goal.enterbeginfungl->letreduce_to_quantified_ind=pf_applyTacred.reduce_to_quantified_indglinlett=pf_get_type_ofglcinlett'=trysnd(reduce_to_quantified_indt)withUserError_->tinleteq_clause=pf_applymake_clenv_bindinggl(c,t')lbindcinleteq_clause'=Clenv.clenv_pose_dependent_evars~with_evarseq_clauseinleteqn=clenv_typeeq_clause'in(* FIXME evar leak *)let(eq,u,eq_args)=pf_applyfind_this_eq_data_decomposegleqninleteq={eq_data=(eq,eq_args);eq_clenv=eq_clause'}intclTHEN(Proofview.Unsafe.tclEVARSeq_clause'.evd)(taceq)endletonNegatedEqualitywith_evarstac=Proofview.Goal.enterbeginfungl->letsigma=Tacmach.New.projectglinletccl=Proofview.Goal.conclglinletenv=Proofview.Goal.envglinmatchEConstr.kindsigma(hnf_constrenvsigmaccl)with|Prod(_,t,u)whenis_empty_typeenvsigmau->tclTHENintrof(onLastHypId(funid->onEqualitywith_evarstac(mkVarid,NoBindings)))|_->letinfo=Exninfo.reify()intclZEROMSG~info(str"Not a negated primitive equality.")endletdiscrSimpleClausewith_evars=function|None->onNegatedEqualitywith_evarsdiscrEq|Someid->onEqualitywith_evarsdiscrEq(mkVarid,NoBindings)letdiscrwith_evars=onEqualitywith_evarsdiscrEqletdiscrClausewith_evars=onClause(discrSimpleClausewith_evars)letdiscrEverywherewith_evars=tclTHEN(Proofview.tclUNIT())(* Delay the interpretation of side-effect *)(tclTHEN(tclREPEATintrof)(tryAllHyps(funid->tclCOMPLETE(discrwith_evars(mkVarid,NoBindings)))))letdiscr_tacwith_evars=function|None->discrEverywherewith_evars|Somec->onInductionArg(funclear_flag->discrwith_evars)cletdiscrConcl=discrClausefalseonConclletdiscrHypid=discrClausefalse(onHypid)(* returns the sigma type (sigS, sigT) with the respective
constructor depending on the sort *)(* J.F.: correction du bug #1167 en accord avec Hugo. *)letfind_sigma_dataenvs=build_sigma_type()(* [make_tuple env sigma (rterm,rty) lind] assumes [lind] is the lesser
index bound in [rty]
Then we build the term
[(existT A P (mkRel lind) rterm)] of type [(sigS A P)]
where [A] is the type of [mkRel lind] and [P] is [\na:A.rty{1/lind}]
*)letmake_tupleenvsigma(rterm,rty)lind=assert(not(noccurnsigmalindrty));letsigdata=find_sigma_dataenv(get_sort_ofenvsigmarty)inletsigma,a=type_of~refresh:trueenvsigma(mkRellind)inleta=simplenvsigmaainletna=Context.Rel.Declaration.get_annot(lookup_rellindenv)in(* We move [lind] to [1] and lift other rels > [lind] by 1 *)letrty=lift(1-lind)(liftnlind(lind+1)rty)in(* Now [lind] is [mkRel 1] and we abstract on (na:a) *)letp=mkLambda(na,a,rty)inletsigma,exist_term=Evd.fresh_globalenvsigmasigdata.introinletsigma,sig_term=Evd.fresh_globalenvsigmasigdata.typinsigma,(applist(exist_term,[a;p;(mkRellind);rterm]),applist(sig_term,[a;p]))(* check that the free-references of the type of [c] are contained in
the free-references of the normal-form of that type. Strictly
computing the exact set of free rels would require full
normalization but this is not reasonable (e.g. in presence of
records that contains proofs). We restrict ourself to a "simpl"
normalization *)letminimal_free_relsenvsigma(c,cty)=letcty_rels=free_relssigmactyinletcty'=simplenvsigmactyinletrels'=free_relssigmacty'inifInt.Set.subsetcty_relsrels'then(cty,cty_rels)else(cty',rels')(* Like the above, but recurse over all the rels found until there are
no more rels to be found *)letminimal_free_rels_recenvsigma=letrecminimalrec_free_rels_recprev_rels(c,cty)=let(cty,direct_rels)=minimal_free_relsenvsigma(c,cty)inletcombined_rels=Int.Set.unionprev_relsdirect_relsinletfolderrelsi=snd(minimalrec_free_rels_recrels(c,get_type_ofenvsigma(mkReli)))in(cty,List.fold_leftfoldercombined_rels(Int.Set.elements(Int.Set.diffdirect_relsprev_rels)))inminimalrec_free_rels_recInt.Set.empty(* [sig_clausal_form siglen ty]
Will explode [siglen] [sigS,sigT ]'s on [ty] (depending on the
type of ty), and return:
(1) a pattern, with meta-variables in it for various arguments,
which, when the metavariables are replaced with appropriate
terms, will have type [ty]
(2) an integer, which is the last argument - the one which we just
returned.
(3) a pattern, for the type of that last meta
(4) a typing for each patvar
WARNING: No checking is done to make sure that the
sigS(or sigT)'s are actually there.
- Only homogeneous pairs are built i.e. pairs where all the
dependencies are of the same sort
[sig_clausal_form] proceed as follows: the default tuple is
constructed by taking the tuple-type, exploding the first [tuplen]
[sigS]'s, and replacing at each step the binder in the
right-hand-type by a fresh metavariable. In addition, on the way
back out, we will construct the pattern for the tuple which uses
these meta-vars.
This gives us a pattern, which we use to match against the type of
[dflt]; if that fails, then against the S(TRONG)NF of that type. If
both fail, then we just cannot construct our tuple. If one of
those succeed, then we can construct our value easily - we just use
the tuple-pattern.
*)letsig_clausal_formenvsigmasort_of_tysiglentydflt=letsigdata=find_sigma_dataenvsort_of_tyinletrecsigrec_clausal_formsigmasiglenp_i=ifInt.equalsiglen0then(* is the default value typable with the expected type *)letsigma,dflt_typ=type_ofenvsigmadfltintryletsigma=Evarconv.unify_leq_delayenvsigmadflt_typp_iinletsigma=Evarconv.solve_unif_constraints_with_heuristicsenvsigmainsigma,dfltwithEvarconv.UnableToUnify_->user_errPp.(str"Cannot solve a unification problem.")elselet(a,p_i_minus_1)=matchwhd_beta_stackenvsigmap_iwith|(_sigS,[a;p])->(a,p)|_->anomaly~label:"sig_clausal_form"(Pp.str"should be a sigma type.")inletsigma,ev=Evarutil.new_evarenvsigmaainletrty=beta_applistsigma(p_i_minus_1,[ev])inletsigma,tuple_tail=sigrec_clausal_formsigma(siglen-1)rtyinifEConstr.isEvarsigmaevthen(* This at least happens if what has been detected as a
dependency is not one; use an evasive error message;
even if the problem is upwards: unification should be
tried in the first place in make_iterated_tuple instead
of approximatively computing the free rels; then
unsolved evars would mean not binding rel *)user_errPp.(str"Cannot solve a unification problem.")elseletsigma,exist_term=Evd.fresh_globalenvsigmasigdata.introinsigma,applist(exist_term,[a;p_i_minus_1;ev;tuple_tail])inletsigma=Evd.clear_metassigmainletsigma,scf=sigrec_clausal_formsigmasiglentyinsigma,Evarutil.nf_evarsigmascf(* The problem is to build a destructor (a generalization of the
predecessor) which, when applied to a term made of constructors
(say [Ci(e1,Cj(e2,Ck(...,term,...),...),...)]), returns a given
subterm of the term (say [term]).
Let [typ] be the type of [term]. If [term] has no dependencies in
the [e1], [e2], etc, then all is simple. If not, then we need to
encapsulated the dependencies into a dependent tuple in such a way
that the destructor has not a dependent type and rewriting can then
be applied. The destructor has the form
[e]Cases e of
| ...
| Ci (x1,x2,...) =>
Cases x2 of
| ...
| Cj (y1,y2,...) =>
Cases y2 of
| ...
| Ck (...,z,...) => z
| ... end
| ... end
| ... end
and the dependencies is expressed by the fact that [z] has a type
dependent in the x1, y1, ...
Assume [z] is typed as follows: env |- z:zty
If [zty] has no dependencies, this is simple. Otherwise, assume
[zty] has free (de Bruijn) variables in,...i1 then the role of
[make_iterated_tuple env sigma (term,typ) (z,zty)] is to build the
tuple
[existT [xn]Pn Rel(in) .. (existT [x2]P2 Rel(i2) (existT [x1]P1 Rel(i1) z))]
where P1 is zty[i1/x1], P2 is {x1 | P1[i2/x2]} etc.
To do this, we find the free (relative) references of the strong NF
of [z]'s type, gather them together in left-to-right order
(i.e. highest-numbered is farthest-left), and construct a big
iterated pair out of it. This only works when the references are
all themselves to members of [Set]s, because we use [sigS] to
construct the tuple.
Suppose now that our constructed tuple is of length [tuplen]. We
need also to construct a default value for the other branches of
the destructor. As default value, we take a tuple of the form
[existT [xn]Pn ?n (... existT [x2]P2 ?2 (existT [x1]P1 ?1 term))]
but for this we have to solve the following unification problem:
typ = zty[i1/?1;...;in/?n]
This is done by [sig_clausal_form].
*)letmake_iterated_tupleenvsigmadflt(z,zty)=let(zty,rels)=minimal_free_rels_recenvsigma(z,zty)inletsort_of_zty=get_sort_ofenvsigmaztyinletsorted_rels=Int.Set.elementsrelsinletsigma,(tuple,tuplety)=List.fold_left(fun(sigma,t)->make_tupleenvsigmat)(sigma,(z,zty))sorted_relsinassert(closed0sigmatuplety);letn=List.lengthsorted_relsinletsigma,dfltval=sig_clausal_formenvsigmasort_of_ztyntupletydfltinsigma,(tuple,tuplety,dfltval)letrecbuild_injrecenvsigmadfltc=function|[]->make_iterated_tupleenvsigmadflt(c,get_type_ofenvsigmac)|((sp,cnum),argnum)::l->trylet(cnum_nlams,cnum_env,kont)=descend_thenenvsigmaccnuminletnewc=mkRel(cnum_nlams-argnum)inletsigma,(subval,tuplety,dfltval)=build_injreccnum_envsigmadfltnewclinletres=kontsigmasubval(dfltval,tuplety)insigma,(res,tuplety,dfltval)withUserError_->failwith"caught"letbuild_injectorenvsigmadfltccpath=letsigma,(injcode,resty,_)=build_injrecenvsigmadfltccpathinsigma,(injcode,resty)leteq_dec_scheme_kind_name=ref(fun_->failwith"eq_dec_scheme undefined")letset_eq_dec_scheme_kindk=eq_dec_scheme_kind_name:=(fun_->k)letinject_if_homogenous_dependent_pairty=Proofview.Goal.enterbeginfungl->tryletsigma=Tacmach.New.projectglinleteq,u,(t,t1,t2)=pf_applyfind_this_eq_data_decomposegltyin(* fetch the informations of the pair *)letsigTconstr=Coqlib.(lib_ref"core.sigT.type")inletexistTconstr=Coqlib.lib_ref"core.sigT.intro"in(* check whether the equality deals with dep pairs or not *)leteqTypeDest=fst(decompose_appsigmat)inifnot(isRefXsigmasigTconstreqTypeDest)thenraiseExit;lethd1,ar1=decompose_app_vectsigmat1andhd2,ar2=decompose_app_vectsigmat2inifnot(isRefXsigmaexistTconstrhd1)thenraiseExit;ifnot(isRefXsigmaexistTconstrhd2)thenraiseExit;let(ind,_),_=trypf_applyfind_mrectypeglar1.(0)withNot_found->raiseExitin(* check if the user has declared the dec principle *)(* and compare the fst arguments of the dep pair *)(* Note: should work even if not an inductive type, but the table only *)(* knows inductive types *)ifnot(Option.has_some(Ind_tables.lookup_scheme(!eq_dec_scheme_kind_name())ind)&&pf_applyis_convglar1.(2)ar2.(2))thenraiseExit;check_required_library["Coq";"Logic";"Eqdep_dec"];letnew_eq_args=[|pf_get_type_ofglar1.(3);ar1.(3);ar2.(3)|]inletinj2=lib_ref"core.eqdep_dec.inj_pair2"infind_scheme(!eq_dec_scheme_kind_name())ind>>=func->(* cut with the good equality and prove the requested goal *)tclTHENLIST[intro;onLastHyp(funhyp->Tacticals.New.pf_constr_of_globalCoqlib.(lib_ref"core.eq.type")>>=funceq->tclTHENS(cut(mkApp(ceq,new_eq_args)))[clear[destVarsigmahyp];Tacticals.New.pf_constr_of_globalinj2>>=funinj2->Logic.refiner~check:trueEConstr.Unsafe.(to_constr(mkApp(inj2,[|ar1.(0);mkConstc;ar1.(1);ar1.(2);ar1.(3);ar2.(3);hyp|])))])]withExit->Proofview.tclUNIT()end(* Given t1=t2 Inj calculates the whd normal forms of t1 and t2 and it
expands then only when the whdnf has a constructor of an inductive type
in hd position, otherwise delta expansion is not done *)letsimplify_argsenvsigmat=(* Quick hack to reduce in arguments of eq only *)matchdecompose_appsigmatwith|eq,[t;c1;c2]->applist(eq,[t;simplenvsigmac1;simplenvsigmac2])|eq,[t1;c1;t2;c2]->applist(eq,[t1;simplenvsigmac1;t2;simplenvsigmac2])|_->tletinject_at_positionsenvsigmal2reqposnstac=let{eq_data=(eq,(t,t1,t2));eq_clenv=eq_clause}=eqinlete=next_ident_awayeq_baseid(vars_of_envenv)inlete_env=push_named(LocalAssum(make_annoteSorts.Relevant,t))envinletevdref=refsigmainletfilter(cpath,t1',t2')=try(* arbitrarily take t1' as the injector default value *)letsigma,(injbody,resty)=build_injectore_env!evdreft1'(mkVare)cpathinletinjfun=mkNamedLambda(make_annoteSorts.Relevant)tinjbodyinletsigma,congr=Evd.fresh_globalenvsigmaeq.congrin(* pf : eq t t1 t2 -> eq resty (injfun t1) (injfun t2) *)letpf=mkApp(congr,[|t;resty;injfun;t1;t2|])inletsigma,pf_typ=Typing.type_ofenvsigmapfinletpf_typ=Vars.subst1mkProp(pi3@@destProdsigmapf_typ)inletpf=mkApp(pf,[|Clenv.clenv_valueeq_clause|])inletty=simplify_argsenvsigmapf_typinevdref:=sigma;Some(pf,ty)withFailure_->Noneinletinjectors=List.map_filterfilterposnsinifList.is_emptyinjectorsthentclZEROMSG(str"Failed to decompose the equality.")elseProofview.tclTHEN(Proofview.Unsafe.tclEVARS!evdref)(Tacticals.New.tclTHENFIRST(Proofview.tclIGNORE(Proofview.Monad.List.map(fun(pf,ty)->tclTHENS(cutty)[inject_if_homogenous_dependent_pairty;Logic.refiner~check:trueEConstr.Unsafe.(to_constrpf)])(ifl2rthenList.revinjectorselseinjectors)))(tac(List.lengthinjectors)))exceptionNothingToInjectlet()=CErrors.register_handler(function|NothingToInject->Some(Pp.str"Nothing to inject.")|_->None)letinjEqThenkeep_proofstacl2reql=let{eq_data=(eq,(t,t1,t2));eq_clenv=eq_clause}=eqlinletsigma=eq_clause.evdinletenv=eq_clause.envinmatchfind_positionsenvsigma~keep_proofs~no_discr:truet1t2with|Inl_->assertfalse|Inr[]->letsuggestion=ifkeep_proofsthen""else" You can try to use option Set Keep Proof Equalities."intclZEROMSG(strbrk("No information can be deduced from this equality and the injectivity of constructors. This may be because the terms are convertible, or due to pattern matching restrictions in the sort Prop."^suggestion))|Inr[([],_,_)]->Proofview.tclZERONothingToInject|Inrposns->inject_at_positionsenvsigmal2reqlposns(tac(clenv_valueeq_clause))letget_previous_hyp_positionidgl=letenv,sigma=Proofview.Goal.(envgl,sigmagl)inletrecauxdest=function|[]->raise(RefinerError(env,sigma,NoSuchHypid))|d::right->lethyp=Context.Named.Declaration.get_iddinifId.equalhypidthendestelseaux(MoveAfterhyp)rightinauxMoveLast(Proofview.Goal.hypsgl)letinjEqflags?(old=false)with_evarsclear_flagipats=(* Decide which compatibility mode to use *)letipats_style,l2r,dft_clear_flag,bounded_intro=matchipatswith|Nonewhennotold&&use_injection_in_contextflags->Some[],true,true,true|None->None,false,false,false|_->letb=use_injection_pattern_l2r_orderflagsinipats,b,b,bin(* Built the post tactic depending on compatibility mode *)letpost_taccn=matchipats_stylewith|Someipats->Proofview.Goal.enterbeginfungl->letsigma=projectglinletdestopt=matchEConstr.kindsigmacwith|Varid->get_previous_hyp_positionidgl|_->MoveLastinletclear_tac=tclTRY(apply_clear_requestclear_flagdft_clear_flagc)in(* Try should be removal if dependency were treated *)letintro_tac=ifbounded_introthenintro_patterns_bound_towith_evarsndestoptipatselseintro_patterns_towith_evarsdestoptipatsintclTHENclear_tacintro_tacend|None->tclIDTACininjEqThen(keep_proof_equalitiesflags)post_tacl2rletinjflagsipatswith_evarsclear_flag=onEqualitywith_evars(injEqflagswith_evarsclear_flagipats)letinjClauseflagsipatswith_evars=function|None->onNegatedEqualitywith_evars(injEqflagswith_evarsNoneipats)|Somec->onInductionArg(injflagsipatswith_evars)cletsimpleInjClauseflagswith_evars=function|None->onNegatedEqualitywith_evars(injEqflags~old:truewith_evarsNoneNone)|Somec->onInductionArg(funclear_flag->onEqualitywith_evars(injEqflags~old:truewith_evarsclear_flagNone))cletinjConclflags=injClauseflagsNonefalseNoneletinjHypflagsclear_flagid=injClauseflagsNonefalse(Some(clear_flag,ElimOnIdentCAst.(makeid)))letdecompEqThenkeep_proofsntaceq=let{eq_data=(_,(_,t1,t2)asu);eq_clenv=clause}=eqinProofview.Goal.enterbeginfungl->letsigma=clause.evdinletenv=Proofview.Goal.envglinmatchfind_positionsenvsigma~keep_proofs~no_discr:falset1t2with|Inl(cpath,(_,dirn),_)->discr_positionsenvsigmaeqcpathdirn|Inr[]->(* Change: do not fail, simplify clear this trivial hyp *)ntac(clenv_valueclause)0|Inrposns->inject_at_positionsenvsigmatrueeqposns(ntac(clenv_valueclause))endletdEqThen~keep_proofswith_evarsntac=function|None->onNegatedEqualitywith_evars(decompEqThen(use_keep_proofskeep_proofs)(ntacNone))|Somec->onInductionArg(funclear_flag->onEqualitywith_evars(decompEqThen(use_keep_proofskeep_proofs)(ntacclear_flag)))cletdEq~keep_proofswith_evars=dEqThen~keep_proofswith_evars(funclear_flagcx->(apply_clear_requestclear_flag(use_clear_hyp_by_default())c))letintro_decomp_eqtac(eq,_,data)(c,t)=Proofview.Goal.enterbeginfungl->letcl=pf_applymake_clenv_bindinggl(c,t)NoBindingsinleteq={eq_data=(eq,data);eq_clenv=cl}indecompEqThen!keep_proof_equalities_for_injection(fun_->tac)eqendlet()=declare_intro_decomp_eqintro_decomp_eq(* [subst_tuple_term dep_pair B]
Given that dep_pair looks like:
(existT e1 (existT e2 ... (existT en en+1) ... ))
of type {x1:T1 & {x2:T2(x1) & ... {xn:Tn(x1..xn-1) & en+1 } } }
and B might contain instances of the ei, we will return the term:
([x1:ty1]...[xn+1:tyn+1]B
(projT1 (mkRel 1))
(projT1 (projT2 (mkRel 1)))
...
(projT1 (projT2^n (mkRel 1)))
(projT2 (projT2^n (mkRel 1)))
That is, we will abstract out the terms e1...en+1 of types
t1 (=_beta T1), ..., tn+1 (=_beta Tn+1(e1..en)) as usual, but
will then produce a term in which the abstraction is on a single
term - the debruijn index [mkRel 1], which will be of the same type
as dep_pair (note that the abstracted body may not be typable!).
ALGORITHM for abstraction:
We have a list of terms, [e1]...[en+1], which we want to abstract
out of [B]. For each term [ei], going backwards from [n+1], we
just do a [subst_term], and then do a lambda-abstraction to the
type of the [ei].
*)letdecomp_tuple_termenvsigmact=letrecdecomprecinner_codeexexty=letiterated_decomp=trylet({proj1=p1;proj2=p2}),(i,a,p,car,cdr)=find_sigma_data_decomposeenvsigmaexinletcar_code=applist(mkConstU(destConstRefp1,i),[a;p;inner_code])andcdr_code=applist(mkConstU(destConstRefp2,i),[a;p;inner_code])inletcdrtyp=beta_applistsigma(p,[car])inList.map(funl->((car,a),car_code)::l)(decompreccdr_codecdrcdrtyp)withConstr_matching.PatternMatchingFailure->[]in[((ex,exty),inner_code)]::iterated_decompindecomprec(mkRel1)ctletsubst_tuple_termenvsigmadep_pair1dep_pair2b=lettyp=get_type_ofenvsigmadep_pair1in(* We find all possible decompositions *)letdecomps1=decomp_tuple_termenvsigmadep_pair1typinletdecomps2=decomp_tuple_termenvsigmadep_pair2typin(* We adjust to the shortest decomposition *)letn=min(List.lengthdecomps1)(List.lengthdecomps2)inletdecomp1=List.nthdecomps1(n-1)inletdecomp2=List.nthdecomps2(n-1)in(* We rewrite dep_pair1 ... *)lete1_list,proj_list=List.splitdecomp1in(* ... and use dep_pair2 to compute the expected goal *)lete2_list,_=List.splitdecomp2in(* We build the expected goal *)letabst_B=List.fold_right(fun(e,t)body->lambda_createenvsigma(Sorts.Relevant,t,subst_termsigmaebody))e1_listbinletpred_body=beta_applistsigma(abst_B,proj_list)inletbody=mkApp(lambda_createenvsigma(Sorts.Relevant,typ,pred_body),[|dep_pair1|])inletexpected_goal=beta_applistsigma(abst_B,List.mapfste2_list)in(* Simulate now the normalisation treatment made by Logic.mk_refgoals *)letexpected_goal=nf_betaiotaenvsigmaexpected_goalin(* Retype to get universes right *)letsigma,expected_goal_ty=Typing.type_ofenvsigmaexpected_goalinletsigma,_=Typing.type_ofenvsigmabodyin(sigma,(body,expected_goal))(* Like "replace" but decompose dependent equalities *)(* i.e. if equality is "exists t v = exists u w", and goal is "phi(t,u)", *)(* then it uses the predicate "\x.phi(proj1_sig x,proj2_sig x)", and so *)(* on for further iterated sigma-tuples *)letcutSubstInConcll2reqn=Proofview.Goal.enterbeginfungl->letenv=Proofview.Goal.envglinletsigma=Proofview.Goal.sigmaglinlet(lbeq,u,(t,e1,e2))=pf_applyfind_eq_data_decomposegleqninlettyp=pf_conclglinlet(e1,e2)=ifl2rthen(e1,e2)else(e2,e1)inlet(sigma,(typ,expected))=subst_tuple_termenvsigmae1e2typintclTHEN(Proofview.Unsafe.tclEVARSsigma)(tclTHENFIRST(tclTHENLIST[(change_concltyp);(* Put in pattern form *)(replace_coreonConcll2reqn)])(change_conclexpected))(* Put in normalized form *)endletcutSubstInHypl2reqnid=Proofview.Goal.enterbeginfungl->letenv=Proofview.Goal.envglinletsigma=Proofview.Goal.sigmaglinlet(lbeq,u,(t,e1,e2))=pf_applyfind_eq_data_decomposegleqninlettyp=pf_get_hyp_typidglinlet(e1,e2)=ifl2rthen(e1,e2)else(e2,e1)inlet(sigma,(typ,expected))=subst_tuple_termenvsigmae1e2typintclTHEN(Proofview.Unsafe.tclEVARSsigma)(tclTHENFIRST(tclTHENLIST[(change_in_hyp~check:trueNone(make_change_argtyp)(id,InHypTypeOnly));(replace_core(onHypid)l2reqn)])(change_in_hyp~check:trueNone(make_change_argexpected)(id,InHypTypeOnly)))endlettry_rewritetac=Proofview.tclORELSEtacbeginfunction(e,info)->matchewith|Constr_matching.PatternMatchingFailure->tclZEROMSG~info(str"Not a primitive equality here.")|e->(* XXX: absorbing anomalies?? *)tclZEROMSG~info(strbrk"Cannot find a well-typed generalization of the goal that makes the proof progress.")endletcutSubstClausel2reqncls=matchclswith|None->cutSubstInConcll2reqn|Someid->cutSubstInHypl2reqnidletwarn_deprecated_cutrewrite=CWarnings.create~name:"deprecated-cutrewrite"~category:"deprecated"(fun()->strbrk"\"cutrewrite\" is deprecated. Use \"replace\" instead.")letcutRewriteClausel2reqncls=warn_deprecated_cutrewrite();try_rewrite(cutSubstClausel2reqncls)letcutRewriteInHypl2reqnid=cutRewriteClausel2reqn(Someid)letcutRewriteInConcll2reqn=cutRewriteClausel2reqnNoneletsubstClausel2rccls=Proofview.Goal.enterbeginfungl->leteq=pf_applyget_type_ofglcintclTHENS(cutSubstClausel2reqcls)[Proofview.tclUNIT();exact_no_checkc]endletrewriteClausel2rccls=try_rewrite(substClausel2rccls)letrewriteInHypl2rcid=rewriteClausel2rc(Someid)letrewriteInConcll2rc=rewriteClausel2rcNone(* Naming scheme for rewrite and cutrewrite tactics
give equality give proof of equality
/ cutSubstClause substClause
raw | cutSubstInHyp substInHyp
\ cutSubstInConcl substInConcl
/ cutRewriteClause rewriteClause
user| cutRewriteInHyp rewriteInHyp
\ cutRewriteInConcl rewriteInConcl
raw = raise typing error or PatternMatchingFailure
user = raise user error specific to rewrite
*)(**********************************************************************)(* Substitutions tactics (JCF) *)letregular_subst_tactic=reftruelet()=declare_bool_option{optdepr=true;optkey=["Regular";"Subst";"Tactic"];optread=(fun()->!regular_subst_tactic);optwrite=(:=)regular_subst_tactic}letrestrict_to_eq_and_identityeq=(* compatibility *)ifnot(Constr.isRefX(lib_ref"core.eq.type")eq)&¬(Constr.isRefX(lib_ref"core.identity.type")eq)thenraiseConstr_matching.PatternMatchingFailureexceptionFoundHypof(Id.t*constr*bool)(* tests whether hyp [c] is [x = t] or [t = x], [x] not occurring in [t] *)letis_eq_xglxd=letid=NamedDecl.get_iddintryletis_varidc=matchEConstr.kind(projectgl)cwith|Varid'->Id.equalidid'|_->falseinletc=pf_nf_evargl(NamedDecl.get_typed)inlet(_,lhs,rhs)=pi3(pf_applyfind_eq_data_decomposeglc)inif(is_varxlhs)&¬(local_occur_var(projectgl)xrhs)thenraise(FoundHyp(id,rhs,true));if(is_varxrhs)&¬(local_occur_var(projectgl)xlhs)thenraise(FoundHyp(id,lhs,false))withConstr_matching.PatternMatchingFailure->()exceptionFoundDepInGlobalofId.toption*GlobRef.tlettest_non_indirectly_dependent_section_variableglx=letenv=Proofview.Goal.envglinletsigma=Tacmach.New.projectglinlethyps=Proofview.Goal.hypsglinletconcl=Proofview.Goal.conclglinList.iter(fundecl->NamedDecl.iter_constr(func->matchoccur_var_indirectlyenvsigmaxcwith|Somegr->raise(FoundDepInGlobal(Some(NamedDecl.get_iddecl),gr))|None->())decl)hyps;matchoccur_var_indirectlyenvsigmaxconclwith|Somegr->raise(FoundDepInGlobal(None,gr))|None->()letcheck_non_indirectly_dependent_section_variableglx=trytest_non_indirectly_dependent_section_variableglxwithFoundDepInGlobal(pos,gr)->letwhere=matchposwith|Someid->str"hypothesis "++Id.printid|None->str"the conclusion of the goal"inuser_err~hdr:"Subst"(strbrk"Section variable "++Id.printx++strbrk" occurs implicitly in global declaration "++Printer.pr_globalgr++strbrk" present in "++where++strbrk".")letis_non_indirectly_dependent_section_variableglz=trytest_non_indirectly_dependent_section_variableglz;truewithFoundDepInGlobal(pos,gr)->false(* Rewrite "hyp:x=rhs" or "hyp:rhs=x" (if dir=false) everywhere and
erase hyp and x; proceed by generalizing all dep hyps *)letsubst_onedep_proof_okx(hyp,rhs,dir)=Proofview.Goal.enterbeginfungl->letsigma=Tacmach.New.projectglinlethyps=Proofview.Goal.hypsglinletconcl=Proofview.Goal.conclglin(* The set of hypotheses using x *)letdephyps=List.rev(pi3(List.fold_right(fundcl(dest,deps,allhyps)->letid=NamedDecl.get_iddclinifnot(Id.equalidhyp)&&List.exists(funy->local_occur_var_in_declsigmaydcl)depsthenletid_dest=if!regular_subst_tacticthendestelseMoveLastin(dest,id::deps,(id_dest,id)::allhyps)else(MoveBeforeid,deps,allhyps))hyps(MoveBeforex,[x],[])))in(* In practice, no dep hyps before x, so MoveBefore x is good enough *)(* Decides if x appears in conclusion *)letdepconcl=local_occur_varsigmaxconclinletneed_rewrite=not(List.is_emptydephyps)||depconclintclTHENLIST((ifneed_rewritethen[revert(List.mapsnddephyps);general_rewrite~where:None~l2r:dirAtLeastOneOccurrence~freeze:true~dep:dep_proof_ok~with_evars:false(mkVarhyp,NoBindings);(tclMAP(fun(dest,id)->intro_move(Someid)dest)dephyps)]else[Proofview.tclUNIT()])@[tclTRY(clear[x;hyp])])end(* Look for an hypothesis hyp of the form "x=rhs" or "rhs=x", rewrite
it everywhere, and erase hyp and x; proceed by generalizing all dep hyps *)letsubst_one_vardep_proof_okx=Proofview.Goal.enterbeginfungl->letdecl=pf_get_hypxglin(* If x has a body, simply replace x with body and clear x *)ifis_local_defdeclthentclTHEN(unfold_bodyx)(clear[x])else(* Find a non-recursive definition for x *)letres=try(* [is_eq_x] ensures nf_evar on its side *)lethyps=Proofview.Goal.hypsglinlettesthyp_=is_eq_xglxhypinContext.Named.fold_outsidetest~init:()hyps;user_err~hdr:"Subst"(str"Cannot find any non-recursive equality over "++Id.printx++str".")withFoundHypres->resinifis_section_variablexthencheck_non_indirectly_dependent_section_variableglx;subst_onedep_proof_okxresendletsubst_gendep_proof_okids=tclMAP(subst_one_vardep_proof_ok)ids(* For every x, look for an hypothesis hyp of the form "x=rhs" or "rhs=x",
rewrite it everywhere, and erase hyp and x; proceed by generalizing
all dep hyps *)letsubst=subst_gentruetypesubst_tactic_flags={only_leibniz:bool;rewrite_dependent_proof:bool}letdefault_subst_tactic_flags={only_leibniz=false;rewrite_dependent_proof=true}letwarn_deprecated_simple_subst=CWarnings.create~name:"deprecated-simple-subst"~category:"deprecated"(fun()->strbrk"\"simple subst\" is deprecated")letsubst_all?(flags=default_subst_tactic_flags)()=let()=ifflags.only_leibniz||notflags.rewrite_dependent_proofthenwarn_deprecated_simple_subst()inif!regular_subst_tacticthen(* Find hypotheses to treat in linear time *)letprocesshyp=Proofview.Goal.enterbeginfungl->letenv=Proofview.Goal.envglinletsigma=projectglinletc=pf_get_hyphypgl|>NamedDecl.get_typeintryletlbeq,u,(_,x,y)=pf_applyfind_eq_data_decomposeglcinletu=EInstance.kindsigmauinleteq=Constr.mkRef(lbeq.eq,u)inifflags.only_leibnizthenrestrict_to_eq_and_identityeq;matchEConstr.kindsigmax,EConstr.kindsigmaywith|Varx,VarywhenId.equalxy->Proofview.tclUNIT()|Varx',_whennot(Termops.local_occur_varsigmax'y)&¬(is_evaluableenv(EvalVarRefx'))&&is_non_indirectly_dependent_section_variableglx'->subst_oneflags.rewrite_dependent_proofx'(hyp,y,true)|_,Vary'whennot(Termops.local_occur_varsigmay'x)&¬(is_evaluableenv(EvalVarRefy'))&&is_non_indirectly_dependent_section_variablegly'->subst_oneflags.rewrite_dependent_proofy'(hyp,x,false)|_->Proofview.tclUNIT()withConstr_matching.PatternMatchingFailure->Proofview.tclUNIT()endinProofview.Goal.enterbeginfungl->tclMAPprocess(List.rev(List.mapNamedDecl.get_id(Proofview.Goal.hypsgl)))endelse(* Old implementation, not able to manage configurations like a=b, a=t,
or situations like "a = S b, b = S a", or also accidentally unfolding
let-ins *)Proofview.Goal.enterbeginfungl->letsigma=projectglinletfind_eq_data_decompose=pf_applyfind_eq_data_decomposeglinlettest(_,c)=tryletlbeq,u,(_,x,y)=find_eq_data_decomposecinletu=EInstance.kindsigmauinleteq=Constr.mkRef(lbeq.eq,u)inifflags.only_leibnizthenrestrict_to_eq_and_identityeq;(* J.F.: added to prevent failure on goal containing x=x as an hyp *)ifEConstr.eq_constrsigmaxythenfailwith"caught";matchEConstr.kindsigmaxwithVarx->x|_->matchEConstr.kindsigmaywithVary->y|_->failwith"caught"withConstr_matching.PatternMatchingFailure->failwith"caught"inlettestp=trySome(testp)withFailure_->Noneinlethyps=pf_hyps_typesglinletids=List.map_filtertesthypsinletids=List.uniquizeidsinsubst_genflags.rewrite_dependent_proofidsend(* Rewrite the first assumption for which a condition holds
and gives the direction of the rewrite *)letcond_eq_term_leftctgl=trylet(_,x,_)=pi3(pf_applyfind_eq_data_decomposeglt)inifpf_conv_xglcxthentrueelsefailwith"not convertible"withConstr_matching.PatternMatchingFailure->failwith"not an equality"letcond_eq_term_rightctgl=trylet(_,_,x)=pi3(pf_applyfind_eq_data_decomposeglt)inifpf_conv_xglcxthenfalseelsefailwith"not convertible"withConstr_matching.PatternMatchingFailure->failwith"not an equality"letcond_eq_termctgl=trylet(_,x,y)=pi3(pf_applyfind_eq_data_decomposeglt)inifpf_conv_xglcxthentrueelseifpf_conv_xglcythenfalseelsefailwith"not convertible"withConstr_matching.PatternMatchingFailure->failwith"not an equality"letrewrite_assumption_condcond_eq_termcl=letrecarechypsgl=matchhypswith|[]->user_errPp.(str"No such assumption.")|hyp::rest->letid=NamedDecl.get_idhypinbegintryletdir=cond_eq_term(NamedDecl.get_typehyp)glingeneral_rewrite_clausedirfalse(mkVarid,NoBindings)clwith|Failure_|UserError_->arecrestglendinProofview.Goal.enterbeginfungl->lethyps=Proofview.Goal.hypsglinarechypsglend(* Generalize "subst x" to substitution of subterm appearing as an
equation in the context, but not clearing the hypothesis *)letreplace_termdir_optc=letcond_eq_fun=matchdir_optwith|None->cond_eq_termc|Sometrue->cond_eq_term_leftc|Somefalse->cond_eq_term_rightcinrewrite_assumption_condcond_eq_fun(* Declare rewriting tactic for intro patterns "<-" and "->" *)let()=letgmrl2rwith_evarstacc=general_rewrite_clausel2rwith_evarstaccinHook.setTactics.general_rewrite_clausegmrlet()=Hook.setTactics.subst_onesubst_one