Skip to content

Commit

Permalink
Avoid duplicate hashconsing of kerpairs
Browse files Browse the repository at this point in the history
  • Loading branch information
SkySkimmer committed Jul 1, 2024
1 parent b4b8f68 commit 76089d6
Show file tree
Hide file tree
Showing 3 changed files with 31 additions and 7 deletions.
19 changes: 19 additions & 0 deletions clib/hashset.ml
Original file line number Diff line number Diff line change
Expand Up @@ -34,6 +34,7 @@ module type S = sig
val create : int -> t
val clear : t -> unit
val repr : int -> elt -> t -> elt
val lookup : (elt -> elt -> bool) -> int -> elt -> t -> elt option
val stats : t -> statistics
end

Expand Down Expand Up @@ -180,6 +181,24 @@ module Make (E : EqType) =

external unsafe_weak_get : 'a Weak.t -> int -> 'a option = "caml_weak_get"

let lookup eq h d t =
let table = t.table in
let index = get_index table h in
let bucket = Array.unsafe_get table index in
let hashes = Array.unsafe_get t.hashes index in
let sz = Weak.length bucket in
let pos = ref 0 in
let ans = ref None in
while !pos < sz && !ans == None do
let i = !pos in
if Int.equal h (Array.unsafe_get hashes i) then begin
match unsafe_weak_get bucket i with
| Some v as res when eq v d -> ans := res
| _ -> incr pos
end else incr pos
done;
!ans

let repr h d t =
let table = t.table in
let index = get_index table h in
Expand Down
2 changes: 2 additions & 0 deletions clib/hashset.mli
Original file line number Diff line number Diff line change
Expand Up @@ -48,6 +48,8 @@ module type S = sig
[constr] is stored in [set] and will be used as the canonical
representation of this value in the future. *)

val lookup : (elt -> elt -> bool) -> int -> elt -> t -> elt option

val stats : t -> statistics
(** Recover statistics on the table. *)
end
Expand Down
17 changes: 10 additions & 7 deletions kernel/names.ml
Original file line number Diff line number Diff line change
Expand Up @@ -565,7 +565,6 @@ module KerPair = struct
module Self_Hashcons =
struct
type t = kernel_pair
type u = KerName.t -> KerName.t
let hashcons hkn = function
| Same kn -> Same (hkn kn)
| Dual (knu,knc) -> make (hkn knu) (hkn knc)
Expand All @@ -579,15 +578,19 @@ module KerPair = struct
the same canonical part is a logical invariant of the system, it
is not necessarily an invariant in memory, so we treat kernel
names as they are syntactically for hash-consing) *)
let hash = function
| Same kn -> KerName.hash kn
| Dual (knu, knc) ->
Hashset.Combine.combine (KerName.hash knu) (KerName.hash knc)
end

module HashKP = Hashcons.Make(Self_Hashcons)
module Tbl = Hashset.Make(Self_Hashcons)

let tab = Tbl.create 97

let hcons = Hashcons.simple_hcons HashKP.generate HashKP.hcons KerName.hcons
let hcons c =
let h = SyntacticOrd.hash c in
match Tbl.lookup SyntacticOrd.equal h c tab with
| Some v -> v
| None ->
let v = Self_Hashcons.hashcons KerName.hcons c in
Tbl.repr h v tab

end

Expand Down

0 comments on commit 76089d6

Please sign in to comment.