1234567891011121314151617181920212223242526272829303132333435363738394041424344454647484950515253545556575859606162636465666768697071727374757677787980818283848586878889909192939495969798991001011021031041051061071081091101111121131141151161171181191201211221231241251261271281291301311321331341351361371381391401411421431441451461471481491501511521531541551561571581591601611621631641651661671681691701711721731741751761771781791801811821831841851861871881891901911921931941951961971981992002012022032042052062072082092102112122132142152162172182192202212222232242252262272282292302312322332342352362372382392402412422432442452462472482492502512522532542552562572582592602612622632642652662672682692702712722732742752762772782792802812822832842852862872882892902912922932942952962972982993003013023033043053063073083093103113123133143153163173183193203213223233243253263273283293303313323333343353363373383393403413423433443453463473483493503513523533543553563573583593603613623633643653663673683693703713723733743753763773783793803813823833843853863873883893903913923933943953963973983994004014024034044054064074084094104114124134144154164174184194204214224234244254264274284294304314324334344354364374384394404414424434444454464474484494504514524534544554564574584594604614624634644654664674684694704714724734744754764774784794804814824834844854864874884894904914924934944954964974984995005015025035045055065075085095105115125135145155165175185195205215225235245255265275285295305315325335345355365375385395405415425435445455465475485495505515525535545555565575585595605615625635645655665675685695705715725735745755765775785795805815825835845855865875885895905915925935945955965975985996006016026036046056066076086096106116126136146156166176186196206216226236246256266276286296306316326336346356366376386396406416426436446456466476486496506516526536546556566576586596606616626636646656666676686696706716726736746756766776786796806816826836846856866876886896906916926936946956966976986997007017027037047057067077087097107117127137147157167177187197207217227237247257267277287297307317327337347357367377387397407417427437447457467477487497507517527537547557567577587597607617627637647657667677687697707717727737747757767777787797807817827837847857867877887897907917927937947957967977987998008018028038048058068078088098108118128138148158168178188198208218228238248258268278288298308318328338348358368378388398408418428438448458468478488498508518528538548558568578588598608618628638648658668678688698708718728738748758768778788798808818828838848858868878888898908918928938948958968978988999009019029039049059069079089099109119129139149159169179189199209219229239249259269279289299309319329339349359369379389399409419429439449459469479489499509519529539549559569579589599609619629639649659669679689699709719729739749759769779789799809819829839849859869879889899909919929939949959969979989991000100110021003100410051006100710081009101010111012101310141015101610171018101910201021102210231024102510261027(************************************************************************)(* * 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) *)(************************************************************************)(* This file is about the automatic generation of schemes about
decidable equality, created by Vincent Siles, Oct 2007 *)openCErrorsopenUtilopenPpopenTermopenConstropenContextopenVarsopenTermopsopenDeclarationsopenNamesopenInductiveopsopenTacticsopenInd_tablesopenNamegenopenTactypesopenProofview.NotationsmoduleRelDecl=Context.Rel.Declaration(**********************************************************************)(* Generic synthesis of boolean equality *)letquick_chopnl=letreckick_last=function|t::[]->[]|t::q->t::(kick_lastq)|[]->failwith"kick_last"andaux=function|(0,l')->l'|(n,h::t)->aux(n-1,t)|_->failwith"quick_chop"inifn>(List.lengthl)thenfailwith"quick_chop args"elsekick_last(aux(n,l))letdeconstruct_typet=letl,r=decompose_prodtin(List.rev_mapsndl)@[r]exceptionEqNotFoundofinductive*inductiveexceptionEqUnknownofstringexceptionUndefinedCstofstringexceptionInductiveWithProductexceptionInductiveWithSortexceptionParameterWithoutEqualityofGlobRef.texceptionNonSingletonPropofinductiveexceptionDecidabilityMutualNotSupportedexceptionNoDecidabilityCoInductiveexceptionConstructorWithNonParametricInductiveTypeofinductiveexceptionDecidabilityIndicesNotSupported(* Some pre declaration of constant we are going to use *)letandb_prop=fun_->UnivGen.constr_of_monomorphic_global(Global.env())(Coqlib.lib_ref"core.bool.andb_prop")letandb_true_intro=fun_->UnivGen.constr_of_monomorphic_global(Global.env())(Coqlib.lib_ref"core.bool.andb_true_intro")(* We avoid to use lazy as the binding of constants can change *)letbb()=UnivGen.constr_of_monomorphic_global(Global.env())(Coqlib.lib_ref"core.bool.type")lettt()=UnivGen.constr_of_monomorphic_global(Global.env())(Coqlib.lib_ref"core.bool.true")letff()=UnivGen.constr_of_monomorphic_global(Global.env())(Coqlib.lib_ref"core.bool.false")leteq()=UnivGen.constr_of_monomorphic_global(Global.env())(Coqlib.lib_ref"core.eq.type")letsumbool()=UnivGen.constr_of_monomorphic_global(Global.env())(Coqlib.lib_ref"core.sumbool.type")letandb=fun_->UnivGen.constr_of_monomorphic_global(Global.env())(Coqlib.lib_ref"core.bool.andb")letinduct_onc=inductionfalseNonecNoneNoneletdestruct_onc=destructfalseNonecNoneNoneletdestruct_on_usingcid=destructfalseNonec(Some(CAst.make@@IntroOrPattern[[CAst.make@@IntroNamingIntroAnonymous];[CAst.make@@IntroNaming(IntroIdentifierid)]]))Noneletdestruct_on_ascl=destructfalseNonec(Some(CAst.makel))Noneletinj_flags=Some{Equality.keep_proof_equalities=true;(* necessary *)Equality.injection_pattern_l2r_order=true;(* does not matter here *)}letmy_discr_tac=Equality.discr_tacfalseNoneletmy_inj_tacx=Equality.injinj_flagsNonefalseNone(EConstr.mkVarx,NoBindings)(* reconstruct the inductive with the correct de Bruijn indexes *)letmkFullIndenv(ind,u)n=letmib=Environ.lookup_mind(fstind)envinletnparams=mib.mind_nparamsinletnparrec=mib.mind_nparams_recin(* params context divided *)letlnonparrec,lnamesparrec=context_chop(nparams-nparrec)mib.mind_params_ctxtinifnparrec>0thenmkApp(mkIndU(ind,u),Array.of_list(Context.Rel.instance_listmkRel(nparrec+n)lnamesparrec))elsemkIndU(ind,u)letcheck_bool_is_defined()=ifnot(Coqlib.has_ref"core.bool.type")thenraise(UndefinedCst"bool")letcheck_no_indicesmib=ifArray.exists(funmip->mip.mind_nrealargs<>0)mib.mind_packetsthenraiseDecidabilityIndicesNotSupportedletget_schemehandlekind=matchlocal_lookup_schemehandlekindwith|None->assertfalse|Somec->cletbeq_scheme_kind_aux=ref(fun_->failwith"Undefined")letget_inductive_depsenvkn=(* fetching the mutual inductive body *)letmib=Environ.lookup_mindknenvin(* number of inductives in the mutual *)letnb_ind=Array.lengthmib.mind_packetsin(* number of params in the type *)letnparrec=mib.mind_nparams_recincheck_no_indicesmib;letmake_one_eqaccui=(* This function is only trying to recursively compute the inductive types
appearing as arguments of the constructors. This is done to support
equality decision over hereditarily first-order types. It could be
perfomed in a much cleaner way, e.g. using the kernel normal form of
constructor types and kernel whd_all for the argument types. *)letrecauxaccuc=let(c,a)=Reductionops.whd_betaiota_stackenvEvd.emptyEConstr.(of_constrc)inlet(c,a)=EConstr.Unsafe.(to_constrc,List.mapto_constra)inmatchConstr.kindcwith|Cast(x,_,_)->auxaccu(Term.applist(x,a))|App_->assertfalse|Ind((kn',_),_)->ifEnviron.QMutInd.equalenvknkn'thenaccuelseList.fold_leftaux(kn'::accu)a|Const(kn,u)->(matchEnviron.constant_opt_value_inenv(kn,u)with|Somec->auxaccu(Term.applist(c,a))|None->accu)|Rel_|Var_|Sort_|Prod_|Lambda_|LetIn_|Proj_|Construct_|Case_|CoFix_|Fix_|Meta_|Evar_|Int_|Float_|Array_->accuinletu=Univ.Instance.emptyinletconstrsn=get_constructorsenv(make_ind_family(((kn,i),u),Context.Rel.instance_listmkRel(n+nb_ind-1)mib.mind_params_ctxt))inletconstrsi=constrs(3+nparrec)inletfoldiaccuarg=letfoldaccuc=auxaccu(RelDecl.get_typec)inList.fold_leftfoldaccuarg.cs_argsinArray.fold_left_ifoldaccuconstrsiinArray.fold_left_i(funiaccu_->make_one_eqaccui)[]mib.mind_packetsletbuild_beq_scheme_depsenvkn=letinds=get_inductive_depsenvkninList.map(funind->SchemeMutualDep(ind,!beq_scheme_kind_aux()))indsletbuild_beq_schemeenvhandlekn=check_bool_is_defined();(* fetching the mutual inductive body *)letmib=Environ.lookup_mindknenvin(* number of inductives in the mutual *)letnb_ind=Array.lengthmib.mind_packetsin(* number of params in the type *)letnparams=mib.mind_nparamsinletnparrec=mib.mind_nparams_recincheck_no_indicesmib;(* params context divided *)letlnonparrec,lnamesparrec=context_chop(nparams-nparrec)mib.mind_params_ctxtin(* predef coq's boolean type *)(* rec name *)letrec_namei=(Id.to_string(Array.getmib.mind_packetsi).mind_typename)^"_eqrec"in(* construct the "fun A B ... N, eqA eqB eqC ... N => fixpoint" part *)letcreate_inputc=letmyArrowuv=mkArrowuSorts.Relevant(lift1v)andeqName=function|Names->Id.of_string("eq_"^(Id.to_strings))|Anonymous->Id.of_string"eq_A"inletext_rel_list=Context.Rel.instance_listmkRel0lnamesparrecinletlift_cnt=ref0inleteqs_typ=List.map(funaa->leta=lift!lift_cntaainincrlift_cnt;myArrowa(myArrowa(bb())))ext_rel_listinleteq_input=List.fold_left2(funabdecl->(* mkLambda(n,b,a) ) *)(* here I leave the Naming thingy so that the type of
the function is more readable for the user *)mkNamedLambda(map_annoteqName(RelDecl.get_annotdecl))ba)c(List.reveqs_typ)lnamesparrecinList.fold_left(funadecl->(* mkLambda(n,t,a)) eq_input rel_list *)(* Same here , hoping the auto renaming will do something good ;) *)letx=map_annot(functionNames->s|Anonymous->Id.of_string"A")(RelDecl.get_annotdecl)inmkNamedLambdax(RelDecl.get_typedecl)a)eq_inputlnamesparrecinletmake_one_eqcur=letu=Univ.Instance.emptyinletind=(kn,cur),u(* FIXME *)in(* current inductive we are working on *)letcur_packet=mib.mind_packets.(snd(fstind))in(* Inductive toto : [rettyp] := *)letrettyp=Inductive.type_of_inductive((mib,cur_packet),u)in(* split rettyp in a list without the non rec params and the last ->
e.g. Inductive vec (A:Set) : nat -> Set := ... will do [nat] *)letrettyp_l=quick_chopnparrec(deconstruct_typerettyp)in(* give a type A, this function tries to find the equality on A declared
previously *)(* nlist = the number of args (A , B , ... )
eqA = the de Bruijn index of the first eq param
ndx = how much to translate due to the 2nd Case
*)letcompute_A_equalityrel_listnlisteqAndxt=letlifti=ndxinletrecauxc=let(c,a)=Reductionops.whd_betaiota_stackenvEvd.emptyEConstr.(of_constrc)inlet(c,a)=EConstr.Unsafe.(to_constrc,List.mapto_constra)inmatchConstr.kindcwith|Relx->mkRel(x-nlist+ndx)|Varx->(* Support for working in a context with "eq_x : x -> x -> bool" *)leteid=Id.of_string("eq_"^(Id.to_stringx))inlet()=tryignore(Environ.lookup_namedeidenv)withNot_found->raise(ParameterWithoutEquality(GlobRef.VarRefx))inmkVareid|Cast(x,_,_)->aux(Term.applist(x,a))|App_->assertfalse|Ind((kn',iasind'),u)(*FIXME: universes *)->ifEnviron.QMutInd.equalenvknkn'thenmkRel(eqA-nlist-i+nb_ind-1)elsebegintryletc=get_schemehandle(!beq_scheme_kind_aux())ind'inleteq=mkConstcinleteqa=Array.of_list@@List.mapauxainletargs=Array.append(Array.of_list(List.map(funx->liftliftix)a))eqainifInt.equal(Array.lengthargs)0theneqelsemkApp(eq,args)withNot_found->raise(EqNotFound(ind',fstind))end|Sort_->raiseInductiveWithSort|Prod_->raiseInductiveWithProduct|Lambda_->raise(EqUnknown"abstraction")|LetIn_->raise(EqUnknown"let-in")|Const(kn,u)->(matchEnviron.constant_opt_value_inenv(kn,u)with|Somec->aux(Term.applist(c,a))|None->(* Support for working in a context with "eq_x : x -> x -> bool" *)(* Needs Hints, see test suite *)leteq_lbl=Label.make("eq_"^Label.to_string(Constant.labelkn))inletkneq=Constant.change_labelkneq_lblinifEnviron.mem_constantkneqenvthenlet_=Environ.constant_opt_value_inenv(kneq,u)inTerm.applist(mkConstkneq,a)elseraise(ParameterWithoutEquality(GlobRef.ConstRefkn)))|Proj_->raise(EqUnknown"projection")|Construct_->raise(EqUnknown"constructor")|Case_->raise(EqUnknown"match")|CoFix_->raise(EqUnknown"cofix")|Fix_->raise(EqUnknown"fix")|Meta_->raise(EqUnknown"meta-variable")|Evar_->raise(EqUnknown"existential variable")|Int_->raise(EqUnknown"int")|Float_->raise(EqUnknown"float")|Array_->raise(EqUnknown"array")inauxtin(* construct the predicate for the Case part*)letdo_predicaterel_listn=List.fold_left(funab->mkLambda(make_annotAnonymousSorts.Relevant,b,a))(mkLambda(make_annotAnonymousSorts.Relevant,mkFullIndenvind(n+3+(List.lengthrettyp_l)+nb_ind-1),(bb())))(List.revrettyp_l)in(* make_one_eq *)(* do the [| C1 ... => match Y with ... end
...
Cn => match Y with ... end |] part *)letrci=Sorts.Relevantin(* TODO relevance *)letci=make_case_infoenv(fstind)rciMatchStyleinletconstrsn=letparams=Context.Rel.instance_listmkRel(n+nb_ind-1)mib.mind_params_ctxtinget_constructorsenv(make_ind_family(ind,params))inletconstrsi=constrs(3+nparrec)inletn=Array.lengthconstrsiinletar=Array.initn(funi->letnb_cstr_args=List.lengthconstrsi.(i).cs_argsinletconstrsj=constrs(3+nparrec+nb_cstr_args)inletar2=Array.initn(funj->ifInt.equalijthenletcc=matchnb_cstr_argswith|0->tt()|_->leteqs=Array.initnb_cstr_args(funndx->letcc=RelDecl.get_type(List.nthconstrsi.(i).cs_argsndx)inleteqA=compute_A_equalityrel_listnparrec(nparrec+3+2*nb_cstr_args)(nb_cstr_args+ndx+1)ccinmkApp(eqA,[|mkRel(ndx+1+nb_cstr_args);mkRel(ndx+1)|]))inArray.fold_left(funab->mkApp(andb(),[|b;a|]))eqs.(0)(Array.subeqs1(nb_cstr_args-1))inList.fold_left(funadecl->mkLambda(RelDecl.get_annotdecl,RelDecl.get_typedecl,a))ccconstrsj.(j).cs_argselseList.fold_left(funadecl->mkLambda(RelDecl.get_annotdecl,RelDecl.get_typedecl,a))(ff())(constrsj.(j).cs_args))inletpred=EConstr.of_constr(do_predicaterel_listnb_cstr_args)inletcase=simple_make_case_or_projectenv(Evd.from_envenv)cipredNoInvert(EConstr.mkVar(Id.of_string"Y"))(EConstr.of_constr_arrayar2)inList.fold_left(funadecl->mkLambda(RelDecl.get_annotdecl,RelDecl.get_typedecl,a))(EConstr.Unsafe.to_constrcase)(constrsi.(i).cs_args))inletpred=EConstr.of_constr(do_predicaterel_list0)inletcase=simple_make_case_or_projectenv(Evd.from_envenv)cipredNoInvert(EConstr.mkVar(Id.of_string"X"))(EConstr.of_constr_arrayar)inmkNamedLambda(make_annot(Id.of_string"X")Sorts.Relevant)(mkFullIndenvind(nb_ind-1+1))(mkNamedLambda(make_annot(Id.of_string"Y")Sorts.Relevant)(mkFullIndenvind(nb_ind-1+2))((EConstr.Unsafe.to_constrcase)))in(* build_beq_scheme *)letnames=Array.makenb_ind(make_annotAnonymousSorts.Relevant)andtypes=Array.makenb_indmkSetandcores=Array.makenb_indmkSetinletu=Univ.Instance.emptyinfori=0to(nb_ind-1)donames.(i)<-make_annot(Name(Id.of_string(rec_namei)))Sorts.Relevant;types.(i)<-mkArrow(mkFullIndenv((kn,i),u)0)Sorts.Relevant(mkArrow(mkFullIndenv((kn,i),u)1)Sorts.Relevant(bb()));letc=make_one_eqiincores.(i)<-c;done;letres=Array.initnb_ind(funi->letkelim=Inductive.elim_sort(mib,mib.mind_packets.(i))inifnot(Sorts.family_leqInSetkelim)thenraise(NonSingletonProp(kn,i));letfix=matchmib.mind_finitewith|CoFinite->raiseNoDecidabilityCoInductive;|Finite->mkFix(((Array.makenb_ind0),i),(names,types,cores))|BiFinite->(* If the inductive type is not recursive, the fixpoint is
not used, so let's replace it with garbage *)letsubst=List.initnb_ind(fun_->mkProp)inVars.substlsubstcores.(i)increate_inputfix)inres,UState.from_envenvletbeq_scheme_kind=declare_mutual_scheme_object"_beq"~deps:build_beq_scheme_depsbuild_beq_schemelet_=beq_scheme_kind_aux:=fun()->beq_scheme_kind(* This function tryies to get the [inductive] between a constr
the constr should be Ind i or App(Ind i,[|args|])
*)letdestruct_indenvsigmac=letopenEConstrinlet(c,v)=Reductionops.whd_all_stackenvsigmacindestIndsigmac,Array.of_listvletbl_scheme_kind_aux=ref(fun()->failwith"Undefined")letlb_scheme_kind_aux=ref(fun()->failwith"Undefined")(*
In the following, avoid is the list of names to avoid.
If the args of the Inductive type are A1 ... An
then avoid should be
[| lb_An ... lb _A1 (resp. bl_An ... bl_A1)
eq_An .... eq_A1 An ... A1 |]
so from Ai we can find the correct eq_Ai bl_ai or lb_ai
*)(* used in the leib -> bool side*)letdo_replace_lbhandleaavoidnargpq=letopenEConstrinletavoid=Array.of_listaavoidinletdo_argenvsigmahdvoffset=matchkindsigmavwith|Vars->letx=narg*offsetinletn=Array.lengthavoidinletrecfindi=ifId.equalavoid.(n-i)sthenavoid.(n-i-x)else(ifi<nthenfind(i+1)elseuser_err(str"Var "++Id.prints++str" seems unknown."))inmkVar(find1)|Const(cst,_)->(* Works in specific situations where the args have to be already declared as a
Parameter (see example "J" in test file SchemeEquality.v) *)letlbl=Label.to_string(Constant.labelcst)inletnewlbl=ifInt.equaloffset1then("eq_"^lbl)else(lbl^"_lb")inletnewcst=Constant.change_labelcst(Label.makenewlbl)inifEnviron.mem_constantnewcstenvthenmkConstnewcstelseraise(ConstructorWithNonParametricInductiveType(fsthd))|_->raise(ConstructorWithNonParametricInductiveType(fsthd))inProofview.Goal.enterbeginfungl->lettype_of_pq=Tacmach.pf_get_type_ofglpinletsigma=Tacmach.projectglinletenv=Tacmach.pf_envglinletu,v=destruct_indenvsigmatype_of_pqinletc=get_schemehandle(!lb_scheme_kind_aux())(fstu)inletlb_type_of_p=mkConstcinProofview.tclEVARMAP>>=funsigma->letlb_args=Array.append(Array.appendv(Array.Smart.map(funx->do_argenvsigmaux1)v))(Array.Smart.map(funx->do_argenvsigmaux2)v)inletapp=ifArray.is_emptylb_argsthenlb_type_of_pelsemkApp(lb_type_of_p,lb_args)inTacticals.tclTHENLIST[Equality.replacepq;applyapp;Auto.default_auto]end(* used in the bool -> leb side *)letdo_replace_blhandle(ind,uasindu)aavoidnarglftrgt=letopenEConstrinletavoid=Array.of_listaavoidinletdo_argenvsigmahdvoffset=matchkindsigmavwith|Vars->letx=narg*offsetinletn=Array.lengthavoidinletrecfindi=ifId.equalavoid.(n-i)sthenavoid.(n-i-x)else(ifi<nthenfind(i+1)elseuser_err(str"Var "++Id.prints++str" seems unknown."))inmkVar(find1)|Const(cst,_)->(* Works in specific situations where the args have to be already declared as a
Parameter (see example "J" in test file SchemeEquality.v) *)letlbl=Label.to_string(Constant.labelcst)inletnewlbl=ifInt.equaloffset1then("eq_"^lbl)else(lbl^"_bl")inletnewcst=Constant.change_labelcst(Label.makenewlbl)inifEnviron.mem_constantnewcstenvthenmkConstnewcstelseraise(ConstructorWithNonParametricInductiveType(fsthd))|_->raise(ConstructorWithNonParametricInductiveType(fsthd))inletrecauxl1l2=match(l1,l2)with|(t1::q1,t2::q2)->Proofview.Goal.enterbeginfungl->letsigma=Tacmach.projectglinletenv=Tacmach.pf_envglinifEConstr.eq_constrsigmat1t2thenauxq1q2else(lettt1=Tacmach.pf_get_type_ofglt1inletu,v=trydestruct_indenvsigmatt1(* trick so that the good sequence is returned*)withewhenCErrors.noncriticale->indu,[||]inifInd.CanOrd.equal(fstu)indthenTacticals.tclTHENLIST[Equality.replacet1t2;Auto.default_auto;auxq1q2]else(letc=get_schemehandle(!bl_scheme_kind_aux())(fstu)inletbl_t1=mkConstcinletbl_args=Array.append(Array.appendv(Array.Smart.map(funx->do_argenvsigmaux1)v))(Array.Smart.map(funx->do_argenvsigmaux2)v)inletapp=ifArray.is_emptybl_argsthenbl_t1elsemkApp(bl_t1,bl_args)inTacticals.tclTHENLIST[Equality.replace_byt1t2(Tacticals.tclTHEN(applyapp)(Auto.default_auto));auxq1q2]))end|([],[])->Proofview.tclUNIT()|_->Tacticals.tclZEROMSG(str"Both side of the equality must have the same arity.")inProofview.tclEVARMAP>>=funsigma->begintryProofview.tclUNIT(destAppsigmalft)withDestKO->Tacticals.tclZEROMSG(str"replace failed.")end>>=fun(ind1,ca1)->begintryProofview.tclUNIT(destAppsigmargt)withDestKO->Tacticals.tclZEROMSG(str"replace failed.")end>>=fun(ind2,ca2)->begintryProofview.tclUNIT(fst(destIndsigmaind1))withDestKO->begintryProofview.tclUNIT(fst(fst(destConstructsigmaind1)))withDestKO->Tacticals.tclZEROMSG(str"The expected type is an inductive one.")endend>>=fun(sp1,i1)->begintryProofview.tclUNIT(fst(destIndsigmaind2))withDestKO->begintryProofview.tclUNIT(fst(fst(destConstructsigmaind2)))withDestKO->Tacticals.tclZEROMSG(str"The expected type is an inductive one.")endend>>=fun(sp2,i2)->Proofview.tclENV>>=funenv->ifnot(Environ.QMutInd.equalenvsp1sp2)||not(Int.equali1i2)thenTacticals.tclZEROMSG(str"Eq should be on the same type")elseaux(Array.to_listca1)(Array.to_listca2)(*
create, from a list of ids [i1,i2,...,in] the list
[(in,eq_in,in_bl,in_al),,...,(i1,eq_i1,i1_bl_i1_al )]
*)letlist_idl=List.fold_left(funadecl->lets'=matchRelDecl.get_namedeclwithNames->Id.to_strings|Anonymous->"A"in(Id.of_strings',Id.of_string("eq_"^s'),Id.of_string(s'^"_bl"),Id.of_string(s'^"_lb"))::a)[]lletavoid_of_list_idlist_id=List.fold_left(funavoid(s,seq,sbl,slb)->List.fold_left(funavoidid->Id.Set.addidavoid)avoid[s;seq;sbl;slb])Id.Set.emptylist_id(*
build the right eq_I A B.. N eq_A .. eq_N
*)leteqIhandleindlist_id=leteA=Array.of_list((List.map(fun(s,_,_,_)->mkVars)list_id)@(List.map(fun(_,seq,_,_)->mkVarseq)list_id))ande=mkConst(get_schemehandlebeq_scheme_kindind)inmkApp(e,eA)(**********************************************************************)(* Boolean->Leibniz *)openNamegenletcompute_bl_goalenvhandleindlnamesparrecnparrec=letlist_id=list_idlnamesparrecinleteqI=eqIhandleindlist_idinletavoid=avoid_of_list_idlist_idinletx=next_ident_away(Id.of_string"x")avoidinlety=next_ident_away(Id.of_string"y")(Id.Set.addxavoid)inletcreate_inputc=letbl_typ=List.map(fun(s,seq,_,_)->mkNamedProd(make_annotxSorts.Relevant)(mkVars)(mkNamedProd(make_annotySorts.Relevant)(mkVars)(mkArrow(mkApp(eq(),[|bb();mkApp(mkVarseq,[|mkVarx;mkVary|]);tt()|]))Sorts.Relevant(mkApp(eq(),[|mkVars;mkVarx;mkVary|])))))list_idinletbl_input=List.fold_left2(funa(s,_,sbl,_)b->mkNamedProd(make_annotsblSorts.Relevant)ba)c(List.revlist_id)(List.revbl_typ)inleteqs_typ=List.map(fun(s,_,_,_)->mkProd(make_annotAnonymousSorts.Relevant,mkVars,mkProd(make_annotAnonymousSorts.Relevant,mkVars,(bb()))))list_idinleteq_input=List.fold_left2(funa(s,seq,_,_)b->mkNamedProd(make_annotseqSorts.Relevant)ba)bl_input(List.revlist_id)(List.reveqs_typ)inList.fold_left(funadecl->letx=map_annot(functionNames->s|Anonymous->next_ident_away(Id.of_string"A")avoid)(RelDecl.get_annotdecl)inmkNamedProdx(RelDecl.get_typedecl)a)eq_inputlnamesparrecinletu=Univ.Instance.emptyincreate_input(mkNamedProd(make_annotxSorts.Relevant)(mkFullIndenv(ind,u)nparrec)(mkNamedProd(make_annotySorts.Relevant)(mkFullIndenv(ind,u)(nparrec+1))(mkArrow(mkApp(eq(),[|bb();mkApp(eqI,[|mkVarx;mkVary|]);tt()|]))Sorts.Relevant(mkApp(eq(),[|mkFullIndenv(ind,u)(nparrec+3);mkVarx;mkVary|])))))letcompute_bl_tacthandleindlnamesparrecnparrec=letlist_id=list_idlnamesparrecinletfirst_intros=(List.map(fun(s,_,_,_)->s)list_id)@(List.map(fun(_,seq,_,_)->seq)list_id)@(List.map(fun(_,_,sbl,_)->sbl)list_id)inintros_using_thenfirst_introsbeginfunfresh_first_intros->Tacticals.tclTHENLIST[intro_using_then(Id.of_string"x")(funfreshn->induct_on(EConstr.mkVarfreshn));intro_using_then(Id.of_string"y")(funfreshm->destruct_on(EConstr.mkVarfreshm));intro_using_then(Id.of_string"Z")beginfunfreshz->Tacticals.tclTHENLIST[intros;Tacticals.tclTRY(Tacticals.tclORELSEreflexivitymy_discr_tac);simpl_in_hyp(freshz,Locus.InHyp);(*
repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]).
*)Tacticals.tclREPEAT(Tacticals.tclTHENLIST[Simple.apply_infreshz(EConstr.of_constr(andb_prop()));destruct_on_as(EConstr.mkVarfreshz)(IntroOrPattern[[CAst.make@@IntroNaming(IntroFresh(Id.of_string"Z"));CAst.make@@IntroNaming(IntroIdentifierfreshz)]])]);(*
Ci a1 ... an = Ci b1 ... bn
replace bi with ai; auto || replace bi with ai by apply typeofbi_prod ; auto
*)Proofview.Goal.enterbeginfungl->letconcl=Proofview.Goal.conclglinletsigma=Tacmach.projectglinmatchEConstr.kindsigmaconclwith|App(c,ca)->(matchEConstr.kindsigmacwith|Ind(indeq,u)->ifGlobRef.equal(GlobRef.IndRefindeq)Coqlib.(lib_ref"core.eq.type")thenTacticals.tclTHEN(do_replace_blhandleind(List.revfresh_first_intros)nparrec(ca.(2))(ca.(1)))Auto.default_autoelseTacticals.tclZEROMSG(str"Failure while solving Boolean->Leibniz.")|_->Tacticals.tclZEROMSG(str" Failure while solving Boolean->Leibniz."))|_->Tacticals.tclZEROMSG(str"Failure while solving Boolean->Leibniz.")end]end]endletmake_bl_schemeenvhandlemind=letmib=Environ.lookup_mindmindenvinifnot(Int.equal(Array.lengthmib.mind_packets)1)thenuser_err(str"Automatic building of boolean->Leibniz lemmas not supported");letind=(mind,0)inletnparams=mib.mind_nparamsinletnparrec=mib.mind_nparams_recinletlnonparrec,lnamesparrec=(* TODO subst *)context_chop(nparams-nparrec)mib.mind_params_ctxtinletbl_goal=compute_bl_goalenvhandleindlnamesparrecnparrecinletuctx=UState.from_envenvinletbl_goal=EConstr.of_constrbl_goalinlet(ans,_,_,_,ctx)=Declare.build_by_tactic~poly:false~side_eff:falseenv~uctx~typ:bl_goal(compute_bl_tacthandle(ind,EConstr.EInstance.empty)lnamesparrecnparrec)in([|ans|],ctx)letmake_bl_scheme_depsenvind=letinds=get_inductive_depsenvindinletmapind=SchemeMutualDep(ind,!bl_scheme_kind_aux())inSchemeMutualDep(ind,beq_scheme_kind)::List.mapmapindsletbl_scheme_kind=declare_mutual_scheme_object"_dec_bl"~deps:make_bl_scheme_depsmake_bl_schemelet_=bl_scheme_kind_aux:=fun()->bl_scheme_kind(**********************************************************************)(* Leibniz->Boolean *)letcompute_lb_goalenvhandleindlnamesparrecnparrec=letlist_id=list_idlnamesparrecinleteq=eq()andtt=tt()andbb=bb()inletavoid=avoid_of_list_idlist_idinleteqI=eqIhandleindlist_idinletx=next_ident_away(Id.of_string"x")avoidinlety=next_ident_away(Id.of_string"y")(Id.Set.addxavoid)inletcreate_inputc=letlb_typ=List.map(fun(s,seq,_,_)->mkNamedProd(make_annotxSorts.Relevant)(mkVars)(mkNamedProd(make_annotySorts.Relevant)(mkVars)(mkArrow(mkApp(eq,[|mkVars;mkVarx;mkVary|]))Sorts.Relevant(mkApp(eq,[|bb;mkApp(mkVarseq,[|mkVarx;mkVary|]);tt|])))))list_idinletlb_input=List.fold_left2(funa(s,_,_,slb)b->mkNamedProd(make_annotslbSorts.Relevant)ba)c(List.revlist_id)(List.revlb_typ)inleteqs_typ=List.map(fun(s,_,_,_)->mkProd(make_annotAnonymousSorts.Relevant,mkVars,mkProd(make_annotAnonymousSorts.Relevant,mkVars,bb)))list_idinleteq_input=List.fold_left2(funa(s,seq,_,_)b->mkNamedProd(make_annotseqSorts.Relevant)ba)lb_input(List.revlist_id)(List.reveqs_typ)inList.fold_left(funadecl->letx=map_annot(functionNames->s|Anonymous->Id.of_string"A")(RelDecl.get_annotdecl)inmkNamedProdx(RelDecl.get_typedecl)a)eq_inputlnamesparrecinletu=Univ.Instance.emptyincreate_input(mkNamedProd(make_annotxSorts.Relevant)(mkFullIndenv(ind,u)nparrec)(mkNamedProd(make_annotySorts.Relevant)(mkFullIndenv(ind,u)(nparrec+1))(mkArrow(mkApp(eq,[|mkFullIndenv(ind,u)(nparrec+2);mkVarx;mkVary|]))Sorts.Relevant(mkApp(eq,[|bb;mkApp(eqI,[|mkVarx;mkVary|]);tt|])))))letcompute_lb_tacthandleindlnamesparrecnparrec=letlist_id=list_idlnamesparrecinletfirst_intros=(List.map(fun(s,_,_,_)->s)list_id)@(List.map(fun(_,seq,_,_)->seq)list_id)@(List.map(fun(_,_,_,slb)->slb)list_id)inintros_using_thenfirst_introsbeginfunfresh_first_intros->Tacticals.tclTHENLIST[intro_using_then(Id.of_string"x")(funfreshn->induct_on(EConstr.mkVarfreshn));intro_using_then(Id.of_string"y")(funfreshm->destruct_on(EConstr.mkVarfreshm));intro_using_then(Id.of_string"Z")beginfunfreshz->Tacticals.tclTHENLIST[intros;Tacticals.tclTRY(Tacticals.tclORELSEreflexivitymy_discr_tac);my_inj_tacfreshz;intros;simpl_in_concl;Auto.default_auto;Tacticals.tclREPEAT(Tacticals.tclTHENLIST[apply(EConstr.of_constr(andb_true_intro()));simplest_split;Auto.default_auto]);Proofview.Goal.enterbeginfungls->letconcl=Proofview.Goal.conclglsinletsigma=Tacmach.projectglsin(* assume the goal to be eq (eq_type ...) = true *)matchEConstr.kindsigmaconclwith|App(c,ca)->(match(EConstr.kindsigmaca.(1))with|App(c',ca')->letn=Array.lengthca'indo_replace_lbhandle(List.revfresh_first_intros)nparrecca'.(n-2)ca'.(n-1)|_->Tacticals.tclZEROMSG(str"Failure while solving Leibniz->Boolean."))|_->Tacticals.tclZEROMSG(str"Failure while solving Leibniz->Boolean.")end]end]endletmake_lb_schemeenvhandlemind=letmib=Environ.lookup_mindmindenvinifnot(Int.equal(Array.lengthmib.mind_packets)1)thenuser_err(str"Automatic building of Leibniz->boolean lemmas not supported");letind=(mind,0)inletnparams=mib.mind_nparamsinletnparrec=mib.mind_nparams_recinletlnonparrec,lnamesparrec=context_chop(nparams-nparrec)mib.mind_params_ctxtinletlb_goal=compute_lb_goalenvhandleindlnamesparrecnparrecinletuctx=UState.from_envenvinletlb_goal=EConstr.of_constrlb_goalinlet(ans,_,_,_,ctx)=Declare.build_by_tactic~poly:false~side_eff:falseenv~uctx~typ:lb_goal(compute_lb_tacthandleindlnamesparrecnparrec)in([|ans|],ctx)letmake_lb_scheme_depsenvind=letinds=get_inductive_depsenvindinletmapind=SchemeMutualDep(ind,!lb_scheme_kind_aux())inSchemeMutualDep(ind,beq_scheme_kind)::List.mapmapindsletlb_scheme_kind=declare_mutual_scheme_object"_dec_lb"~deps:make_lb_scheme_depsmake_lb_schemelet_=lb_scheme_kind_aux:=fun()->lb_scheme_kind(**********************************************************************)(* Decidable equality *)letcheck_not_is_defined()=ifnot(Coqlib.has_ref"core.not.type")thenraise(UndefinedCst"not")(* {n=m}+{n<>m} part *)letcompute_dec_goalenvindlnamesparrecnparrec=check_not_is_defined();leteq=eq()andtt=tt()andbb=bb()inletlist_id=list_idlnamesparrecinletavoid=avoid_of_list_idlist_idinletx=next_ident_away(Id.of_string"x")avoidinlety=next_ident_away(Id.of_string"y")(Id.Set.addxavoid)inletcreate_inputc=letlb_typ=List.map(fun(s,seq,_,_)->mkNamedProd(make_annotxSorts.Relevant)(mkVars)(mkNamedProd(make_annotySorts.Relevant)(mkVars)(mkArrow(mkApp(eq,[|mkVars;mkVarx;mkVary|]))Sorts.Relevant(mkApp(eq,[|bb;mkApp(mkVarseq,[|mkVarx;mkVary|]);tt|])))))list_idinletbl_typ=List.map(fun(s,seq,_,_)->mkNamedProd(make_annotxSorts.Relevant)(mkVars)(mkNamedProd(make_annotySorts.Relevant)(mkVars)(mkArrow(mkApp(eq,[|bb;mkApp(mkVarseq,[|mkVarx;mkVary|]);tt|]))Sorts.Relevant(mkApp(eq,[|mkVars;mkVarx;mkVary|])))))list_idinletlb_input=List.fold_left2(funa(s,_,_,slb)b->mkNamedProd(make_annotslbSorts.Relevant)ba)c(List.revlist_id)(List.revlb_typ)inletbl_input=List.fold_left2(funa(s,_,sbl,_)b->mkNamedProd(make_annotsblSorts.Relevant)ba)lb_input(List.revlist_id)(List.revbl_typ)inleteqs_typ=List.map(fun(s,_,_,_)->mkProd(make_annotAnonymousSorts.Relevant,mkVars,mkProd(make_annotAnonymousSorts.Relevant,mkVars,bb)))list_idinleteq_input=List.fold_left2(funa(s,seq,_,_)b->mkNamedProd(make_annotseqSorts.Relevant)ba)bl_input(List.revlist_id)(List.reveqs_typ)inList.fold_left(funadecl->letx=map_annot(functionNames->s|Anonymous->Id.of_string"A")(RelDecl.get_annotdecl)inmkNamedProdx(RelDecl.get_typedecl)a)eq_inputlnamesparrecinleteqnm=mkApp(eq,[|mkFullIndenvind(2*nparrec+2);mkVarx;mkVary|])increate_input(mkNamedProd(make_annotxSorts.Relevant)(mkFullIndenvind(2*nparrec))(mkNamedProd(make_annotySorts.Relevant)(mkFullIndenvind(2*nparrec+1))(mkApp(sumbool(),[|eqnm;mkApp(UnivGen.constr_of_monomorphic_global(Global.env())@@Coqlib.lib_ref"core.not.type",[|eqnm|])|]))))letcompute_dec_tacthandleindlnamesparrecnparrec=leteq=eq()andtt=tt()andff=ff()andbb=bb()inletlist_id=list_idlnamesparrecinlet_=get_schemehandlebeq_scheme_kindindin(* This is just an assertion? *)let_non_fresh_eqI=eqIhandleindlist_idinleteqtruex=mkApp(eq,[|bb;x;tt|])inleteqfalsex=mkApp(eq,[|bb;x;ff|])inletfirst_intros=(List.map(fun(s,_,_,_)->s)list_id)@(List.map(fun(_,seq,_,_)->seq)list_id)@(List.map(fun(_,_,sbl,_)->sbl)list_id)@(List.map(fun(_,_,_,slb)->slb)list_id)inletfresh_idsgl=fresh_id_in_env(Id.Set.empty)s(Proofview.Goal.envgl)inintros_using_thenfirst_introsbeginfunfresh_first_intros->leteqI=leta=Array.of_listfresh_first_introsinletn=List.lengthlist_idinassert(Int.equal(Array.lengtha)(4*n));letfresh_list_id=List.initn(funi->(Array.getai,Array.geta(i+n),Array.geta(i+2*n),Array.geta(i+3*n)))ineqIhandleindfresh_list_idinintro_using_then(Id.of_string"x")beginfunfreshn->intro_using_then(Id.of_string"y")beginfunfreshm->Proofview.Goal.enterbeginfungl->letfreshH=fresh_id(Id.of_string"H")glinleteqbnm=mkApp(eqI,[|mkVarfreshn;mkVarfreshm|])inletarfresh=Array.of_listfresh_first_introsinletxargs=Array.subarfresh0(2*nparrec)inletc=get_schemehandlebl_scheme_kindindinletblI=mkConstcinletc=get_schemehandlelb_scheme_kindindinletlbI=mkConstcinTacticals.tclTHENLIST[(*we do this so we don't have to prove the same goal twice *)assert_by(NamefreshH)(EConstr.of_constr(mkApp(sumbool(),[|eqtrueeqbnm;eqfalseeqbnm|])))(Tacticals.tclTHEN(destruct_on(EConstr.of_constreqbnm))Auto.default_auto);Proofview.Goal.enterbeginfungl->letfreshH2=fresh_id(Id.of_string"H")glinTacticals.tclTHENS(destruct_on_using(EConstr.mkVarfreshH)freshH2)[(* left *)Tacticals.tclTHENLIST[simplest_left;apply(EConstr.of_constr(mkApp(blI,Array.mapmkVarxargs)));Auto.default_auto];(*right *)Proofview.Goal.enterbeginfungl->letfreshH3=fresh_id(Id.of_string"H")glinTacticals.tclTHENLIST[simplest_right;unfold_constr(Coqlib.lib_ref"core.not.type");intro;Equality.subst_all();assert_by(NamefreshH3)(EConstr.of_constr(mkApp(eq,[|bb;mkApp(eqI,[|mkVarfreshm;mkVarfreshm|]);tt|])))(Tacticals.tclTHENLIST[apply(EConstr.of_constr(mkApp(lbI,Array.mapmkVarxargs)));Auto.default_auto]);Equality.general_rewrite~where:(SomefreshH3)~l2r:trueLocus.AllOccurrences~freeze:true~dep:false~with_evars:true((EConstr.mkVarfreshH2),NoBindings);my_discr_tac]end]end]endendendendletmake_eq_decidabilityenvhandlemind=letmib=Environ.lookup_mindmindenvinifnot(Int.equal(Array.lengthmib.mind_packets)1)thenraiseDecidabilityMutualNotSupported;letind=(mind,0)inletnparams=mib.mind_nparamsinletnparrec=mib.mind_nparams_recinletu=Univ.Instance.emptyinletuctx=UState.from_envenvinletlnonparrec,lnamesparrec=context_chop(nparams-nparrec)mib.mind_params_ctxtinlet(ans,_,_,_,ctx)=Declare.build_by_tactic~poly:false~side_eff:falseenv~uctx~typ:(EConstr.of_constr(compute_dec_goalenv(ind,u)lnamesparrecnparrec))(compute_dec_tacthandleindlnamesparrecnparrec)in([|ans|],ctx)leteq_dec_scheme_kind=declare_mutual_scheme_object"_eq_dec"~deps:(fun_ind->[SchemeMutualDep(ind,bl_scheme_kind);SchemeMutualDep(ind,lb_scheme_kind)])make_eq_decidability(* The eq_dec_scheme proofs depend on the equality and discr tactics
but the inj tactics, that comes with discr, depends on the
eq_dec_scheme... *)let_=Equality.set_eq_dec_scheme_kindeq_dec_scheme_kind