aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--.gitattributes2
-rw-r--r--.gitlab-ci.yml2
-rw-r--r--API/API.ml2
-rw-r--r--API/API.mli55
-rw-r--r--CHANGES2
-rw-r--r--Makefile.build10
-rw-r--r--dev/Bugzilla_Coq_autolink.user.js25
-rw-r--r--dev/ci/ci-basic-overlay.sh4
-rw-r--r--dev/doc/changes.md6
-rw-r--r--dev/doc/xml-protocol.md8
-rw-r--r--doc/common/styles/html/coqremote/cover.html15
-rw-r--r--doc/common/styles/html/simple/cover.html15
-rw-r--r--engine/evd.mli2
-rw-r--r--engine/namegen.ml14
-rw-r--r--intf/vernacexpr.ml2
-rw-r--r--kernel/mod_typing.ml14
-rw-r--r--kernel/safe_typing.ml7
-rw-r--r--kernel/safe_typing.mli4
-rw-r--r--kernel/term_typing.ml7
-rw-r--r--kernel/typeops.ml10
-rw-r--r--kernel/typeops.mli8
-rw-r--r--lib/flags.ml5
-rw-r--r--lib/flags.mli9
-rw-r--r--lib/pp.ml19
-rw-r--r--lib/system.ml8
-rw-r--r--lib/system.mli6
-rw-r--r--lib/unicode.ml97
-rw-r--r--lib/unicode.mli15
-rw-r--r--library/library.ml3
-rw-r--r--library/nametab.ml47
-rw-r--r--library/nametab.mli9
-rw-r--r--man/coqdep.12
-rw-r--r--parsing/cLexer.ml416
-rw-r--r--parsing/g_vernac.ml43
-rw-r--r--plugins/extraction/extract_env.ml24
-rw-r--r--plugins/extraction/extract_env.mli3
-rw-r--r--plugins/extraction/miniml.mli2
-rw-r--r--plugins/extraction/mlutil.ml28
-rw-r--r--plugins/extraction/table.ml5
-rw-r--r--plugins/extraction/table.mli3
-rw-r--r--plugins/ltac/g_ltac.ml45
-rw-r--r--plugins/ltac/ltac_plugin.mlpack4
-rw-r--r--plugins/ltac/pptactic.ml2
-rw-r--r--plugins/ltac/pptactic.mli2
-rw-r--r--plugins/ltac/tacentries.ml35
-rw-r--r--plugins/ltac/tacentries.mli3
-rw-r--r--plugins/ltac/tacenv.ml44
-rw-r--r--plugins/ltac/tacenv.mli10
-rw-r--r--plugins/ltac/tacexpr.mli3
-rw-r--r--plugins/ltac/tacintern.ml6
-rw-r--r--plugins/ssr/ssrcommon.ml4
-rw-r--r--plugins/ssr/ssrparser.ml44
-rw-r--r--pretyping/nativenorm.ml50
-rw-r--r--pretyping/retyping.ml1
-rw-r--r--printing/ppvernac.ml2
-rw-r--r--printing/prettyp.ml76
-rw-r--r--printing/prettyp.mli26
-rw-r--r--proofs/proof_bullet.ml40
-rw-r--r--stm/stm.ml81
-rw-r--r--stm/stm.mli23
-rw-r--r--tactics/tactics.ml10
-rw-r--r--test-suite/bugs/closed/4852.v54
-rw-r--r--test-suite/bugs/closed/5765.v3
-rw-r--r--test-suite/bugs/closed/5769.v20
-rw-r--r--test-suite/output/Notations3.out4
-rw-r--r--test-suite/output/Notations3.v7
-rw-r--r--tools/TimeFileMaker.py4
-rw-r--r--toplevel/coqinit.ml7
-rw-r--r--toplevel/coqinit.mli1
-rw-r--r--toplevel/coqtop.ml194
-rw-r--r--toplevel/coqtop.mli1
-rw-r--r--toplevel/vernac.ml52
-rw-r--r--toplevel/vernac.mli6
-rw-r--r--vernac/command.ml3
-rw-r--r--vernac/metasyntax.ml13
-rw-r--r--vernac/vernacentries.ml2
76 files changed, 863 insertions, 457 deletions
diff --git a/.gitattributes b/.gitattributes
index 6af0a106b..00f78b449 100644
--- a/.gitattributes
+++ b/.gitattributes
@@ -1,5 +1,3 @@
-.dir-locals.el export-ignore
.gitattributes export-ignore
.gitignore export-ignore
.mailmap export-ignore
-TODO export-ignore
diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml
index ae55302d1..d2e335d45 100644
--- a/.gitlab-ci.yml
+++ b/.gitlab-ci.yml
@@ -60,6 +60,7 @@ before_script:
paths:
- _install_ci
- config/Makefile
+ - test-suite/misc/universes/all_stdlib.v
expire_in: 1 week
script:
- set -e
@@ -70,6 +71,7 @@ before_script:
- echo 'start:coq.build'
- make -j ${NJOBS}
+ - make test-suite/misc/universes/all_stdlib.v
- echo 'end:coq:build'
- echo 'start:coq.install'
diff --git a/API/API.ml b/API/API.ml
index 46ad36d36..bf99d0feb 100644
--- a/API/API.ml
+++ b/API/API.ml
@@ -212,7 +212,7 @@ module Pputils = Pputils
module Ppconstr = Ppconstr
module Printer = Printer
(* module Printmod *)
-(* module Prettyp *)
+module Prettyp = Prettyp
module Ppvernac = Ppvernac
(******************************************************************************)
diff --git a/API/API.mli b/API/API.mli
index 3ed326ff0..3bd047043 100644
--- a/API/API.mli
+++ b/API/API.mli
@@ -1795,6 +1795,7 @@ sig
val pr_path : full_path -> Pp.t
val make_path : Names.DirPath.t -> Names.Id.t -> full_path
val eq_full_path : full_path -> full_path -> bool
+ val repr_path : full_path -> Names.DirPath.t * Names.Id.t
val dirpath : full_path -> Names.DirPath.t
val path_of_string : string -> full_path
@@ -1935,24 +1936,19 @@ module Nametab :
sig
exception GlobalizationError of Libnames.qualid
- type ltac_constant = Names.KerName.t
-
val global : Libnames.reference -> Globnames.global_reference
val global_of_path : Libnames.full_path -> Globnames.global_reference
val shortest_qualid_of_global : Names.Id.Set.t -> Globnames.global_reference -> Libnames.qualid
val path_of_global : Globnames.global_reference -> Libnames.full_path
val locate_extended : Libnames.qualid -> Globnames.extended_global_reference
val full_name_module : Libnames.qualid -> Names.DirPath.t
- val locate_tactic : Libnames.qualid -> Names.KerName.t
val pr_global_env : Names.Id.Set.t -> Globnames.global_reference -> Pp.t
- val shortest_qualid_of_tactic : Names.KerName.t -> Libnames.qualid
val basename_of_global : Globnames.global_reference -> Names.Id.t
type visibility =
| Until of int
| Exactly of int
- val push_tactic : visibility -> Libnames.full_path -> Names.KerName.t -> unit
val error_global_not_found : ?loc:Loc.t -> Libnames.qualid -> 'a
val shortest_qualid_of_module : Names.ModPath.t -> Libnames.qualid
val dirpath_of_module : Names.ModPath.t -> Names.DirPath.t
@@ -1960,6 +1956,40 @@ sig
val dirpath_of_global : Globnames.global_reference -> Names.DirPath.t
val locate : Libnames.qualid -> Globnames.global_reference
val locate_constant : Libnames.qualid -> Names.Constant.t
+
+ (** NOT FOR PUBLIC USE YET. Plugin writers, please do not rely on this API. *)
+
+ module type UserName = sig
+ type t
+ val equal : t -> t -> bool
+ val to_string : t -> string
+ val repr : t -> Names.Id.t * Names.Id.t list
+ end
+
+ module type EqualityType =
+ sig
+ type t
+ val equal : t -> t -> bool
+ end
+
+ module type NAMETREE = sig
+ type elt
+ type t
+ type user_name
+
+ val empty : t
+ val push : visibility -> user_name -> elt -> t -> t
+ val locate : Libnames.qualid -> t -> elt
+ val find : user_name -> t -> elt
+ val exists : user_name -> t -> bool
+ val user_name : Libnames.qualid -> t -> user_name
+ val shortest_qualid : Names.Id.Set.t -> user_name -> t -> Libnames.qualid
+ val find_prefixes : Libnames.qualid -> t -> elt list
+ end
+
+ module Make (U : UserName) (E : EqualityType) :
+ NAMETREE with type user_name = U.t and type elt = E.t
+
end
module Global :
@@ -4975,6 +5005,21 @@ sig
val pr_transparent_state : Names.transparent_state -> Pp.t
end
+module Prettyp :
+sig
+ type 'a locatable_info = {
+ locate : Libnames.qualid -> 'a option;
+ locate_all : Libnames.qualid -> 'a list;
+ shortest_qualid : 'a -> Libnames.qualid;
+ name : 'a -> Pp.t;
+ print : 'a -> Pp.t;
+ about : 'a -> Pp.t;
+ }
+
+ val register_locatable : string -> 'a locatable_info -> unit
+ val print_located_other : string -> Libnames.reference -> Pp.t
+end
+
(************************************************************************)
(* End of modules from printing/ *)
(************************************************************************)
diff --git a/CHANGES b/CHANGES
index 08c609cfa..39f119e7c 100644
--- a/CHANGES
+++ b/CHANGES
@@ -5,6 +5,8 @@ Notations
- Recursive notations with the recursive pattern repeating on the
right (e.g. "( x ; .. ; y ; z )") now supported.
+- Notations with a specific level for the leftmost nonterminal,
+ when printing-only, are supported.
Tactics
diff --git a/Makefile.build b/Makefile.build
index ecaaccaaf..7b2e209a2 100644
--- a/Makefile.build
+++ b/Makefile.build
@@ -428,12 +428,18 @@ tools: $(TOOLS) $(OCAMLLIBDEP) $(COQDEPBOOT)
# may still be missing or not taken in account yet by make when coqdep_boot
# is being built.
+# Remember to update the dependencies below when you add files!
+
COQDEPBOOTSRC := lib/minisys.cmo \
lib/segmenttree.cmo lib/unicodetable.cmo lib/unicode.cmo \
tools/coqdep_lexer.cmo tools/coqdep_common.cmo tools/coqdep_boot.cmo
-tools/coqdep_lexer.cmo : tools/coqdep_lexer.cmi
-tools/coqdep_lexer.cmx : tools/coqdep_lexer.cmi
+lib/segmenttree.cmo : lib/segmenttree.cmi
+lib/segmenttree.cmx : lib/segmenttree.cmi
+lib/unicode.cmo : lib/unicodetable.cmo lib/segmenttree.cmi lib/unicode.cmi
+lib/unicode.cmx : lib/unicodetable.cmx lib/segmenttree.cmx lib/unicode.cmi
+tools/coqdep_lexer.cmo : lib/unicode.cmi tools/coqdep_lexer.cmi
+tools/coqdep_lexer.cmx : lib/unicode.cmx tools/coqdep_lexer.cmi
tools/coqdep_common.cmo : lib/minisys.cmo tools/coqdep_lexer.cmi tools/coqdep_common.cmi
tools/coqdep_common.cmx : lib/minisys.cmx tools/coqdep_lexer.cmx tools/coqdep_common.cmi
tools/coqdep_boot.cmo : tools/coqdep_common.cmi
diff --git a/dev/Bugzilla_Coq_autolink.user.js b/dev/Bugzilla_Coq_autolink.user.js
new file mode 100644
index 000000000..ed056021b
--- /dev/null
+++ b/dev/Bugzilla_Coq_autolink.user.js
@@ -0,0 +1,25 @@
+// ==UserScript==
+// @name Bugzilla Coq autolink
+// @namespace CoqScript
+// @include https://coq.inria.fr/bugs/*
+// @description Makes #XXXX into links to Github Coq PRs
+// @version 1
+// @grant none
+// ==/UserScript==
+
+var regex = /#(\d+)/g;
+var substr = '<a href="https://github.com/coq/coq/pull/$1">$&</a>';
+
+function doNode(node)
+{
+ node.innerHTML = node.innerHTML.replace(regex,substr);
+}
+
+var comments = document.getElementsByClassName("bz_comment_table")[0];
+var pars = comments.getElementsByClassName("bz_comment_text");
+
+for(var j=0; j<pars.length; j++)
+{
+ doNode(pars[j]);
+}
+
diff --git a/dev/ci/ci-basic-overlay.sh b/dev/ci/ci-basic-overlay.sh
index 8e7265969..43525dcd4 100644
--- a/dev/ci/ci-basic-overlay.sh
+++ b/dev/ci/ci-basic-overlay.sh
@@ -127,5 +127,5 @@
########################################################################
# Bignums
########################################################################
-: ${bignums_CI_BRANCH:=fix-thunk-printer}
-: ${bignums_CI_GITURL:=https://github.com/ppedrot/bignums.git}
+: ${bignums_CI_BRANCH:=master}
+: ${bignums_CI_GITURL:=https://github.com/coq/bignums.git}
diff --git a/dev/doc/changes.md b/dev/doc/changes.md
index 5ed74917a..b5e19f33c 100644
--- a/dev/doc/changes.md
+++ b/dev/doc/changes.md
@@ -641,6 +641,12 @@ The main search functions now take a function iterating over the
results. This allows for clients to use streaming or more economic
printing.
+### XML Protocol
+
+- In several places, flat text wrapped in `<string>` tags now appears as structured text inside `<richpp>` tags.
+
+- The "errormsg" feedback has been replaced by a "message" feedback which contains `<feedback\_content>` tag, with a message_level attribute of "error".
+
## Changes between Coq 8.4 and Coq 8.5
### Refactoring : more mli interfaces and simpler grammar.cma
diff --git a/dev/doc/xml-protocol.md b/dev/doc/xml-protocol.md
index cf7d205d8..18f6288f6 100644
--- a/dev/doc/xml-protocol.md
+++ b/dev/doc/xml-protocol.md
@@ -1,4 +1,4 @@
-#Coq XML Protocol for Coq 8.6#
+# Coq XML Protocol
This document is based on documentation originally written by CJ Bell
for his [vscoq](https://github.com/siegebell/vscoq/) project.
@@ -12,11 +12,7 @@ A somewhat out-of-date description of the async state machine is
[documented here](https://github.com/ejgallego/jscoq/blob/master/etc/notes/coq-notes.md).
OCaml types for the protocol can be found in the [`ide/interface.mli` file](/ide/interface.mli).
-# CHANGES
-## Changes from 8.5:
- * In several places, flat text wrapped in <string> tags now appears as structured text inside <richpp> tags
- * The "errormsg" feedback has been replaced by a "message" feedback which contains
- <feedback\_content> tag, with a message_level attribute of "error"
+Changes to the XML protocol are documented as part of [`dev/doc/changes.txt`](/dev/doc/changes.txt).
* [Commands](#commands)
- [About](#command-about)
diff --git a/doc/common/styles/html/coqremote/cover.html b/doc/common/styles/html/coqremote/cover.html
index 1c415eca6..5d151381f 100644
--- a/doc/common/styles/html/coqremote/cover.html
+++ b/doc/common/styles/html/coqremote/cover.html
@@ -52,20 +52,7 @@
<h2 style="text-align:center; font-size: 150%">The Coq Development Team</h2>
<br /><br /><br />
-<div style="text-align: left; font-size: 80%; text-indent: 0pt">
-<ul style="list-style: none; margin-left: 0pt">
- <li>V7.x © INRIA 1999-2004</li>
- <li>V8.0 © INRIA 2004-2008</li>
- <li>V8.1 © INRIA 2006-2011</li>
- <li>V8.2 © INRIA 2008-2011</li>
- <li>V8.3 © INRIA 2010-2011</li>
- <li>V8.4 © INRIA 2012-2014</li>
- <li>V8.5 © INRIA 2015-2016</li>
- <li>V8.6 © INRIA 2016</li>
-</ul>
-
-<p style="text-indent:0pt">This research was partly supported by IST
- working group ``Types''</p>
+<p style="text-indent:0pt">Copyright © INRIA 1999-2017</p>
<p style="text-indent:0pt">This material may be distributed only subject to the terms and conditions set forth in the Open Publication License, v1.0 or later (the latest version is presently available at <a href="http://www.opencontent.org/openpub">http://www.opencontent.org/openpub</a>). Options A and B are not elected.</p>
diff --git a/doc/common/styles/html/simple/cover.html b/doc/common/styles/html/simple/cover.html
index 25fb56320..605313104 100644
--- a/doc/common/styles/html/simple/cover.html
+++ b/doc/common/styles/html/simple/cover.html
@@ -30,20 +30,7 @@
<br /><br /><br />
-<div style="text-align: left; font-size: 80%; text-indent: 0pt">
-<ul style="list-style: none; margin-left: 0pt">
- <li>V7.x © INRIA 1999-2004</li>
- <li>V8.0 © INRIA 2004-2008</li>
- <li>V8.1 © INRIA 2006-2011</li>
- <li>V8.2 © INRIA 2008-2011</li>
- <li>V8.3 © INRIA 2010-2011</li>
- <li>V8.4 © INRIA 2012-2014</li>
- <li>V8.5 © INRIA 2015-2016</li>
- <li>V8.6 © INRIA 2016</li>
-</ul>
-
-<p style="text-indent:0pt">This research was partly supported by IST
- working group ``Types''</p>
+<p style="text-indent:0pt">Copyright © INRIA 1999-2017</p>
<p style="text-indent: 0pt">This material may be distributed only subject to the terms and conditions set forth in the Open Publication License, v1.0 or later (the latest version is presently available at <a href="http://www.opencontent.org/openpub">http://www.opencontent.org/openpub</a>). Options A and B are not elected.</p>
diff --git a/engine/evd.mli b/engine/evd.mli
index 9055dcc86..96e4b6acc 100644
--- a/engine/evd.mli
+++ b/engine/evd.mli
@@ -31,7 +31,7 @@ open Environ
(** {6 Evars} *)
type evar = existential_key
-(** Existential variables. TODO: Should be made opaque one day. *)
+(** Existential variables. *)
val string_of_existential : evar -> string
diff --git a/engine/namegen.ml b/engine/namegen.ml
index 1dd29e6ea..2e62b8901 100644
--- a/engine/namegen.ml
+++ b/engine/namegen.ml
@@ -43,6 +43,8 @@ let default_non_dependent_ident = Id.of_string default_non_dependent_string
let default_dependent_ident = Id.of_string "x"
+let default_generated_non_letter_string = "x"
+
(**********************************************************************)
(* Globality of identifiers *)
@@ -107,7 +109,17 @@ let head_name sigma c = (* Find the head constant of a constr if any *)
hdrec c
let lowercase_first_char id = (* First character of a constr *)
- Unicode.lowercase_first_char (Id.to_string id)
+ let s = Id.to_string id in
+ match Unicode.split_at_first_letter s with
+ | None ->
+ (* General case: nat -> n *)
+ Unicode.lowercase_first_char s
+ | Some (s,s') ->
+ if String.length s' = 0 then
+ (* No letter, e.g. __, or __'_, etc. *)
+ default_generated_non_letter_string
+ else
+ s ^ Unicode.lowercase_first_char s'
let sort_hdchar = function
| Prop(_) -> "P"
diff --git a/intf/vernacexpr.ml b/intf/vernacexpr.ml
index 03e8ea43d..4a471d4a2 100644
--- a/intf/vernacexpr.ml
+++ b/intf/vernacexpr.ml
@@ -91,7 +91,7 @@ type locatable =
| LocateTerm of reference or_by_notation
| LocateLibrary of reference
| LocateModule of reference
- | LocateTactic of reference
+ | LocateOther of string * reference
| LocateFile of string
type showable =
diff --git a/kernel/mod_typing.ml b/kernel/mod_typing.ml
index d2b41aae9..8568bf14b 100644
--- a/kernel/mod_typing.ml
+++ b/kernel/mod_typing.ml
@@ -166,16 +166,10 @@ let rec check_with_mod env struc (idl,mp1) mp equiv =
let mb_mp1 = lookup_module mp1 env in
let mtb_mp1 = module_type_of_module mb_mp1 in
let cst = match old.mod_expr with
- | Abstract ->
- begin
- try
- let mtb_old = module_type_of_module old in
- let chk_cst = Subtyping.check_subtypes env' mtb_mp1 mtb_old in
- Univ.ContextSet.add_constraints chk_cst old.mod_constraints
- with Failure _ ->
- (* TODO: where can a Failure come from ??? *)
- error_incorrect_with_constraint lab
- end
+ | Abstract ->
+ let mtb_old = module_type_of_module old in
+ let chk_cst = Subtyping.check_subtypes env' mtb_mp1 mtb_old in
+ Univ.ContextSet.add_constraints chk_cst old.mod_constraints
| Algebraic (NoFunctor (MEident(mp'))) ->
check_modpath_equiv env' mp1 mp';
old.mod_constraints
diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml
index ad622b07d..fd024b215 100644
--- a/kernel/safe_typing.ml
+++ b/kernel/safe_typing.ml
@@ -683,12 +683,15 @@ let build_module_body params restype senv =
with one extra component and some updated fields
(constraints, required, etc) *)
+let allow_delayed_constants = ref false
+
let propagate_senv newdef newenv newresolver senv oldsenv =
let now_cst, later_cst = List.partition Future.is_val senv.future_cst in
(* This asserts that after Paral-ITP, standard vo compilation is behaving
* exctly as before: the same universe constraints are added to modules *)
- if !Flags.compilation_mode = Flags.BuildVo &&
- !Flags.async_proofs_mode = Flags.APoff then assert(later_cst = []);
+ if not !allow_delayed_constants && later_cst <> [] then
+ CErrors.anomaly ~label:"safe_typing"
+ Pp.(str "True Future.t were created for opaque constants even if -async-proofs is off");
{ oldsenv with
env = newenv;
modresolver = newresolver;
diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli
index 752fdd793..f0f273f35 100644
--- a/kernel/safe_typing.mli
+++ b/kernel/safe_typing.mli
@@ -158,6 +158,10 @@ val add_module_parameter :
MBId.t -> Entries.module_struct_entry -> Declarations.inline ->
Mod_subst.delta_resolver safe_transformer
+(** Traditional mode: check at end of module that no future was
+ created. *)
+val allow_delayed_constants : bool ref
+
(** The optional result type is given without its functorial part *)
val end_module :
diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml
index 3f42c348f..0c0b910e1 100644
--- a/kernel/term_typing.ml
+++ b/kernel/term_typing.ml
@@ -428,13 +428,12 @@ let build_constant_declaration kn env result =
let expr =
!suggest_proof_using (Constant.to_string kn)
env vars ids_typ context_ids in
- if !Flags.compilation_mode = Flags.BuildVo then
- record_aux env ids_typ vars expr;
+ if !Flags.record_aux_file then record_aux env ids_typ vars expr;
vars
in
keep_hyps env (Idset.union ids_typ ids_def), def
| None ->
- if !Flags.compilation_mode = Flags.BuildVo then
+ if !Flags.record_aux_file then
record_aux env Id.Set.empty Id.Set.empty "";
[], def (* Empty section context: no need to check *)
| Some declared ->
@@ -614,7 +613,7 @@ let translate_local_def mb env id centry =
let open Cooking in
let decl = infer_declaration ~trust:mb env None (DefinitionEntry centry) in
let typ = decl.cook_type in
- if Option.is_empty decl.cook_context && !Flags.compilation_mode = Flags.BuildVo then begin
+ if Option.is_empty decl.cook_context && !Flags.record_aux_file then begin
match decl.cook_body with
| Undef _ -> ()
| Def _ -> ()
diff --git a/kernel/typeops.ml b/kernel/typeops.ml
index 044877e82..b40badd7c 100644
--- a/kernel/typeops.ml
+++ b/kernel/typeops.ml
@@ -65,6 +65,10 @@ let type_of_type u =
let uu = Universe.super u in
mkType uu
+let type_of_sort = function
+ | Prop c -> type1
+ | Type u -> type_of_type u
+
(*s Type of a de Bruijn index. *)
let type_of_relative env n =
@@ -323,11 +327,7 @@ let rec execute env cstr =
let open Context.Rel.Declaration in
match kind_of_term cstr with
(* Atomic terms *)
- | Sort (Prop c) ->
- type1
-
- | Sort (Type u) ->
- type_of_type u
+ | Sort s -> type_of_sort s
| Rel n ->
type_of_relative env n
diff --git a/kernel/typeops.mli b/kernel/typeops.mli
index a8f7fba9a..96be6c14a 100644
--- a/kernel/typeops.mli
+++ b/kernel/typeops.mli
@@ -37,15 +37,19 @@ val assumption_of_judgment : env -> unsafe_judgment -> types
val type_judgment : env -> unsafe_judgment -> unsafe_type_judgment
(** {6 Type of sorts. } *)
+val type1 : types
+val type_of_sort : Sorts.t -> types
val judge_of_prop : unsafe_judgment
val judge_of_set : unsafe_judgment
val judge_of_prop_contents : contents -> unsafe_judgment
val judge_of_type : universe -> unsafe_judgment
(** {6 Type of a bound variable. } *)
+val type_of_relative : env -> int -> types
val judge_of_relative : env -> int -> unsafe_judgment
(** {6 Type of variables } *)
+val type_of_variable : env -> variable -> types
val judge_of_variable : env -> variable -> unsafe_judgment
(** {6 type of a constant } *)
@@ -66,9 +70,9 @@ val judge_of_abstraction :
env -> Name.t -> unsafe_type_judgment -> unsafe_judgment
-> unsafe_judgment
-val sort_of_product : env -> sorts -> sorts -> sorts
-
(** {6 Type of a product. } *)
+val sort_of_product : env -> sorts -> sorts -> sorts
+val type_of_product : env -> Name.t -> sorts -> sorts -> types
val judge_of_product :
env -> Name.t -> unsafe_type_judgment -> unsafe_type_judgment
-> unsafe_judgment
diff --git a/lib/flags.ml b/lib/flags.ml
index d4be81c61..a178eb755 100644
--- a/lib/flags.ml
+++ b/lib/flags.ml
@@ -43,11 +43,8 @@ let with_extra_values o l f x =
let boot = ref false
let load_init = ref true
-let batch_mode = ref false
-type compilation_mode = BuildVo | BuildVio | Vio2Vo
-let compilation_mode = ref BuildVo
-let compilation_output_name = ref None
+let record_aux_file = ref false
let test_mode = ref false
diff --git a/lib/flags.mli b/lib/flags.mli
index 3024c6039..8ec2a8073 100644
--- a/lib/flags.mli
+++ b/lib/flags.mli
@@ -13,12 +13,9 @@
val boot : bool ref
val load_init : bool ref
-(* Will affect STM caching *)
-val batch_mode : bool ref
-
-type compilation_mode = BuildVo | BuildVio | Vio2Vo
-val compilation_mode : compilation_mode ref
-val compilation_output_name : string option ref
+(** Set by coqtop to tell the kernel to output to the aux file; will
+ be eventually removed by cleanups such as PR#1103 *)
+val record_aux_file : bool ref
(* Flag set when the test-suite is called. Its only effect to display
verbose information for `Fail` *)
diff --git a/lib/pp.ml b/lib/pp.ml
index 88ddcb35b..c3338688d 100644
--- a/lib/pp.ml
+++ b/lib/pp.ml
@@ -82,10 +82,21 @@ let utf8_length s =
done ;
!cnt
-let app s1 s2 = match s1, s2 with
- | Ppcmd_empty, s
- | s, Ppcmd_empty -> s
- | s1, s2 -> Ppcmd_glue [s1; s2]
+let rec app d1 d2 = match d1, d2 with
+ | Ppcmd_empty, d
+ | d, Ppcmd_empty -> d
+
+ (* Optimizations *)
+ | Ppcmd_glue [l1;l2], Ppcmd_glue l3 -> Ppcmd_glue (l1 :: l2 :: l3)
+ | Ppcmd_glue [l1;l2], d2 -> Ppcmd_glue [l1 ; l2 ; d2]
+ | d1, Ppcmd_glue l2 -> Ppcmd_glue (d1 :: l2)
+
+ | Ppcmd_tag(t1,d1), Ppcmd_tag(t2,d2)
+ when t1 = t2 -> Ppcmd_tag(t1,app d1 d2)
+ | d1, d2 -> Ppcmd_glue [d1; d2]
+ (* Optimizations deemed too costly *)
+ (* | Ppcmd_glue l1, Ppcmd_glue l2 -> Ppcmd_glue (l1 @ l2) *)
+ (* | Ppcmd_string s1, Ppcmd_string s2 -> Ppcmd_string (s1 ^ s2) *)
let seq s = Ppcmd_glue s
diff --git a/lib/system.ml b/lib/system.ml
index 0b64b237d..4b5066ef4 100644
--- a/lib/system.ml
+++ b/lib/system.ml
@@ -54,6 +54,8 @@ let make_dir_table dir =
let filter_dotfiles s f = if f.[0] = '.' then s else StrSet.add f s in
Array.fold_left filter_dotfiles StrSet.empty (Sys.readdir dir)
+let trust_file_cache = ref true
+
let exists_in_dir_respecting_case dir bf =
let cache_dir dir =
let contents = make_dir_table dir in
@@ -62,10 +64,10 @@ let exists_in_dir_respecting_case dir bf =
let contents, fresh =
try
(* in batch mode, assume the directory content is still fresh *)
- StrMap.find dir !dirmap, !Flags.batch_mode
+ StrMap.find dir !dirmap, !trust_file_cache
with Not_found ->
(* in batch mode, we are not yet sure the directory exists *)
- if !Flags.batch_mode && not (exists_dir dir) then StrSet.empty, true
+ if !trust_file_cache && not (exists_dir dir) then StrSet.empty, true
else cache_dir dir, true in
StrSet.mem bf contents ||
not fresh &&
@@ -80,7 +82,7 @@ let file_exists_respecting_case path f =
let df = Filename.dirname f in
(String.equal df "." || aux df)
&& exists_in_dir_respecting_case (Filename.concat path df) bf
- in (!Flags.batch_mode || Sys.file_exists (Filename.concat path f)) && aux f
+ in (!trust_file_cache || Sys.file_exists (Filename.concat path f)) && aux f
let rec search paths test =
match paths with
diff --git a/lib/system.mli b/lib/system.mli
index 7281de97c..aa964abeb 100644
--- a/lib/system.mli
+++ b/lib/system.mli
@@ -54,6 +54,12 @@ val where_in_path_rex :
val find_file_in_path :
?warn:bool -> CUnix.load_path -> string -> CUnix.physical_path * string
+val trust_file_cache : bool ref
+(** [trust_file_cache] indicates whether we trust the underlying
+ mapped file-system not to change along the execution of Coq. This
+ assumption greatly speds up file search, but it is often
+ inconvenient in interactive mode *)
+
val file_exists_respecting_case : string -> string -> bool
(** {6 I/O functions } *)
diff --git a/lib/unicode.ml b/lib/unicode.ml
index 8eb2eb45d..f193c4e0f 100644
--- a/lib/unicode.ml
+++ b/lib/unicode.ml
@@ -8,13 +8,14 @@
(** Unicode utilities *)
-type status = Letter | IdentPart | Symbol | Unknown
+type status = Letter | IdentPart | Symbol | IdentSep | Unknown
(* The following table stores classes of Unicode characters that
- are used by the lexer. There are 3 different classes so 2 bits are
- allocated for each character. We only use 16 bits over the 31 bits
- to simplify the masking process. (This choice seems to be a good
- trade-off between speed and space after some benchmarks.) *)
+ are used by the lexer. There are 5 different classes so 3 bits
+ are allocated for each character. We encode the masks of 8
+ characters per word, thus using 24 bits over the 31 available
+ bits. (This choice seems to be a good trade-off between speed
+ and space after some benchmarks.) *)
(* A 256 KiB table, initially filled with zeros. *)
let table = Array.make (1 lsl 17) 0
@@ -24,14 +25,15 @@ let table = Array.make (1 lsl 17) 0
define the position of the pattern in the word.
Notice that pattern "00" means "undefined". *)
let mask i = function
- | Letter -> 1 lsl ((i land 7) lsl 1) (* 01 *)
- | IdentPart -> 2 lsl ((i land 7) lsl 1) (* 10 *)
- | Symbol -> 3 lsl ((i land 7) lsl 1) (* 11 *)
- | Unknown -> 0 lsl ((i land 7) lsl 1) (* 00 *)
+ | Letter -> 1 lsl ((i land 7) * 3) (* 001 *)
+ | IdentPart -> 2 lsl ((i land 7) * 3) (* 010 *)
+ | Symbol -> 3 lsl ((i land 7) * 3) (* 011 *)
+ | IdentSep -> 4 lsl ((i land 7) * 3) (* 100 *)
+ | Unknown -> 0 lsl ((i land 7) * 3) (* 000 *)
-(* Helper to reset 2 bits in a word. *)
+(* Helper to reset 3 bits in a word. *)
let reset_mask i =
- lnot (3 lsl ((i land 7) lsl 1))
+ lnot (7 lsl ((i land 7) * 3))
(* Initialize the lookup table from a list of segments, assigning
a status to every character of each segment. The order of these
@@ -50,13 +52,14 @@ let mk_lookup_table_from_unicode_tables_for status tables =
(* Look up into the table and interpret the found pattern. *)
let lookup x =
- let v = (table.(x lsr 3) lsr ((x land 7) lsl 1)) land 3 in
+ let v = (table.(x lsr 3) lsr ((x land 7) * 3)) land 7 in
if v = 1 then Letter
else if v = 2 then IdentPart
else if v = 3 then Symbol
+ else if v = 4 then IdentSep
else Unknown
-(* [classify] discriminates between 3 different kinds of
+(* [classify] discriminates between 5 different kinds of
symbols based on the standard unicode classification (extracted from
Camomile). *)
let classify =
@@ -67,13 +70,13 @@ let classify =
Unicodetable.sm; (* Symbol, maths. *)
Unicodetable.sc; (* Symbol, currency. *)
Unicodetable.so; (* Symbol, modifier. *)
- Unicodetable.pd; (* Punctation, dash. *)
- Unicodetable.pc; (* Punctation, connector. *)
- Unicodetable.pe; (* Punctation, open. *)
- Unicodetable.ps; (* Punctation, close. *)
- Unicodetable.pi; (* Punctation, initial quote. *)
- Unicodetable.pf; (* Punctation, final quote. *)
- Unicodetable.po; (* Punctation, other. *)
+ Unicodetable.pd; (* Punctuation, dash. *)
+ Unicodetable.pc; (* Punctuation, connector. *)
+ Unicodetable.pe; (* Punctuation, open. *)
+ Unicodetable.ps; (* Punctution, close. *)
+ Unicodetable.pi; (* Punctuation, initial quote. *)
+ Unicodetable.pf; (* Punctuation, final quote. *)
+ Unicodetable.po; (* Punctuation, other. *)
];
mk_lookup_table_from_unicode_tables_for Letter
[
@@ -107,14 +110,14 @@ let classify =
[(0x02074, 0x02079)]; (* Superscript 4-9. *)
single 0x0002E; (* Dot. *)
];
- mk_lookup_table_from_unicode_tables_for Letter
+ mk_lookup_table_from_unicode_tables_for IdentSep
[
single 0x005F; (* Underscore. *)
single 0x00A0; (* Non breaking space. *)
];
mk_lookup_table_from_unicode_tables_for IdentPart
[
- single 0x0027; (* Special space. *)
+ single 0x0027; (* Single quote. *)
];
(* Lookup *)
lookup
@@ -198,22 +201,40 @@ let escaped_if_non_utf8 s =
(* Check the well-formedness of an identifier *)
+let is_valid_ident_initial = function
+ | Letter | IdentSep -> true
+ | IdentPart | Symbol | Unknown -> false
+
let initial_refutation j n s =
- match classify n with
- | Letter -> None
- | _ ->
+ if is_valid_ident_initial (classify n) then None
+ else
let c = String.sub s 0 j in
Some (false,
"Invalid character '"^c^"' at beginning of identifier \""^s^"\".")
+let is_valid_ident_trailing = function
+ | Letter | IdentSep | IdentPart -> true
+ | Symbol | Unknown -> false
+
let trailing_refutation i j n s =
- match classify n with
- | Letter | IdentPart -> None
- | _ ->
+ if is_valid_ident_trailing (classify n) then None
+ else
let c = String.sub s i j in
Some (false,
"Invalid character '"^c^"' in identifier \""^s^"\".")
+let is_unknown = function
+ | Unknown -> true
+ | Letter | IdentSep | IdentPart | Symbol -> false
+
+let is_ident_part = function
+ | IdentPart -> true
+ | Letter | IdentSep | Symbol | Unknown -> false
+
+let is_ident_sep = function
+ | IdentSep -> true
+ | Letter | IdentPart | Symbol | Unknown -> false
+
let ident_refutation s =
if s = ".." then None else try
let j, n = next_utf8 s 0 in
@@ -247,6 +268,26 @@ let lowercase_first_char s =
let j, n = next_utf8 s 0 in
utf8_of_unicode (lowercase_unicode n)
+let split_at_first_letter s =
+ let n, v = next_utf8 s 0 in
+ if ((* optim *) n = 1 && s.[0] != '_') || not (is_ident_sep (classify v)) then None
+ else begin
+ let n = ref n in
+ let p = ref 0 in
+ while !n < String.length s &&
+ let n', v = next_utf8 s !n in
+ p := n';
+ (* Test if not letter *)
+ ((* optim *) n' = 1 && (s.[!n] = '_' || s.[!n] = '\''))
+ || let st = classify v in
+ is_ident_sep st || is_ident_part st
+ do n := !n + !p
+ done;
+ let s1 = String.sub s 0 !n in
+ let s2 = String.sub s !n (String.length s - !n) in
+ Some (s1,s2)
+ end
+
(** For extraction, we need to encode unicode character into ascii ones *)
let is_basic_ascii s =
diff --git a/lib/unicode.mli b/lib/unicode.mli
index b8e7c33ad..32ffbb8e9 100644
--- a/lib/unicode.mli
+++ b/lib/unicode.mli
@@ -8,7 +8,7 @@
(** Unicode utilities *)
-type status = Letter | IdentPart | Symbol | Unknown
+type status
(** Classify a unicode char into 3 classes or unknown. *)
val classify : int -> status
@@ -17,10 +17,23 @@ val classify : int -> status
Return [Some (b,s)] otherwise, where [s] is an explanation and [b] is severity. *)
val ident_refutation : string -> (bool * string) option
+(** Tells if a valid initial character for an identifier *)
+val is_valid_ident_initial : status -> bool
+
+(** Tells if a valid non-initial character for an identifier *)
+val is_valid_ident_trailing : status -> bool
+
+(** Tells if a character is unclassified *)
+val is_unknown : status -> bool
+
(** First char of a string, converted to lowercase
@raise Assert_failure if the input string is empty. *)
val lowercase_first_char : string -> string
+(** Split a string supposed to be an ident at the first letter;
+ as an optimization, return None if the first character is a letter *)
+val split_at_first_letter : string -> (string * string) option
+
(** Return [true] if all UTF-8 characters in the input string are just plain
ASCII characters. Returns [false] otherwise. *)
val is_basic_ascii : string -> bool
diff --git a/library/library.ml b/library/library.ml
index 1da2c591d..e2832ecdc 100644
--- a/library/library.ml
+++ b/library/library.ml
@@ -683,7 +683,8 @@ let error_recursively_dependent_library dir =
let save_library_to ?todo dir f otab =
let except = match todo with
| None ->
- assert(!Flags.compilation_mode = Flags.BuildVo);
+ (* XXX *)
+ (* assert(!Flags.compilation_mode = Flags.BuildVo); *)
assert(Filename.check_suffix f ".vo");
Future.UUIDSet.empty
| Some (l,_) ->
diff --git a/library/nametab.ml b/library/nametab.ml
index 68fdbb4f3..0ec4a37cd 100644
--- a/library/nametab.ml
+++ b/library/nametab.ml
@@ -19,10 +19,6 @@ exception GlobalizationError of qualid
let error_global_not_found ?loc q =
Loc.raise ?loc (GlobalizationError q)
-(* Kinds of global names *)
-
-type ltac_constant = kernel_name
-
(* The visibility can be registered either
- for all suffixes not shorter then a given int - when the object
is loaded inside a module
@@ -274,19 +270,14 @@ struct
end
module ExtRefEqual = ExtRefOrdered
-module KnEqual = Names.KerName
module MPEqual = Names.ModPath
module ExtRefTab = Make(FullPath)(ExtRefEqual)
-module KnTab = Make(FullPath)(KnEqual)
module MPTab = Make(FullPath)(MPEqual)
type ccitab = ExtRefTab.t
let the_ccitab = ref (ExtRefTab.empty : ccitab)
-type kntab = KnTab.t
-let the_tactictab = ref (KnTab.empty : kntab)
-
type mptab = MPTab.t
let the_modtypetab = ref (MPTab.empty : mptab)
@@ -327,10 +318,6 @@ let the_modrevtab = ref (MPmap.empty : mprevtab)
type mptrevtab = full_path MPmap.t
let the_modtyperevtab = ref (MPmap.empty : mptrevtab)
-type knrevtab = full_path KNmap.t
-let the_tacticrevtab = ref (KNmap.empty : knrevtab)
-
-
(* Push functions *********************************************************)
(* This is for permanent constructions (never discharged -- but with
@@ -368,13 +355,6 @@ let push_modtype vis sp kn =
the_modtypetab := MPTab.push vis sp kn !the_modtypetab;
the_modtyperevtab := MPmap.add kn sp !the_modtyperevtab
-(* This is for tactic definition names *)
-
-let push_tactic vis sp kn =
- the_tactictab := KnTab.push vis sp kn !the_tactictab;
- the_tacticrevtab := KNmap.add kn sp !the_tacticrevtab
-
-
(* This is to remember absolute Section/Module names and to avoid redundancy *)
let push_dir vis dir dir_ref =
the_dirtab := DirTab.push vis dir dir_ref !the_dirtab;
@@ -402,8 +382,6 @@ let locate_syndef qid = match locate_extended qid with
let locate_modtype qid = MPTab.locate qid !the_modtypetab
let full_name_modtype qid = MPTab.user_name qid !the_modtypetab
-let locate_tactic qid = KnTab.locate qid !the_tactictab
-
let locate_dir qid = DirTab.locate qid !the_dirtab
let locate_module qid =
@@ -428,8 +406,6 @@ let locate_all qid =
let locate_extended_all qid = ExtRefTab.find_prefixes qid !the_ccitab
-let locate_extended_all_tactic qid = KnTab.find_prefixes qid !the_tactictab
-
let locate_extended_all_dir qid = DirTab.find_prefixes qid !the_dirtab
let locate_extended_all_modtype qid = MPTab.find_prefixes qid !the_modtypetab
@@ -471,8 +447,6 @@ let exists_module = exists_dir
let exists_modtype sp = MPTab.exists sp !the_modtypetab
-let exists_tactic kn = KnTab.exists kn !the_tactictab
-
(* Reverse locate functions ***********************************************)
let path_of_global ref =
@@ -492,9 +466,6 @@ let path_of_syndef kn =
let dirpath_of_module mp =
MPmap.find mp !the_modrevtab
-let path_of_tactic kn =
- KNmap.find kn !the_tacticrevtab
-
let path_of_modtype mp =
MPmap.find mp !the_modtyperevtab
@@ -519,10 +490,6 @@ let shortest_qualid_of_modtype kn =
let sp = MPmap.find kn !the_modtyperevtab in
MPTab.shortest_qualid Id.Set.empty sp !the_modtypetab
-let shortest_qualid_of_tactic kn =
- let sp = KNmap.find kn !the_tacticrevtab in
- KnTab.shortest_qualid Id.Set.empty sp !the_tactictab
-
let pr_global_env env ref =
try pr_qualid (shortest_qualid_of_global env ref)
with Not_found as e ->
@@ -541,28 +508,24 @@ let global_inductive r =
(********************************************************************)
(* Registration of tables as a global table and rollback *)
-type frozen = ccitab * dirtab * mptab * kntab
- * globrevtab * mprevtab * mptrevtab * knrevtab
+type frozen = ccitab * dirtab * mptab
+ * globrevtab * mprevtab * mptrevtab
let freeze _ : frozen =
!the_ccitab,
!the_dirtab,
!the_modtypetab,
- !the_tactictab,
!the_globrevtab,
!the_modrevtab,
- !the_modtyperevtab,
- !the_tacticrevtab
+ !the_modtyperevtab
-let unfreeze (ccit,dirt,mtyt,tact,globr,modr,mtyr,tacr) =
+let unfreeze (ccit,dirt,mtyt,globr,modr,mtyr) =
the_ccitab := ccit;
the_dirtab := dirt;
the_modtypetab := mtyt;
- the_tactictab := tact;
the_globrevtab := globr;
the_modrevtab := modr;
- the_modtyperevtab := mtyr;
- the_tacticrevtab := tacr
+ the_modtyperevtab := mtyr
let _ =
Summary.declare_summary "names"
diff --git a/library/nametab.mli b/library/nametab.mli
index 025a63b1c..3a380637c 100644
--- a/library/nametab.mli
+++ b/library/nametab.mli
@@ -78,10 +78,6 @@ val push_modtype : visibility -> full_path -> module_path -> unit
val push_dir : visibility -> DirPath.t -> global_dir_reference -> unit
val push_syndef : visibility -> full_path -> syndef_name -> unit
-type ltac_constant = kernel_name
-val push_tactic : visibility -> full_path -> ltac_constant -> unit
-
-
(** {6 The following functions perform globalization of qualified names } *)
(** These functions globalize a (partially) qualified name or fail with
@@ -95,7 +91,6 @@ val locate_modtype : qualid -> module_path
val locate_dir : qualid -> global_dir_reference
val locate_module : qualid -> module_path
val locate_section : qualid -> DirPath.t
-val locate_tactic : qualid -> ltac_constant
(** These functions globalize user-level references into global
references, like [locate] and co, but raise a nice error message
@@ -109,7 +104,6 @@ val global_inductive : reference -> inductive
val locate_all : qualid -> global_reference list
val locate_extended_all : qualid -> extended_global_reference list
-val locate_extended_all_tactic : qualid -> ltac_constant list
val locate_extended_all_dir : qualid -> global_dir_reference list
val locate_extended_all_modtype : qualid -> module_path list
@@ -125,7 +119,6 @@ val exists_modtype : full_path -> bool
val exists_dir : DirPath.t -> bool
val exists_section : DirPath.t -> bool (** deprecated synonym of [exists_dir] *)
val exists_module : DirPath.t -> bool (** deprecated synonym of [exists_dir] *)
-val exists_tactic : full_path -> bool (** deprecated synonym of [exists_dir] *)
(** {6 These functions locate qualids into full user names } *)
@@ -144,7 +137,6 @@ val path_of_syndef : syndef_name -> full_path
val path_of_global : global_reference -> full_path
val dirpath_of_module : module_path -> DirPath.t
val path_of_modtype : module_path -> full_path
-val path_of_tactic : ltac_constant -> full_path
(** Returns in particular the dirpath or the basename of the full path
associated to global reference *)
@@ -166,7 +158,6 @@ val shortest_qualid_of_global : Id.Set.t -> global_reference -> qualid
val shortest_qualid_of_syndef : Id.Set.t -> syndef_name -> qualid
val shortest_qualid_of_modtype : module_path -> qualid
val shortest_qualid_of_module : module_path -> qualid
-val shortest_qualid_of_tactic : ltac_constant -> qualid
(** Deprecated synonyms *)
diff --git a/man/coqdep.1 b/man/coqdep.1
index 81f7e1e0d..ed727db7c 100644
--- a/man/coqdep.1
+++ b/man/coqdep.1
@@ -82,7 +82,7 @@ Prints the dependencies of Caml modules.
\" the standard output. No dependency is computed with this option.
.TP
.BI \-I/\-Q/\-R \ options
-Have the same effects on load path and modules names than for other
+Have the same effects on load path and modules names as for other
coq commands (coqtop, coqc).
.TP
.BI \-coqlib \ directory
diff --git a/parsing/cLexer.ml4 b/parsing/cLexer.ml4
index 9c9189ffe..f26398fa9 100644
--- a/parsing/cLexer.ml4
+++ b/parsing/cLexer.ml4
@@ -245,8 +245,8 @@ let check_ident str =
loop_id true s
| [< s >] ->
match unlocated lookup_utf8 Ploc.dummy s with
- | Utf8Token (Unicode.Letter, n) -> njunk n s; loop_id true s
- | Utf8Token (Unicode.IdentPart, n) when intail ->
+ | Utf8Token (st, n) when not intail && Unicode.is_valid_ident_initial st -> njunk n s; loop_id true s
+ | Utf8Token (st, n) when intail && Unicode.is_valid_ident_trailing st ->
njunk n s;
loop_id true s
| EmptyStream -> ()
@@ -311,9 +311,9 @@ let rec ident_tail loc len = parser
ident_tail loc (store len c) s
| [< s >] ->
match lookup_utf8 loc s with
- | Utf8Token ((Unicode.IdentPart | Unicode.Letter), n) ->
+ | Utf8Token (st, n) when Unicode.is_valid_ident_trailing st ->
ident_tail loc (nstore n len s) s
- | Utf8Token (Unicode.Unknown, n) ->
+ | Utf8Token (st, n) when Unicode.is_unknown st ->
let id = get_buff len in
let u = String.concat "" (List.map (String.make 1) (Stream.npeek n s)) in
warn_unrecognized_unicode ~loc:!@loc (u,id); len
@@ -539,7 +539,7 @@ let parse_after_dot loc c bp =
(try find_keyword loc ("."^field) s with Not_found -> FIELD field)
| [< s >] ->
match lookup_utf8 loc s with
- | Utf8Token (Unicode.Letter, n) ->
+ | Utf8Token (st, n) when Unicode.is_valid_ident_initial st ->
let len = ident_tail loc (nstore n 0 s) s in
let field = get_buff len in
(try find_keyword loc ("."^field) s with Not_found -> FIELD field)
@@ -553,7 +553,7 @@ let parse_after_qmark loc bp s =
| None -> KEYWORD "?"
| _ ->
match lookup_utf8 loc s with
- | Utf8Token (Unicode.Letter, _) -> LEFTQMARK
+ | Utf8Token (st, _) when Unicode.is_valid_ident_initial st -> LEFTQMARK
| AsciiChar | Utf8Token _ | EmptyStream ->
fst (process_chars loc bp '?' s)
@@ -618,13 +618,13 @@ let rec next_token loc = parser bp
comment_stop bp; between_commands := new_between_commands; t
| [< s >] ->
match lookup_utf8 loc s with
- | Utf8Token (Unicode.Letter, n) ->
+ | Utf8Token (st, n) when Unicode.is_valid_ident_initial st ->
let len = ident_tail loc (nstore n 0 s) s in
let id = get_buff len in
let ep = Stream.count s in
comment_stop bp;
(try find_keyword loc id s with Not_found -> IDENT id), set_loc_pos loc bp ep
- | AsciiChar | Utf8Token ((Unicode.Symbol | Unicode.IdentPart | Unicode.Unknown), _) ->
+ | AsciiChar | Utf8Token _ ->
let t = process_chars loc bp (Stream.next s) s in
comment_stop bp; t
| EmptyStream ->
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index 819d236cd..3e0b85b54 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -1020,8 +1020,7 @@ GEXTEND Gram
| IDENT "Term"; qid = smart_global -> LocateTerm qid
| IDENT "File"; f = ne_string -> LocateFile f
| IDENT "Library"; qid = global -> LocateLibrary qid
- | IDENT "Module"; qid = global -> LocateModule qid
- | IDENT "Ltac"; qid = global -> LocateTactic qid ] ]
+ | IDENT "Module"; qid = global -> LocateModule qid ] ]
;
option_value:
[ [ n = integer -> IntValue (Some n)
diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml
index f503c572d..3c46d5c43 100644
--- a/plugins/extraction/extract_env.ml
+++ b/plugins/extraction/extract_env.ml
@@ -342,7 +342,7 @@ let rec extract_structure env mp reso ~all = function
and extract_mexpr env mp = function
| MEwith _ -> assert false (* no 'with' syntax for modules *)
- | me when lang () != Ocaml ->
+ | me when lang () != Ocaml || Table.is_extrcompute () ->
(* In Haskell/Scheme, we expand everything.
For now, we also extract everything, dead code will be removed later
(see [Modutil.optimize_struct]. *)
@@ -570,11 +570,12 @@ let print_structure_to_file (fn,si,mo) dry struc =
let reset () =
Visit.reset (); reset_tables (); reset_renaming_tables Everything
-let init modular library =
+let init ?(compute=false) modular library =
check_inside_section (); check_inside_module ();
set_keywords (descr ()).keywords;
set_modular modular;
set_library library;
+ set_extrcompute compute;
reset ();
if modular && lang () == Scheme then error_scheme ()
@@ -684,8 +685,22 @@ let extraction_library is_rec m =
List.iter print struc;
reset ()
+(** For extraction compute, we flatten all the module structure,
+ getting rid of module types or unapplied functors *)
+
+let flatten_structure struc =
+ let rec flatten_elem (lab,elem) = match elem with
+ |SEdecl d -> [d]
+ |SEmodtype _ -> []
+ |SEmodule m -> match m.ml_mod_expr with
+ |MEfunctor _ -> []
+ |MEident _ | MEapply _ -> assert false (* should be expanded *)
+ |MEstruct (_,elems) -> flatten_elems elems
+ and flatten_elems l = List.flatten (List.map flatten_elem l)
+ in flatten_elems (List.flatten (List.map snd struc))
+
let structure_for_compute c =
- init false false;
+ init false false ~compute:true;
let env = Global.env () in
let ast, mlt = Extraction.extract_constr env c in
let ast = Mlutil.normalize ast in
@@ -694,8 +709,7 @@ let structure_for_compute c =
let () = ast_iter_references add_ref add_ref add_ref ast in
let refs = Refset.elements !refs in
let struc = optimize_struct (refs,[]) (mono_environment refs []) in
- let flatstruc = List.map snd (List.flatten (List.map snd struc)) in
- flatstruc, ast, mlt
+ (flatten_structure struc), ast, mlt
(* For the test-suite :
extraction to a temporary file + run ocamlc on it *)
diff --git a/plugins/extraction/extract_env.mli b/plugins/extraction/extract_env.mli
index 5769ff117..7bbb825b1 100644
--- a/plugins/extraction/extract_env.mli
+++ b/plugins/extraction/extract_env.mli
@@ -34,5 +34,4 @@ val print_one_decl :
(* Used by Extraction Compute *)
val structure_for_compute :
- Term.constr ->
- Miniml.ml_flat_structure * Miniml.ml_ast * Miniml.ml_type
+ Term.constr -> (Miniml.ml_decl list) * Miniml.ml_ast * Miniml.ml_type
diff --git a/plugins/extraction/miniml.mli b/plugins/extraction/miniml.mli
index edebba49d..5e967ef37 100644
--- a/plugins/extraction/miniml.mli
+++ b/plugins/extraction/miniml.mli
@@ -187,8 +187,6 @@ type ml_structure = (ModPath.t * ml_module_structure) list
type ml_signature = (ModPath.t * ml_module_sig) list
-type ml_flat_structure = ml_structure_elem list
-
type unsafe_needs = {
mldummy : bool;
tdummy : bool;
diff --git a/plugins/extraction/mlutil.ml b/plugins/extraction/mlutil.ml
index a4c2bcd88..b01b0198d 100644
--- a/plugins/extraction/mlutil.ml
+++ b/plugins/extraction/mlutil.ml
@@ -127,11 +127,15 @@ let rec mgu = function
| Taxiom, Taxiom -> ()
| _ -> raise Impossible
-let needs_magic p = try mgu p; false with Impossible -> true
+let skip_typing () = lang () == Scheme || is_extrcompute ()
-let put_magic_if b a = if b && lang () != Scheme then MLmagic a else a
+let needs_magic p =
+ if skip_typing () then false
+ else try mgu p; false with Impossible -> true
-let put_magic p a = if needs_magic p && lang () != Scheme then MLmagic a else a
+let put_magic_if b a = if b then MLmagic a else a
+
+let put_magic p a = if needs_magic p then MLmagic a else a
let generalizable a =
lang () != Ocaml ||
@@ -769,6 +773,20 @@ let eta_red e =
else e
| _ -> e
+(* Performs an eta-reduction when the core is atomic,
+ or otherwise returns None *)
+
+let atomic_eta_red e =
+ let ids,t = collect_lams e in
+ let n = List.length ids in
+ match t with
+ | MLapp (f,a) when test_eta_args_lift 0 n a ->
+ (match f with
+ | MLrel k when k>n -> Some (MLrel (k-n))
+ | MLglob _ | MLexn _ | MLdummy _ -> Some f
+ | _ -> None)
+ | _ -> None
+
(*s Computes all head linear beta-reductions possible in [(t a)].
Non-linear head beta-redex become let-in. *)
@@ -1053,6 +1071,10 @@ let rec simpl o = function
simpl o (MLcase(typ,e,br'))
| MLmagic(MLdummy _ as e) when lang () == Haskell -> e
| MLmagic(MLexn _ as e) -> e
+ | MLlam _ as e ->
+ (match atomic_eta_red e with
+ | Some e' -> e'
+ | None -> ast_map (simpl o) e)
| a -> ast_map (simpl o) a
(* invariant : list [a] of arguments is non-empty *)
diff --git a/plugins/extraction/table.ml b/plugins/extraction/table.ml
index 30e3b520f..995d5fd19 100644
--- a/plugins/extraction/table.ml
+++ b/plugins/extraction/table.ml
@@ -250,6 +250,11 @@ let modular () = !modular_ref
let set_library b = library_ref := b
let library () = !library_ref
+let extrcompute = ref false
+
+let set_extrcompute b = extrcompute := b
+let is_extrcompute () = !extrcompute
+
(*s Printing. *)
(* The following functions work even on objects not in [Global.env ()].
diff --git a/plugins/extraction/table.mli b/plugins/extraction/table.mli
index 7e47d0bc8..cc93f294b 100644
--- a/plugins/extraction/table.mli
+++ b/plugins/extraction/table.mli
@@ -165,6 +165,9 @@ val modular : unit -> bool
val set_library : bool -> unit
val library : unit -> bool
+val set_extrcompute : bool -> unit
+val is_extrcompute : unit -> bool
+
(*s Table for custom inlining *)
val to_inline : global_reference -> bool
diff --git a/plugins/ltac/g_ltac.ml4 b/plugins/ltac/g_ltac.ml4
index 86c983bdd..d639dd0e1 100644
--- a/plugins/ltac/g_ltac.ml4
+++ b/plugins/ltac/g_ltac.ml4
@@ -482,6 +482,11 @@ VERNAC COMMAND EXTEND VernacPrintLtac CLASSIFIED AS QUERY
[ Feedback.msg_notice (Tacintern.print_ltac (snd (Libnames.qualid_of_reference r))) ]
END
+VERNAC COMMAND EXTEND VernacLocateLtac CLASSIFIED AS QUERY
+| [ "Locate" "Ltac" reference(r) ] ->
+ [ Tacentries.print_located_tactic r ]
+END
+
let pr_ltac_ref = Libnames.pr_reference
let pr_tacdef_body tacdef_body =
diff --git a/plugins/ltac/ltac_plugin.mlpack b/plugins/ltac/ltac_plugin.mlpack
index 12b4c81fc..3972b7aac 100644
--- a/plugins/ltac/ltac_plugin.mlpack
+++ b/plugins/ltac/ltac_plugin.mlpack
@@ -1,9 +1,9 @@
Tacarg
+Tacsubst
+Tacenv
Pptactic
Pltac
Taccoerce
-Tacsubst
-Tacenv
Tactic_debug
Tacintern
Tacentries
diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml
index d8bd16620..d588c888c 100644
--- a/plugins/ltac/pptactic.ml
+++ b/plugins/ltac/pptactic.ml
@@ -336,7 +336,7 @@ type 'a extra_genarg_printer =
let pr_ltac_constant kn =
if !Flags.in_debugger then KerName.print kn
else try
- pr_qualid (Nametab.shortest_qualid_of_tactic kn)
+ pr_qualid (Tacenv.shortest_qualid_of_tactic kn)
with Not_found -> (* local tactic not accessible anymore *)
str "<" ++ KerName.print kn ++ str ">"
diff --git a/plugins/ltac/pptactic.mli b/plugins/ltac/pptactic.mli
index c79d5b389..d9da954fe 100644
--- a/plugins/ltac/pptactic.mli
+++ b/plugins/ltac/pptactic.mli
@@ -93,7 +93,7 @@ val pr_alias_key : Names.KerName.t -> Pp.t
val pr_alias : (Val.t -> Pp.t) ->
int -> Names.KerName.t -> Val.t list -> Pp.t
-val pr_ltac_constant : Nametab.ltac_constant -> Pp.t
+val pr_ltac_constant : ltac_constant -> Pp.t
val pr_raw_tactic : raw_tactic_expr -> Pp.t
diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml
index a8d518fbd..bcd5b388a 100644
--- a/plugins/ltac/tacentries.ml
+++ b/plugins/ltac/tacentries.ml
@@ -409,7 +409,7 @@ let create_ltac_quotation name cast (e, l) =
type tacdef_kind =
| NewTac of Id.t
- | UpdateTac of Nametab.ltac_constant
+ | UpdateTac of Tacexpr.ltac_constant
let is_defined_tac kn =
try ignore (Tacenv.interp_ltac kn); true with Not_found -> false
@@ -441,7 +441,7 @@ let register_ltac local tacl =
| Tacexpr.TacticRedefinition (ident, body) ->
let loc = loc_of_reference ident in
let kn =
- try Nametab.locate_tactic (snd (qualid_of_reference ident))
+ try Tacenv.locate_tactic (snd (qualid_of_reference ident))
with Not_found ->
CErrors.user_err ?loc
(str "There is no Ltac named " ++ pr_reference ident ++ str ".")
@@ -464,7 +464,7 @@ let register_ltac local tacl =
let defs () =
(** Register locally the tactic to handle recursivity. This function affects
the whole environment, so that we transactify it afterwards. *)
- let iter_rec (sp, kn) = Nametab.push_tactic (Nametab.Until 1) sp kn in
+ let iter_rec (sp, kn) = Tacenv.push_tactic (Nametab.Until 1) sp kn in
let () = List.iter iter_rec recvars in
List.map map rfun
in
@@ -475,7 +475,7 @@ let register_ltac local tacl =
Flags.if_verbose Feedback.msg_info (Id.print id ++ str " is defined")
| UpdateTac kn ->
Tacenv.redefine_ltac local kn tac;
- let name = Nametab.shortest_qualid_of_tactic kn in
+ let name = Tacenv.shortest_qualid_of_tactic kn in
Flags.if_verbose Feedback.msg_info (Libnames.pr_qualid name ++ str " is redefined")
in
List.iter iter defs
@@ -488,7 +488,7 @@ let print_ltacs () =
let entries = List.sort sort entries in
let map (kn, entry) =
let qid =
- try Some (Nametab.shortest_qualid_of_tactic kn)
+ try Some (Tacenv.shortest_qualid_of_tactic kn)
with Not_found -> None
in
match qid with
@@ -506,6 +506,31 @@ let print_ltacs () =
in
Feedback.msg_notice (prlist_with_sep fnl pr_entry entries)
+let locatable_ltac = "Ltac"
+
+let () =
+ let open Prettyp in
+ let locate qid = try Some (Tacenv.locate_tactic qid) with Not_found -> None in
+ let locate_all = Tacenv.locate_extended_all_tactic in
+ let shortest_qualid = Tacenv.shortest_qualid_of_tactic in
+ let name kn = str "Ltac" ++ spc () ++ pr_path (Tacenv.path_of_tactic kn) in
+ let print kn =
+ let qid = qualid_of_path (Tacenv.path_of_tactic kn) in
+ Tacintern.print_ltac qid
+ in
+ let about = name in
+ register_locatable locatable_ltac {
+ locate;
+ locate_all;
+ shortest_qualid;
+ name;
+ print;
+ about;
+ }
+
+let print_located_tactic qid =
+ Feedback.msg_notice (Prettyp.print_located_other locatable_ltac qid)
+
(** Grammar *)
let () =
diff --git a/plugins/ltac/tacentries.mli b/plugins/ltac/tacentries.mli
index aa8f4efe6..ab2c6b307 100644
--- a/plugins/ltac/tacentries.mli
+++ b/plugins/ltac/tacentries.mli
@@ -62,3 +62,6 @@ val create_ltac_quotation : string ->
val print_ltacs : unit -> unit
(** Display the list of ltac definitions currently available. *)
+
+val print_located_tactic : Libnames.reference -> unit
+(** Display the absolute name of a tactic. *)
diff --git a/plugins/ltac/tacenv.ml b/plugins/ltac/tacenv.ml
index 13b44f0e2..8c59a36fa 100644
--- a/plugins/ltac/tacenv.ml
+++ b/plugins/ltac/tacenv.ml
@@ -11,6 +11,42 @@ open Pp
open Names
open Tacexpr
+(** Nametab for tactics *)
+
+(** TODO: Share me somewhere *)
+module FullPath =
+struct
+ open Libnames
+ type t = full_path
+ let equal = eq_full_path
+ let to_string = string_of_path
+ let repr sp =
+ let dir,id = repr_path sp in
+ id, (DirPath.repr dir)
+end
+
+module KnTab = Nametab.Make(FullPath)(KerName)
+
+let tactic_tab = Summary.ref ~name:"LTAC-NAMETAB" (KnTab.empty, KNmap.empty)
+
+let push_tactic vis sp kn =
+ let (tab, revtab) = !tactic_tab in
+ let tab = KnTab.push vis sp kn tab in
+ let revtab = KNmap.add kn sp revtab in
+ tactic_tab := (tab, revtab)
+
+let locate_tactic qid = KnTab.locate qid (fst !tactic_tab)
+
+let locate_extended_all_tactic qid = KnTab.find_prefixes qid (fst !tactic_tab)
+
+let exists_tactic kn = KnTab.exists kn (fst !tactic_tab)
+
+let path_of_tactic kn = KNmap.find kn (snd !tactic_tab)
+
+let shortest_qualid_of_tactic kn =
+ let sp = KNmap.find kn (snd !tactic_tab) in
+ KnTab.shortest_qualid Id.Set.empty sp (fst !tactic_tab)
+
(** Tactic notations (TacAlias) *)
type alias = KerName.t
@@ -103,19 +139,19 @@ let replace kn path t =
let load_md i ((sp, kn), (local, id, b, t)) = match id with
| None ->
- let () = if not local then Nametab.push_tactic (Until i) sp kn in
+ let () = if not local then push_tactic (Until i) sp kn in
add kn b t
| Some kn0 -> replace kn0 kn t
let open_md i ((sp, kn), (local, id, b, t)) = match id with
| None ->
- let () = if not local then Nametab.push_tactic (Exactly i) sp kn in
+ let () = if not local then push_tactic (Exactly i) sp kn in
add kn b t
| Some kn0 -> replace kn0 kn t
let cache_md ((sp, kn), (local, id ,b, t)) = match id with
| None ->
- let () = Nametab.push_tactic (Until 1) sp kn in
+ let () = push_tactic (Until 1) sp kn in
add kn b t
| Some kn0 -> replace kn0 kn t
@@ -128,7 +164,7 @@ let subst_md (subst, (local, id, b, t)) =
let classify_md (local, _, _, _ as o) = Substitute o
-let inMD : bool * Nametab.ltac_constant option * bool * glob_tactic_expr -> obj =
+let inMD : bool * ltac_constant option * bool * glob_tactic_expr -> obj =
declare_object {(default_object "TAC-DEFINITION") with
cache_function = cache_md;
load_function = load_md;
diff --git a/plugins/ltac/tacenv.mli b/plugins/ltac/tacenv.mli
index 958109e5a..4ecc978fe 100644
--- a/plugins/ltac/tacenv.mli
+++ b/plugins/ltac/tacenv.mli
@@ -7,11 +7,21 @@
(************************************************************************)
open Names
+open Libnames
open Tacexpr
open Geninterp
(** This module centralizes the various ways of registering tactics. *)
+(** {5 Tactic naming} *)
+
+val push_tactic : Nametab.visibility -> full_path -> ltac_constant -> unit
+val locate_tactic : qualid -> ltac_constant
+val locate_extended_all_tactic : qualid -> ltac_constant list
+val exists_tactic : full_path -> bool
+val path_of_tactic : ltac_constant -> full_path
+val shortest_qualid_of_tactic : ltac_constant -> qualid
+
(** {5 Tactic notations} *)
type alias = KerName.t
diff --git a/plugins/ltac/tacexpr.mli b/plugins/ltac/tacexpr.mli
index 64da097de..2c36faeff 100644
--- a/plugins/ltac/tacexpr.mli
+++ b/plugins/ltac/tacexpr.mli
@@ -10,13 +10,14 @@ open Loc
open Names
open Constrexpr
open Libnames
-open Nametab
open Genredexpr
open Genarg
open Pattern
open Misctypes
open Locus
+type ltac_constant = KerName.t
+
type direction_flag = bool (* true = Left-to-right false = right-to-right *)
type lazy_flag =
| General (* returns all possible successes *)
diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml
index fc6ee6aab..99d7684d3 100644
--- a/plugins/ltac/tacintern.ml
+++ b/plugins/ltac/tacintern.ml
@@ -118,7 +118,7 @@ let intern_constr_reference strict ist = function
let intern_isolated_global_tactic_reference r =
let (loc,qid) = qualid_of_reference r in
- TacCall (Loc.tag ?loc (ArgArg (loc,locate_tactic qid),[]))
+ TacCall (Loc.tag ?loc (ArgArg (loc,Tacenv.locate_tactic qid),[]))
let intern_isolated_tactic_reference strict ist r =
(* An ltac reference *)
@@ -137,7 +137,7 @@ let intern_isolated_tactic_reference strict ist r =
let intern_applied_global_tactic_reference r =
let (loc,qid) = qualid_of_reference r in
- ArgArg (loc,locate_tactic qid)
+ ArgArg (loc,Tacenv.locate_tactic qid)
let intern_applied_tactic_reference ist r =
(* An ltac reference *)
@@ -722,7 +722,7 @@ let pr_ltac_fun_arg n = spc () ++ Name.print n
let print_ltac id =
try
- let kn = Nametab.locate_tactic id in
+ let kn = Tacenv.locate_tactic id in
let entries = Tacenv.ltac_entries () in
let tac = KNmap.find kn entries in
let filter mp =
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
index cf5fdf318..d37c676e3 100644
--- a/plugins/ssr/ssrcommon.ml
+++ b/plugins/ssr/ssrcommon.ml
@@ -814,8 +814,8 @@ let ssr_n_tac seed n gl =
let name = if n = -1 then seed else ("ssr" ^ seed ^ string_of_int n) in
let fail msg = CErrors.user_err (Pp.str msg) in
let tacname =
- try Nametab.locate_tactic (Libnames.qualid_of_ident (Id.of_string name))
- with Not_found -> try Nametab.locate_tactic (ssrqid name)
+ try Tacenv.locate_tactic (Libnames.qualid_of_ident (Id.of_string name))
+ with Not_found -> try Tacenv.locate_tactic (ssrqid name)
with Not_found ->
if n = -1 then fail "The ssreflect library was not loaded"
else fail ("The tactic "^name^" was not found") in
diff --git a/plugins/ssr/ssrparser.ml4 b/plugins/ssr/ssrparser.ml4
index 1e1a986da..7b591fead 100644
--- a/plugins/ssr/ssrparser.ml4
+++ b/plugins/ssr/ssrparser.ml4
@@ -1554,8 +1554,8 @@ END
let ssrautoprop gl =
try
let tacname =
- try Nametab.locate_tactic (qualid_of_ident (Id.of_string "ssrautoprop"))
- with Not_found -> Nametab.locate_tactic (ssrqid "ssrautoprop") in
+ try Tacenv.locate_tactic (qualid_of_ident (Id.of_string "ssrautoprop"))
+ with Not_found -> Tacenv.locate_tactic (ssrqid "ssrautoprop") in
let tacexpr = Loc.tag @@ Tacexpr.Reference (ArgArg (Loc.tag @@ tacname)) in
Proofview.V82.of_tactic (eval_tactic (Tacexpr.TacArg tacexpr)) gl
with Not_found -> Proofview.V82.of_tactic (Auto.full_trivial []) gl
diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml
index 39c2ceeba..1038945c1 100644
--- a/pretyping/nativenorm.ml
+++ b/pretyping/nativenorm.ml
@@ -5,13 +5,11 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Pp
open CErrors
open Term
open Vars
open Environ
open Reduction
-open Univ
open Declarations
open Names
open Inductive
@@ -20,8 +18,6 @@ open Nativecode
open Nativevalues
open Context.Rel.Declaration
-module RelDecl = Context.Rel.Declaration
-
(** This module implements normalization by evaluation to OCaml code *)
exception Find_at of int
@@ -177,44 +173,6 @@ let build_case_type dep p realargs c =
if dep then mkApp(mkApp(p, realargs), [|c|])
else mkApp(p, realargs)
-(* TODO move this function *)
-let type_of_rel env n =
- env |> lookup_rel n |> RelDecl.get_type |> lift n
-
-let type_of_prop = mkSort type1_sort
-
-let type_of_sort s =
- match s with
- | Prop _ -> type_of_prop
- | Type u -> mkType (Univ.super u)
-
-let type_of_var env id =
- let open Context.Named.Declaration in
- try env |> lookup_named id |> get_type
- with Not_found ->
- anomaly ~label:"type_of_var" (str "variable " ++ Id.print id ++ str " unbound.")
-
-let sort_of_product env domsort rangsort =
- match (domsort, rangsort) with
- (* Product rule (s,Prop,Prop) *)
- | (_, Prop Null) -> rangsort
- (* Product rule (Prop/Set,Set,Set) *)
- | (Prop _, Prop Pos) -> rangsort
- (* Product rule (Type,Set,?) *)
- | (Type u1, Prop Pos) ->
- if is_impredicative_set env then
- (* Rule is (Type,Set,Set) in the Set-impredicative calculus *)
- rangsort
- else
- (* Rule is (Type_i,Set,Type_i) in the Set-predicative calculus *)
- Type (sup u1 type0_univ)
- (* Product rule (Prop,Type_i,Type_i) *)
- | (Prop Pos, Type u2) -> Type (sup type0_univ u2)
- (* Product rule (Prop,Type_i,Type_i) *)
- | (Prop Null, Type _) -> rangsort
- (* Product rule (Type_i,Type_i,Type_i) *)
- | (Type u1, Type u2) -> Type (sup u1 u2)
-
(* normalisation of values *)
let branch_of_switch lvl ans bs =
@@ -328,15 +286,15 @@ and nf_atom_type env sigma atom =
match atom with
| Arel i ->
let n = (nb_rel env - i) in
- mkRel n, type_of_rel env n
+ mkRel n, Typeops.type_of_relative env n
| Aconstant cst ->
mkConstU cst, Typeops.type_of_constant_in env cst
| Aind ind ->
mkIndU ind, Inductiveops.type_of_inductive env ind
| Asort s ->
- mkSort s, type_of_sort s
+ mkSort s, Typeops.type_of_sort s
| Avar id ->
- mkVar id, type_of_var env id
+ mkVar id, Typeops.type_of_variable env id
| Acase(ans,accu,p,bs) ->
let a,ta = nf_accu_type env sigma accu in
let ((mind,_),u as ind),allargs = find_rectype_a env ta in
@@ -387,7 +345,7 @@ and nf_atom_type env sigma atom =
let vn = mk_rel_accu (nb_rel env) in
let env = push_rel (LocalAssum (n,dom)) env in
let codom,s2 = nf_type_sort env sigma (codom vn) in
- mkProd(n,dom,codom), mkSort (sort_of_product env s1 s2)
+ mkProd(n,dom,codom), Typeops.type_of_product env n s1 s2
| Aevar(ev,ty) ->
let ty = nf_type env sigma ty in
mkEvar ev, ty
diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml
index 079524f34..56f8b33e0 100644
--- a/pretyping/retyping.ml
+++ b/pretyping/retyping.ml
@@ -214,7 +214,6 @@ let type_of_global_reference_knowing_conclusion env sigma c conclty =
type_of_inductive_knowing_conclusion env sigma (spec, EInstance.kind sigma u) conclty
| Const (cst, u) ->
let t = constant_type_in env (cst, EInstance.kind sigma u) in
- (* TODO *)
sigma, EConstr.of_constr t
| Var id -> sigma, type_of_var env id
| Construct (cstr, u) -> sigma, EConstr.of_constr (type_of_constructor env (cstr, EInstance.kind sigma u))
diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml
index 10dd42ea9..a1cdfdbaa 100644
--- a/printing/ppvernac.ml
+++ b/printing/ppvernac.ml
@@ -1156,7 +1156,7 @@ open Decl_kinds
| LocateFile f -> keyword "File" ++ spc() ++ qs f
| LocateLibrary qid -> keyword "Library" ++ spc () ++ pr_module qid
| LocateModule qid -> keyword "Module" ++ spc () ++ pr_module qid
- | LocateTactic qid -> keyword "Ltac" ++ spc () ++ pr_ltac_ref qid
+ | LocateOther (s, qid) -> keyword s ++ spc () ++ pr_ltac_ref qid
in
return (keyword "Locate" ++ spc() ++ pr_locate loc)
| VernacRegister (id, RegisterInline) ->
diff --git a/printing/prettyp.ml b/printing/prettyp.ml
index 09859157c..2077526db 100644
--- a/printing/prettyp.ml
+++ b/printing/prettyp.ml
@@ -304,14 +304,33 @@ let print_inductive_argument_scopes =
(*********************)
(* "Locate" commands *)
+type 'a locatable_info = {
+ locate : qualid -> 'a option;
+ locate_all : qualid -> 'a list;
+ shortest_qualid : 'a -> qualid;
+ name : 'a -> Pp.t;
+ print : 'a -> Pp.t;
+ about : 'a -> Pp.t;
+}
+
+type locatable = Locatable : 'a locatable_info -> locatable
+
type logical_name =
| Term of global_reference
| Dir of global_dir_reference
| Syntactic of kernel_name
| ModuleType of module_path
- | Tactic of Nametab.ltac_constant
+ | Other : 'a * 'a locatable_info -> logical_name
| Undefined of qualid
+(** Generic table for objects that are accessible through a name. *)
+let locatable_map : locatable String.Map.t ref = ref String.Map.empty
+
+let register_locatable name f =
+ locatable_map := String.Map.add name (Locatable f) !locatable_map
+
+exception ObjFound of logical_name
+
let locate_any_name ref =
let (loc,qid) = qualid_of_reference ref in
try Term (Nametab.locate qid)
@@ -321,7 +340,13 @@ let locate_any_name ref =
try Dir (Nametab.locate_dir qid)
with Not_found ->
try ModuleType (Nametab.locate_modtype qid)
- with Not_found -> Undefined qid
+ with Not_found ->
+ let iter _ (Locatable info) = match info.locate qid with
+ | None -> ()
+ | Some ans -> raise (ObjFound (Other (ans, info)))
+ in
+ try String.Map.iter iter !locatable_map; Undefined qid
+ with ObjFound obj -> obj
let pr_located_qualid = function
| Term ref ->
@@ -344,8 +369,7 @@ let pr_located_qualid = function
str s ++ spc () ++ pr_dirpath dir
| ModuleType mp ->
str "Module Type" ++ spc () ++ pr_path (Nametab.path_of_modtype mp)
- | Tactic kn ->
- str "Ltac" ++ spc () ++ pr_path (Nametab.path_of_tactic kn)
+ | Other (obj, info) -> info.name obj
| Undefined qid ->
pr_qualid qid ++ spc () ++ str "not a defined object."
@@ -383,10 +407,6 @@ let locate_term qid =
in
List.map expand (Nametab.locate_extended_all qid)
-let locate_tactic qid =
- let all = Nametab.locate_extended_all_tactic qid in
- List.map (fun kn -> (Tactic kn, Nametab.shortest_qualid_of_tactic kn)) all
-
let locate_module qid =
let all = Nametab.locate_extended_all_dir qid in
let map dir = match dir with
@@ -408,13 +428,30 @@ let locate_modtype qid =
in
modtypes @ List.map_filter map all
+let locate_other s qid =
+ let Locatable info = String.Map.find s !locatable_map in
+ let ans = info.locate_all qid in
+ let map obj = (Other (obj, info), info.shortest_qualid obj) in
+ List.map map ans
+
+type locatable_kind =
+| LocTerm
+| LocModule
+| LocOther of string
+| LocAny
+
let print_located_qualid name flags ref =
let (loc,qid) = qualid_of_reference ref in
- let located = [] in
- let located = if List.mem `LTAC flags then locate_tactic qid @ located else located in
- let located = if List.mem `MODTYPE flags then locate_modtype qid @ located else located in
- let located = if List.mem `MODULE flags then locate_module qid @ located else located in
- let located = if List.mem `TERM flags then locate_term qid @ located else located in
+ let located = match flags with
+ | LocTerm -> locate_term qid
+ | LocModule -> locate_modtype qid @ locate_module qid
+ | LocOther s -> locate_other s qid
+ | LocAny ->
+ locate_term qid @
+ locate_modtype qid @
+ locate_module qid @
+ String.Map.fold (fun s _ accu -> locate_other s qid @ accu) !locatable_map []
+ in
match located with
| [] ->
let (dir,id) = repr_qualid qid in
@@ -432,10 +469,10 @@ let print_located_qualid name flags ref =
else mt ()) ++
display_alias o)) l
-let print_located_term ref = print_located_qualid "term" [`TERM] ref
-let print_located_tactic ref = print_located_qualid "tactic" [`LTAC] ref
-let print_located_module ref = print_located_qualid "module" [`MODULE; `MODTYPE] ref
-let print_located_qualid ref = print_located_qualid "object" [`TERM; `LTAC; `MODULE; `MODTYPE] ref
+let print_located_term ref = print_located_qualid "term" LocTerm ref
+let print_located_other s ref = print_located_qualid s (LocOther s) ref
+let print_located_module ref = print_located_qualid "module" LocModule ref
+let print_located_qualid ref = print_located_qualid "object" LocAny ref
(******************************************)
(**** Printing declarations and judgments *)
@@ -765,7 +802,7 @@ let print_any_name = function
| Dir (DirModule(dirpath,(mp,_))) -> print_module (printable_body dirpath) mp
| Dir _ -> mt ()
| ModuleType mp -> print_modtype mp
- | Tactic kn -> mt () (** TODO *)
+ | Other (obj, info) -> info.print obj
| Undefined qid ->
try (* Var locale de but, pas var de section... donc pas d'implicits *)
let dir,str = repr_qualid qid in
@@ -822,8 +859,9 @@ let print_about_any ?loc k =
v 0 (
print_syntactic_def kn ++ fnl () ++
hov 0 (str "Expands to: " ++ pr_located_qualid k))
- | Dir _ | ModuleType _ | Tactic _ | Undefined _ ->
+ | Dir _ | ModuleType _ | Undefined _ ->
hov 0 (pr_located_qualid k)
+ | Other (obj, info) -> hov 0 (info.about obj)
let print_about = function
| ByNotation (loc,(ntn,sc)) ->
diff --git a/printing/prettyp.mli b/printing/prettyp.mli
index f4277b6c5..dbd101159 100644
--- a/printing/prettyp.mli
+++ b/printing/prettyp.mli
@@ -50,12 +50,34 @@ val print_all_instances : unit -> Pp.t
val inspect : int -> Pp.t
-(** Locate *)
+(** {5 Locate} *)
+
+type 'a locatable_info = {
+ locate : qualid -> 'a option;
+ (** Locate the most precise object with the provided name if any. *)
+ locate_all : qualid -> 'a list;
+ (** Locate all objects whose name is a suffix of the provided name *)
+ shortest_qualid : 'a -> qualid;
+ (** Return the shortest name in the current context *)
+ name : 'a -> Pp.t;
+ (** Data as printed by the Locate command *)
+ print : 'a -> Pp.t;
+ (** Data as printed by the Print command *)
+ about : 'a -> Pp.t;
+ (** Data as printed by the About command *)
+}
+(** Generic data structure representing locatable objects. *)
+
+val register_locatable : string -> 'a locatable_info -> unit
+(** Define a new type of locatable objects that can be reached via the
+ corresponding generic vernacular commands. The string should be a unique
+ name describing the kind of objects considered and that is added as a
+ grammar command prefix for vernacular commands Locate. *)
val print_located_qualid : reference -> Pp.t
val print_located_term : reference -> Pp.t
-val print_located_tactic : reference -> Pp.t
val print_located_module : reference -> Pp.t
+val print_located_other : string -> reference -> Pp.t
type object_pr = {
print_inductive : mutual_inductive -> Pp.t;
diff --git a/proofs/proof_bullet.ml b/proofs/proof_bullet.ml
index f80cb7cc6..4f575ab4b 100644
--- a/proofs/proof_bullet.ml
+++ b/proofs/proof_bullet.ml
@@ -110,10 +110,6 @@ module Strict = struct
let push (b:t) pr =
focus bullet_cond (b::get_bullets pr) 1 pr
- (* Used only in the next function.
- TODO: use a recursive function instead? *)
- exception SuggestFound of t
-
let suggest_bullet (prf : proof): suggestion =
if is_done prf then ProofFinished
else if not (no_focused_goal prf)
@@ -122,24 +118,24 @@ module Strict = struct
| b::_ -> Unfinished b
| _ -> NoBulletInUse
else (* There is no goal under focus but some are unfocussed,
- let us look at the bullet needed. If no *)
- let pcobaye = ref prf in
- try
- while true do
- let pcobaye', b = pop !pcobaye in
- (* pop went well, this means that there are no more goals
- *under this* bullet b, see if a new b can be pushed. *)
- (try let _ = push b pcobaye' in (* push didn't fail so a new b can be pushed. *)
- raise (SuggestFound b)
- with SuggestFound _ as e -> raise e
- | _ -> ()); (* b could not be pushed, so we must look for a outer bullet *)
- pcobaye := pcobaye'
- done;
- assert false
- with SuggestFound b -> Suggest b
- | _ -> NeedClosingBrace (* No push was possible, but there are still
- subgoals somewhere: there must be a "}" to use. *)
-
+ let us look at the bullet needed. *)
+ let rec loop prf =
+ match pop prf with
+ | prf, b ->
+ (* pop went well, this means that there are no more goals
+ *under this* bullet b, see if a new b can be pushed. *)
+ begin
+ try ignore (push b prf); Suggest b
+ with _ ->
+ (* b could not be pushed, so we must look for a outer bullet *)
+ loop prf
+ end
+ | exception _ ->
+ (* No pop was possible, but there are still
+ subgoals somewhere: there must be a "}" to use. *)
+ NeedClosingBrace
+ in
+ loop prf
let rec pop_until (prf : proof) bul : proof =
let prf', b = pop prf in
diff --git a/stm/stm.ml b/stm/stm.ml
index 984a87429..2c5e9ed81 100644
--- a/stm/stm.ml
+++ b/stm/stm.ml
@@ -64,11 +64,6 @@ let call_process_error_once =
end
-(* During interactive use we cache more states so that Undoing is fast *)
-let interactive () =
- if !Flags.ide_slave || not !Flags.batch_mode then `Yes
- else `No
-
let async_proofs_workers_extra_env = ref [||]
type aast = {
@@ -253,6 +248,12 @@ end (* }}} *)
(*************************** THE DOCUMENT *************************************)
(******************************************************************************)
+(* The main document type associated to a VCS *)
+type stm_doc_type =
+ | VoDoc of Names.DirPath.t
+ | VioDoc of Names.DirPath.t
+ | Interactive of Names.DirPath.t
+
(* Imperative wrap around VCS to obtain _the_ VCS that is the
* representation of the document Coq is currently processing *)
module VCS : sig
@@ -269,7 +270,10 @@ module VCS : sig
type vcs = (branch_type, transaction, vcs state_info, box) Vcs_.t
- val init : id -> unit
+ val init : stm_doc_type -> id -> unit
+ (* val get_type : unit -> stm_doc_type *)
+ val is_interactive : unit -> [`Yes | `No | `Shallow]
+ val is_vio_doc : unit -> bool
val current_branch : unit -> Branch.t
val checkout : Branch.t -> unit
@@ -451,10 +455,25 @@ end = struct (* {{{ *)
type vcs = (branch_type, transaction, vcs state_info, box) t
let vcs : vcs ref = ref (empty Stateid.dummy)
- let init id =
+ let doc_type = ref (Interactive (Names.DirPath.make []))
+
+ let init dt id =
+ doc_type := dt;
vcs := empty id;
vcs := set_info !vcs id (default_info ())
+ (* let get_type () = !doc_type *)
+
+ let is_interactive () =
+ match !doc_type with
+ | Interactive _ -> `Yes
+ | _ -> `No
+
+ let is_vio_doc () =
+ match !doc_type with
+ | VioDoc _ -> true
+ | _ -> false
+
let current_branch () = current_branch !vcs
let checkout head = vcs := checkout !vcs head
@@ -765,7 +784,7 @@ end = struct (* {{{ *)
| { state = Error ie } -> cur_id := id; Exninfo.iraise ie
| _ ->
(* coqc has a 1 slot cache and only for valid states *)
- if interactive () = `No && Stateid.equal id !cur_id then ()
+ if VCS.is_interactive () = `No && Stateid.equal id !cur_id then ()
else anomaly Pp.(str "installing a non cached state.")
let get_cached id =
@@ -1054,7 +1073,7 @@ end = struct (* {{{ *)
" the \"-async-proofs-cache force\" option to Coq."))
let undo_vernac_classifier v =
- if !Flags.batch_mode && !Flags.async_proofs_cache <> Some Flags.Force
+ if VCS.is_interactive () = `No && !Flags.async_proofs_cache <> Some Flags.Force
then undo_costly_in_batch_mode v;
try
match v with
@@ -1106,7 +1125,7 @@ end = struct (* {{{ *)
VtBack (true, oid), VtLater
| VernacBacktrack (id,_,_)
| VernacBackTo id ->
- VtBack (not !Flags.batch_mode, Stateid.of_int id), VtNow
+ VtBack (VCS.is_interactive () = `Yes, Stateid.of_int id), VtNow
| _ -> VtUnknown, VtNow
with
| Not_found ->
@@ -1364,7 +1383,7 @@ end = struct (* {{{ *)
let build_proof_here ?loc ~drop_pt (id,valid) eop =
Future.create (State.exn_on id ~valid) (fun () ->
let wall_clock1 = Unix.gettimeofday () in
- if !Flags.batch_mode then Reach.known_state ~cache:`No eop
+ if VCS.is_interactive () = `No then Reach.known_state ~cache:`No eop
else Reach.known_state ~cache:`Shallow eop;
let wall_clock2 = Unix.gettimeofday () in
Aux_file.record_in_aux_at ?loc "proof_build_time"
@@ -1501,7 +1520,6 @@ end = struct (* {{{ *)
module TaskQueue = AsyncTaskQueue.MakeQueue(ProofTask)
let queue = ref None
-
let init () =
if Flags.async_proofs_is_master () then
queue := Some (TaskQueue.create !Flags.async_proofs_n_workers)
@@ -1637,7 +1655,7 @@ end = struct (* {{{ *)
let id, valid as t_exn_info = exn_info in
let cancel_switch = ref false in
if TaskQueue.n_workers (Option.get !queue) = 0 then
- if !Flags.compilation_mode = Flags.BuildVio then begin
+ if VCS.is_vio_doc () then begin
let f,assign =
Future.create_delegate ~blocking:true ~name:pname (State.exn_on id ~valid) in
let t_uuid = Future.uuid f in
@@ -1962,14 +1980,14 @@ let pstate = summary_pstate
let async_policy () =
let open Flags in
if is_universe_polymorphism () then false
- else if interactive () = `Yes then
+ else if VCS.is_interactive () = `Yes then
(async_proofs_is_master () || !async_proofs_mode = APonLazy)
else
- (!compilation_mode = BuildVio || !async_proofs_mode <> APoff)
+ (VCS.is_vio_doc () || !async_proofs_mode <> APoff)
let delegate name =
get_hint_bp_time name >= !Flags.async_proofs_delegation_threshold
- || !Flags.compilation_mode = Flags.BuildVio
+ || VCS.is_vio_doc ()
|| !Flags.async_proofs_full
let warn_deprecated_nested_proofs =
@@ -2030,7 +2048,7 @@ let collect_proof keep cur hd brkind id =
`ASync (parent last,accn,name,delegate name)
| `Fork((_, hd', GuaranteesOpacity, ids), _) when
has_proof_no_using last && not (State.is_cached_and_valid (parent last)) &&
- !Flags.compilation_mode = Flags.BuildVio ->
+ VCS.is_vio_doc () ->
assert (VCS.Branch.equal hd hd'||VCS.Branch.equal hd VCS.edit_branch);
(try
let name, hint = name ids, get_hint_ctx loc in
@@ -2339,9 +2357,20 @@ end (* }}} *)
(********************************* STM API ************************************)
(******************************************************************************)
-let init () =
- VCS.init Stateid.initial;
+type stm_init_options = {
+ doc_type : stm_doc_type;
+(*
+ fb_handler : Feedback.feedback -> unit;
+ iload_path : (string list * string * bool) list;
+ require_libs : (Names.DirPath.t * string * bool option) list;
+ implicit_std : bool;
+*)
+}
+
+let init { doc_type } =
+ VCS.init doc_type Stateid.initial;
set_undo_classifier Backtrack.undo_vernac_classifier;
+ (* we declare the library here XXX: *)
State.define ~cache:`Yes (fun () -> ()) Stateid.initial;
Backtrack.record ();
Slaves.init ();
@@ -2362,7 +2391,7 @@ let init () =
let observe id =
let vcs = VCS.backup () in
try
- Reach.known_state ~cache:(interactive ()) id;
+ Reach.known_state ~cache:(VCS.is_interactive ()) id;
VCS.print ()
with e ->
let e = CErrors.push e in
@@ -2518,7 +2547,7 @@ let process_transaction ?(newtip=Stateid.fresh ())
stm_prerr_endline (fun () -> "undo to state " ^ Stateid.to_string id);
Backtrack.backto id;
VCS.checkout_shallowest_proof_branch ();
- Reach.known_state ~cache:(interactive ()) id; `Ok
+ Reach.known_state ~cache:(VCS.is_interactive ()) id; `Ok
| VtBack (false, id), VtLater ->
anomaly(str"classifier: VtBack + VtLater must imply part_of_script.")
@@ -2536,7 +2565,7 @@ let process_transaction ?(newtip=Stateid.fresh ())
let id = VCS.new_node ~id:newtip () in
let queue =
if !Flags.async_proofs_full then `QueryQueue (ref false)
- else if Flags.(!compilation_mode = BuildVio) &&
+ else if VCS.is_vio_doc () &&
VCS.((get_branch head).kind = `Master) &&
may_pierce_opaque x
then `SkipQueue
@@ -2626,7 +2655,7 @@ let process_transaction ?(newtip=Stateid.fresh ())
let step () =
VCS.checkout VCS.Branch.master;
let mid = VCS.get_branch_pos VCS.Branch.master in
- Reach.known_state ~cache:(interactive ()) mid;
+ Reach.known_state ~cache:(VCS.is_interactive ()) mid;
stm_vernac_interp id x;
(* Vernac x may or may not start a proof *)
if not in_proof && Proof_global.there_are_pending_proofs () then
@@ -2834,7 +2863,7 @@ let edit_at id =
VCS.edit_branch (`Edit (mode, qed_id, master_id, keep, old_branch));
VCS.delete_boxes_of id;
cancel_switch := true;
- Reach.known_state ~cache:(interactive ()) id;
+ Reach.known_state ~cache:(VCS.is_interactive ()) id;
VCS.checkout_shallowest_proof_branch ();
`Focus { stop = qed_id; start = master_id; tip } in
let no_edit = function
@@ -2857,7 +2886,7 @@ let edit_at id =
VCS.gc ();
VCS.print ();
if not !Flags.async_proofs_full then
- Reach.known_state ~cache:(interactive ()) id;
+ Reach.known_state ~cache:(VCS.is_interactive ()) id;
VCS.checkout_shallowest_proof_branch ();
`NewTip in
try
@@ -2886,7 +2915,7 @@ let edit_at id =
| true, None, _ ->
if on_cur_branch id then begin
VCS.reset_branch (VCS.current_branch ()) id;
- Reach.known_state ~cache:(interactive ()) id;
+ Reach.known_state ~cache:(VCS.is_interactive ()) id;
VCS.checkout_shallowest_proof_branch ();
`NewTip
end else if is_ancestor_of_cur_branch id then begin
diff --git a/stm/stm.mli b/stm/stm.mli
index 3f01fca01..ea8aecaed 100644
--- a/stm/stm.mli
+++ b/stm/stm.mli
@@ -10,6 +10,26 @@ open Names
(** state-transaction-machine interface *)
+(** The STM doc type determines some properties such as what
+ uncompleted proofs are allowed and recording of aux files. *)
+type stm_doc_type =
+ | VoDoc of DirPath.t
+ | VioDoc of DirPath.t
+ | Interactive of DirPath.t
+
+(* Main initalization routine *)
+type stm_init_options = {
+ doc_type : stm_doc_type;
+(*
+ fb_handler : Feedback.feedback -> unit;
+ iload_path : (string list * string * bool) list;
+ require_libs : (Names.DirPath.t * string * bool option) list;
+ implicit_std : bool;
+*)
+}
+
+val init : stm_init_options -> unit
+
(* [parse_sentence sid pa] Reads a sentence from [pa] with parsing
state [sid] Returns [End_of_input] if the stream ends *)
val parse_sentence : Stateid.t -> Pcoq.Gram.coq_parsable ->
@@ -83,9 +103,6 @@ val finish_tasks : string ->
(* Id of the tip of the current branch *)
val get_current_state : unit -> Stateid.t
-(* Misc *)
-val init : unit -> unit
-
(* This returns the node at that position *)
val get_ast : Stateid.t -> (Vernacexpr.vernac_expr Loc.located) option
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index d6c24e9cc..6f7e61f6b 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -1579,11 +1579,11 @@ let elimination_in_clause_scheme with_evars ?(flags=elim_flags ())
let elimclause = make_clenv_binding env sigma (elim, elimty) bindings in
let indmv = destMeta sigma (nth_arg sigma i elimclause.templval.rebus) in
let hypmv =
- try match List.remove Int.equal indmv (clenv_independent elimclause) with
- | [a] -> a
- | _ -> failwith ""
- with Failure _ -> user_err ~hdr:"elimination_clause"
- (str "The type of elimination clause is not well-formed.") in
+ match List.remove Int.equal indmv (clenv_independent elimclause) with
+ | [a] -> a
+ | _ -> user_err ~hdr:"elimination_clause"
+ (str "The type of elimination clause is not well-formed.")
+ in
let elimclause' = clenv_fchain ~flags indmv elimclause indclause in
let hyp = mkVar id in
let hyp_typ = Retyping.get_type_of env sigma hyp in
diff --git a/test-suite/bugs/closed/4852.v b/test-suite/bugs/closed/4852.v
new file mode 100644
index 000000000..5068ed9b9
--- /dev/null
+++ b/test-suite/bugs/closed/4852.v
@@ -0,0 +1,54 @@
+(** BZ 4852 : unsatisfactory Extraction Implicit for a fixpoint defined via wf *)
+
+Require Import Coq.Lists.List.
+Import ListNotations.
+Require Import Omega.
+
+Definition wfi_lt := well_founded_induction_type Wf_nat.lt_wf.
+
+Tactic Notation "wfinduction" constr(term) "on" ne_hyp_list(Hs) "as" ident(H) :=
+ let R := fresh in
+ let E := fresh in
+ remember term as R eqn:E;
+ revert E; revert Hs;
+ induction R as [R H] using wfi_lt;
+ intros; subst R.
+
+Hint Rewrite @app_comm_cons @app_assoc @app_length : app_rws.
+
+Ltac solve_nat := autorewrite with app_rws in *; cbn in *; omega.
+
+Notation "| x |" := (length x) (at level 11, no associativity, format "'|' x '|'").
+
+Definition split_acc (ls : list nat) : forall acc1 acc2,
+ (|acc1| = |acc2| \/ |acc1| = S (|acc2|)) ->
+ { lss : list nat * list nat |
+ let '(ls1, ls2) := lss in |ls1++ls2| = |ls++acc1++acc2| /\ (|ls1| = |ls2| \/ |ls1| = S (|ls2|))}.
+Proof.
+ induction ls as [|a ls IHls]. all:intros acc1 acc2 H.
+ { exists (acc1, acc2). cbn. intuition reflexivity. }
+ destruct (IHls (a::acc2) acc1) as [[ls1 ls2] (H1 & H2)]. 1:solve_nat.
+ exists (ls1, ls2). cbn. intuition solve_nat.
+Defined.
+
+Definition join(ls : list nat) : { rls : list nat | |rls| = |ls| }.
+Proof.
+ wfinduction (|ls|) on ls as IH.
+ case (split_acc ls [] []). 1:solve_nat.
+ intros (ls1 & ls2) (H1 & H2).
+ destruct ls2 as [|a ls2].
+ - exists ls1. solve_nat.
+ - unshelve eelim (IH _ _ ls1 eq_refl). 1:solve_nat. intros rls1 H3.
+ unshelve eelim (IH _ _ ls2 eq_refl). 1:solve_nat. intros rls2 H4.
+ exists (a :: rls1 ++ rls2). solve_nat.
+Defined.
+
+Require Import ExtrOcamlNatInt.
+Extract Inlined Constant length => "List.length".
+Extract Inlined Constant app => "List.append".
+
+Extraction Inline wfi_lt.
+Extraction Implicit wfi_lt [1 3].
+Recursive Extraction join. (* was: Error: An implicit occurs after extraction *)
+Extraction TestCompile join.
+
diff --git a/test-suite/bugs/closed/5765.v b/test-suite/bugs/closed/5765.v
new file mode 100644
index 000000000..343ab4935
--- /dev/null
+++ b/test-suite/bugs/closed/5765.v
@@ -0,0 +1,3 @@
+(* 'pat binder not (yet?) allowed in parameters of inductive types *)
+
+Fail Inductive X '(a,b) := x.
diff --git a/test-suite/bugs/closed/5769.v b/test-suite/bugs/closed/5769.v
new file mode 100644
index 000000000..42573aad8
--- /dev/null
+++ b/test-suite/bugs/closed/5769.v
@@ -0,0 +1,20 @@
+(* Check a few naming heuristics based on types *)
+(* was buggy for types names _something *)
+
+Inductive _foo :=.
+Lemma bob : (sigT (fun x : nat => _foo)) -> _foo.
+destruct 1.
+exact _f.
+Abort.
+
+Inductive _'Foo :=.
+Lemma bob : (sigT (fun x : nat => _'Foo)) -> _'Foo.
+destruct 1.
+exact _'f.
+Abort.
+
+Inductive ____ :=.
+Lemma bob : (sigT (fun x : nat => ____)) -> ____.
+destruct 1.
+exact x0.
+Abort.
diff --git a/test-suite/output/Notations3.out b/test-suite/output/Notations3.out
index 65efe228a..6ef75dd13 100644
--- a/test-suite/output/Notations3.out
+++ b/test-suite/output/Notations3.out
@@ -124,3 +124,7 @@ return (1, 2, 3, 4)
: nat
!!! _ _ : nat, True
: (nat -> Prop) * ((nat -> Prop) * Prop)
+((*1).2).3
+ : nat
+*(1.2)
+ : nat
diff --git a/test-suite/output/Notations3.v b/test-suite/output/Notations3.v
index ea372ad1a..8c7bbe591 100644
--- a/test-suite/output/Notations3.v
+++ b/test-suite/output/Notations3.v
@@ -190,3 +190,10 @@ Check 1+1+1.
(* Test presence of notation variables in the recursive parts (introduced in dfdaf4de) *)
Notation "!!! x .. y , b" := ((fun x => b), .. ((fun y => b), True) ..) (at level 200, x binder).
Check !!! (x y:nat), True.
+
+(* Allow level for leftmost nonterminal when printing-only, BZ#5739 *)
+
+Notation "* x" := (id x) (only printing, at level 15, format "* x").
+Notation "x . y" := (x + y) (only printing, at level 20, x at level 14, left associativity, format "x . y").
+Check (((id 1) + 2) + 3).
+Check (id (1 + 2)).
diff --git a/tools/TimeFileMaker.py b/tools/TimeFileMaker.py
index a207c2171..7298ef5e8 100644
--- a/tools/TimeFileMaker.py
+++ b/tools/TimeFileMaker.py
@@ -28,10 +28,10 @@ def get_times(file_name):
else:
with open(file_name, 'r') as f:
lines = f.read()
- reg = re.compile(r'^([^\s]+) \([^\)]*?user: ([0-9\.]+)[^\)]*?\)$', re.MULTILINE)
+ reg = re.compile(r'^([^\s]+) \([^\)]*?user: ([0-9\.]+)[^\)]*?\)\s*$', re.MULTILINE)
times = reg.findall(lines)
if all(time in ('0.00', '0.01') for name, time in times):
- reg = re.compile(r'^([^\s]*) \([^\)]*?real: ([0-9\.]+)[^\)]*?\)$', re.MULTILINE)
+ reg = re.compile(r'^([^\s]*) \([^\)]*?real: ([0-9\.]+)[^\)]*?\)\s*$', re.MULTILINE)
times = reg.findall(lines)
if all(STRIP_REG.search(name.strip()) for name, time in times):
times = tuple((STRIP_REG.sub(STRIP_REP, name.strip()), time) for name, time in times)
diff --git a/toplevel/coqinit.ml b/toplevel/coqinit.ml
index 5ca886965..8f676c656 100644
--- a/toplevel/coqinit.ml
+++ b/toplevel/coqinit.ml
@@ -32,7 +32,7 @@ let load_rcfile sid =
try
if !rcfile_specified then
if CUnix.file_readable_p !rcfile then
- Vernac.load_vernac false sid !rcfile
+ Vernac.load_vernac ~verbosely:false ~interactive:false ~check:true sid !rcfile
else raise (Sys_error ("Cannot read rcfile: "^ !rcfile))
else
try
@@ -43,7 +43,7 @@ let load_rcfile sid =
Envars.home ~warn / "."^rcdefaultname^"."^Coq_config.version;
Envars.home ~warn / "."^rcdefaultname
] in
- Vernac.load_vernac false sid inferedrc
+ Vernac.load_vernac ~verbosely:false ~interactive:false ~check:true sid inferedrc
with Not_found -> sid
(*
Flags.if_verbose
@@ -114,9 +114,6 @@ let init_load_path () =
(* additional ml directories, given with option -I *)
List.iter Mltop.add_ml_dir (List.rev !ml_includes)
-let init_library_roots () =
- includes := []
-
(* Initialises the Ocaml toplevel before launching it, so that it can
find the "include" file in the *source* directory *)
let init_ocaml_path () =
diff --git a/toplevel/coqinit.mli b/toplevel/coqinit.mli
index 3432e79cc..1f7fd6ed9 100644
--- a/toplevel/coqinit.mli
+++ b/toplevel/coqinit.mli
@@ -21,6 +21,5 @@ val push_include : string -> Names.DirPath.t -> bool -> unit
val push_ml_include : string -> unit
val init_load_path : unit -> unit
-val init_library_roots : unit -> unit
val init_ocaml_path : unit -> unit
diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml
index db7bcb76b..0624c9bed 100644
--- a/toplevel/coqtop.ml
+++ b/toplevel/coqtop.ml
@@ -9,7 +9,6 @@
open Pp
open CErrors
open Libnames
-open Coqinit
let () = at_exit flush_all
@@ -121,6 +120,14 @@ let print_memory_stat () =
let _ = at_exit print_memory_stat
+(******************************************************************************)
+(* Deprecated *)
+(******************************************************************************)
+let _remove_top_ml () = Mltop.remove ()
+
+(******************************************************************************)
+(* Engagement *)
+(******************************************************************************)
let impredicative_set = ref Declarations.PredicativeSet
let set_impredicative_set c = impredicative_set := Declarations.ImpredicativeSet
let set_type_in_type () =
@@ -129,14 +136,18 @@ let set_type_in_type () =
let engage () =
Global.set_engagement !impredicative_set
-let set_batch_mode () = Flags.batch_mode := true
-
+(******************************************************************************)
+(* Interactive toplevel name *)
+(******************************************************************************)
let toplevel_default_name = Names.(DirPath.make [Id.of_string "Top"])
let toplevel_name = ref toplevel_default_name
let set_toplevel_name dir =
if Names.DirPath.is_empty dir then user_err Pp.(str "Need a non empty toplevel module name");
toplevel_name := dir
+(******************************************************************************)
+(* Input/Output State *)
+(******************************************************************************)
let warn_deprecated_inputstate =
CWarnings.create ~name:"deprecated-inputstate" ~category:"deprecated"
(fun () -> strbrk "The inputstate option is deprecated and discouraged.")
@@ -164,21 +175,22 @@ let outputstate () =
let fname = CUnix.make_suffix !outputstate ".coq" in
States.extern_state fname
-let set_include d p implicit =
- let p = dirpath_of_string p in
- push_include d p implicit
-
+(******************************************************************************)
+(* Interactive Load File Simulation *)
+(******************************************************************************)
let load_vernacular_list = ref ([] : (string * bool) list)
+
let add_load_vernacular verb s =
load_vernacular_list := ((CUnix.make_suffix s ".v"),verb) :: !load_vernacular_list
+
let load_vernacular sid =
List.fold_left
- (fun sid (s,v) ->
- let s = Loadpath.locate_file s in
+ (fun sid (f_in, verbosely) ->
+ let s = Loadpath.locate_file f_in in
if !Flags.beautify then
- Flags.(with_option beautify_file (Vernac.load_vernac v sid) s)
+ Flags.with_option Flags.beautify_file (Vernac.load_vernac ~verbosely ~interactive:false ~check:true sid) f_in
else
- Vernac.load_vernac v sid s)
+ Vernac.load_vernac ~verbosely ~interactive:false ~check:true sid s)
sid (List.rev !load_vernacular_list)
let load_vernacular_obj = ref ([] : string list)
@@ -187,6 +199,13 @@ let load_vernac_obj () =
let map dir = Qualid (Loc.tag @@ qualid_of_string dir) in
Vernacentries.vernac_require None None (List.rev_map map !load_vernacular_obj)
+(******************************************************************************)
+(* Required Modules *)
+(******************************************************************************)
+let set_include d p implicit =
+ let p = dirpath_of_string p in
+ Coqinit.push_include d p implicit
+
let require_prelude () =
let vo = Envars.coqlib () / "theories/Init/Prelude.vo" in
let vio = Envars.coqlib () / "theories/Init/Prelude.vio" in
@@ -209,10 +228,23 @@ let add_compat_require v =
| Flags.V8_7 -> add_require "Coq.Compat.Coq87"
| Flags.VOld | Flags.Current -> ()
-let compile_list = ref ([] : (bool * string) list)
+(******************************************************************************)
+(* File Compilation *)
+(******************************************************************************)
let glob_opt = ref false
+let compile_list = ref ([] : (bool * string) list)
+let _compilation_list : Stm.stm_doc_type list ref = ref []
+
+let compilation_mode = ref Vernac.BuildVo
+let compilation_output_name = ref None
+
+let batch_mode = ref false
+let set_batch_mode () =
+ System.trust_file_cache := false;
+ batch_mode := true
+
let add_compile verbose s =
set_batch_mode ();
Flags.quiet := true;
@@ -226,21 +258,66 @@ let add_compile verbose s =
in
compile_list := (verbose,s) :: !compile_list
-let compile_file (v,f) =
+let compile_file mode (verbosely,f_in) =
if !Flags.beautify then
- Flags.(with_option beautify_file (Vernac.compile v) f)
+ Flags.with_option Flags.beautify_file (fun f_in -> Vernac.compile ~verbosely ~mode ~f_in ~f_out:None) f_in
else
- Vernac.compile v f
+ Vernac.compile ~verbosely ~mode ~f_in ~f_out:None
let compile_files () =
if !compile_list == [] then ()
else
let init_state = States.freeze ~marshallable:`No in
+ let mode = !compilation_mode in
List.iter (fun vf ->
States.unfreeze init_state;
- compile_file vf)
+ compile_file mode vf)
(List.rev !compile_list)
+(******************************************************************************)
+(* VIO Dispatching *)
+(******************************************************************************)
+
+let vio_tasks = ref []
+let add_vio_task f =
+ set_batch_mode ();
+ Flags.quiet := true;
+ vio_tasks := f :: !vio_tasks
+let check_vio_tasks () =
+ let rc =
+ List.fold_left (fun acc t -> Vio_checking.check_vio t && acc)
+ true (List.rev !vio_tasks) in
+ if not rc then exit 1
+
+(* vio files *)
+let vio_files = ref []
+let vio_files_j = ref 0
+let vio_checking = ref false
+let add_vio_file f =
+ set_batch_mode ();
+ Flags.quiet := true;
+ vio_files := f :: !vio_files
+
+let set_vio_checking_j opt j =
+ try vio_files_j := int_of_string j
+ with Failure _ ->
+ prerr_endline ("The first argument of " ^ opt ^ " must the number");
+ prerr_endline "of concurrent workers to be used (a positive integer).";
+ prerr_endline "Makefiles generated by coq_makefile should be called";
+ prerr_endline "setting the J variable like in 'make vio2vo J=3'";
+ exit 1
+
+let schedule_vio_checking () =
+ if !vio_files <> [] && !vio_checking then
+ Vio_checking.schedule_vio_checking !vio_files_j !vio_files
+
+let schedule_vio_compilation () =
+ if !vio_files <> [] && not !vio_checking then
+ Vio_checking.schedule_vio_compilation !vio_files_j !vio_files
+
+(******************************************************************************)
+(* UI Options *)
+(******************************************************************************)
(** Options for proof general *)
let set_emacs () =
@@ -296,14 +373,15 @@ let usage_no_coqlib = CWarnings.create ~name:"usage-no-coqlib" ~category:"filesy
(fun () -> Pp.str "cannot guess a path for Coq libraries; dynaminally loaded flags will not be mentioned")
exception NoCoqLib
-let usage () =
+
+let usage batch =
begin
try
Envars.set_coqlib ~fail:(fun x -> raise NoCoqLib);
- init_load_path ();
+ Coqinit.init_load_path ();
with NoCoqLib -> usage_no_coqlib ()
end;
- if !Flags.batch_mode then Usage.print_usage_coqc ()
+ if batch then Usage.print_usage_coqc ()
else begin
Mltop.load_ml_objects_raw_rex
(Str.regexp (if Mltop.is_native then "^.*top.cmxs$" else "^.*top.cma$"));
@@ -389,47 +467,10 @@ let get_error_resilience opt = function
let get_task_list s = List.map int_of_string (Str.split (Str.regexp ",") s)
-let vio_tasks = ref []
-
-let add_vio_task f =
- set_batch_mode ();
- Flags.quiet := true;
- vio_tasks := f :: !vio_tasks
-
-let check_vio_tasks () =
- let rc =
- List.fold_left (fun acc t -> Vio_checking.check_vio t && acc)
- true (List.rev !vio_tasks) in
- if not rc then exit 1
-
-let vio_files = ref []
-let vio_files_j = ref 0
-let vio_checking = ref false
-let add_vio_file f =
- set_batch_mode ();
- Flags.quiet := true;
- vio_files := f :: !vio_files
-
-let set_vio_checking_j opt j =
- try vio_files_j := int_of_string j
- with Failure _ ->
- prerr_endline ("The first argument of " ^ opt ^ " must the number");
- prerr_endline "of concurrent workers to be used (a positive integer).";
- prerr_endline "Makefiles generated by coq_makefile should be called";
- prerr_endline "setting the J variable like in 'make vio2vo J=3'";
- exit 1
-
let is_not_dash_option = function
| Some f when String.length f > 0 && f.[0] <> '-' -> true
| _ -> false
-let schedule_vio_checking () =
- if !vio_files <> [] && !vio_checking then
- Vio_checking.schedule_vio_checking !vio_files_j !vio_files
-let schedule_vio_compilation () =
- if !vio_files <> [] && not !vio_checking then
- Vio_checking.schedule_vio_compilation !vio_files_j !vio_files
-
let get_native_name s =
(* We ignore even critical errors because this mode has to be super silent *)
try
@@ -464,7 +505,7 @@ let parse_args arglist =
(* Complex options with many args *)
|"-I"|"-include" ->
begin match rem with
- | d :: rem -> push_ml_include d; args := rem
+ | d :: rem -> Coqinit.push_ml_include d; args := rem
| [] -> error_missing_arg opt
end
|"-Q" ->
@@ -522,7 +563,7 @@ let parse_args arglist =
|"-dump-glob" -> Dumpglob.dump_into_file (next ()); glob_opt := true
|"-feedback-glob" -> Dumpglob.feedback_glob ()
|"-exclude-dir" -> System.exclude_directory (next ())
- |"-init-file" -> set_rcfile (next ())
+ |"-init-file" -> Coqinit.set_rcfile (next ())
|"-inputstate"|"-is" -> set_inputstate (next ())
|"-load-ml-object" -> Mltop.dir_ml_load (next ())
|"-load-ml-source" -> Mltop.dir_ml_use (next ())
@@ -537,7 +578,9 @@ let parse_args arglist =
|"-with-geoproof" -> Coq_config.with_geoproof := get_bool opt (next ())
|"-main-channel" -> Spawned.main_channel := get_host_port opt (next())
|"-control-channel" -> Spawned.control_channel := get_host_port opt (next())
- |"-vio2vo" -> add_compile false (next ()); Flags.compilation_mode := Flags.Vio2Vo
+ |"-vio2vo" ->
+ add_compile false (next ());
+ compilation_mode := Vernac.Vio2Vo
|"-toploop" -> set_toploop (next ())
|"-w" | "-W" ->
let w = next () in
@@ -545,7 +588,7 @@ let parse_args arglist =
else
let w = CWarnings.get_flags () ^ "," ^ w in
CWarnings.set_flags (CWarnings.normalize_flags_string w)
- |"-o" -> Flags.compilation_output_name := Some (next())
+ |"-o" -> compilation_output_name := Some (next())
(* Options with zero arg *)
|"-async-queries-always-delegate"
@@ -557,15 +600,15 @@ let parse_args arglist =
|"-batch" -> set_batch_mode ()
|"-test-mode" -> Flags.test_mode := true
|"-beautify" -> Flags.beautify := true
- |"-boot" -> Flags.boot := true; no_load_rc ()
+ |"-boot" -> Flags.boot := true; Coqinit.no_load_rc ()
|"-bt" -> Backtrace.record_backtrace true
|"-color" -> set_color (next ())
|"-config"|"--config" -> print_config := true
- |"-debug" -> set_debug ()
+ |"-debug" -> Coqinit.set_debug ()
|"-stm-debug" -> Flags.stm_debug := true
|"-emacs" -> set_emacs ()
|"-filteropts" -> filter_opts := true
- |"-h"|"-H"|"-?"|"-help"|"--help" -> usage ()
+ |"-h"|"-H"|"-?"|"-help"|"--help" -> usage !batch_mode
|"-ideslave" -> set_ideslave ()
|"-impredicative-set" -> set_impredicative_set ()
|"-indices-matter" -> Indtypes.enforce_indices_matter ()
@@ -578,9 +621,11 @@ let parse_args arglist =
else Flags.native_compiler := true
|"-output-context" -> output_context := true
|"-profile-ltac" -> Flags.profile_ltac := true
- |"-q" -> no_load_rc ()
+ |"-q" -> Coqinit.no_load_rc ()
|"-quiet"|"-silent" -> Flags.quiet := true; Flags.make_warn false
- |"-quick" -> Flags.compilation_mode := Flags.BuildVio
+ |"-quick" ->
+ Safe_typing.allow_delayed_constants := true;
+ compilation_mode := Vernac.BuildVio
|"-list-tags" -> print_tags := true
|"-time" -> Flags.time := true
|"-type-in-type" -> set_type_in_type ()
@@ -615,7 +660,7 @@ let init_toplevel arglist =
if !print_config then (Envars.print_config stdout Coq_config.all_src_dirs; exit (exitcode ()));
if !print_tags then (print_style_tags (); exit (exitcode ()));
if !filter_opts then (print_string (String.concat "\n" extras); exit 0);
- init_load_path ();
+ Coqinit.init_load_path ();
Option.iter Mltop.load_ml_object_raw !toploop;
let extras = !toploop_init extras in
if not (CList.is_empty extras) then begin
@@ -627,20 +672,19 @@ let init_toplevel arglist =
inputstate ();
Mltop.init_known_plugins ();
engage ();
- if (not !Flags.batch_mode || CList.is_empty !compile_list)
+ if (not !batch_mode || CList.is_empty !compile_list)
&& Global.env_is_initial ()
then Declaremods.start_library !toplevel_name;
- init_library_roots ();
load_vernac_obj ();
require ();
- (* XXX: This is incorrect in batch mode, as we will initialize
- the STM before having done Declaremods.start_library, thus
- state 1 is invalid. This bug was present in 8.5/8.6. *)
- Stm.init ();
- let sid = load_rcfile (Stm.get_current_state ()) in
+ Stm.(init { doc_type = Interactive Names.DirPath.empty });
+ let sid = Coqinit.load_rcfile (Stm.get_current_state ()) in
(* XXX: We ignore this for now, but should be threaded to the toplevels *)
let _sid = load_vernacular sid in
compile_files ();
+ (* All these tasks use coqtop as a driver to invoke more coqtop,
+ * they should be really orthogonal to coqtop.
+ *)
schedule_vio_checking ();
schedule_vio_compilation ();
check_vio_tasks ();
@@ -648,13 +692,13 @@ let init_toplevel arglist =
with any ->
flush_all();
let extra =
- if !Flags.batch_mode && not Stateid.(equal (Stm.get_current_state ()) dummy)
+ if !batch_mode && not Stateid.(equal (Stm.get_current_state ()) dummy)
then None
else Some (str "Error during initialization: ")
in
fatal_error ?extra any
end;
- if !Flags.batch_mode then begin
+ if !batch_mode then begin
flush_all();
if !output_context then
Feedback.msg_notice Flags.(with_option raw_print Prettyp.print_full_pure_context () ++ fnl ());
diff --git a/toplevel/coqtop.mli b/toplevel/coqtop.mli
index 892d64d91..ac2be7e16 100644
--- a/toplevel/coqtop.mli
+++ b/toplevel/coqtop.mli
@@ -15,7 +15,6 @@ val init_toplevel : string list -> unit
val start : unit -> unit
-
(* For other toploops *)
val toploop_init : (string list -> string list) ref
val toploop_run : (unit -> unit) ref
diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml
index 1b020bc87..cb89dc8ff 100644
--- a/toplevel/vernac.ml
+++ b/toplevel/vernac.ml
@@ -113,13 +113,13 @@ let vernac_error msg =
(* Stm.End_of_input -> true *)
(* | _ -> false *)
-let rec interp_vernac sid (loc,com) =
+let rec interp_vernac ~check ~interactive sid (loc,com) =
let interp = function
| VernacLoad (verbosely, fname) ->
let fname = Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) fname in
let fname = CUnix.make_suffix fname ".v" in
let f = Loadpath.locate_file fname in
- load_vernac verbosely sid f
+ load_vernac ~verbosely ~check ~interactive sid f
| v ->
(* XXX: We need to run this before add as the classification is
@@ -142,17 +142,16 @@ let rec interp_vernac sid (loc,com) =
(* Main STM interaction *)
if ntip <> `NewTip then
anomaly (str "vernac.ml: We got an unfocus operation on the toplevel!");
+
(* Due to bug #5363 we cannot use observe here as we should,
it otherwise reveals bugs *)
(* Stm.observe nsid; *)
-
- let check_proof = Flags.(!compilation_mode = BuildVo || not !batch_mode) in
- if check_proof then Stm.finish ();
+ if check then Stm.finish ();
(* We could use a more refined criteria that depends on the
vernac. For now we imitate the old approach and rely on the
classification. *)
- let print_goals = not !Flags.batch_mode && not !Flags.quiet &&
+ let print_goals = interactive && not !Flags.quiet &&
is_proof_step && Proof_global.there_are_pending_proofs () in
if print_goals then Feedback.msg_notice (pr_open_cur_subgoals ());
@@ -167,7 +166,7 @@ let rec interp_vernac sid (loc,com) =
with reraise ->
(* XXX: In non-interactive mode edit_at seems to do very weird
things, so we better avoid it while we investigate *)
- if not !Flags.batch_mode then ignore(Stm.edit_at sid);
+ if interactive then ignore(Stm.edit_at sid);
let (reraise, info) = CErrors.push reraise in
let info = begin
match Loc.get_loc info with
@@ -176,7 +175,7 @@ let rec interp_vernac sid (loc,com) =
end in iraise (reraise, info)
(* Load a vernac file. CErrors are annotated with file and location *)
-and load_vernac verbosely sid file =
+and load_vernac ~verbosely ~check ~interactive sid file =
let ft_beautify, close_beautify =
if !Flags.beautify_file then
let chan_beautify = open_out (file^beautify_suffix) in
@@ -217,7 +216,7 @@ and load_vernac verbosely sid file =
Option.iter (vernac_echo ?loc) in_echo;
checknav_simple (loc, ast);
- let nsid = Flags.silently (interp_vernac !rsid) (loc, ast) in
+ let nsid = Flags.silently (interp_vernac ~check ~interactive !rsid) (loc, ast) in
rsid := nsid
done;
!rsid
@@ -245,7 +244,7 @@ and load_vernac verbosely sid file =
let process_expr sid loc_ast =
checknav_deep loc_ast;
- interp_vernac sid loc_ast
+ interp_vernac ~interactive:true ~check:true sid loc_ast
let warn_file_no_extension =
CWarnings.create ~name:"file-no-extension" ~category:"filesystem"
@@ -282,8 +281,10 @@ let ensure_exists f =
if not (Sys.file_exists f) then
vernac_error (hov 0 (str "Can't find file" ++ spc () ++ str f))
+type compilation_mode = BuildVo | BuildVio | Vio2Vo
+
(* Compile a vernac file *)
-let compile verbosely f =
+let compile ~verbosely ~mode ~f_in ~f_out=
let check_pending_proofs () =
let pfs = Proof_global.get_all_proof_names () in
if not (List.is_empty pfs) then
@@ -293,12 +294,12 @@ let compile verbosely f =
|> prlist_with_sep pr_comma Names.Id.print)
++ str ".")
in
- match !Flags.compilation_mode with
- | Flags.BuildVo ->
- let long_f_dot_v = ensure_v f in
+ match mode with
+ | BuildVo ->
+ let long_f_dot_v = ensure_v f_in in
ensure_exists long_f_dot_v;
let long_f_dot_vo =
- match !Flags.compilation_output_name with
+ match f_out with
| None -> long_f_dot_v ^ "o"
| Some f -> ensure_vo long_f_dot_v f in
let ldir = Flags.verbosely Library.start_library long_f_dot_vo in
@@ -309,7 +310,7 @@ let compile verbosely f =
Dumpglob.start_dump_glob ~vfile:long_f_dot_v ~vofile:long_f_dot_vo;
Dumpglob.dump_string ("F" ^ Names.DirPath.to_string ldir ^ "\n");
let wall_clock1 = Unix.gettimeofday () in
- let _ = load_vernac verbosely (Stm.get_current_state ()) long_f_dot_v in
+ let _ = load_vernac ~verbosely ~check:true ~interactive:false (Stm.get_current_state ()) long_f_dot_v in
Stm.join ();
let wall_clock2 = Unix.gettimeofday () in
check_pending_proofs ();
@@ -318,32 +319,31 @@ let compile verbosely f =
(Printf.sprintf "%.3f" (wall_clock2 -. wall_clock1));
Aux_file.stop_aux_file ();
Dumpglob.end_dump_glob ()
- | Flags.BuildVio ->
- let long_f_dot_v = ensure_v f in
+ | BuildVio ->
+ let long_f_dot_v = ensure_v f_in in
ensure_exists long_f_dot_v;
let long_f_dot_vio =
- match !Flags.compilation_output_name with
+ match f_out with
| None -> long_f_dot_v ^ "io"
| Some f -> ensure_vio long_f_dot_v f in
let ldir = Flags.verbosely Library.start_library long_f_dot_vio in
Dumpglob.noglob ();
Stm.set_compilation_hints long_f_dot_vio;
- let _ = load_vernac verbosely (Stm.get_current_state ()) long_f_dot_v in
+ let _ = load_vernac ~verbosely ~check:false ~interactive:false (Stm.get_current_state ()) long_f_dot_v in
Stm.finish ();
check_pending_proofs ();
Stm.snapshot_vio ldir long_f_dot_vio;
Stm.reset_task_queue ()
- | Flags.Vio2Vo ->
+ | Vio2Vo ->
let open Filename in
- let open Library in
Dumpglob.noglob ();
- let f = if check_suffix f ".vio" then chop_extension f else f in
- let lfdv, sum, lib, univs, disch, tasks, proofs = load_library_todo f in
+ let f = if check_suffix f_in ".vio" then chop_extension f_in else f_in in
+ let lfdv, sum, lib, univs, disch, tasks, proofs = Library.load_library_todo f in
Stm.set_compilation_hints lfdv;
let univs, proofs = Stm.finish_tasks lfdv univs disch proofs tasks in
Library.save_library_raw lfdv sum lib univs proofs
-let compile v f =
+let compile ~verbosely ~mode ~f_in ~f_out =
ignore(CoqworkmgrApi.get 1);
- compile v f;
+ compile ~verbosely ~mode ~f_in ~f_out;
CoqworkmgrApi.giveback 1
diff --git a/toplevel/vernac.mli b/toplevel/vernac.mli
index bccf560e1..410dcf0d4 100644
--- a/toplevel/vernac.mli
+++ b/toplevel/vernac.mli
@@ -17,7 +17,9 @@ val process_expr : Stateid.t -> Vernacexpr.vernac_expr Loc.located -> Stateid.t
(** [load_vernac echo sid file] Loads [file] on top of [sid], will
echo the commands if [echo] is set. Callers are expected to handle
and print errors in form of exceptions. *)
-val load_vernac : bool -> Stateid.t -> string -> Stateid.t
+val load_vernac : verbosely:bool -> check:bool -> interactive:bool -> Stateid.t -> string -> Stateid.t
+
+type compilation_mode = BuildVo | BuildVio | Vio2Vo
(** Compile a vernac file, (f is assumed without .v suffix) *)
-val compile : bool -> string -> unit
+val compile : verbosely:bool -> mode:compilation_mode -> f_in:string -> f_out:string option -> unit
diff --git a/vernac/command.ml b/vernac/command.ml
index 120f9590f..a1a87d54e 100644
--- a/vernac/command.ml
+++ b/vernac/command.ml
@@ -518,7 +518,8 @@ let check_param = function
| CLocalDef (na, _, _) -> check_named na
| CLocalAssum (nas, Default _, _) -> List.iter check_named nas
| CLocalAssum (nas, Generalized _, _) -> ()
-| CLocalPattern _ -> assert false
+| CLocalPattern (loc,_) ->
+ Loc.raise ?loc (Stream.Error "pattern with quote not allowed here.")
let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite =
check_all_names_different indl;
diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml
index 5298ef2e4..b2d48bb2f 100644
--- a/vernac/metasyntax.ml
+++ b/vernac/metasyntax.ml
@@ -991,7 +991,7 @@ let is_not_printable onlyparse nonreversible = function
(warn_non_reversible_notation (); true)
else onlyparse
-let find_precedence lev etyps symbols =
+let find_precedence lev etyps symbols onlyprint =
let first_symbol =
let rec aux = function
| Break _ :: t -> aux t
@@ -1009,8 +1009,13 @@ let find_precedence lev etyps symbols =
| None -> [],0
| Some (NonTerminal x) ->
(try match List.assoc x etyps with
- | ETConstr _ ->
- user_err Pp.(str "The level of the leftmost non-terminal cannot be changed.")
+ | ETConstr _ ->
+ if onlyprint then
+ if Option.is_empty lev then
+ user_err Pp.(str "Explicit level needed in only-printing mode when the level of the leftmost non-terminal is given.")
+ else [],Option.get lev
+ else
+ user_err Pp.(str "The level of the leftmost non-terminal cannot be changed.")
| ETName | ETBigint | ETReference ->
begin match lev with
| None ->
@@ -1134,7 +1139,7 @@ let compute_syntax_data df modifiers =
let need_squash = not (List.equal Notation.symbol_eq symbols symbols_for_grammar) in
let ntn_for_grammar = if need_squash then make_notation_key symbols_for_grammar else ntn_for_interp in
if not onlyprint then check_rule_productivity symbols_for_grammar;
- let msgs,n = find_precedence mods.level mods.etyps symbols in
+ let msgs,n = find_precedence mods.level mods.etyps symbols onlyprint in
(* To globalize... *)
let etyps = join_auxiliary_recursive_types recvars mods.etyps in
let sy_typs, prec =
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 83296cf58..203d913db 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -1785,7 +1785,7 @@ let vernac_locate = let open Feedback in function
(Constrextern.without_symbols pr_lglob_constr) ntn sc)
| LocateLibrary qid -> print_located_library qid
| LocateModule qid -> msg_notice (print_located_module qid)
- | LocateTactic qid -> msg_notice (print_located_tactic qid)
+ | LocateOther (s, qid) -> msg_notice (print_located_other s qid)
| LocateFile f -> msg_notice (locate_file f)
let vernac_register id r =