aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--CHANGES9
-rw-r--r--CONTRIBUTING.md2
-rw-r--r--checker/subtyping.ml8
-rw-r--r--checker/univ.ml7
-rw-r--r--checker/univ.mli2
-rw-r--r--clib/cList.ml14
-rw-r--r--clib/cList.mli8
-rw-r--r--configure.ml16
-rwxr-xr-xdev/tools/check-eof-newline.sh31
-rwxr-xr-xdev/tools/pre-commit73
-rw-r--r--doc/common/macros.tex1
-rw-r--r--doc/refman/RefMan-lib.tex1
-rw-r--r--doc/refman/RefMan-syn.tex398
-rw-r--r--interp/constrexpr_ops.ml140
-rw-r--r--interp/constrexpr_ops.mli7
-rw-r--r--interp/constrextern.ml115
-rw-r--r--interp/constrintern.ml588
-rw-r--r--interp/constrintern.mli5
-rw-r--r--interp/implicit_quantifiers.ml4
-rw-r--r--interp/notation.ml110
-rw-r--r--interp/notation.mli12
-rw-r--r--interp/notation_ops.ml841
-rw-r--r--interp/notation_ops.mli8
-rw-r--r--interp/ppextend.ml1
-rw-r--r--interp/ppextend.mli1
-rw-r--r--interp/reserve.ml2
-rw-r--r--intf/constrexpr.ml12
-rw-r--r--intf/extend.ml49
-rw-r--r--intf/glob_term.ml2
-rw-r--r--intf/intf.mllib2
-rw-r--r--intf/notation_term.ml24
-rw-r--r--intf/vernacexpr.ml1
-rw-r--r--kernel/subtyping.ml9
-rw-r--r--kernel/univ.ml7
-rw-r--r--kernel/univ.mli3
-rw-r--r--parsing/egramcoq.ml62
-rw-r--r--parsing/g_constr.ml422
-rw-r--r--parsing/g_vernac.ml415
-rw-r--r--parsing/pcoq.ml12
-rw-r--r--parsing/pcoq.mli7
-rw-r--r--plugins/funind/glob_term_to_relation.ml8
-rw-r--r--plugins/funind/indfun.ml34
-rw-r--r--plugins/ltac/g_tactic.ml46
-rw-r--r--plugins/ltac/pptactic.ml7
-rw-r--r--plugins/ltac/tacentries.ml3
-rw-r--r--plugins/ssr/ssrcommon.ml6
-rw-r--r--plugins/ssr/ssrparser.ml422
-rw-r--r--plugins/ssr/ssrvernac.ml44
-rw-r--r--plugins/ssrmatching/ssrmatching.ml410
-rw-r--r--pretyping/cases.ml5
-rw-r--r--pretyping/glob_ops.ml38
-rw-r--r--pretyping/glob_ops.mli10
-rw-r--r--printing/ppconstr.ml91
-rw-r--r--printing/ppconstr.mli8
-rw-r--r--printing/ppvernac.ml40
-rw-r--r--printing/ppvernac.mli2
-rw-r--r--test-suite/bugs/closed/5532.v15
-rw-r--r--test-suite/bugs/closed/6529.v16
-rw-r--r--test-suite/modules/cumpoly.v19
-rw-r--r--test-suite/output/Notations.out2
-rw-r--r--test-suite/output/Notations2.out7
-rw-r--r--test-suite/output/Notations2.v5
-rw-r--r--test-suite/output/Notations3.out97
-rw-r--r--test-suite/output/Notations3.v167
-rw-r--r--test-suite/output/PatternsInBinders.out8
-rw-r--r--test-suite/output/PatternsInBinders.v5
-rw-r--r--test-suite/success/Notations2.v29
-rw-r--r--theories/Init/Logic.v7
-rw-r--r--theories/Init/Notations.v27
-rw-r--r--theories/Init/Specif.v9
-rw-r--r--theories/Unicode/Utf8_core.v15
-rw-r--r--vernac/metasyntax.ml360
-rw-r--r--vernac/metasyntax.mli4
-rw-r--r--vernac/vernacentries.ml8
74 files changed, 2429 insertions, 1296 deletions
diff --git a/CHANGES b/CHANGES
index 8d344d4f0..32a533d14 100644
--- a/CHANGES
+++ b/CHANGES
@@ -11,6 +11,15 @@ Notations
priority is given to latest notations defined in the scopes being
opened rather than to the latest notations defined independently of
whether they are in an opened scope or not.
+- Notations can now refer to the syntactic category of patterns (as in
+ "fun 'pat =>" or "match p with pat => ... end"). Two variants are
+ available, depending on whether a single variable is considered as a
+ pattern or not.
+- Recursive notations now support ".." patterns with several
+ occurrences of the recursive term or binder, possibly mixing terms
+ and binders, possibly in reverse left-to-right order.
+- "Locate" now working also on notations of the form "x + y" (rather
+ than "_ + _").
Specification language
diff --git a/CONTRIBUTING.md b/CONTRIBUTING.md
index 1a769333c..74a138f14 100644
--- a/CONTRIBUTING.md
+++ b/CONTRIBUTING.md
@@ -32,6 +32,8 @@ If your pull request fixes a bug, please consider adding a regression test as we
Don't be alarmed if the pull request process takes some time. It can take a few days to get feedback, approval on the final changes, and then a merge. Coq doesn't release new versions very frequently so it can take a few months for your change to land in a released version. That said, you can start using the latest Coq `master` branch to take advantage of all the new features, improvements, and fixes.
+Whitespace discipline (do not indent using tabs, no trailing spaces, text files end with newlines) is checked by Travis (using `git diff --check`). We ship a [`dev/tools/pre-commit`](/dev/tools/pre-commit) git hook which fixes these errors at commit time. `configure` automatically sets you up to use it, unless you already have a hook at `.git/hooks/pre-commit`.
+
Here are a few tags Coq developers may add to your PR and what they mean. In general feedback and requests for you as the pull request author will be in the comments and tags are only used to organize pull requests.
- [needs: rebase](https://github.com/coq/coq/pulls?utf8=%E2%9C%93&q=is%3Aopen%20is%3Apr%20label%3A%22needs%3A%20rebase%22%20) indicates the PR should be rebased on top of the latest `master` branch. See the [GitHub documentation](https://help.github.com/articles/about-git-rebase/) for a brief introduction to using `git rebase`.
diff --git a/checker/subtyping.ml b/checker/subtyping.ml
index 98a9c8250..77201c25b 100644
--- a/checker/subtyping.ml
+++ b/checker/subtyping.ml
@@ -108,6 +108,14 @@ let check_inductive env mp1 l info1 mib2 spec2 subst1 subst2=
let env = check_polymorphic_instance error env auctx auctx' in
env, Univ.make_abstract_instance auctx'
| Cumulative_ind cumi, Cumulative_ind cumi' ->
+ (** Currently there is no way to control variance of inductive types, but
+ just in case we require that they are in a subtyping relation. *)
+ let () =
+ let v = Univ.ACumulativityInfo.variance cumi in
+ let v' = Univ.ACumulativityInfo.variance cumi' in
+ if not (Array.for_all2 Univ.Variance.check_subtype v' v) then
+ CErrors.anomaly Pp.(str "Variance mismatch for " ++ MutInd.print kn)
+ in
let auctx = Univ.ACumulativityInfo.univ_context cumi in
let auctx' = Univ.ACumulativityInfo.univ_context cumi' in
let env = check_polymorphic_instance error env auctx auctx' in
diff --git a/checker/univ.ml b/checker/univ.ml
index 46b3ce680..ebc37bc10 100644
--- a/checker/univ.ml
+++ b/checker/univ.ml
@@ -1011,6 +1011,13 @@ struct
A'] as opposed to [A' <= A]. *)
type t = Irrelevant | Covariant | Invariant
+ let check_subtype x y = match x, y with
+ | (Irrelevant | Covariant | Invariant), Irrelevant -> true
+ | Irrelevant, Covariant -> false
+ | (Covariant | Invariant), Covariant -> true
+ | (Irrelevant | Covariant), Invariant -> false
+ | Invariant, Invariant -> true
+
let leq_constraint csts variance u u' =
match variance with
| Irrelevant -> csts
diff --git a/checker/univ.mli b/checker/univ.mli
index 8c0685e0b..32e48f593 100644
--- a/checker/univ.mli
+++ b/checker/univ.mli
@@ -225,7 +225,7 @@ sig
case because [forall x : A, B <= forall x : A', B'] requires [A =
A'] as opposed to [A' <= A]. *)
type t = Irrelevant | Covariant | Invariant
-
+ val check_subtype : t -> t -> bool
val leq_constraints : t array -> Instance.t constraint_function
val eq_constraints : t array -> Instance.t constraint_function
end
diff --git a/clib/cList.ml b/clib/cList.ml
index 0ef7c3d8b..627a3e3e0 100644
--- a/clib/cList.ml
+++ b/clib/cList.ml
@@ -62,6 +62,7 @@ sig
val fold_right_and_left :
('a -> 'b -> 'b list -> 'a) -> 'b list -> 'a -> 'a
val fold_left3 : ('a -> 'b -> 'c -> 'd -> 'a) -> 'a -> 'b list -> 'c list -> 'd list -> 'a
+ val fold_left2_set : exn -> ('a -> 'b -> 'c -> 'b list -> 'c list -> 'a) -> 'a -> 'b list -> 'c list -> 'a
val for_all_i : (int -> 'a -> bool) -> int -> 'a list -> bool
val except : 'a eq -> 'a -> 'a list -> 'a list
val remove : 'a eq -> 'a -> 'a list -> 'a list
@@ -472,6 +473,19 @@ let fold_right_and_left f l hd =
| a::l -> let hd = aux (a::tl) l in f hd a tl
in aux [] l
+(* Match sets as lists according to a matching function, also folding a side effect *)
+let rec fold_left2_set e f x l1 l2 =
+ match l1 with
+ | a1::l1 ->
+ let rec find seen = function
+ | [] -> raise e
+ | a2::l2 ->
+ try fold_left2_set e f (f x a1 a2 l1 l2) l1 (List.rev_append seen l2)
+ with e' when e' = e -> find (a2::seen) l2 in
+ find [] l2
+ | [] ->
+ if l2 = [] then x else raise e
+
let iteri f l = fold_left_i (fun i _ x -> f i x) 0 () l
let for_all_i p =
diff --git a/clib/cList.mli b/clib/cList.mli
index f87db04cf..b3ee28548 100644
--- a/clib/cList.mli
+++ b/clib/cList.mli
@@ -121,6 +121,14 @@ sig
val fold_right_and_left :
('a -> 'b -> 'b list -> 'a) -> 'b list -> 'a -> 'a
val fold_left3 : ('a -> 'b -> 'c -> 'd -> 'a) -> 'a -> 'b list -> 'c list -> 'd list -> 'a
+
+ (** Fold sets, i.e. lists up to order; the folding function tells
+ when elements match by returning a value and raising the given
+ exception otherwise; sets should have the same size; raise the
+ given exception if no pairing of the two sets is found;;
+ complexity in O(n^2) *)
+ val fold_left2_set : exn -> ('a -> 'b -> 'c -> 'b list -> 'c list -> 'a) -> 'a -> 'b list -> 'c list -> 'a
+
val for_all_i : (int -> 'a -> bool) -> int -> 'a list -> bool
val except : 'a eq -> 'a -> 'a list -> 'a list
val remove : 'a eq -> 'a -> 'a list -> 'a list
diff --git a/configure.ml b/configure.ml
index 9992e03b8..69db9407a 100644
--- a/configure.ml
+++ b/configure.ml
@@ -451,6 +451,22 @@ let vcs =
else if dir_exists "{arch}" then "gnuarch"
else "none"
+(** * Git Precommit Hook *)
+let _ =
+ let f = ".git/hooks/pre-commit" in
+ if vcs = "git" && dir_exists ".git/hooks" && not (Sys.file_exists f) then begin
+ printf "Creating pre-commit hook in %s\n" f;
+ let o = open_out f in
+ let pr s = fprintf o s in
+ pr "#!/bin/sh\n";
+ pr "\n";
+ pr "if [ -x dev/tools/pre-commit ]; then\n";
+ pr " exec dev/tools/pre-commit\n";
+ pr "fi\n";
+ close_out o;
+ Unix.chmod f 0o775
+ end
+
(** * Browser command *)
let browser =
diff --git a/dev/tools/check-eof-newline.sh b/dev/tools/check-eof-newline.sh
index 9e4c8661d..e244d9ab8 100755
--- a/dev/tools/check-eof-newline.sh
+++ b/dev/tools/check-eof-newline.sh
@@ -1,13 +1,40 @@
#!/usr/bin/env bash
+# Usage: check-eof-newline.sh [--fix] FILES...
+# Detect missing end of file newlines for FILES.
+# Files are skipped if untracked by git and depending on gitattributes.
+# With --fix, automatically append a newline.
+# Exit status:
+# Without --fix: 1 if any file had a missing newline, 0 otherwise.
+# With --fix: 1 if any non writable file had a missing newline, 0 otherwise.
+
+FIX=
+if [ "$1" = --fix ];
+then
+ FIX=1
+ shift
+fi
+
CODE=0
for f in "$@"; do
if git ls-files --error-unmatch "$f" >/dev/null 2>&1 && \
git check-attr whitespace -- "$f" | grep -q -v -e 'unset$' -e 'unspecified$' && \
[ -n "$(tail -c 1 "$f")" ]
then
- echo "No newline at end of file $f!"
- CODE=1
+ if [ -n "$FIX" ];
+ then
+ if [ -w "$f" ];
+ then
+ echo >> "$f"
+ echo "Newline appended to file $f!"
+ else
+ echo "File $f is missing a newline and not writable!"
+ CODE=1
+ fi
+ else
+ echo "No newline at end of file $f!"
+ CODE=1
+ fi
fi
done
diff --git a/dev/tools/pre-commit b/dev/tools/pre-commit
new file mode 100755
index 000000000..c9cdee84a
--- /dev/null
+++ b/dev/tools/pre-commit
@@ -0,0 +1,73 @@
+#!/bin/sh
+
+# configure automatically sets up a wrapper at .git/hooks/pre-commit
+# which calls this script (if it exists).
+
+set -e
+
+if ! git diff --cached --name-only -z | xargs -0 dev/tools/check-eof-newline.sh ||
+ ! git diff-index --check --cached HEAD >/dev/null 2>&1 ;
+then
+ 1>&2 echo "Auto fixing whitespace issues..."
+
+ # We fix whitespace in the index and in the working tree
+ # separately to preserve non-added changes.
+ index=$(mktemp "git-fix-ws-index.XXXXX")
+ fixed_index=$(mktemp "git-fix-ws-index-fixed.XXXXX")
+ tree=$(mktemp "git-fix-ws-tree.XXXXX")
+ 1>&2 echo "Patches are saved in '$index', '$fixed_index' and '$tree'."
+ 1>&2 echo "If an error destroys your changes you can recover using them."
+ 1>&2 echo "(The files are cleaned up on success.)"
+ 1>&2 echo #newline
+
+ git diff-index -p --cached HEAD > "$index"
+ git diff-index -p HEAD > "$tree"
+
+ # reset work tree and index
+ # NB: untracked files which were not added are untouched
+ git apply --cached -R "$index"
+ git apply -R "$tree"
+
+ # Fix index
+ # For end of file newlines we must go through the worktree
+ 1>&2 echo "Fixing staged changes..."
+ git apply --cached --whitespace=fix "$index"
+ git apply --whitespace=fix "$index" 2>/dev/null # no need to repeat yourself
+ git diff --cached --name-only -z | xargs -0 dev/tools/check-eof-newline.sh --fix
+ git add -u
+ 1>&2 echo #newline
+
+ # reset work tree
+ git diff-index -p --cached HEAD > "$fixed_index"
+ # If all changes were bad whitespace changes the patch is empty
+ # making git fail. Don't fail now: we fix the worktree first.
+ if [ -s "$fixed_index" ]
+ then
+ git apply -R "$fixed_index"
+ fi
+
+ # Fix worktree
+ 1>&2 echo "Fixing unstaged changes..."
+ git apply --whitespace=fix "$tree"
+ git diff --name-only -z | xargs -0 dev/tools/check-eof-newline.sh --fix
+ 1>&2 echo #newline
+
+ if ! [ -s "$fixed_index" ]
+ then
+ 1>&2 echo "No changes after fixing whitespace issues!"
+ exit 1
+ fi
+
+ # Check that we did fix whitespace
+ if ! git diff-index --check --cached HEAD;
+ then
+ 1>&2 echo "Auto-fixing whitespace failed: errors remain."
+ 1>&2 echo "This may fix itself if you try again."
+ 1>&2 echo "(Consider whether the number of errors decreases after each run.)"
+ exit 1
+ fi
+ 1>&2 echo "Whitespace issues fixed!"
+
+ # clean up temporary files
+ rm "$index" "$tree" "$fixed_index"
+fi
diff --git a/doc/common/macros.tex b/doc/common/macros.tex
index 81def1674..6a28c5b3d 100644
--- a/doc/common/macros.tex
+++ b/doc/common/macros.tex
@@ -182,6 +182,7 @@
\newcommand{\declnotation}{\nterm{decl\_notation}}
\newcommand{\symbolentry}{\nterm{symbol}}
\newcommand{\modifiers}{\nterm{modifiers}}
+\newcommand{\binderinterp}{\nterm{binder\_interp}}
\newcommand{\localdef}{\nterm{local\_def}}
\newcommand{\localdecls}{\nterm{local\_decls}}
\newcommand{\ident}{\nterm{ident}}
diff --git a/doc/refman/RefMan-lib.tex b/doc/refman/RefMan-lib.tex
index c8e844302..89f5be843 100644
--- a/doc/refman/RefMan-lib.tex
+++ b/doc/refman/RefMan-lib.tex
@@ -55,6 +55,7 @@ Figure~\ref{init-notations}.
\hline
Notation & Precedence & Associativity \\
\hline
+\verb!_ -> _! & 99 & right \\
\verb!_ <-> _! & 95 & no \\
\verb!_ \/ _! & 85 & right \\
\verb!_ /\ _! & 80 & right \\
diff --git a/doc/refman/RefMan-syn.tex b/doc/refman/RefMan-syn.tex
index eecb5ac7c..a4ed2134e 100644
--- a/doc/refman/RefMan-syn.tex
+++ b/doc/refman/RefMan-syn.tex
@@ -3,25 +3,32 @@
In this chapter, we introduce advanced commands to modify the way
{\Coq} parses and prints objects, i.e. the translations between the
-concrete and internal representations of terms and commands. The main
-commands are {\tt Notation} and {\tt Infix} which are described in
-section \ref{Notation}. It also happens that the same symbolic
-notation is expected in different contexts. To achieve this form of
-overloading, {\Coq} offers a notion of interpretation scope. This is
-described in Section~\ref{scopes}.
-
-\Rem The commands {\tt Grammar}, {\tt Syntax} and {\tt Distfix} which
-were present for a while in {\Coq} are no longer available from {\Coq}
-version 8.0. The underlying AST structure is also no longer available.
-The functionalities of the command {\tt Syntactic Definition} are
-still available; see Section~\ref{Abbreviations}.
+concrete and internal representations of terms and commands.
+
+The main commands to provide custom symbolic notations for terms are
+{\tt Notation} and {\tt Infix}. They are described in Section
+\ref{Notation}. There is also a variant of {\tt Notation} which does
+not modify the parser. This provides with a form of abbreviation and
+it is described in Section~\ref{Abbreviations}. It is sometimes
+expected that the same symbolic notation has different meanings in
+different contexts. To achieve this form of overloading, {\Coq} offers
+a notion of interpretation scope. This is described in
+Section~\ref{scopes}.
+
+The main command to provide custom notations for tactics is {\tt
+ Tactic Notation}. It is described in Section~\ref{Tactic-Notation}.
+
+% No need any more to remind this
+%% \Rem The commands {\tt Grammar}, {\tt Syntax} and {\tt Distfix} which
+%% were present for a while in {\Coq} are no longer available from {\Coq}
+%% version 8.0. The underlying AST structure is also no longer available.
\section[Notations]{Notations\label{Notation}
\comindex{Notation}}
\subsection{Basic notations}
-A {\em notation} is a symbolic abbreviation denoting some term
+A {\em notation} is a symbolic expression denoting some term
or term pattern.
A typical notation is the use of the infix symbol \verb=/\= to denote
@@ -37,7 +44,7 @@ string \verb="A /\ B"= (called a {\em notation}) tells how it is
symbolically written.
A notation is always surrounded by double quotes (except when the
-abbreviation is a single identifier; see \ref{Abbreviations}). The
+abbreviation has the form of an ordinary applicative expression; see \ref{Abbreviations}). The
notation is composed of {\em tokens} separated by spaces. Identifiers
in the string (such as \texttt{A} and \texttt{B}) are the {\em
parameters} of the notation. They must occur at least once each in the
@@ -61,7 +68,7 @@ syntactic expression (see \ref{ReservedNotation}), explicit precedences and
associativity rules have to be given.
\Rem The right-hand side of a notation is interpreted at the time the
-notation is given. In particular, implicit arguments (see
+notation is given. In particular, disambiguation of constants, implicit arguments (see
Section~\ref{Implicit Arguments}), coercions (see
Section~\ref{Coercions}), etc. are resolved at the time of the
declaration of the notation.
@@ -105,8 +112,8 @@ parentheses are mandatory (this is a ``no associativity'')\footnote{
which {\Coq} is built, namely {\camlpppp}, currently does not implement the
no-associativity and replaces it by a left associativity; hence it is
the same for {\Coq}: no-associativity is in fact left associativity}.
-We don't know of a special convention of the associativity of
-disjunction and conjunction, so let's apply for instance a right
+We do not know of a special convention of the associativity of
+disjunction and conjunction, so let us apply for instance a right
associativity (which is the choice of {\Coq}).
Precedence levels and associativity rules of notations have to be
@@ -142,7 +149,8 @@ Notation "x = y" := (@eq _ x y) (at level 70, no associativity).
\end{coq_example*}
One can define {\em closed} notations whose both sides are symbols. In
-this case, the default precedence level for inner subexpression is 200.
+this case, the default precedence level for inner subexpression is
+200, and the default level for the notation itself is 0.
\begin{coq_eval}
Set Printing Depth 50.
@@ -150,7 +158,7 @@ Set Printing Depth 50.
(**** an incompatibility with the reserved notation ********)
\end{coq_eval}
\begin{coq_example*}
-Notation "( x , y )" := (@pair _ _ x y) (at level 0).
+Notation "( x , y )" := (@pair _ _ x y).
\end{coq_example*}
One can also define notations for binders.
@@ -161,17 +169,17 @@ Set Printing Depth 50.
(**** an incompatibility with the reserved notation ********)
\end{coq_eval}
\begin{coq_example*}
-Notation "{ x : A | P }" := (sig A (fun x => P)) (at level 0).
+Notation "{ x : A | P }" := (sig A (fun x => P)).
\end{coq_example*}
In the last case though, there is a conflict with the notation for
-type casts. This last notation, as shown by the command {\tt Print Grammar
+type casts. The notation for type casts, as shown by the command {\tt Print Grammar
constr} is at level 100. To avoid \verb=x : A= being parsed as a type cast,
it is necessary to put {\tt x} at a level below 100, typically 99. Hence, a
-correct definition is
+correct definition is the following.
\begin{coq_example*}
-Notation "{ x : A | P }" := (sig A (fun x => P)) (at level 0, x at level 99).
+Notation "{ x : A | P }" := (sig A (fun x => P)) (x at level 99).
\end{coq_example*}
%This change has retrospectively an effect on the notation for notation
@@ -182,14 +190,17 @@ Notation "{ x : A | P }" := (sig A (fun x => P)) (at level 0, x at level 99).
%Notation "{ A } + { B }" := (sumbool A B) (at level 0, A at level 99).
%\end{coq_example*}
-See the next section for more about factorization.
+More generally, it is required that notations are explicitly
+factorized on the left. See the next section for more about
+factorization.
\subsection{Simple factorization rules}
-{\Coq} extensible parsing is performed by Camlp5 which is essentially a
-LL1 parser. Hence, some care has to be taken not to hide already
-existing rules by new rules. Some simple left factorization work has
-to be done. Here is an example.
+{\Coq} extensible parsing is performed by {\camlpppp} which is
+essentially a LL1 parser: it decides which notation to parse by
+looking tokens from left to right. Hence, some care has to be taken
+not to hide already existing rules by new rules. Some simple left
+factorization work has to be done. Here is an example.
\begin{coq_eval}
(********** The next rule for notation _ < _ < _ produces **********)
@@ -242,17 +253,19 @@ on the {\Coq} printer. For example:
Check (and True True).
\end{coq_example}
-However, printing, especially pretty-printing, requires
-more care than parsing. We may want specific indentations,
-line breaks, alignment if on several lines, etc.
+However, printing, especially pretty-printing, also requires some
+care. We may want specific indentations, line breaks, alignment if on
+several lines, etc. For pretty-printing, {\Coq} relies on {\ocaml}
+formatting library, which provides indentation and automatic line
+breaks depending on page width by means of {\em formatting boxes}.
-The default printing of notations is very rudimentary. For printing a
-notation, a {\em formatting box} is opened in such a way that if the
+The default printing of notations is rudimentary. For printing a
+notation, a formatting box is opened in such a way that if the
notation and its arguments cannot fit on a single line, a line break
is inserted before the symbols of the notation and the arguments on
the next lines are aligned with the argument on the first line.
-A first, simple control that a user can have on the printing of a
+A first simple control that a user can have on the printing of a
notation is the insertion of spaces at some places of the
notation. This is performed by adding extra spaces between the symbols
and parameters: each extra space (other than the single space needed
@@ -277,6 +290,13 @@ Notation "'If' c1 'then' c2 'else' c3" := (IF_then_else c1 c2 c3)
\end{coq_example}
\end{small}
+\begin{coq_example}
+Check
+ (IF_then_else (IF_then_else True False True)
+ (IF_then_else True False True)
+ (IF_then_else True False True)).
+\end{coq_example}
+
A {\em format} is an extension of the string denoting the notation with
the possible following elements delimited by single quotes:
@@ -313,22 +333,15 @@ Notations do not survive the end of sections. No typing of the denoted
expression is performed at definition time. Type-checking is done only
at the time of use of the notation.
-\begin{coq_example}
-Check
- (IF_then_else (IF_then_else True False True)
- (IF_then_else True False True)
- (IF_then_else True False True)).
-\end{coq_example}
-
\Rem
Sometimes, a notation is expected only for the parser.
%(e.g. because
%the underlying parser of {\Coq}, namely {\camlpppp}, is LL1 and some extra
%rules are needed to circumvent the absence of factorization).
-To do so, the option {\em only parsing} is allowed in the list of modifiers of
+To do so, the option {\tt only parsing} is allowed in the list of modifiers of
\texttt{Notation}.
-Conversely, the {\em only printing} can be used to declare
+Conversely, the {\tt only printing} can be used to declare
that a notation should only be used for printing and should not declare a
parsing rule. In particular, such notations do not modify the parser.
@@ -339,16 +352,16 @@ The \texttt{Infix} command is a shortening for declaring notations of
infix symbols. Its syntax is
\begin{quote}
-\noindent\texttt{Infix "{\symbolentry}" :=} {\qualid} {\tt (} \nelist{\em modifier}{,} {\tt )}.
+\noindent\texttt{Infix "{\symbolentry}" :=} {\term} {\tt (} \nelist{\em modifier}{,} {\tt )}.
\end{quote}
and it is equivalent to
\begin{quote}
-\noindent\texttt{Notation "x {\symbolentry} y" := ({\qualid} x y) (} \nelist{\em modifier}{,} {\tt )}.
+\noindent\texttt{Notation "x {\symbolentry} y" := ({\term} x y) (} \nelist{\em modifier}{,} {\tt )}.
\end{quote}
-where {\tt x} and {\tt y} are fresh names distinct from {\qualid}. Here is an example.
+where {\tt x} and {\tt y} are fresh names. Here is an example.
\begin{coq_example*}
Infix "/\" := and (at level 80, right associativity).
@@ -380,12 +393,14 @@ reserved. Hence their precedence and associativity cannot be changed.
\comindex{CoFixpoint {\ldots} where {\ldots}}
\comindex{Inductive {\ldots} where {\ldots}}}
-Thanks to reserved notations, the inductive, co-inductive, recursive
-and corecursive definitions can benefit of customized notations. To do
-this, insert a {\tt where} notation clause after the definition of the
-(co)inductive type or (co)recursive term (or after the definition of
-each of them in case of mutual definitions). The exact syntax is given
-on Figure~\ref{notation-syntax}. Here are examples:
+Thanks to reserved notations, the inductive, co-inductive, record,
+recursive and corecursive definitions can benefit of customized
+notations. To do this, insert a {\tt where} notation clause after the
+definition of the (co)inductive type or (co)recursive term (or after
+the definition of each of them in case of mutual definitions). The
+exact syntax is given on Figure~\ref{notation-syntax} for inductive,
+co-inductive, recursive and corecursive definitions and on
+Figure~\ref{record-syntax} for records. Here are examples:
\begin{coq_eval}
Set Printing Depth 50.
@@ -479,20 +494,28 @@ Locate "exists _ .. _ , _".
\\
\\
{\modifiers}
- & ::= & \nelist{\ident}{,} {\tt at level} {\naturalnumber} \\
- & $|$ & \nelist{\ident}{,} {\tt at next level} \\
- & $|$ & {\tt at level} {\naturalnumber} \\
- & $|$ & {\tt left associativity} \\
- & $|$ & {\tt right associativity} \\
- & $|$ & {\tt no associativity} \\
+ & ::= & {\tt at level} {\naturalnumber} \\
+ & $|$ & \nelist{\ident}{,} {\tt at level} {\naturalnumber} \zeroone{\binderinterp}\\
+ & $|$ & \nelist{\ident}{,} {\tt at next level} \zeroone{\binderinterp}\\
+ & $|$ & {\ident} {\binderinterp} \\
& $|$ & {\ident} {\tt ident} \\
- & $|$ & {\ident} {\tt binder} \\
- & $|$ & {\ident} {\tt closed binder} \\
& $|$ & {\ident} {\tt global} \\
& $|$ & {\ident} {\tt bigint} \\
+ & $|$ & {\ident} \zeroone{{\tt strict}} {\tt pattern} \zeroone{{\tt at level} {\naturalnumber}}\\
+ & $|$ & {\ident} {\tt binder} \\
+ & $|$ & {\ident} {\tt closed binder} \\
+ & $|$ & {\tt left associativity} \\
+ & $|$ & {\tt right associativity} \\
+ & $|$ & {\tt no associativity} \\
& $|$ & {\tt only parsing} \\
& $|$ & {\tt only printing} \\
- & $|$ & {\tt format} {\str}
+ & $|$ & {\tt format} {\str} \\
+\\
+\\
+{\binderinterp}
+ & ::= & {\tt as ident} \\
+ & $|$ & {\tt as pattern} \\
+ & $|$ & {\tt as strict pattern} \\
\end{tabular}
\end{centerframe}
\end{small}
@@ -500,9 +523,93 @@ Locate "exists _ .. _ , _".
\label{notation-syntax}
\end{figure}
-\subsection{Notations and simple binders}
+\subsection{Notations and binders}
+
+Notations can include binders. This section lists
+different ways to deal with binders. For further examples, see also
+Section~\ref{RecursiveNotationsWithBinders}.
-Notations can be defined for binders as in the example:
+\subsubsection{Binders bound in the notation and parsed as identifiers}
+
+Here is the basic example of a notation using a binder:
+
+\begin{coq_example*}
+Notation "'sigma' x : A , B" := (sigT (fun x : A => B))
+ (at level 200, x ident, A at level 200, right associativity).
+\end{coq_example*}
+
+The binding variables in the right-hand side that occur as a parameter
+of the notation (here {\tt x}) dynamically bind all the occurrences
+in their respective binding scope after instantiation of the
+parameters of the notation. This means that the term bound to {\tt B} can
+refer to the variable name bound to {\tt x} as shown in the following
+application of the notation:
+
+\begin{coq_example}
+Check sigma z : nat, z = 0.
+\end{coq_example}
+
+Notice the modifier {\tt x ident} in the declaration of the
+notation. It tells to parse {\tt x} as a single identifier.
+
+\subsubsection{Binders bound in the notation and parsed as patterns}
+
+In the same way as patterns can be used as binders, as in {\tt fun
+ '(x,y) => x+y} or {\tt fun '(existT \_ x \_) => x}, notations can be
+defined so that any pattern (in the sense of the entry {\pattern} of
+Figure~\ref{term-syntax-aux}) can be used in place of the
+binder. Here is an example:
+
+\begin{coq_eval}
+Reset Initial.
+\end{coq_eval}
+
+\begin{coq_example*}
+Notation "'subset' ' p , P " := (sig (fun p => P))
+ (at level 200, p pattern, format "'subset' ' p , P").
+\end{coq_example*}
+
+\begin{coq_example}
+Check subset '(x,y), x+y=0.
+\end{coq_example}
+
+The modifier {\tt p pattern} in the declaration of the notation
+tells to parse $p$ as a pattern. Note that a single
+variable is both an identifier and a pattern, so, e.g., the following
+also works:
+
+% Note: we rely on the notation of the standard library which does not
+% print the expected output, so we hide the output.
+\begin{coq_example}
+Check subset 'x, x=0.
+\end{coq_example}
+
+If one wants to prevent such a notation to be used for printing when the
+pattern is reduced to a single identifier, one has to use instead
+the modifier {\tt p strict pattern}. For parsing, however, a {\tt
+ strict pattern} will continue to include the case of a
+variable. Here is an example showing the difference:
+
+\begin{coq_example*}
+Notation "'subset_bis' ' p , P" := (sig (fun p => P))
+ (at level 200, p strict pattern).
+Notation "'subset_bis' p , P " := (sig (fun p => P))
+ (at level 200, p ident).
+\end{coq_example*}
+
+\begin{coq_example}
+Check subset_bis 'x, x=0.
+\end{coq_example}
+
+The default level for a {\tt pattern} is 0. One can use a different level by
+using {\tt pattern at level} $n$ where the scale is the same as the one for
+terms (Figure~\ref{init-notations}).
+
+\subsubsection{Binders bound in the notation and parsed as terms}
+
+Sometimes, for the sake of factorization of rules, a binder has to be
+parsed as a term. This is typically the case for a notation such as
+the following:
\begin{coq_eval}
Set Printing Depth 50.
@@ -510,18 +617,53 @@ Set Printing Depth 50.
(**** an incompatibility with the reserved notation ********)
\end{coq_eval}
\begin{coq_example*}
-Notation "{ x : A | P }" := (sig (fun x : A => P)) (at level 0).
+Notation "{ x : A | P }" := (sig (fun x : A => P))
+ (at level 0, x at level 99 as ident).
+\end{coq_example*}
+
+This is so because the grammar also contains rules starting with
+{\tt \{} and followed by a term, such as the rule for the notation
+ {\tt \{ A \} + \{ B \}} for the constant {\tt
+ sumbool}~(see Section~\ref{sumbool}).
+
+Then, in the rule, {\tt x ident} is replaced by {\tt x at level 99 as
+ ident} meaning that {\tt x} is parsed as a term at level 99 (as done
+in the notation for {\tt sumbool}), but that this term has actually to
+be an identifier.
+
+The notation {\tt \{ x | P \}} is already defined in the standard
+library with the {\tt as ident} modifier. We cannot redefine it but
+one can define an alternative notation, say {\tt \{ p such that P }\},
+using instead {\tt as pattern}.
+
+% Note, this conflicts with the default rule in the standard library, so
+% we don't show the
+\begin{coq_example*}
+Notation "{ p 'such' 'that' P }" := (sig (fun p => P))
+ (at level 0, p at level 99 as pattern).
\end{coq_example*}
-The binding variables in the left-hand-side that occur as a parameter
-of the notation naturally bind all their occurrences appearing in
-their respective scope after instantiation of the parameters of the
-notation.
+Then, the following works:
+\begin{coq_example}
+Check {(x,y) such that x+y=0}.
+\end{coq_example}
+
+To enforce that the pattern should not be used for printing when it
+is just an identifier, one could have said {\tt p at level
+ 99 as strict pattern}.
+
+Note also that in the absence of a {\tt as ident}, {\tt as strict
+ pattern} or {\tt as pattern} modifiers, the default is to consider
+subexpressions occurring in binding position and parsed as terms to be
+{\tt as ident}.
-Contrastingly, the binding variables that are not a parameter of the
-notation do not capture the variables of same name that
-could appear in their scope after instantiation of the
-notation. E.g., for the notation
+\subsubsection{Binders not bound in the notation}
+\label{NotationsWithBinders}
+
+We can also have binders in the right-hand side of a notation which
+are not themselves bound in the notation. In this case, the binders
+are considered up to renaming of the internal binder. E.g., for the
+notation
\begin{coq_example*}
Notation "'exists_different' n" := (exists p:nat, p<>n) (at level 200).
@@ -537,14 +679,6 @@ Set Printing Depth 50.
Fail Check (exists_different p).
\end{coq_example}
-\Rem Binding variables must not necessarily be parsed using the
-{\tt ident} entry. For factorization purposes, they can be said to be
-parsed at another level (e.g. {\tt x} in \verb="{ x : A | P }"= must be
-parsed at level 99 to be factorized with the notation
-\verb="{ A } + { B }"= for which {\tt A} can be any term).
-However, even if parsed as a term, this term must at the end be effectively
-a single identifier.
-
\subsection{Notations with recursive patterns}
\label{RecursiveNotations}
@@ -565,24 +699,22 @@ notation parses any number of time (but at least one time) a sequence
of expressions separated by the sequence of tokens $s$ (in the
example, $s$ is just ``{\tt ;}'').
-In the right-hand side, the term enclosed within {\tt ..} must be a
-pattern with two holes of the form $\phi([~]_E,[~]_I)$ where the first
-hole is occupied either by $x$ or by $y$ and the second hole is
-occupied by an arbitrary term $t$ called the {\it terminating}
-expression of the recursive notation. The subterm {\tt ..} $\phi(x,t)$
-{\tt ..} (or {\tt ..} $\phi(y,t)$ {\tt ..}) must itself occur at
-second position of the same pattern where the first hole is occupied
-by the other variable, $y$ or $x$. Otherwise said, the right-hand side
-must contain a subterm of the form either $\phi(x,${\tt ..}
-$\phi(y,t)$ {\tt ..}$)$ or $\phi(y,${\tt ..} $\phi(x,t)$ {\tt ..}$)$.
-The pattern $\phi$ is the {\em iterator} of the recursive notation
-and, of course, the name $x$ and $y$ can be chosen arbitrarily.
-
-The parsing phase produces a list of expressions which are used to
-fill in order the first hole of the iterating pattern which is
+The right-hand side must contain a subterm of the form either
+$\phi(x,${\tt ..} $\phi(y,t)$ {\tt ..}$)$ or $\phi(y,${\tt ..}
+$\phi(x,t)$ {\tt ..}$)$ where $\phi([~]_E,[~]_I)$, called the {\em
+ iterator} of the recursive notation is an arbitrary expression with
+distinguished placeholders and
+where $t$ is called the {\tt terminating expression} of the recursive
+notation. In the example, we choose the name s$x$ and $y$ but in
+practice they can of course be chosen arbitrarily. Note that the
+placeholder $[~]_I$ has to occur only once but the $[~]_E$ can occur
+several times.
+
+Parsing the notation produces a list of expressions which are used to
+fill the first placeholder of the iterating pattern which itself is
repeatedly nested as many times as the length of the list, the second
-hole being the nesting point. In the innermost occurrence of the
-nested iterating pattern, the second hole is finally filled with the
+placeholder being the nesting point. In the innermost occurrence of the
+nested iterating pattern, the second placeholder is finally filled with the
terminating expression.
In the example above, the iterator $\phi([~]_E,[~]_I)$ is {\tt cons
@@ -609,24 +741,26 @@ notations, they can also be declared within interpretation scopes (see
section \ref{scopes}).
\subsection{Notations with recursive patterns involving binders}
+\label{RecursiveNotationsWithBinders}
Recursive notations can also be used with binders. The basic example is:
\begin{coq_example*}
-Notation "'exists' x .. y , p" := (ex (fun x => .. (ex (fun y => p)) ..))
+Notation "'exists' x .. y , p" :=
+ (ex (fun x => .. (ex (fun y => p)) ..))
(at level 200, x binder, y binder, right associativity).
\end{coq_example*}
The principle is the same as in Section~\ref{RecursiveNotations}
-except that in the iterator $\phi([~]_E,[~]_I)$, the first hole is a
-placeholder occurring at the position of the binding variable of a {\tt
+except that in the iterator $\phi([~]_E,[~]_I)$, the placeholder $[~]_E$ can
+also occur in position of the binding variable of a {\tt
fun} or a {\tt forall}.
To specify that the part ``$x$ {\tt ..} $y$'' of the notation
parses a sequence of binders, $x$ and $y$ must be marked as {\tt
- binder} in the list of modifiers of the notation. Then, the list of
-binders produced at the parsing phase are used to fill in the first
-hole of the iterating pattern which is repeatedly nested as many times
+ binder} in the list of modifiers of the notation. The
+binders of the parsed sequence are used to fill the occurrences of the first
+placeholder of the iterating pattern which is repeatedly nested as many times
as the number of binders generated. If ever the generalization
operator {\tt `} (see Section~\ref{implicit-generalization}) is used
in the binding list, the added binders are taken into account too.
@@ -635,14 +769,14 @@ Binders parsing exist in two flavors. If $x$ and $y$ are marked as
{\tt binder}, then a sequence such as {\tt a b c : T} will be accepted
and interpreted as the sequence of binders {\tt (a:T) (b:T)
(c:T)}. For instance, in the notation above, the syntax {\tt exists
- a b : nat, a = b} is provided.
+ a b : nat, a = b} is valid.
The variables $x$ and $y$ can also be marked as {\tt closed binder} in
which case only well-bracketed binders of the form {\tt (a b c:T)} or
{\tt \{a b c:T\}} etc. are accepted.
With closed binders, the recursive sequence in the left-hand side can
-be of the general form $x$ $s$ {\tt ..} $s$ $y$ where $s$ is an
+be of the more general form $x$ $s$ {\tt ..} $s$ $y$ where $s$ is an
arbitrary sequence of tokens. With open binders though, $s$ has to be
empty. Here is an example of recursive notation with closed binders:
@@ -661,6 +795,40 @@ Notation "'FUNAPP' x .. y , f" :=
(at level 200, x binder, y binder, right associativity).
\end{coq_example*}
+If an occurrence of the $[~]_E$ is not in position of a binding
+variable but of a term, it is the name used in the binding which is
+used. Here is an example:
+
+\begin{coq_example*}
+Notation "'exists_non_null' x .. y , P" :=
+ (ex (fun x => x <> 0 /\ .. (ex (fun y => y <> 0 /\ P)) ..))
+ (at level 200, x binder).
+\end{coq_example*}
+
+\subsection{Predefined entries}
+
+By default, sub-expressions are parsed as terms and the corresponding
+grammar entry is called {\tt constr}. However, one may sometimes want
+to restrict the syntax of terms in a notation. For instance, the
+following notation will accept to parse only global reference in
+position of {\tt x}:
+
+\begin{coq_example*}
+Notation "'apply' f a1 .. an" := (.. (f a1) .. an)
+ (at level 10, f global, a1, an at level 9).
+\end{coq_example*}
+
+In addition to {\tt global}, one can restrict the syntax of a
+sub-expression by using the entry names {\tt ident} or {\tt pattern}
+already seen in Section~\ref{NotationsWithBinders}, even when the
+corresponding expression is not used as a binder in the right-hand
+side. E.g.:
+
+\begin{coq_example*}
+Notation "'apply_id' f a1 .. an" := (.. (f a1) .. an)
+ (at level 10, f ident, a1, an at level 9).
+\end{coq_example*}
+
\subsection{Summary}
\paragraph{Syntax of notations}
@@ -754,7 +922,7 @@ stack by using the command
{\tt Close Scope} {\scope}.
\end{quote}
Notice that this command does not only cancel the last {\tt Open Scope
-{\scope}} but all the invocation of it.
+{\scope}} but all the invocations of it.
\Rem {\tt Open Scope} and {\tt Close Scope} do not survive the end of
sections where they occur. When defined outside of a section, they are
@@ -1108,7 +1276,7 @@ Check reflexive iff.
\end{coq_example}
An abbreviation expects no precedence nor associativity, since it
-follows the usual syntax of application. Abbreviations are used as
+is parsed as usual application. Abbreviations are used as
much as possible by the {\Coq} printers unless the modifier
\verb=(only parsing)= is given.
@@ -1121,7 +1289,7 @@ abbreviation but at the time it is used. Especially, abbreviations can
be bound to terms with holes (i.e. with ``\_''). The general syntax
for abbreviations is
\begin{quote}
-\zeroone{{\tt Local}} \texttt{Notation} {\ident} \sequence{\ident} {\ident} \texttt{:=} {\term}
+\zeroone{{\tt Local}} \texttt{Notation} {\ident} \sequence{\ident}{} \texttt{:=} {\term}
\zeroone{{\tt (only parsing)}}~\verb=.=
\end{quote}
@@ -1147,13 +1315,15 @@ at the time of use of the abbreviation.
%\verb=(only parsing)= is given) while syntactic definitions were not.
\section{Tactic Notations
+\label{Tactic-Notation}
\comindex{Tactic Notation}}
Tactic notations allow to customize the syntax of the tactics of the
-tactic language\footnote{Tactic notations are just a simplification of
-the {\tt Grammar tactic simple\_tactic} command that existed in
-versions prior to version 8.0.}. Tactic notations obey the following
-syntax
+tactic language.
+%% \footnote{Tactic notations are just a simplification of
+%% the {\tt Grammar tactic simple\_tactic} command that existed in
+%% versions prior to version 8.0.}
+Tactic notations obey the following syntax:
\medskip
\noindent
@@ -1196,7 +1366,9 @@ level indicates the parsing precedence of the tactic notation. This
information is particularly relevant for notations of tacticals.
Levels 0 to 5 are available (default is 0).
To know the parsing precedences of the
-existing tacticals, use the command {\tt Print Grammar tactic.}
+existing tacticals, use the command
+\comindex{Print Grammar tactic}
+ {\tt Print Grammar tactic.}
Each type of tactic argument has a specific semantic regarding how it
is parsed and how it is interpreted. The semantic is described in the
diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml
index da04d8786..8aca6e333 100644
--- a/interp/constrexpr_ops.ml
+++ b/interp/constrexpr_ops.ml
@@ -67,7 +67,7 @@ let rec cases_pattern_expr_eq p1 p2 =
if CAst.(p1.v == p2.v) then true
else match CAst.(p1.v, p2.v) with
| CPatAlias(a1,i1), CPatAlias(a2,i2) ->
- Id.equal i1 i2 && cases_pattern_expr_eq a1 a2
+ Name.equal (snd i1) (snd i2) && cases_pattern_expr_eq a1 a2
| CPatCstr(c1,a1,b1), CPatCstr(c2,a2,b2) ->
eq_reference c1 c2 &&
Option.equal (List.equal cases_pattern_expr_eq) a1 a2 &&
@@ -112,10 +112,10 @@ let rec constr_expr_eq e1 e2 =
eq_located Id.equal id1 id2 &&
List.equal cofix_expr_eq fl1 fl2
| CProdN(bl1,a1), CProdN(bl2,a2) ->
- List.equal binder_expr_eq bl1 bl2 &&
+ List.equal local_binder_eq bl1 bl2 &&
constr_expr_eq a1 a2
| CLambdaN(bl1,a1), CLambdaN(bl2,a2) ->
- List.equal binder_expr_eq bl1 bl2 &&
+ List.equal local_binder_eq bl1 bl2 &&
constr_expr_eq a1 a2
| CLetIn((_,na1),a1,t1,b1), CLetIn((_,na2),a2,t2,b2) ->
Name.equal na1 na2 &&
@@ -193,10 +193,6 @@ and branch_expr_eq (_, (p1, e1)) (_, (p2, e2)) =
List.equal (List.equal cases_pattern_expr_eq) p1 p2 &&
constr_expr_eq e1 e2
-and binder_expr_eq ((n1, _, e1) : binder_expr) (n2, _, e2) =
- (** Don't care about the [binder_kind] *)
- List.equal (eq_located Name.equal) n1 n2 && constr_expr_eq e1 e2
-
and fix_expr_eq (id1,(j1, r1),bl1,a1,b1) (id2,(j2, r2),bl2,a2,b2) =
(eq_located Id.equal id1 id2) &&
Option.equal (eq_located Id.equal) j1 j2 &&
@@ -226,9 +222,10 @@ and local_binder_eq l1 l2 = match l1, l2 with
List.equal (eq_located Name.equal) n1 n2 && constr_expr_eq e1 e2
| _ -> false
-and constr_notation_substitution_eq (e1, el1, bl1) (e2, el2, bl2) =
+and constr_notation_substitution_eq (e1, el1, b1, bl1) (e2, el2, b2, bl2) =
List.equal constr_expr_eq e1 e2 &&
List.equal (List.equal constr_expr_eq) el1 el2 &&
+ List.equal cases_pattern_expr_eq b1 b2 &&
List.equal (List.equal local_binder_eq) bl1 bl2
and instance_eq (x1,c1) (x2,c2) =
@@ -272,7 +269,7 @@ let is_constructor id =
let rec cases_pattern_fold_names f a pt = match CAst.(pt.v) with
| CPatRecord l ->
List.fold_left (fun acc (r, cp) -> cases_pattern_fold_names f acc cp) a l
- | CPatAlias (pat,id) -> f id a
+ | CPatAlias (pat,(loc,na)) -> Name.fold_right f na (cases_pattern_fold_names f a pat)
| CPatOr (patl) ->
List.fold_left (cases_pattern_fold_names f) a patl
| CPatCstr (_,patl1,patl2) ->
@@ -307,14 +304,6 @@ let ids_of_cases_tomatch tms =
(Option.fold_right (down_located (Name.fold_right Id.Set.add)) ona l))
tms Id.Set.empty
-let rec fold_constr_expr_binders g f n acc b = function
- | (nal,bk,t)::l ->
- let nal = snd (List.split nal) in
- let n' = List.fold_right (Name.fold_right g) nal n in
- f n (fold_constr_expr_binders g f n' acc b l) t
- | [] ->
- f n acc b
-
let rec fold_local_binders g f n acc b = function
| CLocalAssum (nal,bk,t)::l ->
let nal = snd (List.split nal) in
@@ -331,12 +320,12 @@ let rec fold_local_binders g f n acc b = function
let fold_constr_expr_with_binders g f n acc = CAst.with_val (function
| CAppExpl ((_,_,_),l) -> List.fold_left (f n) acc l
| CApp ((_,t),l) -> List.fold_left (f n) (f n acc t) (List.map fst l)
- | CProdN (l,b) | CLambdaN (l,b) -> fold_constr_expr_binders g f n acc b l
+ | CProdN (l,b) | CLambdaN (l,b) -> fold_local_binders g f n acc b l
| CLetIn (na,a,t,b) ->
f (Name.fold_right g (snd na) n) (Option.fold_left (f n) (f n acc a) t) b
| CCast (a,(CastConv b|CastVM b|CastNative b)) -> f n (f n acc a) b
| CCast (a,CastCoerce) -> f n acc a
- | CNotation (_,(l,ll,bll)) ->
+ | CNotation (_,(l,ll,bl,bll)) ->
(* The following is an approximation: we don't know exactly if
an ident is binding nor to which subterms bindings apply *)
let acc = List.fold_left (f n) acc (l@List.flatten ll) in
@@ -382,12 +371,6 @@ let occur_var_constr_expr id c = Id.Set.mem id (free_vars_of_constr_expr c)
(* Used in correctness and interface *)
let map_binder g e nal = List.fold_right (down_located (Name.fold_right g)) nal e
-let map_binders f g e bl =
- (* TODO: avoid variable capture in [t] by some [na] in [List.tl nal] *)
- let h (e,bl) (nal,bk,t) = (map_binder g e nal,(nal,bk,f e t)::bl) in
- let (e,rbl) = List.fold_left h (e,[]) bl in
- (e, List.rev rbl)
-
let map_local_binders f g e bl =
(* TODO: avoid variable capture in [t] by some [na] in [List.tl nal] *)
let h (e,bl) = function
@@ -406,15 +389,15 @@ let map_constr_expr_with_binders g f e = CAst.map (function
| CApp ((p,a),l) ->
CApp ((p,f e a),List.map (fun (a,i) -> (f e a,i)) l)
| CProdN (bl,b) ->
- let (e,bl) = map_binders f g e bl in CProdN (bl,f e b)
+ let (e,bl) = map_local_binders f g e bl in CProdN (bl,f e b)
| CLambdaN (bl,b) ->
- let (e,bl) = map_binders f g e bl in CLambdaN (bl,f e b)
+ let (e,bl) = map_local_binders f g e bl in CLambdaN (bl,f e b)
| CLetIn (na,a,t,b) ->
CLetIn (na,f e a,Option.map (f e) t,f (Name.fold_right g (snd na) e) b)
| CCast (a,c) -> CCast (f e a, Miscops.map_cast_type (f e) c)
- | CNotation (n,(l,ll,bll)) ->
+ | CNotation (n,(l,ll,bl,bll)) ->
(* This is an approximation because we don't know what binds what *)
- CNotation (n,(List.map (f e) l,List.map (List.map (f e)) ll,
+ CNotation (n,(List.map (f e) l,List.map (List.map (f e)) ll, bl,
List.map (fun bl -> snd (map_local_binders f g e bl)) bll))
| CGeneralization (b,a,c) -> CGeneralization (b,a,f e c)
| CDelimiters (s,a) -> CDelimiters (s,f e a)
@@ -473,9 +456,10 @@ let locs_of_notation ?loc locs ntn =
| (ba,ea)::l -> if Int.equal pos ba then aux ea l else (pos,ba)::aux ea l
in aux bl (List.sort (fun l1 l2 -> fst l1 - fst l2) locs)
-let ntn_loc ?loc (args,argslist,binderslist) =
+let ntn_loc ?loc (args,argslist,binders,binderslist) =
locs_of_notation ?loc
(List.map constr_loc (args@List.flatten argslist)@
+ List.map cases_pattern_expr_loc binders@
List.map local_binders_loc binderslist)
let patntn_loc ?loc (args,argslist) =
@@ -529,9 +513,9 @@ let split_at_annot bl na =
let mkIdentC id = CAst.make @@ CRef (Ident (Loc.tag id),None)
let mkRefC r = CAst.make @@ CRef (r,None)
let mkCastC (a,k) = CAst.make @@ CCast (a,k)
-let mkLambdaC (idl,bk,a,b) = CAst.make @@ CLambdaN ([idl,bk,a],b)
+let mkLambdaC (idl,bk,a,b) = CAst.make @@ CLambdaN ([CLocalAssum (idl,bk,a)],b)
let mkLetInC (id,a,t,b) = CAst.make @@ CLetIn (id,a,t,b)
-let mkProdC (idl,bk,a,b) = CAst.make @@ CProdN ([idl,bk,a],b)
+let mkProdC (idl,bk,a,b) = CAst.make @@ CProdN ([CLocalAssum (idl,bk,a)],b)
let mkAppC (f,l) =
let l = List.map (fun x -> (x,None)) l in
@@ -539,56 +523,11 @@ let mkAppC (f,l) =
| CApp (g,l') -> CAst.make @@ CApp (g, l' @ l)
| _ -> CAst.make @@ CApp ((None, f), l)
-let add_name_in_env env n =
- match snd n with
- | Anonymous -> env
- | Name id -> id :: env
-
-let fresh_var env c =
- Namegen.next_ident_away (Id.of_string "pat")
- (List.fold_left (fun accu id -> Id.Set.add id accu) (free_vars_of_constr_expr c) env)
-
-let expand_binders ?loc mkC bl c =
- let rec loop ?loc bl c =
- match bl with
- | [] -> ([], c)
- | b :: bl ->
- match b with
- | CLocalDef ((loc1,_) as n, oty, b) ->
- let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in
- let env = add_name_in_env env n in
- (env, CAst.make ?loc @@ CLetIn (n,oty,b,c))
- | CLocalAssum ((loc1,_)::_ as nl, bk, t) ->
- let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in
- let env = List.fold_left add_name_in_env env nl in
- (env, mkC ?loc (nl,bk,t) c)
- | CLocalAssum ([],_,_) -> loop ?loc bl c
- | CLocalPattern (loc1, (p, ty)) ->
- let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in
- let ni = fresh_var env c in
- let id = (loc1, Name ni) in
- let ty = match ty with
- | Some ty -> ty
- | None -> CAst.make ?loc:loc1 @@ CHole (None, IntroAnonymous, None)
- in
- let e = CAst.make @@ CRef (Libnames.Ident (loc1, ni), None) in
- let c = CAst.make ?loc @@
- CCases
- (LetPatternStyle, None, [(e,None,None)],
- [(Loc.tag ?loc:loc1 ([[p]], c))])
- in
- (ni :: env, mkC ?loc ([id],Default Explicit,ty) c)
- in
- let (_, c) = loop ?loc bl c in
- c
-
let mkCProdN ?loc bll c =
- let mk ?loc b c = CAst.make ?loc @@ CProdN ([b],c) in
- expand_binders ?loc mk bll c
+ CAst.make ?loc @@ CProdN (bll,c)
let mkCLambdaN ?loc bll c =
- let mk ?loc b c = CAst.make ?loc @@ CLambdaN ([b],c) in
- expand_binders ?loc mk bll c
+ CAst.make ?loc @@ CLambdaN (bll,c)
let coerce_reference_to_id = function
| Ident (_,id) -> id
@@ -608,6 +547,49 @@ let coerce_to_name = function
| { CAst.loc; _ } -> CErrors.user_err ?loc ~hdr:"coerce_to_name"
(str "This expression should be a name.")
+let mkCPatOr ?loc = function
+ | [pat] -> pat
+ | disjpat -> CAst.make ?loc @@ (CPatOr disjpat)
+
+let mkAppPattern ?loc p lp =
+ let open CAst in
+ make ?loc @@ (match p.v with
+ | CPatAtom (Some r) -> CPatCstr (r, None, lp)
+ | CPatCstr (r, None, l2) ->
+ CErrors.user_err ?loc:p.loc ~hdr:"compound_pattern"
+ (Pp.str "Nested applications not supported.")
+ | CPatCstr (r, l1, l2) -> CPatCstr (r, l1 , l2@lp)
+ | CPatNotation (n, s, l) -> CPatNotation (n , s, l@lp)
+ | _ -> CErrors.user_err
+ ?loc:p.loc ~hdr:"compound_pattern"
+ (Pp.str "Such pattern cannot have arguments."))
+
+let rec coerce_to_cases_pattern_expr c = CAst.map_with_loc (fun ?loc -> function
+ | CRef (r,None) ->
+ CPatAtom (Some r)
+ | CHole (None,Misctypes.IntroAnonymous,None) ->
+ CPatAtom None
+ | CLetIn ((loc,Name id),b,None,{ CAst.v = CRef (Ident (_,id'),None) }) when Id.equal id id' ->
+ CPatAlias (coerce_to_cases_pattern_expr b, (loc,Name id))
+ | CApp ((None,p),args) when List.for_all (fun (_,e) -> e=None) args ->
+ (mkAppPattern (coerce_to_cases_pattern_expr p) (List.map (fun (a,_) -> coerce_to_cases_pattern_expr a) args)).CAst.v
+ | CAppExpl ((None,r,i),args) ->
+ CPatCstr (r,Some (List.map coerce_to_cases_pattern_expr args),[])
+ | CNotation (ntn,(c,cl,[],[])) ->
+ CPatNotation (ntn,(List.map coerce_to_cases_pattern_expr c,
+ List.map (List.map coerce_to_cases_pattern_expr) cl),[])
+ | CPrim p ->
+ CPatPrim p
+ | CRecord l ->
+ CPatRecord (List.map (fun (r,p) -> (r,coerce_to_cases_pattern_expr p)) l)
+ | CDelimiters (s,p) ->
+ CPatDelimiters (s,coerce_to_cases_pattern_expr p)
+ | CCast (p,CastConv t) ->
+ CPatCast (coerce_to_cases_pattern_expr p,t)
+ | _ ->
+ CErrors.user_err ?loc ~hdr:"coerce_to_cases_pattern_expr"
+ (str "This expression should be coercible to a pattern.")) c
+
let asymmetric_patterns = ref (false)
let _ = Goptions.declare_bool_option {
Goptions.optdepr = false;
diff --git a/interp/constrexpr_ops.mli b/interp/constrexpr_ops.mli
index 6e5c0f851..0b00b0e4d 100644
--- a/interp/constrexpr_ops.mli
+++ b/interp/constrexpr_ops.mli
@@ -54,6 +54,11 @@ val mkCLambdaN : ?loc:Loc.t -> local_binder_expr list -> constr_expr -> constr_e
val mkCProdN : ?loc:Loc.t -> local_binder_expr list -> constr_expr -> constr_expr
(** Same as [prod_constr_expr], with location *)
+val mkCPatOr : ?loc:Loc.t -> cases_pattern_expr list -> cases_pattern_expr
+
+val mkAppPattern : ?loc:Loc.t -> cases_pattern_expr -> cases_pattern_expr list -> cases_pattern_expr
+(** Apply a list of pattern arguments to a pattern *)
+
(** @deprecated variant of mkCLambdaN *)
val abstract_constr_expr : constr_expr -> local_binder_expr list -> constr_expr
[@@ocaml.deprecated "deprecated variant of mkCLambdaN"]
@@ -73,6 +78,8 @@ val coerce_to_id : constr_expr -> Id.t located
val coerce_to_name : constr_expr -> Name.t located
(** Destruct terms of the form [CRef (Ident _)] or [CHole _]. *)
+val coerce_to_cases_pattern_expr : constr_expr -> cases_pattern_expr
+
(** {6 Binder manipulation} *)
val default_binder_kind : binder_kind
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index 4f7d537d3..9e18966b6 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -257,7 +257,7 @@ let insert_pat_delimiters ?loc p = function
let insert_pat_alias ?loc p = function
| Anonymous -> p
- | Name id -> CAst.make ?loc @@ CPatAlias (p,id)
+ | Name _ as na -> CAst.make ?loc @@ CPatAlias (p,(loc,na))
(**********************************************************************)
(* conversion of references *)
@@ -323,34 +323,35 @@ let is_zero s =
Int.equal (String.length s) i || (s.[i] == '0' && aux (i+1))
in aux 0
-let make_notation_gen loc ntn mknot mkprim destprim l =
+let make_notation_gen loc ntn mknot mkprim destprim l bl =
match ntn,List.map destprim l with
(* Special case to avoid writing "- 3" for e.g. (Z.opp 3) *)
| "- _", [Some (Numeral (p,true))] when not (is_zero p) ->
- mknot (loc,ntn,([mknot (loc,"( _ )",l)]))
+ assert (bl=[]);
+ mknot (loc,ntn,([mknot (loc,"( _ )",l,[])]),[])
| _ ->
match decompose_notation_key ntn, l with
| [Terminal "-"; Terminal x], [] when is_number x ->
mkprim (loc, Numeral (x,false))
| [Terminal x], [] when is_number x ->
mkprim (loc, Numeral (x,true))
- | _ -> mknot (loc,ntn,l)
+ | _ -> mknot (loc,ntn,l,bl)
-let make_notation loc ntn (terms,termlists,binders as subst) =
- if not (List.is_empty termlists) || not (List.is_empty binders) then
+let make_notation loc ntn (terms,termlists,binders,binderlists as subst) =
+ if not (List.is_empty termlists) || not (List.is_empty binderlists) then
CAst.make ?loc @@ CNotation (ntn,subst)
else
make_notation_gen loc ntn
- (fun (loc,ntn,l) -> CAst.make ?loc @@ CNotation (ntn,(l,[],[])))
+ (fun (loc,ntn,l,bl) -> CAst.make ?loc @@ CNotation (ntn,(l,[],bl,[])))
(fun (loc,p) -> CAst.make ?loc @@ CPrim p)
- destPrim terms
+ destPrim terms binders
let make_pat_notation ?loc ntn (terms,termlists as subst) args =
if not (List.is_empty termlists) then (CAst.make ?loc @@ CPatNotation (ntn,subst,args)) else
make_notation_gen loc ntn
- (fun (loc,ntn,l) -> CAst.make ?loc @@ CPatNotation (ntn,(l,[]),args))
+ (fun (loc,ntn,l,_) -> CAst.make ?loc @@ CPatNotation (ntn,(l,[]),args))
(fun (loc,p) -> CAst.make ?loc @@ CPatPrim p)
- destPatPrim terms
+ destPatPrim terms []
let mkPat ?loc qid l = CAst.make ?loc @@
(* Normally irrelevant test with v8 syntax, but let's do it anyway *)
@@ -533,6 +534,10 @@ let occur_name na aty =
| Name id -> occur_var_constr_expr id aty
| Anonymous -> false
+let is_gvar id c = match DAst.get c with
+| GVar id' -> Id.equal id id'
+| _ -> false
+
let is_projection nargs = function
| Some r when not !Flags.in_debugger && not !Flags.raw_print && !print_projections ->
(try
@@ -817,13 +822,11 @@ let rec extern inctx scopes vars r =
| GProd (na,bk,t,c) ->
let t = extern_typ scopes vars t in
- let (idl,c) = factorize_prod scopes (add_vname vars na) na bk t c in
- CProdN ([(Loc.tag na)::idl,Default bk,t],c)
+ factorize_prod scopes (add_vname vars na) na bk t c
| GLambda (na,bk,t,c) ->
let t = extern_typ scopes vars t in
- let (idl,c) = factorize_lambda inctx scopes (add_vname vars na) na bk t c in
- CLambdaN ([(Loc.tag na)::idl,Default bk,t],c)
+ factorize_lambda inctx scopes (add_vname vars na) na bk t c
| GCases (sty,rtntypopt,tml,eqns) ->
let vars' =
@@ -919,24 +922,60 @@ and extern_typ (_,scopes) =
and sub_extern inctx (_,scopes) = extern inctx (None,scopes)
and factorize_prod scopes vars na bk aty c =
- let c = extern_typ scopes vars c in
- match na, c with
- | Name id, { CAst.loc ; v = CProdN ([nal,Default bk',ty],c) }
- when binding_kind_eq bk bk' && constr_expr_eq aty ty
- && not (occur_var_constr_expr id ty) (* avoid na in ty escapes scope *) ->
- nal,c
- | _ ->
- [],c
+ let store, get = set_temporary_memory () in
+ match na, DAst.get c with
+ | Name id, GCases (LetPatternStyle, None, [(e,(Anonymous,None))],(_::_ as eqns))
+ when is_gvar id e && List.length (store (factorize_eqns eqns)) = 1 ->
+ (match get () with
+ | [(_,(ids,disj_of_patl,b))] ->
+ let disjpat = List.map (function [pat] -> pat | _ -> assert false) disj_of_patl in
+ let disjpat = if occur_glob_constr id b then List.map (set_pat_alias id) disjpat else disjpat in
+ let b = extern_typ scopes vars b in
+ let p = mkCPatOr (List.map (extern_cases_pattern_in_scope scopes vars) disjpat) in
+ let binder = CLocalPattern (c.loc,(p,None)) in
+ (match b.v with
+ | CProdN (bl,b) -> CProdN (binder::bl,b)
+ | _ -> CProdN ([binder],b))
+ | _ -> assert false)
+ | _, _ ->
+ let c = extern_typ scopes vars c in
+ match na, c.v with
+ | Name id, CProdN (CLocalAssum(nal,Default bk',ty)::bl,b)
+ when binding_kind_eq bk bk' && constr_expr_eq aty ty
+ && not (occur_var_constr_expr id ty) (* avoid na in ty escapes scope *) ->
+ CProdN (CLocalAssum(Loc.tag na::nal,Default bk,aty)::bl,b)
+ | _, CProdN (bl,b) ->
+ CProdN (CLocalAssum([Loc.tag na],Default bk,aty)::bl,b)
+ | _, _ ->
+ CProdN ([CLocalAssum([Loc.tag na],Default bk,aty)],c)
and factorize_lambda inctx scopes vars na bk aty c =
- let c = sub_extern inctx scopes vars c in
- match c with
- | { CAst.loc; v = CLambdaN ([nal,Default bk',ty],c) }
- when binding_kind_eq bk bk' && constr_expr_eq aty ty
- && not (occur_name na ty) (* avoid na in ty escapes scope *) ->
- nal,c
- | _ ->
- [],c
+ let store, get = set_temporary_memory () in
+ match na, DAst.get c with
+ | Name id, GCases (LetPatternStyle, None, [(e,(Anonymous,None))],(_::_ as eqns))
+ when is_gvar id e && List.length (store (factorize_eqns eqns)) = 1 ->
+ (match get () with
+ | [(_,(ids,disj_of_patl,b))] ->
+ let disjpat = List.map (function [pat] -> pat | _ -> assert false) disj_of_patl in
+ let disjpat = if occur_glob_constr id b then List.map (set_pat_alias id) disjpat else disjpat in
+ let b = sub_extern inctx scopes vars b in
+ let p = mkCPatOr (List.map (extern_cases_pattern_in_scope scopes vars) disjpat) in
+ let binder = CLocalPattern (c.loc,(p,None)) in
+ (match b.v with
+ | CLambdaN (bl,b) -> CLambdaN (binder::bl,b)
+ | _ -> CLambdaN ([binder],b))
+ | _ -> assert false)
+ | _, _ ->
+ let c = sub_extern inctx scopes vars c in
+ match c.v with
+ | CLambdaN (CLocalAssum(nal,Default bk',ty)::bl,b)
+ when binding_kind_eq bk bk' && constr_expr_eq aty ty
+ && not (occur_name na ty) (* avoid na in ty escapes scope *) ->
+ CLambdaN (CLocalAssum(Loc.tag na::nal,Default bk,aty)::bl,b)
+ | CLambdaN (bl,b) ->
+ CLambdaN (CLocalAssum([Loc.tag na],Default bk,aty)::bl,b)
+ | _ ->
+ CLambdaN ([CLocalAssum([Loc.tag na],Default bk,aty)],c)
and extern_local_binder scopes vars = function
[] -> ([],[],[])
@@ -965,7 +1004,7 @@ and extern_local_binder scopes vars = function
| GLocalPattern ((p,_),_,bk,ty) ->
let ty =
if !Flags.raw_print then Some (extern_typ scopes vars ty) else None in
- let p = extern_cases_pattern vars p in
+ let p = mkCPatOr (List.map (extern_cases_pattern vars) p) in
let (assums,ids,l) = extern_local_binder scopes vars l in
(assums,ids, CLocalPattern(Loc.tag @@ (p,ty)) :: l)
@@ -1014,7 +1053,7 @@ and extern_notation (tmp_scope,scopes as allscopes) vars t = function
| _, None -> t, [], [], []
| _ -> raise No_match in
(* Try matching ... *)
- let terms,termlists,binders =
+ let terms,termlists,binders,binderlists =
match_notation_constr !print_universes t pat in
(* Try availability of interpretation ... *)
let e =
@@ -1035,11 +1074,15 @@ and extern_notation (tmp_scope,scopes as allscopes) vars t = function
List.map (fun (c,(scopt,scl)) ->
List.map (extern true (scopt,scl@scopes') vars) c)
termlists in
- let bll =
- List.map (fun (bl,(scopt,scl)) ->
- pi3 (extern_local_binder (scopt,scl@scopes') vars bl))
+ let bl =
+ List.map (fun (bl,(scopt,scl)) ->
+ mkCPatOr (List.map (extern_cases_pattern_in_scope (scopt,scl@scopes') vars) bl))
binders in
- insert_delimiters (make_notation loc ntn (l,ll,bll)) key)
+ let bll =
+ List.map (fun (bl,(scopt,scl)) ->
+ pi3 (extern_local_binder (scopt,scl@scopes') vars bl))
+ binderlists in
+ insert_delimiters (make_notation loc ntn (l,ll,bl,bll)) key)
| SynDefRule kn ->
let l =
List.map (fun (c,(scopt,scl)) ->
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index 4afe301dd..694bec897 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -213,20 +213,20 @@ let expand_notation_string ntn n =
(* This contracts the special case of "{ _ }" for sumbool, sumor notations *)
(* Remark: expansion of squash at definition is done in metasyntax.ml *)
-let contract_notation ntn (l,ll,bll) =
+let contract_curly_brackets ntn (l,ll,bl,bll) =
let ntn' = ref ntn in
let rec contract_squash n = function
| [] -> []
- | { CAst.v = CNotation ("{ _ }",([a],[],[])) } :: l ->
+ | { CAst.v = CNotation ("{ _ }",([a],[],[],[])) } :: l ->
ntn' := expand_notation_string !ntn' n;
contract_squash n (a::l)
| a :: l ->
a::contract_squash (n+1) l in
let l = contract_squash 0 l in
(* side effect; don't inline *)
- !ntn',(l,ll,bll)
+ !ntn',(l,ll,bl,bll)
-let contract_pat_notation ntn (l,ll) =
+let contract_curly_brackets_pat ntn (l,ll) =
let ntn' = ref ntn in
let rec contract_squash n = function
| [] -> []
@@ -271,29 +271,24 @@ let error_expect_binder_notation_type ?loc id =
(Id.print id ++
str " is expected to occur in binding position in the right-hand side.")
-let set_var_scope ?loc id istermvar env ntnvars =
+let set_var_scope ?loc id istermvar (tmp_scope,subscopes as scopes) ntnvars =
try
- let isonlybinding,idscopes,typ = Id.Map.find id ntnvars in
- if istermvar then isonlybinding := false;
+ let used_as_binder,idscopes,typ = Id.Map.find id ntnvars in
+ if not istermvar then used_as_binder := true;
let () = if istermvar then
(* scopes have no effect on the interpretation of identifiers *)
begin match !idscopes with
- | None -> idscopes := Some (env.tmp_scope, env.scopes)
- | Some (tmp, scope) ->
- let s1 = make_current_scope tmp scope in
- let s2 = make_current_scope env.tmp_scope env.scopes in
- if not (List.equal String.equal s1 s2) then error_inconsistent_scope ?loc id s1 s2
+ | None -> idscopes := Some scopes
+ | Some (tmp_scope', subscopes') ->
+ let s' = make_current_scope tmp_scope' subscopes' in
+ let s = make_current_scope tmp_scope subscopes in
+ if not (List.equal String.equal s' s) then error_inconsistent_scope ?loc id s' s
end
in
match typ with
- | NtnInternTypeBinder ->
+ | Notation_term.NtnInternTypeOnlyBinder ->
if istermvar then error_expect_binder_notation_type ?loc id
- | NtnInternTypeConstr ->
- (* We need sometimes to parse idents at a constr level for
- factorization and we cannot enforce this constraint:
- if not istermvar then error_expect_constr_notation_type loc id *)
- ()
- | NtnInternTypeIdent -> ()
+ | Notation_term.NtnInternTypeAny -> ()
with Not_found ->
(* Not in a notation *)
()
@@ -302,15 +297,11 @@ let set_type_scope env = {env with tmp_scope = Notation.current_type_scope_name
let reset_tmp_scope env = {env with tmp_scope = None}
-let rec it_mkGProd ?loc env body =
- match env with
- (loc2, (na, bk, t)) :: tl -> it_mkGProd ?loc:loc2 tl (DAst.make ?loc:(Loc.merge_opt loc loc2) @@ GProd (na, bk, t, body))
- | [] -> body
+let set_env_scopes env (scopt,subscopes) =
+ {env with tmp_scope = scopt; scopes = subscopes @ env.scopes}
-let rec it_mkGLambda ?loc env body =
- match env with
- (loc2, (na, bk, t)) :: tl -> it_mkGLambda ?loc:loc2 tl (DAst.make ?loc:(Loc.merge_opt loc loc2) @@ GLambda (na, bk, t, body))
- | [] -> body
+let mkGProd ?loc (na,bk,t) body = DAst.make ?loc @@ GProd (na, bk, t, body)
+let mkGLambda ?loc (na,bk,t) body = DAst.make ?loc @@ GLambda (na, bk, t, body)
(**********************************************************************)
(* Utilities for binders *)
@@ -378,12 +369,12 @@ let push_name_env ?(global_level=false) ntnvars implargs env =
check_hidden_implicit_parameters ?loc id env.impls ;
if Id.Map.is_empty ntnvars && Id.equal id ldots_var
then error_ldots_var ?loc;
- set_var_scope ?loc id false env ntnvars;
+ set_var_scope ?loc id false (env.tmp_scope,env.scopes) ntnvars;
if global_level then Dumpglob.dump_definition (loc,id) true "var"
else Dumpglob.dump_binding ?loc id;
{env with ids = Id.Set.add id env.ids; impls = Id.Map.add id implargs env.impls}
-let intern_generalized_binder ?(global_level=false) intern_type lvar
+let intern_generalized_binder ?(global_level=false) intern_type ntnvars
env (loc, na) b b' t ty =
let ids = (match na with Anonymous -> fun x -> x | Name na -> Id.Set.add na) env.ids in
let ty, ids' =
@@ -394,7 +385,7 @@ let intern_generalized_binder ?(global_level=false) intern_type lvar
let ty' = intern_type {env with ids = ids; unb = true} ty in
let fvs = Implicit_quantifiers.generalizable_vars_of_glob_constr ~bound:ids ~allowed:ids' ty' in
let env' = List.fold_left
- (fun env (l, x) -> push_name_env ~global_level lvar (Variable,[],[],[])(*?*) env (l, Name x))
+ (fun env (l, x) -> push_name_env ~global_level ntnvars (Variable,[],[],[])(*?*) env (l, Name x))
env fvs in
let bl = List.map
(fun (loc, id) ->
@@ -413,9 +404,9 @@ let intern_generalized_binder ?(global_level=false) intern_type lvar
in Implicit_quantifiers.make_fresh ids' (Global.env ()) id
in Name name
| _ -> na
- in (push_name_env ~global_level lvar (impls_type_list ty')(*?*) env' (loc,na)), (loc,(na,b',ty')) :: List.rev bl
+ in (push_name_env ~global_level ntnvars (impls_type_list ty')(*?*) env' (loc,na)), (loc,(na,b',ty')) :: List.rev bl
-let intern_assumption intern lvar env nal bk ty =
+let intern_assumption intern ntnvars env nal bk ty =
let intern_type env = intern (set_type_scope env) in
match bk with
| Default k ->
@@ -424,11 +415,11 @@ let intern_assumption intern lvar env nal bk ty =
let impls = impls_type_list ty in
List.fold_left
(fun (env, bl) (loc, na as locna) ->
- (push_name_env lvar impls env locna,
+ (push_name_env ntnvars impls env locna,
(Loc.tag ?loc (na,k,locate_if_hole ?loc na ty))::bl))
(env, []) nal
| Generalized (b,b',t) ->
- let env, b = intern_generalized_binder intern_type lvar env (List.hd nal) b b' t ty in
+ let env, b = intern_generalized_binder intern_type ntnvars env (List.hd nal) b b' t ty in
env, b
let glob_local_binder_of_extended = DAst.with_loc_val (fun ?loc -> function
@@ -443,39 +434,48 @@ let glob_local_binder_of_extended = DAst.with_loc_val (fun ?loc -> function
let intern_cases_pattern_fwd = ref (fun _ -> failwith "intern_cases_pattern_fwd")
-let intern_local_binder_aux ?(global_level=false) intern lvar (env,bl) = function
+let intern_letin_binder intern ntnvars env ((loc,na as locna),def,ty) =
+ let term = intern env def in
+ let ty = Option.map (intern env) ty in
+ (push_name_env ntnvars (impls_term_list term) env locna,
+ (na,Explicit,term,ty))
+
+let intern_cases_pattern_as_binder ?loc ntnvars env p =
+ let il,disjpat =
+ let (il, subst_disjpat) = !intern_cases_pattern_fwd ntnvars (None,env.scopes) p in
+ let substl,disjpat = List.split subst_disjpat in
+ if not (List.for_all (fun subst -> Id.Map.equal Id.equal subst Id.Map.empty) substl) then
+ user_err ?loc (str "Unsupported nested \"as\" clause.");
+ il,disjpat
+ in
+ let env = List.fold_right (fun (loc,id) env -> push_name_env ntnvars (Variable,[],[],[]) env (loc,Name id)) il env in
+ let na = alias_of_pat (List.hd disjpat) in
+ let ienv = Name.fold_right Id.Set.remove na env.ids in
+ let id = Namegen.next_name_away_with_default "pat" na ienv in
+ let na = (loc, Name id) in
+ env,((disjpat,il),id),na
+
+let intern_local_binder_aux ?(global_level=false) intern ntnvars (env,bl) = function
| CLocalAssum(nal,bk,ty) ->
- let env, bl' = intern_assumption intern lvar env nal bk ty in
+ let env, bl' = intern_assumption intern ntnvars env nal bk ty in
let bl' = List.map (fun (loc,(na,c,t)) -> DAst.make ?loc @@ GLocalAssum (na,c,t)) bl' in
env, bl' @ bl
| CLocalDef((loc,na as locna),def,ty) ->
- let term = intern env def in
- let ty = Option.map (intern env) ty in
- (push_name_env lvar (impls_term_list term) env locna,
- (DAst.make ?loc @@ GLocalDef (na,Explicit,term,ty)) :: bl)
+ let env,(na,bk,def,ty) = intern_letin_binder intern ntnvars env (locna,def,ty) in
+ env, (DAst.make ?loc @@ GLocalDef (na,bk,def,ty)) :: bl
| CLocalPattern (loc,(p,ty)) ->
let tyc =
match ty with
| Some ty -> ty
| None -> CAst.make ?loc @@ CHole(None,Misctypes.IntroAnonymous,None)
in
- let il,cp =
- match !intern_cases_pattern_fwd (None,env.scopes) p with
- | (il, [(subst,cp)]) ->
- if not (Id.Map.equal Id.equal subst Id.Map.empty) then
- user_err ?loc (str "Unsupported nested \"as\" clause.");
- il,cp
- | _ -> assert false
- in
- let env = {env with ids = List.fold_right Id.Set.add il env.ids} in
- let id = Namegen.next_ident_away (Id.of_string "pat") env.ids in
- let na = (loc, Name id) in
+ let env, ((disjpat,il),id),na = intern_cases_pattern_as_binder ?loc ntnvars env p in
let bk = Default Explicit in
- let _, bl' = intern_assumption intern lvar env [na] bk tyc in
+ let _, bl' = intern_assumption intern ntnvars env [na] bk tyc in
let _,(_,bk,t) = List.hd bl' in
- (env, (DAst.make ?loc @@ GLocalPattern((cp,il),id,bk,t)) :: bl)
+ (env, (DAst.make ?loc @@ GLocalPattern((disjpat,List.map snd il),id,bk,t)) :: bl)
-let intern_generalization intern env lvar loc bk ak c =
+let intern_generalization intern env ntnvars loc bk ak c =
let c = intern {env with unb = true} c in
let fvs = Implicit_quantifiers.generalizable_vars_of_glob_constr ~bound:env.ids c in
let env', c' =
@@ -504,10 +504,26 @@ let intern_generalization intern env lvar loc bk ak c =
GLambda (Name id, bk, DAst.make ?loc:loc' @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc))
in
List.fold_right (fun (loc, id as lid) (env, acc) ->
- let env' = push_name_env lvar (Variable,[],[],[]) env (loc, Name id) in
+ let env' = push_name_env ntnvars (Variable,[],[],[]) env (loc, Name id) in
(env', abs lid acc)) fvs (env,c)
in c'
+let rec expand_binders ?loc mk bl c =
+ match bl with
+ | [] -> c
+ | b :: bl ->
+ match DAst.get b with
+ | GLocalDef (n, bk, b, oty) ->
+ expand_binders ?loc mk bl (DAst.make ?loc @@ GLetIn (n, b, oty, c))
+ | GLocalAssum (n, bk, t) ->
+ expand_binders ?loc mk bl (mk ?loc (n,bk,t) c)
+ | GLocalPattern ((disjpat,ids), id, bk, ty) ->
+ let tm = DAst.make ?loc (GVar id) in
+ (* Distribute the disjunctive patterns over the shared right-hand side *)
+ let eqnl = List.map (fun pat -> (loc,(ids,[pat],c))) disjpat in
+ let c = DAst.make ?loc @@ GCases (Misctypes.LetPatternStyle, None, [tm,(Anonymous,None)], eqnl) in
+ expand_binders ?loc mk bl (mk ?loc (Name id,Explicit,ty) c)
+
(**********************************************************************)
(* Syntax extensions *)
@@ -515,7 +531,7 @@ let option_mem_assoc id = function
| Some (id',c) -> Id.equal id id'
| None -> false
-let find_fresh_name renaming (terms,termlists,binders) avoid id =
+let find_fresh_name renaming (terms,termlists,binders,binderlists) avoid id =
let fold1 _ (c, _) accu = Id.Set.union (free_vars_of_constr_expr c) accu in
let fold2 _ (l, _) accu =
let fold accu c = Id.Set.union (free_vars_of_constr_expr c) accu in
@@ -528,14 +544,53 @@ let find_fresh_name renaming (terms,termlists,binders) avoid id =
(* TODO binders *)
next_ident_away_from id (fun id -> Id.Set.mem id fvs3)
-let traverse_binder (terms,_,_ as subst) avoid (renaming,env) = function
- | Anonymous -> (renaming,env),Anonymous
+let is_var store pat =
+ match DAst.get pat with
+ | PatVar na -> store na; true
+ | _ -> false
+
+let out_var pat =
+ match pat.CAst.v with
+ | CPatAtom (Some (Ident (_,id))) -> Name id
+ | CPatAtom None -> Anonymous
+ | _ -> assert false
+
+let term_of_name = function
+ | Name id -> DAst.make (GVar id)
+ | Anonymous ->
+ let st = Evar_kinds.Define (not (Program.get_proofs_transparency ())) in
+ DAst.make (GHole (Evar_kinds.QuestionMark (st,Anonymous), Misctypes.IntroAnonymous, None))
+
+let traverse_binder intern_pat ntnvars (terms,_,binders,_ as subst) avoid (renaming,env) = function
+ | Anonymous -> (renaming,env), None, Anonymous
| Name id ->
+ let store,get = set_temporary_memory () in
+ try
+ (* We instantiate binder name with patterns which may be parsed as terms *)
+ let pat = coerce_to_cases_pattern_expr (fst (Id.Map.find id terms)) in
+ let env,((disjpat,ids),id),na = intern_pat ntnvars env pat in
+ let pat, na = match disjpat with
+ | [pat] when is_var store pat -> let na = get () in None, na
+ | _ -> Some ((List.map snd ids,disjpat),id), snd na in
+ (renaming,env), pat, na
+ with Not_found ->
try
- (* Binders bound in the notation are considered first-order objects *)
- let _,na as locna = coerce_to_name (fst (Id.Map.find id terms)) in
- let env = push_name_env Id.Map.empty (Variable,[],[],[]) env locna in
- (renaming,env), na
+ (* Trying to associate a pattern *)
+ let pat,(onlyident,scopes) = Id.Map.find id binders in
+ let env = set_env_scopes env scopes in
+ if onlyident then
+ (* Do not try to interpret a variable as a constructor *)
+ let na = out_var pat in
+ let env = push_name_env ntnvars (Variable,[],[],[]) env (pat.CAst.loc, na) in
+ (renaming,env), None, na
+ else
+ (* Interpret as a pattern *)
+ let env,((disjpat,ids),id),na = intern_pat ntnvars env pat in
+ let pat, na =
+ match disjpat with
+ | [pat] when is_var store pat -> let na = get () in None, na
+ | _ -> Some ((List.map snd ids,disjpat),id), snd na in
+ (renaming,env), pat, na
with Not_found ->
(* Binders not bound in the notation do not capture variables *)
(* outside the notation (i.e. in the substitution) *)
@@ -543,49 +598,27 @@ let traverse_binder (terms,_,_ as subst) avoid (renaming,env) = function
let renaming' =
if Id.equal id id' then renaming else Id.Map.add id id' renaming
in
- (renaming',env), Name id'
-
-type letin_param_r =
- | LPLetIn of Name.t * glob_constr * glob_constr option
- | LPCases of (cases_pattern * Id.t list) * Id.t
-(* Unused thus fatal warning *)
-(* and letin_param = letin_param_r Loc.located *)
-
-let make_letins =
- List.fold_right
- (fun a c ->
- match a with
- | loc, LPLetIn (na,b,t) ->
- DAst.make ?loc @@ GLetIn(na,b,t,c)
- | loc, LPCases ((cp,il),id) ->
- let tt = (DAst.make ?loc @@ GVar id, (Name id,None)) in
- DAst.make ?loc @@ GCases(Misctypes.LetPatternStyle,None,[tt],[(loc,(il,[cp],c))]))
-
-let rec subordinate_letins letins l = match l with
- | bnd :: l ->
- let loc = bnd.CAst.loc in
- begin match DAst.get bnd with
- (* binders come in reverse order; the non-let are returned in reverse order together *)
- (* with the subordinated let-in in writing order *)
- | GLocalDef (na,_,b,t) ->
- subordinate_letins ((Loc.tag ?loc @@ LPLetIn (na,b,t))::letins) l
- | GLocalAssum (na,bk,t) ->
- let letins',rest = subordinate_letins [] l in
- letins',((loc,(na,bk,t)),letins)::rest
- | GLocalPattern (u,id,bk,t) ->
- subordinate_letins ((Loc.tag ?loc @@ LPCases (u,id))::letins)
- ([DAst.make ?loc @@ GLocalAssum (Name id,bk,t)] @ l)
- end
- | [] ->
- letins,[]
+ (renaming',env), None, Name id'
+
+type binder_action =
+| AddLetIn of Name.t Loc.located * constr_expr * constr_expr option
+| AddTermIter of (constr_expr * subscopes) Names.Id.Map.t
+| AddPreBinderIter of Id.t * local_binder_expr (* A binder to be internalized *)
+| AddBinderIter of Id.t * extended_glob_local_binder (* A binder already internalized - used for generalized binders *)
let dmap_with_loc f n =
CAst.map_with_loc (fun ?loc c -> f ?loc (DAst.get_thunk c)) n
+let error_cannot_coerce_wildcard_term ?loc () =
+ user_err ?loc Pp.(str "Cannot turn \"_\" into a term.")
+
+let error_cannot_coerce_disjunctive_pattern_term ?loc () =
+ user_err ?loc Pp.(str "Cannot turn a disjunctive pattern into a term.")
+
let terms_of_binders bl =
let rec term_of_pat pt = dmap_with_loc (fun ?loc -> function
| PatVar (Name id) -> CRef (Ident (loc,id), None)
- | PatVar (Anonymous) -> user_err Pp.(str "Cannot turn \"_\" into a term.")
+ | PatVar (Anonymous) -> error_cannot_coerce_wildcard_term ?loc ()
| PatCstr (c,l,_) ->
let r = Qualid (loc,qualid_of_path (path_of_global (ConstructRef c))) in
let hole = CAst.make ?loc @@ CHole (None,Misctypes.IntroAnonymous,None) in
@@ -599,40 +632,62 @@ let terms_of_binders bl =
| GLocalDef (Name id,_,_,_) -> extract_variables l
| GLocalDef (Anonymous,_,_,_)
| GLocalAssum (Anonymous,_,_) -> user_err Pp.(str "Cannot turn \"_\" into a term.")
- | GLocalPattern ((u,_),_,_,_) -> term_of_pat u :: extract_variables l
+ | GLocalPattern (([u],_),_,_,_) -> term_of_pat u :: extract_variables l
+ | GLocalPattern ((_,_),_,_,_) -> error_cannot_coerce_disjunctive_pattern_term ?loc ()
end
| [] -> [] in
extract_variables bl
-let instantiate_notation_constr loc intern ntnvars subst infos c =
- let (terms,termlists,binders) = subst in
+let flatten_generalized_binders_if_any y l =
+ match List.rev l with
+ | [] -> assert false
+ | a::l -> a, List.map (fun a -> AddBinderIter (y,a)) l (* if l not empty, this means we had a generalized binder *)
+
+let flatten_binders bl =
+ let dispatch = function
+ | CLocalAssum (nal,bk,t) -> List.map (fun na -> CLocalAssum ([na],bk,t)) nal
+ | a -> [a] in
+ List.flatten (List.map dispatch bl)
+
+let instantiate_notation_constr loc intern intern_pat ntnvars subst infos c =
+ let (terms,termlists,binders,binderlists) = subst in
(* when called while defining a notation, avoid capturing the private binders
of the expression by variables bound by the notation (see #3892) *)
let avoid = Id.Map.domain ntnvars in
- let rec aux (terms,binderopt,terminopt as subst') (renaming,env) c =
+ let rec aux (terms,binderopt,iteropt as subst') (renaming,env) c =
let subinfos = renaming,{env with tmp_scope = None} in
match c with
- | NVar id when Id.equal id ldots_var -> Option.get terminopt
+ | NVar id when Id.equal id ldots_var ->
+ let rec aux_letin env = function
+ | [],terminator,_ -> aux (terms,None,None) (renaming,env) terminator
+ | AddPreBinderIter (y,binder)::rest,terminator,iter ->
+ let env,binders = intern_local_binder_aux intern ntnvars (env,[]) binder in
+ let binder,extra = flatten_generalized_binders_if_any y binders in
+ aux (terms,Some (y,binder),Some (extra@rest,terminator,iter)) (renaming,env) iter
+ | AddBinderIter (y,binder)::rest,terminator,iter ->
+ aux (terms,Some (y,binder),Some (rest,terminator,iter)) (renaming,env) iter
+ | AddTermIter nterms::rest,terminator,iter ->
+ aux (nterms,None,Some (rest,terminator,iter)) (renaming,env) iter
+ | AddLetIn (na,c,t)::rest,terminator,iter ->
+ let env,(na,_,c,t) = intern_letin_binder intern ntnvars env (na,c,t) in
+ DAst.make ?loc (GLetIn (na,c,t,aux_letin env (rest,terminator,iter))) in
+ aux_letin env (Option.get iteropt)
| NVar id -> subst_var subst' (renaming, env) id
- | NList (x,y,iter,terminator,lassoc) ->
+ | NList (x,y,iter,terminator,revert) ->
let l,(scopt,subscopes) =
(* All elements of the list are in scopes (scopt,subscopes) *)
try
let l,scopes = Id.Map.find x termlists in
- (if lassoc then List.rev l else l),scopes
+ (if revert then List.rev l else l),scopes
with Not_found ->
try
- let (bl,(scopt,subscopes)) = Id.Map.find x binders in
+ let (bl,(scopt,subscopes)) = Id.Map.find x binderlists in
let env,bl' = List.fold_left (intern_local_binder_aux intern ntnvars) (env,[]) bl in
- terms_of_binders (if lassoc then bl' else List.rev bl'),(None,[])
+ terms_of_binders (if revert then bl' else List.rev bl'),(None,[])
with Not_found ->
anomaly (Pp.str "Inconsistent substitution of recursive notation.") in
- let termin = aux (terms,None,None) subinfos terminator in
- let fold a t =
- let nterms = Id.Map.add y (a, (scopt, subscopes)) terms in
- aux (nterms,None,Some t) subinfos iter
- in
- List.fold_right fold l termin
+ let l = List.map (fun a -> AddTermIter ((Id.Map.add y (a,(scopt,subscopes)) terms))) l in
+ aux (terms,None,Some (l,terminator,iter)) subinfos (NVar ldots_var)
| NHole (knd, naming, arg) ->
let knd = match knd with
| Evar_kinds.BinderType (Name id as na) ->
@@ -653,47 +708,57 @@ let instantiate_notation_constr loc intern ntnvars subst infos c =
let gc = intern nenv c in
(gc, Some c)
in
- let bindings = Id.Map.map mk_env terms in
+ let mk_env' (c, (onlyident,(tmp_scope,subscopes))) =
+ let nenv = {env with tmp_scope; scopes = subscopes @ env.scopes} in
+ if onlyident then
+ let na = out_var c in term_of_name na, None
+ else
+ let _,((disjpat,_),_),_ = intern_pat ntnvars nenv c in
+ match disjpat with
+ | [pat] -> (glob_constr_of_cases_pattern pat, None)
+ | _ -> error_cannot_coerce_disjunctive_pattern_term ?loc:c.CAst.loc ()
+ in
+ let terms = Id.Map.map mk_env terms in
+ let binders = Id.Map.map mk_env' binders in
+ let bindings = Id.Map.fold Id.Map.add terms binders in
Some (Genintern.generic_substitute_notation bindings arg)
in
DAst.make ?loc @@ GHole (knd, naming, arg)
- | NBinderList (x,y,iter,terminator) ->
+ | NBinderList (x,y,iter,terminator,revert) ->
(try
(* All elements of the list are in scopes (scopt,subscopes) *)
- let (bl,(scopt,subscopes)) = Id.Map.find x binders in
- let env,bl = List.fold_left (intern_local_binder_aux intern ntnvars) (env,[]) bl in
- let letins,bl = subordinate_letins [] bl in
- let termin = aux (terms,None,None) (renaming,env) terminator in
- let res = List.fold_left (fun t binder ->
- aux (terms,Some(y,binder),Some t) subinfos iter)
- termin bl in
- make_letins letins res
+ let (bl,(scopt,subscopes)) = Id.Map.find x binderlists in
+ (* We flatten binders so that we can interpret them at substitution time *)
+ let bl = flatten_binders bl in
+ let bl = if revert then List.rev bl else bl in
+ (* We isolate let-ins which do not contribute to the repeated pattern *)
+ let l = List.map (function | CLocalDef (na,c,t) -> AddLetIn (na,c,t)
+ | binder -> AddPreBinderIter (y,binder)) bl in
+ (* We stack the binders to iterate or let-ins to insert *)
+ aux (terms,None,Some (l,terminator,iter)) subinfos (NVar ldots_var)
with Not_found ->
anomaly (Pp.str "Inconsistent substitution of recursive notation."))
| NProd (Name id, NHole _, c') when option_mem_assoc id binderopt ->
- let a,letins = snd (Option.get binderopt) in
- let e = make_letins letins (aux subst' infos c') in
- let (_loc,(na,bk,t)) = a in
- DAst.make ?loc @@ GProd (na,bk,t,e)
+ let binder = snd (Option.get binderopt) in
+ expand_binders ?loc mkGProd [binder] (aux subst' (renaming,env) c')
| NLambda (Name id,NHole _,c') when option_mem_assoc id binderopt ->
- let a,letins = snd (Option.get binderopt) in
- let (_loc,(na,bk,t)) = a in
- DAst.make ?loc @@ GLambda (na,bk,t,make_letins letins (aux subst' infos c'))
+ let binder = snd (Option.get binderopt) in
+ expand_binders ?loc mkGLambda [binder] (aux subst' (renaming,env) c')
(* Two special cases to keep binder name synchronous with BinderType *)
| NProd (na,NHole(Evar_kinds.BinderType na',naming,arg),c')
when Name.equal na na' ->
- let subinfos,na = traverse_binder subst avoid subinfos na in
+ let subinfos,disjpat,na = traverse_binder intern_pat ntnvars subst avoid subinfos na in
let ty = DAst.make ?loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in
- DAst.make ?loc @@ GProd (na,Explicit,ty,aux subst' subinfos c')
+ DAst.make ?loc @@ GProd (na,Explicit,ty,Option.fold_right apply_cases_pattern disjpat (aux subst' subinfos c'))
| NLambda (na,NHole(Evar_kinds.BinderType na',naming,arg),c')
when Name.equal na na' ->
- let subinfos,na = traverse_binder subst avoid subinfos na in
+ let subinfos,disjpat,na = traverse_binder intern_pat ntnvars subst avoid subinfos na in
let ty = DAst.make ?loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in
- DAst.make ?loc @@ GLambda (na,Explicit,ty,aux subst' subinfos c')
+ DAst.make ?loc @@ GLambda (na,Explicit,ty,Option.fold_right apply_cases_pattern disjpat (aux subst' subinfos c'))
| t ->
glob_constr_of_notation_constr_with_binders ?loc
- (traverse_binder subst avoid) (aux subst') subinfos t
- and subst_var (terms, _binderopt, _terminopt) (renaming, env) id =
+ (traverse_binder intern_pat ntnvars subst avoid) (aux subst') subinfos t
+ and subst_var (terms, binderopt, _terminopt) (renaming, env) id =
(* subst remembers the delimiters stack in the interpretation *)
(* of the notations *)
try
@@ -701,6 +766,28 @@ let instantiate_notation_constr loc intern ntnvars subst infos c =
intern {env with tmp_scope = scopt;
scopes = subscopes @ env.scopes} a
with Not_found ->
+ try
+ let pat,(onlyident,scopes) = Id.Map.find id binders in
+ let env = set_env_scopes env scopes in
+ (* We deactivate impls to avoid the check on hidden parameters *)
+ (* and since we are only interested in the pattern as a term *)
+ let env = reset_hidden_inductive_implicit_test env in
+ if onlyident then
+ term_of_name (out_var pat)
+ else
+ let env,((disjpat,ids),id),na = intern_pat ntnvars env pat in
+ match disjpat with
+ | [pat] -> glob_constr_of_cases_pattern pat
+ | _ -> user_err Pp.(str "Cannot turn a disjunctive pattern into a term.")
+ with Not_found ->
+ try
+ match binderopt with
+ | Some (x,binder) when Id.equal x id ->
+ let terms = terms_of_binders [binder] in
+ assert (List.length terms = 1);
+ intern env (List.hd terms)
+ | _ -> raise Not_found
+ with Not_found ->
DAst.make ?loc (
try
GVar (Id.Map.find id renaming)
@@ -709,27 +796,80 @@ let instantiate_notation_constr loc intern ntnvars subst infos c =
GVar id)
in aux (terms,None,None) infos c
-let split_by_type ids =
- List.fold_right (fun (x,(scl,typ)) (l1,l2,l3) ->
+(* Turning substitution coming from parsing and based on production
+ into a substitution for interpretation and based on binding/constr
+ distinction *)
+
+let cases_pattern_of_name (loc,na) =
+ let atom = match na with Name id -> Some (Ident (loc,id)) | Anonymous -> None in
+ CAst.make ?loc (CPatAtom atom)
+
+let split_by_type ids subst =
+ let bind id scl l s =
+ match l with
+ | [] -> assert false
+ | a::l -> l, Id.Map.add id (a,scl) s in
+ let (terms,termlists,binders,binderlists),subst =
+ List.fold_left (fun ((terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists')) (id,(scl,typ)) ->
match typ with
- | NtnTypeConstr | NtnTypeOnlyBinder -> ((x,scl)::l1,l2,l3)
- | NtnTypeConstrList -> (l1,(x,scl)::l2,l3)
- | NtnTypeBinderList -> (l1,l2,(x,scl)::l3)) ids ([],[],[])
+ | NtnTypeConstr ->
+ let terms,terms' = bind id scl terms terms' in
+ (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists')
+ | NtnTypeBinder NtnBinderParsedAsConstr (Extend.AsIdentOrPattern | Extend.AsStrictPattern) ->
+ let a,terms = match terms with a::terms -> a,terms | _ -> assert false in
+ let binders' = Id.Map.add id (coerce_to_cases_pattern_expr a,(false,scl)) binders' in
+ (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists')
+ | NtnTypeBinder NtnBinderParsedAsConstr Extend.AsIdent ->
+ let a,terms = match terms with a::terms -> a,terms | _ -> assert false in
+ let binders' = Id.Map.add id (cases_pattern_of_name (coerce_to_name a),(true,scl)) binders' in
+ (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists')
+ | NtnTypeBinder (NtnParsedAsIdent | NtnParsedAsPattern _ as x) ->
+ let onlyident = (x = NtnParsedAsIdent) in
+ let binders,binders' = bind id (onlyident,scl) binders binders' in
+ (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists')
+ | NtnTypeConstrList ->
+ let termlists,termlists' = bind id scl termlists termlists' in
+ (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists')
+ | NtnTypeBinderList ->
+ let binderlists,binderlists' = bind id scl binderlists binderlists' in
+ (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists'))
+ (subst,(Id.Map.empty,Id.Map.empty,Id.Map.empty,Id.Map.empty)) ids in
+ assert (terms = [] && termlists = [] && binders = [] && binderlists = []);
+ subst
+
+let split_by_type_pat ?loc ids subst =
+ let bind id scl l s =
+ match l with
+ | [] -> assert false
+ | a::l -> l, Id.Map.add id (a,scl) s in
+ let (terms,termlists),subst =
+ List.fold_left (fun ((terms,termlists),(terms',termlists')) (id,(scl,typ)) ->
+ match typ with
+ | NtnTypeConstr | NtnTypeBinder _ ->
+ let terms,terms' = bind id scl terms terms' in
+ (terms,termlists),(terms',termlists')
+ | NtnTypeConstrList ->
+ let termlists,termlists' = bind id scl termlists termlists' in
+ (terms,termlists),(terms',termlists')
+ | NtnTypeBinderList -> error_invalid_pattern_notation ?loc ())
+ (subst,(Id.Map.empty,Id.Map.empty)) ids in
+ assert (terms = [] && termlists = []);
+ subst
let make_subst ids l =
let fold accu (id, scl) a = Id.Map.add id (a, scl) accu in
List.fold_left2 fold Id.Map.empty ids l
-let intern_notation intern env lvar loc ntn fullargs =
- let ntn,(args,argslist,bll as fullargs) = contract_notation ntn fullargs in
+let intern_notation intern env ntnvars loc ntn fullargs =
+ (* Adjust to parsing of { } *)
+ let ntn,fullargs = contract_curly_brackets ntn fullargs in
+ (* Recover interpretation { } *)
let ((ids,c),df) = interp_notation ?loc ntn (env.tmp_scope,env.scopes) in
Dumpglob.dump_notation_location (ntn_loc ?loc fullargs ntn) ntn df;
- let ids,idsl,idsbl = split_by_type ids in
- let terms = make_subst ids args in
- let termlists = make_subst idsl argslist in
- let binders = make_subst idsbl bll in
- instantiate_notation_constr loc intern lvar
- (terms, termlists, binders) (Id.Map.empty, env) c
+ (* Dispatch parsing substitution to an interpretation substitution *)
+ let subst = split_by_type ids fullargs in
+ (* Instantiate the notation *)
+ instantiate_notation_constr loc intern intern_cases_pattern_as_binder ntnvars subst (Id.Map.empty, env) c
(**********************************************************************)
(* Discriminating between bound variables and global references *)
@@ -746,17 +886,17 @@ let gvar (loc, id) us = match us with
user_err ?loc (str "Variable " ++ Id.print id ++
str " cannot have a universe instance")
-let intern_var genv (ltacvars,ntnvars) namedctx loc id us =
+let intern_var env (ltacvars,ntnvars) namedctx loc id us =
(* Is [id] a notation variable *)
if Id.Map.mem id ntnvars then
begin
- if not (Id.Map.mem id genv.impls) then set_var_scope ?loc id true genv ntnvars;
+ if not (Id.Map.mem id env.impls) then set_var_scope ?loc id true (env.tmp_scope,env.scopes) ntnvars;
gvar (loc,id) us, [], [], []
end
else
(* Is [id] registered with implicit arguments *)
try
- let ty,expl_impls,impls,argsc = Id.Map.find id genv.impls in
+ let ty,expl_impls,impls,argsc = Id.Map.find id env.impls in
let expl_impls = List.map
(fun id -> CAst.make ?loc @@ CRef (Ident (loc,id),None), Some (loc,ExplByName id)) expl_impls in
let tys = string_of_ty ty in
@@ -764,7 +904,7 @@ let intern_var genv (ltacvars,ntnvars) namedctx loc id us =
gvar (loc,id) us, make_implicits_list impls, argsc, expl_impls
with Not_found ->
(* Is [id] bound in current term or is an ltac var bound to constr *)
- if Id.Set.mem id genv.ids || Id.Set.mem id ltacvars.ltac_vars
+ if Id.Set.mem id env.ids || Id.Set.mem id ltacvars.ltac_vars
then
gvar (loc,id) us, [], [], []
else if Id.equal id ldots_var
@@ -849,7 +989,7 @@ let glob_sort_of_level (level: Misctypes.glob_level) : Misctypes.glob_sort =
| Misctypes.GType info -> Misctypes.GType [sort_info_of_level_info info]
(* Is it a global reference or a syntactic definition? *)
-let intern_qualid loc qid intern env lvar us args =
+let intern_qualid loc qid intern env ntnvars us args =
match intern_extended_global_of_qualid (loc,qid) with
| TrueGlobal ref -> (DAst.make ?loc @@ GRef (ref, us)), true, args
| SynDef sp ->
@@ -859,10 +999,10 @@ let intern_qualid loc qid intern env lvar us args =
let args1,args2 = List.chop nids args in
check_no_explicitation args1;
let terms = make_subst ids (List.map fst args1) in
- let subst = (terms, Id.Map.empty, Id.Map.empty) in
+ let subst = (terms, Id.Map.empty, Id.Map.empty, Id.Map.empty) in
let infos = (Id.Map.empty, env) in
let projapp = match c with NRef _ -> true | _ -> false in
- let c = instantiate_notation_constr loc intern lvar subst infos c in
+ let c = instantiate_notation_constr loc intern intern_cases_pattern_as_binder ntnvars subst infos c in
let loc = c.CAst.loc in
let err () =
user_err ?loc (str "Notation " ++ pr_qualid qid
@@ -888,8 +1028,8 @@ let intern_qualid loc qid intern env lvar us args =
c, projapp, args2
(* Rule out section vars since these should have been found by intern_var *)
-let intern_non_secvar_qualid loc qid intern env lvar us args =
- let c, _, _ as r = intern_qualid loc qid intern env lvar us args in
+let intern_non_secvar_qualid loc qid intern env ntnvars us args =
+ let c, _, _ as r = intern_qualid loc qid intern env ntnvars us args in
match DAst.get c with
| GRef (VarRef _, _) -> raise Not_found
| _ -> r
@@ -929,11 +1069,11 @@ let interp_reference vars r =
(** Private internalization patterns *)
type 'a raw_cases_pattern_expr_r =
- | RCPatAlias of 'a raw_cases_pattern_expr * Id.t
+ | RCPatAlias of 'a raw_cases_pattern_expr * Name.t Loc.located
| RCPatCstr of Globnames.global_reference
* 'a raw_cases_pattern_expr list * 'a raw_cases_pattern_expr list
(** [RCPatCstr (loc, c, l1, l2)] represents ((@c l1) l2) *)
- | RCPatAtom of Id.t option
+ | RCPatAtom of (Id.t Loc.located * (Notation_term.tmp_scope_name option * Notation_term.scope_name list)) option
| RCPatOr of 'a raw_cases_pattern_expr list
and 'a raw_cases_pattern_expr = ('a raw_cases_pattern_expr_r, 'a) DAst.t
@@ -993,7 +1133,7 @@ let check_number_of_pattern loc n l =
if not (Int.equal n p) then raise (InternalizationError (loc,BadPatternsNumber (n,p)))
let check_or_pat_variables loc ids idsl =
- if List.exists (fun ids' -> not (List.eq_set Id.equal ids ids')) idsl then
+ if List.exists (fun ids' -> not (List.eq_set (fun (loc,id) (_,id') -> Id.equal id id') ids ids')) idsl then
user_err ?loc (str
"The components of this disjunctive pattern must bind the same variables.")
@@ -1059,7 +1199,7 @@ let add_implicits_check_length fail nargs nargs_with_letin impls_st len_pl1 pl2
then let (b,out) = aux i (q,[]) in (b,(DAst.make @@ RCPatAtom None)::out)
else fail (remaining_args (len_pl1+i) il)
|imp::q,(hh::tt as l) -> if is_status_implicit imp
- then let (b,out) = aux i (q,l) in (b,(DAst.make @@ RCPatAtom(None))::out)
+ then let (b,out) = aux i (q,l) in (b,(DAst.make @@ RCPatAtom None)::out)
else let (b,out) = aux (succ i) (q,tt) in (b,hh::out)
in aux 0 (impl_list,pl2)
@@ -1232,7 +1372,7 @@ let sort_fields ~complete loc fields completer =
(** {6 Manage multiple aliases} *)
type alias = {
- alias_ids : Id.t list;
+ alias_ids : Id.t Loc.located list;
alias_map : Id.t Id.Map.t;
}
@@ -1243,17 +1383,20 @@ let empty_alias = {
(* [merge_aliases] returns the sets of all aliases encountered at this
point and a substitution mapping extra aliases to the first one *)
-let merge_aliases aliases id =
- let alias_ids = aliases.alias_ids @ [id] in
+let merge_aliases aliases (loc,na) =
+ match na with
+ | Anonymous -> aliases
+ | Name id ->
+ let alias_ids = aliases.alias_ids @ [loc,id] in
let alias_map = match aliases.alias_ids with
| [] -> aliases.alias_map
- | id' :: _ -> Id.Map.add id id' aliases.alias_map
+ | (_,id') :: _ -> Id.Map.add id id' aliases.alias_map
in
{ alias_ids; alias_map; }
let alias_of als = match als.alias_ids with
| [] -> Anonymous
-| id :: _ -> Name id
+| (_,id) :: _ -> Name id
(** {6 Expanding notations }
@@ -1271,6 +1414,8 @@ let is_zero s =
let merge_subst s1 s2 = Id.Map.fold Id.Map.add s1 s2
let product_of_cases_patterns aliases idspl =
+ (* each [pl] is a disjunction of patterns over common identifiers [ids] *)
+ (* We stepwise build a disjunction of patterns [ptaill] over common [ids'] *)
List.fold_right (fun (ids,pl) (ids',ptaill) ->
(ids @ ids',
(* Cartesian prod of the or-pats for the nth arg and the tail args *)
@@ -1281,7 +1426,7 @@ let product_of_cases_patterns aliases idspl =
let rec subst_pat_iterator y t = DAst.(map (function
| RCPatAtom id as p ->
- begin match id with Some x when Id.equal x y -> DAst.get t | _ -> p end
+ begin match id with Some ((_,x),_) when Id.equal x y -> DAst.get t | _ -> p end
| RCPatCstr (id,l1,l2) ->
RCPatCstr (id,List.map (subst_pat_iterator y t) l1,
List.map (subst_pat_iterator y t) l2)
@@ -1310,13 +1455,16 @@ let drop_notations_pattern looked_for genv =
if top then looked_for else function ConstructRef _ -> () | _ -> raise Not_found
in
(** [rcp_of_glob] : from [glob_constr] to [raw_cases_pattern_expr] *)
- let rec rcp_of_glob x = DAst.(map (function
- | GVar id -> RCPatAtom (Some id)
+ let rec rcp_of_glob scopes x = DAst.(map (function
+ | GVar id -> RCPatAtom (Some ((x.CAst.loc,id),scopes))
| GHole (_,_,_) -> RCPatAtom (None)
| GRef (g,_) -> RCPatCstr (g,[],[])
| GApp (r, l) ->
begin match DAst.get r with
- | GRef (g,_) -> RCPatCstr (g, List.map rcp_of_glob l,[])
+ | GRef (g,_) ->
+ let allscs = find_arguments_scope g in
+ let allscs = simple_adjust_scopes (List.length l) allscs in (* TO CHECK *)
+ RCPatCstr (g, List.map2 (fun sc a -> rcp_of_glob (sc,snd scopes) a) allscs l,[])
| _ ->
CErrors.anomaly Pp.(str "Invalid return pattern from Notation.interp_prim_token_cases_pattern_expr.")
end
@@ -1396,27 +1544,25 @@ let drop_notations_pattern looked_for genv =
| CPatNotation ("- _",([a],[]),[]) when is_non_zero_pat a ->
let p = match a.CAst.v with CPatPrim (Numeral (p, _)) -> p | _ -> assert false in
let pat, _df = Notation.interp_prim_token_cases_pattern_expr ?loc (ensure_kind false loc) (Numeral (p,false)) scopes in
- rcp_of_glob pat
+ rcp_of_glob scopes pat
| CPatNotation ("( _ )",([a],[]),[]) ->
in_pat top scopes a
- | CPatNotation (ntn, fullargs,extrargs) ->
- let ntn,(args,argsl as fullargs) = contract_pat_notation ntn fullargs in
+ | CPatNotation (ntn,fullargs,extrargs) ->
+ let ntn,(terms,termlists) = contract_curly_brackets_pat ntn fullargs in
let ((ids',c),df) = Notation.interp_notation ?loc ntn scopes in
- let (ids',idsl',_) = split_by_type ids' in
+ let (terms,termlists) = split_by_type_pat ?loc ids' (terms,termlists) in
Dumpglob.dump_notation_location (patntn_loc ?loc fullargs ntn) ntn df;
- let substlist = make_subst idsl' argsl in
- let subst = make_subst ids' args in
- in_not top loc scopes (subst,substlist) extrargs c
+ in_not top loc scopes (terms,termlists) extrargs c
| CPatDelimiters (key, e) ->
in_pat top (None,find_delimiters_scope ?loc key::snd scopes) e
| CPatPrim p ->
let pat, _df = Notation.interp_prim_token_cases_pattern_expr ?loc (test_kind false) p scopes in
- rcp_of_glob pat
- | CPatAtom Some id ->
+ rcp_of_glob scopes pat
+ | CPatAtom (Some id) ->
begin
match drop_syndef top scopes id [] with
| Some (a,b,c) -> DAst.make ?loc @@ RCPatCstr (a, b, c)
- | None -> DAst.make ?loc @@ RCPatAtom (Some (find_pattern_variable id))
+ | None -> DAst.make ?loc @@ RCPatAtom (Some ((loc,find_pattern_variable id),scopes))
end
| CPatAtom None -> DAst.make ?loc @@ RCPatAtom None
| CPatOr pl -> DAst.make ?loc @@ RCPatOr (List.map (in_pat top scopes) pl)
@@ -1446,7 +1592,7 @@ let drop_notations_pattern looked_for genv =
let (a,(scopt,subscopes)) = Id.Map.find id subst in
in_pat top (scopt,subscopes@snd scopes) a
with Not_found ->
- if Id.equal id ldots_var then DAst.make ?loc @@ RCPatAtom (Some id) else
+ if Id.equal id ldots_var then DAst.make ?loc @@ RCPatAtom (Some ((loc,id),scopes)) else
anomaly (str "Unbound pattern notation variable: " ++ Id.print id ++ str ".")
end
| NRef g ->
@@ -1459,7 +1605,7 @@ let drop_notations_pattern looked_for genv =
let pl = List.map2 (fun x -> in_not false loc (x,snd scopes) fullsubst []) argscs1 pl in
let pl = add_local_defs_and_check_length loc genv g pl args in
DAst.make ?loc @@ RCPatCstr (g, pl @ List.map (in_pat false scopes) args, [])
- | NList (x,y,iter,terminator,lassoc) ->
+ | NList (x,y,iter,terminator,revert) ->
if not (List.is_empty args) then user_err ?loc
(strbrk "Application of arguments to a recursive notation not supported in patterns.");
(try
@@ -1470,7 +1616,7 @@ let drop_notations_pattern looked_for genv =
let nsubst = Id.Map.add y (a, (scopt, subscopes)) subst in
let u = in_not false loc scopes (nsubst, substlist) [] iter in
subst_pat_iterator ldots_var t u)
- (if lassoc then List.rev l else l) termin
+ (if revert then List.rev l else l) termin
with Not_found ->
anomaly (Pp.str "Inconsistent substitution of recursive notation."))
| NHole _ ->
@@ -1479,9 +1625,9 @@ let drop_notations_pattern looked_for genv =
| t -> error_invalid_pattern_notation ?loc ()
in in_pat true
-let rec intern_pat genv aliases pat =
+let rec intern_pat genv ntnvars aliases pat =
let intern_cstr_with_all_args loc c with_letin idslpl1 pl2 =
- let idslpl2 = List.map (intern_pat genv empty_alias) pl2 in
+ let idslpl2 = List.map (intern_pat genv ntnvars empty_alias) pl2 in
let (ids',pll) = product_of_cases_patterns aliases (idslpl1@idslpl2) in
let pl' = List.map (fun (asubst,pl) ->
(asubst, DAst.make ?loc @@ PatCstr (c,chop_params_pattern loc (fst c) pl with_letin,alias_of aliases))) pll in
@@ -1490,7 +1636,7 @@ let rec intern_pat genv aliases pat =
match DAst.get pat with
| RCPatAlias (p, id) ->
let aliases' = merge_aliases aliases id in
- intern_pat genv aliases' p
+ intern_pat genv ntnvars aliases' p
| RCPatCstr (head, expl_pl, pl) ->
if !asymmetric_patterns then
let len = if List.is_empty expl_pl then Some (List.length pl) else None in
@@ -1503,29 +1649,30 @@ let rec intern_pat genv aliases pat =
let with_letin, pl2 =
add_implicits_check_constructor_length genv loc c (List.length idslpl1 + List.length expl_pl) pl in
intern_cstr_with_all_args loc c with_letin idslpl1 (expl_pl@pl2)
- | RCPatAtom (Some id) ->
- let aliases = merge_aliases aliases id in
- (aliases.alias_ids,[aliases.alias_map, DAst.make ?loc @@ PatVar (alias_of aliases)])
+ | RCPatAtom (Some ((loc,id),scopes)) ->
+ let aliases = merge_aliases aliases (loc,Name id) in
+ set_var_scope ?loc id false scopes ntnvars;
+ (aliases.alias_ids,[aliases.alias_map, DAst.make ?loc @@ PatVar (alias_of aliases)]) (* TO CHECK: aura-t-on id? *)
| RCPatAtom (None) ->
let { alias_ids = ids; alias_map = asubst; } = aliases in
(ids, [asubst, DAst.make ?loc @@ PatVar (alias_of aliases)])
| RCPatOr pl ->
assert (not (List.is_empty pl));
- let pl' = List.map (intern_pat genv aliases) pl in
+ let pl' = List.map (intern_pat genv ntnvars aliases) pl in
let (idsl,pl') = List.split pl' in
let ids = List.hd idsl in
check_or_pat_variables loc ids (List.tl idsl);
(ids,List.flatten pl')
-let intern_cases_pattern genv scopes aliases pat =
- intern_pat genv aliases
+let intern_cases_pattern genv ntnvars scopes aliases pat =
+ intern_pat genv ntnvars aliases
(drop_notations_pattern (function ConstructRef _ -> () | _ -> raise Not_found) genv scopes pat)
let _ =
intern_cases_pattern_fwd :=
- fun scopes p -> intern_cases_pattern (Global.env ()) scopes empty_alias p
+ fun ntnvars scopes p -> intern_cases_pattern (Global.env ()) ntnvars scopes empty_alias p
-let intern_ind_pattern genv scopes pat =
+let intern_ind_pattern genv ntnvars scopes pat =
let no_not =
try
drop_notations_pattern (function (IndRef _ | ConstructRef _) -> () | _ -> raise Not_found) genv scopes pat
@@ -1537,10 +1684,9 @@ let intern_ind_pattern genv scopes pat =
let c = (function IndRef ind -> ind | _ -> error_bad_inductive_type ?loc) head in
let with_letin, pl2 = add_implicits_check_ind_length genv loc c
(List.length expl_pl) pl in
- let idslpl1 = List.rev_map (intern_pat genv empty_alias) expl_pl in
- let idslpl2 = List.map (intern_pat genv empty_alias) pl2 in
+ let idslpl = List.map (intern_pat genv ntnvars empty_alias) (expl_pl@pl2) in
(with_letin,
- match product_of_cases_patterns empty_alias (List.rev_append idslpl1 idslpl2) with
+ match product_of_cases_patterns empty_alias idslpl with
| _,[_,pl] -> (c,chop_params_pattern loc c pl with_letin)
| _ -> error_bad_inductive_type ?loc)
| x -> error_bad_inductive_type ?loc
@@ -1697,24 +1843,25 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c =
Array.map (fun (bl,_,_) -> bl) idl,
Array.map (fun (_,ty,_) -> ty) idl,
Array.map (fun (_,_,bd) -> bd) idl)
- | CProdN ([],c2) ->
- intern_type env c2
- | CProdN ((nal,bk,ty)::bll,c2) ->
- iterate_prod ?loc env bk ty (CAst.make ?loc @@ CProdN (bll, c2)) nal
+ | CProdN (bl,c2) ->
+ let (env',bl) = List.fold_left intern_local_binder (env,[]) bl in
+ expand_binders ?loc mkGProd bl (intern_type env' c2)
| CLambdaN ([],c2) ->
+ (* Such a term is built sometimes: it should not change scope *)
intern env c2
- | CLambdaN ((nal,bk,ty)::bll,c2) ->
- iterate_lam loc (reset_tmp_scope env) bk ty (CAst.make ?loc @@ CLambdaN (bll, c2)) nal
+ | CLambdaN (bl,c2) ->
+ let (env',bl) = List.fold_left intern_local_binder (reset_tmp_scope env,[]) bl in
+ expand_binders ?loc mkGLambda bl (intern env' c2)
| CLetIn (na,c1,t,c2) ->
let inc1 = intern (reset_tmp_scope env) c1 in
let int = Option.map (intern_type env) t in
DAst.make ?loc @@
GLetIn (snd na, inc1, int,
intern (push_name_env ntnvars (impls_term_list inc1) env na) c2)
- | CNotation ("- _", ([a],[],[])) when is_non_zero a ->
+ | CNotation ("- _", ([a],[],[],[])) when is_non_zero a ->
let p = match a.CAst.v with CPrim (Numeral (p, _)) -> p | _ -> assert false in
intern env (CAst.make ?loc @@ CPrim (Numeral (p,false)))
- | CNotation ("( _ )",([a],[],[])) -> intern env a
+ | CNotation ("( _ )",([a],[],[],[])) -> intern env a
| CNotation (ntn,args) ->
intern_notation intern env ntnvars loc ntn args
| CGeneralization (b,a,c) ->
@@ -1746,8 +1893,8 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c =
| CRef (ref,us) ->
intern_applied_reference intern env
(Environ.named_context globalenv) lvar us args ref
- | CNotation (ntn,([],[],[])) ->
- let c = intern_notation intern env ntnvars loc ntn ([],[],[]) in
+ | CNotation (ntn,([],[],[],[])) ->
+ let c = intern_notation intern env ntnvars loc ntn ([],[],[],[]) in
let x, impl, scopes, l = find_appl_head_data c in
(x,impl,scopes,l), args
| _ -> (intern env f,[],[],[]), args in
@@ -1849,6 +1996,8 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c =
| None -> None
| Some gen ->
let (ltacvars, ntnvars) = lvar in
+ (* Preventively declare notation variables in ltac as non-bindings *)
+ Id.Map.iter (fun x (used_as_binder,_,_) -> used_as_binder := false) ntnvars;
let ntnvars = Id.Map.domain ntnvars in
let extra = ltacvars.ltac_extra in
let lvars = Id.Set.union ltacvars.ltac_bound ltacvars.ltac_vars in
@@ -1899,7 +2048,7 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c =
(* Expands a multiple pattern into a disjunction of multiple patterns *)
and intern_multiple_pattern env n pl =
- let idsl_pll = List.map (intern_cases_pattern globalenv (None,env.scopes) empty_alias) pl in
+ let idsl_pll = List.map (intern_cases_pattern globalenv ntnvars (None,env.scopes) empty_alias) pl in
let loc = loc_of_multiple_pattern pl in
check_number_of_pattern loc n pl;
product_of_cases_patterns empty_alias idsl_pll
@@ -1917,6 +2066,7 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c =
and intern_eqn n env (loc,(lhs,rhs)) =
let eqn_ids,pll = intern_disjunctive_multiple_pattern env loc n lhs in
(* Linearity implies the order in ids is irrelevant *)
+ let eqn_ids = List.map snd eqn_ids in
check_linearity lhs eqn_ids;
let env_ids = List.fold_right Id.Set.add eqn_ids env.ids in
List.map (fun (asubst,pl) ->
@@ -1938,7 +2088,7 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c =
(* the "in" part *)
let match_td,typ = match t with
| Some t ->
- let with_letin,(ind,l) = intern_ind_pattern globalenv (None,env.scopes) t in
+ let with_letin,(ind,l) = intern_ind_pattern globalenv ntnvars (None,env.scopes) t in
let (mib,mip) = Inductive.lookup_mind_specif globalenv ind in
let nparams = (List.length (mib.Declarations.mind_params_ctxt)) in
(* for "in Vect n", we answer (["n","n"],[(loc,"n")])
@@ -1979,14 +2129,6 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c =
[], None in
(tm',(snd na,typ)), extra_id, match_td
- and iterate_prod ?loc env bk ty body nal =
- let env, bl = intern_assumption intern ntnvars env nal bk ty in
- it_mkGProd ?loc bl (intern_type env body)
-
- and iterate_lam loc env bk ty body nal =
- let env, bl = intern_assumption intern ntnvars env nal bk ty in
- it_mkGLambda ?loc bl (intern env body)
-
and intern_impargs c env l subscopes args =
let eargs, rargs = extract_explicit_arg l args in
if !parsing_explicit then
@@ -2089,7 +2231,7 @@ let intern_type env c = intern_gen IsType env c
let intern_pattern globalenv patt =
try
- intern_cases_pattern globalenv (None,[]) empty_alias patt
+ intern_cases_pattern globalenv Id.Map.empty (None,[]) empty_alias patt
with
InternalizationError (loc,e) ->
user_err ?loc ~hdr:"internalize" (explain_internalization_error e)
@@ -2160,7 +2302,8 @@ let intern_constr_pattern env ?(as_type=false) ?(ltacvars=empty_ltac_sign) c =
let interp_notation_constr env ?(impls=empty_internalization_env) nenv a =
(* [vl] is intended to remember the scope of the free variables of [a] *)
- let vl = Id.Map.map (fun typ -> (ref true, ref None, typ)) nenv.ninterp_var_type in
+ let vl = Id.Map.map (fun typ -> (ref false, ref None, typ)) nenv.ninterp_var_type in
+ let impls = Id.Map.fold (fun id _ impls -> Id.Map.remove id impls) nenv.ninterp_var_type impls in
let c = internalize (Global.env()) {ids = extract_ids env; unb = false;
tmp_scope = None; scopes = []; impls = impls}
false (empty_ltac_sign, vl) a in
@@ -2169,8 +2312,9 @@ let interp_notation_constr env ?(impls=empty_internalization_env) nenv a =
(* Splits variables into those that are binding, bound, or both *)
(* binding and bound *)
let out_scope = function None -> None,[] | Some (a,l) -> a,l in
- let vars = Id.Map.map (fun (isonlybinding, sc, typ) ->
- (!isonlybinding, out_scope !sc, typ)) vl in
+ let unused = match reversible with NonInjective ids -> ids | _ -> [] in
+ let vars = Id.Map.mapi (fun id (used_as_binder, sc, typ) ->
+ (!used_as_binder && not (List.mem_f Id.equal id unused), out_scope !sc)) vl in
(* Returns [a] and the ordered list of variables with their scopes *)
vars, a, reversible
diff --git a/interp/constrintern.mli b/interp/constrintern.mli
index 632b423b0..f09b17a49 100644
--- a/interp/constrintern.mli
+++ b/interp/constrintern.mli
@@ -87,7 +87,7 @@ val intern_gen : typing_constraint -> env ->
constr_expr -> glob_constr
val intern_pattern : env -> cases_pattern_expr ->
- Id.t list * (Id.t Id.Map.t * cases_pattern) list
+ Id.t Loc.located list * (Id.t Id.Map.t * cases_pattern) list
val intern_context : bool -> env -> internalization_env -> local_binder_expr list -> internalization_env * glob_decl list
@@ -184,8 +184,7 @@ val global_reference_in_absolute_module : DirPath.t -> Id.t -> Globnames.global_
guaranteed to have the same domain as the input one. *)
val interp_notation_constr : env -> ?impls:internalization_env ->
notation_interp_env -> constr_expr ->
- (bool * subscopes * notation_var_internalization_type) Id.Map.t *
- notation_constr * reversibility_flag
+ (bool * subscopes) Id.Map.t * notation_constr * reversibility_status
(** Globalization options *)
val parsing_explicit : bool ref
diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml
index 519f2480b..a838d7106 100644
--- a/interp/implicit_quantifiers.ml
+++ b/interp/implicit_quantifiers.ml
@@ -93,8 +93,8 @@ let free_vars_of_constr_expr c ?(bound=Id.Set.empty) l =
in
let rec aux bdvars l c = match CAst.(c.v) with
| CRef (Ident (loc,id),_) -> found loc id bdvars l
- | CNotation ("{ _ : _ | _ }", ({ CAst.v = CRef (Ident (_, id),_) } :: _, [], [])) when not (Id.Set.mem id bdvars) ->
- Constrexpr_ops.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux (Id.Set.add id bdvars) l c
+ | CNotation ("{ _ : _ | _ }", ({ CAst.v = CRef (Ident (_, id),_) } :: _, [], [], [])) when not (Id.Set.mem id bdvars) ->
+ Constrexpr_ops.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux (Id.Set.add id bdvars) l c
| _ -> Constrexpr_ops.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux bdvars l c
in aux bound l c
diff --git a/interp/notation.ml b/interp/notation.ml
index 94ce2a6c8..ea7ef21b1 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -82,18 +82,35 @@ let parenRelation_eq t1 t2 = match t1, t2 with
| Prec l1, Prec l2 -> Int.equal l1 l2
| _ -> false
-let notation_var_internalization_type_eq v1 v2 = match v1, v2 with
-| NtnInternTypeConstr, NtnInternTypeConstr -> true
-| NtnInternTypeBinder, NtnInternTypeBinder -> true
-| NtnInternTypeIdent, NtnInternTypeIdent -> true
-| (NtnInternTypeConstr | NtnInternTypeBinder | NtnInternTypeIdent), _ -> false
-
-let level_eq (l1, t1, u1) (l2, t2, u2) =
- let tolerability_eq (i1, r1) (i2, r2) =
- Int.equal i1 i2 && parenRelation_eq r1 r2
- in
+open Extend
+
+let production_level_eq l1 l2 = true (* (l1 = l2) *)
+
+let production_position_eq pp1 pp2 = true (* pp1 = pp2 *) (* match pp1, pp2 with
+| NextLevel, NextLevel -> true
+| NumLevel n1, NumLevel n2 -> Int.equal n1 n2
+| (NextLevel | NumLevel _), _ -> false *)
+
+let constr_entry_key_eq eq v1 v2 = match v1, v2 with
+| ETName, ETName -> true
+| ETReference, ETReference -> true
+| ETBigint, ETBigint -> true
+| ETBinder b1, ETBinder b2 -> b1 == b2
+| ETConstr lev1, ETConstr lev2 -> eq lev1 lev2
+| ETConstrAsBinder (bk1,lev1), ETConstrAsBinder (bk2,lev2) -> eq lev1 lev2 && bk1 = bk2
+| ETPattern (b1,n1), ETPattern (b2,n2) -> b1 = b2 && Option.equal Int.equal n1 n2
+| ETOther (s1,s1'), ETOther (s2,s2') -> String.equal s1 s2 && String.equal s1' s2'
+| (ETName | ETReference | ETBigint | ETBinder _ | ETConstr _ | ETPattern _ | ETOther _ | ETConstrAsBinder _), _ -> false
+
+let level_eq_gen strict (l1, t1, u1) (l2, t2, u2) =
+ let tolerability_eq (i1, r1) (i2, r2) = Int.equal i1 i2 && parenRelation_eq r1 r2 in
+ let prod_eq (l1,pp1) (l2,pp2) =
+ if strict then production_level_eq l1 l2 && production_position_eq pp1 pp2
+ else production_level_eq l1 l2 in
Int.equal l1 l2 && List.equal tolerability_eq t1 t2
- && List.equal notation_var_internalization_type_eq u1 u2
+ && List.equal (constr_entry_key_eq prod_eq) u1 u2
+
+let level_eq = level_eq_gen false
let declare_scope scope =
try let _ = String.Map.find scope !scope_map in ()
@@ -292,7 +309,7 @@ let cases_pattern_key c = match DAst.get c with
let notation_constr_key = function (* Rem: NApp(NRef ref,[]) stands for @ref *)
| NApp (NRef ref,args) -> RefKey(canonical_gr ref), Some (List.length args)
| NList (_,_,NApp (NRef ref,args),_,_)
- | NBinderList (_,_,NApp (NRef ref,args),_) ->
+ | NBinderList (_,_,NApp (NRef ref,args),_,_) ->
RefKey (canonical_gr ref), Some (List.length args)
| NRef ref -> RefKey(canonical_gr ref), None
| NApp (_,args) -> Oth, Some (List.length args)
@@ -609,12 +626,18 @@ let availability_of_prim_token n printer_scope local_scopes =
let pair_eq f g (x1, y1) (x2, y2) = f x1 x2 && g y1 y2
+let notation_binder_source_eq s1 s2 = match s1, s2 with
+| NtnParsedAsIdent, NtnParsedAsIdent -> true
+| NtnParsedAsPattern b1, NtnParsedAsPattern b2 -> b1 = b2
+| NtnBinderParsedAsConstr bk1, NtnBinderParsedAsConstr bk2 -> bk1 = bk2
+| (NtnParsedAsIdent | NtnParsedAsPattern _ | NtnBinderParsedAsConstr _), _ -> false
+
let ntpe_eq t1 t2 = match t1, t2 with
| NtnTypeConstr, NtnTypeConstr -> true
-| NtnTypeOnlyBinder, NtnTypeOnlyBinder -> true
+| NtnTypeBinder s1, NtnTypeBinder s2 -> notation_binder_source_eq s1 s2
| NtnTypeConstrList, NtnTypeConstrList -> true
| NtnTypeBinderList, NtnTypeBinderList -> true
-| (NtnTypeConstr | NtnTypeOnlyBinder | NtnTypeConstrList | NtnTypeBinderList), _ -> false
+| (NtnTypeConstr | NtnTypeBinder _ | NtnTypeConstrList | NtnTypeBinderList), _ -> false
let var_attributes_eq (_, (sc1, tp1)) (_, (sc2, tp2)) =
pair_eq (Option.equal String.equal) (List.equal String.equal) sc1 sc2 &&
@@ -926,8 +949,63 @@ let factorize_entries = function
(ntn,[c],[]) l in
(ntn,l_of_ntn)::rest
+type symbol_token = WhiteSpace of int | String of string
+
+let split_notation_string str =
+ let push_token beg i l =
+ if Int.equal beg i then l else
+ let s = String.sub str beg (i - beg) in
+ String s :: l
+ in
+ let push_whitespace beg i l =
+ if Int.equal beg i then l else WhiteSpace (i-beg) :: l
+ in
+ let rec loop beg i =
+ if i < String.length str then
+ if str.[i] == ' ' then
+ push_token beg i (loop_on_whitespace (i+1) (i+1))
+ else
+ loop beg (i+1)
+ else
+ push_token beg i []
+ and loop_on_whitespace beg i =
+ if i < String.length str then
+ if str.[i] != ' ' then
+ push_whitespace beg i (loop i (i+1))
+ else
+ loop_on_whitespace beg (i+1)
+ else
+ push_whitespace beg i []
+ in
+ loop 0 0
+
+let rec raw_analyze_notation_tokens = function
+ | [] -> []
+ | String ".." :: sl -> NonTerminal Notation_ops.ldots_var :: raw_analyze_notation_tokens sl
+ | String "_" :: _ -> user_err Pp.(str "_ must be quoted.")
+ | String x :: sl when Id.is_valid x ->
+ NonTerminal (Names.Id.of_string x) :: raw_analyze_notation_tokens sl
+ | String s :: sl ->
+ Terminal (String.drop_simple_quotes s) :: raw_analyze_notation_tokens sl
+ | WhiteSpace n :: sl ->
+ Break n :: raw_analyze_notation_tokens sl
+
+let decompose_raw_notation ntn = raw_analyze_notation_tokens (split_notation_string ntn)
+
+let possible_notations ntn =
+ (* We collect the possible interpretations of a notation string depending on whether it is
+ in "x 'U' y" or "_ U _" format *)
+ let toks = split_notation_string ntn in
+ if List.exists (function String "_" -> true | _ -> false) toks then
+ (* Only "_ U _" format *)
+ [ntn]
+ else
+ let ntn' = make_notation_key (raw_analyze_notation_tokens toks) in
+ if String.equal ntn ntn' then (* Only symbols *) [ntn] else [ntn;ntn']
+
let browse_notation strict ntn map =
- let find ntn' =
+ let ntns = possible_notations ntn in
+ let find ntn' ntn =
if String.contains ntn ' ' then String.equal ntn ntn'
else
let toks = decompose_notation_key ntn' in
@@ -940,7 +1018,7 @@ let browse_notation strict ntn map =
String.Map.fold
(fun scope_name sc ->
String.Map.fold (fun ntn { not_interp = (_, r); not_location = df } l ->
- if find ntn then (ntn,(scope_name,r,df))::l else l) sc.notations)
+ if List.exists (find ntn) ntns then (ntn,(scope_name,r,df))::l else l) sc.notations)
map [] in
List.sort (fun x y -> String.compare (fst x) (fst y)) l
diff --git a/interp/notation.mli b/interp/notation.mli
index 7d055571c..a4c79d6d3 100644
--- a/interp/notation.mli
+++ b/interp/notation.mli
@@ -176,16 +176,20 @@ val scope_class_of_class : Classops.cl_typ -> scope_class
(** Building notation key *)
type symbol =
- | Terminal of string
- | NonTerminal of Id.t
- | SProdList of Id.t * symbol list
- | Break of int
+ | Terminal of string (* an expression including symbols or a simply-quoted ident, e.g. "'U'" or "!" *)
+ | NonTerminal of Id.t (* an identifier "x" *)
+ | SProdList of Id.t * symbol list (* an expression "x sep .. sep y", remembering x (or y) and sep *)
+ | Break of int (* a sequence of blanks > 1, e.g. " " *)
val symbol_eq : symbol -> symbol -> bool
+(** Make/decompose a notation of the form "_ U _" *)
val make_notation_key : symbol list -> notation
val decompose_notation_key : notation -> symbol list
+(** Decompose a notation of the form "a 'U' b" *)
+val decompose_raw_notation : string -> symbol list
+
(** Prints scopes (expects a pure aconstr printer) *)
val pr_scope_class : scope_class -> Pp.t
val pr_scope : (glob_constr -> Pp.t) -> scope_name -> Pp.t
diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml
index 326d05cba..c65f4785e 100644
--- a/interp/notation_ops.ml
+++ b/interp/notation_ops.ml
@@ -42,9 +42,9 @@ let rec eq_notation_constr (vars1,vars2 as vars) t1 t2 = match t1, t2 with
Name.equal na1 na2 && (eq_notation_constr vars) t1 t2 && (eq_notation_constr vars) u1 u2
| NProd (na1, t1, u1), NProd (na2, t2, u2) ->
Name.equal na1 na2 && (eq_notation_constr vars) t1 t2 && (eq_notation_constr vars) u1 u2
-| NBinderList (i1, j1, t1, u1), NBinderList (i2, j2, t2, u2) ->
+| NBinderList (i1, j1, t1, u1, b1), NBinderList (i2, j2, t2, u2, b2) ->
Id.equal i1 i2 && Id.equal j1 j2 && (eq_notation_constr vars) t1 t2 &&
- (eq_notation_constr vars) u1 u2
+ (eq_notation_constr vars) u1 u2 && b1 == b2
| NLetIn (na1, b1, t1, u1), NLetIn (na2, b2, t2, u2) ->
Name.equal na1 na2 && eq_notation_constr vars b1 b2 &&
Option.equal (eq_notation_constr vars) t1 t2 && (eq_notation_constr vars) u1 u2
@@ -101,13 +101,24 @@ let name_to_ident = function
let to_id g e id = let e,na = g e (Name id) in e,name_to_ident na
+let product_of_cases_patterns patl =
+ List.fold_right (fun patl restl ->
+ List.flatten (List.map (fun p -> List.map (fun rest -> p::rest) restl) patl))
+ patl [[]]
+
let rec cases_pattern_fold_map ?loc g e = DAst.with_val (function
| PatVar na ->
- let e',na' = g e na in e', DAst.make ?loc @@ PatVar na'
+ let e',disjpat,na' = g e na in
+ e', (match disjpat with
+ | None -> [DAst.make ?loc @@ PatVar na']
+ | Some ((_,disjpat),_) -> disjpat)
| PatCstr (cstr,patl,na) ->
- let e',na' = g e na in
+ let e',disjpat,na' = g e na in
+ if disjpat <> None then user_err (Pp.str "Unable to instantiate an \"as\" clause with a pattern.");
let e',patl' = List.fold_left_map (cases_pattern_fold_map ?loc g) e patl in
- e', DAst.make ?loc @@ PatCstr (cstr,patl',na')
+ (* Distribute outwards the inner disjunctive patterns *)
+ let disjpatl' = product_of_cases_patterns patl' in
+ e', List.map (fun patl' -> DAst.make ?loc @@ PatCstr (cstr,patl',na')) disjpatl'
)
let subst_binder_type_vars l = function
@@ -136,6 +147,16 @@ let rec subst_glob_vars l gc = DAst.map (function
let ldots_var = Id.of_string ".."
+let protect g e na =
+ let e',disjpat,na = g e na in
+ if disjpat <> None then user_err (Pp.str "Unsupported substitution of an arbitrary pattern.");
+ e',na
+
+let apply_cases_pattern ?loc ((ids,disjpat),id) c =
+ let tm = DAst.make ?loc (GVar id) in
+ let eqns = List.map (fun pat -> (loc,(ids,[pat],c))) disjpat in
+ DAst.make ?loc @@ GCases (LetPatternStyle, None, [tm,(Anonymous,None)], eqns)
+
let glob_constr_of_notation_constr_with_binders ?loc g f e nc =
let lt x = DAst.make ?loc x in lt @@ match nc with
| NVar id -> GVar id
@@ -146,46 +167,51 @@ let glob_constr_of_notation_constr_with_binders ?loc g f e nc =
let inner = lt @@ GApp (lt @@ GVar (ldots_var),[subst_glob_vars innerl it]) in
let outerl = (ldots_var,inner)::(if swap then [x, lt @@ GVar y] else []) in
DAst.get (subst_glob_vars outerl it)
- | NBinderList (x,y,iter,tail) ->
+ | NBinderList (x,y,iter,tail,swap) ->
let t = f e tail in let it = f e iter in
- let innerl = [(ldots_var,t);(x, lt @@ GVar y)] in
+ let innerl = (ldots_var,t)::(if swap then [] else [x, lt @@ GVar y]) in
let inner = lt @@ GApp (lt @@ GVar ldots_var,[subst_glob_vars innerl it]) in
- let outerl = [(ldots_var,inner)] in
+ let outerl = (ldots_var,inner)::(if swap then [x, lt @@ GVar y] else []) in
DAst.get (subst_glob_vars outerl it)
| NLambda (na,ty,c) ->
- let e',na = g e na in GLambda (na,Explicit,f e ty,f e' c)
+ let e',disjpat,na = g e na in GLambda (na,Explicit,f e ty,Option.fold_right (apply_cases_pattern ?loc) disjpat (f e' c))
| NProd (na,ty,c) ->
- let e',na = g e na in GProd (na,Explicit,f e ty,f e' c)
+ let e',disjpat,na = g e na in GProd (na,Explicit,f e ty,Option.fold_right (apply_cases_pattern ?loc) disjpat (f e' c))
| NLetIn (na,b,t,c) ->
- let e',na = g e na in GLetIn (na,f e b,Option.map (f e) t,f e' c)
+ let e',disjpat,na = g e na in
+ (match disjpat with
+ | None -> GLetIn (na,f e b,Option.map (f e) t,f e' c)
+ | Some disjpat -> DAst.get (apply_cases_pattern ?loc disjpat (f e' c)))
| NCases (sty,rtntypopt,tml,eqnl) ->
let e',tml' = List.fold_right (fun (tm,(na,t)) (e',tml') ->
let e',t' = match t with
| None -> e',None
| Some (ind,nal) ->
let e',nal' = List.fold_right (fun na (e',nal) ->
- let e',na' = g e' na in e',na'::nal) nal (e',[]) in
- e',Some (Loc.tag ?loc (ind,nal')) in
- let e',na' = g e' na in
- (e',(f e tm,(na',t'))::tml')) tml (e,[]) in
- let fold (idl,e) na = let (e,na) = g e na in ((Name.cons na idl,e),na) in
+ let e',na' = protect g e' na in
+ e',na'::nal) nal (e',[]) in
+ e',Some (Loc.tag ?loc (ind,nal')) in
+ let e',na' = protect g e' na in
+ (e',(f e tm,(na',t'))::tml')) tml (e,[]) in
+ let fold (idl,e) na = let (e,disjpat,na) = g e na in ((Name.cons na idl,e),disjpat,na) in
let eqnl' = List.map (fun (patl,rhs) ->
- let ((idl,e),patl) =
- List.fold_left_map (cases_pattern_fold_map ?loc fold) ([],e) patl in
- Loc.tag (idl,patl,f e rhs)) eqnl in
- GCases (sty,Option.map (f e') rtntypopt,tml',eqnl')
+ let ((idl,e),patl) =
+ List.fold_left_map (cases_pattern_fold_map ?loc fold) ([],e) patl in
+ let disjpatl = product_of_cases_patterns patl in
+ List.map (fun patl -> Loc.tag (idl,patl,f e rhs)) disjpatl) eqnl in
+ GCases (sty,Option.map (f e') rtntypopt,tml',List.flatten eqnl')
| NLetTuple (nal,(na,po),b,c) ->
- let e',nal = List.fold_left_map g e nal in
- let e'',na = g e na in
+ let e',nal = List.fold_left_map (protect g) e nal in
+ let e'',na = protect g e na in
GLetTuple (nal,(na,Option.map (f e'') po),f e b,f e' c)
| NIf (c,(na,po),b1,b2) ->
- let e',na = g e na in
+ let e',na = protect g e na in
GIf (f e c,(na,Option.map (f e') po),f e b1,f e b2)
| NRec (fk,idl,dll,tl,bl) ->
- let e,dll = Array.fold_left_map (List.fold_left_map (fun e (na,oc,b) ->
- let e,na = g e na in
+ let e,dll = Array.fold_left_map (List.fold_map (fun e (na,oc,b) ->
+ let e,na = protect g e na in
(e,(na,Explicit,Option.map (f e) oc,f e b)))) e dll in
- let e',idl = Array.fold_left_map (to_id g) e idl in
+ let e',idl = Array.fold_left_map (to_id (protect g)) e idl in
GRec (fk,idl,dll,Array.map (f e) tl,Array.map (f e') bl)
| NCast (c,k) -> GCast (f e c,Miscops.map_cast_type (f e) k)
| NSort x -> GSort x
@@ -195,13 +221,19 @@ let glob_constr_of_notation_constr_with_binders ?loc g f e nc =
let glob_constr_of_notation_constr ?loc x =
let rec aux () x =
- glob_constr_of_notation_constr_with_binders ?loc (fun () id -> ((),id)) aux () x
+ glob_constr_of_notation_constr_with_binders ?loc (fun () id -> ((),None,id)) aux () x
in aux () x
(******************************************************************************)
(* Translating a glob_constr into a notation, interpreting recursive patterns *)
-let add_id r id = r := (id :: pi1 !r, pi2 !r, pi3 !r)
+type found_variables = {
+ vars : Id.t list;
+ recursive_term_vars : (Id.t * Id.t) list;
+ recursive_binders_vars : (Id.t * Id.t) list;
+ }
+
+let add_id r id = r := { !r with vars = id :: (!r).vars }
let add_name r = function Anonymous -> () | Name id -> add_id r id
let is_gvar id c = match DAst.get c with
@@ -245,13 +277,25 @@ let check_is_hole id t = match DAst.get t with GHole _ -> () | _ ->
(strbrk "In recursive notation with binders, " ++ Id.print id ++
strbrk " is expected to come without type.")
+let check_pair_matching ?loc x y x' y' revert revert' =
+ if not (Id.equal x x' && Id.equal y y' && revert = revert') then
+ let x,y = if revert then y,x else x,y in
+ let x',y' = if revert' then y',x' else x',y' in
+ (* This is a case where one would like to highlight two locations! *)
+ user_err ?loc
+ (strbrk "Found " ++ Id.print x ++ strbrk " matching " ++ Id.print y ++
+ strbrk " while " ++ Id.print x' ++ strbrk " matching " ++ Id.print y' ++
+ strbrk " was first found.")
+
let pair_equal eq1 eq2 (a,b) (a',b') = eq1 a a' && eq2 b b'
+let mem_recursive_pair (x,y) l = List.mem_f (pair_equal Id.equal Id.equal) (x,y) l
+
type recursive_pattern_kind =
-| RecursiveTerms of bool (* associativity *)
-| RecursiveBinders of glob_constr * glob_constr
+| RecursiveTerms of bool (* in reverse order *)
+| RecursiveBinders of bool (* in reverse order *)
-let compare_recursive_parts found f f' (iterator,subc) =
+let compare_recursive_parts recvars found f f' (iterator,subc) =
let diff = ref None in
let terminator = ref None in
let rec aux c1 c2 = match DAst.get c1, DAst.get c2 with
@@ -270,24 +314,41 @@ let compare_recursive_parts found f f' (iterator,subc) =
List.for_all2eq aux l1 l2
| _ -> mk_glob_constr_eq aux c1 c2
end
- | GVar x, GVar y when not (Id.equal x y) ->
+ | GVar x, GVar y
+ when mem_recursive_pair (x,y) recvars || mem_recursive_pair (y,x) recvars ->
(* We found the position where it differs *)
- let lassoc = match !terminator with None -> false | Some _ -> true in
- let x,y = if lassoc then y,x else x,y in
+ let revert = mem_recursive_pair (y,x) recvars in
+ let x,y = if revert then y,x else x,y in
begin match !diff with
| None ->
- let () = diff := Some (x, y, RecursiveTerms lassoc) in
+ let () = diff := Some (x, y, RecursiveTerms revert) in
+ true
+ | Some (x', y', RecursiveTerms revert')
+ | Some (x', y', RecursiveBinders revert') ->
+ check_pair_matching ?loc:c1.CAst.loc x y x' y' revert revert';
true
- | Some _ -> false
end
| GLambda (Name x,_,t_x,c), GLambda (Name y,_,t_y,term)
- | GProd (Name x,_,t_x,c), GProd (Name y,_,t_y,term) when not (Id.equal x y) ->
+ | GProd (Name x,_,t_x,c), GProd (Name y,_,t_y,term)
+ when mem_recursive_pair (x,y) recvars || mem_recursive_pair (y,x) recvars ->
(* We found a binding position where it differs *)
+ check_is_hole x t_x;
+ check_is_hole y t_y;
+ let revert = mem_recursive_pair (y,x) recvars in
+ let x,y = if revert then y,x else x,y in
begin match !diff with
| None ->
- let () = diff := Some (x, y, RecursiveBinders (t_x,t_y)) in
+ let () = diff := Some (x, y, RecursiveBinders revert) in
aux c term
- | Some _ -> false
+ | Some (x', y', RecursiveBinders revert') ->
+ check_pair_matching ?loc:c1.CAst.loc x y x' y' revert revert';
+ true
+ | Some (x', y', RecursiveTerms revert') ->
+ (* Recursive binders have precedence: they can be coerced to
+ terms but not reciprocally *)
+ check_pair_matching ?loc:c1.CAst.loc x y x' y' revert revert';
+ let () = diff := Some (x, y, RecursiveBinders revert) in
+ true
end
| _ ->
mk_glob_constr_eq aux c1 c2 in
@@ -296,46 +357,36 @@ let compare_recursive_parts found f f' (iterator,subc) =
| None ->
let loc1 = loc_of_glob_constr iterator in
let loc2 = loc_of_glob_constr (Option.get !terminator) in
- (* Here, we would need a loc made of several parts ... *)
- user_err ?loc:(subtract_loc loc1 loc2)
+ (* Here, we would need a loc made of several parts ... *)
+ user_err ?loc:(subtract_loc loc1 loc2)
(str "Both ends of the recursive pattern are the same.")
- | Some (x,y,RecursiveTerms lassoc) ->
- let toadd,x,y,lassoc =
- if List.mem_f (pair_equal Id.equal Id.equal) (x,y) (pi2 !found) ||
- List.mem_f (pair_equal Id.equal Id.equal) (x,y) (pi3 !found)
- then
- None,x,y,lassoc
- else if List.mem_f (pair_equal Id.equal Id.equal) (y,x) (pi2 !found) ||
- List.mem_f (pair_equal Id.equal Id.equal) (y,x) (pi3 !found)
- then
- None,y,x,not lassoc
- else
- Some (x,y),x,y,lassoc in
- let iterator =
- f' (if lassoc then iterator
- else subst_glob_vars [x, DAst.make @@ GVar y] iterator) in
- (* found variables have been collected by compare_constr *)
- found := (List.remove Id.equal y (pi1 !found),
- Option.fold_right (fun a l -> a::l) toadd (pi2 !found),
- pi3 !found);
- NList (x,y,iterator,f (Option.get !terminator),lassoc)
- | Some (x,y,RecursiveBinders (t_x,t_y)) ->
- let iterator = f' (subst_glob_vars [x, DAst.make @@ GVar y] iterator) in
- (* found have been collected by compare_constr *)
- found := (List.remove Id.equal y (pi1 !found), pi2 !found, (x,y) :: pi3 !found);
- check_is_hole x t_x;
- check_is_hole y t_y;
- NBinderList (x,y,iterator,f (Option.get !terminator))
+ | Some (x,y,RecursiveTerms revert) ->
+ (* By arbitrary convention, we use the second variable of the pair
+ as the place-holder for the iterator *)
+ let iterator =
+ f' (if revert then iterator else subst_glob_vars [x, DAst.make @@ GVar y] iterator) in
+ (* found variables have been collected by compare_constr *)
+ found := { !found with vars = List.remove Id.equal y (!found).vars;
+ recursive_term_vars = List.add_set (pair_equal Id.equal Id.equal) (x,y) (!found).recursive_term_vars };
+ NList (x,y,iterator,f (Option.get !terminator),revert)
+ | Some (x,y,RecursiveBinders revert) ->
+ let iterator =
+ f' (if revert then iterator else subst_glob_vars [x, DAst.make @@ GVar y] iterator) in
+ (* found have been collected by compare_constr *)
+ found := { !found with vars = List.remove Id.equal y (!found).vars;
+ recursive_binders_vars = List.add_set (pair_equal Id.equal Id.equal) (x,y) (!found).recursive_binders_vars };
+ NBinderList (x,y,iterator,f (Option.get !terminator),revert)
else
raise Not_found
-let notation_constr_and_vars_of_glob_constr a =
- let found = ref ([],[],[]) in
+let notation_constr_and_vars_of_glob_constr recvars a =
+ let found = ref { vars = []; recursive_term_vars = []; recursive_binders_vars = [] } in
let has_ltac = ref false in
+ (* Turn a glob_constr into a notation_constr by first trying to find a recursive pattern *)
let rec aux c =
let keepfound = !found in
(* n^2 complexity but small and done only once per notation *)
- try compare_recursive_parts found aux aux' (split_at_recursive_part c)
+ try compare_recursive_parts recvars found aux aux' (split_at_recursive_part c)
with Not_found ->
found := keepfound;
match DAst.get c with
@@ -395,8 +446,9 @@ let notation_constr_and_vars_of_glob_constr a =
(* Side effect *)
t, !found, !has_ltac
-let check_variables_and_reversibility nenv (found,foundrec,foundrecbinding) =
- let injective = ref true in
+let check_variables_and_reversibility nenv
+ { vars = found; recursive_term_vars = foundrec; recursive_binders_vars = foundrecbinding } =
+ let injective = ref [] in
let recvars = nenv.ninterp_rec_vars in
let fold _ y accu = Id.Set.add y accu in
let useless_vars = Id.Map.fold fold recvars Id.Set.empty in
@@ -419,33 +471,36 @@ let check_variables_and_reversibility nenv (found,foundrec,foundrecbinding) =
user_err Pp.(str
(Id.to_string x ^
" should not be bound in a recursive pattern of the right-hand side."))
- else injective := false
+ else injective := x :: !injective
in
let check_pair s x y where =
- if not (List.mem_f (pair_equal Id.equal Id.equal) (x,y) where) then
+ if not (mem_recursive_pair (x,y) where) then
user_err (strbrk "in the right-hand side, " ++ Id.print x ++
str " and " ++ Id.print y ++ strbrk " should appear in " ++ str s ++
str " position as part of a recursive pattern.") in
let check_type x typ =
match typ with
- | NtnInternTypeConstr ->
+ | NtnInternTypeAny ->
begin
try check_pair "term" x (Id.Map.find x recvars) foundrec
with Not_found -> check_bound x
end
- | NtnInternTypeBinder ->
+ | NtnInternTypeOnlyBinder ->
begin
try check_pair "binding" x (Id.Map.find x recvars) foundrecbinding
with Not_found -> check_bound x
- end
- | NtnInternTypeIdent -> check_bound x in
+ end in
Id.Map.iter check_type vars;
- !injective
+ List.rev !injective
let notation_constr_of_glob_constr nenv a =
- let a, found, has_ltac = notation_constr_and_vars_of_glob_constr a in
+ let recvars = Id.Map.bindings nenv.ninterp_rec_vars in
+ let a, found, has_ltac = notation_constr_and_vars_of_glob_constr recvars a in
let injective = check_variables_and_reversibility nenv found in
- a, not has_ltac && injective
+ let status = if has_ltac then HasLtac else match injective with
+ | [] -> APrioriReversible
+ | l -> NonInjective l in
+ a, status
(**********************************************************************)
(* Substitution of kernel names, avoiding a list of bound identifiers *)
@@ -501,11 +556,11 @@ let rec subst_notation_constr subst bound raw =
if r1' == r1 && r2' == r2 then raw else
NProd (n,r1',r2')
- | NBinderList (id1,id2,r1,r2) ->
+ | NBinderList (id1,id2,r1,r2,b) ->
let r1' = subst_notation_constr subst bound r1
and r2' = subst_notation_constr subst bound r2 in
if r1' == r1 && r2' == r2 then raw else
- NBinderList (id1,id2,r1',r2')
+ NBinderList (id1,id2,r1',r2',b)
| NLetIn (n,r1,t,r2) ->
let r1' = subst_notation_constr subst bound r1 in
@@ -616,8 +671,20 @@ let is_term_meta id metas =
try match Id.List.assoc id metas with _,(NtnTypeConstr | NtnTypeConstrList) -> true | _ -> false
with Not_found -> false
+let is_onlybinding_strict_meta id metas =
+ try match Id.List.assoc id metas with _,NtnTypeBinder (NtnParsedAsPattern true) -> true | _ -> false
+ with Not_found -> false
+
let is_onlybinding_meta id metas =
- try match Id.List.assoc id metas with _,NtnTypeOnlyBinder -> true | _ -> false
+ try match Id.List.assoc id metas with _,NtnTypeBinder _ -> true | _ -> false
+ with Not_found -> false
+
+let is_onlybinding_pattern_like_meta isvar id metas =
+ try match Id.List.assoc id metas with
+ | _,NtnTypeBinder (NtnBinderParsedAsConstr
+ (Extend.AsIdentOrPattern | Extend.AsStrictPattern)) -> true
+ | _,NtnTypeBinder (NtnParsedAsPattern strict) -> not (strict && isvar)
+ | _ -> false
with Not_found -> false
let is_bindinglist_meta id metas =
@@ -636,7 +703,7 @@ let alpha_rename alpmetas v =
if alpmetas == [] then v
else try rename_glob_vars alpmetas v with UnsoundRenaming -> raise No_match
-let add_env (alp,alpmetas) (terms,onlybinders,termlists,binderlists) var v =
+let add_env (alp,alpmetas) (terms,termlists,binders,binderlists) var v =
(* Check that no capture of binding variables occur *)
(* [alp] is used when matching a pattern "fun x => ... x ... ?var ... x ..."
with an actual term "fun z => ... z ..." when "x" is not bound in the
@@ -664,19 +731,49 @@ let add_env (alp,alpmetas) (terms,onlybinders,termlists,binderlists) var v =
refinement *)
let v = alpha_rename alpmetas v in
(* TODO: handle the case of multiple occs in different scopes *)
- ((var,v)::terms,onlybinders,termlists,binderlists)
+ ((var,v)::terms,termlists,binders,binderlists)
-let add_termlist_env (alp,alpmetas) (terms,onlybinders,termlists,binderlists) var vl =
+let add_termlist_env (alp,alpmetas) (terms,termlists,binders,binderlists) var vl =
if List.exists (fun (id,_) -> List.exists (occur_glob_constr id) vl) alp then raise No_match;
let vl = List.map (alpha_rename alpmetas) vl in
- (terms,onlybinders,(var,vl)::termlists,binderlists)
+ (terms,(var,vl)::termlists,binders,binderlists)
-let add_binding_env alp (terms,onlybinders,termlists,binderlists) var v =
+let add_binding_env alp (terms,termlists,binders,binderlists) var v =
(* TODO: handle the case of multiple occs in different scopes *)
- (terms,(var,v)::onlybinders,termlists,binderlists)
+ (terms,termlists,(var,v)::binders,binderlists)
-let add_bindinglist_env (terms,onlybinders,termlists,binderlists) x bl =
- (terms,onlybinders,termlists,(x,bl)::binderlists)
+let add_bindinglist_env (terms,termlists,binders,binderlists) x bl =
+ (terms,termlists,binders,(x,bl)::binderlists)
+
+let rec map_cases_pattern_name_left f = DAst.map (function
+ | PatVar na -> PatVar (f na)
+ | PatCstr (c,l,na) -> PatCstr (c,List.map_left (map_cases_pattern_name_left f) l,f na)
+ )
+
+let rec fold_cases_pattern_eq f x p p' =
+ let loc = p.CAst.loc in
+ match DAst.get p, DAst.get p' with
+ | PatVar na, PatVar na' -> let x,na = f x na na' in x, DAst.make ?loc @@ PatVar na
+ | PatCstr (c,l,na), PatCstr (c',l',na') when eq_constructor c c' ->
+ let x,l = fold_cases_pattern_list_eq f x l l' in
+ let x,na = f x na na' in
+ x, DAst.make ?loc @@ PatCstr (c,l,na)
+ | _ -> failwith "Not equal"
+
+and fold_cases_pattern_list_eq f x pl pl' = match pl, pl' with
+ | [], [] -> x, []
+ | p::pl, p'::pl' ->
+ let x, p = fold_cases_pattern_eq f x p p' in
+ let x, pl = fold_cases_pattern_list_eq f x pl pl' in
+ x, p :: pl
+ | _ -> assert false
+
+let rec cases_pattern_eq p1 p2 = match DAst.get p1, DAst.get p2 with
+| PatVar na1, PatVar na2 -> Name.equal na1 na2
+| PatCstr (c1, pl1, na1), PatCstr (c2, pl2, na2) ->
+ eq_constructor c1 c2 && List.equal cases_pattern_eq pl1 pl2 &&
+ Name.equal na1 na2
+| _ -> false
let rec pat_binder_of_term t = DAst.map (function
| GVar id -> PatVar (Name id)
@@ -691,39 +788,111 @@ let rec pat_binder_of_term t = DAst.map (function
| _ -> raise No_match
) t
-let bind_term_env alp (terms,onlybinders,termlists,binderlists as sigma) var v =
+let unify_name_upto alp na na' =
+ match na, na' with
+ | Anonymous, na' -> alp, na'
+ | na, Anonymous -> alp, na
+ | Name id, Name id' ->
+ if Id.equal id id' then alp, na'
+ else (fst alp,(id,id')::snd alp), na'
+
+let unify_pat_upto alp p p' =
+ try fold_cases_pattern_eq unify_name_upto alp p p' with Failure _ -> raise No_match
+
+let unify_term alp v v' =
+ match DAst.get v, DAst.get v' with
+ | GHole _, _ -> v'
+ | _, GHole _ -> v
+ | _, _ -> if glob_constr_eq (alpha_rename (snd alp) v) v' then v else raise No_match
+
+let unify_opt_term alp v v' =
+ match v, v' with
+ | Some t, Some t' -> Some (unify_term alp t t')
+ | (Some _ as x), None | None, (Some _ as x) -> x
+ | None, None -> None
+
+let unify_binding_kind bk bk' = if bk == bk' then bk' else raise No_match
+
+let unify_binder_upto alp b b' =
+ let loc, loc' = CAst.(b.loc, b'.loc) in
+ match DAst.get b, DAst.get b' with
+ | GLocalAssum (na,bk,t), GLocalAssum (na',bk',t') ->
+ let alp, na = unify_name_upto alp na na' in
+ alp, DAst.make ?loc @@ GLocalAssum (na, unify_binding_kind bk bk', unify_term alp t t')
+ | GLocalDef (na,bk,c,t), GLocalDef (na',bk',c',t') ->
+ let alp, na = unify_name_upto alp na na' in
+ alp, DAst.make ?loc @@ GLocalDef (na, unify_binding_kind bk bk', unify_term alp c c', unify_opt_term alp t t')
+ | GLocalPattern ((disjpat,ids),id,bk,t), GLocalPattern ((disjpat',_),_,bk',t') when List.length disjpat = List.length disjpat' ->
+ let alp, p = List.fold_left2_map unify_pat_upto alp disjpat disjpat' in
+ alp, DAst.make ?loc @@ GLocalPattern ((p,ids), id, unify_binding_kind bk bk', unify_term alp t t')
+ | _ -> raise No_match
+
+let rec unify_terms alp vl vl' =
+ match vl, vl' with
+ | [], [] -> []
+ | v :: vl, v' :: vl' -> unify_term alp v v' :: unify_terms alp vl vl'
+ | _ -> raise No_match
+
+let rec unify_binders_upto alp bl bl' =
+ match bl, bl' with
+ | [], [] -> alp, []
+ | b :: bl, b' :: bl' ->
+ let alp,b = unify_binder_upto alp b b' in
+ let alp,bl = unify_binders_upto alp bl bl' in
+ alp, b :: bl
+ | _ -> raise No_match
+
+let unify_id alp id na' =
+ match na' with
+ | Anonymous -> Name (rename_var (snd alp) id)
+ | Name id' ->
+ if Id.equal (rename_var (snd alp) id) id' then na' else raise No_match
+
+let unify_pat alp p p' =
+ if cases_pattern_eq (map_cases_pattern_name_left (Name.map (rename_var (snd alp))) p) p' then p'
+ else raise No_match
+
+let unify_term_binder alp c = DAst.(map (fun b' ->
+ match DAst.get c, b' with
+ | GVar id, GLocalAssum (na', bk', t') ->
+ GLocalAssum (unify_id alp id na', bk', t')
+ | _, GLocalPattern (([p'],ids), id, bk', t') ->
+ let p = pat_binder_of_term c in
+ GLocalPattern (([unify_pat alp p p'],ids), id, bk', t')
+ | _ -> raise No_match))
+
+let rec unify_terms_binders alp cl bl' =
+ match cl, bl' with
+ | [], [] -> []
+ | c :: cl, b' :: bl' ->
+ begin match DAst.get b' with
+ | GLocalDef ( _, _, _, t) -> unify_terms_binders alp cl bl'
+ | _ -> unify_term_binder alp c b' :: unify_terms_binders alp cl bl'
+ end
+ | _ -> raise No_match
+
+let bind_term_env alp (terms,termlists,binders,binderlists as sigma) var v =
try
+ (* If already bound to a term, unify with the new term *)
let v' = Id.List.assoc var terms in
- match DAst.get v, DAst.get v' with
- | GHole _, _ -> sigma
- | _, GHole _ ->
- let sigma = Id.List.remove_assoc var terms,onlybinders,termlists,binderlists in
- add_env alp sigma var v
- | _, _ ->
- if glob_constr_eq (alpha_rename (snd alp) v) v' then sigma
- else raise No_match
+ let v'' = unify_term alp v v' in
+ if v'' == v' then sigma else
+ let sigma = (Id.List.remove_assoc var terms,termlists,binders,binderlists) in
+ add_env alp sigma var v
with Not_found -> add_env alp sigma var v
-let bind_termlist_env alp (terms,onlybinders,termlists,binderlists as sigma) var vl =
+let bind_termlist_env alp (terms,termlists,binders,binderlists as sigma) var vl =
try
+ (* If already bound to a list of term, unify with the new terms *)
let vl' = Id.List.assoc var termlists in
- let unify_term v v' =
- match DAst.get v, DAst.get v' with
- | GHole _, _ -> v'
- | _, GHole _ -> v
- | _, _ -> if glob_constr_eq (alpha_rename (snd alp) v) v' then v' else raise No_match in
- let rec unify vl vl' =
- match vl, vl' with
- | [], [] -> []
- | v :: vl, v' :: vl' -> unify_term v v' :: unify vl vl'
- | _ -> raise No_match in
- let vl = unify vl vl' in
- let sigma = (terms,onlybinders,Id.List.remove_assoc var termlists,binderlists) in
+ let vl = unify_terms alp vl vl' in
+ let sigma = (terms,Id.List.remove_assoc var termlists,binders,binderlists) in
add_termlist_env alp sigma var vl
with Not_found -> add_termlist_env alp sigma var vl
-let bind_term_as_binding_env alp (terms,onlybinders,termlists,binderlists as sigma) var id =
+let bind_term_as_binding_env alp (terms,termlists,binders,binderlists as sigma) var id =
try
+ (* If already bound to a term, unify the binder and the term *)
match DAst.get (Id.List.assoc var terms) with
| GVar id' ->
(if not (Id.equal id id') then (fst alp,(id,id')::snd alp) else alp),
@@ -735,142 +904,49 @@ let bind_term_as_binding_env alp (terms,onlybinders,termlists,binderlists as sig
(* TODO: look at the consequences for alp *)
alp, add_env alp sigma var (DAst.make @@ GVar id)
-let bind_binding_as_term_env alp (terms,onlybinders,termlists,binderlists as sigma) var id =
+let force_cases_pattern c =
+ DAst.make ?loc:c.CAst.loc (DAst.get c)
+
+let bind_binding_as_term_env alp (terms,termlists,binders,binderlists as sigma) var c =
+ let pat = try force_cases_pattern (cases_pattern_of_glob_constr Anonymous c) with Not_found -> raise No_match in
try
- let v' = Id.List.assoc var onlybinders in
- match v' with
- | Anonymous ->
- (* Should not occur, since the term has to be bound upwards *)
- let sigma = (terms,Id.List.remove_assoc var onlybinders,termlists,binderlists) in
- add_binding_env alp sigma var (Name id)
- | Name id' ->
- if Id.equal (rename_var (snd alp) id) id' then sigma else raise No_match
- with Not_found -> add_binding_env alp sigma var (Name id)
-
-let bind_binding_env alp (terms,onlybinders,termlists,binderlists as sigma) var v =
+ (* If already bound to a binder, unify the term and the binder *)
+ let patl' = Id.List.assoc var binders in
+ let patl'' = List.map2 (unify_pat alp) [pat] patl' in
+ if patl' == patl'' then sigma
+ else
+ let sigma = (terms,termlists,Id.List.remove_assoc var binders,binderlists) in
+ add_binding_env alp sigma var patl''
+ with Not_found -> add_binding_env alp sigma var [pat]
+
+let bind_binding_env alp (terms,termlists,binders,binderlists as sigma) var disjpat =
try
- let v' = Id.List.assoc var onlybinders in
- match v, v' with
- | Anonymous, _ -> alp, sigma
- | _, Anonymous ->
- let sigma = (terms,Id.List.remove_assoc var onlybinders,termlists,binderlists) in
- alp, add_binding_env alp sigma var v
- | Name id1, Name id2 ->
- if Id.equal id1 id2 then alp,sigma
- else (fst alp,(id1,id2)::snd alp),sigma
- with Not_found -> alp, add_binding_env alp sigma var v
-
-let rec map_cases_pattern_name_left f = DAst.map (function
- | PatVar na -> PatVar (f na)
- | PatCstr (c,l,na) -> PatCstr (c,List.map_left (map_cases_pattern_name_left f) l,f na)
- )
-
-let rec fold_cases_pattern_eq f x p p' =
- let loc = p.CAst.loc in
- match DAst.get p, DAst.get p' with
- | PatVar na, PatVar na' -> let x,na = f x na na' in x, DAst.make ?loc @@ PatVar na
- | PatCstr (c,l,na), PatCstr (c',l',na') when eq_constructor c c' ->
- let x,l = fold_cases_pattern_list_eq f x l l' in
- let x,na = f x na na' in
- x, DAst.make ?loc @@ PatCstr (c,l,na)
- | _ -> failwith "Not equal"
-
-and fold_cases_pattern_list_eq f x pl pl' = match pl, pl' with
- | [], [] -> x, []
- | p::pl, p'::pl' ->
- let x, p = fold_cases_pattern_eq f x p p' in
- let x, pl = fold_cases_pattern_list_eq f x pl pl' in
- x, p :: pl
- | _ -> assert false
-
-let rec cases_pattern_eq p1 p2 = match DAst.get p1, DAst.get p2 with
-| PatVar na1, PatVar na2 -> Name.equal na1 na2
-| PatCstr (c1, pl1, na1), PatCstr (c2, pl2, na2) ->
- eq_constructor c1 c2 && List.equal cases_pattern_eq pl1 pl2 &&
- Name.equal na1 na2
-| _ -> false
-
-let bind_bindinglist_env alp (terms,onlybinders,termlists,binderlists as sigma) var bl =
+ (* If already bound to a binder possibly *)
+ (* generating an alpha-renaming from unifying the new binder *)
+ let disjpat' = Id.List.assoc var binders in
+ let alp, disjpat = List.fold_left2_map unify_pat_upto alp disjpat disjpat' in
+ let sigma = (terms,termlists,Id.List.remove_assoc var binders,binderlists) in
+ alp, add_binding_env alp sigma var disjpat
+ with Not_found -> alp, add_binding_env alp sigma var disjpat
+
+let bind_bindinglist_env alp (terms,termlists,binders,binderlists as sigma) var bl =
let bl = List.rev bl in
try
+ (* If already bound to a list of binders possibly *)
+ (* generating an alpha-renaming from unifying the new binders *)
let bl' = Id.List.assoc var binderlists in
- let unify_name alp na na' =
- match na, na' with
- | Anonymous, na' -> alp, na'
- | na, Anonymous -> alp, na
- | Name id, Name id' ->
- if Id.equal id id' then alp, na'
- else (fst alp,(id,id')::snd alp), na' in
- let unify_pat alp p p' =
- try fold_cases_pattern_eq unify_name alp p p' with Failure _ -> raise No_match in
- let unify_term alp v v' =
- match DAst.get v, DAst.get v' with
- | GHole _, _ -> v'
- | _, GHole _ -> v
- | _, _ -> if glob_constr_eq (alpha_rename (snd alp) v) v' then v else raise No_match in
- let unify_opt_term alp v v' =
- match v, v' with
- | Some t, Some t' -> Some (unify_term alp t t')
- | (Some _ as x), None | None, (Some _ as x) -> x
- | None, None -> None in
- let unify_binding_kind bk bk' = if bk == bk' then bk' else raise No_match in
- let unify_binder alp b b' =
- let loc, loc' = CAst.(b.loc, b'.loc) in
- match DAst.get b, DAst.get b' with
- | GLocalAssum (na,bk,t), GLocalAssum (na',bk',t') ->
- let alp, na = unify_name alp na na' in
- alp, DAst.make ?loc @@ GLocalAssum (na, unify_binding_kind bk bk', unify_term alp t t')
- | GLocalDef (na,bk,c,t), GLocalDef (na',bk',c',t') ->
- let alp, na = unify_name alp na na' in
- alp, DAst.make ?loc @@ GLocalDef (na, unify_binding_kind bk bk', unify_term alp c c', unify_opt_term alp t t')
- | GLocalPattern ((p,ids),id,bk,t), GLocalPattern ((p',_),_,bk',t') ->
- let alp, p = unify_pat alp p p' in
- alp, DAst.make ?loc @@ GLocalPattern ((p,ids), id, unify_binding_kind bk bk', unify_term alp t t')
- | _ -> raise No_match in
- let rec unify alp bl bl' =
- match bl, bl' with
- | [], [] -> alp, []
- | b :: bl, b' :: bl' ->
- let alp,b = unify_binder alp b b' in
- let alp,bl = unify alp bl bl' in
- alp, b :: bl
- | _ -> raise No_match in
- let alp, bl = unify alp bl bl' in
- let sigma = (terms,onlybinders,termlists,Id.List.remove_assoc var binderlists) in
+ let alp, bl = unify_binders_upto alp bl bl' in
+ let sigma = (terms,termlists,binders,Id.List.remove_assoc var binderlists) in
alp, add_bindinglist_env sigma var bl
with Not_found ->
alp, add_bindinglist_env sigma var bl
-let bind_bindinglist_as_term_env alp (terms,onlybinders,termlists,binderlists) var cl =
+let bind_bindinglist_as_termlist_env alp (terms,termlists,binders,binderlists) var cl =
try
+ (* If already bound to a list of binders, unify the terms and binders *)
let bl' = Id.List.assoc var binderlists in
- let unify_id id na' =
- match na' with
- | Anonymous -> Name (rename_var (snd alp) id)
- | Name id' ->
- if Id.equal (rename_var (snd alp) id) id' then na' else raise No_match in
- let unify_pat p p' =
- if cases_pattern_eq (map_cases_pattern_name_left (Name.map (rename_var (snd alp))) p) p' then p'
- else raise No_match in
- let unify_term_binder c = DAst.(map (fun b' ->
- match DAst.get c, b' with
- | GVar id, GLocalAssum (na', bk', t') ->
- GLocalAssum (unify_id id na', bk', t')
- | _, GLocalPattern ((p',ids), id, bk', t') ->
- let p = pat_binder_of_term c in
- GLocalPattern ((unify_pat p p',ids), id, bk', t')
- | _ -> raise No_match )) in
- let rec unify cl bl' =
- match cl, bl' with
- | [], [] -> []
- | c :: cl, b' :: bl' ->
- begin match DAst.get b' with
- | GLocalDef ( _, _, _, t) -> unify cl bl'
- | _ -> unify_term_binder c b' :: unify cl bl'
- end
- | _ -> raise No_match in
- let bl = unify cl bl' in
- let sigma = (terms,onlybinders,termlists,Id.List.remove_assoc var binderlists) in
+ let bl = unify_terms_binders alp cl bl' in
+ let sigma = (terms,termlists,binders,Id.List.remove_assoc var binderlists) in
add_bindinglist_env sigma var bl
with Not_found ->
anomaly (str "There should be a binder list bindings this list of terms.")
@@ -894,8 +970,10 @@ let match_opt f sigma t1 t2 = match (t1,t2) with
| _ -> raise No_match
let match_names metas (alp,sigma) na1 na2 = match (na1,na2) with
+ | (na1,Name id2) when is_onlybinding_strict_meta id2 metas ->
+ raise No_match
| (na1,Name id2) when is_onlybinding_meta id2 metas ->
- bind_binding_env alp sigma id2 na1
+ bind_binding_env alp sigma id2 [DAst.make (PatVar na1)]
| (Name id1,Name id2) when is_term_meta id2 metas ->
(* We let the non-binding occurrence define the rhs and hence reason up to *)
(* alpha-conversion for the given occurrence of the name (see #4592)) *)
@@ -907,54 +985,42 @@ let match_names metas (alp,sigma) na1 na2 = match (na1,na2) with
| (Anonymous,Anonymous) -> alp,sigma
| _ -> raise No_match
-let rec match_cases_pattern_binders metas acc pat1 pat2 =
+let rec match_cases_pattern_binders allow_catchall metas (alp,sigma as acc) pat1 pat2 =
match DAst.get pat1, DAst.get pat2 with
+ | PatVar _, PatVar (Name id2) when is_onlybinding_pattern_like_meta true id2 metas ->
+ bind_binding_env alp sigma id2 [pat1]
+ | _, PatVar (Name id2) when is_onlybinding_pattern_like_meta false id2 metas ->
+ bind_binding_env alp sigma id2 [pat1]
| PatVar na1, PatVar na2 -> match_names metas acc na1 na2
+ | _, PatVar Anonymous when allow_catchall -> acc
| PatCstr (c1,patl1,na1), PatCstr (c2,patl2,na2)
when eq_constructor c1 c2 && Int.equal (List.length patl1) (List.length patl2) ->
- List.fold_left2 (match_cases_pattern_binders metas)
- (match_names metas acc na1 na2) patl1 patl2
+ List.fold_left2 (match_cases_pattern_binders false metas)
+ (match_names metas acc na1 na2) patl1 patl2
| _ -> raise No_match
-let glue_letin_with_decls = true
-
-let rec match_iterated_binders islambda decls bi = DAst.(with_loc_val (fun ?loc -> function
- | GLambda (na,bk,t,b) as b0 ->
- begin match na, DAst.get b with
- | Name p, GCases (LetPatternStyle,None,[(e,_)],[(_,(ids,[cp],b))])
- when islambda && is_gvar p e && not (occur_glob_constr p b) ->
- match_iterated_binders islambda ((DAst.make ?loc @@ GLocalPattern((cp,ids),p,bk,t))::decls) b
- | _, _ when islambda ->
- match_iterated_binders islambda ((DAst.make ?loc @@ GLocalAssum(na,bk,t))::decls) b
- | _ -> (decls, DAst.make ?loc b0)
- end
- | GProd (na,bk,t,b) as b0 ->
- begin match na, DAst.get b with
- | Name p, GCases (LetPatternStyle,None,[(e,_)],[(_,(ids,[cp],b))])
- when not islambda && is_gvar p e && not (occur_glob_constr p b) ->
- match_iterated_binders islambda ((DAst.make ?loc @@ GLocalPattern((cp,ids),p,bk,t))::decls) b
- | Name _, _ when not islambda ->
- match_iterated_binders islambda ((DAst.make ?loc @@ GLocalAssum(na,bk,t))::decls) b
- | _ -> (decls, DAst.make ?loc b0)
- end
- | GLetIn (na,c,t,b) when glue_letin_with_decls ->
- match_iterated_binders islambda
- ((DAst.make ?loc @@ GLocalDef (na,Explicit (*?*), c,t))::decls) b
- | b -> (decls, DAst.make ?loc b)
- )) bi
-
-let remove_sigma x (terms,onlybinders,termlists,binderlists) =
- (Id.List.remove_assoc x terms,onlybinders,termlists,binderlists)
+let remove_sigma x (terms,termlists,binders,binderlists) =
+ (Id.List.remove_assoc x terms,termlists,binders,binderlists)
-let remove_bindinglist_sigma x (terms,onlybinders,termlists,binderlists) =
- (terms,onlybinders,termlists,Id.List.remove_assoc x binderlists)
+let remove_bindinglist_sigma x (terms,termlists,binders,binderlists) =
+ (terms,termlists,binders,Id.List.remove_assoc x binderlists)
let add_ldots_var metas = (ldots_var,((None,[]),NtnTypeConstr))::metas
let add_meta_bindinglist x metas = (x,((None,[]),NtnTypeBinderList))::metas
-let match_binderlist_with_app match_fun alp metas sigma rest x y iter termin =
- let rec aux sigma bl rest =
+(* This tells if letins in the middle of binders should be included in
+ the sequence of binders *)
+let glue_inner_letin_with_decls = true
+
+(* This tells if trailing letins (with no further proper binders)
+ should be included in sequence of binders *)
+let glue_trailing_letin_with_decls = false
+
+exception OnlyTrailingLetIns
+
+let match_binderlist match_fun alp metas sigma rest x y iter termin revert =
+ let rec aux trailing_letins sigma bl rest =
try
let metas = add_ldots_var (add_meta_bindinglist y metas) in
let (terms,_,_,binderlists as sigma) = match_fun alp metas sigma rest iter in
@@ -963,16 +1029,32 @@ let match_binderlist_with_app match_fun alp metas sigma rest x y iter termin =
match Id.List.assoc y binderlists with [b] -> b | _ ->assert false
in
let sigma = remove_bindinglist_sigma y (remove_sigma ldots_var sigma) in
- aux sigma (b::bl) rest
- with No_match when not (List.is_empty bl) ->
- bl, rest, sigma in
- let bl,rest,sigma = aux sigma [] rest in
+ (* In case y is bound not only to a binder but also to a term *)
+ let sigma = remove_sigma y sigma in
+ aux false sigma (b::bl) rest
+ with No_match ->
+ match DAst.get rest with
+ | GLetIn (na,c,t,rest') when glue_inner_letin_with_decls ->
+ let b = DAst.make ?loc:rest.CAst.loc @@ GLocalDef (na,Explicit (*?*), c,t) in
+ (* collect let-in *)
+ (try aux true sigma (b::bl) rest'
+ with OnlyTrailingLetIns
+ when not (trailing_letins && not glue_trailing_letin_with_decls) ->
+ (* renounce to take into account trailing let-ins *)
+ if not (List.is_empty bl) then bl, rest, sigma else raise No_match)
+ | _ ->
+ if trailing_letins && not glue_trailing_letin_with_decls then
+ (* Backtrack to when we tried to glue letins *)
+ raise OnlyTrailingLetIns;
+ if not (List.is_empty bl) then bl, rest, sigma else raise No_match in
+ let bl,rest,sigma = aux false sigma [] rest in
+ let bl = if revert then List.rev bl else bl in
let alp,sigma = bind_bindinglist_env alp sigma x bl in
match_fun alp metas sigma rest termin
let add_meta_term x metas = (x,((None,[]),NtnTypeConstr))::metas
-let match_termlist match_fun alp metas sigma rest x y iter termin lassoc =
+let match_termlist match_fun alp metas sigma rest x y iter termin revert =
let rec aux sigma acc rest =
try
let metas = add_ldots_var (add_meta_term y metas) in
@@ -983,12 +1065,12 @@ let match_termlist match_fun alp metas sigma rest x y iter termin lassoc =
aux sigma (t::acc) rest
with No_match when not (List.is_empty acc) ->
acc, match_fun metas sigma rest termin in
- let l,(terms,onlybinders,termlists,binderlists as sigma) = aux sigma [] rest in
- let l = if lassoc then l else List.rev l in
+ let l,(terms,termlists,binders,binderlists as sigma) = aux sigma [] rest in
+ let l = if revert then l else List.rev l in
if is_bindinglist_meta x metas then
(* This is a recursive pattern for both bindings and terms; it is *)
(* registered for binders *)
- bind_bindinglist_as_term_env alp sigma x l
+ bind_bindinglist_as_termlist_env alp sigma x l
else
bind_termlist_env alp sigma x l
@@ -1023,72 +1105,19 @@ let rec match_ inner u alp metas sigma a1 a2 =
match DAst.get a1, a2 with
(* Matching notation variable *)
| r1, NVar id2 when is_term_meta id2 metas -> bind_term_env alp sigma id2 a1
- | GVar id1, NVar id2 when is_onlybinding_meta id2 metas -> bind_binding_as_term_env alp sigma id2 id1
+ | GVar _, NVar id2 when is_onlybinding_pattern_like_meta true id2 metas -> bind_binding_as_term_env alp sigma id2 a1
+ | r1, NVar id2 when is_onlybinding_pattern_like_meta false id2 metas -> bind_binding_as_term_env alp sigma id2 a1
+ | GVar _, NVar id2 when is_onlybinding_strict_meta id2 metas -> raise No_match
+ | GVar _, NVar id2 when is_onlybinding_meta id2 metas -> bind_binding_as_term_env alp sigma id2 a1
| r1, NVar id2 when is_bindinglist_meta id2 metas -> bind_term_env alp sigma id2 a1
(* Matching recursive notations for terms *)
- | r1, NList (x,y,iter,termin,lassoc) ->
- match_termlist (match_hd u alp) alp metas sigma a1 x y iter termin lassoc
-
- | GLambda (na1, bk, t1, b1), NBinderList (x,y,iter,termin) ->
- begin match na1, DAst.get b1, iter with
- (* "λ p, let 'cp = p in t" -> "λ 'cp, t" *)
- | Name p, GCases (LetPatternStyle,None,[(e,_)],[(_,(ids,[cp],b1))]), NLambda (Name _, _, _)
- when is_gvar p e && not (occur_glob_constr p b1) ->
- let (decls,b) = match_iterated_binders true [DAst.make ?loc @@ GLocalPattern((cp,ids),p,bk,t1)] b1 in
- let alp,sigma = bind_bindinglist_env alp sigma x decls in
- match_in u alp metas sigma b termin
- (* Matching recursive notations for binders: ad hoc cases supporting let-in *)
- | _, _, NLambda (Name _,_,_) ->
- let (decls,b) = match_iterated_binders true [DAst.make ?loc @@ GLocalAssum (na1,bk,t1)] b1 in
- (* TODO: address the possibility that termin is a Lambda itself *)
- let alp,sigma = bind_bindinglist_env alp sigma x decls in
- match_in u alp metas sigma b termin
- (* Matching recursive notations for binders: general case *)
- | _, _, _ ->
- match_binderlist_with_app (match_hd u) alp metas sigma a1 x y iter termin
- end
-
- | GProd (na1, bk, t1, b1), NBinderList (x,y,iter,termin) ->
- (* "∀ p, let 'cp = p in t" -> "∀ 'cp, t" *)
- begin match na1, DAst.get b1, iter, termin with
- | Name p, GCases (LetPatternStyle,None,[(e, _)],[(_,(ids,[cp],b1))]), NProd (Name _,_,_), NVar _
- when is_gvar p e && not (occur_glob_constr p b1) ->
- let (decls,b) = match_iterated_binders true [DAst.make ?loc @@ GLocalPattern ((cp,ids),p,bk,t1)] b1 in
- let alp,sigma = bind_bindinglist_env alp sigma x decls in
- match_in u alp metas sigma b termin
- | _, _, NProd (Name _,_,_), _ when na1 != Anonymous ->
- let (decls,b) = match_iterated_binders false [DAst.make ?loc @@ GLocalAssum (na1,bk,t1)] b1 in
- (* TODO: address the possibility that termin is a Prod itself *)
- let alp,sigma = bind_bindinglist_env alp sigma x decls in
- match_in u alp metas sigma b termin
- (* Matching recursive notations for binders: general case *)
- | _, _, _, _ ->
- match_binderlist_with_app (match_hd u) alp metas sigma a1 x y iter termin
- end
+ | r1, NList (x,y,iter,termin,revert) ->
+ match_termlist (match_hd u alp) alp metas sigma a1 x y iter termin revert
(* Matching recursive notations for binders: general case *)
- | _r, NBinderList (x,y,iter,termin) ->
- match_binderlist_with_app (match_hd u) alp metas sigma a1 x y iter termin
-
- (* Matching individual binders as part of a recursive pattern *)
- | GLambda (na1, bk, t1, b1), NLambda (na2, t2, b2) ->
- begin match na1, DAst.get b1, na2 with
- | Name p, GCases (LetPatternStyle,None,[(e,_)],[(_,(ids,[cp],b1))]), Name id
- when is_gvar p e && is_bindinglist_meta id metas && not (occur_glob_constr p b1) ->
- let alp,sigma = bind_bindinglist_env alp sigma id [DAst.make ?loc @@ GLocalPattern ((cp,ids),p,bk,t1)] in
- match_in u alp metas sigma b1 b2
- | _, _, Name id when is_bindinglist_meta id metas ->
- let alp,sigma = bind_bindinglist_env alp sigma id [DAst.make ?loc @@ GLocalAssum (na1,bk,t1)] in
- match_in u alp metas sigma b1 b2
- | _ ->
- match_binders u alp metas na1 na2 (match_in u alp metas sigma t1 t2) b1 b2
- end
-
- | GProd (na,bk,t,b1), NProd (Name id,_,b2)
- when is_bindinglist_meta id metas && na != Anonymous ->
- let alp,sigma = bind_bindinglist_env alp sigma id [DAst.make ?loc @@ GLocalAssum (na,bk,t)] in
- match_in u alp metas sigma b1 b2
+ | _r, NBinderList (x,y,iter,termin,revert) ->
+ match_binderlist (match_hd u) alp metas sigma a1 x y iter termin revert
(* Matching compositionally *)
| GVar id1, NVar id2 when alpha_var id1 id2 (fst alp) -> sigma
@@ -1104,8 +1133,10 @@ let rec match_ inner u alp metas sigma a1 a2 =
let may_use_eta = does_not_come_from_already_eta_expanded_var f1 in
List.fold_left2 (match_ may_use_eta u alp metas)
(match_in u alp metas sigma f1 f2) l1 l2
- | GProd (na1,_,t1,b1), NProd (na2,t2,b2) ->
- match_binders u alp metas na1 na2 (match_in u alp metas sigma t1 t2) b1 b2
+ | GLambda (na1,bk1,t1,b1), NLambda (na2,t2,b2) ->
+ match_extended_binders false u alp metas na1 na2 bk1 t1 (match_in u alp metas sigma t1 t2) b1 b2
+ | GProd (na1,bk1,t1,b1), NProd (na2,t2,b2) ->
+ match_extended_binders true u alp metas na1 na2 bk1 t1 (match_in u alp metas sigma t1 t2) b1 b2
| GLetIn (na1,b1,_,c1), NLetIn (na2,b2,None,c2)
| GLetIn (na1,b1,None,c1), NLetIn (na2,b2,_,c2) ->
match_binders u alp metas na1 na2 (match_in u alp metas sigma b1 b2) c1 c2
@@ -1113,9 +1144,7 @@ let rec match_ inner u alp metas sigma a1 a2 =
match_binders u alp metas na1 na2
(match_in u alp metas (match_in u alp metas sigma b1 b2) t1 t2) c1 c2
| GCases (sty1,rtno1,tml1,eqnl1), NCases (sty2,rtno2,tml2,eqnl2)
- when sty1 == sty2
- && Int.equal (List.length tml1) (List.length tml2)
- && Int.equal (List.length eqnl1) (List.length eqnl2) ->
+ when sty1 == sty2 && Int.equal (List.length tml1) (List.length tml2) ->
let rtno1' = abstract_return_type_context_glob_constr tml1 rtno1 in
let rtno2' = abstract_return_type_context_notation_constr tml2 rtno2 in
let sigma =
@@ -1125,7 +1154,14 @@ let rec match_ inner u alp metas sigma a1 a2 =
let sigma = List.fold_left2
(fun s (tm1,_) (tm2,_) ->
match_in u alp metas s tm1 tm2) sigma tml1 tml2 in
- List.fold_left2 (match_equations u alp metas) sigma eqnl1 eqnl2
+ (* Try two different strategies for matching clauses *)
+ (try
+ List.fold_left2_set No_match (match_equations u alp metas) sigma eqnl1 eqnl2
+ with
+ No_match ->
+ List.fold_left2_set No_match (match_disjunctive_equations u alp metas) sigma
+ (Detyping.factorize_eqns eqnl1)
+ (List.map (fun (patl,rhs) -> ([patl],rhs)) eqnl2))
| GLetTuple (nal1,(na1,to1),b1,c1), NLetTuple (nal2,(na2,to2),b2,c2)
when Int.equal (List.length nal1) (List.length nal2) ->
let sigma = match_opt (match_binders u alp metas na1 na2) sigma to1 to2 in
@@ -1191,44 +1227,83 @@ and match_in u = match_ true u
and match_hd u = match_ false u
and match_binders u alp metas na1 na2 sigma b1 b2 =
+ (* Match binders which cannot be substituted by a pattern *)
let (alp,sigma) = match_names metas (alp,sigma) na1 na2 in
match_in u alp metas sigma b1 b2
-and match_equations u alp metas sigma (_,(_,patl1,rhs1)) (patl2,rhs2) =
+and match_extended_binders ?loc isprod u alp metas na1 na2 bk t sigma b1 b2 =
+ (* Match binders which can be substituted by a pattern *)
+ let store, get = set_temporary_memory () in
+ match na1, DAst.get b1, na2 with
+ (* Matching individual binders as part of a recursive pattern *)
+ | Name p, GCases (LetPatternStyle,None,[(e,_)],(_::_ as eqns)), Name id
+ when is_gvar p e && is_bindinglist_meta id metas && List.length (store (Detyping.factorize_eqns eqns)) = 1 ->
+ (match get () with
+ | [(_,(ids,disj_of_patl,b1))] ->
+ let disjpat = List.map (function [pat] -> pat | _ -> assert false) disj_of_patl in
+ let disjpat = if occur_glob_constr p b1 then List.map (set_pat_alias p) disjpat else disjpat in
+ let alp,sigma = bind_bindinglist_env alp sigma id [DAst.make ?loc @@ GLocalPattern ((disjpat,ids),p,bk,t)] in
+ match_in u alp metas sigma b1 b2
+ | _ -> assert false)
+ | Name p, GCases (LetPatternStyle,None,[(e,_)],(_::_ as eqns)), Name id
+ when is_gvar p e && is_onlybinding_pattern_like_meta false id metas && List.length (store (Detyping.factorize_eqns eqns)) = 1 ->
+ (match get () with
+ | [(_,(ids,disj_of_patl,b1))] ->
+ let disjpat = List.map (function [pat] -> pat | _ -> assert false) disj_of_patl in
+ let disjpat = if occur_glob_constr p b1 then List.map (set_pat_alias p) disjpat else disjpat in
+ let alp,sigma = bind_binding_env alp sigma id disjpat in
+ match_in u alp metas sigma b1 b2
+ | _ -> assert false)
+ | _, _, Name id when is_bindinglist_meta id metas && (not isprod || na1 != Anonymous)->
+ let alp,sigma = bind_bindinglist_env alp sigma id [DAst.make ?loc @@ GLocalAssum (na1,bk,t)] in
+ match_in u alp metas sigma b1 b2
+ | _, _, _ ->
+ let (alp,sigma) = match_names metas (alp,sigma) na1 na2 in
+ match_in u alp metas sigma b1 b2
+
+and match_equations u alp metas sigma (_,(ids,patl1,rhs1)) (patl2,rhs2) rest1 rest2 =
(* patl1 and patl2 have the same length because they respectively
correspond to some tml1 and tml2 that have the same length *)
+ let allow_catchall = (rest2 = [] && ids = []) in
let (alp,sigma) =
- List.fold_left2 (match_cases_pattern_binders metas)
+ List.fold_left2 (match_cases_pattern_binders allow_catchall metas)
(alp,sigma) patl1 patl2 in
match_in u alp metas sigma rhs1 rhs2
-let term_of_binder bi = DAst.make @@ match bi with
- | Name id -> GVar id
- | Anonymous -> GHole (Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None)
+and match_disjunctive_equations u alp metas sigma (_,(ids,disjpatl1,rhs1)) (disjpatl2,rhs2) _ _ =
+ (* patl1 and patl2 have the same length because they respectively
+ correspond to some tml1 and tml2 that have the same length *)
+ let (alp,sigma) =
+ List.fold_left2_set No_match
+ (fun alp_sigma patl1 patl2 _ _ ->
+ List.fold_left2 (match_cases_pattern_binders false metas) alp_sigma patl1 patl2)
+ (alp,sigma) disjpatl1 disjpatl2 in
+ match_in u alp metas sigma rhs1 rhs2
let match_notation_constr u c (metas,pat) =
- let terms,binders,termlists,binderlists =
+ let terms,termlists,binders,binderlists =
match_ false u ([],[]) metas ([],[],[],[]) c pat in
- (* Reorder canonically the substitution *)
- let find_binder x =
- try term_of_binder (Id.List.assoc x binders)
- with Not_found ->
- (* Happens for binders bound to Anonymous *)
- (* Find a better way to propagate Anonymous... *)
- DAst.make @@GVar x in
- List.fold_right (fun (x,(scl,typ)) (terms',termlists',binders') ->
+ (* Turning substitution based on binding/constr distinction into a
+ substitution based on entry productions *)
+ List.fold_right (fun (x,(scl,typ)) (terms',termlists',binders',binderlists') ->
match typ with
| NtnTypeConstr ->
let term = try Id.List.assoc x terms with Not_found -> raise No_match in
- ((term, scl)::terms',termlists',binders')
- | NtnTypeOnlyBinder ->
- ((find_binder x, scl)::terms',termlists',binders')
+ ((term, scl)::terms',termlists',binders',binderlists')
+ | NtnTypeBinder (NtnBinderParsedAsConstr _) ->
+ (match Id.List.assoc x binders with
+ | [pat] ->
+ let v = glob_constr_of_cases_pattern pat in
+ ((v,scl)::terms',termlists',binders',binderlists')
+ | _ -> raise No_match)
+ | NtnTypeBinder (NtnParsedAsIdent | NtnParsedAsPattern _) ->
+ (terms',termlists',(Id.List.assoc x binders,scl)::binders',binderlists')
| NtnTypeConstrList ->
- (terms',(Id.List.assoc x termlists,scl)::termlists',binders')
+ (terms',(Id.List.assoc x termlists,scl)::termlists',binders',binderlists')
| NtnTypeBinderList ->
let bl = try Id.List.assoc x binderlists with Not_found -> raise No_match in
- (terms',termlists',(bl, scl)::binders'))
- metas ([],[],[])
+ (terms',termlists',binders',(bl, scl)::binderlists'))
+ metas ([],[],[],[])
(* Matching cases pattern *)
@@ -1240,7 +1315,7 @@ let bind_env_cases_pattern (terms,x,termlists,y as sigma) var v =
(* TODO: handle the case of multiple occs in different scopes *)
(var,v)::terms,x,termlists,y
-let match_cases_pattern_list match_fun metas sigma rest x y iter termin lassoc =
+let match_cases_pattern_list match_fun metas sigma rest x y iter termin revert =
let rec aux sigma acc rest =
try
let metas = add_ldots_var (add_meta_term y metas) in
@@ -1251,10 +1326,10 @@ let match_cases_pattern_list match_fun metas sigma rest x y iter termin lassoc =
aux sigma (t::acc) rest
with No_match when not (List.is_empty acc) ->
acc, match_fun metas sigma rest termin in
- let l,(terms,onlybinders,termlists,binderlists as sigma) = aux sigma [] rest in
- (terms,onlybinders,(x,if lassoc then l else List.rev l)::termlists, binderlists)
+ let l,(terms,termlists,binders,binderlists as sigma) = aux sigma [] rest in
+ (terms,(x,if revert then l else List.rev l)::termlists,binders,binderlists)
-let rec match_cases_pattern metas (terms,(),termlists,() as sigma) a1 a2 =
+let rec match_cases_pattern metas (terms,termlists,(),() as sigma) a1 a2 =
match DAst.get a1, a2 with
| r1, NVar id2 when Id.List.mem_assoc id2 metas -> (bind_env_cases_pattern sigma id2 a1),(0,[])
| PatVar Anonymous, NHole _ -> sigma,(0,[])
@@ -1270,10 +1345,10 @@ let rec match_cases_pattern metas (terms,(),termlists,() as sigma) a1 a2 =
raise No_match
else
let l1',more_args = Util.List.chop le2 l1 in
- (List.fold_left2 (match_cases_pattern_no_more_args metas) sigma l1' l2),(le2,more_args)
- | r1, NList (x,y,iter,termin,lassoc) ->
+ (List.fold_left2 (match_cases_pattern_no_more_args metas) sigma l1' l2),(le2,more_args)
+ | r1, NList (x,y,iter,termin,revert) ->
(match_cases_pattern_list (match_cases_pattern_no_more_args)
- metas (terms,(),termlists,()) a1 x y iter termin lassoc),(0,[])
+ metas (terms,termlists,(),()) a1 x y iter termin revert),(0,[])
| _ -> raise No_match
and match_cases_pattern_no_more_args metas sigma a1 a2 =
@@ -1300,15 +1375,15 @@ let reorder_canonically_substitution terms termlists metas =
List.fold_right (fun (x,(scl,typ)) (terms',termlists') ->
match typ with
| NtnTypeConstr -> ((Id.List.assoc x terms, scl)::terms',termlists')
- | NtnTypeOnlyBinder -> assert false
+ | NtnTypeBinder _ -> assert false
| NtnTypeConstrList -> (terms',(Id.List.assoc x termlists,scl)::termlists')
| NtnTypeBinderList -> assert false)
metas ([],[])
let match_notation_constr_cases_pattern c (metas,pat) =
- let (terms,(),termlists,()),more_args = match_cases_pattern metas ([],(),[],()) c pat in
+ let (terms,termlists,(),()),more_args = match_cases_pattern metas ([],[],(),()) c pat in
reorder_canonically_substitution terms termlists metas, more_args
let match_notation_constr_ind_pattern ind args (metas,pat) =
- let (terms,(),termlists,()),more_args = match_ind_pattern metas ([],(),[],()) ind args pat in
+ let (terms,termlists,(),()),more_args = match_ind_pattern metas ([],[],(),()) ind args pat in
reorder_canonically_substitution terms termlists metas, more_args
diff --git a/interp/notation_ops.mli b/interp/notation_ops.mli
index 0904a4ea3..746f52e48 100644
--- a/interp/notation_ops.mli
+++ b/interp/notation_ops.mli
@@ -29,12 +29,15 @@ val ldots_var : Id.t
bound by the notation; also interpret recursive patterns *)
val notation_constr_of_glob_constr : notation_interp_env ->
- glob_constr -> notation_constr * reversibility_flag
+ glob_constr -> notation_constr * reversibility_status
(** Re-interpret a notation as a [glob_constr], taking care of binders *)
+val apply_cases_pattern : ?loc:Loc.t ->
+ (Id.t list * cases_pattern_disjunction) * Id.t -> glob_constr -> glob_constr
+
val glob_constr_of_notation_constr_with_binders : ?loc:Loc.t ->
- ('a -> Name.t -> 'a * Name.t) ->
+ ('a -> Name.t -> 'a * ((Id.t list * cases_pattern_disjunction) * Id.t) option * Name.t) ->
('a -> notation_constr -> glob_constr) ->
'a -> notation_constr -> glob_constr
@@ -49,6 +52,7 @@ exception No_match
val match_notation_constr : bool -> 'a glob_constr_g -> interpretation ->
('a glob_constr_g * subscopes) list * ('a glob_constr_g list * subscopes) list *
+ ('a cases_pattern_disjunction_g * subscopes) list *
('a extended_glob_local_binder_g list * subscopes) list
val match_notation_constr_cases_pattern :
diff --git a/interp/ppextend.ml b/interp/ppextend.ml
index ce19dd8a9..606196fcd 100644
--- a/interp/ppextend.ml
+++ b/interp/ppextend.ml
@@ -33,6 +33,7 @@ let ppcmd_of_cut = function
type unparsing =
| UnpMetaVar of int * parenRelation
+ | UnpBinderMetaVar of int * parenRelation
| UnpListMetaVar of int * parenRelation * unparsing list
| UnpBinderListMetaVar of int * bool * unparsing list
| UnpTerminal of string
diff --git a/interp/ppextend.mli b/interp/ppextend.mli
index 7b62a2074..77823e32a 100644
--- a/interp/ppextend.mli
+++ b/interp/ppextend.mli
@@ -26,6 +26,7 @@ val ppcmd_of_cut : ppcut -> Pp.t
type unparsing =
| UnpMetaVar of int * parenRelation
+ | UnpBinderMetaVar of int * parenRelation
| UnpListMetaVar of int * parenRelation * unparsing list
| UnpBinderListMetaVar of int * bool * unparsing list
| UnpTerminal of string
diff --git a/interp/reserve.ml b/interp/reserve.ml
index 22c5a2f5e..04710282f 100644
--- a/interp/reserve.ml
+++ b/interp/reserve.ml
@@ -71,7 +71,7 @@ let reserve_revtable = Summary.ref KeyMap.empty ~name:"reserved-type-rev"
let notation_constr_key = function (* Rem: NApp(NRef ref,[]) stands for @ref *)
| NApp (NRef ref,args) -> RefKey(canonical_gr ref), Some (List.length args)
| NList (_,_,NApp (NRef ref,args),_,_)
- | NBinderList (_,_,NApp (NRef ref,args),_) -> RefKey (canonical_gr ref), Some (List.length args)
+ | NBinderList (_,_,NApp (NRef ref,args),_,_) -> RefKey (canonical_gr ref), Some (List.length args)
| NRef ref -> RefKey(canonical_gr ref), None
| _ -> Oth, None
diff --git a/intf/constrexpr.ml b/intf/constrexpr.ml
index fbf9e248a..c59828794 100644
--- a/intf/constrexpr.ml
+++ b/intf/constrexpr.ml
@@ -46,7 +46,7 @@ type prim_token =
type instance_expr = Misctypes.glob_level list
type cases_pattern_expr_r =
- | CPatAlias of cases_pattern_expr * Id.t
+ | CPatAlias of cases_pattern_expr * Name.t Loc.located
| CPatCstr of reference
* cases_pattern_expr list option * cases_pattern_expr list
(** [CPatCstr (_, c, Some l1, l2)] represents (@c l1) l2 *)
@@ -70,8 +70,8 @@ and constr_expr_r =
| CRef of reference * instance_expr option
| CFix of Id.t Loc.located * fix_expr list
| CCoFix of Id.t Loc.located * cofix_expr list
- | CProdN of binder_expr list * constr_expr
- | CLambdaN of binder_expr list * constr_expr
+ | CProdN of local_binder_expr list * constr_expr
+ | CLambdaN of local_binder_expr list * constr_expr
| CLetIn of Name.t Loc.located * constr_expr * constr_expr option * constr_expr
| CAppExpl of (proj_flag * reference * instance_expr option) * constr_expr list
| CApp of (proj_flag * constr_expr) *
@@ -107,9 +107,6 @@ and case_expr = constr_expr (* expression that is being matched
and branch_expr =
(cases_pattern_expr list list * constr_expr) Loc.located
-and binder_expr =
- Name.t Loc.located list * binder_kind * constr_expr
-
and fix_expr =
Id.t Loc.located * (Id.t Loc.located option * recursion_order_expr) *
local_binder_expr list * constr_expr * constr_expr
@@ -131,7 +128,8 @@ and local_binder_expr =
and constr_notation_substitution =
constr_expr list * (** for constr subterms *)
constr_expr list list * (** for recursive notations *)
- local_binder_expr list list (** for binders subexpressions *)
+ cases_pattern_expr list * (** for binders *)
+ local_binder_expr list list (** for binder lists (recursive notations) *)
type constr_pattern_expr = constr_expr
diff --git a/intf/extend.ml b/intf/extend.ml
index 5552bed55..78f0aa117 100644
--- a/intf/extend.ml
+++ b/intf/extend.ml
@@ -29,29 +29,48 @@ type production_level =
| NextLevel
| NumLevel of int
-type ('lev,'pos) constr_entry_key_gen =
- | ETName | ETReference | ETBigint
- | ETBinder of bool
- | ETConstr of ('lev * 'pos)
- | ETPattern
+type constr_as_binder_kind =
+ | AsIdent
+ | AsIdentOrPattern
+ | AsStrictPattern
+
+(** User-level types used to tell how to parse or interpret of the non-terminal *)
+
+type 'a constr_entry_key_gen =
+ | ETName
+ | ETReference
+ | ETBigint
+ | ETBinder of bool (* open list of binders if true, closed list of binders otherwise *)
+ | ETConstr of 'a
+ | ETConstrAsBinder of constr_as_binder_kind * 'a
+ | ETPattern of bool * int option (* true = strict pattern, i.e. not a single variable *)
| ETOther of string * string
- | ETConstrList of ('lev * 'pos) * Tok.t list
- | ETBinderList of bool * Tok.t list
-(** Entries level (left-hand-side of grammar rules) *)
+(** Entries level (left-hand side of grammar rules) *)
type constr_entry_key =
- (int,unit) constr_entry_key_gen
-
-(** Entries used in productions (in right-hand-side of grammar rules) *)
-
-type constr_prod_entry_key =
- (production_level,production_position) constr_entry_key_gen
+ (production_level * production_position) constr_entry_key_gen
(** Entries used in productions, vernac side (e.g. "x bigint" or "x ident") *)
type simple_constr_prod_entry_key =
- (production_level,unit) constr_entry_key_gen
+ production_level option constr_entry_key_gen
+
+(** Entries used in productions (in right-hand-side of grammar rules), to parse non-terminals *)
+
+type binder_entry_kind = ETBinderOpen | ETBinderClosed of Tok.t list
+
+type binder_target = ForBinder | ForTerm
+
+type constr_prod_entry_key =
+ | ETProdName (* Parsed as a name (ident or _) *)
+ | ETProdReference (* Parsed as a global reference *)
+ | ETProdBigint (* Parsed as an (unbounded) integer *)
+ | ETProdConstr of (production_level * production_position) (* Parsed as constr or pattern *)
+ | ETProdPattern of int (* Parsed as pattern as a binder (as subpart of a constr) *)
+ | ETProdOther of string * string (* Intended for embedding custom entries in constr or pattern *)
+ | ETProdConstrList of (production_level * production_position) * Tok.t list (* Parsed as non-empty list of constr *)
+ | ETProdBinderList of binder_entry_kind (* Parsed as non-empty list of local binders *)
(** {5 AST for user-provided entries} *)
diff --git a/intf/glob_term.ml b/intf/glob_term.ml
index 61bbe2c26..3f48fa547 100644
--- a/intf/glob_term.ml
+++ b/intf/glob_term.ml
@@ -105,7 +105,7 @@ type cases_pattern_disjunction = [ `any ] cases_pattern_disjunction_g
type 'a extended_glob_local_binder_r =
| GLocalAssum of Name.t * binding_kind * 'a glob_constr_g
| GLocalDef of Name.t * binding_kind * 'a glob_constr_g * 'a glob_constr_g option
- | GLocalPattern of ('a cases_pattern_g * Id.t list) * Id.t * binding_kind * 'a glob_constr_g
+ | GLocalPattern of ('a cases_pattern_disjunction_g * Id.t list) * Id.t * binding_kind * 'a glob_constr_g
and 'a extended_glob_local_binder_g = ('a extended_glob_local_binder_r, 'a) DAst.t
type extended_glob_local_binder = [ `any ] extended_glob_local_binder_g
diff --git a/intf/intf.mllib b/intf/intf.mllib
index 38a2a71cc..2b8960d3f 100644
--- a/intf/intf.mllib
+++ b/intf/intf.mllib
@@ -2,9 +2,9 @@ Constrexpr
Evar_kinds
Genredexpr
Locus
+Extend
Notation_term
Decl_kinds
-Extend
Glob_term
Misctypes
Pattern
diff --git a/intf/notation_term.ml b/intf/notation_term.ml
index cad6f4b82..86f5adbd7 100644
--- a/intf/notation_term.ml
+++ b/intf/notation_term.ml
@@ -25,11 +25,11 @@ type notation_constr =
| NVar of Id.t
| NApp of notation_constr * notation_constr list
| NHole of Evar_kinds.t * Misctypes.intro_pattern_naming_expr * Genarg.glob_generic_argument option
- | NList of Id.t * Id.t * notation_constr * notation_constr * bool
+ | NList of Id.t * Id.t * notation_constr * notation_constr * (* associativity: *) bool
(** Part only in [glob_constr] *)
| NLambda of Name.t * notation_constr * notation_constr
| NProd of Name.t * notation_constr * notation_constr
- | NBinderList of Id.t * Id.t * notation_constr * notation_constr
+ | NBinderList of Id.t * Id.t * notation_constr * notation_constr * (* associativity: *) bool
| NLetIn of Name.t * notation_constr * notation_constr option * notation_constr
| NCases of Constr.case_style * notation_constr option *
(notation_constr * (Name.t * (inductive * Name.t list) option)) list *
@@ -60,21 +60,31 @@ type subscopes = tmp_scope_name option * scope_name list
(** Type of the meta-variables of an notation_constr: in a recursive pattern x..y,
x carries the sequence of objects bound to the list x..y *)
+
+type notation_binder_source =
+ (* This accepts only pattern *)
+ (* NtnParsedAsPattern true means only strict pattern (no single variable) at printing *)
+ | NtnParsedAsPattern of bool
+ (* This accepts only ident *)
+ | NtnParsedAsIdent
+ (* This accepts ident, or pattern, or both *)
+ | NtnBinderParsedAsConstr of Extend.constr_as_binder_kind
+
type notation_var_instance_type =
- | NtnTypeConstr | NtnTypeOnlyBinder | NtnTypeConstrList | NtnTypeBinderList
+ | NtnTypeConstr | NtnTypeBinder of notation_binder_source | NtnTypeConstrList | NtnTypeBinderList
-(** Type of variables when interpreting a constr_expr as an notation_constr:
+(** Type of variables when interpreting a constr_expr as a notation_constr:
in a recursive pattern x..y, both x and y carry the individual type
of each element of the list x..y *)
type notation_var_internalization_type =
- | NtnInternTypeConstr | NtnInternTypeBinder | NtnInternTypeIdent
+ | NtnInternTypeAny | NtnInternTypeOnlyBinder
(** This characterizes to what a notation is interpreted to *)
type interpretation =
(Id.t * (subscopes * notation_var_instance_type)) list *
notation_constr
-type reversibility_flag = bool
+type reversibility_status = APrioriReversible | HasLtac | NonInjective of Id.t list
type notation_interp_env = {
ninterp_var_type : notation_var_internalization_type Id.Map.t;
@@ -95,7 +105,7 @@ type precedence = int
type parenRelation = L | E | Any | Prec of precedence
type tolerability = precedence * parenRelation
-type level = precedence * tolerability list * notation_var_internalization_type list
+type level = precedence * tolerability list * Extend.constr_entry_key list
(** Grammar rules for a notation *)
diff --git a/intf/vernacexpr.ml b/intf/vernacexpr.ml
index 8e0fe54c5..d16c9bb80 100644
--- a/intf/vernacexpr.ml
+++ b/intf/vernacexpr.ml
@@ -214,6 +214,7 @@ type proof_expr =
type syntax_modifier =
| SetItemLevel of string list * Extend.production_level
+ | SetItemLevelAsBinder of string list * Extend.constr_as_binder_kind * Extend.production_level option
| SetLevel of int
| SetAssoc of Extend.gram_assoc
| SetEntryType of string * Extend.simple_constr_prod_entry_key
diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml
index d0d5cb1d5..e95d5d2b5 100644
--- a/kernel/subtyping.ml
+++ b/kernel/subtyping.ml
@@ -118,6 +118,15 @@ let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2
let env = check_polymorphic_instance error env auctx auctx' in
env, Univ.make_abstract_instance auctx'
| Cumulative_ind cumi, Cumulative_ind cumi' ->
+ (** Currently there is no way to control variance of inductive types, but
+ just in case we require that they are in a subtyping relation. *)
+ let () =
+ let v = ACumulativityInfo.variance cumi in
+ let v' = ACumulativityInfo.variance cumi' in
+ if not (Array.for_all2 Variance.check_subtype v' v) then
+ CErrors.anomaly Pp.(str "Variance of " ++ KerName.print kn1 ++
+ str " is not compatible with the one of " ++ KerName.print kn2)
+ in
let auctx = Univ.ACumulativityInfo.univ_context cumi in
let auctx' = Univ.ACumulativityInfo.univ_context cumi' in
let env = check_polymorphic_instance error env auctx auctx' in
diff --git a/kernel/univ.ml b/kernel/univ.ml
index fbb047364..c42b66749 100644
--- a/kernel/univ.ml
+++ b/kernel/univ.ml
@@ -759,6 +759,13 @@ struct
| Invariant, _ | _, Invariant -> Invariant
| Covariant, Covariant -> Covariant
+ let check_subtype x y = match x, y with
+ | (Irrelevant | Covariant | Invariant), Irrelevant -> true
+ | Irrelevant, Covariant -> false
+ | (Covariant | Invariant), Covariant -> true
+ | (Irrelevant | Covariant), Invariant -> false
+ | Invariant, Invariant -> true
+
let pr = function
| Irrelevant -> str "*"
| Covariant -> str "+"
diff --git a/kernel/univ.mli b/kernel/univ.mli
index c45ebe21c..74d1bfd3a 100644
--- a/kernel/univ.mli
+++ b/kernel/univ.mli
@@ -253,6 +253,9 @@ sig
A'] as opposed to [A' <= A]. *)
type t = Irrelevant | Covariant | Invariant
+ (** [check_subtype x y] holds if variance [y] is also an instance of [x] *)
+ val check_subtype : t -> t -> bool
+
val sup : t -> t -> t
val pr : t -> Pp.t
diff --git a/parsing/egramcoq.ml b/parsing/egramcoq.ml
index ea6266dd4..a3d257154 100644
--- a/parsing/egramcoq.ml
+++ b/parsing/egramcoq.ml
@@ -229,11 +229,11 @@ type (_, _) entry =
| TTName : ('self, Name.t Loc.located) entry
| TTReference : ('self, reference) entry
| TTBigint : ('self, Constrexpr.raw_natural_number) entry
-| TTBinder : ('self, local_binder_expr list) entry
| TTConstr : prod_info * 'r target -> ('r, 'r) entry
| TTConstrList : prod_info * Tok.t list * 'r target -> ('r, 'r list) entry
-| TTBinderListT : ('self, local_binder_expr list) entry
-| TTBinderListF : Tok.t list -> ('self, local_binder_expr list list) entry
+| TTPattern : int -> ('self, cases_pattern_expr) entry
+| TTOpenBinderList : ('self, local_binder_expr list) entry
+| TTClosedBinderList : Tok.t list -> ('self, local_binder_expr list list) entry
type _ any_entry = TTAny : ('s, 'r) entry -> 's any_entry
@@ -289,31 +289,24 @@ let symbol_of_entry : type s r. _ -> _ -> (s, r) entry -> (s, r) symbol = fun as
Alist1 (symbol_of_target typ' assoc from forpat)
| TTConstrList (typ', tkl, forpat) ->
Alist1sep (symbol_of_target typ' assoc from forpat, make_sep_rules tkl)
-| TTBinderListF [] -> Alist1 (Aentry Constr.binder)
-| TTBinderListF tkl -> Alist1sep (Aentry Constr.binder, make_sep_rules tkl)
+| TTPattern p -> Aentryl (Constr.pattern, p)
+| TTClosedBinderList [] -> Alist1 (Aentry Constr.binder)
+| TTClosedBinderList tkl -> Alist1sep (Aentry Constr.binder, make_sep_rules tkl)
| TTName -> Aentry Prim.name
-| TTBinder -> Aentry Constr.binder
-| TTBinderListT -> Aentry Constr.open_binders
+| TTOpenBinderList -> Aentry Constr.open_binders
| TTBigint -> Aentry Prim.bigint
| TTReference -> Aentry Constr.global
let interp_entry forpat e = match e with
-| ETName -> TTAny TTName
-| ETReference -> TTAny TTReference
-| ETBigint -> TTAny TTBigint
-| ETBinder true -> anomaly (Pp.str "Should occur only as part of BinderList.")
-| ETBinder false -> TTAny TTBinder
-| ETConstr p -> TTAny (TTConstr (p, forpat))
-| ETPattern -> assert false (** not used *)
-| ETOther _ -> assert false (** not used *)
-| ETConstrList (p, tkl) -> TTAny (TTConstrList (p, tkl, forpat))
-| ETBinderList (true, []) -> TTAny TTBinderListT
-| ETBinderList (true, _) -> assert false
-| ETBinderList (false, tkl) -> TTAny (TTBinderListF tkl)
-
-let constr_expr_of_name (loc,na) = CAst.make ?loc @@ match na with
- | Anonymous -> CHole (None,Misctypes.IntroAnonymous,None)
- | Name id -> CRef (Ident (Loc.tag ?loc id), None)
+| ETProdName -> TTAny TTName
+| ETProdReference -> TTAny TTReference
+| ETProdBigint -> TTAny TTBigint
+| ETProdConstr p -> TTAny (TTConstr (p, forpat))
+| ETProdPattern p -> TTAny (TTPattern p)
+| ETProdOther _ -> assert false (** not used *)
+| ETProdConstrList (p, tkl) -> TTAny (TTConstrList (p, tkl, forpat))
+| ETProdBinderList ETBinderOpen -> TTAny TTOpenBinderList
+| ETProdBinderList (ETBinderClosed tkl) -> TTAny (TTClosedBinderList tkl)
let cases_pattern_expr_of_name (loc,na) = CAst.make ?loc @@ match na with
| Anonymous -> CPatAtom None
@@ -322,7 +315,8 @@ let cases_pattern_expr_of_name (loc,na) = CAst.make ?loc @@ match na with
type 'r env = {
constrs : 'r list;
constrlists : 'r list list;
- binders : (local_binder_expr list * bool) list;
+ binders : cases_pattern_expr list;
+ binderlists : local_binder_expr list list;
}
let push_constr subst v = { subst with constrs = v :: subst.constrs }
@@ -332,12 +326,16 @@ match e with
| TTConstr _ -> push_constr subst v
| TTName ->
begin match forpat with
- | ForConstr -> push_constr subst (constr_expr_of_name v)
+ | ForConstr -> { subst with binders = cases_pattern_expr_of_name v :: subst.binders }
| ForPattern -> push_constr subst (cases_pattern_expr_of_name v)
end
-| TTBinder -> { subst with binders = (v, true) :: subst.binders }
-| TTBinderListT -> { subst with binders = (v, true) :: subst.binders }
-| TTBinderListF _ -> { subst with binders = (List.flatten v, false) :: subst.binders }
+| TTPattern _ ->
+ begin match forpat with
+ | ForConstr -> { subst with binders = v :: subst.binders }
+ | ForPattern -> push_constr subst v
+ end
+| TTOpenBinderList -> { subst with binderlists = v :: subst.binderlists }
+| TTClosedBinderList _ -> { subst with binderlists = List.flatten v :: subst.binderlists }
| TTBigint ->
begin match forpat with
| ForConstr -> push_constr subst (CAst.make @@ CPrim (Numeral (v,true)))
@@ -437,11 +435,9 @@ let rec pure_sublevels : type a b c. int option -> (a, b, c) rule -> int list =
let make_act : type r. r target -> _ -> r gen_eval = function
| ForConstr -> fun notation loc env ->
- let env = (env.constrs, env.constrlists, List.map fst env.binders) in
- CAst.make ~loc @@ CNotation (notation , env)
+ let env = (env.constrs, env.constrlists, env.binders, env.binderlists) in
+ CAst.make ~loc @@ CNotation (notation, env)
| ForPattern -> fun notation loc env ->
- let invalid = List.exists (fun (_, b) -> not b) env.binders in
- let () = if invalid then Constrexpr_ops.error_invalid_pattern_notation ~loc () in
let env = (env.constrs, env.constrlists) in
CAst.make ~loc @@ CPatNotation (notation, env, [])
@@ -457,7 +453,7 @@ let extend_constr state forpat ng =
let needed_levels, state = register_empty_levels state isforpat pure_sublevels in
let (pos,p4assoc,name,reinit), state = find_position state isforpat assoc level in
let empty_rules = List.map (prepare_empty_levels isforpat) needed_levels in
- let empty = { constrs = []; constrlists = []; binders = [] } in
+ let empty = { constrs = []; constrlists = []; binders = []; binderlists = [] } in
let act = ty_eval r (make_act forpat ng.notgram_notation) empty in
let rule = (name, p4assoc, [Rule (symbs, act)]) in
let r = ExtendRule (entry, reinit, (pos, [rule])) in
diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4
index db68a75e0..9f12db649 100644
--- a/parsing/g_constr.ml4
+++ b/parsing/g_constr.ml4
@@ -120,7 +120,7 @@ let name_colon =
| _ -> err ())
| _ -> err ())
-let aliasvar = function { CAst.loc = loc; CAst.v = CPatAlias (_, id) } -> Some (loc,Name id) | _ -> None
+let aliasvar = function { CAst.v = CPatAlias (_, na) } -> Some na | _ -> None
GEXTEND Gram
GLOBAL: binder_constr lconstr constr operconstr universe_level sort sort_family
@@ -216,9 +216,11 @@ GEXTEND Gram
| "("; c = operconstr LEVEL "200"; ")" ->
(match c.CAst.v with
| CPrim (Numeral (n,true)) ->
- CAst.make ~loc:(!@loc) @@ CNotation("( _ )",([c],[],[]))
+ CAst.make ~loc:(!@loc) @@ CNotation("( _ )",([c],[],[],[]))
| _ -> c)
| "{|"; c = record_declaration; "|}" -> c
+ | "{"; c = binder_constr ; "}" ->
+ CAst.make ~loc:(!@loc) @@ CNotation(("{ _ }"),([c],[],[],[]))
| "`{"; c = operconstr LEVEL "200"; "}" ->
CAst.make ~loc:(!@loc) @@ CGeneralization (Implicit, None, c)
| "`("; c = operconstr LEVEL "200"; ")" ->
@@ -385,19 +387,9 @@ GEXTEND Gram
| "99" RIGHTA [ ]
| "90" RIGHTA [ ]
| "10" LEFTA
- [ p = pattern; "as"; id = ident ->
- CAst.make ~loc:!@loc @@ CPatAlias (p, id)
- | p = pattern; lp = LIST1 NEXT ->
- (let open CAst in match p with
- | { v = CPatAtom (Some r) } -> CAst.make ~loc:!@loc @@ CPatCstr (r, None, lp)
- | { v = CPatCstr (r, None, l2); loc } ->
- CErrors.user_err ?loc ~hdr:"compound_pattern"
- (Pp.str "Nested applications not supported.")
- | { v = CPatCstr (r, l1, l2) } -> CAst.make ~loc:!@loc @@ CPatCstr (r, l1 , l2@lp)
- | { v = CPatNotation (n, s, l) } -> CAst.make ~loc:!@loc @@ CPatNotation (n , s, l@lp)
- | _ -> CErrors.user_err
- ?loc:(cases_pattern_expr_loc p) ~hdr:"compound_pattern"
- (Pp.str "Such pattern cannot have arguments."))
+ [ p = pattern; "as"; na = name ->
+ CAst.make ~loc:!@loc @@ CPatAlias (p, na)
+ | p = pattern; lp = LIST1 NEXT -> mkAppPattern ~loc:!@loc p lp
| "@"; r = Prim.reference; lp = LIST0 NEXT ->
CAst.make ~loc:!@loc @@ CPatCstr (r, Some lp, []) ]
| "1" LEFTA
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index d42b5f622..d90fd3eb7 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -1174,6 +1174,8 @@ GEXTEND Gram
| x = IDENT; ","; l = LIST1 [id = IDENT -> id ] SEP ","; "at";
lev = level -> SetItemLevel (x::l,lev)
| x = IDENT; "at"; lev = level -> SetItemLevel ([x],lev)
+ | x = IDENT; "at"; lev = level; b = constr_as_binder_kind -> SetItemLevelAsBinder ([x],b,Some lev)
+ | x = IDENT; b = constr_as_binder_kind -> SetItemLevelAsBinder ([x],b,None)
| x = IDENT; typ = syntax_extension_type -> SetEntryType (x,typ)
] ]
;
@@ -1181,7 +1183,20 @@ GEXTEND Gram
[ [ IDENT "ident" -> ETName | IDENT "global" -> ETReference
| IDENT "bigint" -> ETBigint
| IDENT "binder" -> ETBinder true
+ | IDENT "constr"; n = OPT at_level; b = constr_as_binder_kind -> ETConstrAsBinder (b,n)
+ | IDENT "pattern" -> ETPattern (false,None)
+ | IDENT "pattern"; "at"; IDENT "level"; n = natural -> ETPattern (false,Some n)
+ | IDENT "strict"; IDENT "pattern" -> ETPattern (true,None)
+ | IDENT "strict"; IDENT "pattern"; "at"; IDENT "level"; n = natural -> ETPattern (true,Some n)
| IDENT "closed"; IDENT "binder" -> ETBinder false
] ]
;
+ at_level:
+ [ [ "at"; n = level -> n ] ]
+ ;
+ constr_as_binder_kind:
+ [ [ "as"; IDENT "ident" -> AsIdent
+ | "as"; IDENT "pattern" -> AsIdentOrPattern
+ | "as"; IDENT "strict"; IDENT "pattern" -> AsStrictPattern ] ]
+ ;
END
diff --git a/parsing/pcoq.ml b/parsing/pcoq.ml
index 54e7949ae..ddb26d771 100644
--- a/parsing/pcoq.ml
+++ b/parsing/pcoq.ml
@@ -638,3 +638,15 @@ let () =
Grammar.register0 wit_constr (Constr.constr);
Grammar.register0 wit_red_expr (Vernac_.red_expr);
()
+
+(** Registering extra grammar *)
+
+type any_entry = AnyEntry : 'a Gram.entry -> any_entry
+
+let grammar_names : any_entry list String.Map.t ref = ref String.Map.empty
+
+let register_grammars_by_name name grams =
+ grammar_names := String.Map.add name grams !grammar_names
+
+let find_grammars_by_name name =
+ String.Map.find name !grammar_names
diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli
index 75378d2c6..accb51366 100644
--- a/parsing/pcoq.mli
+++ b/parsing/pcoq.mli
@@ -315,3 +315,10 @@ val (!@) : Ploc.t -> Loc.t
type frozen_t
val parser_summary_tag : frozen_t Summary.Dyn.tag
+
+(** Registering grammars by name *)
+
+type any_entry = AnyEntry : 'a Gram.entry -> any_entry
+
+val register_grammars_by_name : string -> any_entry list -> unit
+val find_grammars_by_name : string -> any_entry list
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index f7639d22d..693ab464d 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -1295,8 +1295,8 @@ let rec rebuild_return_type rt =
CAst.make ?loc @@ Constrexpr.CProdN(n,rebuild_return_type t')
| Constrexpr.CLetIn(na,v,t,t') ->
CAst.make ?loc @@ Constrexpr.CLetIn(na,v,t,rebuild_return_type t')
- | _ -> CAst.make ?loc @@ Constrexpr.CProdN([[Loc.tag Anonymous],
- Constrexpr.Default Decl_kinds.Explicit, rt],
+ | _ -> CAst.make ?loc @@ Constrexpr.CProdN([Constrexpr.CLocalAssum ([Loc.tag Anonymous],
+ Constrexpr.Default Decl_kinds.Explicit, rt)],
CAst.make @@ Constrexpr.CSort(GType []))
let do_build_inductive
@@ -1356,7 +1356,7 @@ let do_build_inductive
acc)
| None ->
CAst.make @@ Constrexpr.CProdN
- ([[(Loc.tag n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t],
+ ([Constrexpr.CLocalAssum([(Loc.tag n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t)],
acc
)
)
@@ -1423,7 +1423,7 @@ let do_build_inductive
acc)
| None ->
CAst.make @@ Constrexpr.CProdN
- ([[(Loc.tag n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t],
+ ([Constrexpr.CLocalAssum([(Loc.tag n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t)],
acc
)
)
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index 58154d310..96e4a94d3 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -591,11 +591,11 @@ and rebuild_nal aux bk bl' nal typ =
match nal,typ with
| _,{ CAst.v = CProdN([],typ) } -> rebuild_nal aux bk bl' nal typ
| [], _ -> rebuild_bl aux bl' typ
- | na::nal,{ CAst.v = CProdN((na'::nal',bk',nal't)::rest,typ') } ->
+ | na::nal,{ CAst.v = CProdN(CLocalAssum(na'::nal',bk',nal't)::rest,typ') } ->
if Name.equal (snd na) (snd na') || Name.is_anonymous (snd na')
then
let assum = CLocalAssum([na],bk,nal't) in
- let new_rest = if nal' = [] then rest else ((nal',bk',nal't)::rest) in
+ let new_rest = if nal' = [] then rest else (CLocalAssum(nal',bk',nal't)::rest) in
rebuild_nal
(assum::aux)
bk
@@ -604,7 +604,7 @@ and rebuild_nal aux bk bl' nal typ =
(CAst.make @@ CProdN(new_rest,typ'))
else
let assum = CLocalAssum([na'],bk,nal't) in
- let new_rest = if nal' = [] then rest else ((nal',bk',nal't)::rest) in
+ let new_rest = if nal' = [] then rest else (CLocalAssum(nal',bk',nal't)::rest) in
rebuild_nal
(assum::aux)
bk
@@ -731,10 +731,14 @@ let rec add_args id new_args = CAst.map (function
end
| CFix _ | CCoFix _ -> anomaly ~label:"add_args " (Pp.str "todo.")
| CProdN(nal,b1) ->
- CProdN(List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal,
+ CProdN(List.map (function CLocalAssum (nal,k,b2) -> CLocalAssum (nal,k,add_args id new_args b2)
+ | CLocalDef (na,b1,t) -> CLocalDef (na,add_args id new_args b1,Option.map (add_args id new_args) t)
+ | CLocalPattern _ -> user_err (Pp.str "pattern with quote not allowed here.")) nal,
add_args id new_args b1)
| CLambdaN(nal,b1) ->
- CLambdaN(List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal,
+ CLambdaN(List.map (function CLocalAssum (nal,k,b2) -> CLocalAssum (nal,k,add_args id new_args b2)
+ | CLocalDef (na,b1,t) -> CLocalDef (na,add_args id new_args b1,Option.map (add_args id new_args) t)
+ | CLocalPattern _ -> user_err (Pp.str "pattern with quote not allowed here.")) nal,
add_args id new_args b1)
| CLetIn(na,b1,t,b2) ->
CLetIn(na,add_args id new_args b1,Option.map (add_args id new_args) t,add_args id new_args b2)
@@ -793,7 +797,7 @@ let rec chop_n_arrow n t =
then t (* If we have already removed all the arrows then return the type *)
else (* If not we check the form of [t] *)
match t.CAst.v with
- | Constrexpr.CProdN(nal_ta',t') -> (* If we have a forall, to result are possible :
+ | Constrexpr.CProdN(nal_ta',t') -> (* If we have a forall, two results are possible :
either we need to discard more than the number of arrows contained
in this product declaration then we just recall [chop_n_arrow] on
the remaining number of arrow to chop and [t'] we discard it and
@@ -805,7 +809,7 @@ let rec chop_n_arrow n t =
let new_n =
let rec aux (n:int) = function
[] -> n
- | (nal,k,t'')::nal_ta' ->
+ | CLocalAssum(nal,k,t'')::nal_ta' ->
let nal_l = List.length nal in
if n >= nal_l
then
@@ -813,9 +817,10 @@ let rec chop_n_arrow n t =
else
let new_t' = CAst.make @@
Constrexpr.CProdN(
- ((snd (List.chop n nal)),k,t'')::nal_ta',t')
+ CLocalAssum((snd (List.chop n nal)),k,t'')::nal_ta',t')
in
raise (Stop new_t')
+ | _ -> anomaly (Pp.str "Not enough products.")
in
aux n nal_ta'
in
@@ -828,16 +833,13 @@ let rec chop_n_arrow n t =
let rec get_args b t : Constrexpr.local_binder_expr list *
Constrexpr.constr_expr * Constrexpr.constr_expr =
match b.CAst.v with
- | Constrexpr.CLambdaN ((nal_ta), b') ->
+ | Constrexpr.CLambdaN (CLocalAssum(nal,k,ta) as d::rest, b') ->
begin
- let n =
- (List.fold_left (fun n (nal,_,_) ->
- n+List.length nal) 0 nal_ta )
- in
- let nal_tas,b'',t'' = get_args b' (chop_n_arrow n t) in
- (List.map (fun (nal,k,ta) ->
- (Constrexpr.CLocalAssum (nal,k,ta))) nal_ta)@nal_tas, b'',t''
+ let n = List.length nal in
+ let nal_tas,b'',t'' = get_args (CAst.make ?loc:b.CAst.loc @@ Constrexpr.CLambdaN (rest,b')) (chop_n_arrow n t) in
+ d :: nal_tas, b'',t''
end
+ | Constrexpr.CLambdaN ([], b) -> [],b,t
| _ -> [],b,t
diff --git a/plugins/ltac/g_tactic.ml4 b/plugins/ltac/g_tactic.ml4
index d792d4ff7..e68140828 100644
--- a/plugins/ltac/g_tactic.ml4
+++ b/plugins/ltac/g_tactic.ml4
@@ -115,10 +115,11 @@ let mk_fix_tac (loc,id,bl,ann,ty) =
match bl,ann with
[([_],_,_)], None -> 1
| _, Some x ->
- let ids = List.map snd (List.flatten (List.map pi1 bl)) in
+ let ids = List.map snd (List.flatten (List.map (fun (nal,_,_) -> nal) bl)) in
(try List.index Names.Name.equal (snd x) ids
with Not_found -> user_err Pp.(str "No such fix variable."))
| _ -> user_err Pp.(str "Cannot guess decreasing argument of fix.") in
+ let bl = List.map (fun (nal,bk,t) -> CLocalAssum (nal,bk,t)) bl in
(id,n, CAst.make ~loc @@ CProdN(bl,ty))
let mk_cofix_tac (loc,id,bl,ann,ty) =
@@ -126,6 +127,7 @@ let mk_cofix_tac (loc,id,bl,ann,ty) =
user_err ~loc:aloc
~hdr:"Constr:mk_cofix_tac"
(Pp.str"Annotation forbidden in cofix expression.")) ann in
+ let bl = List.map (fun (nal,bk,t) -> CLocalAssum (nal,bk,t)) bl in
(id,CAst.make ~loc @@ CProdN(bl,ty))
(* Functions overloaded by quotifier *)
@@ -160,7 +162,7 @@ let mkTacCase with_evar = function
let rec mkCLambdaN_simple_loc ?loc bll c =
match bll with
| ((loc1,_)::_ as idl,bk,t) :: bll ->
- CAst.make ?loc @@ CLambdaN ([idl,bk,t],mkCLambdaN_simple_loc ?loc:(Loc.merge_opt loc1 loc) bll c)
+ CAst.make ?loc @@ CLambdaN ([CLocalAssum (idl,bk,t)],mkCLambdaN_simple_loc ?loc:(Loc.merge_opt loc1 loc) bll c)
| ([],_,_) :: bll -> mkCLambdaN_simple_loc ?loc bll c
| [] -> c
diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml
index e5ff47356..4f430b79e 100644
--- a/plugins/ltac/pptactic.ml
+++ b/plugins/ltac/pptactic.ml
@@ -353,9 +353,10 @@ let string_of_genarg_arg (ArgumentType arg) =
let rec strip_ty acc n ty =
match ty.CAst.v with
Constrexpr.CProdN(bll,a) ->
- let nb =
- List.fold_left (fun i (nal,_,_) -> i + List.length nal) 0 bll in
- let bll = List.map (fun (x, _, y) -> x, y) bll in
+ let bll = List.map (function
+ | CLocalAssum (nal,_,t) -> nal,t
+ | _ -> user_err Pp.(str "Cannot translate fix tactic: not only products")) bll in
+ let nb = List.fold_left (fun i (nal,t) -> i + List.length nal) 0 bll in
if nb >= n then (List.rev (bll@acc)), a
else strip_ty (bll@acc) (n-nb) a
| _ -> user_err Pp.(str "Cannot translate fix tactic: not enough products") in
diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml
index ee84be541..8112cc400 100644
--- a/plugins/ltac/tacentries.ml
+++ b/plugins/ltac/tacentries.ml
@@ -545,11 +545,10 @@ let print_located_tactic qid =
(** Grammar *)
let () =
- let open Metasyntax in
let entries = [
AnyEntry Pltac.tactic_expr;
AnyEntry Pltac.binder_tactic;
AnyEntry Pltac.simple_tactic;
AnyEntry Pltac.tactic_arg;
] in
- register_grammar "tactic" entries
+ register_grammars_by_name "tactic" entries
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
index 9822da1c7..499154d50 100644
--- a/plugins/ssr/ssrcommon.ml
+++ b/plugins/ssr/ssrcommon.ml
@@ -835,9 +835,9 @@ let rec mkCHoles ?loc n =
if n <= 0 then [] else (CAst.make ?loc @@ CHole (None, IntroAnonymous, None)) :: mkCHoles ?loc (n - 1)
let mkCHole loc = CAst.make ?loc @@ CHole (None, IntroAnonymous, None)
let mkCLambda ?loc name ty t = CAst.make ?loc @@
- CLambdaN ([[loc, name], Default Explicit, ty], t)
+ CLambdaN ([CLocalAssum([loc, name], Default Explicit, ty)], t)
let mkCArrow ?loc ty t = CAst.make ?loc @@
- CProdN ([[Loc.tag Anonymous], Default Explicit, ty], t)
+ CProdN ([CLocalAssum([Loc.tag Anonymous], Default Explicit, ty)], t)
let mkCCast ?loc t ty = CAst.make ?loc @@ CCast (t, CastConv ty)
let rec isCHoles = function { CAst.v = CHole _ } :: cl -> isCHoles cl | cl -> cl = []
@@ -855,7 +855,7 @@ let pf_interp_ty ?(resolve_typeclasses=false) ist gl ty =
| _, (_, Some ty) ->
let rec force_type ty = CAst.(map (function
| CProdN (abs, t) ->
- n_binders := !n_binders + List.length (List.flatten (List.map pi1 abs));
+ n_binders := !n_binders + List.length (List.flatten (List.map (function CLocalAssum (nal,_,_) -> nal | CLocalDef (na,_,_) -> [na] | CLocalPattern (_,_) -> (* We count a 'pat for 1; TO BE CHECKED *) [Loc.tag Name.Anonymous]) abs));
CProdN (abs, force_type t)
| CLetIn (n, v, oty, t) -> incr n_binders; CLetIn (n, v, oty, force_type t)
| _ -> (mkCCast ty (mkCType None)).v)) ty in
diff --git a/plugins/ssr/ssrparser.ml4 b/plugins/ssr/ssrparser.ml4
index 46403aef3..211cad3f5 100644
--- a/plugins/ssr/ssrparser.ml4
+++ b/plugins/ssr/ssrparser.ml4
@@ -1023,10 +1023,10 @@ let rec format_local_binders h0 bl0 = match h0, bl0 with
| _ -> []
let rec format_constr_expr h0 c0 = let open CAst in match h0, c0 with
- | BFvar :: h, { v = CLambdaN ([[_, x], _, _], c) } ->
+ | BFvar :: h, { v = CLambdaN ([CLocalAssum([_, x], _, _)], c) } ->
let bs, c' = format_constr_expr h c in
Bvar x :: bs, c'
- | BFdecl _:: h, { v = CLambdaN ([lxs, _, t], c) } ->
+ | BFdecl _:: h, { v = CLambdaN ([CLocalAssum(lxs, _, t)], c) } ->
let bs, c' = format_constr_expr h c in
Bdecl (List.map snd lxs, t) :: bs, c'
| BFdef :: h, { v = CLetIn((_, x), v, oty, c) } ->
@@ -1165,20 +1165,20 @@ ARGUMENT EXTEND ssrbinder TYPED AS ssrfwdfmt * constr PRINTED BY pr_ssrbinder
| [ ssrbvar(bv) ] ->
[ let xloc, _ as x = bvar_lname bv in
(FwdPose, [BFvar]),
- CAst.make ~loc @@ CLambdaN ([[x],Default Explicit,mkCHole xloc],mkCHole (Some loc)) ]
+ CAst.make ~loc @@ CLambdaN ([CLocalAssum([x],Default Explicit,mkCHole xloc)],mkCHole (Some loc)) ]
| [ "(" ssrbvar(bv) ")" ] ->
[ let xloc, _ as x = bvar_lname bv in
(FwdPose, [BFvar]),
- CAst.make ~loc @@ CLambdaN ([[x],Default Explicit,mkCHole xloc],mkCHole (Some loc)) ]
+ CAst.make ~loc @@ CLambdaN ([CLocalAssum([x],Default Explicit,mkCHole xloc)],mkCHole (Some loc)) ]
| [ "(" ssrbvar(bv) ":" lconstr(t) ")" ] ->
[ let x = bvar_lname bv in
(FwdPose, [BFdecl 1]),
- CAst.make ~loc @@ CLambdaN ([[x], Default Explicit, t], mkCHole (Some loc)) ]
+ CAst.make ~loc @@ CLambdaN ([CLocalAssum([x], Default Explicit, t)], mkCHole (Some loc)) ]
| [ "(" ssrbvar(bv) ne_ssrbvar_list(bvs) ":" lconstr(t) ")" ] ->
[ let xs = List.map bvar_lname (bv :: bvs) in
let n = List.length xs in
(FwdPose, [BFdecl n]),
- CAst.make ~loc @@ CLambdaN ([xs, Default Explicit, t], mkCHole (Some loc)) ]
+ CAst.make ~loc @@ CLambdaN ([CLocalAssum (xs, Default Explicit, t)], mkCHole (Some loc)) ]
| [ "(" ssrbvar(id) ":" lconstr(t) ":=" lconstr(v) ")" ] ->
[ (FwdPose,[BFdef]), CAst.make ~loc @@ CLetIn (bvar_lname id, v, Some t, mkCHole (Some loc)) ]
| [ "(" ssrbvar(id) ":=" lconstr(v) ")" ] ->
@@ -1191,7 +1191,7 @@ GEXTEND Gram
[ ["of" | "&"]; c = operconstr LEVEL "99" ->
let loc = !@loc in
(FwdPose, [BFvar]),
- CAst.make ~loc @@ CLambdaN ([[Loc.tag ~loc Anonymous],Default Explicit,c],mkCHole (Some loc)) ]
+ CAst.make ~loc @@ CLambdaN ([CLocalAssum ([Loc.tag ~loc Anonymous],Default Explicit,c)],mkCHole (Some loc)) ]
];
END
@@ -1217,7 +1217,7 @@ let push_binders c2 bs =
| ct -> loop false ct bs
let rec fix_binders = let open CAst in function
- | (_, { v = CLambdaN ([xs, _, t], _) } ) :: bs ->
+ | (_, { v = CLambdaN ([CLocalAssum(xs, _, t)], _) } ) :: bs ->
CLocalAssum (xs, Default Explicit, t) :: fix_binders bs
| (_, { v = CLetIn (x, v, oty, _) } ) :: bs ->
CLocalDef (x, v, oty) :: fix_binders bs
@@ -1325,13 +1325,13 @@ let intro_id_to_binder = List.map (function
| IPatId id ->
let xloc, _ as x = bvar_lname (mkCVar id) in
(FwdPose, [BFvar]),
- CAst.make @@ CLambdaN ([[x], Default Explicit, mkCHole xloc],
+ CAst.make @@ CLambdaN ([CLocalAssum([x], Default Explicit, mkCHole xloc)],
mkCHole None)
| _ -> anomaly "non-id accepted as binder")
let binder_to_intro_id = CAst.(List.map (function
- | (FwdPose, [BFvar]), { v = CLambdaN ([ids,_,_],_) }
- | (FwdPose, [BFdecl _]), { v = CLambdaN ([ids,_,_],_) } ->
+ | (FwdPose, [BFvar]), { v = CLambdaN ([CLocalAssum(ids,_,_)],_) }
+ | (FwdPose, [BFdecl _]), { v = CLambdaN ([CLocalAssum(ids,_,_)],_) } ->
List.map (function (_, Name id) -> IPatId id | _ -> IPatAnon One) ids
| (FwdPose, [BFdef]), { v = CLetIn ((_,Name id),_,_,_) } -> [IPatId id]
| (FwdPose, [BFdef]), { v = CLetIn ((_,Anonymous),_,_,_) } -> [IPatAnon One]
diff --git a/plugins/ssr/ssrvernac.ml4 b/plugins/ssr/ssrvernac.ml4
index d74ad06b3..25d1472f1 100644
--- a/plugins/ssr/ssrvernac.ml4
+++ b/plugins/ssr/ssrvernac.ml4
@@ -74,7 +74,7 @@ let frozen_lexer = CLexer.get_keyword_state () ;;
let no_ct = None, None and no_rt = None in
let aliasvar = function
- | [[{ CAst.v = CPatAlias (_, id); loc }]] -> Some (loc,Name id)
+ | [[{ CAst.v = CPatAlias (_, na); loc }]] -> Some na
| _ -> None in
let mk_cnotype mp = aliasvar mp, None in
let mk_ctype mp t = aliasvar mp, Some t in
@@ -298,7 +298,7 @@ let interp_search_notation ?loc tag okey =
let rec sub () = function
| NVar x when List.mem_assoc x nvars -> DAst.make ?loc @@ GPatVar (FirstOrderPatVar x)
| c ->
- glob_constr_of_notation_constr_with_binders ?loc (fun _ x -> (), x) sub () c in
+ glob_constr_of_notation_constr_with_binders ?loc (fun _ x -> (), None, x) sub () c in
let _, npat = Patternops.pattern_of_glob_constr (sub () body) in
Search.GlobSearchSubPattern npat
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index 7d0584a00..71e8b4ac4 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -137,7 +137,7 @@ let destGLambda c = match DAst.get c with GLambda (Name id, _, _, c) -> (id, c)
let isGHole c = match DAst.get c with GHole _ -> true | _ -> false
let mkCHole ~loc = CAst.make ?loc @@ CHole (None, IntroAnonymous, None)
let mkCLambda ?loc name ty t = CAst.make ?loc @@
- CLambdaN ([[Loc.tag ?loc name], Default Explicit, ty], t)
+ CLambdaN ([CLocalAssum([Loc.tag ?loc name], Default Explicit, ty)], t)
let mkCLetIn ?loc name bo t = CAst.make ?loc @@
CLetIn ((Loc.tag ?loc name), bo, None, t)
let mkCCast ?loc t ty = CAst.make ?loc @@ CCast (t, dC ty)
@@ -916,7 +916,7 @@ let glob_cpattern gs p =
| k, (v, Some t) as orig ->
if k = 'x' then glob_ssrterm gs ('(', (v, Some t)) else
match t.CAst.v with
- | CNotation("( _ in _ )", ([t1; t2], [], [])) ->
+ | CNotation("( _ in _ )", ([t1; t2], [], [], [])) ->
(try match glob t1, glob t2 with
| (r1, None), (r2, None) -> encode k "In" [r1;r2]
| (r1, Some _), (r2, Some _) when isCVar t1 ->
@@ -924,11 +924,11 @@ let glob_cpattern gs p =
| (r1, Some _), (r2, Some _) -> encode k "In" [r1; r2]
| _ -> CErrors.anomaly (str"where are we?.")
with _ when isCVar t1 -> encode k "In" [bind_in t1 t2])
- | CNotation("( _ in _ in _ )", ([t1; t2; t3], [], [])) ->
+ | CNotation("( _ in _ in _ )", ([t1; t2; t3], [], [], [])) ->
check_var t2; encode k "In" [fst (glob t1); bind_in t2 t3]
- | CNotation("( _ as _ )", ([t1; t2], [], [])) ->
+ | CNotation("( _ as _ )", ([t1; t2], [], [], [])) ->
encode k "As" [fst (glob t1); fst (glob t2)]
- | CNotation("( _ as _ in _ )", ([t1; t2; t3], [], [])) ->
+ | CNotation("( _ as _ in _ )", ([t1; t2; t3], [], [], [])) ->
check_var t2; encode k "As" [fst (glob t1); bind_in t2 t3]
| _ -> glob_ssrterm gs orig
;;
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 311c1c09e..a0434f927 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -450,11 +450,6 @@ let current_pattern eqn =
| pat::_ -> pat
| [] -> anomaly (Pp.str "Empty list of patterns.")
-let alias_of_pat = DAst.with_val (function
- | PatVar name -> name
- | PatCstr(_,_,name) -> name
- )
-
let remove_current_pattern eqn =
match eqn.patterns with
| pat::pats ->
diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml
index a21137a05..25817478e 100644
--- a/pretyping/glob_ops.ml
+++ b/pretyping/glob_ops.ml
@@ -19,6 +19,16 @@ open Ltac_pretype
let cases_pattern_loc c = c.CAst.loc
+let alias_of_pat pat = DAst.with_val (function
+ | PatVar name -> name
+ | PatCstr(_,_,name) -> name
+ ) pat
+
+let set_pat_alias id = DAst.map (function
+ | PatVar Anonymous -> PatVar (Name id)
+ | PatCstr (cstr,patl,Anonymous) -> PatCstr (cstr,patl,Name id)
+ | pat -> assert false)
+
let cases_predicate_names tml =
List.flatten (List.map (function
| (tm,(na,None)) -> [na]
@@ -452,6 +462,10 @@ let rec rename_glob_vars l c = force @@ DAst.map_with_loc (fun ?loc -> function
(**********************************************************************)
(* Conversion from glob_constr to cases pattern, if possible *)
+let is_gvar id c = match DAst.get c with
+| GVar id' -> Id.equal id id'
+| _ -> false
+
let rec cases_pattern_of_glob_constr na = DAst.map (function
| GVar id ->
begin match na with
@@ -468,6 +482,9 @@ let rec cases_pattern_of_glob_constr na = DAst.map (function
PatCstr (cstr,List.map (cases_pattern_of_glob_constr Anonymous) l,na)
| _ -> raise Not_found
end
+ | GLetIn (Name id as na',b,None,e) when is_gvar id e && na = Anonymous ->
+ (* A canonical encoding of aliases *)
+ DAst.get (cases_pattern_of_glob_constr na' b)
| _ -> raise Not_found
)
@@ -503,23 +520,34 @@ let add_patterns_for_params_remove_local_defs (ind,j) l =
drop_local_defs typi l in
Util.List.addn nparams (DAst.make @@ PatVar Anonymous) l
+let add_alias ?loc na c =
+ match na with
+ | Anonymous -> c
+ | Name id -> GLetIn (na,DAst.make ?loc c,None,DAst.make ?loc (GVar id))
+
(* Turn a closed cases pattern into a glob_constr *)
-let rec glob_constr_of_closed_cases_pattern_aux x = DAst.map_with_loc (fun ?loc -> function
- | PatCstr (cstr,[],Anonymous) -> GRef (ConstructRef cstr,None)
- | PatCstr (cstr,l,Anonymous) ->
+let rec glob_constr_of_cases_pattern_aux isclosed x = DAst.map_with_loc (fun ?loc -> function
+ | PatCstr (cstr,[],na) -> add_alias ?loc na (GRef (ConstructRef cstr,None))
+ | PatCstr (cstr,l,na) ->
let ref = DAst.make ?loc @@ GRef (ConstructRef cstr,None) in
let l = add_patterns_for_params_remove_local_defs cstr l in
- GApp (ref, List.map glob_constr_of_closed_cases_pattern_aux l)
+ add_alias ?loc na (GApp (ref, List.map (glob_constr_of_cases_pattern_aux isclosed) l))
+ | PatVar (Name id) when not isclosed ->
+ GVar id
+ | PatVar Anonymous when not isclosed ->
+ GHole (Evar_kinds.QuestionMark (Define false,Anonymous),Misctypes.IntroAnonymous,None)
| _ -> raise Not_found
) x
let glob_constr_of_closed_cases_pattern p = match DAst.get p with
| PatCstr (cstr,l,na) ->
let loc = p.CAst.loc in
- na,glob_constr_of_closed_cases_pattern_aux (DAst.make ?loc @@ PatCstr (cstr,l,Anonymous))
+ na,glob_constr_of_cases_pattern_aux true (DAst.make ?loc @@ PatCstr (cstr,l,Anonymous))
| _ ->
raise Not_found
+let glob_constr_of_cases_pattern p = glob_constr_of_cases_pattern_aux false p
+
(**********************************************************************)
(* Interpreting ltac variables *)
diff --git a/pretyping/glob_ops.mli b/pretyping/glob_ops.mli
index 9dd7068cb..0d9fb1f45 100644
--- a/pretyping/glob_ops.mli
+++ b/pretyping/glob_ops.mli
@@ -13,6 +13,10 @@ open Glob_term
val cases_pattern_eq : 'a cases_pattern_g -> 'a cases_pattern_g -> bool
+val alias_of_pat : 'a cases_pattern_g -> Name.t
+
+val set_pat_alias : Id.t -> 'a cases_pattern_g -> 'a cases_pattern_g
+
val cast_type_eq : ('a -> 'a -> bool) ->
'a Misctypes.cast_type -> 'a Misctypes.cast_type -> bool
@@ -78,10 +82,14 @@ val map_pattern : (glob_constr -> glob_constr) ->
Take the current alias as parameter,
@raise Not_found if translation is impossible *)
-val cases_pattern_of_glob_constr : Name.t -> glob_constr -> cases_pattern
+val cases_pattern_of_glob_constr : Name.t -> 'a glob_constr_g -> 'a cases_pattern_g
val glob_constr_of_closed_cases_pattern : 'a cases_pattern_g -> Name.t * 'a glob_constr_g
+(** A canonical encoding of cases pattern into constr such that
+ composed with [cases_pattern_of_glob_constr Anonymous] gives identity *)
+val glob_constr_of_cases_pattern : 'a cases_pattern_g -> 'a glob_constr_g
+
val add_patterns_for_params_remove_local_defs : constructor -> 'a cases_pattern_g list -> 'a cases_pattern_g list
val ltac_interp_name : Ltac_pretype.ltac_var_map -> Name.t -> Name.t
diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml
index 1146b42a0..873505940 100644
--- a/printing/ppconstr.ml
+++ b/printing/ppconstr.ml
@@ -86,8 +86,8 @@ let tag_var = tag Tag.variable
open Notation
- let print_hunks n pr pr_binders (terms, termlists, binders) unps =
- let env = ref terms and envlist = ref termlists and bll = ref binders in
+ let print_hunks n pr pr_patt pr_binders (terms, termlists, binders, binderlists) unps =
+ let env = ref terms and envlist = ref termlists and bl = ref binders and bll = ref binderlists in
let pop r = let a = List.hd !r in r := List.tl !r; a in
let return unp pp1 pp2 = (tag_unparsing unp pp1) ++ pp2 in
(* Warning:
@@ -102,6 +102,11 @@ let tag_var = tag Tag.variable
let pp2 = aux l in
let pp1 = pr (n, prec) c in
return unp pp1 pp2
+ | UnpBinderMetaVar (_, prec) as unp :: l ->
+ let c = pop bl in
+ let pp2 = aux l in
+ let pp1 = pr_patt (n, prec) c in
+ return unp pp1 pp2
| UnpListMetaVar (_, prec, sl) as unp :: l ->
let cl = pop envlist in
let pp1 = prlist_with_sep (fun () -> aux sl) (pr (n,prec)) cl in
@@ -127,9 +132,9 @@ let tag_var = tag Tag.variable
in
aux unps
- let pr_notation pr pr_binders s env =
+ let pr_notation pr pr_patt pr_binders s env =
let unpl, level = find_notation_printing_rule s in
- print_hunks level pr pr_binders env unpl, level
+ print_hunks level pr pr_patt pr_binders env unpl, level
let pr_delimiters key strm =
strm ++ str ("%"^key)
@@ -263,8 +268,8 @@ let tag_var = tag Tag.variable
in
str "{| " ++ prlist_with_sep pr_semicolon pp l ++ str " |}", lpatrec
- | CPatAlias (p, id) ->
- pr_patt mt (las,E) p ++ str " as " ++ pr_id id, las
+ | CPatAlias (p, na) ->
+ pr_patt mt (las,E) p ++ str " as " ++ pr_lname na, las
| CPatCstr (c, None, []) ->
pr_reference c, latom
@@ -292,7 +297,7 @@ let tag_var = tag Tag.variable
pr_patt (fun()->str"(") (max_int,E) p ++ str")", latom
| CPatNotation (s,(l,ll),args) ->
- let strm_not, l_not = pr_notation (pr_patt mt) (fun _ _ _ -> mt()) s (l,ll,[]) in
+ let strm_not, l_not = pr_notation (pr_patt mt) (fun _ _ -> mt ()) (fun _ _ _ -> mt()) s (l,ll,[],[]) in
(if List.is_empty args||prec_less l_not (lapp,L) then strm_not else surround strm_not)
++ prlist (pr_patt spc (lapp,L)) args, if not (List.is_empty args) then lapp else l_not
@@ -394,68 +399,6 @@ let tag_var = tag Tag.variable
if is_open then pr_delimited_binders pr_com_at sep pr_c
else pr_undelimited_binders sep pr_c
- let rec extract_prod_binders = let open CAst in function
- (* | CLetIn (loc,na,b,c) as x ->
- let bl,c = extract_prod_binders c in
- if bl = [] then [], x else CLocalDef (na,b) :: bl, c*)
- | { v = CProdN ([],c) } ->
- extract_prod_binders c
- | { loc; v = CProdN ([[_,Name id],bk,t],
- { v = CCases (LetPatternStyle,None, [{ v = CRef (Ident (_,id'),None)},None,None],[(_,([[p]],b))])} ) }
- when Id.equal id id' && not (Id.Set.mem id (free_vars_of_constr_expr b)) ->
- let bl,c = extract_prod_binders b in
- CLocalPattern (loc, (p,None)) :: bl, c
- | { loc; v = CProdN ((nal,bk,t)::bl,c) } ->
- let bl,c = extract_prod_binders (CAst.make ?loc @@ CProdN(bl,c)) in
- CLocalAssum (nal,bk,t) :: bl, c
- | c -> [], c
-
- let rec extract_lam_binders ce = let open CAst in match ce.v with
- (* | CLetIn (loc,na,b,c) as x ->
- let bl,c = extract_lam_binders c in
- if bl = [] then [], x else CLocalDef (na,b) :: bl, c*)
- | CLambdaN ([],c) ->
- extract_lam_binders c
- | CLambdaN ([[_,Name id],bk,t],
- { v = CCases (LetPatternStyle,None, [{ v = CRef (Ident (_,id'),None)},None,None],[(_,([[p]],b))])} )
- when Id.equal id id' && not (Id.Set.mem id (free_vars_of_constr_expr b)) ->
- let bl,c = extract_lam_binders b in
- CLocalPattern (ce.loc,(p,None)) :: bl, c
- | CLambdaN ((nal,bk,t)::bl,c) ->
- let bl,c = extract_lam_binders (CAst.make ?loc:ce.loc @@ CLambdaN(bl,c)) in
- CLocalAssum (nal,bk,t) :: bl, c
- | _ -> [], ce
-
- let split_lambda = CAst.with_loc_val (fun ?loc -> function
- | CLambdaN ([[na],bk,t],c) -> (na,t,c)
- | CLambdaN (([na],bk,t)::bl,c) -> (na,t, CAst.make ?loc @@ CLambdaN(bl,c))
- | CLambdaN ((na::nal,bk,t)::bl,c) -> (na,t, CAst.make ?loc @@ CLambdaN((nal,bk,t)::bl,c))
- | _ -> anomaly (Pp.str "ill-formed fixpoint body.")
- )
-
- let rename na na' t c =
- match (na,na') with
- | (_,Name id), (_,Name id') ->
- (na',t,replace_vars_constr_expr (Id.Map.singleton id id') c)
- | (_,Name id), (_,Anonymous) -> (na,t,c)
- | _ -> (na',t,c)
-
- let split_product na' = CAst.with_loc_val (fun ?loc -> function
- | CProdN ([[na],bk,t],c) -> rename na na' t c
- | CProdN (([na],bk,t)::bl,c) -> rename na na' t (CAst.make ?loc @@ CProdN(bl,c))
- | CProdN ((na::nal,bk,t)::bl,c) ->
- rename na na' t (CAst.make ?loc @@ CProdN((nal,bk,t)::bl,c))
- | _ -> anomaly (Pp.str "ill-formed fixpoint body.")
- )
-
- let rec split_fix n typ def =
- if Int.equal n 0 then ([],typ,def)
- else
- let (na,_,def) = split_lambda def in
- let (na,t,typ) = split_product na typ in
- let (bl,typ,def) = split_fix (n-1) typ def in
- (CLocalAssum ([na],default_binder_kind,t)::bl,typ,def)
-
let pr_recursive_decl pr pr_dangling dangling_with_for id bl annot t c =
let pr_body =
if dangling_with_for then pr_dangling else pr in
@@ -574,8 +517,7 @@ let tag_var = tag Tag.variable
(pr_cofixdecl (pr mt) (pr_dangling_with_for mt pr)) (snd id) cofix),
lfix
)
- | CProdN _ ->
- let (bl,a) = extract_prod_binders a in
+ | CProdN (bl,a) ->
return (
hov 0 (
hov 2 (pr_delimited_binders pr_forall spc
@@ -583,8 +525,7 @@ let tag_var = tag Tag.variable
str "," ++ pr spc ltop a),
lprod
)
- | CLambdaN _ ->
- let (bl,a) = extract_lam_binders a in
+ | CLambdaN (bl,a) ->
return (
hov 0 (
hov 2 (pr_delimited_binders pr_fun spc
@@ -722,10 +663,10 @@ let tag_var = tag Tag.variable
| CastCoerce -> str ":>"),
lcast
)
- | CNotation ("( _ )",([t],[],[])) ->
+ | CNotation ("( _ )",([t],[],[],[])) ->
return (pr (fun()->str"(") (max_int,L) t ++ str")", latom)
| CNotation (s,env) ->
- pr_notation (pr mt) (pr_binders_gen (pr mt ltop)) s env
+ pr_notation (pr mt) pr_patt (pr_binders_gen (pr mt ltop)) s env
| CGeneralization (bk,ak,c) ->
return (pr_generalization bk ak (pr mt ltop c), latom)
| CPrim p ->
diff --git a/printing/ppconstr.mli b/printing/ppconstr.mli
index 1320cce9b..158906dd2 100644
--- a/printing/ppconstr.mli
+++ b/printing/ppconstr.mli
@@ -17,14 +17,6 @@ open Names
open Misctypes
open Notation_term
-val extract_lam_binders :
- constr_expr -> local_binder_expr list * constr_expr
-val extract_prod_binders :
- constr_expr -> local_binder_expr list * constr_expr
-val split_fix :
- int -> constr_expr -> constr_expr ->
- local_binder_expr list * constr_expr * constr_expr
-
val prec_less : precedence -> tolerability -> bool
val pr_tight_coma : unit -> Pp.t
diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml
index 950246c53..31c0d20f3 100644
--- a/printing/ppvernac.ml
+++ b/printing/ppvernac.ml
@@ -93,16 +93,35 @@ open Decl_kinds
let sep = fun _ -> spc()
let sep_v2 = fun _ -> str"," ++ spc()
- let pr_set_entry_type = function
+ let pr_at_level = function
+ | NumLevel n -> keyword "at" ++ spc () ++ keyword "level" ++ spc () ++ int n
+ | NextLevel -> keyword "at" ++ spc () ++ keyword "next" ++ spc () ++ keyword "level"
+
+ let pr_constr_as_binder_kind = function
+ | AsIdent -> keyword "as ident"
+ | AsIdentOrPattern -> keyword "as pattern"
+ | AsStrictPattern -> keyword "as strict pattern"
+
+ let pr_strict b = if b then str "strict " else mt ()
+
+ let pr_set_entry_type pr = function
| ETName -> str"ident"
| ETReference -> str"global"
- | ETPattern -> str"pattern"
- | ETConstr _ -> str"constr"
+ | ETPattern (b,None) -> pr_strict b ++ str"pattern"
+ | ETPattern (b,Some n) -> pr_strict b ++ str"pattern" ++ spc () ++ pr_at_level (NumLevel n)
+ | ETConstr lev -> str"constr" ++ pr lev
| ETOther (_,e) -> str e
+ | ETConstrAsBinder (bk,lev) -> pr lev ++ spc () ++ pr_constr_as_binder_kind bk
| ETBigint -> str "bigint"
| ETBinder true -> str "binder"
| ETBinder false -> str "closed binder"
- | ETBinderList _ | ETConstrList _ -> failwith "Internal entry type"
+
+ let pr_at_level_opt = function
+ | None -> mt ()
+ | Some n -> spc () ++ pr_at_level n
+
+ let pr_set_simple_entry_type =
+ pr_set_entry_type pr_at_level_opt
let pr_comment pr_c = function
| CommentConstr c -> pr_c c
@@ -360,17 +379,16 @@ open Decl_kinds
let pr_thm_token k = keyword (Kindops.string_of_theorem_kind k)
let pr_syntax_modifier = function
- | SetItemLevel (l,NextLevel) ->
- prlist_with_sep sep_v2 str l ++
- spc() ++ keyword "at next level"
- | SetItemLevel (l,NumLevel n) ->
+ | SetItemLevel (l,n) ->
+ prlist_with_sep sep_v2 str l ++ spc () ++ pr_at_level n
+ | SetItemLevelAsBinder (l,bk,n) ->
prlist_with_sep sep_v2 str l ++
- spc() ++ keyword "at level" ++ spc() ++ int n
- | SetLevel n -> keyword "at level" ++ spc() ++ int n
+ spc() ++ pr_at_level_opt n ++ spc() ++ pr_constr_as_binder_kind bk
+ | SetLevel n -> pr_at_level (NumLevel n)
| SetAssoc LeftA -> keyword "left associativity"
| SetAssoc RightA -> keyword "right associativity"
| SetAssoc NonA -> keyword "no associativity"
- | SetEntryType (x,typ) -> str x ++ spc() ++ pr_set_entry_type typ
+ | SetEntryType (x,typ) -> str x ++ spc() ++ pr_set_simple_entry_type typ
| SetOnlyPrinting -> keyword "only printing"
| SetOnlyParsing -> keyword "only parsing"
| SetCompatVersion v -> keyword("compat \"" ^ Flags.pr_version v ^ "\"")
diff --git a/printing/ppvernac.mli b/printing/ppvernac.mli
index 34b4fb97f..603be6308 100644
--- a/printing/ppvernac.mli
+++ b/printing/ppvernac.mli
@@ -9,6 +9,8 @@
(** This module implements pretty-printers for vernac_expr syntactic
objects and their subcomponents. *)
+val pr_set_entry_type : ('a -> Pp.t) -> 'a Extend.constr_entry_key_gen -> Pp.t
+
(** Prints a fixpoint body *)
val pr_rec_definition : (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) -> Pp.t
diff --git a/test-suite/bugs/closed/5532.v b/test-suite/bugs/closed/5532.v
new file mode 100644
index 000000000..ee5446e54
--- /dev/null
+++ b/test-suite/bugs/closed/5532.v
@@ -0,0 +1,15 @@
+(* A wish granted by the new support for patterns in notations *)
+
+Local Notation mkmatch0 e p
+ := match e with
+ | p => true
+ | _ => false
+ end.
+Local Notation "'mkmatch' [[ e ]] [[ p ]]"
+ := match e with
+ | p => true
+ | _ => false
+ end
+ (at level 0, p pattern).
+Check mkmatch0 _ ((0, 0)%core).
+Check mkmatch [[ _ ]] [[ ((0, 0)%core) ]].
diff --git a/test-suite/bugs/closed/6529.v b/test-suite/bugs/closed/6529.v
new file mode 100644
index 000000000..8d9081999
--- /dev/null
+++ b/test-suite/bugs/closed/6529.v
@@ -0,0 +1,16 @@
+Require Import Vector Program.
+
+Program Definition append_nil_def :=
+ forall A n (ls: t A n), append ls (nil A) = ls. (* Works *)
+
+Lemma append_nil : append_nil_def. (* Works *)
+Proof.
+Admitted.
+
+Program Lemma append_nil' :
+ forall A n (ls: t A n), append ls (nil A) = ls.
+Abort.
+
+Fail Program Lemma append_nil'' :
+ forall A B n (ls: t A n), append ls (nil A) = ls.
+(* Error: Anomaly "Evar ?X25 was not declared." Please report at http://coq.inria.fr/bugs/. *)
diff --git a/test-suite/modules/cumpoly.v b/test-suite/modules/cumpoly.v
new file mode 100644
index 000000000..654b86cb4
--- /dev/null
+++ b/test-suite/modules/cumpoly.v
@@ -0,0 +1,19 @@
+Set Universe Polymorphism.
+
+(** Check that variance subtyping is respected. The signature T is asking for
+ invariance, while M provide an irrelevant implementation, which is deemed
+ legit.
+
+ There is currently no way to go the other way around, so it's not possible
+ to generate a counter-example that should fail with the wrong subtyping.
+*)
+
+Module Type T.
+Parameter t@{i|Set <= i} : Type@{i}.
+Cumulative Inductive I@{i|Set <= i} : Type@{i} := C : t@{i} -> I.
+End T.
+
+Module M : T.
+Definition t@{i|Set <= i} : Type@{i} := nat.
+Cumulative Inductive I@{i|Set <= i} : Type@{i} := C : t@{i} -> I.
+End M.
diff --git a/test-suite/output/Notations.out b/test-suite/output/Notations.out
index 2f0ee765d..891296b0a 100644
--- a/test-suite/output/Notations.out
+++ b/test-suite/output/Notations.out
@@ -41,7 +41,7 @@ fun x : nat => ifn x is succ n then n else 0
-4
: Z
The command has indeed failed with message:
-x should not be bound in a recursive pattern of the right-hand side.
+Cannot find where the recursive pattern starts.
The command has indeed failed with message:
in the right-hand side, y and z should appear in
term position as part of a recursive pattern.
diff --git a/test-suite/output/Notations2.out b/test-suite/output/Notations2.out
index 121a369a9..6ffe56e11 100644
--- a/test-suite/output/Notations2.out
+++ b/test-suite/output/Notations2.out
@@ -17,10 +17,9 @@ fun (P : nat -> nat -> Prop) (x : nat) => exists y, P x y
∃ n p : nat, n + p = 0
: Prop
let a := 0 in
-∃ x y : nat,
-let b := 1 in
-let c := b in
-let d := 2 in ∃ z : nat, let e := 3 in let f := 4 in x + y = z + d
+∃ (x y : nat) (b := 1) (c := b) (d := 2) (z : nat),
+let e := 3 in
+let f := 4 in x + y = z + d
: Prop
∀ n p : nat, n + p = 0
: Prop
diff --git a/test-suite/output/Notations2.v b/test-suite/output/Notations2.v
index 531398bb0..923caedac 100644
--- a/test-suite/output/Notations2.v
+++ b/test-suite/output/Notations2.v
@@ -36,8 +36,9 @@ Check fun P:nat->nat->Prop => fun x:nat => ex (P x).
(* Test notations with binders *)
-Notation "∃ x .. y , P":= (ex (fun x => .. (ex (fun y => P)) ..))
- (x binder, y binder, at level 200, right associativity).
+Notation "∃ x .. y , P":= (ex (fun x => .. (ex (fun y => P)) ..))
+ (x binder, y binder, at level 200, right associativity,
+ format "'[ ' ∃ x .. y ']' , P").
Check (∃ n p, n+p=0).
diff --git a/test-suite/output/Notations3.out b/test-suite/output/Notations3.out
index 1b5725275..e6a6e0288 100644
--- a/test-suite/output/Notations3.out
+++ b/test-suite/output/Notations3.out
@@ -138,3 +138,100 @@ amatch = mmatch 0 (with 0 => 1| 1 => 2 end)
: unit
alist = [0; 1; 2]
: list nat
+! '{{x, y}}, x + y = 0
+ : Prop
+exists x : nat,
+ nat ->
+ exists y : nat,
+ nat ->
+ exists '{{u, t}}, forall z1 : nat, z1 = 0 /\ x + y = 0 /\ u + t = 0
+ : Prop
+exists x : nat,
+ nat ->
+ exists y : nat,
+ nat ->
+ exists '{{z, t}}, forall z2 : nat, z2 = 0 /\ x + y = 0 /\ z + t = 0
+ : Prop
+exists_true '{{x, y}} (u := 0) '{{z, t}}, x + y = 0 /\ z + t = 0
+ : Prop
+exists_true (A : Type) (R : A -> A -> Prop) (_ : Reflexive R),
+(forall x : A, R x x)
+ : Prop
+exists_true (x : nat) (A : Type) (R : A -> A -> Prop)
+(_ : Reflexive R) (y : nat), x + y = 0 -> forall z : A, R z z
+ : Prop
+{{{{True, nat -> True}}, nat -> True}}
+ : Prop * Prop * Prop
+{{D 1, 2}}
+ : nat * nat * (nat * nat * (nat * nat))
+! a b : nat # True #
+ : Prop * (Prop * Prop)
+!!!! a b : nat # True #
+ : Prop * Prop * (Prop * Prop * Prop)
+@@ a b : nat # a = b # b = a #
+ : Prop * Prop
+exists_non_null x y z t : nat , x = y /\ z = t
+ : Prop
+forall_non_null x y z t : nat , x = y /\ z = t
+ : Prop
+{{RL 1, 2}}
+ : nat * (nat * nat)
+{{RR 1, 2}}
+ : nat * nat * nat
+@pair nat (prod nat nat) (S (S O)) (@pair nat nat (S O) O)
+ : prod nat (prod nat nat)
+@pair (prod nat nat) nat (@pair nat nat O (S (S O))) (S O)
+ : prod (prod nat nat) nat
+{{RLRR 1, 2}}
+ : nat * (nat * nat) * (nat * nat * nat) * (nat * (nat * nat)) *
+ (nat * nat * nat)
+pair
+ (pair
+ (pair (pair (S (S O)) (pair (S O) O)) (pair (pair O (S (S O))) (S O)))
+ (pair (S O) (pair (S (S O)) O))) (pair (pair O (S O)) (S (S O)))
+ : prod
+ (prod (prod (prod nat (prod nat nat)) (prod (prod nat nat) nat))
+ (prod nat (prod nat nat))) (prod (prod nat nat) nat)
+fun x : nat => if x is n .+ 1 then n else 1
+ : nat -> nat
+{'{{x, y}} : nat * nat | x + y = 0}
+ : Set
+exists2' {{x, y}}, x = 0 & y = 0
+ : Prop
+myexists2 x : nat * nat,
+ let '{{y, z}} := x in y > z & let '{{y, z}} := x in z > y
+ : Prop
+fun '({{x, y}} as z) => x + y = 0 /\ z = z
+ : nat * nat -> Prop
+myexists ({{x, y}} as z), x + y = 0 /\ z = z
+ : Prop
+exists '({{x, y}} as z), x + y = 0 /\ z = z
+ : Prop
+∀ '({{x, y}} as z), x + y = 0 /\ z = z
+ : Prop
+fun '({{{{x, y}}, true}} | {{{{x, y}}, false}}) => x + y
+ : nat * nat * bool -> nat
+myexists ({{{{x, y}}, true}} | {{{{x, y}}, false}}), x > y
+ : Prop
+exists '({{{{x, y}}, true}} | {{{{x, y}}, false}}), x > y
+ : Prop
+∀ '({{{{x, y}}, true}} | {{{{x, y}}, false}}), x > y
+ : Prop
+fun p : nat => if p is S n then n else 0
+ : nat -> nat
+fun p : comparison => if p is Lt then 1 else 0
+ : comparison -> nat
+fun S : nat => [S | S + S]
+ : nat -> nat * (nat -> nat)
+fun N : nat => [N | N + 0]
+ : nat -> nat * (nat -> nat)
+fun S : nat => [[S | S + S]]
+ : nat -> nat * (nat -> nat)
+{I : nat | I = I}
+ : Set
+{'I : True | I = I}
+ : Prop
+{'{{x, y}} : nat * nat | x + y = 0}
+ : Set
+exists2 '{{y, z}} : nat * nat, y > z & z > y
+ : Prop
diff --git a/test-suite/output/Notations3.v b/test-suite/output/Notations3.v
index a8d6c97fb..c98bfff41 100644
--- a/test-suite/output/Notations3.v
+++ b/test-suite/output/Notations3.v
@@ -253,3 +253,170 @@ Definition alist := [0;1;2].
Print alist.
End B.
+
+(* Test contraction of "forall x, let 'pat := x in ..." into "forall 'pat, ..." *)
+(* for isolated "forall" (was not working already in 8.6) *)
+Notation "! x .. y , A" := (id (forall x, .. (id (forall y, A)) .. )) (at level 200, x binder).
+Check ! '(x,y), x+y=0.
+
+(* Check that the terminator of a recursive pattern is interpreted in
+ the correct environment of bindings *)
+Notation "'exists_mixed' x .. y , P" := (ex (fun x => forall z:nat, .. (ex (fun y => forall z:nat, z=0 /\ P)) ..)) (at level 200, x binder).
+Check exists_mixed x y '(u,t), x+y=0/\u+t=0.
+Check exists_mixed x y '(z,t), x+y=0/\z+t=0.
+
+(* Check that intermediary let-in are inserted inbetween instances of
+ the repeated pattern *)
+Notation "'exists_true' x .. y , P" := (exists x, True /\ .. (exists y, True /\ P) ..) (at level 200, x binder).
+Check exists_true '(x,y) (u:=0) '(z,t), x+y=0/\z+t=0.
+
+(* Check that generalized binders are correctly interpreted *)
+
+Module G.
+Generalizable Variables A R.
+Class Reflexive {A:Type} (R : A->A->Prop) := reflexivity : forall x : A, R x x.
+Check exists_true `{Reflexive A R}, forall x, R x x.
+Check exists_true x `{Reflexive A R} y, x+y=0 -> forall z, R z z.
+End G.
+
+(* Allows recursive patterns for binders to be associative on the left *)
+Notation "!! x .. y # A #" := (.. (A,(forall x, True)) ..,(forall y, True)) (at level 200, x binder).
+Check !! a b : nat # True #.
+
+(* Examples where the recursive pattern refer several times to the recursive variable *)
+
+Notation "{{D x , .. , y }}" := ((x,x), .. ((y,y),(0,0)) ..).
+Check {{D 1, 2 }}.
+
+Notation "! x .. y # A #" :=
+ ((forall x, x=x), .. ((forall y, y=y), A) ..)
+ (at level 200, x binder).
+Check ! a b : nat # True #.
+
+Notation "!!!! x .. y # A #" :=
+ (((forall x, x=x),(forall x, x=0)), .. (((forall y, y=y),(forall y, y=0)), A) ..)
+ (at level 200, x binder).
+Check !!!! a b : nat # True #.
+
+Notation "@@ x .. y # A # B #" :=
+ ((forall x, .. (forall y, A) ..), (forall x, .. (forall y, B) ..))
+ (at level 200, x binder).
+Check @@ a b : nat # a=b # b=a #.
+
+Notation "'exists_non_null' x .. y , P" :=
+ (ex (fun x => x <> 0 /\ .. (ex (fun y => y <> 0 /\ P)) ..))
+ (at level 200, x binder).
+Check exists_non_null x y z t , x=y/\z=t.
+
+Notation "'forall_non_null' x .. y , P" :=
+ (forall x, x <> 0 -> .. (forall y, y <> 0 -> P) ..)
+ (at level 200, x binder).
+Check forall_non_null x y z t , x=y/\z=t.
+
+(* Examples where the recursive pattern is in reverse order *)
+
+Notation "{{RL c , .. , d }}" := (pair d .. (pair c 0) ..).
+Check {{RL 1 , 2}}.
+
+Notation "{{RR c , .. , d }}" := (pair .. (pair 0 d) .. c).
+Check {{RR 1 , 2}}.
+
+Set Printing All.
+Check {{RL 1 , 2}}.
+Check {{RR 1 , 2}}.
+Unset Printing All.
+
+Notation "{{RLRR c , .. , d }}" := (pair d .. (pair c 0) .., pair .. (pair 0 d) .. c, pair c .. (pair d 0) .., pair .. (pair 0 c) .. d).
+Check {{RLRR 1 , 2}}.
+Unset Printing Notations.
+Check {{RLRR 1 , 2}}.
+Set Printing Notations.
+
+(* Check insensitivity of "match" clauses to order *)
+
+Notation "'if' t 'is' n .+ 1 'then' p 'else' q" :=
+ (match t with S n => p | 0 => q end)
+ (at level 200).
+Check fun x => if x is n.+1 then n else 1.
+
+(* Examples with binding patterns *)
+
+Check {'(x,y)|x+y=0}.
+
+Module D.
+Notation "'exists2'' x , p & q" := (ex2 (fun x => p) (fun x => q))
+ (at level 200, x pattern, p at level 200, right associativity,
+ format "'[' 'exists2'' '/ ' x , '/ ' '[' p & '/' q ']' ']'")
+ : type_scope.
+
+Check exists2' (x,y), x=0 & y=0.
+End D.
+
+(* Ensuring for reparsability that printer of notations does not use a
+ pattern where only an ident could be reparsed *)
+
+Module E.
+Inductive myex2 {A:Type} (P Q:A -> Prop) : Prop :=
+ myex_intro2 : forall x:A, P x -> Q x -> myex2 P Q.
+Notation "'myexists2' x : A , p & q" := (myex2 (A:=A) (fun x => p) (fun x => q))
+ (at level 200, x ident, A at level 200, p at level 200, right associativity,
+ format "'[' 'myexists2' '/ ' x : A , '/ ' '[' p & '/' q ']' ']'")
+ : type_scope.
+Check myex2 (fun x => let '(y,z) := x in y>z) (fun x => let '(y,z) := x in z>y).
+End E.
+
+(* A canonical example of a notation with a non-recursive binder *)
+
+Parameter myex : forall {A}, (A -> Prop) -> Prop.
+Notation "'myexists' x , p" := (myex (fun x => p))
+ (at level 200, x pattern, p at level 200, right associativity).
+
+(* A canonical example of a notation with recursive binders *)
+
+Notation "∀ x .. y , P" := (forall x, .. (forall y, P) ..)
+ (at level 200, x binder, y binder, right associativity) : type_scope.
+
+(* Check that printing 'pat uses an "as" when the variable bound to
+ the pattern is dependent. We check it for the three kinds of
+ notations involving bindings of patterns *)
+
+Check fun '((x,y) as z) => x+y=0/\z=z. (* Primitive fun/forall *)
+Check myexists ((x,y) as z), x+y=0/\z=z. (* Isolated binding pattern *)
+Check exists '((x,y) as z), x+y=0/\z=z. (* Applicative recursive binder *)
+Check ∀ '((x,y) as z), x+y=0/\z=z. (* Other example of recursive binder, now treated as the exists case *)
+
+(* Check parsability and printability of irrefutable disjunctive patterns *)
+
+Check fun '(((x,y),true)|((x,y),false)) => x+y.
+Check myexists (((x,y),true)|((x,y),false)), x>y.
+Check exists '(((x,y),true)|((x,y),false)), x>y.
+Check ∀ '(((x,y),true)|((x,y),false)), x>y.
+
+(* Check Georges' printability of a "if is then else" notation *)
+
+Notation "'if' c 'is' p 'then' u 'else' v" :=
+ (match c with p => u | _ => v end)
+ (at level 200, p pattern at level 100).
+Check fun p => if p is S n then n else 0.
+Check fun p => if p is Lt then 1 else 0.
+
+(* Check that mixed binders and terms defaults to ident and not pattern *)
+Module F.
+ (* First without an indirection *)
+Notation "[ n | t ]" := (n, (fun n : nat => t)).
+Check fun S : nat => [ S | S+S ].
+Check fun N : nat => (N, (fun n => n+0)). (* another test in passing *)
+ (* Then with an indirection *)
+Notation "[[ n | p | t ]]" := (n, (fun p : nat => t)).
+Notation "[[ n | t ]]" := [[ n | n | t ]].
+Check fun S : nat => [[ S | S+S ]].
+End F.
+
+(* Check parsability/printability of {x|P} and variants *)
+
+Check {I:nat|I=I}.
+Check {'I:True|I=I}.
+Check {'(x,y)|x+y=0}.
+
+(* Check exists2 with a pattern *)
+Check ex2 (fun x => let '(y,z) := x in y>z) (fun x => let '(y,z) := x in z>y).
diff --git a/test-suite/output/PatternsInBinders.out b/test-suite/output/PatternsInBinders.out
index 95be04c32..8a6d94c73 100644
--- a/test-suite/output/PatternsInBinders.out
+++ b/test-suite/output/PatternsInBinders.out
@@ -31,7 +31,7 @@ exists '(x, y) '(z, w), swap (x, y) = (z, w)
: Prop
both_z =
fun pat : nat * nat =>
-let '(n, p) as pat0 := pat return (F pat0) in (Z n, Z p) : F (n, p)
+let '(n, p) as x := pat return (F x) in (Z n, Z p) : F (n, p)
: forall pat : nat * nat, F pat
fun '(x, y) '(z, t) => swap (x, y) = (z, t)
: A * B -> B * A -> Prop
@@ -39,3 +39,9 @@ forall '(x, y) '(z, t), swap (x, y) = (z, t)
: Prop
fun (pat : nat) '(x, y) => x + y = pat
: nat -> nat * nat -> Prop
+f = fun x : nat => x + x
+ : nat -> nat
+
+Argument scope is [nat_scope]
+fun x : nat => x + x
+ : nat -> nat
diff --git a/test-suite/output/PatternsInBinders.v b/test-suite/output/PatternsInBinders.v
index 0bad472f4..d671053c0 100644
--- a/test-suite/output/PatternsInBinders.v
+++ b/test-suite/output/PatternsInBinders.v
@@ -67,3 +67,8 @@ End Suboptimal.
(** Test risk of collision for internal name *)
Check fun pat => fun '(x,y) => x+y = pat.
+
+(** Test name in degenerate case *)
+Definition f 'x := x+x.
+Print f.
+Check fun 'x => x+x.
diff --git a/test-suite/success/Notations2.v b/test-suite/success/Notations2.v
index 2655b651a..7c2cf3ee5 100644
--- a/test-suite/success/Notations2.v
+++ b/test-suite/success/Notations2.v
@@ -92,16 +92,37 @@ Check ##### 0 _ 0%bool 0%bool : prod' bool bool.
Check fun A (x :prod' bool A) => match x with ##### 0 _ y 0%bool => 2 | _ => 1 end.
(* 10. Check computation of binding variable through other notations *)
-(* i should be detected as binding variable and the scopes not being checked *)
-
+(* it should be detected as binding variable and the scopes not being checked *)
Notation "'FUNNAT' i => t" := (fun i : nat => i = t) (at level 200).
Notation "'Funnat' i => t" := (FUNNAT i => t + i%nat) (at level 200).
(* 11. Notations with needed factorization of a recursive pattern *)
(* See https://github.com/coq/coq/issues/6078#issuecomment-342287412 *)
-Module A.
+Module M11.
Notation "[:: x1 ; .. ; xn & s ]" := (cons x1 .. (cons xn s) ..).
Notation "[:: x1 ; .. ; xn ]" := (cons x1 .. (cons xn nil) ..).
Check [:: 1 ; 2 ; 3 ].
Check [:: 1 ; 2 ; 3 & nil ]. (* was failing *)
-End A.
+End M11.
+
+(* 12. Preventively check that a variable which does not occur can be instantiated *)
+(* by any term. In particular, it should not be restricted to a binder *)
+Module M12.
+Notation "N ++ x" := (S x) (only parsing).
+Check 2 ++ 0.
+End M12.
+
+(* 13. Check that internal data about associativity are not used in comparing levels *)
+Module M13.
+Notation "x ;; z" := (x + z)
+ (at level 100, z at level 200, only parsing, right associativity).
+Notation "x ;; z" := (x * z)
+ (at level 100, z at level 200, only parsing) : foo_scope.
+End M13.
+
+(* 14. Check that a notation with a "ident" binder does not include a pattern *)
+Module M14.
+Notation "'myexists' x , p" := (ex (fun x => p))
+ (at level 200, x ident, p at level 200, right associativity) : type_scope.
+Check myexists I, I = 0. (* Should not be seen as a constructor *)
+End M14.
diff --git a/theories/Init/Logic.v b/theories/Init/Logic.v
index 037d37daf..053ed601f 100644
--- a/theories/Init/Logic.v
+++ b/theories/Init/Logic.v
@@ -267,6 +267,13 @@ Notation "'exists2' x : A , p & q" := (ex2 (A:=A) (fun x => p) (fun x => q))
format "'[' 'exists2' '/ ' x : A , '/ ' '[' p & '/' q ']' ']'")
: type_scope.
+Notation "'exists2' ' x , p & q" := (ex2 (fun x => p) (fun x => q))
+ (at level 200, x strict pattern, p at level 200, right associativity) : type_scope.
+Notation "'exists2' ' x : A , p & q" := (ex2 (A:=A) (fun x => p) (fun x => q))
+ (at level 200, x strict pattern, A at level 200, p at level 200, right associativity,
+ format "'[' 'exists2' '/ ' ' x : A , '/ ' '[' p & '/' q ']' ']'")
+ : type_scope.
+
(** Derived rules for universal quantification *)
Section universal_quantification.
diff --git a/theories/Init/Notations.v b/theories/Init/Notations.v
index 5e8d2faa5..a9051e761 100644
--- a/theories/Init/Notations.v
+++ b/theories/Init/Notations.v
@@ -78,6 +78,33 @@ Reserved Notation "{ x : A | P & Q }" (at level 0, x at level 99).
Reserved Notation "{ x : A & P }" (at level 0, x at level 99).
Reserved Notation "{ x : A & P & Q }" (at level 0, x at level 99).
+Reserved Notation "{ ' pat | P }"
+ (at level 0, pat strict pattern, format "{ ' pat | P }").
+Reserved Notation "{ ' pat | P & Q }"
+ (at level 0, pat strict pattern, format "{ ' pat | P & Q }").
+
+Reserved Notation "{ ' pat : A | P }"
+ (at level 0, pat strict pattern, format "{ ' pat : A | P }").
+Reserved Notation "{ ' pat : A | P & Q }"
+ (at level 0, pat strict pattern, format "{ ' pat : A | P & Q }").
+
+Reserved Notation "{ ' pat : A & P }"
+ (at level 0, pat strict pattern, format "{ ' pat : A & P }").
+Reserved Notation "{ ' pat : A & P & Q }"
+ (at level 0, pat strict pattern, format "{ ' pat : A & P & Q }").
+
+(** Support for Gonthier-Ssreflect's "if c is pat then u else v" *)
+
+Module IfNotations.
+
+Notation "'if' c 'is' p 'then' u 'else' v" :=
+ (match c with p => u | _ => v end)
+ (at level 200, p pattern at level 100).
+
+End IfNotations.
+
+(** Scopes *)
+
Delimit Scope type_scope with type.
Delimit Scope function_scope with function.
Delimit Scope core_scope with core.
diff --git a/theories/Init/Specif.v b/theories/Init/Specif.v
index 3b4f833a3..47e8a7558 100644
--- a/theories/Init/Specif.v
+++ b/theories/Init/Specif.v
@@ -53,6 +53,15 @@ Notation "{ x : A & P }" := (sigT (A:=A) (fun x => P)) : type_scope.
Notation "{ x : A & P & Q }" := (sigT2 (A:=A) (fun x => P) (fun x => Q)) :
type_scope.
+Notation "{ ' pat | P }" := (sig (fun pat => P)) : type_scope.
+Notation "{ ' pat | P & Q }" := (sig2 (fun pat => P) (fun pat => Q)) : type_scope.
+Notation "{ ' pat : A | P }" := (sig (A:=A) (fun pat => P)) : type_scope.
+Notation "{ ' pat : A | P & Q }" := (sig2 (A:=A) (fun pat => P) (fun pat => Q)) :
+ type_scope.
+Notation "{ ' pat : A & P }" := (sigT (A:=A) (fun pat => P)) : type_scope.
+Notation "{ ' pat : A & P & Q }" := (sigT2 (A:=A) (fun pat => P) (fun pat => Q)) :
+ type_scope.
+
Add Printing Let sig.
Add Printing Let sig2.
Add Printing Let sigT.
diff --git a/theories/Unicode/Utf8_core.v b/theories/Unicode/Utf8_core.v
index a0545c0a4..95c8336d2 100644
--- a/theories/Unicode/Utf8_core.v
+++ b/theories/Unicode/Utf8_core.v
@@ -10,10 +10,12 @@
(* Logic *)
-Notation "∀ x .. y , P" := (forall x, .. (forall y, P) ..)
- (at level 200, x binder, y binder, right associativity) : type_scope.
-Notation "∃ x .. y , P" := (exists x, .. (exists y, P) ..)
- (at level 200, x binder, y binder, right associativity) : type_scope.
+Notation "∀ x .. y , P" := (forall x, .. (forall y, P) ..)
+ (at level 200, x binder, y binder, right associativity,
+ format "'[ ' ∀ x .. y ']' , P") : type_scope.
+Notation "∃ x .. y , P" := (exists x, .. (exists y, P) ..)
+ (at level 200, x binder, y binder, right associativity,
+ format "'[ ' ∃ x .. y ']' , P") : type_scope.
Notation "x ∨ y" := (x \/ y) (at level 85, right associativity) : type_scope.
Notation "x ∧ y" := (x /\ y) (at level 80, right associativity) : type_scope.
@@ -25,5 +27,6 @@ Notation "¬ x" := (~x) (at level 75, right associativity) : type_scope.
Notation "x ≠ y" := (x <> y) (at level 70) : type_scope.
(* Abstraction *)
-Notation "'λ' x .. y , t" := (fun x => .. (fun y => t) ..)
- (at level 200, x binder, y binder, right associativity).
+Notation "'λ' x .. y , t" := (fun x => .. (fun y => t) ..)
+ (at level 200, x binder, y binder, right associativity,
+ format "'[ ' 'λ' x .. y ']' , t").
diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml
index 3be357598..bcffe640c 100644
--- a/vernac/metasyntax.ml
+++ b/vernac/metasyntax.ml
@@ -43,13 +43,6 @@ let add_token_obj s = Lib.add_anonymous_leaf (inToken s)
let entry_buf = Buffer.create 64
-type any_entry = AnyEntry : 'a Pcoq.Gram.entry -> any_entry
-
-let grammars : any_entry list String.Map.t ref = ref String.Map.empty
-
-let register_grammar name grams =
- grammars := String.Map.add name grams !grammars
-
let pr_entry e =
let () = Buffer.clear entry_buf in
let ft = Format.formatter_of_buffer entry_buf in
@@ -57,11 +50,11 @@ let pr_entry e =
str (Buffer.contents entry_buf)
let pr_registered_grammar name =
- let gram = try Some (String.Map.find name !grammars) with Not_found -> None in
+ let gram = try Some (Pcoq.find_grammars_by_name name) with Not_found -> None in
match gram with
| None -> user_err Pp.(str "Unknown or unprintable grammar entry.")
| Some entries ->
- let pr_one (AnyEntry e) =
+ let pr_one (Pcoq.AnyEntry e) =
str "Entry " ++ str (Pcoq.Gram.Entry.name e) ++ str " is" ++ fnl () ++
pr_entry e
in
@@ -199,36 +192,6 @@ let parse_format ((loc, str) : lstring) =
(***********************)
(* Analyzing notations *)
-type symbol_token = WhiteSpace of int | String of string
-
-let split_notation_string str =
- let push_token beg i l =
- if Int.equal beg i then l else
- let s = String.sub str beg (i - beg) in
- String s :: l
- in
- let push_whitespace beg i l =
- if Int.equal beg i then l else WhiteSpace (i-beg) :: l
- in
- let rec loop beg i =
- if i < String.length str then
- if str.[i] == ' ' then
- push_token beg i (loop_on_whitespace (i+1) (i+1))
- else
- loop beg (i+1)
- else
- push_token beg i []
- and loop_on_whitespace beg i =
- if i < String.length str then
- if str.[i] != ' ' then
- push_whitespace beg i (loop i (i+1))
- else
- loop_on_whitespace beg (i+1)
- else
- push_whitespace beg i []
- in
- loop 0 0
-
(* Interpret notations with a recursive component *)
let out_nt = function NonTerminal x -> x | _ -> assert false
@@ -284,17 +247,6 @@ let quote_notation_token x =
if (n > 0 && norm) || (n > 2 && x.[0] == '\'') then "'"^x^"'"
else x
-let rec raw_analyze_notation_tokens = function
- | [] -> []
- | String ".." :: sl -> NonTerminal ldots_var :: raw_analyze_notation_tokens sl
- | String "_" :: _ -> user_err Pp.(str "_ must be quoted.")
- | String x :: sl when CLexer.is_ident x ->
- NonTerminal (Names.Id.of_string x) :: raw_analyze_notation_tokens sl
- | String s :: sl ->
- Terminal (String.drop_simple_quotes s) :: raw_analyze_notation_tokens sl
- | WhiteSpace n :: sl ->
- Break n :: raw_analyze_notation_tokens sl
-
let is_numeral symbs =
match List.filter (function Break _ -> false | _ -> true) symbs with
| ([Terminal "-"; Terminal x] | [Terminal x]) ->
@@ -315,8 +267,8 @@ let rec get_notation_vars onlyprint = function
| (Terminal _ | Break _) :: sl -> get_notation_vars onlyprint sl
| SProdList _ :: _ -> assert false
-let analyze_notation_tokens ~onlyprint l =
- let l = raw_analyze_notation_tokens l in
+let analyze_notation_tokens ~onlyprint ntn =
+ let l = decompose_raw_notation ntn in
let vars = get_notation_vars onlyprint l in
let recvars,l = interp_list_parser [] l in
recvars, List.subtract Id.equal vars (List.map snd recvars), l
@@ -333,13 +285,17 @@ let prec_assoc = function
| LeftA -> (E,L)
| NonA -> (L,L)
-let precedence_of_entry_type from = function
- | ETConstr (NumLevel n,BorderProd (_,None)) -> n, Prec n
- | ETConstr (NumLevel n,BorderProd (b,Some a)) ->
+let precedence_of_position_and_level from = function
+ | NumLevel n, BorderProd (_,None) -> n, Prec n
+ | NumLevel n, BorderProd (b,Some a) ->
n, let (lp,rp) = prec_assoc a in if b == Left then lp else rp
- | ETConstr (NumLevel n,InternalProd) -> n, Prec n
- | ETConstr (NextLevel,_) -> from, L
- | _ -> 0, E (* ?? *)
+ | NumLevel n, InternalProd -> n, Prec n
+ | NextLevel, _ -> from, L
+
+let precedence_of_entry_type from = function
+ | ETConstr x | ETConstrAsBinder (_,x) -> precedence_of_position_and_level from x
+ | ETPattern (_,n) -> let n = match n with None -> 0 | Some n -> n in n, Prec n
+ | _ -> 0, E (* should not matter *)
(* Some breaking examples *)
(* "x = y" : "x /1 = y" (breaks before any symbol) *)
@@ -377,8 +333,8 @@ let is_non_terminal = function
| NonTerminal _ | SProdList _ -> true
| _ -> false
-let is_next_non_terminal = function
-| [] -> false
+let is_next_non_terminal b = function
+| [] -> b
| pr :: _ -> is_non_terminal pr
let is_next_terminal = function Terminal _ :: _ -> true | _ -> false
@@ -387,8 +343,9 @@ let is_next_break = function Break _ :: _ -> true | _ -> false
let add_break n l = (None,UnpCut (PpBrk(n,0))) :: l
-let add_break_if_none n = function
- | (((_,UnpCut (PpBrk _)) :: _) | []) as l -> l
+let add_break_if_none n b = function
+ | (_,UnpCut (PpBrk _)) :: _ as l -> l
+ | [] when not b -> []
| l -> (None,UnpCut (PpBrk(n,0))) :: l
let check_open_binder isopen sl m =
@@ -403,45 +360,59 @@ let check_open_binder isopen sl m =
prlist_with_sep spc pr_token sl
++ strbrk "\" is allowed to occur.")
+let unparsing_metavar i from typs =
+ let x = List.nth typs (i-1) in
+ let prec = snd (precedence_of_entry_type from x) in
+ match x with
+ | ETConstr _ | ETConstrAsBinder _ | ETReference | ETBigint ->
+ UnpMetaVar (i,prec)
+ | ETPattern _ ->
+ UnpBinderMetaVar (i,prec)
+ | ETName ->
+ UnpBinderMetaVar (i,Prec 0)
+ | ETBinder isopen ->
+ assert false
+ | ETOther _ -> failwith "TODO"
+
(* Heuristics for building default printing rules *)
let index_id id l = List.index Id.equal id l
let make_hunks etyps symbols from =
let vars,typs = List.split etyps in
- let rec make = function
+ let rec make b = function
| NonTerminal m :: prods ->
let i = index_id m vars in
- let _,prec = precedence_of_entry_type from (List.nth typs (i-1)) in
- let u = UnpMetaVar (i,prec) in
- if is_next_non_terminal prods then
- (None,u) :: add_break_if_none 1 (make prods)
+ let u = unparsing_metavar i from typs in
+ if is_next_non_terminal b prods then
+ (None, u) :: add_break_if_none 1 b (make b prods)
else
- (None,u) :: make_with_space prods
- | Terminal s :: prods when List.exists is_non_terminal prods ->
+ (None, u) :: make_with_space b prods
+ | Terminal s :: prods
+ when (* true to simulate presence of non-terminal *) b || List.exists is_non_terminal prods ->
if (is_comma s || is_operator s) then
(* Always a breakable space after comma or separator *)
- (None,UnpTerminal s) :: add_break_if_none 1 (make prods)
+ (None, UnpTerminal s) :: add_break_if_none 1 b (make b prods)
else if is_right_bracket s && is_next_terminal prods then
(* Always no space after right bracked, but possibly a break *)
- (None,UnpTerminal s) :: add_break_if_none 0 (make prods)
- else if is_left_bracket s && is_next_non_terminal prods then
- (None,UnpTerminal s) :: make prods
+ (None, UnpTerminal s) :: add_break_if_none 0 b (make b prods)
+ else if is_left_bracket s && is_next_non_terminal b prods then
+ (None, UnpTerminal s) :: make b prods
else if not (is_next_break prods) then
(* Add rigid space, no break, unless user asked for something *)
- (None,UnpTerminal (s^" ")) :: make prods
+ (None, UnpTerminal (s^" ")) :: make b prods
else
(* Rely on user spaces *)
- (None,UnpTerminal s) :: make prods
+ (None, UnpTerminal s) :: make b prods
| Terminal s :: prods ->
(* Separate but do not cut a trailing sequence of terminal *)
(match prods with
- | Terminal _ :: _ -> (None,UnpTerminal (s^" ")) :: make prods
- | _ -> (None,UnpTerminal s) :: make prods)
+ | Terminal _ :: _ -> (None,UnpTerminal (s^" ")) :: make b prods
+ | _ -> (None,UnpTerminal s) :: make b prods)
| Break n :: prods ->
- add_break n (make prods)
+ add_break n (make b prods)
| SProdList (m,sl) :: prods ->
let i = index_id m vars in
@@ -451,47 +422,52 @@ let make_hunks etyps symbols from =
(* If no separator: add a break *)
if List.is_empty sl then add_break 1 []
(* We add NonTerminal for simulation but remove it afterwards *)
- else snd (List.sep_last (make (sl@[NonTerminal m]))) in
+ else make true sl in
let hunk = match typ with
| ETConstr _ -> UnpListMetaVar (i,prec,List.map snd sl')
| ETBinder isopen ->
check_open_binder isopen sl m;
UnpBinderListMetaVar (i,isopen,List.map snd sl')
| _ -> assert false in
- (None,hunk) :: make_with_space prods
+ (None, hunk) :: make_with_space b prods
| [] -> []
- and make_with_space prods =
+ and make_with_space b prods =
match prods with
| Terminal s' :: prods'->
if is_operator s' then
(* A rigid space before operator and a breakable after *)
- (None,UnpTerminal (" "^s')) :: add_break_if_none 1 (make prods')
+ (None,UnpTerminal (" "^s')) :: add_break_if_none 1 b (make b prods')
else if is_comma s' then
(* No space whatsoever before comma *)
- make prods
+ make b prods
else if is_right_bracket s' then
- make prods
+ make b prods
else
(* A breakable space between any other two terminals *)
- add_break_if_none 1 (make prods)
+ add_break_if_none 1 b (make b prods)
| (NonTerminal _ | SProdList _) :: _ ->
(* A breakable space before a non-terminal *)
- add_break_if_none 1 (make prods)
+ add_break_if_none 1 b (make b prods)
| Break _ :: _ ->
(* Rely on user wish *)
- make prods
+ make b prods
| [] -> []
- in make symbols
+ in make false symbols
(* Build default printing rules from explicit format *)
let error_format ?loc () = user_err ?loc Pp.(str "The format does not match the notation.")
+let warn_format_break =
+ CWarnings.create ~name:"notation-both-format-and-spaces" ~category:"parsing"
+ (fun () ->
+ strbrk "Discarding format implicitly indicated by multiple spaces in notation because an explicit format modifier is given.")
+
let rec split_format_at_ldots hd = function
- | (loc,UnpTerminal s) :: fmt when String.equal s (Id.to_string ldots_var) -> loc, List.rev hd, fmt
+ | (loc,UnpTerminal s) :: fmt when String.equal s (Id.to_string Notation_ops.ldots_var) -> loc, List.rev hd, fmt
| u :: fmt ->
check_no_ldots_in_box u;
split_format_at_ldots (u::hd) fmt
@@ -538,8 +514,7 @@ let hunks_of_format (from,(vars,typs)) symfmt =
let symbs, l = aux (symbs,fmt) in symbs, UnpTerminal s :: l
| NonTerminal s :: symbs, (_,UnpTerminal s') :: fmt when Id.equal s (Id.of_string s') ->
let i = index_id s vars in
- let _,prec = precedence_of_entry_type from (List.nth typs (i-1)) in
- let symbs, l = aux (symbs,fmt) in symbs, UnpMetaVar (i,prec) :: l
+ let symbs, l = aux (symbs,fmt) in symbs, unparsing_metavar i from typs :: l
| symbs, (_,UnpBox (a,b)) :: fmt ->
let symbs', b' = aux (symbs,b) in
let symbs', l = aux (symbs',fmt) in
@@ -562,6 +537,7 @@ let hunks_of_format (from,(vars,typs)) symfmt =
| _ -> assert false in
symbs, hunk :: l
| symbs, [] -> symbs, []
+ | Break _ :: symbs, fmt -> warn_format_break (); aux (symbs,fmt)
| _, fmt -> error_format ?loc:(fst (List.hd fmt)) ()
in
match aux symfmt with
@@ -574,8 +550,8 @@ let hunks_of_format (from,(vars,typs)) symfmt =
let assoc_of_type n (_,typ) = precedence_of_entry_type n typ
let is_not_small_constr = function
- ETConstr _ -> true
- | ETOther("constr","binder_constr") -> true
+ ETProdConstr _ -> true
+ | ETProdOther("constr","binder_constr") -> true
| _ -> false
let rec define_keywords_aux = function
@@ -607,14 +583,14 @@ let distribute a ll = List.map (fun l -> a @ l) ll
let expand_list_rule typ tkl x n p ll =
let camlp5_message_name = Some (add_suffix x ("_"^string_of_int n)) in
- let main = GramConstrNonTerminal (ETConstr typ, camlp5_message_name) in
+ let main = GramConstrNonTerminal (ETProdConstr typ, camlp5_message_name) in
let tks = List.map (fun x -> GramConstrTerminal x) tkl in
let rec aux i hds ll =
if i < p then aux (i+1) (main :: tks @ hds) ll
else if Int.equal i (p+n) then
let hds =
GramConstrListMark (p+n,true,p) :: hds
- @ [GramConstrNonTerminal (ETConstrList (typ,tkl), Some x)] in
+ @ [GramConstrNonTerminal (ETProdConstrList (typ,tkl), Some x)] in
distribute hds ll
else
distribute (GramConstrListMark (i+1,false,p) :: hds @ [main]) ll @
@@ -623,7 +599,7 @@ let expand_list_rule typ tkl x n p ll =
let is_constr_typ typ x etyps =
match List.assoc x etyps with
- | ETConstr typ' -> typ = typ'
+ | ETConstr typ' | ETConstrAsBinder (_,typ') -> typ = typ'
| _ -> false
let include_possible_similar_trailing_pattern typ etyps sl l =
@@ -636,12 +612,21 @@ let include_possible_similar_trailing_pattern typ etyps sl l =
with Exit -> n,l in
try_aux 0 l
+let prod_entry_type = function
+ | ETName -> ETProdName
+ | ETReference -> ETProdReference
+ | ETBigint -> ETProdBigint
+ | ETBinder _ -> assert false (* See check_binder_type *)
+ | ETConstr p | ETConstrAsBinder (_,p) -> ETProdConstr p
+ | ETPattern (_,n) -> ETProdPattern (match n with None -> 0 | Some n -> n)
+ | ETOther (s,t) -> ETProdOther (s,t)
+
let make_production etyps symbols =
let rec aux = function
| [] -> [[]]
| NonTerminal m :: l ->
let typ = List.assoc m etyps in
- distribute [GramConstrNonTerminal (typ, Some m)] (aux l)
+ distribute [GramConstrNonTerminal (prod_entry_type typ, Some m)] (aux l)
| Terminal s :: l ->
distribute [GramConstrTerminal (CLexer.terminal s)] (aux l)
| Break _ :: l ->
@@ -656,8 +641,10 @@ let make_production etyps symbols =
let p,l' = include_possible_similar_trailing_pattern typ etyps sl l in
expand_list_rule typ tkl x 1 p (aux l')
| ETBinder o ->
- distribute
- [GramConstrNonTerminal (ETBinderList (o,tkl), Some x)] (aux l)
+ check_open_binder o sl x;
+ let typ = if o then (assert (tkl = []); ETBinderOpen) else ETBinderClosed tkl in
+ distribute
+ [GramConstrNonTerminal (ETProdBinderList typ, Some x)] (aux l)
| _ ->
user_err Pp.(str "Components of recursive patterns in notation must be terms or binders.") in
let prods = aux symbols in
@@ -675,6 +662,7 @@ let rec find_symbols c_current c_next c_last = function
let border = function
| (_,ETConstr(_,BorderProd (_,a))) :: _ -> a
+ | (_,(ETConstrAsBinder(_,(_,BorderProd (_,a))))) :: _ -> a
| _ -> None
let recompute_assoc typs =
@@ -688,17 +676,16 @@ let recompute_assoc typs =
(* Registration of syntax extensions (parsing/printing, no interpretation)*)
let pr_arg_level from (lev,typ) =
- let pplev = match lev with
+ let pplev = function
| (n,L) when Int.equal n from -> str "at next level"
| (n,E) -> str "at level " ++ int n
| (n,L) -> str "at level below " ++ int n
| (n,Prec m) when Int.equal m n -> str "at level " ++ int n
| (n,_) -> str "Unknown level" in
- let pptyp = match typ with
- | NtnInternTypeConstr -> mt ()
- | NtnInternTypeBinder -> str " " ++ surround (str "binder")
- | NtnInternTypeIdent -> str " " ++ surround (str "ident") in
- pplev ++ pptyp
+ Ppvernac.pr_set_entry_type (fun _ -> (*TO CHECK*) mt()) typ ++
+ (match typ with
+ | ETConstr _ | ETConstrAsBinder _ | ETPattern _ -> spc () ++ pplev lev
+ | _ -> mt ())
let pr_level ntn (from,args,typs) =
str "at level " ++ int from ++ spc () ++ str "with arguments" ++ spc() ++
@@ -830,15 +817,23 @@ let interp_modifiers modl = let open NotationMods in
interp { acc with etyps = (id,typ) :: acc.etyps; } l
| SetItemLevel ([],n) :: l ->
interp acc l
+ | SetItemLevelAsBinder ([],_,_) :: l ->
+ interp acc l
| SetItemLevel (s::idl,n) :: l ->
let id = Id.of_string s in
if Id.List.mem_assoc id acc.etyps then
user_err ~hdr:"Metasyntax.interp_modifiers"
(str s ++ str " is already assigned to an entry or constr level.");
- let typ = ETConstr (n,()) in
+ let typ = ETConstr (Some n) in
interp { acc with etyps = (id,typ)::acc.etyps; } (SetItemLevel (idl,n)::l)
+ | SetItemLevelAsBinder (s::idl,bk,n) :: l ->
+ let id = Id.of_string s in
+ if Id.List.mem_assoc id acc.etyps then
+ user_err ~hdr:"Metasyntax.interp_modifiers"
+ (str s ++ str " is already assigned to an entry or constr level.");
+ let typ = ETConstrAsBinder (bk,n) in
+ interp { acc with etyps = (id,typ)::acc.etyps; } (SetItemLevelAsBinder (idl,bk,n)::l)
| SetLevel n :: l ->
-
interp { acc with level = Some n; } l
| SetAssoc a :: l ->
if not (Option.is_empty acc.assoc) then user_err Pp.(str "An associativity is given more than once.");
@@ -902,12 +897,17 @@ let get_compat_version mods =
let set_entry_type etyps (x,typ) =
let typ = try
match List.assoc x etyps, typ with
- | ETConstr (n,()), (_,BorderProd (left,_)) ->
+ | ETConstr (Some n), (_,BorderProd (left,_)) ->
ETConstr (n,BorderProd (left,None))
- | ETConstr (n,()), (_,InternalProd) -> ETConstr (n,InternalProd)
- | (ETPattern | ETName | ETBigint | ETOther _ |
- ETReference | ETBinder _ as t), _ -> t
- | (ETBinderList _ |ETConstrList _), _ -> assert false
+ | ETConstr (Some n), (_,InternalProd) -> ETConstr (n,InternalProd)
+ | ETConstrAsBinder (bk, Some n), (_,BorderProd (left,_)) ->
+ ETConstrAsBinder (bk, (n,BorderProd (left,None)))
+ | ETConstrAsBinder (bk, Some n), (_,InternalProd) ->
+ ETConstrAsBinder (bk, (n,InternalProd))
+ | ETPattern (b,n), _ -> ETPattern (b,n)
+ | (ETName | ETBigint | ETReference | ETBinder _ | ETOther _ as x), _ -> x
+ | ETConstr None, _ -> ETConstr typ
+ | ETConstrAsBinder (bk,None), _ -> ETConstrAsBinder (bk,typ)
with Not_found -> ETConstr typ
in (x,typ)
@@ -927,12 +927,9 @@ let join_auxiliary_recursive_types recvars etyps =
recvars etyps
let internalization_type_of_entry_type = function
- | ETConstr _ -> NtnInternTypeConstr
- | ETBigint | ETReference -> NtnInternTypeConstr
- | ETBinder _ -> NtnInternTypeBinder
- | ETName -> NtnInternTypeIdent
- | ETPattern | ETOther _ -> user_err Pp.(str "Not supported.")
- | ETBinderList _ | ETConstrList _ -> assert false
+ | ETBinder _ -> NtnInternTypeOnlyBinder
+ | ETConstr _ | ETConstrAsBinder _ | ETBigint | ETReference
+ | ETName | ETPattern _ | ETOther _ -> NtnInternTypeAny
let set_internalization_type typs =
List.map (fun (_, e) -> internalization_type_of_entry_type e) typs
@@ -943,28 +940,36 @@ let make_internalization_vars recvars mainvars typs =
maintyps @ extratyps
let make_interpretation_type isrec isonlybinding = function
- | NtnInternTypeConstr when isrec -> NtnTypeConstrList
- | NtnInternTypeConstr | NtnInternTypeIdent ->
- if isonlybinding then NtnTypeOnlyBinder else NtnTypeConstr
- | NtnInternTypeBinder when isrec -> NtnTypeBinderList
- | NtnInternTypeBinder -> user_err Pp.(str "Type binder is only for use in recursive notations for binders.")
-
-let make_interpretation_vars recvars allvars =
+ | ETConstr _ ->
+ if isrec then NtnTypeConstrList else
+ if isonlybinding then
+ (* Parsed as constr, but interpreted as a binder: default is to parse it as an ident only *)
+ NtnTypeBinder (NtnBinderParsedAsConstr AsIdent)
+ else NtnTypeConstr
+ | ETConstrAsBinder (bk,_) -> NtnTypeBinder (NtnBinderParsedAsConstr bk)
+ | ETName -> NtnTypeBinder NtnParsedAsIdent
+ | ETPattern (ppstrict,_) -> NtnTypeBinder (NtnParsedAsPattern ppstrict) (* Parsed as ident/pattern, primarily interpreted as binder; maybe strict at printing *)
+ | ETBigint | ETReference | ETOther _ -> NtnTypeConstr
+ | ETBinder _ ->
+ if isrec then NtnTypeBinderList
+ else anomaly Pp.(str "Type binder is only for use in recursive notations for binders.")
+
+let make_interpretation_vars recvars allvars typs =
let eq_subscope (sc1, l1) (sc2, l2) =
Option.equal String.equal sc1 sc2 &&
List.equal String.equal l1 l2
in
let check (x, y) =
- let (_,scope1, _) = Id.Map.find x allvars in
- let (_,scope2, _) = Id.Map.find y allvars in
+ let (_,scope1) = Id.Map.find x allvars in
+ let (_,scope2) = Id.Map.find y allvars in
if not (eq_subscope scope1 scope2) then error_not_same_scope x y
in
let () = List.iter check recvars in
let useless_recvars = List.map snd recvars in
let mainvars =
Id.Map.filter (fun x _ -> not (Id.List.mem x useless_recvars)) allvars in
- Id.Map.mapi (fun x (isonlybinding, sc, typ) ->
- (sc, make_interpretation_type (Id.List.mem_assoc x recvars) isonlybinding typ)) mainvars
+ Id.Map.mapi (fun x (isonlybinding, sc) ->
+ (sc, make_interpretation_type (Id.List.mem_assoc x recvars) isonlybinding (Id.List.assoc x typs))) mainvars
let check_rule_productivity l =
if List.for_all (function NonTerminal _ | Break _ -> true | _ -> false) l then
@@ -979,18 +984,27 @@ let warn_notation_bound_to_variable =
let warn_non_reversible_notation =
CWarnings.create ~name:"non-reversible-notation" ~category:"parsing"
- (fun () ->
- strbrk "This notation will not be used for printing as it is not reversible.")
-
-let is_not_printable onlyparse nonreversible = function
+ (function
+ | APrioriReversible -> assert false
+ | HasLtac ->
+ strbrk "This notation contains Ltac expressions: it will not be used for printing."
+ | NonInjective ids ->
+ let n = List.length ids in
+ strbrk (String.plural n "Variable") ++ spc () ++ pr_enum Id.print ids ++ spc () ++
+ strbrk (if n > 1 then "do" else "does") ++
+ str " not occur in the right-hand side." ++ spc() ++
+ strbrk "The notation will not be used for printing as it is not reversible.")
+
+let is_not_printable onlyparse reversibility = function
| NVar _ ->
if not onlyparse then warn_notation_bound_to_variable ();
true
| _ ->
- if not onlyparse && nonreversible then
- (warn_non_reversible_notation (); true)
+ if not onlyparse && reversibility <> APrioriReversible then
+ (warn_non_reversible_notation reversibility; true)
else onlyparse
+
let find_precedence lev etyps symbols onlyprint =
let first_symbol =
let rec aux = function
@@ -1008,29 +1022,30 @@ let find_precedence lev etyps symbols onlyprint =
match first_symbol with
| None -> [],0
| Some (NonTerminal x) ->
+ let test () =
+ if onlyprint then
+ if Option.is_empty lev then
+ user_err Pp.(str "Explicit level needed in only-printing mode when the level of the leftmost non-terminal is given.")
+ else [],Option.get lev
+ else
+ user_err Pp.(str "The level of the leftmost non-terminal cannot be changed.") in
(try match List.assoc x etyps with
- | ETConstr _ ->
- if onlyprint then
- if Option.is_empty lev then
- user_err Pp.(str "Explicit level needed in only-printing mode when the level of the leftmost non-terminal is given.")
- else [],Option.get lev
- else
- user_err Pp.(str "The level of the leftmost non-terminal cannot be changed.")
- | ETName | ETBigint | ETReference ->
+ | ETConstr _ -> test ()
+ | ETConstrAsBinder (_,Some _) -> test ()
+ | (ETName | ETBigint | ETReference) ->
begin match lev with
| None ->
([Feedback.msg_info ?loc:None ,strbrk "Setting notation at level 0."],0)
| Some 0 ->
([],0)
| _ ->
- user_err Pp.(str "A notation starting with an atomic expression must be at level 0.")
+ user_err Pp.(str "A notation starting with an atomic expression must be at level 0.")
end
- | ETPattern | ETBinder _ | ETOther _ -> (* Give a default ? *)
- if Option.is_empty lev then
- user_err Pp.(str "Need an explicit level.")
- else [],Option.get lev
- | ETConstrList _ | ETBinderList _ ->
- assert false (* internally used in grammar only *)
+ | (ETPattern _ | ETBinder _ | ETOther _ | ETConstrAsBinder _) ->
+ (* Give a default ? *)
+ if Option.is_empty lev then
+ user_err Pp.(str "Need an explicit level.")
+ else [],Option.get lev
with Not_found ->
if Option.is_empty lev then
user_err Pp.(str "A left-recursive notation must have an explicit level.")
@@ -1074,7 +1089,7 @@ let remove_curly_brackets l =
module SynData = struct
- type subentry_types = (Id.t * (production_level, production_position) constr_entry_key_gen) list
+ type subentry_types = (Id.t * (production_level * production_position) constr_entry_key_gen) list
(* XXX: Document *)
type syn_data = {
@@ -1128,8 +1143,7 @@ let compute_syntax_data df modifiers =
let onlyparse = mods.only_parsing in
if onlyprint && onlyparse then user_err (str "A notation cannot be both 'only printing' and 'only parsing'.");
let assoc = Option.append mods.assoc (Some NonA) in
- let toks = split_notation_string df in
- let (recvars,mainvars,symbols) = analyze_notation_tokens ~onlyprint toks in
+ let (recvars,mainvars,symbols) = analyze_notation_tokens ~onlyprint df in
let _ = check_useless_entry_types recvars mainvars mods.etyps in
let _ = check_binder_type recvars mods.etyps in
@@ -1152,7 +1166,7 @@ let compute_syntax_data df modifiers =
let i_typs = set_internalization_type sy_typs in
let pa_sy_data = (sy_typs_for_grammar,symbols_for_grammar) in
let pp_sy_data = (sy_typs,symbols) in
- let sy_fulldata = (ntn_for_grammar,(n,prec_for_grammar,i_typs),need_squash) in
+ let sy_fulldata = (ntn_for_grammar,(n,prec_for_grammar,List.map snd sy_typs_for_grammar),need_squash) in
let df' = ((Lib.library_dp(),Lib.current_dirpath true),df) in
let i_data = ntn_for_interp, df' in
@@ -1171,7 +1185,7 @@ let compute_syntax_data df modifiers =
mainvars;
intern_typs = i_typs;
- level = (n,prec,i_typs);
+ level = (n,prec,List.map snd sy_typs);
pa_syntax_data = pa_sy_data;
pp_syntax_data = pp_sy_data;
not_data = sy_fulldata;
@@ -1329,10 +1343,10 @@ let add_notation_in_scope local df env c mods scope =
ninterp_var_type = to_map i_vars;
ninterp_rec_vars = to_map sd.recvars;
} in
- let (acvars, ac, reversible) = interp_notation_constr env nenv c in
- let interp = make_interpretation_vars sd.recvars acvars in
+ let (acvars, ac, reversibility) = interp_notation_constr env nenv c in
+ let interp = make_interpretation_vars sd.recvars acvars (fst sd.pa_syntax_data) in
let map (x, _) = try Some (x, Id.Map.find x interp) with Not_found -> None in
- let onlyparse = is_not_printable sd.only_parsing (not reversible) ac in
+ let onlyparse = is_not_printable sd.only_parsing reversibility ac in
let notation = {
notobj_local = local;
notobj_scope = scope;
@@ -1350,8 +1364,7 @@ let add_notation_in_scope local df env c mods scope =
sd.info
let add_notation_interpretation_core local df env ?(impls=empty_internalization_env) c scope onlyparse onlyprint compat =
- let dfs = split_notation_string df in
- let (recvars,mainvars,symbs) = analyze_notation_tokens ~onlyprint dfs in
+ let (recvars,mainvars,symbs) = analyze_notation_tokens ~onlyprint df in
(* Recover types of variables and pa/pp rules; redeclare them if needed *)
let i_typs, onlyprint = if not (is_numeral symbs) then begin
let sy = recover_notation_syntax (make_notation_key symbs) in
@@ -1363,15 +1376,15 @@ let add_notation_interpretation_core local df env ?(impls=empty_internalization_
(* Declare interpretation *)
let path = (Lib.library_dp(), Lib.current_dirpath true) in
let df' = (make_notation_key symbs, (path,df)) in
- let i_vars = make_internalization_vars recvars mainvars i_typs in
+ let i_vars = make_internalization_vars recvars mainvars (List.map internalization_type_of_entry_type i_typs) in
let nenv = {
ninterp_var_type = to_map i_vars;
ninterp_rec_vars = to_map recvars;
} in
- let (acvars, ac, reversible) = interp_notation_constr env ~impls nenv c in
- let interp = make_interpretation_vars recvars acvars in
+ let (acvars, ac, reversibility) = interp_notation_constr env ~impls nenv c in
+ let interp = make_interpretation_vars recvars acvars (List.combine mainvars i_typs) in
let map (x, _) = try Some (x, Id.Map.find x interp) with Not_found -> None in
- let onlyparse = is_not_printable onlyparse (not reversible) ac in
+ let onlyparse = is_not_printable onlyparse reversibility ac in
let notation = {
notobj_local = local;
notobj_scope = scope;
@@ -1427,8 +1440,7 @@ let add_notation local env c ((loc,df),modifiers) sc =
let add_notation_extra_printing_rule df k v =
let notk =
- let dfs = split_notation_string df in
- let _,_, symbs = analyze_notation_tokens ~onlyprint:true dfs in
+ let _,_, symbs = analyze_notation_tokens ~onlyprint:true df in
make_notation_key symbs in
Notation.add_notation_extra_printing_rule notk k v
@@ -1499,23 +1511,21 @@ let try_interp_name_alias = function
| _ -> raise Not_found
let add_syntactic_definition env ident (vars,c) local onlyparse =
- let nonprintable = ref false in
- let vars,pat =
- try [], NRef (try_interp_name_alias (vars,c))
+ let vars,reversibility,pat =
+ try [], APrioriReversible, NRef (try_interp_name_alias (vars,c))
with Not_found ->
- let fold accu id = Id.Map.add id NtnInternTypeConstr accu in
+ let fold accu id = Id.Map.add id NtnInternTypeAny accu in
let i_vars = List.fold_left fold Id.Map.empty vars in
let nenv = {
ninterp_var_type = i_vars;
ninterp_rec_vars = Id.Map.empty;
} in
- let nvars, pat, reversible = interp_notation_constr env nenv c in
- let () = nonprintable := not reversible in
- let map id = let (_,sc,_) = Id.Map.find id nvars in (id, sc) in
- List.map map vars, pat
+ let nvars, pat, reversibility = interp_notation_constr env nenv c in
+ let map id = let (_,sc) = Id.Map.find id nvars in (id, sc) in
+ List.map map vars, reversibility, pat
in
let onlyparse = match onlyparse with
- | None when (is_not_printable false !nonprintable pat) -> Some Flags.Current
+ | None when (is_not_printable false reversibility pat) -> Some Flags.Current
| p -> p
in
Syntax_def.declare_syntactic_definition local ident onlyparse (vars,pat)
diff --git a/vernac/metasyntax.mli b/vernac/metasyntax.mli
index e064b570e..9137f7a7e 100644
--- a/vernac/metasyntax.mli
+++ b/vernac/metasyntax.mli
@@ -55,10 +55,6 @@ val add_syntactic_definition : env -> Id.t -> Id.t list * constr_expr ->
val pr_grammar : string -> Pp.t
-type any_entry = AnyEntry : 'a Pcoq.Gram.entry -> any_entry
-
-val register_grammar : string -> any_entry list -> unit
-
val check_infix_modifiers : syntax_modifier list -> unit
val with_syntax_protection : ('a -> 'b) -> 'a -> 'b
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 35ef4bfa4..fa457c895 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -443,11 +443,13 @@ let start_proof_and_print k l hook =
let hook env sigma ev =
let tac = !Obligations.default_tactic in
let evi = Evd.find sigma ev in
+ let evi = Evarutil.nf_evar_info sigma evi in
let env = Evd.evar_filtered_env evi in
try
- let concl = Evarutil.nf_evars_universes sigma evi.Evd.evar_concl in
- let concl = EConstr.of_constr concl in
- if Evarutil.has_undefined_evars sigma concl then raise Exit;
+ let concl = EConstr.of_constr evi.Evd.evar_concl in
+ if not (Evarutil.is_ground_env sigma env &&
+ Evarutil.is_ground_term sigma concl)
+ then raise Exit;
let c, _, ctx =
Pfedit.build_by_tactic env (Evd.evar_universe_context sigma)
concl (Tacticals.New.tclCOMPLETE tac)