aboutsummaryrefslogtreecommitdiffhomepage
path: root/kernel/term.ml
diff options
context:
space:
mode:
authorGravatar ppedrot <ppedrot@85f007b7-540e-0410-9357-904b9bb8a0f7>2012-09-26 20:38:55 +0000
committerGravatar ppedrot <ppedrot@85f007b7-540e-0410-9357-904b9bb8a0f7>2012-09-26 20:38:55 +0000
commitb12466ef14d4bdf13af9c2f772c692ee1760da2d (patch)
treecfd2c437c5836f4132040f2b89bbdfc0f708ea44 /kernel/term.ml
parentfb1eb17d40febf3f4927568921b64b7a8c7ca028 (diff)
Reusing the Hashset data structure in Hashcons. Hopefully, this should
not disrupt anything... git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@15836 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'kernel/term.ml')
-rw-r--r--kernel/term.ml7
1 files changed, 5 insertions, 2 deletions
diff --git a/kernel/term.ml b/kernel/term.ml
index 8ab04af24..42309c37e 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -1236,7 +1236,10 @@ let equals_constr t1 t2 =
(** Note that the following Make has the side effect of creating
once and for all the table we'll use for hash-consing all constr *)
-module H = Hashset.Make(struct type t = constr let equal = equals_constr end)
+module HashsetTerm = Hashset.Make(struct type t = constr let equal = equals_constr end)
+
+let term_table = HashsetTerm.create 19991
+(* The associative table to hashcons terms. *)
open Hashset.Combine
@@ -1320,7 +1323,7 @@ let hcons_term (sh_sort,sh_ci,sh_construct,sh_ind,sh_con,sh_na,sh_id) =
let (y, h) = hash_term t in
(* [h] must be positive. *)
let h = h land 0x3FFFFFFF in
- (H.may_add_and_get h y, h)
+ (HashsetTerm.repr h y term_table, h)
in
(* Make sure our statically allocated Rels (1 to 16) are considered