Skip to content

Commit

Permalink
hconstr v6: hcons in hconstr.of_constr
Browse files Browse the repository at this point in the history
  • Loading branch information
SkySkimmer committed Jun 12, 2024
1 parent 435e5eb commit 449b26f
Show file tree
Hide file tree
Showing 3 changed files with 28 additions and 5 deletions.
27 changes: 25 additions & 2 deletions kernel/hConstr.ml
Original file line number Diff line number Diff line change
Expand Up @@ -406,6 +406,7 @@ type t = {
kind : (t,t,Sorts.t,Instance.t,Sorts.relevance) kind_of_term;
isRel : int (* 0 -> not a rel, otherwise unique identifier of that binder *);
hash : int;
hcons_hash : int;
mutable refcount : int;
}

Expand Down Expand Up @@ -591,26 +592,48 @@ let hcons_inplace f a = Array.iteri (fun i x -> Array.unsafe_set a i (f x)) a
let of_kind_nohashcons kind =
let self = kind_to_constr kind in
let hash = hash_kind hash kind in
let hcons_hash = hash_kind (fun x -> x.hcons_hash) kind in
let () = match kind with
| Rel _ -> assert false
| _ -> ()
in
{ self; kind; hash; isRel = 0; refcount = 1 }
{ self; kind; hash; hcons_hash; isRel = 0; refcount = 1 }

let hcons_array a ka =
let h = hash_array (fun c -> c.hcons_hash) ka in
Array.iteri (fun i c -> Array.unsafe_set a i c.self) ka;
HashsetTermArray.repr h a term_array_table

let hcons_arrays self kind =
match Constr.kind self, kind with
| App (h, args), App (_, kargs) ->
make @@ App (h, hcons_array args kargs)
| Case (v1,v2,pms,v3,v4,v5,v6), Case (_,_,kpms,_,_,_,_) ->
make @@ Case (v1,v2,hcons_array pms kpms,v3,v4,v5,v6)
| Fix (ln,(lna,tl,bl)), Fix (_,(_,ktl,kbl)) ->
make @@ Fix (ln, (lna,hcons_array tl ktl, hcons_array bl kbl))
| CoFix (ln,(lna,tl,bl)), CoFix (_,(_,ktl,kbl)) ->
make @@ CoFix (ln, (lna,hcons_array tl ktl, hcons_array bl kbl))
| Array (u,t,def,ty), Array (_,kt,_,_) -> make @@ Array (u,hcons_array t kt,def,ty)
| _ -> self

let rec of_constr tbl local_env c =
let c =
let kind = of_constr_aux tbl local_env c in
let self = kind_to_constr kind in
let self = if hasheq (Constr.kind self) (Constr.kind c) then c else self in
let hash = hash_kind hash kind in
let hcons_hash = hash_kind (fun x -> x.hcons_hash) kind in
let self = hcons_arrays self kind in
let self = HashsetTerm.repr hcons_hash self term_table in
let isRel, hash = match kind with
| Rel n ->
let uid = Range.get local_env.rels (n-1) in
assert (uid <> 0);
uid, Hashset.Combine.combine uid hash
| _ -> 0, hash
in
{ self; kind; hash; isRel; refcount = 1 }
{ self; kind; hash; hcons_hash; isRel; refcount = 1 }
in
match Tbl.find_opt tbl c with
| Some c' -> c'.refcount <- c'.refcount + 1; c'
Expand Down
2 changes: 1 addition & 1 deletion kernel/safe_typing.ml
Original file line number Diff line number Diff line change
Expand Up @@ -670,7 +670,7 @@ let add_constant_aux senv (kn, cb) =
let l = Constant.label kn in
(* This is the only place where we hashcons the contents of a constant body *)
let senv, cb = push_const_bytecode senv cb in
let cb = if sections_are_opened senv then cb else Declareops.hcons_const_body cb in
let cb = if true || sections_are_opened senv then cb else Declareops.hcons_const_body cb in
let senv' = add_field (l,SFBconst cb) (C kn) senv in
let senv'' = match cb.const_body with
| Undef (Some lev) ->
Expand Down
4 changes: 2 additions & 2 deletions kernel/typeops.ml
Original file line number Diff line number Diff line change
Expand Up @@ -916,7 +916,7 @@ let check_wellformed_universes env c =

let infer env constr =
let () = check_wellformed_universes env constr in
let _constr, t = execute env constr in
let constr, t = execute env constr in
make_judge constr t

let assumption_of_judgment env {uj_val=c; uj_type=t} =
Expand All @@ -928,7 +928,7 @@ let type_judgment env {uj_val=c; uj_type=t} =

let infer_type env constr =
let () = check_wellformed_universes env constr in
let _constr, t = execute env constr in
let constr, t = execute env constr in
let s = check_type env constr t in
{utj_val = constr; utj_type = s}

Expand Down

0 comments on commit 449b26f

Please sign in to comment.