aboutsummaryrefslogtreecommitdiffhomepage
path: root/checker/declarations.mli
diff options
context:
space:
mode:
authorGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2011-04-03 11:23:31 +0000
committerGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2011-04-03 11:23:31 +0000
commit5681594c83c2ba9a2c0e21983cac0f161ff95f02 (patch)
treeea458a8321f71b3e2fba5d67cfc3f79866241d48 /checker/declarations.mli
parentda1e32cbdc78050ea2e89eee896ba2b40db1b5dd (diff)
Lazy loading of opaque proofs: fast as -dont-load-proofs without its drawbacks
The recent experiment with -dont-load-proofs in the stdlib showed that this options isn't fully safe: some axioms were generated (Include ? functor application ? This is still to be fully understood). Instead, I've implemented an idea of Yann: only load opaque proofs when we need them. This is almost as fast as -dont-load-proofs (on the stdlib, we're now 15% faster than before instead of 20% faster with -dont-load-proofs), but fully compatible with Coq standard behavior. Technically, the const_body field of Declarations.constant_body now regroup const_body + const_opaque + const_inline in a ternary type. It is now either: - Undef : an axiom or parameter, with an inline info - Def : a transparent definition, with a constr_substituted - OpaqueDef : an opaque definition, with a lazy constr_substitued Accessing the lazy constr of an OpaqueDef might trigger the read on disk of the final section of a .vo, where opaque proofs are located. Some functions (body_of_constant, is_opaque, constant_has_body) emulate the behavior of the old fields. The rest of Coq (including the checker) has been adapted accordingly, either via direct access to the new const_body or via these new functions. Many places look nicer now (ok, subjective notion). There are now three options: -lazy-load-proofs (default), -force-load-proofs (earlier semantics), -dont-load-proofs. Note that -outputstate now implies -force-load-proofs (otherwise the marshaling fails on some delayed lazy). On the way, I fixed what looked like a bug : a module type (T with Definition x := c) was accepted even when x in T was opaque. I also tried to clarify Subtyping.check_constant. git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@13952 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'checker/declarations.mli')
-rw-r--r--checker/declarations.mli34
1 files changed, 29 insertions, 5 deletions
diff --git a/checker/declarations.mli b/checker/declarations.mli
index b5038996c..90beb326b 100644
--- a/checker/declarations.mli
+++ b/checker/declarations.mli
@@ -29,16 +29,40 @@ type constr_substituted
val force_constr : constr_substituted -> constr
val from_val : constr -> constr_substituted
-type inline = int option (* inlining level, None for no inlining *)
+(** 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
+
+(** Inlining level of parameters at functor applications.
+ This is ignored by the checker. *)
+
+type inline = int option
+
+(** A constant can have no body (axiom/parameter), or a
+ transparent body, or an opaque one *)
+
+type constant_def =
+ | Undef of inline
+ | Def of constr_substituted
+ | OpaqueDef of lazy_constr
type constant_body = {
const_hyps : section_context; (* New: younger hyp at top *)
- const_body : constr_substituted option;
+ const_body : constant_def;
const_type : constant_type;
const_body_code : to_patch_substituted;
- const_constraints : Univ.constraints;
- const_opaque : bool;
- const_inline : inline }
+ const_constraints : Univ.constraints }
+
+val body_of_constant : constant_body -> constr_substituted option
+val constant_has_body : constant_body -> bool
+val is_opaque : constant_body -> bool
(* Mutual inductives *)