123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585(************************************************************************)(* * 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) *)(************************************************************************)openPpopenCErrorsopenUtilopenNamesopenTermopenTermopsopenConstropenContextopenEConstropenVarsopenNamegenopenInductiveopsopenPrinteropenRetypingopenTacmachopenTacticalsopenTacticsopenElimopenEqualityopenTactypesopenProofview.NotationsmoduleNamedDecl=Context.Named.Declarationletvar_occurs_in_pfglid=letenv=Proofview.Goal.envglinletsigma=projectglinoccur_varenvsigmaid(Proofview.Goal.conclgl)||List.exists(occur_var_in_declenvsigmaid)(Proofview.Goal.hypsgl)(* [make_inv_predicate (ity,args) C]
is given the inductive type, its arguments, both the global
parameters and its local arguments, and is expected to produce a
predicate P such that if largs is the "local" part of the
arguments, then (P largs) will be convertible with a conclusion of
the form:
<A1>a1=a1-><A2>a2=a2 ... -> C
Algorithm: suppose length(largs)=n
(1) Push the entire arity, [xbar:Abar], carrying along largs and
the conclusion
(2) Pair up each ai with its respective Rel version: a1==(Rel n),
a2==(Rel n-1), etc.
(3) For each pair, ai,Rel j, if the Ai is dependent - that is, the
type of [Rel j] is an open term, then we construct the iterated
tuple, [make_iterated_tuple] does it, and use that for our equation
Otherwise, we just use <Ai>ai=Rel j
*)typeinversion_status=Depofconstroption|NoDeptypeinversion_kind=|SimpleInversion|FullInversion|FullInversionClearletevd_comb1fevdrefx=let(evd',y)=f!evdrefxinevdref:=evd';yletcompute_eqnenvsigmaniai=(mkRel(n-i),get_type_ofenvsigma(mkRel(n-i)))letmake_inv_predicateenvevdindfrealargsidstatusconcl=letnrealargs=List.lengthrealargsinlet(hyps,concl)=matchstatuswith|NoDep->(* We push the arity and leave concl unchanged *)lethyps_arity=get_arityenvindfin(hyps_arity,concl)|Depdflt_concl->ifnot(occur_varenv!evdidconcl)thenuser_err(str"Current goal does not depend on "++Id.printid++str".");(* We abstract the conclusion of goal with respect to
realargs and c to * be concl in order to rewrite and have
c also rewritten when the case * will be done *)letpred=matchdflt_conclwith|Someconcl->concl(*assumed it's some [x1..xn,H:I(x1..xn)]C*)|None->letsort=get_sort_family_ofenv!evdconclinletsort=evd_comb1Evd.fresh_sort_in_familyevdsortinletp=make_arityenv!evdtrueindfsortinletevd',(p,ptyp)=Unification.abstract_list_allenv!evdpconcl(realargs@[mkVarid])inevd:=evd';pin(* abstract_list_all returns only lambda *)lethyps,bodypred=decompose_lambda_n_assum!evd(nrealargs+1)predin(* We lift to make room for the equations *)(hyps,liftnrealargsbodypred)inletnhyps=Context.Rel.lengthhypsinletenv'=push_rel_contexthypsenvin(* Now the arity is pushed, and we need to construct the pairs
* ai,mkRel(n-i+1) *)(* Now, we can recurse down this list, for each ai,(mkRel k) whether to
push <Ai>(mkRel k)=ai (when Ai is closed).
In any case, we carry along the rest of pairs *)leteq_term,refl_term=tryCoqlib.lib_ref"core.eq.type",Coqlib.lib_ref"core.eq.refl"withCoqlib.NotFoundRef_->tryCoqlib.lib_ref"core.identity.type",Coqlib.lib_ref"core.identity.refl"withCoqlib.NotFoundRef_->user_err(str"No registered equality"++spc()++hov1(str"(needs \"core.eq.type\" and \"core.eq.refl\""++spc()++str"or \"core.identity.type\" and \"core.identy.refl\")."))inletrecbuild_concleqnsargsn=function|[]->it_mkProdconcleqns,Array.rev_of_listargs|ai::restlist->letai=liftnhypsaiinlet(xi,ti)=compute_eqnenv'!evdnhypsnaiinlet(lhs,eqnty,rhs)=ifclosed0!evdtithen(* shortcut *)(xi,ti,ai)elseletopenCombinatorsinletsigma,{telescope_value;telescope_type},default_value=make_iterated_tupleenv'!evd~default:aixitiinevd:=sigma;telescope_value,telescope_type,default_valueinleteq=evd_comb1(Evd.fresh_globalenv)evdeq_terminleteqn=applist(eq,[eqnty;lhs;rhs])inletrefl_term=evd_comb1(Evd.fresh_globalenv)evdrefl_terminletrefl=mkApp(refl_term,[|eqnty;rhs|])inletr=Retyping.relevance_of_termenv!evdreflinleteqns=(make_annotAnonymousr,liftneqn)::eqnsinlet_=evd_comb1(Typing.type_ofenv)evdreflinletargs=refl::argsinbuild_concleqnsargs(succn)restlistinlet(newconcl,args)=build_concl[][]0realargsinletpredicate=it_mkLambda_or_LetInnewconcl(name_contextenv!evdhyps)inlet_=evd_comb1(Typing.type_ofenv)evdpredicatein(* OK - this predicate should now be usable by res_elimination_then to
do elimination on the conclusion. *)predicate,args(* The result of the elimination is a bunch of goals like:
|- (cibar:Cibar)Equands->C
where the cibar are either dependent or not. We are fed a
signature, with "true" for every recursive argument, and false for
every non-recursive one. So we need to do the
sign_branch_len(sign) intros, thinning out all recursive
assumptions. This leaves us with exactly length(sign) assumptions.
We save their names, and then do introductions for all the equands
(there are some number of them, which is the other argument of the
tactic)
This gives us the #neqns equations, whose names we get also, and
the #length(sign) arguments.
Suppose that #nodep of these arguments are non-dependent.
Generalize and thin them.
This gives us #dep = #length(sign)-#nodep arguments which are
dependent.
Now, we want to take each of the equations, and do all possible
injections to get the left-hand-side to be a variable. At the same
time, if we find a lhs/rhs pair which are different, we can
discriminate them to prove false and finish the branch.
Then, we thin away the equations, and do the introductions for the
#nodep arguments which we generalized before.
*)(* Called after the case-assumptions have been killed off, and all the
intros have been done. Given that the clause in question is an
equality (if it isn't we fail), we are responsible for projecting
the equality, using Injection and Discriminate, and applying it to
the concusion *)(* Computes the subset of hypothesis in the local context whose
type depends on t (should be of the form (mkVar id)), then
it generalizes them, applies tac to rewrite all occurrencies of t,
and introduces generalized hypotheis.
Precondition: t=(mkVar id) *)letdependent_hypsenvididlistgl=letrecdep_rec=function|[]->[]|d::l->(* Update the type of id1: it may have been subject to rewriting *)letd=pf_get_hyp(NamedDecl.get_idd)glinifoccur_var_in_declenv(projectgl)iddthend::dep_reclelsedep_reclindep_recidlistletsplit_dep_and_nodephypsgl=List.fold_right(fund(l1,l2)->ifvar_occurs_in_pfgl(NamedDecl.get_idd)then(d::l1,l2)else(l1,d::l2))hyps([],[])(* Computation of dids is late; must have been done in rewrite_equations*)(* Will keep generalizing and introducing back and forth... *)(* Moreover, others hyps depending of dids should have been *)(* generalized; in such a way that [dids] can endly be cleared *)(* Consider for instance this case extracted from Well_Ordering.v
A : Set
B : A ->Set
a0 : A
f : (B a0) ->WO
y : WO
H0 : (le_WO y (sup a0 f))
============================
(Acc WO le_WO y)
Inversion H0 gives
A : Set
B : A ->Set
a0 : A
f : (B a0) ->WO
y : WO
H0 : (le_WO y (sup a0 f))
a1 : A
f0 : (B a1) ->WO
v : (B a1)
H1 : (f0 v)=y
H3 : a1=a0
f1 : (B a0) ->WO
v0 : (B a0)
H4 : (existS A [a:A](B a) ->WO a0 f1)=(existS A [a:A](B a) ->WO a0 f)
============================
(Acc WO le_WO (f1 v0))
while, ideally, we would have expected
A : Set
B : A ->Set
a0 : A
f0 : (B a0)->WO
v : (B a0)
============================
(Acc WO le_WO (f0 v))
obtained from destruction with equalities
A : Set
B : A ->Set
a0 : A
f : (B a0) ->WO
y : WO
H0 : (le_WO y (sup a0 f))
a1 : A
f0 : (B a1)->WO
v : (B a1)
H1 : (f0 v)=y
H2 : (sup a1 f0)=(sup a0 f)
============================
(Acc WO le_WO (f0 v))
by clearing initial hypothesis H0 and its dependency y, clearing H1
(in fact H1 can be avoided using the same trick as for newdestruct),
decomposing H2 to get a1=a0 and (a1,f0)=(a0,f), replacing a1 by a0
everywhere and removing a1 and a1=a0 (in fact it would have been more
regular to replace a0 by a1, avoiding f1 and v0 cannot replace f0 and v),
finally removing H4 (here because f is not used, more generally after using
eq_dep and replacing f by f0) [and finally rename a0, f0 into a,f].
Summary: nine useless hypotheses!
Nota: with Inversion_clear, only four useless hypotheses
*)letgeneralizeRewriteIntrosas_modetacdepidsid=Proofview.tclENV>>=funenv->Proofview.Goal.enterbeginfungl->letdids=dependent_hypsenviddepidsglinletreintros=ifas_modethenintros_replacingelseintros_possibly_replacingin(tclTHENLIST[Generalize.bring_hypsdids;tac;(* may actually fail to replace if dependent in a previous eq *)reintros(ids_of_named_contextdids)])endleterror_too_many_namespats=letloc=Loc.merge_opt(List.hdpats).CAst.loc(List.lastpats).CAst.locinProofview.tclENV>>=funenv->Proofview.tclEVARMAP>>=funsigma->tclZEROMSG?loc(str"Unexpected "++str(String.plural(List.lengthpats)"introduction pattern")++str": "++pr_enum(Miscprint.pr_intro_pattern(func->Printer.pr_econstr_envenvsigma(snd(cenv(Evd.from_envenv)))))pats++str".")letget_names(allow_conj,issimple)({CAst.loc;v=pat}asx)=matchpatwith|IntroNamingIntroAnonymous|IntroForthcoming_->user_errPp.(str"Anonymous pattern not allowed for inversion equations.")|IntroNaming(IntroFresh_)->user_errPp.(str"Fresh pattern not allowed for inversion equations.")|IntroActionIntroWildcard->user_errPp.(str"Discarding pattern not allowed for inversion equations.")|IntroAction(IntroRewrite_)->user_errPp.(str"Rewriting pattern not allowed for inversion equations.")|IntroAction(IntroOrAndPattern(IntroAndPattern[]))whenallow_conj->(None,[])|IntroAction(IntroOrAndPattern(IntroAndPattern({CAst.v=IntroNaming(IntroIdentifierid)}::_asl)|IntroOrPattern[{CAst.v=IntroNaming(IntroIdentifierid)}::_asl]))whenallow_conj->(Someid,l)|IntroAction(IntroOrAndPattern(IntroAndPattern_))->ifissimplethenuser_errPp.(str"Conjunctive patterns not allowed for simple inversion equations.")elseuser_errPp.(str"Nested conjunctive patterns not allowed for inversion equations.")|IntroAction(IntroInjectionl)->user_errPp.(str"Injection patterns not allowed for inversion equations.")|IntroAction(IntroOrAndPattern(IntroOrPattern_))->user_errPp.(str"Disjunctive patterns not allowed for inversion equations.")|IntroAction(IntroApplyOn(c,pat))->user_errPp.(str"Apply patterns not allowed for inversion equations.")|IntroNaming(IntroIdentifierid)->(Someid,[x])letrectclMAP_iallow_conjntacfun=function|[]->tclDOn(tacfun(None,[]))|a::lasl'->ifInt.equaln0thenerror_too_many_namesl'elsetclTHEN(tacfun(get_namesallow_conja))(tclMAP_iallow_conj(n-1)tacfunl)letremember_first_eqidx=if!x==Logic.MoveLastthenx:=Logic.MoveAfterid(* invariant: ProjectAndApply is responsible for erasing the clause
which it is given as input
It simplifies the clause (an equality) to use it as a rewrite rule and then
erases the result of the simplification. *)(* invariant: ProjectAndApplyNoThining simplifies the clause (an equality) .
If it can discriminate then the goal is proved, if not tries to use it as
a rewrite rule. It erases the clause which is given as input *)letdest_nf_eqenvsigmat=matchEConstr.kindsigmatwith|App(r,[|t;x;y|])->letopenReductionopsinletis_global_existsgrc=matchCoqlib.lib_ref_optgrwith|Somegr->isRefXenvsigmagrc|None->falseinletis_eq=is_global_exists"core.eq.type"rinletis_identity=is_global_exists"core.identity.type"rinifis_eq||is_identitythen(t,whd_allenvsigmax,whd_allenvsigmay)elseuser_errPp.(str"Not an equality.")|_->user_errPp.(str"Not an equality.")letprojectAndApplyas_modethinavoidideqnamenamesdepids=letsubst_hypl2rid=tclTHEN(tclTRY(rewriteInConcll2r(EConstr.mkVarid)))(ifthinthenclear[id]else(remember_first_eqideqname;tclIDTAC))inletsubstHypIfVariabletacid=Proofview.Goal.enterbeginfungl->letsigma=projectglin(* We only look at the type of hypothesis "id" *)lethyp=pf_nf_evargl(pf_get_hyp_typidgl)inlet(t,t1,t2)=dest_nf_eq(pf_envgl)sigmahypinmatch(EConstr.kindsigmat1,EConstr.kindsigmat2)with|Varid1,_->generalizeRewriteIntrosas_mode(subst_hyptrueid)depidsid1|_,Varid2->generalizeRewriteIntrosas_mode(subst_hypfalseid)depidsid2|_->tacidendinletdeq_traileridneqns=tclTHENLIST[ifas_modethenclear[id]elsetclIDTAC;(tclMAP_i(false,false)neqns(function(idopt,_)->tclTRY(tclTHEN(intro_move_avoididoptavoidLogic.MoveLast)(* try again to substitute and if still not a variable after *)(* decomposition, arbitrarily try to rewrite RL !? *)(tclTRY(onLastHypId(substHypIfVariable(funid->subst_hypfalseid))))))names);(ifas_modethentclIDTACelseclear[id])](* Doing the above late breaks the computation of dids in
generalizeRewriteIntros, and hence breaks proper intros_replacing
but it is needed for compatibility *)insubstHypIfVariable(* If no immediate variable in the equation, try to decompose it *)(* and apply a trailer which again try to substitute *)(funid->dEqThen~keep_proofs:Nonefalse(deq_trailerid)(Some(None,ElimOnConstr(EConstr.mkVarid,NoBindings))))idletnLastDeclsitac=Proofview.Goal.enterbeginfungl->tac(nLastDeclsgli)end(* Introduction of the equations on arguments
othin: discriminates Simple Inversion, Inversion and Inversion_clear
None: the equations are introduced, but not rewritten
Some thin: the equations are rewritten, and cleared if thin is true *)letrewrite_equationsas_modeothinneqnsnamesba=Proofview.Goal.enterbeginfungl->let(depids,nodepids)=split_dep_and_nodepbaglinletfirst_eq=refLogic.MoveLastinletavoid=ifas_modethenId.Set.of_list(List.mapNamedDecl.get_idnodepids)elseId.Set.emptyinmatchothinwith|Somethin->tclTHENLIST[tclDOneqnsintro;Generalize.bring_hypsnodepids;clear(ids_of_named_contextnodepids);(nLastDeclsneqns(functx->Generalize.bring_hyps(List.revctx)));(nLastDeclsneqns(functx->clear(ids_of_named_contextctx)));tclMAP_i(true,false)neqns(fun(idopt,names)->(tclTHEN(intro_move_avoididoptavoidLogic.MoveLast)(onLastHypId(funid->tclTRY(projectAndApplyas_modethinavoididfirst_eqnamesdepids)))))names;tclMAP(fund->tclIDTAC>>=fun()->(* delay for [first_eq]. *)letidopt=ifas_modethenSome(NamedDecl.get_idd)elseNoneinintro_moveidopt(ifthinthenLogic.MoveLastelse!first_eq))nodepids;(tclMAP(fund->tclTRY(clear[NamedDecl.get_idd]))depids)]|None->(* simple inversion *)ifas_modethentclMAP_i(false,true)neqns(fun(idopt,_)->intro_moveidoptLogic.MoveLast)nameselse(tclTHENLIST[tclDOneqnsintro;Generalize.bring_hypsnodepids;clear(ids_of_named_contextnodepids)])endletinterp_inversion_kind=function|SimpleInversion->None|FullInversion->Somefalse|FullInversionClear->Sometrueletrewrite_equations_tacas_modeothinidneqnsnamesba=letothin=interp_inversion_kindothininlettac=rewrite_equationsas_modeothinneqnsnamesbainmatchothinwith|Sometrue(* if Inversion_clear, clear the hypothesis *)->tclTHENtac(tclTRY(clear[id]))|_->tacletraw_inversioninv_kindidstatusnames=Proofview.Goal.enterbeginfungl->letsigma=Proofview.Goal.sigmaglinletenv=Proofview.Goal.envglinletconcl=Proofview.Goal.conclglinletc=mkVaridinlet((ind,u),t)=trypf_applyTacred.reduce_to_atomic_indgl(pf_get_type_ofglc)withUserError_->letmsg=str"The type of "++Id.printid++str" is not inductive."inCErrors.user_errmsginletIndType(indf,realargs)=find_rectypeenvsigmatinletevdref=refsigmainlet(elim_predicate,args)=make_inv_predicateenvevdrefindfrealargsidstatusconclinletsigma=!evdrefinletdep=status!=NoDep&&(local_occur_varsigmaidconcl)inletcut_concl=ifdepthenapplist(elim_predicate,realargs@[c])elseapplist(elim_predicate,realargs)inletrefinedid=letprf=mkApp(mkVarid,args)inRefine.refine~typecheck:false(funh->(h,prf))inletneqns=List.lengthrealargsinletas_mode=names!=Noneinlet(_,args)=decompose_appsigmatinletsolve_tac=(* We have to change again because assert_before performs βι-reduction *)change_conclcut_concl<*>case_tacdepnames(rewrite_equations_tacas_modeinv_kindidneqns)(ind,u,args)idintclTHEN(Proofview.Unsafe.tclEVARSsigma)(tclTHENS(assert_beforeAnonymouscut_concl)[solve_tac;onLastHypId(funid->tclTHEN(refinedid)reflexivity)])end(* Error messages of the inversion tactics *)letwrap_inv_errorid=function(e,info)->matchewith|Indrec.RecursionSchemeError(_,Indrec.NotAllowedCaseAnalysis(_,(Type_|Setask),i))->Proofview.tclENV>>=funenv->Proofview.tclEVARMAP>>=funsigma->tclZEROMSG((strbrk"Inversion would require case analysis on sort "++pr_sortsigmak++strbrk" which is not allowed for inductive definition "++pr_inductiveenv(fsti)++str"."))|e->Proofview.tclZERO~infoe(* The most general inversion tactic *)letinversioninv_kindstatusnamesid=Proofview.tclORELSE(raw_inversioninv_kindidstatusnames)(wrap_inv_errorid)(* Specializing it... *)letinv_genthinstatusnames=try_intros_until(inversionthinstatusnames)letinvk=inv_genkNoDepletinv_tacid=invFullInversionNone(NamedHyp(CAst.makeid))letinv_clear_tacid=invFullInversionClearNone(NamedHyp(CAst.makeid))letdinvkc=inv_genk(Depc)letdinv_tacid=dinvFullInversionNoneNone(NamedHyp(CAst.makeid))letdinv_clear_tacid=dinvFullInversionClearNoneNone(NamedHyp(CAst.makeid))(* InvIn will bring the specified clauses into the conclusion, and then
* perform inversion on the named hypothesis. After, it will intro them
* back to their places in the hyp-list. *)letinvInknamesidsid=Proofview.Goal.enterbeginfungl->lethyps=List.map(funid->pf_get_hypidgl)idsinletconcl=Proofview.Goal.conclglinletsigma=projectglinletnb_prod_init=nb_prodsigmaconclinletintros_replace_ids=Proofview.Goal.enterbeginfungl->letconcl=pf_conclglinletsigma=projectglinletnb_of_new_hyp=nb_prodsigmaconcl-(List.lengthhyps+nb_prod_init)inifnb_of_new_hyp<1thenintros_replacingidselsetclTHEN(tclDOnb_of_new_hypintro)(intros_replacingids)endinProofview.tclORELSE(tclTHENLIST[Generalize.bring_hypshyps;inversionkNoDepnamesid;intros_replace_ids])(wrap_inv_errorid)endletinvIn_genknamesidl=try_intros_until(invInknamesidl)letinv_clauseknames=function|[]->invknames|idl->invIn_genknamesidl