aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--CHANGES3
-rw-r--r--checker/check.ml28
-rw-r--r--checker/check.mli30
-rw-r--r--checker/checker.ml11
-rw-r--r--doc/refman/RefMan-com.tex5
-rw-r--r--man/coqchk.12
-rw-r--r--tools/CoqMakefile.in2
7 files changed, 60 insertions, 21 deletions
diff --git a/CHANGES b/CHANGES
index 6a7a5a42b..4c83b7c19 100644
--- a/CHANGES
+++ b/CHANGES
@@ -25,6 +25,9 @@ Vernacular Commands
- The deprecated Coercion Local, Open Local Scope, Notation Local syntax
was removed. Use Local as a prefix instead.
+Checker
+- The checker now accepts filenames in addition to logical paths.
+
Changes from 8.7+beta2 to 8.7.0
===============================
diff --git a/checker/check.ml b/checker/check.ml
index 180ca1ece..21fdba1fa 100644
--- a/checker/check.ml
+++ b/checker/check.ml
@@ -22,6 +22,11 @@ let extend_dirpath p id = DirPath.make (id :: DirPath.repr p)
type section_path = {
dirpath : string list ;
basename : string }
+
+type object_file =
+| PhysicalFile of CUnix.physical_path
+| LogicalFile of section_path
+
let dir_of_path p =
DirPath.make (List.map Id.of_string p.dirpath)
let path_of_dirpath dir =
@@ -69,11 +74,6 @@ let libraries_table = ref LibraryMap.empty
let find_library dir =
LibraryMap.find dir !libraries_table
-let try_find_library dir =
- try find_library dir
- with Not_found ->
- user_err Pp.(str ("Unknown library " ^ (DirPath.to_string dir)))
-
let library_full_filename dir = (find_library dir).library_filename
(* If a library is loaded several time, then the first occurrence must
@@ -263,7 +263,17 @@ let try_locate_absolute_library dir =
| LibUnmappedDir -> error_unmapped_dir (path_of_dirpath dir)
| LibNotFound -> error_lib_not_found (path_of_dirpath dir)
-let try_locate_qualified_library qid =
+let try_locate_qualified_library lib = match lib with
+| PhysicalFile f ->
+ let () =
+ if not (System.file_exists_respecting_case "" f) then
+ error_lib_not_found { dirpath = []; basename = f; }
+ in
+ let dir = Filename.dirname f in
+ let base = Filename.chop_extension (Filename.basename f) in
+ let dir = extend_dirpath (find_logical_path dir) (Id.of_string base) in
+ (dir, f)
+| LogicalFile qid ->
try
locate_qualified_library qid
with
@@ -412,9 +422,3 @@ let recheck_library ~norec ~admit ~check =
(fun (dir,_) -> pr_dirpath dir ++ fnl()) needed));
List.iter (check_one_lib nochk) needed;
Flags.if_verbose Feedback.msg_notice (str"Modules were successfully checked")
-
-open Printf
-
-let mem s =
- let m = try_find_library s in
- h 0 (str (sprintf "%dk" (CObj.size_kb m)))
diff --git a/checker/check.mli b/checker/check.mli
new file mode 100644
index 000000000..28ae385b5
--- /dev/null
+++ b/checker/check.mli
@@ -0,0 +1,30 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open CUnix
+open Names
+
+type section_path = {
+ dirpath : string list;
+ basename : string;
+}
+
+type object_file =
+| PhysicalFile of physical_path
+| LogicalFile of section_path
+
+type logical_path = DirPath.t
+
+val default_root_prefix : DirPath.t
+
+val add_load_path : physical_path * logical_path -> unit
+
+val recheck_library :
+ norec:object_file list ->
+ admit:object_file list ->
+ check:object_file list -> unit
diff --git a/checker/checker.ml b/checker/checker.ml
index e960a55fd..b2433ee36 100644
--- a/checker/checker.ml
+++ b/checker/checker.ml
@@ -40,9 +40,10 @@ let dirpath_of_string s =
[] -> Check.default_root_prefix
| dir -> DirPath.make (List.map Id.of_string dir)
let path_of_string s =
- match parse_dir s with
+ if Filename.check_suffix s ".vo" then PhysicalFile s
+ else match parse_dir s with
[] -> invalid_arg "path_of_string"
- | l::dir -> {dirpath=dir; basename=l}
+ | l::dir -> LogicalFile {dirpath=dir; basename=l}
let ( / ) = Filename.concat
@@ -144,15 +145,15 @@ let set_impredicative_set () = impredicative_set := Cic.ImpredicativeSet
let engage () = Safe_typing.set_engagement (!impredicative_set)
-let admit_list = ref ([] : section_path list)
+let admit_list = ref ([] : object_file list)
let add_admit s =
admit_list := path_of_string s :: !admit_list
-let norec_list = ref ([] : section_path list)
+let norec_list = ref ([] : object_file list)
let add_norec s =
norec_list := path_of_string s :: !norec_list
-let compile_list = ref ([] : section_path list)
+let compile_list = ref ([] : object_file list)
let add_compile s =
compile_list := path_of_string s :: !compile_list
diff --git a/doc/refman/RefMan-com.tex b/doc/refman/RefMan-com.tex
index 8b1fc7c8f..b4d9f60eb 100644
--- a/doc/refman/RefMan-com.tex
+++ b/doc/refman/RefMan-com.tex
@@ -299,8 +299,9 @@ The following command-line options are recognized by the commands {\tt
\section{Compiled libraries checker ({\tt coqchk})}
-The {\tt coqchk} command takes a list of library paths as argument.
-The corresponding compiled libraries (.vo files) are searched in the
+The {\tt coqchk} command takes a list of library paths as argument, described
+either by their logical name or by their physical filename, which must end in
+{\tt .vo}. The corresponding compiled libraries (.vo files) are searched in the
path, recursively processing the libraries they depend on. The content
of all these libraries is then type-checked. The effect of {\tt
coqchk} is only to return with normal exit code in case of success,
diff --git a/man/coqchk.1 b/man/coqchk.1
index 76a7cfc5d..a00914eab 100644
--- a/man/coqchk.1
+++ b/man/coqchk.1
@@ -23,7 +23,7 @@ library was not found, corrupted content, type-checking failure, etc.
.IR modules \&
is a list of modules to be checked. Modules can be referred to by a
-short or qualified name.
+short or qualified logical name, or by their filename.
.SH OPTIONS
diff --git a/tools/CoqMakefile.in b/tools/CoqMakefile.in
index 4ee6efec0..87783350a 100644
--- a/tools/CoqMakefile.in
+++ b/tools/CoqMakefile.in
@@ -392,7 +392,7 @@ checkproofs:
.PHONY: checkproofs
validate: $(VOFILES)
- $(TIMER) $(COQCHK) $(COQCHKFLAGS) $(notdir $(^:.vo=))
+ $(TIMER) $(COQCHK) $(COQCHKFLAGS) $^
.PHONY: validate
only: $(TGTS)