123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456(************************************************************************)(* * 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) *)(************************************************************************)openUtilopenConstropenContextopenPpopenNamesopenEnvironopenDeclarationsopenLibnamesopenGoptions(** Note: there is currently two modes for printing modules.
- The "short" one, that just prints the names of the fields.
- The "rich" one, that also tries to print the types of the fields.
The short version used to be the default behavior, but now we print
types by default. The following option allows changing this.
*)moduleTag=structletdefinition="module.definition"letkeyword="module.keyword"endlettagts=Pp.tagtslettag_definitions=tagTag.definitionslettag_keywords=tagTag.keywordstypeshort=OnlyNames|WithContentsletshort=reffalselet()=declare_bool_option{optdepr=false;optkey=["Short";"Module";"Printing"];optread=(fun()->!short);optwrite=((:=)short)}(** Each time we have to print a non-globally visible structure,
we place its elements in a fake fresh namespace. *)letmk_fake_top=letr=ref0infun()->incrr;Id.of_string("FAKETOP"^(string_of_int!r))letdefs=tag_definition(strs)letkeywords=tag_keyword(strs)letget_new_idlocalsid=letrecget_idlid=letdir=DirPath.make[id]inifnot(Nametab.exists_dirdir)thenidelseget_id(Id.Set.addidl)(Namegen.next_ident_awayidl)inletavoid=List.fold_left(funaccu(_,id)->Id.Set.addidaccu)Id.Set.emptylocalsinget_idavoidid(** Inductive declarations *)openReductionletprint_paramsenvsigmaparams=ifList.is_emptyparamsthenmt()elsePrinter.pr_rel_contextenvsigmaparams++brk(1,2)letprint_constructorsenvparsigmanamestypes=letpc=prlist_with_sep(fun()->brk(1,0)++str"| ")(fun(id,c)->Id.printid++str" : "++Printer.pr_lconstr_envenvparsigmac)(Array.to_list(Array.map2(funnt->(n,t))namestypes))inhv0(str" "++pc)letbuild_ind_typemip=Inductive.type_of_inductivemipletprint_one_inductiveenvsigmamib((_,i)asind)=letu=Univ.make_abstract_instance(Declareops.inductive_polymorphic_contextmib)inletmip=mib.mind_packets.(i)inletparams=Inductive.inductive_paramdecls(mib,u)inletnparamdecls=Context.Rel.lengthparamsinletargs=Context.Rel.to_extended_listmkRel0paramsinletarity=hnf_prod_applist_assumenvnparamdecls(build_ind_type((mib,mip),u))argsinletcstrtypes=Inductive.type_of_constructors(ind,u)(mib,mip)inletcstrtypes=Array.map(func->hnf_prod_applist_assumenvnparamdeclscargs)cstrtypesinletenvpar=push_rel_contextparamsenvinletinst=ifDeclareops.inductive_is_polymorphicmibthenPrinter.pr_universe_instancesigmauelsemt()inhov0(Id.printmip.mind_typename++inst++brk(1,4)++print_paramsenvsigmaparams++str": "++Printer.pr_lconstr_envenvparsigmaarity++str" :=")++brk(0,2)++print_constructorsenvparsigmamip.mind_consnamescstrtypesletprint_mutual_inductiveenvmindmibudecl=letinds=List.init(Array.lengthmib.mind_packets)(funx->(mind,x))inletkeyword=letopenDeclarationsinmatchmib.mind_finitewith|Finite->"Inductive"|BiFinite->"Variant"|CoFinite->"CoInductive"inletbl=Printer.universe_binders_with_opt_names(Declareops.inductive_polymorphic_contextmib)udeclinletsigma=Evd.from_ctx(UState.of_bindersbl)inhov0(defkeyword++spc()++prlist_with_sep(fun()->fnl()++str" with ")(print_one_inductiveenvsigmamib)inds++str"."++Printer.pr_universessigma?variance:mib.mind_variancemib.mind_universes)letget_fields=letrecprodec_reclsubstc=matchkindcwith|Prod(na,t,c)->letid=matchna.binder_namewithNameid->id|Anonymous->Id.of_string"_"inprodec_rec((id,true,Vars.substlsubstt)::l)(mkVarid::subst)c|LetIn(na,b,_,c)->letid=matchna.binder_namewithNameid->id|Anonymous->Id.of_string"_"inprodec_rec((id,false,Vars.substlsubstb)::l)(mkVarid::subst)c|_->List.revlinprodec_rec[][]letprint_recordenvmindmibudecl=letu=Univ.make_abstract_instance(Declareops.inductive_polymorphic_contextmib)inletmip=mib.mind_packets.(0)inletparams=Inductive.inductive_paramdecls(mib,u)inletnparamdecls=Context.Rel.lengthparamsinletargs=Context.Rel.to_extended_listmkRel0paramsinletarity=hnf_prod_applist_assumenvnparamdecls(build_ind_type((mib,mip),u))argsinletcstrtypes=Inductive.type_of_constructors((mind,0),u)(mib,mip)inletcstrtype=hnf_prod_applist_assumenvnparamdeclscstrtypes.(0)argsinletfields=get_fieldscstrtypeinletenvpar=push_rel_contextparamsenvinletbl=Printer.universe_binders_with_opt_names(Declareops.inductive_polymorphic_contextmib)udeclinletsigma=Evd.from_ctx(UState.of_bindersbl)inletkeyword=letopenDeclarationsinmatchmib.mind_finitewith|BiFinite->"Record"|Finite->"Inductive"|CoFinite->"CoInductive"inhov0(hov0(defkeyword++spc()++Id.printmip.mind_typename++brk(1,4)++print_paramsenvsigmaparams++str": "++Printer.pr_lconstr_envenvparsigmaarity++brk(1,2)++str":= "++Id.printmip.mind_consnames.(0))++brk(1,2)++hv2(str"{ "++prlist_with_sep(fun()->str";"++brk(2,0))(fun(id,b,c)->Id.printid++str(ifbthen" : "else" := ")++Printer.pr_lconstr_envenvparsigmac)fields)++str" }."++Printer.pr_universessigma?variance:mib.mind_variancemib.mind_universes)letpr_mutual_inductive_bodyenvmindmibudecl=ifmib.mind_record!=NotRecord&¬!Flags.raw_printthenprint_recordenvmindmibudeclelseprint_mutual_inductiveenvmindmibudecl(** Modpaths *)letrecprint_local_modpathlocals=function|MPboundmbid->Id.print(Util.List.assoc_fMBId.equalmbidlocals)|MPdot(mp,l)->print_local_modpathlocalsmp++str"."++Label.printl|MPfile_->raiseNot_foundletprint_modpathlocalsmp=try(* must be with let because streams are lazy! *)letqid=Nametab.shortest_qualid_of_modulempinpr_qualidqidwith|Not_found->print_local_modpathlocalsmpletprint_knlocalskn=tryletqid=Nametab.shortest_qualid_of_modtypekninpr_qualidqidwithNot_found->tryprint_local_modpathlocalsknwithNot_found->print_modpathlocalsknletnametab_register_dirobj_mp=letid=mk_fake_top()inletobj_dir=DirPath.make[id]inNametab.(push_dir(Until1)obj_dir(GlobDirRef.DirModule{obj_dir;obj_mp;}))(** Nota: the [global_reference] we register in the nametab below
might differ from internal ones, since we cannot recreate here
the canonical part of constant and inductive names, but only
the user names. This works nonetheless since we search now
[Nametab.the_globrevtab] modulo user name. *)letnametab_register_bodympdir(l,body)=letpushidref=Nametab.push(Nametab.Until(1+List.length(DirPath.reprdir)))(make_pathdirid)refinmatchbodywith|SFBmodule_->()(* TODO *)|SFBmodtype_->()(* TODO *)|SFBconst_->push(Label.to_idl)(GlobRef.ConstRef(Constant.make2mpl))|SFBmindmib->letmind=MutInd.make2mplinArray.iteri(funimip->pushmip.mind_typename(GlobRef.IndRef(mind,i));Array.iteri(funjid->pushid(GlobRef.ConstructRef((mind,i),j+1)))mip.mind_consnames)mib.mind_packetsletimport_module=Declaremods.import_moduleLibobject.Unfilteredletprocess_module_binding=Declaremods.process_module_bindingletnametab_register_module_bodympstruc=(* If [mp] is a globally visible module, we simply import it *)tryimport_module~export:falsempwithNot_found->(* Otherwise we try to emulate an import by playing with nametab *)nametab_register_dirmp;List.iter(nametab_register_bodympDirPath.empty)strucletget_typ_expr_algmtb=matchmtb.mod_type_algwith|Some(NoFunctorme)->me|_->raiseNot_foundletnametab_register_modparammbidmtb=letid=MBId.to_idmbidinmatchmtb.mod_typewith|MoreFunctor_->id(* functorial param : nothing to register *)|NoFunctorstruc->(* We first try to use the algebraic type expression if any,
via a Declaremods function that converts back to module entries *)trylet()=process_module_bindingmbid(get_typ_expr_algmtb)inidwithewhenCErrors.noncriticale->(* Otherwise, we try to play with the nametab ourselves *)letmp=MPboundmbidinletcheckid=Nametab.exists_dir(DirPath.make[id])inletid=Namegen.next_ident_away_fromidcheckinletdir=DirPath.make[id]innametab_register_dirmp;List.iter(nametab_register_bodympdir)struc;idletprint_bodyis_implextentenvmp(l,body)=letname=Label.printlinhov2(matchbodywith|SFBmodule_->keyword"Module"++spc()++name|SFBmodtype_->keyword"Module Type"++spc()++name|SFBconstcb->letctx=Declareops.constant_polymorphic_contextcbin(matchcb.const_bodywith|Def_->def"Definition"++spc()|OpaqueDef_whenis_impl->def"Theorem"++spc()|_->def"Parameter"++spc())++name++(matchextentwith|OnlyNames->mt()|WithContents->letbl=Printer.universe_binders_with_opt_namesctxNoneinletsigma=Evd.from_ctx(UState.of_bindersbl)instr" :"++spc()++hov0(Printer.pr_ltype_envenvsigmacb.const_type)++(matchcb.const_bodywith|Deflwhenis_impl->spc()++hov2(str":= "++Printer.pr_lconstr_envenvsigmal)|_->mt())++str"."++Printer.pr_abstract_universe_ctxsigmactx)|SFBmindmib->matchextentwith|WithContents->pr_mutual_inductive_bodyenv(MutInd.make2mpl)mibNone|OnlyNames->letkeyword=letopenDeclarationsinmatchmib.mind_finitewith|Finite->def"Inductive"|BiFinite->def"Variant"|CoFinite->def"CoInductive"inkeyword++spc()++name)letprint_structis_implextentenvmpstruc=prlist_with_sepspc(print_bodyis_implextentenvmp)strucletprint_structureis_typeextentenvmplocalsstruc=letenv'=Modops.add_structurempstrucMod_subst.empty_delta_resolverenvinnametab_register_module_bodympstruc;letkwd=ifis_typethen"Sig"else"Struct"inhv2(keywordkwd++spc()++print_structfalseextentenv'mpstruc++brk(1,-2)++keyword"End")letrecflatten_appmexprl=matchmexprwith|MEapply(mexpr,arg)->flatten_appmexpr(arg::l)|MEidentmp->mp::l|MEwith_->assertfalseletrecprint_typ_exprextentenvmplocalsmty=matchmtywith|MEidentkn->print_knlocalskn|MEapply_->letlapp=flatten_appmty[]inletfapp=List.hdlappinletmapp=List.tllappinhov3(str"("++(print_knlocalsfapp)++spc()++prlist_with_sepspc(print_modpathlocals)mapp++str")")|MEwith(me,WithDef(idl,(c,_)))->lets=String.concat"."(List.mapId.to_stringidl)inletbody=matchextentwith|WithContents->letsigma=Evd.from_envenvinspc()++str":="++spc()++Printer.pr_lconstr_envenvsigmac|OnlyNames->mt()inhov2(print_typ_exprextentenvmplocalsme++spc()++str"with"++spc()++def"Definition"++spc()++strs++body)|MEwith(me,WithMod(idl,mp'))->lets=String.concat"."(List.mapId.to_stringidl)inletbody=matchextentwith|WithContents->spc()++str":="++spc()++print_modpathlocalsmp'|OnlyNames->mt()inhov2(print_typ_exprextentenvmplocalsme++spc()++str"with"++spc()++keyword"Module"++spc()++strs++body)letprint_mod_exprenvmplocals=function|MEidentmp->print_modpathlocalsmp|MEapply_asme->letlapp=flatten_appme[]inhov3(str"("++prlist_with_sepspc(print_modpathlocals)lapp++str")")|MEwith_->assertfalse(* No 'with' syntax for modules *)letrecprint_functorftyfatomis_typeextentenvmplocals=function|NoFunctorme->fatomis_typeextentenvmplocalsme|MoreFunctor(mbid,mtb1,me2)->letid=nametab_register_modparammbidmtb1inletmp1=MPboundmbidinletpr_mtb1=ftyextentenvmp1localsmtb1inletenv'=Modops.add_module_typemp1mtb1envinletlocals'=(mbid,get_new_idlocals(MBId.to_idmbid))::localsinletkwd=ifis_typethen"Funsig"else"Functor"inhov2(keywordkwd++spc()++str"("++Id.printid++str":"++pr_mtb1++str")"++spc()++print_functorftyfatomis_typeextentenv'mplocals'me2)letrecprint_expressionx=print_functorprint_modtype(functiontrue->print_typ_expr|false->fun_->print_mod_expr)xandprint_signaturex=print_functorprint_modtypeprint_structurexandprint_modtypeextentenvmplocalsmtb=matchmtb.mod_type_algwith|Someme->print_expressiontrueextentenvmplocalsme|None->print_signaturetrueextentenvmplocalsmtb.mod_typeletrecprintable_bodydir=letdir=pop_dirpathdirinDirPath.is_emptydir||tryletopenNametab.GlobDirRefinmatchNametab.locate_dir(qualid_of_dirpathdir)withDirOpenModtype_->false|DirModule_|DirOpenModule_->printable_bodydir|_->truewithNot_found->true(** Since we might play with nametab above, we should reset to prior
state after the printing *)letprint_expression'is_typeextentenvmpme=Vernacstate.System.protect(fune->print_expressionis_typeextentenvmp[]e)meletprint_signature'is_typeextentenvmpme=Vernacstate.System.protect(fune->print_signatureis_typeextentenvmp[]e)meletunsafe_print_moduleextentenvmpwith_bodymb=letname=print_modpath[]mpinletpr_equals=spc()++str":= "inletbody=matchwith_body,mb.mod_exprwith|false,_|true,Abstract->mt()|_,Algebraicme->pr_equals++print_expression'falseextentenvmpme|_,Structsign->pr_equals++print_signature'falseextentenvmpsign|_,FullStruct->pr_equals++print_signature'falseextentenvmpmb.mod_typeinletmodtype=matchmb.mod_expr,mb.mod_type_algwith|FullStruct,_->mt()|_,Somety->brk(1,1)++str": "++print_expression'trueextentenvmpty|_,_->brk(1,1)++str": "++print_signature'trueextentenvmpmb.mod_typeinhv0(keyword"Module"++spc()++name++modtype++body)exceptionShortPrintingletprint_modulewith_bodymp=letme=Global.lookup_modulempintryif!shortthenraiseShortPrinting;unsafe_print_moduleWithContents(Global.env())mpwith_bodyme++fnl()withewhenCErrors.noncriticale->unsafe_print_moduleOnlyNames(Global.env())mpwith_bodyme++fnl()letprint_modtypekn=letmtb=Global.lookup_modtypekninletname=print_kn[]kninhv1(keyword"Module Type"++spc()++name++str" ="++spc()++tryif!shortthenraiseShortPrinting;print_signature'trueWithContents(Global.env())knmtb.mod_typewithewhenCErrors.noncriticale->print_signature'trueOnlyNames(Global.env())knmtb.mod_type)