123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193(************************************************************************)(* * 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) *)(************************************************************************)openUtilopenPpopenNamesopenTacexpr(** Nametab for tactics *)(** TODO: Share me somewhere *)moduleFullPath=structopenLibnamestypet=full_pathletequal=eq_full_pathletto_string=string_of_pathletreprsp=letdir,id=repr_pathspinid,(DirPath.reprdir)endmoduleKnTab=Nametab.Make(FullPath)(KerName)lettactic_tab=Summary.ref~name:"LTAC-NAMETAB"(KnTab.empty,KNmap.empty)letpush_tacticvisspkn=let(tab,revtab)=!tactic_tabinlettab=KnTab.pushvisspkntabinletrevtab=KNmap.addknsprevtabintactic_tab:=(tab,revtab)letlocate_tacticqid=KnTab.locateqid(fst!tactic_tab)letlocate_extended_all_tacticqid=KnTab.find_prefixesqid(fst!tactic_tab)letexists_tactickn=KnTab.existskn(fst!tactic_tab)letpath_of_tactickn=KNmap.findkn(snd!tactic_tab)letshortest_qualid_of_tactickn=letsp=KNmap.findkn(snd!tactic_tab)inKnTab.shortest_qualidId.Set.emptysp(fst!tactic_tab)(** Tactic notations (TacAlias) *)typealias=KerName.ttypealias_tactic={alias_args:Id.tlist;alias_body:glob_tactic_expr;alias_deprecation:Deprecation.toption;}letalias_map=Summary.ref~name:"tactic-alias"(KNmap.empty:alias_tacticKNmap.t)letregister_aliaskeytac=alias_map:=KNmap.addkeytac!alias_mapletinterp_aliaskey=tryKNmap.findkey!alias_mapwithNot_found->CErrors.anomaly(str"Unknown tactic alias: "++KerName.printkey++str".")letcheck_aliaskey=KNmap.memkey!alias_map(** ML tactic extensions (TacML) *)typeml_tactic=Geninterp.Val.tlist->Geninterp.interp_sign->unitProofview.tacticmoduleMLName=structtypet=ml_tactic_nameletcomparetac1tac2=letc=String.comparetac1.mltac_tactictac2.mltac_tacticinifc=0thenString.comparetac1.mltac_plugintac2.mltac_pluginelsecendmoduleMLTacMap=Map.Make(MLName)letpr_tacnamet=strt.mltac_plugin++str"::"++strt.mltac_tacticlettac_tab=refMLTacMap.emptyletregister_ml_tactic?(overwrite=false)s(t:ml_tacticarray)=let()=ifMLTacMap.mems!tac_tabthenifoverwritethentac_tab:=MLTacMap.removes!tac_tabelseCErrors.anomaly(str"Cannot redeclare tactic "++pr_tacnames++str".")intac_tab:=MLTacMap.addst!tac_tabletinterp_ml_tactic{mltac_name=s;mltac_index=i}=trylettacs=MLTacMap.finds!tac_tabinlet()=ifArray.lengthtacs<=ithenraiseNot_foundintacs.(i)withNot_found->CErrors.user_err(str"The tactic "++pr_tacnames++str" is not installed.")(***************************************************************************)(* Tactic registration *)(* Summary and Object declaration *)openLibobjecttypeltac_entry={tac_for_ml:bool;tac_body:glob_tactic_expr;tac_redef:ModPath.tlist;tac_deprecation:Deprecation.toption}letmactab=Summary.ref(KNmap.empty:ltac_entryKNmap.t)~name:"tactic-definition"letltac_entries()=!mactabletinterp_ltacr=(KNmap.findr!mactab).tac_bodyletis_ltac_for_ml_tacticr=(KNmap.findr!mactab).tac_for_mlletadd~deprecationknbt=letentry={tac_for_ml=b;tac_body=t;tac_redef=[];tac_deprecation=deprecation;}inmactab:=KNmap.addknentry!mactabletreplaceknpatht=letpath=KerName.modpathpathinletentry_e={ewithtac_body=t;tac_redef=path::e.tac_redef}inmactab:=KNmap.modifyknentry!mactablettac_deprecationkn=try(KNmap.findkn!mactab).tac_deprecationwithNot_found->Noneletload_mdi((sp,kn),(local,id,b,t,deprecation))=matchidwith|None->let()=ifnotlocalthenpush_tactic(Nametab.Untili)spkninadd~deprecationknbt|Somekn0->replacekn0kntletopen_mdi((sp,kn),(local,id,b,t,deprecation))=matchidwith|None->let()=ifnotlocalthenpush_tactic(Nametab.Exactlyi)spkninadd~deprecationknbt|Somekn0->replacekn0kntletcache_md((sp,kn),(local,id,b,t,deprecation))=matchidwith|None->let()=push_tactic(Nametab.Until1)spkninadd~deprecationknbt|Somekn0->replacekn0kntletsubst_kindsubstid=matchidwith|None->None|Somekn->Some(Mod_subst.subst_knsubstkn)letsubst_md(subst,(local,id,b,t,deprecation))=(local,subst_kindsubstid,b,Tacsubst.subst_tacticsubstt,deprecation)letclassify_md(local,_,_,_,_aso)=SubstituteoletinMD:bool*ltac_constantoption*bool*glob_tactic_expr*Deprecation.toption->obj=declare_object{(default_object"TAC-DEFINITION")withcache_function=cache_md;load_function=load_md;open_function=simple_openopen_md;subst_function=subst_md;classify_function=classify_md}letregister_ltacfor_mllocal?deprecationidtac=ignore(Lib.add_leafid(inMD(local,None,for_ml,tac,deprecation)))letredefine_ltaclocal?deprecationkntac=Lib.add_anonymous_leaf(inMD(local,Somekn,false,tac,deprecation))