aboutsummaryrefslogtreecommitdiffhomepage
path: root/checker/declarations.mli
diff options
context:
space:
mode:
authorGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2013-04-02 22:08:44 +0000
committerGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2013-04-02 22:08:44 +0000
commitf5ab2e37b0609d8edb8d65dfae49741442a90657 (patch)
tree72bb704f147a824b743566b447c4e98685ab2db6 /checker/declarations.mli
parent5635c35ea4ec172fd81147effed4f33e2f581aaa (diff)
Revised infrastructure for lazy loading of opaque proofs
Get rid of the LightenLibrary hack : no more last-minute collect of opaque terms and Obj.magic tricks. Instead, we make coqc accumulate the opaque terms as soon as constant_bodies are created outside sections. In these cases, the opaque terms are placed in a special table, and some (DirPath.t * int) are used as indexes in constant_body. In an interactive session, the local opaque terms stay directly stored in the constant_body. The structure of .vo file stays similar : magic number, regular library structure, digest of the first part, array of opaque terms. In addition, we now have a final checksum for checking the integrity of the whole .vo file. The other difference is that lazy_constr aren't changed into int indexes in .vo files, but are now coded as (substitution list * DirPath.t * int). In particular this approach allows to refer to opaque terms from another library. This (and accumulating substitutions in lazy_constr) seems to greatly help decreasing the size of opaque tables : -20% of vo size on the standard library :-). The compilation times are slightly better, but that can be statistic noise. The -force-load-proofs isn't active anymore : it behaves now just like -lazy-load-proofs. The -dont-load-proofs mode has slightly changed : opaque terms aren't seen as axioms anymore, but accessing their bodies will raise an error. Btw, API change : Declareops.body_of_constant now produces directly a constr option instead of a constr_substituted option git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@16382 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'checker/declarations.mli')
-rw-r--r--checker/declarations.mli13
1 files changed, 3 insertions, 10 deletions
diff --git a/checker/declarations.mli b/checker/declarations.mli
index 80c895bbe..cc3123ca7 100644
--- a/checker/declarations.mli
+++ b/checker/declarations.mli
@@ -30,16 +30,9 @@ type constr_substituted
val force_constr : constr_substituted -> constr
val from_val : constr -> constr_substituted
-(** Beware! In .vo files, lazy_constr are stored as integers
- used as indexes for a separate table. The actual lazy_constr is restored
- later, by [Safe_typing.LightenLibrary.load]. This allows us
- to use here a different definition of lazy_constr than coqtop:
- since the checker will inspect all proofs parts, even opaque
- ones, no need to use Lazy.t here *)
-
type lazy_constr
-val force_lazy_constr : lazy_constr -> constr
-val lazy_constr_from_val : constr_substituted -> lazy_constr
+
+val indirect_opaque_access : (DirPath.t -> int -> constr) ref
(** Inlining level of parameters at functor applications.
This is ignored by the checker. *)
@@ -63,7 +56,7 @@ type constant_body = {
const_native_name : native_name ref;
const_inline_code : bool }
-val body_of_constant : constant_body -> constr_substituted option
+val body_of_constant : constant_body -> constr option
val constant_has_body : constant_body -> bool
val is_opaque : constant_body -> bool