123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691692693694695696697698699700701702703704705706707708709710711712713714715716717718719720721722723724725726727728729730731732733734735736737738739740741742743744745746747748749750751752753754755756757758759760761762763764765766767768769770771772773774775776777778779780781782783784785786787788789790791792793794795796797798799800801802803804805806807808809810811812813814815816817818819820821822823824825826827828829830831832833834835836837838839840841842843844845846847848849850851852853854855856857858859860861862863864865866867(************************************************************************)(* * 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) *)(************************************************************************)openCErrorsopenUtilopenNamesopenConstropenContextmoduleESorts=structincludeEvd.MiniEConstr.ESortsletequalsigmas1s2=Sorts.equal(kindsigmas1)(kindsigmas2)endmoduleEInstance=structincludeEvd.MiniEConstr.EInstanceletequalsigmai1i2=Univ.Instance.equal(kindsigmai1)(kindsigmai2)endinclude(Evd.MiniEConstr:moduletypeofEvd.MiniEConstrwithmoduleESorts:=ESortsandmoduleEInstance:=EInstance)typetypes=ttypeconstr=ttypeexistential=tpexistentialtypecase_return=tpcase_returntypecase_branch=tpcase_branchtypecase_invert=tpcase_inverttypecase=(t,t,EInstance.t)pcasetypefixpoint=(t,t)pfixpointtypecofixpoint=(t,t)pcofixpointtypeunsafe_judgment=(constr,types)Environ.punsafe_judgmenttypeunsafe_type_judgment=typesEnviron.punsafe_type_judgmenttypenamed_declaration=(constr,types)Context.Named.Declaration.pttyperel_declaration=(constr,types)Context.Rel.Declaration.pttypenamed_context=(constr,types)Context.Named.pttyperel_context=(constr,types)Context.Rel.pttype'apuniverses='a*EInstance.tletin_punivsa=(a,EInstance.empty)letmkSProp=of_kind(Sort(ESorts.makeSorts.sprop))letmkProp=of_kind(Sort(ESorts.makeSorts.prop))letmkSet=of_kind(Sort(ESorts.makeSorts.set))letmkTypeu=of_kind(Sort(ESorts.make(Sorts.sort_of_univu)))letmkReln=of_kind(Reln)letmkVarid=of_kind(Varid)letmkMetan=of_kind(Metan)letmkEvare=of_kind(Evare)letmkSorts=of_kind(Sort(ESorts.makes))letmkCast(b,k,t)=of_kind(Cast(b,k,t))letmkProd(na,t,u)=of_kind(Prod(na,t,u))letmkLambda(na,t,c)=of_kind(Lambda(na,t,c))letmkLetIn(na,b,t,c)=of_kind(LetIn(na,b,t,c))letmkApp(f,arg)=of_kind(App(f,arg))letmkConstUpc=of_kind(Constpc)letmkConstc=of_kind(Const(in_punivsc))letmkIndUpi=of_kind(Indpi)letmkIndi=of_kind(Ind(in_punivsi))letmkConstructUpc=of_kind(Constructpc)letmkConstructc=of_kind(Construct(in_punivsc))letmkConstructUi((ind,u),i)=of_kind(Construct((ind,i),u))letmkCase(ci,u,pms,c,iv,r,p)=of_kind(Case(ci,u,pms,c,iv,r,p))letmkFixf=of_kind(Fixf)letmkCoFixf=of_kind(CoFixf)letmkProj(p,c)=of_kind(Proj(p,c))letmkArrowt1rt2=of_kind(Prod(make_annotAnonymousr,t1,t2))letmkArrowRt1t2=mkArrowt1Sorts.Relevantt2letmkInti=of_kind(Inti)letmkFloatf=of_kind(Floatf)letmkArray(u,t,def,ty)=of_kind(Array(u,t,def,ty))letmkRef(gr,u)=letopenGlobRefinmatchgrwith|ConstRefc->mkConstU(c,u)|IndRefind->mkIndU(ind,u)|ConstructRefc->mkConstructU(c,u)|VarRefx->mkVarxlettype1=mkSortSorts.type1letapplist(f,arg)=mkApp(f,Array.of_listarg)letapplistcfarg=mkApp(f,Array.of_listarg)letisRelsigmac=matchkindsigmacwithRel_->true|_->falseletisVarsigmac=matchkindsigmacwithVar_->true|_->falseletisIndsigmac=matchkindsigmacwithInd_->true|_->falseletisEvarsigmac=matchkindsigmacwithEvar_->true|_->falseletisMetasigmac=matchkindsigmacwithMeta_->true|_->falseletisSortsigmac=matchkindsigmacwithSort_->true|_->falseletisCastsigmac=matchkindsigmacwithCast_->true|_->falseletisAppsigmac=matchkindsigmacwithApp_->true|_->falseletisLambdasigmac=matchkindsigmacwithLambda_->true|_->falseletisLetInsigmac=matchkindsigmacwithLetIn_->true|_->falseletisProdsigmac=matchkindsigmacwithProd_->true|_->falseletisConstsigmac=matchkindsigmacwithConst_->true|_->falseletisConstructsigmac=matchkindsigmacwithConstruct_->true|_->falseletisFixsigmac=matchkindsigmacwithFix_->true|_->falseletisCoFixsigmac=matchkindsigmacwithCoFix_->true|_->falseletisCasesigmac=matchkindsigmacwithCase_->true|_->falseletisProjsigmac=matchkindsigmacwithProj_->true|_->falseletrecisTypesigmac=matchkindsigmacwith|Sorts->(matchESorts.kindsigmaswith|Sorts.Type_->true|_->false)|Cast(c,_,_)->isTypesigmac|_->falseletisVarIdsigmaidc=matchkindsigmacwithVarid'->Id.equalidid'|_->falseletisRelNsigmanc=matchkindsigmacwithReln'->Int.equalnn'|_->falseletisRefsigmac=matchkindsigmacwith|Const_|Ind_|Construct_|Var_->true|_->falseletisRefXsigmaxc=letopenGlobRefinmatchx,kindsigmacwith|ConstRefc,Const(c',_)->Constant.CanOrd.equalcc'|IndRefi,Ind(i',_)->Ind.CanOrd.equalii'|ConstructRefi,Construct(i',_)->Construct.CanOrd.equalii'|VarRefid,Varid'->Id.equalidid'|_->falseletdestRelsigmac=matchkindsigmacwith|Relp->p|_->raiseDestKOletdestVarsigmac=matchkindsigmacwith|Varp->p|_->raiseDestKOletdestIndsigmac=matchkindsigmacwith|Indp->p|_->raiseDestKOletdestEvarsigmac=matchkindsigmacwith|Evarp->p|_->raiseDestKOletdestMetasigmac=matchkindsigmacwith|Metap->p|_->raiseDestKOletdestSortsigmac=matchkindsigmacwith|Sortp->p|_->raiseDestKOletdestCastsigmac=matchkindsigmacwith|Cast(c,k,t)->(c,k,t)|_->raiseDestKOletdestAppsigmac=matchkindsigmacwith|App(f,a)->(f,a)|_->raiseDestKOletdestLambdasigmac=matchkindsigmacwith|Lambda(na,t,c)->(na,t,c)|_->raiseDestKOletdestLetInsigmac=matchkindsigmacwith|LetIn(na,b,t,c)->(na,b,t,c)|_->raiseDestKOletdestProdsigmac=matchkindsigmacwith|Prod(na,t,c)->(na,t,c)|_->raiseDestKOletdestConstsigmac=matchkindsigmacwith|Constp->p|_->raiseDestKOletdestConstructsigmac=matchkindsigmacwith|Constructp->p|_->raiseDestKOletdestFixsigmac=matchkindsigmacwith|Fixp->p|_->raiseDestKOletdestCoFixsigmac=matchkindsigmacwith|CoFixp->p|_->raiseDestKOletdestCasesigmac=matchkindsigmacwith|Case(ci,u,pms,t,iv,c,p)->(ci,u,pms,t,iv,c,p)|_->raiseDestKOletdestProjsigmac=matchkindsigmacwith|Proj(p,c)->(p,c)|_->raiseDestKOletdestRefsigmac=letopenGlobRefinmatchkindsigmacwith|Varx->VarRefx,EInstance.empty|Const(c,u)->ConstRefc,u|Ind(ind,u)->IndRefind,u|Construct(c,u)->ConstructRefc,u|_->raiseDestKOletdecompose_appsigmac=matchkindsigmacwith|App(f,cl)->(f,Array.to_listcl)|_->(c,[])letdecompose_lamsigmac=letreclamdec_reclc=matchkindsigmacwith|Lambda(x,t,c)->lamdec_rec((x,t)::l)c|Cast(c,_,_)->lamdec_reclc|_->l,cinlamdec_rec[]cletdecompose_lam_assumsigmac=letopenRel.Declarationinletreclamdec_reclc=matchkindsigmacwith|Lambda(x,t,c)->lamdec_rec(Context.Rel.add(LocalAssum(x,t))l)c|LetIn(x,b,t,c)->lamdec_rec(Context.Rel.add(LocalDef(x,b,t))l)c|Cast(c,_,_)->lamdec_reclc|_->l,cinlamdec_recContext.Rel.emptycletdecompose_lam_n_assumsigmanc=letopenRel.Declarationinifn<0thenuser_errPp.(str"decompose_lam_n_assum: integer parameter must be positive");letreclamdec_reclnc=ifInt.equaln0thenl,celsematchkindsigmacwith|Lambda(x,t,c)->lamdec_rec(Context.Rel.add(LocalAssum(x,t))l)(n-1)c|LetIn(x,b,t,c)->lamdec_rec(Context.Rel.add(LocalDef(x,b,t))l)nc|Cast(c,_,_)->lamdec_reclnc|c->user_errPp.(str"decompose_lam_n_assum: not enough abstractions")inlamdec_recContext.Rel.emptyncletdecompose_lam_n_declssigman=letopenRel.Declarationinifn<0thenuser_errPp.(str"decompose_lam_n_decls: integer parameter must be positive");letreclamdec_reclnc=ifInt.equaln0thenl,celsematchkindsigmacwith|Lambda(x,t,c)->lamdec_rec(Context.Rel.add(LocalAssum(x,t))l)(n-1)c|LetIn(x,b,t,c)->lamdec_rec(Context.Rel.add(LocalDef(x,b,t))l)(n-1)c|Cast(c,_,_)->lamdec_reclnc|c->user_errPp.(str"decompose_lam_n_decls: not enough abstractions")inlamdec_recContext.Rel.emptynletlamnnenvb=letreclamrec=function|(0,env,b)->b|(n,((v,t)::l),b)->lamrec(n-1,l,mkLambda(v,t,b))|_->assertfalseinlamrec(n,env,b)letcompose_lamlb=lamn(List.lengthl)lbletrecto_lambdasigmanprod=ifInt.equaln0thenprodelsematchkindsigmaprodwith|Prod(na,ty,bd)->mkLambda(na,ty,to_lambdasigma(n-1)bd)|Cast(c,_,_)->to_lambdasigmanc|_->user_err~hdr:"to_lambda"(Pp.mt())letdecompose_prodsigmac=letrecproddec_reclc=matchkindsigmacwith|Prod(x,t,c)->proddec_rec((x,t)::l)c|Cast(c,_,_)->proddec_reclc|_->l,cinproddec_rec[]cletdecompose_prod_assumsigmac=letopenRel.Declarationinletrecproddec_reclc=matchkindsigmacwith|Prod(x,t,c)->proddec_rec(Context.Rel.add(LocalAssum(x,t))l)c|LetIn(x,b,t,c)->proddec_rec(Context.Rel.add(LocalDef(x,b,t))l)c|Cast(c,_,_)->proddec_reclc|_->l,cinproddec_recContext.Rel.emptycletdecompose_prod_n_assumsigmanc=letopenRel.Declarationinifn<0thenuser_errPp.(str"decompose_prod_n_assum: integer parameter must be positive");letrecprodec_reclnc=ifInt.equaln0thenl,celsematchkindsigmacwith|Prod(x,t,c)->prodec_rec(Context.Rel.add(LocalAssum(x,t))l)(n-1)c|LetIn(x,b,t,c)->prodec_rec(Context.Rel.add(LocalDef(x,b,t))l)(n-1)c|Cast(c,_,_)->prodec_reclnc|c->user_errPp.(str"decompose_prod_n_assum: not enough assumptions")inprodec_recContext.Rel.emptyncletexistential_type=Evd.existential_typeletliftnc=of_constr(Vars.liftn(unsafe_to_constrc))letof_branches:Constr.case_brancharray->case_brancharray=matchEvd.MiniEConstr.unsafe_eqwith|Refl->funx->xletunsafe_to_branches:case_brancharray->Constr.case_brancharray=matchEvd.MiniEConstr.unsafe_eqwith|Refl->funx->xletof_return:Constr.case_return->case_return=matchEvd.MiniEConstr.unsafe_eqwith|Refl->funx->xletunsafe_to_return:case_return->Constr.case_return=matchEvd.MiniEConstr.unsafe_eqwith|Refl->funx->xletmap_branchesfbr=letfc=unsafe_to_constr(f(of_constrc))inof_branches(Constr.map_branchesf(unsafe_to_branchesbr))letmap_return_predicatefp=letfc=unsafe_to_constr(f(of_constrc))inof_return(Constr.map_return_predicatef(unsafe_to_returnp))letmapsigmafc=letfc=unsafe_to_constr(f(of_constrc))inof_constr(Constr.mapf(unsafe_to_constr(whd_evarsigmac)))letmap_with_binderssigmagflc=letflc=unsafe_to_constr(fl(of_constrc))inof_constr(Constr.map_with_bindersgfl(unsafe_to_constr(whd_evarsigmac)))letitersigmafc=letfc=f(of_constrc)inConstr.iterf(unsafe_to_constr(whd_evarsigmac))letexpand_caseenv_sigma(ci,u,pms,p,iv,c,bl)=letu=EInstance.unsafe_to_instanceuinletpms=unsafe_to_constr_arraypmsinletp=unsafe_to_returnpinletiv=unsafe_to_case_invertivinletc=unsafe_to_constrcinletbl=unsafe_to_branchesblinlet(ci,p,iv,c,bl)=Inductive.expand_caseenv(ci,u,pms,p,iv,c,bl)inletp=of_constrpinletc=of_constrcinletiv=of_case_invertivinletbl=of_constr_arrayblin(ci,p,iv,c,bl)letannotate_caseenvsigma(ci,u,pms,p,iv,c,blascase)=let(_,p,_,_,bl)=expand_caseenvsigmacaseinletp=(* Too bad we need to fetch this data in the environment, should be in the
case_info instead. *)let(_,mip)=Inductive.lookup_mind_specifenvci.ci_indindecompose_lam_n_declssigma(mip.Declarations.mind_nrealdecls+1)pinletmk_brcn=decompose_lam_n_declssigmancinletbl=Array.map2mk_brblci.ci_cstr_ndeclsin(ci,u,pms,p,iv,c,bl)letexpand_branchenv_sigmaupms(ind,i)(nas,_br)=letopenDeclarationsinletu=EInstance.unsafe_to_instanceuinletpms=unsafe_to_constr_arraypmsinlet(mib,mip)=Inductive.lookup_mind_specifenvindinletparamdecl=Vars.subst_instance_contextumib.mind_params_ctxtinletparamsubst=Vars.subst_of_rel_context_instanceparamdecl(Array.to_listpms)inletsubst=paramsubst@Inductive.ind_subst(fstind)mibuinlet(ctx,_)=mip.mind_nf_lc.(i-1)inlet(ctx,_)=List.chopmip.mind_consnrealdecls.(i-1)ctxinletans=Inductive.instantiate_contextusubstnasctxinletans:rel_context=matchEvd.MiniEConstr.unsafe_eqwithRefl->ansinansletcontract_caseenv_sigma(ci,p,iv,c,bl)=letp=unsafe_to_constrpinletiv=unsafe_to_case_invertivinletc=unsafe_to_constrcinletbl=unsafe_to_constr_arrayblinlet(ci,u,pms,p,iv,c,bl)=Inductive.contract_caseenv(ci,p,iv,c,bl)inletu=EInstance.makeuinletpms=of_constr_arraypmsinletp=of_returnpinletiv=of_case_invertivinletc=of_constrcinletbl=of_branchesblin(ci,u,pms,p,iv,c,bl)letiter_with_full_bindersenvsigmagfnc=letopenContext.Rel.Declarationinmatchkindsigmacwith|(Rel_|Meta_|Var_|Sort_|Const_|Ind_|Construct_|Int_|Float_)->()|Cast(c,_,t)->fnc;fnt|Prod(na,t,c)->fnt;f(g(LocalAssum(na,t))n)c|Lambda(na,t,c)->fnt;f(g(LocalAssum(na,t))n)c|LetIn(na,b,t,c)->fnb;fnt;f(g(LocalDef(na,b,t))n)c|App(c,l)->fnc;Array.Fun1.iterfnl|Evar(_,l)->List.iter(func->fnc)l|Case(ci,u,pms,p,iv,c,bl)->let(ci,_,pms,p,iv,c,bl)=annotate_caseenvsigma(ci,u,pms,p,iv,c,bl)inletf_ctx(ctx,c)=f(List.fold_rightgctxn)cinArray.Fun1.iterfnpms;f_ctxp;iter_invert(fn)iv;fnc;Array.iterf_ctxbl|Proj(p,c)->fnc|Fix(_,(lna,tl,bl))->Array.iter(fn)tl;letn'=Array.fold_left2_i(funinnat->g(LocalAssum(na,liftit))n)nlnatlinArray.iter(fn')bl|CoFix(_,(lna,tl,bl))->Array.iter(fn)tl;letn'=Array.fold_left2_i(funinnat->g(LocalAssum(na,liftit))n)nlnatlinArray.iter(fn')bl|Array(_u,t,def,ty)->Array.Fun1.iterfnt;fndef;fntyletiter_with_binderssigmagfnc=letflc=fl(of_constrc)inConstr.iter_with_bindersgfn(unsafe_to_constr(whd_evarsigmac))letfoldsigmafaccc=letfaccc=facc(of_constrc)inConstr.foldfacc(unsafe_to_constr(whd_evarsigmac))letcompare_genkeq_insteq_sorteq_constrnargsc1c2=(c1==c2)||Constr.compare_head_gen_withkkeq_insteq_sorteq_constrnargsc1c2leteq_constrsigmac1c2=letkindc=kindsigmacinleteq_inst_i1i2=EInstance.equalsigmai1i2inleteq_sortss1s2=ESorts.equalsigmas1s2inletreceq_constrnargsc1c2=compare_genkindeq_insteq_sortseq_constrnargsc1c2ineq_constr0c1c2leteq_constr_nounivssigmac1c2=letkindc=kindsigmacinletreceq_constrnargsc1c2=compare_genkind(fun___->true)(fun__->true)eq_constrnargsc1c2ineq_constr0c1c2letcompare_constrsigmacmpc1c2=letkindc=kindsigmacinleteq_inst_i1i2=EInstance.equalsigmai1i2inleteq_sortss1s2=ESorts.equalsigmas1s2inletcmpnargsc1c2=cmpc1c2incompare_genkindeq_insteq_sortscmp0c1c2letcompare_cumulative_instancescv_pbnargs_okvariancesuu'cstrs=letopenUnivProbleminifnotnargs_okthenenforce_eq_instances_univsfalseuu'cstrselseCArray.fold_left3(funcstrsvuu'->letopenUniv.Varianceinmatchvwith|Irrelevant->Set.add(UWeak(u,u'))cstrs|Covariant->letu=Univ.Universe.makeuinletu'=Univ.Universe.makeu'in(matchcv_pbwith|Reduction.CONV->Set.add(UEq(u,u'))cstrs|Reduction.CUMUL->Set.add(ULe(u,u'))cstrs)|Invariant->letu=Univ.Universe.makeuinletu'=Univ.Universe.makeu'inSet.add(UEq(u,u'))cstrs)cstrsvariances(Univ.Instance.to_arrayu)(Univ.Instance.to_arrayu')letcmp_inductivescv_pb(mind,indasspec)nargsu1u2cstrs=letopenUnivProbleminmatchmind.Declarations.mind_variancewith|None->enforce_eq_instances_univsfalseu1u2cstrs|Somevariances->letnum_param_arity=Reduction.inductive_cumulativity_argumentsspecincompare_cumulative_instancescv_pb(Int.equalnum_param_aritynargs)variancesu1u2cstrsletcmp_constructors(mind,ind,cnsasspec)nargsu1u2cstrs=letopenUnivProbleminmatchmind.Declarations.mind_variancewith|None->enforce_eq_instances_univsfalseu1u2cstrs|Some_->letnum_cnstr_args=Reduction.constructor_cumulativity_argumentsspecinifnot(Int.equalnum_cnstr_argsnargs)thenenforce_eq_instances_univsfalseu1u2cstrselseArray.fold_left2(funcstrsu1u2->UnivProblem.(Set.add(UWeak(u1,u2))cstrs))cstrs(Univ.Instance.to_arrayu1)(Univ.Instance.to_arrayu2)leteq_universesenvsigmacstrscv_pbrefargsll'=ifEInstance.is_emptylthen(assert(EInstance.is_emptyl');true)elseletl=EInstance.kindsigmalandl'=EInstance.kindsigmal'inletopenGlobRefinletopenUnivProbleminmatchrefargswith|Some(ConstRefc,1)whenEnviron.is_array_typeenvc->cstrs:=compare_cumulative_instancescv_pbtrue[|Univ.Variance.Irrelevant|]ll'!cstrs;true|None|Some(ConstRef_,_)->cstrs:=enforce_eq_instances_univstruell'!cstrs;true|Some(VarRef_,_)->assertfalse(* variables don't have instances *)|Some(IndRefind,nargs)->letmind=Environ.lookup_mind(fstind)envincstrs:=cmp_inductivescv_pb(mind,sndind)nargsll'!cstrs;true|Some(ConstructRef((mi,ind),ctor),nargs)->letmind=Environ.lookup_mindmienvincstrs:=cmp_constructors(mind,ind,ctor)nargsll'!cstrs;truelettest_constr_universesenvsigmaleqmn=letopenUnivProbleminletkindc=kindsigmacinifm==nthenSomeSet.emptyelseletcstrs=refSet.emptyinletcv_pb=ifleqthenReduction.CUMULelseReduction.CONVinleteq_universesrefargsll'=eq_universesenvsigmacstrsReduction.CONVrefargsll'andleq_universesrefargsll'=eq_universesenvsigmacstrscv_pbrefargsll'inleteq_sortss1s2=lets1=ESorts.kindsigmas1inlets2=ESorts.kindsigmas2inifSorts.equals1s2thentrueelse(cstrs:=Set.add(UEq(Sorts.univ_of_sorts1,Sorts.univ_of_sorts2))!cstrs;true)inletleq_sortss1s2=lets1=ESorts.kindsigmas1inlets2=ESorts.kindsigmas2inifSorts.equals1s2thentrueelse(cstrs:=Set.add(ULe(Sorts.univ_of_sorts1,Sorts.univ_of_sorts2))!cstrs;true)inletreceq_constr'nargsmn=compare_genkindeq_universeseq_sortseq_constr'nargsmninletres=ifleqthenletreccompare_leqnargsmn=Constr.compare_head_gen_leq_withkindkindleq_universesleq_sortseq_constr'leq_constr'nargsmnandleq_constr'nargsmn=m==n||compare_leqnargsmnincompare_leq0mnelseConstr.compare_head_gen_withkindkindeq_universeseq_sortseq_constr'0mninifresthenSome!cstrselseNoneleteq_constr_universesenvsigmamn=test_constr_universesenvsigmafalsemnletleq_constr_universesenvsigmamn=test_constr_universesenvsigmatruemnletcompare_head_gen_projenvsigmaequeqseqc'nargsmn=letkindc=kindsigmacinmatchkindm,kindnwith|Proj(p,c),App(f,args)|App(f,args),Proj(p,c)->(matchkindfwith|Const(p',u)whenEnviron.QConstant.equalenv(Projection.constantp)p'->letnpars=Projection.nparspinifArray.lengthargs==npars+1theneqc'0cargs.(npars)elsefalse|_->false)|_->Constr.compare_head_gen_withkindkindequeqseqc'nargsmnleteq_constr_universes_projenvsigmamn=letopenUnivProbleminifm==nthenSomeSet.emptyelseletcstrs=refSet.emptyinleteq_universesrefll'=eq_universesenvsigmacstrsReduction.CONVrefll'inleteq_sortss1s2=lets1=ESorts.kindsigmas1inlets2=ESorts.kindsigmas2inifSorts.equals1s2thentrueelse(cstrs:=Set.add(UEq(Sorts.univ_of_sorts1,Sorts.univ_of_sorts2))!cstrs;true)inletreceq_constr'nargsmn=m==n||compare_head_gen_projenvsigmaeq_universeseq_sortseq_constr'nargsmninletres=eq_constr'0mninifresthenSome!cstrselseNoneletuniverses_of_constrsigmac=letopenUnivinletrecauxsc=matchkindsigmacwith|Const(c,u)->LSet.foldLSet.add(Instance.levels(EInstance.kindsigmau))s|Ind((mind,_),u)|Construct(((mind,_),_),u)->LSet.foldLSet.add(Instance.levels(EInstance.kindsigmau))s|Sortu->letsort=ESorts.kindsigmauinifSorts.is_smallsortthenselseletu=Sorts.univ_of_sortsortinLSet.foldLSet.add(Universe.levelsu)s|Evar(k,args)->letconcl=Evd.evar_concl(Evd.findsigmak)infoldsigmaaux(auxsconcl)c|Array(u,_,_,_)->lets=LSet.foldLSet.add(Instance.levels(EInstance.kindsigmau))sinfoldsigmaauxsc|Case(_,u,_,_,_,_,_)->lets=LSet.foldLSet.add(Instance.levels(EInstance.kindsigmau))sinfoldsigmaauxsc|_->foldsigmaauxscinauxLSet.emptycopenContextopenEnvironletcast_list:typeab.(a,b)eq->alist->blist=funReflx->xletcast_list_snd:typeab.(a,b)eq->('c*a)list->('c*b)list=funReflx->xletcast_rel_decl:typeab.(a,b)eq->(a,a)Rel.Declaration.pt->(b,b)Rel.Declaration.pt=funReflx->xletcast_rel_context:typeab.(a,b)eq->(a,a)Rel.pt->(b,b)Rel.pt=funReflx->xletcast_rec_decl:typeab.(a,b)eq->(a,a)Constr.prec_declaration->(b,b)Constr.prec_declaration=funReflx->xletcast_named_decl:typeab.(a,b)eq->(a,a)Named.Declaration.pt->(b,b)Named.Declaration.pt=funReflx->xletcast_named_context:typeab.(a,b)eq->(a,a)Named.pt->(b,b)Named.pt=funReflx->xmoduleVars=structexceptionLocalOccurletto_constr=unsafe_to_constrletto_rel_decl=unsafe_to_rel_decltypesubstl=tlist(** Operations that commute with evar-normalization *)letlift=liftletliftnnmc=of_constr(Vars.liftnnm(to_constrc))letsubstnlsubstnc=of_constr(Vars.substnl(cast_listunsafe_eqsubst)n(to_constrc))letsubstlsubstc=of_constr(Vars.substl(cast_listunsafe_eqsubst)(to_constrc))letsubst1cr=of_constr(Vars.subst1(to_constrc)(to_constrr))letsubstnl_declsubstnd=of_rel_decl(Vars.substnl_decl(cast_listunsafe_eqsubst)n(to_rel_decld))letsubstl_declsubstd=of_rel_decl(Vars.substl_decl(cast_listunsafe_eqsubst)(to_rel_decld))letsubst1_declcd=of_rel_decl(Vars.subst1_decl(to_constrc)(to_rel_decld))letreplace_varssubstc=of_constr(Vars.replace_vars(cast_list_sndunsafe_eqsubst)(to_constrc))letsubstn_varsnsubstc=of_constr(Vars.substn_varsnsubst(to_constrc))letsubst_varssubstc=of_constr(Vars.subst_varssubst(to_constrc))letsubst_varsubstc=of_constr(Vars.subst_varsubst(to_constrc))letsubst_univs_level_constrsubstc=of_constr(Vars.subst_univs_level_constrsubst(to_constrc))(** Operations that dot NOT commute with evar-normalization *)letnoccurnsigmanterm=letrecoccur_recnc=matchkindsigmacwith|Relm->ifInt.equalmnthenraiseLocalOccur|_->iter_with_binderssigmasuccoccur_recncintryoccur_recnterm;truewithLocalOccur->falseletnoccur_betweensigmanmterm=letrecoccur_recnc=matchkindsigmacwith|Relp->ifn<=p&&p<n+mthenraiseLocalOccur|_->iter_with_binderssigmasuccoccur_recncintryoccur_recnterm;truewithLocalOccur->falseletclosednsigmanc=letrecclosed_recnc=matchkindsigmacwith|Relm->ifm>nthenraiseLocalOccur|_->iter_with_binderssigmasuccclosed_recncintryclosed_recnc;truewithLocalOccur->falseletclosed0sigmac=closednsigma0cletsubst_of_rel_context_instancectxsubst=cast_list(symunsafe_eq)(Vars.subst_of_rel_context_instance(cast_rel_contextunsafe_eqctx)(cast_listunsafe_eqsubst))endletrecisAritysigmac=matchkindsigmacwith|Prod(_,_,c)->isAritysigmac|LetIn(_,b,_,c)->isAritysigma(Vars.subst1bc)|Cast(c,_,_)->isAritysigmac|Sort_->true|_->falsetypearity=rel_context*ESorts.tletdestAritysigma=letopenContext.Rel.Declarationinletrecprodec_reclc=matchkindsigmacwith|Prod(x,t,c)->prodec_rec(LocalAssum(x,t)::l)c|LetIn(x,b,t,c)->prodec_rec(LocalDef(x,b,t)::l)c|Cast(c,_,_)->prodec_reclc|Sorts->l,s|_->anomaly~label:"destArity"(Pp.str"not an arity.")inprodec_rec[]letmkProd_or_LetIndeclc=letopenContext.Rel.Declarationinmatchdeclwith|LocalAssum(na,t)->mkProd(na,t,c)|LocalDef(na,b,t)->mkLetIn(na,b,t,c)letmkLambda_or_LetIndeclc=letopenContext.Rel.Declarationinmatchdeclwith|LocalAssum(na,t)->mkLambda(na,t,c)|LocalDef(na,b,t)->mkLetIn(na,b,t,c)letmkNamedProdidtypc=mkProd(map_annotName.mk_nameid,typ,Vars.subst_varid.binder_namec)letmkNamedLambdaidtypc=mkLambda(map_annotName.mk_nameid,typ,Vars.subst_varid.binder_namec)letmkNamedLetInidc1tc2=mkLetIn(map_annotName.mk_nameid,c1,t,Vars.subst_varid.binder_namec2)letmkNamedProd_or_LetIndeclc=letopenContext.Named.Declarationinmatchdeclwith|LocalAssum(id,t)->mkNamedProdidtc|LocalDef(id,b,t)->mkNamedLetInidbtcletmkNamedLambda_or_LetIndeclc=letopenContext.Named.Declarationinmatchdeclwith|LocalAssum(id,t)->mkNamedLambdaidtc|LocalDef(id,b,t)->mkNamedLetInidbtcletit_mkProd_or_LetIntctx=List.fold_left(funcd->mkProd_or_LetIndc)tctxletit_mkLambda_or_LetIntctx=List.fold_left(funcd->mkLambda_or_LetIndc)tctxletpush_relde=push_rel(cast_rel_declunsafe_eqd)eletpush_rel_contextde=push_rel_context(cast_rel_contextunsafe_eqd)eletpush_rec_typesde=push_rec_types(cast_rec_declunsafe_eqd)eletpush_namedde=push_named(cast_named_declunsafe_eqd)eletpush_named_contextde=push_named_context(cast_named_contextunsafe_eqd)eletpush_named_context_valde=push_named_context_val(cast_named_declunsafe_eqd)eletrel_contexte=cast_rel_context(symunsafe_eq)(rel_contexte)letnamed_contexte=cast_named_context(symunsafe_eq)(named_contexte)letval_of_named_contexte=val_of_named_context(cast_named_contextunsafe_eqe)letnamed_context_of_vale=cast_named_context(symunsafe_eq)(named_context_of_vale)letof_existential:Constr.existential->existential=letgen:typeab.(a,b)eq->'c*blist->'c*alist=funReflx->xingenunsafe_eqletlookup_relie=cast_rel_decl(symunsafe_eq)(lookup_relie)letlookup_namedne=cast_named_decl(symunsafe_eq)(lookup_namedne)letlookup_named_valne=cast_named_decl(symunsafe_eq)(lookup_named_ctxtne)letmap_rel_context_in_envfenvsign=letrecauxenvacc=function|d::sign->aux(push_reldenv)(Context.Rel.Declaration.map_constr(fenv)d::acc)sign|[]->accinauxenv[](List.revsign)letmatch_named_context_val:named_context_val->(named_declaration*lazy_val*named_context_val)option=matchunsafe_eqwith|Refl->match_named_context_valletidentity_subst_val:named_context_val->tlist=matchunsafe_eqwithRefl->functx->ctx.env_named_varletfresh_global?loc?rigid?namesenvsigmareference=let(evd,t)=Evd.fresh_global?loc?rigid?namesenvsigmareferenceinevd,tletis_global=isRefX(** Kind of type *)typekind_of_type=|SortTypeofESorts.t|CastTypeoftypes*t|ProdTypeofName.tContext.binder_annot*t*t|LetInTypeofName.tContext.binder_annot*t*t*t|AtomicTypeoft*tarrayletkind_of_typesigmat=matchkindsigmatwith|Sorts->SortTypes|Cast(c,_,t)->CastType(c,t)|Prod(na,t,c)->ProdType(na,t,c)|LetIn(na,b,t,c)->LetInType(na,b,t,c)|App(c,l)->AtomicType(c,l)|(Rel_|Meta_|Var_|Evar_|Const_|Proj_|Case_|Fix_|CoFix_|Ind_)->AtomicType(t,[||])|(Lambda_|Construct_|Int_|Float_|Array_)->failwith"Not a type"moduleUnsafe=structletto_sorts=ESorts.unsafe_to_sortsletto_instance=EInstance.unsafe_to_instanceletto_constr=unsafe_to_constrletto_constr_array=unsafe_to_constr_arrayletto_rel_decl=unsafe_to_rel_declletto_named_decl=unsafe_to_named_declletto_named_context=letgen:typeab.(a,b)eq->(a,a)Context.Named.pt->(b,b)Context.Named.pt=funReflx->xingenunsafe_eqletto_case_invert=unsafe_to_case_invertleteq=unsafe_eqend