123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691692693694695696697698699700701702703704705706707708709710711712713714715716717718719720721722723724725726727728729730731732733734735736737738739740741742743744745746747748749750751752753754755756757758759760761762763764765766767768769770771772773774775776777778779780781782783784785786787788789790791792793794795796797798799800801802803804805806807808809810811812813814815816817818819820821822823824825826827828829830831832833834835836837838839840841842843844845846847848849850851852853854855856857858859860861862863864865866867868869870871872873874875876877878879880881882883884885886887888889890891892893894895896897898899900901902903904905906907908909910911912913914915916917918919920921922923924925926927928929930931932933934935936937938939940941942943944945946947948949950951952953954955956957958959960961962963964965966967968969970971972973974975976977978979980981982983984985986987988989990991992993994995996997998999100010011002100310041005100610071008100910101011101210131014101510161017101810191020102110221023102410251026102710281029103010311032103310341035103610371038103910401041104210431044104510461047104810491050105110521053105410551056105710581059106010611062106310641065106610671068106910701071107210731074107510761077107810791080108110821083108410851086108710881089109010911092109310941095109610971098109911001101110211031104110511061107110811091110111111121113111411151116111711181119112011211122112311241125112611271128112911301131113211331134113511361137113811391140114111421143114411451146114711481149115011511152115311541155115611571158115911601161116211631164116511661167116811691170117111721173117411751176117711781179118011811182118311841185118611871188118911901191119211931194119511961197119811991200120112021203120412051206120712081209121012111212121312141215121612171218121912201221122212231224122512261227122812291230123112321233123412351236123712381239124012411242124312441245124612471248124912501251125212531254125512561257125812591260126112621263126412651266126712681269127012711272127312741275127612771278127912801281128212831284128512861287128812891290129112921293129412951296129712981299130013011302130313041305130613071308130913101311131213131314131513161317131813191320132113221323132413251326132713281329133013311332133313341335133613371338133913401341134213431344134513461347134813491350135113521353135413551356135713581359136013611362136313641365136613671368136913701371137213731374137513761377137813791380138113821383138413851386138713881389139013911392139313941395139613971398139914001401140214031404140514061407140814091410141114121413141414151416141714181419142014211422142314241425142614271428142914301431143214331434143514361437143814391440144114421443144414451446144714481449145014511452145314541455145614571458145914601461146214631464146514661467146814691470147114721473147414751476147714781479148014811482148314841485148614871488148914901491149214931494149514961497149814991500150115021503150415051506150715081509151015111512151315141515151615171518151915201521152215231524152515261527152815291530153115321533153415351536153715381539154015411542154315441545154615471548154915501551155215531554155515561557155815591560156115621563156415651566156715681569157015711572157315741575157615771578157915801581(************************************************************************)(* * 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) *)(************************************************************************)openPpopenCErrorsopenSortsopenUtilopenNamesopenNameopsopenConstropenVarsopenEnviron(* module RelDecl = Context.Rel.Declaration *)moduleNamedDecl=Context.Named.Declarationtypeeconstr=constrtypeetypes=types(** Generic filters *)moduleFilter:sigtypetvalequal:t->t->boolvalidentity:tvalfilter_list:t->'alist->'alistvalfilter_array:t->'aarray->'aarrayvalextend:int->t->tvalcompose:t->t->tvalapply_subfilter:t->boollist->tvalrestrict_upon:t->int->(int->bool)->toptionvalmap_along:(bool->'a->bool)->t->'alist->tvalmake:boollist->tvalrepr:t->boollistoptionend=structtypet=boollistoption(** We guarantee through the interface that if a filter is [Some _] then it
contains at least one [false] somewhere. *)letidentity=Noneletrecequall1l2=matchl1,l2with|[],[]->true|h1::l1,h2::l2->(ifh1thenh2elsenoth2)&&equall1l2|_->falseletequall1l2=matchl1,l2with|None,None->true|Some_,None|None,Some_->false|Somel1,Somel2->equall1l2letrecis_identity=function|[]->true|true::l->is_identityl|false::_->falseletnormalizef=ifis_identityfthenNoneelseSomefletfilter_listfl=matchfwith|None->l|Somef->CList.filter_withflletfilter_arrayfv=matchfwith|None->v|Somef->CArray.filter_withfvletrecextendnl=ifn=0thenlelseextend(predn)(true::l)letextendn=function|None->None|Somef->Some(extendnf)letcomposef1f2=matchf1with|None->f2|Somef1->matchf2with|None->None|Somef2->normalize(CList.filter_withf1f2)letapply_subfilter_arrayfiltersubfilter=(* In both cases we statically know that the argument will contain at
least one [false] *)matchfilterwith|None->Some(Array.to_listsubfilter)|Somef->letlen=Array.lengthsubfilterinletfoldb(i,ans)=ifbthenlet()=assert(0<=i)in(predi,Array.unsafe_getsubfilteri::ans)else(i,false::ans)inSome(snd(List.fold_rightfoldf(predlen,[])))letapply_subfilterfiltersubfilter=apply_subfilter_arrayfilter(Array.of_listsubfilter)letrestrict_uponflenp=letnewfilter=Array.initlenpinifArray.for_all(funid->id)newfilterthenNoneelseSome(apply_subfilter_arrayfnewfilter)letmap_alongffltl=letans=matchfltwith|None->List.map(funx->ftruex)l|Someflt->List.map2ffltlinnormalizeansletmakel=normalizelletreprf=fendmoduleAbstraction=structtypeabstraction=|Abstract|Imitatetypet=abstractionlistletidentity=[]letabstract_lastl=Abstract::lendmoduleIdentity:sigtypetvalmake:econstrlist->tvalnone:unit->tvalrepr:named_context_val->Filter.t->t->econstrlistvalis_identity:econstrlist->t->boolend=structtypet=econstrlistoptionrefletmakes=ref(Somes)letnone()=refNoneletreprsignfilters=match!swith|None->letans=Filter.filter_listfiltersign.env_named_varinlet()=s:=Someansinans|Somes->sletis_identityls=match!swith|None->false|Somes->s==lend(* The kinds of existential variables are now defined in [Evar_kinds] *)(* The type of mappings for existential variables *)moduleStore=Store.Make()letstring_of_existentialevk="?X"^string_of_int(Evar.reprevk)typeevar_body=|Evar_empty|Evar_definedofconstrtypeevar_info={evar_concl:constr;evar_hyps:named_context_val;evar_body:evar_body;evar_filter:Filter.t;evar_abstract_arguments:Abstraction.t;evar_source:Evar_kinds.tLoc.located;evar_candidates:constrlistoption;(* if not None, list of allowed instances *)evar_identity:Identity.t;}letmake_evarhypsccl={evar_concl=ccl;evar_hyps=hyps;evar_body=Evar_empty;evar_filter=Filter.identity;evar_abstract_arguments=Abstraction.identity;evar_source=Loc.tag@@Evar_kinds.InternalHole;evar_candidates=None;evar_identity=Identity.none();}letinstance_mismatch()=anomaly(Pp.str"Signature and its instance do not match.")letevar_conclevi=evi.evar_conclletevar_filterevi=evi.evar_filterletevar_bodyevi=evi.evar_bodyletevar_contextevi=named_context_of_valevi.evar_hypsletevar_filtered_contextevi=Filter.filter_list(evar_filterevi)(evar_contextevi)letevar_candidatesevi=evi.evar_candidatesletevar_hypsevi=evi.evar_hypsletevar_filtered_hypsevi=matchFilter.repr(evar_filterevi)with|None->evar_hypsevi|Somefilter->letrecmake_hypsfilterctxt=matchfilter,ctxtwith|[],[]->empty_named_context_val|false::filter,_::ctxt->make_hypsfilterctxt|true::filter,decl::ctxt->lethyps=make_hypsfilterctxtinpush_named_context_valdeclhyps|_->instance_mismatch()inmake_hypsfilter(evar_contextevi)letevar_envenvevi=Environ.reset_with_named_contextevi.evar_hypsenvletevar_filtered_envenvevi=matchFilter.repr(evar_filterevi)with|None->evar_envenvevi|Somefilter->letrecmake_envfilterctxt=matchfilter,ctxtwith|[],[]->reset_contextenv|false::filter,_::ctxt->make_envfilterctxt|true::filter,decl::ctxt->letenv=make_envfilterctxtinpush_nameddeclenv|_->instance_mismatch()inmake_envfilter(evar_contextevi)letevar_identity_substevi=Identity.reprevi.evar_hypsevi.evar_filterevi.evar_identityletmap_evar_bodyf=function|Evar_empty->Evar_empty|Evar_definedd->Evar_defined(fd)letmap_evar_infofevi={eviwithevar_body=map_evar_bodyfevi.evar_body;evar_hyps=map_named_val(fund->NamedDecl.map_constrfd)evi.evar_hyps;evar_concl=fevi.evar_concl;evar_candidates=Option.map(List.mapf)evi.evar_candidates}(* This exception is raised by *.existential_value *)exceptionNotInstantiatedEvar(* Note: let-in contributes to the instance *)letevar_instance_arraytest_idinfoargs=letrecinstrecfilterctxtargs=matchfilter,ctxt,argswith|[],[],[]->[]|false::filter,_::ctxt,args->instrecfilterctxtargs|true::filter,d::ctxt,c::args->iftest_iddctheninstrecfilterctxtargselse(NamedDecl.get_idd,c)::instrecfilterctxtargs|_->instance_mismatch()inmatchFilter.repr(evar_filterinfo)with|None->letrecinstancectxtargs=matchctxt,argswith|[],[]->[]|d::ctxt,c::args->iftest_iddctheninstancectxtargselse(NamedDecl.get_idd,c)::instancectxtargs|_->instance_mismatch()ininstance(evar_contextinfo)args|Somefilter->instrecfilter(evar_contextinfo)argsletmake_evar_instance_arrayinfoargs=ifIdentity.is_identityargsinfo.evar_identitythen[]elseevar_instance_array(NamedDecl.get_id%>isVarId)infoargsletinstantiate_evar_arrayinfocargs=letinst=make_evar_instance_arrayinfoargsinmatchinstwith|[]->c|_->replace_varsinstctype'ain_evar_universe_context='a*UState.t(*******************************************************************)(* Metamaps *)(*******************************************************************)(* Constraints for existential variables *)(*******************************************************************)type'afreelisted={rebus:'a;freemetas:Int.Set.t}(* Collects all metavars appearing in a constr *)letmetavars_ofc=letreccollrecaccc=matchkindcwith|Metamv->Int.Set.addmvacc|_->Constr.foldcollrecacccincollrecInt.Set.emptycletmk_freelistedc={rebus=c;freemetas=metavars_ofc}letmap_flfcfl={cflwithrebus=fcfl.rebus}(* Status of an instance found by unification wrt to the meta it solves:
- a supertype of the meta (e.g. the solution to ?X <= T is a supertype of ?X)
- a subtype of the meta (e.g. the solution to T <= ?X is a supertype of ?X)
- a term that can be eta-expanded n times while still being a solution
(e.g. the solution [P] to [?X u v = P u v] can be eta-expanded twice)
*)typeinstance_constraint=IsSuperType|IsSubType|Convleteq_instance_constraintc1c2=c1==c2(* Status of the unification of the type of an instance against the type of
the meta it instantiates:
- CoerceToType means that the unification of types has not been done
and that a coercion can still be inserted: the meta should not be
substituted freely (this happens for instance given via the
"with" binding clause).
- TypeProcessed means that the information obtainable from the
unification of types has been extracted.
- TypeNotProcessed means that the unification of types has not been
done but it is known that no coercion may be inserted: the meta
can be substituted freely.
*)typeinstance_typing_status=CoerceToType|TypeNotProcessed|TypeProcessed(* Status of an instance together with the status of its type unification *)typeinstance_status=instance_constraint*instance_typing_status(* Clausal environments *)typeclbinding=|CltypofName.t*constrfreelisted|ClvalofName.t*(constrfreelisted*instance_status)*constrfreelistedletmap_clbf=function|Cltyp(na,cfl)->Cltyp(na,map_flfcfl)|Clval(na,(cfl1,pb),cfl2)->Clval(na,(map_flfcfl1,pb),map_flfcfl2)(* name of defined is erased (but it is pretty-printed) *)letclb_name=functionCltyp(na,_)->(na,false)|Clval(na,_,_)->(na,true)(***********************)moduleMetaset=Int.SetmoduleMetamap=Int.Mapletmetamap_to_listm=Metamap.fold(funnvl->(n,v)::l)m[](*************************)(* Unification state *)typeconv_pb=Reduction.conv_pbtypeevar_constraint=conv_pb*Environ.env*constr*constrmoduleEvMap=Evar.MapmoduleEvNames:sigtypetvalempty:tvaladd_name_undefined:Id.toption->Evar.t->evar_info->t->tvalremove_name_defined:Evar.t->t->tvalrename:Evar.t->Id.t->t->tvalreassign_name_defined:Evar.t->Evar.t->t->tvalident:Evar.t->t->Id.toptionvalkey:Id.t->t->Evar.tend=structtypet=Id.tEvMap.t*Evar.tId.Map.tletempty=(EvMap.empty,Id.Map.empty)letadd_name_newly_undefinedidevkevi(evtoid,idtoevasnames)=matchidwith|None->names|Someid->ifId.Map.memididtoevthenuser_err(str"Already an existential evar of name "++Id.printid);(EvMap.addevkidevtoid,Id.Map.addidevkidtoev)letadd_name_undefinednamingevkevi(evtoid,idtoevasevar_names)=ifEvMap.memevkevtoidthenevar_nameselseadd_name_newly_undefinednamingevkevievar_namesletremove_name_definedevk(evtoid,idtoevasnames)=letid=trySome(EvMap.findevkevtoid)withNot_found->Noneinmatchidwith|None->names|Someid->(EvMap.removeevkevtoid,Id.Map.removeididtoev)letrenameevkid(evtoid,idtoev)=letid'=trySome(EvMap.findevkevtoid)withNot_found->Noneinmatchid'with|None->(EvMap.addevkidevtoid,Id.Map.addidevkidtoev)|Someid'->ifId.Map.memididtoevthenanomaly(str"Evar name already in use.");(EvMap.setevkidevtoid(* overwrite old name *),Id.Map.addidevk(Id.Map.removeid'idtoev))letreassign_name_definedevkevk'(evtoid,idtoevasnames)=letid=trySome(EvMap.findevkevtoid)withNot_found->Noneinmatchidwith|None->names(* evk' must not be defined *)|Someid->(EvMap.addevk'id(EvMap.removeevkevtoid),Id.Map.addidevk'(Id.Map.removeididtoev))letidentevk(evtoid,_)=trySome(EvMap.findevkevtoid)withNot_found->Noneletkeyid(_,idtoev)=Id.Map.findididtoevendtypeevar_flags={obligation_evars:Evar.Set.t;aliased_evars:Evar.tEvar.Map.t;typeclass_evars:Evar.Set.t}typeside_effect_role=|Schemaofinductive*stringtypeside_effects={seff_private:Safe_typing.private_constants;seff_roles:side_effect_roleCmap.t;}moduleFutureGoals:sigtypet=private{comb:Evar.tlist;principal:Evar.toption;(** if [Some e], [e] must be
contained in
[comb]. The evar
[e] will inherit
properties (now: the
name) of the evar which
will be instantiated with
a term containing [e]. *)}valmap_filter:(Evar.t->Evar.toption)->t->t(** Applies a function on the future goals *)valfilter:(Evar.t->bool)->t->t(** Applies a filter on the future goals *)typestackvalempty_stack:stackvalpush:stack->stackvalpop:stack->t*stackvaladd:principal:bool->Evar.t->stack->stackvalremove:Evar.t->stack->stackvalfold:('a->Evar.t->'a)->'a->stack->'avalpr_stack:stack->Pp.tend=structtypet={comb:Evar.tlist;principal:Evar.toption;(** if [Some e], [e] must be
contained in
[comb]. The evar
[e] will inherit
properties (now: the
name) of the evar which
will be instantiated with
a term containing [e]. *)}typestack=tlistletsetf=function|[]->anomalyPp.(str"future_goals stack should not be empty")|hd::tl->fhd::tlletadd~principalevkstack=letaddfgl=letcomb=evk::fgl.combinletprincipal=ifprincipalthenmatchfgl.principalwith|Some_->CErrors.user_errPp.(str"Only one main goal per instantiation.")|None->Someevkelsefgl.principalin{comb;principal}insetaddstackletremoveestack=letremovefgl=letfiltere'=not(Evar.equalee')inletprincipal=Option.filterfilterfgl.principalinletcomb=List.filterfilterfgl.combin{principal;comb}inList.mapremovestackletempty={principal=None;comb=[];}letempty_stack=[empty]letpushstack=empty::stackletpopstack=matchstackwith|[]->anomalyPp.(str"future_goals stack should not be empty")|hd::tl->hd,tlletfoldfaccstack=letfuture_goals=List.hdstackinList.fold_leftfaccfuture_goals.combletfilterffgl=letcomb=List.filterffgl.combinletprincipal=Option.filterffgl.principalin{comb;principal}letmap_filterffgl=letcomb=List.map_filterffgl.combinletprincipal=Option.bindfgl.principalfin{comb;principal}letpr_stackstack=letopenPpinletpr_future_goalsfgl=prlist_with_sepspcEvar.printfgl.comb++pr_opt(funev->str"(principal: "++Evar.printev++str")")fgl.principalinifList.is_emptystackthenstr"(empty stack)"elseprlist_with_sep(fun()->str"||")pr_future_goalsstackendtypeevar_map={(* Existential variables *)defn_evars:evar_infoEvMap.t;undf_evars:evar_infoEvMap.t;evar_names:EvNames.t;(** Universes *)universes:UState.t;(** Conversion problems *)conv_pbs:evar_constraintlist;last_mods:Evar.Set.t;(** Metas *)metas:clbindingMetamap.t;evar_flags:evar_flags;(** Interactive proofs *)effects:side_effects;future_goals:FutureGoals.stack;(** list of newly created evars, to be
eventually turned into goals if not solved.*)given_up:Evar.Set.t;shelf:Evar.tlistlist;extras:Store.t;}letget_is_maybe_typeclass,(is_maybe_typeclass_hook:(evar_map->constr->bool)Hook.t)=Hook.make~default:(funevdc->false)()letis_maybe_typeclasssigmac=Hook.getget_is_maybe_typeclasssigmac(*** Lifting primitive from Evar.Map. ***)letrenameevkidevd={evdwithevar_names=EvNames.renameevkidevd.evar_names}letadd_with_name?name?(typeclass_candidate=true)dei=matchi.evar_bodywith|Evar_empty->letevar_names=EvNames.add_name_undefinednameeid.evar_namesinletevar_flags=iftypeclass_candidate&&is_maybe_typeclassdi.evar_conclthenletflags=d.evar_flagsin{flagswithtypeclass_evars=Evar.Set.addeflags.typeclass_evars}elsed.evar_flagsin{dwithundf_evars=EvMap.addeid.undf_evars;evar_names;evar_flags}|Evar_defined_->letevar_names=EvNames.remove_name_defineded.evar_namesin{dwithdefn_evars=EvMap.addeid.defn_evars;evar_names}(** Evd.add is a low-level function mainly used to update the evar_info
associated to an evar, so we prevent registering its typeclass status. *)letadddei=add_with_name~typeclass_candidate:falsedei(*** Evar flags: typeclasses, aliased or obligation flag *)letget_typeclass_evarsevd=evd.evar_flags.typeclass_evarsletset_typeclass_evarsevdtcs=letflags=evd.evar_flagsin{evdwithevar_flags={flagswithtypeclass_evars=tcs}}letis_typeclass_evarevdevk=letflags=evd.evar_flagsinEvar.Set.memevkflags.typeclass_evarsletget_obligation_evarsevd=evd.evar_flags.obligation_evarsletset_obligation_evarevdevk=letflags=evd.evar_flagsinletevar_flags={flagswithobligation_evars=Evar.Set.addevkflags.obligation_evars}in{evdwithevar_flags}letis_obligation_evarevdevk=letflags=evd.evar_flagsinEvar.Set.memevkflags.obligation_evars(** Inheritance of flags: for evar-evar and restriction cases *)letinherit_evar_flagsevar_flagsevkevk'=letevk_typeclass=Evar.Set.memevkevar_flags.typeclass_evarsinletevk_obligation=Evar.Set.memevkevar_flags.obligation_evarsinletaliased_evars=Evar.Map.addevkevk'evar_flags.aliased_evarsinlettypeclass_evars=ifevk_typeclassthenlettypeclass_evars=Evar.Set.removeevkevar_flags.typeclass_evarsinEvar.Set.addevk'typeclass_evarselseevar_flags.typeclass_evarsinletobligation_evars=ifevk_obligationthenletobligation_evars=Evar.Set.removeevkevar_flags.obligation_evarsinEvar.Set.addevk'obligation_evarselseevar_flags.obligation_evarsin{obligation_evars;aliased_evars;typeclass_evars}(** Removal: in all other cases of definition *)letremove_evar_flagsevkevar_flags={typeclass_evars=Evar.Set.removeevkevar_flags.typeclass_evars;obligation_evars=Evar.Set.removeevkevar_flags.obligation_evars;(* Aliasing information is kept. *)aliased_evars=evar_flags.aliased_evars}(** New evars *)letevar_counter_summary_name="evar counter"(* Generator of existential names *)letevar_ctr,evar_counter_summary_tag=Summary.ref_tag0~name:evar_counter_summary_nameletnew_untyped_evar()=increvar_ctr;Evar.unsafe_of_int!evar_ctrletnew_evarevd?name?typeclass_candidateevi=letevk=new_untyped_evar()inletevd=add_with_nameevd?name?typeclass_candidateevkeviin(evd,evk)letremovede=letundf_evars=EvMap.removeed.undf_evarsinletdefn_evars=EvMap.removeed.defn_evarsinletfuture_goals=FutureGoals.removeed.future_goalsinletevar_flags=remove_evar_flagsed.evar_flagsin{dwithundf_evars;defn_evars;future_goals;evar_flags}letfindde=tryEvMap.finded.undf_evarswithNot_found->EvMap.finded.defn_evarsletfind_undefinedde=EvMap.finded.undf_evarsletmemde=EvMap.memed.undf_evars||EvMap.memed.defn_evarsletundefined_mapd=d.undf_evarsletdrop_all_definedd={dwithdefn_evars=EvMap.empty}(* spiwack: not clear what folding over an evar_map, for now we shall
simply fold over the inner evar_map. *)letfoldfda=EvMap.foldfd.defn_evars(EvMap.foldfd.undf_evarsa)letfold_undefinedfda=EvMap.foldfd.undf_evarsaletraw_mapfd=letfevkinfo=letans=fevkinfoinlet()=matchinfo.evar_body,ans.evar_bodywith|Evar_defined_,Evar_empty|Evar_empty,Evar_defined_->anomaly(str"Unrespectful mapping function.")|_->()inansinletdefn_evars=EvMap.Smart.mapifd.defn_evarsinletundf_evars=EvMap.Smart.mapifd.undf_evarsin{dwithdefn_evars;undf_evars;}letraw_map_undefinedfd=letfevkinfo=letans=fevkinfoinlet()=matchans.evar_bodywith|Evar_defined_->anomaly(str"Unrespectful mapping function.")|_->()inansin{dwithundf_evars=EvMap.Smart.mapifd.undf_evars;}letis_evar=memletis_definedde=EvMap.memed.defn_evarsletis_undefinedde=EvMap.memed.undf_evarsletexistential_opt_valued(n,args)=matchEvMap.find_optnd.defn_evarswith|None->None|Someinfo->matchevar_bodyinfowith|Evar_definedc->Some(instantiate_evar_arrayinfocargs)|Evar_empty->None(* impossible but w/e *)letexistential_valuedev=matchexistential_opt_valuedevwith|None->raiseNotInstantiatedEvar|Somev->vletexistential_value0=existential_valueletexistential_opt_value0=existential_opt_valueletexistential_typed(n,args)=letinfo=tryfinddnwithNot_found->anomaly(str"Evar "++str(string_of_existentialn)++str" was not declared.")ininstantiate_evar_arrayinfoinfo.evar_conclargsletexistential_type0=existential_typeletadd_constraintsdc={dwithuniverses=UState.add_constraintsd.universesc}letadd_universe_constraintsdc={dwithuniverses=UState.add_universe_constraintsd.universesc}(*** /Lifting... ***)(* evar_map are considered empty disregarding histories *)letis_emptyd=EvMap.is_emptyd.defn_evars&&EvMap.is_emptyd.undf_evars&&List.is_emptyd.conv_pbs&&Metamap.is_emptyd.metasletcmapfevd={evdwithmetas=Metamap.map(map_clbf)evd.metas;defn_evars=EvMap.map(map_evar_infof)evd.defn_evars;undf_evars=EvMap.map(map_evar_infof)evd.undf_evars}(* spiwack: deprecated *)letcreate_evar_defssigma={sigmawithconv_pbs=[];last_mods=Evar.Set.empty;metas=Metamap.empty}letempty_evar_flags={obligation_evars=Evar.Set.empty;aliased_evars=Evar.Map.empty;typeclass_evars=Evar.Set.empty}letempty_side_effects={seff_private=Safe_typing.empty_private_constants;seff_roles=Cmap.empty;}letempty={defn_evars=EvMap.empty;undf_evars=EvMap.empty;universes=UState.empty;conv_pbs=[];last_mods=Evar.Set.empty;evar_flags=empty_evar_flags;metas=Metamap.empty;effects=empty_side_effects;evar_names=EvNames.empty;(* id<->key for undefined evars *)future_goals=FutureGoals.empty_stack;given_up=Evar.Set.empty;shelf=[[]];extras=Store.empty;}letfrom_env?binderse={emptywithuniverses=UState.from_env?binderse}letfrom_ctxuctx={emptywithuniverses=uctx}lethas_undefinedevd=not(EvMap.is_emptyevd.undf_evars)lethas_given_upevd=not(Evar.Set.is_emptyevd.given_up)lethas_shelvedevd=not(List.for_allList.is_emptyevd.shelf)letevars_reset_evd?(with_conv_pbs=false)?(with_univs=true)evdd=letconv_pbs=ifwith_conv_pbsthenevd.conv_pbselsed.conv_pbsinletlast_mods=ifwith_conv_pbsthenevd.last_modselsed.last_modsinletuniverses=ifnotwith_univsthenevd.universeselseUState.unionevd.universesd.universesin{evdwithmetas=d.metas;last_mods;conv_pbs;universes}letmerge_universe_contextevductx'={evdwithuniverses=UState.unionevd.universesuctx'}letset_universe_contextevductx'={evdwithuniverses=uctx'}(* TODO: make unique *)letadd_conv_pb?(tail=false)pbd=iftailthen{dwithconv_pbs=d.conv_pbs@[pb]}else{dwithconv_pbs=pb::d.conv_pbs}letconv_pbsd=d.conv_pbsletevar_sourceevkd=(finddevk).evar_sourceletevar_identevkevd=EvNames.identevkevd.evar_namesletevar_keyidevd=EvNames.keyidevd.evar_namesletget_aliased_evarsevd=evd.evar_flags.aliased_evarsletis_aliased_evarevdevk=trySome(Evar.Map.findevkevd.evar_flags.aliased_evars)withNot_found->Noneletdowncastevkcclevd=letevar_info=EvMap.findevkevd.undf_evarsinletevar_info'={evar_infowithevar_concl=ccl}in{evdwithundf_evars=EvMap.addevkevar_info'evd.undf_evars}(* extracts conversion problems that satisfy predicate p *)(* Note: conv_pbs not satisying p are stored back in reverse order *)letextract_conv_pbsevdp=let(pbs,pbs1)=List.fold_left(fun(pbs,pbs1)pb->ifppbthen(pb::pbs,pbs1)else(pbs,pb::pbs1))([],[])evd.conv_pbsin{evdwithconv_pbs=pbs1;last_mods=Evar.Set.empty},pbsletextract_changed_conv_pbsevdp=extract_conv_pbsevd(funpb->pevd.last_modspb)letextract_all_conv_pbsevd=extract_conv_pbsevd(fun_->true)letloc_of_conv_pbevd(pbty,env,t1,t2)=matchkind(fst(decompose_appt1))with|Evar(evk1,_)->fst(evar_sourceevk1evd)|_->matchkind(fst(decompose_appt2))with|Evar(evk2,_)->fst(evar_sourceevk2evd)|_->None(**********************************************************)(* Sort variables *)typerigid=UState.rigid=|UnivRigid|UnivFlexibleofbool(** Is substitution by an algebraic ok? *)letuniv_rigid=UnivRigidletuniv_flexible=UnivFlexiblefalseletuniv_flexible_alg=UnivFlexibletrueletevar_universe_contextd=d.universesletuniverse_context_setd=UState.context_setd.universesletto_universe_contextevd=UState.contextevd.universesletuniv_entry~polyevd=UState.univ_entry~polyevd.universesletcheck_univ_decl~polyevddecl=UState.check_univ_decl~polyevd.universesdeclletrestrict_universe_contextevdvars={evdwithuniverses=UState.restrictevd.universesvars}letuniverse_substevd=UState.substevd.universesletmerge_context_set?loc?(sideff=false)rigidevdctx'={evdwithuniverses=UState.merge?loc~sideffrigidevd.universesctx'}letwith_context_set?locrigidd(a,ctx)=(merge_context_set?locrigiddctx,a)letnew_univ_level_variable?loc?namerigidevd=letuctx',u=UState.new_univ_variable?locrigidnameevd.universesin({evdwithuniverses=uctx'},u)letnew_univ_variable?loc?namerigidevd=letuctx',u=UState.new_univ_variable?locrigidnameevd.universesin({evdwithuniverses=uctx'},Univ.Universe.makeu)letnew_sort_variable?loc?namerigidd=let(d',u)=new_univ_variable?locrigid?namedin(d',Sorts.sort_of_univu)letadd_global_univdu={dwithuniverses=UState.add_global_univd.universesu}letmake_flexible_variableevd~algebraicu={evdwithuniverses=UState.make_flexible_variableevd.universes~algebraicu}letmake_nonalgebraic_variableevdu={evdwithuniverses=UState.make_nonalgebraic_variableevd.universesu}(****************************************)(* Operations on constants *)(****************************************)letfresh_sort_in_family?loc?(rigid=univ_flexible)evds=with_context_set?locrigidevd(UnivGen.fresh_sort_in_familys)letfresh_constant_instance?loc?(rigid=univ_flexible)envevdc=with_context_set?locrigidevd(UnivGen.fresh_constant_instanceenvc)letfresh_inductive_instance?loc?(rigid=univ_flexible)envevdi=with_context_set?locrigidevd(UnivGen.fresh_inductive_instanceenvi)letfresh_constructor_instance?loc?(rigid=univ_flexible)envevdc=with_context_set?locrigidevd(UnivGen.fresh_constructor_instanceenvc)letfresh_array_instance?loc?(rigid=univ_flexible)envevd=with_context_set?locrigidevd(UnivGen.fresh_array_instanceenv)letfresh_global?loc?(rigid=univ_flexible)?namesenvevdgr=with_context_set?locrigidevd(UnivGen.fresh_global_instance?loc?namesenvgr)letis_sort_variableevds=UState.is_sort_variableevd.universessletis_flexible_levelevdl=letuctx=evd.universesinUniv.LMap.meml(UState.substuctx)letis_eq_sorts1s2=ifSorts.equals1s2thenNoneelseletu1=univ_of_sorts1andu2=univ_of_sorts2inifUniv.Universe.equalu1u2thenNoneelseSome(u1,u2)(* Precondition: l is not defined in the substitution *)letuniverse_rigidityevdl=letuctx=evd.universesinifUniv.LSet.meml(Univ.ContextSet.levels(UState.context_setuctx))thenUnivFlexible(Univ.LSet.meml(UState.algebraicsuctx))elseUnivRigidletnormalize_universeevd=letvars=UState.substevd.universesinletnormalize=UnivSubst.normalize_universe_opt_substvarsinnormalizeletnormalize_universe_instanceevdl=letvars=UState.substevd.universesinletnormalize=UnivSubst.level_subst_of(UnivSubst.normalize_univ_variable_opt_substvars)inUniv.Instance.subst_fnnormalizelletnormalize_sortevarss=matchswith|SProp|Prop|Set->s|Typeu->letu'=normalize_universeevarsuinifu'==uthenselseSorts.sort_of_univu'(* FIXME inefficient *)letset_eq_sortenvds1s2=lets1=normalize_sortds1ands2=normalize_sortds2inmatchis_eq_sorts1s2with|None->d|Some(u1,u2)->ifnot(type_in_typeenv)thenadd_universe_constraintsd(UnivProblem.Set.singleton(UnivProblem.UEq(u1,u2)))elsedletset_eq_leveldu1u2=add_constraintsd(Univ.enforce_eq_levelu1u2Univ.Constraint.empty)letset_leq_leveldu1u2=add_constraintsd(Univ.enforce_leq_levelu1u2Univ.Constraint.empty)letset_eq_instances?(flex=false)du1u2=add_universe_constraintsd(UnivProblem.enforce_eq_instances_univsflexu1u2UnivProblem.Set.empty)letset_leq_sortenvevds1s2=lets1=normalize_sortevds1ands2=normalize_sortevds2inmatchis_eq_sorts1s2with|None->evd|Some(u1,u2)->ifnot(type_in_typeenv)thenadd_universe_constraintsevd(UnivProblem.Set.singleton(UnivProblem.ULe(u1,u2)))elseevdletcheck_eqevdss'=UGraph.check_eq(UState.ugraphevd.universes)ss'letcheck_leqevdss'=UGraph.check_leq(UState.ugraphevd.universes)ss'letcheck_constraintsevdcsts=UGraph.check_constraintscsts(UState.ugraphevd.universes)letfix_undefined_variablesevd={evdwithuniverses=UState.fix_undefined_variablesevd.universes}letnf_univ_variablesevd=letuctx=UState.normalize_variablesevd.universesin{evdwithuniverses=uctx}letminimize_universesevd=letuctx'=UState.normalize_variablesevd.universesinletuctx'=UState.minimizeuctx'in{evdwithuniverses=uctx'}letuniverse_of_nameevds=UState.universe_of_nameevd.universessletuniverse_bindersevd=UState.universe_bindersevd.universesletuniversesevd=UState.ugraphevd.universesletupdate_sigma_univsugraphevd={evdwithuniverses=UState.update_sigma_univsevd.universesugraph}exceptionUniversesDiffer=UState.UniversesDiffer(**********************************************************)(* Side effects *)letconcat_side_effectseffeff'={seff_private=Safe_typing.concat_privateeff.seff_privateeff'.seff_private;seff_roles=Cmap.foldCmap.addeff.seff_roleseff'.seff_roles;}letemit_side_effectseffevd=leteffects=concat_side_effectseffevd.effectsin{evdwitheffects;universes=UState.emit_side_effectseff.seff_privateevd.universes}letdrop_side_effectsevd={evdwitheffects=empty_side_effects;}leteval_side_effectsevd=evd.effects(* Future goals *)letdeclare_future_goalevkevd=letfuture_goals=FutureGoals.add~principal:falseevkevd.future_goalsin{evdwithfuture_goals}letdeclare_principal_goalevkevd=letfuture_goals=FutureGoals.add~principal:trueevkevd.future_goalsin{evdwithfuture_goals}letpush_future_goalsevd={evdwithfuture_goals=FutureGoals.pushevd.future_goals}letpop_future_goalsevd=lethd,future_goals=FutureGoals.popevd.future_goalsinhd,{evdwithfuture_goals}letfold_future_goalsfsigma=FutureGoals.foldfsigmasigma.future_goalsletremove_future_goalevdevk={evdwithfuture_goals=FutureGoals.removeevkevd.future_goals}letpr_future_goals_stackevd=FutureGoals.pr_stackevd.future_goalsletgive_upevevd={evdwithgiven_up=Evar.Set.addevevd.given_up}letpush_shelfevd={evdwithshelf=[]::evd.shelf}letpop_shelfevd=matchevd.shelfwith|[]->anomalyPp.(str"shelf stack should not be empty")|hd::tl->hd,{evdwithshelf=tl}letfilter_shelffevd={evdwithshelf=List.map(List.filterf)evd.shelf}letshelveevdl=matchevd.shelfwith|[]->anomalyPp.(str"shelf stack should not be empty")|hd::tl->{evdwithshelf=(hd@l)::tl}letunshelveevdl={evdwithshelf=List.map(List.filter(funev->not(CList.mem_fEvar.equalevl)))evd.shelf}letgiven_upevd=evd.given_upletshelfevd=List.flattenevd.shelfletpr_shelfevd=letopenPpinifList.is_emptyevd.shelfthenstr"(empty stack)"elseprlist_with_sep(fun()->str"||")(prlist_with_sepspcEvar.print)evd.shelfletdefine_auxdefundefevkbody=letoldinfo=tryEvMap.findevkundefwithNot_found->ifEvMap.memevkdefthenanomaly~label:"Evd.define"(Pp.str"cannot define an evar twice.")elseanomaly~label:"Evd.define"(Pp.str"cannot define undeclared evar.")inlet()=assert(oldinfo.evar_body==Evar_empty)inletnewinfo={oldinfowithevar_body=Evar_definedbody}inEvMap.addevknewinfodef,EvMap.removeevkundef(* define the existential of section path sp as the constr body *)letdefine_genevkbodyevdevar_flags=let(defn_evars,undf_evars)=define_auxevd.defn_evarsevd.undf_evarsevkbodyinletlast_mods=matchevd.conv_pbswith|[]->evd.last_mods|_->Evar.Set.addevkevd.last_modsinletevar_names=EvNames.remove_name_definedevkevd.evar_namesin{evdwithdefn_evars;undf_evars;last_mods;evar_names;evar_flags}(** By default, the obligation and evar tag of the evar is removed *)letdefineevkbodyevd=letevar_flags=remove_evar_flagsevkevd.evar_flagsindefine_genevkbodyevdevar_flags(** In case of an evar-evar solution, the flags are inherited *)letdefine_with_evarevkbodyevd=letevk'=fst(destEvarbody)inletevar_flags=inherit_evar_flagsevd.evar_flagsevkevk'inletevd=unshelveevd[evk]inletfuture_goals=FutureGoals.removeevkevd.future_goalsinletevd={evdwithfuture_goals}indefine_genevkbodyevdevar_flags(* In case of restriction, we declare the aliasing and inherit the obligation
and typeclass flags. *)letrestrictevkfilter?candidates?srcevd=letevk'=new_untyped_evar()inletevar_info=EvMap.findevkevd.undf_evarsinletid_inst=Filter.filter_listfilterevar_info.evar_hyps.env_named_varinletevar_info'={evar_infowithevar_filter=filter;evar_candidates=candidates;evar_source=(matchsrcwithNone->evar_info.evar_source|Somesrc->src);evar_identity=Identity.makeid_inst;}inletlast_mods=matchevd.conv_pbswith|[]->evd.last_mods|_->Evar.Set.addevkevd.last_modsinletevar_names=EvNames.reassign_name_definedevkevk'evd.evar_namesinletbody=mkEvar(evk',id_inst)inlet(defn_evars,undf_evars)=define_auxevd.defn_evarsevd.undf_evarsevkbodyinletevar_flags=inherit_evar_flagsevd.evar_flagsevkevk'inletevd={evdwithundf_evars=EvMap.addevk'evar_info'undf_evars;defn_evars;last_mods;evar_names;evar_flags}in(* Mark new evar as future goal, removing previous one,
circumventing Proofview.advance but making Proof.run_tactic catch these. *)letevd=unshelveevd[evk]inletevd=remove_future_goalevdevkinletevd=declare_future_goalevk'evdin(evd,evk')letupdate_sourceevdevksrc=letevar_info=EvMap.findevkevd.undf_evarsinletevar_info'={evar_infowithevar_source=src}in{evdwithundf_evars=EvMap.addevkevar_info'evd.undf_evars}(**********************************************************)(* Accessing metas *)(** We use this function to overcome OCaml compiler limitations and to prevent
the use of costly in-place modifications. *)letset_metasevdmetas={defn_evars=evd.defn_evars;undf_evars=evd.undf_evars;universes=evd.universes;conv_pbs=evd.conv_pbs;last_mods=evd.last_mods;evar_flags=evd.evar_flags;metas;effects=evd.effects;evar_names=evd.evar_names;future_goals=evd.future_goals;given_up=evd.given_up;shelf=evd.shelf;extras=evd.extras;}letmeta_listevd=metamap_to_listevd.metasletundefined_metasevd=letfilter=function|(n,Clval(_,_,typ))->None|(n,Cltyp(_,typ))->Someninletm=List.map_filterfilter(meta_listevd)inList.sortInt.comparemletmap_metas_fvaluefevd=letmap=function|Clval(id,(c,s),typ)->Clval(id,(mk_freelisted(fc.rebus),s),typ)|x->xinset_metasevd(Metamap.Smart.mapmapevd.metas)letmap_metasfevd=letmapcl=map_clbfclinset_metasevd(Metamap.Smart.mapmapevd.metas)letmeta_opt_fvalueevdmv=matchMetamap.findmvevd.metaswith|Clval(_,b,_)->Someb|Cltyp_->Noneletmeta_definedevdmv=matchMetamap.findmvevd.metaswith|Clval_->true|Cltyp_->falselettry_meta_fvalueevdmv=matchMetamap.findmvevd.metaswith|Clval(_,b,_)->b|Cltyp_->raiseNot_foundletmeta_fvalueevdmv=trytry_meta_fvalueevdmvwithNot_found->anomaly~label:"meta_fvalue"(Pp.str"meta has no value.")letmeta_valueevdmv=(fst(try_meta_fvalueevdmv)).rebusletmeta_ftypeevdmv=matchMetamap.findmvevd.metaswith|Cltyp(_,b)->b|Clval(_,_,b)->bletmeta_typeevdmv=(meta_ftypeevdmv).rebusletmeta_type0=meta_typeletmeta_declaremvv?(name=Anonymous)evd=letmetas=Metamap.addmv(Cltyp(name,mk_freelistedv))evd.metasinset_metasevdmetas(* If the meta is defined then forget its name *)letmeta_nameevdmv=tryfst(clb_name(Metamap.findmvevd.metas))withNot_found->Anonymousletevar_source_of_metamvevd=matchmeta_nameevdmvwith|Anonymous->Loc.tagEvar_kinds.GoalEvar|Nameid->Loc.tag@@Evar_kinds.VarInstanceidletuse_meta_sourceevdmvv=matchConstr.kindvwith|Evar(evk,_)->letf=function|None->None|Someeviasx->matchevi.evar_sourcewith|None,Evar_kinds.GoalEvar->Some{eviwithevar_source=evar_source_of_metamvevd}|_->xin{evdwithundf_evars=EvMap.updateevkfevd.undf_evars}|_->evdletmeta_assignmv(v,pb)evd=letmodify_=function|Cltyp(na,ty)->Clval(na,(mk_freelistedv,pb),ty)|_->anomaly~label:"meta_assign"(Pp.str"already defined.")inletmetas=Metamap.modifymvmodifyevd.metasinletevd=use_meta_sourceevdmvvinset_metasevdmetasletmeta_reassignmv(v,pb)evd=letmodify_=function|Clval(na,_,ty)->Clval(na,(mk_freelistedv,pb),ty)|_->anomaly~label:"meta_reassign"(Pp.str"not yet defined.")inletmetas=Metamap.modifymvmodifyevd.metasinset_metasevdmetasletclear_metasevd={evdwithmetas=Metamap.empty}letmeta_merge?(with_univs=true)evd1evd2=letmetas=Metamap.foldMetamap.addevd1.metasevd2.metasinletuniverses=ifwith_univsthenUState.unionevd2.universesevd1.universeselseevd2.universesin{evd2withuniverses;metas;}typemetabinding=metavariable*constr*instance_statusletretract_coercible_metasevd=letmc=ref[]inletmapnv=matchvwith|Clval(na,(b,(Conv,CoerceToTypeass)),typ)->let()=mc:=(n,b.rebus,s)::!mcinCltyp(na,typ)|v->vinletmetas=Metamap.Smart.mapimapevd.metasin!mc,set_metasevdmetasletdependent_evar_identevevd=letevi=findevdevinmatchevi.evar_sourcewith|(_,Evar_kinds.VarInstanceid)->id|_->anomaly(str"Not an evar resulting of a dependent binding.")(**********************************************************)(* Extra data *)letget_extra_dataevd=evd.extrasletset_extra_dataextrasevd={evdwithextras}(*******************************************************************)typeopen_constr=evar_map*constr(*******************************************************************)(* The type constructor ['a sigma] adds an evar map to an object of
type ['a] *)type'asigma={it:'a;sigma:evar_map}letsig_itx=x.itletsig_sigx=x.sigmaleton_sigsf=letsigma',v=fs.sigmain{swithsigma=sigma'},v(*******************************************************************)(* The state monad with state an evar map. *)moduleMonadR=Monad.Make(structtype+'at=evar_map->evar_map*'aletreturna=funs->(s,a)let(>>=)xf=funs->let(s',a)=xsinfas'let(>>)xy=funs->let(s',())=xsinys'letmapfx=funs->on_sndf(xs)end)moduleMonad=Monad.Make(structtype+'at=evar_map->'a*evar_mapletreturna=funs->(a,s)let(>>=)xf=funs->let(a,s')=xsinfas'let(>>)xy=funs->let((),s')=xsinys'letmapfx=funs->on_fstf(xs)end)(**********************************************************)(* Failure explanation *)typeunsolvability_explanation=SeveralInstancesFoundofintmoduleMiniEConstr=structmoduleESorts=structtypet=Sorts.tletmakes=sletkindsigma=function|Sorts.Typeu->Sorts.sort_of_univ(normalize_universesigmau)|s->sletunsafe_to_sortss=sendmoduleEInstance=structtypet=Univ.Instance.tletmakei=iletkindsigmai=ifUniv.Instance.is_emptyithenielsenormalize_universe_instancesigmailetempty=Univ.Instance.emptyletis_empty=Univ.Instance.is_emptyletunsafe_to_instancet=tendtypet=econstrletsafe_evar_valuesigmaev=trySome(existential_valuesigmaev)withNotInstantiatedEvar|Not_found->Noneletrecwhd_evarsigmac=matchConstr.kindcwith|Evarev->beginmatchsafe_evar_valuesigmaevwith|Somec->whd_evarsigmac|None->cend|App(f,args)whenisEvarf->(* Enforce smart constructor invariant on applications *)letev=destEvarfinbeginmatchsafe_evar_valuesigmaevwith|None->c|Somef->whd_evarsigma(mkApp(f,args))end|Cast(c0,k,t)whenisEvarc0->(* Enforce smart constructor invariant on casts. *)letev=destEvarc0inbeginmatchsafe_evar_valuesigmaevwith|None->c|Somec->whd_evarsigma(mkCast(c,k,t))end|_->cletkindsigmac=Constr.kind(whd_evarsigmac)letkind_upto=kindletof_kind=Constr.of_kindletof_constrc=cletof_constr_arrayv=vletunsafe_to_constrc=cletunsafe_to_constr_arrayv=vletunsafe_eq=Reflletto_constr?(abort_on_undefined_evars=true)sigmac=letevar_value=ifnotabort_on_undefined_evarsthenfunev->safe_evar_valuesigmaevelsefunev->matchsafe_evar_valuesigmaevwith|Some_asv->v|None->anomaly~label:"econstr"Pp.(str"grounding a non evar-free term")inUnivSubst.nf_evars_and_universes_opt_substevar_value(universe_substsigma)cletto_constr_optsigmac=letevar_valueev=Some(existential_valuesigmaev)intrySome(UnivSubst.nf_evars_and_universes_opt_substevar_value(universe_substsigma)c)withNotInstantiatedEvar->Noneletof_named_decld=dletunsafe_to_named_decld=dletof_rel_decld=dletunsafe_to_rel_decld=dletto_rel_declsigmad=Context.Rel.Declaration.map_constr(to_constrsigma)dletunsafe_to_case_invertx=xletof_case_invertx=xend(** The following functions return the set of evars immediately
contained in the object *)(* excluding defined evars *)letevars_of_termevdc=letrecevrecaccc=letc=MiniEConstr.whd_evarevdcinmatchkindcwith|Evar(n,l)->Evar.Set.addn(List.fold_leftevrecaccl)|_->Constr.foldevrecacccinevrecEvar.Set.emptycletevar_nodes_of_termc=letrecevrecaccc=matchkindcwith|Evar(n,l)->Evar.Set.addn(List.fold_leftevrecaccl)|_->Constr.foldevrecacccinevrecEvar.Set.emptycletevars_of_named_contextevdnc=Context.Named.fold_outside(NamedDecl.fold_constr(funconstrs->Evar.Set.unions(evars_of_termevdconstr)))nc~init:Evar.Set.emptyletevars_of_filtered_evar_infoevdevi=Evar.Set.union(evars_of_termevdevi.evar_concl)(Evar.Set.union(matchevi.evar_bodywith|Evar_empty->Evar.Set.empty|Evar_definedb->evars_of_termevdb)(evars_of_named_contextevd(evar_filtered_contextevi)))