123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259(************************************************************************)(* * 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) *)(************************************************************************)openConstrletbt_lib_constrn=lazy(UnivGen.constr_of_monomorphic_global@@Coqlib.lib_refn)letdecomp_termsigma(c:Constr.t)=Constr.kind(EConstr.Unsafe.to_constr(Termops.strip_outer_castsigma(EConstr.of_constrc)))letlappcv=Constr.mkApp(Lazy.forcec,v)let(===)=Constr.equalmoduleCoqList=structlet_nil=bt_lib_constr"core.list.nil"let_cons=bt_lib_constr"core.list.cons"letconstyht=lapp_cons[|ty;h;t|]letnilty=lapp_nil[|ty|]letrecof_listty=function|[]->nilty|t::q->constyt(of_listtyq)endmoduleCoqPositive=structlet_xH=bt_lib_constr"num.pos.xH"let_xO=bt_lib_constr"num.pos.xO"let_xI=bt_lib_constr"num.pos.xI"(* A coq nat from an int *)letrecof_intn=ifn<=1thenLazy.force_xHelseletans=of_int(n/2)inifnmod2=0thenlapp_xO[|ans|]elselapp_xI[|ans|]endmoduleEnv=structmoduleConstrHashtbl=Hashtbl.Make(Constr)typet=(intConstrHashtbl.t*intref)letadd(tbl,off)(t:Constr.t)=tryConstrHashtbl.findtbltwith|Not_found->leti=!offinlet()=ConstrHashtbl.addtbltiinlet()=incroffiniletempty()=(ConstrHashtbl.create16,ref1)letto_list(env,_)=(* we need to get an ordered list *)letfoldconstrkeyaccu=(key,constr)::accuinletl=ConstrHashtbl.foldfoldenv[]inletsorted_l=List.sort(funp1p2->Int.compare(fstp1)(fstp2))linList.mapsndsorted_lendmoduleBool=structletind=lazy(Globnames.destIndRef(Coqlib.lib_ref"core.bool.type"))lettyp=bt_lib_constr"core.bool.type"lettrueb=bt_lib_constr"core.bool.true"letfalseb=bt_lib_constr"core.bool.false"letandb=bt_lib_constr"core.bool.andb"letorb=bt_lib_constr"core.bool.orb"letxorb=bt_lib_constr"core.bool.xorb"letnegb=bt_lib_constr"core.bool.negb"typet=|Varofint|Constofbool|Andboft*t|Orboft*t|Xorboft*t|Negboft|Ifboft*t*tletquote(env:Env.t)sigma(c:Constr.t):t=lettrueb=Lazy.forcetruebinletfalseb=Lazy.forcefalsebinletandb=Lazy.forceandbinletorb=Lazy.forceorbinletxorb=Lazy.forcexorbinletnegb=Lazy.forcenegbinletrecauxc=matchdecomp_termsigmacwith|App(head,args)->ifhead===andb&&Array.lengthargs=2thenAndb(auxargs.(0),auxargs.(1))elseifhead===orb&&Array.lengthargs=2thenOrb(auxargs.(0),auxargs.(1))elseifhead===xorb&&Array.lengthargs=2thenXorb(auxargs.(0),auxargs.(1))elseifhead===negb&&Array.lengthargs=1thenNegb(auxargs.(0))elseVar(Env.addenvc)|Case(info,_,_,_,_,arg,pats)->letis_bool=leti=info.ci_indinNames.Ind.CanOrd.equali(Lazy.forceind)inifis_boolthenIfb((auxarg),(aux(sndpats.(0))),(aux(sndpats.(1))))elseVar(Env.addenvc)|_->ifc===falsebthenConstfalseelseifc===truebthenConsttrueelseVar(Env.addenvc)inauxcendmoduleBtauto=structopenPpleteq=bt_lib_constr"core.eq.type"letf_var=bt_lib_constr"plugins.btauto.f_var"letf_btm=bt_lib_constr"plugins.btauto.f_btm"letf_top=bt_lib_constr"plugins.btauto.f_top"letf_cnj=bt_lib_constr"plugins.btauto.f_cnj"letf_dsj=bt_lib_constr"plugins.btauto.f_dsj"letf_neg=bt_lib_constr"plugins.btauto.f_neg"letf_xor=bt_lib_constr"plugins.btauto.f_xor"letf_ifb=bt_lib_constr"plugins.btauto.f_ifb"leteval=bt_lib_constr"plugins.btauto.eval"letwitness=bt_lib_constr"plugins.btauto.witness"letsoundness=bt_lib_constr"plugins.btauto.soundness"letrecconvert=function|Bool.Varn->lappf_var[|CoqPositive.of_intn|]|Bool.Consttrue->Lazy.forcef_top|Bool.Constfalse->Lazy.forcef_btm|Bool.Andb(b1,b2)->lappf_cnj[|convertb1;convertb2|]|Bool.Orb(b1,b2)->lappf_dsj[|convertb1;convertb2|]|Bool.Negbb->lappf_neg[|convertb|]|Bool.Xorb(b1,b2)->lappf_xor[|convertb1;convertb2|]|Bool.Ifb(b1,b2,b3)->lappf_ifb[|convertb1;convertb2;convertb3|]letconvert_envenv:Constr.t=CoqList.of_list(Lazy.forceBool.typ)envletreifyenvt=lappeval[|convert_envenv;convertt|]letprint_counterexampleppenv=Proofview.Goal.enterbeginfungl->letenv=Proofview.Goal.envglinletsigma=Proofview.Goal.sigmaglinletvar=lappwitness[|p|]inletvar=EConstr.of_constrvarin(* Compute an assignment that dissatisfies the goal *)letredfun,_=Redexpr.reduction_of_red_exprenvGenredexpr.(CbvVmNone)inlet_,var=redfunenvsigmavarinletvar=EConstr.Unsafe.to_constrvarinletrecto_listl=matchdecomp_termsigmalwith|App(c,_)whenc===(Lazy.forceCoqList._nil)->[]|App(c,[|_;h;t|])whenc===(Lazy.forceCoqList._cons)->ifh===(Lazy.forceBool.trueb)then(true::to_listt)elseifh===(Lazy.forceBool.falseb)then(false::to_listt)elseinvalid_arg"to_list"|_->invalid_arg"to_list"inletconcatsep=function|[]->mt()|h::t->letrecaux=function|[]->mt()|x::t->(sep++x++auxt)inh++auxtinletmsg=tryletvar=to_listvarinletassign=List.combinepenvvarinletmap_msg(key,v)=letb=ifvthenstr"true"elsestr"false"inletterm=Printer.pr_constr_envenvsigmakeyinterm++spc()++str":="++spc()++binletassign=List.mapmap_msgassigninletl=str"["++(concat(str";"++spc())assign)++str"]"instr"Not a tautology:"++spc()++lwithewhenCErrors.noncriticale->(str"Not a tautology")inTacticals.New.tclFAIL0msgendlettry_unificationenv=Proofview.Goal.enterbeginfungl->letconcl=Proofview.Goal.conclglinleteq=Lazy.forceeqinletconcl=EConstr.Unsafe.to_constrconclinlett=decomp_term(Tacmach.New.projectgl)conclinmatchtwith|App(c,[|typ;p;_|])whenc===eq->(* should be an equality [@eq poly ?p (Cst false)] *)lettac=Tacticals.New.tclORELSE0Tactics.reflexivity(print_counterexamplepenv)intac|_->letmsg=str"Btauto: Internal error"inTacticals.New.tclFAIL0msgendlettac=Proofview.Goal.enterbeginfungl->letconcl=Proofview.Goal.conclglinletconcl=EConstr.Unsafe.to_constrconclinletsigma=Tacmach.New.projectglinleteq=Lazy.forceeqinletbool=Lazy.forceBool.typinlett=decomp_termsigmaconclinmatchtwith|App(c,[|typ;tl;tr|])whentyp===bool&&c===eq->letenv=Env.empty()inletfl=Bool.quoteenvsigmatlinletfr=Bool.quoteenvsigmatrinletenv=Env.to_listenvinletfl=reifyenvflinletfr=reifyenvfrinletchanged_gl=Constr.mkApp(c,[|typ;fl;fr|])inletchanged_gl=EConstr.of_constrchanged_glinTacticals.New.tclTHENLIST[Tactics.change_conclchanged_gl;Tactics.apply(EConstr.of_constr(Lazy.forcesoundness));Tactics.normalise_vm_in_concl;try_unificationenv]|_->letmsg=str"Cannot recognize a boolean equality"inTacticals.New.tclFAIL0msgendend