123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691(************************************************************************)(* * 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) *)(************************************************************************)openPpopenCErrorsopenSortsopenUtilopenContextopenEnvironopenNamesopenLibnamesopenConstrexpropenConstrexpr_opsopenConstrinternopenType_errorsopenPretypingopenContext.Rel.DeclarationopenEntriesmoduleRelDecl=Context.Rel.Declaration(* 3b| Mutual inductive definitions *)letwarn_auto_template=CWarnings.create~name:"auto-template"~category:"vernacular"~default:CWarnings.Disabled(funid->Pp.(strbrk"Automatically declaring "++Id.printid++strbrk" as template polymorphic. Use attributes or "++strbrk"disable Auto Template Polymorphism to avoid this warning."))letshould_auto_template=letopenGoptionsinletauto=reftrueinlet()=declare_bool_option{optdepr=false;optkey=["Auto";"Template";"Polymorphism"];optread=(fun()->!auto);optwrite=(funb->auto:=b);}infunidwould_auto->letb=!auto&&would_autoinifbthenwarn_auto_templateid;bletpush_typesenvidlrltl=List.fold_left3(funenvidrt->EConstr.push_rel(LocalAssum(make_annot(Nameid)r,t))env)envidlrltltypestructured_one_inductive_expr={ind_name:Id.t;ind_arity:constr_expr;ind_lc:(Id.t*constr_expr)list}exceptionSameofId.tletcheck_all_names_differentindl=letrecelements=function|[]->Id.Set.empty|id::l->lets=elementslinifId.Set.memidsthenraise(Sameid)elseId.Set.addidsinletind_names=List.map(funind->ind.ind_name)indlinletcstr_names=List.map_append(funind->List.mapfstind.ind_lc)indlinletind_names=matchelementsind_nameswith|s->s|exception(Samet)->raise(InductiveError(SameNamesTypest))inletcstr_names=matchelementscstr_nameswith|s->s|exception(Samec)->raise(InductiveError(SameNamesConstructorsc))inletl=Id.Set.interind_namescstr_namesinifnot(Id.Set.is_emptyl)thenraise(InductiveError(SameNamesOverlap(Id.Set.elementsl)))(** Make the arity conclusion flexible to avoid generating an upper bound universe now,
only if the universe does not appear anywhere else.
This is really a hack to stay compatible with the semantics of template polymorphic
inductives which are recognized when a "Type" appears at the end of the conlusion in
the source syntax. *)letreccheck_type_conclusionind=letopenGlob_terminmatchDAst.getindwith|GSort(UAnonymous{rigid=true})->(Sometrue)|GSort(UNamed_)->(Somefalse)|GProd(_,_,_,e)|GLetIn(_,_,_,e)|GLambda(_,_,_,e)|GApp(e,_)|GCast(e,_)->check_type_conclusione|_->Noneletmake_anonymous_conclusion_flexiblesigma=function|None->sigma|Some(false,_)->sigma|Some(true,s)->(matchEConstr.ESorts.kindsigmaswith|Typeu->(matchUniv.universe_leveluwith|Someu->Evd.make_flexible_variablesigma~algebraic:trueu|None->sigma)|_->sigma)letintern_ind_arityenvsigmaind=letc=intern_genIsTypeenvsigmaind.ind_arityinletimpls=Implicit_quantifiers.implicits_of_glob_constr~with_products:truecinletpseudo_poly=check_type_conclusioncin(constr_locind.ind_arity,c,impls,pseudo_poly)letpretype_ind_arityenvsigma(loc,c,impls,pseudo_poly)=letsigma,t=understand_tccenvsigma~expected_type:IsTypecinmatchReductionops.sort_of_arityenvsigmatwith|exceptionReduction.NotArity->user_err?loc(str"Not an arity")|s->letconcl=matchpseudo_polywith|Someb->Some(b,s)|None->Noneinsigma,(t,Retyping.relevance_of_sorts,concl,impls)(* ind_rel is the Rel for this inductive in the context without params.
n is how many arguments there are in the constructor. *)letmodel_conclusionenvsigmaind_relparamsnarity_indices=letmodel_head=EConstr.mkRel(n+Context.Rel.lengthparams+ind_rel)inletmodel_params=Context.Rel.to_extended_vectEConstr.mkRelnparamsinletsigma,model_indices=List.fold_right(fun(_,t)(sigma,subst)->lett=EConstr.Vars.substlsubst(EConstr.Vars.liftnn(List.lengthsubst+1)t)inletsigma,c=Evarutil.new_evarenvsigmatinsigma,c::subst)arity_indices(sigma,[])insigma,EConstr.mkApp(EConstr.mkApp(model_head,model_params),Array.of_list(List.revmodel_indices))letinterp_cstrsenv(sigma,ind_rel)implsparamsindarity=letcnames,ctyps=List.splitind.ind_lcinletarity_indices,cstr_sort=Reductionops.splay_arityenvsigmaarityin(* Interpret the constructor types *)letinterp_cstrsigmactyp=letflags=Pretyping.{all_no_fail_flagswithuse_typeclasses=UseTCForConv;solve_unification_constraints=false}inletsigma,(ctyp,cimpl)=interp_type_evars_impls~flagsenvsigma~implsctypinletctx,concl=Reductionops.splay_prod_assumenvsigmactypinletconcl_env=EConstr.push_rel_contextctxenvinletsigma_with_model_evars,model=model_conclusionconcl_envsigmaind_relparams(Context.Rel.lengthctx)arity_indicesin(* unify the expected with the provided conclusion *)letsigma=tryEvarconv.unifyconcl_envsigma_with_model_evarsReduction.CONVconclmodelwithEvarconv.UnableToUnify(sigma,e)->user_err(Himsg.explain_pretype_errorconcl_envsigma(Pretype_errors.CannotUnify(concl,model,(Somee))))insigma,(ctyp,cimpl)inletsigma,(ctyps,cimpls)=on_sndList.split@@List.fold_left_mapinterp_cstrsigmactypsin(sigma,predind_rel),(cnames,ctyps,cimpls)letsign_levelenvevdsign=fst(List.fold_right(fund(lev,env)->matchdwith|LocalDef_->lev,push_reldenv|LocalAssum_->lets=Retyping.get_sort_ofenvevd(EConstr.of_constr(RelDecl.get_typed))inletu=univ_of_sortsin(Univ.supulev,push_reldenv))sign(Univ.Universe.sprop,env))letsup_listmin=List.fold_leftUniv.supminletextract_levelenvevdmintys=letsorts=List.map(funty->letctx,concl=Reduction.dest_prod_assumenvtyinsign_levelenvevd(LocalAssum(make_annotAnonymousSorts.Relevant,concl)::ctx))tysinsup_listminsortsletis_flexible_sortevdu=matchUniv.Universe.leveluwith|Somel->Evd.is_flexible_levelevdl|None->false(**********************************************************************)(* Tools for template polymorphic inductive types *)(* Miscellaneous functions to remove or test local univ assumed to
occur only in the le constraints *)(*
Solve a system of universe constraint of the form
u_s11, ..., u_s1p1, w1 <= u1
...
u_sn1, ..., u_snpn, wn <= un
where
- the ui (1 <= i <= n) are universe variables,
- the sjk select subsets of the ui for each equations,
- the wi are arbitrary complex universes that do not mention the ui.
*)letis_direct_sort_constraintsv=matchswith|Someu->Univ.univ_level_memuv|None->falseletsolve_constraints_systemlevelslevel_bounds=letopenUnivinletlevels=Array.mapi(funio->matchowith|Someu->(matchUniverse.leveluwith|Someu->Someu|_->level_bounds.(i)<-Universe.suplevel_bounds.(i)u;None)|None->None)levelsinletv=Array.copylevel_boundsinletnind=Array.lengthvinletclos=Array.map(fun_->Int.Set.empty)levelsin(* First compute the transitive closure of the levels dependencies *)fori=0tonind-1doforj=0tonind-1doifnot(Int.equalij)&&is_direct_sort_constraintlevels.(j)v.(i)thenclos.(i)<-Int.Set.addjclos.(i);done;done;letrecclosure()=letcontinue=reffalseinArray.iteri(funideps->letdeps'=Int.Set.fold(funjacc->Int.Set.unionaccclos.(j))depsdepsinifInt.Set.equaldepsdeps'then()else(clos.(i)<-deps';continue:=true))clos;if!continuethenclosure()else()inclosure();fori=0tonind-1doforj=0tonind-1doifnot(Int.equalij)&&Int.Set.memjclos.(i)then(v.(i)<-Universe.supv.(i)level_bounds.(j));done;done;vletinductive_levelsenvevdaritiesinds=letdestarities=List.map(funx->x,Reduction.dest_arityenvx)aritiesinletlevels=List.map(fun(x,(ctx,a))->ifSorts.is_propa||Sorts.is_spropathenNoneelseSome(univ_of_sorta))destaritiesinletcstrs_levels,sizes=CList.split(List.map2(fun(_,tys)(arity,(ctx,du))->letlen=List.lengthtysinletminlev=Sorts.univ_of_sortduinletminlev=iflen>1&¬(is_impredicative_sortenvdu)thenUniv.supminlevUniv.type0_univelseminlevinletminlev=(* Indices contribute. *)ifindices_matterenvthenbeginletilev=sign_levelenvevdctxinUniv.supilevminlevendelseminlevinletclev=extract_levelenvevdminlevtysin(clev,len))indsdestarities)in(* Take the transitive closure of the system of constructors *)(* level constraints and remove the recursive dependencies *)letlevels'=solve_constraints_system(Array.of_listlevels)(Array.of_listcstrs_levels)inletevd,arities=CList.fold_left3(fun(evd,arities)cu(arity,(ctx,du))len->ifis_impredicative_sortenvduthen(* Any product is allowed here. *)evd,(false,arity)::aritieselse(* If in a predicative sort, or asked to infer the type,
we take the max of:
- indices (if in indices-matter mode)
- constructors
- Type(1) if there is more than 1 constructor
*)(* Constructors contribute. *)letevd=ifSorts.is_setduthenifnot(Evd.check_leqevdcuUniv.type0_univ)thenraise(InductiveErrorLargeNonPropInductiveNotInType)elseevdelseevdinletevd=iflen>=2&&Univ.is_type0m_univcuthen(* "Polymorphic" type constraint and more than one constructor,
should not land in Prop. Add constraint only if it would
land in Prop directly (no informative arguments as well). *)Evd.set_leq_sortenvevdSorts.setduelseevdinletduu=Sorts.univ_of_sortduinlettemplate_prop,evd=ifnot(Univ.is_small_univduu)&&Univ.Universe.equalcuduuthenifis_flexible_sortevdduu&¬(Evd.check_leqevdUniv.type0_univduu)thenifTerm.isArityarity(* If not a syntactic arity, the universe may be used in a
polymorphic instance and so cannot be lowered to Prop.
See #13300. *)thentrue,Evd.set_eq_sortenvevdSorts.propduelsefalse,Evd.set_eq_sortenvevdSorts.setduelsefalse,evdelsefalse,Evd.set_eq_sortenvevd(sort_of_univcu)duin(evd,(template_prop,arity)::arities))(evd,[])(Array.to_listlevels')destaritiessizesinevd,List.revaritiesletcheck_named{CAst.loc;v=na}=matchnawith|Name_->()|Anonymous->letmsg=str"Parameters must be named."inuser_err?locmsglettemplate_polymorphism_candidate~ctor_levelsuctxparamsconcl=matchuctxwith|Entries.Monomorphic_entryuctx->letconcltemplate=Option.cata(funs->not(Sorts.is_smalls))falseconclinifnotconcltemplatethenfalseelseletconclu=Option.cataSorts.univ_of_sortUniv.type0m_univconclinOption.has_some@@IndTyping.template_polymorphic_univs~ctor_levelsuctxparamsconclu|Entries.Polymorphic_entry_->falseletcheck_param=function|CLocalDef(na,_,_)->check_namedna|CLocalAssum(nas,Default_,_)->List.itercheck_namednas|CLocalAssum(nas,Generalized_,_)->()|CLocalPattern{CAst.loc}->Loc.raise?loc(Stream.Error"pattern with quote not allowed here")letrestrict_inductive_universessigmactx_paramsaritiesconstructors=letmerge_universes_of_constrc=Univ.LSet.union(EConstr.universes_of_constrsigma(EConstr.of_constrc))inletuvars=Univ.LSet.emptyinletuvars=Context.Rel.(fold_outside(Declaration.fold_constrmerge_universes_of_constr)ctx_params~init:uvars)inletuvars=List.fold_rightmerge_universes_of_constraritiesuvarsinletuvars=List.fold_right(fun(_,ctypes)->List.fold_rightmerge_universes_of_constrctypes)constructorsuvarsinEvd.restrict_universe_contextsigmauvarsletcheck_trivial_variancesvariances=Array.iter(function|None|SomeUniv.Variance.Invariant->()|Some_->CErrors.user_errPp.(strbrk"Universe variance was specified but this inductive will not be cumulative."))variancesletvariance_of_entry~cumulative~variancesuctx=matchuctxwith|Monomorphic_entry_->check_trivial_variancesvariances;None|Polymorphic_entry(nas,_)->ifnotcumulativethenbegincheck_trivial_variancesvariances;Noneendelseletlvs=Array.lengthvariancesinletlus=Array.lengthnasinassert(lvs<=lus);Some(Array.appendvariances(Array.make(lus-lvs)None))letinterp_mutual_inductive_constr~sigma~template~udecl~variances~ctx_params~indnames~arities~arityconcl~constructors~env_ar_params~cumulative~poly~private_ind~finite=(* Compute renewed arities *)letsigma=Evd.minimize_universessigmainletnf=Evarutil.nf_evars_universessigmainletconstructors=List.map(on_snd(List.mapnf))constructorsinletarities=List.mapEConstr.(to_constrsigma)aritiesinletsigma=List.fold_leftmake_anonymous_conclusion_flexiblesigmaarityconclinletsigma,arities=inductive_levelsenv_ar_paramssigmaaritiesconstructorsinletsigma=Evd.minimize_universessigmainletnf=Evarutil.nf_evars_universessigmainletarities=List.map(on_sndnf)aritiesinletconstructors=List.map(on_snd(List.mapnf))constructorsinletctx_params=List.mapTermops.(map_rel_decl(EConstr.to_constrsigma))ctx_paramsinletarityconcl=List.map(Option.map(fun(_anon,s)->EConstr.ESorts.kindsigmas))arityconclinletsigma=restrict_inductive_universessigmactx_params(List.mapsndarities)constructorsinletuctx=Evd.check_univ_decl~polysigmaudeclin(* Build the inductive entries *)letentries=List.map4(funindname(templatearity,arity)concl(cnames,ctypes)->{mind_entry_typename=indname;mind_entry_arity=arity;mind_entry_consnames=cnames;mind_entry_lc=ctypes})indnamesaritiesarityconclconstructorsinlettemplate=List.map4(funindname(templatearity,_)concl(_,ctypes)->lettemplate_candidate()=templatearity||letctor_levels=letadd_levelsclevels=Univ.LSet.unionlevels(Vars.universes_of_constrc)inletparam_levels=List.fold_left(funlevelsd->matchdwith|LocalAssum_->levels|LocalDef(_,b,t)->add_levelsb(add_levelstlevels))Univ.LSet.emptyctx_paramsinList.fold_left(funlevelsc->add_levelsclevels)param_levelsctypesintemplate_polymorphism_candidate~ctor_levelsuctxctx_paramsconclinmatchtemplatewith|Sometemplate->ifpoly&&templatethenuser_errPp.(strbrk"Template-polymorphism and universe polymorphism are not compatible.");template|None->should_auto_templateindname(template_candidate()))indnamesaritiesarityconclconstructorsinletis_template=List.for_all(funt->t)templatein(* Build the mutual inductive entry *)letmind_ent={mind_entry_params=ctx_params;mind_entry_record=None;mind_entry_finite=finite;mind_entry_inds=entries;mind_entry_private=ifprivate_indthenSomefalseelseNone;mind_entry_universes=uctx;mind_entry_template=is_template;mind_entry_variance=variance_of_entry~cumulative~variancesuctx;}inmind_ent,Evd.universe_binderssigmaletinterp_paramsenvudecluparamslparamsl=letsigma,udecl,variances=interp_cumul_univ_decl_optenvudeclinletsigma,(uimpls,((env_uparams,ctx_uparams),useruimpls))=interp_context_evars~program_mode:falseenvsigmauparamslinletsigma,(impls,((env_params,ctx_params),userimpls))=interp_context_evars~program_mode:false~impl_env:uimplsenv_uparamssigmaparamslin(* Names of parameters as arguments of the inductive type (defs removed) *)sigma,env_params,(ctx_params,env_uparams,ctx_uparams,userimpls,useruimpls,impls,udecl,variances)(* When a hole remains for a param, pretend the param is uniform and
do the unification.
[env_ar_par] is [uparams; inds; params]
*)letmaybe_unify_params_inenv_ar_parsigma~ninds~nparams~binders:kc=letis_indsigmakc=matchEConstr.kindsigmacwith|Constr.Reln->(* env is [uparams; inds; params; k other things] *)n>k+nparams&&n<=k+nparams+ninds|_->falseinletrecaux(env,kasenvk)sigmac=matchEConstr.kindsigmacwith|Constr.App(h,args)whenis_indsigmakh->Array.fold_left_i(funisigmaarg->ifi>=nparams||not(EConstr.isEvarsigmaarg)thensigmaelsebegintryEvarconv.unify_delayenvsigmaarg(EConstr.mkRel(k+nparams-i))withEvarconv.UnableToUnify_->(* ignore errors, we will get a "Cannot infer ..." error instead *)sigmaend)sigmaargs|_->Termops.fold_constr_with_full_bindersenvsigma(fund(env,k)->EConstr.push_reldenv,k+1)auxenvksigmacinaux(env_ar_par,k)sigmacletinterp_mutual_inductive_genenv0~templateudecl(uparamsl,paramsl,indl)notations~cumulative~poly~private_indfinite=check_all_names_differentindl;List.itercheck_paramparamsl;ifnot(List.is_emptyuparamsl)&¬(List.is_emptynotations)thenuser_err(str"Inductives with uniform parameters may not have attached notations.");letindnames=List.map(funind->ind.ind_name)indlinletninds=List.lengthindlin(* In case of template polymorphism, we need to compute more constraints *)letenv0=ifpolythenenv0elseEnviron.set_universes_lboundenv0UGraph.Bound.Propinletsigma,env_params,(ctx_params,env_uparams,ctx_uparams,userimpls,useruimpls,impls,udecl,variances)=interp_paramsenv0udecluparamslparamslin(* Interpret the arities *)letarities=List.map(intern_ind_arityenv_paramssigma)indlinletsigma,arities=List.fold_left_map(pretype_ind_arityenv_params)sigmaaritiesinletarities,relevances,arityconcl,indimpls=List.split4aritiesinletlift_ctxnctx=lett=EConstr.it_mkProd_or_LetInEConstr.mkPropctxinlett=EConstr.Vars.liftntinletctx,_=EConstr.decompose_prod_assumsigmatinctxinletctx_params_lifted,fullarities=lift_ctxnindsctx_params,CList.map_i(funic->EConstr.Vars.lifti(EConstr.it_mkProd_or_LetIncctx_params))0aritiesinletenv_ar=push_typesenv_uparamsindnamesrelevancesfullaritiesinletenv_ar_params=EConstr.push_rel_contextctx_params_liftedenv_arin(* Compute interpretation metadatas *)letindimpls=List.map(funimpls->userimpls@impls)indimplsinletimpls=compute_internalization_envenv_uparamssigma~implsInductiveindnamesfullaritiesindimplsinletntn_impls=compute_internalization_envenv_uparamssigmaInductiveindnamesfullaritiesindimplsinlet(sigma,_),constructors=Metasyntax.with_syntax_protection(fun()->(* Temporary declaration of notations and scopes *)List.iter(Metasyntax.set_notation_for_interpretationenv_paramsntn_impls)notations;(* Interpret the constructor types *)List.fold_left2_map(fun(sigma,ind_rel)indarity->interp_cstrsenv_ar_params(sigma,ind_rel)implsctx_params_liftedind(EConstr.Vars.liftnninds(Rel.lengthctx_params+1)arity))(sigma,ninds)indlarities)()inletnparams=Context.Rel.lengthctx_paramsinletsigma=List.fold_left(funsigma(_,ctyps,_)->List.fold_left(funsigmactyp->maybe_unify_params_inenv_ar_paramssigma~ninds~nparams~binders:0ctyp)sigmactyps)sigmaconstructorsin(* generalize over the uniform parameters *)letnuparams=Context.Rel.lengthctx_uparamsinletuargs=Context.Rel.to_extended_vectEConstr.mkRel0ctx_uparamsinletuparam_subst=List.initnindsEConstr.(funi->mkApp(mkRel(i+1+nuparams),uargs))@List.initnuparamsEConstr.(funi->mkRel(i+1))inletgeneralize_constructorc=EConstr.Unsafe.to_constr(EConstr.Vars.substnluparam_substnparamsc)inletcimpls=List.mappi3constructorsinletconstructors=List.map(fun(cnames,ctypes,cimpls)->(cnames,List.mapgeneralize_constructorctypes))constructorsinletctx_params=ctx_params@ctx_uparamsinletuserimpls=useruimpls@userimplsinletindimpls=List.map(funiimpl->useruimpls@iimpl)indimplsinletfullarities=List.map(func->EConstr.it_mkProd_or_LetIncctx_uparams)fullaritiesinletenv_ar=push_typesenv0indnamesrelevancesfullaritiesinletenv_ar_params=EConstr.push_rel_contextctx_paramsenv_arin(* Try further to solve evars, and instantiate them *)letsigma=solve_remaining_evarsall_and_fail_flagsenv_paramssigmainletimpls=List.map2(funindimplscimpls->indimpls,List.map(funimpls->userimpls@impls)cimpls)indimplscimplsinletmie,pl=interp_mutual_inductive_constr~template~sigma~ctx_params~udecl~variances~arities~arityconcl~constructors~env_ar_params~poly~finite~cumulative~private_ind~indnamesin(mie,pl,impls)(* Very syntactical equality *)leteq_local_bindersbl1bl2=List.equallocal_binder_eqbl1bl2leteq_params(up1,p1)(up2,p2)=eq_local_bindersup1up2&&Option.equaleq_local_bindersp1p2letextract_coercionsindl=letmkqid(_,({CAst.v=id},_))=qualid_of_identidinletextractlc=List.filter(fun(iscoe,_)->iscoe)lcinList.mapmkqid(List.flatten(List.map(fun(_,_,_,lc)->extractlc)indl))letextract_paramsindl=letparamsl=List.map(fun(_,params,_,_)->params)indlinmatchparamslwith|[]->anomaly(Pp.str"empty list of inductive types.")|params::paramsl->ifnot(List.for_all(eq_paramsparams)paramsl)thenuser_errPp.(str"Parameters should be syntactically the same for each inductive type.");paramsletextract_inductiveindl=List.map(fun({CAst.v=indname},_,ar,lc)->{ind_name=indname;ind_arity=Option.cata(funx->x)(CAst.make@@CSort(Glob_term.UAnonymous{rigid=true}))ar;ind_lc=List.map(fun(_,({CAst.v=id},t))->(id,t))lc})indlletextract_mutual_inductive_declaration_componentsindl=letindl,ntnl=List.splitindlinletparams=extract_paramsindlinletcoes=extract_coercionsindlinletindl=extract_inductiveindlin(params,indl),coes,List.flattenntnltypeuniform_inductive_flag=|UniformParameters|NonUniformParametersletdo_mutual_inductive~templateudeclindl~cumulative~poly?typing_flags~private_ind~uniformfinite=let(params,indl),coes,ntns=extract_mutual_inductive_declaration_componentsindlin(* Interpret the types *)letindl=matchparamswith|uparams,Someparams->(uparams,params,indl)|params,None->matchuniformwith|UniformParameters->(params,[],indl)|NonUniformParameters->([],params,indl)inletenv=Global.env()inletenv=Environ.update_typing_flags?typing_flagsenvinletmie,pl,impls=interp_mutual_inductive_genenv~templateudeclindlntns~cumulative~poly~private_indfinitein(* Declare the mutual inductive block with its associated schemes *)ignore(DeclareInd.declare_mutual_inductive_with_eliminations?typing_flagsmieplimpls);(* Declare the possible notations of inductive types *)List.iter(Metasyntax.add_notation_interpretation(Global.env()))ntns;(* Declare the coercions *)List.iter(funqid->ComCoercion.try_add_new_coercion(Nametab.locateqid)~local:false~poly)coes(** Prepare a "match" template for a given inductive type.
For each branch of the match, we list the constructor name
followed by enough pattern variables.
[Not_found] is raised if the given string isn't the qualid of
a known inductive type. *)(*
HH notes in PR #679:
The Show Match could also be made more robust, for instance in the
presence of let in the branch of a constructor. A
decompose_prod_assum would probably suffice for that, but then, it
is a Context.Rel.Declaration.t which needs to be matched and not
just a pair (name,type).
Otherwise, this is OK. After all, the API on inductive types is not
so canonical in general, and in this simple case, working at the
low-level of mind_nf_lc seems reasonable (compared to working at the
higher-level of Inductiveops).
*)letmake_casesind=letopenDeclarationsinletmib,mip=Global.lookup_inductiveindinUtil.Array.fold_right_i(funi(ctx,_)l->letal=Util.List.skipn(List.lengthmib.mind_params_ctxt)(List.revctx)inletrecrenameavoid=function|[]->[]|RelDecl.LocalDef_::l->"_"::renameavoidl|RelDecl.LocalAssum(n,_)::l->letn'=Namegen.next_name_away_with_default(Id.to_stringNamegen.default_dependent_ident)n.Context.binder_nameavoidinId.to_stringn'::rename(Id.Set.addn'avoid)linletal'=renameId.Set.emptyalinletconsref=GlobRef.ConstructRef(ith_constructor_of_inductiveind(i+1))in(Libnames.string_of_qualid(Nametab.shortest_qualid_of_globalId.Set.emptyconsref)::al')::l)mip.mind_nf_lc[]