123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637(************************************************************************)(* * 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) *)(************************************************************************)openUtilopenCErrorsopenNamesopenLibnamesopenConstrexpropenExtendopenNotation_gramopenPcoq(**********************************************************************)(* This determines (depending on the associativity of the current
level and on the expected associativity) if a reference to constr_n is
a reference to the current level (to be translated into "SELF" on the
left border and into "constr LEVEL n" elsewhere), to the level below
(to be translated into "NEXT") or to an below wrt associativity (to be
translated in camlp5 into "constr" without level) or to another level
(to be translated into "constr LEVEL n")
The boolean is true if the entry was existing _and_ empty; this to
circumvent a weakness of camlp5 whose undo mechanism is not the
converse of the extension mechanism *)letconstr_level=string_of_intletdefault_levels=[200,Gramlib.Gramext.RightA,false;100,Gramlib.Gramext.RightA,false;99,Gramlib.Gramext.RightA,true;90,Gramlib.Gramext.RightA,true;10,Gramlib.Gramext.LeftA,false;9,Gramlib.Gramext.RightA,false;8,Gramlib.Gramext.RightA,true;1,Gramlib.Gramext.LeftA,false;0,Gramlib.Gramext.RightA,false]letdefault_pattern_levels=[200,Gramlib.Gramext.RightA,true;100,Gramlib.Gramext.RightA,false;99,Gramlib.Gramext.RightA,true;90,Gramlib.Gramext.RightA,true;10,Gramlib.Gramext.LeftA,false;1,Gramlib.Gramext.LeftA,false;0,Gramlib.Gramext.RightA,false]letdefault_constr_levels=(default_levels,default_pattern_levels)letfind_levelslevels=function|InConstrEntry->levels,String.Map.find"constr"levels|InCustomEntrys->trylevels,String.Map.findslevelswithNot_found->String.Map.adds([],[])levels,([],[])letsave_levelslevelscustomlev=lets=matchcustomwithInConstrEntry->"constr"|InCustomEntrys->sinString.Map.addslevlevels(* At a same level, LeftA takes precedence over RightA and NoneA *)(* In case, several associativity exists for a level, we make two levels, *)(* first LeftA, then RightA and NoneA together *)letadmissible_assoc=function|Gramlib.Gramext.LeftA,Some(Gramlib.Gramext.RightA|Gramlib.Gramext.NonA)->false|Gramlib.Gramext.RightA,SomeGramlib.Gramext.LeftA->false|_->trueletcreate_assoc=function|None->Gramlib.Gramext.RightA|Somea->aleterror_level_assocpcurrentexpected=letopenPpinletpr_assoc=function|Gramlib.Gramext.LeftA->str"left"|Gramlib.Gramext.RightA->str"right"|Gramlib.Gramext.NonA->str"non"inuser_err(str"Level "++intp++str" is already declared "++pr_assoccurrent++str" associative while it is now expected to be "++pr_assocexpected++str" associative.")typeposition=NewFirst|NewAfterofint|ReuseFirst|ReuseLevelofintletcreate_pos=function|None->NewFirst|Somelev->NewAfterlevletfind_position_gencurrentensureassoclev=matchlevwith|None->current,(ReuseFirst,None,None,None)|Somen->letafter=refNoneinletinit=refNoneinletrecadd_levelq=function|(p,_,_aspa)::lwhenp>n->pa::add_level(Somep)l|(p,a,reinit)::lwhenInt.equalpn->ifreinitthenleta'=create_assocassocin(init:=Some(a',q);(p,a',false)::l)elseifadmissible_assoc(a,assoc)thenraiseExitelseerror_level_assocpa(Option.getassoc)|l->after:=q;(n,create_assocassoc,ensure)::lintryletupdated=add_levelNonecurrentinletassoc=create_assocassocinbeginmatch!initwith|None->(* Create the entry *)updated,(create_pos!after,Someassoc,Some(constr_leveln),None)|_->(* The reinit flag has been updated *)updated,(ReuseLeveln,None,None,!init)endwith(* Nothing has changed *)Exit->(* Just inherit the existing associativity and name (None) *)current,(ReuseLeveln,None,None,None)letreclist_mem_assoc_triplex=function|[]->false|(a,b,c)::l->Int.equalax||list_mem_assoc_triplexlletregister_empty_levelsaccuforpatlevels=letrecfilteraccu=function|[]->([],accu)|(where,n)::rem->letrem,accu=filteraccureminletaccu,(clev,plev)=find_levelsaccuwhereinletlevels=ifforpatthenplevelseclevinifnot(list_mem_assoc_triplenlevels)thenletnlev,ans=find_position_genlevelstrueNone(Somen)inletnlev=ifforpatthen(clev,nlev)else(nlev,plev)in(where,ans)::rem,save_levelsaccuwherenlevelserem,accuinlet(l,accu)=filteracculevelsinList.revl,acculetfind_positionaccucustomforpatassoclevel=letaccu,(clev,plev)=find_levelsaccucustominletlevels=ifforpatthenplevelseclevinletnlev,ans=find_position_genlevelsfalseassoclevelinletnlev=ifforpatthen(clev,nlev)else(nlev,plev)in(ans,save_levelsaccucustomnlev)(**************************************************************************)(*
* --- Note on the mapping of grammar productions to camlp5 actions ---
*
* Translation of environments: a production
* [ nt1(x1) ... nti(xi) ] -> act(x1..xi)
* is written (with camlp5 conventions):
* (fun vi -> .... (fun v1 -> act(v1 .. vi) )..)
* where v1..vi are the values generated by non-terminals nt1..nti.
* Since the actions are executed by substituting an environment,
* the make_*_action family build the following closure:
*
* ((fun env ->
* (fun vi ->
* (fun env -> ...
*
* (fun v1 ->
* (fun env -> gram_action .. env act)
* ((x1,v1)::env))
* ...)
* ((xi,vi)::env)))
* [])
*)(**********************************************************************)(** Declare Notations grammar rules *)(**********************************************************************)(* Binding constr entry keys to entries *)(* Camlp5 levels do not treat NonA: use RightA with a NEXT on the left *)letcamlp5_assoc=letopenGramlib.Gramextinfunction|SomeNonA|SomeRightA->RightA|None|SomeLeftA->LeftAletassoc_eqalar=letopenGramlib.Gramextinmatchal,arwith|NonA,NonA|RightA,RightA|LeftA,LeftA->true|_,_->false(** [adjust_level assoc from prod] where [assoc] and [from] are the name
and associativity of the level where to add the rule; the meaning of
the result is
DefaultLevel = entry name
NextLevel = NEXT
NumLevel n = constr LEVEL n *)letadjust_levelcustomassoc(custom',from)p=letopenGramlib.Gramextinmatchpwith(* If a level in a different grammar, no other choice than denoting it by absolute level *)|(NumLeveln,_)whennot(Notation.notation_entry_eqcustomcustom')->NumLeveln(* If a default level in a different grammar, the entry name is ok *)|(DefaultLevel,InternalProd)->ifNotation.notation_entry_eqcustomInConstrEntrythenNumLevel200elseDefaultLevel|(DefaultLevel,BorderProd_)whennot(Notation.notation_entry_eqcustomcustom')->ifNotation.notation_entry_eqcustomInConstrEntrythenNumLevel200elseDefaultLevel(* Associativity is None means force the level *)|(NumLeveln,BorderProd(_,None))->NumLeveln|(DefaultLevel,BorderProd(_,None))->assertfalse(* Compute production name on the right side *)(* If NonA or LeftA on the right-hand side, set to NEXT *)|((NumLevel_|DefaultLevel),BorderProd(Right,Some(NonA|LeftA)))->NextLevel(* If RightA on the right-hand side, set to the explicit (current) level *)|(NumLeveln,BorderProd(Right,SomeRightA))->NumLeveln|(DefaultLevel,BorderProd(Right,SomeRightA))->NumLevelfrom(* Compute production name on the left side *)(* If NonA on the left-hand side, adopt the current assoc ?? *)|((NumLevel_|DefaultLevel),BorderProd(Left,SomeNonA))->DefaultLevel(* If the expected assoc is the current one, set to SELF *)|((NumLevel_|DefaultLevel),BorderProd(Left,Somea))whenassoc_eqa(camlp5_assocassoc)->DefaultLevel(* Otherwise, force the level, n or n-1, according to expected assoc *)|(NumLeveln,BorderProd(Left,SomeLeftA))->NumLeveln|((NumLevel_|DefaultLevel),BorderProd(Left,Some_))->NextLevel(* None means NEXT *)|(NextLevel,_)->assert(Notation.notation_entry_eqcustomcustom');NextLevel(* Compute production name elsewhere *)|(NumLeveln,InternalProd)->iffrom=n+1thenNextLevelelseNumLevelntype_target=|ForConstr:constr_exprtarget|ForPattern:cases_pattern_exprtargettypeprod_info=production_level*production_positiontype(_,_)entry=|TTIdent:('self,lident)entry|TTName:('self,lname)entry|TTReference:('self,qualid)entry|TTBigint:('self,string)entry|TTBinder:bool->('self,kinded_cases_pattern_expr)entry|TTConstr:notation_entry*prod_info*'rtarget->('r,'r)entry|TTConstrList:notation_entry*prod_info*stringTok.plist*'rtarget->('r,'rlist)entry|TTPattern:int->('self,cases_pattern_expr)entry|TTOpenBinderList:('self,local_binder_exprlist)entry|TTClosedBinderList:stringTok.plist->('self,local_binder_exprlistlist)entrytype_any_entry=TTAny:('s,'r)entry->'sany_entryletconstr_custom_entry:(string,Constrexpr.constr_expr)entry_command=create_entry_command"constr"(funsst->[s],st)letpattern_custom_entry:(string,Constrexpr.cases_pattern_expr)entry_command=create_entry_command"pattern"(funsst->[s],st)letcustom_entry_locality=Summary.ref~name:"LOCAL-CUSTOM-ENTRY"String.Set.empty(** If the entry is present then local *)letcreate_custom_entry~locals=ifList.mems["constr";"pattern";"ident";"global";"binder";"bigint"]thenuser_errPp.(quote(strs)++str" is a reserved entry name.");letsc="custom:"^sinletsp="custom_pattern:"^sinlet_=extend_entry_commandconstr_custom_entryscinlet_=extend_entry_commandpattern_custom_entryspinlet()=iflocalthencustom_entry_locality:=String.Set.adds!custom_entry_localityin()letfind_custom_entrys=letsc="custom:"^sinletsp="custom_pattern:"^sintry(find_custom_entryconstr_custom_entrysc,find_custom_entrypattern_custom_entrysp)withNot_found->user_errPp.(str"Undeclared custom entry: "++strs++str".")letexists_custom_entrys=matchfind_custom_entryswith|_->true|exception_->falseletlocality_of_custom_entrys=String.Set.mems!custom_entry_locality(* This computes the name of the level where to add a new rule *)letinterp_constr_entry_key:typer._->rtarget->int->rEntry.t*intoption=funcustomforpatlevel->matchcustomwith|InCustomEntrys->(let(entry_for_constr,entry_for_patttern)=find_custom_entrysinmatchforpatwith|ForConstr->entry_for_constr,Somelevel|ForPattern->entry_for_patttern,Somelevel)|InConstrEntry->matchforpatwith|ForConstr->iflevel=200thenConstr.binder_constr,NoneelseConstr.term,Somelevel|ForPattern->Constr.pattern,Somelevellettarget_entry:types.notation_entry->starget->sEntry.t=function|InConstrEntry->(function|ForConstr->Constr.term|ForPattern->Constr.pattern)|InCustomEntrys->let(entry_for_constr,entry_for_patttern)=find_custom_entrysinfunction|ForConstr->entry_for_constr|ForPattern->entry_for_pattternletis_selfcustom(custom',from)e=Notation.notation_entry_eqcustomcustom'&&matchewith|(NumLeveln,BorderProd(Right,_(* Some(NonA|LeftA) *)))->false|(NumLeveln,BorderProd(Left,_))->Int.equalfromn|_->falseletis_binder_levelcustom(custom',from)e=matchewith|(NumLevel200,(BorderProd(Right,_)|InternalProd))->custom=InConstrEntry&&custom'=InConstrEntry&&from=200|_->falseletmake_sep_rules=function|[tk]->Pcoq.Symbol.tokentk|tkl->letr=Pcoq.mk_rule(List.revtkl)inPcoq.Symbol.rules[r]type('s,'a)mayrec_symbol=|MayRecNo:('s,Gramlib.Grammar.norec,'a)Symbol.t->('s,'a)mayrec_symbol|MayRecMay:('s,Gramlib.Grammar.mayrec,'a)Symbol.t->('s,'a)mayrec_symbolletsymbol_of_target:types._->_->_->_->starget->(s,s)mayrec_symbol=funcustompassocfromforpat->ifis_binder_levelcustomfrompthen(* Prevent self *)MayRecNo(Pcoq.Symbol.nterml(target_entrycustomforpat)"200")elseifis_selfcustomfrompthenMayRecMayPcoq.Symbol.selfelseletg=target_entrycustomforpatinletlev=adjust_levelcustomassocfrompinbeginmatchlevwith|DefaultLevel->MayRecNo(Pcoq.Symbol.ntermg)|NextLevel->MayRecMayPcoq.Symbol.next|NumLevellev->MayRecNo(Pcoq.Symbol.ntermlg(string_of_intlev))endletsymbol_of_entry:typesr._->_->(s,r)entry->(s,r)mayrec_symbol=funassocfromtyp->matchtypwith|TTConstr(s,p,forpat)->symbol_of_targetspassocfromforpat|TTConstrList(s,typ',[],forpat)->beginmatchsymbol_of_targetstyp'assocfromforpatwith|MayRecNos->MayRecNo(Pcoq.Symbol.list1s)|MayRecMays->MayRecMay(Pcoq.Symbol.list1s)end|TTConstrList(s,typ',tkl,forpat)->beginmatchsymbol_of_targetstyp'assocfromforpatwith|MayRecNos->MayRecNo(Pcoq.Symbol.list1seps(make_sep_rulestkl)false)|MayRecMays->MayRecMay(Pcoq.Symbol.list1seps(make_sep_rulestkl)false)end|TTPatternp->MayRecNo(Pcoq.Symbol.ntermlConstr.pattern(string_of_intp))|TTClosedBinderList[]->MayRecNo(Pcoq.Symbol.list1(Pcoq.Symbol.ntermConstr.binder))|TTClosedBinderListtkl->MayRecNo(Pcoq.Symbol.list1sep(Pcoq.Symbol.ntermConstr.binder)(make_sep_rulestkl)false)|TTIdent->MayRecNo(Pcoq.Symbol.ntermPrim.identref)|TTName->MayRecNo(Pcoq.Symbol.ntermPrim.name)|TTBindertrue->MayRecNo(Pcoq.Symbol.ntermConstr.one_open_binder)|TTBinderfalse->MayRecNo(Pcoq.Symbol.ntermConstr.one_closed_binder)|TTOpenBinderList->MayRecNo(Pcoq.Symbol.ntermConstr.open_binders)|TTBigint->MayRecNo(Pcoq.Symbol.ntermPrim.bignat)|TTReference->MayRecNo(Pcoq.Symbol.ntermConstr.global)letinterp_entryforpate=matchewith|ETProdIdent->TTAnyTTIdent|ETProdName->TTAnyTTName|ETProdReference->TTAnyTTReference|ETProdBigint->TTAnyTTBigint|ETProdOneBindero->TTAny(TTBindero)|ETProdConstr(s,p)->TTAny(TTConstr(s,p,forpat))|ETProdPatternp->TTAny(TTPatternp)|ETProdConstrList(s,p,tkl)->TTAny(TTConstrList(s,p,tkl,forpat))|ETProdBinderListETBinderOpen->TTAnyTTOpenBinderList|ETProdBinderList(ETBinderClosedtkl)->TTAny(TTClosedBinderListtkl)letcases_pattern_expr_of_id{CAst.loc;v=id}=CAst.make?loc@@CPatAtom(Some(qualid_of_ident?locid))letcases_pattern_expr_of_name{CAst.loc;v=na}=CAst.make?loc@@matchnawith|Anonymous->CPatAtomNone|Nameid->CPatAtom(Some(qualid_of_ident?locid))type'renv={constrs:'rlist;constrlists:'rlistlist;binders:kinded_cases_pattern_exprlist;binderlists:local_binder_exprlistlist;}letpush_constrsubstv={substwithconstrs=v::subst.constrs}letpush_item:typesr.starget->(s,r)entry->senv->r->senv=funforpatesubstv->matchewith|TTConstr_->push_constrsubstv|TTIdent->beginmatchforpatwith|ForConstr->{substwithbinders=(cases_pattern_expr_of_idv,Glob_term.Explicit)::subst.binders}|ForPattern->push_constrsubst(cases_pattern_expr_of_idv)end|TTName->beginmatchforpatwith|ForConstr->{substwithbinders=(cases_pattern_expr_of_namev,Glob_term.Explicit)::subst.binders}|ForPattern->push_constrsubst(cases_pattern_expr_of_namev)end|TTPattern_->beginmatchforpatwith|ForConstr->{substwithbinders=(v,Glob_term.Explicit)::subst.binders}|ForPattern->push_constrsubstvend|TTBindero->{substwithbinders=v::subst.binders}|TTOpenBinderList->{substwithbinderlists=v::subst.binderlists}|TTClosedBinderList_->{substwithbinderlists=List.flattenv::subst.binderlists}|TTBigint->beginmatchforpatwith|ForConstr->push_constrsubst(CAst.make@@CPrim(Number(NumTok.Signed.of_int_stringv)))|ForPattern->push_constrsubst(CAst.make@@CPatPrim(Number(NumTok.Signed.of_int_stringv)))end|TTReference->beginmatchforpatwith|ForConstr->push_constrsubst(CAst.make@@CRef(v,None))|ForPattern->push_constrsubst(CAst.make@@CPatAtom(Somev))end|TTConstrList_->{substwithconstrlists=v::subst.constrlists}type(_,_)ty_symbol=|TyTerm:stringTok.p->('s,string)ty_symbol|TyNonTerm:'starget*('s,'a)entry*('s,'a)mayrec_symbol*bool->('s,'a)ty_symboltype('self,_,'r)ty_rule=|TyStop:('self,'r,'r)ty_rule|TyNext:('self,'a,'r)ty_rule*('self,'b)ty_symbol->('self,'b->'a,'r)ty_rule|TyMark:int*bool*int*('self,'a,'r)ty_rule->('self,'a,'r)ty_ruletype'rgen_eval=Loc.t->'renv->'rletrecty_eval:typesa.(s,a,Loc.t->s)ty_rule->sgen_eval->senv->a=function|TyStop->funfenvloc->flocenv|TyNext(rem,TyTerm_)->funfenv_->ty_evalremfenv|TyNext(rem,TyNonTerm(_,_,_,false))->funfenv_->ty_evalremfenv|TyNext(rem,TyNonTerm(forpat,e,_,true))->funfenvv->ty_evalremf(push_itemforpateenvv)|TyMark(n,b,p,rem)->funfenv->letheads,constrs=List.chopnenv.constrsinletconstrlists,constrs=ifbthen(* We rearrange constrs = c1..cn rem and constrlists = [d1..dr e1..ep] rem' into
constrs = e1..ep rem and constrlists [c1..cn d1..dr] rem' *)letconstrlist=List.hdenv.constrlistsinletconstrlist,tail=List.chop(List.lengthconstrlist-p)constrlistin(heads@constrlist)::List.tlenv.constrlists,tail@constrselse(* We rearrange constrs = c1..cn e1..ep rem into
constrs = e1..ep rem and add a constr list [c1..cn] *)letconstrlist,tail=List.chop(n-p)headsinconstrlist::env.constrlists,tail@constrsinty_evalremf{envwithconstrs;constrlists;}type('s,'a,'r)mayrec_rule=|MayRecRNo:('s,Gramlib.Grammar.norec,'a,'r)Rule.t->('s,'a,'r)mayrec_rule|MayRecRMay:('s,Gramlib.Grammar.mayrec,'a,'r)Rule.t->('s,'a,'r)mayrec_ruleletrecty_erase:typesar.(s,a,r)ty_rule->(s,a,r)mayrec_rule=function|TyStop->MayRecRNoRule.stop|TyMark(_,_,_,r)->ty_eraser|TyNext(rem,TyTermtok)->beginmatchty_eraseremwith|MayRecRNorem->MayRecRMay(Rule.nextrem(Symbol.tokentok))|MayRecRMayrem->MayRecRMay(Rule.nextrem(Symbol.tokentok))end|TyNext(rem,TyNonTerm(_,_,s,_))->beginmatchty_eraserem,swith|MayRecRNorem,MayRecNos->MayRecRMay(Rule.nextrems)|MayRecRNorem,MayRecMays->MayRecRMay(Rule.nextrems)|MayRecRMayrem,MayRecNos->MayRecRMay(Rule.nextrems)|MayRecRMayrem,MayRecMays->MayRecRMay(Rule.nextrems)endtype('self,'r)any_ty_rule=|AnyTyRule:('self,'act,Loc.t->'r)ty_rule->('self,'r)any_ty_ruleletmake_ty_ruleassocfromforpatprods=letrecmake_ty_rule=function|[]->AnyTyRuleTyStop|GramConstrTerminaltok::rem->letAnyTyRuler=make_ty_rulereminAnyTyRule(TyNext(r,TyTermtok))|GramConstrNonTerminal(e,var)::rem->letAnyTyRuler=make_ty_rulereminletTTAnye=interp_entryforpateinlets=symbol_of_entryassocfromeinletbind=matchvarwithNone->false|Some_->trueinAnyTyRule(TyNext(r,TyNonTerm(forpat,e,s,bind)))|GramConstrListMark(n,b,p)::rem->letAnyTyRuler=make_ty_rulereminAnyTyRule(TyMark(n,b,p,r))inmake_ty_rule(List.revprods)lettarget_to_bool:typer.rtarget->bool=function|ForConstr->false|ForPattern->trueletprepare_empty_levelsforpat(where,(pos,p4assoc,name,reinit))=letempty=matchposwith|ReuseFirst->Pcoq.Reuse(None,[])|ReuseLeveln->Pcoq.Reuse(Some(constr_leveln),[])|NewFirst->Pcoq.Fresh(Gramlib.Gramext.First,[(name,p4assoc,[])])|NewAftern->Pcoq.Fresh(Gramlib.Gramext.After(constr_leveln),[(name,p4assoc,[])])inmatchreinitwith|None->ExtendRule(target_entrywhereforpat,empty)|Some(assoc,pos)->letpos=matchposwithNone->Gramlib.Gramext.First|Somen->Gramlib.Gramext.After(constr_leveln)inletreinit=(assoc,pos)inExtendRuleReinit(target_entrywhereforpat,reinit,empty)letdifferent_levels(custom,opt_level)(custom',string_level)=matchopt_levelwith|None->true|Somelevel->not(Notation.notation_entry_eqcustomcustom')||level<>int_of_stringstring_levelletrecpure_sublevels'assocfromforpatlevel=function|[]->[]|GramConstrNonTerminal(e,_)::rem->letrem=pure_sublevels'assocfromforpatlevelreminletpushwhereprem=matchsymbol_of_targetwherepassocfromforpatwith|MayRecNosym->(matchPcoq.level_of_nontermsymwith|None->rem|Somei->ifdifferent_levels(fstfrom,level)(where,i)then(where,int_of_stringi)::remelserem)|_->remin(matchewith|ETProdPatterni->pushInConstrEntry(NumLeveli,InternalProd)rem|ETProdConstr(s,p)->pushsprem|_->rem)|(GramConstrTerminal_|GramConstrListMark_)::rem->pure_sublevels'assocfromforpatlevelremletmake_act:typer.rtarget->_->rgen_eval=function|ForConstr->funnotationlocenv->letenv=(env.constrs,env.constrlists,env.binders,env.binderlists)inCAst.make~loc@@CNotation(None,notation,env)|ForPattern->funnotationlocenv->letenv=(env.constrs,env.constrlists)inCAst.make~loc@@CPatNotation(None,notation,env,[])letextend_constrstateforpatng=letcustom,n,_=ng.notgram_levelinletassoc=ng.notgram_associnlet(entry,level)=interp_constr_entry_keycustomforpatninletfold(accu,state)pt=letAnyTyRuler=make_ty_ruleassoc(custom,n)forpatptinletpure_sublevels=pure_sublevels'assoc(custom,n)forpatlevelptinletisforpat=target_to_boolforpatinletneeded_levels,state=register_empty_levelsstateisforpatpure_sublevelsinlet(pos,p4assoc,name,reinit),state=find_positionstatecustomisforpatassoclevelinletempty_rules=List.map(prepare_empty_levelsforpat)needed_levelsinletempty={constrs=[];constrlists=[];binders=[];binderlists=[]}inletact=ty_evalr(make_actforpatng.notgram_notation)emptyinletrule=letr=matchty_eraserwith|MayRecRNosymbs->Pcoq.Production.makesymbsact|MayRecRMaysymbs->Pcoq.Production.makesymbsactinletrule=name,p4assoc,[r]inmatchposwith|NewFirst->Pcoq.Fresh(Gramlib.Gramext.First,[rule])|NewAftern->Pcoq.Fresh(Gramlib.Gramext.After(constr_leveln),[rule])|ReuseFirst->Pcoq.Reuse(None,[r])|ReuseLeveln->Pcoq.Reuse(Some(constr_leveln),[r])inletr=matchreinitwith|None->ExtendRule(entry,rule)|Some(assoc,pos)->letpos=matchposwithNone->Gramlib.Gramext.First|Somen->Gramlib.Gramext.After(constr_leveln)inletreinit=(assoc,pos)inExtendRuleReinit(entry,reinit,rule)in(accu@empty_rules@[r],state)inList.fold_leftfold([],state)ng.notgram_prodsletconstr_levels=GramState.field()letis_disjunctive_pattern_ruleng=String.is_sub"( _ | "(sndng.notgram_notation)0letwarn_disj_pattern_notation=letopenPpinletppng=str"Use of "++Notation.pr_notationng.notgram_notation++str" Notation is deprecated as it is inconsistent with pattern syntax."inCWarnings.create~name:"disj-pattern-notation"~category:"notation"~default:CWarnings.Disabledppletextend_constr_notationngstate=letlevels=matchGramState.getstateconstr_levelswith|None->String.Map.add"constr"default_constr_levelsString.Map.empty|Somelev->levin(* Add the notation in constr *)let(r,levels)=extend_constrlevelsForConstrngin(* Add the notation in cases_pattern, unless it would disrupt *)(* parsing nested disjunctive patterns. *)let(r',levels)=ifis_disjunctive_pattern_rulengthenbeginwarn_disj_pattern_notationng;([],levels)endelseextend_constrlevelsForPatternnginletstate=GramState.setstateconstr_levelslevelsin(r@r',state)letconstr_grammar:one_notation_grammargrammar_command=create_grammar_command"Notation"extend_constr_notationletextend_constr_grammarntn=extend_grammar_commandconstr_grammarntn