From 5681594c83c2ba9a2c0e21983cac0f161ff95f02 Mon Sep 17 00:00:00 2001 From: letouzey Date: Sun, 3 Apr 2011 11:23:31 +0000 Subject: 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 --- kernel/declarations.ml | 67 +++++++++++++++++++++++++++++++++++++++++--------- 1 file changed, 56 insertions(+), 11 deletions(-) (limited to 'kernel/declarations.ml') diff --git a/kernel/declarations.ml b/kernel/declarations.ml index fc8064d59..f092c8261 100644 --- a/kernel/declarations.ml +++ b/kernel/declarations.ml @@ -49,17 +49,64 @@ let force = force subst_mps let subst_constr_subst = subst_substituted -type inline = int option (* inlining level, None for no inlining *) +(** Opaque proof terms are not loaded immediately, but are there + in a lazy form. Forcing this lazy may trigger some unmarshal of + the necessary structure. The ['a substituted] type isn't really great + here, so we store "manually" a substitution list, the younger one at top. +*) + +type lazy_constr = constr_substituted Lazy.t * substitution list + +let force_lazy_constr (c,l) = + List.fold_right subst_constr_subst l (Lazy.force c) + +let lazy_constr_is_val (c,_) = Lazy.lazy_is_val c + +let make_lazy_constr c = (c, []) + +let force_opaque lc = force (force_lazy_constr lc) + +let opaque_from_val c = (Lazy.lazy_from_val (from_val c), []) + +let subst_lazy_constr sub (c,l) = (c,sub::l) + +(** Inlining level of parameters at functor applications. + None means no inlining *) + +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 + +let subst_constant_def sub = function + | Undef inl -> Undef inl + | Def c -> Def (subst_constr_subst sub c) + | OpaqueDef lc -> OpaqueDef (subst_lazy_constr sub lc) 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 : Cemitcodes.to_patch_substituted; - (* const_type_code : Cemitcodes.to_patch; *) - const_constraints : constraints; - const_opaque : bool; - const_inline : inline } + const_constraints : constraints } + +let body_of_constant cb = match cb.const_body with + | Undef _ -> None + | Def c -> Some c + | OpaqueDef lc -> Some (force_lazy_constr lc) + +let constant_has_body cb = match cb.const_body with + | Undef _ -> false + | Def _ | OpaqueDef _ -> true + +let is_opaque cb = match cb.const_body with + | OpaqueDef _ -> true + | Undef _ | Def _ -> false (*s Inductive types (internal representation with redundant information). *) @@ -218,14 +265,12 @@ let subst_arity sub arity = (* TODO: should be changed to non-coping after Term.subst_mps *) let subst_const_body sub cb = { const_hyps = (assert (cb.const_hyps=[]); []); - const_body = Option.map (subst_constr_subst sub) cb.const_body; + const_body = subst_constant_def sub cb.const_body; const_type = subst_arity sub cb.const_type; const_body_code = Cemitcodes.subst_to_patch_subst sub cb.const_body_code; (*const_type_code = Cemitcodes.subst_to_patch sub cb.const_type_code;*) - const_constraints = cb.const_constraints; - const_opaque = cb.const_opaque; - const_inline = cb.const_inline} - + const_constraints = cb.const_constraints} + let subst_arity sub = function | Monomorphic s -> Monomorphic { -- cgit v1.2.3