aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--doc/changelog/05-tactic-language/11241-master+bug-cofix-with-8.10.rst4
-rw-r--r--plugins/ltac/g_tactic.mlg2
-rw-r--r--test-suite/success/RecTutorial.v4
-rw-r--r--test-suite/success/cofixtac.v10
4 files changed, 15 insertions, 5 deletions
diff --git a/doc/changelog/05-tactic-language/11241-master+bug-cofix-with-8.10.rst b/doc/changelog/05-tactic-language/11241-master+bug-cofix-with-8.10.rst
new file mode 100644
index 0000000000..462ba4a7b1
--- /dev/null
+++ b/doc/changelog/05-tactic-language/11241-master+bug-cofix-with-8.10.rst
@@ -0,0 +1,4 @@
+- **Fixed:**
+ Syntax of tactic `cofix ... with ...` was broken from Coq 8.10.
+ (`#11241 <https://github.com/coq/coq/pull/11241>`_,
+ by Hugo Herbelin).
diff --git a/plugins/ltac/g_tactic.mlg b/plugins/ltac/g_tactic.mlg
index d82eadcfc7..f0d6258cd1 100644
--- a/plugins/ltac/g_tactic.mlg
+++ b/plugins/ltac/g_tactic.mlg
@@ -135,7 +135,7 @@ let mk_cofix_tac (loc,id,bl,ann,ty) =
~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))
+ (id,if bl = [] then ty else CAst.make ~loc @@ CProdN(bl,ty))
(* Functions overloaded by quotifier *)
let destruction_arg_of_constr (c,lbind as clbind) = match lbind with
diff --git a/test-suite/success/RecTutorial.v b/test-suite/success/RecTutorial.v
index 6370cab6b2..4fac798f76 100644
--- a/test-suite/success/RecTutorial.v
+++ b/test-suite/success/RecTutorial.v
@@ -1210,7 +1210,3 @@ Proof.
constructor.
trivial.
Qed.
-
-
-
-
diff --git a/test-suite/success/cofixtac.v b/test-suite/success/cofixtac.v
new file mode 100644
index 0000000000..ae75ee770f
--- /dev/null
+++ b/test-suite/success/cofixtac.v
@@ -0,0 +1,10 @@
+CoInductive stream :=
+| C : content -> stream
+with content :=
+| D : nat -> stream -> content.
+
+Lemma one : stream.
+cofix c with (d : content).
+- constructor. apply d.
+- constructor. exact 1. apply c.
+Defined.