123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309(************************************************************************)(* * 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) *)(************************************************************************)openEqualityopenNamesopenPpopenConstropenTermopsopenCErrorsopenUtilopenMod_substopenLocus(* Rewriting rules *)typerew_rule={rew_lemma:constr;rew_type:types;rew_pat:constr;rew_ctx:Univ.ContextSet.t;rew_l2r:bool;rew_tac:Genarg.glob_generic_argumentoption}letsubst_hintsubsthint=letcst'=subst_mpssubsthint.rew_lemmainlettyp'=subst_mpssubsthint.rew_typeinletpat'=subst_mpssubsthint.rew_patinlett'=Option.Smart.map(Genintern.generic_substitutesubst)hint.rew_tacinifhint.rew_lemma==cst'&&hint.rew_type==typ'&&hint.rew_tac==t'thenhintelse{hintwithrew_lemma=cst';rew_type=typ';rew_pat=pat';rew_tac=t'}moduleHintIdent=structtypet=int*rew_ruleletcompare(i,t)(j,t')=i-jletsubsts(i,t)=(i,subst_hintst)letconstr_of(i,t)=t.rew_patendmoduleHintOpt=structletreducec=cletdirection=trueendmoduleHintDN=Term_dnet.Make(HintIdent)(HintOpt)(* Summary and Object declaration *)letrewtab=Summary.ref(String.Map.empty:HintDN.tString.Map.t)~name:"autorewrite"letraw_find_basebas=String.Map.findbas!rewtabletfind_basebas=tryraw_find_basebaswithNot_found->user_err~hdr:"AutoRewrite"(str"Rewriting base "++strbas++str" does not exist.")letfind_rewritesbas=List.rev_mapsnd(HintDN.find_all(find_basebas))letfind_matchesbaspat=letbase=find_basebasinletres=HintDN.search_patternbasepatinList.mapsndresletprint_rewrite_hintdbbas=letenv=Global.env()inletsigma=Evd.from_envenvin(str"Database "++strbas++fnl()++prlist_with_sepfnl(funh->str(ifh.rew_l2rthen"rewrite -> "else"rewrite <- ")++Printer.pr_lconstr_envenvsigmah.rew_lemma++str" of type "++Printer.pr_lconstr_envenvsigmah.rew_type++Option.cata(funtac->str" then use tactic "++Pputils.pr_glb_genericenvsigmatac)(mt())h.rew_tac)(find_rewritesbas))typeraw_rew_rule=(constrUniv.in_universe_context_set*bool*Genarg.raw_generic_argumentoption)CAst.t(* Applies all the rules of one base *)letone_basegeneral_rewrite_maybe_intac_mainbas=letlrul=find_rewritesbasinlettry_rewritedirctxctc=Proofview.Goal.enterbeginfungl->letsigma=Proofview.Goal.sigmaglinletsubst,ctx'=UnivGen.fresh_universe_context_set_instancectxinletc'=Vars.subst_univs_level_constrsubstcinletsigma=Evd.merge_context_setEvd.univ_flexiblesigmactx'inProofview.tclTHEN(Proofview.Unsafe.tclEVARSsigma)(general_rewrite_maybe_indirc'tc)endinletopenProofview.NotationsinProofview.tclProofInfo[@ocaml.warning"-3"]>>=fun(_name,poly)->letlrul=List.map(funh->lettac=matchh.rew_tacwith|None->Proofview.tclUNIT()|Some(Genarg.GenArg(Genarg.Glbwitwit,tac))->letist={Geninterp.lfun=Id.Map.empty;poly;extra=Geninterp.TacStore.empty}inFtactic.run(Geninterp.interpwitisttac)(fun_->Proofview.tclUNIT())in(h.rew_ctx,h.rew_lemma,h.rew_l2r,tac))lrulinTacticals.New.tclREPEAT_MAIN(Proofview.tclPROGRESS(List.fold_left(funtac(ctx,csr,dir,tc)->Tacticals.New.tclTHENtac(Tacticals.New.tclREPEAT_MAIN(Tacticals.New.tclTHENFIRST(try_rewritedirctxcsrtc)tac_main)))(Proofview.tclUNIT())lrul))(* The AutoRewrite tactic *)letautorewrite?(conds=Naive)tac_mainlbas=Tacticals.New.tclREPEAT_MAIN(Proofview.tclPROGRESS(List.fold_left(funtacbas->Tacticals.New.tclTHENtac(one_base(fundirctac->lettac=(tac,conds)ingeneral_rewrite~where:None~l2r:dirAllOccurrences~freeze:true~dep:false~with_evars:false~tac(EConstr.of_constrc,Tactypes.NoBindings))tac_mainbas))(Proofview.tclUNIT())lbas))letautorewrite_multi_in?(conds=Naive)idltac_mainlbas=Proofview.Goal.enterbeginfungl->(* let's check at once if id exists (to raise the appropriate error) *)let_=List.map(funid->Tacmach.New.pf_get_hypidgl)idlinletgeneral_rewrite_iniddircstrtac=letcstr=EConstr.of_constrcstringeneral_rewrite~where:(Someid)~l2r:dirAllOccurrences~freeze:true~dep:false~with_evars:false~tac:(tac,conds)(cstr,Tactypes.NoBindings)inTacticals.New.tclMAP(funid->Tacticals.New.tclREPEAT_MAIN(Proofview.tclPROGRESS(List.fold_left(funtacbas->Tacticals.New.tclTHENtac(one_base(general_rewrite_inid)tac_mainbas))(Proofview.tclUNIT())lbas)))idlendletautorewrite_in?(conds=Naive)id=autorewrite_multi_in~conds[id]letgen_auto_multi_rewritecondstac_mainlbascl=lettry_do_hypstreat_idl=autorewrite_multi_in~conds(List.maptreat_idl)tac_mainlbasinifnot(Locusops.is_all_occurrencescl.concl_occs)&&cl.concl_occs!=NoOccurrencesthenletinfo=Exninfo.reify()inTacticals.New.tclZEROMSG~info(str"The \"at\" syntax isn't available yet for the autorewrite tactic.")elseletcompose_tact1t2=matchcl.onhypswith|Some[]->t1|_->Tacticals.New.tclTHENFIRSTt1t2incompose_tac(ifcl.concl_occs!=NoOccurrencesthenautorewrite~condstac_mainlbaselseProofview.tclUNIT())(matchcl.onhypswith|Somel->try_do_hyps(fun((_,id),_)->id)l|None->(* try to rewrite in all hypothesis
(except maybe the rewritten one) *)Proofview.Goal.enterbeginfungl->letids=Tacmach.New.pf_ids_of_hypsglintry_do_hyps(funid->id)idsend)letauto_multi_rewrite?(conds=Naive)lemscl=Proofview.V82.wrap_exceptions(fun()->gen_auto_multi_rewriteconds(Proofview.tclUNIT())lemscl)letauto_multi_rewrite_with?(conds=Naive)tac_mainlbascl=letonconcl=matchcl.Locus.concl_occswithNoOccurrences->false|_->trueinmatchonconcl,cl.Locus.onhypswith|false,Some[_]|true,Some[]|false,Some[]->(* autorewrite with .... in clause using tac n'est sur que
si clause represente soit le but soit UNE hypothese
*)Proofview.V82.wrap_exceptions(fun()->gen_auto_multi_rewritecondstac_mainlbascl)|_->letinfo=Exninfo.reify()inTacticals.New.tclZEROMSG~info(strbrk"autorewrite .. in .. using can only be used either with a unique hypothesis or on the conclusion.")(* Functions necessary to the library object declaration *)letcache_hintrewrite(_,(rbase,lrl))=letbase=tryraw_find_baserbasewithNot_found->HintDN.emptyinletmax=tryfst(Util.List.last(HintDN.find_allbase))withFailure_->0inletlrl=HintDN.refresh_metaslrlinletlrl=HintDN.map(fun(i,h)->(i+max,h))lrlinrewtab:=String.Map.addrbase(HintDN.unionlrlbase)!rewtabletsubst_hintrewrite(subst,(rbase,listasnode))=letlist'=HintDN.substsubstlistiniflist'==listthennodeelse(rbase,list')(* Declaration of the Hint Rewrite library object *)letinGlobalHintRewrite:string*HintDN.t->Libobject.obj=letopenLibobjectindeclare_object@@superglobal_object_nodischarge"HINT_REWRITE_GLOBAL"~cache:cache_hintrewrite~subst:(Somesubst_hintrewrite)letinExportHintRewrite:string*HintDN.t->Libobject.obj=letopenLibobjectindeclare_object@@global_object_nodischarge"HINT_REWRITE_EXPORT"~cache:cache_hintrewrite~subst:(Somesubst_hintrewrite)typehypinfo={hyp_ty:EConstr.types;hyp_pat:EConstr.constr;}letdecompose_applied_relationenvsigmacctypeleft2right=letfind_relty=(* FIXME: this is nonsense, we generate evars and then we drop the
corresponding evarmap. This sometimes works because [Term_dnet] performs
evar surgery via [Termops.filtering]. *)letsigma,ty=Clenv.make_evar_clauseenvsigmatyinlet(_,args)=Termops.decompose_app_vectsigmaty.Clenv.cl_conclinletlen=Array.lengthargsinif2<=lenthenletc1=args.(len-2)inletc2=args.(len-1)inSome(ifleft2rightthenc1elsec2)elseNoneinmatchfind_relctypewith|Somec->Some{hyp_pat=c;hyp_ty=ctype}|None->letctx,t'=Reductionops.splay_prod_assumenvsigmactypein(* Search for underlying eq *)letctype=it_mkProd_or_LetInt'ctxinmatchfind_relctypewith|Somec->Some{hyp_pat=c;hyp_ty=ctype}|None->Noneletfind_applied_relation?locenvsigmacleft2right=letctype=Retyping.get_type_ofenvsigma(EConstr.of_constrc)inmatchdecompose_applied_relationenvsigmacctypeleft2rightwith|Somec->c|None->user_err?loc~hdr:"decompose_applied_relation"(str"The type"++spc()++Printer.pr_econstr_envenvsigmactype++spc()++str"of this term does not end with an applied relation.")letwarn_deprecated_hint_rewrite_without_locality=CWarnings.create~name:"deprecated-hint-rewrite-without-locality"~category:"deprecated"(fun()->strbrk"The default value for rewriting hint locality is currently \
\"local\" in a section and \"global\" otherwise, but is scheduled to change \
in a future release. For the time being, adding rewriting hints outside of sections \
without specifying an explicit locality attribute is therefore deprecated. It is \
recommended to use \"export\" whenever possible. Use the attributes \
#[local], #[global] and #[export] depending on your choice. For example: \
\"#[export] Hint Rewrite foo : bar.\"")(* To add rewriting rules to a base *)letadd_rew_rules~localitybaselrul=letcounter=ref0inletenv=Global.env()inletsigma=Evd.from_envenvinletist=Genintern.empty_glob_sign(Global.env())inletinterntac=snd(Genintern.generic_internisttac)inletlrul=List.fold_left(fundn{CAst.loc;v=((c,ctx),b,t)}->letsigma=Evd.merge_context_setEvd.univ_rigidsigmactxinletinfo=find_applied_relation?locenvsigmacbinletpat=EConstr.Unsafe.to_constrinfo.hyp_patinletrul={rew_lemma=c;rew_type=EConstr.Unsafe.to_constrinfo.hyp_ty;rew_pat=pat;rew_ctx=ctx;rew_l2r=b;rew_tac=Option.mapinternt}inincrcounter;HintDN.addpat(!counter,rul)dn)HintDN.emptylrulinletopenGoptionsinmatchlocalitywith|OptLocal->cache_hintrewrite((),(base,lrul))|OptDefault->let()=ifnot@@Global.sections_are_opened()thenwarn_deprecated_hint_rewrite_without_locality()inLib.add_anonymous_leaf(inGlobalHintRewrite(base,lrul))|OptGlobal->let()=ifGlobal.sections_are_opened()thenCErrors.user_errPp.(str"This command does not support the global attribute in sections.");inLib.add_anonymous_leaf(inGlobalHintRewrite(base,lrul))|OptExport->let()=ifGlobal.sections_are_opened()thenCErrors.user_errPp.(str"This command does not support the export attribute in sections.");inLib.add_anonymous_leaf(inExportHintRewrite(base,lrul))