123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691692693694695696697698699700701702703704705706707708709710711712713714715716717718719720721722723724725726727728729730731732733734735736737738739740741742743744745746747748749750751752753754755756757758759760761762763764765766767768769770771772773774775776777778779780781782783784785786787788789790791792793794795796797798799800801802803804805806807808809810811812813814815816817818819820821822823824825826827828829830831832833834835836837838839840841842843844845846847848849850851852853854855856857858859860861862863864865866867868869870871872873874875876877878879880881882883884885886887888889890891892893894895896897898899900901902903904905906907908909910911912913914915916917918919920921922923924925926927928929930931932933934935936937938939940941942943944945946947948949950951952953954955956957958959960961962963964965966967968969970971972973974975976977978979980981982983984985986987988989990991992993994995996997998999100010011002100310041005100610071008100910101011101210131014101510161017101810191020102110221023102410251026102710281029103010311032103310341035103610371038103910401041104210431044104510461047104810491050105110521053105410551056105710581059106010611062106310641065106610671068106910701071107210731074107510761077107810791080108110821083108410851086108710881089109010911092109310941095109610971098109911001101110211031104110511061107110811091110111111121113111411151116111711181119112011211122112311241125112611271128112911301131113211331134113511361137113811391140114111421143114411451146114711481149115011511152115311541155115611571158115911601161116211631164116511661167116811691170117111721173117411751176117711781179118011811182118311841185118611871188118911901191119211931194119511961197119811991200120112021203120412051206120712081209121012111212121312141215121612171218121912201221122212231224122512261227122812291230123112321233123412351236123712381239124012411242124312441245124612471248124912501251(************************************************************************)(* * 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) *)(************************************************************************)(* Created in Caml by Gérard Huet for CoC 4.8 [Dec 1988] *)(* Functional code by Jean-Christophe Filliâtre for Coq V7.0 [1999] *)(* Extension with algebraic universes by HH for Coq V7.0 [Sep 2001] *)(* Additional support for sort-polymorphic inductive types by HH [Mar 2006] *)(* Support for universe polymorphism by MS [2014] *)(* Revisions by Bruno Barras, Hugo Herbelin, Pierre Letouzey, Matthieu
Sozeau, Pierre-Marie Pédrot *)openPpopenCErrorsopenUtil(* Universes are stratified by a partial ordering $\le$.
Let $\~{}$ be the associated equivalence. We also have a strict ordering
$<$ between equivalence classes, and we maintain that $<$ is acyclic,
and contained in $\le$ in the sense that $[U]<[V]$ implies $U\le V$.
At every moment, we have a finite number of universes, and we
maintain the ordering in the presence of assertions $U<V$ and $U\le V$.
The equivalence $\~{}$ is represented by a tree structure, as in the
union-find algorithm. The assertions $<$ and $\le$ are represented by
adjacency lists *)moduleRawLevel=structopenNamesmoduleUGlobal=structtypet={library:DirPath.t;process:string;uid:int;}letmakelibraryprocessuid={library;process;uid}letreprx=(x.library,x.process,x.uid)letequalu1u2=Int.equalu1.uidu2.uid&&DirPath.equalu1.libraryu2.library&&String.equalu1.processu2.processlethashu=Hashset.Combine.combine3u.uid(String.hashu.process)(DirPath.hashu.library)letcompareu1u2=letc=Int.compareu1.uidu2.uidinifc<>0thencelseletc=DirPath.compareu1.libraryu2.libraryinifc<>0thencelseString.compareu1.processu2.processendtypet=|SProp|Prop|Set|LevelofUGlobal.t|Varofint(* Hash-consing *)letequalxy=x==y||matchx,ywith|SProp,SProp->true|Prop,Prop->true|Set,Set->true|Levell,Levell'->UGlobal.equalll'|Varn,Varn'->Int.equalnn'|_->falseletcompareuv=matchu,vwith|SProp,SProp->0|SProp,_->-1|_,SProp->1|Prop,Prop->0|Prop,_->-1|_,Prop->1|Set,Set->0|Set,_->-1|_,Set->1|Levell1,Levell2->UGlobal.comparel1l2|Level_,_->-1|_,Level_->1|Varn,Varm->Int.comparenmlethequalxy=x==y||matchx,ywith|SProp,SProp->true|Prop,Prop->true|Set,Set->true|UGlobal.(Level{library=d;process=s;uid=n},Level{library=d';process=s';uid=n'})->n==n'&&s==s'&&d==d'|Varn,Varn'->n==n'|_->falselethcons=function|SPropasx->x|Propasx->x|Setasx->x|UGlobal.(Level{library=d;process=s;uid=n})asx->lets'=CString.hconssinletd'=Names.DirPath.hconsdinifs'==s&&d'==dthenxelseLevel(UGlobal.maked's'n)|Var_nasx->xopenHashset.Combinelethash=function|SProp->combinesmall10|Prop->combinesmall11|Set->combinesmall12|Varn->combinesmall2n|Levell->combinesmall3(UGlobal.hashl)endmoduleLevel=structmoduleUGlobal=RawLevel.UGlobaltyperaw_level=RawLevel.t=|SProp|Prop|Set|LevelofUGlobal.t|Varofint(** Embed levels with their hash value *)typet={hash:int;data:RawLevel.t}letequalxy=x==y||Int.equalx.hashy.hash&&RawLevel.equalx.datay.datalethashx=x.hashletdatax=x.data(** Hashcons on levels + their hash *)moduleSelf=structtypenonrect=ttypeu=unitleteqxy=x.hash==y.hash&&RawLevel.hequalx.datay.datalethashx=x.hashlethashcons()x=letdata'=RawLevel.hconsx.datainifx.data==data'thenxelse{xwithdata=data'}endlethcons=letmoduleH=Hashcons.Make(Self)inHashcons.simple_hconsH.generateH.hcons()letmakel=hcons{hash=RawLevel.hashl;data=l}letset=makeSetletprop=makePropletsprop=makeSPropletis_smallx=matchdataxwith|Level_->false|Var_->false|SProp->true|Prop->true|Set->trueletis_propx=matchdataxwith|Prop->true|_->falseletis_setx=matchdataxwith|Set->true|_->falseletis_spropx=matchdataxwith|SProp->true|_->falseletcompareuv=ifu==vthen0elseRawLevel.compare(datau)(datav)letto_stringx=matchdataxwith|SProp->"SProp"|Prop->"Prop"|Set->"Set"|UGlobal.(Level{library=d;process=s;uid=n})->Names.DirPath.to_stringd^(ifCString.is_emptysthen""else"."^s)^"."^string_of_intn|Varn->"Var("^string_of_intn^")"letpru=str(to_stringu)letvars=Array.init20(funi->make(Vari))letvarn=ifn<20thenvars.(n)elsemake(Varn)letvar_indexu=matchdatauwith|Varn->Somen|_->Noneletmakeqid=make(Levelqid)letnameu=matchdatauwith|Levell->Somel|_->Noneend(** Level maps *)moduleLMap=structmoduleM=HMap.Make(Level)includeMletlunionlr=union(fun_kl_r->Somel)lrletsubst_unionlr=union(fun_klr->matchl,rwith|Some_,_->Somel|None,None->Somel|_,_->Somer)lrletdiffextorig=fold(funuvacc->ifmemuorigthenaccelseadduvacc)extemptyletprfm=h(prlist_with_sepfnl(fun(u,v)->Level.pru++fv)(bindingsm))endmoduleLSet=structincludeLMap.Setletprprls=str"{"++prlist_with_sepspcprl(elementss)++str"}"letof_arrayl=Array.fold_left(funaccx->addxacc)emptylendtype'auniverse_map='aLMap.ttypeuniverse_level=Level.ttypeuniverse_level_subst_fn=universe_level->universe_leveltypeuniverse_set=LSet.t(* An algebraic universe [universe] is either a universe variable
[Level.t] or a formal universe known to be greater than some
universe variables and strictly greater than some (other) universe
variables
Universes variables denote universes initially present in the term
to type-check and non variable algebraic universes denote the
universes inferred while type-checking: it is either the successor
of a universe present in the initial term to type-check or the
maximum of two algebraic universes
*)moduleUniverse=struct(* Invariants: non empty, sorted and without duplicates *)moduleExpr=structtypet=Level.t*int(* Hashing of expressions *)moduleExprHash=structtypet=Level.t*inttypeu=Level.t->Level.tlethashconshdir(b,nasx)=letb'=hdirbinifb'==bthenxelse(b',n)leteql1l2=l1==l2||matchl1,l2with|(b,n),(b',n')->b==b'&&n==n'lethash(x,n)=n+Level.hashxendmoduleH=Hashcons.Make(ExprHash)lethcons=Hashcons.simple_hconsH.generateH.hconsLevel.hconsletmakel=(l,0)letcompareuv=ifu==vthen0elselet(x,n)=uand(x',n')=vinletc=Int.comparenn'inifInt.equal0cthenLevel.comparexx'elsecletsprop=hcons(Level.sprop,0)letprop=hcons(Level.prop,0)letset=hcons(Level.set,0)lettype1=hcons(Level.set,1)letis_small=function|(l,0)->Level.is_smalll|_->falseletequalxy=x==y||(let(u,n)=xand(v,n')=yinInt.equalnn'&&Level.equaluv)lethash=ExprHash.hashletleq(u,n)(v,n')=letcmp=Level.compareuvinifInt.equalcmp0thenn<=n'elseifn<=n'then(Level.is_propu&¬(Level.is_spropv))elsefalseletsuccessor(u,nase)=ifis_smallethentype1else(u,n+1)letaddnk(u,nasx)=ifk=0thenxelseifLevel.is_smalluthen(Level.set,n+k)else(u,n+k)typesuper_result=SuperSameofbool(* The level expressions are in cumulativity relation. boolean
indicates if left is smaller than right? *)|SuperDiffofint(* The level expressions are unrelated, the comparison result
is canonical *)(** [super u v] compares two level expressions,
returning [SuperSame] if they refer to the same level at potentially different
increments or [SuperDiff] if they are different. The booleans indicate if the
left expression is "smaller" than the right one in both cases. *)letsuper(u,n)(v,n')=letcmp=Level.compareuvinifInt.equalcmp0thenSuperSame(n<n')elseletopenRawLevelinmatchLevel.datau,n,Level.datav,n'with|SProp,_,SProp,_|Prop,_,Prop,_->SuperSame(n<n')|SProp,0,Prop,0->SuperSametrue|Prop,0,SProp,0->SuperSamefalse|(SProp|Prop),0,_,_->SuperSametrue|_,_,(SProp|Prop),0->SuperSamefalse|_,_,_,_->SuperDiffcmpletto_string(v,n)=ifInt.equaln0thenLevel.to_stringvelseLevel.to_stringv^"+"^string_of_intnletprx=str(to_stringx)letpr_withf(v,n)=ifInt.equaln0thenfvelsefv++str"+"++intnletis_level=function|(_v,0)->true|_->falseletlevel=function|(v,0)->Somev|_->Noneletget_level(v,_n)=vletmapf(v,nasx)=letv'=fvinifv'==vthenxelseifLevel.is_propv'&&n!=0then(Level.set,n)else(v',n)endtypet=Expr.tlistlettipl=[l]letconsxl=x::lletrechash=function|[]->0|e::l->Hashset.Combine.combinesmall(Expr.ExprHash.hashe)(hashl)letequalxy=x==y||List.equalExpr.equalxyletcomparexy=ifx==ythen0elseList.compareExpr.comparexymoduleHuniv=Hashcons.Hlist(Expr)lethcons=Hashcons.recursive_hconsHuniv.generateHuniv.hconsExpr.hconsmoduleSelf=structtypenonrect=tletcompare=compareendmoduleMap=CMap.Make(Self)moduleSet=CSet.Make(Self)letmakel=tip(Expr.makel)lettipx=tipxletprl=matchlwith|[u]->Expr.pru|_->str"max("++hov0(prlist_with_seppr_commaExpr.prl)++str")"letpr_withfl=matchlwith|[u]->Expr.pr_withfu|_->str"max("++hov0(prlist_with_seppr_comma(Expr.pr_withf)l)++str")"letis_levell=matchlwith|[l]->Expr.is_levell|_->falseletrecis_levelsl=matchlwith|l::r->Expr.is_levell&&is_levelsr|[]->trueletlevell=matchlwith|[l]->Expr.levell|_->Noneletlevelsl=List.fold_left(funaccx->LSet.add(Expr.get_levelx)acc)LSet.emptylletis_smallu=matchuwith|[l]->Expr.is_smalll|_->falseletsprop=tipExpr.sprop(* The lower predicative level of the hierarchy that contains (impredicative)
Prop and singleton inductive types *)lettype0m=tipExpr.prop(* The level of sets *)lettype0=tipExpr.set(* When typing [Prop] and [Set], there is no constraint on the level,
hence the definition of [type1_univ], the type of [Prop] *)lettype1=tipExpr.type1letis_spropx=equalspropxletis_type0mx=equaltype0mxletis_type0x=equaltype0x(* Returns the formal universe that lies just above the universe variable u.
Used to type the sort u. *)letsuperl=ifis_smalllthentype1elseList.Smart.map(funx->Expr.successorx)lletaddnnl=List.Smart.map(funx->Expr.addnnx)lletrecmerge_univsl1l2=matchl1,l2with|[],_->l2|_,[]->l1|h1::t1,h2::t2->letopenExprin(matchsuperh1h2with|SuperSametrue(* h1 < h2 *)->merge_univst1l2|SuperSamefalse->merge_univsl1t2|SuperDiffc->ifc<=0(* h1 < h2 is name order *)thenconsh1(merge_univst1l2)elseconsh2(merge_univsl1t2))letsortu=letrecauxal=matchlwith|b::l'->letopenExprin(matchsuperabwith|SuperSamefalse->auxal'|SuperSametrue->l|SuperDiffc->ifc<=0thenconsalelseconsb(auxal'))|[]->consalinList.fold_right(funaacc->auxaacc)u[](* Returns the formal universe that is greater than the universes u and v.
Used to type the products. *)letsupxy=merge_univsxyletexists=List.existsletfor_all=List.for_allletreprx:t=xendtypeuniverse=Universe.t(* The level of predicative Set *)lettype0m_univ=Universe.type0mlettype0_univ=Universe.type0lettype1_univ=Universe.type1letis_type0m_univ=Universe.is_type0mletis_type0_univ=Universe.is_type0letis_univ_variablel=Universe.levell!=Noneletis_small_univ=Universe.is_smallletpr_uni=Universe.prletsup=Universe.supletsuper=Universe.superletuniverse_level=Universe.leveltypeconstraint_type=AcyclicGraph.constraint_type=Lt|Le|Eqtypeexplanation=(constraint_type*Level.t)listletconstraint_type_ordc1c2=matchc1,c2with|Lt,Lt->0|Lt,_->-1|Le,Lt->1|Le,Le->0|Le,Eq->-1|Eq,Eq->0|Eq,_->1(* Constraints and sets of constraints. *)typeuniv_constraint=Level.t*constraint_type*Level.tletpr_constraint_typeop=letop_str=matchopwith|Lt->" < "|Le->" <= "|Eq->" = "instrop_strmoduleUConstraintOrd=structtypet=univ_constraintletcompare(u,c,v)(u',c',v')=leti=constraint_type_ordcc'inifnot(Int.equali0)thenielseleti'=Level.compareuu'inifnot(Int.equali'0)theni'elseLevel.comparevv'endmoduleConstraint=structmoduleS=Set.Make(UConstraintOrd)includeSletprprlc=v0(prlist_with_sepspc(fun(u1,op,u2)->hov0(prlu1++pr_constraint_typeop++prlu2))(elementsc))endletempty_constraint=Constraint.emptyletunion_constraint=Constraint.unionleteq_constraint=Constraint.equaltypeconstraints=Constraint.tmoduleHconstraint=Hashcons.Make(structtypet=univ_constrainttypeu=universe_level->universe_levellethashconshul(l1,k,l2)=(hull1,k,hull2)leteq(l1,k,l2)(l1',k',l2')=l1==l1'&&k==k'&&l2==l2'lethash=Hashtbl.hashend)moduleHconstraints=Hashcons.Make(structtypet=constraintstypeu=univ_constraint->univ_constraintlethashconshucs=Constraint.fold(funx->Constraint.add(hucx))sConstraint.emptyleteqss'=List.for_all2eq(==)(Constraint.elementss)(Constraint.elementss')lethash=Hashtbl.hashend)lethcons_constraint=Hashcons.simple_hconsHconstraint.generateHconstraint.hconsLevel.hconslethcons_constraints=Hashcons.simple_hconsHconstraints.generateHconstraints.hconshcons_constraint(** A value with universe constraints. *)type'aconstrained='a*constraintsletconstraints_of(_,cst)=cst(** Constraint functions. *)type'aconstraint_function='a->'a->constraints->constraintsletenforce_eq_leveluvc=(* We discard trivial constraints like u=u *)ifLevel.equaluvthencelseConstraint.add(u,Eq,v)cletenforce_equvc=matchUniverse.levelu,Universe.levelvwith|Someu,Somev->enforce_eq_leveluvc|_->anomaly(Pp.str"A universe comparison can only happen between variables.")letcheck_univ_equv=Universe.equaluvletenforce_equvc=ifcheck_univ_equvthencelseenforce_equvcletconstraint_add_leqvuc=(* We just discard trivial constraints like u<=u *)ifUniverse.Expr.equalvuthencelsematchv,uwith|(x,n),(y,m)->letj=m-ninifj=-1(* n = m+1, v+1 <= u <-> v < u *)thenConstraint.add(x,Lt,y)celseifj<=-1(* n = m+k, v+k <= u and k>0 *)thenifLevel.equalxythen(* u+k <= u with k>0 *)Constraint.add(x,Lt,x)celseanomaly(Pp.str"Unable to handle arbitrary u+k <= v constraints.")elseifj=0thenConstraint.add(x,Le,y)celse(* j >= 1 *)(* m = n + k, u <= v+k *)ifLevel.equalxythenc(* u <= u+k, trivial *)elseifLevel.is_smallxthenc(* Prop,Set <= u+S k, trivial *)elseConstraint.add(x,Le,y)c(* u <= v implies u <= v+k *)letcheck_univ_leq_oneuv=Universe.exists(Universe.Expr.lequ)vletcheck_univ_lequv=Universe.for_all(funu->check_univ_leq_oneuv)uletenforce_lequvc=matchUniverse.is_spropu,Universe.is_spropvwith|true,true->c|true,false->Constraint.add(Level.sprop,Le,Level.prop)c|false,true->Constraint.add(Level.prop,Le,Level.sprop)c|false,false->List.fold_left(funcv->(List.fold_left(funcu->constraint_add_lequvc)cu))cvletenforce_lequvc=ifcheck_univ_lequvthencelseenforce_lequvcletenforce_leq_leveluvc=ifLevel.equaluvthencelseConstraint.add(u,Le,v)c(* Miscellaneous functions to remove or test local univ assumed to
occur in a universe *)letuniv_level_memuv=List.exists(fun(l,n)->Int.equaln0&&Level.equalul)vletuniv_level_remuvmin=matchUniverse.levelvwith|Someu'->ifLevel.equaluu'thenminelsev|None->List.filter(fun(l,n)->not(Int.equaln0&&Level.equalul))v(* Is u mentioned in v (or equals to v) ? *)(**********************************************************************)(** Universe polymorphism *)(**********************************************************************)(** A universe level substitution, note that no algebraic universes are
involved *)typeuniverse_level_subst=universe_leveluniverse_map(** A full substitution might involve algebraic universes *)typeuniverse_subst=universeuniverse_mapmoduleVariance=struct(** A universe position in the instance given to a cumulative
inductive can be the following. Note there is no Contravariant
case because [forall x : A, B <= forall x : A', B'] requires [A =
A'] as opposed to [A' <= A]. *)typet=Irrelevant|Covariant|Invariantletsupxy=matchx,ywith|Irrelevant,s|s,Irrelevant->s|Invariant,_|_,Invariant->Invariant|Covariant,Covariant->Covariantletequalab=matcha,bwith|Irrelevant,Irrelevant|Covariant,Covariant|Invariant,Invariant->true|(Irrelevant|Covariant|Invariant),_->falseletcheck_subtypexy=matchx,ywith|(Irrelevant|Covariant|Invariant),Irrelevant->true|Irrelevant,Covariant->false|(Covariant|Invariant),Covariant->true|(Irrelevant|Covariant),Invariant->false|Invariant,Invariant->trueletpr=function|Irrelevant->str"*"|Covariant->str"+"|Invariant->str"="letleq_constraintcstsvarianceuu'=matchvariancewith|Irrelevant->csts|Covariant->enforce_leq_leveluu'csts|Invariant->enforce_eq_leveluu'cstsleteq_constraintcstsvarianceuu'=matchvariancewith|Irrelevant->csts|Covariant|Invariant->enforce_eq_leveluu'cstsletleq_constraintsvarianceuu'csts=letlen=Array.lengthuinassert(len=Array.lengthu'&&len=Array.lengthvariance);Array.fold_left3leq_constraintcstsvarianceuu'leteq_constraintsvarianceuu'csts=letlen=Array.lengthuinassert(len=Array.lengthu'&&len=Array.lengthvariance);Array.fold_left3eq_constraintcstsvarianceuu'endmoduleInstance:sigtypet=Level.tarrayvalempty:tvalis_empty:t->boolvalof_array:Level.tarray->tvalto_array:t->Level.tarrayvalappend:t->t->tvalequal:t->t->boolvallength:t->intvalhcons:t->tvalhash:t->intvalshare:t->t*intvalsubst_fn:universe_level_subst_fn->t->tvalpr:(Level.t->Pp.t)->?variance:Variance.tarray->t->Pp.tvallevels:t->LSet.tend=structtypet=Level.tarrayletempty:t=[||]moduleHInstancestruct=structtypenonrect=ttypeu=Level.t->Level.tlethashconshuniva=letlen=Array.lengthainifInt.equallen0thenemptyelsebeginfori=0tolen-1doletx=Array.unsafe_getaiinletx'=hunivxinifx==x'then()elseArray.unsafe_setaix'done;aendleteqt1t2=t1==t2||(Int.equal(Array.lengtht1)(Array.lengtht2)&&letrecauxi=(Int.equali(Array.lengtht1))||(t1.(i)==t2.(i)&&aux(i+1))inaux0)lethasha=letaccu=ref0infori=0toArray.lengtha-1doletl=Array.unsafe_getaiinleth=Level.hashlinaccu:=Hashset.Combine.combine!accuh;done;(* [h] must be positive. *)leth=!acculand0x3FFFFFFFinhendmoduleHInstance=Hashcons.Make(HInstancestruct)lethcons=Hashcons.simple_hconsHInstance.generateHInstance.hconsLevel.hconslethash=HInstancestruct.hashletsharea=(hconsa,hasha)letempty=hcons[||]letis_emptyx=Int.equal(Array.lengthx)0letappendxy=ifArray.lengthx=0thenyelseifArray.lengthy=0thenxelseArray.appendxyletof_arraya=assert(Array.for_all(funx->not(Level.is_propx||Level.is_spropx))a);aletto_arraya=aletlengtha=Array.lengthaletsubst_fnfnt=lett'=CArray.Smart.mapfntinift'==tthentelseof_arrayt'letlevelsx=LSet.of_arrayxletprprl?variance=letppuiu=letv=Option.map(funv->v.(i))varianceinpr_opt_no_spcVariance.prv++prluinprvecti_with_sepspcppuletequaltu=t==u||(Array.is_emptyt&&Array.is_emptyu)||(CArray.for_all2Level.equaltu(* Necessary as universe instances might come from different modules and
unmarshalling doesn't preserve sharing *))endletenforce_eq_instancesxy=letax=Instance.to_arrayxanday=Instance.to_arrayyinifArray.lengthax!=Array.lengthaythenanomaly(Pp.(++)(Pp.str"Invalid argument: enforce_eq_instances called with")(Pp.str" instances of different lengths."));CArray.fold_right2enforce_eq_levelaxayletenforce_eq_variance_instances=Variance.eq_constraintsletenforce_leq_variance_instances=Variance.leq_constraintsletsubst_instance_levelsl=matchl.Level.datawith|Level.Varn->s.(n)|_->lletsubst_instance_instancesi=Array.Smart.map(funl->subst_instance_levelsl)iletsubst_instance_universesu=letfx=Universe.Expr.map(funu->subst_instance_levelsu)xinletu'=List.Smart.mapfuinifu==u'thenuelseUniverse.sortu'letsubst_instance_constraints(u,d,vasc)=letu'=subst_instance_levelsuinletv'=subst_instance_levelsvinifu'==u&&v'==vthencelse(u',d,v')letsubst_instance_constraintsscsts=Constraint.fold(funccsts->Constraint.add(subst_instance_constraintsc)csts)cstsConstraint.emptytype'apuniverses='a*Instance.tletout_punivs(x,_y)=xletin_punivsx=(x,Instance.empty)leteq_puniversesf(x,u)(y,u')=fxy&&Instance.equaluu'(** A context of universe levels with universe constraints,
representing local universe variables and constraints *)moduleUContext=structtypet=Instance.tconstrainedletmakex=x(** Universe contexts (variables as a list) *)letempty=(Instance.empty,Constraint.empty)letis_empty(univs,cst)=Instance.is_emptyunivs&&Constraint.is_emptycstletprprl?variance(univs,cstasctx)=ifis_emptyctxthenmt()elseh(Instance.prprl?varianceunivs++str" |= ")++h(v0(Constraint.prprlcst))lethcons(univs,cst)=(Instance.hconsunivs,hcons_constraintscst)letinstance(univs,_cst)=univsletconstraints(_univs,cst)=cstletunion(univs,cst)(univs',cst')=Instance.appendunivsunivs',Constraint.unioncstcst'letdestx=xletsize(x,_)=Instance.lengthxendtypeuniverse_context=UContext.tlethcons_universe_context=UContext.hconsmoduleAUContext=structtypet=Names.Name.tarrayconstrainedletmakenamescsts:t=names,cstsletrepr(inst,cst)=(Array.init(Array.lengthinst)(funi->Level.vari),cst)letprf?variancectx=UContext.prf?variance(reprctx)letinstantiateinst(u,cst)=assert(Array.lengthu=Array.lengthinst);subst_instance_constraintsinstcstletnames(nas,_)=naslethcons(univs,cst)=(Array.mapNames.Name.hconsunivs,hcons_constraintscst)letempty=([||],Constraint.empty)letis_empty(nas,cst)=Array.is_emptynas&&Constraint.is_emptycstletunion(nas,cst)(nas',cst')=(Array.appendnasnas',Constraint.unioncstcst')letsize(nas,_)=Array.lengthnasendtype'auniv_abstracted={univ_abstracted_value:'a;univ_abstracted_binder:AUContext.t;}letmap_univ_abstractedf{univ_abstracted_value;univ_abstracted_binder}=letuniv_abstracted_value=funiv_abstracted_valuein{univ_abstracted_value;univ_abstracted_binder}lethcons_abstract_universe_context=AUContext.hcons(** A set of universes with universe constraints.
We linearize the set to a list after typechecking.
Beware, representation could change.
*)moduleContextSet=structtypet=universe_setconstrainedletempty=(LSet.empty,Constraint.empty)letis_empty(univs,cst)=LSet.is_emptyunivs&&Constraint.is_emptycstletequal(univs,cstasx)(univs',cst'asy)=x==y||(LSet.equalunivsunivs'&&Constraint.equalcstcst')letof_sets=(s,Constraint.empty)letsingletonl=of_set(LSet.singletonl)letof_instancei=of_set(Instance.levelsi)letunion(univs,cstasx)(univs',cst'asy)=ifx==ythenxelseLSet.unionunivsunivs',Constraint.unioncstcst'letappend(univs,cst)(univs',cst')=letunivs=LSet.foldLSet.addunivsunivs'inletcst=Constraint.foldConstraint.addcstcst'in(univs,cst)letdiff(univs,cst)(univs',cst')=LSet.diffunivsunivs',Constraint.diffcstcst'letadd_universeu(univs,cst)=LSet.adduunivs,cstletadd_constraintscst'(univs,cst)=univs,Constraint.unioncstcst'letadd_instanceinst(univs,cst)=letv=Instance.to_arrayinstinletfoldaccuu=LSet.adduaccuinletunivs=Array.fold_leftfoldunivsvin(univs,cst)letsort_levelsa=Array.sortLevel.comparea;aletto_context(ctx,cst)=(Instance.of_array(sort_levels(Array.of_list(LSet.elementsctx))),cst)letof_context(ctx,cst)=(Instance.levelsctx,cst)letprprl(univs,cstasctx)=ifis_emptyctxthenmt()elseh(LSet.prprlunivs++str" |= ")++h(v0(Constraint.prprlcst))letconstraints(_univs,cst)=cstletlevels(univs,_cst)=univsletsize(univs,_)=LSet.cardinalunivsendtypeuniverse_context_set=ContextSet.t(** A value in a universe context (resp. context set). *)type'ain_universe_context='a*universe_contexttype'ain_universe_context_set='a*universe_context_setletextend_in_context_set(a,ctx)ctx'=(a,ContextSet.unionctxctx')(** Substitutions. *)letempty_level_subst=LMap.emptyletis_empty_level_subst=LMap.is_empty(** Substitution functions *)(** With level to level substitutions. *)letsubst_univs_level_levelsubstl=tryLMap.findlsubstwithNot_found->lletsubst_univs_level_universesubstu=letfx=Universe.Expr.map(funu->subst_univs_level_levelsubstu)xinletu'=List.Smart.mapfuinifu==u'thenuelseUniverse.sortu'letsubst_univs_level_instancesubsti=leti'=Instance.subst_fn(subst_univs_level_levelsubst)iinifi==i'thenielsei'letsubst_univs_level_constraintsubst(u,d,v)=letu'=subst_univs_level_levelsubstuandv'=subst_univs_level_levelsubstvinifd!=Lt&&Level.equalu'v'thenNoneelseSome(u',d,v')letsubst_univs_level_constraintssubstcsts=Constraint.fold(func->Option.fold_rightConstraint.add(subst_univs_level_constraintsubstc))cstsConstraint.emptyletsubst_univs_level_abstract_universe_contextsubst(inst,csts)=inst,subst_univs_level_constraintssubstcsts(** With level to universe substitutions. *)typeuniverse_subst_fn=universe_level->universeletmake_substsubst=funl->LMap.findlsubstletsubst_univs_expr_optfn(l,n)=Universe.addnn(fnl)letsubst_univs_universefnul=letsubst,nosubst=List.fold_right(funu(subst,nosubst)->tryleta'=subst_univs_expr_optfnuin(a'::subst,nosubst)withNot_found->(subst,u::nosubst))ul([],[])inifCList.is_emptysubstthenulelseletsubsts=List.fold_leftUniverse.merge_univs[]substinList.fold_left(funaccu->Universe.merge_univsacc(Universe.tipu))substsnosubstletmake_instance_substi=letarr=Instance.to_arrayiinArray.fold_left_i(funiaccl->LMap.addl(Level.vari)acc)LMap.emptyarrletmake_inverse_instance_substi=letarr=Instance.to_arrayiinArray.fold_left_i(funiaccl->LMap.add(Level.vari)lacc)LMap.emptyarrletmake_abstract_instance(ctx,_)=Array.init(Array.lengthctx)(funi->Level.vari)letabstract_universesnasctx=letinstance=UContext.instancectxinlet()=assert(Int.equal(Array.lengthnas)(Instance.lengthinstance))inletsubst=make_instance_substinstanceinletcstrs=subst_univs_level_constraintssubst(UContext.constraintsctx)inletctx=(nas,cstrs)ininstance,ctxletreccompact_univsvarsiu=matchuwith|[]->(s,List.revvars)|(lvl,_)::u->matchLevel.var_indexlvlwith|Somekwhennot(LMap.memlvls)->letlvl'=Level.variincompact_univ(LMap.addlvllvl's)(k::vars)(i+1)u|_->compact_univsvarsiuletcompact_univu=let(s,s')=compact_univLMap.empty[]0uin(subst_univs_level_universesu,s')(** Pretty-printing *)letpr_constraintsprl=Constraint.prprlletpr_universe_context=UContext.prletpr_abstract_universe_context=AUContext.prletpr_universe_context_set=ContextSet.prletpr_universe_subst=LMap.pr(funu->str" := "++Universe.pru++spc())letpr_universe_level_subst=LMap.pr(funu->str" := "++Level.pru++spc())moduleHuniverse_set=Hashcons.Make(structtypet=universe_settypeu=universe_level->universe_levellethashconshucs=LSet.fold(funx->LSet.add(hucx))sLSet.emptyleteqss'=LSet.equalss'lethash=Hashtbl.hashend)lethcons_universe_set=Hashcons.simple_hconsHuniverse_set.generateHuniverse_set.hconsLevel.hconslethcons_universe_context_set(v,c)=(hcons_universe_setv,hcons_constraintsc)lethcons_univx=Universe.hconsx(* Universe inconsistency: error raised when trying to enforce a relation
that would create a cycle in the graph of universes. *)typeuniv_inconsistency=constraint_type*universe*universe*explanationLazy.toption(* Do not use in this file as we may be type-in-type *)exceptionUniverseInconsistencyofuniv_inconsistencyletexplain_universe_inconsistencyprl(o,u,v,p:univ_inconsistency)=letpr_uni=Universe.pr_withprlinletpr_rel=function|Eq->str"="|Lt->str"<"|Le->str"<="inletreason=matchpwith|None->mt()|Somep->letp=Lazy.forcepinifp=[]thenmt()elsestr" because"++spc()++pr_univ++prlist(fun(r,v)->spc()++pr_relr++str" "++prlv)p++(ifUniverse.equal(Universe.make(snd(List.lastp)))uthenmt()else(spc()++str"= "++pr_uniu))instr"Cannot enforce"++spc()++pr_uniu++spc()++pr_relo++spc()++pr_univ++reason