aboutsummaryrefslogtreecommitdiffhomepage
path: root/library
diff options
context:
space:
mode:
Diffstat (limited to 'library')
-rw-r--r--library/decls.ml10
-rw-r--r--library/goptions.ml4
-rw-r--r--library/impargs.ml22
-rw-r--r--library/impargs.mli12
-rw-r--r--library/lib.ml27
-rw-r--r--library/lib.mli3
-rw-r--r--library/library.ml18
-rw-r--r--library/summary.ml25
-rw-r--r--library/summary.mli11
9 files changed, 79 insertions, 53 deletions
diff --git a/library/decls.ml b/library/decls.ml
index 6e21880f1..2952c258a 100644
--- a/library/decls.ml
+++ b/library/decls.ml
@@ -14,6 +14,8 @@ open Names
open Decl_kinds
open Libnames
+module NamedDecl = Context.Named.Declaration
+
(** Datas associated to section variables and local definitions *)
type variable_data =
@@ -46,20 +48,18 @@ let constant_kind kn = Cmap.find kn !csttab
(** Miscellaneous functions. *)
-open Context.Named.Declaration
-
let initialize_named_context_for_proof () =
let sign = Global.named_context () in
List.fold_right
(fun d signv ->
- let id = get_id d in
- let d = if variable_opacity id then LocalAssum (id, get_type d) else d in
+ let id = NamedDecl.get_id d in
+ let d = if variable_opacity id then NamedDecl.LocalAssum (id, NamedDecl.get_type d) else d in
Environ.push_named_context_val d signv) sign Environ.empty_named_context_val
let last_section_hyps dir =
Context.Named.fold_outside
(fun d sec_ids ->
- let id = get_id d in
+ let id = NamedDecl.get_id d in
try if DirPath.equal dir (variable_path id) then id::sec_ids else sec_ids
with Not_found -> sec_ids)
(Environ.named_context (Global.env()))
diff --git a/library/goptions.ml b/library/goptions.ml
index db8933d28..813bf30a3 100644
--- a/library/goptions.ml
+++ b/library/goptions.ml
@@ -365,8 +365,8 @@ let set_string_option_value = set_string_option_value_gen None
let msg_option_value (name,v) =
match v with
- | BoolValue true -> str "true"
- | BoolValue false -> str "false"
+ | BoolValue true -> str "on"
+ | BoolValue false -> str "off"
| IntValue (Some n) -> int n
| IntValue None -> str "undefined"
| StringValue s -> str s
diff --git a/library/impargs.ml b/library/impargs.ml
index c9d4afc79..ea2805b67 100644
--- a/library/impargs.ml
+++ b/library/impargs.ml
@@ -22,6 +22,9 @@ open Constrexpr
open Termops
open Namegen
open Decl_kinds
+open Context.Named.Declaration
+
+module NamedDecl = Context.Named.Declaration
(*s Flags governing the computation of implicit arguments *)
@@ -68,10 +71,9 @@ let is_reversible_pattern_implicit_args () = !implicit_args.reversible_pattern
let is_contextual_implicit_args () = !implicit_args.contextual
let is_maximal_implicit_args () = !implicit_args.maximal
-let with_implicits flags f x =
+let with_implicit_protection f x =
let oflags = !implicit_args in
try
- implicit_args := flags;
let rslt = f x in
implicit_args := oflags;
rslt
@@ -165,7 +167,6 @@ let update pos rig (na,st) =
(* modified is_rigid_reference with a truncated env *)
let is_flexible_reference env bound depth f =
- let open Context.Named.Declaration in
match kind_of_term f with
| Rel n when n >= bound+depth -> (* inductive type *) false
| Rel n when n >= depth -> (* previous argument *) true
@@ -174,7 +175,7 @@ let is_flexible_reference env bound depth f =
let cb = Environ.lookup_constant kn env in
(match cb.const_body with Def _ -> true | _ -> false)
| Var id ->
- Environ.lookup_named id env |> is_local_def
+ env |> Environ.lookup_named id |> is_local_def
| Ind _ | Construct _ -> false
| _ -> true
@@ -450,8 +451,7 @@ let compute_all_mib_implicits flags manual kn =
let compute_var_implicits flags manual id =
let env = Global.env () in
- let open Context.Named.Declaration in
- compute_semi_auto_implicits env flags manual (get_type (lookup_named id env))
+ compute_semi_auto_implicits env flags manual (NamedDecl.get_type (lookup_named id env))
(* Implicits of a global reference. *)
@@ -516,15 +516,11 @@ let subst_implicits (subst,(req,l)) =
(ImplLocal,List.smartmap (subst_implicits_decl subst) l)
let impls_of_context ctx =
- let map (id, impl, _, _) = match impl with
- | Implicit -> Some (id, Manual, (true, true))
+ let map (decl, impl) = match impl with
+ | Implicit -> Some (NamedDecl.get_id decl, Manual, (true, true))
| _ -> None
in
- let is_set (_, _, b, _) = match b with
- | None -> true
- | Some _ -> false
- in
- List.rev_map map (List.filter is_set ctx)
+ List.rev_map map (List.filter (fst %> is_local_assum) ctx)
let adjust_side_condition p = function
| LessArgsThan n -> LessArgsThan (n+p)
diff --git a/library/impargs.mli b/library/impargs.mli
index 34e529ca2..3919a519c 100644
--- a/library/impargs.mli
+++ b/library/impargs.mli
@@ -29,8 +29,7 @@ val is_reversible_pattern_implicit_args : unit -> bool
val is_contextual_implicit_args : unit -> bool
val is_maximal_implicit_args : unit -> bool
-type implicits_flags
-val with_implicits : implicits_flags -> ('a -> 'b) -> 'a -> 'b
+val with_implicit_protection : ('a -> 'b) -> 'a -> 'b
(** {6 ... } *)
(** An [implicits_list] is a list of positions telling which arguments
@@ -136,14 +135,5 @@ val select_impargs_size : int -> implicits_list list -> implicit_status list
val select_stronger_impargs : implicits_list list -> implicit_status list
-type implicit_interactive_request
-
-type implicit_discharge_request =
- | ImplLocal
- | ImplConstant of constant * implicits_flags
- | ImplMutualInductive of mutual_inductive * implicits_flags
- | ImplInteractive of global_reference * implicits_flags *
- implicit_interactive_request
-
val explicitation_eq : Constrexpr.explicitation -> Constrexpr.explicitation -> bool
(** Equality on [explicitation]. *)
diff --git a/library/lib.ml b/library/lib.ml
index 376643ebb..749cc4ff3 100644
--- a/library/lib.ml
+++ b/library/lib.ml
@@ -13,6 +13,9 @@ open Libnames
open Globnames
open Nameops
open Libobject
+open Context.Named.Declaration
+
+module NamedDecl = Context.Named.Declaration
type is_type = bool (* Module Type or just Module *)
type export = bool option (* None for a Module Type *)
@@ -388,7 +391,7 @@ let find_opening_node id =
- the list of substitution to do at section closing
*)
-type variable_info = Names.Id.t * Decl_kinds.binding_kind * Term.constr option * Term.types
+type variable_info = Context.Named.Declaration.t * Decl_kinds.binding_kind
type variable_context = variable_info list
type abstr_info = variable_context * Univ.universe_level_subst * Univ.UContext.t
@@ -428,12 +431,10 @@ let add_section_context ctx =
sectab := (Context ctx :: vars,repl,abs)::sl
let extract_hyps (secs,ohyps) =
- let open Context.Named.Declaration in
let rec aux = function
- | (Variable (id,impl,poly,ctx)::idl, decl::hyps) when Names.Id.equal id (get_id decl) ->
- let (id',b,t) = to_tuple decl in
+ | (Variable (id,impl,poly,ctx)::idl, decl::hyps) when Names.Id.equal id (NamedDecl.get_id decl) ->
let l, r = aux (idl,hyps) in
- (id',impl,b,t) :: l, if poly then Univ.ContextSet.union r ctx else r
+ (decl,impl) :: l, if poly then Univ.ContextSet.union r ctx else r
| (Variable (_,_,poly,ctx)::idl,hyps) ->
let l, r = aux (idl,hyps) in
l, if poly then Univ.ContextSet.union r ctx else r
@@ -443,17 +444,11 @@ let extract_hyps (secs,ohyps) =
| [], _ -> [],Univ.ContextSet.empty
in aux (secs,ohyps)
-let instance_from_variable_context sign =
- let rec inst_rec = function
- | (id,b,None,_) :: sign -> id :: inst_rec sign
- | _ :: sign -> inst_rec sign
- | [] -> [] in
- Array.of_list (inst_rec sign)
-
-let named_of_variable_context ctx = let open Context.Named.Declaration in
- List.map (function id,_,None,t -> LocalAssum (id,t)
- | id,_,Some b,t -> LocalDef (id,b,t))
- ctx
+let instance_from_variable_context =
+ List.map fst %> List.filter is_local_assum %> List.map NamedDecl.get_id %> Array.of_list
+
+let named_of_variable_context =
+ List.map fst
let add_section_replacement f g poly hyps =
match !sectab with
diff --git a/library/lib.mli b/library/lib.mli
index 7080b5dba..092643c2d 100644
--- a/library/lib.mli
+++ b/library/lib.mli
@@ -162,8 +162,7 @@ val xml_open_section : (Names.Id.t -> unit) Hook.t
val xml_close_section : (Names.Id.t -> unit) Hook.t
(** {6 Section management for discharge } *)
-type variable_info = Names.Id.t * Decl_kinds.binding_kind *
- Term.constr option * Term.types
+type variable_info = Context.Named.Declaration.t * Decl_kinds.binding_kind
type variable_context = variable_info list
type abstr_info = variable_context * Univ.universe_level_subst * Univ.UContext.t
diff --git a/library/library.ml b/library/library.ml
index c3d0485d3..3086e3d18 100644
--- a/library/library.ml
+++ b/library/library.ml
@@ -452,13 +452,13 @@ let intern_from_file f =
module DPMap = Map.Make(DirPath)
let rec intern_library (needed, contents) (dir, f) from =
- Feedback.feedback(Feedback.FileDependency (from, DirPath.to_string dir));
(* Look if in the current logical environment *)
try (find_library dir).libsum_digests, (needed, contents)
with Not_found ->
(* Look if already listed and consequently its dependencies too *)
try (DPMap.find dir contents).library_digests, (needed, contents)
with Not_found ->
+ Feedback.feedback(Feedback.FileDependency (from, DirPath.to_string dir));
(* [dir] is an absolute name which matches [f] which must be in loadpath *)
let f = match f with Some f -> f | None -> try_locate_absolute_library dir in
let m = intern_from_file f in
@@ -468,16 +468,18 @@ let rec intern_library (needed, contents) (dir, f) from =
pr_dirpath m.library_name ++ spc () ++ str "and not library" ++
spc() ++ pr_dirpath dir);
Feedback.feedback (Feedback.FileLoaded(DirPath.to_string dir, f));
- m.library_digests, intern_library_deps (needed, contents) dir m (Some f)
+ m.library_digests, intern_library_deps (needed, contents) dir m f
and intern_library_deps libs dir m from =
let needed, contents = Array.fold_left (intern_mandatory_library dir from) libs m.library_deps in
(dir :: needed, DPMap.add dir m contents )
and intern_mandatory_library caller from libs (dir,d) =
- let digest, libs = intern_library libs (dir, None) from in
+ let digest, libs = intern_library libs (dir, None) (Some from) in
if not (Safe_typing.digest_match ~actual:digest ~required:d) then
- user_err (str "Compiled library " ++ pr_dirpath caller ++ str ".vo makes inconsistent assumptions over library " ++ pr_dirpath dir);
+ user_err (str "Compiled library " ++ pr_dirpath caller ++
+ str " (in file " ++ str from ++ str ") makes inconsistent assumptions \
+ over library " ++ pr_dirpath dir);
libs
let rec_intern_library libs (dir, f) =
@@ -551,12 +553,20 @@ let in_require : require_obj -> obj =
let (f_xml_require, xml_require) = Hook.make ~default:ignore ()
+let warn_require_in_module =
+ CWarnings.create ~name:"require-in-module" ~category:"deprecated"
+ (fun () -> strbrk "Require inside a module is" ++
+ strbrk " deprecated and strongly discouraged. " ++
+ strbrk "You can Require a module at toplevel " ++
+ strbrk "and optionally Import it inside another one.")
+
let require_library_from_dirpath modrefl export =
let needed, contents = List.fold_left rec_intern_library ([], DPMap.empty) modrefl in
let needed = List.rev_map (fun dir -> DPMap.find dir contents) needed in
let modrefl = List.map fst modrefl in
if Lib.is_module_or_modtype () then
begin
+ warn_require_in_module ();
add_anonymous_leaf (in_require (needed,modrefl,None));
Option.iter (fun exp ->
add_anonymous_leaf (in_import_library (modrefl,exp)))
diff --git a/library/summary.ml b/library/summary.ml
index 4fef06438..6efa07f38 100644
--- a/library/summary.ml
+++ b/library/summary.ml
@@ -186,4 +186,29 @@ let ref ?(freeze=fun _ r -> r) ~name x =
init_function = (fun () -> r := x) };
r
+module Local = struct
+
+type 'a local_ref = ('a CEphemeron.key * string) ref
+
+let (:=) r v = r := (CEphemeron.create v, snd !r)
+
+let (!) r =
+ let key, name = !r in
+ try CEphemeron.get key
+ with CEphemeron.InvalidKey ->
+ let _, { init_function } =
+ Int.Map.find (String.hash (mangle name)) !summaries in
+ init_function ();
+ CEphemeron.get (fst !r)
+
+let ref ?(freeze=fun x -> x) ~name init =
+ let r = Pervasives.ref (CEphemeron.create init, name) in
+ declare_summary name
+ { freeze_function = (fun _ -> freeze !r);
+ unfreeze_function = ((:=) r);
+ init_function = (fun () -> r := init) };
+ r
+
+end
+
let dump = Dyn.dump
diff --git a/library/summary.mli b/library/summary.mli
index 27889cab2..1b57613cb 100644
--- a/library/summary.mli
+++ b/library/summary.mli
@@ -42,6 +42,17 @@ val declare_summary : string -> 'a summary_declaration -> unit
val ref : ?freeze:(marshallable -> 'a -> 'a) -> name:string -> 'a -> 'a ref
+(* As [ref] but the value is local to a process, i.e. not sent to, say, proof
+ * workers. It is useful to implement a local cache for example. *)
+module Local : sig
+
+ type 'a local_ref
+ val ref : ?freeze:('a -> 'a) -> name:string -> 'a -> 'a local_ref
+ val (:=) : 'a local_ref -> 'a -> unit
+ val (!) : 'a local_ref -> 'a
+
+end
+
(** Special name for the summary of ML modules. This summary entry is
special because its unfreeze may load ML code and hence add summary
entries. Thus is has to be recognizable, and handled appropriately *)