aboutsummaryrefslogtreecommitdiffhomepage
path: root/checker
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
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')
-rw-r--r--checker/check.ml17
-rw-r--r--checker/check_stat.ml2
-rw-r--r--checker/declarations.ml69
-rw-r--r--checker/declarations.mli34
-rw-r--r--checker/environ.ml6
-rw-r--r--checker/mod_checking.ml21
-rw-r--r--checker/modops.ml19
-rw-r--r--checker/safe_typing.ml42
-rw-r--r--checker/safe_typing.mli10
-rw-r--r--checker/subtyping.ml40
10 files changed, 157 insertions, 103 deletions
diff --git a/checker/check.ml b/checker/check.ml
index 3dc510dcc..bb42b949d 100644
--- a/checker/check.ml
+++ b/checker/check.ml
@@ -270,10 +270,10 @@ let with_magic_number_check f a =
(************************************************************************)
(* Internalise libraries *)
-let mk_library md f get_table digest = {
+let mk_library md f table digest = {
library_name = md.md_name;
library_filename = f;
- library_compiled = Safe_typing.LightenLibrary.load ~load_proof:true get_table md.md_compiled;
+ library_compiled = Safe_typing.LightenLibrary.load table md.md_compiled;
library_deps = md.md_deps;
library_digest = digest }
@@ -287,24 +287,21 @@ let depgraph = ref LibraryMap.empty
let intern_from_file (dir, f) =
Flags.if_verbose msg (str"[intern "++str f++str" ...");
- let (md,get_table,close,digest) =
+ let (md,table,digest) =
try
let ch = with_magic_number_check raw_intern_library f in
let (md:library_disk) = System.marshal_in ch in
let digest = System.marshal_in ch in
- let get_table () =
- (System.marshal_in ch : Safe_typing.LightenLibrary.table)
- in
+ let table = (System.marshal_in ch : Safe_typing.LightenLibrary.table) in
+ close_in ch;
if dir <> md.md_name then
errorlabstrm "load_physical_library"
(name_clash_message dir md.md_name f);
Flags.if_verbose msgnl(str" done]");
- md,get_table,(fun () -> close_in ch),digest
+ md,table,digest
with e -> Flags.if_verbose msgnl(str" failed!]"); raise e in
depgraph := LibraryMap.add md.md_name md.md_deps !depgraph;
- let library = mk_library md f get_table digest in
- close ();
- library
+ mk_library md f table digest
let get_deps (dir, f) =
try LibraryMap.find dir !depgraph
diff --git a/checker/check_stat.ml b/checker/check_stat.ml
index 7ca406fec..5f28269ee 100644
--- a/checker/check_stat.ml
+++ b/checker/check_stat.ml
@@ -37,7 +37,7 @@ let cst_filter f csts =
(fun c ce acc -> if f c ce then c::acc else acc)
csts []
-let is_ax _ cb = cb.const_body = None
+let is_ax _ cb = not (constant_has_body cb)
let pr_ax csts =
let axs = cst_filter is_ax csts in
diff --git a/checker/declarations.ml b/checker/declarations.ml
index 084b4eb54..890996d10 100644
--- a/checker/declarations.ml
+++ b/checker/declarations.ml
@@ -478,27 +478,67 @@ let val_cstr_subst = val_substituted val_constr
let subst_constr_subst = subst_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 = constr_substituted
+let subst_lazy_constr = subst_substituted
+let force_lazy_constr = force_constr
+let lazy_constr_from_val c = c
+let val_lazy_constr = val_cstr_subst
+
+(** 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
+
+let val_cst_def =
+ val_sum "constant_def" 0
+ [|[|val_opt val_int|]; [|val_cstr_subst|]; [|val_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 : to_patch_substituted;
- (* const_type_code : Cemitcodes.to_patch; *)
- const_constraints : Univ.constraints;
- const_opaque : bool;
- const_inline : inline }
+ const_constraints : Univ.constraints }
+
+let body_of_constant cb = match cb.const_body with
+ | Undef _ -> None
+ | Def c -> Some c
+ | OpaqueDef c -> Some c
+
+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
+ | Def _ | Undef _ -> false
let val_cb = val_tuple ~name:"constant_body"
[|val_nctxt;
- val_opt val_cstr_subst;
+ val_cst_def;
val_cst_type;
no_val;
- val_cstrs;
- val_bool;
- val_opt val_int |]
-
+ val_cstrs|]
let subst_rel_declaration sub (id,copt,t as x) =
let copt' = Option.smartmap (subst_mps sub) copt in
@@ -662,13 +702,10 @@ let subst_arity sub = function
(* 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 ->
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 *)
diff --git a/checker/environ.ml b/checker/environ.ml
index d960e2a7c..d79cb118b 100644
--- a/checker/environ.ml
+++ b/checker/environ.ml
@@ -124,10 +124,10 @@ exception NotEvaluableConst of const_evaluation_result
let constant_value env kn =
let cb = lookup_constant kn env in
- if cb.const_opaque then raise (NotEvaluableConst Opaque);
match cb.const_body with
- | Some l_body -> force_constr l_body
- | None -> raise (NotEvaluableConst NoBody)
+ | Def l_body -> force_constr l_body
+ | OpaqueDef _ -> raise (NotEvaluableConst Opaque)
+ | Undef _ -> raise (NotEvaluableConst NoBody)
(* A global const is evaluable if it is defined and not opaque *)
let evaluable_constant cst env =
diff --git a/checker/mod_checking.ml b/checker/mod_checking.ml
index 30149331e..e6582e918 100644
--- a/checker/mod_checking.ml
+++ b/checker/mod_checking.ml
@@ -33,7 +33,7 @@ let check_constant_declaration env kn cb =
let ty, cu = refresh_arity ty in
let envty = add_constraints cu env' in
let _ = infer_type envty ty in
- (match cb.const_body with
+ (match body_of_constant cb with
| Some bd ->
let j = infer env' (force_constr bd) in
conv_leq envty j ty
@@ -110,14 +110,17 @@ let check_definition_sub env cb1 cb2 =
let typ1 = Typeops.type_of_constant_type env cb1.const_type in
let typ2 = Typeops.type_of_constant_type env cb2.const_type in
check_type env typ1 typ2;
- (match cb2 with
- | {const_body=Some lc2;const_opaque=false} ->
- let c2 = force_constr lc2 in
- let c1 = match cb1.const_body with
- | Some lc1 -> force_constr lc1
- | None -> assert false in
- Reduction.conv env c1 c2
- | _ -> ())
+ (match cb2.const_body with
+ | Undef _ -> ()
+ | Def lc2 ->
+ (match cb1.const_body with
+ | Def lc1 ->
+ let c1 = force_constr lc1 in
+ let c2 = force_constr lc2 in
+ Reduction.conv env c1 c2
+ (* Coq only places transparent cb in With_definition_body *)
+ | _ -> assert false)
+ | _ -> ()) (* Pierre L: shouldn't this case raise an error ? *)
let lookup_modtype mp env =
try Environ.lookup_modtype mp env
diff --git a/checker/modops.ml b/checker/modops.ml
index b18b58a9f..4e2bbc210 100644
--- a/checker/modops.ml
+++ b/checker/modops.ml
@@ -95,19 +95,14 @@ and add_module mb env =
let strengthen_const env mp_from l cb resolver =
- match cb.const_opaque, cb.const_body with
- | false, Some _ -> cb
- | true, Some _
- | _, None ->
+ match cb.const_body with
+ | Def _ -> cb
+ | _ ->
let con = make_con mp_from empty_dirpath l in
- (* let con = constant_of_delta resolver con in*)
- let const = Const con in
- let const_subs = Some (Declarations.from_val const) in
- {cb with
- const_body = const_subs;
- const_opaque = false;
- }
-
+ (* let con = constant_of_delta resolver con in*)
+ let const = Const con in
+ let def = Def (Declarations.from_val const) in
+ { cb with const_body = def }
let rec strengthen_mod env mp_from mp_to mb =
if Declarations.mp_in_delta mb.mod_mp mb.mod_delta then
diff --git a/checker/safe_typing.ml b/checker/safe_typing.ml
index eaf2aae80..bc067dc5f 100644
--- a/checker/safe_typing.ml
+++ b/checker/safe_typing.ml
@@ -77,15 +77,18 @@ type compiled_library =
module LightenLibrary : sig
type table
type lightened_compiled_library
- val load : load_proof:bool -> (unit -> table)
- -> lightened_compiled_library -> compiled_library
+ val load : table -> lightened_compiled_library -> compiled_library
end = struct
(* The table is implemented as an array of [constr_substituted].
- Thus, its keys are integers which can be easily embedded inside
- [constr_substituted]. This way the [compiled_library] type does
- not have to be changed. *)
+ Keys are hence integers. To avoid changing the [compiled_library]
+ type, we brutally encode integers into [lazy_constr]. This isn't
+ pretty, but shouldn't be dangerous since the produced structure
+ [lightened_compiled_library] is abstract and only meant for writing
+ to .vo via Marshal (which doesn't care about types).
+ *)
type table = constr_substituted array
+ let key_of_lazy_constr (c:lazy_constr) = (Obj.magic c : int)
(* To avoid any future misuse of the lightened library that could
interpret encoded keys as real [constr_substituted], we hide
@@ -115,8 +118,8 @@ end = struct
}
and traverse_struct struc =
let traverse_body (l,body) = (l,match body with
- | SFBconst ({const_opaque=true} as x) ->
- SFBconst {x with const_body = on_opaque_const_body x.const_body }
+ | (SFBconst cb) when is_opaque cb ->
+ SFBconst {cb with const_body = on_opaque_const_body cb.const_body}
| (SFBconst _ | SFBmind _ ) as x ->
x
| SFBmodule m ->
@@ -148,21 +151,16 @@ end = struct
[constr_substituted]. Otherwise, we set the [const_body] field
to [None].
*)
- let load ~load_proof (get_table : unit -> table) lightened_library =
- let decode_key : constr_substituted option -> constr_substituted option =
- if load_proof then
- let table = get_table () in
- function Some cterm ->
- Some (table.(
- try
- match Declarations.force_constr cterm with
- | Term.Rel key -> key
- | _ -> assert false
- with _ -> assert false
- ))
- | _ -> None
- else
- fun _ -> None
+ let load table lightened_library =
+ let decode_key = function
+ | Undef _ | Def _ -> assert false
+ | OpaqueDef k ->
+ let k = key_of_lazy_constr k in
+ let body =
+ try table.(k)
+ with _ -> error "Error while retrieving an opaque body"
+ in
+ OpaqueDef (lazy_constr_from_val body)
in
traverse_library decode_key lightened_library
diff --git a/checker/safe_typing.mli b/checker/safe_typing.mli
index 986393b29..cd2c06d20 100644
--- a/checker/safe_typing.mli
+++ b/checker/safe_typing.mli
@@ -33,12 +33,8 @@ sig
type table
type lightened_compiled_library
- (** [load lpf get_table lcl] builds a compiled library from a
+ (** [load table lcl] builds a compiled library from a
lightened library [lcl] by remplacing every index by its related
- opaque terms inside the table obtained by [get_table ()].
- If [lpf] is unset then the table is considered empty, which
- implies that [get_table] is not evaluated and every index
- is replaced by [None] inside the compiled library. *)
- val load : load_proof:bool -> (unit -> table)
- -> lightened_compiled_library -> compiled_library
+ opaque terms inside [table]. *)
+ val load : table -> lightened_compiled_library -> compiled_library
end
diff --git a/checker/subtyping.ml b/checker/subtyping.ml
index 670f2783b..e00671b03 100644
--- a/checker/subtyping.ml
+++ b/checker/subtyping.ml
@@ -237,22 +237,26 @@ let check_constant env mp1 l info1 cb2 spec2 subst1 subst2 =
match info1 with
| Constant cb1 ->
assert (cb1.const_hyps=[] && cb2.const_hyps=[]) ;
- (*Start by checking types*)
- let cb1 = subst_const_body subst1 cb1 in
- let cb2 = subst_const_body subst2 cb2 in
- let typ1 = Typeops.type_of_constant_type env cb1.const_type in
- let typ2 = Typeops.type_of_constant_type env cb2.const_type in
- check_type env typ1 typ2;
- let con = make_con mp1 empty_dirpath l in
- (match cb2 with
- | {const_body=Some lc2;const_opaque=false} ->
- let c2 = force_constr lc2 in
- let c1 = match cb1.const_body with
- | Some lc1 -> force_constr lc1
- | None -> Const con
- in
- check_conv conv env c1 c2
- | _ -> ())
+ let cb1 = subst_const_body subst1 cb1 in
+ let cb2 = subst_const_body subst2 cb2 in
+ (*Start by checking types*)
+ let typ1 = Typeops.type_of_constant_type env cb1.const_type in
+ let typ2 = Typeops.type_of_constant_type env cb2.const_type in
+ check_type env typ1 typ2;
+ (* Now we check the bodies *)
+ (match cb2.const_body with
+ | Undef _ -> ()
+ | Def lc2 ->
+ (* Only a transparent cb1 can implement a transparent cb2.
+ NB: cb1 might have been strengthened and appear as transparent.
+ Anyway [check_conv] will handle that afterwards. *)
+ (match cb1.const_body with
+ | Undef _ | OpaqueDef _ -> error ()
+ | Def lc1 ->
+ let c1 = force_constr lc1 in
+ let c2 = force_constr lc2 in
+ check_conv conv env c1 c2)
+ | OpaqueDef lc2 -> ()) (* Pierre L. : this looks really strange ?! *)
| IndType ((kn,i),mind1) ->
ignore (Util.error (
"The kernel does not recognize yet that a parameter can be " ^
@@ -260,7 +264,7 @@ let check_constant env mp1 l info1 cb2 spec2 subst1 subst2 =
"inductive type and give a definition to map the old name to the new " ^
"name."));
assert (mind1.mind_hyps=[] && cb2.const_hyps=[]) ;
- if cb2.const_body <> None then error () ;
+ if constant_has_body cb2 then error () ;
let arity1 = type_of_inductive env (mind1,mind1.mind_packets.(i)) in
let typ2 = Typeops.type_of_constant_type env cb2.const_type in
check_conv conv_leq env arity1 typ2
@@ -271,7 +275,7 @@ let check_constant env mp1 l info1 cb2 spec2 subst1 subst2 =
"constructor and give a definition to map the old name to the new " ^
"name."));
assert (mind1.mind_hyps=[] && cb2.const_hyps=[]) ;
- if cb2.const_body <> None then error () ;
+ if constant_has_body cb2 then error () ;
let ty1 = type_of_constructor cstr (mind1,mind1.mind_packets.(i)) in
let ty2 = Typeops.type_of_constant_type env cb2.const_type in
check_conv conv env ty1 ty2