aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--CHANGES2
-rw-r--r--doc/sphinx/language/gallina-extensions.rst1
-rw-r--r--doc/sphinx/proof-engine/tactics.rst25
-rw-r--r--plugins/ltac/extratactics.ml46
-rw-r--r--test-suite/bugs/closed/4882.v50
-rw-r--r--test-suite/success/ImplicitTactic.v16
-rw-r--r--theories/Logic/Diaconescu.v2
7 files changed, 24 insertions, 78 deletions
diff --git a/CHANGES b/CHANGES
index 5b5885487..e8718bbab 100644
--- a/CHANGES
+++ b/CHANGES
@@ -23,6 +23,8 @@ Tactics
- Option "Ltac Debug" now applies also to terms built using Ltac functions.
+- Deprecated the Implicit Tactic family of commands.
+
Tools
- Coq_makefile lets one override or extend the following variables from
diff --git a/doc/sphinx/language/gallina-extensions.rst b/doc/sphinx/language/gallina-extensions.rst
index 6ea1c162f..ff5d352c9 100644
--- a/doc/sphinx/language/gallina-extensions.rst
+++ b/doc/sphinx/language/gallina-extensions.rst
@@ -2240,6 +2240,7 @@ This option (off by default) activates the full display of how the
context of an existential variable is instantiated at each of the
occurrences of the existential variable.
+.. _tactics-in-terms:
Solving existential variables using tactics
~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
diff --git a/doc/sphinx/proof-engine/tactics.rst b/doc/sphinx/proof-engine/tactics.rst
index 29e0b34bc..90ca0da43 100644
--- a/doc/sphinx/proof-engine/tactics.rst
+++ b/doc/sphinx/proof-engine/tactics.rst
@@ -3703,19 +3703,24 @@ Setting implicit automation tactics
time the term argument of a tactic has one of its holes not fully
resolved.
- .. example::
+ .. deprecated:: 8.9
- .. coqtop:: all
+ This command is deprecated. Use :ref:`typeclasses <typeclasses>` or
+ :ref:`tactics-in-terms <tactics-in-terms>` instead.
- Parameter quo : nat -> forall n:nat, n<>0 -> nat.
- Notation "x // y" := (quo x y _) (at level 40).
- Declare Implicit Tactic assumption.
- Goal forall n m, m<>0 -> { q:nat & { r | q * m + r = n } }.
- intros.
- exists (n // m).
+ .. example::
+
+ .. coqtop:: all
+
+ Parameter quo : nat -> forall n:nat, n<>0 -> nat.
+ Notation "x // y" := (quo x y _) (at level 40).
+ Declare Implicit Tactic assumption.
+ Goal forall n m, m<>0 -> { q:nat & { r | q * m + r = n } }.
+ intros.
+ exists (n // m).
- The tactic ``exists (n // m)`` did not fail. The hole was solved
- by ``assumption`` so that it behaved as ``exists (quo n m H)``.
+ The tactic ``exists (n // m)`` did not fail. The hole was solved
+ by ``assumption`` so that it behaved as ``exists (quo n m H)``.
.. _decisionprocedures:
diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4
index 8813c7764..cb7183638 100644
--- a/plugins/ltac/extratactics.ml4
+++ b/plugins/ltac/extratactics.ml4
@@ -523,10 +523,16 @@ let inImplicitTactic : glob_tactic_expr option -> obj =
subst_function = subst_implicit_tactic;
classify_function = (fun o -> Dispose)}
+let warn_deprecated_implicit_tactic =
+ CWarnings.create ~name:"deprecated-implicit-tactic" ~category:"deprecated"
+ (fun () -> strbrk "Implicit tactics are deprecated")
+
let declare_implicit_tactic tac =
+ let () = warn_deprecated_implicit_tactic () in
Lib.add_anonymous_leaf (inImplicitTactic (Some (Tacintern.glob_tactic tac)))
let clear_implicit_tactic () =
+ let () = warn_deprecated_implicit_tactic () in
Lib.add_anonymous_leaf (inImplicitTactic None)
VERNAC COMMAND EXTEND ImplicitTactic CLASSIFIED AS SIDEFF
diff --git a/test-suite/bugs/closed/4882.v b/test-suite/bugs/closed/4882.v
deleted file mode 100644
index 8c26af708..000000000
--- a/test-suite/bugs/closed/4882.v
+++ /dev/null
@@ -1,50 +0,0 @@
-
-Definition Foo {T}{a : T} : T := a.
-
-Module A.
-
- Declare Implicit Tactic eauto.
-
- Goal forall A (x : A), A.
- intros.
- apply Foo. (* Check defined evars are normalized *)
- (* Qed. *)
- Abort.
-
-End A.
-
-Module B.
-
- Definition Foo {T}{a : T} : T := a.
-
- Declare Implicit Tactic eassumption.
-
- Goal forall A (x : A), A.
- intros.
- apply Foo.
- (* Qed. *)
- Abort.
-
-End B.
-
-Module C.
-
- Declare Implicit Tactic first [exact True|assumption].
-
- Goal forall (x : True), True.
- intros.
- apply (@Foo _ _).
- Qed.
-
-End C.
-
-Module D.
-
- Declare Implicit Tactic assumption.
-
- Goal forall A (x : A), A.
- intros.
- exact _.
- Qed.
-
-End D.
diff --git a/test-suite/success/ImplicitTactic.v b/test-suite/success/ImplicitTactic.v
deleted file mode 100644
index d8fa3043d..000000000
--- a/test-suite/success/ImplicitTactic.v
+++ /dev/null
@@ -1,16 +0,0 @@
-(* A Wiedijk-Cruz-Filipe style tactic for solving implicit arguments *)
-
-(* Declare a term expression with a hole *)
-Parameter quo : nat -> forall n:nat, n<>0 -> nat.
-Notation "x / y" := (quo x y _) : nat_scope.
-
-(* Declare the tactic for resolving implicit arguments still
- unresolved after type-checking; it must complete the subgoal to
- succeed *)
-Declare Implicit Tactic assumption.
-
-Goal forall n d, d<>0 -> { q:nat & { r:nat | d * q + r = n }}.
-intros.
-(* Here, assumption is used to solve the implicit argument of quo *)
-exists (n / d).
-
diff --git a/theories/Logic/Diaconescu.v b/theories/Logic/Diaconescu.v
index 3317766c9..66e82ddbf 100644
--- a/theories/Logic/Diaconescu.v
+++ b/theories/Logic/Diaconescu.v
@@ -234,8 +234,6 @@ Qed.
(** An alternative more concise proof can be done by directly using
the guarded relational choice *)
-Declare Implicit Tactic auto.
-
Lemma proof_irrel_rel_choice_imp_eq_dec' : a1=a2 \/ ~a1=a2.
Proof.
assert (decide: forall x:A, x=a1 \/ x=a2 ->