123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691692693694695696697698699700701702703704705706707708709710711712713714715716717718719720721722723724725726727728729730731732733734735736737738739740741742743744745746747748749750751752753754755756757758759760761762763764765766767768769770771772773774775776777778779780781782783784785786787788789790791792793794795796797798799800801802803804805806807808809810811812813814815816817818819820821822823824825826827828829830831832833834835836837838839840841842843844845846847848849850851852853854855856857858859860861862863864865866867868869870871872873874875876877878879880881882883884885886887888889890891892893894895896897898899900901902903904905906907908909910911912913914915916917918919920921922923924925926927928929930931932933934935936937938939940941942943944945946947948949950951952953954955956957958959960961962963964965966967968969970971972973974975976977978979980981982983984985986987988989990991992993994995996997998999100010011002100310041005100610071008100910101011101210131014101510161017101810191020102110221023102410251026102710281029103010311032103310341035103610371038103910401041104210431044104510461047104810491050105110521053105410551056105710581059106010611062106310641065106610671068106910701071107210731074107510761077107810791080108110821083108410851086108710881089109010911092109310941095109610971098109911001101110211031104110511061107110811091110111111121113111411151116111711181119112011211122112311241125112611271128112911301131113211331134113511361137113811391140114111421143114411451146114711481149115011511152115311541155115611571158115911601161116211631164116511661167116811691170117111721173117411751176117711781179118011811182118311841185118611871188118911901191119211931194119511961197119811991200120112021203120412051206120712081209121012111212121312141215121612171218121912201221122212231224122512261227122812291230123112321233123412351236123712381239124012411242124312441245124612471248124912501251125212531254125512561257125812591260126112621263126412651266126712681269127012711272127312741275127612771278127912801281128212831284128512861287128812891290129112921293129412951296129712981299130013011302130313041305130613071308130913101311131213131314131513161317131813191320132113221323132413251326132713281329133013311332133313341335133613371338133913401341134213431344134513461347134813491350135113521353135413551356135713581359136013611362136313641365136613671368136913701371137213731374137513761377137813791380138113821383138413851386138713881389139013911392139313941395139613971398139914001401140214031404140514061407140814091410141114121413141414151416141714181419142014211422142314241425142614271428142914301431143214331434143514361437143814391440144114421443144414451446144714481449145014511452145314541455145614571458145914601461146214631464146514661467146814691470147114721473147414751476147714781479148014811482148314841485148614871488148914901491149214931494149514961497149814991500150115021503150415051506150715081509151015111512151315141515151615171518151915201521152215231524152515261527152815291530153115321533153415351536153715381539154015411542154315441545154615471548154915501551155215531554155515561557155815591560156115621563156415651566156715681569157015711572157315741575157615771578157915801581158215831584158515861587158815891590159115921593159415951596159715981599160016011602160316041605160616071608160916101611161216131614161516161617161816191620162116221623162416251626162716281629163016311632163316341635163616371638163916401641164216431644164516461647164816491650165116521653165416551656165716581659166016611662166316641665166616671668166916701671167216731674167516761677167816791680168116821683168416851686168716881689169016911692169316941695169616971698169917001701170217031704170517061707170817091710171117121713171417151716171717181719172017211722172317241725172617271728172917301731173217331734173517361737173817391740174117421743174417451746174717481749175017511752175317541755175617571758175917601761176217631764176517661767176817691770177117721773177417751776177717781779178017811782178317841785178617871788178917901791179217931794179517961797179817991800180118021803180418051806180718081809181018111812181318141815181618171818181918201821182218231824182518261827182818291830183118321833183418351836183718381839184018411842184318441845184618471848184918501851185218531854185518561857185818591860186118621863186418651866186718681869187018711872187318741875187618771878187918801881188218831884188518861887(************************************************************************)(* * 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) *)(************************************************************************)openPpopenCErrorsopenUtilopenNamesopenConstrexpropenConstrexpr_opsopenNotation_termopenNotation_gramopenNotation_opsopenPpextendopenExtendopenLibobjectopenConstrinternopenVernacexpropenLibnamesopenNotationopenNameops(**********************************************************************)(* Tokens *)letcache_token(_,s)=CLexer.add_keywordsletinToken:string->obj=declare_object@@global_object_nodischarge"TOKEN"~cache:cache_token~subst:(SomeLibobject.ident_subst_function)letadd_token_objs=Lib.add_anonymous_leaf(inTokens)(**********************************************************************)(* Printing grammar entries *)letentry_buf=Buffer.create64letpr_entrye=let()=Buffer.clearentry_bufinletft=Format.formatter_of_bufferentry_bufinlet()=Pcoq.Entry.printfteinstr(Buffer.contentsentry_buf)letpr_registered_grammarname=letgram=Pcoq.find_grammars_by_namenameinmatchgramwith|[]->user_errPp.(str"Unknown or unprintable grammar entry.")|entries->letpr_one(Pcoq.AnyEntrye)=str"Entry "++str(Pcoq.Entry.namee)++str" is"++fnl()++pr_entryeinprlistpr_oneentriesletpr_grammar=function|"constr"|"term"|"binder_constr"->str"Entry constr is"++fnl()++pr_entryPcoq.Constr.constr++str"and lconstr is"++fnl()++pr_entryPcoq.Constr.lconstr++str"where binder_constr is"++fnl()++pr_entryPcoq.Constr.binder_constr++str"and term is"++fnl()++pr_entryPcoq.Constr.term|"pattern"->pr_entryPcoq.Constr.pattern|"vernac"->str"Entry vernac_control is"++fnl()++pr_entryPvernac.Vernac_.vernac_control++str"Entry command is"++fnl()++pr_entryPvernac.Vernac_.command++str"Entry syntax is"++fnl()++pr_entryPvernac.Vernac_.syntax++str"Entry gallina is"++fnl()++pr_entryPvernac.Vernac_.gallina++str"Entry gallina_ext is"++fnl()++pr_entryPvernac.Vernac_.gallina_ext|name->pr_registered_grammarnameletpr_custom_grammarname=pr_registered_grammar("custom:"^name)(**********************************************************************)(* Parse a format (every terminal starting with a letter or a single
quote (except a single quote alone) must be quoted) *)letparse_format({CAst.loc;v=str}:lstring)=letlen=String.lengthstrin(* TODO: update the line of the location when the string contains newlines *)letmake_locij=Option.map(Loc.shift_loc(i+1)(j-len))locinletpush_tokenloca=function|(i,cur)::l->(i,(loc,a)::cur)::l|[]->assertfalseinletpush_whiteinl=ifInt.equaln0thenlelsepush_token(make_loci(i+n))(UnpTerminal(String.maken' '))linletclose_boxstartstopb=function|(_,a)::(_::_asl)->push_token(make_locstartstop)(UnpBox(b,a))l|[a]->user_err?loc:(make_locstartstop)Pp.(str"Non terminated box in format.")|[]->assertfalseinletclose_quotationstarti=ifi<len&&str.[i]=='\''thenif(Int.equal(i+1)len||str.[i+1]==' ')theni+1elseuser_err?loc:(make_loc(i+1)(i+1))Pp.(str"Space expected after quoted expression.")elseuser_err?loc:(make_locstart(i-1))Pp.(str"Beginning of quoted expression expected to be ended by a quote.")inletrecspacesni=ifi<len&&str.[i]==' 'thenspaces(n+1)(i+1)elseninletrecnonspacesquotedni=ifi<len&&str.[i]!=' 'thenifstr.[i]=='\''&"ed&&(i+1>=len||str.[i+1]==' ')thenifInt.equaln0thenuser_err?loc:(make_loc(i-1)i)Pp.(str"Empty quoted token.")elsenelsenonspacesquoted(n+1)(i+1)elseifquotedthenuser_err?loc:(make_locii)Pp.(str"Spaces are not allowed in (quoted) symbols.")elseninletrecparse_non_formati=letn=nonspacesfalse0iinpush_token(make_loci(i+n-1))(UnpTerminal(String.substrin))(parse_token1(i+n))andparse_quotednki=ifi<lenthenmatchstr.[i]with(* Parse " // " *)|'/'wheni+1<len&&str.[i+1]=='/'->(* We discard the useless n spaces... *)push_token(make_loc(i-n)(i+1))(UnpCutPpFnl)(parse_token1(close_quotationi(i+2)))(* Parse " .. / .. " *)|'/'wheni+1<len->letp=spaces0(i+1)inpush_token(make_loc(i-n)(i+p))(UnpCut(PpBrk(n,p)))(parse_token1(close_quotationi(i+p+1)))|c->(* The spaces are real spaces *)push_white(i-n-1-k)n(matchcwith|'['->ifi+1<lenthenmatchstr.[i+1]with(* Parse " [h .. ", *)|'h'wheni+1<=len&&str.[i+2]=='v'->(parse_boxi(funn->PpHVBn)(i+3))(* Parse " [v .. ", *)|'v'->parse_boxi(funn->PpVBn)(i+2)(* Parse " [ .. ", *)|' '|'\''->parse_boxi(funn->PpHOVBn)(i+1)|_->user_err?loc:(make_locii)Pp.(str"\"v\", \"hv\", \" \" expected after \"[\" in format.")elseuser_err?loc:(make_locii)Pp.(str"\"v\", \"hv\" or \" \" expected after \"[\" in format.")(* Parse "]" *)|']'->((i,[])::parse_token1(close_quotationi(i+1)))(* Parse a non formatting token *)|c->letn=nonspacestrue0iinpush_token(make_loci(i+n-1))(UnpTerminal(String.substr(i-1)(n+2)))(parse_token1(close_quotationi(i+n))))elseifInt.equaln0then[]elseuser_err?loc:(make_loc(len-n)len)Pp.(str"Ending spaces non part of a format annotation.")andparse_boxstartboxi=letn=spaces0iinclose_boxstart(i+n-1)(boxn)(parse_token1(close_quotationi(i+n)))andparse_tokenki=letn=spaces0iinleti=i+ninifi<lenthenmatchstr.[i]with(* Parse a ' *)|'\''wheni+1>=len||str.[i+1]==' '->push_white(i-n)(n-k)(push_token(make_loci(i+1))(UnpTerminal"'")(parse_token1(i+1)))(* Parse the beginning of a quoted expression *)|'\''->parse_quoted(n-k)k(i+1)(* Otherwise *)|_->push_white(i-n)(n-k)(parse_non_formati)elsepush_white(i-n)n[(len,[])]inifnot(String.is_emptystr)thenmatchparse_token00with|[_,l]->l|(i,_)::_->user_err?loc:(make_locii)Pp.(str"Box closed without being opened.")|[]->assertfalseelse[](***********************)(* Analyzing notations *)(* Find non-terminal tokens of notation *)(* To protect alphabetic tokens and quotes from being seen as variables *)letquote_notation_tokenx=letn=String.lengthxinletnorm=CLexer.is_identxinif(n>0&&norm)||(n>2&&x.[0]=='\'')then"'"^x^"'"elsexletis_numeral_in_constrentrysymbs=matchentry,List.filter(functionBreak_->false|_->true)symbswith|InConstrEntry,([Terminal"-";Terminalx]|[Terminalx])->NumTok.Unsigned.parse_stringx<>None|_->falseletanalyze_notation_tokens~onlyprintdf=let(recvars,mainvars,symbolsasres)=decompose_raw_notationdfin(* don't check for nonlinearity if printing only, see Bug 5526 *)(ifnotonlyprintthenmatchList.duplicatesId.equal(mainvars@List.mapsndrecvars)with|id::_->user_err~hdr:"Metasyntax.get_notation_vars"(str"Variable "++Id.printid++str" occurs more than once.")|_->());resleterror_not_same_scopexy=user_err~hdr:"Metasyntax.error_not_name_scope"(str"Variables "++Id.printx++str" and "++Id.printy++str" must be in the same scope.")(**********************************************************************)(* Build pretty-printing rules *)letpr_notation_entry=function|InConstrEntry->str"constr"|InCustomEntrys->str"custom "++strsletprecedence_of_position_and_levelfrom_level=function|NumLeveln,BorderProd(b,Somea)->(letopenGramlib.Gramextinmatcha,bwith|RightA,Left->LevelLtn|RightA,Right->LevelLen|LeftA,Left->LevelLen|LeftA,Right->LevelLtn|NonA,_->LevelLtn),Someb|NumLeveln,_->LevelLen,None|NextLevel,_->LevelLtfrom_level,None|DefaultLevel,_->LevelSome,None(** Computing precedences of subentries for parsing *)letprecedence_of_entry_type(from_custom,from_level)=function|ETConstr(custom,_,x)whennotation_entry_eqcustomfrom_custom->fst(precedence_of_position_and_levelfrom_levelx)|ETConstr(custom,_,(NumLeveln,_))->LevelLen|ETConstr(custom,_,(NextLevel,_))->user_err(strbrk"\"next level\" is only for sub-expressions in the same entry as where the notation is ("++quote(pr_notation_entrycustom)++strbrk" is different from "++quote(pr_notation_entryfrom_custom)++str").")|ETPattern(_,n)->letn=matchnwithNone->0|Somen->ninLevelLen|_->LevelSome(* should not matter *)(** Computing precedences for future insertion of parentheses at
the time of printing using hard-wired constr levels *)letunparsing_precedence_of_entry_typefrom_level=function|ETConstr(InConstrEntry,_,x)->(* Possible insertion of parentheses at printing time to deal
with precedence in a constr entry is managed using [prec_less]
in [ppconstr.ml] *)precedence_of_position_and_levelfrom_levelx|ETConstr(custom,_,_)->(* Precedence of printing for a custom entry is managed using
explicit insertion of entry coercions at the time of building
a [constr_expr] *)LevelSome,None|ETPattern(_,n)->(* in constr *)LevelLe(matchnwithSomen->n|None->0),None|_->LevelSome,None(* should not matter *)(* Some breaking examples *)(* "x = y" : "x /1 = y" (breaks before any symbol) *)(* "x =S y" : "x /1 =S /1 y" (protect from confusion; each side for symmetry)*)(* "+ {" : "+ {" may breaks reversibility without space but oth. not elegant *)(* "x y" : "x spc y" *)(* "{ x } + { y }" : "{ x } / + { y }" *)(* "< x , y > { z , t }" : "< x , / y > / { z , / t }" *)letstarts_with_left_brackets=letl=String.lengthsinnot(Int.equall0)&&(s.[0]=='{'||s.[0]=='['||s.[0]=='(')letends_with_right_brackets=letl=String.lengthsinnot(Int.equall0)&&(s.[l-1]=='}'||s.[l-1]==']'||s.[l-1]==')')letis_left_brackets=starts_with_left_brackets&¬(ends_with_right_brackets)letis_right_brackets=not(starts_with_left_brackets)&&ends_with_right_bracketsletis_commas=letl=String.lengthsinnot(Int.equall0)&&(s.[0]==','||s.[0]==';')letis_operators=letl=String.lengthsinnot(Int.equall0)&&(s.[0]=='+'||s.[0]=='*'||s.[0]=='='||s.[0]=='-'||s.[0]=='/'||s.[0]=='<'||s.[0]=='>'||s.[0]=='@'||s.[0]=='\\'||s.[0]=='&'||s.[0]=='~'||s.[0]=='$')letis_non_terminal=function|NonTerminal_|SProdList_->true|_->falseletis_next_non_terminalb=function|[]->b|pr::_->is_non_terminalprletis_next_terminal=functionTerminal_::_->true|_->falseletis_next_break=functionBreak_::_->true|_->falseletadd_breaknl=(None,UnpCut(PpBrk(n,0)))::lletadd_break_if_nonenb=function|(_,UnpCut(PpBrk_))::_asl->l|[]whennotb->[]|l->(None,UnpCut(PpBrk(n,0)))::lletcheck_open_binderisopenslm=letpr_token=function|Terminals->strs|Breakn->str"␣"|_->assertfalseinifisopen&¬(List.is_emptysl)thenuser_err(str"as "++Id.printm++str" is a non-closed binder, no such \""++prlist_with_sepspcpr_tokensl++strbrk"\" is allowed to occur.")letunparsing_metavarifromtyps=letx=List.nthtyps(i-1)inletprec,side=unparsing_precedence_of_entry_typefromxinmatchxwith|ETConstr_|ETGlobal|ETBigint->UnpMetaVar(prec,side)|ETPattern_|ETName_|ETIdent->UnpBinderMetaVar(prec,NotQuotedPattern)|ETBinderisopen->UnpBinderMetaVar(prec,QuotedPattern)(* Heuristics for building default printing rules *)letindex_ididl=List.indexId.equalidlletmake_hunksetypssymbolsfrom_level=letvars,typs=List.splitetypsinletrecmakeb=function|NonTerminalm::prods->leti=index_idmvarsinletu=unparsing_metavarifrom_leveltypsinifis_next_non_terminalbprodsthen(None,u)::add_break_if_none1b(makebprods)else(None,u)::make_with_spacebprods|Terminals::prodswhen(* true to simulate presence of non-terminal *)b||List.existsis_non_terminalprods->if(is_commas||is_operators)then(* Always a breakable space after comma or separator *)(None,UnpTerminals)::add_break_if_none1b(makebprods)elseifis_right_brackets&&is_next_terminalprodsthen(* Always no space after right bracked, but possibly a break *)(None,UnpTerminals)::add_break_if_none0b(makebprods)elseifis_left_brackets&&is_next_non_terminalbprodsthen(None,UnpTerminals)::makebprodselseifnot(is_next_breakprods)then(* Add rigid space, no break, unless user asked for something *)(None,UnpTerminal(s^" "))::makebprodselse(* Rely on user spaces *)(None,UnpTerminals)::makebprods|Terminals::prods->(* Separate but do not cut a trailing sequence of terminal *)(matchprodswith|Terminal_::_->(None,UnpTerminal(s^" "))::makebprods|_->(None,UnpTerminals)::makebprods)|Breakn::prods->add_breakn(makebprods)|SProdList(m,sl)::prods->leti=index_idmvarsinlettyp=List.nthtyps(i-1)inletprec,side=unparsing_precedence_of_entry_typefrom_leveltypinletsl'=(* If no separator: add a break *)ifList.is_emptyslthenadd_break1[](* We add NonTerminal for simulation but remove it afterwards *)elsemaketrueslinlethunk=matchtypwith|ETConstr_->UnpListMetaVar(prec,List.mapsndsl',side)|ETBinderisopen->check_open_binderisopenslm;UnpBinderListMetaVar(isopen,List.mapsndsl')|_->assertfalsein(None,hunk)::make_with_spacebprods|[]->[]andmake_with_spacebprods=matchprodswith|Terminals'::prods'->ifis_operators'then(* A rigid space before operator and a breakable after *)(None,UnpTerminal(" "^s'))::add_break_if_none1b(makebprods')elseifis_commas'then(* No space whatsoever before comma *)makebprodselseifis_right_brackets'thenmakebprodselse(* A breakable space between any other two terminals *)add_break_if_none1b(makebprods)|(NonTerminal_|SProdList_)::_->(* A breakable space before a non-terminal *)add_break_if_none1b(makebprods)|Break_::_->(* Rely on user wish *)makebprods|[]->[]inmakefalsesymbols(* Build default printing rules from explicit format *)leterror_format?loc()=user_err?locPp.(str"The format does not match the notation.")letwarn_format_break=CWarnings.create~name:"notation-both-format-and-spaces"~category:"parsing"(fun()->strbrk"Discarding format implicitly indicated by multiple spaces in notation because an explicit format modifier is given.")lethas_ldotsl=List.exists(function(_,UnpTerminals)->String.equals(Id.to_stringNotation_ops.ldots_var)|_->false)lletrecsplit_format_at_ldotshd=function|(loc,UnpTerminals)::fmtwhenString.equals(Id.to_stringNotation_ops.ldots_var)->loc,List.revhd,fmt|u::fmt->check_no_ldots_in_boxu;split_format_at_ldots(u::hd)fmt|[]->raiseExitandcheck_no_ldots_in_box=function|(_,UnpBox(_,fmt))->(tryletloc,_,_=split_format_at_ldots[]fmtinuser_err?locPp.(str("The special symbol \"..\" must occur at the same formatting depth than the variables of which it is the ellipse."))withExit->())|_->()leterror_not_same?loc()=user_err?locPp.(str"The format is not the same on the right- and left-hand sides of the special token \"..\".")letfind_prod_list_locsfmtfmt=(* [fmt] is some [UnpTerminal x :: sfmt @ UnpTerminal ".." :: sfmt @ UnpTerminal y :: rest] *)ifList.is_emptysfmtthen(* No separators; we highlight the sequence "x .." *)Loc.merge_opt(fst(List.hdfmt))(fst(List.hd(List.tlfmt)))else(* A separator; we highlight the separating sequence *)Loc.merge_opt(fst(List.hdsfmt))(fst(List.lastsfmt))letis_blanks=letn=String.lengthsinletrecauxis=i>=n||s.[i]=' '&&aux(i+1)sinaux0sletis_formatting=function|(_,UnpCut_)->true|(_,UnpTerminals)->is_blanks|_->falseletrecis_var_in_recursive_format=function|(_,UnpTerminals)whennot(is_blanks)->true|(loc,UnpBox(b,l))->(matchList.filter(funa->not(is_formattinga))lwith|[a]->is_var_in_recursive_formata|_->error_not_same?loc())|_->falseletreccheck_eq_var_upto_name=function|(_,UnpTerminals1),(_,UnpTerminals2)whennot(is_blanks1&&is_blanks2)||s1=s2->()|(_,UnpBox(b1,l1)),(_,UnpBox(b2,l2))whenb1=b2->List.itercheck_eq_var_upto_name(List.combinel1l2)|(_,UnpCutb1),(_,UnpCutb2)whenb1=b2->()|_,(loc,_)->error_not_same?loc()letskip_var_in_recursive_format=function|a::slwhenis_var_in_recursive_formata->a,sl|(loc,_)::_->error_not_same?loc()|[]->assertfalseletread_recursive_formatslfmt=(* Turn [[UnpTerminal s :: some-list @ UnpTerminal ".." :: same-some-list @ UnpTerminal s' :: rest] *)(* into [(some-list,rest)] *)letget_headfmt=letvar,sl=skip_var_in_recursive_formatfmtintryvar,split_format_at_ldots[]slwithExit->error_not_same?loc:(fst(List.last(ifsl=[]thenfmtelsesl)))()inletrecget_tail=function|(loc,a)::sepfmt,(_,b)::fmtwhen(=)ab->get_tail(sepfmt,fmt)(* FIXME *)|[],tail->skip_var_in_recursive_formattail|(loc,_)::_,([]|(_,UnpTerminal_)::_)->error_not_same?loc()|_,(loc,_)::_->error_not_same?loc()inletvar1,(loc,slfmt,fmt)=get_headfmtinletvar2,res=get_tail(slfmt,fmt)incheck_eq_var_upto_name(var1,var2);(* To do, though not so important: check that the names match
the names in the notation *)slfmt,reslethunks_of_format(from_level,(vars,typs))symfmt=letrecaux=function|symbs,(_,(UnpTerminals'asu))::fmtwhenString.equals'(String.make(String.lengths')' ')->letsymbs,l=aux(symbs,fmt)insymbs,u::l|Terminals::symbs,(_,UnpTerminals')::fmtwhenString.equals(String.drop_simple_quotess')->letsymbs,l=aux(symbs,fmt)insymbs,UnpTerminals::l|NonTerminals::symbs,(_,UnpTerminals')::fmtwhenId.equals(Id.of_strings')->leti=index_idsvarsinletsymbs,l=aux(symbs,fmt)insymbs,unparsing_metavarifrom_leveltyps::l|symbs,(_,(UnpCut_asu))::fmt->letsymbs,l=aux(symbs,fmt)insymbs,u::l|SProdList(m,sl)::symbs,fmtwhenhas_ldotsfmt->leti=index_idmvarsinlettyp=List.nthtyps(i-1)inletprec,side=unparsing_precedence_of_entry_typefrom_leveltypinletloc_slfmt,rfmt=read_recursive_formatslfmtinletsl,slfmt=aux(sl,loc_slfmt)inifnot(List.is_emptysl)thenerror_format?loc:(find_prod_list_locloc_slfmtfmt)();letsymbs,l=aux(symbs,rfmt)inlethunk=matchtypwith|ETConstr_->UnpListMetaVar(prec,slfmt,side)|ETBinderisopen->check_open_binderisopenslm;UnpBinderListMetaVar(isopen,slfmt)|_->assertfalseinsymbs,hunk::l|symbs,(_,UnpBox(a,b))::fmt->letsymbs',b'=aux(symbs,b)inletsymbs',l=aux(symbs',fmt)insymbs',UnpBox(a,List.map(funx->(None,x))b')::l|symbs,[]->symbs,[]|Break_::symbs,fmt->warn_format_break();aux(symbs,fmt)|_,fmt->error_format?loc:(fst(List.hdfmt))()inmatchauxsymfmtwith|[],l->l|_->error_format()(**********************************************************************)(* Build parsing rules *)letassoc_of_typefromn(_,typ)=precedence_of_entry_type(from,n)typletis_not_small_constr=functionETProdConstr_->true|_->falseletrecdefine_keywords_aux=function|GramConstrNonTerminal(e,Some_)asn1::GramConstrTerminal(Tok.PIDENT(Somek))::lwhenis_not_small_constre->Flags.if_verboseFeedback.msg_info(str"Identifier '"++strk++str"' now a keyword");CLexer.add_keywordk;n1::GramConstrTerminal(Tok.PKEYWORDk)::define_keywords_auxl|n::l->n::define_keywords_auxl|[]->[](* Ensure that IDENT articulation terminal symbols are keywords *)letdefine_keywords=function|GramConstrTerminal(Tok.PIDENT(Somek))::l->Flags.if_verboseFeedback.msg_info(str"Identifier '"++strk++str"' now a keyword");CLexer.add_keywordk;GramConstrTerminal(Tok.PKEYWORDk)::define_keywords_auxl|l->define_keywords_auxlletdistributeall=List.map(funl->a@l)ll(* Expand LIST1(t,sep);sep;t;...;t (with the trailing pattern
occurring p times, possibly p=0) into the combination of
t;sep;t;...;t;sep;t (p+1 times)
t;sep;t;...;t;sep;t;sep;t (p+2 times)
...
t;sep;t;...;t;sep;t;...;t;sep;t (p+n times)
t;sep;t;...;t;sep;t;...;t;sep;t;LIST1(t,sep) *)letexpand_list_rulestyptklxnpll=letcamlp5_message_name=Some(add_suffixx("_"^string_of_intn))inletmain=GramConstrNonTerminal(ETProdConstr(s,typ),camlp5_message_name)inlettks=List.map(funx->GramConstrTerminalx)tklinletrecauxihdsll=ifi<pthenaux(i+1)(main::tks@hds)llelseifInt.equali(p+n)thenlethds=GramConstrListMark(p+n,true,p)::hds@[GramConstrNonTerminal(ETProdConstrList(s,typ,tkl),Somex)]indistributehdsllelsedistribute(GramConstrListMark(i+1,false,p)::hds@[main])ll@aux(i+1)(main::tks@hds)llinaux0[]llletis_constr_typ(s,lev)xetyps=matchList.assocxetypswith(* TODO: factorize these rules with the ones computing the effective
sublevel sent to camlp5, so as to include the case of
DefaultLevel which are valid *)|ETConstr(s',_,(lev',InternalProd|(NumLevel_|NextLevelaslev'),_))->Notation.notation_entry_eqss'&&production_level_eqlevlev'|_->falseletinclude_possible_similar_trailing_patterntypetypssll=letrecauxn=function|Terminals::sl,Terminals'::l'whens=s'->auxn(sl,l')|[],NonTerminalx::l'whenis_constr_typtypxetyps->try_auxnl'|Break_::sl,l->auxn(sl,l)|sl,Break_::l->auxn(sl,l)|_->raiseExitandtry_auxnl=tryaux(n+1)(sl,l)withExit->n,lintry_aux0lletprod_entry_type=function|ETIdent->ETProdIdent|ETName_->ETProdName|ETGlobal->ETProdReference|ETBigint->ETProdBigint|ETBindero->ETProdOneBindero|ETConstr(s,_,p)->ETProdConstr(s,p)|ETPattern(_,n)->ETProdPattern(matchnwithNone->0|Somen->n)letmake_productionetypssymbols=letrecaux=function|[]->[[]]|NonTerminalm::l->lettyp=List.assocmetypsindistribute[GramConstrNonTerminal(prod_entry_typetyp,Somem)](auxl)|Terminals::l->distribute[GramConstrTerminal(CLexer.terminals)](auxl)|Break_::l->auxl|SProdList(x,sl)::l->lettkl=List.flatten(List.map(functionTerminals->[CLexer.terminals]|Break_->[]|_->anomaly(Pp.str"Found a non terminal token in recursive notation separator."))sl)inmatchList.assocxetypswith|ETConstr(s,_,(lev,_astyp))->letp,l'=include_possible_similar_trailing_pattern(s,lev)etypssllinexpand_list_rulestyptklx1p(auxl')|ETBindero->check_open_binderoslx;lettyp=ifothen(assert(tkl=[]);ETBinderOpen)elseETBinderClosedtklindistribute[GramConstrNonTerminal(ETProdBinderListtyp,Somex)](auxl)|_->user_errPp.(str"Components of recursive patterns in notation must be terms or binders.")inletprods=auxsymbolsinList.mapdefine_keywordsprodsletrecfind_symbolsc_currentc_nextc_last=function|[]->[]|NonTerminalid::sl->letprec=ifnot(List.is_emptysl)thenc_currentelsec_lastin(id,prec)::(find_symbolsc_nextc_nextc_lastsl)|Terminals::sl->find_symbolsc_nextc_nextc_lastsl|Breakn::sl->find_symbolsc_currentc_nextc_lastsl|SProdList(x,_)::sl'->(x,c_next)::(find_symbolsc_nextc_nextc_lastsl')letborder=function|(_,(ETConstr(_,_,(_,BorderProd(_,a)))))::_->a|_->Noneletrecompute_assoctyps=letopenGramlib.Gramextinmatchbordertyps,border(List.revtyps)with|SomeLeftA,SomeRightA->assertfalse|SomeLeftA,_->SomeLeftA|_,SomeRightA->SomeRightA|_->None(**************************************************************************)(* Registration of syntax extensions (parsing/printing, no interpretation)*)letpr_arg_levelfrom(lev,typ)=letpplev=function|LevelLtnwhenInt.equalnfrom->spc()++str"at next level"|LevelLen->spc()++str"at level "++intn|LevelLtn->spc()++str"at level below "++intn|LevelSome->mt()inPpvernac.pr_set_entry_type(fun_->(*TO CHECK*)mt())typ++pplevlevletpr_levelntn(from,fromlevel,args)typs=(matchfromwithInConstrEntry->mt()|InCustomEntrys->str"in "++strs++spc())++str"at level "++intfromlevel++spc()++str"with arguments"++spc()++prlist_with_seppr_comma(pr_arg_levelfromlevel)(List.combineargstyps)leterror_incompatible_levelntnoldprecoldtypsprectyps=user_err(str"Notation "++pr_notationntn++str" is already defined"++spc()++pr_levelntnoldprecoldtyps++spc()++str"while it is now required to be"++spc()++pr_levelntnprectyps++str".")leterror_parsing_incompatible_levelntnntn'oldprecoldtypsprectyps=user_err(str"Notation "++pr_notationntn++str" relies on a parsing rule for "++pr_notationntn'++spc()++str" which is already defined"++spc()++pr_levelntnoldprecoldtyps++spc()++str"while it is now required to be"++spc()++pr_levelntnprectyps++str".")letwarn_incompatible_format=CWarnings.create~name:"notation-incompatible-format"~category:"parsing"(fun(specific,ntn)->lethead,scope=matchspecificwith|None->str"Notation",mt()|SomeLastLonelyNotation->str"Lonely notation",mt()|Some(NotationInScopesc)->str"Notation",strbrk(" in scope "^sc)inhead++spc()++pr_notationntn++strbrk" was already defined with a different format"++scope++str".")typesyntax_parsing_extension={synext_level:Notation.level;synext_notation:notation;synext_notgram:notation_grammaroption;synext_nottyps:Extend.constr_entry_keylist;}typesyntax_printing_extension={synext_reserved:bool;synext_unparsing:unparsing_rule;synext_extra:(string*string)list;}letgeneric_format_to_declarentn{synext_unparsing=(rules,_);synext_extra=extra_rules}=trylet(generic_rules,_),reserved,generic_extra_rules=Ppextend.find_generic_notation_printing_rulentninifreserved&&(not(List.for_all2equnparsing_eqrulesgeneric_rules)||extra_rules<>generic_extra_rules)then(warn_incompatible_format(None,ntn);true)elsefalsewithNot_found->trueletcheck_reserved_formatntn=function|None->()|Somesy_pp_rules->let_=generic_format_to_declarentnsy_pp_rulesin()letspecific_format_to_declare(specific,ntnasspecific_ntn){synext_unparsing=(rules,_);synext_extra=extra_rules}=trylet(specific_rules,_),specific_extra_rules=Ppextend.find_specific_notation_printing_rulespecific_ntninifnot(List.for_all2equnparsing_eqrulesspecific_rules)||extra_rules<>specific_extra_rulesthen(warn_incompatible_format(Somespecific,ntn);true)elsefalsewithNot_found->truetypesyntax_extension_obj=locality_flag*(syntax_parsing_extension*syntax_printing_extensionoption)letcheck_and_extend_constr_grammarntnrule=tryletntn_for_grammar=rule.notgram_notationinifnotation_eqntnntn_for_grammarthenraiseNot_found;letprec=rule.notgram_levelinlettyps=rule.notgram_typsinletoldprec=Notation.level_of_notationntn_for_grammarinletoldparsing=trySome(Notgram_ops.grammar_of_notationntn_for_grammar)withNot_found->Noneinletoldtyps=Notgram_ops.subentries_of_notationntn_for_grammarinifnot(Notation.level_eqprecoldprec)&&oldparsing<>Nonethenerror_parsing_incompatible_levelntnntn_for_grammaroldprecoldtypsprectyps;ifoldparsing=NonethenraiseNot_foundwithNot_found->Egramcoq.extend_constr_grammarruleletcache_one_syntax_extension(pa_se,pp_se)=letntn=pa_se.synext_notationinletprec=pa_se.synext_levelin(* Check and ensure that the level and the precomputed parsing rule is declared *)letoldparsing=tryletoldprec=Notation.level_of_notationntninletoldparsing=trySome(Notgram_ops.grammar_of_notationntn)withNot_found->Noneinletoldtyps=Notgram_ops.subentries_of_notationntninifnot(Notation.level_eqprecoldprec)&&(oldparsing<>None||pa_se.synext_notgram=None)thenerror_incompatible_levelntnoldprecoldtypsprecpa_se.synext_nottyps;oldparsingwithNot_found->(* Declare the level and the precomputed parsing rule *)let()=Notation.declare_notation_levelntnprecinlet()=Notgram_ops.declare_notation_subentriesntnpa_se.synext_nottypsinlet()=Option.iter(Notgram_ops.declare_notation_grammarntn)pa_se.synext_notgraminNonein(* Declare the parsing rule *)beginmatcholdparsing,pa_se.synext_notgramwith|None,Somegrams->List.iter(check_and_extend_constr_grammarntn)grams|_->(* The grammars rules are canonically derived from the string and the precedence*)()end;(* Printing *)matchpp_sewith|None->()|Somepp_se->(* Check compatibility of format in case of two Reserved Notation *)(* and declare or redeclare printing rule *)ifgeneric_format_to_declarentnpp_sethendeclare_generic_notation_printing_rulesntn~extra:pp_se.synext_extra~reserved:pp_se.synext_reservedpp_se.synext_unparsingletcache_syntax_extension(_,(_,sy))=cache_one_syntax_extensionsyletsubst_parsing_rulesubstx=xletsubst_printing_rulesubstx=xletsubst_syntax_extension(subst,(local,(pa_sy,pp_sy)))=(local,({pa_sywithsynext_notgram=Option.map(List.map(subst_parsing_rulesubst))pa_sy.synext_notgram},Option.map(funpp_sy->{pp_sywithsynext_unparsing=subst_printing_rulesubstpp_sy.synext_unparsing})pp_sy))letclassify_syntax_definition(local,_aso)=iflocalthenDisposeelseSubstituteoletopen_syntax_extensionio=ifInt.equali1thencache_syntax_extensionoletinSyntaxExtension:syntax_extension_obj->obj=declare_object{(default_object"SYNTAX-EXTENSION")withopen_function=simple_openopen_syntax_extension;cache_function=cache_syntax_extension;subst_function=subst_syntax_extension;classify_function=classify_syntax_definition}(**************************************************************************)(* Precedences *)(* Interpreting user-provided modifiers *)(* XXX: We could move this to the parser itself *)moduleSyndefMods=structtypet={only_parsing:bool;scopes:(Id.t*scope_name)list;}letdefault={only_parsing=false;scopes=[];}endletinterp_syndef_modifiersmodl=letopenSyndefModsinletrecinterpskippedacc=function|[]->List.revskipped,acc|SetOnlyParsing::l->interpskipped{accwithonly_parsing=true;}l|SetItemScope([],_)::l->interpskippedaccl|SetItemScope(s::ids,k)::l->letscopes=acc.scopesinletid=Id.of_stringsinifList.mem_associdscopesthenuser_err(str"Notation scope for argument "++strs++str" can be specified only once.");interpskipped{accwithscopes=(id,k)::scopes}(SetItemScope(ids,s)::l)|x::l->interp(x::skipped)acclininterp[]defaultmodlmoduleNotationMods=structtypenotation_modifier={assoc:Gramlib.Gramext.g_assocoption;level:intoption;custom:notation_entry;etyps:(Id.t*simple_constr_prod_entry_key)list;subtyps:(Id.t*production_level)list;(* common to syn_data below *)only_printing:bool;format:lstringoption;extra:(string*string)list;}letdefault={assoc=None;level=None;custom=InConstrEntry;etyps=[];subtyps=[];only_printing=false;format=None;extra=[];}end(* To be turned into a fatal warning in 8.14 *)letwarn_deprecated_ident_entry=CWarnings.create~name:"deprecated-ident-entry"~category:"deprecated"(fun()->strbrk"grammar entry \"ident\" permitted \"_\" in addition to proper identifiers; this use is deprecated and its meaning will change in the future; use \"name\" instead.")letinterp_modifiersmodl=letopenNotationModsinletrecinterpsubtypsacc=function|[]->subtyps,acc|SetEntryType(s,typ)::l->letid=Id.of_stringsinifId.List.mem_associdacc.etypsthenuser_err~hdr:"Metasyntax.interp_modifiers"(strs++str" is already assigned to an entry or constr level.");interpsubtyps{accwithetyps=(id,typ)::acc.etyps;}l|SetItemLevel([],bko,n)::l->interpsubtypsaccl|SetItemLevel(s::idl,bko,n)::l->letid=Id.of_stringsinifId.List.mem_associdacc.etypsthenuser_err~hdr:"Metasyntax.interp_modifiers"(strs++str" is already assigned to an entry or constr level.");interp((id,bko,n)::subtyps)acc(SetItemLevel(idl,bko,n)::l)|SetLeveln::l->(matchacc.customwith|InCustomEntrys->ifacc.level<>Nonethenuser_err(str("isolated \"at level "^string_of_intn^"\" unexpected."))elseuser_err(str("use \"in custom "^s^" at level "^string_of_intn^"\"")++spc()++str"rather than"++spc()++str("\"at level "^string_of_intn^"\"")++spc()++str"isolated.")|InConstrEntry->ifacc.level<>Nonethenuser_err(str"A level is already assigned.");interpsubtyps{accwithlevel=Somen;}l)|SetCustomEntry(s,n)::l->ifacc.level<>Nonethen(ifn=Nonethenuser_err(str("use \"in custom "^s^" at level "^string_of_int(Option.getacc.level)^"\"")++spc()++str"rather than"++spc()++str("\"at level "^string_of_int(Option.getacc.level)^"\"")++spc()++str"isolated.")elseuser_err(str("isolated \"at level "^string_of_int(Option.getacc.level)^"\" unexpected.")));ifacc.custom<>InConstrEntrythenuser_err(str"Entry is already assigned to custom "++strs++(matchacc.levelwithNone->mt()|Somelev->str" at level "++intlev)++str".");interpsubtyps{accwithcustom=InCustomEntrys;level=n}l|SetAssoca::l->ifnot(Option.is_emptyacc.assoc)thenuser_errPp.(str"An associativity is given more than once.");interpsubtyps{accwithassoc=Somea;}l|SetOnlyPrinting::l->interpsubtyps{accwithonly_printing=true;}l|SetFormat("text",s)::l->ifnot(Option.is_emptyacc.format)thenuser_errPp.(str"A format is given more than once.");interpsubtyps{accwithformat=Somes;}l|SetFormat(k,s)::l->interpsubtyps{accwithextra=(k,s.CAst.v)::acc.extra;}l|(SetOnlyParsing|SetItemScope_)::_->assertfalseinletmodl,syndef_mods=interp_syndef_modifiersmodlinletsubtyps,mods=interp[]defaultmodlin(* interpret item levels wrt to main entry *)letextra_etyps=List.map(fun(id,bko,n)->(id,ETConstr(mods.custom,bko,n)))subtypsin(* Temporary hack: "ETName false" (i.e. "ident" in deprecation phase) means "ETIdent" for custom entries *)letmods={modswithetyps=List.map(function|(id,ETNamefalse)->ifmods.custom=InConstrEntrythen(warn_deprecated_ident_entry();(id,ETNametrue))else(id,ETIdent)|x->x)mods.etyps}insyndef_mods,{modswithetyps=extra_etyps@mods.etyps}letcheck_infix_modifiersmodifiers=let_,mods=interp_modifiersmodifiersinlett=mods.NotationMods.etypsinletu=mods.NotationMods.subtypsinifnot(List.is_emptyt)||not(List.is_emptyu)thenuser_errPp.(str"Explicit entry level or type unexpected in infix notation.")letcheck_useless_entry_typesrecvarsmainvarsetyps=letvars=let(l1,l2)=List.splitrecvarsinl1@l2@mainvarsinmatchList.filter(fun(x,etyp)->not(List.memxvars))etypswith|(x,_)::_->user_err~hdr:"Metasyntax.check_useless_entry_types"(Id.printx++str" is unbound in the notation.")|_->()letinterp_non_syntax_modifiersmods=letsetmodif(only_parsing,only_printing,entry)=matchmodifwith|SetOnlyParsing->Some(true,only_printing,entry)|SetOnlyPrinting->Some(only_parsing,true,entry)|SetCustomEntry(entry,None)->Some(only_parsing,only_printing,InCustomEntryentry)|_->NoneinList.fold_left(funstmodif->Option.bindst@@setmodif)(Some(false,false,InConstrEntry))mods(* Check if an interpretation can be used for printing a cases printing *)lethas_no_binders_type=List.for_all(fun(_,(_,typ))->matchtypwith|NtnTypeBinder_|NtnTypeBinderList->false|NtnTypeConstr|NtnTypeConstrList->true)(* Compute precedences from modifiers (or find default ones) *)letset_entry_typefromnetyps(x,typ)=letmake_levns=matchtypwith|BorderProd_->NumLeveln|InternalProd->DefaultLevelinlettyp=trymatchList.assocxetyps,typwith|ETConstr(s,bko,DefaultLevel),_->ifnotation_entry_eqfromsthenETConstr(s,bko,(make_levns,typ))elseETConstr(s,bko,(DefaultLevel,typ))|ETConstr(s,bko,n),BorderProd(left,_)->ETConstr(s,bko,(n,BorderProd(left,None)))|ETConstr(s,bko,n),InternalProd->ETConstr(s,bko,(n,InternalProd))|ETPattern(b,n),_->ETPattern(b,n)|(ETIdent|ETName_|ETBigint|ETGlobal|ETBinder_asx),_->xwithNot_found->ETConstr(from,None,(make_levnfrom,typ))in(x,typ)letjoin_auxiliary_recursive_typesrecvarsetyps=List.fold_right(fun(x,y)typs->letxtyp=trySome(List.assocxetyps)withNot_found->Noneinletytyp=trySome(List.assocyetyps)withNot_found->Noneinmatchxtyp,ytypwith|None,None->typs|Some_,None->typs|None,Someytyp->(x,ytyp)::typs|Somextyp,Someytypwhen(=)xtypytyp->typs(* FIXME *)|Somextyp,Someytyp->user_err(strbrk"In "++Id.printx++str" .. "++Id.printy++strbrk", both ends have incompatible types."))recvarsetypsletinternalization_type_of_entry_type=function|ETBinder_->NtnInternTypeOnlyBinder|ETConstr_|ETBigint|ETGlobal|ETIdent|ETName_|ETPattern_->NtnInternTypeAnyNoneletset_internalization_typetyps=List.map(fun(_,e)->internalization_type_of_entry_typee)typsletmake_internalization_varsrecvarsmainvarstyps=letmaintyps=List.combinemainvarstypsinletextratyps=List.map(fun(x,y)->(y,List.assocxmaintyps))recvarsinmaintyps@extratypsletmake_interpretation_typeisrecisonlybindingdefault_if_binding=function(* Parsed as constr list *)|ETConstr(_,None,_)whenisrec->NtnTypeConstrList(* Parsed as constr, but interpreted as a binder *)|ETConstr(_,Somebk,_)->NtnTypeBinder(NtnBinderParsedAsConstrbk)|ETConstr(_,None,_)whenisonlybinding->NtnTypeBinder(NtnBinderParsedAsConstrdefault_if_binding)(* Parsed as constr, interpreted as constr *)|ETConstr(_,None,_)->NtnTypeConstr(* Others *)|ETIdent->NtnTypeBinderNtnParsedAsIdent|ETName_->NtnTypeBinderNtnParsedAsName|ETPattern(ppstrict,_)->NtnTypeBinder(NtnParsedAsPatternppstrict)(* Parsed as ident/pattern, primarily interpreted as binder; maybe strict at printing *)|ETBigint|ETGlobal->NtnTypeConstr|ETBinder_->ifisrecthenNtnTypeBinderListelseNtnTypeBinderNtnParsedAsBinderletsubentry_of_constr_prod_entryfrom_level=function(* Specific 8.2 approximation *)|ETConstr(InCustomEntrys,_,x)->letn=matchfst(precedence_of_position_and_levelfrom_levelx)with|LevelLtn->n-1|LevelLen->n|LevelSome->max_intinInCustomEntryLevel(s,n)(* level and use of parentheses for coercion is hard-wired for "constr";
we don't remember the level *)|ETConstr(InConstrEntry,_,_)->InConstrEntrySomeLevel|_->InConstrEntrySomeLevelletmake_interpretation_vars(* For binders, default is to parse only as an ident *)?(default_if_binding=AsName)recvarslevelallvarstyps=leteq_subscope(sc1,l1)(sc2,l2)=Option.equalString.equalsc1sc2&&List.equalString.equall1l2inletcheck(x,y)=let(_,scope1)=Id.Map.findxallvarsinlet(_,scope2)=Id.Map.findyallvarsinifnot(eq_subscopescope1scope2)thenerror_not_same_scopexyinlet()=List.itercheckrecvarsinletuseless_recvars=List.mapsndrecvarsinletmainvars=Id.Map.filter(funx_->not(Id.List.memxuseless_recvars))allvarsinId.Map.mapi(funx(isonlybinding,sc)->lettyp=Id.List.assocxtypsin((subentry_of_constr_prod_entryleveltyp,sc),make_interpretation_type(Id.List.mem_assocxrecvars)isonlybindingdefault_if_bindingtyp))mainvarsletcheck_rule_productivityl=ifList.for_all(functionNonTerminal_|Break_->true|_->false)lthenuser_errPp.(str"A notation must include at least one symbol.");if(matchlwithSProdList_::_->true|_->false)thenuser_errPp.(str"A recursive notation must start with at least one symbol.")letwarn_notation_bound_to_variable=CWarnings.create~name:"notation-bound-to-variable"~category:"parsing"(fun()->strbrk"This notation will not be used for printing as it is bound to a single variable.")letwarn_non_reversible_notation=CWarnings.create~name:"non-reversible-notation"~category:"parsing"(function|APrioriReversible->assertfalse|HasLtac->strbrk"This notation contains Ltac expressions: it will not be used for printing."|NonInjectiveids->letn=List.lengthidsinstrbrk(String.pluraln"Variable")++spc()++pr_enumId.printids++spc()++strbrk(ifn>1then"do"else"does")++str" not occur in the right-hand side."++spc()++strbrk"The notation will not be used for printing as it is not reversible.")letis_coercionleveltyps=matchlevel,typswith|Some(custom,n,_),[e]->(matche,customwith|ETConstr_,_->letcustomkey=make_notation_entry_levelcustomninletsubentry=subentry_of_constr_prod_entryneinifnotation_entry_level_eqsubentrycustomkeythenNoneelseSome(IsEntryCoercionsubentry)|ETGlobal,InCustomEntrys->Some(IsEntryGlobal(s,n))|ETIdent,InCustomEntrys->Some(IsEntryIdent(s,n))|_->None)|Some_,_->assertfalse|None,_->Noneletprintabilityleveltypsonlyparsereversibility=function|NVar_whenreversibility=APrioriReversible->letcoe=is_coercionleveltypsinifnotonlyparse&&coe=Nonethenwarn_notation_bound_to_variable();true,coe|_->(ifnotonlyparse&&reversibility<>APrioriReversiblethen(warn_non_reversible_notationreversibility;true)elseonlyparse),Noneletfind_precedencecustomlevetypssymbolsonlyprint=letfirst_symbol=letrecaux=function|Break_::t->auxt|h::t->Someh|[]->Noneinauxsymbolsinletlast_is_terminal()=letrecauxb=function|Break_::t->auxbt|Terminal_::t->auxtruet|_::t->auxfalset|[]->binauxfalsesymbolsinmatchfirst_symbolwith|None->[],0|Some(NonTerminalx)->lettest()=ifonlyprintthenifOption.is_emptylevthenuser_errPp.(str"Explicit level needed in only-printing mode when the level of the leftmost non-terminal is given.")else[],Option.getlevelseuser_errPp.(str"The level of the leftmost non-terminal cannot be changed.")in(trymatchList.assocxetyps,customwith|ETConstr(s,_,(NumLevel_|NextLevel)),s'whens=s'->test()|(ETIdent|ETName_|ETBigint|ETGlobal),_->beginmatchlevwith|None->([fun()->Flags.if_verbose(Feedback.msg_info?loc:None)(strbrk"Setting notation at level 0.")],0)|Some0->([],0)|_->user_errPp.(str"A notation starting with an atomic expression must be at level 0.")end|(ETPattern_|ETBinder_),InConstrEntrywhennotonlyprint->(* Don't know exactly if we can make sense of this case *)user_errPp.(str"Binders or patterns not supported in leftmost position.")|(ETPattern_|ETBinder_|ETConstr_),_->(* Give a default ? *)ifOption.is_emptylevthenuser_errPp.(str"Need an explicit level.")else[],Option.getlevwithNot_found->ifOption.is_emptylevthenuser_errPp.(str"A left-recursive notation must have an explicit level.")else[],Option.getlev)|Some(Terminal_)whenlast_is_terminal()->ifOption.is_emptylevthen([fun()->Flags.if_verbose(Feedback.msg_info?loc:None)(strbrk"Setting notation at level 0.")],0)else[],Option.getlev|Some_->ifOption.is_emptylevthenuser_errPp.(str"Cannot determine the level.");[],Option.getlevletcheck_curly_brackets_notation_exists()=trylet_=Notation.level_of_notation(InConstrEntry,"{ _ }")in()withNot_found->user_errPp.(str"Notations involving patterns of the form \"{ _ }\" are treated \n\
specially and require that the notation \"{ _ }\" is already reserved.")(* Remove patterns of the form "{ _ }", unless it is the "{ _ }" notation *)letremove_curly_bracketsl=letrecskip_breakacc=function|Break_asbr::l->skip_break(br::acc)l|l->List.revacc,linletrecauxdeb=function|[]->[]|Terminal"{"ast1::l->letbr,next=skip_break[]lin(matchnextwith|NonTerminal_asx::l'->letbr',next'=skip_break[]l'in(matchnext'with|Terminal"}"ast2::l''->ifdeb&&List.is_emptyl''then[t1;x;t2]elsebegincheck_curly_brackets_notation_exists();x::auxfalsel''end|l1->t1::br@x::br'@auxfalsel1)|l0->t1::auxfalsel0)|x::l->x::auxfalselinauxtruelmoduleSynData=structtypesubentry_types=(Id.t*constr_entry_key)list(* XXX: Document *)typesyn_data={(* Notation name and location *)info:notation*notation_location;(* Fields coming from the vernac-level modifiers *)only_parsing:bool;only_printing:bool;deprecation:Deprecation.toption;format:lstringoption;extra:(string*string)list;(* XXX: Callback to printing, must remove *)msgs:(unit->unit)list;(* Fields for internalization *)recvars:(Id.t*Id.t)list;mainvars:Id.List.eltlist;intern_typs:notation_var_internalization_typelist;(* Notation data for parsing *)level:level;subentries:constr_entry_keylist;pa_syntax_data:subentry_types*symbollist;pp_syntax_data:subentry_types*symbollist;not_data:notation*(* notation *)level*(* level, precedence *)constr_entry_keylist*bool;(* needs_squash *)}endletfind_subentry_typesfromnassocetypssymbols=lettyps=find_symbols(BorderProd(Left,assoc))(InternalProd)(BorderProd(Right,assoc))symbolsinletsy_typs=List.map(set_entry_typefromnetyps)typsinletprec=List.map(assoc_of_typefromn)sy_typsinsy_typs,precletcheck_locality_compatibilitylocalcustomi_typs=ifnotlocalthenletsubcustom=List.map_filter(function_,ETConstr(InCustomEntrys,_,_)->Somes|_->None)i_typsinletallcustoms=matchcustomwithInCustomEntrys->s::subcustom|_->subcustominList.iter(funs->ifEgramcoq.locality_of_custom_entrysthenuser_err(strbrk"Notation has to be declared local as it depends on custom entry "++strs++strbrk" which is local."))(List.uniquizeallcustoms)letcompute_syntax_data~localdeprecationdfmodifiers=letopenSynDatainletopenSyndefModsinletopenNotationModsinletsyndef_mods,mods=interp_modifiersmodifiersinletonly_printing=mods.only_printinginletonly_parsing=syndef_mods.only_parsinginifonly_printing&&only_parsingthenuser_err(str"A notation cannot be both 'only printing' and 'only parsing'.");ifsyndef_mods.scopes<>[]thenuser_err(str"General notations don't support 'in scope'.");letassoc=Option.appendmods.assoc(SomeGramlib.Gramext.NonA)inlet(recvars,mainvars,symbols)=analyze_notation_tokens~onlyprint:only_printingdfinlet_=check_useless_entry_typesrecvarsmainvarsmods.etypsin(* Notations for interp and grammar *)letmsgs,n=find_precedencemods.custommods.levelmods.etypssymbolsonly_printinginletntn_for_interp=make_notation_keymods.customsymbolsinletsymbols_for_grammar=ifmods.custom=InConstrEntrythenremove_curly_bracketssymbolselsesymbolsinletneed_squash=not(List.equalNotation.symbol_eqsymbolssymbols_for_grammar)inletntn_for_grammar=ifneed_squashthenmake_notation_keymods.customsymbols_for_grammarelsentn_for_interpinifmods.custom=InConstrEntry&¬only_printingthencheck_rule_productivitysymbols_for_grammar;(* To globalize... *)letetyps=join_auxiliary_recursive_typesrecvarsmods.etypsinletsy_typs,prec=find_subentry_typesmods.customnassocetypssymbolsinletsy_typs_for_grammar,prec_for_grammar=ifneed_squashthenfind_subentry_typesmods.customnassocetypssymbols_for_grammarelsesy_typs,precinleti_typs=set_internalization_typesy_typsincheck_locality_compatibilitylocalmods.customsy_typs;letpa_sy_data=(sy_typs_for_grammar,symbols_for_grammar)inletpp_sy_data=(sy_typs,symbols)inletsy_fulldata=(ntn_for_grammar,(mods.custom,n,prec_for_grammar),List.mapsndsy_typs_for_grammar,need_squash)inletdf'=((Lib.library_dp(),Lib.current_dirpathtrue),df)inleti_data=ntn_for_interp,df'in(* Return relevant data for interpretation and for parsing/printing *){info=i_data;only_parsing;only_printing;deprecation;format=mods.format;extra=mods.extra;msgs;recvars;mainvars;intern_typs=i_typs;level=(mods.custom,n,prec);subentries=List.mapsndsy_typs;pa_syntax_data=pa_sy_data;pp_syntax_data=pp_sy_data;not_data=sy_fulldata;}letwarn_only_parsing_reserved_notation=CWarnings.create~name:"irrelevant-reserved-notation-only-parsing"~category:"parsing"(fun()->strbrk"The only parsing modifier has no effect in Reserved Notation.")letcompute_pure_syntax_data~localdfmods=letopenSynDatainletsd=compute_syntax_data~localNonedfmodsinifsd.only_parsingthenletmsgs=(fun()->warn_only_parsing_reserved_notation?loc:None())::sd.msgsin{sdwithmsgs;only_parsing=false}elsesd(**********************************************************************)(* Registration of notations interpretation *)typenotation_obj={notobj_local:bool;notobj_scope:scope_nameoption;notobj_interp:interpretation;notobj_coercion:entry_coercion_kindoption;notobj_use:notation_useoption;notobj_deprecation:Deprecation.toption;notobj_notation:notation*notation_location;notobj_specific_pp_rules:syntax_printing_extensionoption;notobj_also_in_cases_pattern:bool;}letload_notation_commonsilently_define_scope_if_undefined_(_,nobj)=(* When the default shall be to require that a scope already exists *)(* the call to ensure_scope will have to be removed *)ifsilently_define_scope_if_undefinedthen(* Don't warn if the scope is not defined: *)(* there was already a warning at "cache" time *)Option.iterNotation.declare_scopenobj.notobj_scopeelseOption.iterNotation.ensure_scopenobj.notobj_scopeletload_notation=load_notation_commontrueletopen_notationi(_,nobj)=ifInt.equali1thenbeginletscope=nobj.notobj_scopeinlet(ntn,df)=nobj.notobj_notationinletpat=nobj.notobj_interpinletdeprecation=nobj.notobj_deprecationinletscope=matchscopewithNone->LastLonelyNotation|Somesc->NotationInScopescinletalso_in_cases_pattern=nobj.notobj_also_in_cases_patternin(* Declare the notation *)(matchnobj.notobj_usewith|Someuse->Notation.declare_notation(scope,ntn)patdf~use~also_in_cases_patternnobj.notobj_coerciondeprecation|None->());(* Declare specific format if any *)(matchnobj.notobj_specific_pp_ruleswith|Somepp_sy->ifspecific_format_to_declare(scope,ntn)pp_sythenPpextend.declare_specific_notation_printing_rules(scope,ntn)~extra:pp_sy.synext_extrapp_sy.synext_unparsing|None->())endletcache_notationo=load_notation_commonfalse1o;open_notation1oletsubst_notation(subst,nobj)={nobjwithnotobj_interp=subst_interpretationsubstnobj.notobj_interp;}letclassify_notationnobj=ifnobj.notobj_localthenDisposeelseSubstitutenobjletinNotation:notation_obj->obj=declare_object{(default_object"NOTATION")withopen_function=simple_openopen_notation;cache_function=cache_notation;subst_function=subst_notation;load_function=load_notation;classify_function=classify_notation}(**********************************************************************)letwith_lib_stk_protectionfx=letfs=Lib.freeze()intryleta=fxinLib.unfreezefs;awithreraise->letreraise=Exninfo.capturereraiseinlet()=Lib.unfreezefsinExninfo.iraisereraiseletwith_syntax_protectionfx=with_lib_stk_protection(Pcoq.with_grammar_rule_protection(with_notation_protectionf))x(**********************************************************************)(* Recovering existing syntax *)exceptionNoSyntaxRuleletrecover_notation_syntaxntn=letpa=tryletprec=Notation.level_of_notationntninletpa_typs=Notgram_ops.subentries_of_notationntninletpa_rule=trySome(Notgram_ops.grammar_of_notationntn)withNot_found->Nonein{synext_level=prec;synext_notation=ntn;synext_notgram=pa_rule;synext_nottyps=pa_typs;}withNot_found->raiseNoSyntaxRuleinletpp=tryletpp_rule,reserved,pp_extra_rules=find_generic_notation_printing_rulentninSome{synext_reserved=reserved;synext_unparsing=pp_rule;synext_extra=pp_extra_rules;}withNot_found->Noneinpa,ppletrecover_squash_syntaxsy=letsq,_=recover_notation_syntax(InConstrEntry,"{ _ }")inmatchsq.synext_notgramwith|Somegram->sy::gram|None->raiseNoSyntaxRule(**********************************************************************)(* Main entry point for building parsing and printing rules *)letmake_pa_rulelevelentries(typs,symbols)ntnneed_squash=letassoc=recompute_assoctypsinletprod=make_productiontypssymbolsinletsy={notgram_level=level;notgram_assoc=assoc;notgram_notation=ntn;notgram_prods=prod;notgram_typs=entries;}in(* By construction, the rule for "{ _ }" is declared, but we need to
redeclare it because the file where it is declared needs not be open
when the current file opens (especially in presence of -nois) *)ifneed_squashthenrecover_squash_syntaxsyelse[sy]letmake_pp_rulelevel(typs,symbols)fmt=matchfmtwith|None->lethunks=make_hunkstypssymbolslevelinifList.exists(function_,(UnpCut(PpBrk_)|UnpListMetaVar_)->true|_->false)hunksthen[UnpBox(PpHOVB0,hunks)]else(* Optimization to work around what seems an ocaml Format bug (see Mantis #7804/#7807) *)List.mapsndhunks(* drop locations which are dummy *)|Somefmt->hunks_of_format(level,List.splittyps)(symbols,parse_formatfmt)letmake_parsing_rules(sd:SynData.syn_data)=letopenSynDatainletntn_for_grammar,prec_for_grammar,typs_for_grammar,need_squash=sd.not_datainletpa_rule=ifsd.only_printingthenNoneelseSome(make_pa_ruleprec_for_grammartyps_for_grammarsd.pa_syntax_datantn_for_grammarneed_squash)in{synext_level=sd.level;synext_notation=fstsd.info;synext_notgram=pa_rule;synext_nottyps=typs_for_grammar;}letwarn_irrelevant_format=CWarnings.create~name:"irrelevant-format-only-parsing"~category:"parsing"(fun()->str"The format modifier is irrelevant for only parsing rules.")letmake_printing_rulesreserved(sd:SynData.syn_data)=letopenSynDatainletcustom,level,_=sd.levelinletformat=ifsd.only_parsing&&sd.format<>Nonethen(warn_irrelevant_format();None)elsesd.formatinletpp_rule=make_pp_rulelevelsd.pp_syntax_dataformatin(* We produce a generic rule even if this precise notation is only parsing *)Some{synext_reserved=reserved;synext_unparsing=(pp_rule,level);synext_extra=sd.extra;}letwarn_unused_interpretation=CWarnings.create~name:"unused-notation"~category:"parsing"(funb->strbrk"interpretation is used neither for printing nor for parsing, "++(ifbthenstrbrk"the declaration could be replaced by \"Reserved Notation\"."elsestrbrk"the declaration could be removed."))letmake_usereservedonlyparseonlyprint=matchonlyparse,onlyprintwith|false,false->SomeParsingAndPrinting|true,false->SomeOnlyParsing|false,true->SomeOnlyPrinting|true,true->warn_unused_interpretationreserved;None(**********************************************************************)(* Main functions about notations *)letto_mapl=letfoldaccu(x,v)=Id.Map.addxvaccuinList.fold_leftfoldId.Map.emptylletadd_notation_in_scope~localdeprecationdfenvcmodsscope=letopenSynDatainletsd=compute_syntax_data~localdeprecationdfmodsin(* Prepare the parsing and printing rules *)letsy_pa_rules=make_parsing_rulessdinletsy_pp_rules,gen_sy_pp_rules=matchsd.only_parsing,Ppextend.has_generic_notation_printing_rule(fstsd.info)with|true,true->None,None|onlyparse,has_generic->letrules=make_printing_rulesfalsesdinlet_=check_reserved_format(fstsd.info)rulesin(ifonlyparsethenNoneelserules),(ifhas_genericthenNoneelse(* We use the format of this notation as the default *)rules)in(* Prepare the interpretation *)leti_vars=make_internalization_varssd.recvarssd.mainvarssd.intern_typsinletnenv={ninterp_var_type=to_mapi_vars;ninterp_rec_vars=to_mapsd.recvars;}inlet(acvars,ac,reversibility)=interp_notation_constrenvnenvcinletinterp=make_interpretation_varssd.recvars(pi2sd.level)acvars(fstsd.pa_syntax_data)inletmap(x,_)=trySome(x,Id.Map.findxinterp)withNot_found->Noneinletvars=List.map_filtermapi_varsin(* Order of elements is important here! *)letalso_in_cases_pattern=has_no_binders_typevarsinletonlyparse,coe=printability(Somesd.level)sd.subentriessd.only_parsingreversibilityacinletnotation,location=sd.infoinletuse=make_usetrueonlyparsesd.only_printinginletnotation={notobj_local=local;notobj_scope=scope;notobj_use=use;notobj_interp=(vars,ac);notobj_coercion=coe;notobj_deprecation=sd.deprecation;notobj_notation=(notation,location);notobj_specific_pp_rules=sy_pp_rules;notobj_also_in_cases_pattern=also_in_cases_pattern;}in(* Ready to change the global state *)List.iter(funf->f())sd.msgs;Lib.add_anonymous_leaf(inSyntaxExtension(local,(sy_pa_rules,gen_sy_pp_rules)));Lib.add_anonymous_leaf(inNotationnotation);sd.infoletadd_notation_interpretation_core~localdfenv?(impls=empty_internalization_env)entrycscopeonlyparseonlyprintdeprecation=let(recvars,mainvars,symbs)=analyze_notation_tokens~onlyprintdfin(* Recover types of variables and pa/pp rules; redeclare them if needed *)letnotation_key=make_notation_keyentrysymbsinletlevel,i_typs,onlyprint,pp_sy=ifnot(is_numeral_in_constrentrysymbs)thenbeginlet(pa_sy,pp_syassy)=recover_notation_syntaxnotation_keyinlet()=Lib.add_anonymous_leaf(inSyntaxExtension(local,sy))in(* If the only printing flag has been explicitly requested, put it back *)letonlyprint=onlyprint||pa_sy.synext_notgram=Noneinlettyps=pa_sy.synext_nottypsinSomepa_sy.synext_level,typs,onlyprint,pp_syendelseNone,[],false,Nonein(* Declare interpretation *)letpath=(Lib.library_dp(),Lib.current_dirpathtrue)inletdf'=notation_key,(path,df)inleti_vars=make_internalization_varsrecvarsmainvars(List.mapinternalization_type_of_entry_typei_typs)inletnenv={ninterp_var_type=to_mapi_vars;ninterp_rec_vars=to_maprecvars;}inlet(acvars,ac,reversibility)=interp_notation_constrenv~implsnenvcinletplevel=matchlevelwithSome(from,level,l)->level|None(* numeral: irrelevant )*)->0inletinterp=make_interpretation_varsrecvarsplevelacvars(List.combinemainvarsi_typs)inletmap(x,_)=trySome(x,Id.Map.findxinterp)withNot_found->Noneinletvars=List.map_filtermapi_varsin(* Order of elements is important here! *)letalso_in_cases_pattern=has_no_binders_typevarsinletonlyparse,coe=printabilityleveli_typsonlyparsereversibilityacinletuse=make_usefalseonlyparseonlyprintinletnotation={notobj_local=local;notobj_scope=scope;notobj_use=use;notobj_interp=(vars,ac);notobj_coercion=coe;notobj_deprecation=deprecation;notobj_notation=df';notobj_specific_pp_rules=pp_sy;notobj_also_in_cases_pattern=also_in_cases_pattern;}inLib.add_anonymous_leaf(inNotationnotation);df'(* Notations without interpretation (Reserved Notation) *)letadd_syntax_extension~local({CAst.loc;v=df},mods)=letopenSynDatainletpsd={(compute_pure_syntax_data~localdfmods)withdeprecation=None}inletpa_rules=make_parsing_rulespsdinletpp_rules=make_printing_rulestruepsdinList.iter(funf->f())psd.msgs;Lib.add_anonymous_leaf(inSyntaxExtension(local,(pa_rules,pp_rules)))(* Notations with only interpretation *)letadd_notation_interpretationenvdecl_ntn=let{decl_ntn_string={CAst.loc;v=df};decl_ntn_interp=c;decl_ntn_modifiers=modifiers;decl_ntn_scope=sc;}=decl_ntninmatchinterp_non_syntax_modifiersmodifierswith|None->CErrors.user_err(str"Only modifiers not affecting parsing are supported here.")|Some(only_parsing,only_printing,entry)->letdf'=add_notation_interpretation_core~local:falsedfenventrycsconly_parsingfalseNoneinDumpglob.dump_notation(loc,df')sctrueletset_notation_for_interpretationenvimplsdecl_ntn=let{decl_ntn_string={CAst.v=df};decl_ntn_interp=c;decl_ntn_modifiers=modifiers;decl_ntn_scope=sc;}=decl_ntninmatchinterp_non_syntax_modifiersmodifierswith|None->CErrors.user_err(str"Only modifiers not affecting parsing are supported here")|Some(only_parsing,only_printing,entry)->(tryignore(Flags.silently(fun()->add_notation_interpretation_core~local:falsedfenv~implsentrycsconly_parsingfalseNone)());withNoSyntaxRule->user_errPp.(str"Parsing rule for this notation has to be previously declared."));Option.iter(funsc->Notation.open_close_scope(false,true,sc))sc(* Main entry point *)letadd_notation~localdeprecationenvc({CAst.loc;v=df},modifiers)sc=letdf'=matchinterp_non_syntax_modifiersmodifierswith|Some(only_parsing,only_printing,entry)->(* No syntax data: try to rely on a previously declared rule *)begintryadd_notation_interpretation_core~localdfenventrycsconly_parsingonly_printingdeprecationwithNoSyntaxRule->(* Try to determine a default syntax rule *)add_notation_in_scope~localdeprecationdfenvcmodifiersscend|None->(* Declare both syntax and interpretation *)add_notation_in_scope~localdeprecationdfenvcmodifiersscinDumpglob.dump_notation(loc,df')sctrueletadd_notation_extra_printing_ruledfkv=letnotk=let_,_,symbs=analyze_notation_tokens~onlyprint:truedfinmake_notation_keyInConstrEntrysymbsinadd_notation_extra_printing_rulenotkkv(* Infix notations *)letinject_varx=CAst.make@@CRef(qualid_of_identx,None)letadd_infix~localdeprecationenv({CAst.loc;v=inf},modifiers)prsc=check_infix_modifiersmodifiers;(* check the precedence *)letvars=names_of_constr_exprprinletx=Namegen.next_ident_away(Id.of_string"x")varsinlety=Namegen.next_ident_away(Id.of_string"y")(Id.Set.addxvars)inletmetas=[inject_varx;inject_vary]inletc=mkAppC(pr,metas)inletdf=CAst.make?loc@@Id.to_stringx^" "^(quote_notation_tokeninf)^" "^Id.to_stringyinadd_notation~localdeprecationenvc(df,modifiers)sc(**********************************************************************)(* Scopes, delimiters and classes bound to scopes *)typescope_command=|ScopeDeclare|ScopeDelimAddofstring|ScopeDelimRemove|ScopeClassesofscope_classlistletload_scope_command_commonsilently_define_scope_if_undefined_(_,(local,scope,o))=letdeclare_scope_if_needed=ifsilently_define_scope_if_undefinedthenNotation.declare_scopeelseNotation.ensure_scopeinmatchowith|ScopeDeclare->Notation.declare_scopescope(* When the default shall be to require that a scope already exists *)(* the call to declare_scope_if_needed will have to be removed below *)|ScopeDelimAdddlm->declare_scope_if_neededscope|ScopeDelimRemove->declare_scope_if_neededscope|ScopeClassescl->declare_scope_if_neededscopeletload_scope_command=load_scope_command_commontrueletopen_scope_commandi(_,(local,scope,o))=ifInt.equali1thenmatchowith|ScopeDeclare->()|ScopeDelimAdddlm->Notation.declare_delimitersscopedlm|ScopeDelimRemove->Notation.remove_delimitersscope|ScopeClassescl->List.iter(Notation.declare_scope_classscope)clletcache_scope_commando=load_scope_command_commonfalse1o;open_scope_command1oletsubst_scope_command(subst,(local,scope,oasx))=matchowith|ScopeClassescl->letenv=Global.env()inletcl'=List.map_filter(subst_scope_classenvsubst)clinletcl'=ifList.for_all2eq(==)clcl'thenclelsecl'inlocal,scope,ScopeClassescl'|_->xletclassify_scope_command(local,_,_aso)=iflocalthenDisposeelseSubstituteoletinScopeCommand:locality_flag*scope_name*scope_command->obj=declare_object{(default_object"DELIMITERS")withcache_function=cache_scope_command;open_function=simple_openopen_scope_command;load_function=load_scope_command;subst_function=subst_scope_command;classify_function=classify_scope_command}letdeclare_scopelocalscope=Lib.add_anonymous_leaf(inScopeCommand(local,scope,ScopeDeclare))letadd_delimiterslocalscopekey=Lib.add_anonymous_leaf(inScopeCommand(local,scope,ScopeDelimAddkey))letremove_delimiterslocalscope=Lib.add_anonymous_leaf(inScopeCommand(local,scope,ScopeDelimRemove))letadd_class_scopelocalscopecl=Lib.add_anonymous_leaf(inScopeCommand(local,scope,ScopeClassescl))letadd_syntactic_definition~localdeprecationenvident(vars,c)modl=letopenSyndefModsinletskipped,{only_parsing;scopes}=interp_syndef_modifiersmodlinifskipped<>[]thenuser_err(str"Simple notations don't support "++Ppvernac.pr_syntax_modifier(List.hdskipped));letvars=List.map(funv->v,List.assoc_optvscopes)varsinletacvars,pat,reversibility=matchvars,intern_name_aliascwith|[],Some(r,u)->Id.Map.empty,NRef(r,u),APrioriReversible|_->letfoldaccu(id,scope)=Id.Map.addid(NtnInternTypeAnyscope)accuinleti_vars=List.fold_leftfoldId.Map.emptyvarsinletnenv={ninterp_var_type=i_vars;ninterp_rec_vars=Id.Map.empty;}ininterp_notation_constrenvnenvcinletin_pat(id,_)=(id,ETConstr(Constrexpr.InConstrEntry,None,(NextLevel,InternalProd)))inletinterp=make_interpretation_vars~default_if_binding:AsNameOrPattern[]0acvars(List.mapin_patvars)inletvars=List.map(fun(x,_)->(x,Id.Map.findxinterp))varsinletalso_in_cases_pattern=has_no_binders_typevarsinletonlyparsing=only_parsing||fst(printabilityNone[]falsereversibilitypat)inSyntax_def.declare_syntactic_definition~local~also_in_cases_patterndeprecationident~onlyparsing(vars,pat)(**********************************************************************)(* Declaration of custom entry *)letwarn_custom_entry=CWarnings.create~name:"custom-entry-overridden"~category:"parsing"(funs->strbrk"Custom entry "++strs++strbrk" has been overridden.")letload_custom_entry_(_,(local,s))=ifEgramcoq.exists_custom_entrysthenwarn_custom_entryselseEgramcoq.create_custom_entry~localsletcache_custom_entryo=load_custom_entry1oletsubst_custom_entry(subst,x)=xletclassify_custom_entry(local,saso)=iflocalthenDisposeelseSubstituteoletinCustomEntry:locality_flag*string->obj=declare_object{(default_object"CUSTOM-ENTRIES")withcache_function=cache_custom_entry;load_function=load_custom_entry;subst_function=subst_custom_entry;classify_function=classify_custom_entry}letdeclare_custom_entrylocals=ifEgramcoq.exists_custom_entrysthenuser_errPp.(str"Custom entry "++strs++str" already exists.")elseLib.add_anonymous_leaf(inCustomEntry(local,s))