aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--config/coq_config.mli14
-rw-r--r--configure.ml122
-rw-r--r--ide/ideutils.ml2
-rw-r--r--ide/minilib.ml4
-rw-r--r--ide/preferences.ml4
-rw-r--r--lib/envars.ml56
-rw-r--r--lib/envars.mli13
-rw-r--r--tools/coq_makefile.ml8
-rw-r--r--tools/coqdoc/cdglobals.ml24
9 files changed, 135 insertions, 112 deletions
diff --git a/config/coq_config.mli b/config/coq_config.mli
index c171bd355..2b3bc2c25 100644
--- a/config/coq_config.mli
+++ b/config/coq_config.mli
@@ -8,11 +8,19 @@
val local : bool (* local use (no installation) *)
-val coqlib : string option (* where the std library is installed *)
-val configdir : string option (* where configuration files are installed *)
-val datadir : string option (* where extra data files are installed *)
+(* The fields below are absolute paths *)
+val coqlib : string (* where the std library is installed *)
+val configdir : string (* where configuration files are installed *)
+val datadir : string (* where extra data files are installed *)
val docdir : string (* where the doc is installed *)
+(* The fields below are paths relative to the installation prefix *)
+(* However, if an absolute path, it means discarding the actual prefix *)
+val coqlibsuffix : string (* std library relative to installation prefix *)
+val configdirsuffix : string (* config files relative to installation prefix *)
+val datadirsuffix : string (* data files relative to installation prefix *)
+val docdirsuffix : string (* doc directory relative to installation prefix *)
+
val ocaml : string (* names of ocaml binaries *)
val ocamlfind : string
val ocamllex : string
diff --git a/configure.ml b/configure.ml
index a2aabb49c..a5204d5b5 100644
--- a/configure.ml
+++ b/configure.ml
@@ -425,11 +425,11 @@ let arch = match !Prefs.arch with
else if arch <> "" then arch
else try_archs arch_progs
-(** NB: [arch_win32] is broader than [os_type_win32], cf. cygwin *)
+(** NB: [arch_is_win32] is broader than [os_type_win32], cf. cygwin *)
-let arch_win32 = (arch = "win32")
+let arch_is_win32 = (arch = "win32")
-let exe = exe := if arch_win32 then ".exe" else ""; !exe
+let exe = exe := if arch_is_win32 then ".exe" else ""; !exe
let dll = if os_type_win32 then ".dll" else ".so"
(** * VCS
@@ -449,7 +449,7 @@ let vcs =
let browser =
match !Prefs.browser with
| Some b -> b
- | None when arch_win32 -> "start %s"
+ | None when arch_is_win32 -> "start %s"
| None when arch = "Darwin" -> "open %s"
| _ -> "firefox -remote \"OpenURL(%s,new-tab)\" || firefox %s &"
@@ -855,73 +855,91 @@ let withdoc = check_doc ()
let coqtop = Sys.getcwd ()
-let unix = os_type_cygwin || not arch_win32
+let unix = os_type_cygwin || not arch_is_win32
(** Variable name, description, ref in Prefs, default dir, prefix-relative *)
+type path_style =
+ | Absolute of string (* Should start with a "/" *)
+ | Relative of string (* Should not start with a "/" *)
+
let install = [
"BINDIR", "the Coq binaries", Prefs.bindir,
- (if unix then "/usr/local/bin" else "C:/coq/bin"),
- "/bin";
+ Relative "bin", Relative "bin", Relative "bin";
"COQLIBINSTALL", "the Coq library", Prefs.libdir,
- (if unix then "/usr/local/lib/coq" else "C:/coq/lib"),
- (if arch_win32 then "" else "/lib/coq");
+ Relative "lib", Relative "lib/coq", Relative "";
"CONFIGDIR", "the Coqide configuration files", Prefs.configdir,
- (if unix then "/etc/xdg/coq" else "C:/coq/config"),
- (if arch_win32 then "/config" else "/etc/xdg/coq");
+ Relative "config", Absolute "/etc/xdg/coq", Relative "ide";
"DATADIR", "the Coqide data files", Prefs.datadir,
- (if unix then "/usr/local/share/coq" else "C:/coq/share"),
- "/share/coq";
+ Relative "share", Relative "share/coq", Relative "ide";
"MANDIR", "the Coq man pages", Prefs.mandir,
- (if unix then "/usr/local/share/man" else "C:/coq/man"),
- "/share/man";
+ Relative "man", Relative "share/man", Relative "man";
"DOCDIR", "the Coq documentation", Prefs.docdir,
- (if unix then "/usr/local/share/doc/coq" else "C:/coq/doc"),
- "/share/doc/coq";
+ Relative "doc", Relative "share/doc/coq", Relative "doc";
"EMACSLIB", "the Coq Emacs mode", Prefs.emacslib,
- (if unix then "/usr/local/share/emacs/site-lisp" else "C:/coq/emacs"),
- (if arch_win32 then "/emacs" else "/share/emacs/site-lisp");
+ Relative "emacs", Relative "share/emacs/site-lisp", Relative "tools";
"COQDOCDIR", "the Coqdoc LaTeX files", Prefs.coqdocdir,
- (if unix then "/usr/local/share/texmf/tex/latex/misc" else "C:/coq/latex"),
- (if arch_win32 then "/latex" else "/share/emacs/site-lisp");
+ Relative "latex", Relative "share/texmf/tex/latex/misc", Relative "tools/coqdoc";
]
-let do_one_instdir (var,msg,r,dflt,suff) =
- let dir = match !r, !Prefs.prefix with
- | Some d, _ -> d
- | _, Some p -> p^suff
- | _ ->
+let strip_trailing_slash_if_any p =
+ if p.[String.length p - 1] = '/' then String.sub p 0 (String.length p - 1) else p
+
+let use_suffix prefix = function
+ | Relative "" -> prefix
+ | Relative suff -> prefix ^ "/" ^ suff
+ | Absolute path -> path
+
+let relativize = function
+ (* Turn a global layout based on some prefix to a relative layout *)
+ | Relative _ as suffix -> suffix
+ | Absolute path -> Relative (String.sub path 1 (String.length path - 1))
+
+let find_suffix prefix path = match prefix with
+ | None -> Absolute path
+ | Some p ->
+ let p = strip_trailing_slash_if_any p in
+ let lpath = String.length path in
+ let lp = String.length p in
+ if lpath > lp && String.sub path 0 lp = p then
+ Relative (String.sub path (lp+1) (lpath - lp - 1))
+ else
+ Absolute path
+
+let do_one_instdir (var,msg,uservalue,selfcontainedlayout,unixlayout,locallayout) =
+ let dir,suffix =
+ if !Prefs.local then (use_suffix coqtop locallayout,locallayout)
+ else match !uservalue, !Prefs.prefix with
+ | Some d, p -> d,find_suffix p d
+ | _, Some p ->
+ let suffix = if arch_is_win32 then selfcontainedlayout else relativize unixlayout in
+ use_suffix p suffix, suffix
+ | _, p ->
+ let suffix = if unix then unixlayout else selfcontainedlayout in
+ let base = if unix then "/usr/local" else "C:/coq" in
+ let dflt = use_suffix base suffix in
let () = printf "Where should I install %s [%s]? " msg dflt in
let line = read_line () in
- if line = "" then dflt else line
- in (var,msg,dir,dir<>dflt)
-
-let do_one_noinst (var,msg,_,_,_) =
- if var="CONFIGDIR" || var="DATADIR" then (var,msg,coqtop^"/ide",true)
- else (var,msg,coqtop^"/doc",false)
+ if line = "" then (dflt,suffix) else (line,find_suffix p line)
+ in (var,msg,dir,suffix)
-let install_dirs =
- let f = if !Prefs.local then do_one_noinst else do_one_instdir in
- List.map f install
+let install_dirs = List.map do_one_instdir install
let select var = List.find (fun (v,_,_,_) -> v=var) install_dirs
-let libdir = let (_,_,d,_) = select "COQLIBINSTALL" in d
-
-let docdir = let (_,_,d,_) = select "DOCDIR" in d
+let coqlib,coqlibsuffix = let (_,_,d,s) = select "COQLIBINSTALL" in d,s
-let configdir =
- let (_,_,d,b) = select "CONFIGDIR" in if b then Some d else None
+let docdir,docdirsuffix = let (_,_,d,s) = select "DOCDIR" in d,s
-let datadir =
- let (_,_,d,b) = select "DATADIR" in if b then Some d else None
+let configdir,configdirsuffix = let (_,_,d,s) = select "CONFIGDIR" in d,s
+let datadir,datadirsuffix = let (_,_,d,s) = select "DATADIR" in d,s
(** * OCaml runtime flags *)
(** Do we use -custom (yes by default on Windows and MacOS) *)
-let custom_os = arch_win32 || arch = "Darwin"
+let custom_os = arch_is_win32 || arch = "Darwin"
let use_custom = match !Prefs.custom with
| Some b -> b
@@ -941,7 +959,7 @@ let config_runtime () =
| _ ->
let ld="CAML_LD_LIBRARY_PATH" in
build_loadpath := sprintf "export %s:='%s/kernel/byterun':$(%s)" ld coqtop ld;
- ["-dllib";"-lcoqrun";"-dllpath";libdir/"kernel/byterun"]
+ ["-dllib";"-lcoqrun";"-dllpath";coqlib/"kernel/byterun"]
let vmbyteflags = config_runtime ()
@@ -1019,18 +1037,22 @@ let write_configml f =
let pr_s = pr "let %s = %S\n" in
let pr_b = pr "let %s = %B\n" in
let pr_i = pr "let %s = %d\n" in
- let pr_o s o = pr "let %s = %s\n" s
- (match o with None -> "None" | Some d -> sprintf "Some %S" d)
+ let pr_p s o = pr "let %s = %S\n" s
+ (match o with Relative s -> s | Absolute s -> s)
in
pr "(* DO NOT EDIT THIS FILE: automatically generated by ../configure *)\n";
pr "(* Exact command that generated this file: *)\n";
pr "(* %s *)\n\n" (String.concat " " (Array.to_list Sys.argv));
pr_b "local" !Prefs.local;
pr "let vmbyteflags = ["; List.iter (pr "%S;") vmbyteflags; pr "]\n";
- pr_o "coqlib" (if !Prefs.local then None else Some libdir);
- pr_o "configdir" configdir;
- pr_o "datadir" datadir;
+ pr_s "coqlib" coqlib;
+ pr_s "configdir" configdir;
+ pr_s "datadir" datadir;
pr_s "docdir" docdir;
+ pr_p "coqlibsuffix" coqlibsuffix;
+ pr_p "configdirsuffix" configdirsuffix;
+ pr_p "datadirsuffix" datadirsuffix;
+ pr_p "docdirsuffix" docdirsuffix;
pr_s "ocaml" camlexec.top;
pr_s "ocamlfind" camlexec.find;
pr_s "ocamllex" camlexec.lex;
@@ -1049,7 +1071,7 @@ let write_configml f =
pr_s "date" short_date;
pr_s "compile_date" full_date;
pr_s "arch" arch;
- pr_b "arch_is_win32" arch_win32;
+ pr_b "arch_is_win32" arch_is_win32;
pr_s "exec_extension" exe;
pr_s "coqideincl" !lablgtkincludes;
pr_s "has_coqide" coqide;
diff --git a/ide/ideutils.ml b/ide/ideutils.ml
index a08ab07b5..8a0e507c0 100644
--- a/ide/ideutils.ml
+++ b/ide/ideutils.ml
@@ -422,7 +422,7 @@ let browse prerr url =
let doc_url () =
if doc_url#get = use_default_doc_url || doc_url#get = ""
then
- let addr = List.fold_left Filename.concat (Coq_config.docdir)
+ let addr = List.fold_left Filename.concat (Envars.docdir ())
["html";"refman";"index.html"]
in
if Sys.file_exists addr then "file://"^addr else Coq_config.wwwrefman
diff --git a/ide/minilib.ml b/ide/minilib.ml
index 2c24e46f8..2b278fac6 100644
--- a/ide/minilib.ml
+++ b/ide/minilib.ml
@@ -54,12 +54,12 @@ let coqide_config_home () =
let coqide_data_dirs () =
coqify (Glib.get_user_data_dir ())
:: List.map coqify (Glib.get_system_data_dirs ())
- @ Option.List.cons Coq_config.datadir []
+ @ [Envars.datadir ()]
let coqide_config_dirs () =
coqide_config_home ()
:: List.map coqify (Glib.get_system_config_dirs ())
- @ Option.List.cons Coq_config.configdir []
+ @ [Envars.configdir ()]
let is_prefix_of pre s =
let i = ref 0 in
diff --git a/ide/preferences.ml b/ide/preferences.ml
index 44a89edf9..08739d013 100644
--- a/ide/preferences.ml
+++ b/ide/preferences.ml
@@ -922,7 +922,7 @@ let configure ?(apply=(fun () -> ())) () =
in
let doc_url =
let predefined = [
- "file://"^(List.fold_left Filename.concat (Coq_config.docdir) ["refman";"html"]);
+ "file://"^(List.fold_left Filename.concat (Envars.docdir ()) ["refman";"html"]);
Coq_config.wwwrefman;
use_default_doc_url
] in
@@ -935,7 +935,7 @@ let configure ?(apply=(fun () -> ())) () =
doc_url#get in
let library_url =
let predefined = [
- "file://"^(List.fold_left Filename.concat (Coq_config.docdir) ["stdlib";"html"]);
+ "file://"^(List.fold_left Filename.concat (Envars.docdir ()) ["stdlib";"html"]);
Coq_config.wwwstdlib
] in
combo
diff --git a/lib/envars.ml b/lib/envars.ml
index 79516bb1b..bc8012297 100644
--- a/lib/envars.ml
+++ b/lib/envars.ml
@@ -23,8 +23,6 @@ let ( / ) a b =
let coqify d = d / "coq"
-let opt2list = function None -> [] | Some x -> [x]
-
let home ~warn =
getenv_else "HOME" (fun () ->
try (Sys.getenv "HOMEDRIVE")^(Sys.getenv "HOMEPATH") with Not_found ->
@@ -81,9 +79,6 @@ let expand_path_macros ~warn s =
(** {2 Coq paths} *)
-let relative_base =
- Filename.dirname (Filename.dirname Sys.executable_name)
-
let coqbin =
CUnix.canonical_path_name (Filename.dirname Sys.executable_name)
@@ -98,25 +93,26 @@ let _ =
if Coq_config.arch_is_win32 then
Unix.putenv "PATH" (coqbin ^ ";" ^ getenv_else "PATH" (fun () -> ""))
+(** Add a local installation suffix (unless the suffix is itself
+ absolute in which case the prefix does not matter) *)
+let use_suffix prefix suffix =
+ if String.length suffix > 0 && suffix.[0] = '/' then suffix else prefix / suffix
+
(** [check_file_else ~dir ~file oth] checks if [file] exists in
- the installation directory [dir] given relatively to [coqroot].
- If this Coq is only locally built, then [file] must be in [coqroot].
+ the installation directory [dir] given relatively to [coqroot],
+ which maybe has been relocated.
If the check fails, then [oth ()] is evaluated.
Using file system equality seems well enough for this heuristic *)
let check_file_else ~dir ~file oth =
- let path = if Coq_config.local then coqroot else coqroot / dir in
+ let path = use_suffix coqroot dir in
if Sys.file_exists (path / file) then path else oth ()
let guess_coqlib fail =
let prelude = "theories/Init/Prelude.vo" in
- let dir = if Coq_config.arch_is_win32 then "lib" else "lib/coq" in
- check_file_else ~dir ~file:prelude
+ check_file_else ~dir:Coq_config.coqlibsuffix ~file:prelude
(fun () ->
- let coqlib = match Coq_config.coqlib with
- | Some coqlib -> coqlib
- | None -> coqroot
- in
- if Sys.file_exists (coqlib / prelude) then coqlib
+ if not Coq_config.local && Sys.file_exists (Coq_config.coqlib / prelude)
+ then Coq_config.coqlib
else
fail "cannot guess a path for Coq libraries; please use -coqlib option")
@@ -130,8 +126,19 @@ let set_coqlib ~fail =
let coqlib () = !Flags.coqlib
let docdir () =
- let dir = if Coq_config.arch_is_win32 then "doc" else "share/doc/coq" in
- check_file_else ~dir ~file:"html" (fun () -> Coq_config.docdir)
+ (* This assumes implicitly that the suffix is non-trivial *)
+ let path = use_suffix coqroot Coq_config.docdirsuffix in
+ if Sys.file_exists path then path else Coq_config.docdir
+
+let datadir () =
+ (* This assumes implicitly that the suffix is non-trivial *)
+ let path = use_suffix coqroot Coq_config.datadirsuffix in
+ if Sys.file_exists path then path else Coq_config.datadir
+
+let configdir () =
+ (* This assumes implicitly that the suffix is non-trivial *)
+ let path = use_suffix coqroot Coq_config.configdirsuffix in
+ if Sys.file_exists path then path else Coq_config.configdir
let coqpath =
let coqpath = getenv_else "COQPATH" (fun () -> "") in
@@ -186,20 +193,9 @@ let xdg_data_dirs warn =
try
List.map coqify (path_to_list (Sys.getenv "XDG_DATA_DIRS"))
with
- | Not_found when String.equal Sys.os_type "Win32" -> [relative_base / "share"]
- | Not_found -> ["/usr/local/share/coq";"/usr/share/coq"]
- in
- xdg_data_home warn :: sys_dirs @ opt2list Coq_config.datadir
-
-let xdg_config_dirs warn =
- let sys_dirs =
- try
- List.map coqify (path_to_list (Sys.getenv "XDG_CONFIG_DIRS"))
- with
- | Not_found when String.equal Sys.os_type "Win32" -> [relative_base / "config"]
- | Not_found -> ["/etc/xdg/coq"]
+ | Not_found -> [datadir ()]
in
- xdg_config_home warn :: sys_dirs @ opt2list Coq_config.configdir
+ xdg_data_home warn :: sys_dirs
let xdg_dirs ~warn =
List.filter Sys.file_exists (xdg_data_dirs warn)
diff --git a/lib/envars.mli b/lib/envars.mli
index b164e789d..c8bbf17d9 100644
--- a/lib/envars.mli
+++ b/lib/envars.mli
@@ -27,12 +27,18 @@ val home : warn:(string -> unit) -> string
(** [coqlib] is the path to the Coq library. *)
val coqlib : unit -> string
+(** [docdir] is the path to the installed documentation. *)
+val docdir : unit -> string
+
+(** [datadir] is the path to the installed data directory. *)
+val datadir : unit -> string
+
+(** [configdir] is the path to the installed config directory. *)
+val configdir : unit -> string
+
(** [set_coqlib] must be runned once before any access to [coqlib] *)
val set_coqlib : fail:(string -> string) -> unit
-(** [docdir] is the path to the Coq documentation. *)
-val docdir : unit -> string
-
(** [coqbin] is the name of the current executable. *)
val coqbin : string
@@ -66,7 +72,6 @@ val camlp4 : unit -> string
*)
val xdg_config_home : (string -> unit) -> string
val xdg_data_home : (string -> unit) -> string
-val xdg_config_dirs : (string -> unit) -> string list
val xdg_data_dirs : (string -> unit) -> string list
val xdg_dirs : warn : (string -> unit) -> string list
diff --git a/tools/coq_makefile.ml b/tools/coq_makefile.ml
index 78c92e68b..8e2f75fc9 100644
--- a/tools/coq_makefile.ml
+++ b/tools/coq_makefile.ml
@@ -119,14 +119,8 @@ let read_whole_file s =
Buffer.contents b
let makefile_template =
- let open Filename in
let template = "/tools/CoqMakefile.in" in
- if Coq_config.local then
- let coqbin = CUnix.canonical_path_name (dirname Sys.executable_name) in
- dirname coqbin ^ template
- else match Coq_config.coqlib with
- | None -> assert false
- | Some dir -> dir ^ template
+ Coq_config.coqlib ^ template
let quote s = if String.contains s ' ' then "\"" ^ s ^ "\"" else s
diff --git a/tools/coqdoc/cdglobals.ml b/tools/coqdoc/cdglobals.ml
index 5d48473d8..ef203960b 100644
--- a/tools/coqdoc/cdglobals.ml
+++ b/tools/coqdoc/cdglobals.ml
@@ -70,23 +70,21 @@ let normalize_filename f =
let dirname = Filename.dirname f in
normalize_path dirname, basename
+(** Add a local installation suffix (unless the suffix is itself
+ absolute in which case the prefix does not matter) *)
+let use_suffix prefix suffix =
+ if String.length suffix > 0 && suffix.[0] = '/' then suffix else prefix / suffix
+
(** A weaker analog of the function in Envars *)
let guess_coqlib () =
let file = "theories/Init/Prelude.vo" in
- match Coq_config.coqlib with
- | Some coqlib when Sys.file_exists (coqlib / file) -> coqlib
- | Some _ | None ->
- let coqbin = normalize_path (Filename.dirname Sys.executable_name) in
- let prefix = Filename.dirname coqbin in
- let rpath =
- if Coq_config.local then []
- else if Coq_config.arch_is_win32 then ["lib"]
- else ["lib/coq"]
- in
- let coqlib = List.fold_left (/) prefix rpath in
- if Sys.file_exists (coqlib / file) then coqlib
- else prefix
+ let coqbin = normalize_path (Filename.dirname Sys.executable_name) in
+ let prefix = Filename.dirname coqbin in
+ let coqlib = use_suffix prefix Coq_config.coqlibsuffix in
+ if Sys.file_exists (coqlib / file) then coqlib else
+ if not Coq_config.local && Sys.file_exists (Coq_config.coqlib / file)
+ then Coq_config.coqlib else prefix
let header_trailer = ref true
let header_file = ref ""