123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691692693694695696697698699700701702703704705706707708709710711712713714715716717718719720721722723724725726727728729730731732733734735736737738739740741742743744745746747748749750751752753754755756757758759760761762763764765766767768769770771772773774775776777778779780781782783784785786787788789790791792793794795796797798799800801802803804805806807808809810811812813814815816817818819820821822823824825826827828829830831832833834835836837838839840841842843844845(************************************************************************)(* * 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) *)(************************************************************************)openCErrorsopenUtilopenNamesopenUnivtypeuniverses_entry=|Monomorphic_entryofUniv.ContextSet.t|Polymorphic_entryofUniv.UContext.tmoduleUNameMap=Names.Id.Maptypeuinfo={uname:Id.toption;uloc:Loc.toption;}moduleUPairSet=UnivMinim.UPairSet(* 2nd part used to check consistency on the fly. *)typet={names:UnivNames.universe_binders*uinfoLevel.Map.t;(** Printing/location information *)local:ContextSet.t;(** The local graph of universes (variables and constraints) *)seff_univs:Level.Set.t;(** Local universes used through private constants *)univ_variables:UnivSubst.universe_opt_subst;(** The local universes that are unification variables *)univ_algebraic:Level.Set.t;(** The subset of unification variables that can be instantiated with
algebraic universes as they appear in inferred types only. *)universes:UGraph.t;(** The current graph extended with the local constraints *)universes_lbound:UGraph.Bound.t;(** The lower bound on universes (e.g. Set or Prop) *)initial_universes:UGraph.t;(** The graph at the creation of the evar_map *)minim_extra:UnivMinim.extra;}letinitial_sprop_cumulative=UGraph.set_cumulative_sproptrueUGraph.initial_universesletempty={names=UNameMap.empty,Level.Map.empty;local=ContextSet.empty;seff_univs=Level.Set.empty;univ_variables=Level.Map.empty;univ_algebraic=Level.Set.empty;universes=initial_sprop_cumulative;universes_lbound=UGraph.Bound.Set;initial_universes=initial_sprop_cumulative;minim_extra=UnivMinim.empty_extra;}letelaboration_sprop_cumul=Goptions.declare_bool_option_and_ref~depr:false~key:["Elaboration";"StrictProp";"Cumulativity"]~value:trueletmake~lboundunivs=letunivs=UGraph.set_cumulative_sprop(elaboration_sprop_cumul())univsin{emptywithuniverses=univs;universes_lbound=lbound;initial_universes=univs}letis_emptyuctx=ContextSet.is_emptyuctx.local&&Level.Map.is_emptyuctx.univ_variablesletuname_unionst=ifs==tthenselseUNameMap.merge(funklr->matchl,rwith|Some_,_->l|_,_->r)stletunionuctxuctx'=ifuctx==uctx'thenuctxelseifis_emptyuctx'thenuctxelseletlocal=ContextSet.unionuctx.localuctx'.localinletseff=Level.Set.unionuctx.seff_univsuctx'.seff_univsinletnames=uname_union(fstuctx.names)(fstuctx'.names)inletnames_rev=Level.Map.lunion(snductx.names)(snductx'.names)inletnewus=Level.Set.diff(ContextSet.levelsuctx'.local)(ContextSet.levelsuctx.local)inletnewus=Level.Set.diffnewus(Level.Map.domainuctx.univ_variables)inletextra=UnivMinim.extra_unionuctx.minim_extrauctx'.minim_extrainletdeclarenewg=Level.Set.fold(funug->UGraph.add_universeu~lbound:uctx.universes_lbound~strict:falseg)newusgin{names=(names,names_rev);local=local;seff_univs=seff;univ_variables=Level.Map.subst_unionuctx.univ_variablesuctx'.univ_variables;univ_algebraic=Level.Set.unionuctx.univ_algebraicuctx'.univ_algebraic;initial_universes=declarenewuctx.initial_universes;universes=(iflocal==uctx.localthenuctx.universeselseletcstrsr=ContextSet.constraintsuctx'.localinUGraph.merge_constraintscstrsr(declarenewuctx.universes));universes_lbound=uctx.universes_lbound;minim_extra=extra}letcontext_setuctx=uctx.localletconstraintsuctx=snductx.localletcompute_instance_bindersrbindersinst=letmaplvl=tryName(Option.get(Level.Map.findlvlrbinders).uname)withOption.IsNone|Not_found->AnonymousinArray.mapmap(Instance.to_arrayinst)letcontextuctx=let(_,rbinders)=uctx.namesinContextSet.to_context(compute_instance_bindersrbinders)uctx.localtypenamed_universes_entry=universes_entry*UnivNames.universe_bindersletuniv_entry~polyuctx=let(binders,_)=uctx.namesinletentry=ifpolythenPolymorphic_entry(contextuctx)elseMonomorphic_entry(context_setuctx)inentry,bindersletof_context_setlocal={emptywithlocal}typeuniverse_opt_subst=UnivSubst.universe_opt_substletsubstuctx=uctx.univ_variablesletnf_universesuctxc=UnivSubst.nf_evars_and_universes_opt_subst(fun_->None)(substuctx)cletugraphuctx=uctx.universesletinitial_graphuctx=uctx.initial_universesletalgebraicsuctx=uctx.univ_algebraicletadd_names?locsl(names,names_rev)=ifUNameMap.memsnamesthenuser_err?locPp.(str"Universe "++Names.Id.prints++str" already bound.");(UNameMap.addslnames,Level.Map.addl{uname=Somes;uloc=loc}names_rev)letadd_loclloc(names,names_rev)=matchlocwith|None->(names,names_rev)|Some_->(names,Level.Map.addl{uname=None;uloc=loc}names_rev)letof_bindersnames=letrev_map=UNameMap.fold(funidlrmap->Level.Map.addl{uname=Someid;uloc=None}rmap)namesLevel.Map.emptyin{emptywithnames=(names,rev_map)}letuniverse_of_nameuctxs=UNameMap.finds(fstuctx.names)letuniverse_bindersuctx=letnamed,_=uctx.namesinnamedletinstantiate_variablelbv=tryv:=Level.Map.setl(Someb)!vwithNot_found->assertfalseexceptionUniversesDifferletdrop_weak_constraints=Goptions.declare_bool_option_and_ref~depr:false~key:["Cumulativity";"Weak";"Constraints"]~value:falseletsort_inconsistencycstlr=raise(UGraph.UniverseInconsistency(cst,l,r,None))letlevel_inconsistencycstlr=letmku=Sorts.sort_of_univ@@Universe.makeuinraise(UGraph.UniverseInconsistency(cst,mkl,mkr,None))letsubst_univs_sortnormalizes=matchswith|Sorts.Set|Sorts.Prop|Sorts.SProp->s|Sorts.Typeu->Sorts.sort_of_univ(UnivSubst.subst_univs_universenormalizeu)typesmall_universe=USet|UProp|USPropletis_uset=functionUSet->true|UProp|USProp->falsetypesort_classification=|USmallofsmall_universe(* Set, Prop or SProp *)|ULevelofLevel.t(* Var or Global *)|UMaxofUniverse.t*Level.Set.t(* Max of Set, Var, Global without increments *)|UAlgebraicofUniverse.t(* Arbitrary algebraic expression *)letclassifys=matchswith|Sorts.Prop->USmallUProp|Sorts.SProp->USmallUSProp|Sorts.Set->USmallUSet|Sorts.Typeu->ifUniverse.is_levelsuthenmatchUniverse.leveluwith|None->UMax(u,Universe.levelsu)|Someu->ULeveluelseUAlgebraicutypelocal={local_cst:Constraints.t;local_above_prop:Level.Set.t;local_weak:UPairSet.t;}letadd_localcstlocal={localwithlocal_cst=Constraints.addcstlocal.local_cst}(* Constraint with algebraic on the left and a single level on the right *)letenforce_leq_upuvlocal={localwithlocal_cst=UnivSubst.enforce_lequ(Universe.makev)local.local_cst}letprocess_universe_constraintsuctxcstrs=letopenUnivSubstinletopenUnivProbleminletunivs=uctx.universesinletvars=refuctx.univ_variablesinletcumulative_sprop=UGraph.cumulative_spropunivsinletnormalizeu=normalize_univ_variable_opt_subst!varsuinletnf_constraint=function|ULub(u,v)->ULub(level_subst_ofnormalizeu,level_subst_ofnormalizev)|UWeak(u,v)->UWeak(level_subst_ofnormalizeu,level_subst_ofnormalizev)|UEq(u,v)->UEq(subst_univs_sortnormalizeu,subst_univs_sortnormalizev)|ULe(u,v)->ULe(subst_univs_sortnormalizeu,subst_univs_sortnormalizev)inletis_locall=Level.Map.meml!varsinletequalize_smalllslocal=letls=matchlwith|USProp->Sorts.sprop|UProp->Sorts.prop|USet->Sorts.setinifUGraph.check_eq_sortunivslssthenlocalelseifis_usetlthenmatchclassifyswith|USmall_->sort_inconsistencyEqSorts.sets|ULevelr->ifis_localrthenlet()=instantiate_variablerUniverse.type0varsinadd_local(Level.set,Eq,r)localelsesort_inconsistencyEqSorts.sets|UMax(u,_)|UAlgebraicu->ifuniv_level_memLevel.setuthenletinst=univ_level_remLevel.setuuinenforce_leq_upinstLevel.setlocalelsesort_inconsistencyEqlsselsesort_inconsistencyEqlssinletequalize_variablesfol'r'local=let()=ifis_locall'theninstantiate_variablel'(Universe.maker')varselseifis_localr'theninstantiate_variabler'(Universe.makel')varselseifnot(UnivProblem.check_eq_levelunivsl'r')then(* Two rigid/global levels, none of them being local,
one of them being Prop/Set, disallow *)ifLevel.is_setl'||Level.is_setr'thenlevel_inconsistencyEql'r'elseiffothenraiseUniversesDifferinifLevel.equall'r'thenlocalelseadd_local(l',Eq,r')localinletequalize_algebraiclrulocal=letalg=Level.Set.memluctx.univ_algebraicinletinst=univ_level_remlruruinifalg&¬(Level.Set.meml(Universe.levelsinst))thenlet()=instantiate_variablelinstvarsinlocalelseifuniv_level_memlruthenenforce_leq_upinstllocalelsesort_inconsistencyEq(Sorts.sort_of_univ(Universe.makel))(Sorts.sort_of_univru)inletequalize_universeslrlocal=matchclassifyl,classifyrwith|USmalll',(USmall_|ULevel_|UMax_|UAlgebraic_)->equalize_smalll'rlocal|(ULevel_|UMax_|UAlgebraic_),USmallr'->equalize_smallr'llocal|ULevell',ULevelr'->equalize_variablesfalsel'r'local|ULevell',(UAlgebraicr|UMax(r,_))|(UAlgebraicr|UMax(r,_)),ULevell'->equalize_algebraicl'rlocal|(UAlgebraic_|UMax_),(UAlgebraic_|UMax_)->(* both are algebraic *)ifUGraph.check_eq_sortunivslrthenlocalelsesort_inconsistencyEqlrinletunify_universescstlocal=letcst=nf_constraintcstinifUnivProblem.is_trivialcstthenlocalelsematchcstwith|ULe(l,r)->beginmatchclassifyrwith|UAlgebraic_|UMax_->ifUGraph.check_leq_sortunivslrthenlocalelseuser_errPp.(str"Algebraic universe on the right")|USmallr'->(* Invariant: there are no universes u <= Set in the graph. Except for
template levels, Set <= u anyways. Otherwise, for template
levels, any constraint u <= Set is turned into u := Set. *)ifUGraph.type_in_typeunivsthenlocalelsebeginmatchclassifylwith|UAlgebraic_->(* l contains a +1 and r=r' small so l <= r impossible *)sort_inconsistencyLelr|USmalll'->ifUGraph.check_leq_sortunivslrthenlocalelsesort_inconsistencyLelr|ULevell'->ifis_usetr'&&is_locall'then(* Unbounded universe constrained from above, we equalize it *)let()=instantiate_variablel'Universe.type0varsinadd_local(l',Eq,Level.set)localelsesort_inconsistencyLelr|UMax(_,levels)->ifis_usetr'thenletfoldl'local=letl=Sorts.sort_of_univ@@Universe.makel'inifLevel.is_setl'||is_locall'thenequalize_variablesfalsel'Level.setlocalelsesort_inconsistencyLelrinLevel.Set.foldfoldlevelslocalelsesort_inconsistencyLelrend|ULevelr'->(* We insert the constraint in the graph even if the graph
already contains it. Indeed, checking the existence of the
constraint is costly when the constraint does not already
exist directly as a single edge in the graph, but adding an
edge in the graph which is implied by others is cheap.
Hence, by doing this, we avoid a costly check here, and
make further checks of this constraint easier since it will
exist directly in the graph. *)matchclassifylwith|USmallUProp->{localwithlocal_above_prop=Level.Set.addr'local.local_above_prop}|USmallUSProp->ifUGraph.type_in_typeunivs||cumulative_spropthenlocalelsesort_inconsistencyLelr|USmallUSet->add_local(Level.set,Le,r')local|ULevell'->add_local(l',Le,r')local|UAlgebraicl->enforce_leq_uplr'local|UMax(_,l)->Univ.Level.Set.fold(funl'accu->add_local(l',Le,r')accu)llocalend|ULub(l,r)->equalize_variablestruelrlocal|UWeak(l,r)->ifnot(drop_weak_constraints())then{localwithlocal_weak=UPairSet.add(l,r)local.local_weak}elselocal|UEq(l,r)->equalize_universeslrlocalinletunify_universescstlocal=ifnot(UGraph.type_in_typeunivs)thenunify_universescstlocalelsetryunify_universescstlocalwithUGraph.UniverseInconsistency_->localinletlocal={local_cst=Constraints.empty;local_weak=uctx.minim_extra.UnivMinim.weak_constraints;local_above_prop=uctx.minim_extra.UnivMinim.above_prop;}inletlocal=UnivProblem.Set.foldunify_universescstrslocalinletextra={UnivMinim.above_prop=local.local_above_prop;UnivMinim.weak_constraints=local.local_weak}in!vars,extra,local.local_cstletadd_constraintsuctxcstrs=letunivs,old_cstrs=uctx.localinletcstrs'=Constraints.fold(fun(l,d,r)acc->letl=Universe.makelandr=Sorts.sort_of_univ@@Universe.makerinletcstr'=letopenUnivProbleminmatchdwith|Lt->ULe(Sorts.sort_of_univ@@Universe.superl,r)|Le->ULe(Sorts.sort_of_univl,r)|Eq->UEq(Sorts.sort_of_univl,r)inUnivProblem.Set.addcstr'acc)cstrsUnivProblem.Set.emptyinletvars,extra,cstrs'=process_universe_constraintsuctxcstrs'in{uctxwithlocal=(univs,Constraints.unionold_cstrscstrs');univ_variables=vars;universes=UGraph.merge_constraintscstrs'uctx.universes;minim_extra=extra;}letadd_universe_constraintsuctxcstrs=letunivs,local=uctx.localinletvars,extra,local'=process_universe_constraintsuctxcstrsin{uctxwithlocal=(univs,Constraints.unionlocallocal');univ_variables=vars;universes=UGraph.merge_constraintslocal'uctx.universes;minim_extra=extra;}letconstrain_variablesdiffuctx=letunivs,local=uctx.localinletunivs,vars,local=Level.Set.fold(funl(univs,vars,cstrs)->trymatchLevel.Map.findlvarswith|Someu->(Level.Set.addlunivs,Level.Map.removelvars,Constraints.add(l,Eq,Option.get(Universe.levelu))cstrs)|None->(univs,vars,cstrs)withNot_found|Option.IsNone->(univs,vars,cstrs))diff(univs,uctx.univ_variables,local)in{uctxwithlocal=(univs,local);univ_variables=vars}letid_of_leveluctxl=trySome(Option.get(Level.Map.findl(snductx.names)).uname)withNot_found|Option.IsNone->Noneletqualid_of_leveluctxl=letmap,map_rev=uctx.namesintrySome(Libnames.qualid_of_ident(Option.get(Level.Map.findlmap_rev).uname))withNot_found|Option.IsNone->UnivNames.qualid_of_levelmaplletpr_uctx_leveluctxl=matchqualid_of_leveluctxlwith|Someqid->Libnames.pr_qualidqid|None->Level.prltype('a,'b)gen_universe_decl={univdecl_instance:'a;(* Declared universes *)univdecl_extensible_instance:bool;(* Can new universes be added *)univdecl_constraints:'b;(* Declared constraints *)univdecl_extensible_constraints:bool(* Can new constraints be added *)}typeuniverse_decl=(lidentlist,Constraints.t)gen_universe_declletdefault_univ_decl={univdecl_instance=[];univdecl_extensible_instance=true;univdecl_constraints=Constraints.empty;univdecl_extensible_constraints=true}letpr_error_unbound_universesleftuctx=letopenPpinletn=Level.Set.cardinalleftinletprlevu=letinfo=Level.Map.find_optu(snductx.names)inh(pr_uctx_leveluctxu++(matchinfowith|None|Some{uloc=None}->mt()|Some{uloc=Someloc}->spc()++str"("++Loc.prloc++str")"))in(hv0(str(CString.pluraln"Universe")++spc()++(prlist_with_sepspcprlev(Level.Set.elementsleft))++spc()++str(CString.conjugate_verb_to_ben)++str" unbound."))exceptionUnboundUnivsofLevel.Set.t*t(* Deliberately using no location as the location of the univs
doesn't correspond to the failing command. *)leterror_unbound_universesleftuctx=raise(UnboundUnivs(left,uctx))let_=CErrors.register_handler(function|UnboundUnivs(left,uctx)->Some(pr_error_unbound_universesleftuctx)|_->None)letuniverse_context~names~extensibleuctx=letlevels=ContextSet.levelsuctx.localinletnewinst,left=List.fold_right(fun{CAst.loc;v=id}(newinst,acc)->letl=tryuniverse_of_nameuctxidwithNot_found->assertfalsein(l::newinst,Level.Set.removelacc))names([],levels)inifnotextensible&¬(Level.Set.is_emptyleft)thenerror_unbound_universesleftuctxelseletleft=ContextSet.sort_levels(Array.of_list(Level.Set.elementsleft))inletinst=Array.append(Array.of_listnewinst)leftinletinst=Instance.of_arrayinstin(inst,ContextSet.constraintsuctx.local)letcheck_universe_context_set~names~extensibleuctx=ifextensiblethen()elseletleft=List.fold_left(funleft{CAst.loc;v=id}->letl=tryuniverse_of_nameuctxidwithNot_found->assertfalseinLevel.Set.removelleft)(ContextSet.levelsuctx.local)namesinifnot(Level.Set.is_emptyleft)thenerror_unbound_universesleftuctxletcheck_implicationuctxcstrscstrs'=letgr=initial_graphuctxinletgrext=UGraph.merge_constraintscstrsgrinletcstrs'=Constraints.filter(func->not(UGraph.check_constraintgrextc))cstrs'inifConstraints.is_emptycstrs'then()elseCErrors.user_errPp.(str"Universe constraints are not implied by the ones declared: "++pr_constraints(pr_uctx_leveluctx)cstrs')letcheck_mono_univ_decluctxdecl=let()=letnames=decl.univdecl_instanceinletextensible=decl.univdecl_extensible_instanceincheck_universe_context_set~names~extensibleuctxinifnotdecl.univdecl_extensible_constraintsthencheck_implicationuctxdecl.univdecl_constraints(ContextSet.constraintsuctx.local);uctx.localletcheck_univ_decl~polyuctxdecl=ifnotdecl.univdecl_extensible_constraintsthencheck_implicationuctxdecl.univdecl_constraints(ContextSet.constraintsuctx.local);letnames=decl.univdecl_instanceinletextensible=decl.univdecl_extensible_instanceinlet(binders,rbinders)=uctx.namesinifpolythenletinst,csts=universe_context~names~extensibleuctxinletnas=compute_instance_bindersrbindersinstinletuctx=UContext.makenas(inst,csts)inPolymorphic_entryuctx,binderselselet()=check_universe_context_set~names~extensibleuctxinMonomorphic_entryuctx.local,bindersletis_boundllbound=matchlboundwith|UGraph.Bound.Prop->false|UGraph.Bound.Set->Level.is_setlletrestrict_universe_context~lbound(univs,csts)keep=letremoved=Level.Set.diffunivskeepinifLevel.Set.is_emptyremovedthenunivs,cstselseletallunivs=Constraints.fold(fun(u,_,v)all->Level.Set.addu(Level.Set.addvall))cstsunivsinletg=UGraph.initial_universesinletg=Level.Set.fold(funvg->ifLevel.is_setvthengelseUGraph.add_universev~lbound~strict:falseg)allunivsginletg=UGraph.merge_constraintscstsginletallkept=Level.Set.union(UGraph.domainUGraph.initial_universes)(Level.Set.diffallunivsremoved)inletcsts=UGraph.constraints_for~kept:allkeptginletcsts=Constraints.filter(fun(l,d,r)->not(is_boundllbound&&d==Le))cstsin(Level.Set.interunivskeep,csts)letrestrictuctxvars=letvars=Level.Set.unionvarsuctx.seff_univsinletvars=Names.Id.Map.fold(funnalvars->Level.Set.addlvars)(fstuctx.names)varsinletuctx'=restrict_universe_context~lbound:uctx.universes_lbounductx.localvarsin{uctxwithlocal=uctx'}typerigid=|UnivRigid|UnivFlexibleofbool(** Is substitution by an algebraic ok? *)letuniv_rigid=UnivRigidletuniv_flexible=UnivFlexiblefalseletuniv_flexible_alg=UnivFlexibletrue(** ~sideff indicates that it is ok to redeclare a universe.
~extend also merges the universe context in the local constraint structures
and not only in the graph. This depends if the
context we merge comes from a side effect that is already inlined
or defined separately. In the later case, there is no extension,
see [emit_side_effects] for example. *)letmerge?loc~sideffrigiductxuctx'=letlevels=ContextSet.levelsuctx'inletuctx=matchrigidwith|UnivRigid->uctx|UnivFlexibleb->letfolduaccu=ifLevel.Map.memuaccuthenaccuelseLevel.Map.adduNoneaccuinletuvars'=Level.Set.foldfoldlevelsuctx.univ_variablesinifbthen{uctxwithuniv_variables=uvars';univ_algebraic=Level.Set.unionuctx.univ_algebraiclevels}else{uctxwithuniv_variables=uvars'}inletlocal=ContextSet.appenductx'uctx.localinletdeclareg=Level.Set.fold(funug->tryUGraph.add_universe~lbound:uctx.universes_lbound~strict:falseugwithUGraph.AlreadyDeclaredwhensideff->g)levelsginletnames=letfolduaccu=letmodify_info=matchinfo.ulocwith|None->{infowithuloc=loc}|Some_->infointryLevel.Map.modifyumodifyaccuwithNot_found->Level.Map.addu{uname=None;uloc=loc}accuin(fstuctx.names,Level.Set.foldfoldlevels(snductx.names))inletinitial=declareuctx.initial_universesinletunivs=declareuctx.universesinletuniverses=UGraph.merge_constraints(ContextSet.constraintsuctx')univsin{uctxwithnames;local;universes;initial_universes=initial}letmerge_substuctxs={uctxwithuniv_variables=Level.Map.subst_unionuctx.univ_variabless}letdemote_seff_univsunivsuctx=letseff=Level.Set.unionuctx.seff_univsunivsin{uctxwithseff_univs=seff}letdemote_global_univsenvuctx=letenv_ugraph=Environ.universesenvinletglobal_univs=UGraph.domainenv_ugraphinletglobal_constraints,_=UGraph.constraints_of_universesenv_ugraphinletpromoted_uctx=ContextSet.(of_setglobal_univs|>add_constraintsglobal_constraints)in{uctxwithlocal=ContextSet.diffuctx.localpromoted_uctx}letmerge_seffuctxuctx'=letlevels=ContextSet.levelsuctx'inletdeclareg=Level.Set.fold(funug->tryUGraph.add_universe~lbound:uctx.universes_lbound~strict:falseugwithUGraph.AlreadyDeclared->g)levelsginletinitial_universes=declareuctx.initial_universesinletunivs=declareuctx.universesinletuniverses=UGraph.merge_constraints(ContextSet.constraintsuctx')univsin{uctxwithuniverses;initial_universes}letemit_side_effectseffu=letuctx=Safe_typing.universes_of_privateeffinletu=demote_seff_univs(fstuctx)uinmerge_seffuuctxletupdate_sigma_univsuctxugraph=letunivs=UGraph.set_cumulative_sprop(elaboration_sprop_cumul())ugraphinleteunivs={uctxwithinitial_universes=univs;universes=univs}inmerge_seffeunivseunivs.localletadd_universe?locnamestrictlbounductxu=letinitial_universes=UGraph.add_universe~lbound~strictuuctx.initial_universesinletuniverses=UGraph.add_universe~lbound~strictuuctx.universesinletlocal=ContextSet.add_universeuuctx.localinletnames=matchnamewith|Somen->add_names?locnuuctx.names|None->add_loculocuctx.namesin{uctxwithnames;local;initial_universes;universes}letnew_univ_variable?locrigidnameuctx=letu=UnivGen.fresh_level()inletuctx=matchrigidwith|UnivRigid->uctx|UnivFlexibleallow_alg->letuniv_variables=Level.Map.adduNoneuctx.univ_variablesinifallow_algthenletuniv_algebraic=Level.Set.adduuctx.univ_algebraicin{uctxwithuniv_variables;univ_algebraic}else{uctxwithuniv_variables}inletuctx=add_universe?locnamefalseuctx.universes_lbounductxuinuctx,uletadd_global_univuctxu=add_universeNonetrueUGraph.Bound.Setuctxuletmake_with_initial_binders~lboundunivsus=letuctx=make~lboundunivsinList.fold_left(funuctx{CAst.loc;v=id}->fst(new_univ_variable?locuniv_rigid(Someid)uctx))uctxusletfrom_env?(binders=[])env=make_with_initial_binders~lbound:(Environ.universes_lboundenv)(Environ.universesenv)bindersletmake_flexible_variableuctx~algebraicu=let{local=cstrs;univ_variables=uvars;univ_algebraic=avars;universes=g;}=uctxinassert(tryLevel.Map.finduuvars==NonewithNot_found->true);matchUGraph.choose(funv->not(Level.equaluv)&&(algebraic||not(Level.Set.memvavars)))guwith|Somev->letuvars'=Level.Map.addu(Some(Universe.makev))uvarsin{uctxwithuniv_variables=uvars';}|None->letuvars'=Level.Map.adduNoneuvarsinletavars'=ifalgebraicthenletuu=Universe.makeuinletsubstu_not_algu'v=Option.cata(funvu->Universe.equaluuvu&¬(Level.Set.memu'avars))falsevinlethas_upper_constraint()=Constraints.exists(fun(l,d,r)->d==Lt&&Level.equallu)(ContextSet.constraintscstrs)inifnot(Level.Map.existssubstu_not_alguvars||has_upper_constraint())thenLevel.Set.adduavarselseavarselseavarsin{uctxwithuniv_variables=uvars';univ_algebraic=avars'}letmake_nonalgebraic_variableuctxu={uctxwithuniv_algebraic=Level.Set.removeuuctx.univ_algebraic}letmake_flexible_nonalgebraicuctx={uctxwithuniv_algebraic=Level.Set.empty}letis_sort_variableuctxs=matchswith|Sorts.Typeu->(matchUniverse.leveluwith|Somelasx->ifLevel.Set.meml(ContextSet.levelsuctx.local)thenxelseNone|None->None)|_->Noneletsubst_univs_context_with_defdefusubst(uctx,cst)=(Level.Set.diffuctxdef,UnivSubst.subst_univs_constraintsusubstcst)letrefresh_constraintsunivs(ctx,cstrs)=letcstrs',univs'=Constraints.fold(func(cstrs',univs)->(Constraints.addccstrs',UGraph.enforce_constraintcunivs))cstrs(Constraints.empty,univs)in((ctx,cstrs'),univs')letnormalize_variablesuctx=letnormalized_variables,def,subst=UnivSubst.normalize_univ_variablesuctx.univ_variablesinletmake_substsubstl=Level.Map.findlsubstinletuctx_local=subst_univs_context_with_defdef(make_substsubst)uctx.localinletuctx_local',univs=refresh_constraintsuctx.initial_universesuctx_localin{uctxwithlocal=uctx_local';univ_variables=normalized_variables;universes=univs}letabstract_undefined_variablesuctx=letvars'=Level.Map.fold(funuvacc->ifv==NonethenLevel.Set.removeuaccelseacc)uctx.univ_variablesuctx.univ_algebraicin{uctxwithlocal=ContextSet.empty;univ_algebraic=vars'}letfix_undefined_variablesuctx=letalgs',vars'=Level.Map.fold(funuv(algs,varsasacc)->ifv==Nonethen(Level.Set.removeualgs,Level.Map.removeuvars)elseacc)uctx.univ_variables(uctx.univ_algebraic,uctx.univ_variables)in{uctxwithuniv_variables=vars';univ_algebraic=algs'}letminimizeuctx=letopenUnivMiniminletlbound=uctx.universes_lboundinlet((vars',algs'),us')=normalize_context_set~lbounductx.universesuctx.localuctx.univ_variablesuctx.univ_algebraicuctx.minim_extrainifContextSet.equalus'uctx.localthenuctxelseletus',universes=refresh_constraintsuctx.initial_universesus'in{names=uctx.names;local=us';seff_univs=uctx.seff_univs;(* not sure about this *)univ_variables=vars';univ_algebraic=algs';universes=universes;universes_lbound=lbound;initial_universes=uctx.initial_universes;minim_extra=UnivMinim.empty_extra;(* weak constraints are consumed *)}(* XXX print above_prop too *)letpr_weakprl{minim_extra={UnivMinim.weak_constraints=weak}}=letopenPpinprlist_with_sepfnl(fun(u,v)->prlu++str" ~ "++prlv)(UPairSet.elementsweak)letpr_universe_body=function|None->Pp.mt()|Somex->Pp.(str" := "++Univ.Universe.prx)letpr_universe_opt_subst=Univ.Level.Map.prpr_universe_body