aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorPierre-Marie Pédrot2020-11-15 00:20:36 +0100
committerPierre-Marie Pédrot2020-11-15 00:20:36 +0100
commita118b906b3da7cb2e03a72f7a8079a7fc99c6f84 (patch)
tree9db2c41ab6ef97e9d80375b7dca1ff1ed1c5f111
parenta237a3d2c9de704873e1e20fa38282cf3562cddf (diff)
parent8ed2d808b2c9caf022b5e22bb43f2ca6876ebd1b (diff)
Merge PR #13339: In -noinit mode, add support for Proof using, using is not a keyword.
Ack-by: SkySkimmer Reviewed-by: herbelin Ack-by: ppedrot
-rw-r--r--doc/changelog/07-commands-and-options/13339-proof-using-noinit.rst5
-rw-r--r--plugins/ltac/g_ltac.mlg8
-rw-r--r--test-suite/success/proof_using_noinit.v9
-rw-r--r--vernac/g_proofs.mlg2
-rw-r--r--vernac/g_vernac.mlg3
5 files changed, 22 insertions, 5 deletions
diff --git a/doc/changelog/07-commands-and-options/13339-proof-using-noinit.rst b/doc/changelog/07-commands-and-options/13339-proof-using-noinit.rst
new file mode 100644
index 0000000000..9ae759be56
--- /dev/null
+++ b/doc/changelog/07-commands-and-options/13339-proof-using-noinit.rst
@@ -0,0 +1,5 @@
+- **Added:**
+ The :cmd:`Proof using` command can now be used without loading the
+ Ltac plugin (`-noinit` mode)
+ (`#13339 <https://github.com/coq/coq/pull/13339>`_,
+ by Théo Zimmermann).
diff --git a/plugins/ltac/g_ltac.mlg b/plugins/ltac/g_ltac.mlg
index c54f8ffa78..c2e95c45f9 100644
--- a/plugins/ltac/g_ltac.mlg
+++ b/plugins/ltac/g_ltac.mlg
@@ -329,11 +329,11 @@ GRAMMAR EXTEND Gram
;
command:
[ [ IDENT "Proof"; "with"; ta = Pltac.tactic;
- l = OPT [ "using"; l = G_vernac.section_subset_expr -> { l } ] ->
+ l = OPT [ IDENT "using"; l = G_vernac.section_subset_expr -> { l } ] ->
{ Vernacexpr.VernacProof (Some (in_tac ta), l) }
- | IDENT "Proof"; "using"; l = G_vernac.section_subset_expr;
- ta = OPT [ "with"; ta = Pltac.tactic -> { in_tac ta } ] ->
- { Vernacexpr.VernacProof (ta,Some l) } ] ]
+ | IDENT "Proof"; IDENT "using"; l = G_vernac.section_subset_expr;
+ "with"; ta = Pltac.tactic ->
+ { Vernacexpr.VernacProof (Some (in_tac ta),Some l) } ] ]
;
hint:
[ [ IDENT "Extern"; n = natural; c = OPT Constr.constr_pattern ; "=>";
diff --git a/test-suite/success/proof_using_noinit.v b/test-suite/success/proof_using_noinit.v
new file mode 100644
index 0000000000..f99b49619c
--- /dev/null
+++ b/test-suite/success/proof_using_noinit.v
@@ -0,0 +1,9 @@
+(* -*- coq-prog-args: ("-noinit"); -*- *)
+
+Section A.
+Variable A : Prop.
+Hypothesis a : A.
+Lemma b : A.
+Proof using a.
+Admitted.
+End A.
diff --git a/vernac/g_proofs.mlg b/vernac/g_proofs.mlg
index ebec720ce2..5b80ed6794 100644
--- a/vernac/g_proofs.mlg
+++ b/vernac/g_proofs.mlg
@@ -56,6 +56,8 @@ GRAMMAR EXTEND Gram
[ [ IDENT "Goal"; c = lconstr ->
{ VernacDefinition (Decls.(NoDischarge, Definition), ((CAst.make ~loc Names.Anonymous), None), ProveBody ([], c)) }
| IDENT "Proof" -> { VernacProof (None,None) }
+ | IDENT "Proof"; IDENT "using"; l = G_vernac.section_subset_expr ->
+ { VernacProof (None,Some l) }
| IDENT "Proof" ; IDENT "Mode" ; mn = string -> { VernacProofMode mn }
| IDENT "Proof"; c = lconstr -> { VernacExactProof c }
| IDENT "Abort" -> { VernacAbort None }
diff --git a/vernac/g_vernac.mlg b/vernac/g_vernac.mlg
index f192d67624..1c80d71ea5 100644
--- a/vernac/g_vernac.mlg
+++ b/vernac/g_vernac.mlg
@@ -114,7 +114,8 @@ GRAMMAR EXTEND Gram
;
attribute:
[ [ k = ident ; v = attr_value -> { Names.Id.to_string k, v }
- | "using" ; v = attr_value -> { "using", v } ]
+ (* Required because "ident" is declared a keyword when loading Ltac. *)
+ | IDENT "using" ; v = attr_value -> { "using", v } ]
]
;
attr_value: