Legend:
Page
Library
Module
Module type
Parameter
Class
Class type
Source
Source file comCoercion.ml
123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403(************************************************************************)(* * 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) *)(************************************************************************)openCErrorsopenUtilopenPpopenNamesopenTermopenConstropenContextopenVarsopenTermopsopenEnvironopenCoercionopsopenDeclareopenLibobjectletstrength_minl=ifList.mem`LOCALlthen`LOCALelse`GLOBALletloc_of_boolb=ifbthen`LOCALelse`GLOBAL(* Errors *)typecoercion_error_kind=|AlreadyExists|NotAFunction|NoSourceofcl_typoption|ForbiddenSourceClassofcl_typ|NoTarget|WrongTargetofcl_typ*cl_typ|NotAClassofGlobRef.texceptionCoercionErrorofcoercion_error_kindletexplain_coercion_errorg=function|AlreadyExists->(Printer.pr_globalg++str" is already a coercion")|NotAFunction->(Printer.pr_globalg++str" is not a function")|NoSource(Somecl)->(str"Cannot recognize "++pr_classcl++str" as a source class of "++Printer.pr_globalg)|NoSourceNone->(str": cannot find the source class of "++Printer.pr_globalg)|ForbiddenSourceClasscl->pr_classcl++str" cannot be a source class"|NoTarget->(str"Cannot find the target class")|WrongTarget(clt,cl)->(str"Found target class "++pr_classcl++str" instead of "++pr_classclt)|NotAClassref->(str"Type of "++Printer.pr_globalref++str" does not end with a sort")(* Verifications pour l'ajout d'une classe *)letcheck_reference_arityref=letenv=Global.env()inletc,_=Typeops.type_of_global_in_contextenvrefinifnot(Reductionops.is_arityenv(Evd.from_envenv)(EConstr.of_constrc))(* FIXME *)thenraise(CoercionError(NotAClassref))letcheck_arity=function|CL_FUN|CL_SORT->()|CL_CONSTcst->check_reference_arity(GlobRef.ConstRefcst)|CL_PROJp->check_reference_arity(GlobRef.ConstRef(Projection.Repr.constantp))|CL_SECVARid->check_reference_arity(GlobRef.VarRefid)|CL_INDkn->check_reference_arity(GlobRef.IndRefkn)(* Coercions *)(* check that the computed target is the provided one *)letcheck_targetclt=function|Someclwhennot(cl_typ_eqclclt)->raise(CoercionError(WrongTarget(clt,cl)))|_->()(* condition d'heritage uniforme *)letuniform_condsigmactxlt=List.for_all2eq(EConstr.eq_constrsigma)lt(Context.Rel.instance_listEConstr.mkRel0ctx)letclass_of_global=function|GlobRef.ConstRefsp->(matchStructures.PrimitiveProjections.find_optspwith|Somep->CL_PROJp|None->CL_CONSTsp)|GlobRef.IndRefsp->CL_INDsp|GlobRef.VarRefid->CL_SECVARid|GlobRef.ConstructRef_asc->user_err(str"Constructors, such as "++Printer.pr_globalc++str", cannot be used as a class.")(*
lp est la liste (inverse'e) des arguments de la coercion
ids est le nom de la classe source
sps_opt est le sp de la classe source dans le cas des structures
retourne:
la classe source
nbre d'arguments de la classe
le constr de la class
la liste des variables dont depend la classe source
l'indice de la classe source dans la liste lp
*)letget_sourceenvlpsource=letopenContext.Rel.Declarationinmatchsourcewith|None->(* Take the latest non let-in argument *)letrecaux=function|[]->raiseNot_found|LocalDef_::lt->auxlt|LocalAssum(_,t1)::lt->letcl1,u1,lv1=find_class_type(push_rel_contextltenv)Evd.empty(EConstr.of_constrt1)incl1,lt,lv1,1inauxlp|Somecl->(* Take the first argument that matches *)letrecauxenvacc=function|[]->raiseNot_found|LocalDef_asdecl::lt->aux(push_reldeclenv)(decl::acc)lt|LocalAssum(_,t1)asdecl::lt->tryletcl1,u1,lv1=find_class_typeenvEvd.empty(EConstr.of_constrt1)inifcl_typ_eqclcl1thencl1,acc,lv1,Context.Rel.nhypslt+1elseraiseNot_foundwithNot_found->aux(push_reldeclenv)(decl::acc)ltinauxenv[](List.revlp)letget_targetenvlptind=if(ind>1)thenCL_FUNelsematchpi1(find_class_type(push_rel_contextlpenv)Evd.empty(EConstr.of_constrt))with|CL_CONSTpwhenStructures.PrimitiveProjections.memp->CL_PROJ(Option.get@@Structures.PrimitiveProjections.find_optp)|x->xletstrength_of_cl=function|CL_CONSTkn->`GLOBAL|CL_SECVARid->`LOCAL|_->`GLOBALletstrength_of_global=function|GlobRef.VarRef_->`LOCAL|_->`GLOBALletget_strengthstrerefclsclt=letstres=strength_of_clclsinletstret=strength_of_clcltinletstref=strength_of_globalrefinstrength_min[stre;stres;stret;stref]letident_key_of_class=function|CL_FUN->"Funclass"|CL_SORT->"Sortclass"|CL_CONSTsp->Label.to_string(Constant.labelsp)|CL_PROJsp->Label.to_string(Projection.Repr.labelsp)|CL_IND(sp,_)->Label.to_string(MutInd.labelsp)|CL_SECVARid->Id.to_stringid(* Identity coercion *)leterror_not_transparentsource=user_err(pr_classsource++str" must be a transparent constant.")letbuild_id_coercionidf_optsourcepoly=letenv=Global.env()inletsigma=Evd.from_envenvinletsigma,vs=matchsourcewith|CL_CONSTsp->Evd.fresh_globalenvsigma(GlobRef.ConstRefsp)|_->error_not_transparentsourceinletvs=EConstr.Unsafe.to_constrvsinletc=matchconstant_opt_value_inenv(destConstvs)with|Somec->c|None->error_not_transparentsourceinletlams,t=decompose_lam_assumcinletval_f=it_mkLambda_or_LetIn(mkLambda(make_annot(NameNamegen.default_dependent_ident)Sorts.Relevant,applistcvs(Context.Rel.instance_listmkRel0lams),mkRel1))lamsinlettyp_f=List.fold_left(fundc->Term.mkProd_wo_LetIncd)(mkProd(make_annotAnonymousSorts.Relevant,applistcvs(Context.Rel.instance_listmkRel0lams),lift1t))lamsin(* juste pour verification *)letsigma,val_t=Typing.type_ofenvsigma(EConstr.of_constrval_f)inlet()=ifnot(Reductionops.is_conv_leqenvsigmaval_t(EConstr.of_constrtyp_f))thenuser_err(strbrk"Cannot be defined as coercion (maybe a bad number of arguments).")inletname=matchidf_optwith|Someidf->idf|None->letcl,u,_=find_class_typeenvsigma(EConstr.of_constrt)inId.of_string("Id_"^(ident_key_of_classsource)^"_"^(ident_key_of_classcl))inletunivs=Evd.univ_entry~polysigmainletconstr_entry=(* Cast is necessary to express [val_f] is identity *)DefinitionEntry(definition_entry~types:typ_f~univs~inline:true(mkCast(val_f,DEFAULTcast,typ_f)))inletkind=Decls.(IsDefinitionIdentityCoercion)inletkn=declare_constant~name~kindconstr_entryinGlobRef.ConstRefknletcheck_source=function|Some(CL_FUNass)->raise(CoercionError(ForbiddenSourceClasss))|_->()letcache_coercion?(update=false)c=letenv=Global.env()inletsigma=Evd.from_envenvinCoercionops.declare_coercionenvsigma~updatecletopen_coercionio=ifInt.equali1thencache_coercionoletdischarge_coercionc=ifc.coe_localthenNoneelseletn=tryArray.length(Lib.section_instancec.coe_value)withNot_found->0inletnc={cwithcoe_param=n+c.coe_param;coe_is_projection=Option.mapLib.discharge_proj_reprc.coe_is_projection;}inSomencletrebuild_coercionc={cwithcoe_typ=fst(Typeops.type_of_global_in_context(Global.env())c.coe_value)}letclassify_coercionobj=ifobj.coe_localthenDisposeelseSubstituteletcoe_cat=create_category"coercions"letinCoercion:coe_info_typ->obj=declare_object{(default_object"COERCION")withopen_function=simple_open~cat:coe_catopen_coercion;cache_function=cache_coercion;subst_function=(fun(subst,c)->subst_coercionsubstc);classify_function=classify_coercion;discharge_function=discharge_coercion;rebuild_function=rebuild_coercion}letdeclare_coercioncoeftyp?(local=false)~reversible~isid~src:cls~target:clt~params:ps()=letisproj=matchcoefwith|GlobRef.ConstRefc->Structures.PrimitiveProjections.find_optc|_->Noneinletc={coe_value=coef;coe_typ=typ;coe_local=local;coe_reversible=reversible;coe_is_identity=isid;coe_is_projection=isproj;coe_source=cls;coe_target=clt;coe_param=ps;}inLib.add_leaf(inCoercionc)(*
nom de la fonction coercion
strength de f
nom de la classe source (optionnel)
sp de la classe source (dans le cas des structures)
nom de la classe target (optionnel)
booleen "coercion identite'?"
lorque source est None alors target est None aussi.
*)letwarn_uniform_inheritance=CWarnings.create~name:"uniform-inheritance"~category:"typechecker"(fung->Printer.pr_globalg++strbrk" does not respect the uniform inheritance condition.")letadd_new_coercion_corecoefstrepoly~nonuniform~reversiblesourcetargetisid:unit=check_sourcesource;letenv=Global.env()inlett,_=Typeops.type_of_global_in_contextenvcoefinifcoercion_existscoefthenraise(CoercionErrorAlreadyExists);letlp,tg=decompose_prod_assumtinletllp=List.lengthlpinifInt.equalllp0thenraise(CoercionErrorNotAFunction);let(cls,ctx,lvs,ind)=tryget_sourceenvlpsourcewithNot_found->raise(CoercionError(NoSourcesource))incheck_source(Somecls);ifnot(nonuniform||uniform_condEvd.empty(* FIXME - for when possibly called with unresolved evars in the future *)ctxlvs)thenwarn_uniform_inheritancecoef;letclt=tryget_targetenvlptgindwithNot_found->raise(CoercionErrorNoTarget)incheck_targetclttarget;check_aritycls;check_arityclt;letlocal=matchget_strengthstrecoefclscltwith|`LOCAL->true|`GLOBAL->falseinletparams=List.length(Context.Rel.instance_listEConstr.mkRel0ctx)indeclare_coercioncoeft~local~reversible~isid~src:cls~target:clt~params()lettry_add_new_coercion_coreref~localc~nonuniform~reversibledef=tryadd_new_coercion_coreref(loc_of_boollocal)c~nonuniform~reversibledefwithCoercionErrore->user_err(explain_coercion_errorrefe++str".")lettry_add_new_coercionref~local~poly~nonuniform~reversible=try_add_new_coercion_coreref~localpoly~nonuniform~reversibleNoneNonefalselettry_add_new_coercion_subclasscl~local~poly~nonuniform~reversible=letcoe_ref=build_id_coercionNoneclpolyintry_add_new_coercion_corecoe_ref~localpoly~nonuniform~reversible(Somecl)Nonetruelettry_add_new_coercion_with_targetref~local~poly~nonuniform~reversible~source~target=try_add_new_coercion_coreref~localpoly~nonuniform~reversible(Somesource)(Sometarget)falselettry_add_new_identity_coercionid~local~poly~source~target=letref=build_id_coercion(Someid)sourcepolyintry_add_new_coercion_coreref~localpoly~nonuniform:false~reversible:true(Somesource)(Sometarget)truelettry_add_new_coercion_with_sourceref~local~poly~nonuniform~reversible~source=try_add_new_coercion_coreref~localpoly~nonuniform~reversible(Somesource)Nonefalseletadd_coercion_hookpolynonuniformreversible{Declare.Hook.S.scope;dref;_}=letopenLocalityinletlocal=matchscopewith|Discharge->assertfalse(* Local Coercion in section behaves like Local Definition *)|GlobalImportNeedQualified->true|GlobalImportDefaultBehavior->falseinlet()=try_add_new_coerciondref~local~poly~nonuniform~reversibleinletmsg=Nametab.pr_global_envId.Set.emptydref++str" is now a coercion"inFlags.if_verboseFeedback.msg_infomsgletadd_coercion_hook~poly~nonuniform~reversible=Declare.Hook.make(add_coercion_hookpolynonuniformreversible)letadd_subclass_hook~poly{Declare.Hook.S.scope;dref;_}=letopenLocalityinletstre=matchscopewith|Discharge->assertfalse(* Local Subclass in section behaves like Local Definition *)|GlobalImportNeedQualified->true|GlobalImportDefaultBehavior->falseinletcl=class_of_globaldrefintry_add_new_coercion_subclasscl~local:stre~poly~nonuniform:falseletnonuniform=Attributes.bool_attribute~name:"nonuniform"letadd_subclass_hook~poly~reversible=Declare.Hook.make(add_subclass_hook~poly~reversible)letwarn_reverse_no_change=CWarnings.create~name:"reversible-no-change"~category:"coercions"(fun()->str"The reversible attribute is unchanged.")letchange_reverseref~reversible=ifnot(coercion_existsref)thenuser_err(Printer.pr_globalref++str" is not a coercion.");letcoe_info=coercion_inforefinifreversible=coe_info.coe_reversiblethenwarn_reverse_no_change()elsecache_coercion~update:true{coe_infowithcoe_reversible=reversible}