aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2016-10-06 17:34:12 +0200
committerGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2017-02-17 11:52:38 +0100
commitebc0870ca932acf0ceea5fe513e2ca40e74c2a02 (patch)
tree4d7d3257b5858bcd2728f678f5cb1f937b2d282e
parent653de32139ecee3114779a5ee2961c58793889e5 (diff)
Moving the Ltac plugin to a pack-based one.
This is cumbersome, because now code may fail at link time if it's not referring to the correct module name. Therefore, one has to add corresponding open statements a the top of every file depending on a Ltac module. This includes seemingly unrelated files that use EXTEND statements.
-rw-r--r--plugins/btauto/g_btauto.ml42
-rw-r--r--plugins/cc/g_congruence.ml41
-rw-r--r--plugins/decl_mode/decl_interp.ml1
-rw-r--r--plugins/decl_mode/decl_proof_instr.ml1
-rw-r--r--plugins/decl_mode/g_decl_mode.ml41
-rw-r--r--plugins/decl_mode/ppdecl_proof.ml1
-rw-r--r--plugins/extraction/g_extraction.ml41
-rw-r--r--plugins/firstorder/g_ground.ml41
-rw-r--r--plugins/firstorder/ground.ml1
-rw-r--r--plugins/fourier/g_fourier.ml41
-rw-r--r--plugins/funind/g_indfun.ml41
-rw-r--r--plugins/funind/invfun.ml1
-rw-r--r--plugins/ltac/ltac_plugin.mlpack (renamed from plugins/ltac/ltac_plugin.mllib)0
-rw-r--r--plugins/micromega/g_micromega.ml41
-rw-r--r--plugins/nsatz/g_nsatz.ml44
-rw-r--r--plugins/omega/g_omega.ml41
-rw-r--r--plugins/quote/g_quote.ml41
-rw-r--r--plugins/romega/g_romega.ml41
-rw-r--r--plugins/rtauto/g_rtauto.ml42
-rw-r--r--plugins/rtauto/refl_tauto.ml1
-rw-r--r--plugins/setoid_ring/g_newring.ml41
-rw-r--r--plugins/setoid_ring/newring.ml1
-rw-r--r--plugins/ssrmatching/ssrmatching.ml41
-rw-r--r--test-suite/bugs/closed/3612.v3
-rw-r--r--test-suite/bugs/closed/3649.v2
-rw-r--r--test-suite/bugs/closed/4121.v4
-rw-r--r--test-suite/bugs/closed/4527.v1
-rw-r--r--test-suite/bugs/closed/4533.v3
-rw-r--r--test-suite/bugs/closed/4544.v3
29 files changed, 38 insertions, 5 deletions
diff --git a/plugins/btauto/g_btauto.ml4 b/plugins/btauto/g_btauto.ml4
index f3e2c99f4..298027448 100644
--- a/plugins/btauto/g_btauto.ml4
+++ b/plugins/btauto/g_btauto.ml4
@@ -8,6 +8,8 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
+open Ltac_plugin
+
DECLARE PLUGIN "btauto_plugin"
TACTIC EXTEND btauto
diff --git a/plugins/cc/g_congruence.ml4 b/plugins/cc/g_congruence.ml4
index 6f6811334..7e76854b1 100644
--- a/plugins/cc/g_congruence.ml4
+++ b/plugins/cc/g_congruence.ml4
@@ -8,6 +8,7 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
+open Ltac_plugin
open Cctac
open Stdarg
diff --git a/plugins/decl_mode/decl_interp.ml b/plugins/decl_mode/decl_interp.ml
index f68c01b18..2b63ed6d6 100644
--- a/plugins/decl_mode/decl_interp.ml
+++ b/plugins/decl_mode/decl_interp.ml
@@ -6,6 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Ltac_plugin
open CErrors
open Util
open Names
diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml
index e19dc86c4..deb2ede1d 100644
--- a/plugins/decl_mode/decl_proof_instr.ml
+++ b/plugins/decl_mode/decl_proof_instr.ml
@@ -6,6 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Ltac_plugin
open CErrors
open Util
open Pp
diff --git a/plugins/decl_mode/g_decl_mode.ml4 b/plugins/decl_mode/g_decl_mode.ml4
index 18a35c6cf..a71d20f0d 100644
--- a/plugins/decl_mode/g_decl_mode.ml4
+++ b/plugins/decl_mode/g_decl_mode.ml4
@@ -10,6 +10,7 @@
DECLARE PLUGIN "decl_mode_plugin"
+open Ltac_plugin
open Compat
open Pp
open Decl_expr
diff --git a/plugins/decl_mode/ppdecl_proof.ml b/plugins/decl_mode/ppdecl_proof.ml
index 59a0bb5a2..f5de638ed 100644
--- a/plugins/decl_mode/ppdecl_proof.ml
+++ b/plugins/decl_mode/ppdecl_proof.ml
@@ -6,6 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Ltac_plugin
open CErrors
open Pp
open Decl_expr
diff --git a/plugins/extraction/g_extraction.ml4 b/plugins/extraction/g_extraction.ml4
index e1d6bb4a8..3ed959cf2 100644
--- a/plugins/extraction/g_extraction.ml4
+++ b/plugins/extraction/g_extraction.ml4
@@ -12,6 +12,7 @@ DECLARE PLUGIN "extraction_plugin"
(* ML names *)
+open Ltac_plugin
open Genarg
open Stdarg
open Pcoq.Prim
diff --git a/plugins/firstorder/g_ground.ml4 b/plugins/firstorder/g_ground.ml4
index 260e86ad6..e28d6aa62 100644
--- a/plugins/firstorder/g_ground.ml4
+++ b/plugins/firstorder/g_ground.ml4
@@ -8,6 +8,7 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
+open Ltac_plugin
open Formula
open Sequent
open Ground
diff --git a/plugins/firstorder/ground.ml b/plugins/firstorder/ground.ml
index 628af4e71..d6cd7e2a0 100644
--- a/plugins/firstorder/ground.ml
+++ b/plugins/firstorder/ground.ml
@@ -6,6 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Ltac_plugin
open Formula
open Sequent
open Rules
diff --git a/plugins/fourier/g_fourier.ml4 b/plugins/fourier/g_fourier.ml4
index 7c665ae7b..1960fa835 100644
--- a/plugins/fourier/g_fourier.ml4
+++ b/plugins/fourier/g_fourier.ml4
@@ -8,6 +8,7 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
+open Ltac_plugin
open FourierR
DECLARE PLUGIN "fourier_plugin"
diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4
index 6603a95a8..368b23be3 100644
--- a/plugins/funind/g_indfun.ml4
+++ b/plugins/funind/g_indfun.ml4
@@ -6,6 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
(*i camlp4deps: "grammar/grammar.cma" i*)
+open Ltac_plugin
open Compat
open Util
open Term
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index c8b4e4833..70333b063 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -6,6 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Ltac_plugin
open Tacexpr
open Declarations
open CErrors
diff --git a/plugins/ltac/ltac_plugin.mllib b/plugins/ltac/ltac_plugin.mlpack
index af1c7149d..af1c7149d 100644
--- a/plugins/ltac/ltac_plugin.mllib
+++ b/plugins/ltac/ltac_plugin.mlpack
diff --git a/plugins/micromega/g_micromega.ml4 b/plugins/micromega/g_micromega.ml4
index 79020ed03..ccb6daa11 100644
--- a/plugins/micromega/g_micromega.ml4
+++ b/plugins/micromega/g_micromega.ml4
@@ -16,6 +16,7 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
+open Ltac_plugin
open Stdarg
open Tacarg
diff --git a/plugins/nsatz/g_nsatz.ml4 b/plugins/nsatz/g_nsatz.ml4
index 5f906a8da..195dec362 100644
--- a/plugins/nsatz/g_nsatz.ml4
+++ b/plugins/nsatz/g_nsatz.ml4
@@ -1,5 +1,3 @@
-DECLARE PLUGIN "nsatz_plugin"
-
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
@@ -10,6 +8,8 @@ DECLARE PLUGIN "nsatz_plugin"
(*i camlp4deps: "grammar/grammar.cma" i*)
+open Ltac_plugin
+
DECLARE PLUGIN "nsatz_plugin"
TACTIC EXTEND nsatz_compute
diff --git a/plugins/omega/g_omega.ml4 b/plugins/omega/g_omega.ml4
index 27115abec..6b711a176 100644
--- a/plugins/omega/g_omega.ml4
+++ b/plugins/omega/g_omega.ml4
@@ -17,6 +17,7 @@
DECLARE PLUGIN "omega_plugin"
+open Ltac_plugin
open Names
open Coq_omega
open Stdarg
diff --git a/plugins/quote/g_quote.ml4 b/plugins/quote/g_quote.ml4
index e7e6ecef9..f2c021f59 100644
--- a/plugins/quote/g_quote.ml4
+++ b/plugins/quote/g_quote.ml4
@@ -8,6 +8,7 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
+open Ltac_plugin
open Names
open Misctypes
open Tacexpr
diff --git a/plugins/romega/g_romega.ml4 b/plugins/romega/g_romega.ml4
index 2f38688d1..9a54ad778 100644
--- a/plugins/romega/g_romega.ml4
+++ b/plugins/romega/g_romega.ml4
@@ -10,6 +10,7 @@
DECLARE PLUGIN "romega_plugin"
+open Ltac_plugin
open Names
open Refl_omega
open Stdarg
diff --git a/plugins/rtauto/g_rtauto.ml4 b/plugins/rtauto/g_rtauto.ml4
index d27b04834..7e58ef9a3 100644
--- a/plugins/rtauto/g_rtauto.ml4
+++ b/plugins/rtauto/g_rtauto.ml4
@@ -8,6 +8,8 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
+open Ltac_plugin
+
DECLARE PLUGIN "rtauto_plugin"
TACTIC EXTEND rtauto
diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml
index 367a13333..35d6768c1 100644
--- a/plugins/rtauto/refl_tauto.ml
+++ b/plugins/rtauto/refl_tauto.ml
@@ -8,6 +8,7 @@
module Search = Explore.Make(Proof_search)
+open Ltac_plugin
open CErrors
open Util
open Term
diff --git a/plugins/setoid_ring/g_newring.ml4 b/plugins/setoid_ring/g_newring.ml4
index 0987c44ae..707ff79a6 100644
--- a/plugins/setoid_ring/g_newring.ml4
+++ b/plugins/setoid_ring/g_newring.ml4
@@ -8,6 +8,7 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
+open Ltac_plugin
open Pp
open Util
open Libnames
diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml
index 657efe175..59f23a637 100644
--- a/plugins/setoid_ring/newring.ml
+++ b/plugins/setoid_ring/newring.ml
@@ -6,6 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Ltac_plugin
open Pp
open CErrors
open Util
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index fc988a2c5..f4f6efa4a 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -15,6 +15,7 @@ let frozen_lexer = CLexer.freeze () ;;
(*i camlp4use: "pa_extend.cmo" i*)
(*i camlp4deps: "grammar/grammar.cma" i*)
+open Ltac_plugin
open Names
open Pp
open Pcoq
diff --git a/test-suite/bugs/closed/3612.v b/test-suite/bugs/closed/3612.v
index a54768507..4b4f81dbc 100644
--- a/test-suite/bugs/closed/3612.v
+++ b/test-suite/bugs/closed/3612.v
@@ -38,8 +38,11 @@ Axiom path_path_sigma : forall {A : Type} (P : A -> Type) (u v : sigT P)
(s : transport (fun x => transport P x u.2 = v.2) r p..2 = q..2),
p = q.
+Declare ML Module "ltac_plugin".
Declare ML Module "coretactics".
+Set Default Proof Mode "Classic".
+
Goal forall (A : Type) (B : forall _ : A, Type) (x : @sigT A (fun x : A => B x))
(xx : @paths (@sigT A (fun x0 : A => B x0)) x x),
@paths (@paths (@sigT A (fun x0 : A => B x0)) x x) xx
diff --git a/test-suite/bugs/closed/3649.v b/test-suite/bugs/closed/3649.v
index fc4c171e2..8687eaab0 100644
--- a/test-suite/bugs/closed/3649.v
+++ b/test-suite/bugs/closed/3649.v
@@ -2,7 +2,9 @@
(* File reduced by coq-bug-finder from original input, then from 9518 lines to 404 lines, then from 410 lines to 208 lines, then from 162 lines to 77 lines *)
(* coqc version trunk (September 2014) compiled on Sep 18 2014 21:0:5 with OCaml 4.01.0
coqtop version cagnode16:/afs/csail.mit.edu/u/j/jgross/coq-trunk,trunk (07e4438bd758c2ced8caf09a6961ccd77d84e42b) *)
+Declare ML Module "ltac_plugin".
Declare ML Module "coretactics".
+Set Default Proof Mode "Classic".
Reserved Notation "x -> y" (at level 99, right associativity, y at level 200).
Reserved Notation "x = y" (at level 70, no associativity).
Delimit Scope type_scope with type.
diff --git a/test-suite/bugs/closed/4121.v b/test-suite/bugs/closed/4121.v
index d34a2b8b1..816bc845f 100644
--- a/test-suite/bugs/closed/4121.v
+++ b/test-suite/bugs/closed/4121.v
@@ -4,6 +4,8 @@ Unset Strict Universe Declaration.
(* coqc version 8.5beta1 (March 2015) compiled on Mar 11 2015 18:51:36 with OCaml 4.01.0
coqtop version cagnode15:/afs/csail.mit.edu/u/j/jgross/coq-8.5,v8.5 (8dbfee5c5f897af8186cb1bdfb04fd4f88eca677) *)
+Declare ML Module "ltac_plugin".
+
Set Universe Polymorphism.
Class Contr_internal (A : Type) := BuildContr { center : A }.
Arguments center A {_}.
@@ -13,4 +15,4 @@ Definition contr_paths_contr0 {A} `{Contr A} : Contr A := {| center := center A
Instance contr_paths_contr1 {A} `{Contr A} : Contr A := {| center := center A |}.
Check @contr_paths_contr0@{i}.
Check @contr_paths_contr1@{i}. (* Error: Universe instance should have length 2 *)
-(** It should have length 1, just like contr_paths_contr0 *) \ No newline at end of file
+(** It should have length 1, just like contr_paths_contr0 *)
diff --git a/test-suite/bugs/closed/4527.v b/test-suite/bugs/closed/4527.v
index 08628377f..c6fcc24b6 100644
--- a/test-suite/bugs/closed/4527.v
+++ b/test-suite/bugs/closed/4527.v
@@ -5,6 +5,7 @@ then from 269 lines to 255 lines *)
(* coqc version 8.5 (January 2016) compiled on Jan 23 2016 16:15:22 with OCaml
4.01.0
coqtop version 8.5 (January 2016) *)
+Declare ML Module "ltac_plugin".
Inductive False := .
Axiom proof_admitted : False.
Tactic Notation "admit" := case proof_admitted.
diff --git a/test-suite/bugs/closed/4533.v b/test-suite/bugs/closed/4533.v
index ae17fb145..64c7fd8eb 100644
--- a/test-suite/bugs/closed/4533.v
+++ b/test-suite/bugs/closed/4533.v
@@ -5,6 +5,7 @@ then from 285 lines to 271 lines *)
(* coqc version 8.5 (January 2016) compiled on Jan 23 2016 16:15:22 with OCaml
4.01.0
coqtop version 8.5 (January 2016) *)
+Declare ML Module "ltac_plugin".
Inductive False := .
Axiom proof_admitted : False.
Tactic Notation "admit" := case proof_admitted.
@@ -223,4 +224,4 @@ v = _) r,
| [ |- p2 @ p0 @ p1 @ eissect (to O A) (g x) = r ] => idtac "good"
| [ |- ?G ] => fail 1 "bad" G
end.
- Fail rewrite concat_p_pp. \ No newline at end of file
+ Fail rewrite concat_p_pp.
diff --git a/test-suite/bugs/closed/4544.v b/test-suite/bugs/closed/4544.v
index da140c931..64dd8c304 100644
--- a/test-suite/bugs/closed/4544.v
+++ b/test-suite/bugs/closed/4544.v
@@ -2,6 +2,7 @@
(* File reduced by coq-bug-finder from original input, then from 2553 lines to 1932 lines, then from 1946 lines to 1932 lines, then from 2467 lines to 1002 lines, then from 1016 lines to 1002 lines *)
(* coqc version 8.5 (January 2016) compiled on Jan 23 2016 16:15:22 with OCaml 4.01.0
coqtop version 8.5 (January 2016) *)
+Declare ML Module "ltac_plugin".
Inductive False := .
Axiom proof_admitted : False.
Tactic Notation "admit" := case proof_admitted.
@@ -1004,4 +1005,4 @@ Proof.
Fail Timeout 1 Time rewrite !loops_functor_group.
(* 0.004 s in 8.5rc1, 8.677 s in 8.5 *)
Timeout 1 do 3 rewrite loops_functor_group.
-Abort. \ No newline at end of file
+Abort.