aboutsummaryrefslogtreecommitdiffhomepage
path: root/kernel/cooking.ml
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 /kernel/cooking.ml
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 'kernel/cooking.ml')
-rw-r--r--kernel/cooking.ml24
1 files changed, 16 insertions, 8 deletions
diff --git a/kernel/cooking.ml b/kernel/cooking.ml
index 87474b863..02330339d 100644
--- a/kernel/cooking.ml
+++ b/kernel/cooking.ml
@@ -117,16 +117,24 @@ type recipe = {
d_abstract : named_context;
d_modlist : work_list }
-let on_body f =
- Option.map (fun c -> Declarations.from_val (f (Declarations.force c)))
+let on_body f = function
+ | Undef inl -> Undef inl
+ | Def cs -> Def (Declarations.from_val (f (Declarations.force cs)))
+ | OpaqueDef lc ->
+ OpaqueDef (Declarations.opaque_from_val (f (Declarations.force_opaque lc)))
+
+let constr_of_def = function
+ | Undef _ -> assert false
+ | Def cs -> Declarations.force cs
+ | OpaqueDef lc -> Declarations.force_opaque lc
let cook_constant env r =
let cb = r.d_from in
let hyps = Sign.map_named_context (expmod_constr r.d_modlist) r.d_abstract in
- let body =
- on_body (fun c ->
- abstract_constant_body (expmod_constr r.d_modlist c) hyps)
- cb.const_body in
+ let body = on_body
+ (fun c -> abstract_constant_body (expmod_constr r.d_modlist c) hyps)
+ cb.const_body
+ in
let typ = match cb.const_type with
| NonPolymorphicType t ->
let typ = abstract_constant_type (expmod_constr r.d_modlist t) hyps in
@@ -134,7 +142,7 @@ let cook_constant env r =
| PolymorphicArity (ctx,s) ->
let t = mkArity (ctx,Type s.poly_level) in
let typ = abstract_constant_type (expmod_constr r.d_modlist t) hyps in
- let j = make_judge (force (Option.get body)) typ in
+ let j = make_judge (constr_of_def body) typ in
Typeops.make_polymorphic_if_constant_for_ind env j
in
- (body, typ, cb.const_constraints, cb.const_opaque, None)
+ (body, typ, cb.const_constraints)