aboutsummaryrefslogtreecommitdiffhomepage
path: root/parsing
diff options
context:
space:
mode:
authorGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2014-05-22 00:58:56 +0200
committerGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2014-05-22 01:03:44 +0200
commit9c927ebdd7e72bb4ae39b549f55f375d66683be5 (patch)
tree3952c62a17e4b73986d5f2bb52034cd8ae57fa6e /parsing
parentcfd8ec82784a5c893b63f3c82736eb76fe487ad7 (diff)
Removing useless use of metaids in tactic AST.
Diffstat (limited to 'parsing')
-rw-r--r--parsing/g_ltac.ml42
-rw-r--r--parsing/g_tactic.ml45
-rw-r--r--parsing/pcoq.mli2
3 files changed, 3 insertions, 6 deletions
diff --git a/parsing/g_ltac.ml4 b/parsing/g_ltac.ml4
index 58def67b6..f70a865d7 100644
--- a/parsing/g_ltac.ml4
+++ b/parsing/g_ltac.ml4
@@ -215,7 +215,7 @@ GEXTEND Gram
| "|"; mrl = LIST1 match_rule SEP "|" -> mrl ] ]
;
message_token:
- [ [ id = identref -> MsgIdent (AI id)
+ [ [ id = identref -> MsgIdent id
| s = STRING -> MsgString s
| n = integer -> MsgInt n ] ]
;
diff --git a/parsing/g_tactic.ml4 b/parsing/g_tactic.ml4
index 2003ec7f9..314252202 100644
--- a/parsing/g_tactic.ml4
+++ b/parsing/g_tactic.ml4
@@ -228,10 +228,7 @@ GEXTEND Gram
;
(* An identifier or a quotation meta-variable *)
id_or_meta:
- [ [ id = identref -> AI id
-
- (* This is used in quotations *)
- | id = METAIDENT -> MetaId (!@loc, id) ] ]
+ [ [ id = identref -> id ] ]
;
open_constr:
[ [ c = constr -> ((),c) ] ]
diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli
index 810a652dc..7d138efa3 100644
--- a/parsing/pcoq.mli
+++ b/parsing/pcoq.mli
@@ -218,7 +218,7 @@ module Tactic :
val red_expr : raw_red_expr Gram.entry
val simple_tactic : raw_atomic_tactic_expr Gram.entry
val simple_intropattern : intro_pattern_expr located Gram.entry
- val clause_dft_concl : Names.Id.t Loc.located Tacexpr.or_metaid Locus.clause_expr Gram.entry
+ val clause_dft_concl : Names.Id.t Loc.located Locus.clause_expr Gram.entry
val tactic_arg : raw_tactic_arg Gram.entry
val tactic_expr : raw_tactic_expr Gram.entry
val binder_tactic : raw_tactic_expr Gram.entry