aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--.gitlab-ci.yml19
-rw-r--r--checker/indtypes.ml12
-rw-r--r--checker/inductive.ml10
-rw-r--r--checker/modops.ml22
-rw-r--r--checker/subtyping.ml4
-rw-r--r--default.nix80
-rw-r--r--dev/ci/ci-common.sh2
-rw-r--r--dev/ci/user-overlays/07746-cleanup-unused-various.sh6
-rw-r--r--dev/doc/changes.md83
-rw-r--r--doc/sphinx/language/cic.rst231
-rw-r--r--engine/evarutil.ml4
-rw-r--r--engine/evarutil.mli4
-rw-r--r--engine/evd.ml6
-rw-r--r--engine/evd.mli4
-rw-r--r--engine/uState.ml2
-rw-r--r--engine/univGen.ml4
-rw-r--r--engine/univGen.mli2
-rw-r--r--engine/univSubst.ml12
-rw-r--r--engine/univSubst.mli2
-rw-r--r--engine/universes.mli4
-rw-r--r--grammar/coqpp_ast.mli21
-rw-r--r--grammar/coqpp_lex.mll4
-rw-r--r--grammar/coqpp_main.ml237
-rw-r--r--grammar/coqpp_parse.mly76
-rw-r--r--grammar/tacextend.mlp9
-rw-r--r--ide/idetop.ml5
-rw-r--r--kernel/cooking.ml2
-rw-r--r--kernel/cooking.mli3
-rw-r--r--kernel/indtypes.ml15
-rw-r--r--kernel/inductive.ml10
-rw-r--r--kernel/modops.ml4
-rw-r--r--kernel/nativecode.ml14
-rw-r--r--kernel/nativecode.mli2
-rw-r--r--kernel/nativelibrary.ml2
-rw-r--r--kernel/subtyping.ml4
-rw-r--r--kernel/term_typing.ml6
-rw-r--r--lib/cWarnings.ml6
-rw-r--r--lib/util.ml4
-rw-r--r--library/libobject.ml16
-rw-r--r--library/library.ml36
-rw-r--r--plugins/btauto/g_btauto.mlg (renamed from plugins/btauto/g_btauto.ml4)6
-rw-r--r--plugins/cc/g_congruence.mlg (renamed from plugins/cc/g_congruence.ml4)14
-rw-r--r--plugins/fourier/g_fourier.mlg (renamed from plugins/fourier/g_fourier.ml4)6
-rw-r--r--plugins/funind/functional_principles_types.ml11
-rw-r--r--plugins/ltac/coretactics.mlg (renamed from plugins/ltac/coretactics.ml4)184
-rw-r--r--plugins/ltac/g_eqdecide.mlg (renamed from plugins/ltac/g_eqdecide.ml4)8
-rw-r--r--plugins/ltac/pptactic.ml51
-rw-r--r--plugins/ltac/rewrite.ml2
-rw-r--r--plugins/ltac/taccoerce.ml23
-rw-r--r--plugins/ltac/taccoerce.mli14
-rw-r--r--plugins/ltac/tacentries.ml30
-rw-r--r--plugins/ltac/tacentries.mli4
-rw-r--r--plugins/ltac/tacinterp.ml14
-rw-r--r--plugins/micromega/g_micromega.mlg (renamed from plugins/micromega/g_micromega.ml4)38
-rw-r--r--plugins/nsatz/g_nsatz.mlg (renamed from plugins/nsatz/g_nsatz.ml4)6
-rw-r--r--plugins/omega/coq_omega.ml10
-rw-r--r--plugins/omega/g_omega.mlg (renamed from plugins/omega/g_omega.ml4)9
-rw-r--r--plugins/quote/g_quote.mlg (renamed from plugins/quote/g_quote.ml4)16
-rw-r--r--plugins/romega/g_romega.mlg (renamed from plugins/romega/g_romega.ml4)12
-rw-r--r--plugins/rtauto/g_rtauto.mlg (renamed from plugins/rtauto/g_rtauto.ml4)5
-rw-r--r--pretyping/glob_ops.ml4
-rw-r--r--pretyping/indrec.ml9
-rw-r--r--pretyping/typing.ml2
-rw-r--r--printing/pputils.ml2
-rw-r--r--shell.nix4
-rw-r--r--tactics/elimschemes.ml2
-rw-r--r--tactics/eqschemes.ml6
-rw-r--r--tactics/inv.ml2
-rw-r--r--tactics/leminv.ml2
-rw-r--r--tools/CoqMakefile.in10
-rw-r--r--tools/coq_makefile.ml4
-rw-r--r--tools/coqdoc/index.ml2
-rw-r--r--tools/coqdoc/index.mli2
-rw-r--r--tools/coqdoc/output.ml4
74 files changed, 875 insertions, 632 deletions
diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml
index 8469d6119..773a89a46 100644
--- a/.gitlab-ci.yml
+++ b/.gitlab-ci.yml
@@ -45,6 +45,9 @@ before_script:
- opam list
- opam config list
+after_script:
+ - echo "The build completed normally (not a runner failure)."
+
################ GITLAB CACHING ######################
# - use artifacts between jobs #
######################################################
@@ -233,6 +236,22 @@ windows32:
variables:
ARCH: "32"
+pkg:nix:
+ image: nixorg/nix:latest # Minimal NixOS image which doesn't even contain git
+ stage: test
+ variables:
+ GIT_STRATEGY: none
+ dependencies: [] # We don't need to download build artifacts
+ before_script: [] # We don't want to use the shared 'before_script'
+ script:
+ - export TMPDIR=$PWD
+ - nix-build "$CI_PROJECT_URL/-/archive/$CI_COMMIT_SHA.tar.gz" -K
+ artifacts:
+ name: "$CI_JOB_NAME.logs"
+ when: on_failure
+ paths:
+ - nix-build-coq.drv-0/*/test-suite/logs
+
warnings:base:
<<: *warnings-template
diff --git a/checker/indtypes.ml b/checker/indtypes.ml
index bcb71fe55..8f11e01c3 100644
--- a/checker/indtypes.ml
+++ b/checker/indtypes.ml
@@ -221,7 +221,7 @@ let allowed_sorts issmall isunit s =
-let compute_elim_sorts env_ar params mib arity lc =
+let compute_elim_sorts env_ar params arity lc =
let inst = extended_rel_list 0 params in
let env_params = push_rel_context params env_ar in
let lc = Array.map
@@ -239,7 +239,7 @@ let compute_elim_sorts env_ar params mib arity lc =
allowed_sorts small unit s
-let typecheck_one_inductive env params mib mip =
+let typecheck_one_inductive env params mip =
(* mind_typename and mind_consnames not checked *)
(* mind_reloc_tbl, mind_nb_constant, mind_nb_args not checked (VM) *)
(* mind_arity_ctxt, mind_arity, mind_nrealargs DONE (typecheck_arity) *)
@@ -256,7 +256,7 @@ let typecheck_one_inductive env params mib mip =
Array.iter2 check_cons_args mip.mind_nf_lc mip.mind_consnrealdecls;
(* mind_kelim: checked by positivity criterion ? *)
let sorts =
- compute_elim_sorts env params mib mip.mind_arity mip.mind_nf_lc in
+ compute_elim_sorts env params mip.mind_arity mip.mind_nf_lc in
let reject_sort s = not (List.mem_f family_equal s sorts) in
if List.exists reject_sort mip.mind_kelim then
failwith "elimination not allowed";
@@ -355,7 +355,7 @@ let lambda_implicit_lift n a =
(* This removes global parameters of the inductive types in lc (for
nested inductive types only ) *)
-let abstract_mind_lc env ntyps npars lc =
+let abstract_mind_lc ntyps npars lc =
if npars = 0 then
lc
else
@@ -448,7 +448,7 @@ let check_positivity_one (env, _,ntypes,_ as ienv) hyps nrecp (_,i as ind) indlc
let auxntyp = mib.mind_ntypes in
if auxntyp <> 1 then raise (IllFormedInd (LocalNonPos n));
(* The nested inductive type with parameters removed *)
- let auxlcvect = abstract_mind_lc env auxntyp auxnpar mip.mind_nf_lc in
+ let auxlcvect = abstract_mind_lc auxntyp auxnpar mip.mind_nf_lc in
(* Extends the environment with a variable corresponding to
the inductive def *)
let (env',_,_,_ as ienv') = ienv_push_inductive ienv ((mi,u),lpar) in
@@ -625,7 +625,7 @@ let check_inductive env kn mib =
(* - check arities *)
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;
+ Array.iter (typecheck_one_inductive env_ar params) mib.mind_packets;
(* check the inferred subtyping relation *)
let () =
match mib.mind_universes with
diff --git a/checker/inductive.ml b/checker/inductive.ml
index b1edec556..d15380643 100644
--- a/checker/inductive.ml
+++ b/checker/inductive.ml
@@ -797,7 +797,7 @@ let rec subterm_specif renv stack t =
| Lambda (x,a,b) ->
assert (l=[]);
- let spec,stack' = extract_stack renv a stack in
+ let spec,stack' = extract_stack stack in
subterm_specif (push_var renv (x,a,spec)) stack' b
(* Metas and evars are considered OK *)
@@ -813,7 +813,7 @@ and stack_element_specif = function
|SClosure (h_renv,h) -> lazy_subterm_specif h_renv [] h
|SArg x -> x
-and extract_stack renv a = function
+and extract_stack = function
| [] -> Lazy.from_val Not_subterm , []
| h::t -> stack_element_specif h, t
@@ -845,7 +845,7 @@ let error_illegal_rec_call renv fx (arg_renv,arg) =
let error_partial_apply renv fx =
raise (FixGuardError (renv.env,NotEnoughArgumentsForFixCall fx))
-let filter_stack_domain env ci p stack =
+let filter_stack_domain env p stack =
let absctx, ar = dest_lam_assum env p in
(* Optimization: if the predicate is not dependent, no restriction is needed
and we avoid building the recargs tree. *)
@@ -925,7 +925,7 @@ let check_one_fix renv recpos trees def =
let case_spec = branches_specif renv
(lazy_subterm_specif renv [] c_0) ci in
let stack' = push_stack_closures renv l stack in
- let stack' = filter_stack_domain renv.env ci p stack' in
+ let stack' = filter_stack_domain renv.env p stack' in
Array.iteri (fun k br' ->
let stack_br = push_stack_args case_spec.(k) stack' in
check_rec_call renv stack_br br') lrest
@@ -968,7 +968,7 @@ let check_one_fix renv recpos trees def =
| Lambda (x,a,b) ->
assert (l = []);
check_rec_call renv [] a ;
- let spec, stack' = extract_stack renv a stack in
+ let spec, stack' = extract_stack stack in
check_rec_call (push_var renv (x,a,spec)) stack' b
| Prod (x,a,b) ->
diff --git a/checker/modops.ml b/checker/modops.ml
index c7ad0977a..b92d7bbf1 100644
--- a/checker/modops.ml
+++ b/checker/modops.ml
@@ -80,7 +80,7 @@ and add_module mb env =
let add_module_type mp mtb env = add_module (module_body_of_type mp mtb) env
-let strengthen_const mp_from l cb resolver =
+let strengthen_const mp_from l cb =
match cb.const_body with
| Def _ -> cb
| _ ->
@@ -104,34 +104,34 @@ and strengthen_body : 'a. (_ -> 'a) -> _ -> _ -> 'a generic_module_body -> 'a ge
match mb.mod_type with
| MoreFunctor _ -> mb
| NoFunctor sign ->
- let resolve_out,sign_out = strengthen_sig mp_from sign mp_to mb.mod_delta
+ let resolve_out,sign_out = strengthen_sig mp_from sign mp_to
in
{ mb with
mod_expr = mk_expr mp_to;
mod_type = NoFunctor sign_out;
mod_delta = resolve_out }
-and strengthen_sig mp_from sign mp_to resolver =
+and strengthen_sig mp_from sign mp_to =
match sign with
| [] -> empty_delta_resolver,[]
| (l,SFBconst cb) :: rest ->
- let item' = l,SFBconst (strengthen_const mp_from l cb resolver) in
- let resolve_out,rest' = strengthen_sig mp_from rest mp_to resolver in
+ let item' = l,SFBconst (strengthen_const mp_from l cb) in
+ let resolve_out,rest' = strengthen_sig mp_from rest mp_to in
resolve_out,item'::rest'
| (_,SFBmind _ as item):: rest ->
- let resolve_out,rest' = strengthen_sig mp_from rest mp_to resolver in
+ let resolve_out,rest' = strengthen_sig mp_from rest mp_to in
resolve_out,item::rest'
| (l,SFBmodule mb) :: rest ->
let mp_from' = MPdot (mp_from,l) in
let mp_to' = MPdot(mp_to,l) in
let mb_out = strengthen_mod mp_from' mp_to' mb in
let item' = l,SFBmodule (mb_out) in
- let resolve_out,rest' = strengthen_sig mp_from rest mp_to resolver in
+ let resolve_out,rest' = strengthen_sig mp_from rest mp_to in
resolve_out (*add_delta_resolver resolve_out mb.mod_delta*),
- item':: rest'
- | (l,SFBmodtype mty as item) :: rest ->
- let resolve_out,rest' = strengthen_sig mp_from rest mp_to resolver in
- resolve_out,item::rest'
+ item':: rest'
+ | (_,SFBmodtype _ as item) :: rest ->
+ let resolve_out,rest' = strengthen_sig mp_from rest mp_to in
+ resolve_out,item::rest'
let strengthen mtb mp =
strengthen_body ignore mtb.mod_mp mp mtb
diff --git a/checker/subtyping.ml b/checker/subtyping.ml
index 6d0d6f6c6..3f7f84470 100644
--- a/checker/subtyping.ml
+++ b/checker/subtyping.ml
@@ -217,7 +217,7 @@ let check_inductive env mp1 l info1 mib2 spec2 subst1 subst2=
let _ = Array.map2_i check_cons_types mib1.mind_packets mib2.mind_packets
in ()
-let check_constant env mp1 l info1 cb2 spec2 subst1 subst2 =
+let check_constant env l info1 cb2 spec2 subst1 subst2 =
let error () = error_not_match l spec2 in
let check_conv f = check_conv_error error f in
let check_type env t1 t2 = check_conv conv_leq env t1 t2 in
@@ -281,7 +281,7 @@ and check_signatures env mp1 sig1 sig2 subst1 subst2 =
let check_one_body (l,spec2) =
match spec2 with
| SFBconst cb2 ->
- check_constant env mp1 l (get_obj mp1 map1 l)
+ check_constant env l (get_obj mp1 map1 l)
cb2 spec2 subst1 subst2
| SFBmind mib2 ->
check_inductive env mp1 l (get_obj mp1 map1 l)
diff --git a/default.nix b/default.nix
index a2645f4fc..1be274081 100644
--- a/default.nix
+++ b/default.nix
@@ -9,23 +9,27 @@
# nix-shell supports the --arg option (see Nix doc) that allows you for
# instance to do this:
-# $ nix-shell --arg ocamlPackages "(import <nixpkgs> {}).ocamlPackages_latest" --arg buildIde false
+# $ nix-shell --arg ocamlPackages "(import <nixpkgs> {}).ocaml-ng.ocamlPackages_4_05" --arg buildIde false
# You can also compile Coq and "install" it by running:
# $ make clean # (only needed if you have left-over compilation files)
# $ nix-build
# at the root of the Coq repository.
# nix-build also supports the --arg option, so you will be able to do:
-# $ nix-build --arg doCheck false
+# $ nix-build --arg doInstallCheck false
# if you want to speed up things by not running the test-suite.
# Once the build is finished, you will find, in the current directory,
# a symlink to where Coq was installed.
-{ pkgs ? (import <nixpkgs> {})
+{ pkgs ?
+ (import (fetchTarball {
+ url = "https://github.com/NixOS/nixpkgs/archive/060a98e9f4ad879492e48d63e887b0b6db26299e.tar.gz";
+ sha256 = "1lzvp3md0hf6kp2bvc6dbzh40navlyd51qlns9wbkz6lqk3lgf6j";
+ }) {})
, ocamlPackages ? pkgs.ocaml-ng.ocamlPackages_4_06
, buildIde ? true
, buildDoc ? true
-, doCheck ? true
+, doInstallCheck ? true
}:
with pkgs;
@@ -36,64 +40,50 @@ stdenv.mkDerivation rec {
name = "coq";
buildInputs = [
-
- # Coq dependencies
hostname
- ] ++ (with ocamlPackages; [
- ocaml
- findlib
- camlp5_strict
- num
-
- ]) ++ (if buildIde then [
-
- # CoqIDE dependencies
- ocamlPackages.lablgtk
-
- ] else []) ++ (if buildDoc then [
-
+ python2 time # coq-makefile timing tools
+ ]
+ ++ (with ocamlPackages; [ ocaml findlib camlp5_strict num ])
+ ++ optional buildIde ocamlPackages.lablgtk
+ ++ optionals buildDoc [
# Sphinx doc dependencies
pkgconfig (python3.withPackages
(ps: [ ps.sphinx ps.sphinx_rtd_theme ps.pexpect ps.beautifulsoup4
ps.antlr4-python3-runtime ps.sphinxcontrib-bibtex ]))
- antlr4
-
- ] else []) ++ (if doCheck then
-
+ antlr4
+ ]
+ ++ optionals doInstallCheck (
# Test-suite dependencies
# ncurses is required to build an OCaml REPL
optional (!versionAtLeast ocaml.version "4.07") ncurses
- ++ [
- python
- rsync
- which
- ocamlPackages.ounit
-
- ] else []) ++ (if lib.inNixShell then [
- ocamlPackages.merlin
- ocamlPackages.ocp-indent
- ocamlPackages.ocp-index
-
- # Dependencies of the merging script
- jq
- curl
- git
- gnupg
-
- ] else []);
+ ++ [ ocamlPackages.ounit rsync which ]
+ )
+ ++ optionals lib.inNixShell (
+ [ jq curl git gnupg ] # Dependencies of the merging script
+ ++ (with ocamlPackages; [ merlin ocp-indent ocp-index ]) # Dev tools
+ );
src =
if lib.inNixShell then null
else
with builtins; filterSource
- (path: _: !elem (baseNameOf path) [".git" "result" "bin"]) ./.;
+ (path: _:
+ !elem (baseNameOf path) [".git" "result" "bin" "_build_ci"]) ./.;
prefixKey = "-prefix ";
- buildFlags = [ "world" ] ++ optional buildDoc "doc-html";
+ buildFlags = [ "world" "byte" ] ++ optional buildDoc "doc-html";
+
+ installTargets =
+ [ "install" "install-byte" ] ++ optional buildDoc "install-doc-html";
+
+ inherit doInstallCheck;
- installTargets = [ "install" ] ++ optional buildDoc "install-doc-html";
+ preInstallCheck = ''
+ patchShebangs tools/
+ patchShebangs test-suite/
+ '';
- inherit doCheck;
+ installCheckTarget = [ "check" ];
}
diff --git a/dev/ci/ci-common.sh b/dev/ci/ci-common.sh
index 85df249d3..a68cd0933 100644
--- a/dev/ci/ci-common.sh
+++ b/dev/ci/ci-common.sh
@@ -69,7 +69,7 @@ git_checkout()
if [ ! -d .git ] ; then git clone "${_DEPTH[@]}" "${_URL}" . ; fi && \
echo "Checking out ${_DEST}" && \
git fetch "${_URL}" "${_BRANCH}" && \
- git checkout "${_COMMIT}" && \
+ git -c advice.detachedHead=false checkout "${_COMMIT}" && \
echo "${_DEST}: $(git log -1 --format='%s | %H | %cd | %aN')" )
}
diff --git a/dev/ci/user-overlays/07746-cleanup-unused-various.sh b/dev/ci/user-overlays/07746-cleanup-unused-various.sh
new file mode 100644
index 000000000..8688b0d53
--- /dev/null
+++ b/dev/ci/user-overlays/07746-cleanup-unused-various.sh
@@ -0,0 +1,6 @@
+#!/bin/sh
+
+if [ "$CI_PULL_REQUEST" = "7746" ] || [ "$CI_BRANCH" = "cleanup-unused-various" ]; then
+ Equations_CI_BRANCH="adapt-unused"
+ Equations_CI_GITURL="https://github.com/SkySkimmer/Coq-Equations.git"
+fi
diff --git a/dev/doc/changes.md b/dev/doc/changes.md
index 4177513a1..6d5023405 100644
--- a/dev/doc/changes.md
+++ b/dev/doc/changes.md
@@ -95,19 +95,73 @@ Primitive number parsers
### Transitioning away from Camlp5
-In an effort to reduce dependency on Camlp5, the use of Camlp5 GEXTEND macro
-is discouraged. Most plugin writers do not need this low-level interface, but
-for those who do, the transition path is somewhat straightforward. To convert
-a ml4 file containing only standard OCaml with GEXTEND statements, the following
-should be enough:
-- replace its "ml4" extension with "mlg"
-- replace GEXTEND with GRAMMAR EXTEND
-- wrap every occurrence of OCaml code into braces { }
+In an effort to reduce dependency on camlp5, the use of several grammar macros
+is discouraged. Coq is now shipped with its own preprocessor, called coqpp,
+which serves the same purpose as camlp5.
+
+To perform the transition to coqpp macros, one first needs to change the
+extension of a macro file from `.ml4` to `.mlg`. Not all camlp5 macros are
+handled yet.
+
+Due to parsing constraints, the syntax of the macros is slightly different, but
+updating the source code is mostly a matter of straightforward
+search-and-replace. The main differences are summarized below.
+
+#### OCaml code
+
+Every piece of toplevel OCaml code needs to be wrapped into braces.
For instance, code of the form
```
let myval = 0
+```
+should be turned into
+```
+{
+let myval = 0
+}
+```
+
+#### TACTIC EXTEND
+
+Steps to perform:
+- replace the brackets enclosing OCaml code in actions with braces
+- if not there yet, add a leading `|̀ to the first rule
+
+For instance, code of the form:
+```
+TACTIC EXTEND my_tac
+ [ "tac1" int_or_var(i) tactic(t) ] -> [ mytac1 ist i t ]
+| [ "tac2" tactic(t) ] -> [ mytac2 t ]
+END
+```
+should be turned into
+```
+TACTIC EXTEND my_tac
+| [ "tac1" int_or_var(i) tactic(t) ] -> { mytac1 ist i t }
+| [ "tac2" tactic(t) ] -> { mytac2 t }
+END
+```
+
+#### VERNAC EXTEND
+
+Not handled yet.
+
+#### ARGUMENT EXTEND
+
+Not handled yet.
+
+#### GEXTEND
+Most plugin writers do not need this low-level interface, but for the sake of
+completeness we document it.
+
+Steps to perform are:
+- replace GEXTEND with GRAMMAR EXTEND
+- wrap every occurrence of OCaml code in actions into braces { }
+
+For instance, code of the form
+```
GEXTEND Gram
GLOBAL: my_entry;
@@ -119,10 +173,6 @@ END
```
should be turned into
```
-{
-let myval = 0
-}
-
GRAMMAR EXTEND Gram
GLOBAL: my_entry;
@@ -133,9 +183,12 @@ my_entry:
END
```
-Currently mlg files can only contain GRAMMAR EXTEND statements. They do not
-support TACTIC, VERNAC nor ARGUMENT EXTEND. In case you have a file containing
-both kinds at the same time, it is recommended splitting it in two.
+Caveats:
+- No `GLOBAL` entries mean that they are all local, while camlp5 special-cases
+ this as a shorthand for all global entries. Solution: always define a `GLOBAL`
+ section.
+- No complex patterns allowed in token naming. Solution: match on it inside the
+ OCaml quotation.
## Changes between Coq 8.7 and Coq 8.8
diff --git a/doc/sphinx/language/cic.rst b/doc/sphinx/language/cic.rst
index f6bab0267..b01a4ef0f 100644
--- a/doc/sphinx/language/cic.rst
+++ b/doc/sphinx/language/cic.rst
@@ -721,67 +721,68 @@ called the *context of parameters*. Furthermore, we must have that
each :math:`T` in :math:`(t:T)∈Γ_I` can be written as: :math:`∀Γ_P,∀Γ_{\mathit{Arr}(t)}, S` where
:math:`Γ_{\mathit{Arr}(t)}` is called the *Arity* of the inductive type t and :math:`S` is called
the sort of the inductive type t (not to be confused with :math:`\Sort` which is the set of sorts).
-** Examples** The declaration for parameterized lists is:
-
-.. math::
- \ind{1}{[\List:\Set→\Set]}{\left[\begin{array}{rcl}
- \Nil & : & \forall A:\Set,\List~A \\
- \cons & : & \forall A:\Set, A→ \List~A→ \List~A
- \end{array}
- \right]}
-
-which corresponds to the result of the |Coq| declaration:
.. example::
- .. coqtop:: in
-
- Inductive list (A:Set) : Set :=
- | nil : list A
- | cons : A -> list A -> list A.
+ The declaration for parameterized lists is:
-The declaration for a mutual inductive definition of tree and forest
-is:
+ .. math::
+ \ind{1}{[\List:\Set→\Set]}{\left[\begin{array}{rcl}
+ \Nil & : & \forall A:\Set,\List~A \\
+ \cons & : & \forall A:\Set, A→ \List~A→ \List~A
+ \end{array}
+ \right]}
-.. math::
- \ind{~}{\left[\begin{array}{rcl}\tree&:&\Set\\\forest&:&\Set\end{array}\right]}
- {\left[\begin{array}{rcl}
- \node &:& \forest → \tree\\
- \emptyf &:& \forest\\
- \consf &:& \tree → \forest → \forest\\
- \end{array}\right]}
+ which corresponds to the result of the |Coq| declaration:
-which corresponds to the result of the |Coq| declaration:
+ .. coqtop:: in
+
+ Inductive list (A:Set) : Set :=
+ | nil : list A
+ | cons : A -> list A -> list A.
.. example::
- .. coqtop:: in
+ The declaration for a mutual inductive definition of tree and forest
+ is:
- Inductive tree : Set :=
- | node : forest -> tree
- with forest : Set :=
- | emptyf : forest
- | consf : tree -> forest -> forest.
+ .. math::
+ \ind{~}{\left[\begin{array}{rcl}\tree&:&\Set\\\forest&:&\Set\end{array}\right]}
+ {\left[\begin{array}{rcl}
+ \node &:& \forest → \tree\\
+ \emptyf &:& \forest\\
+ \consf &:& \tree → \forest → \forest\\
+ \end{array}\right]}
-The declaration for a mutual inductive definition of even and odd is:
+ which corresponds to the result of the |Coq| declaration:
-.. math::
- \ind{1}{\left[\begin{array}{rcl}\even&:&\nat → \Prop \\
- \odd&:&\nat → \Prop \end{array}\right]}
- {\left[\begin{array}{rcl}
- \evenO &:& \even~0\\
- \evenS &:& \forall n, \odd~n -> \even~(\kw{S}~n)\\
- \oddS &:& \forall n, \even~n -> \odd~(\kw{S}~n)
- \end{array}\right]}
+ .. coqtop:: in
-which corresponds to the result of the |Coq| declaration:
+ Inductive tree : Set :=
+ | node : forest -> tree
+ with forest : Set :=
+ | emptyf : forest
+ | consf : tree -> forest -> forest.
.. example::
- .. coqtop:: in
+ The declaration for a mutual inductive definition of even and odd is:
+
+ .. math::
+ \ind{1}{\left[\begin{array}{rcl}\even&:&\nat → \Prop \\
+ \odd&:&\nat → \Prop \end{array}\right]}
+ {\left[\begin{array}{rcl}
+ \evenO &:& \even~0\\
+ \evenS &:& \forall n, \odd~n -> \even~(\kw{S}~n)\\
+ \oddS &:& \forall n, \even~n -> \odd~(\kw{S}~n)
+ \end{array}\right]}
+
+ which corresponds to the result of the |Coq| declaration:
- Inductive even : nat -> prop :=
- | even_O : even 0
- | even_S : forall n, odd n -> even (S n)
- with odd : nat -> prop :=
- | odd_S : forall n, even n -> odd (S n).
+ .. coqtop:: in
+
+ Inductive even : nat -> prop :=
+ | even_O : even 0
+ | even_S : forall n, odd n -> even (S n)
+ with odd : nat -> prop :=
+ | odd_S : forall n, even n -> odd (S n).
@@ -838,8 +839,8 @@ rules, we need a few definitions:
Arity of a given sort
+++++++++++++++++++++
-A type :math:`T` is an *arity of sort s* if it converts to the sort s or to a
-product :math:`∀ x:T,U` with :math:`U` an arity of sort s.
+A type :math:`T` is an *arity of sort* :math:`s` if it converts to the sort :math:`s` or to a
+product :math:`∀ x:T,U` with :math:`U` an arity of sort :math:`s`.
.. example::
@@ -850,7 +851,7 @@ product :math:`∀ x:T,U` with :math:`U` an arity of sort s.
Arity
+++++
A type :math:`T` is an *arity* if there is a :math:`s∈ \Sort` such that :math:`T` is an arity of
-sort s.
+sort :math:`s`.
.. example::
@@ -921,29 +922,29 @@ condition* for a constant :math:`X` in the following cases:
For instance, if one considers the following variant of a tree type
branching over the natural numbers:
- .. coqtop:: in
+ .. coqtop:: in
- Inductive nattree (A:Type) : Type :=
- | leaf : nattree A
- | node : A -> (nat -> nattree A) -> nattree A.
- End TreeExample.
-
- Then every instantiated constructor of ``nattree A`` satisfies the nested positivity
- condition for ``nattree``:
+ Inductive nattree (A:Type) : Type :=
+ | leaf : nattree A
+ | node : A -> (nat -> nattree A) -> nattree A.
+ End TreeExample.
+
+ Then every instantiated constructor of ``nattree A`` satisfies the nested positivity
+ condition for ``nattree``:
- + Type ``nattree A`` of constructor ``leaf`` satisfies the positivity condition for
- ``nattree`` because ``nattree`` does not appear in any (real) arguments of the
- type of that constructor (primarily because ``nattree`` does not have any (real)
- arguments) ... (bullet 1)
+ + Type ``nattree A`` of constructor ``leaf`` satisfies the positivity condition for
+ ``nattree`` because ``nattree`` does not appear in any (real) arguments of the
+ type of that constructor (primarily because ``nattree`` does not have any (real)
+ arguments) ... (bullet 1)
- + Type ``A → (nat → nattree A) → nattree A`` of constructor ``node`` satisfies the
- positivity condition for ``nattree`` because:
+ + Type ``A → (nat → nattree A) → nattree A`` of constructor ``node`` satisfies the
+ positivity condition for ``nattree`` because:
- - ``nattree`` occurs only strictly positively in ``A`` ... (bullet 3)
+ - ``nattree`` occurs only strictly positively in ``A`` ... (bullet 3)
- - ``nattree`` occurs only strictly positively in ``nat → nattree A`` ... (bullet 3 + 2)
+ - ``nattree`` occurs only strictly positively in ``nat → nattree A`` ... (bullet 3 + 2)
- - ``nattree`` satisfies the positivity condition for ``nattree A`` ... (bullet 1)
+ - ``nattree`` satisfies the positivity condition for ``nattree A`` ... (bullet 1)
.. _Correctness-rules:
@@ -978,35 +979,33 @@ provided that the following side conditions hold:
One can remark that there is a constraint between the sort of the
arity of the inductive type and the sort of the type of its
constructors which will always be satisfied for the impredicative
-sortProp but may fail to define inductive definition on sort Set and
+sort :math:`\Prop` but may fail to define inductive definition on sort :math:`\Set` and
generate constraints between universes for inductive definitions in
the Type hierarchy.
-**Examples**. It is well known that the existential quantifier can be encoded as an
-inductive definition. The following declaration introduces the second-
-order existential quantifier :math:`∃ X.P(X)`.
-
.. example::
+ It is well known that the existential quantifier can be encoded as an
+ inductive definition. The following declaration introduces the second-
+ order existential quantifier :math:`∃ X.P(X)`.
+
.. coqtop:: in
-
+
Inductive exProp (P:Prop->Prop) : Prop :=
| exP_intro : forall X:Prop, P X -> exProp P.
-The same definition on Set is not allowed and fails:
+ The same definition on :math:`\Set` is not allowed and fails:
-.. example::
.. coqtop:: all
Fail Inductive exSet (P:Set->Prop) : Set :=
exS_intro : forall X:Set, P X -> exSet P.
-It is possible to declare the same inductive definition in the
-universe Type. The exType inductive definition has type
-:math:`(\Type(i)→\Prop)→\Type(j)` with the constraint that the parameter :math:`X` of :math:`\kw{exT_intro}`
-has type :math:`\Type(k)` with :math:`k<j` and :math:`k≤ i`.
+ It is possible to declare the same inductive definition in the
+ universe :math:`\Type`. The :g:`exType` inductive definition has type
+ :math:`(\Type(i)→\Prop)→\Type(j)` with the constraint that the parameter :math:`X` of :math:`\kw{exT_intro}`
+ has type :math:`\Type(k)` with :math:`k<j` and :math:`k≤ i`.
-.. example::
.. coqtop:: all
Inductive exType (P:Type->Prop) : Type :=
@@ -1019,9 +1018,9 @@ has type :math:`\Type(k)` with :math:`k<j` and :math:`k≤ i`.
Template polymorphism
+++++++++++++++++++++
-Inductive types declared in Type are polymorphic over their arguments
-in Type. If :math:`A` is an arity of some sort and s is a sort, we write :math:`A_{/s}`
-for the arity obtained from :math:`A` by replacing its sort with s.
+Inductive types declared in :math:`\Type` are polymorphic over their arguments
+in :math:`\Type`. If :math:`A` is an arity of some sort and math:`s` is a sort, we write :math:`A_{/s}`
+for the arity obtained from :math:`A` by replacing its sort with :math:`s`.
Especially, if :math:`A` is well-typed in some global environment and local
context, then :math:`A_{/s}` is typable by typability of all products in the
Calculus of Inductive Constructions. The following typing rule is
@@ -1382,7 +1381,7 @@ this type.
[I:Prop|I→ s]
A *singleton definition* has only one constructor and all the
-arguments of this constructor have type Prop. In that case, there is a
+arguments of this constructor have type :math:`\Prop`. In that case, there is a
canonical way to interpret the informative extraction on an object in
that type, such that the elimination on any sort :math:`s` is legal. Typical
examples are the conjunction of non-informative propositions and the
@@ -1421,45 +1420,45 @@ corresponding to the :math:`c:C` constructor.
We write :math:`\{c\}^P` for :math:`\{c:C\}^P` with :math:`C` the type of :math:`c`.
-**Example.**
-The following term in concrete syntax::
+.. example::
+ The following term in concrete syntax::
- match t as l return P' with
- | nil _ => t1
- | cons _ hd tl => t2
- end
+ match t as l return P' with
+ | nil _ => t1
+ | cons _ hd tl => t2
+ end
-can be represented in abstract syntax as
+ can be represented in abstract syntax as
-.. math::
- \case(t,P,f 1 | f 2 )
+ .. math::
+ \case(t,P,f 1 | f 2 )
-where
+ where
-.. math::
- \begin{eqnarray*}
- P & = & \lambda~l~.~P^\prime\\
- f_1 & = & t_1\\
- f_2 & = & \lambda~(hd:\nat)~.~\lambda~(tl:\List~\nat)~.~t_2
- \end{eqnarray*}
+ .. math::
+ \begin{eqnarray*}
+ P & = & \lambda~l~.~P^\prime\\
+ f_1 & = & t_1\\
+ f_2 & = & \lambda~(hd:\nat)~.~\lambda~(tl:\List~\nat)~.~t_2
+ \end{eqnarray*}
-According to the definition:
+ According to the definition:
-.. math::
- \{(\kw{nil}~\nat)\}^P ≡ \{(\kw{nil}~\nat) : (\List~\nat)\}^P ≡ (P~(\kw{nil}~\nat))
+ .. math::
+ \{(\kw{nil}~\nat)\}^P ≡ \{(\kw{nil}~\nat) : (\List~\nat)\}^P ≡ (P~(\kw{nil}~\nat))
-.. math::
-
- \begin{array}{rl}
- \{(\kw{cons}~\nat)\}^P & ≡\{(\kw{cons}~\nat) : (\nat→\List~\nat→\List~\nat)\}^P \\
- & ≡∀ n:\nat, \{(\kw{cons}~\nat~n) : \List~\nat→\List~\nat)\}^P \\
- & ≡∀ n:\nat, ∀ l:\List~\nat, \{(\kw{cons}~\nat~n~l) : \List~\nat)\}^P \\
- & ≡∀ n:\nat, ∀ l:\List~\nat,(P~(\kw{cons}~\nat~n~l)).
- \end{array}
+ .. math::
+
+ \begin{array}{rl}
+ \{(\kw{cons}~\nat)\}^P & ≡\{(\kw{cons}~\nat) : (\nat→\List~\nat→\List~\nat)\}^P \\
+ & ≡∀ n:\nat, \{(\kw{cons}~\nat~n) : \List~\nat→\List~\nat)\}^P \\
+ & ≡∀ n:\nat, ∀ l:\List~\nat, \{(\kw{cons}~\nat~n~l) : \List~\nat)\}^P \\
+ & ≡∀ n:\nat, ∀ l:\List~\nat,(P~(\kw{cons}~\nat~n~l)).
+ \end{array}
-Given some :math:`P` then :math:`\{(\kw{nil}~\nat)\}^P` represents the expected type of :math:`f_1` ,
-and :math:`\{(\kw{cons}~\nat)\}^P` represents the expected type of :math:`f_2`.
+ Given some :math:`P` then :math:`\{(\kw{nil}~\nat)\}^P` represents the expected type of :math:`f_1` ,
+ and :math:`\{(\kw{cons}~\nat)\}^P` represents the expected type of :math:`f_2`.
.. _Typing-rule:
@@ -1819,7 +1818,7 @@ while it will type-check, if one uses instead the `coqtop`
``-impredicative-set`` option..
The major change in the theory concerns the rule for product formation
-in the sort Set, which is extended to a domain in any sort:
+in the sort :math:`\Set`, which is extended to a domain in any sort:
.. inference:: ProdImp
@@ -1832,11 +1831,11 @@ in the sort Set, which is extended to a domain in any sort:
This extension has consequences on the inductive definitions which are
allowed. In the impredicative system, one can build so-called *large
inductive definitions* like the example of second-order existential
-quantifier (exSet).
+quantifier (:g:`exSet`).
There should be restrictions on the eliminations which can be
performed on such definitions. The eliminations rules in the
-impredicative system for sort Set become:
+impredicative system for sort :math:`\Set` become:
diff --git a/engine/evarutil.ml b/engine/evarutil.ml
index 0c044f20d..b77bf55d8 100644
--- a/engine/evarutil.ml
+++ b/engine/evarutil.ml
@@ -495,12 +495,12 @@ let e_new_type_evar env evdref ?src ?filter ?naming ?principal ?hypnaming rigid
evdref := evd;
c
-let new_Type ?(rigid=Evd.univ_flexible) env evd =
+let new_Type ?(rigid=Evd.univ_flexible) evd =
let open EConstr in
let (evd, s) = new_sort_variable rigid evd in
(evd, mkSort s)
-let e_new_Type ?(rigid=Evd.univ_flexible) env evdref =
+let e_new_Type ?(rigid=Evd.univ_flexible) evdref =
let evd', s = new_sort_variable rigid !evdref in
evdref := evd'; EConstr.mkSort s
diff --git a/engine/evarutil.mli b/engine/evarutil.mli
index 135aa73fc..0ad323ac4 100644
--- a/engine/evarutil.mli
+++ b/engine/evarutil.mli
@@ -63,7 +63,7 @@ val new_type_evar :
env -> evar_map -> rigid ->
evar_map * (constr * Sorts.t)
-val new_Type : ?rigid:rigid -> env -> evar_map -> evar_map * constr
+val new_Type : ?rigid:rigid -> evar_map -> evar_map * constr
(** Polymorphic constants *)
@@ -287,7 +287,7 @@ val e_new_type_evar : env -> evar_map ref ->
?principal:bool -> ?hypnaming:naming_mode -> rigid -> constr * Sorts.t
[@@ocaml.deprecated "Use [Evarutil.new_type_evar]"]
-val e_new_Type : ?rigid:rigid -> env -> evar_map ref -> constr
+val e_new_Type : ?rigid:rigid -> evar_map ref -> constr
[@@ocaml.deprecated "Use [Evarutil.new_Type]"]
val e_new_global : evar_map ref -> GlobRef.t -> constr
diff --git a/engine/evd.ml b/engine/evd.ml
index 761ae7983..d1c7fef73 100644
--- a/engine/evd.ml
+++ b/engine/evd.ml
@@ -805,8 +805,8 @@ let make_flexible_variable evd ~algebraic u =
(* Operations on constants *)
(****************************************)
-let fresh_sort_in_family ?loc ?(rigid=univ_flexible) env evd s =
- with_context_set ?loc rigid evd (UnivGen.fresh_sort_in_family env s)
+let fresh_sort_in_family ?loc ?(rigid=univ_flexible) evd s =
+ with_context_set ?loc rigid evd (UnivGen.fresh_sort_in_family s)
let fresh_constant_instance ?loc env evd c =
with_context_set ?loc univ_flexible evd (UnivGen.fresh_constant_instance env c)
@@ -820,8 +820,6 @@ let fresh_constructor_instance ?loc env evd c =
let fresh_global ?loc ?(rigid=univ_flexible) ?names env evd gr =
with_context_set ?loc rigid evd (UnivGen.fresh_global_instance ?names env gr)
-let whd_sort_variable evd t = t
-
let is_sort_variable evd s = UState.is_sort_variable evd.universes s
let is_flexible_level evd l =
diff --git a/engine/evd.mli b/engine/evd.mli
index d638bb2d3..db2bd4eed 100644
--- a/engine/evd.mli
+++ b/engine/evd.mli
@@ -340,8 +340,6 @@ val shelve_on_future_goals : Evar.t list -> future_goals -> future_goals
Evar maps also keep track of the universe constraints defined at a given
point. This section defines the relevant manipulation functions. *)
-val whd_sort_variable : evar_map -> econstr -> econstr
-
exception UniversesDiffer
val add_universe_constraints : evar_map -> UnivProblem.Set.t -> evar_map
@@ -598,7 +596,7 @@ val update_sigma_env : evar_map -> env -> evar_map
(** Polymorphic universes *)
-val fresh_sort_in_family : ?loc:Loc.t -> ?rigid:rigid -> env -> evar_map -> Sorts.family -> evar_map * Sorts.t
+val fresh_sort_in_family : ?loc:Loc.t -> ?rigid:rigid -> evar_map -> Sorts.family -> evar_map * Sorts.t
val fresh_constant_instance : ?loc:Loc.t -> env -> evar_map -> Constant.t -> evar_map * pconstant
val fresh_inductive_instance : ?loc:Loc.t -> env -> evar_map -> inductive -> evar_map * pinductive
val fresh_constructor_instance : ?loc:Loc.t -> env -> evar_map -> constructor -> evar_map * pconstructor
diff --git a/engine/uState.ml b/engine/uState.ml
index 81ab3dd66..0791e4c27 100644
--- a/engine/uState.ml
+++ b/engine/uState.ml
@@ -583,7 +583,7 @@ let refresh_constraints univs (ctx, cstrs) =
in ((ctx, cstrs'), univs')
let normalize_variables uctx =
- let normalized_variables, undef, def, subst =
+ let normalized_variables, def, subst =
UnivSubst.normalize_univ_variables uctx.uctx_univ_variables
in
let ctx_local = subst_univs_context_with_def def (Univ.make_subst subst) uctx.uctx_local in
diff --git a/engine/univGen.ml b/engine/univGen.ml
index 796a1bcc1..b07d4848f 100644
--- a/engine/univGen.ml
+++ b/engine/univGen.ml
@@ -215,7 +215,7 @@ let type_of_reference env r =
let type_of_global t = type_of_reference (Global.env ()) t
-let fresh_sort_in_family env = function
+let fresh_sort_in_family = function
| InProp -> Sorts.prop, ContextSet.empty
| InSet -> Sorts.set, ContextSet.empty
| InType ->
@@ -223,7 +223,7 @@ let fresh_sort_in_family env = function
Type (Univ.Universe.make u), ContextSet.singleton u
let new_sort_in_family sf =
- fst (fresh_sort_in_family (Global.env ()) sf)
+ fst (fresh_sort_in_family sf)
let extend_context (a, ctx) (ctx') =
(a, ContextSet.union ctx ctx')
diff --git a/engine/univGen.mli b/engine/univGen.mli
index 8169dbda4..439424934 100644
--- a/engine/univGen.mli
+++ b/engine/univGen.mli
@@ -39,7 +39,7 @@ val fresh_instance_from_context : AUContext.t ->
val fresh_instance_from : AUContext.t -> Instance.t option ->
Instance.t in_universe_context_set
-val fresh_sort_in_family : env -> Sorts.family ->
+val fresh_sort_in_family : Sorts.family ->
Sorts.t in_universe_context_set
val fresh_constant_instance : env -> Constant.t ->
pconstant in_universe_context_set
diff --git a/engine/univSubst.ml b/engine/univSubst.ml
index 6a433d9fb..2f59a3fa8 100644
--- a/engine/univSubst.ml
+++ b/engine/univSubst.ml
@@ -162,13 +162,13 @@ let subst_opt_univs_constr s =
let normalize_univ_variables ctx =
let ctx = normalize_opt_subst ctx in
- let undef, def, subst =
- Univ.LMap.fold (fun u v (undef, def, subst) ->
+ let def, subst =
+ Univ.LMap.fold (fun u v (def, subst) ->
match v with
- | None -> (Univ.LSet.add u undef, def, subst)
- | Some b -> (undef, Univ.LSet.add u def, Univ.LMap.add u b subst))
- ctx (Univ.LSet.empty, Univ.LSet.empty, Univ.LMap.empty)
- in ctx, undef, def, subst
+ | None -> (def, subst)
+ | Some b -> (Univ.LSet.add u def, Univ.LMap.add u b subst))
+ ctx (Univ.LSet.empty, Univ.LMap.empty)
+ in ctx, def, subst
let pr_universe_body = function
| None -> mt ()
diff --git a/engine/univSubst.mli b/engine/univSubst.mli
index 26e8d1db9..e76d25333 100644
--- a/engine/univSubst.mli
+++ b/engine/univSubst.mli
@@ -23,7 +23,7 @@ val make_opt_subst : universe_opt_subst -> universe_subst_fn
val subst_opt_univs_constr : universe_opt_subst -> constr -> constr
val normalize_univ_variables : universe_opt_subst ->
- universe_opt_subst * LSet.t * LSet.t * universe_subst
+ universe_opt_subst * LSet.t * universe_subst
val normalize_univ_variable :
find:(Level.t -> Universe.t) ->
diff --git a/engine/universes.mli b/engine/universes.mli
index 29673de1e..ad937471e 100644
--- a/engine/universes.mli
+++ b/engine/universes.mli
@@ -86,7 +86,7 @@ val fresh_instance_from : AUContext.t -> Instance.t option ->
Instance.t in_universe_context_set
[@@ocaml.deprecated "Use [UnivGen.fresh_instance_from]"]
-val fresh_sort_in_family : env -> Sorts.family ->
+val fresh_sort_in_family : Sorts.family ->
Sorts.t in_universe_context_set
[@@ocaml.deprecated "Use [UnivGen.fresh_sort_in_family]"]
@@ -154,7 +154,7 @@ val subst_opt_univs_constr : universe_opt_subst -> constr -> constr
[@@ocaml.deprecated "Use [UnivSubst.subst_opt_univs_constr]"]
val normalize_univ_variables : universe_opt_subst ->
- universe_opt_subst * LSet.t * LSet.t * universe_subst
+ universe_opt_subst * LSet.t * universe_subst
[@@ocaml.deprecated "Use [UnivSubst.normalize_univ_variables]"]
val normalize_univ_variable :
diff --git a/grammar/coqpp_ast.mli b/grammar/coqpp_ast.mli
index 245e530ae..39b4d2ab3 100644
--- a/grammar/coqpp_ast.mli
+++ b/grammar/coqpp_ast.mli
@@ -13,14 +13,22 @@ type loc = {
type code = { code : string }
+type user_symbol =
+| Ulist1 of user_symbol
+| Ulist1sep of user_symbol * string
+| Ulist0 of user_symbol
+| Ulist0sep of user_symbol * string
+| Uopt of user_symbol
+| Uentry of string
+| Uentryl of string * int
+
type token_data =
| TokNone
| TokName of string
-| TokNameSep of string * string
type ext_token =
| ExtTerminal of string
-| ExtNonTerminal of string * token_data
+| ExtNonTerminal of user_symbol * token_data
type tactic_rule = {
tac_toks : ext_token list;
@@ -70,11 +78,18 @@ type grammar_ext = {
gramext_entries : grammar_entry list;
}
+type tactic_ext = {
+ tacext_name : string;
+ tacext_level : int option;
+ tacext_rules : tactic_rule list;
+}
+
type node =
| Code of code
| Comment of string
+| DeclarePlugin of string
| GramExt of grammar_ext
| VernacExt
-| TacticExt of string * tactic_rule list
+| TacticExt of tactic_ext
type t = node list
diff --git a/grammar/coqpp_lex.mll b/grammar/coqpp_lex.mll
index f35766b16..6c6562c20 100644
--- a/grammar/coqpp_lex.mll
+++ b/grammar/coqpp_lex.mll
@@ -83,6 +83,7 @@ let alphanum = ['_' 'a'-'z' 'A'-'Z' '0'-'9' '\'']
let ident = letterlike alphanum*
let qualid = ident ('.' ident)*
let space = [' ' '\t' '\r']
+let number = [ '0'-'9' ]
rule extend = parse
| "(*" { start_comment (); comment lexbuf }
@@ -92,6 +93,8 @@ rule extend = parse
| "TACTIC" { TACTIC }
| "EXTEND" { EXTEND }
| "END" { END }
+| "DECLARE" { DECLARE }
+| "PLUGIN" { PLUGIN }
(** Camlp5 specific keywords *)
| "GLOBAL" { GLOBAL }
| "FIRST" { FIRST }
@@ -105,6 +108,7 @@ rule extend = parse
(** Standard *)
| ident { IDENT (Lexing.lexeme lexbuf) }
| qualid { QUALID (Lexing.lexeme lexbuf) }
+| number { INT (int_of_string (Lexing.lexeme lexbuf)) }
| space { extend lexbuf }
| '\"' { string lexbuf }
| '\n' { newline lexbuf; extend lexbuf }
diff --git a/grammar/coqpp_main.ml b/grammar/coqpp_main.ml
index 23a5104e2..ef591bc99 100644
--- a/grammar/coqpp_main.ml
+++ b/grammar/coqpp_main.ml
@@ -55,6 +55,8 @@ let string_split s =
in
if len == 0 then [] else split 0
+let plugin_name = "__coq_plugin_name"
+
module GramExt :
sig
@@ -82,42 +84,42 @@ let get_local_entries ext =
in
uniquize StringSet.empty local
-let print_local chan ext =
+let print_local fmt ext =
let locals = get_local_entries ext in
match locals with
| [] -> ()
| e :: locals ->
- let mk_e chan e = fprintf chan "%s.Entry.create \"%s\"" ext.gramext_name e in
- let () = fprintf chan "@[<hv 2>let %s =@ @[%a@]@]@ " e mk_e e in
- let iter e = fprintf chan "@[<hv 2>and %s =@ @[%a@]@]@ " e mk_e e in
+ let mk_e fmt e = fprintf fmt "%s.Entry.create \"%s\"" ext.gramext_name e in
+ let () = fprintf fmt "@[<hv 2>let %s =@ @[%a@]@]@ " e mk_e e in
+ let iter e = fprintf fmt "@[<hv 2>and %s =@ @[%a@]@]@ " e mk_e e in
let () = List.iter iter locals in
- fprintf chan "in@ "
+ fprintf fmt "in@ "
-let print_string chan s = fprintf chan "\"%s\"" s
+let print_string fmt s = fprintf fmt "\"%s\"" s
-let print_list chan pr l =
- let rec prl chan = function
+let print_list fmt pr l =
+ let rec prl fmt = function
| [] -> ()
- | [x] -> fprintf chan "%a" pr x
- | x :: l -> fprintf chan "%a;@ %a" pr x prl l
+ | [x] -> fprintf fmt "%a" pr x
+ | x :: l -> fprintf fmt "%a;@ %a" pr x prl l
in
- fprintf chan "@[<hv>[%a]@]" prl l
+ fprintf fmt "@[<hv>[%a]@]" prl l
-let print_opt chan pr = function
-| None -> fprintf chan "None"
-| Some x -> fprintf chan "Some@ (%a)" pr x
+let print_opt fmt pr = function
+| None -> fprintf fmt "None"
+| Some x -> fprintf fmt "Some@ (%a)" pr x
-let print_position chan pos = match pos with
-| First -> fprintf chan "Extend.First"
-| Last -> fprintf chan "Extend.Last"
-| Before s -> fprintf chan "Extend.Before@ \"%s\"" s
-| After s -> fprintf chan "Extend.After@ \"%s\"" s
-| Level s -> fprintf chan "Extend.Level@ \"%s\"" s
+let print_position fmt pos = match pos with
+| First -> fprintf fmt "Extend.First"
+| Last -> fprintf fmt "Extend.Last"
+| Before s -> fprintf fmt "Extend.Before@ \"%s\"" s
+| After s -> fprintf fmt "Extend.After@ \"%s\"" s
+| Level s -> fprintf fmt "Extend.Level@ \"%s\"" s
-let print_assoc chan = function
-| LeftA -> fprintf chan "Extend.LeftA"
-| RightA -> fprintf chan "Extend.RightA"
-| NonA -> fprintf chan "Extend.NonA"
+let print_assoc fmt = function
+| LeftA -> fprintf fmt "Extend.LeftA"
+| RightA -> fprintf fmt "Extend.RightA"
+| NonA -> fprintf fmt "Extend.NonA"
type symb =
| SymbToken of string * string option
@@ -171,111 +173,166 @@ let rec parse_tokens = function
and parse_token tkn = parse_tokens [tkn]
-let print_fun chan (vars, body) =
+let print_fun fmt (vars, body) =
let vars = List.rev vars in
let iter = function
- | None -> fprintf chan "_@ "
- | Some id -> fprintf chan "%s@ " id
+ | None -> fprintf fmt "_@ "
+ | Some id -> fprintf fmt "%s@ " id
in
- let () = fprintf chan "fun@ " in
+ let () = fprintf fmt "fun@ " in
let () = List.iter iter vars in
(** FIXME: use Coq locations in the macros *)
- let () = fprintf chan "loc ->@ @[%s@]" body.code in
+ let () = fprintf fmt "loc ->@ @[%s@]" body.code in
()
(** Meta-program instead of calling Tok.of_pattern here because otherwise
violates value restriction *)
-let print_tok chan = function
-| "", s -> fprintf chan "Tok.KEYWORD %a" print_string s
-| "IDENT", s -> fprintf chan "Tok.IDENT %a" print_string s
-| "PATTERNIDENT", s -> fprintf chan "Tok.PATTERNIDENT %a" print_string s
-| "FIELD", s -> fprintf chan "Tok.FIELD %a" print_string s
-| "INT", s -> fprintf chan "Tok.INT %a" print_string s
-| "STRING", s -> fprintf chan "Tok.STRING %a" print_string s
-| "LEFTQMARK", _ -> fprintf chan "Tok.LEFTQMARK"
-| "BULLET", s -> fprintf chan "Tok.BULLET %a" print_string s
-| "EOI", _ -> fprintf chan "Tok.EOI"
+let print_tok fmt = function
+| "", s -> fprintf fmt "Tok.KEYWORD %a" print_string s
+| "IDENT", s -> fprintf fmt "Tok.IDENT %a" print_string s
+| "PATTERNIDENT", s -> fprintf fmt "Tok.PATTERNIDENT %a" print_string s
+| "FIELD", s -> fprintf fmt "Tok.FIELD %a" print_string s
+| "INT", s -> fprintf fmt "Tok.INT %a" print_string s
+| "STRING", s -> fprintf fmt "Tok.STRING %a" print_string s
+| "LEFTQMARK", _ -> fprintf fmt "Tok.LEFTQMARK"
+| "BULLET", s -> fprintf fmt "Tok.BULLET %a" print_string s
+| "EOI", _ -> fprintf fmt "Tok.EOI"
| _ -> failwith "Tok.of_pattern: not a constructor"
-let rec print_prod chan p =
+let rec print_prod fmt p =
let (vars, tkns) = List.split p.gprod_symbs in
let f = (vars, p.gprod_body) in
let tkn = List.rev_map parse_tokens tkns in
- fprintf chan "@[Extend.Rule@ (@[%a@],@ @[(%a)@])@]" print_symbols tkn print_fun f
+ fprintf fmt "@[Extend.Rule@ (@[%a@],@ @[(%a)@])@]" print_symbols tkn print_fun f
-and print_symbols chan = function
-| [] -> fprintf chan "Extend.Stop"
+and print_symbols fmt = function
+| [] -> fprintf fmt "Extend.Stop"
| tkn :: tkns ->
- fprintf chan "Extend.Next @[(%a,@ %a)@]" print_symbols tkns print_symbol tkn
+ fprintf fmt "Extend.Next @[(%a,@ %a)@]" print_symbols tkns print_symbol tkn
-and print_symbol chan tkn = match tkn with
+and print_symbol fmt tkn = match tkn with
| SymbToken (t, s) ->
let s = match s with None -> "" | Some s -> s in
- fprintf chan "(Extend.Atoken (%a))" print_tok (t, s)
+ fprintf fmt "(Extend.Atoken (%a))" print_tok (t, s)
| SymbEntry (e, None) ->
- fprintf chan "(Extend.Aentry %s)" e
+ fprintf fmt "(Extend.Aentry %s)" e
| SymbEntry (e, Some l) ->
- fprintf chan "(Extend.Aentryl (%s, %a))" e print_string l
+ fprintf fmt "(Extend.Aentryl (%s, %a))" e print_string l
| SymbSelf ->
- fprintf chan "Extend.Aself"
+ fprintf fmt "Extend.Aself"
| SymbNext ->
- fprintf chan "Extend.Anext"
+ fprintf fmt "Extend.Anext"
| SymbList0 (s, None) ->
- fprintf chan "(Extend.Alist0 %a)" print_symbol s
+ fprintf fmt "(Extend.Alist0 %a)" print_symbol s
| SymbList0 (s, Some sep) ->
- fprintf chan "(Extend.Alist0sep (%a, %a))" print_symbol s print_symbol sep
+ fprintf fmt "(Extend.Alist0sep (%a, %a))" print_symbol s print_symbol sep
| SymbList1 (s, None) ->
- fprintf chan "(Extend.Alist1 %a)" print_symbol s
+ fprintf fmt "(Extend.Alist1 %a)" print_symbol s
| SymbList1 (s, Some sep) ->
- fprintf chan "(Extend.Alist1sep (%a, %a))" print_symbol s print_symbol sep
+ fprintf fmt "(Extend.Alist1sep (%a, %a))" print_symbol s print_symbol sep
| SymbOpt s ->
- fprintf chan "(Extend.Aopt %a)" print_symbol s
+ fprintf fmt "(Extend.Aopt %a)" print_symbol s
| SymbRules rules ->
- let pr chan (r, body) =
+ let pr fmt (r, body) =
let (vars, tkn) = List.split r in
let tkn = List.rev tkn in
- fprintf chan "Extend.Rules @[({ Extend.norec_rule = %a },@ (%a))@]" print_symbols tkn print_fun (vars, body)
+ fprintf fmt "Extend.Rules @[({ Extend.norec_rule = %a },@ (%a))@]" print_symbols tkn print_fun (vars, body)
in
- let pr chan rules = print_list chan pr rules in
- fprintf chan "(Extend.Arules %a)" pr (List.rev rules)
-
-let print_rule chan r =
- let pr_lvl chan lvl = print_opt chan print_string lvl in
- let pr_asc chan asc = print_opt chan print_assoc asc in
- let pr_prd chan prd = print_list chan print_prod prd in
- fprintf chan "@[(%a,@ %a,@ %a)@]" pr_lvl r.grule_label pr_asc r.grule_assoc pr_prd (List.rev r.grule_prods)
-
-let print_entry chan gram e =
- let print_position_opt chan pos = print_opt chan print_position pos in
- let print_rules chan rules = print_list chan print_rule rules in
- fprintf chan "let () =@ @[%s.safe_extend@ %s@ @[(%a, %a)@]@]@ in@ "
+ let pr fmt rules = print_list fmt pr rules in
+ fprintf fmt "(Extend.Arules %a)" pr (List.rev rules)
+
+let print_rule fmt r =
+ let pr_lvl fmt lvl = print_opt fmt print_string lvl in
+ let pr_asc fmt asc = print_opt fmt print_assoc asc in
+ let pr_prd fmt prd = print_list fmt print_prod prd in
+ fprintf fmt "@[(%a,@ %a,@ %a)@]" pr_lvl r.grule_label pr_asc r.grule_assoc pr_prd (List.rev r.grule_prods)
+
+let print_entry fmt gram e =
+ let print_position_opt fmt pos = print_opt fmt print_position pos in
+ let print_rules fmt rules = print_list fmt print_rule rules in
+ fprintf fmt "let () =@ @[%s.safe_extend@ %s@ @[(%a, %a)@]@]@ in@ "
gram e.gentry_name print_position_opt e.gentry_pos print_rules e.gentry_rules
-let print_ast chan ext =
- let () = fprintf chan "let _ = @[" in
- let () = fprintf chan "@[<v>%a@]" print_local ext in
- let () = List.iter (fun e -> print_entry chan ext.gramext_name e) ext.gramext_entries in
- let () = fprintf chan "()@]\n" in
+let print_ast fmt ext =
+ let () = fprintf fmt "let _ = @[" in
+ let () = fprintf fmt "@[<v>%a@]" print_local ext in
+ let () = List.iter (fun e -> print_entry fmt ext.gramext_name e) ext.gramext_entries in
+ let () = fprintf fmt "()@]@\n" in
()
end
-let pr_ast chan = function
-| Code s -> fprintf chan "%s@\n" s.code
-| Comment s -> fprintf chan "%s@\n" s
-| GramExt gram -> fprintf chan "%a@\n" GramExt.print_ast gram
-| VernacExt -> fprintf chan "VERNACEXT@\n"
-| TacticExt (id, rules) ->
- let pr_tok = function
- | ExtTerminal s -> Printf.sprintf "%s" s
- | ExtNonTerminal (s, _) -> Printf.sprintf "%s" s
- in
- let pr_tacrule r =
- let toks = String.concat " " (List.map pr_tok r.tac_toks) in
- Printf.sprintf "[%s] -> {%s}" toks r.tac_body.code
+module TacticExt :
+sig
+
+val print_ast : Format.formatter -> tactic_ext -> unit
+
+end =
+struct
+
+let rec print_symbol fmt = function
+| Ulist1 s ->
+ fprintf fmt "@[Extend.TUlist1 (%a)@]" print_symbol s
+| Ulist1sep (s, sep) ->
+ fprintf fmt "@[Extend.TUlist1sep (%a, \"%s\")@]" print_symbol s sep
+| Ulist0 s ->
+ fprintf fmt "@[Extend.TUlist0 (%a)@]" print_symbol s
+| Ulist0sep (s, sep) ->
+ fprintf fmt "@[Extend.TUlist0sep (%a, \"%s\")@]" print_symbol s sep
+| Uopt s ->
+ fprintf fmt "@[Extend.TUopt (%a)@]" print_symbol s
+| Uentry e ->
+ fprintf fmt "@[Extend.TUentry (Genarg.get_arg_tag wit_%s)@]" e
+| Uentryl (e, l) ->
+ assert (e = "tactic");
+ fprintf fmt "@[Extend.TUentryl (Genarg.get_arg_tag wit_%s, %i)@]" e l
+
+let rec print_clause fmt = function
+| [] -> fprintf fmt "@[TyNil@]"
+| ExtTerminal s :: cl -> fprintf fmt "@[TyIdent (\"%s\", %a)@]" s print_clause cl
+| ExtNonTerminal (g, TokNone) :: cl ->
+ fprintf fmt "@[TyAnonArg (%a, %a)@]"
+ print_symbol g print_clause cl
+| ExtNonTerminal (g, TokName id) :: cl ->
+ fprintf fmt "@[TyArg (%a, \"%s\", %a)@]"
+ print_symbol g id print_clause cl
+
+let rec print_binders fmt = function
+| [] -> fprintf fmt "ist@ "
+| (ExtTerminal _ | ExtNonTerminal (_, TokNone)) :: rem -> print_binders fmt rem
+| (ExtNonTerminal (_, TokName id)) :: rem ->
+ fprintf fmt "%s@ %a" id print_binders rem
+
+let print_rule fmt r =
+ fprintf fmt "@[TyML (%a, @[fun %a -> %s@])@]"
+ print_clause r.tac_toks print_binders r.tac_toks r.tac_body.code
+
+let rec print_rules fmt = function
+| [] -> ()
+| [r] -> fprintf fmt "(%a)@\n" print_rule r
+| r :: rem -> fprintf fmt "(%a);@\n%a" print_rule r print_rules rem
+
+let print_rules fmt rules =
+ fprintf fmt "Tacentries.([@[<v>%a@]])" print_rules rules
+
+let print_ast fmt ext =
+ let pr fmt () =
+ let level = match ext.tacext_level with None -> 0 | Some i -> i in
+ fprintf fmt "Tacentries.tactic_extend %s \"%s\" ~level:%i %a"
+ plugin_name ext.tacext_name level print_rules ext.tacext_rules
in
- let rules = String.concat ", " (List.map pr_tacrule rules) in
- fprintf chan "TACTICEXT (%s, [%s])@\n" id rules
+ let () = fprintf fmt "let () = @[%a@]\n" pr () in
+ ()
+
+end
+
+let pr_ast fmt = function
+| Code s -> fprintf fmt "%s@\n" s.code
+| Comment s -> fprintf fmt "%s@\n" s
+| DeclarePlugin name -> fprintf fmt "let %s = \"%s\"@\n" plugin_name name
+| GramExt gram -> fprintf fmt "%a@\n" GramExt.print_ast gram
+| VernacExt -> fprintf fmt "VERNACEXT@\n"
+| TacticExt tac -> fprintf fmt "%a@\n" TacticExt.print_ast tac
let () =
let () =
diff --git a/grammar/coqpp_parse.mly b/grammar/coqpp_parse.mly
index aa22d2717..baafd633c 100644
--- a/grammar/coqpp_parse.mly
+++ b/grammar/coqpp_parse.mly
@@ -10,13 +10,59 @@
open Coqpp_ast
+let starts s pat =
+ let len = String.length s in
+ let patlen = String.length pat in
+ if patlen <= len && String.sub s 0 patlen = pat then
+ Some (String.sub s patlen (len - patlen))
+ else None
+
+let ends s pat =
+ let len = String.length s in
+ let patlen = String.length pat in
+ if patlen <= len && String.sub s (len - patlen) patlen = pat then
+ Some (String.sub s 0 (len - patlen))
+ else None
+
+let between s pat1 pat2 = match starts s pat1 with
+| None -> None
+| Some s -> ends s pat2
+
+let without_sep k sep r =
+ if sep <> "" then raise Parsing.Parse_error else k r
+
+let parse_user_entry s sep =
+ let table = [
+ "ne_", "_list", without_sep (fun r -> Ulist1 r);
+ "ne_", "_list_sep", (fun sep r -> Ulist1sep (r, sep));
+ "", "_list", without_sep (fun r -> Ulist0 r);
+ "", "_list_sep", (fun sep r -> Ulist0sep (r, sep));
+ "", "_opt", without_sep (fun r -> Uopt r);
+ ] in
+ let rec parse s sep = function
+ | [] ->
+ let () = without_sep ignore sep () in
+ begin match starts s "tactic" with
+ | Some ("0"|"1"|"2"|"3"|"4"|"5") -> Uentryl ("tactic", int_of_string s)
+ | Some _ | None -> Uentry s
+ end
+ | (pat1, pat2, k) :: rem ->
+ match between s pat1 pat2 with
+ | None -> parse s sep rem
+ | Some s ->
+ let r = parse s "" table in
+ k sep r
+ in
+ parse s sep table
+
%}
%token <Coqpp_ast.code> CODE
%token <string> COMMENT
%token <string> IDENT QUALID
%token <string> STRING
-%token VERNAC TACTIC GRAMMAR EXTEND END
+%token <int> INT
+%token VERNAC TACTIC GRAMMAR EXTEND END DECLARE PLUGIN
%token LBRACKET RBRACKET PIPE ARROW COMMA EQUAL
%token LPAREN RPAREN COLON SEMICOLON
%token GLOBAL FIRST LAST BEFORE AFTER LEVEL LEFTA RIGHTA NONA
@@ -42,11 +88,16 @@ nodes:
node:
| CODE { Code $1 }
| COMMENT { Comment $1 }
+| declare_plugin { $1 }
| grammar_extend { $1 }
| vernac_extend { $1 }
| tactic_extend { $1 }
;
+declare_plugin:
+| DECLARE PLUGIN STRING { DeclarePlugin $3 }
+;
+
grammar_extend:
| GRAMMAR EXTEND qualid_or_ident globals gram_entries END
{ GramExt { gramext_name = $3; gramext_globals = $4; gramext_entries = $5 } }
@@ -57,7 +108,13 @@ vernac_extend:
;
tactic_extend:
-| TACTIC EXTEND IDENT tactic_rules END { TacticExt ($3, $4) }
+| TACTIC EXTEND IDENT tactic_level tactic_rules END
+ { TacticExt { tacext_name = $3; tacext_level = $4; tacext_rules = $5 } }
+;
+
+tactic_level:
+| { None }
+| LEVEL INT { Some $2 }
;
tactic_rules:
@@ -77,9 +134,18 @@ ext_tokens:
ext_token:
| STRING { ExtTerminal $1 }
-| IDENT { ExtNonTerminal ($1, TokNone) }
-| IDENT LPAREN IDENT RPAREN { ExtNonTerminal ($1, TokName $3) }
-| IDENT LPAREN IDENT COMMA STRING RPAREN { ExtNonTerminal ($1, TokNameSep ($3, $5)) }
+| IDENT {
+ let e = parse_user_entry $1 "" in
+ ExtNonTerminal (e, TokNone)
+ }
+| IDENT LPAREN IDENT RPAREN {
+ let e = parse_user_entry $1 "" in
+ ExtNonTerminal (e, TokName $3)
+ }
+| IDENT LPAREN IDENT COMMA STRING RPAREN {
+ let e = parse_user_entry $1 $5 in
+ ExtNonTerminal (e, TokName $3)
+}
;
qualid_or_ident:
diff --git a/grammar/tacextend.mlp b/grammar/tacextend.mlp
index 525be6432..cea0bed59 100644
--- a/grammar/tacextend.mlp
+++ b/grammar/tacextend.mlp
@@ -15,11 +15,6 @@ open Argextend
let plugin_name = <:expr< __coq_plugin_name >>
-let mlexpr_of_ident id =
- (** Workaround for badly-designed generic arguments lacking a closure *)
- let id = "$" ^ id in
- <:expr< Names.Id.of_string_soft $str:id$ >>
-
let rec mlexpr_of_symbol = function
| Ulist1 s -> <:expr< Extend.TUlist1 $mlexpr_of_symbol s$ >>
| Ulist1sep (s,sep) -> <:expr< Extend.TUlist1sep $mlexpr_of_symbol s$ $str:sep$ >>
@@ -38,9 +33,9 @@ let rec mlexpr_of_clause = function
| [] -> <:expr< TyNil >>
| ExtTerminal s :: cl -> <:expr< TyIdent($str:s$, $mlexpr_of_clause cl$) >>
| ExtNonTerminal(g,None) :: cl ->
- <:expr< TyAnonArg(Loc.tag($mlexpr_of_symbol g$), $mlexpr_of_clause cl$) >>
+ <:expr< TyAnonArg($mlexpr_of_symbol g$, $mlexpr_of_clause cl$) >>
| ExtNonTerminal(g,Some id) :: cl ->
- <:expr< TyArg(Loc.tag($mlexpr_of_symbol g$, $mlexpr_of_ident id$), $mlexpr_of_clause cl$) >>
+ <:expr< TyArg($mlexpr_of_symbol g$, $mlexpr_of_string id$, $mlexpr_of_clause cl$) >>
let rec binders_of_clause e = function
| [] -> <:expr< fun ist -> $e$ >>
diff --git a/ide/idetop.ml b/ide/idetop.ml
index 7abbf239b..6fb004061 100644
--- a/ide/idetop.ml
+++ b/ide/idetop.ml
@@ -151,7 +151,7 @@ let hyp_next_tac sigma env decl =
("inversion clear "^id_s), ("inversion_clear "^id_s^".")
]
-let concl_next_tac sigma concl =
+let concl_next_tac =
let expand s = (s,s^".") in
List.map expand ([
"intro";
@@ -230,10 +230,9 @@ let hints () =
| [] -> None
| g :: _ ->
let env = Goal.V82.env sigma g in
- let hint_goal = concl_next_tac sigma g in
let get_hint_hyp env d accu = hyp_next_tac sigma env d :: accu in
let hint_hyps = List.rev (Environ.fold_named_context get_hint_hyp env ~init: []) in
- Some (hint_hyps, hint_goal)
+ Some (hint_hyps, concl_next_tac)
with Proof_global.NoCurrentProof -> None
diff --git a/kernel/cooking.ml b/kernel/cooking.ml
index 521a7d890..094609b96 100644
--- a/kernel/cooking.ml
+++ b/kernel/cooking.ml
@@ -204,7 +204,7 @@ let lift_univs cb subst auctx0 =
let auctx' = Univ.subst_univs_level_abstract_universe_context (Univ.make_instance_subst subst) auctx in
subst, (Polymorphic_const (AUContext.union auctx0 auctx'))
-let cook_constant ~hcons env { from = cb; info } =
+let cook_constant ~hcons { from = cb; info } =
let { Opaqueproof.modlist; abstract } = info in
let cache = RefTable.create 13 in
let abstract, usubst, abs_ctx = abstract in
diff --git a/kernel/cooking.mli b/kernel/cooking.mli
index 4c254d2ea..6ebe691b8 100644
--- a/kernel/cooking.mli
+++ b/kernel/cooking.mli
@@ -10,7 +10,6 @@
open Constr
open Declarations
-open Environ
(** {6 Cooking the constants. } *)
@@ -26,7 +25,7 @@ type result = {
cook_context : Constr.named_context option;
}
-val cook_constant : hcons:bool -> env -> recipe -> result
+val cook_constant : hcons:bool -> recipe -> result
val cook_constr : Opaqueproof.cooking_info -> constr -> constr
(** {6 Utility functions used in module [Discharge]. } *)
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index 61b71df31..5d45c2c1a 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -120,16 +120,6 @@ let mind_check_names mie =
(* Typing the arities and constructor types *)
-(* An inductive definition is a "unit" if it has only one constructor
- and that all arguments expected by this constructor are
- logical, this is the case for equality, conjunction of logical properties
-*)
-let is_unit constrsinfos =
- match constrsinfos with (* One info = One constructor *)
- | [level] -> is_type0m_univ level
- | [] -> (* type without constructors *) true
- | _ -> false
-
let infos_and_sort env t =
let rec aux env t max =
let t = whd_all env t in
@@ -174,10 +164,9 @@ let infer_constructor_packet env_ar_par params lc =
let lc'' = Array.map (fun j -> Term.it_mkProd_or_LetIn j.utj_val params) jlc in
(* compute the max of the sorts of the products of the constructors types *)
let levels = List.map (infos_and_sort env_ar_par) lc in
- let isunit = is_unit levels in
let min = if Array.length jlc > 1 then Universe.type0 else Universe.type0m in
let level = List.fold_left (fun max l -> Universe.sup max l) min levels in
- (lc'', (isunit, level))
+ (lc'', level)
(* If indices matter *)
let cumulate_arity_large_levels env sign =
@@ -354,7 +343,7 @@ let typecheck_inductive env mie =
(* Compute/check the sorts of the inductive types *)
let inds =
- Array.map (fun ((id,full_arity,sign,expltype,def_level,inf_level),cn,lc,(is_unit,clev)) ->
+ Array.map (fun ((id,full_arity,sign,expltype,def_level,inf_level),cn,lc,clev) ->
let infu =
(** Inferred level, with parameters and constructors. *)
match inf_level with
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index 584c1af03..88b00600e 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -785,7 +785,7 @@ let rec subterm_specif renv stack t =
| Lambda (x,a,b) ->
let () = assert (List.is_empty l) in
- let spec,stack' = extract_stack renv a stack in
+ let spec,stack' = extract_stack stack in
subterm_specif (push_var renv (x,a,spec)) stack' b
(* Metas and evars are considered OK *)
@@ -817,7 +817,7 @@ and stack_element_specif = function
|SClosure (h_renv,h) -> lazy_subterm_specif h_renv [] h
|SArg x -> x
-and extract_stack renv a = function
+and extract_stack = function
| [] -> Lazy.from_val Not_subterm , []
| h::t -> stack_element_specif h, t
@@ -848,7 +848,7 @@ let error_illegal_rec_call renv fx (arg_renv,arg) =
let error_partial_apply renv fx =
raise (FixGuardError (renv.env,NotEnoughArgumentsForFixCall fx))
-let filter_stack_domain env ci p stack =
+let filter_stack_domain env p stack =
let absctx, ar = dest_lam_assum env p in
(* Optimization: if the predicate is not dependent, no restriction is needed
and we avoid building the recargs tree. *)
@@ -933,7 +933,7 @@ let check_one_fix renv recpos trees def =
let case_spec = branches_specif renv
(lazy_subterm_specif renv [] c_0) ci in
let stack' = push_stack_closures renv l stack in
- let stack' = filter_stack_domain renv.env ci p stack' in
+ let stack' = filter_stack_domain renv.env p stack' in
Array.iteri (fun k br' ->
let stack_br = push_stack_args case_spec.(k) stack' in
check_rec_call renv stack_br br') lrest
@@ -976,7 +976,7 @@ let check_one_fix renv recpos trees def =
| Lambda (x,a,b) ->
let () = assert (List.is_empty l) in
check_rec_call renv [] a ;
- let spec, stack' = extract_stack renv a stack in
+ let spec, stack' = extract_stack stack in
check_rec_call (push_var renv (x,a,spec)) stack' b
| Prod (x,a,b) ->
diff --git a/kernel/modops.ml b/kernel/modops.ml
index 02bab581a..98a997311 100644
--- a/kernel/modops.ml
+++ b/kernel/modops.ml
@@ -265,7 +265,7 @@ let subst_structure subst = subst_structure subst do_delta_codom
(* spiwack: here comes the function which takes care of importing
the retroknowledge declared in the library *)
(* lclrk : retroknowledge_action list, rkaction : retroknowledge action *)
-let add_retroknowledge mp =
+let add_retroknowledge =
let perform rkaction env = match rkaction with
| Retroknowledge.RKRegister (f, e) when (isConst e || isInd e) ->
Environ.register env f e
@@ -309,7 +309,7 @@ and add_module mb linkinfo env =
let env = Environ.shallow_add_module mb env in
match mb.mod_type with
|NoFunctor struc ->
- add_retroknowledge mp mb.mod_retroknowledge
+ add_retroknowledge mb.mod_retroknowledge
(add_structure mp struc mb.mod_delta linkinfo env)
|MoreFunctor _ -> env
diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml
index 1748e98a4..39f7de942 100644
--- a/kernel/nativecode.ml
+++ b/kernel/nativecode.ml
@@ -1649,15 +1649,15 @@ let pp_mllam fmt l =
and pp_letrec fmt defs =
let len = Array.length defs in
- let pp_one_rec i (fn, argsn, body) =
+ let pp_one_rec (fn, argsn, body) =
Format.fprintf fmt "%a%a =@\n %a"
pp_lname fn
pp_ldecls argsn pp_mllam body in
Format.fprintf fmt "@[let rec ";
- pp_one_rec 0 defs.(0);
+ pp_one_rec defs.(0);
for i = 1 to len - 1 do
Format.fprintf fmt "@\nand ";
- pp_one_rec i defs.(i)
+ pp_one_rec defs.(i)
done;
and pp_blam fmt l =
@@ -1941,7 +1941,7 @@ let is_code_loaded ~interactive name =
let param_name = Name (Id.of_string "params")
let arg_name = Name (Id.of_string "arg")
-let compile_mind prefix ~interactive mb mind stack =
+let compile_mind mb mind stack =
let u = Declareops.inductive_polymorphic_context mb in
(** Generate data for every block *)
let f i stack ob =
@@ -2020,7 +2020,7 @@ let compile_mind_deps env prefix ~interactive
then init
else
let comp_stack =
- compile_mind prefix ~interactive mib mind comp_stack
+ compile_mind mib mind comp_stack
in
let name =
if interactive then LinkedInteractive prefix
@@ -2092,9 +2092,9 @@ let compile_constant_field env prefix con acc cb =
in
gl@acc
-let compile_mind_field prefix mp l acc mb =
+let compile_mind_field mp l acc mb =
let mind = MutInd.make2 mp l in
- compile_mind prefix ~interactive:false mb mind acc
+ compile_mind mb mind acc
let mk_open s = Gopen s
diff --git a/kernel/nativecode.mli b/kernel/nativecode.mli
index 684983a87..96efa7faa 100644
--- a/kernel/nativecode.mli
+++ b/kernel/nativecode.mli
@@ -67,7 +67,7 @@ val register_native_file : string -> unit
val compile_constant_field : env -> string -> Constant.t ->
global list -> constant_body -> global list
-val compile_mind_field : string -> ModPath.t -> Label.t ->
+val compile_mind_field : ModPath.t -> Label.t ->
global list -> mutual_inductive_body -> global list
val mk_conv_code : env -> evars -> string -> constr -> constr -> linkable_code
diff --git a/kernel/nativelibrary.ml b/kernel/nativelibrary.ml
index 8bff43632..edce9367f 100644
--- a/kernel/nativelibrary.ml
+++ b/kernel/nativelibrary.ml
@@ -37,7 +37,7 @@ and translate_field prefix mp env acc (l,x) =
let id = mb.mind_packets.(0).mind_typename in
let msg = Printf.sprintf "Compiling inductive %s..." (Id.to_string id) in
Feedback.msg_debug (Pp.str msg));
- compile_mind_field prefix mp l acc mb
+ compile_mind_field mp l acc mb
| SFBmodule md ->
let mp = md.mod_mp in
(if !Flags.debug then
diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml
index 1e58f5c24..74042f9e0 100644
--- a/kernel/subtyping.ml
+++ b/kernel/subtyping.ml
@@ -224,7 +224,7 @@ let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2
cst
-let check_constant cst env mp1 l info1 cb2 spec2 subst1 subst2 =
+let check_constant cst env l info1 cb2 spec2 subst1 subst2 =
let error why = error_signature_mismatch l spec2 why in
let check_conv cst poly f = check_conv_error error cst poly f in
let check_type poly cst env t1 t2 =
@@ -292,7 +292,7 @@ and check_signatures cst env mp1 sig1 mp2 sig2 subst1 subst2 reso1 reso2=
let check_one_body cst (l,spec2) =
match spec2 with
| SFBconst cb2 ->
- check_constant cst env mp1 l (get_obj mp1 map1 l)
+ check_constant cst env l (get_obj mp1 map1 l)
cb2 spec2 subst1 subst2
| SFBmind mib2 ->
check_inductive cst env mp1 l (get_obj mp1 map1 l)
diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml
index bad449731..1f7ee145a 100644
--- a/kernel/term_typing.ml
+++ b/kernel/term_typing.ml
@@ -378,7 +378,7 @@ let build_constant_declaration kn env result =
str "Proof using " ++ declared_vars ++ fnl () ++
str "to" ++ fnl () ++
str "Proof using " ++ inferred_vars) in
- let sort evn l =
+ let sort l =
List.filter (fun decl ->
let id = NamedDecl.get_id decl in
List.exists (NamedDecl.get_id %> Names.Id.equal id) l)
@@ -411,7 +411,7 @@ let build_constant_declaration kn env result =
[], def (* Empty section context: no need to check *)
| Some declared ->
(* We use the declared set and chain a check of correctness *)
- sort env declared,
+ sort declared,
match def with
| Undef _ as x -> x (* nothing to check *)
| Def cs as x ->
@@ -554,7 +554,7 @@ let translate_recipe env kn r =
be useless. It is detected by the dirpath of the constant being empty. *)
let (_, dir, _) = Constant.repr3 kn in
let hcons = DirPath.is_empty dir in
- build_constant_declaration kn env (Cooking.cook_constant ~hcons env r)
+ build_constant_declaration kn env (Cooking.cook_constant ~hcons r)
let translate_local_def env id centry =
let open Cooking in
diff --git a/lib/cWarnings.ml b/lib/cWarnings.ml
index fda25a0a6..0cf989e49 100644
--- a/lib/cWarnings.ml
+++ b/lib/cWarnings.ml
@@ -120,7 +120,7 @@ let uniquize_flags_rev flags =
(** [normalize_flags] removes redundant warnings. Unknown warnings are kept
because they may be declared in a plugin that will be linked later. *)
-let normalize_flags ~silent warnings =
+let normalize_flags warnings =
let warnings = cut_before_all_rev warnings in
uniquize_flags_rev warnings
@@ -130,7 +130,7 @@ let normalize_flags_string s =
if is_none_keyword s then s
else
let flags = flags_of_string s in
- let flags = normalize_flags ~silent:false flags in
+ let flags = normalize_flags flags in
string_of_flags flags
let parse_warnings items =
@@ -146,7 +146,7 @@ let parse_flags s =
else begin
Flags.make_warn true;
let flags = flags_of_string s in
- let flags = normalize_flags ~silent:true flags in
+ let flags = normalize_flags flags in
parse_warnings flags;
string_of_flags flags
end
diff --git a/lib/util.ml b/lib/util.ml
index 7d7d380b2..38d73d345 100644
--- a/lib/util.ml
+++ b/lib/util.ml
@@ -38,8 +38,8 @@ let is_blank = function
module Empty =
struct
- type t
- let abort (x : t) = assert false
+ type t = { abort : 'a. 'a }
+ let abort (x : t) = x.abort
end
(* Strings *)
diff --git a/library/libobject.ml b/library/libobject.ml
index c5cd01525..79a3fed1b 100644
--- a/library/libobject.ml
+++ b/library/libobject.ml
@@ -98,7 +98,7 @@ let declare_object_full odecl = declare_object_full odecl
(* this function describes how the cache, load, open, and export functions
are triggered. *)
-let apply_dyn_fun deflt f lobj =
+let apply_dyn_fun f lobj =
let tag = object_tag lobj in
let dodecl =
try Hashtbl.find cache_tab tag
@@ -107,24 +107,24 @@ let apply_dyn_fun deflt f lobj =
f dodecl
let cache_object ((_,lobj) as node) =
- apply_dyn_fun () (fun d -> d.dyn_cache_function node) lobj
+ apply_dyn_fun (fun d -> d.dyn_cache_function node) lobj
let load_object i ((_,lobj) as node) =
- apply_dyn_fun () (fun d -> d.dyn_load_function i node) lobj
+ apply_dyn_fun (fun d -> d.dyn_load_function i node) lobj
let open_object i ((_,lobj) as node) =
- apply_dyn_fun () (fun d -> d.dyn_open_function i node) lobj
+ apply_dyn_fun (fun d -> d.dyn_open_function i node) lobj
let subst_object ((_,lobj) as node) =
- apply_dyn_fun lobj (fun d -> d.dyn_subst_function node) lobj
+ apply_dyn_fun (fun d -> d.dyn_subst_function node) lobj
let classify_object lobj =
- apply_dyn_fun Dispose (fun d -> d.dyn_classify_function lobj) lobj
+ apply_dyn_fun (fun d -> d.dyn_classify_function lobj) lobj
let discharge_object ((_,lobj) as node) =
- apply_dyn_fun None (fun d -> d.dyn_discharge_function node) lobj
+ apply_dyn_fun (fun d -> d.dyn_discharge_function node) lobj
let rebuild_object lobj =
- apply_dyn_fun lobj (fun d -> d.dyn_rebuild_function lobj) lobj
+ apply_dyn_fun (fun d -> d.dyn_rebuild_function lobj) lobj
let dump = Dyn.dump
diff --git a/library/library.ml b/library/library.ml
index 400f3dcf1..0ff82d7cc 100644
--- a/library/library.ml
+++ b/library/library.ml
@@ -167,7 +167,7 @@ let opened_libraries () = !libraries_imports_list
let register_loaded_library m =
let libname = m.libsum_name in
- let link m =
+ let link () =
let dirname = Filename.dirname (library_full_filename libname) in
let prefix = Nativecode.mod_uid_of_dirpath libname ^ "." in
let f = prefix ^ "cmo" in
@@ -176,7 +176,7 @@ let register_loaded_library m =
Nativelib.link_library ~prefix ~dirname ~basename:f
in
let rec aux = function
- | [] -> link m; [libname]
+ | [] -> link (); [libname]
| m'::_ as l when DirPath.equal m' libname -> l
| m'::l' -> m' :: aux l' in
libraries_loaded_list := aux !libraries_loaded_list;
@@ -288,16 +288,15 @@ let locate_absolute_library dir =
try
let name = Id.to_string base ^ ext in
let _, file = System.where_in_path ~warn:false loadpath name in
- [file]
- with Not_found -> [] in
- match find ".vo" @ find ".vio" with
- | [] -> raise LibNotFound
- | [file] -> file
- | [vo;vi] when Unix.((stat vo).st_mtime < (stat vi).st_mtime) ->
+ Some file
+ with Not_found -> None in
+ match find ".vo", find ".vio" with
+ | None, None -> raise LibNotFound
+ | Some file, None | None, Some file -> file
+ | Some vo, Some vi when Unix.((stat vo).st_mtime < (stat vi).st_mtime) ->
warn_several_object_files (vi, vo);
vi
- | [vo;vi] -> vo
- | _ -> assert false
+ | Some vo, Some _ -> vo
let locate_qualified_library ?root ?(warn = true) qid =
(* Search library in loadpath *)
@@ -309,18 +308,17 @@ let locate_qualified_library ?root ?(warn = true) qid =
let name = Id.to_string base ^ ext in
let lpath, file =
System.where_in_path ~warn (List.map fst loadpath) name in
- [lpath, file]
- with Not_found -> [] in
+ Some (lpath, file)
+ with Not_found -> None in
let lpath, file =
- match find ".vo" @ find ".vio" with
- | [] -> raise LibNotFound
- | [lpath, file] -> lpath, file
- | [lpath_vo, vo; lpath_vi, vi]
+ match find ".vo", find ".vio" with
+ | None, None -> raise LibNotFound
+ | Some res, None | None, Some res -> res
+ | Some (_, vo), Some (_, vi as resvi)
when Unix.((stat vo).st_mtime < (stat vi).st_mtime) ->
warn_several_object_files (vi, vo);
- lpath_vi, vi
- | [lpath_vo, vo; _ ] -> lpath_vo, vo
- | _ -> assert false
+ resvi
+ | Some resvo, Some _ -> resvo
in
let dir = add_dirpath_suffix (String.List.assoc lpath loadpath) base in
(* Look if loaded *)
diff --git a/plugins/btauto/g_btauto.ml4 b/plugins/btauto/g_btauto.mlg
index 3ae0f45cb..312ef1e55 100644
--- a/plugins/btauto/g_btauto.ml4
+++ b/plugins/btauto/g_btauto.mlg
@@ -8,11 +8,15 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
+{
+
open Ltac_plugin
+}
+
DECLARE PLUGIN "btauto_plugin"
TACTIC EXTEND btauto
-| [ "btauto" ] -> [ Refl_btauto.Btauto.tac ]
+| [ "btauto" ] -> { Refl_btauto.Btauto.tac }
END
diff --git a/plugins/cc/g_congruence.ml4 b/plugins/cc/g_congruence.mlg
index fb013ac13..685059294 100644
--- a/plugins/cc/g_congruence.ml4
+++ b/plugins/cc/g_congruence.mlg
@@ -8,22 +8,26 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
+{
+
open Ltac_plugin
open Cctac
open Stdarg
+}
+
DECLARE PLUGIN "cc_plugin"
(* Tactic registration *)
TACTIC EXTEND cc
- [ "congruence" ] -> [ congruence_tac 1000 [] ]
- |[ "congruence" integer(n) ] -> [ congruence_tac n [] ]
- |[ "congruence" "with" ne_constr_list(l) ] -> [ congruence_tac 1000 l ]
+| [ "congruence" ] -> { congruence_tac 1000 [] }
+| [ "congruence" integer(n) ] -> { congruence_tac n [] }
+| [ "congruence" "with" ne_constr_list(l) ] -> { congruence_tac 1000 l }
|[ "congruence" integer(n) "with" ne_constr_list(l) ] ->
- [ congruence_tac n l ]
+ { congruence_tac n l }
END
TACTIC EXTEND f_equal
- [ "f_equal" ] -> [ f_equal ]
+| [ "f_equal" ] -> { f_equal }
END
diff --git a/plugins/fourier/g_fourier.ml4 b/plugins/fourier/g_fourier.mlg
index 44560ac18..703e29f96 100644
--- a/plugins/fourier/g_fourier.ml4
+++ b/plugins/fourier/g_fourier.mlg
@@ -8,11 +8,15 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
+{
+
open Ltac_plugin
open FourierR
+}
+
DECLARE PLUGIN "fourier_plugin"
TACTIC EXTEND fourier
- [ "fourierz" ] -> [ fourier () ]
+| [ "fourierz" ] -> { fourier () }
END
diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml
index 31496513a..b2a528a1f 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -322,8 +322,7 @@ let generate_functional_principle (evd: Evd.evar_map ref)
try
let f = funs.(i) in
- let env = Global.env () in
- let type_sort = Evarutil.evd_comb1 (Evd.fresh_sort_in_family env) evd InType in
+ let type_sort = Evarutil.evd_comb1 Evd.fresh_sort_in_family evd InType in
let new_sorts =
match sorts with
| None -> Array.make (Array.length funs) (type_sort)
@@ -344,7 +343,7 @@ let generate_functional_principle (evd: Evd.evar_map ref)
(* let id_of_f = Label.to_id (con_label f) in *)
let register_with_sort fam_sort =
let evd' = Evd.from_env (Global.env ()) in
- let evd',s = Evd.fresh_sort_in_family env evd' fam_sort in
+ let evd',s = Evd.fresh_sort_in_family evd' fam_sort in
let name = Indrec.make_elimination_ident base_new_princ_name fam_sort in
let evd',value = change_property_sort evd' s new_principle_type new_princ_name in
let evd' = fst (Typing.type_of ~refresh:true (Global.env ()) evd' (EConstr.of_constr value)) in
@@ -354,7 +353,7 @@ let generate_functional_principle (evd: Evd.evar_map ref)
Evd.const_univ_entry ~poly evd'
in
let ce = Declare.definition_entry ~univs value in
- ignore(
+ ignore(
Declare.declare_constant
name
(DefinitionEntry ce,
@@ -508,8 +507,8 @@ let make_scheme evd (fas : (pconstant*Sorts.family) list) : Safe_typing.private_
let i = ref (-1) in
let sorts =
List.rev_map (fun (_,x) ->
- Evarutil.evd_comb1 (Evd.fresh_sort_in_family env) evd x
- )
+ Evarutil.evd_comb1 Evd.fresh_sort_in_family evd x
+ )
fas
in
(* We create the first priciple by tactic *)
diff --git a/plugins/ltac/coretactics.ml4 b/plugins/ltac/coretactics.mlg
index 61525cb49..6388906f5 100644
--- a/plugins/ltac/coretactics.ml4
+++ b/plugins/ltac/coretactics.mlg
@@ -8,6 +8,8 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
+{
+
open Util
open Locus
open Tactypes
@@ -18,147 +20,153 @@ open Tacarg
open Names
open Logic
+let wit_hyp = wit_var
+
+}
+
DECLARE PLUGIN "ltac_plugin"
(** Basic tactics *)
TACTIC EXTEND reflexivity
- [ "reflexivity" ] -> [ Tactics.intros_reflexivity ]
+| [ "reflexivity" ] -> { Tactics.intros_reflexivity }
END
TACTIC EXTEND exact
- [ "exact" casted_constr(c) ] -> [ Tactics.exact_no_check c ]
+| [ "exact" casted_constr(c) ] -> { Tactics.exact_no_check c }
END
TACTIC EXTEND assumption
- [ "assumption" ] -> [ Tactics.assumption ]
+| [ "assumption" ] -> { Tactics.assumption }
END
TACTIC EXTEND etransitivity
- [ "etransitivity" ] -> [ Tactics.intros_transitivity None ]
+| [ "etransitivity" ] -> { Tactics.intros_transitivity None }
END
TACTIC EXTEND cut
- [ "cut" constr(c) ] -> [ Tactics.cut c ]
+| [ "cut" constr(c) ] -> { Tactics.cut c }
END
TACTIC EXTEND exact_no_check
- [ "exact_no_check" constr(c) ] -> [ Tactics.exact_no_check c ]
+| [ "exact_no_check" constr(c) ] -> { Tactics.exact_no_check c }
END
TACTIC EXTEND vm_cast_no_check
- [ "vm_cast_no_check" constr(c) ] -> [ Tactics.vm_cast_no_check c ]
+| [ "vm_cast_no_check" constr(c) ] -> { Tactics.vm_cast_no_check c }
END
TACTIC EXTEND native_cast_no_check
- [ "native_cast_no_check" constr(c) ] -> [ Tactics.native_cast_no_check c ]
+| [ "native_cast_no_check" constr(c) ] -> { Tactics.native_cast_no_check c }
END
TACTIC EXTEND casetype
- [ "casetype" constr(c) ] -> [ Tactics.case_type c ]
+| [ "casetype" constr(c) ] -> { Tactics.case_type c }
END
TACTIC EXTEND elimtype
- [ "elimtype" constr(c) ] -> [ Tactics.elim_type c ]
+| [ "elimtype" constr(c) ] -> { Tactics.elim_type c }
END
TACTIC EXTEND lapply
- [ "lapply" constr(c) ] -> [ Tactics.cut_and_apply c ]
+| [ "lapply" constr(c) ] -> { Tactics.cut_and_apply c }
END
TACTIC EXTEND transitivity
- [ "transitivity" constr(c) ] -> [ Tactics.intros_transitivity (Some c) ]
+| [ "transitivity" constr(c) ] -> { Tactics.intros_transitivity (Some c) }
END
(** Left *)
TACTIC EXTEND left
- [ "left" ] -> [ Tactics.left_with_bindings false NoBindings ]
+| [ "left" ] -> { Tactics.left_with_bindings false NoBindings }
END
TACTIC EXTEND eleft
- [ "eleft" ] -> [ Tactics.left_with_bindings true NoBindings ]
+| [ "eleft" ] -> { Tactics.left_with_bindings true NoBindings }
END
TACTIC EXTEND left_with
- [ "left" "with" bindings(bl) ] -> [
+| [ "left" "with" bindings(bl) ] -> {
Tacticals.New.tclDELAYEDWITHHOLES false bl (fun bl -> Tactics.left_with_bindings false bl)
- ]
+ }
END
TACTIC EXTEND eleft_with
- [ "eleft" "with" bindings(bl) ] -> [
+| [ "eleft" "with" bindings(bl) ] -> {
Tacticals.New.tclDELAYEDWITHHOLES true bl (fun bl -> Tactics.left_with_bindings true bl)
- ]
+ }
END
(** Right *)
TACTIC EXTEND right
- [ "right" ] -> [ Tactics.right_with_bindings false NoBindings ]
+| [ "right" ] -> { Tactics.right_with_bindings false NoBindings }
END
TACTIC EXTEND eright
- [ "eright" ] -> [ Tactics.right_with_bindings true NoBindings ]
+| [ "eright" ] -> { Tactics.right_with_bindings true NoBindings }
END
TACTIC EXTEND right_with
- [ "right" "with" bindings(bl) ] -> [
+| [ "right" "with" bindings(bl) ] -> {
Tacticals.New.tclDELAYEDWITHHOLES false bl (fun bl -> Tactics.right_with_bindings false bl)
- ]
+ }
END
TACTIC EXTEND eright_with
- [ "eright" "with" bindings(bl) ] -> [
+| [ "eright" "with" bindings(bl) ] -> {
Tacticals.New.tclDELAYEDWITHHOLES true bl (fun bl -> Tactics.right_with_bindings true bl)
- ]
+ }
END
(** Constructor *)
TACTIC EXTEND constructor
- [ "constructor" ] -> [ Tactics.any_constructor false None ]
-| [ "constructor" int_or_var(i) ] -> [
+| [ "constructor" ] -> { Tactics.any_constructor false None }
+| [ "constructor" int_or_var(i) ] -> {
Tactics.constructor_tac false None i NoBindings
- ]
-| [ "constructor" int_or_var(i) "with" bindings(bl) ] -> [
+ }
+| [ "constructor" int_or_var(i) "with" bindings(bl) ] -> {
let tac bl = Tactics.constructor_tac false None i bl in
Tacticals.New.tclDELAYEDWITHHOLES false bl tac
- ]
+ }
END
TACTIC EXTEND econstructor
- [ "econstructor" ] -> [ Tactics.any_constructor true None ]
-| [ "econstructor" int_or_var(i) ] -> [
+| [ "econstructor" ] -> { Tactics.any_constructor true None }
+| [ "econstructor" int_or_var(i) ] -> {
Tactics.constructor_tac true None i NoBindings
- ]
-| [ "econstructor" int_or_var(i) "with" bindings(bl) ] -> [
+ }
+| [ "econstructor" int_or_var(i) "with" bindings(bl) ] -> {
let tac bl = Tactics.constructor_tac true None i bl in
Tacticals.New.tclDELAYEDWITHHOLES true bl tac
- ]
+ }
END
(** Specialize *)
TACTIC EXTEND specialize
- [ "specialize" constr_with_bindings(c) ] -> [
+| [ "specialize" constr_with_bindings(c) ] -> {
Tacticals.New.tclDELAYEDWITHHOLES false c (fun c -> Tactics.specialize c None)
- ]
-| [ "specialize" constr_with_bindings(c) "as" intropattern(ipat) ] -> [
+ }
+| [ "specialize" constr_with_bindings(c) "as" intropattern(ipat) ] -> {
Tacticals.New.tclDELAYEDWITHHOLES false c (fun c -> Tactics.specialize c (Some ipat))
- ]
+ }
END
TACTIC EXTEND symmetry
- [ "symmetry" ] -> [ Tactics.intros_symmetry {onhyps=Some[];concl_occs=AllOccurrences} ]
+| [ "symmetry" ] -> { Tactics.intros_symmetry {onhyps=Some[];concl_occs=AllOccurrences} }
END
TACTIC EXTEND symmetry_in
-| [ "symmetry" "in" in_clause(cl) ] -> [ Tactics.intros_symmetry cl ]
+| [ "symmetry" "in" in_clause(cl) ] -> { Tactics.intros_symmetry cl }
END
(** Split *)
+{
+
let rec delayed_list = function
| [] -> fun _ sigma -> (sigma, [])
| x :: l ->
@@ -167,147 +175,159 @@ let rec delayed_list = function
let (sigma, l) = delayed_list l env sigma in
(sigma, x :: l)
+}
+
TACTIC EXTEND split
- [ "split" ] -> [ Tactics.split_with_bindings false [NoBindings] ]
+| [ "split" ] -> { Tactics.split_with_bindings false [NoBindings] }
END
TACTIC EXTEND esplit
- [ "esplit" ] -> [ Tactics.split_with_bindings true [NoBindings] ]
+| [ "esplit" ] -> { Tactics.split_with_bindings true [NoBindings] }
END
TACTIC EXTEND split_with
- [ "split" "with" bindings(bl) ] -> [
+| [ "split" "with" bindings(bl) ] -> {
Tacticals.New.tclDELAYEDWITHHOLES false bl (fun bl -> Tactics.split_with_bindings false [bl])
- ]
+ }
END
TACTIC EXTEND esplit_with
- [ "esplit" "with" bindings(bl) ] -> [
+| [ "esplit" "with" bindings(bl) ] -> {
Tacticals.New.tclDELAYEDWITHHOLES true bl (fun bl -> Tactics.split_with_bindings true [bl])
- ]
+ }
END
TACTIC EXTEND exists
- [ "exists" ] -> [ Tactics.split_with_bindings false [NoBindings] ]
-| [ "exists" ne_bindings_list_sep(bll, ",") ] -> [
+| [ "exists" ] -> { Tactics.split_with_bindings false [NoBindings] }
+| [ "exists" ne_bindings_list_sep(bll, ",") ] -> {
Tacticals.New.tclDELAYEDWITHHOLES false (delayed_list bll) (fun bll -> Tactics.split_with_bindings false bll)
- ]
+ }
END
TACTIC EXTEND eexists
- [ "eexists" ] -> [ Tactics.split_with_bindings true [NoBindings] ]
-| [ "eexists" ne_bindings_list_sep(bll, ",") ] -> [
+| [ "eexists" ] -> { Tactics.split_with_bindings true [NoBindings] }
+| [ "eexists" ne_bindings_list_sep(bll, ",") ] -> {
Tacticals.New.tclDELAYEDWITHHOLES true (delayed_list bll) (fun bll -> Tactics.split_with_bindings true bll)
- ]
+ }
END
(** Intro *)
TACTIC EXTEND intros_until
- [ "intros" "until" quantified_hypothesis(h) ] -> [ Tactics.intros_until h ]
+| [ "intros" "until" quantified_hypothesis(h) ] -> { Tactics.intros_until h }
END
TACTIC EXTEND intro
-| [ "intro" ] -> [ Tactics.intro_move None MoveLast ]
-| [ "intro" ident(id) ] -> [ Tactics.intro_move (Some id) MoveLast ]
-| [ "intro" ident(id) "at" "top" ] -> [ Tactics.intro_move (Some id) MoveFirst ]
-| [ "intro" ident(id) "at" "bottom" ] -> [ Tactics.intro_move (Some id) MoveLast ]
-| [ "intro" ident(id) "after" hyp(h) ] -> [ Tactics.intro_move (Some id) (MoveAfter h) ]
-| [ "intro" ident(id) "before" hyp(h) ] -> [ Tactics.intro_move (Some id) (MoveBefore h) ]
-| [ "intro" "at" "top" ] -> [ Tactics.intro_move None MoveFirst ]
-| [ "intro" "at" "bottom" ] -> [ Tactics.intro_move None MoveLast ]
-| [ "intro" "after" hyp(h) ] -> [ Tactics.intro_move None (MoveAfter h) ]
-| [ "intro" "before" hyp(h) ] -> [ Tactics.intro_move None (MoveBefore h) ]
+| [ "intro" ] -> { Tactics.intro_move None MoveLast }
+| [ "intro" ident(id) ] -> { Tactics.intro_move (Some id) MoveLast }
+| [ "intro" ident(id) "at" "top" ] -> { Tactics.intro_move (Some id) MoveFirst }
+| [ "intro" ident(id) "at" "bottom" ] -> { Tactics.intro_move (Some id) MoveLast }
+| [ "intro" ident(id) "after" hyp(h) ] -> { Tactics.intro_move (Some id) (MoveAfter h) }
+| [ "intro" ident(id) "before" hyp(h) ] -> { Tactics.intro_move (Some id) (MoveBefore h) }
+| [ "intro" "at" "top" ] -> { Tactics.intro_move None MoveFirst }
+| [ "intro" "at" "bottom" ] -> { Tactics.intro_move None MoveLast }
+| [ "intro" "after" hyp(h) ] -> { Tactics.intro_move None (MoveAfter h) }
+| [ "intro" "before" hyp(h) ] -> { Tactics.intro_move None (MoveBefore h) }
END
(** Move *)
TACTIC EXTEND move
- [ "move" hyp(id) "at" "top" ] -> [ Tactics.move_hyp id MoveFirst ]
-| [ "move" hyp(id) "at" "bottom" ] -> [ Tactics.move_hyp id MoveLast ]
-| [ "move" hyp(id) "after" hyp(h) ] -> [ Tactics.move_hyp id (MoveAfter h) ]
-| [ "move" hyp(id) "before" hyp(h) ] -> [ Tactics.move_hyp id (MoveBefore h) ]
+| [ "move" hyp(id) "at" "top" ] -> { Tactics.move_hyp id MoveFirst }
+| [ "move" hyp(id) "at" "bottom" ] -> { Tactics.move_hyp id MoveLast }
+| [ "move" hyp(id) "after" hyp(h) ] -> { Tactics.move_hyp id (MoveAfter h) }
+| [ "move" hyp(id) "before" hyp(h) ] -> { Tactics.move_hyp id (MoveBefore h) }
END
(** Rename *)
TACTIC EXTEND rename
-| [ "rename" ne_rename_list_sep(ids, ",") ] -> [ Tactics.rename_hyp ids ]
+| [ "rename" ne_rename_list_sep(ids, ",") ] -> { Tactics.rename_hyp ids }
END
(** Revert *)
TACTIC EXTEND revert
- [ "revert" ne_hyp_list(hl) ] -> [ Tactics.revert hl ]
+| [ "revert" ne_hyp_list(hl) ] -> { Tactics.revert hl }
END
(** Simple induction / destruct *)
+{
+
let simple_induct h =
Tacticals.New.tclTHEN (Tactics.intros_until h)
(Tacticals.New.onLastHyp Tactics.simplest_elim)
+}
+
TACTIC EXTEND simple_induction
- [ "simple" "induction" quantified_hypothesis(h) ] -> [ simple_induct h ]
+| [ "simple" "induction" quantified_hypothesis(h) ] -> { simple_induct h }
END
+{
+
let simple_destruct h =
Tacticals.New.tclTHEN (Tactics.intros_until h)
(Tacticals.New.onLastHyp Tactics.simplest_case)
+}
+
TACTIC EXTEND simple_destruct
- [ "simple" "destruct" quantified_hypothesis(h) ] -> [ simple_destruct h ]
+| [ "simple" "destruct" quantified_hypothesis(h) ] -> { simple_destruct h }
END
(** Double induction *)
TACTIC EXTEND double_induction
- [ "double" "induction" quantified_hypothesis(h1) quantified_hypothesis(h2) ] ->
- [ Elim.h_double_induction h1 h2 ]
+| [ "double" "induction" quantified_hypothesis(h1) quantified_hypothesis(h2) ] ->
+ { Elim.h_double_induction h1 h2 }
END
(* Admit *)
TACTIC EXTEND admit
- [ "admit" ] -> [ Proofview.give_up ]
+|[ "admit" ] -> { Proofview.give_up }
END
(* Fix *)
TACTIC EXTEND fix
- [ "fix" ident(id) natural(n) ] -> [ Tactics.fix id n ]
+| [ "fix" ident(id) natural(n) ] -> { Tactics.fix id n }
END
(* Cofix *)
TACTIC EXTEND cofix
- [ "cofix" ident(id) ] -> [ Tactics.cofix id ]
+| [ "cofix" ident(id) ] -> { Tactics.cofix id }
END
(* Clear *)
TACTIC EXTEND clear
- [ "clear" hyp_list(ids) ] -> [
+| [ "clear" hyp_list(ids) ] -> {
if List.is_empty ids then Tactics.keep []
else Tactics.clear ids
- ]
-| [ "clear" "-" ne_hyp_list(ids) ] -> [ Tactics.keep ids ]
+ }
+| [ "clear" "-" ne_hyp_list(ids) ] -> { Tactics.keep ids }
END
(* Clearbody *)
TACTIC EXTEND clearbody
- [ "clearbody" ne_hyp_list(ids) ] -> [ Tactics.clear_body ids ]
+| [ "clearbody" ne_hyp_list(ids) ] -> { Tactics.clear_body ids }
END
(* Generalize dependent *)
TACTIC EXTEND generalize_dependent
- [ "generalize" "dependent" constr(c) ] -> [ Tactics.generalize_dep c ]
+| [ "generalize" "dependent" constr(c) ] -> { Tactics.generalize_dep c }
END
(* Table of "pervasives" macros tactics (e.g. auto, simpl, etc.) *)
+{
+
open Tacexpr
let initial_atomic () =
@@ -364,3 +384,5 @@ let initial_tacticals () =
]
let () = Mltop.declare_cache_obj initial_tacticals "ltac_plugin"
+
+}
diff --git a/plugins/ltac/g_eqdecide.ml4 b/plugins/ltac/g_eqdecide.mlg
index 2251a6620..e57afe3e3 100644
--- a/plugins/ltac/g_eqdecide.ml4
+++ b/plugins/ltac/g_eqdecide.mlg
@@ -14,15 +14,19 @@
(* by Eduardo Gimenez *)
(************************************************************************)
+{
+
open Eqdecide
open Stdarg
+}
+
DECLARE PLUGIN "ltac_plugin"
TACTIC EXTEND decide_equality
-| [ "decide" "equality" ] -> [ decideEqualityGoal ]
+| [ "decide" "equality" ] -> { decideEqualityGoal }
END
TACTIC EXTEND compare
-| [ "compare" constr(c1) constr(c2) ] -> [ compare c1 c2 ]
+| [ "compare" constr(c1) constr(c2) ] -> { compare c1 c2 }
END
diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml
index 09179dad3..4357689ee 100644
--- a/plugins/ltac/pptactic.ml
+++ b/plugins/ltac/pptactic.ml
@@ -115,7 +115,7 @@ let string_of_genarg_arg (ArgumentType arg) =
let keyword x = tag_keyword (str x)
let primitive x = tag_primitive (str x)
- let has_type (Val.Dyn (tag, x)) t = match Val.eq tag t with
+ let has_type (Val.Dyn (tag, _)) t = match Val.eq tag t with
| None -> false
| Some _ -> true
@@ -188,7 +188,7 @@ let string_of_genarg_arg (ArgumentType arg) =
| AN v -> f v
| ByNotation (s,sc) -> qs s ++ pr_opt (fun sc -> str "%" ++ str sc) sc)
- let pr_located pr (loc,x) = pr x
+ let pr_located pr (_,x) = pr x
let pr_evaluable_reference = function
| EvalVarRef id -> pr_id id
@@ -240,7 +240,7 @@ let string_of_genarg_arg (ArgumentType arg) =
in
pr_sequence (fun x -> x) l
- let pr_extend_gen pr_gen lev { mltac_name = s; mltac_index = i } l =
+ let pr_extend_gen pr_gen _ { mltac_name = s; mltac_index = i } l =
let name =
str s.mltac_plugin ++ str "::" ++ str s.mltac_tactic ++
str "@" ++ int i
@@ -260,7 +260,7 @@ let string_of_genarg_arg (ArgumentType arg) =
| Extend.Uentry tag ->
let ArgT.Any tag = tag in
ArgT.repr tag
- | Extend.Uentryl (tkn, lvl) -> "tactic" ^ string_of_int lvl
+ | Extend.Uentryl (_, lvl) -> "tactic" ^ string_of_int lvl
let pr_alias_key key =
try
@@ -288,7 +288,7 @@ let string_of_genarg_arg (ArgumentType arg) =
let p = pr_tacarg_using_rule pr_gen prods in
if pp.pptac_level > lev then surround p else p
with Not_found ->
- let pr arg = str "_" in
+ let pr _ = str "_" in
KerName.print key ++ spc() ++ pr_sequence pr l ++ str" (* Generic printer *)"
let pr_farg prtac arg = prtac (1, Any) (TacArg (Loc.tag arg))
@@ -341,14 +341,14 @@ let string_of_genarg_arg (ArgumentType arg) =
pr_any_arg pr symb arg
| _ -> str "ltac:(" ++ prtac (1, Any) arg ++ str ")"
- let pr_raw_extend_rec prc prlc prtac prpat =
+ let pr_raw_extend_rec prtac =
pr_extend_gen (pr_farg prtac)
- let pr_glob_extend_rec prc prlc prtac prpat =
+ let pr_glob_extend_rec prtac =
pr_extend_gen (pr_farg prtac)
- let pr_raw_alias prc prlc prtac prpat lev key args =
+ let pr_raw_alias prtac lev key args =
pr_alias_gen (pr_targ (fun l a -> prtac l (TacArg (Loc.tag a)))) lev key args
- let pr_glob_alias prc prlc prtac prpat lev key args =
+ let pr_glob_alias prtac lev key args =
pr_alias_gen (pr_targ (fun l a -> prtac l (TacArg (Loc.tag a)))) lev key args
(**********************************************************************)
@@ -743,7 +743,7 @@ let pr_goal_selector ~toplevel s =
(* Main tactic printer *)
and pr_atom1 a = tag_atom a (match a with
(* Basic tactics *)
- | TacIntroPattern (ev,[]) as t ->
+ | TacIntroPattern (_,[]) as t ->
pr_atom0 t
| TacIntroPattern (ev,(_::_ as p)) ->
hov 1 (primitive (if ev then "eintros" else "intros") ++
@@ -1054,7 +1054,7 @@ let pr_goal_selector ~toplevel s =
primitive "fresh" ++ pr_fresh_ids l, latom
| TacArg(_,TacGeneric arg) ->
pr.pr_generic arg, latom
- | TacArg(_,TacCall(loc,(f,[]))) ->
+ | TacArg(_,TacCall(_,(f,[]))) ->
pr.pr_reference f, latom
| TacArg(_,TacCall(loc,(f,l))) ->
pr_with_comments ?loc (hov 1 (
@@ -1112,8 +1112,8 @@ let pr_goal_selector ~toplevel s =
pr_reference = pr_qualid;
pr_name = pr_lident;
pr_generic = (fun arg -> Pputils.pr_raw_generic (Global.env ()) arg);
- pr_extend = pr_raw_extend_rec pr_constr_expr pr_lconstr_expr pr_raw_tactic_level pr_constr_pattern_expr;
- pr_alias = pr_raw_alias pr_constr_expr pr_lconstr_expr pr_raw_tactic_level pr_constr_pattern_expr;
+ pr_extend = pr_raw_extend_rec pr_raw_tactic_level;
+ pr_alias = pr_raw_alias pr_raw_tactic_level;
} in
make_pr_tac
pr raw_printers
@@ -1142,12 +1142,8 @@ let pr_goal_selector ~toplevel s =
pr_reference = pr_ltac_or_var (pr_located pr_ltac_constant);
pr_name = pr_lident;
pr_generic = (fun arg -> Pputils.pr_glb_generic (Global.env ()) arg);
- pr_extend = pr_glob_extend_rec
- (pr_and_constr_expr (pr_glob_constr_env env)) (pr_and_constr_expr (pr_lglob_constr_env env))
- prtac (pr_pat_and_constr_expr (pr_glob_constr_env env));
- pr_alias = pr_glob_alias
- (pr_and_constr_expr (pr_glob_constr_env env)) (pr_and_constr_expr (pr_lglob_constr_env env))
- prtac (pr_pat_and_constr_expr (pr_glob_constr_env env));
+ pr_extend = pr_glob_extend_rec prtac;
+ pr_alias = pr_glob_alias prtac;
} in
make_pr_tac
pr glob_printers
@@ -1168,8 +1164,8 @@ let pr_goal_selector ~toplevel s =
| _ -> user_err Pp.(str "Cannot translate fix tactic: not enough products") in
strip_ty [] n ty
- let pr_atomic_tactic_level env sigma n t =
- let prtac n (t:atomic_tactic_expr) =
+ let pr_atomic_tactic_level env sigma t =
+ let prtac (t:atomic_tactic_expr) =
let pr = {
pr_tactic = (fun _ _ -> str "<tactic>");
pr_constr = (fun c -> pr_econstr_env env sigma c);
@@ -1188,18 +1184,15 @@ let pr_goal_selector ~toplevel s =
in
pr_atom pr strip_prod_binders_constr tag_atomic_tactic_expr t
in
- prtac n t
+ prtac t
let pr_raw_generic = Pputils.pr_raw_generic
let pr_glb_generic = Pputils.pr_glb_generic
- let pr_raw_extend env = pr_raw_extend_rec
- pr_constr_expr pr_lconstr_expr pr_raw_tactic_level pr_constr_pattern_expr
+ let pr_raw_extend _ = pr_raw_extend_rec pr_raw_tactic_level
- let pr_glob_extend env = pr_glob_extend_rec
- (pr_and_constr_expr (pr_glob_constr_env env)) (pr_and_constr_expr (pr_lglob_constr_env env))
- (pr_glob_tactic_level env) (pr_pat_and_constr_expr (pr_glob_constr_env env))
+ let pr_glob_extend env = pr_glob_extend_rec (pr_glob_tactic_level env)
let pr_alias pr lev key args =
pr_alias_gen (fun _ arg -> pr arg) lev key args
@@ -1207,14 +1200,14 @@ let pr_goal_selector ~toplevel s =
let pr_extend pr lev ml args =
pr_extend_gen pr lev ml args
- let pr_atomic_tactic env sigma c = pr_atomic_tactic_level env sigma ltop c
+ let pr_atomic_tactic env sigma c = pr_atomic_tactic_level env sigma c
let declare_extra_genarg_pprule wit
(f : 'a raw_extra_genarg_printer)
(g : 'b glob_extra_genarg_printer)
(h : 'c extra_genarg_printer) =
begin match wit with
- | ExtraArg s -> ()
+ | ExtraArg _ -> ()
| _ -> user_err Pp.(str "Can declare a pretty-printing rule only for extra argument types.")
end;
let f x =
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index 01c52c413..9f8cd2fc4 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -409,7 +409,7 @@ module TypeGlobal = struct
let inverse env (evd,cstrs) car rel =
- let (evd, sort) = Evarutil.new_Type ~rigid:Evd.univ_flexible env evd in
+ let (evd, sort) = Evarutil.new_Type ~rigid:Evd.univ_flexible evd in
app_poly_check env (evd,cstrs) coq_inverse [| car ; car; sort; rel |]
end
diff --git a/plugins/ltac/taccoerce.ml b/plugins/ltac/taccoerce.ml
index 84baea964..026c00b84 100644
--- a/plugins/ltac/taccoerce.ml
+++ b/plugins/ltac/taccoerce.ml
@@ -165,8 +165,7 @@ let coerce_var_to_ident fresh env sigma v =
(* Interprets, if possible, a constr to an identifier which may not
be fresh but suitable to be given to the fresh tactic. Works for
vars, constants, inductive, constructors and sorts. *)
-let coerce_to_ident_not_fresh env sigma v =
-let g = sigma in
+let coerce_to_ident_not_fresh sigma v =
let id_of_name = function
| Name.Anonymous -> Id.of_string "x"
| Name.Name x -> x in
@@ -183,9 +182,9 @@ let id_of_name = function
| Some c ->
match EConstr.kind sigma c with
| Var id -> id
- | Meta m -> id_of_name (Evd.meta_name g m)
+ | Meta m -> id_of_name (Evd.meta_name sigma m)
| Evar (kn,_) ->
- begin match Evd.evar_ident kn g with
+ begin match Evd.evar_ident kn sigma with
| None -> fail ()
| Some id -> id
end
@@ -208,7 +207,7 @@ let id_of_name = function
| _ -> fail()
-let coerce_to_intro_pattern env sigma v =
+let coerce_to_intro_pattern sigma v =
if has_type v (topwit wit_intro_pattern) then
(out_gen (topwit wit_intro_pattern) v).CAst.v
else if has_type v (topwit wit_var) then
@@ -221,8 +220,8 @@ let coerce_to_intro_pattern env sigma v =
IntroNaming (IntroIdentifier (destVar sigma c))
| _ -> raise (CannotCoerceTo "an introduction pattern")
-let coerce_to_intro_pattern_naming env sigma v =
- match coerce_to_intro_pattern env sigma v with
+let coerce_to_intro_pattern_naming sigma v =
+ match coerce_to_intro_pattern sigma v with
| IntroNaming pat -> pat
| _ -> raise (CannotCoerceTo "a naming introduction pattern")
@@ -255,7 +254,7 @@ let coerce_to_constr env v =
(try [], constr_of_id env id with Not_found -> fail ())
else fail ()
-let coerce_to_uconstr env v =
+let coerce_to_uconstr v =
if has_type v (topwit wit_uconstr) then
out_gen (topwit wit_uconstr) v
else
@@ -299,11 +298,11 @@ let coerce_to_constr_list env v =
List.map map l
| None -> raise (CannotCoerceTo "a term list")
-let coerce_to_intro_pattern_list ?loc env sigma v =
+let coerce_to_intro_pattern_list ?loc sigma v =
match Value.to_list v with
| None -> raise (CannotCoerceTo "an intro pattern list")
| Some l ->
- let map v = CAst.make ?loc @@ coerce_to_intro_pattern env sigma v in
+ let map v = CAst.make ?loc @@ coerce_to_intro_pattern sigma v in
List.map map l
let coerce_to_hyp env sigma v =
@@ -328,7 +327,7 @@ let coerce_to_hyp_list env sigma v =
| None -> raise (CannotCoerceTo "a variable list")
(* Interprets a qualified name *)
-let coerce_to_reference env sigma v =
+let coerce_to_reference sigma v =
match Value.to_constr v with
| Some c ->
begin
@@ -356,7 +355,7 @@ let coerce_to_quantified_hypothesis sigma v =
(* Quantified named or numbered hypothesis or hypothesis in context *)
(* (as in Inversion) *)
-let coerce_to_decl_or_quant_hyp env sigma v =
+let coerce_to_decl_or_quant_hyp sigma v =
if has_type v (topwit wit_int) then
AnonHyp (out_gen (topwit wit_int) v)
else
diff --git a/plugins/ltac/taccoerce.mli b/plugins/ltac/taccoerce.mli
index 56f881684..d2ae92f6c 100644
--- a/plugins/ltac/taccoerce.mli
+++ b/plugins/ltac/taccoerce.mli
@@ -51,12 +51,12 @@ val coerce_to_constr_context : Value.t -> constr
val coerce_var_to_ident : bool -> Environ.env -> Evd.evar_map -> Value.t -> Id.t
-val coerce_to_ident_not_fresh : Environ.env -> Evd.evar_map -> Value.t -> Id.t
+val coerce_to_ident_not_fresh : Evd.evar_map -> Value.t -> Id.t
-val coerce_to_intro_pattern : Environ.env -> Evd.evar_map -> Value.t -> Tacexpr.delayed_open_constr intro_pattern_expr
+val coerce_to_intro_pattern : Evd.evar_map -> Value.t -> Tacexpr.delayed_open_constr intro_pattern_expr
val coerce_to_intro_pattern_naming :
- Environ.env -> Evd.evar_map -> Value.t -> Namegen.intro_pattern_naming_expr
+ Evd.evar_map -> Value.t -> Namegen.intro_pattern_naming_expr
val coerce_to_hint_base : Value.t -> string
@@ -64,7 +64,7 @@ val coerce_to_int : Value.t -> int
val coerce_to_constr : Environ.env -> Value.t -> Ltac_pretype.constr_under_binders
-val coerce_to_uconstr : Environ.env -> Value.t -> Ltac_pretype.closed_glob_constr
+val coerce_to_uconstr : Value.t -> Ltac_pretype.closed_glob_constr
val coerce_to_closed_constr : Environ.env -> Value.t -> constr
@@ -74,17 +74,17 @@ val coerce_to_evaluable_ref :
val coerce_to_constr_list : Environ.env -> Value.t -> constr list
val coerce_to_intro_pattern_list :
- ?loc:Loc.t -> Environ.env -> Evd.evar_map -> Value.t -> Tacexpr.intro_patterns
+ ?loc:Loc.t -> Evd.evar_map -> Value.t -> Tacexpr.intro_patterns
val coerce_to_hyp : Environ.env -> Evd.evar_map -> Value.t -> Id.t
val coerce_to_hyp_list : Environ.env -> Evd.evar_map -> Value.t -> Id.t list
-val coerce_to_reference : Environ.env -> Evd.evar_map -> Value.t -> GlobRef.t
+val coerce_to_reference : Evd.evar_map -> Value.t -> GlobRef.t
val coerce_to_quantified_hypothesis : Evd.evar_map -> Value.t -> quantified_hypothesis
-val coerce_to_decl_or_quant_hyp : Environ.env -> Evd.evar_map -> Value.t -> quantified_hypothesis
+val coerce_to_decl_or_quant_hyp : Evd.evar_map -> Value.t -> quantified_hypothesis
val coerce_to_int_or_var_list : Value.t -> int Locus.or_var list
diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml
index 876e6f320..fac464a62 100644
--- a/plugins/ltac/tacentries.ml
+++ b/plugins/ltac/tacentries.ml
@@ -554,13 +554,18 @@ let () =
] in
register_grammars_by_name "tactic" entries
+let get_identifier id =
+ (** Workaround for badly-designed generic arguments lacking a closure *)
+ Names.Id.of_string_soft ("$" ^ id)
+
+
type _ ty_sig =
| TyNil : (Geninterp.interp_sign -> unit Proofview.tactic) ty_sig
| TyIdent : string * 'r ty_sig -> 'r ty_sig
| TyArg :
- (('a, 'b, 'c) Extend.ty_user_symbol * Id.t) Loc.located * 'r ty_sig -> ('c -> 'r) ty_sig
+ ('a, 'b, 'c) Extend.ty_user_symbol * string * 'r ty_sig -> ('c -> 'r) ty_sig
| TyAnonArg :
- ('a, 'b, 'c) Extend.ty_user_symbol Loc.located * 'r ty_sig -> 'r ty_sig
+ ('a, 'b, 'c) Extend.ty_user_symbol * 'r ty_sig -> 'r ty_sig
type ty_ml = TyML : 'r ty_sig * 'r -> ty_ml
@@ -578,10 +583,11 @@ let rec clause_of_sign : type a. a ty_sig -> Genarg.ArgT.any Extend.user_symbol
fun sign -> match sign with
| TyNil -> []
| TyIdent (s, sig') -> TacTerm s :: clause_of_sign sig'
- | TyArg ((loc,(a,id)),sig') ->
- TacNonTerm (loc,(untype_user_symbol a,Some id)) :: clause_of_sign sig'
- | TyAnonArg ((loc,a),sig') ->
- TacNonTerm (loc,(untype_user_symbol a,None)) :: clause_of_sign sig'
+ | TyArg (a, id, sig') ->
+ let id = get_identifier id in
+ TacNonTerm (None,(untype_user_symbol a,Some id)) :: clause_of_sign sig'
+ | TyAnonArg (a, sig') ->
+ TacNonTerm (None,(untype_user_symbol a,None)) :: clause_of_sign sig'
let clause_of_ty_ml = function
| TyML (t,_) -> clause_of_sign t
@@ -604,7 +610,7 @@ let rec eval_sign : type a. a ty_sig -> a -> Geninterp.Val.t list -> Geninterp.i
| _ :: _ -> assert false
end
| TyIdent (s, sig') -> eval_sign sig' tac
- | TyArg ((_loc,(a,id)), sig') ->
+ | TyArg (a, _, sig') ->
let f = eval_sign sig' in
begin fun tac vals ist -> match vals with
| [] -> assert false
@@ -612,7 +618,7 @@ let rec eval_sign : type a. a ty_sig -> a -> Geninterp.Val.t list -> Geninterp.i
let v' = Taccoerce.Value.cast (topwit (prj a)) v in
f (tac v') vals ist
end tac
- | TyAnonArg ((_loc,a), sig') -> eval_sign sig' tac
+ | TyAnonArg (a, sig') -> eval_sign sig' tac
let eval : ty_ml -> Geninterp.Val.t list -> Geninterp.interp_sign -> unit Proofview.tactic = function
| TyML (t,tac) -> eval_sign t tac
@@ -624,14 +630,14 @@ let is_constr_entry = function
let rec only_constr : type a. a ty_sig -> bool = function
| TyNil -> true
| TyIdent(_,_) -> false
-| TyArg((_,(u,_)),s) -> if is_constr_entry u then only_constr s else false
-| TyAnonArg((_,u),s) -> if is_constr_entry u then only_constr s else false
+| TyArg (u, _, s) -> if is_constr_entry u then only_constr s else false
+| TyAnonArg (u, s) -> if is_constr_entry u then only_constr s else false
let rec mk_sign_vars : type a. a ty_sig -> Name.t list = function
| TyNil -> []
| TyIdent (_,s) -> mk_sign_vars s
-| TyArg((_,(_,name)),s) -> Name name :: mk_sign_vars s
-| TyAnonArg((_,_),s) -> Anonymous :: mk_sign_vars s
+| TyArg (_, name, s) -> Name (get_identifier name) :: mk_sign_vars s
+| TyAnonArg (_, s) -> Anonymous :: mk_sign_vars s
let dummy_id = Id.of_string "_"
diff --git a/plugins/ltac/tacentries.mli b/plugins/ltac/tacentries.mli
index 2bfbbe2e1..9bba9ba71 100644
--- a/plugins/ltac/tacentries.mli
+++ b/plugins/ltac/tacentries.mli
@@ -72,9 +72,9 @@ type _ ty_sig =
| TyNil : (Geninterp.interp_sign -> unit Proofview.tactic) ty_sig
| TyIdent : string * 'r ty_sig -> 'r ty_sig
| TyArg :
- (('a, 'b, 'c) Extend.ty_user_symbol * Names.Id.t) Loc.located * 'r ty_sig -> ('c -> 'r) ty_sig
+ ('a, 'b, 'c) Extend.ty_user_symbol * string * 'r ty_sig -> ('c -> 'r) ty_sig
| TyAnonArg :
- ('a, 'b, 'c) Extend.ty_user_symbol Loc.located * 'r ty_sig -> 'r ty_sig
+ ('a, 'b, 'c) Extend.ty_user_symbol * 'r ty_sig -> 'r ty_sig
type ty_ml = TyML : 'r ty_sig * 'r -> ty_ml
diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml
index 9d1cc1643..d9ac96d89 100644
--- a/plugins/ltac/tacinterp.ml
+++ b/plugins/ltac/tacinterp.ml
@@ -312,11 +312,11 @@ let interp_name ist env sigma = function
| Name id -> Name (interp_ident ist env sigma id)
let interp_intro_pattern_var loc ist env sigma id =
- try try_interp_ltac_var (coerce_to_intro_pattern env sigma) ist (Some (env,sigma)) (make ?loc id)
+ try try_interp_ltac_var (coerce_to_intro_pattern sigma) ist (Some (env,sigma)) (make ?loc id)
with Not_found -> IntroNaming (IntroIdentifier id)
let interp_intro_pattern_naming_var loc ist env sigma id =
- try try_interp_ltac_var (coerce_to_intro_pattern_naming env sigma) ist (Some (env,sigma)) (make ?loc id)
+ try try_interp_ltac_var (coerce_to_intro_pattern_naming sigma) ist (Some (env,sigma)) (make ?loc id)
with Not_found -> IntroIdentifier id
let interp_int ist ({loc;v=id} as locid) =
@@ -357,7 +357,7 @@ let interp_hyp_list ist env sigma l =
let interp_reference ist env sigma = function
| ArgArg (_,r) -> r
| ArgVar {loc;v=id} ->
- try try_interp_ltac_var (coerce_to_reference env sigma) ist (Some (env,sigma)) (make ?loc id)
+ try try_interp_ltac_var (coerce_to_reference sigma) ist (Some (env,sigma)) (make ?loc id)
with Not_found ->
try
VarRef (get_id (Environ.lookup_named id env))
@@ -451,7 +451,7 @@ let default_fresh_id = Id.of_string "H"
let interp_fresh_id ist env sigma l =
let extract_ident ist env sigma id =
- try try_interp_ltac_var (coerce_to_ident_not_fresh env sigma)
+ try try_interp_ltac_var (coerce_to_ident_not_fresh sigma)
ist (Some (env,sigma)) (make id)
with Not_found -> id in
let ids = List.map_filter (function ArgVar {v=id} -> Some id | _ -> None) l in
@@ -474,7 +474,7 @@ let interp_fresh_id ist env sigma l =
(* Extract the uconstr list from lfun *)
let extract_ltac_constr_context ist env sigma =
let add_uconstr id v map =
- try Id.Map.add id (coerce_to_uconstr env v) map
+ try Id.Map.add id (coerce_to_uconstr v) map
with CannotCoerceTo _ -> map
in
let add_constr id v map =
@@ -799,7 +799,7 @@ and interp_or_and_intro_pattern ist env sigma = function
and interp_intro_pattern_list_as_list ist env sigma = function
| [{loc;v=IntroNaming (IntroIdentifier id)}] as l ->
- (try sigma, coerce_to_intro_pattern_list ?loc env sigma (Id.Map.find id ist.lfun)
+ (try sigma, coerce_to_intro_pattern_list ?loc sigma (Id.Map.find id ist.lfun)
with Not_found | CannotCoerceTo _ ->
List.fold_left_map (interp_intro_pattern ist env) sigma l)
| l -> List.fold_left_map (interp_intro_pattern ist env) sigma l
@@ -842,7 +842,7 @@ let interp_declared_or_quantified_hypothesis ist env sigma = function
| AnonHyp n -> AnonHyp n
| NamedHyp id ->
try try_interp_ltac_var
- (coerce_to_decl_or_quant_hyp env sigma) ist (Some (env,sigma)) (make id)
+ (coerce_to_decl_or_quant_hyp sigma) ist (Some (env,sigma)) (make id)
with Not_found -> NamedHyp id
let interp_binding ist env sigma {loc;v=(b,c)} =
diff --git a/plugins/micromega/g_micromega.ml4 b/plugins/micromega/g_micromega.mlg
index 81140a46a..21f0414e9 100644
--- a/plugins/micromega/g_micromega.ml4
+++ b/plugins/micromega/g_micromega.mlg
@@ -16,70 +16,74 @@
(* *)
(************************************************************************)
+{
+
open Ltac_plugin
open Stdarg
open Tacarg
+}
+
DECLARE PLUGIN "micromega_plugin"
TACTIC EXTEND RED
-| [ "myred" ] -> [ Tactics.red_in_concl ]
+| [ "myred" ] -> { Tactics.red_in_concl }
END
TACTIC EXTEND PsatzZ
-| [ "psatz_Z" int_or_var(i) tactic(t) ] -> [ (Coq_micromega.psatz_Z i
+| [ "psatz_Z" int_or_var(i) tactic(t) ] -> { (Coq_micromega.psatz_Z i
(Tacinterp.tactic_of_value ist t))
- ]
-| [ "psatz_Z" tactic(t)] -> [ (Coq_micromega.psatz_Z (-1)) (Tacinterp.tactic_of_value ist t) ]
+ }
+| [ "psatz_Z" tactic(t)] -> { (Coq_micromega.psatz_Z (-1)) (Tacinterp.tactic_of_value ist t) }
END
TACTIC EXTEND Lia
-[ "xlia" tactic(t) ] -> [ (Coq_micromega.xlia (Tacinterp.tactic_of_value ist t)) ]
+| [ "xlia" tactic(t) ] -> { (Coq_micromega.xlia (Tacinterp.tactic_of_value ist t)) }
END
TACTIC EXTEND Nia
-[ "xnlia" tactic(t) ] -> [ (Coq_micromega.xnlia (Tacinterp.tactic_of_value ist t)) ]
+| [ "xnlia" tactic(t) ] -> { (Coq_micromega.xnlia (Tacinterp.tactic_of_value ist t)) }
END
TACTIC EXTEND NRA
-[ "xnra" tactic(t) ] -> [ (Coq_micromega.nra (Tacinterp.tactic_of_value ist t))]
+| [ "xnra" tactic(t) ] -> { (Coq_micromega.nra (Tacinterp.tactic_of_value ist t))}
END
TACTIC EXTEND NQA
-[ "xnqa" tactic(t) ] -> [ (Coq_micromega.nqa (Tacinterp.tactic_of_value ist t))]
+| [ "xnqa" tactic(t) ] -> { (Coq_micromega.nqa (Tacinterp.tactic_of_value ist t))}
END
TACTIC EXTEND Sos_Z
-| [ "sos_Z" tactic(t) ] -> [ (Coq_micromega.sos_Z (Tacinterp.tactic_of_value ist t)) ]
+| [ "sos_Z" tactic(t) ] -> { (Coq_micromega.sos_Z (Tacinterp.tactic_of_value ist t)) }
END
TACTIC EXTEND Sos_Q
-| [ "sos_Q" tactic(t) ] -> [ (Coq_micromega.sos_Q (Tacinterp.tactic_of_value ist t)) ]
+| [ "sos_Q" tactic(t) ] -> { (Coq_micromega.sos_Q (Tacinterp.tactic_of_value ist t)) }
END
TACTIC EXTEND Sos_R
-| [ "sos_R" tactic(t) ] -> [ (Coq_micromega.sos_R (Tacinterp.tactic_of_value ist t)) ]
+| [ "sos_R" tactic(t) ] -> { (Coq_micromega.sos_R (Tacinterp.tactic_of_value ist t)) }
END
TACTIC EXTEND LRA_Q
-[ "lra_Q" tactic(t) ] -> [ (Coq_micromega.lra_Q (Tacinterp.tactic_of_value ist t)) ]
+| [ "lra_Q" tactic(t) ] -> { (Coq_micromega.lra_Q (Tacinterp.tactic_of_value ist t)) }
END
TACTIC EXTEND LRA_R
-[ "lra_R" tactic(t) ] -> [ (Coq_micromega.lra_R (Tacinterp.tactic_of_value ist t)) ]
+| [ "lra_R" tactic(t) ] -> { (Coq_micromega.lra_R (Tacinterp.tactic_of_value ist t)) }
END
TACTIC EXTEND PsatzR
-| [ "psatz_R" int_or_var(i) tactic(t) ] -> [ (Coq_micromega.psatz_R i (Tacinterp.tactic_of_value ist t)) ]
-| [ "psatz_R" tactic(t) ] -> [ (Coq_micromega.psatz_R (-1) (Tacinterp.tactic_of_value ist t)) ]
+| [ "psatz_R" int_or_var(i) tactic(t) ] -> { (Coq_micromega.psatz_R i (Tacinterp.tactic_of_value ist t)) }
+| [ "psatz_R" tactic(t) ] -> { (Coq_micromega.psatz_R (-1) (Tacinterp.tactic_of_value ist t)) }
END
TACTIC EXTEND PsatzQ
-| [ "psatz_Q" int_or_var(i) tactic(t) ] -> [ (Coq_micromega.psatz_Q i (Tacinterp.tactic_of_value ist t)) ]
-| [ "psatz_Q" tactic(t) ] -> [ (Coq_micromega.psatz_Q (-1) (Tacinterp.tactic_of_value ist t)) ]
+| [ "psatz_Q" int_or_var(i) tactic(t) ] -> { (Coq_micromega.psatz_Q i (Tacinterp.tactic_of_value ist t)) }
+| [ "psatz_Q" tactic(t) ] -> { (Coq_micromega.psatz_Q (-1) (Tacinterp.tactic_of_value ist t)) }
END
diff --git a/plugins/nsatz/g_nsatz.ml4 b/plugins/nsatz/g_nsatz.mlg
index 4ac49adb9..16ff512e8 100644
--- a/plugins/nsatz/g_nsatz.ml4
+++ b/plugins/nsatz/g_nsatz.mlg
@@ -8,11 +8,15 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
+{
+
open Ltac_plugin
open Stdarg
+}
+
DECLARE PLUGIN "nsatz_plugin"
TACTIC EXTEND nsatz_compute
-| [ "nsatz_compute" constr(lt) ] -> [ Nsatz.nsatz_compute (EConstr.Unsafe.to_constr lt) ]
+| [ "nsatz_compute" constr(lt) ] -> { Nsatz.nsatz_compute (EConstr.Unsafe.to_constr lt) }
END
diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml
index 6f4138828..e14c4e2ec 100644
--- a/plugins/omega/coq_omega.ml
+++ b/plugins/omega/coq_omega.ml
@@ -38,15 +38,9 @@ open OmegaSolver
(* Added by JCF, 09/03/98 *)
-let elim_id id =
- Proofview.Goal.enter begin fun gl ->
- simplest_elim (mkVar id)
- end
-let resolve_id id = Proofview.Goal.enter begin fun gl ->
- apply (mkVar id)
-end
+let elim_id id = simplest_elim (mkVar id)
-let timing timer_name f arg = f arg
+let resolve_id id = apply (mkVar id)
let display_time_flag = ref false
let display_system_flag = ref false
diff --git a/plugins/omega/g_omega.ml4 b/plugins/omega/g_omega.mlg
index 170b937c9..c3d063cff 100644
--- a/plugins/omega/g_omega.ml4
+++ b/plugins/omega/g_omega.mlg
@@ -18,6 +18,8 @@
DECLARE PLUGIN "omega_plugin"
+{
+
open Ltac_plugin
open Names
open Coq_omega
@@ -43,14 +45,15 @@ let omega_tactic l =
(Tacticals.New.tclREPEAT (Tacticals.New.tclTHENLIST tacs))
(omega_solver)
+}
TACTIC EXTEND omega
-| [ "omega" ] -> [ omega_tactic [] ]
+| [ "omega" ] -> { omega_tactic [] }
END
TACTIC EXTEND omega'
| [ "omega" "with" ne_ident_list(l) ] ->
- [ omega_tactic (List.map Names.Id.to_string l) ]
-| [ "omega" "with" "*" ] -> [ omega_tactic ["nat";"positive";"N";"Z"] ]
+ { omega_tactic (List.map Names.Id.to_string l) }
+| [ "omega" "with" "*" ] -> { omega_tactic ["nat";"positive";"N";"Z"] }
END
diff --git a/plugins/quote/g_quote.ml4 b/plugins/quote/g_quote.mlg
index 09209dc22..749903c3a 100644
--- a/plugins/quote/g_quote.ml4
+++ b/plugins/quote/g_quote.mlg
@@ -8,6 +8,8 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
+{
+
open Ltac_plugin
open Names
open Tacexpr
@@ -16,8 +18,12 @@ open Quote
open Stdarg
open Tacarg
+}
+
DECLARE PLUGIN "quote_plugin"
+{
+
let cont = Id.of_string "cont"
let x = Id.of_string "x"
@@ -27,12 +33,14 @@ let make_cont (k : Val.t) (c : EConstr.t) =
let ist = { lfun = Id.Map.add cont k (Id.Map.singleton x c); extra = TacStore.empty; } in
Tacinterp.eval_tactic_ist ist (TacArg (Loc.tag tac))
+}
+
TACTIC EXTEND quote
- [ "quote" ident(f) ] -> [ quote f [] ]
-| [ "quote" ident(f) "[" ne_ident_list(lc) "]"] -> [ quote f lc ]
+| [ "quote" ident(f) ] -> { quote f [] }
+| [ "quote" ident(f) "[" ne_ident_list(lc) "]"] -> { quote f lc }
| [ "quote" ident(f) "in" constr(c) "using" tactic(k) ] ->
- [ gen_quote (make_cont k) c f [] ]
+ { gen_quote (make_cont k) c f [] }
| [ "quote" ident(f) "[" ne_ident_list(lc) "]"
"in" constr(c) "using" tactic(k) ] ->
- [ gen_quote (make_cont k) c f lc ]
+ { gen_quote (make_cont k) c f lc }
END
diff --git a/plugins/romega/g_romega.ml4 b/plugins/romega/g_romega.mlg
index 5b77d08de..c1ce30027 100644
--- a/plugins/romega/g_romega.ml4
+++ b/plugins/romega/g_romega.mlg
@@ -9,6 +9,8 @@
DECLARE PLUGIN "romega_plugin"
+{
+
open Ltac_plugin
open Names
open Refl_omega
@@ -39,13 +41,15 @@ let romega_tactic unsafe l =
(Tactics.intros)
(total_reflexive_omega_tactic unsafe))
+}
+
TACTIC EXTEND romega
-| [ "romega" ] -> [ romega_tactic false [] ]
-| [ "unsafe_romega" ] -> [ romega_tactic true [] ]
+| [ "romega" ] -> { romega_tactic false [] }
+| [ "unsafe_romega" ] -> { romega_tactic true [] }
END
TACTIC EXTEND romega'
| [ "romega" "with" ne_ident_list(l) ] ->
- [ romega_tactic false (List.map Names.Id.to_string l) ]
-| [ "romega" "with" "*" ] -> [ romega_tactic false ["nat";"positive";"N";"Z"] ]
+ { romega_tactic false (List.map Names.Id.to_string l) }
+| [ "romega" "with" "*" ] -> { romega_tactic false ["nat";"positive";"N";"Z"] }
END
diff --git a/plugins/rtauto/g_rtauto.ml4 b/plugins/rtauto/g_rtauto.mlg
index aa6757634..9c9fdcfa2 100644
--- a/plugins/rtauto/g_rtauto.ml4
+++ b/plugins/rtauto/g_rtauto.mlg
@@ -8,12 +8,15 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
+{
open Ltac_plugin
+}
+
DECLARE PLUGIN "rtauto_plugin"
TACTIC EXTEND rtauto
- [ "rtauto" ] -> [ Proofview.V82.tactic (Refl_tauto.rtauto_tac) ]
+| [ "rtauto" ] -> { Proofview.V82.tactic (Refl_tauto.rtauto_tac) }
END
diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml
index ba193da60..4dfa789ba 100644
--- a/pretyping/glob_ops.ml
+++ b/pretyping/glob_ops.ml
@@ -112,7 +112,7 @@ let fix_kind_eq f k1 k2 = match k1, k2 with
let eq (i1, o1) (i2, o2) =
Option.equal Int.equal i1 i2 && fix_recursion_order_eq f o1 o2
in
- Int.equal i1 i2 && Array.equal eq a1 a1
+ Int.equal i1 i2 && Array.equal eq a1 a2
| GCoFix i1, GCoFix i2 -> Int.equal i1 i2
| (GFix _ | GCoFix _), _ -> false
@@ -452,7 +452,7 @@ let rec rename_glob_vars l c = force @@ DAst.map_with_loc (fun ?loc -> function
else r
| GProd (na,bk,t,c) ->
let na',l' = update_subst na l in
- GProd (na,bk,rename_glob_vars l t,rename_glob_vars l' c)
+ GProd (na',bk,rename_glob_vars l t,rename_glob_vars l' c)
| GLambda (na,bk,t,c) ->
let na',l' = update_subst na l in
GLambda (na',bk,rename_glob_vars l t,rename_glob_vars l' c)
diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml
index 4ab932723..551cc67b6 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -86,7 +86,7 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind =
if not (Sorts.List.mem kind (elim_sorts specif)) then
raise
(RecursionSchemeError
- (NotAllowedCaseAnalysis (false, fst (UnivGen.fresh_sort_in_family env kind), pind)))
+ (NotAllowedCaseAnalysis (false, fst (UnivGen.fresh_sort_in_family kind), pind)))
in
let ndepar = mip.mind_nrealdecls + 1 in
@@ -136,7 +136,7 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind =
mkLambda_string "f" t
(add_branch (push_rel (LocalAssum (Anonymous, t)) env) (k+1))
in
- let (sigma, s) = Evd.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg env sigma kind in
+ let (sigma, s) = Evd.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg sigma kind in
let typP = make_arity env' sigma dep indf s in
let typP = EConstr.Unsafe.to_constr typP in
let c =
@@ -455,7 +455,7 @@ let mis_make_indrec env sigma ?(force_mutual=false) listdepkind mib u =
| ((indi,u),_,_,dep,kinds)::rest ->
let indf = make_ind_family ((indi,u), Context.Rel.to_extended_list mkRel i lnamesparrec) in
let s =
- Evarutil.evd_comb1 (Evd.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg env)
+ Evarutil.evd_comb1 (Evd.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg)
evdref kinds
in
let typP = make_arity env !evdref dep indf s in
@@ -550,8 +550,7 @@ let check_arities env listdepkind =
let kelim = elim_sorts (mibi,mipi) in
if not (Sorts.List.mem kind kelim) then raise
(RecursionSchemeError
- (NotAllowedCaseAnalysis (true, fst (UnivGen.fresh_sort_in_family env
- kind),(mind,u))))
+ (NotAllowedCaseAnalysis (true, fst (UnivGen.fresh_sort_in_family kind),(mind,u))))
else if Int.List.mem ni ln then raise
(RecursionSchemeError (NotMutualInScheme (mind,mind)))
else ni::ln)
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index a66ecaaac..ca2702d74 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -138,7 +138,7 @@ let is_correct_arity env sigma c pj ind specif params =
then error ()
else sigma
| Evar (ev,_), [] ->
- let sigma, s = Evd.fresh_sort_in_family env sigma (max_sort allowed_sorts) in
+ let sigma, s = Evd.fresh_sort_in_family sigma (max_sort allowed_sorts) in
let sigma = Evd.define ev (mkSort s) sigma in
sigma
| _, (LocalDef _ as d)::ar' ->
diff --git a/printing/pputils.ml b/printing/pputils.ml
index c6b8d5022..59e5f68f2 100644
--- a/printing/pputils.ml
+++ b/printing/pputils.ml
@@ -68,7 +68,7 @@ let pr_short_red_flag pr r =
let pr_red_flag pr r =
try pr_short_red_flag pr r
- with complexRedFlags ->
+ with ComplexRedFlag ->
(if r.rBeta then pr_arg str "beta" else mt ()) ++
(if r.rMatch && r.rFix && r.rCofix then pr_arg str "iota" else
(if r.rMatch then pr_arg str "match" else mt ()) ++
diff --git a/shell.nix b/shell.nix
new file mode 100644
index 000000000..45070b2ba
--- /dev/null
+++ b/shell.nix
@@ -0,0 +1,4 @@
+# Some developers don't want a pinned nix-shell by default.
+# If you want to use the pin nix-shell or a more sophisticated set of arguments:
+# $ nix-shell default.nix
+import ./default.nix { pkgs = import <nixpkgs> {}; }
diff --git a/tactics/elimschemes.ml b/tactics/elimschemes.ml
index 70f73df5c..3b69d9922 100644
--- a/tactics/elimschemes.ml
+++ b/tactics/elimschemes.ml
@@ -44,7 +44,7 @@ let optimize_non_type_induction_scheme kind dep sort _ ind =
mib.mind_nparams_rec
else
mib.mind_nparams in
- let sigma, sort = Evd.fresh_sort_in_family env sigma sort in
+ let sigma, sort = Evd.fresh_sort_in_family sigma sort in
let sigma, t', c' = weaken_sort_scheme env sigma false sort npars c t in
let sigma = Evd.minimize_universes sigma in
(Evarutil.nf_evars_universes sigma c', Evd.evar_universe_context sigma), eff
diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml
index ad5239116..ea5ff4a6c 100644
--- a/tactics/eqschemes.ml
+++ b/tactics/eqschemes.ml
@@ -397,7 +397,7 @@ let build_l2r_rew_scheme dep env ind kind =
rel_vect (nrealargs+4) nrealargs;
rel_vect 1 nrealargs;
[|mkRel 1|]]) in
- let s, ctx' = UnivGen.fresh_sort_in_family (Global.env ()) kind in
+ let s, ctx' = UnivGen.fresh_sort_in_family kind in
let ctx = Univ.ContextSet.union ctx ctx' in
let s = mkSort s in
let ci = make_case_info (Global.env()) ind RegularStyle in
@@ -500,7 +500,7 @@ let build_l2r_forward_rew_scheme dep env ind kind =
name_context env ((LocalAssum (Name varH,applied_ind))::realsign) in
let realsign_ind_P n aP =
name_context env ((LocalAssum (Name varH,aP))::realsign_P n) in
- let s, ctx' = UnivGen.fresh_sort_in_family (Global.env ()) kind in
+ let s, ctx' = UnivGen.fresh_sort_in_family kind in
let ctx = Univ.ContextSet.union ctx ctx' in
let s = mkSort s in
let ci = make_case_info (Global.env()) ind RegularStyle in
@@ -578,7 +578,7 @@ let build_r2l_forward_rew_scheme dep env ind kind =
let applied_ind = build_dependent_inductive indu specif in
let realsign_ind =
name_context env ((LocalAssum (Name varH,applied_ind))::realsign) in
- let s, ctx' = UnivGen.fresh_sort_in_family (Global.env ()) kind in
+ let s, ctx' = UnivGen.fresh_sort_in_family kind in
let ctx = Univ.ContextSet.union ctx ctx' in
let s = mkSort s in
let ci = make_case_info (Global.env()) ind RegularStyle in
diff --git a/tactics/inv.ml b/tactics/inv.ml
index e467eacd9..43786c8e1 100644
--- a/tactics/inv.ml
+++ b/tactics/inv.ml
@@ -94,7 +94,7 @@ let make_inv_predicate env evd indf realargs id status concl =
| Some concl -> concl (*assumed it's some [x1..xn,H:I(x1..xn)]C*)
| None ->
let sort = get_sort_family_of env !evd concl in
- let sort = Evarutil.evd_comb1 (Evd.fresh_sort_in_family env) evd sort in
+ let sort = Evarutil.evd_comb1 Evd.fresh_sort_in_family evd sort in
let p = make_arity env !evd true indf sort in
let evd',(p,ptyp) = Unification.abstract_list_all env
!evd p concl (realargs@[mkVar id])
diff --git a/tactics/leminv.ml b/tactics/leminv.ml
index 10937322e..caf4c1eca 100644
--- a/tactics/leminv.ml
+++ b/tactics/leminv.ml
@@ -251,7 +251,7 @@ let add_inversion_lemma_exn ~poly na com comsort bool tac =
let env = Global.env () in
let sigma = Evd.from_env env in
let sigma, c = Constrintern.interp_type_evars env sigma com in
- let sigma, sort = Evd.fresh_sort_in_family ~rigid:univ_rigid env sigma comsort in
+ let sigma, sort = Evd.fresh_sort_in_family ~rigid:univ_rigid sigma comsort in
try
add_inversion_lemma ~poly na env sigma c sort bool tac
with
diff --git a/tools/CoqMakefile.in b/tools/CoqMakefile.in
index 8e60d3932..7aa2d0a5a 100644
--- a/tools/CoqMakefile.in
+++ b/tools/CoqMakefile.in
@@ -65,20 +65,20 @@ VERBOSE ?=
# Time the Coq process (set to non empty), and how (see default value)
TIMED?=
TIMECMD?=
-# Use /usr/bin/env time on linux, gtime on Mac OS
+# Use command time on linux, gtime on Mac OS
TIMEFMT?="$* (real: %e, user: %U, sys: %S, mem: %M ko)"
ifneq (,$(TIMED))
-ifeq (0,$(shell /usr/bin/env time -f $(TIMEFMT) true >/dev/null 2>/dev/null; echo $$?))
-STDTIME?=/usr/bin/env time -f $(TIMEFMT)
+ifeq (0,$(shell command time -f $(TIMEFMT) true >/dev/null 2>/dev/null; echo $$?))
+STDTIME?=command time -f $(TIMEFMT)
else
ifeq (0,$(shell gtime -f $(TIMEFMT) true >/dev/null 2>/dev/null; echo $$?))
STDTIME?=gtime -f $(TIMEFMT)
else
-STDTIME?=time
+STDTIME?=command time
endif
endif
else
-STDTIME?=/usr/bin/env time -f $(TIMEFMT)
+STDTIME?=command time -f $(TIMEFMT)
endif
# Coq binaries
diff --git a/tools/coq_makefile.ml b/tools/coq_makefile.ml
index 6f11ee397..ad489da82 100644
--- a/tools/coq_makefile.ml
+++ b/tools/coq_makefile.ml
@@ -218,7 +218,7 @@ let windrive s =
else ""
;;
-let generate_conf_coq_config oc args =
+let generate_conf_coq_config oc =
section oc "Coq configuration.";
let src_dirs = Coq_config.all_src_dirs in
Envars.print_config ~prefix_var_name:"COQMF_" oc src_dirs;
@@ -282,7 +282,7 @@ let generate_conf oc project args =
fprintf oc "# %s\n\n" (String.concat " " (List.map quote args));
generate_conf_files oc project;
generate_conf_includes oc project;
- generate_conf_coq_config oc args;
+ generate_conf_coq_config oc;
generate_conf_defs oc project;
generate_conf_doc oc project;
generate_conf_extra_target oc project.extra_targets;
diff --git a/tools/coqdoc/index.ml b/tools/coqdoc/index.ml
index df493fdf7..885324aa0 100644
--- a/tools/coqdoc/index.ml
+++ b/tools/coqdoc/index.ml
@@ -77,7 +77,7 @@ let add_ref m loc m' sp id ty =
let find m l = Hashtbl.find reftable (m, l)
-let find_string m s = let (m,s,t) = Hashtbl.find byidtable s in Ref (m,s,t)
+let find_string s = let (m,s,t) = Hashtbl.find byidtable s in Ref (m,s,t)
(* Coq modules *)
diff --git a/tools/coqdoc/index.mli b/tools/coqdoc/index.mli
index 5cd301389..7c9aad67f 100644
--- a/tools/coqdoc/index.mli
+++ b/tools/coqdoc/index.mli
@@ -41,7 +41,7 @@ type index_entry =
val find : coq_module -> loc -> index_entry
(* Find what data is referred to by some string in some coq module *)
-val find_string : coq_module -> string -> index_entry
+val find_string : string -> index_entry
val add_module : coq_module -> unit
diff --git a/tools/coqdoc/output.ml b/tools/coqdoc/output.ml
index d25227002..c640167ac 100644
--- a/tools/coqdoc/output.ml
+++ b/tools/coqdoc/output.ml
@@ -431,7 +431,7 @@ module Latex = struct
else if !Cdglobals.interpolate && !in_doc (* always a var otherwise *)
then
try
- let tag = Index.find_string (get_module false) s in
+ let tag = Index.find_string s in
reference (translate s) tag
with _ -> Tokens.output_tagged_ident_string s
else Tokens.output_tagged_ident_string s
@@ -706,7 +706,7 @@ module Html = struct
else if is_keyword s then
printf "<span class=\"id\" title=\"keyword\">%s</span>" (translate s)
else if !Cdglobals.interpolate && !in_doc (* always a var otherwise *) then
- try reference (translate s) (Index.find_string (get_module false) s)
+ try reference (translate s) (Index.find_string s)
with Not_found -> Tokens.output_tagged_ident_string s
else
Tokens.output_tagged_ident_string s