aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorEmilio Jesus Gallego Arias2018-11-08 13:17:14 +0100
committerEmilio Jesus Gallego Arias2018-11-21 01:25:23 +0100
commit002a974b66bc5b8524c8c045d6b9ec4f57aa7734 (patch)
tree573aec1d41d6dc4f8cb61f8fa0826ed9f27e6709
parent968be14b3788e112425eedf696f2e5e35d35ba17 (diff)
[gramlib] [build] Switch make-based system to packed gramlib
This makes the make-based build system stop linking to Camlp5's gramlib and instead links to our own gramlib. We use the style done in the packing of `Stdlib` in OCaml 4.07. As to introduce a minimal amount of noise in history we use an autogenerated `gramlib__pack` directory. Co-authored-by: Gaƫtan Gilbert <gaetan.gilbert@skyskimmer.net>
-rw-r--r--.gitignore3
-rw-r--r--.merlin.in2
-rw-r--r--META.coq.in16
-rw-r--r--Makefile12
-rw-r--r--Makefile.build78
-rw-r--r--Makefile.common3
-rw-r--r--Makefile.install2
-rw-r--r--configure.ml2
-rw-r--r--dev/ci/user-overlays/08985-ejgallego-build+pack_gramlib.sh6
-rw-r--r--parsing/cLexer.ml1
-rw-r--r--parsing/cLexer.mli2
-rw-r--r--parsing/dune1
-rw-r--r--parsing/extend.ml2
-rw-r--r--parsing/pcoq.ml1
-rw-r--r--parsing/pcoq.mli1
-rw-r--r--plugins/funind/plugin_base.dune1
-rw-r--r--plugins/ltac/plugin_base.dune1
-rw-r--r--plugins/ssr/ssrcommon.ml4
-rw-r--r--plugins/ssr/ssrelim.ml2
-rw-r--r--plugins/ssrmatching/plugin_base.dune1
-rw-r--r--printing/dune1
-rw-r--r--printing/proof_diffs.ml2
-rw-r--r--tools/CoqMakefile.in2
-rw-r--r--toplevel/coqloop.ml2
-rw-r--r--toplevel/dune1
-rw-r--r--vernac/dune1
-rw-r--r--vernac/explainErr.ml2
27 files changed, 108 insertions, 44 deletions
diff --git a/.gitignore b/.gitignore
index e513837445..da675309e5 100644
--- a/.gitignore
+++ b/.gitignore
@@ -165,6 +165,9 @@ user-contrib
plugins/ssr/ssrparser.ml
plugins/ssr/ssrvernac.ml
+# gramlib__pack
+gramlib__pack
+
# ocaml dev files
.merlin
META.coq
diff --git a/.merlin.in b/.merlin.in
index 404a7e7935..eade4d19f9 100644
--- a/.merlin.in
+++ b/.merlin.in
@@ -40,6 +40,8 @@ S API
B API
S ide
B ide
+S gramlib__pack
+B gramlib__pack
S tools
B tools
diff --git a/META.coq.in b/META.coq.in
index 16928587cb..bca37ea74b 100644
--- a/META.coq.in
+++ b/META.coq.in
@@ -4,7 +4,7 @@ description = "The Coq Proof Assistant Plugin API"
version = "8.10"
directory = ""
-requires = "camlp5"
+requires = ""
package "grammar" (
@@ -153,12 +153,24 @@ package "proofs" (
)
+package "gramlib" (
+
+ description = "Coq Parsing Library"
+ version = "8.10"
+
+ requires = ""
+ directory = "gramlib__pack"
+
+ archive(byte) = "gramlib.cma"
+ archive(native) = "gramlib.cmxa"
+)
+
package "parsing" (
description = "Coq Parsing Engine"
version = "8.10"
- requires = "camlp5.gramlib, coq.proofs"
+ requires = "coq.gramlib, coq.proofs"
directory = "parsing"
archive(byte) = "parsing.cma"
diff --git a/Makefile b/Makefile
index e0ab169eda..023578f94a 100644
--- a/Makefile
+++ b/Makefile
@@ -94,6 +94,10 @@ EXISTINGMLI := $(call find, '*.mli')
## Files that will be generated
GENMLGFILES:= $(MLGFILES:.mlg=.ml)
+# GRAMFILES must be in linking order
+export GRAMFILES=$(addprefix gramlib__pack/gramlib__,Ploc Plexing Gramext Grammar)
+export GRAMMLFILES := $(addsuffix .ml, $(GRAMFILES)) $(addsuffix .mli, $(GRAMFILES))
+export GENGRAMFILES := $(GRAMMLFILES) gramlib__pack/gramlib.ml
export GENMLFILES:=$(LEXFILES:.mll=.ml) $(YACCFILES:.mly=.ml) $(GENMLGFILES) ide/coqide_os_specific.ml kernel/copcodes.ml
export GENHFILES:=kernel/byterun/coq_jumptbl.h
export GENFILES:=$(GENMLFILES) $(GENMLIFILES) $(GENHFILES)
@@ -190,12 +194,16 @@ META.coq: META.coq.in
.PHONY: clean cleankeepvo objclean cruftclean indepclean docclean archclean optclean clean-ide mlgclean depclean cleanconfig distclean voclean timingclean alienclean
-clean: objclean cruftclean depclean docclean camldevfilesclean
+clean: objclean cruftclean depclean docclean camldevfilesclean gramlibclean
cleankeepvo: indepclean clean-ide optclean cruftclean depclean docclean
objclean: archclean indepclean
+.PHONY: gramlibclean
+gramlibclean:
+ rm -rf gramlib__pack/
+
cruftclean: mlgclean
find . \( -name '*~' -o -name '*.annot' \) -exec rm -f {} +
rm -f gmon.out core
@@ -284,7 +292,7 @@ KNOWNML:=$(EXISTINGML) $(GENMLFILES) $(MLPACKFILES:.mlpack=.ml) \
$(patsubst %.mlp,%.ml,$(wildcard grammar/*.mlp))
KNOWNOBJS:=$(KNOWNML:.ml=.cmo) $(KNOWNML:.ml=.cmx) $(KNOWNML:.ml=.cmi) \
$(MLIFILES:.mli=.cmi) \
- $(MLLIBFILES:.mllib=.cma) $(MLLIBFILES:.mllib=.cmxa) grammar/grammar.cma
+ gramlib__pack/gramlib.cma gramlib__pack/gramlib.cmxa $(MLLIBFILES:.mllib=.cma) $(MLLIBFILES:.mllib=.cmxa) grammar/grammar.cma
ALIENOBJS:=$(filter-out $(KNOWNOBJS),$(EXISTINGOBJS))
alienclean:
diff --git a/Makefile.build b/Makefile.build
index ee856aae8e..dedd36cef2 100644
--- a/Makefile.build
+++ b/Makefile.build
@@ -206,7 +206,7 @@ OCAMLOPT := $(OCAMLFIND) opt $(CAMLFLAGS)
BYTEFLAGS=$(CAMLDEBUG) $(USERFLAGS)
OPTFLAGS=$(CAMLDEBUGOPT) $(CAMLTIMEPROF) $(USERFLAGS) $(FLAMBDA_FLAGS)
-DEPFLAGS=$(LOCALINCLUDES)$(if $(filter plugins/%,$@),, -I ide -I ide/protocol)
+DEPFLAGS=$(LOCALINCLUDES) -map gramlib__pack/gramlib.ml $(if $(filter plugins/%,$@),, -I ide -I ide/protocol)
# On MacOS, the binaries are signed, except our private ones
ifeq ($(shell which codesign > /dev/null 2>&1 && echo $(ARCH)),Darwin)
@@ -264,9 +264,6 @@ PR_O := $(if $(READABLE_ML4),pr_o.cmo,pr_dump.cmo)
# Main packages linked by Coq.
SYSMOD:=-package num,str,unix,dynlink,threads
-# We do not repeat the dependencies already in SYSMOD here
-P4CMA:=gramlib.cma
-
###########################################################################
# Infrastructure for the rest of the Makefile
###########################################################################
@@ -418,7 +415,7 @@ $(COQTOPEXE): $(TOPBINOPT:.opt=.$(BEST))
bin/%.opt$(EXE): topbin/%_bin.ml $(LINKCMX) $(LIBCOQRUN)
$(SHOW)'COQMKTOP -o $@'
$(HIDE)$(OCAMLOPT) -linkall -linkpkg $(MLINCLUDES) \
- $(SYSMOD) -package camlp5.gramlib \
+ $(SYSMOD) \
$(LINKCMX) $(OPTFLAGS) $(LINKMETADATA) $< -o $@
$(STRIP_HIDE) $@
$(CODESIGN_HIDE) $@
@@ -427,7 +424,7 @@ bin/%.byte$(EXE): topbin/%_bin.ml $(LINKCMO) $(LIBCOQRUN)
$(SHOW)'COQMKTOP -o $@'
$(HIDE)$(OCAMLC) -linkall -linkpkg $(MLINCLUDES) \
-I kernel/byterun/ -cclib -lcoqrun $(VMBYTEFLAGS) \
- $(SYSMOD) -package camlp5.gramlib \
+ $(SYSMOD) \
$(LINKCMO) $(BYTEFLAGS) $< -o $@
COQTOP_BYTE=topbin/coqtop_byte_bin.ml
@@ -438,7 +435,7 @@ $(COQTOPBYTE): $(LINKCMO) $(LIBCOQRUN) $(COQTOP_BYTE)
$(SHOW)'COQMKTOP -o $@'
$(HIDE)$(OCAMLC) -linkall -linkpkg -I lib -I vernac -I toplevel \
-I kernel/byterun/ -cclib -lcoqrun $(VMBYTEFLAGS) \
- $(SYSMOD) -package camlp5.gramlib,compiler-libs.toplevel \
+ $(SYSMOD) -package compiler-libs.toplevel \
$(LINKCMO) $(BYTEFLAGS) $(COQTOP_BYTE) -o $@
# For coqc
@@ -630,15 +627,40 @@ test-suite: world byte $(ALLSTDLIB).v
# Default rules for compiling ML code
###########################################################################
-# Target for libraries .cma and .cmxa
+gramlib__pack:
+ mkdir -p $@
-# The dependency over the .mllib is somewhat artificial, since
-# ocamlc -a won't use this file, hence the $(filter-out ...) below.
-# But this ensures that the .cm(x)a is rebuilt when needed,
-# (especially when removing a module in the .mllib).
-# We used to have a "order-only" dependency over .mllib.d here,
-# but the -include mechanism should already ensure that we have
-# up-to-date dependencies.
+# gramlib.ml contents
+gramlib__pack/gramlib.ml: | gramlib__pack
+ echo " \
+module Ploc = Gramlib__Ploc \
+module Plexing = Gramlib__Plexing \
+module Gramext = Gramlib__Gramext \
+module Grammar = Gramlib__Grammar" > $@
+
+gramlib__pack/gramlib__P%: gramlib/p% | gramlib__pack
+ cp -a $< $@
+ sed -e "1i # 1 \"$<\"" -i $@
+gramlib__pack/gramlib__G%: gramlib/g% | gramlib__pack
+ cp -a $< $@
+ sed -e "1i # 1 \"$<\"" -i $@
+
+# Specific rules for gramlib to pack it Dune / OCaml 4.08 style
+GRAMOBJS=$(addsuffix .cmo, $(GRAMFILES))
+
+gramlib__pack/%: COND_BYTEFLAGS+=-no-alias-deps -w -49
+gramlib__pack/%: COND_OPTFLAGS+=-no-alias-deps -w -49
+
+gramlib__pack/gramlib.%: COND_OPENFLAGS=
+gramlib__pack/gramlib__%: COND_OPENFLAGS=-open Gramlib
+
+gramlib__pack/gramlib.cma: $(GRAMOBJS) gramlib__pack/gramlib.cmo
+ $(SHOW)'OCAMLC -a -o $@'
+ $(HIDE)$(OCAMLC) $(MLINCLUDES) $(BYTEFLAGS) -a -o $@ $^
+
+gramlib__pack/gramlib.cmxa: $(GRAMOBJS:.cmo=.cmx) gramlib__pack/gramlib.cmx
+ $(SHOW)'OCAMLOPT -a -o $@'
+ $(HIDE)$(OCAMLOPT) $(MLINCLUDES) $(OPTFLAGS) -a -o $@ $^
# Specific rule for kernel.cma, with $(VMBYTEFLAGS).
# This helps loading dllcoqrun.so during an ocamldebug
@@ -651,6 +673,16 @@ kernel/kernel.cmxa: kernel/kernel.mllib
$(SHOW)'OCAMLOPT -a -o $@'
$(HIDE)$(OCAMLOPT) $(MLINCLUDES) $(OPTFLAGS) -I kernel/byterun/ -cclib -lcoqrun -a -o $@ $(filter-out %.mllib, $^)
+# Target for libraries .cma and .cmxa
+
+# The dependency over the .mllib is somewhat artificial, since
+# ocamlc -a won't use this file, hence the $(filter-out ...) below.
+# But this ensures that the .cm(x)a is rebuilt when needed,
+# (especially when removing a module in the .mllib).
+# We used to have a "order-only" dependency over .mllib.d here,
+# but the -include mechanism should already ensure that we have
+# up-to-date dependencies.
+
%.cma: %.mllib
$(SHOW)'OCAMLC -a -o $@'
$(HIDE)$(OCAMLC) $(MLINCLUDES) $(BYTEFLAGS) -a -o $@ $(filter-out %.mllib, $^)
@@ -672,11 +704,14 @@ kernel/kernel.cmxa: kernel/kernel.mllib
COND_IDEFLAGS=$(if $(filter ide/fake_ide% tools/coq_makefile%,$<), -I ide -I ide/protocol,)
COND_PRINTERFLAGS=$(if $(filter dev/%,$<), -I dev,)
+# For module packing
+COND_OPENFLAGS=
+
COND_BYTEFLAGS= \
- $(COND_IDEFLAGS) $(COND_PRINTERFLAGS) $(MLINCLUDES) $(BYTEFLAGS)
+ $(COND_IDEFLAGS) $(COND_PRINTERFLAGS) $(MLINCLUDES) $(BYTEFLAGS) $(COND_OPENFLAGS)
COND_OPTFLAGS= \
- $(COND_IDEFLAGS) $(MLINCLUDES) $(OPTFLAGS)
+ $(COND_IDEFLAGS) $(MLINCLUDES) $(OPTFLAGS) $(COND_OPENFLAGS)
plugins/micromega/%.cmi: plugins/micromega/%.mli
$(SHOW)'OCAMLC $<'
@@ -780,12 +815,13 @@ kernel/%.cmx: COND_OPTFLAGS+=-w +a-4-44-50
# Ocamldep is now used directly again (thanks to -ml-synonym in OCaml >= 3.12)
OCAMLDEP = $(OCAMLFIND) ocamldep -slash -ml-synonym .mlpack
-MAINMLFILES := $(filter-out checker/% plugins/%, $(MLFILES) $(MLIFILES))
-MAINMLLIBFILES := $(filter-out checker/% plugins/%, $(MLLIBFILES) $(MLPACKFILES))
+MAINMLFILES := $(filter-out gramlib__pack/% checker/% plugins/%, $(MLFILES) $(MLIFILES))
+MAINMLLIBFILES := $(filter-out gramlib__pack/% checker/% plugins/%, $(MLLIBFILES) $(MLPACKFILES))
-$(MLDFILE).d: $(D_DEPEND_BEFORE_SRC) $(MAINMLFILES) $(D_DEPEND_AFTER_SRC) $(GENFILES)
+$(MLDFILE).d: $(D_DEPEND_BEFORE_SRC) $(MAINMLFILES) $(D_DEPEND_AFTER_SRC) $(GENFILES) $(GENGRAMFILES)
$(SHOW)'OCAMLDEP MLFILES MLIFILES'
- $(HIDE)$(OCAMLDEP) $(DEPFLAGS) $(MAINMLFILES) $(TOTARGET)
+ $(HIDE)$(OCAMLDEP) $(DEPFLAGS) -passrest $(MAINMLFILES) -open Gramlib $(GRAMMLFILES) $(TOTARGET)
+#NB: -passrest is needed to avoid ocamlfind reordering the -open Gramlib
$(MLLIBDFILE).d: $(D_DEPEND_BEFORE_SRC) $(MAINMLLIBFILES) $(D_DEPEND_AFTER_SRC) $(OCAMLLIBDEP) $(GENFILES)
$(SHOW)'OCAMLLIBDEP MLLIBFILES MLPACKFILES'
diff --git a/Makefile.common b/Makefile.common
index f2a11ee4b4..dd91f89ffb 100644
--- a/Makefile.common
+++ b/Makefile.common
@@ -91,7 +91,7 @@ MKDIR:=install -d
CORESRCDIRS:=\
coqpp \
config clib lib kernel kernel/byterun library \
- engine pretyping interp proofs parsing printing \
+ engine pretyping interp proofs gramlib__pack parsing printing \
tactics vernac stm toplevel
PLUGINDIRS:=\
@@ -119,6 +119,7 @@ BYTERUN:=$(addprefix kernel/byterun/, \
CORECMA:=config/config.cma clib/clib.cma lib/lib.cma kernel/kernel.cma library/library.cma \
engine/engine.cma pretyping/pretyping.cma interp/interp.cma proofs/proofs.cma \
+ gramlib__pack/gramlib.cma \
parsing/parsing.cma printing/printing.cma tactics/tactics.cma vernac/vernac.cma \
stm/stm.cma toplevel/toplevel.cma
diff --git a/Makefile.install b/Makefile.install
index be6fe54933..8233807e03 100644
--- a/Makefile.install
+++ b/Makefile.install
@@ -93,7 +93,7 @@ install-tools:
INSTALLCMI = $(sort \
$(filter-out checker/% ide/% tools/%, $(MLIFILES:.mli=.cmi)) \
$(foreach lib,$(CORECMA), $(addsuffix .cmi,$($(lib:.cma=_MLLIB_DEPENDENCIES))))) \
- $(PLUGINS:.cmo=.cmi)
+ $(PLUGINS:.cmo=.cmi) gramlib__pack/gramlib.cmi
INSTALLCMX = $(sort $(filter-out checker/% ide/% tools/% dev/% \
configure.cmx toplevel/coqtop_byte_bin.cmx plugins/extraction/big.cmx, \
diff --git a/configure.ml b/configure.ml
index 47f7633ae8..714efd1eb1 100644
--- a/configure.ml
+++ b/configure.ml
@@ -1212,7 +1212,7 @@ let write_configml f =
pr_b "native_compiler" !prefs.nativecompiler;
let core_src_dirs = [ "config"; "lib"; "clib"; "kernel"; "library";
- "engine"; "pretyping"; "interp"; "parsing"; "proofs";
+ "engine"; "pretyping"; "interp"; "gramlib__pack"; "parsing"; "proofs";
"tactics"; "toplevel"; "printing"; "ide"; "stm"; "vernac" ] in
let core_src_dirs = List.fold_left (fun acc core_src_subdir -> acc ^ " \"" ^ core_src_subdir ^ "\";\n")
""
diff --git a/dev/ci/user-overlays/08985-ejgallego-build+pack_gramlib.sh b/dev/ci/user-overlays/08985-ejgallego-build+pack_gramlib.sh
new file mode 100644
index 0000000000..d7130cc67a
--- /dev/null
+++ b/dev/ci/user-overlays/08985-ejgallego-build+pack_gramlib.sh
@@ -0,0 +1,6 @@
+if [ "$CI_PULL_REQUEST" = "8985" ] || [ "$CI_BRANCH" = "build+pack_gramlib" ]; then
+
+ elpi_CI_REF=use_coq_gramlib
+ elpi_CI_GITURL=https://github.com/ejgallego/coq-elpi
+
+fi
diff --git a/parsing/cLexer.ml b/parsing/cLexer.ml
index 619718f723..d81ee475b5 100644
--- a/parsing/cLexer.ml
+++ b/parsing/cLexer.ml
@@ -11,6 +11,7 @@
open Pp
open Util
open Tok
+open Gramlib
(** Location utilities *)
let ploc_file_of_coq_file = function
diff --git a/parsing/cLexer.mli b/parsing/cLexer.mli
index e4aa8debc1..c0ebdd45ef 100644
--- a/parsing/cLexer.mli
+++ b/parsing/cLexer.mli
@@ -40,7 +40,7 @@ where
tok_text : pattern -> string;
tok_comm : mutable option (list location) }
*)
-include Grammar.GLexerType with type te = Tok.t
+include Gramlib.Grammar.GLexerType with type te = Tok.t
module Error : sig
type t
diff --git a/parsing/dune b/parsing/dune
index 0669e3a3c2..e91740650f 100644
--- a/parsing/dune
+++ b/parsing/dune
@@ -2,7 +2,6 @@
(name parsing)
(public_name coq.parsing)
(wrapped false)
- (flags :standard -open Gramlib)
(libraries coq.gramlib proofs))
(rule
diff --git a/parsing/extend.ml b/parsing/extend.ml
index 6fe2956643..5caeab535a 100644
--- a/parsing/extend.ml
+++ b/parsing/extend.ml
@@ -10,7 +10,7 @@
(** Entry keys for constr notations *)
-type 'a entry = 'a Grammar.GMake(CLexer).Entry.e
+type 'a entry = 'a Gramlib.Grammar.GMake(CLexer).Entry.e
type side = Left | Right
diff --git a/parsing/pcoq.ml b/parsing/pcoq.ml
index d4aa598fd8..445338b786 100644
--- a/parsing/pcoq.ml
+++ b/parsing/pcoq.ml
@@ -12,6 +12,7 @@ open CErrors
open Util
open Extend
open Genarg
+open Gramlib
let curry f x y = f (x, y)
let uncurry f (x,y) = f x y
diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli
index c05229d576..593cf59341 100644
--- a/parsing/pcoq.mli
+++ b/parsing/pcoq.mli
@@ -13,6 +13,7 @@ open Extend
open Genarg
open Constrexpr
open Libnames
+open Gramlib
(** The parser of Coq *)
diff --git a/plugins/funind/plugin_base.dune b/plugins/funind/plugin_base.dune
index 9f583234d8..002eb28eea 100644
--- a/plugins/funind/plugin_base.dune
+++ b/plugins/funind/plugin_base.dune
@@ -2,5 +2,4 @@
(name recdef_plugin)
(public_name coq.plugins.recdef)
(synopsis "Coq's functional induction plugin")
- (flags :standard -open Gramlib)
(libraries coq.plugins.extraction))
diff --git a/plugins/ltac/plugin_base.dune b/plugins/ltac/plugin_base.dune
index 1b31655310..5611f5ba16 100644
--- a/plugins/ltac/plugin_base.dune
+++ b/plugins/ltac/plugin_base.dune
@@ -3,7 +3,6 @@
(public_name coq.plugins.ltac)
(synopsis "Coq's LTAC tactic language")
(modules :standard \ tauto)
- (flags :standard -open Gramlib)
(libraries coq.stm))
(library
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
index fa58a1c39a..ff323ddf47 100644
--- a/plugins/ssr/ssrcommon.ml
+++ b/plugins/ssr/ssrcommon.ml
@@ -1092,8 +1092,8 @@ let tclDO n tac =
let _, info = CErrors.push e in
let e' = CErrors.UserError (l, prefix i ++ s) in
Util.iraise (e', info)
- | Ploc.Exc(loc, CErrors.UserError (l, s)) ->
- raise (Ploc.Exc(loc, CErrors.UserError (l, prefix i ++ s))) in
+ | Gramlib.Ploc.Exc(loc, CErrors.UserError (l, s)) ->
+ raise (Gramlib.Ploc.Exc(loc, CErrors.UserError (l, prefix i ++ s))) in
let rec loop i gl =
if i = n then tac_err_at i gl else
(tclTHEN (tac_err_at i) (loop (i + 1))) gl in
diff --git a/plugins/ssr/ssrelim.ml b/plugins/ssr/ssrelim.ml
index d09b81593e..3f87ccd2bf 100644
--- a/plugins/ssr/ssrelim.ml
+++ b/plugins/ssr/ssrelim.ml
@@ -404,7 +404,7 @@ let equality_inj l b id c gl =
let msg = ref "" in
try Proofview.V82.of_tactic (Equality.inj None l b None c) gl
with
- | Ploc.Exc(_,CErrors.UserError (_,s))
+ | Gramlib.Ploc.Exc(_,CErrors.UserError (_,s))
| CErrors.UserError (_,s)
when msg := Pp.string_of_ppcmds s;
!msg = "Not a projectable equality but a discriminable one." ||
diff --git a/plugins/ssrmatching/plugin_base.dune b/plugins/ssrmatching/plugin_base.dune
index 1450a94de1..06f67c3774 100644
--- a/plugins/ssrmatching/plugin_base.dune
+++ b/plugins/ssrmatching/plugin_base.dune
@@ -2,5 +2,4 @@
(name ssrmatching_plugin)
(public_name coq.plugins.ssrmatching)
(synopsis "Coq ssrmatching plugin")
- (flags :standard -open Gramlib)
(libraries coq.plugins.ltac))
diff --git a/printing/dune b/printing/dune
index 837ac48009..3392342165 100644
--- a/printing/dune
+++ b/printing/dune
@@ -2,6 +2,5 @@
(name printing)
(synopsis "Coq's Term Pretty Printing Library")
(public_name coq.printing)
- (flags :standard -open Gramlib)
(wrapped false)
(libraries parsing proofs))
diff --git a/printing/proof_diffs.ml b/printing/proof_diffs.ml
index 0b630b39b5..a2d7c715bc 100644
--- a/printing/proof_diffs.ml
+++ b/printing/proof_diffs.ml
@@ -88,7 +88,7 @@ let tokenize_string s =
let st = CLexer.get_lexer_state () in
try
let istr = Stream.of_string s in
- let lex = CLexer.lexer.Plexing.tok_func istr in
+ let lex = CLexer.lexer.Gramlib.Plexing.tok_func istr in
let toks = stream_tok [] (fst lex) in
CLexer.set_lexer_state st;
toks
diff --git a/tools/CoqMakefile.in b/tools/CoqMakefile.in
index 92cc820483..0d1629c3d5 100644
--- a/tools/CoqMakefile.in
+++ b/tools/CoqMakefile.in
@@ -99,7 +99,7 @@ BEFORE ?=
AFTER ?=
# FIXME this should be generated by Coq (modules already linked by Coq)
-CAMLDONTLINK=camlp5.gramlib,unix,str
+CAMLDONTLINK=unix,str
# OCaml binaries
CAMLC ?= "$(OCAMLFIND)" ocamlc -c
diff --git a/toplevel/coqloop.ml b/toplevel/coqloop.ml
index cbc5c124c8..6d5f049176 100644
--- a/toplevel/coqloop.ml
+++ b/toplevel/coqloop.ml
@@ -255,7 +255,7 @@ let rec discard_to_dot () =
try
Pcoq.Entry.parse parse_to_dot top_buffer.tokens
with
- | Plexing.Error _ | CLexer.Error.E _ -> discard_to_dot ()
+ | Gramlib.Plexing.Error _ | CLexer.Error.E _ -> discard_to_dot ()
| Stm.End_of_input -> raise Stm.End_of_input
| e when CErrors.noncritical e -> ()
diff --git a/toplevel/dune b/toplevel/dune
index c2f9cd662e..f51e50aaa3 100644
--- a/toplevel/dune
+++ b/toplevel/dune
@@ -3,7 +3,6 @@
(public_name coq.toplevel)
(synopsis "Coq's Interactive Shell [terminal-based]")
(wrapped false)
- (flags :standard -open Gramlib)
(libraries num coq.stm))
; Coqlevel provides the `Num` library to plugins, we could also use
; -linkall in the plugins file, to be discussed.
diff --git a/vernac/dune b/vernac/dune
index 4673251002..45b567d631 100644
--- a/vernac/dune
+++ b/vernac/dune
@@ -3,7 +3,6 @@
(synopsis "Coq's Vernacular Language")
(public_name coq.vernac)
(wrapped false)
- (flags :standard -open Gramlib)
(libraries tactics parsing))
(rule
diff --git a/vernac/explainErr.ml b/vernac/explainErr.ml
index e6803443b3..befb4d7ccf 100644
--- a/vernac/explainErr.ml
+++ b/vernac/explainErr.ml
@@ -29,7 +29,7 @@ exception EvaluatedError of Pp.t * exn option
let explain_exn_default = function
(* Basic interaction exceptions *)
| Stream.Error txt -> hov 0 (str "Syntax error: " ++ str txt ++ str ".")
- | Plexing.Error txt -> hov 0 (str "Syntax error: " ++ str txt ++ str ".")
+ | Gramlib.Plexing.Error txt -> hov 0 (str "Syntax error: " ++ str txt ++ str ".")
| CLexer.Error.E err -> hov 0 (str (CLexer.Error.to_string err))
| Sys_error msg -> hov 0 (str "System error: " ++ guill msg)
| Out_of_memory -> hov 0 (str "Out of memory.")