123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691692693694695696697698699700701702703704705706707708709710711712713714715716717718719720721722723724725726727728729(************************************************************************)(* * 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) *)(************************************************************************)openNumCompatopenQ.NotationsopenUtilopenPolynomialopenVectletdebug=falseletcompare_float(p:float)q=pervasives_comparepq(** Implementation of intervals *)openItvtypevector=Vect.t(** 'cstr' is the type of constraints.
{coeffs = v ; bound = (l,r) } models the constraints l <= v <= r
**)moduleISet=Set.Make(Int)moduleSystem=Hashtbl.Make(Vect)typeproof=Assumofint|Elimofvar*proof*proof|Andofproof*prooftypesystem={sys:cstr_inforefSystem.t;vars:ISet.t}andcstr_info={bound:interval;prf:proof;pos:int;neg:int}(** A system of constraints has the form [\{sys = s ; vars = v\}].
[s] is a hashtable mapping a normalised vector to a [cstr_info] record where
- [bound] is an interval
- [prf_idx] is the set of hypothesis indexes (i.e. constraints in the initial system) used to obtain the current constraint.
In the initial system, each constraint is given an unique singleton proof_idx.
When a new constraint c is computed by a function f(c1,...,cn), its proof_idx is ISet.fold union (List.map (fun x -> x.proof_idx) [c1;...;cn]
- [pos] is the number of positive values of the vector
- [neg] is the number of negative values of the vector
( [neg] + [pos] is therefore the length of the vector)
[v] is an upper-bound of the set of variables which appear in [s].
*)(** To be thrown when a system has no solution *)exceptionSystemContradictionofproof(** Pretty printing *)letrecpp_proofoprf=matchprfwith|Assumi->Printf.fprintfo"H%i"i|Elim(v,prf1,prf2)->Printf.fprintfo"E(%i,%a,%a)"vpp_proofprf1pp_proofprf2|And(prf1,prf2)->Printf.fprintfo"A(%a,%a)"pp_proofprf1pp_proofprf2letpp_cstro(vect,bnd)=letl,r=bndin(matchlwith|None->()|Somen->Printf.fprintfo"%s <= "(Q.to_stringn));Vect.ppovect;matchrwith|None->output_stringo"\n"|Somen->Printf.fprintfo"<=%s\n"(Q.to_stringn)letpp_systemosys=System.iter(funvectibnd->pp_cstro(vect,!ibnd.bound))sys(** [merge_cstr_info] takes:
- the intersection of bounds and
- the union of proofs
- [pos] and [neg] fields should be identical *)letmerge_cstr_infoi1i2=let{pos=p1;neg=n1;bound=i1;prf=prf1}=i1and{pos=p2;neg=n2;bound=i2;prf=prf2}=i2inassert(Int.equalp1p2&&Int.equaln1n2);matchinteri1i2with|None->None(* Could directly raise a system contradiction exception *)|Somebnd->Some{pos=p1;neg=n1;bound=bnd;prf=And(prf1,prf2)}(** [xadd_cstr vect cstr_info] loads an constraint into the system.
The constraint is neither redundant nor contradictory.
@raise SystemContradiction if [cstr_info] returns [None]
*)letxadd_cstrvectcstr_infosys=tryletinfo=System.findsysvectinmatchmerge_cstr_infocstr_info!infowith|None->raise(SystemContradiction(And(cstr_info.prf,!info.prf)))|Someinfo'->info:=info'withNot_found->System.replacesysvect(refcstr_info)exceptionTimeOutletxadd_cstrvectcstr_infosys=ifdebug&&Int.equal(System.lengthsysmod1000)0then(print_string"*";flushstdout);ifSystem.lengthsys<!max_nb_cstrthenxadd_cstrvectcstr_infosyselseraiseTimeOuttypecstr_ext=|Contradiction(** The constraint is contradictory.
Typically, a [SystemContradiction] exception will be raised. *)|Redundant(** The constrain is redundant.
Typically, the constraint will be dropped *)|Cstrofvector*cstr_info(** Taken alone, the constraint is neither contradictory nor redundant.
Typically, it will be added to the constraint system. *)(** [normalise_cstr] : vector -> cstr_info -> cstr_ext *)letnormalise_cstrvectcinfo=matchnorm_itvcinfo.boundwith|None->Contradiction|Some(l,r)->(matchVect.choosevectwith|None->ifItv.in_bound(l,r)Q.zerothenRedundantelseContradiction|Some(_,n,_)->Cstr(Vect.divnvect,letdivnx=x//ninifInt.equal(Q.signn)1then{cinfowithbound=(Option.mapdivnl,Option.mapdivnr)}else{cinfowithpos=cinfo.neg;neg=cinfo.pos;bound=(Option.mapdivnr,Option.mapdivnl)}))(** For compatibility, there is an external representation of constraints *)letcountv=Vect.fold(fun(n,p)_vl->letsg=Q.signvlinassert(sg<>0);ifInt.equalsg1then(n,p+1)else(n+1,p))(0,0)vletnorm_cstr{coeffs=v;op=o;cst=c}idx=letn,p=countvinnormalise_cstrv{pos=p;neg=n;bound=(matchowith|Eq->(Somec,Somec)|Ge->(Somec,None)|Gt->raisePolynomial.Strict);prf=Assumidx}(** [load_system l] takes a list of constraints of type [cstr_compat]
@return a system of constraints
@raise SystemContradiction if a contradiction is found
*)letload_systeml=letsys=System.create1000inletli=List.mapi(funie->(e,i))linletvars=List.fold_left(funvrs(cstr,i)->matchnorm_cstrcstriwith|Contradiction->raise(SystemContradiction(Assumi))|Redundant->vrs|Cstr(vect,info)->xadd_cstrvectinfosys;Vect.fold(funsv_->ISet.addvs)vrscstr.coeffs)ISet.emptyliin{sys;vars}letsystem_listsys=let{sys=s;vars=v}=sysinSystem.fold(funkbil->(k,!bi)::l)s[](** [add (v1,c1) (v2,c2) ]
precondition: (c1 <>/ Q.zero && c2 <>/ Q.zero)
@return a pair [(v,ln)] such that
[v] is the sum of vector [v1] divided by [c1] and vector [v2] divided by [c2]
Note that the resulting vector is not normalised.
*)letadd(v1,c1)(v2,c2)=assert(c1<>/Q.zero&&c2<>/Q.zero);(* XXX Can use Q.inv now *)letres=mul_add(Q.one//c1)v1(Q.one//c2)v2in(res,countres)letadd(v1,c1)(v2,c2)=letres=add(v1,c1)(v2,c2)in(* Printf.printf "add(%a,%s,%a,%s) -> %a\n" pp_vect v1 (Q.to_string c1) pp_vect v2 (Q.to_string c2) pp_vect (fst res) ;*)res(** To perform Fourier elimination, constraints are categorised depending on the sign of the variable to eliminate. *)(** [split x vect info (l,m,r)]
@param v is the variable to eliminate
@param l contains constraints such that (e + a*x) // a >= c / a
@param r contains constraints such that (e + a*x) // - a >= c / -a
@param m contains constraints which do not mention [x]
*)letsplitx(vect:vector)info(l,m,r)=letvl=getxvectinifQ.zero=/vlthen(* The constraint does not mention [x], store it in m *)(l,(vect,info)::m,r)else(* otherwise *)letcons_boundlstbd=matchbdwith|None->lst|Somebnd->(vl,vect,{infowithbound=(Somebnd,None)})::lstinletlb,rb=info.boundinifInt.equal(Q.signvl)1then(cons_boundllb,m,cons_boundrrb)else(* sign_num vl = -1 *)(cons_boundlrb,m,cons_boundrlb)(** [project vr sys] projects system [sys] over the set of variables [ISet.remove vr sys.vars ].
This is a one step Fourier elimination.
*)letprojectvrsys=letl,m,r=System.fold(funvectrfl_m_r->splitvrvect!rfl_m_r)sys.sys([],[],[])inletnew_sys=System.create(System.lengthsys.sys)in(* Constraints in [m] belong to the projection - for those [vr] is already projected out *)List.iter(fun(vect,info)->System.replacenew_sysvect(refinfo))m;letelim(v1,vect1,info1)(v2,vect2,info2)=let{neg=n1;pos=p1;bound=bound1;prf=prf1}=info1and{neg=n2;pos=p2;bound=bound2;prf=prf2}=info2inletbnd1=Option.get(fstbound1)andbnd2=Option.get(fstbound2)inletbound=(bnd1//v1)+/(bnd2//Q.negv2)inletvres,(n,p)=add(vect1,v1)(vect2,Q.negv2)in(vres,{neg=n;pos=p;bound=(Somebound,None);prf=Elim(vr,info1.prf,info2.prf)})inList.iter(funl_elem->List.iter(funr_elem->letvect,info=eliml_elemr_eleminmatchnormalise_cstrvectinfowith|Redundant->()|Contradiction->raise(SystemContradictioninfo.prf)|Cstr(vect,info)->xadd_cstrvectinfonew_sys)r)l;{sys=new_sys;vars=ISet.removevrsys.vars}(** [project_using_eq] performs elimination by pivoting using an equation.
This is the counter_part of the [elim] sub-function of [!project].
@param vr is the variable to be used as pivot
@param c is the coefficient of variable [vr] in vector [vect]
@param len is the length of the equation
@param bound is the bound of the equation
@param prf is the proof of the equation
*)letproject_using_eqvrcvectboundprf(vect',info')=letc2=getvrvect'inifQ.zero=/c2then(vect',info')elseletc1=ifc2>=/Q.zerothenQ.negcelsecinletc2=Q.absc2inletvres,(n,p)=add(vect,c1)(vect',c2)inletcst=bound//c1inletbndres=letfx=cst+/(x//c2)inletl,r=info'.boundin(Option.mapfl,Option.mapfr)in(vres,{neg=n;pos=p;bound=bndres;prf=Elim(vr,prf,info'.prf)})letelim_var_using_eqvrvectcstprfsys=letc=getvrvectinletelim_var=project_using_eqvrcvectcstprfinletnew_sys=System.create(System.lengthsys.sys)inSystem.iter(funvectiref->letvect',info'=elim_var(vect,!iref)inmatchnormalise_cstrvect'info'with|Redundant->()|Contradiction->raise(SystemContradictioninfo'.prf)|Cstr(vect,info')->xadd_cstrvectinfo'new_sys)sys.sys;{sys=new_sys;vars=ISet.removevrsys.vars}(** [size sys] computes the number of entries in the system of constraints *)letsizesys=System.fold(funvirefs->s+!iref.neg+!iref.pos)sys0moduleIMap=CMap.Make(Int)(** [eval_vect map vect] evaluates vector [vect] using the values of [map].
If [map] binds all the variables of [vect], we get
[eval_vect map [(x1,v1);...;(xn,vn)] = (IMap.find x1 map * v1) + ... + (IMap.find xn map) * vn , []]
The function returns as second argument, a sub-vector consisting in the variables that are not in [map]. *)leteval_vectmapvect=Vect.fold(fun(sum,rst)vvl->tryletval_v=IMap.findvmapin(sum+/(val_v*/vl),rst)withNot_found->(sum,Vect.setvvlrst))(Q.zero,Vect.null)vect(** [restrict_bound n sum itv] returns the interval of [x]
given that (fst itv) <= x * n + sum <= (snd itv) *)letrestrict_boundnsum(itv:interval)=letfx=(x-/sum)//ninletl,r=itvinmatchQ.signnwith|0->ifin_bounditvsumthen(None,None)(* redundant *)elsefailwith"SystemContradiction"|1->(Option.mapfl,Option.mapfr)|_->(Option.mapfr,Option.mapfl)(** [bound_of_variable map v sys] computes the interval of [v] in
[sys] given a mapping [map] binding all the other variables *)letbound_of_variablemapvsys=System.fold(funvectirefbnd->letsum,rst=eval_vectmapvectinletvl=Vect.getvrstinmatchinterbnd(restrict_boundvlsum!iref.bound)with|None->Printf.fprintfstdout"bound_of_variable: eval_vecr %a = %s,%a\n"Vect.ppvect(Q.to_stringsum)Vect.pprst;Printf.fprintfstdout"current interval: %a\n"Itv.pp!iref.bound;failwith"bound_of_variable: impossible"|Someitv->itv)sys(None,None)(** [pick_small_value bnd] picks a value being closed to zero within the interval *)letpick_small_valuebnd=matchbndwith|None,None->Q.zero|None,Somei->ifQ.zero<=/Q.floorithenQ.zeroelseQ.floori|Somei,None->ifi<=/Q.zerothenQ.zeroelseQ.ceilingi|Somei,Somej->ifi<=/Q.zero&&Q.zero<=/jthenQ.zeroelseifQ.ceilingi<=/Q.floorjthenQ.ceilingi(* why not *)elsei(** [solution s1 sys_l = Some(sn,\[(vn-1,sn-1);...; (v1,s1)\]\@sys_l)]
then [sn] is a system which contains only [black_v] -- if it existed in [s1]
and [sn+1] is obtained by projecting [vn] out of [sn]
@raise SystemContradiction if system [s] has no solution
*)letsolve_sysblack_vchoose_eqchoose_variablesyssys_l=letrecsolve_syssyssys_l=ifdebugthenPrintf.printf"S #%i size %i\n"(System.lengthsys.sys)(sizesys.sys);ifdebugthenPrintf.printf"solve_sys :\n %a"pp_systemsys.sys;leteqs=choose_eqsysintryletv,vect,cst,ln=fst(List.find(fun((v,_,_,_),_)->v<>black_v)eqs)inifdebugthen(Printf.printf"\nE %a = %s variable %i\n"Vect.ppvect(Q.to_stringcst)v;flushstdout);letsys'=elim_var_using_eqvvectcstlnsysinsolve_syssys'((v,sys)::sys_l)withNot_found->(letvars=choose_variablesysintryletv,est=List.find(fun(v,_)->v<>black_v)varsinifdebugthen(Printf.printf"\nV : %i estimate %f\n"vest;flushstdout);letsys'=projectvsysinsolve_syssys'((v,sys)::sys_l)withNot_found->(* we are done *)Inl(sys,sys_l))insolve_syssyssys_lletsolveblack_vchoose_eqchoose_variablecstrs=tryletsys=load_systemcstrsinifdebugthenPrintf.printf"solve :\n %a"pp_systemsys.sys;solve_sysblack_vchoose_eqchoose_variablesys[]withSystemContradictionprf->Inrprf(** The purpose of module [EstimateElimVar] is to try to estimate the cost of eliminating a variable.
The output is an ordered list of (variable,cost).
*)moduleEstimateElimVar=structtypesys_list=(vector*cstr_info)listletabstract_partition(v:int)(l:sys_list)=letrecxpart(l:sys_list)(ltl:sys_list)(n:intlist)(z:int)(p:intlist)=matchlwith|[]->(ltl,n,z,p)|(l1,info)::rl->(matchVect.choosel1with|None->xpartrl((Vect.null,info)::ltl)n(info.neg+info.pos+z)p|Some(vr,vl,rl1)->ifInt.equalvvrthenletcons_boundlstbd=matchbdwith|None->lst|Somebnd->(info.neg+info.pos)::lstinletlb,rb=info.boundinifInt.equal(Q.signvl)1thenxpartrl((rl1,info)::ltl)(cons_boundnlb)z(cons_boundprb)elsexpartrl((rl1,info)::ltl)(cons_boundnrb)z(cons_boundplb)else(* the variable is greater *)xpartrl((l1,info)::ltl)n(info.neg+info.pos+z)p)inletsys',n,z,p=xpartl[][]0[]inletln=float_of_int(List.lengthn)inletsn=float_of_int(List.fold_left(+)0n)inletlp=float_of_int(List.lengthp)inletsp=float_of_int(List.fold_left(+)0p)in(sys',float_of_intz+.(lp*.sn)+.(ln*.sp)-.(lp*.ln))letchoose_variablesys=let{sys=s;vars=v}=sysinletsl=system_listsysinletevals=fst(ISet.fold(funv(eval,s)->letts,vl=abstract_partitionvsin((v,vl)::eval,ts))v([],sl))inList.sort(funxy->compare_float(sndx)(sndy))evalsendopenEstimateElimVar(** The module [EstimateElimEq] is similar to [EstimateElimVar] but it orders equations.
*)moduleEstimateElimEq=structletitv_pointbnd=matchbndwithSomea,Someb->a=/b|_->falseletrecunroll_untilvl=matchVect.chooselwith|None->(false,Vect.null)|Some(i,_,rl)->ifInt.equalivthen(true,rl)elseifi<vthenunroll_untilvrlelse(false,l)letrecchoose_simple_equationeqs=matcheqswith|[]->None|(vect,a,prf,ln)::eqs->(matchVect.choosevectwith|Some(i,v,rst)->ifVect.is_nullrstthenSome(i,vect,a,prf,ln)elsechoose_simple_equationeqs|_->choose_simple_equationeqs)letchoose_primal_equationeqs(sys_l:(Vect.t*cstr_info)list)=(* Counts the number of equations referring to variable [v] --
It looks like nb_cst is dead...
*)letis_primal_equation_varv=List.fold_left(funnb_eq(vect,info)->iffst(unroll_untilvvect)thenifitv_pointinfo.boundthennb_eq+1elsenb_eqelsenb_eq)0sys_linletrecfind_varvect=matchVect.choosevectwith|None->None|Some(i,_,vect)->letnb_eq=is_primal_equation_variinifInt.equalnb_eq2thenSomeielsefind_varvectinletrecfind_eq_vareqs=matcheqswith|[]->None|(vect,a,prf,ln)::l->(matchfind_varvectwith|None->find_eq_varl|Somer->Some(r,vect,a,prf,ln))inmatchchoose_simple_equationeqswith|None->find_eq_vareqs|Someres->Someresletchoose_equality_varsys=letsys_l=system_listsysinletequalities=List.fold_left(funl(vect,info)->matchinfo.boundwith|Somea,Someb->ifa=/bthen(* This an equation *)(vect,a,info.prf,info.neg+info.pos)::lelsel|_->l)[]sys_linletrecestimate_costvctsyslacctlsys=matchsyslwith|[]->(acc,tlsys)|(l,info)::rsys->(letln=info.pos+info.neginletb,l=unroll_untilvlinmatchbwith|true->ifitv_pointinfo.boundthenestimate_costvctrsys(acc+ln)((l,info)::tlsys)(* this is free *)elseestimate_costvctrsys(acc+ln+ct)((l,info)::tlsys)(* should be more ? *)|false->estimate_costvctrsys(acc+ln)((l,info)::tlsys))inmatchchoose_primal_equationequalitiessys_lwith|None->letcost_eqeqconstprflnacc_costs=letreccost_eqeqrsyslcosts=matchVect.chooseeqrwith|None->costs|Some(v,_,eqr)->letcst,tlsys=estimate_costv(ln-1)sysl0[]incost_eqeqrtlsys(((v,eq,const,prf),cst)::costs)incost_eqeqsys_lacc_costsinletall_costs=List.fold_left(funall_costs(vect,const,prf,ln)->cost_eqvectconstprflnall_costs)[]equalitiesin(* pp_list (fun o ((v,eq,_,_),cst) -> Printf.fprintf o "((%i,%a),%i)\n" v pp_vect eq cst) stdout all_costs ; *)List.sort(funxy->Int.compare(sndx)(sndy))all_costs|Some(v,vect,const,prf,_)->[((v,vect,const,prf),0)]endopenEstimateElimEqmoduleFourier=structletoptimisevectl=(* We add a dummy (fresh) variable for vector *)letfresh=List.fold_left(funfrc->maxfr(Vect.freshc.coeffs))0linletcstr={coeffs=Vect.setfreshQ.minus_onevect;op=Eq;cst=Q.zero}inmatchsolvefreshchoose_equality_varchoose_variable(cstr::l)with|Inrprf->None(* This is an unsatisfiability proof *)|Inl(s,_)->(trySome(bound_of_variableIMap.emptyfreshs.sys)withxwhenCErrors.noncriticalx->Printf.printf"optimise Exception : %s"(Printexc.to_stringx);None)letfind_pointcstrs=matchsolvemax_intchoose_equality_varchoose_variablecstrswith|Inrprf->Inrprf|Inl(_,l)->letrecrebuild_solutionlmap=matchlwith|[]->map|(v,e)::l->letitv=bound_of_variablemapve.sysinletmap=IMap.addv(pick_small_valueitv)mapinrebuild_solutionlmapinletmap=rebuild_solutionlIMap.emptyinletvect=IMap.fold(funvivect->Vect.setvivect)mapVect.nullinifdebugthenPrintf.printf"SOLUTION %a"Vect.ppvect;letres=InlvectinresendmoduleProof=struct(** A proof term in the sense of a ZMicromega.RatProof is a positive combination of the hypotheses which leads to a contradiction.
The proofs constructed by Fourier elimination are more like execution traces:
- certain facts are recorded but are useless
- certain inferences are implicit.
The following code implements proof reconstruction.
*)letaddxy=fst(addxy)letforall_pairsfl1l2=List.fold_left(funacce1->List.fold_left(funacce2->matchfe1e2withNone->acc|Somev->v::acc)accl2)[]l1letadd_opxy=match(x,y)withEq,Eq->Eq|_->Geletpivotv(p1,c1)(p2,c2)=let{coeffs=v1;op=op1;cst=n1}=c1and{coeffs=v2;op=op2;cst=n2}=c2inleta,b=(Vect.getvv1,Vect.getvv2)inifQ.zero=/a||Q.zero=/bthenNoneelseifInt.equal(Q.signa*Q.signb)(-1)thenSome(add(p1,Q.absa)(p2,Q.absb),{coeffs=add(v1,Q.absa)(v2,Q.absb);op=add_opop1op2;cst=(n1//Q.absa)+/(n2//Q.absb)})elseifop1==EqthenSome(add(p1,Q.neg(a//b))(p2,Q.one),{coeffs=add(v1,Q.neg(a//b))(v2,Q.one);op=add_opop1op2;cst=(n1//Q.neg(a//b))+/(n2//Q.one)})elseifop2==EqthenSome(add(p2,Q.neg(b//a))(p1,Q.one),{coeffs=add(v2,Q.neg(b//a))(v1,Q.one);op=add_opop1op2;cst=(n2//Q.neg(b//a))+/(n1//Q.one)})elseNone(* op2 could be Eq ... this might happen *)letnormalise_proofsl=List.fold_left(funacc(prf,cstr)->matchaccwith|Inr_->acc(* I already found a contradiction *)|Inlacc->(matchnorm_cstrcstr0with|Redundant->Inlacc|Contradiction->Inr(prf,cstr)|Cstr(v,info)->Inl((prf,cstr,v,info)::acc)))(Inl[])ltypeoproof=(vector*cstr*Q.t)optionletmerge_proof(oleft:oproof)(prf,cstr,v,info)(oright:oproof)=letl,r=info.boundinletkeeppobbd=match(ob,bd)with|None,None->None|None,Someb->Some(prf,cstr,b)|Some_,None->ob|Some(prfl,cstrl,bl),Someb->ifpblbthenSome(prf,cstr,b)elseobinletoleft=keep(<=/)oleftlinletoright=keep(>=/)orightrin(* Now, there might be a contradiction *)match(oleft,oright)with|None,_|_,None->Inl(oleft,oright)|Some(prfl,cstrl,l),Some(prfr,cstrr,r)->(ifl<=/rthenInl(oleft,oright)else(* There is a contradiction - it should show up by scaling up the vectors - any pivot should do*)matchVect.choosecstrr.coeffswith|None->Inr(add(prfl,Q.one)(prfr,Q.one),cstrr)(* this is wrong *)|Some(v,_,_)->(matchpivotv(prfl,cstrl)(prfr,cstrr)with|None->failwith"merge_proof : pivot is not possible"|Somex->Inrx))letmk_proofhypsprf=(* I am keeping list - I might have a proof for the left bound and a proof for the right bound.
If I perform aggressive elimination of redundancies, I expect the list to be of length at most 2.
For each proof list, all the vectors should be of the form a.v for different constants a.
*)letrecmk_proofprf=matchprfwith|Assumi->[(Vect.setiQ.oneVect.null,List.nthhypsi)]|Elim(v,prf1,prf2)->letprfsl=mk_proofprf1andprfsr=mk_proofprf2in(* I take only the pairs for which the elimination is meaningful *)forall_pairs(pivotv)prfslprfsr|And(prf1,prf2)->(letprfsl1=mk_proofprf1andprfsl2=mk_proofprf2in(* detect trivial redundancies and contradictions *)matchnormalise_proofs(prfsl1@prfsl2)with|Inrx->[x](* This is a contradiction - this should be the end of the proof *)|Inll->((* All the vectors are the same *)letprfs=List.fold_left(funacce->matchaccwith|Inr_->acc(* I have a contradiction *)|Inl(oleft,oright)->merge_proofolefteoright)(Inl(None,None))linmatchprfswith|Inrx->[x]|Inl(oleft,oright)->(match(oleft,oright)with|None,None->[]|None,Some(prf,cstr,_)|Some(prf,cstr,_),None->[(prf,cstr)]|Some(prf1,cstr1,_),Some(prf2,cstr2,_)->[(prf1,cstr1);(prf2,cstr2)])))inmk_proofprfend