aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--.github/PULL_REQUEST_TEMPLATE.md2
-rw-r--r--CHANGES13
-rw-r--r--CONTRIBUTING.md2
-rw-r--r--META.coq310
-rw-r--r--Makefile.build7
-rw-r--r--Makefile.checker9
-rw-r--r--Makefile.ide4
-rw-r--r--checker/checker.mli9
-rw-r--r--checker/cic.mli12
-rw-r--r--checker/closure.ml1
-rw-r--r--checker/closure.mli2
-rw-r--r--checker/environ.ml15
-rw-r--r--checker/environ.mli5
-rw-r--r--checker/indtypes.ml6
-rw-r--r--checker/main.mli10
-rw-r--r--checker/mod_checking.ml11
-rw-r--r--checker/print.mli11
-rw-r--r--checker/reduction.ml30
-rw-r--r--checker/validate.ml2
-rw-r--r--checker/validate.mli9
-rw-r--r--checker/values.ml36
-rw-r--r--checker/values.mli26
-rw-r--r--checker/votour.ml6
-rw-r--r--checker/votour.mli10
-rw-r--r--clib/clib.mllib1
-rw-r--r--clib/range.ml91
-rw-r--r--clib/range.mli37
-rw-r--r--configure.ml19
-rw-r--r--dev/base_include2
-rw-r--r--dev/ci/README.md2
-rw-r--r--dev/ci/user-overlays/06482-ppedrot-check-poly-effects.sh4
-rw-r--r--dev/db95
-rw-r--r--dev/doc/setup.txt24
-rw-r--r--dev/include1
-rw-r--r--dev/ocamldebug-coq.run2
-rw-r--r--dev/set_raw_db1
-rw-r--r--dev/top_printers.ml15
-rw-r--r--dev/top_printers.mli173
-rw-r--r--doc/refman/RefMan-pro.tex9
-rw-r--r--engine/eConstr.ml5
-rw-r--r--engine/eConstr.mli1
-rw-r--r--engine/termops.ml12
-rw-r--r--engine/termops.mli10
-rw-r--r--engine/uState.ml20
-rw-r--r--ide/config_lexer.mli10
-rw-r--r--ide/coq_commands.mli11
-rw-r--r--ide/coq_lex.mli11
-rw-r--r--ide/coq_lex.mll4
-rw-r--r--ide/coqide_main.ml43
-rw-r--r--ide/coqide_main.mli10
-rw-r--r--ide/coqide_ui.mli10
-rw-r--r--ide/gtk_parsing.ml109
-rw-r--r--ide/gtk_parsing.mli26
-rw-r--r--ide/ide_slave.mli10
-rw-r--r--ide/macos_prehook.mli10
-rw-r--r--ide/nanoPG.mli11
-rw-r--r--ide/utf8_convert.mli9
-rw-r--r--ide/wg_Find.ml61
-rw-r--r--intf/vernacexpr.ml32
-rw-r--r--kernel/cClosure.ml36
-rw-r--r--kernel/csymtable.ml2
-rw-r--r--kernel/declarations.ml1
-rw-r--r--kernel/declareops.ml3
-rw-r--r--kernel/declareops.mli2
-rw-r--r--kernel/environ.ml31
-rw-r--r--kernel/names.mli9
-rw-r--r--kernel/nativecode.ml6
-rw-r--r--kernel/nativelambda.ml2
-rw-r--r--kernel/pre_env.ml52
-rw-r--r--kernel/pre_env.mli16
-rw-r--r--kernel/term_typing.ml28
-rw-r--r--parsing/g_vernac.ml432
-rw-r--r--plugins/firstorder/formula.ml3
-rw-r--r--plugins/funind/g_indfun.ml42
-rw-r--r--plugins/funind/glob_term_to_relation.ml4
-rw-r--r--plugins/ltac/g_ltac.ml43
-rw-r--r--plugins/ssr/ssrvernac.ml44
-rw-r--r--pretyping/constr_matching.ml10
-rw-r--r--pretyping/evarconv.ml5
-rw-r--r--pretyping/typing.ml5
-rw-r--r--printing/ppvernac.ml57
-rw-r--r--proofs/proof_global.ml11
-rw-r--r--stm/proofBlockDelimiter.ml4
-rw-r--r--stm/stm.ml9
-rw-r--r--stm/vernac_classifier.ml14
-rw-r--r--tactics/hipattern.ml54
-rw-r--r--test-suite/Makefile13
-rw-r--r--test-suite/README.md75
-rw-r--r--test-suite/bugs/closed/6297.v8
-rw-r--r--test-suite/bugs/closed/6490.v4
-rw-r--r--test-suite/bugs/closed/6617.v34
-rw-r--r--test-suite/success/BracketsWithGoalSelector.v16
-rw-r--r--test-suite/success/abstract_poly.v2
-rw-r--r--test-suite/success/dtauto-let-deps.v24
-rw-r--r--tools/coqdep.ml2
-rw-r--r--tools/md5sum.ml24
-rw-r--r--vernac/vernacentries.ml67
-rw-r--r--vernac/vernacinterp.ml1
-rw-r--r--vernac/vernacinterp.mli1
-rw-r--r--vernac/vernacprop.ml8
100 files changed, 1562 insertions, 521 deletions
diff --git a/.github/PULL_REQUEST_TEMPLATE.md b/.github/PULL_REQUEST_TEMPLATE.md
index c7437c4a4..a9230042a 100644
--- a/.github/PULL_REQUEST_TEMPLATE.md
+++ b/.github/PULL_REQUEST_TEMPLATE.md
@@ -13,4 +13,4 @@ Fixes / closes #????
<!-- If this is a feature pull request / breaks compatibility: -->
<!-- (Otherwise, remove these lines.) -->
- [ ] Corresponding documentation was added / updated.
-- [ ] Entry added in [CHANGES](/CHANGES).
+- [ ] Entry added in CHANGES.
diff --git a/CHANGES b/CHANGES
index ead751147..7fe5b91f5 100644
--- a/CHANGES
+++ b/CHANGES
@@ -40,6 +40,14 @@ Tactics
- Added tactic optimize_heap, analogous to the Vernacular Optimize
Heap, which performs a major garbage collection and heap compaction
in the OCaml run-time system.
+- The tactics "dtauto", "dintuition", "firstorder" now handle inductive types
+ with let bindings in the parameters.
+
+Focusing
+
+- Focusing bracket `{` now supports single-numbered goal selector,
+ e.g. `2: {` will focus on the second sub-goal. As usual, unfocus
+ with `}` once the sub-goal is fully solved.
Vernacular Commands
@@ -60,6 +68,11 @@ Checker
- The checker now accepts filenames in addition to logical paths.
+CoqIDE
+
+- Find and Replace All report the number of occurrences found; Find indicates
+ when it wraps.
+
Documentation
- The Coq FAQ, formerly located at https://coq.inria.fr/faq, has been
diff --git a/CONTRIBUTING.md b/CONTRIBUTING.md
index 37b556c55..1a769333c 100644
--- a/CONTRIBUTING.md
+++ b/CONTRIBUTING.md
@@ -28,6 +28,8 @@ Please make pull requests against the `master` branch.
It's helpful to run the Coq test suite with `make test-suite` before submitting your change. Travis CI runs this test suite and a much larger one including external Coq developments on every pull request, but these results take significantly longer to come back (on the order of a few hours). Running the test suite locally will take somewhere around 10-15 minutes. Refer to [`dev/ci/README.md`](/dev/ci/README.md#information-for-developers) for more information on Travis CI tests.
+If your pull request fixes a bug, please consider adding a regression test as well. See [`test-suite/README.md`](/test-suite/README.md) for how to do so.
+
Don't be alarmed if the pull request process takes some time. It can take a few days to get feedback, approval on the final changes, and then a merge. Coq doesn't release new versions very frequently so it can take a few months for your change to land in a released version. That said, you can start using the latest Coq `master` branch to take advantage of all the new features, improvements, and fixes.
Here are a few tags Coq developers may add to your PR and what they mean. In general feedback and requests for you as the pull request author will be in the comments and tags are only used to organize pull requests.
diff --git a/META.coq b/META.coq
index f7998db8e..d180820e8 100644
--- a/META.coq
+++ b/META.coq
@@ -240,19 +240,6 @@ package "stm" (
)
-package "ltac" (
-
- description = "Coq LTAC Plugin"
- version = "8.7"
-
- requires = "coq.stm"
- directory = "plugins/ltac"
-
- archive(byte) = "ltac_plugin.cmo"
- archive(native) = "ltac_plugin.cmx"
-
-)
-
package "toplevel" (
description = "Coq Toplevel"
@@ -293,3 +280,300 @@ package "ide" (
archive(native) = "ide.cmxa"
)
+
+package "plugins" (
+
+ description = "Coq built-in plugins"
+ version = "8.7"
+
+ directory = "plugins"
+
+ package "ltac" (
+
+ description = "Coq LTAC Plugin"
+ version = "8.7"
+
+ requires = "coq.stm"
+ directory = "ltac"
+
+ archive(byte) = "ltac_plugin.cmo"
+ archive(native) = "ltac_plugin.cmx"
+
+ )
+
+ package "tauto" (
+
+ description = "Coq tauto plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.ltac"
+ directory = "ltac"
+
+ archive(byte) = "tauto_plugin.cmo"
+ archive(native) = "tauto_plugin.cmx"
+ )
+
+ package "omega" (
+
+ description = "Coq omega plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.ltac"
+ directory = "omega"
+
+ archive(byte) = "omega_plugin.cmo"
+ archive(native) = "omega_plugin.cmx"
+ )
+
+ package "romega" (
+
+ description = "Coq romega plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.omega"
+ directory = "romega"
+
+ archive(byte) = "romega_plugin.cmo"
+ archive(native) = "romega_plugin.cmx"
+ )
+
+ package "micromega" (
+
+ description = "Coq micromega plugin"
+ version = "8.7"
+
+ requires = "num,coq.plugins.ltac"
+ directory = "micromega"
+
+ archive(byte) = "micromega_plugin.cmo"
+ archive(native) = "micromega_plugin.cmx"
+ )
+
+ package "quote" (
+
+ description = "Coq quote plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.ltac"
+ directory = "quote"
+
+ archive(byte) = "quote_plugin.cmo"
+ archive(native) = "quote_plugin.cmx"
+ )
+
+ package "newring" (
+
+ description = "Coq newring plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.quote"
+ directory = "setoid_ring"
+
+ archive(byte) = "newring_plugin.cmo"
+ archive(native) = "newring_plugin.cmx"
+ )
+
+ package "fourier" (
+
+ description = "Coq fourier plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.ltac"
+ directory = "fourier"
+
+ archive(byte) = "fourier_plugin.cmo"
+ archive(native) = "fourier_plugin.cmx"
+ )
+
+ package "extraction" (
+
+ description = "Coq extraction plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.ltac"
+ directory = "extraction"
+
+ archive(byte) = "extraction_plugin.cmo"
+ archive(native) = "extraction_plugin.cmx"
+ )
+
+ package "cc" (
+
+ description = "Coq cc plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.ltac"
+ directory = "cc"
+
+ archive(byte) = "cc_plugin.cmo"
+ archive(native) = "cc_plugin.cmx"
+ )
+
+ package "ground" (
+
+ description = "Coq ground plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.ltac"
+ directory = "firstorder"
+
+ archive(byte) = "ground_plugin.cmo"
+ archive(native) = "ground_plugin.cmx"
+ )
+
+ package "rtauto" (
+
+ description = "Coq rtauto plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.ltac"
+ directory = "rtauto"
+
+ archive(byte) = "rtauto_plugin.cmo"
+ archive(native) = "rtauto_plugin.cmx"
+ )
+
+ package "btauto" (
+
+ description = "Coq btauto plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.ltac"
+ directory = "btauto"
+
+ archive(byte) = "btauto_plugin.cmo"
+ archive(native) = "btauto_plugin.cmx"
+ )
+
+ package "recdef" (
+
+ description = "Coq recdef plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.extraction"
+ directory = "funind"
+
+ archive(byte) = "recdef_plugin.cmo"
+ archive(native) = "recdef_plugin.cmx"
+ )
+
+ package "nsatz" (
+
+ description = "Coq nsatz plugin"
+ version = "8.7"
+
+ requires = "num,coq.plugins.ltac"
+ directory = "nsatz"
+
+ archive(byte) = "nsatz_plugin.cmo"
+ archive(native) = "nsatz_plugin.cmx"
+ )
+
+ package "natsyntax" (
+
+ description = "Coq natsyntax plugin"
+ version = "8.7"
+
+ requires = ""
+ directory = "syntax"
+
+ archive(byte) = "nat_syntax_plugin.cmo"
+ archive(native) = "nat_syntax_plugin.cmx"
+ )
+
+ package "zsyntax" (
+
+ description = "Coq zsyntax plugin"
+ version = "8.7"
+
+ requires = ""
+ directory = "syntax"
+
+ archive(byte) = "z_syntax_plugin.cmo"
+ archive(native) = "z_syntax_plugin.cmx"
+ )
+
+ package "rsyntax" (
+
+ description = "Coq rsyntax plugin"
+ version = "8.7"
+
+ requires = ""
+ directory = "syntax"
+
+ archive(byte) = "r_syntax_plugin.cmo"
+ archive(native) = "r_syntax_plugin.cmx"
+ )
+
+ package "int31syntax" (
+
+ description = "Coq int31syntax plugin"
+ version = "8.7"
+
+ requires = ""
+ directory = "syntax"
+
+ archive(byte) = "int31_syntax_plugin.cmo"
+ archive(native) = "int31_syntax_plugin.cmx"
+ )
+
+ package "asciisyntax" (
+
+ description = "Coq asciisyntax plugin"
+ version = "8.7"
+
+ requires = ""
+ directory = "syntax"
+
+ archive(byte) = "ascii_syntax_plugin.cmo"
+ archive(native) = "ascii_syntax_plugin.cmx"
+ )
+
+ package "stringsyntax" (
+
+ description = "Coq stringsyntax plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.asciisyntax"
+ directory = "syntax"
+
+ archive(byte) = "string_syntax_plugin.cmo"
+ archive(native) = "string_syntax_plugin.cmx"
+ )
+
+ package "derive" (
+
+ description = "Coq derive plugin"
+ version = "8.7"
+
+ requires = ""
+ directory = "derive"
+
+ archive(byte) = "derive_plugin.cmo"
+ archive(native) = "derive_plugin.cmx"
+ )
+
+ package "ssrmatching" (
+
+ description = "Coq ssrmatching plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.ltac"
+ directory = "ssrmatching"
+
+ archive(byte) = "ssrmatching_plugin.cmo"
+ archive(native) = "ssrmatching_plugin.cmx"
+ )
+
+ package "ssreflect" (
+
+ description = "Coq ssreflect plugin"
+ version = "8.7"
+
+ requires = "coq.plugins.ssrmatching"
+ directory = "ssr"
+
+ archive(byte) = "ssreflect_plugin.cmo"
+ archive(native) = "ssreflect_plugin.cmx"
+ )
+)
diff --git a/Makefile.build b/Makefile.build
index d8474ae13..f0dd46b0f 100644
--- a/Makefile.build
+++ b/Makefile.build
@@ -594,11 +594,14 @@ kernel/kernel.cma: kernel/kernel.mllib
$(SHOW)'OCAMLOPT -pack -o $@'
$(HIDE)$(OCAMLOPT) $(MLINCLUDES) $(OPTFLAGS) -pack -o $@ $(filter-out %.mlpack, $^)
+COND_IDEFLAGS=$(if $(filter tools/fake_ide% tools/coq_makefile%,$<), -I ide,)
+COND_PRINTERFLAGS=$(if $(filter dev/%,$<), -I dev,)
+
COND_BYTEFLAGS= \
- $(if $(filter tools/fake_ide% tools/coq_makefile%,$<), -I ide,) $(MLINCLUDES) $(BYTEFLAGS)
+ $(COND_IDEFLAGS) $(COND_PRINTERFLAGS) $(MLINCLUDES) $(BYTEFLAGS)
COND_OPTFLAGS= \
- $(if $(filter tools/fake_ide% tools/coq_makefile%,$<), -I ide,) $(MLINCLUDES) $(OPTFLAGS)
+ $(COND_IDEFLAGS) $(MLINCLUDES) $(OPTFLAGS)
plugins/micromega/%.cmi: plugins/micromega/%.mli
$(SHOW)'OCAMLC $<'
diff --git a/Makefile.checker b/Makefile.checker
index 5b5855782..0e429fe86 100644
--- a/Makefile.checker
+++ b/Makefile.checker
@@ -33,7 +33,7 @@ CHECKERDEPS := $(addsuffix .d, $(CHECKMLDFILE) $(CHECKMLLIBFILE))
-include $(CHECKERDEPS)
ifeq ($(BEST),opt)
-$(CHICKEN): checker/check.cmxa checker/main.ml
+$(CHICKEN): checker/check.cmxa checker/main.mli checker/main.ml
$(SHOW)'OCAMLOPT -o $@'
$(HIDE)$(OCAMLOPT) -linkpkg $(SYSMOD) $(CHKLIBS) $(OPTFLAGS) $(LINKMETADATA) -o $@ $^
$(STRIP) $@
@@ -43,7 +43,7 @@ $(CHICKEN): $(CHICKENBYTE)
cp $< $@
endif
-$(CHICKENBYTE): checker/check.cma checker/main.ml
+$(CHICKENBYTE): checker/check.cma checker/main.mli checker/main.ml
$(SHOW)'OCAMLC -o $@'
$(HIDE)$(OCAMLC) -linkpkg $(SYSMOD) $(CHKLIBS) $(BYTEFLAGS) $(CUSTOM) -o $@ $^
@@ -77,9 +77,8 @@ checker/%.cmx: checker/%.ml
md5chk:
$(SHOW)'MD5SUM cic.mli'
- $(HIDE)v=`tr -d "\r" < checker/cic.mli | $(MD5SUM) | sed -n -e 's/ .*//' -e '/^/p'`; \
- if grep -q "$$v" checker/values.ml; \
- then true; else echo "Error: outdated checker/values.ml: $$v" >&2; false; fi
+ $(HIDE)if grep -q "^MD5 $$($(OCAML) tools/md5sum.ml checker/cic.mli)$$" checker/values.ml; \
+ then true; else echo "Error: outdated checker/values.ml" >&2; false; fi
.PHONY: md5chk
diff --git a/Makefile.ide b/Makefile.ide
index 08829b554..09eef1f6b 100644
--- a/Makefile.ide
+++ b/Makefile.ide
@@ -45,8 +45,8 @@ IDEDEPS:=clib/clib.cma lib/lib.cma
IDECMA:=ide/ide.cma
IDETOPLOOPCMA=ide/coqidetop.cma
-LINKIDE:=$(IDEDEPS) $(IDECDEPS) $(IDECMA) ide/coqide_main.ml
-LINKIDEOPT:=$(IDEOPTCDEPS) $(patsubst %.cma,%.cmxa,$(IDEDEPS:.cmo=.cmx)) $(IDECMA:.cma=.cmxa) ide/coqide_main.ml
+LINKIDE:=$(IDEDEPS) $(IDECDEPS) $(IDECMA) ide/coqide_main.mli ide/coqide_main.ml
+LINKIDEOPT:=$(IDEOPTCDEPS) $(patsubst %.cma,%.cmxa,$(IDEDEPS:.cmo=.cmx)) $(IDECMA:.cma=.cmxa) ide/coqide_main.mli ide/coqide_main.ml
IDEFILES=$(wildcard ide/*.lang) ide/coq_style.xml ide/coq.png ide/MacOS/default_accel_map
diff --git a/checker/checker.mli b/checker/checker.mli
new file mode 100644
index 000000000..ceab13774
--- /dev/null
+++ b/checker/checker.mli
@@ -0,0 +1,9 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+val start : unit -> unit
diff --git a/checker/cic.mli b/checker/cic.mli
index 4a0e706aa..95dd18f5f 100644
--- a/checker/cic.mli
+++ b/checker/cic.mli
@@ -170,6 +170,17 @@ type set_predicativity = ImpredicativeSet | PredicativeSet
type engagement = set_predicativity
+(** {6 Conversion oracle} *)
+
+type level = Expand | Level of int | Opaque
+
+type oracle = {
+ var_opacity : level Id.Map.t;
+ cst_opacity : level Cmap.t;
+ var_trstate : Id.Pred.t;
+ cst_trstate : Cpred.t;
+}
+
(** {6 Representation of constants (Definition/Axiom) } *)
@@ -219,6 +230,7 @@ type typing_flags = {
check_guarded : bool; (** If [false] then fixed points and co-fixed
points are assumed to be total. *)
check_universes : bool; (** If [false] universe constraints are not checked *)
+ conv_oracle : oracle; (** Unfolding strategies for conversion *)
}
type constant_body = {
diff --git a/checker/closure.ml b/checker/closure.ml
index 3a56bba01..98f8c4a82 100644
--- a/checker/closure.ml
+++ b/checker/closure.ml
@@ -822,6 +822,7 @@ type clos_infos = fconstr infos
let infos_env x = x.i_env
let infos_flags x = x.i_flags
+let oracle_of_infos x = x.i_env.env_conv_oracle
let create_clos_infos flgs env =
create (fun _ -> inject) flgs env
diff --git a/checker/closure.mli b/checker/closure.mli
index 02d8b22fa..ce8c64e30 100644
--- a/checker/closure.mli
+++ b/checker/closure.mli
@@ -147,6 +147,8 @@ type clos_infos
val create_clos_infos : reds -> env -> clos_infos
val infos_env : clos_infos -> env
val infos_flags : clos_infos -> reds
+val oracle_of_infos : clos_infos -> oracle
+
(* Reduction function *)
diff --git a/checker/environ.ml b/checker/environ.ml
index 9db0d60e8..3830cd0dc 100644
--- a/checker/environ.ml
+++ b/checker/environ.ml
@@ -21,7 +21,15 @@ type env = {
env_globals : globals;
env_rel_context : rel_context;
env_stratification : stratification;
- env_imports : Cic.vodigest MPmap.t }
+ env_imports : Cic.vodigest MPmap.t;
+ env_conv_oracle : oracle; }
+
+let empty_oracle = {
+ var_opacity = Id.Map.empty;
+ cst_opacity = Cmap.empty;
+ var_trstate = Id.Pred.empty;
+ cst_trstate = Cpred.empty;
+}
let empty_env = {
env_globals =
@@ -34,7 +42,8 @@ let empty_env = {
env_stratification =
{ env_universes = Univ.initial_universes;
env_engagement = PredicativeSet };
- env_imports = MPmap.empty }
+ env_imports = MPmap.empty;
+ env_conv_oracle = empty_oracle }
let engagement env = env.env_stratification.env_engagement
let universes env = env.env_stratification.env_universes
@@ -51,6 +60,8 @@ let set_engagement (impr_set as c) env =
{ env with env_stratification =
{ env.env_stratification with env_engagement = c } }
+let set_oracle env oracle = { env with env_conv_oracle = oracle }
+
(* Digests *)
let add_digest env dp digest =
diff --git a/checker/environ.mli b/checker/environ.mli
index 6bda838f8..ba62ed519 100644
--- a/checker/environ.mli
+++ b/checker/environ.mli
@@ -18,6 +18,7 @@ type env = {
env_rel_context : rel_context;
env_stratification : stratification;
env_imports : Cic.vodigest MPmap.t;
+ env_conv_oracle : Cic.oracle;
}
val empty_env : env
@@ -25,6 +26,10 @@ val empty_env : env
val engagement : env -> Cic.engagement
val set_engagement : Cic.engagement -> env -> env
+(** Oracle *)
+
+val set_oracle : env -> Cic.oracle -> env
+
(* Digests *)
val add_digest : env -> DirPath.t -> Cic.vodigest -> env
val lookup_digest : env -> DirPath.t -> Cic.vodigest
diff --git a/checker/indtypes.ml b/checker/indtypes.ml
index 22c843812..bb0db8cfe 100644
--- a/checker/indtypes.ml
+++ b/checker/indtypes.ml
@@ -586,6 +586,8 @@ let check_inductive env kn mib =
Univ.AUContext.repr (Univ.ACumulativityInfo.univ_context cumi)
in
let env = Environ.push_context ind_ctx env in
+ (** Locally set the oracle for further typechecking *)
+ let env0 = Environ.set_oracle env mib.mind_typing_flags.conv_oracle in
(* check mind_record : TODO ? check #constructor = 1 ? *)
(* check mind_finite : always OK *)
(* check mind_ntypes *)
@@ -593,13 +595,13 @@ let check_inductive env kn mib =
user_err Pp.(str "not the right number of packets");
(* check mind_params_ctxt *)
let params = mib.mind_params_ctxt in
- let _ = check_ctxt env params in
+ let _ = check_ctxt env0 params in
(* check mind_nparams *)
if rel_context_nhyps params <> mib.mind_nparams then
user_err Pp.(str "number the right number of parameters");
(* mind_packets *)
(* - check arities *)
- let env_ar = typecheck_arity env params mib.mind_packets in
+ let env_ar = typecheck_arity env0 params mib.mind_packets in
(* - check constructor types *)
Array.iter (typecheck_one_inductive env_ar params mib) mib.mind_packets;
(* check the inferred subtyping relation *)
diff --git a/checker/main.mli b/checker/main.mli
new file mode 100644
index 000000000..e1555ba2e
--- /dev/null
+++ b/checker/main.mli
@@ -0,0 +1,10 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This empty file avoids a race condition that occurs when compiling a .ml file
+ that does not have a corresponding .mli file *)
diff --git a/checker/mod_checking.ml b/checker/mod_checking.ml
index 4357a690e..7685863ea 100644
--- a/checker/mod_checking.ml
+++ b/checker/mod_checking.ml
@@ -26,6 +26,9 @@ let refresh_arity ar =
let check_constant_declaration env kn cb =
Flags.if_verbose Feedback.msg_notice (str " checking cst:" ++ prcon kn);
+ (** Locally set the oracle for further typechecking *)
+ let oracle = env.env_conv_oracle in
+ let env = Environ.set_oracle env cb.const_typing_flags.conv_oracle in
(** [env'] contains De Bruijn universe variables *)
let env' =
match cb.const_universes with
@@ -53,8 +56,12 @@ let check_constant_declaration env kn cb =
conv_leq envty j ty)
| None -> ()
in
- if constant_is_polymorphic cb then add_constant kn cb env
- else add_constant kn cb env'
+ let env =
+ if constant_is_polymorphic cb then add_constant kn cb env
+ else add_constant kn cb env'
+ in
+ (** Reset the value of the oracle *)
+ Environ.set_oracle env oracle
(** {6 Checking modules } *)
diff --git a/checker/print.mli b/checker/print.mli
new file mode 100644
index 000000000..3b2715de9
--- /dev/null
+++ b/checker/print.mli
@@ -0,0 +1,11 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Cic
+
+val print_pure_constr : constr -> unit
diff --git a/checker/reduction.ml b/checker/reduction.ml
index 9b8eac04c..d4b258f58 100644
--- a/checker/reduction.ml
+++ b/checker/reduction.ml
@@ -276,11 +276,29 @@ let in_whnf (t,stk) =
| (FFlex _ | FProd _ | FEvar _ | FInd _ | FAtom _ | FRel _ | FProj _) -> true
| FLOCKED -> assert false
-let oracle_order fl1 fl2 =
- match fl1,fl2 with
- ConstKey c1, ConstKey c2 -> (*height c1 > height c2*)false
- | _, ConstKey _ -> true
- | _ -> false
+let default_level = Level 0
+
+let get_strategy { var_opacity; cst_opacity } = function
+ | VarKey id ->
+ (try Names.Id.Map.find id var_opacity
+ with Not_found -> default_level)
+ | ConstKey (c, _) ->
+ (try Names.Cmap.find c cst_opacity
+ with Not_found -> default_level)
+ | RelKey _ -> Expand
+
+let oracle_order infos l2r k1 k2 =
+ let o = Closure.oracle_of_infos infos in
+ match get_strategy o k1, get_strategy o k2 with
+ | Expand, Expand -> l2r
+ | Expand, (Opaque | Level _) -> true
+ | (Opaque | Level _), Expand -> false
+ | Opaque, Opaque -> l2r
+ | Level _, Opaque -> true
+ | Opaque, Level _ -> false
+ | Level n1, Level n2 ->
+ if Int.equal n1 n2 then l2r
+ else n1 < n2
let unfold_projection infos p c =
let pb = lookup_projection p (infos_env infos) in
@@ -339,7 +357,7 @@ and eqappr univ cv_pb infos (lft1,st1) (lft2,st2) =
with NotConvertible ->
(* else the oracle tells which constant is to be expanded *)
let (app1,app2) =
- if oracle_order fl1 fl2 then
+ if oracle_order infos false fl1 fl2 then
match unfold_reference infos fl1 with
| Some def1 -> ((lft1, whd_stack infos def1 v1), appr2)
| None ->
diff --git a/checker/validate.ml b/checker/validate.ml
index 820040587..2624e6d49 100644
--- a/checker/validate.ml
+++ b/checker/validate.ml
@@ -49,8 +49,6 @@ let (/) (ctx:error_context) s : error_context = s::ctx
exception ValidObjError of string * error_context * Obj.t
let fail ctx o s = raise (ValidObjError(s,ctx,o))
-type func = error_context -> Obj.t -> unit
-
(* Check that object o is a block with tag t *)
let val_tag t ctx o =
if Obj.is_block o && Obj.tag o = t then ()
diff --git a/checker/validate.mli b/checker/validate.mli
new file mode 100644
index 000000000..7eed692a0
--- /dev/null
+++ b/checker/validate.mli
@@ -0,0 +1,9 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+val validate : bool -> Values.value -> 'a -> unit
diff --git a/checker/values.ml b/checker/values.ml
index 5a371164c..313067cb6 100644
--- a/checker/values.ml
+++ b/checker/values.ml
@@ -13,7 +13,7 @@
To ensure this file is up-to-date, 'make' now compares the md5 of cic.mli
with a copy we maintain here:
-MD5 56ac4cade33eff3d26ed5cdadb580c7e checker/cic.mli
+MD5 483493b20fe91cc1bea4350a2db2f82d checker/cic.mli
*)
@@ -70,6 +70,8 @@ let v_map vk vd =
let v_hset v = v_map Int (v_set v)
let v_hmap vk vd = v_map Int (v_map vk vd)
+let v_pred v = v_pair v_bool (v_set v)
+
(* lib/future *)
let v_computation f =
Annot ("Future.computation",
@@ -199,6 +201,17 @@ let v_lazy_constr =
let v_impredicative_set = v_enum "impr-set" 2
let v_engagement = v_impredicative_set
+let v_conv_level =
+ v_sum "conv_level" 2 [|[|Int|]|]
+
+let v_oracle =
+ v_tuple "oracle" [|
+ v_map v_id v_conv_level;
+ v_hmap v_cst v_conv_level;
+ v_pred v_id;
+ v_pred v_cst;
+ |]
+
let v_pol_arity =
v_tuple "polymorphic_arity" [|List(Opt v_level);v_univ|]
@@ -213,7 +226,7 @@ let v_projbody =
v_constr|]
let v_typing_flags =
- v_tuple "typing_flags" [|v_bool; v_bool|]
+ v_tuple "typing_flags" [|v_bool; v_bool; v_oracle|]
let v_const_univs = v_sum "constant_universes" 0 [|[|v_context_set|]; [|v_abs_context|]|]
@@ -372,22 +385,3 @@ let v_lib =
let v_opaques = Array (v_computation v_constr)
let v_univopaques =
Opt (Tuple ("univopaques",[|Array (v_computation v_context_set);v_context_set;v_bool|]))
-
-(** Registering dynamic values *)
-
-module IntOrd =
-struct
- type t = int
- let compare (x : t) (y : t) = compare x y
-end
-
-module IntMap = Map.Make(IntOrd)
-
-let dyn_table : value IntMap.t ref = ref IntMap.empty
-
-let register_dyn name t =
- dyn_table := IntMap.add name t !dyn_table
-
-let find_dyn name =
- try IntMap.find name !dyn_table
- with Not_found -> Any
diff --git a/checker/values.mli b/checker/values.mli
new file mode 100644
index 000000000..aad8fd5f4
--- /dev/null
+++ b/checker/values.mli
@@ -0,0 +1,26 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+type value =
+ | Any
+ | Fail of string
+ | Tuple of string * value array
+ | Sum of string * int * value array array
+ | Array of value
+ | List of value
+ | Opt of value
+ | Int
+ | String
+ | Annot of string * value
+ | Dyn
+
+val v_univopaques : value
+val v_libsum : value
+val v_lib : value
+val v_opaques : value
+val v_stm_seg : value
diff --git a/checker/votour.ml b/checker/votour.ml
index 77c9999c4..8cb97a2b1 100644
--- a/checker/votour.ml
+++ b/checker/votour.ml
@@ -210,7 +210,6 @@ let access_list v o pos =
let access_block o = match Repr.repr o with
| BLOCK (tag, os) -> (tag, os)
| _ -> raise Exit
-let access_int o = match Repr.repr o with INT i -> i | _ -> raise Exit
(** raises Exit if the object has not the expected structure *)
exception Forbidden
@@ -249,8 +248,7 @@ let rec get_children v o pos = match v with
|Dyn ->
begin match Repr.repr o with
| BLOCK (0, [|id; o|]) ->
- let n = access_int id in
- let tpe = find_dyn n in
+ let tpe = Any in
[|(Int, id, 0 :: pos); (tpe, o, 1 :: pos)|]
| _ -> raise Exit
end
@@ -395,7 +393,7 @@ let visit_vo f =
| None -> ()
done
-let main =
+let () =
if not !Sys.interactive then
Arg.parse [] visit_vo
("votour: guided tour of a Coq .vo or .vi file\n"^
diff --git a/checker/votour.mli b/checker/votour.mli
new file mode 100644
index 000000000..e1555ba2e
--- /dev/null
+++ b/checker/votour.mli
@@ -0,0 +1,10 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This empty file avoids a race condition that occurs when compiling a .ml file
+ that does not have a corresponding .mli file *)
diff --git a/clib/clib.mllib b/clib/clib.mllib
index bd5ddb152..0b5d9826f 100644
--- a/clib/clib.mllib
+++ b/clib/clib.mllib
@@ -12,6 +12,7 @@ CString
CStack
Int
+Range
HMap
Bigint
diff --git a/clib/range.ml b/clib/range.ml
new file mode 100644
index 000000000..86a078633
--- /dev/null
+++ b/clib/range.ml
@@ -0,0 +1,91 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+type 'a tree =
+| Leaf of 'a
+| Node of 'a * 'a tree * 'a tree
+
+type 'a t = Nil | Cons of int * 'a tree * 'a t
+
+let oob () = invalid_arg "index out of bounds"
+
+let empty = Nil
+
+let cons x l = match l with
+| Cons (h1, t1, Cons (h2, t2, rem)) ->
+ if Int.equal h1 h2 then Cons (1 + h1 + h2, Node (x, t1, t2), rem)
+ else Cons (1, Leaf x, l)
+| _ -> Cons (1, Leaf x, l)
+
+let is_empty = function
+| Nil -> true
+| _ -> false
+
+let rec tree_get h t i = match t with
+| Leaf x ->
+ if i = 0 then x else oob ()
+| Node (x, t1, t2) ->
+ if i = 0 then x
+ else
+ let h = h / 2 in
+ if i <= h then tree_get h t1 (i - 1) else tree_get h t2 (i - h - 1)
+
+let rec get l i = match l with
+| Nil -> oob ()
+| Cons (h, t, rem) ->
+ if i < h then tree_get h t i else get rem (i - h)
+
+let length l =
+ let rec length accu = function
+ | Nil -> accu
+ | Cons (h, _, l) -> length (h + accu) l
+ in
+ length 0 l
+
+let rec tree_map f = function
+| Leaf x -> Leaf (f x)
+| Node (x, t1, t2) -> Node (f x, tree_map f t1, tree_map f t2)
+
+let rec map f = function
+| Nil -> Nil
+| Cons (h, t, l) -> Cons (h, tree_map f t, map f l)
+
+let rec tree_fold_left f accu = function
+| Leaf x -> f accu x
+| Node (x, t1, t2) ->
+ tree_fold_left f (tree_fold_left f (f accu x) t1) t2
+
+let rec fold_left f accu = function
+| Nil -> accu
+| Cons (_, t, l) -> fold_left f (tree_fold_left f accu t) l
+
+let rec tree_fold_right f t accu = match t with
+| Leaf x -> f x accu
+| Node (x, t1, t2) ->
+ f x (tree_fold_right f t1 (tree_fold_right f t2 accu))
+
+let rec fold_right f l accu = match l with
+| Nil -> accu
+| Cons (_, t, l) -> tree_fold_right f t (fold_right f l accu)
+
+let hd = function
+| Nil -> failwith "hd"
+| Cons (_, Leaf x, _) -> x
+| Cons (_, Node (x, _, _), _) -> x
+
+let tl = function
+| Nil -> failwith "tl"
+| Cons (_, Leaf _, l) -> l
+| Cons (h, Node (_, t1, t2), l) ->
+ let h = h / 2 in
+ Cons (h, t1, Cons (h, t2, l))
+
+let rec skipn n l =
+ if n = 0 then l
+ else if is_empty l then failwith "List.skipn"
+ else skipn (pred n) (tl l)
diff --git a/clib/range.mli b/clib/range.mli
new file mode 100644
index 000000000..ae7684ffa
--- /dev/null
+++ b/clib/range.mli
@@ -0,0 +1,37 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Skewed lists
+
+ This is a purely functional datastructure isomorphic to usual lists, except
+ that it features a O(log n) lookup while preserving the O(1) cons operation.
+
+*)
+
+(** {5 Constructors} *)
+
+type +'a t
+
+val empty : 'a t
+val cons : 'a -> 'a t -> 'a t
+
+(** {5 List operations} *)
+
+val is_empty : 'a t -> bool
+val length : 'a t -> int
+val map : ('a -> 'b) -> 'a t -> 'b t
+val fold_left : ('a -> 'b -> 'a) -> 'a -> 'b t -> 'a
+val fold_right : ('a -> 'b -> 'b) -> 'a t -> 'b -> 'b
+val hd : 'a t -> 'a
+val tl : 'a t -> 'a t
+
+val skipn : int -> 'a t -> 'a t
+
+(** {5 Indexing operations} *)
+
+val get : 'a t -> int -> 'a
diff --git a/configure.ml b/configure.ml
index a3644982c..06a7dd822 100644
--- a/configure.ml
+++ b/configure.ml
@@ -767,14 +767,16 @@ let check_lablgtk_version src dir = match src with
let v, _ = tryrun camlexec.find ["query"; "-format"; "%v"; "lablgtk2"] in
try
let vi = List.map s2i (numeric_prefix_list v) in
- if vi = [2; 18; 0] then
+ if vi < [2; 16; 0] then
+ (false, v)
+ else if vi < [2; 18; 3] then
begin
- (* Version 2.18.3 is known to report incorrectly as 2.18.0 *)
- printf "Warning: could not check the version of lablgtk2.\nMake sure your version is at least 2.18.3.\n";
+ (* Version 2.18.3 is known to report incorrectly as 2.18.0, and Launchpad packages report as version 2.16.0 due to a misconfigured META file; see https://bugs.launchpad.net/ubuntu/+source/lablgtk2/+bug/1577236 *)
+ printf "Warning: Your installed lablgtk reports as %s.\n It is possible that the installed version is actually more recent\n but reports an incorrect version. If the installed version is\n actually more recent than 2.18.3, that's fine; if it is not,\n CoqIDE will compile but may be very unstable.\n" v;
(true, "an unknown version")
end
else
- ([2; 18; 3] <= vi, v)
+ (true, v)
with _ -> (false, v)
let pr_ide = function No -> "no" | Byte -> "only bytecode" | Opt -> "native"
@@ -864,13 +866,6 @@ let strip =
if strip = "" then "strip" else strip
end
-(** * md5sum command *)
-
-let md5sum =
- select_command "Don’t know how to compute MD5 checksums…" [
- "md5sum", [], [ "--version" ];
- "md5", ["-q"], [ "-s" ; "''" ];
- ]
(** * Documentation : do we have latex, hevea, ... *)
@@ -1249,8 +1244,6 @@ let write_makefile f =
pr "# Unix systems and profiling: true\n";
pr "# Unix systems and no profiling: strip\n";
pr "STRIP=%s\n\n" strip;
- pr "#the command md5sum\n";
- pr "MD5SUM=%s\n\n" md5sum;
pr "# LablGTK\n";
pr "COQIDEINCLUDES=%s\n\n" !lablgtkincludes;
pr "# CoqIde (no/byte/opt)\n";
diff --git a/dev/base_include b/dev/base_include
index 2a4ad4a15..472c0c605 100644
--- a/dev/base_include
+++ b/dev/base_include
@@ -52,7 +52,7 @@
#install_printer ppvblock;;
#install_printer (* bigint *) ppbigint;;
#install_printer (* loc *) pploc;;
-#install_printer (* substitution *) prsubst;;
+#install_printer (* substitution *) ppsubst;;
(* Open main files *)
diff --git a/dev/ci/README.md b/dev/ci/README.md
index f4423558c..bb13587e9 100644
--- a/dev/ci/README.md
+++ b/dev/ci/README.md
@@ -103,6 +103,8 @@ The process to merge your PR is then to submit PRs to the external
development repositories, merge the latter first (if the fixes are
backward-compatible), drop the overlay commit and merge the PR on Coq then.
+See also [`test-suite/README.md`](/test-suite/README.md) for information about adding new tests to the test-suite.
+
Travis specific information
---------------------------
diff --git a/dev/ci/user-overlays/06482-ppedrot-check-poly-effects.sh b/dev/ci/user-overlays/06482-ppedrot-check-poly-effects.sh
new file mode 100644
index 000000000..78789a6fc
--- /dev/null
+++ b/dev/ci/user-overlays/06482-ppedrot-check-poly-effects.sh
@@ -0,0 +1,4 @@
+if [ "$TRAVIS_PULL_REQUEST" = "6483" ] || [ "$TRAVIS_BRANCH" = "check-poly-effects" ]; then
+ HoTT_CI_BRANCH=check-poly-effects
+ HoTT_CI_GITURL=https://github.com/ppedrot/HoTT.git
+fi
diff --git a/dev/db b/dev/db
index 24ae3957e..2f8c13485 100644
--- a/dev/db
+++ b/dev/db
@@ -1,37 +1,67 @@
source core.dbg
load_printer top_printers.cmo
+install_printer Top_printers.pP
install_printer Top_printers.ppfuture
-
install_printer Top_printers.ppid
-install_printer Top_printers.ppidset
-install_printer Top_printers.ppevar
-install_printer Top_printers.ppevarsubst
-install_printer Top_printers.ppexistentialfilter
-install_printer Top_printers.ppexistentialset
-install_printer Top_printers.ppintset
install_printer Top_printers.pplab
-install_printer Top_printers.ppdir
install_printer Top_printers.ppmbid
+install_printer Top_printers.ppdir
install_printer Top_printers.ppmp
-install_printer Top_printers.ppkn
install_printer Top_printers.ppcon
-install_printer Top_printers.ppwf_paths
+install_printer Top_printers.ppproj
+install_printer Top_printers.ppkn
install_printer Top_printers.ppmind
+install_printer Top_printers.ppind
install_printer Top_printers.ppsp
install_printer Top_printers.ppqualid
install_printer Top_printers.ppclindex
-install_printer Top_printers.ppbigint
-install_printer Top_printers.pp_transparent_state
-
-install_printer Top_printers.pppattern
-install_printer Top_printers.ppglob_constr
-
+install_printer Top_printers.ppscheme
+install_printer Top_printers.ppwf_paths
+install_printer Top_printers.ppevar
install_printer Top_printers.ppconstr
+install_printer Top_printers.ppsconstr
install_printer Top_printers.ppeconstr
+install_printer Top_printers.ppconstr_expr
+install_printer Top_printers.ppglob_constr
+install_printer Top_printers.pppattern
+install_printer Top_printers.ppfconstr
+install_printer Top_printers.ppbigint
+install_printer Top_printers.ppintset
+install_printer Top_printers.ppidset
+install_printer Top_printers.ppidmapgen
+install_printer Top_printers.ppididmap
+install_printer Top_printers.ppconstrunderbindersidmap
+install_printer Top_printers.ppevarsubst
+install_printer Top_printers.ppunbound_ltac_var_map
+install_printer Top_printers.ppclosure
+install_printer Top_printers.ppclosedglobconstr
+install_printer Top_printers.ppclosedglobconstridmap
+install_printer Top_printers.ppglobal
+install_printer Top_printers.ppconst
+install_printer Top_printers.ppvar
+install_printer Top_printers.ppj
+install_printer Top_printers.ppsubst
+install_printer Top_printers.ppdelta
+install_printer Top_printers.pp_idpred
+install_printer Top_printers.pp_cpred
+install_printer Top_printers.pp_transparent_state
+install_printer Top_printers.pp_stack_t
+install_printer Top_printers.pp_cst_stack_t
+install_printer Top_printers.pp_state_t
+install_printer Top_printers.ppmetas
+install_printer Top_printers.ppevm
+install_printer Top_printers.ppexistentialset
+install_printer Top_printers.ppexistentialfilter
+install_printer Top_printers.ppclenv
+install_printer Top_printers.ppgoalgoal
+install_printer Top_printers.ppgoal
+install_printer Top_printers.pphintdb
+install_printer Top_printers.ppproofview
+install_printer Top_printers.ppopenconstr
+install_printer Top_printers.pproof
install_printer Top_printers.ppuni
-install_printer Top_printers.ppuniverses
-install_printer Top_printers.ppconstraints
+install_printer Top_printers.ppuni_level
install_printer Top_printers.ppuniverse_set
install_printer Top_printers.ppuniverse_instance
install_printer Top_printers.ppuniverse_context
@@ -40,34 +70,19 @@ install_printer Top_printers.ppuniverse_subst
install_printer Top_printers.ppuniverse_opt_subst
install_printer Top_printers.ppuniverse_level_subst
install_printer Top_printers.ppevar_universe_context
+install_printer Top_printers.ppconstraints
+install_printer Top_printers.ppuniverseconstraints
+install_printer Top_printers.ppuniverse_context_future
install_printer Top_printers.ppcumulativity_info
install_printer Top_printers.ppabstract_cumulativity_info
-install_printer Top_printers.pptype
-install_printer Top_printers.ppj
-install_printer Top_printers.ppenv
+install_printer Top_printers.ppuniverses
install_printer Top_printers.ppnamedcontextval
-install_printer Top_printers.pp_stack_t
-install_printer Top_printers.pp_cst_stack_t
-
-install_printer Top_printers.ppmetas
-install_printer Top_printers.ppevm
-install_printer Top_printers.ppgoalgoal
-install_printer Top_printers.ppgoal
-install_printer Top_printers.ppproofview
-install_printer Top_printers.pphintdb
-
+install_printer Top_printers.ppenv
install_printer Top_printers.pptac
install_printer Top_printers.ppobj
install_printer Top_printers.pploc
-install_printer Top_printers.prsubst
-install_printer Top_printers.prdelta
-install_printer Top_printers.ppfconstr
+install_printer Top_printers.pp_argument_type
+install_printer Top_printers.pp_generic_argument
install_printer Top_printers.ppgenarginfo
install_printer Top_printers.ppgenargargt
install_printer Top_printers.ppist
-install_printer Top_printers.ppconstrunderbindersidmap
-install_printer Top_printers.ppunbound_ltac_var_map
-install_printer Top_printers.ppididmap
-install_printer Top_printers.ppidmapgen
-install_printer Top_printers.ppclosure
-install_printer Top_printers.ppclosedglobconstr
diff --git a/dev/doc/setup.txt b/dev/doc/setup.txt
index 26f3d0ddc..0003a2c21 100644
--- a/dev/doc/setup.txt
+++ b/dev/doc/setup.txt
@@ -58,30 +58,12 @@ behave as expected.
A note about rlwrap
-------------------
-Running "coqtop" under "rlwrap" is possible, but (on Debian) there is a catch. If you try:
-
- cd ~/git/coq
- rlwrap bin/coqtop
-
-you will get an error:
+When using "rlwrap coqtop" make sure the version of rlwrap is at least
+0.42, otherwise you will get
rlwrap: error: Couldn't read completions from /usr/share/rlwrap/completions/coqtop: No such file or directory
-This is a known issue:
-
- https://bugs.debian.org/cgi-bin/bugreport.cgi?bug=779692
-
-It was fixed upstream in version 0.42, and in a Debian package that, at the time of writing, is not part of Debian stable/testing/sid archives but only of Debian experimental.
-
- https://packages.debian.org/experimental/rlwrap
-
-The quick solution is to grab it from there, since it installs fine on Debian stable (jessie).
-
- cd /tmp
- wget http://ftp.us.debian.org/debian/pool/main/r/rlwrap/rlwrap_0.42-1_amd64.deb
- sudo dpkg -i rlwrap_0.42-1_amd64.deb
-
-After that, "rlwrap" works fine with "coqtop".
+If this happens either update or use an alternate readline wrapper like "ledit".
How to install and configure Merlin (for Emacs)
diff --git a/dev/include b/dev/include
index 0d34595f4..b982f4c9f 100644
--- a/dev/include
+++ b/dev/include
@@ -36,7 +36,6 @@
#install_printer (* constraints *) ppconstraints;;
#install_printer (* univ constraints *) ppuniverseconstraints;;
#install_printer (* universe *) ppuni;;
-#install_printer (* universes *) ppuniverse;;
#install_printer (* universes *) ppuniverses;;
#install_printer (* univ level *) ppuni_level;;
#install_printer (* univ context *) ppuniverse_context;;
diff --git a/dev/ocamldebug-coq.run b/dev/ocamldebug-coq.run
index 3850c05fd..3cbccab44 100644
--- a/dev/ocamldebug-coq.run
+++ b/dev/ocamldebug-coq.run
@@ -17,7 +17,7 @@ export CAML_LD_LIBRARY_PATH=$COQTOP/kernel/byterun:$CAML_LD_LIBRARY_PATH
exec $OCAMLDEBUG \
-I $CAMLP4LIB -I +threads \
-I $COQTOP \
- -I $COQTOP/config -I $COQTOP/printing -I $COQTOP/grammar \
+ -I $COQTOP/config -I $COQTOP/printing -I $COQTOP/grammar -I $COQTOP/clib \
-I $COQTOP/lib -I $COQTOP/intf -I $COQTOP/kernel -I $COQTOP/kernel/byterun \
-I $COQTOP/library -I $COQTOP/engine \
-I $COQTOP/pretyping -I $COQTOP/parsing -I $COQTOP/vernac \
diff --git a/dev/set_raw_db b/dev/set_raw_db
deleted file mode 100644
index 5caff7e5d..000000000
--- a/dev/set_raw_db
+++ /dev/null
@@ -1 +0,0 @@
-install_printer Top_printers.ppconstrdb
diff --git a/dev/top_printers.ml b/dev/top_printers.ml
index ff3825787..af38ce4b8 100644
--- a/dev/top_printers.ml
+++ b/dev/top_printers.ml
@@ -50,13 +50,13 @@ let ppqualid qid = pp(pr_qualid qid)
let ppclindex cl = pp(Classops.pr_cl_index cl)
let ppscheme k = pp (Ind_tables.pr_scheme_kind k)
-let pprecarg = function
+let prrecarg = function
| Declarations.Norec -> str "Norec"
| Declarations.Mrec (mind,i) ->
str "Mrec[" ++ MutInd.print mind ++ pr_comma () ++ int i ++ str "]"
| Declarations.Imbr (mind,i) ->
str "Imbr[" ++ MutInd.print mind ++ pr_comma () ++ int i ++ str "]"
-let ppwf_paths x = pp (Rtree.pp_tree pprecarg x)
+let ppwf_paths x = pp (Rtree.pp_tree prrecarg x)
(* term printers *)
let envpp pp = let sigma,env = Pfedit.get_current_context () in pp env sigma
@@ -65,8 +65,6 @@ let ppevar evk = pp (Evar.print evk)
let ppconstr x = pp (Termops.print_constr (EConstr.of_constr x))
let ppeconstr x = pp (Termops.print_constr x)
let ppconstr_expr x = pp (Ppconstr.pr_constr_expr x)
-let ppconstrdb x = pp(Flags.with_option rawdebug Termops.print_constr (EConstr.of_constr x))
-let ppterm = ppconstr
let ppsconstr x = ppconstr (Mod_subst.force_constr x)
let ppconstr_univ x = Constrextern.with_universes ppconstr x
let ppglob_constr = (fun x -> pp(pr_lglob_constr_env (Global.env()) x))
@@ -111,7 +109,7 @@ let prconstrunderbindersidmap = pridmap (fun _ (l,c) ->
let ppconstrunderbindersidmap l = pp (prconstrunderbindersidmap l)
let ppunbound_ltac_var_map l = ppidmap (fun _ arg ->
- str"<genarg:" ++ pr_argument_type(genarg_tag arg) ++ str">")
+ str"<genarg:" ++ pr_argument_type(genarg_tag arg) ++ str">") l
open Ltac_pretype
let rec pr_closure {idents=idents;typed=typed;untyped=untyped} =
@@ -149,8 +147,8 @@ let genppj f j = let (c,t) = f j in (c ++ str " : " ++ t)
let ppj j = pp (genppj (envpp pr_ljudge_env) j)
-let prsubst s = pp (Mod_subst.debug_pr_subst s)
-let prdelta s = pp (Mod_subst.debug_pr_delta s)
+let ppsubst s = pp (Mod_subst.debug_pr_subst s)
+let ppdelta s = pp (Mod_subst.debug_pr_delta s)
let pp_idpred s = pp (pr_idpred s)
let pp_cpred s = pp (pr_cpred s)
@@ -200,9 +198,8 @@ let pppftreestate p = pp(print_pftreestate p)
let pproof p = pp(Proof.pr_proof p)
-let ppuni u = pp(pr_uni u)
+let ppuni u = pp(Universe.pr u)
let ppuni_level u = pp (Level.pr u)
-let ppuniverse u = pp (str"[" ++ Universe.pr u ++ str"]")
let prlev = Universes.pr_with_global_universes
let ppuniverse_set l = pp (LSet.pr prlev l)
diff --git a/dev/top_printers.mli b/dev/top_printers.mli
new file mode 100644
index 000000000..7b5e4a0b6
--- /dev/null
+++ b/dev/top_printers.mli
@@ -0,0 +1,173 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Printers for the ocaml toplevel. *)
+
+val pp : Pp.t -> unit
+val pP : Pp.t -> unit (* with surrounding box *)
+
+val ppfuture : 'a Future.computation -> unit
+
+val ppid : Names.Id.t -> unit
+val pplab : Names.Label.t -> unit
+val ppmbid : Names.MBId.t -> unit
+val ppdir : Names.DirPath.t -> unit
+val ppmp : Names.ModPath.t -> unit
+val ppcon : Names.Constant.t -> unit
+val ppproj : Names.Projection.t -> unit
+val ppkn : Names.KerName.t -> unit
+val ppmind : Names.MutInd.t -> unit
+val ppind : Names.inductive -> unit
+
+val ppsp : Libnames.full_path -> unit
+val ppqualid : Libnames.qualid -> unit
+
+val ppclindex : Classops.cl_index -> unit
+
+val ppscheme : 'a Ind_tables.scheme_kind -> unit
+
+val prrecarg : Declarations.recarg -> Pp.t
+val ppwf_paths : Declarations.recarg Rtree.t -> unit
+
+val pr_evar : Evar.t -> Pp.t
+val ppevar : Evar.t -> unit
+
+(* Multiple printers for Constr.t *)
+val ppconstr : Constr.t -> unit (* by Termops printer *)
+val ppconstr_univ : Constr.t -> unit
+
+(* Extern as type *)
+val pptype : Constr.types -> unit
+
+val ppsconstr : Constr.constr Mod_subst.substituted -> unit
+val ppeconstr : EConstr.constr -> unit (* Termops printer *)
+val ppconstr_expr : Constrexpr.constr_expr -> unit
+val ppglob_constr : 'a Glob_term.glob_constr_g -> unit
+val pppattern : Pattern.constr_pattern -> unit
+val ppfconstr : CClosure.fconstr -> unit
+
+val ppbigint : Bigint.bigint -> unit
+
+val ppintset : Int.Set.t -> unit
+val ppidset : Names.Id.Set.t -> unit
+
+val pridmap : (Names.Id.Map.key -> 'a -> Pp.t) -> 'a Names.Id.Map.t -> Pp.t
+val ppidmap : (Names.Id.Map.key -> 'a -> Pp.t) -> 'a Names.Id.Map.t -> unit
+
+val pridmapgen : 'a Names.Id.Map.t -> Pp.t
+val ppidmapgen : 'a Names.Id.Map.t -> unit
+
+val prididmap : Names.Id.t Names.Id.Map.t -> Pp.t
+val ppididmap : Names.Id.t Names.Id.Map.t -> unit
+
+val prconstrunderbindersidmap :
+ (Names.Id.t list * EConstr.constr) Names.Id.Map.t -> Pp.t
+val ppconstrunderbindersidmap :
+ (Names.Id.t list * EConstr.constr) Names.Id.Map.t -> unit
+
+val ppevarsubst :
+ (Constr.t * Constr.t option * Names.Id.Map.key) list Names.Id.Map.t -> unit
+
+val ppunbound_ltac_var_map :
+ 'a Genarg.generic_argument Names.Id.Map.t -> unit
+
+val pr_closure : Ltac_pretype.closure -> Pp.t
+val pr_closed_glob_constr_idmap :
+ Ltac_pretype.closed_glob_constr Names.Id.Map.t -> Pp.t
+val pr_closed_glob_constr : Ltac_pretype.closed_glob_constr -> Pp.t
+val ppclosure : Ltac_pretype.closure -> unit
+val ppclosedglobconstr : Ltac_pretype.closed_glob_constr -> unit
+val ppclosedglobconstridmap :
+ Ltac_pretype.closed_glob_constr Names.Id.Map.t -> unit
+
+val ppglobal : Globnames.global_reference -> unit
+
+val ppconst :
+ Names.KerName.t * (Constr.constr, 'a) Environ.punsafe_judgment -> unit
+val ppvar : Names.Id.t * Constr.constr -> unit
+
+val genppj : ('a -> Pp.t * Pp.t) -> 'a -> Pp.t
+val ppj : EConstr.unsafe_judgment -> unit
+
+val ppsubst : Mod_subst.substitution -> unit
+val ppdelta : Mod_subst.delta_resolver -> unit
+
+val pp_idpred : Names.Id.Pred.t -> unit
+val pp_cpred : Names.Cpred.t -> unit
+val pp_transparent_state : Names.transparent_state -> unit
+
+val pp_stack_t : Constr.t Reductionops.Stack.t -> unit
+val pp_cst_stack_t : Reductionops.Cst_stack.t -> unit
+val pp_state_t : Reductionops.state -> unit
+
+val ppmetas : Evd.Metaset.t -> unit
+val ppevm : Evd.evar_map -> unit
+val ppevmall : Evd.evar_map -> unit
+
+val pr_existentialset : Evar.Set.t -> Pp.t
+val ppexistentialset : Evar.Set.t -> unit
+
+val ppexistentialfilter : Evd.Filter.t -> unit
+
+val ppclenv : Clenv.clausenv -> unit
+
+val ppgoalgoal : Goal.goal -> unit
+
+val ppgoal : Proof_type.goal Evd.sigma -> unit
+(* also print evar map *)
+val ppgoalsigma : Proof_type.goal Evd.sigma -> unit
+
+val pphintdb : Hints.Hint_db.t -> unit
+val ppproofview : Proofview.proofview -> unit
+val ppopenconstr : Evd.open_constr -> unit
+
+val pproof : Proof.t -> unit
+
+(* Universes *)
+val ppuni : Univ.Universe.t -> unit
+val ppuni_level : Univ.Level.t -> unit (* raw *)
+val prlev : Univ.Level.t -> Pp.t (* with global names (does this work?) *)
+val ppuniverse_set : Univ.LSet.t -> unit
+val ppuniverse_instance : Univ.Instance.t -> unit
+val ppuniverse_context : Univ.UContext.t -> unit
+val ppuniverse_context_set : Univ.ContextSet.t -> unit
+val ppuniverse_subst : Univ.universe_subst -> unit
+val ppuniverse_opt_subst : Universes.universe_opt_subst -> unit
+val ppuniverse_level_subst : Univ.universe_level_subst -> unit
+val ppevar_universe_context : UState.t -> unit
+val ppconstraints : Univ.Constraint.t -> unit
+val ppuniverseconstraints : Universes.Constraints.t -> unit
+val ppuniverse_context_future : Univ.UContext.t Future.computation -> unit
+val ppcumulativity_info : Univ.CumulativityInfo.t -> unit
+val ppabstract_cumulativity_info : Univ.ACumulativityInfo.t -> unit
+val ppuniverses : UGraph.t -> unit
+
+val ppnamedcontextval : Environ.named_context_val -> unit
+val ppenv : Environ.env -> unit
+val ppenvwithcst : Environ.env -> unit
+
+val pptac : Tacexpr.glob_tactic_expr -> unit
+
+val ppobj : Libobject.obj -> unit
+
+(* Some super raw printers *)
+val cast_kind_display : Constr.cast_kind -> string
+val constr_display : Constr.constr -> unit
+val print_pure_constr : Constr.types -> unit
+
+val pploc : Loc.t -> unit
+
+val pp_argument_type : Genarg.argument_type -> unit
+val pp_generic_argument : 'a Genarg.generic_argument -> unit
+
+val prgenarginfo : Geninterp.Val.t -> Pp.t
+val ppgenarginfo : Geninterp.Val.t -> unit
+
+val ppgenargargt : ('a, 'b, 'c) Genarg.ArgT.tag -> unit
+
+val ppist : Geninterp.interp_sign -> unit
diff --git a/doc/refman/RefMan-pro.tex b/doc/refman/RefMan-pro.tex
index 3d4c885b3..6b24fdde7 100644
--- a/doc/refman/RefMan-pro.tex
+++ b/doc/refman/RefMan-pro.tex
@@ -320,10 +320,19 @@ Note that when a focused goal is proved a message is displayed
together with a suggestion about the right bullet or {\tt \}} to
unfocus it or focus the next one.
+\begin{Variants}
+
+\item {\tt {\num}: \{}\\
+This focuses on the $\num^{th}$ subgoal to prove.
+
+\end{Variants}
+
\begin{ErrMsgs}
\item \errindex{This proof is focused, but cannot be unfocused
this way} You are trying to use {\tt \}} but the current subproof
has not been fully solved.
+\item \errindex{No such goal}
+\item \errindex{Brackets only support the single numbered goal selector}
\item see also error message about bullets below.
\end{ErrMsgs}
diff --git a/engine/eConstr.ml b/engine/eConstr.ml
index a65b3941e..9ac16b5b4 100644
--- a/engine/eConstr.ml
+++ b/engine/eConstr.ml
@@ -695,6 +695,10 @@ let cast_rel_context :
type a b. (a,b) eq -> (a, a) Rel.pt -> (b, b) Rel.pt =
fun Refl x -> x
+let cast_rec_decl :
+ type a b. (a,b) eq -> (a, a) Constr.prec_declaration -> (b, b) Constr.prec_declaration =
+ fun Refl x -> x
+
let cast_named_decl :
type a b. (a,b) eq -> (a, a) Named.Declaration.pt -> (b, b) Named.Declaration.pt =
fun Refl x -> x
@@ -817,6 +821,7 @@ let it_mkLambda_or_LetIn t ctx = List.fold_left (fun c d -> mkLambda_or_LetIn d
let push_rel d e = push_rel (cast_rel_decl unsafe_eq d) e
let push_rel_context d e = push_rel_context (cast_rel_context unsafe_eq d) e
+let push_rec_types d e = push_rec_types (cast_rec_decl unsafe_eq d) e
let push_named d e = push_named (cast_named_decl unsafe_eq d) e
let push_named_context d e = push_named_context (cast_named_context unsafe_eq d) e
let push_named_context_val d e = push_named_context_val (cast_named_decl unsafe_eq d) e
diff --git a/engine/eConstr.mli b/engine/eConstr.mli
index 30de748a1..6fa338c73 100644
--- a/engine/eConstr.mli
+++ b/engine/eConstr.mli
@@ -251,6 +251,7 @@ end
val push_rel : rel_declaration -> env -> env
val push_rel_context : rel_context -> env -> env
+val push_rec_types : (t, t) Constr.prec_declaration -> env -> env
val push_named : named_declaration -> env -> env
val push_named_context : named_context -> env -> env
diff --git a/engine/termops.ml b/engine/termops.ml
index a71bdff31..40b3d0d8b 100644
--- a/engine/termops.ml
+++ b/engine/termops.ml
@@ -1463,6 +1463,18 @@ let prod_applist sigma c l =
| _ -> anomaly (Pp.str "Not enough prod's.") in
app [] c l
+let prod_applist_assum sigma n c l =
+ let open EConstr in
+ let rec app n subst c l =
+ if Int.equal n 0 then
+ if l == [] then Vars.substl subst c
+ else anomaly (Pp.str "Not enough arguments.")
+ else match EConstr.kind sigma c, l with
+ | Prod(_,_,c), arg::l -> app (n-1) (arg::subst) c l
+ | LetIn(_,b,_,c), _ -> app (n-1) (Vars.substl subst b::subst) c l
+ | _ -> anomaly (Pp.str "Not enough prod/let's.") in
+ app n [] c l
+
(* Combinators on judgments *)
let on_judgment f j = { uj_val = f j.uj_val; uj_type = f j.uj_type }
diff --git a/engine/termops.mli b/engine/termops.mli
index c1600abe8..a3559a693 100644
--- a/engine/termops.mli
+++ b/engine/termops.mli
@@ -159,8 +159,18 @@ val eta_reduce_head : Evd.evar_map -> constr -> constr
(** Flattens application lists *)
val collapse_appl : Evd.evar_map -> constr -> constr
+(** [prod_applist] [forall (x1:B1;...;xn:Bn), B] [a1...an] @return [B[a1...an]] *)
val prod_applist : Evd.evar_map -> constr -> constr list -> constr
+(** In [prod_applist_assum n c args], [c] is supposed to have the
+ form [∀Γ.c] with [Γ] of length [m] and possibly with let-ins; it
+ returns [c] with the assumptions of [Γ] instantiated by [args] and
+ the local definitions of [Γ] expanded.
+ Note that [n] counts both let-ins and prods, while the length of [args]
+ only counts prods. In other words, varying [n] changes how many
+ trailing let-ins are expanded. *)
+val prod_applist_assum : Evd.evar_map -> int -> constr -> constr list -> constr
+
(** Remove recursively the casts around a term i.e.
[strip_outer_cast (Cast (Cast ... (Cast c, t) ... ))] is [c]. *)
val strip_outer_cast : Evd.evar_map -> constr -> constr
diff --git a/engine/uState.ml b/engine/uState.ml
index 6f2b3c4b2..4b650c9c9 100644
--- a/engine/uState.ml
+++ b/engine/uState.ml
@@ -201,14 +201,18 @@ let process_universe_constraints ctx cstrs =
| None -> user_err Pp.(str "Algebraic universe on the right")
| Some r' ->
if Univ.Level.is_small r' then
- let levels = Univ.Universe.levels l in
- let fold l' local =
- let l = Univ.Universe.make l' in
- if Univ.Level.is_small l' || is_local l' then
- equalize_variables false l l' r r' local
- else raise (Univ.UniverseInconsistency (Univ.Le, l, r, None))
- in
- Univ.LSet.fold fold levels local
+ if not (Univ.Universe.is_levels l)
+ then
+ raise (Univ.UniverseInconsistency (Univ.Le, l, r, None))
+ else
+ let levels = Univ.Universe.levels l in
+ let fold l' local =
+ let l = Univ.Universe.make l' in
+ if Univ.Level.is_small l' || is_local l' then
+ equalize_variables false l l' r r' local
+ else raise (Univ.UniverseInconsistency (Univ.Le, l, r, None))
+ in
+ Univ.LSet.fold fold levels local
else
Univ.enforce_leq l r local
end
diff --git a/ide/config_lexer.mli b/ide/config_lexer.mli
new file mode 100644
index 000000000..0c0c5d1e7
--- /dev/null
+++ b/ide/config_lexer.mli
@@ -0,0 +1,10 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+val print_file : string -> string list Util.String.Map.t -> unit
+val load_file : string -> string list Util.String.Map.t
diff --git a/ide/coq_commands.mli b/ide/coq_commands.mli
new file mode 100644
index 000000000..53026be38
--- /dev/null
+++ b/ide/coq_commands.mli
@@ -0,0 +1,11 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+val tactics : string list list
+val commands : string list list
+val state_preserving : string list
diff --git a/ide/coq_lex.mli b/ide/coq_lex.mli
new file mode 100644
index 000000000..417e0a76f
--- /dev/null
+++ b/ide/coq_lex.mli
@@ -0,0 +1,11 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+val delimit_sentences : (int -> GText.tag -> unit) -> string -> unit
+
+exception Unterminated
diff --git a/ide/coq_lex.mll b/ide/coq_lex.mll
index 8bfd937e3..a7e9003db 100644
--- a/ide/coq_lex.mll
+++ b/ide/coq_lex.mll
@@ -17,7 +17,9 @@
let space = [' ' '\n' '\r' '\t' '\012'] (* '\012' is form-feed *)
-let undotted_sep = '{' | '}' | '-'+ | '+'+ | '*'+
+let number = [ '0'-'9' ]+
+
+let undotted_sep = (number space* ':' space*)? '{' | '}' | '-'+ | '+'+ | '*'+
let dot_sep = '.' (space | eof)
diff --git a/ide/coqide_main.ml4 b/ide/coqide_main.ml4
index 8d99cc3e6..6e330c62b 100644
--- a/ide/coqide_main.ml4
+++ b/ide/coqide_main.ml4
@@ -55,6 +55,8 @@ let os_specific_init () = ()
(** Win32 *)
+IFDEF WIN32 THEN
+
(* On win32, we add the directory of coqide to the PATH at launch-time
(this used to be done in a .bat script). *)
@@ -86,7 +88,6 @@ let reroute_stdout_stderr () =
(* We also provide specific kill and interrupt functions. *)
-IFDEF WIN32 THEN
external win32_kill : int -> unit = "win32_kill"
external win32_interrupt : int -> unit = "win32_interrupt"
let () =
diff --git a/ide/coqide_main.mli b/ide/coqide_main.mli
new file mode 100644
index 000000000..e1555ba2e
--- /dev/null
+++ b/ide/coqide_main.mli
@@ -0,0 +1,10 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This empty file avoids a race condition that occurs when compiling a .ml file
+ that does not have a corresponding .mli file *)
diff --git a/ide/coqide_ui.mli b/ide/coqide_ui.mli
new file mode 100644
index 000000000..9f6fa5635
--- /dev/null
+++ b/ide/coqide_ui.mli
@@ -0,0 +1,10 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+val init : unit -> unit
+val ui_m : GAction.ui_manager
diff --git a/ide/gtk_parsing.ml b/ide/gtk_parsing.ml
index f0575e325..7c0a7495a 100644
--- a/ide/gtk_parsing.ml
+++ b/ide/gtk_parsing.ml
@@ -7,11 +7,7 @@
(************************************************************************)
let underscore = Glib.Utf8.to_unichar "_" ~pos:(ref 0)
-let arobase = Glib.Utf8.to_unichar "@" ~pos:(ref 0)
let prime = Glib.Utf8.to_unichar "'" ~pos:(ref 0)
-let bn = Glib.Utf8.to_unichar "\n" ~pos:(ref 0)
-let space = Glib.Utf8.to_unichar " " ~pos:(ref 0)
-let tab = Glib.Utf8.to_unichar "\t" ~pos:(ref 0)
(* TODO: avoid num and prime at the head of a word *)
@@ -30,17 +26,6 @@ let ends_word (it:GText.iter) =
not (is_word_char c)
)
-
-let inside_word (it:GText.iter) =
- let c = it#char in
- not (starts_word it) &&
- not (ends_word it) &&
- is_word_char c
-
-
-let is_on_word_limit (it:GText.iter) = inside_word it || ends_word it
-
-
let find_word_start (it:GText.iter) =
let rec step_to_start it =
Minilib.log "Find word start";
@@ -72,100 +57,6 @@ let get_word_around (it:GText.iter) =
let stop = find_word_end it in
start,stop
-
-let rec complete_backward w (it:GText.iter) =
- Minilib.log "Complete backward...";
- match it#backward_search w with
- | None -> (Minilib.log "backward_search failed";None)
- | Some (start,stop) ->
- Minilib.log ("complete_backward got a match:"^(string_of_int start#offset)^(string_of_int stop#offset));
- if starts_word start then
- let ne = find_word_end stop in
- if ne#compare stop = 0
- then complete_backward w start
- else Some (start,stop,ne)
- else complete_backward w start
-
-
-let rec complete_forward w (it:GText.iter) =
- Minilib.log "Complete forward...";
- match it#forward_search w with
- | None -> None
- | Some (start,stop) ->
- if starts_word start then
- let ne = find_word_end stop in
- if ne#compare stop = 0 then
- complete_forward w stop
- else Some (stop,stop,ne)
- else complete_forward w stop
-
-
-let find_comment_end (start:GText.iter) =
- let rec find_nested_comment (search_start:GText.iter) (search_end:GText.iter) (comment_end:GText.iter) =
- match (search_start#forward_search ~limit:search_end "(*"),(comment_end#forward_search "*)") with
- | None,_ -> comment_end
- | Some _, None -> raise Not_found
- | Some (_,next_search_start),Some (next_search_end,next_comment_end) ->
- find_nested_comment next_search_start next_search_end next_comment_end
- in
- match start#forward_search "*)" with
- | None -> raise Not_found
- | Some (search_end,comment_end) -> find_nested_comment start search_end comment_end
-
-
-let rec find_string_end (start:GText.iter) =
- let dblquote = int_of_char '"' in
- let rec escaped_dblquote c =
- (c#char = dblquote) && not (escaped_dblquote c#backward_char)
- in
- match start#forward_search "\"" with
- | None -> raise Not_found
- | Some (stop,next_start) ->
- if escaped_dblquote stop#backward_char
- then find_string_end next_start
- else next_start
-
-
-let rec find_next_sentence (from:GText.iter) =
- match (from#forward_search ".") with
- | None -> raise Not_found
- | Some (non_vernac_search_end,next_sentence) ->
- match from#forward_search ~limit:non_vernac_search_end "(*",from#forward_search ~limit:non_vernac_search_end "\"" with
- | None,None ->
- if Glib.Unichar.isspace next_sentence#char || next_sentence#compare next_sentence#forward_char == 0
- then next_sentence else find_next_sentence next_sentence
- | None,Some (_,string_search_start) -> find_next_sentence (find_string_end string_search_start)
- | Some (_,comment_search_start),None -> find_next_sentence (find_comment_end comment_search_start)
- | Some (_,comment_search_start),Some (_,string_search_start) ->
- find_next_sentence (
- if comment_search_start#compare string_search_start < 0
- then find_comment_end comment_search_start
- else find_string_end string_search_start)
-
-
-let find_nearest_forward (cursor:GText.iter) targets =
- let fold_targets acc target =
- match cursor#forward_search target,acc with
- | Some (t_start,_),Some nearest when (t_start#compare nearest < 0) -> Some t_start
- | Some (t_start,_),None -> Some t_start
- | _ -> acc
- in
- match List.fold_left fold_targets None targets with
- | None -> raise Not_found
- | Some nearest -> nearest
-
-
-let find_nearest_backward (cursor:GText.iter) targets =
- let fold_targets acc target =
- match cursor#backward_search target,acc with
- | Some (t_start,_),Some nearest when (t_start#compare nearest > 0) -> Some t_start
- | Some (t_start,_),None -> Some t_start
- | _ -> acc
- in
- match List.fold_left fold_targets None targets with
- | None -> raise Not_found
- | Some nearest -> nearest
-
(** On double-click on a view, select the whole word. This is a workaround for
a deficient word handling in TextView. *)
let fix_double_click self =
diff --git a/ide/gtk_parsing.mli b/ide/gtk_parsing.mli
new file mode 100644
index 000000000..b54f731b3
--- /dev/null
+++ b/ide/gtk_parsing.mli
@@ -0,0 +1,26 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+val fix_double_click :
+ < buffer : < get_iter : [> `INSERT ] -> GText.iter;
+ move_mark : [> `INSERT | `SEL_BOUND ] ->
+ where:GText.iter -> unit;
+ .. >;
+ event : < connect :
+ < button_press :
+ callback:([> `TWO_BUTTON_PRESS ] Gdk.event ->
+ bool) ->
+ 'a;
+ .. >;
+ .. >;
+ .. > ->
+ unit
+val starts_word : GText.iter -> bool
+val ends_word : GText.iter -> bool
+val find_word_start : GText.iter -> GText.iter
+val find_word_end : GText.iter -> GText.iter
diff --git a/ide/ide_slave.mli b/ide/ide_slave.mli
new file mode 100644
index 000000000..e1555ba2e
--- /dev/null
+++ b/ide/ide_slave.mli
@@ -0,0 +1,10 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This empty file avoids a race condition that occurs when compiling a .ml file
+ that does not have a corresponding .mli file *)
diff --git a/ide/macos_prehook.mli b/ide/macos_prehook.mli
new file mode 100644
index 000000000..e1555ba2e
--- /dev/null
+++ b/ide/macos_prehook.mli
@@ -0,0 +1,10 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This empty file avoids a race condition that occurs when compiling a .ml file
+ that does not have a corresponding .mli file *)
diff --git a/ide/nanoPG.mli b/ide/nanoPG.mli
new file mode 100644
index 000000000..3ad8435b5
--- /dev/null
+++ b/ide/nanoPG.mli
@@ -0,0 +1,11 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+val get_documentation : unit -> string
+val init : GWindow.window -> Session.session Wg_Notebook.typed_notebook ->
+ GAction.action_group list -> unit
diff --git a/ide/utf8_convert.mli b/ide/utf8_convert.mli
new file mode 100644
index 000000000..06a131a68
--- /dev/null
+++ b/ide/utf8_convert.mli
@@ -0,0 +1,9 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+val f : string -> string
diff --git a/ide/wg_Find.ml b/ide/wg_Find.ml
index a62ff2de5..cb182465a 100644
--- a/ide/wg_Find.ml
+++ b/ide/wg_Find.ml
@@ -84,8 +84,10 @@ class finder name (view : GText.view) =
method private backward_search starti =
let text = view#buffer#start_iter#get_text ~stop:starti in
let regexp = self#regex in
- try
- let i = Str.search_backward regexp text (String.length text - 1) in
+ let offs = (String.length text - 1) in
+ if offs < 0 then None
+ else try
+ let i = Str.search_backward regexp text offs in
let j = Str.match_end () in
Some(view#buffer#start_iter#forward_chars (b2c text i),
view#buffer#start_iter#forward_chars (b2c text j))
@@ -101,24 +103,33 @@ class finder name (view : GText.view) =
with Not_found -> None
method replace_all () =
- let rec replace_at (iter : GText.iter) =
+ let rec replace_at (iter : GText.iter) ct tot =
let found = self#forward_search iter in
match found with
- | None -> ()
+ | None ->
+ let tot_str = if Int.equal ct tot then "" else " of " ^ string_of_int tot in
+ let occ_str = CString.plural tot "occurrence" in
+ let _ = Ideutils.flash_info ("Replaced " ^ string_of_int ct ^ tot_str ^ " " ^ occ_str) in
+ ()
| Some (start, stop) ->
let text = iter#get_text ~stop:view#buffer#end_iter in
let start_mark = view#buffer#create_mark start in
let stop_mark = view#buffer#create_mark ~left_gravity:false stop in
+ let mod_save = view#buffer#modified in
+ let _ = view#buffer#set_modified false in
let _ = view#buffer#delete_interactive ~start ~stop () in
let iter = view#buffer#get_iter_at_mark (`MARK start_mark) in
- let _ = view#buffer#insert_interactive ~iter (self#replacement text)in
+ let _ = view#buffer#insert_interactive ~iter (self#replacement text) in
+ let edited = view#buffer#modified in
+ let _ = view#buffer#set_modified (edited || mod_save) in
let next = view#buffer#get_iter_at_mark (`MARK stop_mark) in
let () = view#buffer#delete_mark (`MARK start_mark) in
let () = view#buffer#delete_mark (`MARK stop_mark) in
- replace_at next
+ let next_ct = if edited then ct + 1 else ct in
+ replace_at next next_ct (tot + 1)
in
let () = view#buffer#begin_user_action () in
- let () = replace_at view#buffer#start_iter in
+ let () = replace_at view#buffer#start_iter 0 0 in
view#buffer#end_user_action ()
method private set_not_found () =
@@ -130,22 +141,52 @@ class finder name (view : GText.view) =
method private set_normal () =
find_entry#misc#modify_base [`NORMAL, `NAME "white"]
- method private find_from backward (starti : GText.iter) =
+ method private find_from backward ?(wrapped=false) (starti : GText.iter) =
let found =
if backward then self#backward_search starti
else self#forward_search starti in
match found with
| None ->
if not backward && not (starti#equal view#buffer#start_iter) then
- self#find_from backward view#buffer#start_iter
+ self#find_from backward ~wrapped:true view#buffer#start_iter
else if backward && not (starti#equal view#buffer#end_iter) then
- self#find_from backward view#buffer#end_iter
+ self#find_from backward ~wrapped:true view#buffer#end_iter
else
+ let _ = Ideutils.flash_info "String not found" in
self#set_not_found ()
| Some (start, stop) ->
+ let text = view#buffer#start_iter#get_text ~stop:view#buffer#end_iter in
+ let rec find_all offs accum =
+ if offs > String.length text then
+ List.rev accum
+ else try
+ let i = Str.search_forward self#regex text offs in
+ let j = Str.match_end () in
+ find_all (j + 1) (i :: accum)
+ with Not_found -> List.rev accum
+ in
+ let occurs = find_all 0 [] in
+ let num_occurs = List.length occurs in
+ (* assoc table of offset, occurrence index pairs *)
+ let occur_tbl = List.mapi (fun ndx occ -> (occ,ndx+1)) occurs in
let _ = view#buffer#select_range start stop in
let scroll = `MARK (view#buffer#create_mark stop) in
let _ = view#scroll_to_mark ~use_align:false scroll in
+ let _ =
+ try
+ let occ_ndx = List.assoc start#offset occur_tbl in
+ let occ_str = CString.plural num_occurs "occurrence" in
+ let wrap_str = if wrapped then
+ if backward then " (wrapped backwards)"
+ else " (wrapped)"
+ else ""
+ in
+ Ideutils.flash_info
+ (string_of_int occ_ndx ^ " of " ^ string_of_int num_occurs ^
+ " " ^ occ_str ^ wrap_str)
+ with Not_found ->
+ CErrors.anomaly (Pp.str "Occurrence of Find string not in table")
+ in
self#set_found ()
method find_forward () =
diff --git a/intf/vernacexpr.ml b/intf/vernacexpr.ml
index 5106e513b..8bd2da2f1 100644
--- a/intf/vernacexpr.ml
+++ b/intf/vernacexpr.ml
@@ -314,7 +314,16 @@ type cumulative_inductive_parsing_flag =
(** {6 The type of vernacular expressions} *)
-type vernac_expr =
+type vernac_implicit_status = Implicit | MaximallyImplicit | NotImplicit
+
+type vernac_argument_status = {
+ name : Name.t;
+ recarg_like : bool;
+ notation_scope : string Loc.located option;
+ implicit_status : vernac_implicit_status;
+}
+
+type nonrec vernac_expr =
| VernacLoad of verbose_flag * string
(* Syntax *)
@@ -451,7 +460,7 @@ type vernac_expr =
| VernacUnfocus
| VernacUnfocused
| VernacBullet of bullet
- | VernacSubproof of int option
+ | VernacSubproof of goal_selector option
| VernacEndSubproof
| VernacShow of showable
| VernacCheckGuard
@@ -463,22 +472,13 @@ type vernac_expr =
(* For extension *)
| VernacExtend of extend_name * Genarg.raw_generic_argument list
- (* Flags *)
- | VernacProgram of vernac_expr
- | VernacPolymorphic of bool * vernac_expr
- | VernacLocal of bool * vernac_expr
-
-and vernac_implicit_status = Implicit | MaximallyImplicit | NotImplicit
-
-and vernac_argument_status = {
- name : Name.t;
- recarg_like : bool;
- notation_scope : string Loc.located option;
- implicit_status : vernac_implicit_status;
-}
+type nonrec vernac_flag =
+ | VernacProgram
+ | VernacPolymorphic of bool
+ | VernacLocal of bool
type vernac_control =
- | VernacExpr of vernac_expr
+ | VernacExpr of vernac_flag list * vernac_expr
(* boolean is true when the `-time` batch-mode command line flag was set.
the flag is used to print differently in `-time` vs `Time foo` *)
| VernacTime of bool * vernac_control located
diff --git a/kernel/cClosure.ml b/kernel/cClosure.ml
index 11616da7b..b1181157e 100644
--- a/kernel/cClosure.ml
+++ b/kernel/cClosure.ml
@@ -258,7 +258,7 @@ type 'a infos_cache = {
i_repr : 'a infos -> constr -> 'a;
i_env : env;
i_sigma : existential -> constr option;
- i_rels : constr option array;
+ i_rels : (Context.Rel.Declaration.t * Pre_env.lazy_val) Range.t;
i_tab : 'a KeyTable.t }
and 'a infos = {
@@ -282,13 +282,16 @@ let ref_value_cache ({i_cache = cache} as infos) ref =
let body =
match ref with
| RelKey n ->
- let len = Array.length cache.i_rels in
- let i = n - 1 in
- let () = if i < 0 || len <= i then raise Not_found in
- begin match Array.unsafe_get cache.i_rels i with
- | None -> raise Not_found
- | Some t -> lift n t
- end
+ let open Context.Rel.Declaration in
+ let i = n - 1 in
+ let (d, _) =
+ try Range.get cache.i_rels i
+ with Invalid_argument _ -> raise Not_found
+ in
+ begin match d with
+ | LocalAssum _ -> raise Not_found
+ | LocalDef (_, t, _) -> lift n t
+ end
| VarKey id -> assoc_defined id cache.i_env
| ConstKey cst -> constant_value_in cache.i_env cst
in
@@ -303,26 +306,13 @@ let ref_value_cache ({i_cache = cache} as infos) ref =
let evar_value cache ev =
cache.i_sigma ev
-let defined_rels flags env =
-(* if red_local_const (snd flags) then*)
- let ctx = rel_context env in
- let len = List.length ctx in
- let ans = Array.make len None in
- let open Context.Rel.Declaration in
- let iter i = function
- | LocalAssum _ -> ()
- | LocalDef (_,b,_) -> Array.unsafe_set ans i (Some b)
- in
- let () = List.iteri iter ctx in
- ans
-(* else (0,[])*)
-
let create mk_cl flgs env evars =
+ let open Pre_env in
let cache =
{ i_repr = mk_cl;
i_env = env;
i_sigma = evars;
- i_rels = defined_rels flgs env;
+ i_rels = (Environ.pre_env env).env_rel_context.env_rel_map;
i_tab = KeyTable.create 17 }
in { i_flags = flgs; i_cache = cache }
diff --git a/kernel/csymtable.ml b/kernel/csymtable.ml
index 2ffe36fcf..712c66f11 100644
--- a/kernel/csymtable.ml
+++ b/kernel/csymtable.ml
@@ -198,7 +198,7 @@ and slot_for_fv env fv =
let rv = Pre_env.lookup_rel_val i env in
begin match force_lazy_val rv with
| None ->
- env.env_rel_context |> Context.Rel.lookup i |> RelDecl.get_value |> fill_fv_cache rv i val_of_rel env_of_rel
+ env |> Pre_env.lookup_rel i |> RelDecl.get_value |> fill_fv_cache rv i val_of_rel env_of_rel
| Some (v, _) -> v
end
| FVuniv_var idu ->
diff --git a/kernel/declarations.ml b/kernel/declarations.ml
index 7f4b85fd0..5b9e1a141 100644
--- a/kernel/declarations.ml
+++ b/kernel/declarations.ml
@@ -74,6 +74,7 @@ type typing_flags = {
check_guarded : bool; (** If [false] then fixed points and co-fixed
points are assumed to be total. *)
check_universes : bool; (** If [false] universe constraints are not checked *)
+ conv_oracle : Conv_oracle.oracle; (** Unfolding strategies for conversion *)
}
(* some contraints are in constant_constraints, some other may be in
diff --git a/kernel/declareops.ml b/kernel/declareops.ml
index d8768a0fc..9eed9efcb 100644
--- a/kernel/declareops.ml
+++ b/kernel/declareops.ml
@@ -15,9 +15,10 @@ module RelDecl = Context.Rel.Declaration
(** Operations concernings types in [Declarations] :
[constant_body], [mutual_inductive_body], [module_body] ... *)
-let safe_flags = {
+let safe_flags oracle = {
check_guarded = true;
check_universes = true;
+ conv_oracle = oracle;
}
(** {6 Arities } *)
diff --git a/kernel/declareops.mli b/kernel/declareops.mli
index 198831848..0eed11f49 100644
--- a/kernel/declareops.mli
+++ b/kernel/declareops.mli
@@ -67,7 +67,7 @@ val inductive_is_cumulative : mutual_inductive_body -> bool
(** {6 Kernel flags} *)
(** A default, safe set of flags for kernel type-checking *)
-val safe_flags : typing_flags
+val safe_flags : Conv_oracle.oracle -> typing_flags
(** {6 Hash-consing} *)
diff --git a/kernel/environ.ml b/kernel/environ.ml
index 1afab453a..93dc2f9a7 100644
--- a/kernel/environ.ml
+++ b/kernel/environ.ml
@@ -37,8 +37,10 @@ type env = Pre_env.env
let pre_env env = env
let env_of_pre_env env = env
-let oracle env = env.env_conv_oracle
-let set_oracle env o = { env with env_conv_oracle = o }
+let oracle env = env.env_typing_flags.conv_oracle
+let set_oracle env o =
+ let env_typing_flags = { env.env_typing_flags with conv_oracle = o } in
+ { env with env_typing_flags }
let empty_named_context_val = empty_named_context_val
@@ -58,18 +60,17 @@ let deactivated_guard env = not (typing_flags env).check_guarded
let universes env = env.env_stratification.env_universes
let named_context env = env.env_named_context.env_named_ctx
let named_context_val env = env.env_named_context
-let rel_context env = env.env_rel_context
+let rel_context env = env.env_rel_context.env_rel_ctx
let opaque_tables env = env.indirect_pterms
let set_opaque_tables env indirect_pterms = { env with indirect_pterms }
let empty_context env =
- match env.env_rel_context, env.env_named_context.env_named_ctx with
+ match env.env_rel_context.env_rel_ctx, env.env_named_context.env_named_ctx with
| [], [] -> true
| _ -> false
(* Rel context *)
-let lookup_rel n env =
- Context.Rel.lookup n env.env_rel_context
+let lookup_rel = lookup_rel
let evaluable_rel n env =
is_local_def (lookup_rel n env)
@@ -86,13 +87,12 @@ let push_rec_types (lna,typarray,_) env =
let fold_rel_context f env ~init =
let rec fold_right env =
- match env.env_rel_context with
- | [] -> init
- | rd::rc ->
+ match match_rel_context_val env.env_rel_context with
+ | None -> init
+ | Some (rd, _, rc) ->
let env =
{ env with
env_rel_context = rc;
- env_rel_val = List.tl env.env_rel_val;
env_nb_rel = env.env_nb_rel - 1 } in
f env rd (fold_right env)
in fold_right env
@@ -142,16 +142,21 @@ let evaluable_named id env =
let reset_with_named_context ctxt env =
{ env with
env_named_context = ctxt;
- env_rel_context = Context.Rel.empty;
- env_rel_val = [];
+ env_rel_context = empty_rel_context_val;
env_nb_rel = 0 }
let reset_context = reset_with_named_context empty_named_context_val
let pop_rel_context n env =
+ let rec skip n ctx =
+ if Int.equal n 0 then ctx
+ else match match_rel_context_val ctx with
+ | None -> invalid_arg "List.skipn"
+ | Some (_, _, ctx) -> skip (pred n) ctx
+ in
let ctxt = env.env_rel_context in
{ env with
- env_rel_context = List.skipn n ctxt;
+ env_rel_context = skip n ctxt;
env_nb_rel = env.env_nb_rel - n }
let fold_named_context f env ~init =
diff --git a/kernel/names.mli b/kernel/names.mli
index 709ebeb7f..b1e8efd8d 100644
--- a/kernel/names.mli
+++ b/kernel/names.mli
@@ -40,19 +40,16 @@ sig
(** Hash over identifiers. *)
val is_valid : string -> bool
- (** Check that a string may be converted to an identifier.
- @raise Unicode.Unsupported if the provided string contains unsupported UTF-8 characters. *)
+ (** Check that a string may be converted to an identifier. *)
val of_bytes : bytes -> t
val of_string : string -> t
(** Converts a string into an identifier.
- @raise UserError if the string is invalid as an identifier.
- @raise Unicode.Unsupported if the provided string contains unsupported UTF-8 characters. *)
+ @raise UserError if the string is invalid as an identifier. *)
val of_string_soft : string -> t
(** Same as {!of_string} except that any string made of supported UTF-8 characters is accepted.
- @raise UserError if the string is invalid as an UTF-8 string.
- @raise Unicode.Unsupported if the provided string contains unsupported UTF-8 characters. *)
+ @raise UserError if the string is invalid as an UTF-8 string. *)
val to_string : t -> string
(** Converts a identifier into an string. *)
diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml
index c558e9ed0..ffe19510a 100644
--- a/kernel/nativecode.ml
+++ b/kernel/nativecode.ml
@@ -1830,7 +1830,7 @@ and apply_fv env sigma univ (fv_named,fv_rel) auxdefs ml =
in
let auxdefs = List.fold_right get_rel_val fv_rel auxdefs in
let auxdefs = List.fold_right get_named_val fv_named auxdefs in
- let lvl = Context.Rel.length env.env_rel_context in
+ let lvl = Context.Rel.length env.env_rel_context.env_rel_ctx in
let fv_rel = List.map (fun (n,_) -> MLglobal (Grel (lvl-n))) fv_rel in
let fv_named = List.map (fun (id,_) -> MLglobal (Gnamed id)) fv_named in
let aux_name = fresh_lname Anonymous in
@@ -1838,8 +1838,8 @@ and apply_fv env sigma univ (fv_named,fv_rel) auxdefs ml =
and compile_rel env sigma univ auxdefs n =
let open Context.Rel.Declaration in
- let decl = Context.Rel.lookup n env.env_rel_context in
- let n = Context.Rel.length env.env_rel_context - n in
+ let decl = Pre_env.lookup_rel n env in
+ let n = List.length env.env_rel_context.env_rel_ctx - n in
match decl with
| LocalDef (_,t,_) ->
let code = lambda_of_constr env sigma t in
diff --git a/kernel/nativelambda.ml b/kernel/nativelambda.ml
index 160a90dc2..b333b0fb9 100644
--- a/kernel/nativelambda.ml
+++ b/kernel/nativelambda.ml
@@ -639,7 +639,7 @@ let optimize lam =
let lambda_of_constr env sigma c =
set_global_env env;
let env = Renv.make () in
- let ids = List.rev_map RelDecl.get_name !global_env.env_rel_context in
+ let ids = List.rev_map RelDecl.get_name !global_env.env_rel_context.env_rel_ctx in
Renv.push_rels env (Array.of_list ids);
let lam = lambda_of_constr env sigma c in
(* if Flags.vm_draw_opt () then begin
diff --git a/kernel/pre_env.ml b/kernel/pre_env.ml
index c5254b453..3ef15105a 100644
--- a/kernel/pre_env.ml
+++ b/kernel/pre_env.ml
@@ -67,15 +67,18 @@ type named_context_val = {
env_named_map : (Context.Named.Declaration.t * lazy_val) Id.Map.t;
}
+type rel_context_val = {
+ env_rel_ctx : Context.Rel.t;
+ env_rel_map : (Context.Rel.Declaration.t * lazy_val) Range.t;
+}
+
type env = {
env_globals : globals; (* globals = constants + inductive types + modules + module-types *)
env_named_context : named_context_val; (* section variables *)
- env_rel_context : Context.Rel.t;
- env_rel_val : lazy_val list;
+ env_rel_context : rel_context_val;
env_nb_rel : int;
env_stratification : stratification;
env_typing_flags : typing_flags;
- env_conv_oracle : Conv_oracle.oracle;
retroknowledge : Retroknowledge.retroknowledge;
indirect_pterms : Opaqueproof.opaquetab;
}
@@ -85,6 +88,11 @@ let empty_named_context_val = {
env_named_map = Id.Map.empty;
}
+let empty_rel_context_val = {
+ env_rel_ctx = [];
+ env_rel_map = Range.empty;
+}
+
let empty_env = {
env_globals = {
env_constants = Cmap_env.empty;
@@ -92,14 +100,12 @@ let empty_env = {
env_modules = MPmap.empty;
env_modtypes = MPmap.empty};
env_named_context = empty_named_context_val;
- env_rel_context = Context.Rel.empty;
- env_rel_val = [];
+ env_rel_context = empty_rel_context_val;
env_nb_rel = 0;
env_stratification = {
env_universes = UGraph.initial_universes;
env_engagement = PredicativeSet };
- env_typing_flags = Declareops.safe_flags;
- env_conv_oracle = Conv_oracle.empty;
+ env_typing_flags = Declareops.safe_flags Conv_oracle.empty;
retroknowledge = Retroknowledge.initial_retroknowledge;
indirect_pterms = Opaqueproof.empty_opaquetab }
@@ -108,21 +114,39 @@ let empty_env = {
let nb_rel env = env.env_nb_rel
+let push_rel_context_val d ctx = {
+ env_rel_ctx = Context.Rel.add d ctx.env_rel_ctx;
+ env_rel_map = Range.cons (d, ref VKnone) ctx.env_rel_map;
+}
+
+let match_rel_context_val ctx = match ctx.env_rel_ctx with
+| [] -> None
+| decl :: rem ->
+ let (_, lval) = Range.hd ctx.env_rel_map in
+ let ctx = { env_rel_ctx = rem; env_rel_map = Range.tl ctx.env_rel_map } in
+ Some (decl, lval, ctx)
+
let push_rel d env =
- let rval = ref VKnone in
{ env with
- env_rel_context = Context.Rel.add d env.env_rel_context;
- env_rel_val = rval :: env.env_rel_val;
+ env_rel_context = push_rel_context_val d env.env_rel_context;
env_nb_rel = env.env_nb_rel + 1 }
+let lookup_rel n env =
+ try fst (Range.get env.env_rel_context.env_rel_map (n - 1))
+ with Invalid_argument _ -> raise Not_found
+
let lookup_rel_val n env =
- try List.nth env.env_rel_val (n - 1)
- with Failure _ -> raise Not_found
+ try snd (Range.get env.env_rel_context.env_rel_map (n - 1))
+ with Invalid_argument _ -> raise Not_found
+
+let rel_skipn n ctx = {
+ env_rel_ctx = Util.List.skipn n ctx.env_rel_ctx;
+ env_rel_map = Range.skipn n ctx.env_rel_map;
+}
let env_of_rel n env =
{ env with
- env_rel_context = Util.List.skipn n env.env_rel_context;
- env_rel_val = Util.List.skipn n env.env_rel_val;
+ env_rel_context = rel_skipn n env.env_rel_context;
env_nb_rel = env.env_nb_rel - n
}
diff --git a/kernel/pre_env.mli b/kernel/pre_env.mli
index 054ae1743..335ca1057 100644
--- a/kernel/pre_env.mli
+++ b/kernel/pre_env.mli
@@ -45,15 +45,18 @@ type named_context_val = private {
env_named_map : (Context.Named.Declaration.t * lazy_val) Id.Map.t;
}
+type rel_context_val = private {
+ env_rel_ctx : Context.Rel.t;
+ env_rel_map : (Context.Rel.Declaration.t * lazy_val) Range.t;
+}
+
type env = {
env_globals : globals;
env_named_context : named_context_val;
- env_rel_context : Context.Rel.t;
- env_rel_val : lazy_val list;
+ env_rel_context : rel_context_val;
env_nb_rel : int;
env_stratification : stratification;
env_typing_flags : typing_flags;
- env_conv_oracle : Conv_oracle.oracle;
retroknowledge : Retroknowledge.retroknowledge;
indirect_pterms : Opaqueproof.opaquetab;
}
@@ -64,8 +67,15 @@ val empty_env : env
(** Rel context *)
+val empty_rel_context_val : rel_context_val
+val push_rel_context_val :
+ Context.Rel.Declaration.t -> rel_context_val -> rel_context_val
+val match_rel_context_val :
+ rel_context_val -> (Context.Rel.Declaration.t * lazy_val * rel_context_val) option
+
val nb_rel : env -> int
val push_rel : Context.Rel.Declaration.t -> env -> env
+val lookup_rel : int -> env -> Context.Rel.Declaration.t
val lookup_rel_val : int -> env -> lazy_val
val env_of_rel : int -> env -> env
diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml
index cbc4ee2ec..5f501bff1 100644
--- a/kernel/term_typing.ml
+++ b/kernel/term_typing.ml
@@ -301,21 +301,29 @@ let infer_declaration (type a) ~(trust : a trust) env (dcl : a constant_entry) =
let { const_entry_type = typ; const_entry_opaque = opaque } = c in
let { const_entry_body = body; const_entry_feedback = feedback_id } = c in
let (body, ctx), side_eff = Future.join body in
- let poly, univsctx = match c.const_entry_universes with
- | Monomorphic_const_entry univs -> false, univs
- | Polymorphic_const_entry univs -> true, Univ.ContextSet.of_context univs
- in
- let ctx = Univ.ContextSet.union univsctx ctx in
let body, ctx, _ = match trust with
| Pure -> body, ctx, []
| SideEffects _ -> inline_side_effects env body ctx side_eff
in
- let env = push_context_set ~strict:(not poly) ctx env in
- let ctx = if poly
- then Polymorphic_const_entry (Univ.ContextSet.to_context ctx)
- else Monomorphic_const_entry ctx
+ let env, usubst, univs = match c.const_entry_universes with
+ | Monomorphic_const_entry univs ->
+ let ctx = Univ.ContextSet.union univs ctx in
+ let env = push_context_set ~strict:true ctx env in
+ env, Univ.empty_level_subst, Monomorphic_const ctx
+ | Polymorphic_const_entry uctx ->
+ (** Ensure not to generate internal constraints in polymorphic mode.
+ The only way for this to happen would be that either the body
+ contained deferred universes, or that it contains monomorphic
+ side-effects. The first property is ruled out by upper layers,
+ and the second one is ensured by the fact we currently
+ unconditionally export side-effects from polymorphic definitions,
+ i.e. [trust] is always [Pure]. *)
+ let () = assert (Univ.ContextSet.is_empty ctx) in
+ let env = push_context ~strict:false uctx env in
+ let sbst, auctx = Univ.abstract_universes uctx in
+ let sbst = Univ.make_instance_subst sbst in
+ env, sbst, Polymorphic_const auctx
in
- let usubst, univs = abstract_constant_universes ctx in
let j = infer env body in
let typ = match typ with
| None ->
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index 92e60f465..d498bda34 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -72,36 +72,36 @@ GEXTEND Gram
| IDENT "Redirect"; s = ne_string; c = located_vernac -> VernacRedirect (s, c)
| IDENT "Timeout"; n = natural; v = vernac_control -> VernacTimeout(n,v)
| IDENT "Fail"; v = vernac_control -> VernacFail v
- | v = vernac -> VernacExpr(v) ]
+ | (f, v) = vernac -> VernacExpr(f, v) ]
]
;
vernac:
- [ [ IDENT "Local"; v = vernac_poly -> VernacLocal (true, v)
- | IDENT "Global"; v = vernac_poly -> VernacLocal (false, v)
+ [ [ IDENT "Local"; (f, v) = vernac_poly -> (VernacLocal true :: f, v)
+ | IDENT "Global"; (f, v) = vernac_poly -> (VernacLocal false :: f, v)
| v = vernac_poly -> v ]
]
;
vernac_poly:
- [ [ IDENT "Polymorphic"; v = vernac_aux -> VernacPolymorphic (true, v)
- | IDENT "Monomorphic"; v = vernac_aux -> VernacPolymorphic (false, v)
+ [ [ IDENT "Polymorphic"; (f, v) = vernac_aux -> (VernacPolymorphic true :: f, v)
+ | IDENT "Monomorphic"; (f, v) = vernac_aux -> (VernacPolymorphic false :: f, v)
| v = vernac_aux -> v ]
]
;
vernac_aux:
(* Better to parse "." here: in case of failure (e.g. in coerce_to_var), *)
(* "." is still in the stream and discard_to_dot works correctly *)
- [ [ IDENT "Program"; g = gallina; "." -> VernacProgram g
- | IDENT "Program"; g = gallina_ext; "." -> VernacProgram g
- | g = gallina; "." -> g
- | g = gallina_ext; "." -> g
- | c = command; "." -> c
- | c = syntax; "." -> c
- | c = subprf -> c
+ [ [ IDENT "Program"; g = gallina; "." -> ([VernacProgram], g)
+ | IDENT "Program"; g = gallina_ext; "." -> ([VernacProgram], g)
+ | g = gallina; "." -> ([], g)
+ | g = gallina_ext; "." -> ([], g)
+ | c = command; "." -> ([], c)
+ | c = syntax; "." -> ([], c)
+ | c = subprf -> ([], c)
] ]
;
vernac_aux: LAST
- [ [ prfcom = command_entry -> prfcom ] ]
+ [ [ prfcom = command_entry -> ([], prfcom) ] ]
;
noedit_mode:
[ [ c = query_command -> c None] ]
@@ -621,11 +621,9 @@ GEXTEND Gram
VernacCanonical (AN qid)
| IDENT "Canonical"; IDENT "Structure"; ntn = by_notation ->
VernacCanonical (ByNotation ntn)
- | IDENT "Canonical"; IDENT "Structure"; qid = global;
- d = def_body ->
+ | IDENT "Canonical"; IDENT "Structure"; qid = global; d = def_body ->
let s = coerce_reference_to_id qid in
- VernacLocal(false,
- VernacDefinition ((NoDischarge,CanonicalStructure),((Loc.tag s),None),d))
+ VernacDefinition ((NoDischarge,CanonicalStructure),((Loc.tag s),None),d)
(* Coercions *)
| IDENT "Coercion"; qid = global; d = def_body ->
diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml
index c55040df0..4c59996aa 100644
--- a/plugins/firstorder/formula.ml
+++ b/plugins/firstorder/formula.ml
@@ -55,7 +55,8 @@ let ind_hyps env sigma nevar ind largs =
let types= Inductiveops.arities_of_constructors env ind in
let myhyps t =
let t = EConstr.of_constr t in
- let t1=Termops.prod_applist sigma t largs in
+ let nparam_decls = Context.Rel.length (fst (Global.lookup_inductive (fst ind))).mind_params_ctxt in
+ let t1=Termops.prod_applist_assum sigma nparam_decls t largs in
let t2=snd (decompose_prod_n_assum sigma nevar t1) in
fst (decompose_prod_assum sigma t2) in
Array.map myhyps types
diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4
index 2fd6f53c4..4b828a702 100644
--- a/plugins/funind/g_indfun.ml4
+++ b/plugins/funind/g_indfun.ml4
@@ -154,7 +154,7 @@ VERNAC COMMAND EXTEND Function
| _,((_,(_,CStructRec),_,_,_),_) -> false) recsl in
match
Vernac_classifier.classify_vernac
- (Vernacexpr.(VernacExpr(VernacFixpoint(Decl_kinds.NoDischarge, List.map snd recsl))))
+ (Vernacexpr.(VernacExpr([], VernacFixpoint(Decl_kinds.NoDischarge, List.map snd recsl))))
with
| Vernacexpr.VtSideff ids, _ when hard ->
Vernacexpr.(VtStartProof ("Classic", GuaranteesOpacity, ids), VtLater)
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index b4e17c5d1..0b929b8ca 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -1509,7 +1509,7 @@ let do_build_inductive
in
let msg =
str "while trying to define"++ spc () ++
- Ppvernac.pr_vernac Vernacexpr.(VernacExpr(VernacInductive(GlobalNonCumulativity,false,Declarations.Finite,repacked_rel_inds)))
+ Ppvernac.pr_vernac Vernacexpr.(VernacExpr([], VernacInductive(GlobalNonCumulativity,false,Declarations.Finite,repacked_rel_inds)))
++ fnl () ++
msg
in
@@ -1524,7 +1524,7 @@ let do_build_inductive
in
let msg =
str "while trying to define"++ spc () ++
- Ppvernac.pr_vernac Vernacexpr.(VernacExpr(VernacInductive(GlobalNonCumulativity,false,Declarations.Finite,repacked_rel_inds)))
+ Ppvernac.pr_vernac Vernacexpr.(VernacExpr([], VernacInductive(GlobalNonCumulativity,false,Declarations.Finite,repacked_rel_inds)))
++ fnl () ++
CErrors.print reraise
in
diff --git a/plugins/ltac/g_ltac.ml4 b/plugins/ltac/g_ltac.ml4
index ebf6e450b..cc7ce339b 100644
--- a/plugins/ltac/g_ltac.ml4
+++ b/plugins/ltac/g_ltac.ml4
@@ -327,7 +327,8 @@ GEXTEND Gram
| IDENT "all"; ":" -> SelectAll ] ]
;
tactic_mode:
- [ [ g = OPT toplevel_selector; tac = G_vernac.query_command -> tac g ] ]
+ [ [ g = OPT toplevel_selector; tac = G_vernac.query_command -> tac g
+ | g = OPT toplevel_selector; "{" -> Vernacexpr.VernacSubproof g ] ]
;
command:
[ [ IDENT "Proof"; "with"; ta = Pltac.tactic;
diff --git a/plugins/ssr/ssrvernac.ml4 b/plugins/ssr/ssrvernac.ml4
index 4f530a0ae..c0479dd24 100644
--- a/plugins/ssr/ssrvernac.ml4
+++ b/plugins/ssr/ssrvernac.ml4
@@ -551,9 +551,9 @@ GEXTEND Gram
| IDENT "Canonical"; qid = Constr.global;
d = G_vernac.def_body ->
let s = coerce_reference_to_id qid in
- Vernacexpr.VernacLocal(false,Vernacexpr.VernacDefinition
+ Vernacexpr.VernacDefinition
((Decl_kinds.NoDischarge,Decl_kinds.CanonicalStructure),
- ((Loc.tag s),None),(d )))
+ ((Loc.tag s),None), d)
]];
END
diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml
index ec7c3077f..c3a221944 100644
--- a/pretyping/constr_matching.ml
+++ b/pretyping/constr_matching.ml
@@ -462,19 +462,21 @@ let sub_match ?(closed=true) env sigma pat c =
in
let sub = (env, c1) :: (env, hd) :: subargs env lc in
try_aux sub next_mk_ctx next
- | Fix (indx,(names,types,bodies)) ->
+ | Fix (indx,(names,types,bodies as recdefs)) ->
let nb_fix = Array.length types in
let next_mk_ctx le =
let (ntypes,nbodies) = CList.chop nb_fix le in
mk_ctx (mkFix (indx,(names, Array.of_list ntypes, Array.of_list nbodies))) in
- let sub = subargs env types @ subargs env bodies in
+ let env' = push_rec_types recdefs env in
+ let sub = subargs env types @ subargs env' bodies in
try_aux sub next_mk_ctx next
- | CoFix (i,(names,types,bodies)) ->
+ | CoFix (i,(names,types,bodies as recdefs)) ->
let nb_fix = Array.length types in
let next_mk_ctx le =
let (ntypes,nbodies) = CList.chop nb_fix le in
mk_ctx (mkCoFix (i,(names, Array.of_list ntypes, Array.of_list nbodies))) in
- let sub = subargs env types @ subargs env bodies in
+ let env' = push_rec_types recdefs env in
+ let sub = subargs env types @ subargs env' bodies in
try_aux sub next_mk_ctx next
| Proj (p,c') ->
begin try
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index 788e4d268..41c4616f7 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -276,11 +276,6 @@ let rec ise_app_stack2 env f evd sk1 sk2 =
end
| _, _ -> (sk1,sk2), Success evd
-let push_rec_types pfix env =
- let (i, c, t) = pfix in
- let inj c = EConstr.Unsafe.to_constr c in
- push_rec_types (i, Array.map inj c, Array.map inj t) env
-
(* This function tries to unify 2 stacks element by element. It works
from the end to the beginning. If it unifies a non empty suffix of
stacks but not the entire stacks, the first part of the answer is
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index 43066c809..3132d2ad5 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -23,11 +23,6 @@ open Arguments_renaming
open Pretype_errors
open Context.Rel.Declaration
-let push_rec_types pfix env =
- let (i, c, t) = pfix in
- let inj c = EConstr.Unsafe.to_constr c in
- push_rec_types (i, Array.map inj c, Array.map inj t) env
-
let meta_type evd mv =
let ty =
try Evd.meta_ftype evd mv
diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml
index e88284fb1..96e39e89a 100644
--- a/printing/ppvernac.ml
+++ b/printing/ppvernac.ml
@@ -91,7 +91,7 @@ open Decl_kinds
let sep_end = function
| VernacBullet _
- | VernacSubproof None
+ | VernacSubproof _
| VernacEndSubproof -> str""
| _ -> str"."
@@ -535,17 +535,25 @@ open Decl_kinds
| SsFwdClose e -> "("^aux e^")*"
in Pp.str (aux e)
- let rec pr_vernac_expr v =
+ let pr_extend s cl =
+ let pr_arg a =
+ try pr_gen a
+ with Failure _ -> str "<error in " ++ str (fst s) ++ str ">" in
+ try
+ let rl = Egramml.get_extend_vernac_rule s in
+ let rec aux rl cl =
+ match rl, cl with
+ | Egramml.GramNonTerminal _ :: rl, arg :: cl -> pr_arg arg :: aux rl cl
+ | Egramml.GramTerminal s :: rl, cl -> str s :: aux rl cl
+ | [], [] -> []
+ | _ -> assert false in
+ hov 1 (pr_sequence identity (aux rl cl))
+ with Not_found ->
+ hov 1 (str "TODO(" ++ str (fst s) ++ spc () ++ prlist_with_sep sep pr_arg cl ++ str ")")
+
+ let pr_vernac_expr v =
let return = tag_vernac v in
match v with
- | VernacPolymorphic (poly, v) ->
- let s = if poly then keyword "Polymorphic" else keyword "Monomorphic" in
- return (s ++ spc () ++ pr_vernac_expr v)
- | VernacProgram v ->
- return (keyword "Program" ++ spc() ++ pr_vernac_expr v)
- | VernacLocal (local, v) ->
- return (pr_locality local ++ spc() ++ pr_vernac_expr v)
-
| VernacLoad (f,s) ->
return (
keyword "Load"
@@ -1199,30 +1207,25 @@ open Decl_kinds
| VernacSubproof None ->
return (str "{")
| VernacSubproof (Some i) ->
- return (keyword "BeginSubproof" ++ spc () ++ int i)
+ return (Proof_bullet.pr_goal_selector i ++ str ":" ++ spc () ++ str "{")
| VernacEndSubproof ->
return (str "}")
- and pr_extend s cl =
- let pr_arg a =
- try pr_gen a
- with Failure _ -> str "<error in " ++ str (fst s) ++ str ">" in
- try
- let rl = Egramml.get_extend_vernac_rule s in
- let rec aux rl cl =
- match rl, cl with
- | Egramml.GramNonTerminal _ :: rl, arg :: cl -> pr_arg arg :: aux rl cl
- | Egramml.GramTerminal s :: rl, cl -> str s :: aux rl cl
- | [], [] -> []
- | _ -> assert false in
- hov 1 (pr_sequence identity (aux rl cl))
- with Not_found ->
- hov 1 (str "TODO(" ++ str (fst s) ++ spc () ++ prlist_with_sep sep pr_arg cl ++ str ")")
+let pr_vernac_flag =
+ function
+ | VernacPolymorphic true -> keyword "Polymorphic"
+ | VernacPolymorphic false -> keyword "Monomorphic"
+ | VernacProgram -> keyword "Program"
+ | VernacLocal local -> pr_locality local
let rec pr_vernac_control v =
let return = tag_vernac v in
match v with
- | VernacExpr v' -> pr_vernac_expr v' ++ sep_end v'
+ | VernacExpr (f, v') ->
+ List.fold_right
+ (fun f a -> pr_vernac_flag f ++ spc() ++ a)
+ f
+ (pr_vernac_expr v' ++ sep_end v')
| VernacTime (_,(_,v)) ->
return (keyword "Time" ++ spc() ++ pr_vernac_control v)
| VernacRedirect (s, (_,v)) ->
diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml
index 167d6bda0..d04bdb652 100644
--- a/proofs/proof_global.ml
+++ b/proofs/proof_global.ml
@@ -343,16 +343,15 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now
if poly || now then
let make_body t (c, eff) =
let body = c in
- let typ =
- if not (keep_body_ucst_separate || not (Safe_typing.empty_private_constants = eff)) then
- nf t
- else t
+ let allow_deferred =
+ not poly && (keep_body_ucst_separate ||
+ not (Safe_typing.empty_private_constants = eff))
in
+ let typ = if allow_deferred then t else nf t in
let env = Global.env () in
let used_univs_body = Univops.universes_of_constr env body in
let used_univs_typ = Univops.universes_of_constr env typ in
- if keep_body_ucst_separate ||
- not (Safe_typing.empty_private_constants = eff) then
+ if allow_deferred then
let initunivs = UState.const_univ_entry ~poly initial_euctx in
let ctx = constrain_variables universes in
(* For vi2vo compilation proofs are computed now but we need to
diff --git a/stm/proofBlockDelimiter.ml b/stm/proofBlockDelimiter.ml
index 01b5b9a01..bebc4d5d5 100644
--- a/stm/proofBlockDelimiter.ml
+++ b/stm/proofBlockDelimiter.ml
@@ -96,7 +96,7 @@ let dynamic_bullet doc { dynamic_switch = id; carry_on_data = b } =
`ValidBlock {
base_state = id;
goals_to_admit = focused;
- recovery_command = Some (Vernacexpr.VernacExpr(Vernacexpr.VernacBullet (to_bullet_val b)))
+ recovery_command = Some (Vernacexpr.VernacExpr([], Vernacexpr.VernacBullet (to_bullet_val b)))
}
| `Not -> `Leaks
@@ -125,7 +125,7 @@ let dynamic_curly_brace doc { dynamic_switch = id } =
`ValidBlock {
base_state = id;
goals_to_admit = focused;
- recovery_command = Some (Vernacexpr.VernacExpr Vernacexpr.VernacEndSubproof)
+ recovery_command = Some (Vernacexpr.VernacExpr ([], Vernacexpr.VernacEndSubproof))
}
| `Not -> `Leaks
diff --git a/stm/stm.ml b/stm/stm.ml
index 7045df0ed..5f4fe6565 100644
--- a/stm/stm.ml
+++ b/stm/stm.ml
@@ -1498,7 +1498,7 @@ end = struct (* {{{ *)
stm_vernac_interp stop
~proof:(pobject, terminator) st
{ verbose = false; loc; indentation = 0; strlen = 0;
- expr = VernacExpr (VernacEndProof (Proved (Opaque,None))) }) in
+ expr = VernacExpr ([], VernacEndProof (Proved (Opaque,None))) }) in
ignore(Future.join checked_proof);
end;
(* STATE: Restore the state XXX: handle exn *)
@@ -1646,7 +1646,7 @@ end = struct (* {{{ *)
let st = Vernacstate.freeze_interp_state `No in
ignore(stm_vernac_interp stop ~proof st
{ verbose = false; loc; indentation = 0; strlen = 0;
- expr = VernacExpr (VernacEndProof (Proved (Opaque,None))) });
+ expr = VernacExpr ([], VernacEndProof (Proved (Opaque,None))) });
`OK proof
end
with e ->
@@ -2177,7 +2177,7 @@ let collect_proof keep cur hd brkind id =
(try
let name, hint = name ids, get_hint_ctx loc in
let t, v = proof_no_using last in
- v.expr <- VernacExpr(VernacProof(t, Some hint));
+ v.expr <- VernacExpr([], VernacProof(t, Some hint));
`ASync (parent last,accn,name,delegate name)
with Not_found ->
let name = name ids in
@@ -2872,10 +2872,9 @@ let process_transaction ?(newtip=Stateid.fresh ()) ?(part_of_script=true)
if not in_proof && Proof_global.there_are_pending_proofs () then
begin
let bname = VCS.mk_branch_name x in
- let rec opacity_of_produced_term = function
+ let opacity_of_produced_term = function
(* This AST is ambiguous, hence we check it dynamically *)
| VernacInstance (false, _,_ , None, _) -> GuaranteesOpacity
- | VernacLocal (_,e) -> opacity_of_produced_term e
| _ -> Doesn'tGuaranteeOpacity in
VCS.commit id (Fork (x,bname,opacity_of_produced_term (Vernacprop.under_control x.expr),[]));
let proof_mode = default_proof_mode () in
diff --git a/stm/vernac_classifier.ml b/stm/vernac_classifier.ml
index 2fa47ba43..99b56d484 100644
--- a/stm/vernac_classifier.ml
+++ b/stm/vernac_classifier.ml
@@ -49,17 +49,13 @@ let declare_vernac_classifier
classifiers := !classifiers @ [s,f]
let classify_vernac e =
- let rec static_classifier ~poly e = match e with
+ let static_classifier ~poly e = match e with
(* Univ poly compatibility: we run it now, so that we can just
* look at Flags in stm.ml. Would be nicer to have the stm
* look at the entire dag to detect this option. *)
| ( VernacSetOption (l,_) | VernacUnsetOption l)
when CList.equal String.equal l Vernacentries.universe_polymorphism_option_name ->
VtSideff [], VtNow
- (* Nested vernac exprs *)
- | VernacProgram e -> static_classifier ~poly e
- | VernacLocal (_,e) -> static_classifier ~poly e
- | VernacPolymorphic (b, e) -> static_classifier ~poly:b e
(* Qed *)
| VernacAbort _ -> VtQed VtDrop, VtLater
| VernacEndProof Admitted -> VtQed VtKeepAsAxiom, VtLater
@@ -193,7 +189,13 @@ let classify_vernac e =
with Not_found -> anomaly(str"No classifier for"++spc()++str (fst s)++str".")
in
let rec static_control_classifier ~poly = function
- | VernacExpr e -> static_classifier ~poly e
+ | VernacExpr (f, e) ->
+ let poly = List.fold_left (fun poly f ->
+ match f with
+ | VernacPolymorphic b -> b
+ | (VernacProgram | VernacLocal _) -> poly
+ ) poly f in
+ static_classifier ~poly e
| VernacTimeout (_,e) -> static_control_classifier ~poly e
| VernacTime (_,(_,e)) | VernacRedirect (_, (_,e)) ->
static_control_classifier ~poly e
diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml
index a3a3e0a9e..11ac16680 100644
--- a/tactics/hipattern.ml
+++ b/tactics/hipattern.ml
@@ -111,7 +111,8 @@ let match_with_one_constructor sigma style onlybinary allow_rec t =
Some (hdapp,args)
else None
else
- let ctyp = Termops.prod_applist sigma (EConstr.of_constr mip.mind_nf_lc.(0)) args in
+ let ctyp = Termops.prod_applist_assum sigma (Context.Rel.length mib.mind_params_ctxt)
+ (EConstr.of_constr mip.mind_nf_lc.(0)) args in
let cargs = List.map RelDecl.get_type (prod_assum sigma ctyp) in
if not (is_lax_conjunction style) || has_nodep_prod sigma ctyp then
(* Record or non strict conjunction *)
@@ -364,36 +365,39 @@ let is_forall_term sigma c = op2bool (match_with_forall_term sigma c)
let match_with_nodep_ind sigma t =
let (hdapp,args) = decompose_app sigma t in
- match EConstr.kind sigma hdapp with
- | Ind (ind, _) ->
- let (mib,mip) = Global.lookup_inductive ind in
- if Array.length (mib.mind_packets)>1 then None else
- let nodep_constr c = has_nodep_prod_after mib.mind_nparams sigma (EConstr.of_constr c) in
- if Array.for_all nodep_constr mip.mind_nf_lc then
- let params=
- if Int.equal mip.mind_nrealargs 0 then args else
- fst (List.chop mib.mind_nparams args) in
- Some (hdapp,params,mip.mind_nrealargs)
- else
- None
- | _ -> None
+ match EConstr.kind sigma hdapp with
+ | Ind (ind, _) ->
+ let (mib,mip) = Global.lookup_inductive ind in
+ if Array.length (mib.mind_packets)>1 then None else
+ let nodep_constr c =
+ has_nodep_prod_after (Context.Rel.length mib.mind_params_ctxt) sigma (EConstr.of_constr c) in
+ if Array.for_all nodep_constr mip.mind_nf_lc then
+ let params=
+ if Int.equal mip.mind_nrealargs 0 then args else
+ fst (List.chop mib.mind_nparams args) in
+ Some (hdapp,params,mip.mind_nrealargs)
+ else
+ None
+ | _ -> None
let is_nodep_ind sigma t = op2bool (match_with_nodep_ind sigma t)
let match_with_sigma_type sigma t =
let (hdapp,args) = decompose_app sigma t in
match EConstr.kind sigma hdapp with
- | Ind (ind, _) ->
- let (mib,mip) = Global.lookup_inductive ind in
- if Int.equal (Array.length (mib.mind_packets)) 1 &&
- (Int.equal mip.mind_nrealargs 0) &&
- (Int.equal (Array.length mip.mind_consnames)1) &&
- has_nodep_prod_after (mib.mind_nparams+1) sigma (EConstr.of_constr mip.mind_nf_lc.(0)) then
- (*allowing only 1 existential*)
- Some (hdapp,args)
- else
- None
- | _ -> None
+ | Ind (ind, _) ->
+ let (mib,mip) = Global.lookup_inductive ind in
+ if Int.equal (Array.length (mib.mind_packets)) 1
+ && (Int.equal mip.mind_nrealargs 0)
+ && (Int.equal (Array.length mip.mind_consnames)1)
+ && has_nodep_prod_after (Context.Rel.length mib.mind_params_ctxt + 1) sigma
+ (EConstr.of_constr mip.mind_nf_lc.(0))
+ then
+ (*allowing only 1 existential*)
+ Some (hdapp,args)
+ else
+ None
+ | _ -> None
let is_sigma_type sigma t = op2bool (match_with_sigma_type sigma t)
diff --git a/test-suite/Makefile b/test-suite/Makefile
index de669218c..16a56f440 100644
--- a/test-suite/Makefile
+++ b/test-suite/Makefile
@@ -40,6 +40,7 @@ coqtopload := $(coqtop) -top Top -async-proofs-cache force -load-vernac-source
coqtopcompile := $(coqtop) -compile
coqdep := $(BIN)coqdep -coqlib $(LIB)
+VERBOSE?=
SHOW := $(if $(VERBOSE),@true,@echo)
HIDE := $(if $(VERBOSE),,@)
REDIR := $(if $(VERBOSE),,> /dev/null 2>&1)
@@ -174,10 +175,20 @@ summary.log:
# if not on travis we can get the log files (they're just there for a
# local build, and downloadable on GitLab)
+PRINT_LOGS?=
+TRAVIS?= # special because we want to print travis_fold directives
+ifdef APPVEYOR
+PRINT_LOGS:=APPVEYOR
+else
+ifdef CIRCLECI
+PRINT_LOGS:=CIRCLECI
+endif #CIRCLECI
+endif #APPVEYOR
+
report: summary.log
$(HIDE)bash save-logs.sh
$(HIDE)if [ -n "${TRAVIS}" ]; then find logs/ -name '*.log' -not -name 'summary.log' -exec 'bash' '-c' 'echo "travis_fold:start:coq.logs.$$(echo '{}' | sed s,/,.,g)"' ';' -exec cat '{}' ';' -exec 'bash' '-c' 'echo "travis_fold:end:coq.logs.$$(echo '{}' | sed s,/,.,g)"' ';'; fi
- $(HIDE)if [ "(" -n "${APPVEYOR}" ")" -o "(" -n "${CIRCLECI}" ")" ]; then find logs/ -name '*.log' -not -name 'summary.log' -exec 'bash' '-c' 'echo {}' ';' -exec cat '{}' ';' -exec 'bash' '-c' 'echo' ';'; fi
+ $(HIDE)if [ -n "${PRINT_LOGS}" ]; then find logs/ -name '*.log' -not -name 'summary.log' -exec 'bash' '-c' 'echo {}' ';' -exec cat '{}' ';' -exec 'bash' '-c' 'echo' ';'; fi
$(HIDE)if grep -q -F 'Error!' summary.log ; then echo FAILURES; grep -F 'Error!' summary.log; false; else echo NO FAILURES; fi
#######################################################################
diff --git a/test-suite/README.md b/test-suite/README.md
new file mode 100644
index 000000000..1d1195646
--- /dev/null
+++ b/test-suite/README.md
@@ -0,0 +1,75 @@
+# Coq Test Suite
+
+The test suite can be run from the Coq root directory by `make test-suite`.
+This does a clean step first, so if you've already run it, then change something,
+you'll have to do a lot of work again.
+
+If you run `make` from the `test-suite` directory, there is no clean step.
+You can also run `make aaa/bbb/ccc.v.log` to build the log for one test,
+or `make ddd` where `ddd` is on of the sub-directories of `test-suite`
+to just build the logs for that directory.
+In these cases, a summary is not printed, but can be generated by `make summary`.
+
+`make -B` can be used to rerun tests ( -B meaning always remake).
+
+From the `test-suite` directory, `make report` (included in `make
+all`) prints a summary of which tests failed using the produced log
+files (this still works when only some tests are built as described
+above). Setting the `PRINT_LOGS` variable will make it print the logs
+of the failing tests.
+
+For instance, running the following in the `test-suite` directory:
+
+```bash
+$ echo Fail. > success/fail.v # make some failing test
+
+$ make
+TEST prerequisite/make_local.v
+...
+TEST success/fail.v
+...
+BUILDING SUMMARY FILE
+FAILURES
+ success/fail.v...Error! (should be accepted)
+Makefile:189: recipe for target 'all failed
+make: *** [report] Error 1
+
+$ make report PRINT_LOGS=1
+BUILDING SUMMARY FILE
+logs/success/fail.v.log
+==========> TESTING success/fail.v <==========
+Welcome to Coq (version information)
+Skipping rcfile loading.
+File "/path/to/success/fail.v", line 1, characters 4-5:
+Error:
+Syntax error: [vernac:Vernac.vernac_control] expected after 'Fail' (in [vernac:Vernac.vernac_control]).
+
+0m0.000000s 0m0.000000s
+0m0.040000s 0m0.000000s
+==========> FAILURE <==========
+ success/fail.v...Error! (should be accepted)
+
+FAILURES
+ success/fail.v...Error! (should be accepted)
+Makefile:189: recipe for target 'report' failed
+make: *** [report] Error 1
+
+$ echo 'Comments "foo".' > success/fail.v
+
+$ make
+TEST success/fail.v
+BUILDING SUMMARY FILE
+NO FAILURES
+```
+
+See [`test-suite/Makefile`](/test-suite/Makefile) for more information.
+
+## Adding a test
+
+Regression tests for closed bugs should be added to `test-suite/bugs/closed`, as `1234.v` where `1234` is the bug number.
+Files in this directory are tested for successful compilation.
+When you fix a bug, you should usually add a regression test here as well.
+
+The error "(bug seems to be opened, please check)" when running `make test-suite` means that a test in `bugs/closed` failed to compile.
+
+There are also output tests in `test-suite/output` which consist of a `.v` file and a `.out` file with the expected output.
diff --git a/test-suite/bugs/closed/6297.v b/test-suite/bugs/closed/6297.v
new file mode 100644
index 000000000..a28607058
--- /dev/null
+++ b/test-suite/bugs/closed/6297.v
@@ -0,0 +1,8 @@
+Set Printing Universes.
+
+(* Error: Anomaly "Uncaught exception "Anomaly: Incorrect universe Set
+ declared for inductive type, inferred level is max(Prop, Set+1)."."
+ Please report at http://coq.inria.fr/bugs/. *)
+Fail Record LTS: Set :=
+ lts { St: Set;
+ init: St }.
diff --git a/test-suite/bugs/closed/6490.v b/test-suite/bugs/closed/6490.v
new file mode 100644
index 000000000..dcf9ff29e
--- /dev/null
+++ b/test-suite/bugs/closed/6490.v
@@ -0,0 +1,4 @@
+Inductive Foo (A' := I) (B : Type) := foo : Foo B.
+
+Goal Foo True. dtauto. Qed.
+Goal Foo True. firstorder. Qed.
diff --git a/test-suite/bugs/closed/6617.v b/test-suite/bugs/closed/6617.v
new file mode 100644
index 000000000..9cabd62d4
--- /dev/null
+++ b/test-suite/bugs/closed/6617.v
@@ -0,0 +1,34 @@
+Definition MR {T M : Type} :=
+fun (R : M -> M -> Prop) (m : T -> M) (x y : T) => R (m x) (m y).
+
+Set Primitive Projections.
+
+Record sigma {A : Type} {B : A -> Type} : Type := sigmaI
+ { pr1 : A; pr2 : B pr1 }.
+
+Axiom F : forall {A : Type} {R : A -> A -> Prop},
+ (forall x, (forall y, R y x -> unit) -> unit) -> forall (x : A), unit.
+
+Definition foo (A : Type) (l : list A) :=
+ let y := {| pr1 := A; pr2 := l |} in
+ let bar := MR lt (fun p : sigma =>
+ (fix Ffix (x : list (pr1 p)) : nat :=
+ match x with
+ | nil => 0
+ | cons _ x1 => S (Ffix x1)
+ end) (pr2 p)) in
+fun (_ : bar y y) =>
+F (fun (r : sigma)
+ (X : forall q : sigma, bar q r -> unit) => tt).
+
+Definition fooT (A : Type) (l : list A) : Type :=
+ ltac:(let ty := type of (foo A l) in exact ty).
+Parameter P : forall A l, fooT A l -> Prop.
+
+Goal forall A l, P A l (foo A l).
+Proof.
+ intros; unfold foo.
+ Fail match goal with
+ | [ |- context [False]] => idtac
+ end.
+Admitted.
diff --git a/test-suite/success/BracketsWithGoalSelector.v b/test-suite/success/BracketsWithGoalSelector.v
new file mode 100644
index 000000000..ed035f521
--- /dev/null
+++ b/test-suite/success/BracketsWithGoalSelector.v
@@ -0,0 +1,16 @@
+Goal forall A B, B \/ A -> A \/ B.
+Proof.
+ intros * [HB | HA].
+ 2: {
+ left.
+ exact HA.
+ Fail right. (* No such goal. Try unfocusing with "}". *)
+ }
+ Fail 2: { (* Non-existent goal. *)
+ idtac. (* The idtac is to get a dot, so that IDEs know to stop there. *)
+ 1:{ (* Syntactic test: no space before bracket. *)
+ right.
+ exact HB.
+Fail Qed.
+ }
+Qed.
diff --git a/test-suite/success/abstract_poly.v b/test-suite/success/abstract_poly.v
index b736b734f..aa8da5336 100644
--- a/test-suite/success/abstract_poly.v
+++ b/test-suite/success/abstract_poly.v
@@ -17,4 +17,4 @@ intros m n P e p.
abstract (rewrite e in p; exact p).
Defined.
-Check bar_subproof@{Set Set Set}.
+Check bar_subproof@{Set Set}.
diff --git a/test-suite/success/dtauto-let-deps.v b/test-suite/success/dtauto-let-deps.v
new file mode 100644
index 000000000..094b2f8b3
--- /dev/null
+++ b/test-suite/success/dtauto-let-deps.v
@@ -0,0 +1,24 @@
+(*
+This test is sensitive to changes in which let-ins are expanded when checking
+for dependencies in constructors.
+If the (x := X) is not reduced, Foo1 won't be recognized as a conjunction,
+and if the (y := X) is reduced, Foo2 will be recognized as a conjunction.
+
+This tests the behavior of engine/termops.ml : prod_applist_assum,
+which is currently specified to reduce exactly the parameters.
+
+If dtauto is changed to reduce lets in constructors before checking dependency,
+this test will need to be changed.
+*)
+
+Context (P Q : Type).
+Inductive Foo1 (X : Type) (x := X) := foo1 : let y := X in P -> Q -> Foo1 x.
+Inductive Foo2 (X : Type) (x := X) := foo2 : let y := X in P -> Q -> Foo2 y.
+
+Goal P -> Q -> Foo1 nat.
+solve [dtauto].
+Qed.
+
+Goal P -> Q -> Foo2 nat.
+Fail solve [dtauto].
+Abort.
diff --git a/tools/coqdep.ml b/tools/coqdep.ml
index fd4be08b1..2433cb1d0 100644
--- a/tools/coqdep.ml
+++ b/tools/coqdep.ml
@@ -455,7 +455,7 @@ let usage () =
eprintf " -coqlib dir : set the coq standard library directory\n";
eprintf " -suffix s : \n";
eprintf " -slash : deprecated, no effect\n";
- eprintf " -dyndep (opt|byte|both|no|var) : set how dependencies over ML modules are printed";
+ eprintf " -dyndep (opt|byte|both|no|var) : set how dependencies over ML modules are printed\n";
exit 1
let split_period = Str.split (Str.regexp (Str.quote "."))
diff --git a/tools/md5sum.ml b/tools/md5sum.ml
new file mode 100644
index 000000000..2fdcacc83
--- /dev/null
+++ b/tools/md5sum.ml
@@ -0,0 +1,24 @@
+let get_content file =
+ let ic = open_in_bin file in
+ let buf = Buffer.create 2048 in
+ let rec fill () =
+ match input_char ic with
+ | '\r' -> fill () (* NOTE: handles the case on Windows where the
+ git checkout has included return characters.
+ See: https://github.com/coq/coq/pull/6305 *)
+ | c -> Buffer.add_char buf c; fill ()
+ | exception End_of_file -> close_in ic; Buffer.contents buf
+ in
+ fill ()
+
+let () =
+ match Sys.argv with
+ | [|_; file|] ->
+ let content = get_content file in
+ let md5 = Digest.to_hex (Digest.string content) in
+ print_string (md5 ^ " " ^ file)
+ | _ ->
+ prerr_endline "Error: This program needs exactly one parameter.";
+ prerr_endline "Usage: ocaml md5sum.ml <FILE>";
+ prerr_endline "Print MD5 (128-bit) checksum of the file content modulo \\r.";
+ exit 1
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index ded0d97e6..3358951f4 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -1846,11 +1846,8 @@ let vernac_unfocused () =
user_err Pp.(str "The proof is not fully unfocused.")
-(* BeginSubproof / EndSubproof.
- BeginSubproof (vernac_subproof) focuses on the first goal, or the goal
- given as argument.
- EndSubproof (vernac_end_subproof) unfocuses from a BeginSubproof, provided
- that the proof of the goal has been completed.
+(* "{" focuses on the first goal, "n: {" focuses on the n-th goal
+ "}" unfocuses, provided that the proof of the goal has been completed.
*)
let subproof_kind = Proof.new_focus_kind ()
let subproof_cond = Proof.done_cond subproof_kind
@@ -1859,7 +1856,9 @@ let vernac_subproof gln =
Proof_global.simple_with_current_proof (fun _ p ->
match gln with
| None -> Proof.focus subproof_cond () 1 p
- | Some n -> Proof.focus subproof_cond () n p)
+ | Some (SelectNth n) -> Proof.focus subproof_cond () n p
+ | _ -> user_err ~hdr:"bracket_selector"
+ (str "Brackets only support the single numbered goal selector."))
let vernac_end_subproof () =
Proof_global.simple_with_current_proof (fun _ p ->
@@ -1959,11 +1958,6 @@ let interp ?proof ~atts ~st c =
(* This one is possible to handle here *)
| VernacAbort id -> CErrors.user_err (str "Abort cannot be used through the Load command")
- (* Handled elsewhere *)
- | VernacProgram _
- | VernacPolymorphic _
- | VernacLocal _ -> assert false
-
(* Syntax *)
| VernacSyntaxExtension (infix, sl) ->
vernac_syntax_extension atts infix sl
@@ -2200,10 +2194,30 @@ let with_fail st b f =
let interp ?(verbosely=true) ?proof ~st (loc,c) =
let orig_univ_poly = Flags.is_universe_polymorphism () in
let orig_program_mode = Flags.is_program_mode () in
+ let flags f atts =
+ List.fold_left
+ (fun (polymorphism, atts) f ->
+ match f with
+ | VernacProgram when not atts.program ->
+ (polymorphism, { atts with program = true })
+ | VernacProgram ->
+ user_err Pp.(str "Program mode specified twice")
+ | VernacPolymorphic b when polymorphism = None ->
+ (Some b, atts)
+ | VernacPolymorphic _ ->
+ user_err Pp.(str "Polymorphism specified twice")
+ | VernacLocal b when Option.is_empty atts.locality ->
+ (polymorphism, { atts with locality = Some b })
+ | VernacLocal _ ->
+ user_err Pp.(str "Locality specified twice")
+ )
+ (None, atts)
+ f
+ in
let rec control = function
- | VernacExpr v ->
- let atts = { loc; locality = None; polymorphic = false; } in
- aux ~atts orig_program_mode v
+ | VernacExpr (f, v) ->
+ let (polymorphism, atts) = flags f { loc; locality = None; polymorphic = false; program = orig_program_mode; } in
+ aux ~polymorphism ~atts v
| VernacFail v -> with_fail st true (fun () -> control v)
| VernacTimeout (n,v) ->
current_timeout := Some n;
@@ -2213,25 +2227,8 @@ let interp ?(verbosely=true) ?proof ~st (loc,c) =
| VernacTime (batch, (_loc,v)) ->
System.with_time ~batch control v;
- and aux ?polymorphism ~atts isprogcmd = function
-
- | VernacProgram c when not isprogcmd ->
- aux ?polymorphism ~atts true c
-
- | VernacProgram _ ->
- user_err Pp.(str "Program mode specified twice")
-
- | VernacPolymorphic (b, c) when polymorphism = None ->
- aux ~polymorphism:b ~atts:atts isprogcmd c
-
- | VernacPolymorphic (b, c) ->
- user_err Pp.(str "Polymorphism specified twice")
-
- | VernacLocal (b, c) when Option.is_empty atts.locality ->
- aux ?polymorphism ~atts:{atts with locality = Some b} isprogcmd c
-
- | VernacLocal _ ->
- user_err Pp.(str "Locality specified twice")
+ and aux ~polymorphism ~atts : _ -> unit =
+ function
| VernacLoad (_,fname) -> vernac_load control fname
@@ -2240,7 +2237,7 @@ let interp ?(verbosely=true) ?proof ~st (loc,c) =
check_vernac_supports_polymorphism c polymorphism;
let polymorphic = Option.default (Flags.is_universe_polymorphism ()) polymorphism in
Flags.make_universe_polymorphism polymorphic;
- Obligations.set_program_mode isprogcmd;
+ Obligations.set_program_mode atts.program;
try
vernac_timeout begin fun () ->
let atts = { atts with polymorphic } in
@@ -2249,7 +2246,7 @@ let interp ?(verbosely=true) ?proof ~st (loc,c) =
else Flags.silently (interp ?proof ~atts ~st) c;
(* If the command is `(Un)Set Program Mode` or `(Un)Set Universe Polymorphism`,
we should not restore the previous state of the flag... *)
- if orig_program_mode || not !Flags.program_mode || isprogcmd then
+ if orig_program_mode || not !Flags.program_mode || atts.program then
Flags.program_mode := orig_program_mode;
if (Flags.is_universe_polymorphism() = polymorphic) then
Flags.make_universe_polymorphism orig_univ_poly;
diff --git a/vernac/vernacinterp.ml b/vernac/vernacinterp.ml
index c0b93c163..c40ca27db 100644
--- a/vernac/vernacinterp.ml
+++ b/vernac/vernacinterp.ml
@@ -16,6 +16,7 @@ type atts = {
loc : Loc.t option;
locality : bool option;
polymorphic : bool;
+ program : bool;
}
type 'a vernac_command = 'a -> atts:atts -> st:Vernacstate.t -> Vernacstate.t
diff --git a/vernac/vernacinterp.mli b/vernac/vernacinterp.mli
index ab3d4bfc5..c5e610f89 100644
--- a/vernac/vernacinterp.mli
+++ b/vernac/vernacinterp.mli
@@ -14,6 +14,7 @@ type atts = {
loc : Loc.t option;
locality : bool option;
polymorphic : bool;
+ program : bool;
}
type 'a vernac_command = 'a -> atts:atts -> st:Vernacstate.t -> Vernacstate.t
diff --git a/vernac/vernacprop.ml b/vernac/vernacprop.ml
index 4fdc78dd2..3932d1c7b 100644
--- a/vernac/vernacprop.ml
+++ b/vernac/vernacprop.ml
@@ -12,14 +12,14 @@
open Vernacexpr
let rec under_control = function
- | VernacExpr c -> c
+ | VernacExpr (_, c) -> c
| VernacRedirect (_,(_,c))
| VernacTime (_,(_,c))
| VernacFail c
| VernacTimeout (_,c) -> under_control c
let rec has_Fail = function
- | VernacExpr c -> false
+ | VernacExpr _ -> false
| VernacRedirect (_,(_,c))
| VernacTime (_,(_,c))
| VernacTimeout (_,c) -> has_Fail c
@@ -45,8 +45,8 @@ let rec is_deep_navigation_vernac = function
(* NB: Reset is now allowed again as asked by A. Chlipala *)
let is_reset = function
- | VernacExpr VernacResetInitial
- | VernacExpr (VernacResetName _) -> true
+ | VernacExpr ( _, VernacResetInitial)
+ | VernacExpr (_, VernacResetName _) -> true
| _ -> false
let is_debug cmd = match under_control cmd with