123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691692693694695696697698699700701702703704705706707708709710711712713714715716717718719720721722723724725726727728729730731732733734735736737738739740741742743744745746747748749750751752753754755756757758759760761762763764765766767768769770771772773774775776777778779780781782783784785786787788789790791792793794795796797798799800801802803(************************************************************************)(* * 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) *)(************************************************************************)typeconstraint_type=Lt|Le|EqmoduletypePoint=sigtypetmoduleSet:CSig.SetSwithtypeelt=tmoduleMap:CMap.ExtSwithtypekey=tandmoduleSet:=SetmoduleConstraint:CSet.Swithtypeelt=(t*constraint_type*t)valequal:t->t->boolvalcompare:t->t->inttypeexplanation=(constraint_type*t)listvalerror_inconsistency:constraint_type->t->t->explanationlazy_toption->'avalpr:t->Pp.tendmoduleMake(Point:Point)=struct(* Created in Caml by Gérard Huet for CoC 4.8 [Dec 1988] *)(* Functional code by Jean-Christophe Filliâtre for Coq V7.0 [1999] *)(* Extension with algebraic universes by HH for Coq V7.0 [Sep 2001] *)(* Additional support for sort-polymorphic inductive types by HH [Mar 2006] *)(* Support for universe polymorphism by MS [2014] *)(* Revisions by Bruno Barras, Hugo Herbelin, Pierre Letouzey, Matthieu
Sozeau, Pierre-Marie Pédrot, Jacques-Henri Jourdan *)(* Points are stratified by a partial ordering $\le$.
Let $\~{}$ be the associated equivalence. We also have a strict ordering
$<$ between equivalence classes, and we maintain that $<$ is acyclic,
and contained in $\le$ in the sense that $[U]<[V]$ implies $U\le V$.
At every moment, we have a finite number of points, and we
maintain the ordering in the presence of assertions $U<V$ and $U\le V$.
The equivalence $\~{}$ is represented by a tree structure, as in the
union-find algorithm. The assertions $<$ and $\le$ are represented by
adjacency lists.
We use the algorithm described in the paper:
Bender, M. A., Fineman, J. T., Gilbert, S., & Tarjan, R. E. (2011). A
new approach to incremental cycle detection and related
problems. arXiv preprint arXiv:1112.0784.
*)moduleIndex:sigtypetvalequal:t->t->boolmoduleSet:CSig.SetSwithtypeelt=tmoduleMap:CMap.ExtSwithtypekey=tandmoduleSet:=Settypetablevalempty:tablevalfresh:Point.t->table->t*tablevalmem:Point.t->table->boolvalfind:Point.t->table->tvalrepr:t->table->Point.tend=structtypet=intletequal=Int.equalmoduleSet=Int.SetmoduleMap=Int.Maptypetable={tab_len:int;tab_fwd:Point.tInt.Map.t;tab_bwd:intPoint.Map.t}letempty={tab_len=0;tab_fwd=Int.Map.empty;tab_bwd=Point.Map.empty;}letmemxt=Point.Map.memxt.tab_bwdletfindxt=Point.Map.findxt.tab_bwdletreprnt=Int.Map.findnt.tab_fwdletfreshxt=let()=assert(not@@memxt)inletn=t.tab_leninn,{tab_len=n+1;tab_fwd=Int.Map.addnxt.tab_fwd;tab_bwd=Point.Map.addxnt.tab_bwd;}endmodulePMap=Index.MapmodulePSet=Index.SetmoduleConstraint=Point.Constrainttypestatus=NoMark|Visited|WeakVisited|ToMerge(* Comparison on this type is pointer equality *)typecanonical_node={canon:Index.t;ltle:boolPMap.t;(* true: strict (lt) constraint.
false: weak (le) constraint. *)gtge:PSet.t;rank:int;klvl:int;ilvl:int;mutablestatus:status}(* A Point.t is either an alias for another one, or a canonical one,
for which we know the points that are above *)typeentry=|Canonicalofcanonical_node|EquivofIndex.ttypet={entries:entryPMap.t;index:int;n_nodes:int;n_edges:int;table:Index.table}(** Used to cleanup mutable marks if a traversal function is
interrupted before it has the opportunity to do it itself. *)letunsafe_cleanup_marksg=letiter_n=matchnwith|Equiv_->()|Canonicaln->n.status<-NoMarkinPMap.iteriterg.entriesletreccleanup_marksg=tryunsafe_cleanup_marksgwithe->(* The only way unsafe_cleanup_marks may raise an exception is when
a serious error (stack overflow, out of memory) occurs, or a signal is
sent. In this unlikely event, we relaunch the cleanup until we finally
succeed. *)cleanup_marksg;raisee(* Every Point.t has a unique canonical arc representative *)(* Low-level function : makes u an alias for v.
Does not removes edges from n_edges, but decrements n_nodes.
u should be entered as canonical before. *)letenter_equivguv={entries=PMap.modifyu(fun_a->matchawith|Canonicaln->n.status<-NoMark;Equivv|_->assertfalse)g.entries;index=g.index;n_nodes=g.n_nodes-1;n_edges=g.n_edges;table=g.table}(* Low-level function : changes data associated with a canonical node.
Resets the mutable fields in the old record, in order to avoid breaking
invariants for other users of this record.
n.canon should already been inserted as a canonical node. *)letchange_nodegn={gwithentries=PMap.modifyn.canon(fun_a->matchawith|Canonicaln'->n'.status<-NoMark;Canonicaln|_->assertfalse)g.entries}(* canonical representative : we follow the Equiv links *)letrecreprgu=matchPMap.findug.entrieswith|Equivv->reprgv|Canonicalarc->arcletrepr_nodegu=tryreprg(Index.findug.table)withNot_found->CErrors.anomaly~label:"Univ.repr"Pp.(str"Universe "++Point.pru++str" undefined.")exceptionAlreadyDeclared(* Reindexes the given point, using the next available index. *)letuse_indexgu=letu=reprguinletg=change_nodeg{uwithilvl=g.index}inassert(g.index>min_int);{gwithindex=g.index-1}(* Returns 1 if u is higher than v in topological order.
-1 lower
0 if u = v *)lettopo_compareuv=ifu.klvl>v.klvlthen1elseifu.klvl<v.klvlthen-1elseifu.ilvl>v.ilvlthen1elseifu.ilvl<v.ilvlthen-1else(assert(u==v);0)(* Checks most of the invariants of the graph. For debugging purposes. *)letcheck_invariants~required_canonicalg=letrequired_canonicalu=required_canonical(Index.reprug.table)inletn_edges=ref0inletn_nodes=ref0inPMap.iter(funlu->matchuwith|Canonicalu->PMap.iter(funv_strict->incrn_edges;letv=reprgvinassert(topo_compareuv=-1);ifu.klvl=v.klvlthenassert(PSet.memu.canonv.gtge||PSet.exists(funl->u==reprgl)v.gtge))u.ltle;PSet.iter(funv->letv=reprgvinassert(v.klvl=u.klvl&&(PMap.memu.canonv.ltle||PMap.exists(funl_->u==reprgl)v.ltle)))u.gtge;assert(u.status=NoMark);assert(Index.equallu.canon);assert(u.ilvl>g.index);assert(not(PMap.memu.canonu.ltle));incrn_nodes|Equiv_->assert(not(required_canonicall)))g.entries;assert(!n_edges=g.n_edges);assert(!n_nodes=g.n_nodes)letclean_ltlegltle=PMap.fold(funustrictacc->letuu=(reprgu).canoninifIndex.equaluuuthenaccelse(letacc=PMap.removeu(fstacc)inifnotstrict&&PMap.memuuaccthen(acc,true)else(PMap.adduustrictacc,true)))ltle(ltle,false)letclean_gtgeggtge=PSet.fold(funuacc->letuu=(reprgu).canoninifIndex.equaluuuthenaccelsePSet.adduu(PSet.removeu(fstacc)),true)gtge(gtge,false)(* [get_ltle] and [get_gtge] return ltle and gtge arcs.
Moreover, if one of these lists is dirty (e.g. points to a
non-canonical node), these functions clean this node in the
graph by removing some duplicate edges *)letget_ltlegu=letltle,chgt_ltle=clean_ltlegu.ltleinifnotchgt_ltlethenu.ltle,u,gelseletsz=PMap.cardinalu.ltleinletsz2=PMap.cardinalltleinletu={uwithltle}inletg=change_nodeguinletg={gwithn_edges=g.n_edges+sz2-sz}inu.ltle,u,gletget_gtgegu=letgtge,chgt_gtge=clean_gtgegu.gtgeinifnotchgt_gtgethenu.gtge,u,gelseletu={uwithgtge}inletg=change_nodeguinu.gtge,u,g(* [revert_graph] rollbacks the changes made to mutable fields in
nodes in the graph.
[to_revert] contains the touched nodes. *)letrevert_graphto_revertg=List.iter(funt->matchPMap.findtg.entrieswith|Equiv_->()|Canonicalt->t.status<-NoMark)to_revertexceptionAbortBackwardoftexceptionCycleDetected(* Implementation of the algorithm described in § 5.1 of the following paper:
Bender, M. A., Fineman, J. T., Gilbert, S., & Tarjan, R. E. (2011). A
new approach to incremental cycle detection and related
problems. arXiv preprint arXiv:1112.0784.
The "STEP X" comments contained in this file refers to the
corresponding step numbers of the algorithm described in Section
5.1 of this paper. *)letrecbackward_traverseto_revertb_traversedcountgx=letcount=count-1inifcount<0thenbeginrevert_graphto_revertg;raise(AbortBackwardg)end;ifx.status=NoMarkthenbeginx.status<-Visited;letto_revert=x.canon::to_revertinletgtge,x,g=get_gtgegxinletto_revert,b_traversed,count,g=PSet.fold(funy(to_revert,b_traversed,count,g)->lety=reprgyinbackward_traverseto_revertb_traversedcountgy)gtge(to_revert,b_traversed,count,g)into_revert,x.canon::b_traversed,count,gendelseto_revert,b_traversed,count,gletrecforward_traversef_traversedgv_klvlxy=lety=reprgyinify.klvl<v_klvlthenbeginlety={ywithklvl=v_klvl;gtge=ifx==ythenPSet.emptyelsePSet.singletonx.canon}inletg=change_nodegyinletltle,y,g=get_ltlegyinletf_traversed,g=PMap.fold(funz_(f_traversed,g)->forward_traversef_traversedgv_klvlyz)ltle(f_traversed,g)iny.canon::f_traversed,gendelseify.klvl=v_klvl&&x!=ythenletg=change_nodeg{ywithgtge=PSet.addx.canony.gtge}inf_traversed,gelsef_traversed,gletrecfind_to_mergeto_revertgxv=letx=reprgxinmatchx.statuswith|Visited->false,to_revert|ToMerge->true,to_revert|NoMark->letto_revert=x::to_revertinifIndex.equalx.canonvthenbeginx.status<-ToMerge;true,to_revertendelsebeginletmerge,to_revert=PSet.fold(funy(merge,to_revert)->letmerge',to_revert=find_to_mergeto_revertgyvinmerge'||merge,to_revert)x.gtge(false,to_revert)inx.status<-ifmergethenToMergeelseVisited;merge,to_revertend|_->assertfalseletget_new_edgesgto_merge=(* Computing edge sets. *)letltle=letfoldaccn=letfoldustrictacc=matchPMap.finduaccwith|true->acc|false->ifstrictthenPMap.addutrueaccelseacc|exceptionNot_found->PMap.addustrictaccinPMap.foldfoldn.ltleaccinmatchto_mergewith|[]->assertfalse|hd::tl->List.fold_leftfoldhd.ltletlinletltle,_=clean_ltlegltleinletfoldaccua=matchPMap.finda.canonltlewith|true->(* There is a lt edge inside the new component. This is a
"bad cycle". *)raiseCycleDetected|false->PMap.removea.canonaccu|exceptionNot_found->accuinletltle=List.fold_leftfoldltleto_mergeinletgtge=List.fold_left(funaccn->PSet.unionaccn.gtge)PSet.emptyto_mergeinletgtge,_=clean_gtgeggtgeinletgtge=List.fold_left(funaccn->PSet.removen.canonacc)gtgeto_mergein(ltle,gtge)letreorderguv=(* STEP 2: backward search in the k-level of u. *)(* [v_klvl] is the chosen future level for u, v and all
traversed nodes. *)letb_traversed,v_klvl,g=letu=reprguintryletto_revert,b_traversed,_,g=backward_traverse[][](u.klvl+1)guinrevert_graphto_revertg;letv_klvl=u.klvlinb_traversed,v_klvl,gwithAbortBackwardg->(* Backward search was too long, use the next k-level. *)letv_klvl=u.klvl+1in[],v_klvl,ginletf_traversed,g=(* STEP 3: forward search. Contrary to what is described in
the paper, we do not test whether v_klvl = u.klvl nor we assign
v_klvl to v.klvl. Indeed, the first call to forward_traverse
will do all that. *)forward_traverse[]gv_klvl(reprgv)vin(* STEP 4: merge nodes if needed. *)letto_merge,b_reindex,f_reindex=if(reprgu).klvl=v_klvlthenbeginletmerge,to_revert=find_to_merge[]guvinletr=ifmergethenList.filter(funu->u.status=ToMerge)to_revert,List.filter(funu->(reprgu).status<>ToMerge)b_traversed,List.filter(funu->(reprgu).status<>ToMerge)f_traversedelse[],b_traversed,f_traversedinList.iter(funu->u.status<-NoMark)to_revert;rendelse[],b_traversed,f_traversedinletto_reindex,g=matchto_mergewith|[]->List.rev_appendf_reindexb_reindex,g|n0::q0->(* Computing new root. *)letroot,rank_rest=List.fold_left(fun((best,_rank_rest)asacc)n->ifn.rank>=best.rankthenn,best.rankelseacc)(n0,min_int)q0inletltle,gtge=get_new_edgesgto_mergein(* Inserting the new root. *)letg=change_nodeg{rootwithltle;gtge;rank=maxroot.rank(rank_rest+1);}in(* Inserting shortcuts for old nodes. *)letg=List.fold_left(fungn->ifIndex.equaln.canonroot.canonthengelseenter_equivgn.canonroot.canon)gto_mergein(* Updating g.n_edges *)letoldsz=List.fold_left(funszu->sz+PMap.cardinalu.ltle)0to_mergeinletsz=PMap.cardinalltleinletg={gwithn_edges=g.n_edges+sz-oldsz}in(* Not clear in the paper: we have to put the newly
created component just between B and F. *)List.rev_appendf_reindex(root.canon::b_reindex),gin(* STEP 5: reindex traversed nodes. *)List.fold_leftuse_indexgto_reindex(* Assumes [u] and [v] are already in the graph. *)(* Does NOT assume that ucan != vcan. *)letinsert_edgestrictucanvcang=tryletu=ucan.canonandv=vcan.canonin(* STEP 1: do we need to reorder nodes ? *)letg=iftopo_compareucanvcan<=0thengelsereorderguvin(* STEP 6: insert the new edge in the graph. *)letu=reprguinletv=reprgvinifu==vthenifstrictthenraiseCycleDetectedelsegelseletg=tryletoldstrict=PMap.findv.canonu.ltleinifstrict&¬oldstrictthenchange_nodeg{uwithltle=PMap.addv.canontrueu.ltle}elsegwithNot_found->{(change_nodeg{uwithltle=PMap.addv.canonstrictu.ltle})withn_edges=g.n_edges+1}inifu.klvl<>v.klvl||PSet.memu.canonv.gtgethengelseletv={vwithgtge=PSet.addu.canonv.gtge}inchange_nodegvwith|CycleDetectedase->raisee|e->(* Unlikely event: fatal error or signal *)let()=cleanup_marksginraiseeletadd?(rank=0)vg=ifIndex.memvg.tablethenraiseAlreadyDeclaredelselet()=assert(g.index>min_int)inletv,table=Index.freshvg.tableinletnode={canon=v;ltle=PMap.empty;gtge=PSet.empty;rank;klvl=0;ilvl=g.index;status=NoMark;}inletentries=PMap.addv(Canonicalnode)g.entriesin{entries;index=g.index-1;n_nodes=g.n_nodes+1;n_edges=g.n_edges;table}exceptionUndeclaredofPoint.tletcheck_declaredgus=letcheckl=ifnot(Index.memlg.table)thenraise(Undeclaredl)inPoint.Set.itercheckusexceptionFound_explanationof(constraint_type*Point.t)listletget_explanationstrictuvg=letu=Index.findug.tableinletv=repr_nodegvinletvisited_strict=refPMap.emptyinletrectraversestrictu=ifu==vthenifstrictthenNoneelseSome[]elseiftopo_compareuv=1thenNoneelseletvisited=trynot(PMap.findu.canon!visited_strict)||strictwithNot_found->falseinifvisitedthenNoneelsebeginvisited_strict:=PMap.addu.canonstrict!visited_strict;tryPMap.iter(funu'strictu'->matchtraverse(strict&¬strictu')(reprgu')with|None->()|Someexp->lettyp=ifstrictu'thenLtelseLeinletu'=Index.repru'g.tableinraise(Found_explanation((typ,u')::exp)))u.ltle;NonewithFound_explanationexp->Someexpendinletu=reprguinifu==vthen[(Eq,Index.reprv.canong.table)]elsematchtraversestrictuwithSomeexp->exp|None->assertfalseletget_explanationstrictuvg=Some(lazy(get_explanationstrictuvg))(* To compare two nodes, we simply do a forward search.
We implement two improvements:
- we ignore nodes that are higher than the destination;
- we do a BFS rather than a DFS because we expect to have a short
path (typically, the shortest path has length 1)
*)exceptionFoundofcanonical_nodelistletsearch_pathstrictuvg=letrecloopto_reverttodonext_todo=matchtodo,next_todowith|[],[]->to_revert(* No path found *)|[],_->loopto_revertnext_todo[]|(u,strict)::todo,_->ifu.status=Visited||(u.status=WeakVisited&&strict)thenloopto_reverttodonext_todoelseletto_revert=ifu.status=NoMarkthenu::to_revertelseto_revertinu.status<-ifstrictthenWeakVisitedelseVisited;iftryPMap.findv.canonu.ltle||notstrictwithNot_found->falsethenraise(Foundto_revert)elsebeginletnext_todo=PMap.fold(funustrictunext_todo->letstrict=notstrictu&&strictinletu=reprguinifu==v&¬strictthenraise(Foundto_revert)elseiftopo_compareuv=1thennext_todoelse(u,strict)::next_todo)u.ltlenext_todoinloopto_reverttodonext_todoendinifu==vthennotstrictelsetryletres,to_revert=tryfalse,loop[][u,strict][]withFoundto_revert->true,to_revertinList.iter(funu->u.status<-NoMark)to_revert;reswithe->(* Unlikely event: fatal error or signal *)let()=cleanup_marksginraisee(** Uncomment to debug the cycle detection algorithm. *)(*let insert_edge strict ucan vcan g =
let check_invariants = check_invariants ~required_canonical:(fun _ -> false) in
check_invariants g;
let g = insert_edge strict ucan vcan g in
check_invariants g;
let ucan = repr g ucan.canon in
let vcan = repr g vcan.canon in
assert (search_path strict ucan vcan g);
g*)(** User interface *)type'acheck_function=t->'a->'a->boolletcheck_eqguv=u==v||letarcu=repr_nodeguandarcv=repr_nodegvinarcu==arcvletcheck_smallergstrictuv=search_pathstrict(repr_nodegu)(repr_nodegv)gletcheck_leqguv=check_smallergfalseuvletcheck_ltguv=check_smallergtrueuv(* enforce_eq g u v will force u=v if possible, will fail otherwise *)letenforce_equvg=letucan=repr_nodeguinletvcan=repr_nodegvinifucan==vcanthengelseiftopo_compareucanvcan=1thenletucan=vcanandvcan=ucaninletg=insert_edgefalseucanvcangin(* Cannot fail *)tryinsert_edgefalsevcanucangwithCycleDetected->Point.error_inconsistencyEqvu(get_explanationtruevug)elseletg=insert_edgefalseucanvcangin(* Cannot fail *)tryinsert_edgefalsevcanucangwithCycleDetected->Point.error_inconsistencyEqvu(get_explanationtrueuvg)(* enforce_leq g u v will force u<=v if possible, will fail otherwise *)letenforce_lequvg=letucan=repr_nodeguinletvcan=repr_nodegvintryinsert_edgefalseucanvcangwithCycleDetected->Point.error_inconsistencyLeuv(get_explanationtruevug)(* enforce_lt u v will force u<v if possible, will fail otherwise *)letenforce_ltuvg=letucan=repr_nodeguinletvcan=repr_nodegvintryinsert_edgetrueucanvcangwithCycleDetected->Point.error_inconsistencyLtuv(get_explanationfalsevug)letempty={entries=PMap.empty;index=0;n_nodes=0;n_edges=0;table=Index.empty}(* Normalization *)letconstraints_ofg=letmoduleUF=Unionfind.Make(Point.Set)(Point.Map)inletuf=UF.create()inletconstraints_ofuvacc=matchvwith|Canonical{canon=u;ltle;_}->PMap.fold(funvstrictacc->lettyp=ifstrictthenLtelseLeinletu=Index.reprug.tableinletv=Index.reprvg.tableinConstraint.add(u,typ,v)acc)ltleacc|Equivv->letu=Index.reprug.tableinletv=Index.reprvg.tableinUF.unionuvuf;accinletcsts=PMap.foldconstraints_ofg.entriesConstraint.emptyincsts,UF.partitionuf(* domain g.entries = kept + removed *)letconstraints_for~keptg=(* rmap: partial map from canonical points to kept points *)letadd_cstukndvcst=Constraint.add(Index.reprug.table,knd,Index.reprvg.table)cstinletkept=Point.Set.fold(funuaccu->PSet.add(Index.findug.table)accu)keptPSet.emptyinletrmap,csts=PSet.fold(funu(rmap,csts)->letarcu=reprguinifPSet.memarcu.canonkeptthenletcsts=ifIndex.equaluarcu.canonthencstselseadd_cstuEqarcu.canoncstsinPMap.addarcu.canonarcu.canonrmap,cstselsematchPMap.findarcu.canonrmapwith|v->rmap,add_cstuEqvcsts|exceptionNot_found->PMap.addarcu.canonurmap,csts)kept(PMap.empty,Constraint.empty)inletrecadd_fromucststodo=matchtodowith|[]->csts|(v,strict)::todo->letv=reprgvin(matchPMap.findv.canonrmapwith|v->letd=ifstrictthenLtelseLeinletcsts=add_cstudvcstsinadd_fromucststodo|exceptionNot_found->(* v is not equal to any kept point *)lettodo=PMap.fold(funv'strict'todo->(v',strict||strict')::todo)v.ltletodoinadd_fromucststodo)inPSet.fold(funucsts->letarc=reprguinPMap.fold(funvstrictcsts->add_fromucsts[v,strict])arc.ltlecsts)keptcstsletdomaing=letfoldu_accu=Point.Set.add(Index.reprug.table)accuinPMap.foldfoldg.entriesPoint.Set.emptyletchoosepgu=letexceptionFoundofPoint.tinletru=(repr_nodegu).canoninletruv=Index.reprrug.tableinifpruvthenSomeruvelsetryPMap.iter(funv->function|Canonical_->()(* we already tried [p ru] *)|Equivv'->letrv=(reprgv').canoninifrv==ruthenletv=Index.reprvg.tableinifpvthenraise(Foundv)(* NB: we could also try [p v'] but it will come up in the
rest of the iteration regardless. *))g.entries;NonewithFoundv->Somevtypenode=AliasofPoint.t|NodeofboolPoint.Map.ttyperepr=nodePoint.Map.tletreprg=letfoldunaccu=letn=matchnwith|Canonicaln->letfoldultaccu=Point.Map.add(Index.reprug.table)ltaccuinletltle=PMap.foldfoldn.ltlePoint.Map.emptyinNodeltle|Equivu->Alias(Index.reprug.table)inPoint.Map.add(Index.reprug.table)naccuinPMap.foldfoldg.entriesPoint.Map.emptyend