123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408(************************************************************************)(* * 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) *)(************************************************************************)openUnivopenUnivSubst(* To disallow minimization to Set *)letget_set_minimization=Goptions.declare_bool_option_and_ref~depr:false~key:["Universe";"Minimization";"ToSet"]~value:true(** Simplification *)letadd_list_maputmap=tryletl=LMap.findumapinLMap.setu(t::l)mapwithNot_found->LMap.addu[t]map(** Precondition: flexible <= ctx *)letchoose_canonicalctxflexiblealgss=letglobal=LSet.diffsctxinletflexible,rigid=LSet.partitionflexible(LSet.intersctx)in(* If there is a global universe in the set, choose it *)ifnot(LSet.is_emptyglobal)thenletcanon=LSet.chooseglobalincanon,(LSet.removecanonglobal,rigid,flexible)else(* No global in the equivalence class, choose a rigid one *)ifnot(LSet.is_emptyrigid)thenletcanon=LSet.chooserigidincanon,(global,LSet.removecanonrigid,flexible)else(* There are only flexible universes in the equivalence
class, choose a non-algebraic. *)letalgs,nonalgs=LSet.partition(funx->LSet.memxalgs)flexibleinifnot(LSet.is_emptynonalgs)thenletcanon=LSet.choosenonalgsincanon,(global,rigid,LSet.removecanonflexible)elseletcanon=LSet.choosealgsincanon,(global,rigid,LSet.removecanonflexible)(* Eq < Le < Lt *)letcompare_constraint_typedd'=matchd,d'with|Eq,Eq->0|Eq,_->-1|_,Eq->1|Le,Le->0|Le,_->-1|_,Le->1|Lt,Lt->0typelowermap=constraint_typeLMap.tletlower_union=letmergekab=matcha,bwith|Some_,None->a|None,Some_->b|None,None->None|Somel,Somer->ifcompare_constraint_typelr>=0thenaelsebinLMap.mergemergeletlower_addlcm=tryletc'=LMap.findlminifcompare_constraint_typecc'>0thenLMap.addlcmelsemwithNot_found->LMap.addlcmletlower_of_listl=List.fold_left(funacc(d,l)->LMap.addldacc)LMap.emptyltypelbound={enforce:bool;alg:bool;lbound:Universe.t;lower:lowermap}moduleLBMap:sigtypet=private{lbmap:lboundLMap.t;lbrev:(Level.t*lowermap)Universe.Map.t}valempty:tvaladd:Level.t->lbound->t->tend=structtypet={lbmap:lboundLMap.t;lbrev:(Level.t*lowermap)Universe.Map.t}(* lbrev is uniquely given from lbmap as a partial reverse mapping *)letempty={lbmap=LMap.empty;lbrev=Universe.Map.empty}letaddubndm=letlbmap=LMap.addubndm.lbmapinletlbrev=ifnotbnd.alg&&bnd.enforcethenmatchUniverse.Map.findbnd.lboundm.lbrevwith|(v,_)->ifLevel.compareuv<=0thenUniverse.Map.addbnd.lbound(u,bnd.lower)m.lbrevelsem.lbrev|exceptionNot_found->Universe.Map.addbnd.lbound(u,bnd.lower)m.lbrevelsem.lbrevin{lbmap;lbrev}endletfind_instinstsv=Universe.Map.findvinsts.LBMap.lbrevletcompute_lboundleft=(* The universe variable was not fixed yet.
Compute its level using its lower bound. *)letsupllbound=matchlboundwith|None->Somel|Somel'->Some(Universe.supll')inList.fold_left(funlbound(d,l)->ifd==Le(* l <= ?u *)thensupllboundelse(* l < ?u *)(assert(d==Lt);ifnot(Universe.levell==None)thensup(Universe.superl)lboundelseNone))Noneleftletinstantiate_with_lboundulboundlower~alg~enforce(ctx,us,algs,insts,cstrs)=ifenforcethenletinst=Universe.makeuinletcstrs'=enforce_leqlboundinstcstrsin(ctx,us,LSet.removeualgs,LBMap.addu{enforce;alg;lbound;lower}insts,cstrs'),{enforce;alg;lbound=inst;lower}else(* Actually instantiate *)(Univ.LSet.removeuctx,Univ.LMap.addu(Somelbound)us,algs,LBMap.addu{enforce;alg;lbound;lower}insts,cstrs),{enforce;alg;lbound;lower}typeconstraints_map=(Univ.constraint_type*Univ.LMap.key)listUniv.LMap.tlet_pr_constraints_map(cmap:constraints_map)=letopenPpinLMap.fold(funlcstrsacc->Level.prl++str" => "++prlist_with_sepspc(fun(d,r)->pr_constraint_typed++Level.prr)cstrs++fnl()++acc)cmap(mt())letremove_algl(ctx,us,algs,insts,cstrs)=(ctx,us,LSet.removelalgs,insts,cstrs)letnot_lowerlower(d,l)=(* We're checking if (d,l) is already implied by the lower
constraints on some level u. If it represents l < u (d is Lt
or d is Le and i > 0, the i < 0 case is impossible due to
invariants of Univ), and the lower constraints only have l <=
u then it is not implied. *)Univ.Universe.exists(fun(l,i)->letd=ifi==0thendelsematchdwith|Le->Lt|d->dintryletd'=LMap.findllowerin(* If d is stronger than the already implied lower
* constraints we must keep it. *)compare_constraint_typedd'>0withNot_found->(* No constraint existing on l *)true)lexceptionUpperBoundedAlg(** [enforce_uppers upper lbound cstrs] interprets [upper] as upper
constraints to [lbound], adding them to [cstrs].
@raise UpperBoundedAlg if any [upper] constraints are strict and
[lbound] algebraic. *)letenforce_uppersupperlboundcstrs=List.fold_left(funcstrs(d,r)->ifd==Univ.Lethenenforce_leqlbound(Universe.maker)cstrselsematchUniverse.levellboundwith|Somelev->Constraint.add(lev,d,r)cstrs|None->raiseUpperBoundedAlg)cstrsupperletminimize_univ_variablesctxusalgsleftrightcstrs=letleft,lbounds=Univ.LMap.fold(funrlower(left,lboundsasacc)->ifUniv.LMap.memrus||not(Univ.LSet.memrctx)thenaccelse(* Fixed universe, just compute its glb for sharing *)letlbounds=matchcompute_lbound(List.map(fun(d,l)->d,Universe.makel)lower)with|None->lbounds|Somelbound->LBMap.addr{enforce=true;alg=false;lbound;lower=lower_of_listlower}lboundsin(Univ.LMap.removerleft,lbounds))left(left,LBMap.empty)inletrecinstance(ctx,us,algs,insts,cstrsasacc)u=letacc,left,lower=matchLMap.finduleftwith|exceptionNot_found->acc,[],LMap.empty|l->letacc,left,newlow,lower=List.fold_left(fun(acc,left,newlow,lower')(d,l)->letacc',{enforce=enf;alg;lbound=l';lower}=auxacclinletl'=ifenfthenUniverse.makelelsel'inacc',(d,l')::left,lower_addldnewlow,lower_unionlowerlower')(acc,[],LMap.empty,LMap.empty)linletleft=CList.uniquize(List.filter(not_lowerlower)left)in(acc,left,LMap.lunionnewlowlower)inletinstantiate_lboundlbound=letalg=LSet.memualgsinifalgthen(* u is algebraic: we instantiate it with its lower bound, if any,
or enforce the constraints if it is bounded from the top. *)letlower=LSet.foldLMap.remove(Universe.levelslbound)lowerininstantiate_with_lboundulboundlower~alg:true~enforce:falseaccelse(* u is non algebraic *)matchUniverse.levellboundwith|Somel->(* The lowerbound is directly a level *)(* u is not algebraic but has no upper bounds,
we instantiate it with its lower bound if it is a
different level, otherwise we keep it. *)letlower=LMap.removellowerinifnot(Level.equallu)then(* Should check that u does not
have upper constraints that are not already in right *)letacc=remove_alglaccininstantiate_with_lboundulboundlower~alg:false~enforce:falseaccelseacc,{enforce=true;alg=false;lbound;lower}|None->beginmatchfind_instinstslboundwith|can,lower->(* Another universe represents the same lower bound,
we can share them with no harm. *)letlower=LMap.removecanlowerininstantiate_with_lboundu(Universe.makecan)lower~alg:false~enforce:falseacc|exceptionNot_found->(* We set u as the canonical universe representing lbound *)instantiate_with_lboundulboundlower~alg:false~enforce:trueaccendinletenforce_uppers((ctx,us,algs,insts,cstrs),basacc)=matchLMap.findurightwith|exceptionNot_found->acc|upper->letupper=List.filter(fun(d,r)->not(LMap.memrus))upperinletcstrs=enforce_uppersupperb.lboundcstrsin(ctx,us,algs,insts,cstrs),binifnot(LSet.memuctx)thenenforce_uppers(acc,{enforce=true;alg=false;lbound=Universe.makeu;lower})elseletlbound=compute_lboundleftinmatchlboundwith|None->(* Nothing to do *)enforce_uppers(acc,{enforce=true;alg=false;lbound=Universe.makeu;lower})|Somelbound->tryenforce_uppers(instantiate_lboundlbound)withUpperBoundedAlg->enforce_uppers(acc,{enforce=true;alg=false;lbound=Universe.makeu;lower})andaux(ctx,us,algs,seen,cstrsasacc)u=tryacc,LMap.finduseen.LBMap.lbmapwithNot_found->instanceaccuinLMap.fold(funuv(ctx,us,algs,seen,cstrsasacc)->ifv==Nonethenfst(auxaccu)elseLSet.removeuctx,us,LSet.removeualgs,seen,cstrs)us(ctx,us,algs,lbounds,cstrs)moduleUPairs=OrderedType.UnorderedPair(Univ.Level)moduleUPairSet=Set.Make(UPairs)letis_boundllbound=matchlboundwith|UGraph.Bound.Prop->Level.is_propl|UGraph.Bound.Set->Level.is_setl(* if [is_minimal u] then constraints [u <= v] may be dropped and get
used only for set_minimization. *)letis_minimal~lboundu=Level.is_spropu||Level.is_propu||is_boundulbound(* TODO check is_small/sprop *)letnormalize_context_set~lboundgctxusalgsweak=let(ctx,csts)=ContextSet.levelsctx,ContextSet.constraintsctxin(* Keep the Prop/Set <= i constraints separate for minimization *)letsmallles,csts=Constraint.partition(fun(l,d,r)->d==Le&&is_minimal~lboundl)cstsinletsmallles=ifget_set_minimization()thenConstraint.filter(fun(l,d,r)->LMap.memrus&¬(Level.is_spropl))smallleselseConstraint.emptyinletsmallles=Constraint.map(fun(_,_,r)->Level.set,Le,r)smalllesinletcsts,partition=(* We first put constraints in a normal-form: all self-loops are collapsed
to equalities. *)letg=UGraph.initial_universes_withginletg=LSet.fold(funvg->UGraph.add_universe~lbound~strict:falsevg)ctxginletadd_softug=ifnot(Level.is_smallu||LSet.memuctx)thentryUGraph.add_universe~lbound~strict:falseugwithUGraph.AlreadyDeclared->gelseginletg=Constraint.fold(fun(l,d,r)g->add_softr(add_softlg))cstsginletg=UGraph.merge_constraintscstsginUGraph.constraints_of_universesgin(* We ignore the trivial Prop/Set <= i constraints. *)letnoneqs=Constraint.filter(fun(l,d,r)->not((d==Le&&is_boundllbound)||(Level.is_propl&&d==Lt&&Level.is_setr)))cstsinletnoneqs=Constraint.unionnoneqssmalllesinletflexx=LMap.memxusinletctx,us,eqs=List.fold_left(fun(ctx,us,cstrs)s->letcanon,(global,rigid,flexible)=choose_canonicalctxflexalgssin(* Add equalities for globals which can't be merged anymore. *)letcstrs=LSet.fold(fungcst->Constraint.add(canon,Eq,g)cst)globalcstrsin(* Also add equalities for rigid variables *)letcstrs=LSet.fold(fungcst->Constraint.add(canon,Eq,g)cst)rigidcstrsinletcanonu=Some(Universe.makecanon)inletus=LSet.fold(funf->LMap.addfcanonu)flexibleusin(LSet.diffctxflexible,us,cstrs))(ctx,us,Constraint.empty)partitionin(* Process weak constraints: when one side is flexible and the 2
universes are unrelated unify them. *)letctx,us,g=UPairSet.fold(fun(u,v)(ctx,us,gasacc)->letnorm=level_subst_of(normalize_univ_variable_opt_substus)inletu=normuandv=normvinletset_toab=(LSet.removeactx,LMap.adda(Some(Universe.makeb))us,UGraph.enforce_constraint(a,Eq,b)g)inifUGraph.check_constraintg(u,Le,v)||UGraph.check_constraintg(v,Le,u)thenaccelseifLMap.memuusthenset_touvelseifLMap.memvusthenset_tovuelseacc)weak(ctx,us,g)in(* Noneqs is now in canonical form w.r.t. equality constraints,
and contains only inequality constraints. *)letnoneqs=letnorm=level_subst_of(normalize_univ_variable_opt_substus)inConstraint.fold(fun(u,d,v)noneqs->letu=normuandv=normvinifd!=Lt&&Level.equaluvthennoneqselseConstraint.add(u,d,v)noneqs)noneqsConstraint.emptyin(* Compute the left and right set of flexible variables, constraints
mentioning other variables remain in noneqs. *)letnoneqs,ucstrsl,ucstrsr=Constraint.fold(fun(l,d,rascstr)(noneq,ucstrsl,ucstrsr)->letlus=LMap.memlusandrus=LMap.memrusinletucstrsl'=iflusthenadd_list_mapl(d,r)ucstrslelseucstrslanducstrsr'=add_list_mapr(d,l)ucstrsrinletnoneqs=iflus||rusthennoneqelseConstraint.addcstrnoneqin(noneqs,ucstrsl',ucstrsr'))noneqs(Constraint.empty,LMap.empty,LMap.empty)in(* Now we construct the instantiation of each variable. *)letctx',us,algs,inst,noneqs=minimize_univ_variablesctxusalgsucstrsrucstrslnoneqsinletus=normalize_opt_substusin(us,algs),(ctx',Constraint.unionnoneqseqs)