123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691692693694695696697698699700701702703704705706707708709710711712713714715716717718719720721722723724725726727728729730731732733734735736737738739740741742743744745746747748749750751752753754755756757758759760761762763764765766767768769770771772773774775776777778779780781782783784785786787788789790791792793794795796797798799800801802803804805806807808809810811812813814815816817818819820821822823824825826827828829830831832833834835836837838839840841842843844845846847848849850851852853854855856857858859860861862863864865866867868869870871872873874875876877878879880881882883884885886887888889890891892893894895896897898899900901902903904905906907908909910911912913914915916917918919920921922923924925926927928929930931932933934935936937938939940941942943944945946947948949950951952953954955956957958959960961962963964965966967968969970971972973974975976977978979980981982983984985986987988989990991992993994995996997998999100010011002100310041005100610071008100910101011101210131014101510161017101810191020102110221023102410251026102710281029103010311032103310341035103610371038103910401041104210431044104510461047104810491050105110521053105410551056105710581059106010611062106310641065106610671068106910701071107210731074107510761077107810791080108110821083108410851086108710881089109010911092109310941095109610971098109911001101110211031104110511061107110811091110111111121113111411151116111711181119112011211122112311241125112611271128112911301131113211331134113511361137113811391140114111421143114411451146114711481149115011511152115311541155115611571158115911601161116211631164116511661167116811691170117111721173117411751176117711781179118011811182118311841185118611871188118911901191119211931194119511961197119811991200120112021203120412051206120712081209121012111212121312141215121612171218121912201221122212231224122512261227122812291230123112321233123412351236123712381239124012411242124312441245124612471248124912501251125212531254125512561257125812591260126112621263126412651266126712681269127012711272127312741275127612771278127912801281128212831284128512861287128812891290129112921293129412951296129712981299130013011302130313041305130613071308130913101311131213131314131513161317131813191320132113221323132413251326132713281329133013311332133313341335133613371338133913401341134213431344134513461347134813491350135113521353135413551356135713581359136013611362136313641365136613671368136913701371137213731374137513761377137813791380138113821383138413851386138713881389139013911392139313941395139613971398139914001401140214031404140514061407140814091410141114121413141414151416141714181419142014211422142314241425142614271428142914301431143214331434143514361437143814391440144114421443144414451446144714481449145014511452145314541455145614571458145914601461146214631464146514661467146814691470147114721473147414751476147714781479148014811482148314841485148614871488148914901491149214931494149514961497149814991500150115021503150415051506150715081509151015111512151315141515151615171518151915201521152215231524152515261527152815291530153115321533153415351536153715381539154015411542154315441545154615471548154915501551155215531554155515561557155815591560156115621563156415651566156715681569157015711572157315741575157615771578157915801581158215831584158515861587158815891590159115921593159415951596159715981599160016011602160316041605160616071608160916101611161216131614161516161617161816191620162116221623162416251626162716281629163016311632163316341635163616371638163916401641164216431644164516461647164816491650165116521653165416551656165716581659166016611662166316641665166616671668166916701671167216731674167516761677167816791680168116821683168416851686168716881689169016911692169316941695169616971698169917001701170217031704170517061707170817091710171117121713171417151716171717181719172017211722172317241725172617271728172917301731173217331734173517361737173817391740174117421743174417451746174717481749175017511752175317541755175617571758175917601761176217631764176517661767176817691770177117721773177417751776177717781779178017811782178317841785178617871788178917901791179217931794179517961797179817991800180118021803180418051806180718081809181018111812181318141815181618171818181918201821182218231824182518261827182818291830183118321833183418351836183718381839184018411842184318441845184618471848184918501851185218531854185518561857185818591860186118621863186418651866186718681869(************************************************************************)(* * 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) *)(************************************************************************)openUtilopenPpopenNamesopenGenargopenTac2envopenTac2expropenTac2entries.PltacopenProofview.Notationsletconstr_flags=letopenPretypingin{use_coercions=true;use_typeclasses=Pretyping.UseTC;solve_unification_constraints=true;fail_evar=true;expand_evars=true;program_mode=false;polymorphic=false;}letopen_constr_no_classes_flags=letopenPretypingin{use_coercions=true;use_typeclasses=Pretyping.NoUseTC;solve_unification_constraints=true;fail_evar=false;expand_evars=true;program_mode=false;polymorphic=false;}(** Standard values *)moduleValue=Tac2ffiopenValueletval_format=Tac2print.val_formatletformat=repr_extval_formatletcore_prefixpathn=KerName.makepath(Label.of_id(Id.of_string_softn))letstd_coren=core_prefixTac2env.std_prefixnletcoq_coren=core_prefixTac2env.coq_prefixnletltac1_coren=core_prefixTac2env.ltac1_prefixnmoduleCore=structlett_int=coq_core"int"lett_string=coq_core"string"lett_array=coq_core"array"lett_unit=coq_core"unit"lett_list=coq_core"list"lett_constr=coq_core"constr"lett_pattern=coq_core"pattern"lett_ident=coq_core"ident"lett_option=coq_core"option"lett_exn=coq_core"exn"lett_reference=std_core"reference"lett_ltac1=ltac1_core"t"letc_nil=coq_core"[]"letc_cons=coq_core"::"letc_none=coq_core"None"letc_some=coq_core"Some"letc_true=coq_core"true"letc_false=coq_core"false"endopenCoreletv_unit=Value.of_unit()letv_blk=Valexpr.make_blockletof_binderb=Value.of_extValue.val_binderbletto_binderb=Value.to_extValue.val_binderbletof_instanceu=letu=Univ.Instance.to_array(EConstr.Unsafe.to_instanceu)inValue.of_array(funv->Value.of_extValue.val_univv)uletto_instanceu=letu=Value.to_array(funv->Value.to_extValue.val_univv)uinEConstr.EInstance.make(Univ.Instance.of_arrayu)letof_rec_declaration(nas,ts,cs)=letbinders=Array.map2(funnat->(na,t))nastsin(Value.of_arrayof_binderbinders,Value.of_arrayValue.of_constrcs)letto_rec_declaration(nas,cs)=letnas=Value.to_arrayto_bindernasin(Array.mapfstnas,Array.mapsndnas,Value.to_arrayValue.to_constrcs)letof_case_invert=letopenConstrinfunction|NoInvert->ValInt0|CaseInvert{indices}->v_blk0[|of_arrayof_constrindices|]letto_case_invert=letopenConstrinfunction|ValInt0->NoInvert|ValBlk(0,[|indices|])->letindices=to_arrayto_constrindicesinCaseInvert{indices}|_->CErrors.anomalyPp.(str"unexpected value shape")letof_resultf=function|Inlc->v_blk0[|fc|]|Inre->v_blk1[|Value.of_exne|](** Stdlib exceptions *)leterr_notfocussed=Tac2interp.LtacError(coq_core"Not_focussed",[||])leterr_outofbounds=Tac2interp.LtacError(coq_core"Out_of_bounds",[||])leterr_notfound=Tac2interp.LtacError(coq_core"Not_found",[||])leterr_matchfailure=Tac2interp.LtacError(coq_core"Match_failure",[||])leterr_division_by_zero=Tac2interp.LtacError(coq_core"Division_by_zero",[||])(** Helper functions *)letthawf=Tac2ffi.applyf[v_unit]letfatal_flag:unitExninfo.t=Exninfo.make()letset_btinfo=if!Tac2interp.print_ltac2_backtracethenTac2interp.get_backtrace>>=funbt->Proofview.tclUNIT(Exninfo.addinfoTac2entries.backtracebt)elseProofview.tclUNITinfoletthrow?(info=Exninfo.null)e=set_btinfo>>=funinfo->letinfo=Exninfo.addinfofatal_flag()inProofview.tclLIFT(Proofview.NonLogical.raise(e,info))letfail?(info=Exninfo.null)e=set_btinfo>>=funinfo->Proofview.tclZERO~infoeletreturnx=Proofview.tclUNITxletpnames={mltac_plugin="coq-core.plugins.ltac2";mltac_tactic=s}letwrapf=return()>>=fun()->return(f())letwrap_unitf=return()>>=fun()->f();returnv_unitletassert_focussed=Proofview.Goal.goals>>=fungls->matchglswith|[_]->Proofview.tclUNIT()|[]|_::_::_->throwerr_notfocussedletpf_applyf=Proofview.Goal.goals>>=function|[]->Proofview.tclENV>>=funenv->Proofview.tclEVARMAP>>=funsigma->fenvsigma|[gl]->gl>>=fungl->f(Proofview.Goal.envgl)(Tacmach.projectgl)|_::_::_->throwerr_notfocussed(** Primitives *)letdefine_primitivenamearityf=Tac2env.define_primitive(pnamename)(mk_closurearityf)letdefine0namef=define_primitivenamearity_one(fun_->f)letdefine1namer0f=define_primitivenamearity_onebeginfunx->f(Value.repr_tor0x)endletdefine2namer0r1f=define_primitivename(arity_sucarity_one)beginfunxy->f(Value.repr_tor0x)(Value.repr_tor1y)endletdefine3namer0r1r2f=define_primitivename(arity_suc(arity_sucarity_one))beginfunxyz->f(Value.repr_tor0x)(Value.repr_tor1y)(Value.repr_tor2z)endletdefine4namer0r1r2r3f=define_primitivename(arity_suc(arity_suc(arity_sucarity_one)))beginfunx0x1x2x3->f(Value.repr_tor0x0)(Value.repr_tor1x1)(Value.repr_tor2x2)(Value.repr_tor3x3)endletdefine5namer0r1r2r3r4f=define_primitivename(arity_suc(arity_suc(arity_suc(arity_sucarity_one))))beginfunx0x1x2x3x4->f(Value.repr_tor0x0)(Value.repr_tor1x1)(Value.repr_tor2x2)(Value.repr_tor3x3)(Value.repr_tor4x4)end(** Printing *)let()=define1"print"ppbeginfunpp->wrap_unit(fun()->Feedback.msg_noticepp)endlet()=define1"message_of_int"intbeginfunn->return(Value.of_pp(Pp.intn))endlet()=define1"message_of_string"stringbeginfuns->return(Value.of_pp(str(Bytes.to_strings)))endlet()=define1"message_of_constr"constrbeginfunc->pf_applybeginfunenvsigma->letpp=Printer.pr_econstr_envenvsigmacinreturn(Value.of_pppp)endendlet()=define1"message_of_ident"identbeginfunc->letpp=Id.printcinreturn(Value.of_pppp)endlet()=define1"message_of_exn"valexprbeginfunv->Proofview.tclENV>>=funenv->Proofview.tclEVARMAP>>=funsigma->letpp=Tac2print.pr_valexprenvsigmav(GTypRef(OtherCore.t_exn,[]))inreturn(Value.of_pppp)endlet()=define2"message_concat"ppppbeginfunm1m2->return(Value.of_pp(Pp.appm1m2))endlet()=define0"format_stop"beginreturn(Value.of_extval_format[])endlet()=define1"format_string"formatbeginfuns->return(Value.of_extval_format(Tac2print.FmtString::s))endlet()=define1"format_int"formatbeginfuns->return(Value.of_extval_format(Tac2print.FmtInt::s))endlet()=define1"format_constr"formatbeginfuns->return(Value.of_extval_format(Tac2print.FmtConstr::s))endlet()=define1"format_ident"formatbeginfuns->return(Value.of_extval_format(Tac2print.FmtIdent::s))endlet()=define2"format_literal"stringformatbeginfunlits->return(Value.of_extval_format(Tac2print.FmtLiteral(Bytes.to_stringlit)::s))endlet()=define1"format_alpha"formatbeginfuns->return(Value.of_extval_format(Tac2print.FmtAlpha::s))endlet()=define2"format_kfprintf"closureformatbeginfunkfmt->letopenTac2printinletfoldaccu=function|FmtLiteral_->accu|FmtString|FmtInt|FmtConstr|FmtIdent->1+accu|FmtAlpha->2+accuinletpop1l=matchlwith[]->assertfalse|x::l->(x,l)inletpop2l=matchlwith[]|[_]->assertfalse|x::y::l->(x,y,l)inletarity=List.fold_leftfold0fmtinletrecevalaccuargsfmt=matchfmtwith|[]->applyk[of_ppaccu]|tag::fmt->matchtagwith|FmtLiterals->eval(Pp.appaccu(Pp.strs))argsfmt|FmtString->let(s,args)=pop1argsinletpp=Pp.str(Bytes.to_string(to_strings))ineval(Pp.appaccupp)argsfmt|FmtInt->let(i,args)=pop1argsinletpp=Pp.int(to_inti)ineval(Pp.appaccupp)argsfmt|FmtConstr->let(c,args)=pop1argsinletc=to_constrcinpf_applybeginfunenvsigma->letpp=Printer.pr_econstr_envenvsigmacineval(Pp.appaccupp)argsfmtend|FmtIdent->let(i,args)=pop1argsinletpp=Id.print(to_identi)ineval(Pp.appaccupp)argsfmt|FmtAlpha->let(f,x,args)=pop2argsinTac2ffi.apply(to_closuref)[of_unit();x]>>=funpp->eval(Pp.appaccu(to_pppp))argsfmtinletevalv=eval(Pp.mt())vfmtinifInt.equalarity0theneval[]elsereturn(Tac2ffi.of_closure(Tac2ffi.abstractarityeval))end(** Array *)let()=define0"array_empty"beginreturn(v_blk0(Array.of_list[]))endlet()=define2"array_make"intvalexprbeginfunnx->ifn<0||n>Sys.max_array_lengththenthrowerr_outofboundselsewrap(fun()->v_blk0(Array.makenx))endlet()=define1"array_length"blockbeginfun(_,v)->return(Value.of_int(Array.lengthv))endlet()=define3"array_set"blockintvalexprbeginfun(_,v)nx->ifn<0||n>=Array.lengthvthenthrowerr_outofboundselsewrap_unit(fun()->v.(n)<-x)endlet()=define2"array_get"blockintbeginfun(_,v)n->ifn<0||n>=Array.lengthvthenthrowerr_outofboundselsewrap(fun()->v.(n))endlet()=define5"array_blit"blockintblockintintbeginfun(_,v0)s0(_,v1)s1l->ifs0<0||s0+l>Array.lengthv0||s1<0||s1+l>Array.lengthv1||l<0thenthrowerr_outofboundselsewrap_unit(fun()->Array.blitv0s0v1s1l)endlet()=define4"array_fill"blockintintvalexprbeginfun(_,d)slv->ifs<0||s+l>Array.lengthd||l<0thenthrowerr_outofboundselsewrap_unit(fun()->Array.filldslv)endlet()=define1"array_concat"(listblock)beginfunl->wrap(fun()->v_blk0(Array.concat(List.mapsndl)))end(** Ident *)let()=define2"ident_equal"identidentbeginfunid1id2->return(Value.of_bool(Id.equalid1id2))endlet()=define1"ident_to_string"identbeginfunid->return(Value.of_string(Bytes.of_string(Id.to_stringid)))endlet()=define1"ident_of_string"stringbeginfuns->letid=trySome(Id.of_string(Bytes.to_strings))with_->Noneinreturn(Value.of_optionValue.of_identid)end(** Int *)let()=define2"int_equal"intintbeginfunmn->return(Value.of_bool(m==n))endletunopnf=define1nintbeginfunm->return(Value.of_int(fm))endletbinopnf=define2nintintbeginfunmn->return(Value.of_int(fmn))endlet()=binop"int_compare"Int.comparelet()=binop"int_add"(+)let()=binop"int_sub"(-)let()=binop"int_mul"(*)let()=define2"int_div"intintbeginfunmn->ifn==0thenthrowerr_division_by_zeroelsereturn(Value.of_int(m/n))endlet()=define2"int_mod"intintbeginfunmn->ifn==0thenthrowerr_division_by_zeroelsereturn(Value.of_int(mmodn))endlet()=unop"int_neg"(~-)let()=unop"int_abs"abslet()=binop"int_asr"(asr)let()=binop"int_lsl"(lsl)let()=binop"int_lsr"(lsr)let()=binop"int_land"(land)let()=binop"int_lor"(lor)let()=binop"int_lxor"(lxor)let()=unop"int_lnot"lnot(** Char *)let()=define1"char_of_int"intbeginfunn->wrap(fun()->Value.of_char(Char.chrn))endlet()=define1"char_to_int"charbeginfunn->wrap(fun()->Value.of_int(Char.coden))end(** String *)let()=define2"string_make"intcharbeginfunnc->ifn<0||n>Sys.max_string_lengththenthrowerr_outofboundselsewrap(fun()->Value.of_string(Bytes.makenc))endlet()=define1"string_length"stringbeginfuns->return(Value.of_int(Bytes.lengths))endlet()=define3"string_set"stringintcharbeginfunsnc->ifn<0||n>=Bytes.lengthsthenthrowerr_outofboundselsewrap_unit(fun()->Bytes.setsnc)endlet()=define2"string_get"stringintbeginfunsn->ifn<0||n>=Bytes.lengthsthenthrowerr_outofboundselsewrap(fun()->Value.of_char(Bytes.getsn))end(** Terms *)(** constr -> constr *)let()=define1"constr_type"constrbeginfunc->letget_typeenvsigma=Proofview.wrap_exceptionsbeginfun()->let(sigma,t)=Typing.type_ofenvsigmacinlett=Value.of_constrtinProofview.Unsafe.tclEVARSsigma<*>Proofview.tclUNITtendinpf_applyget_typeend(** constr -> constr *)let()=define2"constr_equal"constrconstrbeginfunc1c2->Proofview.tclEVARMAP>>=funsigma->letb=EConstr.eq_constrsigmac1c2inProofview.tclUNIT(Value.of_boolb)endlet()=define1"constr_kind"constrbeginfunc->letopenConstrinProofview.tclEVARMAP>>=funsigma->Proofview.tclENV>>=funenv->returnbeginmatchEConstr.kindsigmacwith|Reln->v_blk0[|Value.of_intn|]|Varid->v_blk1[|Value.of_identid|]|Metan->v_blk2[|Value.of_intn|]|Evar(evk,args)->v_blk3[|Value.of_int(Evar.reprevk);Value.of_arrayValue.of_constr(Array.of_listargs);|]|Sorts->v_blk4[|Value.of_extValue.val_sorts|]|Cast(c,k,t)->v_blk5[|Value.of_constrc;Value.of_extValue.val_castk;Value.of_constrt;|]|Prod(na,t,u)->v_blk6[|of_binder(na,t);Value.of_constru;|]|Lambda(na,t,c)->v_blk7[|of_binder(na,t);Value.of_constrc;|]|LetIn(na,b,t,c)->v_blk8[|of_binder(na,t);Value.of_constrb;Value.of_constrc;|]|App(c,cl)->v_blk9[|Value.of_constrc;Value.of_arrayValue.of_constrcl;|]|Const(cst,u)->v_blk10[|Value.of_constantcst;of_instanceu;|]|Ind(ind,u)->v_blk11[|Value.of_extValue.val_inductiveind;of_instanceu;|]|Construct(cstr,u)->v_blk12[|Value.of_extValue.val_constructorcstr;of_instanceu;|]|Case(ci,u,pms,c,iv,t,bl)->(* FIXME: also change representation Ltac2-side? *)let(ci,c,iv,t,bl)=EConstr.expand_caseenvsigma(ci,u,pms,c,iv,t,bl)inv_blk13[|Value.of_extValue.val_caseci;Value.of_constrc;of_case_invertiv;Value.of_constrt;Value.of_arrayValue.of_constrbl;|]|Fix((recs,i),def)->let(nas,cs)=of_rec_declarationdefinv_blk14[|Value.of_arrayValue.of_intrecs;Value.of_inti;nas;cs;|]|CoFix(i,def)->let(nas,cs)=of_rec_declarationdefinv_blk15[|Value.of_inti;nas;cs;|]|Proj(p,c)->v_blk16[|Value.of_extValue.val_projectionp;Value.of_constrc;|]|Intn->v_blk17[|Value.of_uint63n|]|Floatf->v_blk18[|Value.of_floatf|]|Array(u,t,def,ty)->v_blk19[|of_instanceu;Value.of_arrayValue.of_constrt;Value.of_constrdef;Value.of_constrty|]endendlet()=define1"constr_make"valexprbeginfunknd->Proofview.tclEVARMAP>>=funsigma->Proofview.tclENV>>=funenv->letc=matchTac2ffi.to_blockkndwith|(0,[|n|])->letn=Value.to_intninEConstr.mkReln|(1,[|id|])->letid=Value.to_identidinEConstr.mkVarid|(2,[|n|])->letn=Value.to_intninEConstr.mkMetan|(3,[|evk;args|])->letevk=Evar.unsafe_of_int(Value.to_intevk)inletargs=Value.to_arrayValue.to_constrargsinEConstr.mkEvar(evk,Array.to_listargs)|(4,[|s|])->lets=Value.to_extValue.val_sortsinEConstr.mkSort(EConstr.Unsafe.to_sortss)|(5,[|c;k;t|])->letc=Value.to_constrcinletk=Value.to_extValue.val_castkinlett=Value.to_constrtinEConstr.mkCast(c,k,t)|(6,[|na;u|])->let(na,t)=to_bindernainletu=Value.to_construinEConstr.mkProd(na,t,u)|(7,[|na;c|])->let(na,t)=to_bindernainletu=Value.to_constrcinEConstr.mkLambda(na,t,u)|(8,[|na;b;c|])->let(na,t)=to_bindernainletb=Value.to_constrbinletc=Value.to_constrcinEConstr.mkLetIn(na,b,t,c)|(9,[|c;cl|])->letc=Value.to_constrcinletcl=Value.to_arrayValue.to_constrclinEConstr.mkApp(c,cl)|(10,[|cst;u|])->letcst=Value.to_constantcstinletu=to_instanceuinEConstr.mkConstU(cst,u)|(11,[|ind;u|])->letind=Value.to_extValue.val_inductiveindinletu=to_instanceuinEConstr.mkIndU(ind,u)|(12,[|cstr;u|])->letcstr=Value.to_extValue.val_constructorcstrinletu=to_instanceuinEConstr.mkConstructU(cstr,u)|(13,[|ci;c;iv;t;bl|])->letci=Value.to_extValue.val_caseciinletc=Value.to_constrcinletiv=to_case_invertivinlett=Value.to_constrtinletbl=Value.to_arrayValue.to_constrblinEConstr.mkCase(EConstr.contract_caseenvsigma(ci,c,iv,t,bl))|(14,[|recs;i;nas;cs|])->letrecs=Value.to_arrayValue.to_intrecsinleti=Value.to_intiinletdef=to_rec_declaration(nas,cs)inEConstr.mkFix((recs,i),def)|(15,[|i;nas;cs|])->leti=Value.to_intiinletdef=to_rec_declaration(nas,cs)inEConstr.mkCoFix(i,def)|(16,[|p;c|])->letp=Value.to_extValue.val_projectionpinletc=Value.to_constrcinEConstr.mkProj(p,c)|(17,[|n|])->letn=Value.to_uint63ninEConstr.mkIntn|(18,[|f|])->letf=Value.to_floatfinEConstr.mkFloatf|(19,[|u;t;def;ty|])->lett=Value.to_arrayValue.to_constrtinletdef=Value.to_constrdefinletty=Value.to_constrtyinletu=to_instanceuinEConstr.mkArray(u,t,def,ty)|_->assertfalseinreturn(Value.of_constrc)endlet()=define1"constr_check"constrbeginfunc->pf_applybeginfunenvsigma->trylet(sigma,_)=Typing.type_ofenvsigmacinProofview.Unsafe.tclEVARSsigma>>=fun()->return(of_resultValue.of_constr(Inlc))withewhenCErrors.noncriticale->lete=Exninfo.captureeinreturn(of_resultValue.of_constr(Inre))endendlet()=define3"constr_substnl"(listconstr)intconstrbeginfunsubstkc->letans=EConstr.Vars.substnlsubstkcinreturn(Value.of_constrans)endlet()=define3"constr_closenl"(listident)intconstrbeginfunidskc->letans=EConstr.Vars.substn_varskidscinreturn(Value.of_constrans)endlet()=define1"constr_case"(repr_extval_inductive)beginfunind->Proofview.tclENV>>=funenv->tryletans=Inductiveops.make_case_infoenvindSorts.RelevantConstr.RegularStyleinreturn(Value.of_extValue.val_caseans)withewhenCErrors.noncriticale->throwerr_notfoundendlet()=define2"constr_constructor"(repr_extval_inductive)intbeginfun(ind,i)k->Proofview.tclENV>>=funenv->tryletopenDeclarationsinletans=Environ.lookup_mindindenvinlet_=ans.mind_packets.(i).mind_consnames.(k)inreturn(Value.of_extval_constructor((ind,i),(k+1)))withewhenCErrors.noncriticale->throwerr_notfoundendlet()=define3"constr_in_context"identconstrclosurebeginfunidtc->Proofview.Goal.goals>>=function|[gl]->gl>>=fungl->letenv=Proofview.Goal.envglinletsigma=Proofview.Goal.sigmaglinlethas_var=trylet_=Environ.lookup_named_validenvintruewithNot_found->falseinifhas_varthenTacticals.tclZEROMSG(str"Variable already exists")elseletopenContext.Named.Declarationinletnenv=EConstr.push_named(LocalAssum(Context.make_annotidSorts.Relevant,t))envinlet(sigma,(evt,_))=Evarutil.new_type_evarnenvsigmaEvd.univ_flexibleinlet(sigma,evk)=Evarutil.new_pure_evar(Environ.named_context_valnenv)sigmaevtinProofview.Unsafe.tclEVARSsigma>>=fun()->Proofview.Unsafe.tclSETGOALS[Proofview.with_empty_stateevk]>>=fun()->thawc>>=fun_->Proofview.Unsafe.tclSETGOALS[Proofview.with_empty_state(Proofview.Goal.goalgl)]>>=fun()->letargs=List.map(fund->EConstr.mkVar(get_idd))(EConstr.named_contextenv)inletargs=EConstr.mkRel1::argsinletans=EConstr.mkEvar(evk,args)inletans=EConstr.mkLambda(Context.make_annot(Nameid)Sorts.Relevant,t,ans)inreturn(Value.of_constrans)|_->throwerr_notfocussedend(** preterm -> constr *)let()=define1"constr_pretype"(repr_extval_preterm)beginfunc->letopenPretypinginletopenLtac_pretypeinletpretypeenvsigma=Proofview.wrap_exceptionsbeginfun()->(* For now there are no primitives to create preterms with a non-empty
closure. I do not know whether [closed_glob_constr] is really the type
we want but it does not hurt in the meantime. *)let{closure;term}=cinletvars={ltac_constrs=closure.typed;ltac_uconstrs=closure.untyped;ltac_idents=closure.idents;ltac_genargs=Id.Map.empty;}inletflags=constr_flagsinletsigma,t=understand_ltacflagsenvsigmavarsWithoutTypeConstraintterminlett=Value.of_constrtinProofview.Unsafe.tclEVARSsigma<*>Proofview.tclUNITtendinpf_applypretypeendlet()=define2"constr_binder_make"(optionident)constrbeginfunnaty->pf_applybeginfunenvsigma->letrel=Retyping.relevance_of_typeenvsigmatyinletna=matchnawithNone->Anonymous|Someid->Nameidinreturn(Value.of_extval_binder(Context.make_annotnarel,ty))endendlet()=define1"constr_binder_name"(repr_extval_binder)beginfun(bnd,_)->letna=matchbnd.Context.binder_namewithAnonymous->None|Nameid->Someidinreturn(Value.of_optionValue.of_identna)endlet()=define1"constr_binder_type"(repr_extval_binder)beginfun(bnd,ty)->return(of_constrty)endlet()=define1"constr_has_evar"constrbeginfunc->Proofview.tclEVARMAP>>=funsigma->letb=Evarutil.has_undefined_evarssigmacinProofview.tclUNIT(Value.of_boolb)end(** Patterns *)letempty_context=Constr_matching.empty_contextlet()=define0"pattern_empty_context"beginreturn(Value.of_extval_matching_contextempty_context)endlet()=define2"pattern_matches"patternconstrbeginfunpatc->pf_applybeginfunenvsigma->letans=trySome(Constr_matching.matchesenvsigmapatc)withConstr_matching.PatternMatchingFailure->Noneinbeginmatchanswith|None->failerr_matchfailure|Someans->letans=Id.Map.bindingsansinletof_pair(id,c)=Value.of_tuple[|Value.of_identid;Value.of_constrc|]inreturn(Value.of_listof_pairans)endendendlet()=define2"pattern_matches_subterm"patternconstrbeginfunpatc->letopenConstr_matchinginletrecof_anss=matchIStream.peekswith|IStream.Nil->failerr_matchfailure|IStream.Cons({m_sub=(_,sub);m_ctx},s)->letans=Id.Map.bindingssubinletof_pair(id,c)=Value.of_tuple[|Value.of_identid;Value.of_constrc|]inletans=Value.of_tuple[|Value.of_extval_matching_contextm_ctx;Value.of_listof_pairans|]inProofview.tclOR(returnans)(fun_->of_anss)inpf_applybeginfunenvsigma->letpat=Constr_matching.instantiate_patternenvsigmaId.Map.emptypatinletans=Constr_matching.match_subtermenvsigma(Id.Set.empty,pat)cinof_ansansendendlet()=define2"pattern_matches_vect"patternconstrbeginfunpatc->pf_applybeginfunenvsigma->letans=trySome(Constr_matching.matchesenvsigmapatc)withConstr_matching.PatternMatchingFailure->Noneinbeginmatchanswith|None->failerr_matchfailure|Someans->letans=Id.Map.bindingsansinletans=Array.map_of_listsndansinreturn(Value.of_arrayValue.of_constrans)endendendlet()=define2"pattern_matches_subterm_vect"patternconstrbeginfunpatc->letopenConstr_matchinginletrecof_anss=matchIStream.peekswith|IStream.Nil->failerr_matchfailure|IStream.Cons({m_sub=(_,sub);m_ctx},s)->letans=Id.Map.bindingssubinletans=Array.map_of_listsndansinletans=Value.of_tuple[|Value.of_extval_matching_contextm_ctx;Value.of_arrayValue.of_constrans|]inProofview.tclOR(returnans)(fun_->of_anss)inpf_applybeginfunenvsigma->letpat=Constr_matching.instantiate_patternenvsigmaId.Map.emptypatinletans=Constr_matching.match_subtermenvsigma(Id.Set.empty,pat)cinof_ansansendendlet()=define3"pattern_matches_goal"bool(list(pairboolpattern))(pairboolpattern)beginfunrevhpcp->assert_focussed>>=fun()->Proofview.Goal.enter_onebeginfungl->letenv=Proofview.Goal.envglinletsigma=Proofview.Goal.sigmaglinletconcl=Proofview.Goal.conclglinletmk_pattern(b,pat)=ifbthenTac2match.MatchPatternpatelseTac2match.MatchContextpatinletr=(List.mapmk_patternhp,mk_patterncp)inTac2match.match_goalenvsigmaconcl~revr>>=fun(hyps,ctx,subst)->letof_ctxoptctx=Value.of_extval_matching_context(Option.defaultempty_contextctx)inlethids=Value.of_arrayValue.of_ident(Array.map_of_listfsthyps)inlethctx=Value.of_arrayof_ctxopt(Array.map_of_listsndhyps)inletsubs=Value.of_arrayValue.of_constr(Array.map_of_listsnd(Id.Map.bindingssubst))inletcctx=of_ctxoptctxinletans=Value.of_tuple[|hids;hctx;subs;cctx|]inProofview.tclUNITansendendlet()=define2"pattern_instantiate"(repr_extval_matching_context)constrbeginfunctxc->letans=Constr_matching.instantiate_contextctxcinreturn(Value.of_constrans)end(** Error *)let()=define1"throw"exnbeginfun(e,info)->throw~infoeend(** Control *)(** exn -> 'a *)let()=define1"zero"exnbeginfun(e,info)->fail~infoeend(** (unit -> 'a) -> (exn -> 'a) -> 'a *)let()=define2"plus"closureclosurebeginfunxk->Proofview.tclOR(thawx)(fune->Tac2ffi.applyk[Value.of_exne])end(** (unit -> 'a) -> 'a *)let()=define1"once"closurebeginfunf->Proofview.tclONCE(thawf)end(** (unit -> unit) list -> unit *)let()=define1"dispatch"(listclosure)beginfunl->letl=List.map(funf->Proofview.tclIGNORE(thawf))linProofview.tclDISPATCHl>>=fun()->returnv_unitend(** (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit *)let()=define3"extend"(listclosure)closure(listclosure)beginfunlfttacrgt->letlft=List.map(funf->Proofview.tclIGNORE(thawf))lftinlettac=Proofview.tclIGNORE(thawtac)inletrgt=List.map(funf->Proofview.tclIGNORE(thawf))rgtinProofview.tclEXTENDlfttacrgt>>=fun()->returnv_unitend(** (unit -> unit) -> unit *)let()=define1"enter"closurebeginfunf->letf=Proofview.tclIGNORE(thawf)inProofview.tclINDEPENDENTf>>=fun()->returnv_unitend(** (unit -> 'a) -> ('a * ('exn -> 'a)) result *)let()=define1"case"closurebeginfunf->Proofview.tclCASE(thawf)>>=beginfunction|Proofview.Next(x,k)->letk=Tac2ffi.mk_closurearity_onebeginfune->let(e,info)=Value.to_exneinset_btinfo>>=funinfo->k(e,info)endinreturn(v_blk0[|Value.of_tuple[|x;Value.of_closurek|]|])|Proofview.Faile->return(v_blk1[|Value.of_exne|])endend(** int -> int -> (unit -> 'a) -> 'a *)let()=define3"focus"intintclosurebeginfunijtac->Proofview.tclFOCUSij(thawtac)end(** unit -> unit *)let()=define0"shelve"beginProofview.shelve>>=fun()->returnv_unitend(** unit -> unit *)let()=define0"shelve_unifiable"beginProofview.shelve_unifiable>>=fun()->returnv_unitendlet()=define1"new_goal"intbeginfunev->letev=Evar.unsafe_of_intevinProofview.tclEVARMAP>>=funsigma->ifEvd.memsigmaevthenProofview.Unsafe.tclNEWGOALS[Proofview.with_empty_stateev]<*>Proofview.tclUNITv_unitelsethrowerr_notfoundend(** unit -> constr *)let()=define0"goal"beginassert_focussed>>=fun()->Proofview.Goal.enter_onebeginfungl->letconcl=Tacmach.pf_nf_conclglinreturn(Value.of_constrconcl)endend(** ident -> constr *)let()=define1"hyp"identbeginfunid->pf_applybeginfunenv_->letmem=tryignore(Environ.lookup_namedidenv);truewithNot_found->falseinifmemthenreturn(Value.of_constr(EConstr.mkVarid))elseTacticals.tclZEROMSG(str"Hypothesis "++quote(Id.printid)++str" not found")(* FIXME: Do something more sensible *)endendlet()=define0"hyps"beginpf_applybeginfunenv_->letopenContextinletopenNamed.Declarationinlethyps=List.rev(Environ.named_contextenv)inletmap=function|LocalAssum(id,t)->lett=EConstr.of_constrtinValue.of_tuple[|Value.of_identid.binder_name;Value.of_optionValue.of_constrNone;Value.of_constrt|]|LocalDef(id,c,t)->letc=EConstr.of_constrcinlett=EConstr.of_constrtinValue.of_tuple[|Value.of_identid.binder_name;Value.of_optionValue.of_constr(Somec);Value.of_constrt|]inreturn(Value.of_listmaphyps)endend(** (unit -> constr) -> unit *)let()=define1"refine"closurebeginfunc->letc=thawc>>=func->Proofview.tclUNIT((),Value.to_constrc)inProofview.Goal.enterbeginfungl->Refine.generic_refine~typecheck:truecglend>>=fun()->returnv_unitendlet()=define2"with_holes"closureclosurebeginfunxf->Proofview.tclEVARMAP>>=funsigma0->thawx>>=funans->Proofview.tclEVARMAP>>=funsigma->Proofview.Unsafe.tclEVARSsigma0>>=fun()->Tacticals.tclWITHHOLESfalse(Tac2ffi.applyf[ans])sigmaendlet()=define1"progress"closurebeginfunf->Proofview.tclPROGRESS(thawf)endlet()=define2"abstract"(optionident)closurebeginfunidf->Abstract.tclABSTRACTid(Proofview.tclIGNORE(thawf))>>=fun()->returnv_unitendlet()=define2"time"(optionstring)closurebeginfunsf->lets=Option.mapBytes.to_stringsinProofview.tclTIMEs(thawf)endlet()=define0"check_interrupt"beginProofview.tclCHECKINTERRUPT>>=fun()->returnv_unitend(** Fresh *)let()=define2"fresh_free_union"(repr_extval_free)(repr_extval_free)beginfunset1set2->letans=Id.Set.unionset1set2inreturn(Value.of_extValue.val_freeans)endlet()=define1"fresh_free_of_ids"(listident)beginfunids->letfree=List.fold_rightId.Set.addidsId.Set.emptyinreturn(Value.of_extValue.val_freefree)endlet()=define1"fresh_free_of_constr"constrbeginfunc->Proofview.tclEVARMAP>>=funsigma->letrecfoldaccuc=matchEConstr.kindsigmacwith|Constr.Varid->Id.Set.addidaccu|_->EConstr.foldsigmafoldaccucinletans=foldId.Set.emptycinreturn(Value.of_extValue.val_freeans)endlet()=define2"fresh_fresh"(repr_extval_free)identbeginfunavoidid->letnid=Namegen.next_ident_away_fromid(funid->Id.Set.memidavoid)inreturn(Value.of_identnid)end(** Env *)let()=define1"env_get"(listident)beginfunids->letr=matchidswith|[]->None|_::_asids->let(id,path)=List.sep_lastidsinletpath=DirPath.make(List.revpath)inletfp=Libnames.make_pathpathidintrySome(Nametab.global_of_pathfp)withNot_found->Noneinreturn(Value.of_optionValue.of_referencer)endlet()=define1"env_expand"(listident)beginfunids->letr=matchidswith|[]->[]|_::_asids->let(id,path)=List.sep_lastidsinletpath=DirPath.make(List.revpath)inletqid=Libnames.make_qualidpathidinNametab.locate_allqidinreturn(Value.of_listValue.of_referencer)endlet()=define1"env_path"referencebeginfunr->matchNametab.path_of_globalrwith|fp->let(path,id)=Libnames.repr_pathfpinletpath=DirPath.reprpathinreturn(Value.of_listValue.of_ident(List.rev_appendpath[id]))|exceptionNot_found->throwerr_notfoundendlet()=define1"env_instantiate"referencebeginfunr->Proofview.tclENV>>=funenv->Proofview.tclEVARMAP>>=funsigma->let(sigma,c)=Evd.fresh_globalenvsigmarinProofview.Unsafe.tclEVARSsigma>>=fun()->return(Value.of_constrc)end(** Ind *)let()=define2"ind_equal"(repr_extval_inductive)(repr_extval_inductive)beginfunind1ind2->return(Value.of_bool(Ind.UserOrd.equalind1ind2))endlet()=define1"ind_data"(repr_extval_inductive)beginfunind->Proofview.tclENV>>=funenv->ifEnviron.mem_mind(fstind)envthenletmib=Environ.lookup_mind(fstind)envinreturn(Value.of_extval_ind_data(ind,mib))elsethrowerr_notfoundendlet()=define1"ind_repr"(repr_extval_ind_data)beginfun(ind,_)->return(Value.of_extval_inductiveind)endlet()=define1"ind_index"(repr_extval_inductive)beginfun(ind,n)->return(Value.of_intn)endlet()=define1"ind_nblocks"(repr_extval_ind_data)beginfun(ind,mib)->return(Value.of_int(Array.lengthmib.Declarations.mind_packets))endlet()=define1"ind_nconstructors"(repr_extval_ind_data)beginfun((_,n),mib)->letopenDeclarationsinreturn(Value.of_int(Array.lengthmib.mind_packets.(n).mind_consnames))endlet()=define2"ind_get_block"(repr_extval_ind_data)intbeginfun(ind,mib)n->if0<=n&&n<Array.lengthmib.Declarations.mind_packetsthenreturn(Value.of_extval_ind_data((fstind,n),mib))elsethrowerr_notfoundendlet()=define2"ind_get_constructor"(repr_extval_ind_data)intbeginfun((mind,n),mib)i->letopenDeclarationsinletncons=Array.lengthmib.mind_packets.(n).mind_consnamesinif0<=i&&i<nconsthen(* WARNING: In the ML API constructors are indexed from 1 for historical
reasons, but Ltac2 uses 0-indexing instead. *)return(Value.of_extval_constructor((mind,n),i+1))elsethrowerr_notfoundend(** Ltac1 in Ltac2 *)letltac1=Tac2ffi.repr_extValue.val_ltac1letof_ltac1v=Value.of_extValue.val_ltac1vlet()=define1"ltac1_ref"(listident)beginfunids->letopenLtac_plugininletr=matchidswith|[]->raiseNot_found|_::_asids->let(id,path)=List.sep_lastidsinletpath=DirPath.make(List.revpath)inletfp=Libnames.make_pathpathidinifTacenv.exists_tacticfpthenList.hd(Tacenv.locate_extended_all_tactic(Libnames.qualid_of_pathfp))elseraiseNot_foundinlettac=Tacinterp.Value.of_closure(Tacinterp.default_ist())(Tacenv.interp_ltacr)inreturn(Value.of_extval_ltac1tac)endlet()=define1"ltac1_run"ltac1beginfunv->letopenLtac_plugininTacinterp.tactic_of_value(Tacinterp.default_ist())v>>=fun()->returnv_unitendlet()=define3"ltac1_apply"ltac1(listltac1)closurebeginfunfargsk->letopenLtac_plugininletopenTacexprinletopenLocusinletkret=Proofview.tclIGNORE(Tac2ffi.applyk[Value.of_extval_ltac1ret])inletfoldarg(i,vars,lfun)=letid=Id.of_string("x"^string_of_inti)inletx=Reference(ArgVarCAst.(makeid))in(succi,x::vars,Id.Map.addidarglfun)inlet(_,args,lfun)=List.fold_rightfoldargs(0,[],Id.Map.empty)inletlfun=Id.Map.add(Id.of_string"F")flfuninletist={(Tacinterp.default_ist())withTacinterp.lfun=lfun;}inlettac=CAst.make@@TacArg(TacCall(CAst.make(ArgVarCAst.(make@@Id.of_string"F"),args)))inTacinterp.val_interpisttack>>=fun()->returnv_unitendlet()=define1"ltac1_of_constr"constrbeginfunc->letopenLtac_plugininreturn(Value.of_extval_ltac1(Tacinterp.Value.of_constrc))endlet()=define1"ltac1_to_constr"ltac1beginfunv->letopenLtac_plugininreturn(Value.of_optionValue.of_constr(Tacinterp.Value.to_constrv))endlet()=define1"ltac1_of_ident"identbeginfunc->letopenLtac_plugininreturn(Value.of_extval_ltac1(Taccoerce.Value.of_identc))endlet()=define1"ltac1_to_ident"ltac1beginfunv->letopenLtac_plugininreturn(Value.of_optionValue.of_ident(Taccoerce.Value.to_identv))endlet()=define1"ltac1_of_list"(listltac1)beginfunl->letopenGeninterp.Valinreturn(Value.of_extval_ltac1(inject(Basetyp_list)l))endlet()=define1"ltac1_to_list"ltac1beginfunv->letopenLtac_plugininreturn(Value.of_option(Value.of_listof_ltac1)(Tacinterp.Value.to_listv))end(** ML types *)(** Embed all Ltac2 data into Values *)letto_lvarist=letopenGlob_opsinletlfun=Tac2interp.set_envistId.Map.emptyin{empty_lvarwithLtac_pretype.ltac_genargs=lfun}letgtyprefkn=GTypRef(Otherkn,[])letintern_constrselfistc=let(_,(c,_))=Genintern.internStdarg.wit_constristcin(GlbValc,gtypreft_constr)letcatchable_exception=function|Logic_monad.Exception_->false|e->CErrors.noncriticaleletinterp_constrflagsistc=letopenPretypinginletist=to_lvaristinpf_applybeginfunenvsigma->trylet(sigma,c)=understand_ltacflagsenvsigmaistWithoutTypeConstraintcinletc=Value.of_constrcinProofview.Unsafe.tclEVARSsigma>>=fun()->Proofview.tclUNITcwithewhencatchable_exceptione->let(e,info)=Exninfo.captureeinset_btinfo>>=funinfo->matchExninfo.getinfofatal_flagwith|None->Proofview.tclZERO~infoe|Some()->throw~infoeendlet()=letintern=intern_constrinletinterpistc=interp_constrconstr_flagsistcinletprintenvsigmac=str"constr:("++Printer.pr_lglob_constr_envenvsigmac++str")"inletsubstsubstc=Detyping.subst_glob_constr(Global.env())substcinletobj={ml_intern=intern;ml_subst=subst;ml_interp=interp;ml_print=print;}indefine_ml_objectTac2quote.wit_constrobjlet()=letintern=intern_constrinletinterpistc=interp_constropen_constr_no_classes_flagsistcinletprintenvsigmac=str"open_constr:("++Printer.pr_lglob_constr_envenvsigmac++str")"inletsubstsubstc=Detyping.subst_glob_constr(Global.env())substcinletobj={ml_intern=intern;ml_subst=subst;ml_interp=interp;ml_print=print;}indefine_ml_objectTac2quote.wit_open_constrobjlet()=letinterp_id=return(Value.of_identid)inletprint__id=str"ident:("++Id.printid++str")"inletobj={ml_intern=(fun__id->GlbValid,gtypreft_ident);ml_interp=interp;ml_subst=(fun_id->id);ml_print=print;}indefine_ml_objectTac2quote.wit_identobjlet()=letinternselfistc=letenv=ist.Genintern.genvinletsigma=Evd.from_envenvinletwarn=if!Ltac_plugin.Tacintern.strict_checkthenfunx->xelseConstrintern.for_grammarinlet_,pat=warn(fun()->Constrintern.intern_constr_patternenvsigma~as_type:falsec)()inGlbValpat,gtypreft_patterninletsubstsubstc=letenv=Global.env()inletsigma=Evd.from_envenvinPatternops.subst_patternenvsigmasubstcinletprintenvsigmapat=str"pat:("++Printer.pr_lconstr_pattern_envenvsigmapat++str")"inletinterp_c=return(Value.of_patternc)inletobj={ml_intern=intern;ml_interp=interp;ml_subst=subst;ml_print=print;}indefine_ml_objectTac2quote.wit_patternobjlet()=letinterp_c=letopenLtac_pretypeinletclosure={idents=Id.Map.empty;typed=Id.Map.empty;untyped=Id.Map.empty;genargs=Id.Map.empty;}inletc={closure;term=c}inreturn(Value.of_extval_pretermc)inletsubstsubstc=Detyping.subst_glob_constr(Global.env())substcinletprintenvsigmac=str"preterm:("++Printer.pr_lglob_constr_envenvsigmac++str")"inletobj={ml_intern=(fun__e->Empty.aborte);ml_interp=interp;ml_subst=subst;ml_print=print;}indefine_ml_objectTac2quote.wit_pretermobjlet()=letinternselfistref=matchref.CAst.vwith|Tac2qexpr.QHypothesisid->GlbVal(GlobRef.VarRefid),gtypreft_reference|Tac2qexpr.QReferenceqid->letgr=tryNametab.locateqidwithNot_foundasexn->let_,info=Exninfo.captureexninNametab.error_global_not_found~infoqidinGlbValgr,gtypreft_referenceinletsubstsc=Globnames.subst_global_referencescinletinterp_gr=return(Value.of_referencegr)inletprint__=function|GlobRef.VarRefid->str"reference:("++str"&"++Id.printid++str")"|r->str"reference:("++Printer.pr_globalr++str")"inletobj={ml_intern=intern;ml_subst=subst;ml_interp=interp;ml_print=print;}indefine_ml_objectTac2quote.wit_referenceobjlet()=letinternselfist(ids,tac)=letmap{CAst.v=id}=idinletids=List.mapmapidsin(* Prevent inner calls to Ltac2 values *)letextra=Tac2intern.drop_ltac2_envist.Genintern.extrainletltacvars=List.fold_rightId.Set.addidsist.Genintern.ltacvarsinletist={istwithGenintern.extra;ltacvars}inlet_,tac=Genintern.internLtac_plugin.Tacarg.wit_tacticisttacinletfoldty_=GTypArrow(gtypreft_ltac1,ty)inletty=List.fold_leftfold(gtypreft_unit)idsinGlbVal(ids,tac),tyinletinterp_(ids,tac)=letclosargs=letaddlfunidv=letv=Tac2ffi.to_extval_ltac1vinId.Map.addidvlfuninletlfun=List.fold_left2addId.Map.emptyidsargsinletist={env_ist=Id.Map.empty}inletlfun=Tac2interp.set_envistlfuninletist=Ltac_plugin.Tacinterp.default_ist()inletist={istwithGeninterp.lfun=lfun}inlettac=(Ltac_plugin.Tacinterp.eval_tactic_ististtac:unitProofview.tactic)inletwrap(e,info)=set_btinfo>>=funinfo->Proofview.tclZERO~infoeinProofview.tclORtacwrap>>=fun()->returnv_unitinletlen=List.lengthidsinifInt.equallen0thenclos[]elsereturn(Tac2ffi.of_closure(Tac2ffi.abstractlenclos))inletsubsts(ids,tac)=(ids,Genintern.substituteLtac_plugin.Tacarg.wit_tacticstac)inletprintenvsigma(ids,tac)=letids=ifList.is_emptyidsthenmt()elsepr_sequenceId.printids++spc()++str"|-"++spc()instr"ltac1:("++ids++Ltac_plugin.Pptactic.pr_glob_tacticenvtac++str")"inletobj={ml_intern=intern;ml_subst=subst;ml_interp=interp;ml_print=print;}indefine_ml_objectTac2quote.wit_ltac1objlet()=letopenLtac_plugininletinternselfist(ids,tac)=letmap{CAst.v=id}=idinletids=List.mapmapidsin(* Prevent inner calls to Ltac2 values *)letextra=Tac2intern.drop_ltac2_envist.Genintern.extrainletltacvars=List.fold_rightId.Set.addidsist.Genintern.ltacvarsinletist={istwithGenintern.extra;ltacvars}inlet_,tac=Genintern.internLtac_plugin.Tacarg.wit_tacticisttacinletfoldty_=GTypArrow(gtypreft_ltac1,ty)inletty=List.fold_leftfold(gtypreft_ltac1)idsinGlbVal(ids,tac),tyinletinterp_(ids,tac)=letclosargs=letaddlfunidv=letv=Tac2ffi.to_extval_ltac1vinId.Map.addidvlfuninletlfun=List.fold_left2addId.Map.emptyidsargsinletist={env_ist=Id.Map.empty}inletlfun=Tac2interp.set_envistlfuninletist=Ltac_plugin.Tacinterp.default_ist()inletist={istwithGeninterp.lfun=lfun}inreturn(Value.of_extval_ltac1(Tacinterp.Value.of_closureisttac))inletlen=List.lengthidsinifInt.equallen0thenclos[]elsereturn(Tac2ffi.of_closure(Tac2ffi.abstractlenclos))inletsubsts(ids,tac)=(ids,Genintern.substituteTacarg.wit_tacticstac)inletprintenvsigma(ids,tac)=letids=ifList.is_emptyidsthenmt()elsepr_sequenceId.printids++str" |- "instr"ltac1val:("++ids++Ltac_plugin.Pptactic.pr_glob_tacticenvtac++str")"inletobj={ml_intern=intern;ml_subst=subst;ml_interp=interp;ml_print=print;}indefine_ml_objectTac2quote.wit_ltac1valobj(** Ltac2 in terms *)let()=letinterp?loc~polyenvsigmatycon(ids,tac)=(* Syntax prevents bound notation variables in constr quotations *)let()=assert(Id.Set.is_emptyids)inletist=Tac2interp.get_env@@GlobEnv.lfunenvinlettac=Proofview.tclIGNORE(Tac2interp.interpisttac)inletname,poly=Id.of_string"ltac2",polyinletsigma,concl=matchtyconwith|Somety->sigma,ty|None->GlobEnv.new_type_evarenvsigma~src:(loc,Evar_kinds.InternalHole)inletc,sigma=Proof.refine_by_tactic~name~poly(GlobEnv.renamed_envenv)sigmaconcltacinletj={Environ.uj_val=EConstr.of_constrc;Environ.uj_type=concl}in(j,sigma)inGlobEnv.register_constr_interp0wit_ltac2_constrinterplet()=letinterp?loc~polyenvsigmatyconid=letist=Tac2interp.get_env@@GlobEnv.lfunenvinletenv=GlobEnv.renamed_envenvinletc=Id.Map.findidist.env_istinletc=Value.to_constrcinlett=Retyping.get_type_ofenvsigmacinletj={Environ.uj_val=c;uj_type=t}inmatchtyconwith|None->j,sigma|Somety->letsigma=tryEvarconv.unify_leq_delayenvsigmattywithEvarconv.UnableToUnify(sigma,e)->Pretype_errors.error_actual_type?locenvsigmajtyein{jwithEnviron.uj_type=ty},sigmainGlobEnv.register_constr_interp0wit_ltac2_quotationinterplet()=letpr_rawid=Genprint.PrinterBasic(fun_env_sigma->mt())inletpr_glbid=Genprint.PrinterBasic(fun_env_sigma->str"$"++Id.printid)inletpr_top_=Genprint.TopPrinterBasicmtinGenprint.register_print0wit_ltac2_quotationpr_rawpr_glbpr_toplet()=letsubsglobs(ids,tac)=(* Let-bind the notation terms inside the tactic *)letfoldid(c,_)(rem,accu)=letc=GTacExt(Tac2quote.wit_preterm,c)inletrem=Id.Set.removeidreminrem,(Nameid,c)::accuinletrem,bnd=Id.Map.foldfoldglobs(ids,[])inlet()=ifnot@@Id.Set.is_emptyremthen(* FIXME: provide a reasonable middle-ground with the behaviour
introduced by 8d9b66b. We should be able to pass mere syntax to
term notation without facing the wrath of the internalization. *)letplural=ifId.Set.cardinalrem<=1then" "else"s "inCErrors.user_err(str"Missing notation term for variable"++strplural++pr_sequenceId.print(Id.Set.elementsrem)++str", probably an ill-typed expression")inlettac=ifList.is_emptybndthentacelseGTacLet(false,bnd,tac)in(Id.Set.empty,tac)inGenintern.register_ntn_subst0wit_ltac2_constrsubs(** Ltac2 in Ltac1 *)let()=lete=Tac2entries.Pltac.tac2expr_in_envinletinject(loc,v)=Ltac_plugin.Tacexpr.TacGeneric(Some"ltac2",in_gen(rawwitwit_ltac2)v)inLtac_plugin.Tacentries.create_ltac_quotation"ltac2"inject(e,None)(* Ltac1 runtime representation of Ltac2 closures. *)lettyp_ltac2:valexprGeninterp.Val.typ=Geninterp.Val.create"ltac2:ltac2_eval"letcast_typ(typea)(tag:aGeninterp.Val.typ)(v:Geninterp.Val.t):a=letGeninterp.Val.Dyn(tag',v)=vinmatchGeninterp.Val.eqtagtag'with|None->assertfalse|SomeRefl->vlet()=letopenLtac_pluginin(* This is a hack similar to Tacentries.ml_val_tactic_extend *)letintern_fun_e=Empty.aborteinletsubst_funsv=vinlet()=Genintern.register_intern0wit_ltac2_valintern_funinlet()=Genintern.register_subst0wit_ltac2_valsubst_funin(* These are bound names and not relevant *)lettac_id=Id.of_string"F"inletarg_id=Id.of_string"X"inletinterp_funist()=lettac=cast_typtyp_ltac2@@Id.Map.gettac_idist.Tacinterp.lfuninletarg=Id.Map.getarg_idist.Tacinterp.lfuninlettac=Tac2ffi.to_closuretacinTac2ffi.applytac[of_ltac1arg]>>=funans->letans=Tac2ffi.to_extval_ltac1ansinFtactic.returnansinlet()=Geninterp.register_interp0wit_ltac2_valinterp_funindefine1"ltac1_lambda"valexprbeginfunf->letbody=Tacexpr.TacGeneric(Some"coq-core.plugins.ltac2",in_gen(glbwitwit_ltac2_val)())inletclos=CAst.make(Tacexpr.TacFun([Namearg_id],CAst.make(Tacexpr.TacArgbody)))inletf=Geninterp.Val.inject(Geninterp.Val.Basetyp_ltac2)finletlfun=Id.Map.singletontac_idfinletist={(Tacinterp.default_ist())withTacinterp.lfun}inProofview.tclUNIT(of_ltac1(Tacinterp.Value.of_closureistclos))endletltac2_eval=letopenLtac_plugininletml_name={Tacexpr.mltac_plugin="coq-core.plugins.ltac2";mltac_tactic="ltac2_eval";}inleteval_funargsist=matchargswith|[]->assertfalse|tac::args->(* By convention the first argument is the tactic being applied, the rest
being the arguments it should be fed with *)lettac=cast_typtyp_ltac2tacinlettac=Tac2ffi.to_closuretacinletargs=List.map(funarg->Tac2ffi.of_extval_ltac1arg)argsinProofview.tclIGNORE(Tac2ffi.applytacargs)inlet()=Tacenv.register_ml_tacticml_name[|eval_fun|]in{Tacexpr.mltac_name=ml_name;mltac_index=0}let()=letopenLtac_plugininletopenTacinterpinletinterpist(ids,tac)=matchidswith|[]->(* Evaluate the Ltac2 quotation eagerly *)letidtac=Value.of_closure{istwithlfun=Id.Map.empty}(CAst.make(Tacexpr.TacId[]))inletist={env_ist=Id.Map.empty}inTac2interp.interpisttac>>=fun_->Ftactic.returnidtac|_::_->(* Return a closure [@f := {blob} |- fun ids => ltac2_eval(f, ids) ] *)(* This name cannot clash with Ltac2 variables which are all lowercase *)letself_id=Id.of_string"F"inletnas=List.map(funid->Nameid)idsinletmk_argid=Tacexpr.Reference(Locus.ArgVar(CAst.makeid))inletargs=List.mapmk_argidsinletclos=CAst.make(Tacexpr.TacFun(nas,CAst.make(Tacexpr.TacML(ltac2_eval,mk_argself_id::args))))inletself=GTacFun(List.map(funid->Nameid)ids,tac)inletself=Tac2interp.interp_value{env_ist=Id.Map.empty}selfinletself=Geninterp.Val.inject(Geninterp.Val.Basetyp_ltac2)selfinletist={istwithlfun=Id.Map.singletonself_idself}inFtactic.return(Value.of_closureistclos)inGeninterp.register_interp0wit_ltac2interplet()=letpr_raw_=Genprint.PrinterBasic(fun_env_sigma->mt())inletpr_glb(ids,e)=letids=ifList.is_emptyidsthenmt()elsepr_sequenceId.printids++str" |- "inGenprint.PrinterBasicPp.(fun_env_sigma->ids++Tac2print.pr_glbexpre)inletpr_top_=Genprint.TopPrinterBasicmtinGenprint.register_print0wit_ltac2pr_rawpr_glbpr_top(** Built-in notation scopes *)letadd_scopesf=Tac2entries.register_scope(Id.of_strings)fletrecpr_scope=letopenCAstinfunction|SexprStr{v=s}->qstrings|SexprInt{v=n}->Pp.intn|SexprRec(_,{v=na},args)->letna=matchnawith|None->str"_"|Someid->Id.printidinna++str"("++prlist_with_sep(fun()->str", ")pr_scopeargs++str")"letscope_failsargs=letargs=str"("++prlist_with_sep(fun()->str", ")pr_scopeargs++str")"inCErrors.user_err(str"Invalid arguments "++args++str" in scope "++strs)letq_unit=CAst.make@@CTacCst(AbsKn(Tuple0))letadd_generic_scopesentryarg=letparse=function|[]->letscope=Pcoq.Symbol.ntermentryinletactx=CAst.make@@CTacExt(arg,x)inTac2entries.ScopeRule(scope,act)|arg->scope_failsarginadd_scopesparseopenCAstlet()=add_scope"keyword"beginfunction|[SexprStr{loc;v=s}]->letscope=Pcoq.Symbol.token(Tok.PKEYWORDs)inTac2entries.ScopeRule(scope,(fun_->q_unit))|arg->scope_fail"keyword"argendlet()=add_scope"terminal"beginfunction|[SexprStr{loc;v=s}]->letscope=Pcoq.Symbol.token(CLexer.terminals)inTac2entries.ScopeRule(scope,(fun_->q_unit))|arg->scope_fail"terminal"argendlet()=add_scope"list0"beginfunction|[tok]->letTac2entries.ScopeRule(scope,act)=Tac2entries.parse_scopetokinletscope=Pcoq.Symbol.list0scopeinletactl=Tac2quote.of_listactlinTac2entries.ScopeRule(scope,act)|[tok;SexprStr{v=str}]->letTac2entries.ScopeRule(scope,act)=Tac2entries.parse_scopetokinletsep=Pcoq.Symbol.tokens[Pcoq.TPattern(CLexer.terminalstr)]inletscope=Pcoq.Symbol.list0sepscopesepfalseinletactl=Tac2quote.of_listactlinTac2entries.ScopeRule(scope,act)|arg->scope_fail"list0"argendlet()=add_scope"list1"beginfunction|[tok]->letTac2entries.ScopeRule(scope,act)=Tac2entries.parse_scopetokinletscope=Pcoq.Symbol.list1scopeinletactl=Tac2quote.of_listactlinTac2entries.ScopeRule(scope,act)|[tok;SexprStr{v=str}]->letTac2entries.ScopeRule(scope,act)=Tac2entries.parse_scopetokinletsep=Pcoq.Symbol.tokens[Pcoq.TPattern(CLexer.terminalstr)]inletscope=Pcoq.Symbol.list1sepscopesepfalseinletactl=Tac2quote.of_listactlinTac2entries.ScopeRule(scope,act)|arg->scope_fail"list1"argendlet()=add_scope"opt"beginfunction|[tok]->letTac2entries.ScopeRule(scope,act)=Tac2entries.parse_scopetokinletscope=Pcoq.Symbol.optscopeinletactopt=matchoptwith|None->CAst.make@@CTacCst(AbsKn(OtherCore.c_none))|Somex->CAst.make@@CTacApp(CAst.make@@CTacCst(AbsKn(OtherCore.c_some)),[actx])inTac2entries.ScopeRule(scope,act)|arg->scope_fail"opt"argendlet()=add_scope"self"beginfunction|[]->letscope=Pcoq.Symbol.selfinletacttac=tacinTac2entries.ScopeRule(scope,act)|arg->scope_fail"self"argendlet()=add_scope"next"beginfunction|[]->letscope=Pcoq.Symbol.nextinletacttac=tacinTac2entries.ScopeRule(scope,act)|arg->scope_fail"next"argendlet()=add_scope"tactic"beginfunction|[]->(* Default to level 5 parsing *)letscope=Pcoq.Symbol.ntermlltac2_expr"5"inletacttac=tacinTac2entries.ScopeRule(scope,act)|[SexprInt{loc;v=n}]asarg->let()=ifn<0||n>6thenscope_fail"tactic"arginletscope=Pcoq.Symbol.ntermlltac2_expr(string_of_intn)inletacttac=tacinTac2entries.ScopeRule(scope,act)|arg->scope_fail"tactic"argendlet()=add_scope"thunk"beginfunction|[tok]->letTac2entries.ScopeRule(scope,act)=Tac2entries.parse_scopetokinletacte=Tac2quote.thunk(acte)inTac2entries.ScopeRule(scope,act)|arg->scope_fail"thunk"argendlet()=add_scope"constr"(funarg->letdelimiters=List.map(function|SexprRec(_,{v=Somes},[])->s|_->scope_fail"constr"arg)arginletacte=Tac2quote.of_constr~delimiterseinTac2entries.ScopeRule(Pcoq.Symbol.ntermPcoq.Constr.constr,act))let()=add_scope"open_constr"(funarg->letdelimiters=List.map(function|SexprRec(_,{v=Somes},[])->s|_->scope_fail"open_constr"arg)arginletacte=Tac2quote.of_open_constr~delimiterseinTac2entries.ScopeRule(Pcoq.Symbol.ntermPcoq.Constr.constr,act))letadd_expr_scopenameentryf=add_scopenamebeginfunction|[]->Tac2entries.ScopeRule(Pcoq.Symbol.ntermentry,f)|arg->scope_failnameargendlet()=add_expr_scope"ident"q_ident(funid->Tac2quote.of_antiTac2quote.of_identid)let()=add_expr_scope"bindings"q_bindingsTac2quote.of_bindingslet()=add_expr_scope"with_bindings"q_with_bindingsTac2quote.of_bindingslet()=add_expr_scope"intropattern"q_intropatternTac2quote.of_intro_patternlet()=add_expr_scope"intropatterns"q_intropatternsTac2quote.of_intro_patternslet()=add_expr_scope"destruction_arg"q_destruction_argTac2quote.of_destruction_arglet()=add_expr_scope"induction_clause"q_induction_clauseTac2quote.of_induction_clauselet()=add_expr_scope"conversion"q_conversionTac2quote.of_conversionlet()=add_expr_scope"rewriting"q_rewritingTac2quote.of_rewritinglet()=add_expr_scope"clause"q_clauseTac2quote.of_clauselet()=add_expr_scope"hintdb"q_hintdbTac2quote.of_hintdblet()=add_expr_scope"occurrences"q_occurrencesTac2quote.of_occurrenceslet()=add_expr_scope"dispatch"q_dispatchTac2quote.of_dispatchlet()=add_expr_scope"strategy"q_strategy_flagTac2quote.of_strategy_flaglet()=add_expr_scope"reference"q_referenceTac2quote.of_referencelet()=add_expr_scope"move_location"q_move_locationTac2quote.of_move_locationlet()=add_expr_scope"pose"q_poseTac2quote.of_poselet()=add_expr_scope"assert"q_assertTac2quote.of_assertionlet()=add_expr_scope"constr_matching"q_constr_matchingTac2quote.of_constr_matchinglet()=add_expr_scope"goal_matching"q_goal_matchingTac2quote.of_goal_matchinglet()=add_expr_scope"format"Pcoq.Prim.lstringTac2quote.of_formatlet()=add_generic_scope"pattern"Pcoq.Constr.constrTac2quote.wit_pattern(** seq scope, a bit hairy *)openPcoqtype_converter=|CvNil:(Loc.t->raw_tacexpr)converter|CvCns:'actconverter*('a->raw_tacexpr)option->('a->'act)converterletrecapply:typea.aconverter->raw_tacexprlist->a=function|CvNil->funacculoc->Tac2quote.of_tuple~locaccu|CvCns(c,None)->funaccux->applycaccu|CvCns(c,Somef)->funaccux->applyc(fx::accu)typeseqrule=|Seqrule:(Tac2expr.raw_tacexpr,Gramlib.Grammar.norec,'act,Loc.t->raw_tacexpr)Rule.t*'actconverter->seqruleletrecmake_seq_rule=function|[]->Seqrule(Pcoq.Rule.stop,CvNil)|tok::rem->letTac2entries.ScopeRule(scope,f)=Tac2entries.parse_scopetokinletscope=matchPcoq.generalize_symbolscopewith|None->CErrors.user_err(str"Recursive symbols (self / next) are not allowed in local rules")|Somescope->scopeinletSeqrule(r,c)=make_seq_rulereminletr=Pcoq.Rule.next_norecrscopeinletf=matchtokwith|SexprStr_->None(* Leave out mere strings *)|_->SomefinSeqrule(r,CvCns(c,f))let()=add_scope"seq"beginfuntoks->letscope=letSeqrule(r,c)=make_seq_rule(List.revtoks)inPcoq.(Symbol.rules[Rules.maker(applyc[])])inTac2entries.ScopeRule(scope,(fune->e))end