From 12965209478bd99dfbe57f07d5b525e51b903f22 Mon Sep 17 00:00:00 2001 From: coq Date: Fri, 2 Aug 2002 17:17:42 +0000 Subject: Modules dans COQ\!\!\!\! git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@2957 85f007b7-540e-0410-9357-904b9bb8a0f7 --- .cvsignore | 2 + .depend | 1602 ++++++++++++++++++++---------------- .depend.camlp4 | 1 + .depend.coq | 4 +- Makefile | 72 +- contrib/correctness/ptactic.ml | 2 +- contrib/extraction/common.ml | 13 +- contrib/extraction/common.mli | 2 +- contrib/extraction/extract_env.ml | 11 +- contrib/extraction/extract_env.mli | 2 +- contrib/extraction/extraction.ml | 106 +-- contrib/extraction/extraction.mli | 3 +- contrib/extraction/miniml.mli | 2 +- contrib/extraction/mlutil.ml | 33 +- contrib/extraction/mlutil.mli | 8 +- contrib/extraction/table.ml | 42 +- contrib/extraction/table.mli | 2 +- contrib/field/field.ml4 | 29 +- contrib/fourier/fourierR.ml | 180 ++-- contrib/interface/centaur.ml4 | 26 +- contrib/interface/ctast.ml | 5 +- contrib/interface/dad.ml | 1 + contrib/interface/name_to_ast.ml | 37 +- contrib/interface/name_to_ast.mli | 4 +- contrib/interface/parse.ml | 12 +- contrib/interface/pbp.ml | 4 +- contrib/interface/showproof.ml | 1 + contrib/interface/xlate.ml | 33 +- contrib/omega/coq_omega.ml | 19 +- contrib/ring/quote.ml | 4 +- contrib/ring/ring.ml | 132 ++- contrib/romega/const_omega.ml | 18 +- contrib/romega/refl_omega.ml | 2 +- contrib/xml/xmlcommand.ml4 | 114 +-- contrib/xml/xmlcommand.mli | 10 +- contrib/xml/xmlentries.ml4 | 2 +- dev/.cvsignore | 2 + dev/Makefile.common | 52 ++ dev/Makefile.devel | 46 ++ dev/Makefile.dir | 102 +++ dev/Makefile.subdir | 7 + dev/base_include | 12 +- dev/changements.txt | 185 +++++ dev/db | 7 + dev/include | 1 + dev/objects.el | 137 +++ dev/ocamldebug-v7.template | 2 +- dev/top_printers.ml | 45 +- doc/library.dep.ps | 276 ++++--- kernel/closure.ml | 56 +- kernel/closure.mli | 6 +- kernel/conv_oracle.ml | 20 +- kernel/conv_oracle.mli | 6 +- kernel/cooking.ml | 10 +- kernel/declarations.ml | 83 ++ kernel/declarations.mli | 61 +- kernel/entries.ml | 98 +++ kernel/entries.mli | 98 +++ kernel/environ.ml | 136 ++- kernel/environ.mli | 18 +- kernel/indtypes.ml | 20 +- kernel/indtypes.mli | 26 +- kernel/inductive.ml | 26 +- kernel/inductive.mli | 4 + kernel/mod_typing.ml | 206 +++++ kernel/mod_typing.mli | 25 + kernel/modops.ml | 151 ++++ kernel/modops.mli | 79 ++ kernel/names.ml | 230 +++++- kernel/names.mli | 123 ++- kernel/reduction.ml | 5 + kernel/reduction.mli | 3 + kernel/safe_typing.ml | 533 +++++++++--- kernel/safe_typing.mli | 77 +- kernel/sign.ml | 9 + kernel/sign.mli | 3 + kernel/subtyping.ml | 241 ++++++ kernel/subtyping.mli | 19 + kernel/term.ml | 64 +- kernel/term.mli | 12 +- kernel/term_typing.ml | 111 +++ kernel/term_typing.mli | 34 + kernel/typeops.ml | 13 +- kernel/typeops.mli | 5 +- lib/rtree.ml | 20 + lib/rtree.mli | 5 + lib/util.ml | 59 ++ lib/util.mli | 16 + library/declare.ml | 179 ++-- library/declare.mli | 20 +- library/declaremods.ml | 630 ++++++++++++++ library/declaremods.mli | 85 ++ library/global.ml | 85 +- library/global.mli | 67 +- library/goptions.ml | 41 +- library/goptions.mli | 2 + library/impargs.ml | 130 +-- library/impargs.mli | 1 + library/lib.ml | 447 +++++++--- library/lib.mli | 107 ++- library/libnames.ml | 212 +++++ library/libnames.mli | 118 +++ library/libobject.ml | 98 ++- library/libobject.mli | 93 ++- library/library.ml | 502 +++++------ library/library.mli | 68 +- library/nameops.ml | 85 +- library/nameops.mli | 36 +- library/nametab.ml | 409 +++++---- library/nametab.mli | 111 +-- library/summary.ml | 22 +- library/summary.mli | 3 + parsing/.cvsignore | 1 - parsing/ast.ml | 40 +- parsing/ast.mli | 12 +- parsing/astmod.ml | 147 ++++ parsing/astmod.mli | 30 + parsing/astterm.ml | 22 +- parsing/astterm.mli | 9 +- parsing/coqast.ml | 31 +- parsing/coqast.mli | 7 +- parsing/coqlib.ml | 4 + parsing/coqlib.mli | 5 +- parsing/egrammar.ml | 4 + parsing/egrammar.mli | 3 + parsing/esyntax.ml | 10 +- parsing/esyntax.mli | 2 +- parsing/extend.ml | 50 +- parsing/extend.mli | 31 +- parsing/g_module.ml4 | 81 ++ parsing/g_prim.ml4 | 7 +- parsing/g_vernac.ml4 | 48 +- parsing/genarg.mli | 6 +- parsing/pcoq.ml4 | 13 + parsing/pcoq.mli | 11 +- parsing/ppconstr.ml | 23 +- parsing/ppconstr.mli | 2 +- parsing/pptactic.ml | 4 +- parsing/prettyp.ml | 74 +- parsing/prettyp.mli | 15 +- parsing/printer.ml | 1 + parsing/printer.mli | 1 + parsing/q_coqast.ml4 | 16 +- parsing/search.ml | 39 +- parsing/search.mli | 1 + parsing/termast.ml | 1 + parsing/termast.mli | 3 +- pretyping/classops.ml | 69 +- pretyping/classops.mli | 5 +- pretyping/detyping.ml | 7 +- pretyping/indrec.ml | 42 +- pretyping/pattern.ml | 1 + pretyping/pattern.mli | 1 + pretyping/pretyping.ml | 2 +- pretyping/rawterm.ml | 86 ++ pretyping/rawterm.mli | 5 +- pretyping/recordops.ml | 75 +- pretyping/recordops.mli | 9 +- pretyping/syntax_def.ml | 44 +- pretyping/syntax_def.mli | 5 +- pretyping/tacred.ml | 33 +- pretyping/tacred.mli | 4 +- pretyping/termops.ml | 23 +- pretyping/termops.mli | 4 +- proofs/pfedit.ml | 3 +- proofs/pfedit.mli | 4 +- proofs/proof_trees.ml | 9 +- proofs/proof_type.ml | 3 +- proofs/proof_type.mli | 3 +- proofs/tacexpr.ml | 1 + proofs/tacmach.ml | 2 +- tactics/auto.ml | 272 +++--- tactics/auto.mli | 2 +- tactics/autorewrite.ml | 9 +- tactics/dhyp.ml | 7 +- tactics/elim.ml | 3 +- tactics/equality.ml | 9 +- tactics/leminv.ml | 5 +- tactics/setoid_replace.ml | 94 ++- tactics/tacinterp.ml | 23 +- tactics/tacticals.ml | 3 +- tactics/tactics.ml | 18 +- tactics/tactics.mli | 1 + tactics/tauto.ml4 | 1 + tactics/termdn.ml | 1 + test-suite/modules/Nametab.v | 100 +++ test-suite/modules/Nat.v | 68 ++ test-suite/modules/PO.v | 89 ++ test-suite/modules/Przyklad.v | 193 +++++ test-suite/modules/fun_objects.v | 32 + test-suite/modules/grammar.v | 15 + test-suite/modules/modul.v | 39 + test-suite/modules/modulik.v | 5 + test-suite/modules/obj.v | 26 + test-suite/modules/objects.v | 35 + test-suite/modules/pliczek.v | 3 + test-suite/modules/plik.v | 4 + test-suite/modules/sub_objects.v | 35 + tools/coq_makefile.ml4 | 5 +- toplevel/cerrors.ml | 4 +- toplevel/class.ml | 101 ++- toplevel/class.mli | 1 + toplevel/command.ml | 64 +- toplevel/command.mli | 12 +- toplevel/coqtop.ml | 5 +- toplevel/discharge.ml | 88 +- toplevel/himsg.ml | 6 + toplevel/metasyntax.ml | 69 +- toplevel/metasyntax.mli | 2 +- toplevel/mltop.ml4 | 9 +- toplevel/record.ml | 18 +- toplevel/recordobj.ml | 3 +- toplevel/recordobj.mli | 4 +- toplevel/vernac.ml | 28 +- toplevel/vernacentries.ml | 208 ++++- toplevel/vernacentries.mli | 2 +- toplevel/vernacexpr.ml | 12 +- toplevel/vernacinterp.ml | 1 + 218 files changed, 9349 insertions(+), 3255 deletions(-) create mode 100644 dev/Makefile.common create mode 100644 dev/Makefile.devel create mode 100644 dev/Makefile.dir create mode 100644 dev/Makefile.subdir create mode 100644 dev/objects.el create mode 100644 kernel/entries.ml create mode 100644 kernel/entries.mli create mode 100644 kernel/mod_typing.ml create mode 100644 kernel/mod_typing.mli create mode 100644 kernel/modops.ml create mode 100644 kernel/modops.mli create mode 100644 kernel/subtyping.ml create mode 100644 kernel/subtyping.mli create mode 100644 kernel/term_typing.ml create mode 100644 kernel/term_typing.mli create mode 100644 library/declaremods.ml create mode 100644 library/declaremods.mli create mode 100644 library/libnames.ml create mode 100644 library/libnames.mli create mode 100644 parsing/astmod.ml create mode 100644 parsing/astmod.mli create mode 100644 parsing/g_module.ml4 create mode 100644 test-suite/modules/Nametab.v create mode 100644 test-suite/modules/Nat.v create mode 100644 test-suite/modules/PO.v create mode 100644 test-suite/modules/Przyklad.v create mode 100644 test-suite/modules/fun_objects.v create mode 100644 test-suite/modules/grammar.v create mode 100644 test-suite/modules/modul.v create mode 100644 test-suite/modules/modulik.v create mode 100644 test-suite/modules/obj.v create mode 100644 test-suite/modules/objects.v create mode 100644 test-suite/modules/pliczek.v create mode 100644 test-suite/modules/plik.v create mode 100644 test-suite/modules/sub_objects.v diff --git a/.cvsignore b/.cvsignore index 0089c9a7c..612daa284 100644 --- a/.cvsignore +++ b/.cvsignore @@ -3,3 +3,5 @@ coqtop minicoq coqtop.opt glob.dump +tmp +.depend.devel diff --git a/.depend b/.depend index a660d5660..1eae2938e 100644 --- a/.depend +++ b/.depend @@ -5,100 +5,124 @@ kernel/cooking.cmi: kernel/declarations.cmi kernel/environ.cmi \ kernel/names.cmi kernel/term.cmi kernel/univ.cmi kernel/declarations.cmi: kernel/names.cmi lib/rtree.cmi kernel/sign.cmi \ kernel/term.cmi kernel/univ.cmi +kernel/entries.cmi: kernel/names.cmi kernel/sign.cmi kernel/term.cmi \ + kernel/univ.cmi kernel/environ.cmi: kernel/declarations.cmi kernel/names.cmi kernel/sign.cmi \ kernel/term.cmi kernel/univ.cmi kernel/esubst.cmi: lib/util.cmi -kernel/indtypes.cmi: kernel/declarations.cmi kernel/environ.cmi \ - kernel/names.cmi kernel/term.cmi kernel/typeops.cmi kernel/univ.cmi +kernel/indtypes.cmi: kernel/declarations.cmi kernel/entries.cmi \ + kernel/environ.cmi kernel/names.cmi kernel/term.cmi kernel/typeops.cmi \ + kernel/univ.cmi kernel/inductive.cmi: kernel/declarations.cmi kernel/environ.cmi \ kernel/names.cmi kernel/term.cmi kernel/univ.cmi -kernel/names.cmi: lib/predicate.cmi +kernel/modops.cmi: kernel/declarations.cmi kernel/entries.cmi \ + kernel/environ.cmi kernel/names.cmi kernel/univ.cmi +kernel/mod_typing.cmi: kernel/declarations.cmi kernel/entries.cmi \ + kernel/environ.cmi +kernel/names.cmi: lib/pp.cmi lib/predicate.cmi kernel/reduction.cmi: kernel/environ.cmi kernel/sign.cmi kernel/term.cmi \ kernel/univ.cmi kernel/safe_typing.cmi: kernel/cooking.cmi kernel/declarations.cmi \ - kernel/environ.cmi kernel/indtypes.cmi kernel/names.cmi kernel/term.cmi \ + kernel/entries.cmi kernel/environ.cmi kernel/names.cmi kernel/term.cmi \ kernel/univ.cmi kernel/sign.cmi: kernel/names.cmi kernel/term.cmi +kernel/subtyping.cmi: kernel/declarations.cmi kernel/environ.cmi \ + kernel/univ.cmi kernel/term.cmi: kernel/names.cmi kernel/univ.cmi +kernel/term_typing.cmi: kernel/cooking.cmi kernel/declarations.cmi \ + kernel/entries.cmi kernel/environ.cmi kernel/inductive.cmi \ + kernel/names.cmi kernel/term.cmi kernel/typeops.cmi kernel/univ.cmi kernel/type_errors.cmi: kernel/environ.cmi kernel/names.cmi kernel/term.cmi -kernel/typeops.cmi: kernel/environ.cmi kernel/names.cmi kernel/sign.cmi \ - kernel/term.cmi kernel/univ.cmi +kernel/typeops.cmi: kernel/entries.cmi kernel/environ.cmi kernel/names.cmi \ + kernel/sign.cmi kernel/term.cmi kernel/univ.cmi kernel/univ.cmi: kernel/names.cmi lib/pp.cmi lib/pp.cmi: lib/pp_control.cmi library/declare.cmi: kernel/cooking.cmi kernel/declarations.cmi \ - kernel/environ.cmi kernel/indtypes.cmi library/libobject.cmi \ - library/library.cmi kernel/names.cmi library/nametab.cmi \ - kernel/safe_typing.cmi kernel/sign.cmi kernel/term.cmi kernel/univ.cmi -library/global.cmi: kernel/declarations.cmi kernel/environ.cmi \ - kernel/indtypes.cmi kernel/names.cmi library/nametab.cmi \ - kernel/safe_typing.cmi kernel/sign.cmi kernel/term.cmi kernel/univ.cmi -library/goptions.cmi: kernel/names.cmi library/nametab.cmi lib/pp.cmi \ - kernel/term.cmi lib/util.cmi -library/impargs.cmi: kernel/environ.cmi kernel/names.cmi library/nametab.cmi \ - kernel/term.cmi -library/lib.cmi: library/libobject.cmi kernel/names.cmi library/summary.cmi -library/libobject.cmi: kernel/names.cmi -library/library.cmi: library/lib.cmi library/libobject.cmi kernel/names.cmi \ - library/nametab.cmi lib/pp.cmi lib/system.cmi lib/util.cmi + kernel/entries.cmi kernel/indtypes.cmi library/libnames.cmi \ + library/libobject.cmi library/library.cmi kernel/names.cmi \ + library/nametab.cmi kernel/safe_typing.cmi kernel/sign.cmi \ + kernel/term.cmi kernel/univ.cmi +library/declaremods.cmi: kernel/entries.cmi library/lib.cmi \ + library/libnames.cmi library/libobject.cmi kernel/names.cmi lib/pp.cmi \ + kernel/safe_typing.cmi +library/global.cmi: kernel/declarations.cmi kernel/entries.cmi \ + kernel/environ.cmi kernel/indtypes.cmi library/libnames.cmi \ + kernel/names.cmi kernel/safe_typing.cmi kernel/sign.cmi kernel/term.cmi \ + kernel/univ.cmi +library/goptions.cmi: library/libnames.cmi kernel/names.cmi \ + library/nametab.cmi lib/pp.cmi kernel/term.cmi lib/util.cmi +library/impargs.cmi: kernel/environ.cmi library/libnames.cmi kernel/names.cmi \ + library/nametab.cmi kernel/term.cmi +library/lib.cmi: library/libnames.cmi library/libobject.cmi kernel/names.cmi \ + library/summary.cmi +library/libnames.cmi: kernel/names.cmi lib/pp.cmi lib/predicate.cmi +library/libobject.cmi: library/libnames.cmi kernel/names.cmi +library/library.cmi: library/libnames.cmi library/libobject.cmi \ + kernel/names.cmi lib/pp.cmi lib/system.cmi lib/util.cmi library/nameops.cmi: kernel/environ.cmi kernel/names.cmi lib/pp.cmi \ kernel/term.cmi -library/nametab.cmi: kernel/environ.cmi kernel/names.cmi lib/pp.cmi \ - lib/util.cmi -library/summary.cmi: kernel/names.cmi +library/nametab.cmi: library/libnames.cmi kernel/names.cmi lib/pp.cmi \ + kernel/sign.cmi lib/util.cmi +library/summary.cmi: library/libnames.cmi kernel/names.cmi lib/rtree.cmi: lib/pp.cmi lib/system.cmi: lib/pp.cmi lib/util.cmi: lib/pp.cmi parsing/ast.cmi: parsing/coqast.cmi lib/dyn.cmi parsing/genarg.cmi \ - kernel/names.cmi library/nametab.cmi lib/pp.cmi proofs/tacexpr.cmo \ + library/libnames.cmi kernel/names.cmi lib/pp.cmi proofs/tacexpr.cmo \ lib/util.cmi +parsing/astmod.cmi: parsing/coqast.cmi kernel/declarations.cmi \ + kernel/entries.cmi kernel/environ.cmi pretyping/evd.cmi kernel/names.cmi parsing/astterm.cmi: parsing/coqast.cmi kernel/environ.cmi pretyping/evd.cmi \ - library/impargs.cmi kernel/names.cmi library/nametab.cmi \ + library/impargs.cmi library/libnames.cmi kernel/names.cmi \ pretyping/pattern.cmi pretyping/rawterm.cmi kernel/sign.cmi \ kernel/term.cmi lib/util.cmi -parsing/coqast.cmi: lib/dyn.cmi kernel/names.cmi -parsing/coqlib.cmi: kernel/names.cmi library/nametab.cmi \ +parsing/coqast.cmi: lib/dyn.cmi library/libnames.cmi kernel/names.cmi +parsing/coqlib.cmi: library/libnames.cmi kernel/names.cmi library/nametab.cmi \ pretyping/pattern.cmi kernel/term.cmi parsing/egrammar.cmi: parsing/ast.cmi parsing/coqast.cmi parsing/extend.cmi \ - parsing/genarg.cmi proofs/tacexpr.cmo toplevel/vernacexpr.cmo + parsing/genarg.cmi kernel/names.cmi proofs/tacexpr.cmo \ + toplevel/vernacexpr.cmo parsing/esyntax.cmi: parsing/ast.cmi parsing/coqast.cmi parsing/extend.cmi \ - kernel/names.cmi lib/pp.cmi toplevel/vernacexpr.cmo -parsing/extend.cmi: parsing/ast.cmi parsing/coqast.cmi library/nametab.cmi \ - lib/pp.cmi + library/libnames.cmi lib/pp.cmi toplevel/vernacexpr.cmo +parsing/extend.cmi: parsing/ast.cmi parsing/coqast.cmi library/libnames.cmi \ + kernel/names.cmi lib/pp.cmi parsing/genarg.cmi: kernel/closure.cmi parsing/coqast.cmi pretyping/evd.cmi \ - kernel/names.cmi library/nametab.cmi pretyping/rawterm.cmi \ + library/libnames.cmi kernel/names.cmi pretyping/rawterm.cmi \ kernel/term.cmi lib/util.cmi parsing/g_minicoq.cmi: kernel/environ.cmi kernel/names.cmi lib/pp.cmi \ kernel/term.cmi parsing/g_zsyntax.cmi: parsing/coqast.cmi parsing/pcoq.cmi: parsing/ast.cmi parsing/coqast.cmi parsing/genarg.cmi \ - kernel/names.cmi library/nametab.cmi pretyping/rawterm.cmi \ + library/libnames.cmi kernel/names.cmi pretyping/rawterm.cmi \ proofs/tacexpr.cmo lib/util.cmi toplevel/vernacexpr.cmo parsing/ppconstr.cmi: parsing/coqast.cmi kernel/environ.cmi \ - parsing/extend.cmi parsing/genarg.cmi library/nametab.cmi \ + parsing/extend.cmi parsing/genarg.cmi library/libnames.cmi \ parsing/pcoq.cmi lib/pp.cmi pretyping/rawterm.cmi proofs/tacexpr.cmo \ kernel/term.cmi parsing/pptactic.cmi: parsing/egrammar.cmi parsing/genarg.cmi lib/pp.cmi \ proofs/proof_type.cmi proofs/tacexpr.cmo parsing/prettyp.cmi: pretyping/classops.cmi kernel/environ.cmi \ - library/lib.cmi kernel/names.cmi library/nametab.cmi lib/pp.cmi \ - pretyping/reductionops.cmi kernel/safe_typing.cmi kernel/sign.cmi \ - kernel/term.cmi pretyping/termops.cmi lib/util.cmi -parsing/printer.cmi: kernel/environ.cmi kernel/names.cmi library/nametab.cmi \ - pretyping/pattern.cmi lib/pp.cmi pretyping/rawterm.cmi kernel/sign.cmi \ + library/lib.cmi library/libnames.cmi kernel/names.cmi library/nametab.cmi \ + lib/pp.cmi pretyping/reductionops.cmi kernel/safe_typing.cmi \ + kernel/sign.cmi kernel/term.cmi pretyping/termops.cmi lib/util.cmi +parsing/printer.cmi: kernel/environ.cmi library/libnames.cmi kernel/names.cmi \ + library/nametab.cmi pretyping/pattern.cmi lib/pp.cmi \ + pretyping/rawterm.cmi kernel/sign.cmi kernel/term.cmi \ + pretyping/termops.cmi +parsing/search.cmi: kernel/environ.cmi library/libnames.cmi kernel/names.cmi \ + library/nametab.cmi pretyping/pattern.cmi lib/pp.cmi kernel/term.cmi +parsing/termast.cmi: parsing/coqast.cmi kernel/environ.cmi \ + library/libnames.cmi kernel/names.cmi library/nametab.cmi \ + pretyping/pattern.cmi pretyping/rawterm.cmi kernel/sign.cmi \ kernel/term.cmi pretyping/termops.cmi -parsing/search.cmi: kernel/environ.cmi kernel/names.cmi library/nametab.cmi \ - pretyping/pattern.cmi lib/pp.cmi kernel/term.cmi -parsing/termast.cmi: parsing/coqast.cmi kernel/environ.cmi kernel/names.cmi \ - library/nametab.cmi pretyping/pattern.cmi pretyping/rawterm.cmi \ - kernel/sign.cmi kernel/term.cmi pretyping/termops.cmi pretyping/cases.cmi: kernel/environ.cmi pretyping/evarutil.cmi \ pretyping/evd.cmi pretyping/inductiveops.cmi kernel/names.cmi \ pretyping/rawterm.cmi kernel/term.cmi pretyping/cbv.cmi: kernel/closure.cmi kernel/environ.cmi kernel/esubst.cmi \ kernel/names.cmi kernel/term.cmi pretyping/classops.cmi: library/declare.cmi kernel/environ.cmi \ - pretyping/evd.cmi library/libobject.cmi kernel/names.cmi \ - library/nametab.cmi lib/pp.cmi kernel/term.cmi + pretyping/evd.cmi library/libnames.cmi library/libobject.cmi \ + kernel/names.cmi library/nametab.cmi lib/pp.cmi kernel/term.cmi pretyping/coercion.cmi: kernel/environ.cmi pretyping/evarutil.cmi \ pretyping/evd.cmi kernel/names.cmi pretyping/rawterm.cmi kernel/sign.cmi \ kernel/term.cmi @@ -120,31 +144,32 @@ pretyping/instantiate.cmi: kernel/environ.cmi pretyping/evd.cmi \ kernel/names.cmi kernel/sign.cmi kernel/term.cmi pretyping/multcase.cmi: kernel/environ.cmi pretyping/evarutil.cmi \ pretyping/evd.cmi kernel/names.cmi pretyping/rawterm.cmi kernel/term.cmi -pretyping/pattern.cmi: kernel/environ.cmi pretyping/evd.cmi kernel/names.cmi \ - library/nametab.cmi pretyping/rawterm.cmi kernel/sign.cmi kernel/term.cmi +pretyping/pattern.cmi: kernel/environ.cmi pretyping/evd.cmi \ + library/libnames.cmi kernel/names.cmi library/nametab.cmi \ + pretyping/rawterm.cmi kernel/sign.cmi kernel/term.cmi pretyping/pretype_errors.cmi: kernel/environ.cmi pretyping/evd.cmi \ pretyping/inductiveops.cmi kernel/names.cmi lib/pp.cmi \ pretyping/rawterm.cmi kernel/sign.cmi kernel/term.cmi pretyping/pretyping.cmi: lib/dyn.cmi kernel/environ.cmi \ pretyping/evarutil.cmi pretyping/evd.cmi kernel/names.cmi \ pretyping/rawterm.cmi kernel/sign.cmi kernel/term.cmi -pretyping/rawterm.cmi: lib/dyn.cmi kernel/names.cmi library/nametab.cmi \ - kernel/sign.cmi kernel/term.cmi kernel/univ.cmi -pretyping/recordops.cmi: pretyping/classops.cmi library/libobject.cmi \ - library/library.cmi kernel/names.cmi library/nametab.cmi kernel/term.cmi +pretyping/rawterm.cmi: lib/dyn.cmi library/libnames.cmi kernel/names.cmi \ + library/nametab.cmi kernel/sign.cmi kernel/term.cmi kernel/univ.cmi +pretyping/recordops.cmi: pretyping/classops.cmi library/libnames.cmi \ + library/libobject.cmi library/library.cmi kernel/names.cmi \ + library/nametab.cmi kernel/term.cmi pretyping/reductionops.cmi: kernel/closure.cmi kernel/environ.cmi \ pretyping/evd.cmi kernel/names.cmi kernel/sign.cmi kernel/term.cmi \ kernel/univ.cmi pretyping/retyping.cmi: kernel/environ.cmi pretyping/evd.cmi kernel/names.cmi \ kernel/term.cmi -pretyping/syntax_def.cmi: kernel/names.cmi library/nametab.cmi \ +pretyping/syntax_def.cmi: library/libnames.cmi kernel/names.cmi \ pretyping/rawterm.cmi pretyping/tacred.cmi: kernel/closure.cmi kernel/environ.cmi pretyping/evd.cmi \ kernel/names.cmi pretyping/rawterm.cmi pretyping/reductionops.cmi \ kernel/term.cmi -pretyping/termops.cmi: kernel/environ.cmi kernel/names.cmi \ - library/nametab.cmi lib/pp.cmi kernel/sign.cmi kernel/term.cmi \ - kernel/univ.cmi lib/util.cmi +pretyping/termops.cmi: kernel/environ.cmi kernel/names.cmi lib/pp.cmi \ + kernel/sign.cmi kernel/term.cmi kernel/univ.cmi lib/util.cmi pretyping/typing.cmi: kernel/environ.cmi pretyping/evd.cmi kernel/term.cmi proofs/clenv.cmi: kernel/environ.cmi proofs/evar_refiner.cmi \ pretyping/evd.cmi kernel/names.cmi lib/pp.cmi pretyping/pretyping.cmi \ @@ -155,15 +180,17 @@ proofs/evar_refiner.cmi: parsing/coqast.cmi kernel/environ.cmi \ proofs/refiner.cmi kernel/sign.cmi kernel/term.cmi proofs/logic.cmi: kernel/environ.cmi pretyping/evd.cmi kernel/names.cmi \ lib/pp.cmi proofs/proof_type.cmi kernel/sign.cmi kernel/term.cmi -proofs/pfedit.cmi: parsing/coqast.cmi kernel/environ.cmi pretyping/evd.cmi \ - kernel/names.cmi library/nametab.cmi lib/pp.cmi proofs/proof_type.cmi \ - kernel/safe_typing.cmi kernel/sign.cmi proofs/tacmach.cmi kernel/term.cmi +proofs/pfedit.cmi: parsing/coqast.cmi kernel/entries.cmi kernel/environ.cmi \ + pretyping/evd.cmi library/libnames.cmi kernel/names.cmi \ + library/nametab.cmi lib/pp.cmi proofs/proof_type.cmi kernel/sign.cmi \ + proofs/tacmach.cmi kernel/term.cmi proofs/proof_trees.cmi: kernel/environ.cmi pretyping/evd.cmi kernel/names.cmi \ lib/pp.cmi proofs/proof_type.cmi kernel/sign.cmi kernel/term.cmi \ lib/util.cmi proofs/proof_type.cmi: kernel/closure.cmi kernel/environ.cmi \ - pretyping/evd.cmi parsing/genarg.cmi kernel/names.cmi library/nametab.cmi \ - pretyping/rawterm.cmi proofs/tacexpr.cmo kernel/term.cmi lib/util.cmi + pretyping/evd.cmi parsing/genarg.cmi library/libnames.cmi \ + kernel/names.cmi pretyping/rawterm.cmi proofs/tacexpr.cmo kernel/term.cmi \ + lib/util.cmi proofs/refiner.cmi: pretyping/evd.cmi lib/pp.cmi proofs/proof_trees.cmi \ proofs/proof_type.cmi kernel/sign.cmi proofs/tacexpr.cmo kernel/term.cmi proofs/tacmach.cmi: kernel/closure.cmi parsing/coqast.cmi kernel/environ.cmi \ @@ -174,7 +201,7 @@ proofs/tacmach.cmi: kernel/closure.cmi parsing/coqast.cmi kernel/environ.cmi \ proofs/tactic_debug.cmi: kernel/environ.cmi kernel/names.cmi \ proofs/proof_type.cmi proofs/tacexpr.cmo kernel/term.cmi tactics/auto.cmi: tactics/btermdn.cmi proofs/clenv.cmi kernel/environ.cmi \ - pretyping/evd.cmi kernel/names.cmi library/nametab.cmi \ + pretyping/evd.cmi library/libnames.cmi kernel/names.cmi \ pretyping/pattern.cmi lib/pp.cmi proofs/proof_type.cmi \ pretyping/rawterm.cmi kernel/sign.cmi proofs/tacexpr.cmo \ proofs/tacmach.cmi kernel/term.cmi lib/util.cmi toplevel/vernacexpr.cmo @@ -222,20 +249,21 @@ tactics/tacticals.cmi: proofs/clenv.cmi kernel/names.cmi \ kernel/sign.cmi proofs/tacmach.cmi kernel/term.cmi tactics/wcclausenv.cmi tactics/tactics.cmi: proofs/clenv.cmi kernel/closure.cmi parsing/coqast.cmi \ kernel/environ.cmi proofs/evar_refiner.cmi pretyping/evd.cmi \ - kernel/names.cmi library/nametab.cmi proofs/proof_type.cmi \ - pretyping/rawterm.cmi kernel/reduction.cmi kernel/sign.cmi \ - proofs/tacexpr.cmo proofs/tacmach.cmi pretyping/tacred.cmi \ - tactics/tacticals.cmi kernel/term.cmi + library/libnames.cmi kernel/names.cmi library/nametab.cmi \ + proofs/proof_type.cmi pretyping/rawterm.cmi kernel/reduction.cmi \ + kernel/sign.cmi proofs/tacexpr.cmo proofs/tacmach.cmi \ + pretyping/tacred.cmi tactics/tacticals.cmi kernel/term.cmi tactics/termdn.cmi: pretyping/pattern.cmi kernel/term.cmi tactics/wcclausenv.cmi: proofs/clenv.cmi kernel/environ.cmi \ proofs/evar_refiner.cmi pretyping/evd.cmi kernel/names.cmi \ proofs/proof_type.cmi kernel/sign.cmi proofs/tacmach.cmi kernel/term.cmi toplevel/cerrors.cmi: parsing/coqast.cmi lib/pp.cmi toplevel/class.cmi: pretyping/classops.cmi library/declare.cmi \ - kernel/names.cmi library/nametab.cmi proofs/proof_type.cmi \ - kernel/term.cmi -toplevel/command.cmi: parsing/coqast.cmi kernel/environ.cmi pretyping/evd.cmi \ - kernel/indtypes.cmi library/library.cmi kernel/names.cmi \ + library/libnames.cmi kernel/names.cmi library/nametab.cmi \ + proofs/proof_type.cmi kernel/term.cmi +toplevel/command.cmi: parsing/coqast.cmi library/declare.cmi \ + kernel/entries.cmi kernel/environ.cmi pretyping/evd.cmi \ + library/libnames.cmi library/library.cmi kernel/names.cmi \ library/nametab.cmi proofs/proof_type.cmi pretyping/tacred.cmi \ kernel/term.cmi lib/util.cmi toplevel/vernacexpr.cmo toplevel/coqinit.cmi: kernel/names.cmi @@ -246,18 +274,18 @@ toplevel/himsg.cmi: pretyping/cases.cmi kernel/environ.cmi \ kernel/indtypes.cmi proofs/logic.cmi kernel/names.cmi lib/pp.cmi \ pretyping/pretype_errors.cmi kernel/type_errors.cmi toplevel/metasyntax.cmi: parsing/coqast.cmi parsing/extend.cmi \ - library/nametab.cmi proofs/tacexpr.cmo lib/util.cmi \ + library/libnames.cmi proofs/tacexpr.cmo lib/util.cmi \ toplevel/vernacexpr.cmo toplevel/mltop.cmi: library/libobject.cmi kernel/names.cmi toplevel/protectedtoplevel.cmi: lib/pp.cmi toplevel/record.cmi: parsing/genarg.cmi kernel/names.cmi kernel/sign.cmi \ kernel/term.cmi toplevel/vernacexpr.cmo -toplevel/recordobj.cmi: library/nametab.cmi proofs/proof_type.cmi +toplevel/recordobj.cmi: library/libnames.cmi proofs/proof_type.cmi toplevel/searchisos.cmi: library/libobject.cmi kernel/names.cmi \ kernel/term.cmi toplevel/toplevel.cmi: parsing/pcoq.cmi lib/pp.cmi toplevel/vernacentries.cmi: parsing/coqast.cmi kernel/environ.cmi \ - pretyping/evd.cmi kernel/names.cmi library/nametab.cmi kernel/term.cmi \ + pretyping/evd.cmi library/libnames.cmi kernel/names.cmi kernel/term.cmi \ lib/util.cmi toplevel/vernacexpr.cmo toplevel/vernacinterp.cmi toplevel/vernacinterp.cmi: proofs/tacexpr.cmo toplevel/vernac.cmi: parsing/coqast.cmi parsing/pcoq.cmi \ @@ -306,25 +334,25 @@ contrib/correctness/putil.cmi: kernel/names.cmi contrib/correctness/past.cmi \ contrib/correctness/pwp.cmi: contrib/correctness/peffect.cmi \ contrib/correctness/penv.cmi contrib/correctness/prename.cmi \ kernel/term.cmi -contrib/extraction/common.cmi: contrib/extraction/miniml.cmi \ - contrib/extraction/mlutil.cmi kernel/names.cmi library/nametab.cmi \ - lib/pp.cmi -contrib/extraction/extract_env.cmi: kernel/names.cmi library/nametab.cmi \ +contrib/extraction/common.cmi: library/libnames.cmi \ + contrib/extraction/miniml.cmi contrib/extraction/mlutil.cmi \ + kernel/names.cmi lib/pp.cmi +contrib/extraction/extract_env.cmi: library/libnames.cmi kernel/names.cmi \ lib/util.cmi -contrib/extraction/extraction.cmi: contrib/extraction/miniml.cmi \ - library/nametab.cmi +contrib/extraction/extraction.cmi: kernel/environ.cmi library/libnames.cmi \ + contrib/extraction/miniml.cmi contrib/extraction/haskell.cmi: contrib/extraction/miniml.cmi \ kernel/names.cmi library/nametab.cmi lib/pp.cmi -contrib/extraction/miniml.cmi: kernel/names.cmi library/nametab.cmi \ +contrib/extraction/miniml.cmi: library/libnames.cmi kernel/names.cmi \ lib/pp.cmi kernel/term.cmi -contrib/extraction/mlutil.cmi: contrib/extraction/miniml.cmi kernel/names.cmi \ - library/nametab.cmi kernel/term.cmi +contrib/extraction/mlutil.cmi: library/libnames.cmi \ + contrib/extraction/miniml.cmi kernel/names.cmi kernel/term.cmi contrib/extraction/ocaml.cmi: contrib/extraction/miniml.cmi kernel/names.cmi \ library/nametab.cmi lib/pp.cmi contrib/extraction/table.cmi \ kernel/term.cmi contrib/extraction/scheme.cmi: contrib/extraction/miniml.cmi kernel/names.cmi \ library/nametab.cmi lib/pp.cmi -contrib/extraction/table.cmi: kernel/names.cmi library/nametab.cmi \ +contrib/extraction/table.cmi: library/libnames.cmi kernel/names.cmi \ lib/util.cmi toplevel/vernacinterp.cmi contrib/interface/blast.cmi: proofs/proof_type.cmi proofs/tacexpr.cmo \ proofs/tacmach.cmi @@ -332,7 +360,7 @@ contrib/interface/dad.cmi: contrib/interface/ctast.cmo proofs/proof_type.cmi \ proofs/tacexpr.cmo proofs/tacmach.cmi contrib/interface/debug_tac.cmi: proofs/proof_type.cmi proofs/tacexpr.cmo \ proofs/tacmach.cmi -contrib/interface/name_to_ast.cmi: parsing/coqast.cmi library/nametab.cmi \ +contrib/interface/name_to_ast.cmi: parsing/coqast.cmi library/libnames.cmi \ toplevel/vernacexpr.cmo contrib/interface/pbp.cmi: kernel/names.cmi proofs/proof_type.cmi \ proofs/tacexpr.cmo proofs/tacmach.cmi @@ -355,23 +383,25 @@ contrib/jprover/jall.cmi: contrib/jprover/jlogic.cmi \ contrib/jprover/jterm.cmi contrib/jprover/opname.cmi contrib/jprover/jlogic.cmi: contrib/jprover/jterm.cmi contrib/jprover/jterm.cmi: contrib/jprover/opname.cmi -contrib/xml/xmlcommand.cmi: kernel/names.cmi library/nametab.cmi lib/util.cmi +contrib/xml/xmlcommand.cmi: library/libnames.cmi kernel/names.cmi \ + lib/util.cmi +tmp/Discharge/libnames.cmi: kernel/names.cmi lib/pp.cmi lib/predicate.cmi config/coq_config.cmo: config/coq_config.cmi config/coq_config.cmx: config/coq_config.cmi dev/db_printers.cmo: kernel/names.cmi lib/pp.cmi dev/db_printers.cmx: kernel/names.cmx lib/pp.cmx dev/top_printers.cmo: parsing/ast.cmi toplevel/cerrors.cmi proofs/clenv.cmi \ - kernel/environ.cmi pretyping/evd.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/pp.cmi parsing/pptactic.cmi parsing/printer.cmi \ - proofs/proof_trees.cmi proofs/refiner.cmi kernel/sign.cmi lib/system.cmi \ - proofs/tacmach.cmi kernel/term.cmi parsing/termast.cmi \ - pretyping/termops.cmi kernel/univ.cmi + kernel/environ.cmi pretyping/evd.cmi library/libnames.cmi \ + library/libobject.cmi library/nameops.cmi kernel/names.cmi lib/pp.cmi \ + parsing/pptactic.cmi parsing/printer.cmi proofs/proof_trees.cmi \ + proofs/refiner.cmi kernel/sign.cmi lib/system.cmi proofs/tacmach.cmi \ + kernel/term.cmi parsing/termast.cmi pretyping/termops.cmi kernel/univ.cmi dev/top_printers.cmx: parsing/ast.cmx toplevel/cerrors.cmx proofs/clenv.cmx \ - kernel/environ.cmx pretyping/evd.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/pp.cmx parsing/pptactic.cmx parsing/printer.cmx \ - proofs/proof_trees.cmx proofs/refiner.cmx kernel/sign.cmx lib/system.cmx \ - proofs/tacmach.cmx kernel/term.cmx parsing/termast.cmx \ - pretyping/termops.cmx kernel/univ.cmx + kernel/environ.cmx pretyping/evd.cmx library/libnames.cmx \ + library/libobject.cmx library/nameops.cmx kernel/names.cmx lib/pp.cmx \ + parsing/pptactic.cmx parsing/printer.cmx proofs/proof_trees.cmx \ + proofs/refiner.cmx kernel/sign.cmx lib/system.cmx proofs/tacmach.cmx \ + kernel/term.cmx parsing/termast.cmx pretyping/termops.cmx kernel/univ.cmx kernel/closure.cmo: kernel/environ.cmi kernel/esubst.cmi kernel/names.cmi \ lib/pp.cmi kernel/term.cmi lib/util.cmi kernel/closure.cmi kernel/closure.cmx: kernel/environ.cmx kernel/esubst.cmx kernel/names.cmx \ @@ -387,29 +417,47 @@ kernel/cooking.cmx: kernel/declarations.cmx kernel/environ.cmx \ kernel/names.cmx lib/pp.cmx kernel/reduction.cmx kernel/sign.cmx \ kernel/term.cmx lib/util.cmx kernel/cooking.cmi kernel/declarations.cmo: kernel/names.cmi lib/rtree.cmi kernel/sign.cmi \ - kernel/term.cmi kernel/univ.cmi kernel/declarations.cmi + kernel/term.cmi kernel/univ.cmi lib/util.cmi kernel/declarations.cmi kernel/declarations.cmx: kernel/names.cmx lib/rtree.cmx kernel/sign.cmx \ - kernel/term.cmx kernel/univ.cmx kernel/declarations.cmi + kernel/term.cmx kernel/univ.cmx lib/util.cmx kernel/declarations.cmi +kernel/entries.cmo: kernel/names.cmi kernel/sign.cmi kernel/term.cmi \ + kernel/univ.cmi kernel/entries.cmi +kernel/entries.cmx: kernel/names.cmx kernel/sign.cmx kernel/term.cmx \ + kernel/univ.cmx kernel/entries.cmi kernel/environ.cmo: kernel/declarations.cmi kernel/names.cmi kernel/sign.cmi \ kernel/term.cmi kernel/univ.cmi lib/util.cmi kernel/environ.cmi kernel/environ.cmx: kernel/declarations.cmx kernel/names.cmx kernel/sign.cmx \ kernel/term.cmx kernel/univ.cmx lib/util.cmx kernel/environ.cmi kernel/esubst.cmo: lib/util.cmi kernel/esubst.cmi kernel/esubst.cmx: lib/util.cmx kernel/esubst.cmi -kernel/indtypes.cmo: kernel/declarations.cmi kernel/environ.cmi \ - kernel/inductive.cmi kernel/names.cmi kernel/reduction.cmi lib/rtree.cmi \ - kernel/sign.cmi kernel/term.cmi kernel/typeops.cmi kernel/univ.cmi \ - lib/util.cmi kernel/indtypes.cmi -kernel/indtypes.cmx: kernel/declarations.cmx kernel/environ.cmx \ - kernel/inductive.cmx kernel/names.cmx kernel/reduction.cmx lib/rtree.cmx \ - kernel/sign.cmx kernel/term.cmx kernel/typeops.cmx kernel/univ.cmx \ - lib/util.cmx kernel/indtypes.cmi +kernel/indtypes.cmo: kernel/declarations.cmi kernel/entries.cmi \ + kernel/environ.cmi kernel/inductive.cmi kernel/names.cmi \ + kernel/reduction.cmi lib/rtree.cmi kernel/sign.cmi kernel/term.cmi \ + kernel/typeops.cmi kernel/univ.cmi lib/util.cmi kernel/indtypes.cmi +kernel/indtypes.cmx: kernel/declarations.cmx kernel/entries.cmx \ + kernel/environ.cmx kernel/inductive.cmx kernel/names.cmx \ + kernel/reduction.cmx lib/rtree.cmx kernel/sign.cmx kernel/term.cmx \ + kernel/typeops.cmx kernel/univ.cmx lib/util.cmx kernel/indtypes.cmi kernel/inductive.cmo: kernel/declarations.cmi kernel/environ.cmi \ kernel/names.cmi kernel/reduction.cmi kernel/sign.cmi kernel/term.cmi \ kernel/type_errors.cmi kernel/univ.cmi lib/util.cmi kernel/inductive.cmi kernel/inductive.cmx: kernel/declarations.cmx kernel/environ.cmx \ kernel/names.cmx kernel/reduction.cmx kernel/sign.cmx kernel/term.cmx \ kernel/type_errors.cmx kernel/univ.cmx lib/util.cmx kernel/inductive.cmi +kernel/modops.cmo: kernel/declarations.cmi kernel/entries.cmi \ + kernel/environ.cmi kernel/names.cmi kernel/term.cmi lib/util.cmi \ + kernel/modops.cmi +kernel/modops.cmx: kernel/declarations.cmx kernel/entries.cmx \ + kernel/environ.cmx kernel/names.cmx kernel/term.cmx lib/util.cmx \ + kernel/modops.cmi +kernel/mod_typing.cmo: kernel/declarations.cmi kernel/entries.cmi \ + kernel/environ.cmi kernel/modops.cmi kernel/names.cmi \ + kernel/subtyping.cmi kernel/term_typing.cmi kernel/univ.cmi lib/util.cmi \ + kernel/mod_typing.cmi +kernel/mod_typing.cmx: kernel/declarations.cmx kernel/entries.cmx \ + kernel/environ.cmx kernel/modops.cmx kernel/names.cmx \ + kernel/subtyping.cmx kernel/term_typing.cmx kernel/univ.cmx lib/util.cmx \ + kernel/mod_typing.cmi kernel/names.cmo: lib/hashcons.cmi lib/pp.cmi lib/predicate.cmi lib/util.cmi \ kernel/names.cmi kernel/names.cmx: lib/hashcons.cmx lib/pp.cmx lib/predicate.cmx lib/util.cmx \ @@ -423,37 +471,59 @@ kernel/reduction.cmx: kernel/closure.cmx kernel/conv_oracle.cmx \ kernel/names.cmx kernel/sign.cmx kernel/term.cmx kernel/univ.cmx \ lib/util.cmx kernel/reduction.cmi kernel/safe_typing.cmo: kernel/cooking.cmi kernel/declarations.cmi \ - kernel/environ.cmi kernel/indtypes.cmi kernel/inductive.cmi \ - kernel/names.cmi kernel/reduction.cmi kernel/sign.cmi kernel/term.cmi \ + kernel/entries.cmi kernel/environ.cmi kernel/indtypes.cmi \ + kernel/inductive.cmi kernel/mod_typing.cmi kernel/modops.cmi \ + kernel/names.cmi kernel/reduction.cmi kernel/sign.cmi \ + kernel/subtyping.cmi kernel/term.cmi kernel/term_typing.cmi \ kernel/type_errors.cmi kernel/typeops.cmi kernel/univ.cmi lib/util.cmi \ kernel/safe_typing.cmi kernel/safe_typing.cmx: kernel/cooking.cmx kernel/declarations.cmx \ - kernel/environ.cmx kernel/indtypes.cmx kernel/inductive.cmx \ - kernel/names.cmx kernel/reduction.cmx kernel/sign.cmx kernel/term.cmx \ + kernel/entries.cmx kernel/environ.cmx kernel/indtypes.cmx \ + kernel/inductive.cmx kernel/mod_typing.cmx kernel/modops.cmx \ + kernel/names.cmx kernel/reduction.cmx kernel/sign.cmx \ + kernel/subtyping.cmx kernel/term.cmx kernel/term_typing.cmx \ kernel/type_errors.cmx kernel/typeops.cmx kernel/univ.cmx lib/util.cmx \ kernel/safe_typing.cmi kernel/sign.cmo: kernel/names.cmi kernel/term.cmi lib/util.cmi \ kernel/sign.cmi kernel/sign.cmx: kernel/names.cmx kernel/term.cmx lib/util.cmx \ kernel/sign.cmi +kernel/subtyping.cmo: kernel/declarations.cmi kernel/environ.cmi \ + kernel/inductive.cmi kernel/modops.cmi kernel/names.cmi \ + kernel/reduction.cmi kernel/term.cmi kernel/univ.cmi lib/util.cmi \ + kernel/subtyping.cmi +kernel/subtyping.cmx: kernel/declarations.cmx kernel/environ.cmx \ + kernel/inductive.cmx kernel/modops.cmx kernel/names.cmx \ + kernel/reduction.cmx kernel/term.cmx kernel/univ.cmx lib/util.cmx \ + kernel/subtyping.cmi kernel/term.cmo: kernel/esubst.cmi lib/hashcons.cmi kernel/names.cmi \ lib/pp.cmi kernel/univ.cmi lib/util.cmi kernel/term.cmi kernel/term.cmx: kernel/esubst.cmx lib/hashcons.cmx kernel/names.cmx \ lib/pp.cmx kernel/univ.cmx lib/util.cmx kernel/term.cmi +kernel/term_typing.cmo: kernel/cooking.cmi kernel/declarations.cmi \ + kernel/entries.cmi kernel/environ.cmi kernel/indtypes.cmi \ + kernel/inductive.cmi kernel/names.cmi kernel/reduction.cmi \ + kernel/sign.cmi kernel/term.cmi kernel/type_errors.cmi kernel/typeops.cmi \ + kernel/univ.cmi lib/util.cmi kernel/term_typing.cmi +kernel/term_typing.cmx: kernel/cooking.cmx kernel/declarations.cmx \ + kernel/entries.cmx kernel/environ.cmx kernel/indtypes.cmx \ + kernel/inductive.cmx kernel/names.cmx kernel/reduction.cmx \ + kernel/sign.cmx kernel/term.cmx kernel/type_errors.cmx kernel/typeops.cmx \ + kernel/univ.cmx lib/util.cmx kernel/term_typing.cmi kernel/type_errors.cmo: kernel/environ.cmi kernel/names.cmi \ kernel/reduction.cmi kernel/sign.cmi kernel/term.cmi \ kernel/type_errors.cmi kernel/type_errors.cmx: kernel/environ.cmx kernel/names.cmx \ kernel/reduction.cmx kernel/sign.cmx kernel/term.cmx \ kernel/type_errors.cmi -kernel/typeops.cmo: kernel/declarations.cmi kernel/environ.cmi \ - kernel/inductive.cmi kernel/names.cmi kernel/reduction.cmi \ - kernel/sign.cmi kernel/term.cmi kernel/type_errors.cmi kernel/univ.cmi \ - lib/util.cmi kernel/typeops.cmi -kernel/typeops.cmx: kernel/declarations.cmx kernel/environ.cmx \ - kernel/inductive.cmx kernel/names.cmx kernel/reduction.cmx \ - kernel/sign.cmx kernel/term.cmx kernel/type_errors.cmx kernel/univ.cmx \ - lib/util.cmx kernel/typeops.cmi +kernel/typeops.cmo: kernel/declarations.cmi kernel/entries.cmi \ + kernel/environ.cmi kernel/inductive.cmi kernel/names.cmi \ + kernel/reduction.cmi kernel/sign.cmi kernel/term.cmi \ + kernel/type_errors.cmi kernel/univ.cmi lib/util.cmi kernel/typeops.cmi +kernel/typeops.cmx: kernel/declarations.cmx kernel/entries.cmx \ + kernel/environ.cmx kernel/inductive.cmx kernel/names.cmx \ + kernel/reduction.cmx kernel/sign.cmx kernel/term.cmx \ + kernel/type_errors.cmx kernel/univ.cmx lib/util.cmx kernel/typeops.cmi kernel/univ.cmo: lib/hashcons.cmi kernel/names.cmi lib/pp.cmi lib/util.cmi \ kernel/univ.cmi kernel/univ.cmx: lib/hashcons.cmx kernel/names.cmx lib/pp.cmx lib/util.cmx \ @@ -486,82 +556,96 @@ lib/predicate.cmo: lib/predicate.cmi lib/predicate.cmx: lib/predicate.cmi lib/profile.cmo: lib/profile.cmi lib/profile.cmx: lib/profile.cmi -library/declare.cmo: kernel/declarations.cmi kernel/environ.cmi \ - library/global.cmi library/impargs.cmi kernel/indtypes.cmi \ - kernel/inductive.cmi library/lib.cmi library/libobject.cmi \ - library/library.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/pp.cmi kernel/reduction.cmi \ - kernel/safe_typing.cmi kernel/sign.cmi library/summary.cmi \ - kernel/term.cmi kernel/type_errors.cmi kernel/typeops.cmi kernel/univ.cmi \ - lib/util.cmi library/declare.cmi -library/declare.cmx: kernel/declarations.cmx kernel/environ.cmx \ - library/global.cmx library/impargs.cmx kernel/indtypes.cmx \ - kernel/inductive.cmx library/lib.cmx library/libobject.cmx \ - library/library.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/pp.cmx kernel/reduction.cmx \ - kernel/safe_typing.cmx kernel/sign.cmx library/summary.cmx \ - kernel/term.cmx kernel/type_errors.cmx kernel/typeops.cmx kernel/univ.cmx \ - lib/util.cmx library/declare.cmi -library/global.cmo: kernel/environ.cmi kernel/inductive.cmi kernel/names.cmi \ - library/nametab.cmi kernel/safe_typing.cmi kernel/sign.cmi \ - library/summary.cmi kernel/term.cmi lib/util.cmi library/global.cmi -library/global.cmx: kernel/environ.cmx kernel/inductive.cmx kernel/names.cmx \ - library/nametab.cmx kernel/safe_typing.cmx kernel/sign.cmx \ - library/summary.cmx kernel/term.cmx lib/util.cmx library/global.cmi -library/goptions.cmo: library/global.cmi library/lib.cmi \ +library/declare.cmo: kernel/declarations.cmi kernel/entries.cmi \ + kernel/environ.cmi library/global.cmi library/impargs.cmi \ + kernel/indtypes.cmi kernel/inductive.cmi library/lib.cmi \ + library/libnames.cmi library/libobject.cmi library/library.cmi \ + library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/pp.cmi \ + kernel/reduction.cmi kernel/safe_typing.cmi kernel/sign.cmi \ + library/summary.cmi kernel/term.cmi kernel/type_errors.cmi \ + kernel/typeops.cmi kernel/univ.cmi lib/util.cmi library/declare.cmi +library/declare.cmx: kernel/declarations.cmx kernel/entries.cmx \ + kernel/environ.cmx library/global.cmx library/impargs.cmx \ + kernel/indtypes.cmx kernel/inductive.cmx library/lib.cmx \ + library/libnames.cmx library/libobject.cmx library/library.cmx \ + library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/pp.cmx \ + kernel/reduction.cmx kernel/safe_typing.cmx kernel/sign.cmx \ + library/summary.cmx kernel/term.cmx kernel/type_errors.cmx \ + kernel/typeops.cmx kernel/univ.cmx lib/util.cmx library/declare.cmi +library/declaremods.cmo: kernel/declarations.cmi kernel/entries.cmi \ + library/global.cmi library/lib.cmi library/libnames.cmi \ + library/libobject.cmi kernel/names.cmi library/nametab.cmi lib/pp.cmi \ + library/summary.cmi lib/util.cmi library/declaremods.cmi +library/declaremods.cmx: kernel/declarations.cmx kernel/entries.cmx \ + library/global.cmx library/lib.cmx library/libnames.cmx \ + library/libobject.cmx kernel/names.cmx library/nametab.cmx lib/pp.cmx \ + library/summary.cmx lib/util.cmx library/declaremods.cmi +library/global.cmo: kernel/environ.cmi kernel/inductive.cmi \ + library/libnames.cmi kernel/names.cmi kernel/safe_typing.cmi \ + kernel/sign.cmi library/summary.cmi kernel/term.cmi lib/util.cmi \ + library/global.cmi +library/global.cmx: kernel/environ.cmx kernel/inductive.cmx \ + library/libnames.cmx kernel/names.cmx kernel/safe_typing.cmx \ + kernel/sign.cmx library/summary.cmx kernel/term.cmx lib/util.cmx \ + library/global.cmi +library/goptions.cmo: library/lib.cmi library/libnames.cmi \ library/libobject.cmi kernel/names.cmi library/nametab.cmi lib/pp.cmi \ library/summary.cmi kernel/term.cmi lib/util.cmi library/goptions.cmi -library/goptions.cmx: library/global.cmx library/lib.cmx \ +library/goptions.cmx: library/lib.cmx library/libnames.cmx \ library/libobject.cmx kernel/names.cmx library/nametab.cmx lib/pp.cmx \ library/summary.cmx kernel/term.cmx lib/util.cmx library/goptions.cmi library/impargs.cmo: kernel/declarations.cmi kernel/environ.cmi \ library/global.cmi kernel/inductive.cmi library/lib.cmi \ - library/libobject.cmi kernel/names.cmi library/nametab.cmi \ - kernel/reduction.cmi library/summary.cmi kernel/term.cmi lib/util.cmi \ - library/impargs.cmi + library/libnames.cmi library/libobject.cmi kernel/names.cmi \ + library/nametab.cmi kernel/reduction.cmi library/summary.cmi \ + kernel/term.cmi lib/util.cmi library/impargs.cmi library/impargs.cmx: kernel/declarations.cmx kernel/environ.cmx \ library/global.cmx kernel/inductive.cmx library/lib.cmx \ - library/libobject.cmx kernel/names.cmx library/nametab.cmx \ - kernel/reduction.cmx library/summary.cmx kernel/term.cmx lib/util.cmx \ - library/impargs.cmi -library/lib.cmo: library/libobject.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/pp.cmi library/summary.cmi lib/util.cmi \ - library/lib.cmi -library/lib.cmx: library/libobject.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/pp.cmx library/summary.cmx lib/util.cmx \ - library/lib.cmi -library/libobject.cmo: lib/dyn.cmi kernel/names.cmi lib/util.cmi \ - library/libobject.cmi -library/libobject.cmx: lib/dyn.cmx kernel/names.cmx lib/util.cmx \ - library/libobject.cmi -library/library.cmo: kernel/environ.cmi library/global.cmi library/lib.cmi \ - library/libobject.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/options.cmi lib/pp.cmi library/summary.cmi \ - lib/system.cmi lib/util.cmi library/library.cmi -library/library.cmx: kernel/environ.cmx library/global.cmx library/lib.cmx \ - library/libobject.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/options.cmx lib/pp.cmx library/summary.cmx \ - lib/system.cmx lib/util.cmx library/library.cmi + library/libnames.cmx library/libobject.cmx kernel/names.cmx \ + library/nametab.cmx kernel/reduction.cmx library/summary.cmx \ + kernel/term.cmx lib/util.cmx library/impargs.cmi +library/lib.cmo: library/libnames.cmi library/libobject.cmi \ + library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/pp.cmi \ + library/summary.cmi lib/util.cmi library/lib.cmi +library/lib.cmx: library/libnames.cmx library/libobject.cmx \ + library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/pp.cmx \ + library/summary.cmx lib/util.cmx library/lib.cmi +library/libnames.cmo: kernel/names.cmi lib/pp.cmi lib/predicate.cmi \ + lib/util.cmi library/libnames.cmi +library/libnames.cmx: kernel/names.cmx lib/pp.cmx lib/predicate.cmx \ + lib/util.cmx library/libnames.cmi +library/libobject.cmo: lib/dyn.cmi library/libnames.cmi kernel/names.cmi \ + lib/util.cmi library/libobject.cmi +library/libobject.cmx: lib/dyn.cmx library/libnames.cmx kernel/names.cmx \ + lib/util.cmx library/libobject.cmi +library/library.cmo: library/declaremods.cmi library/lib.cmi \ + library/libnames.cmi library/libobject.cmi library/nameops.cmi \ + kernel/names.cmi library/nametab.cmi lib/options.cmi lib/pp.cmi \ + kernel/safe_typing.cmi library/summary.cmi lib/system.cmi lib/util.cmi \ + library/library.cmi +library/library.cmx: library/declaremods.cmx library/lib.cmx \ + library/libnames.cmx library/libobject.cmx library/nameops.cmx \ + kernel/names.cmx library/nametab.cmx lib/options.cmx lib/pp.cmx \ + kernel/safe_typing.cmx library/summary.cmx lib/system.cmx lib/util.cmx \ + library/library.cmi library/nameops.cmo: kernel/declarations.cmi kernel/environ.cmi \ kernel/names.cmi lib/pp.cmi kernel/term.cmi lib/util.cmi \ library/nameops.cmi library/nameops.cmx: kernel/declarations.cmx kernel/environ.cmx \ kernel/names.cmx lib/pp.cmx kernel/term.cmx lib/util.cmx \ library/nameops.cmi -library/nametab.cmo: kernel/declarations.cmi kernel/environ.cmi \ - library/nameops.cmi kernel/names.cmi lib/pp.cmi library/summary.cmi \ - lib/util.cmi library/nametab.cmi -library/nametab.cmx: kernel/declarations.cmx kernel/environ.cmx \ - library/nameops.cmx kernel/names.cmx lib/pp.cmx library/summary.cmx \ - lib/util.cmx library/nametab.cmi +library/nametab.cmo: kernel/declarations.cmi library/libnames.cmi \ + library/nameops.cmi kernel/names.cmi lib/pp.cmi kernel/sign.cmi \ + library/summary.cmi lib/util.cmi library/nametab.cmi +library/nametab.cmx: kernel/declarations.cmx library/libnames.cmx \ + library/nameops.cmx kernel/names.cmx lib/pp.cmx kernel/sign.cmx \ + library/summary.cmx lib/util.cmx library/nametab.cmi library/states.cmo: library/lib.cmi library/library.cmi library/summary.cmi \ lib/system.cmi library/states.cmi library/states.cmx: library/lib.cmx library/library.cmx library/summary.cmx \ lib/system.cmx library/states.cmi -library/summary.cmo: lib/dyn.cmi kernel/names.cmi lib/pp.cmi lib/util.cmi \ - library/summary.cmi -library/summary.cmx: lib/dyn.cmx kernel/names.cmx lib/pp.cmx lib/util.cmx \ - library/summary.cmi +library/summary.cmo: lib/dyn.cmi lib/pp.cmi lib/util.cmi library/summary.cmi +library/summary.cmx: lib/dyn.cmx lib/pp.cmx lib/util.cmx library/summary.cmi lib/rtree.cmo: lib/pp.cmi lib/util.cmi lib/rtree.cmi lib/rtree.cmx: lib/pp.cmx lib/util.cmx lib/rtree.cmi lib/stamps.cmo: lib/stamps.cmi @@ -579,15 +663,21 @@ parsing/argextend.cmx: parsing/ast.cmx parsing/genarg.cmx parsing/pcoq.cmx \ parsing/q_coqast.cmx parsing/q_util.cmx lib/util.cmx \ toplevel/vernacexpr.cmx parsing/ast.cmo: parsing/coqast.cmi lib/dyn.cmi parsing/genarg.cmi \ - library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/pp.cmi \ - proofs/tacexpr.cmo lib/util.cmi parsing/ast.cmi + library/libnames.cmi kernel/names.cmi lib/pp.cmi proofs/tacexpr.cmo \ + lib/util.cmi parsing/ast.cmi parsing/ast.cmx: parsing/coqast.cmx lib/dyn.cmx parsing/genarg.cmx \ - library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/pp.cmx \ - proofs/tacexpr.cmx lib/util.cmx parsing/ast.cmi + library/libnames.cmx kernel/names.cmx lib/pp.cmx proofs/tacexpr.cmx \ + lib/util.cmx parsing/ast.cmi +parsing/astmod.cmo: parsing/astterm.cmi parsing/coqast.cmi kernel/entries.cmi \ + library/lib.cmi library/libnames.cmi kernel/modops.cmi kernel/names.cmi \ + library/nametab.cmi lib/pp.cmi lib/util.cmi parsing/astmod.cmi +parsing/astmod.cmx: parsing/astterm.cmx parsing/coqast.cmx kernel/entries.cmx \ + library/lib.cmx library/libnames.cmx kernel/modops.cmx kernel/names.cmx \ + library/nametab.cmx lib/pp.cmx lib/util.cmx parsing/astmod.cmi parsing/astterm.cmo: parsing/ast.cmi parsing/coqast.cmi library/declare.cmi \ lib/dyn.cmi kernel/environ.cmi pretyping/evarutil.cmi pretyping/evd.cmi \ - library/global.cmi library/impargs.cmi library/nameops.cmi \ - kernel/names.cmi library/nametab.cmi lib/options.cmi \ + library/global.cmi library/impargs.cmi library/libnames.cmi \ + library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/options.cmi \ pretyping/pattern.cmi lib/pp.cmi pretyping/pretype_errors.cmi \ pretyping/pretyping.cmi pretyping/rawterm.cmi pretyping/reductionops.cmi \ pretyping/retyping.cmi kernel/sign.cmi pretyping/syntax_def.cmi \ @@ -595,21 +685,23 @@ parsing/astterm.cmo: parsing/ast.cmi parsing/coqast.cmi library/declare.cmi \ pretyping/typing.cmi lib/util.cmi parsing/astterm.cmi parsing/astterm.cmx: parsing/ast.cmx parsing/coqast.cmx library/declare.cmx \ lib/dyn.cmx kernel/environ.cmx pretyping/evarutil.cmx pretyping/evd.cmx \ - library/global.cmx library/impargs.cmx library/nameops.cmx \ - kernel/names.cmx library/nametab.cmx lib/options.cmx \ + library/global.cmx library/impargs.cmx library/libnames.cmx \ + library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/options.cmx \ pretyping/pattern.cmx lib/pp.cmx pretyping/pretype_errors.cmx \ pretyping/pretyping.cmx pretyping/rawterm.cmx pretyping/reductionops.cmx \ pretyping/retyping.cmx kernel/sign.cmx pretyping/syntax_def.cmx \ kernel/term.cmx parsing/termast.cmx pretyping/termops.cmx \ pretyping/typing.cmx lib/util.cmx parsing/astterm.cmi -parsing/coqast.cmo: lib/dyn.cmi lib/hashcons.cmi kernel/names.cmi \ - parsing/coqast.cmi -parsing/coqast.cmx: lib/dyn.cmx lib/hashcons.cmx kernel/names.cmx \ - parsing/coqast.cmi -parsing/coqlib.cmo: library/declare.cmi kernel/names.cmi library/nametab.cmi \ - pretyping/pattern.cmi kernel/term.cmi lib/util.cmi parsing/coqlib.cmi -parsing/coqlib.cmx: library/declare.cmx kernel/names.cmx library/nametab.cmx \ - pretyping/pattern.cmx kernel/term.cmx lib/util.cmx parsing/coqlib.cmi +parsing/coqast.cmo: lib/dyn.cmi lib/hashcons.cmi library/libnames.cmi \ + kernel/names.cmi lib/util.cmi parsing/coqast.cmi +parsing/coqast.cmx: lib/dyn.cmx lib/hashcons.cmx library/libnames.cmx \ + kernel/names.cmx lib/util.cmx parsing/coqast.cmi +parsing/coqlib.cmo: library/declare.cmi library/libnames.cmi kernel/names.cmi \ + library/nametab.cmi pretyping/pattern.cmi kernel/term.cmi lib/util.cmi \ + parsing/coqlib.cmi +parsing/coqlib.cmx: library/declare.cmx library/libnames.cmx kernel/names.cmx \ + library/nametab.cmx pretyping/pattern.cmx kernel/term.cmx lib/util.cmx \ + parsing/coqlib.cmi parsing/egrammar.cmo: parsing/ast.cmi parsing/coqast.cmi parsing/extend.cmi \ parsing/genarg.cmi parsing/lexer.cmi parsing/pcoq.cmi lib/pp.cmi \ proofs/tacexpr.cmo lib/util.cmi toplevel/vernacexpr.cmo \ @@ -619,18 +711,18 @@ parsing/egrammar.cmx: parsing/ast.cmx parsing/coqast.cmx parsing/extend.cmx \ proofs/tacexpr.cmx lib/util.cmx toplevel/vernacexpr.cmx \ parsing/egrammar.cmi parsing/esyntax.cmo: parsing/ast.cmi parsing/coqast.cmi parsing/extend.cmi \ - parsing/genarg.cmi library/global.cmi lib/gmap.cmi lib/gmapl.cmi \ + parsing/genarg.cmi lib/gmap.cmi lib/gmapl.cmi library/libnames.cmi \ kernel/names.cmi library/nametab.cmi lib/options.cmi lib/pp.cmi \ lib/util.cmi toplevel/vernacexpr.cmo parsing/esyntax.cmi parsing/esyntax.cmx: parsing/ast.cmx parsing/coqast.cmx parsing/extend.cmx \ - parsing/genarg.cmx library/global.cmx lib/gmap.cmx lib/gmapl.cmx \ + parsing/genarg.cmx lib/gmap.cmx lib/gmapl.cmx library/libnames.cmx \ kernel/names.cmx library/nametab.cmx lib/options.cmx lib/pp.cmx \ lib/util.cmx toplevel/vernacexpr.cmx parsing/esyntax.cmi parsing/extend.cmo: parsing/ast.cmi parsing/coqast.cmi parsing/genarg.cmi \ - parsing/lexer.cmi library/nametab.cmi lib/pp.cmi lib/util.cmi \ + parsing/lexer.cmi library/libnames.cmi lib/pp.cmi lib/util.cmi \ parsing/extend.cmi parsing/extend.cmx: parsing/ast.cmx parsing/coqast.cmx parsing/genarg.cmx \ - parsing/lexer.cmx library/nametab.cmx lib/pp.cmx lib/util.cmx \ + parsing/lexer.cmx library/libnames.cmx lib/pp.cmx lib/util.cmx \ parsing/extend.cmi parsing/g_basevernac.cmo: parsing/ast.cmi parsing/coqast.cmi \ parsing/extend.cmi library/goptions.cmi parsing/pcoq.cmi lib/pp.cmi \ @@ -638,10 +730,10 @@ parsing/g_basevernac.cmo: parsing/ast.cmi parsing/coqast.cmi \ parsing/g_basevernac.cmx: parsing/ast.cmx parsing/coqast.cmx \ parsing/extend.cmx library/goptions.cmx parsing/pcoq.cmx lib/pp.cmx \ toplevel/vernacexpr.cmx -parsing/g_cases.cmo: parsing/coqast.cmi kernel/names.cmi parsing/pcoq.cmi \ - lib/pp.cmi lib/util.cmi -parsing/g_cases.cmx: parsing/coqast.cmx kernel/names.cmx parsing/pcoq.cmx \ - lib/pp.cmx lib/util.cmx +parsing/g_cases.cmo: parsing/coqast.cmi library/libnames.cmi kernel/names.cmi \ + parsing/pcoq.cmi lib/pp.cmi lib/util.cmi +parsing/g_cases.cmx: parsing/coqast.cmx library/libnames.cmx kernel/names.cmx \ + parsing/pcoq.cmx lib/pp.cmx lib/util.cmx parsing/g_constr.cmo: parsing/ast.cmi parsing/coqast.cmi kernel/names.cmi \ parsing/pcoq.cmi parsing/g_constr.cmx: parsing/ast.cmx parsing/coqast.cmx kernel/names.cmx \ @@ -664,6 +756,10 @@ parsing/g_minicoq.cmo: kernel/environ.cmi parsing/lexer.cmi kernel/names.cmi \ parsing/g_minicoq.cmx: kernel/environ.cmx parsing/lexer.cmx kernel/names.cmx \ lib/pp.cmx kernel/term.cmx kernel/univ.cmx lib/util.cmx \ parsing/g_minicoq.cmi +parsing/g_module.cmo: parsing/ast.cmi parsing/coqast.cmi kernel/names.cmi \ + parsing/pcoq.cmi lib/pp.cmi lib/util.cmi +parsing/g_module.cmx: parsing/ast.cmx parsing/coqast.cmx kernel/names.cmx \ + parsing/pcoq.cmx lib/pp.cmx lib/util.cmx parsing/g_natsyntax.cmo: parsing/ast.cmi parsing/coqast.cmi \ parsing/coqlib.cmi parsing/esyntax.cmi parsing/extend.cmi \ kernel/names.cmi parsing/pcoq.cmi lib/pp.cmi parsing/termast.cmi \ @@ -672,10 +768,10 @@ parsing/g_natsyntax.cmx: parsing/ast.cmx parsing/coqast.cmx \ parsing/coqlib.cmx parsing/esyntax.cmx parsing/extend.cmx \ kernel/names.cmx parsing/pcoq.cmx lib/pp.cmx parsing/termast.cmx \ lib/util.cmx parsing/g_natsyntax.cmi -parsing/g_prim.cmo: parsing/coqast.cmi kernel/names.cmi library/nametab.cmi \ - parsing/pcoq.cmi proofs/tacexpr.cmo -parsing/g_prim.cmx: parsing/coqast.cmx kernel/names.cmx library/nametab.cmx \ - parsing/pcoq.cmx proofs/tacexpr.cmx +parsing/g_prim.cmo: parsing/coqast.cmi library/libnames.cmi kernel/names.cmi \ + library/nametab.cmi parsing/pcoq.cmi proofs/tacexpr.cmo +parsing/g_prim.cmx: parsing/coqast.cmx library/libnames.cmx kernel/names.cmx \ + library/nametab.cmx parsing/pcoq.cmx proofs/tacexpr.cmx parsing/g_proofs.cmo: parsing/coqast.cmi parsing/genarg.cmi parsing/pcoq.cmi \ lib/pp.cmi proofs/tacexpr.cmo lib/util.cmi toplevel/vernacexpr.cmo parsing/g_proofs.cmx: parsing/coqast.cmx parsing/genarg.cmx parsing/pcoq.cmx \ @@ -694,11 +790,11 @@ parsing/g_tactic.cmx: parsing/ast.cmx parsing/genarg.cmx kernel/names.cmx \ lib/util.cmx parsing/g_vernac.cmo: parsing/ast.cmi toplevel/class.cmi parsing/coqast.cmi \ library/declare.cmi parsing/genarg.cmi library/goptions.cmi \ - library/nametab.cmi lib/options.cmi parsing/pcoq.cmi lib/pp.cmi \ + library/libnames.cmi lib/options.cmi parsing/pcoq.cmi lib/pp.cmi \ toplevel/recordobj.cmi lib/util.cmi toplevel/vernacexpr.cmo parsing/g_vernac.cmx: parsing/ast.cmx toplevel/class.cmx parsing/coqast.cmx \ library/declare.cmx parsing/genarg.cmx library/goptions.cmx \ - library/nametab.cmx lib/options.cmx parsing/pcoq.cmx lib/pp.cmx \ + library/libnames.cmx lib/options.cmx parsing/pcoq.cmx lib/pp.cmx \ toplevel/recordobj.cmx lib/util.cmx toplevel/vernacexpr.cmx parsing/g_zsyntax.cmo: parsing/ast.cmi parsing/astterm.cmi parsing/coqast.cmi \ parsing/esyntax.cmi parsing/extend.cmi kernel/names.cmi parsing/pcoq.cmi \ @@ -716,76 +812,78 @@ parsing/pcoq.cmx: parsing/ast.cmx parsing/coqast.cmx parsing/genarg.cmx \ lib/util.cmx toplevel/vernacexpr.cmx parsing/pcoq.cmi parsing/ppconstr.cmo: parsing/ast.cmi parsing/coqast.cmi lib/dyn.cmi \ parsing/esyntax.cmi parsing/extend.cmi parsing/genarg.cmi \ - library/global.cmi library/nameops.cmi kernel/names.cmi \ + library/libnames.cmi library/nameops.cmi kernel/names.cmi \ library/nametab.cmi lib/pp.cmi pretyping/rawterm.cmi parsing/termast.cmi \ lib/util.cmi parsing/ppconstr.cmi parsing/ppconstr.cmx: parsing/ast.cmx parsing/coqast.cmx lib/dyn.cmx \ parsing/esyntax.cmx parsing/extend.cmx parsing/genarg.cmx \ - library/global.cmx library/nameops.cmx kernel/names.cmx \ + library/libnames.cmx library/nameops.cmx kernel/names.cmx \ library/nametab.cmx lib/pp.cmx pretyping/rawterm.cmx parsing/termast.cmx \ lib/util.cmx parsing/ppconstr.cmi parsing/pptactic.cmo: kernel/closure.cmi lib/dyn.cmi parsing/egrammar.cmi \ - parsing/extend.cmi parsing/genarg.cmi library/nameops.cmi \ - kernel/names.cmi library/nametab.cmi lib/pp.cmi parsing/ppconstr.cmi \ + parsing/extend.cmi parsing/genarg.cmi library/libnames.cmi \ + library/nameops.cmi kernel/names.cmi lib/pp.cmi parsing/ppconstr.cmi \ parsing/printer.cmi pretyping/rawterm.cmi proofs/tacexpr.cmo lib/util.cmi \ parsing/pptactic.cmi parsing/pptactic.cmx: kernel/closure.cmx lib/dyn.cmx parsing/egrammar.cmx \ - parsing/extend.cmx parsing/genarg.cmx library/nameops.cmx \ - kernel/names.cmx library/nametab.cmx lib/pp.cmx parsing/ppconstr.cmx \ + parsing/extend.cmx parsing/genarg.cmx library/libnames.cmx \ + library/nameops.cmx kernel/names.cmx lib/pp.cmx parsing/ppconstr.cmx \ parsing/printer.cmx pretyping/rawterm.cmx proofs/tacexpr.cmx lib/util.cmx \ parsing/pptactic.cmi parsing/prettyp.cmo: pretyping/classops.cmi kernel/declarations.cmi \ library/declare.cmi kernel/environ.cmi pretyping/evd.cmi \ library/global.cmi library/impargs.cmi kernel/inductive.cmi \ pretyping/inductiveops.cmi pretyping/instantiate.cmi library/lib.cmi \ - library/libobject.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/pp.cmi parsing/printer.cmi kernel/reduction.cmi \ - kernel/safe_typing.cmi kernel/sign.cmi pretyping/syntax_def.cmi \ - kernel/term.cmi pretyping/termops.cmi lib/util.cmi parsing/prettyp.cmi + library/libnames.cmi library/libobject.cmi library/nameops.cmi \ + kernel/names.cmi library/nametab.cmi lib/pp.cmi parsing/printer.cmi \ + kernel/reduction.cmi kernel/safe_typing.cmi kernel/sign.cmi \ + pretyping/syntax_def.cmi kernel/term.cmi pretyping/termops.cmi \ + lib/util.cmi parsing/prettyp.cmi parsing/prettyp.cmx: pretyping/classops.cmx kernel/declarations.cmx \ library/declare.cmx kernel/environ.cmx pretyping/evd.cmx \ library/global.cmx library/impargs.cmx kernel/inductive.cmx \ pretyping/inductiveops.cmx pretyping/instantiate.cmx library/lib.cmx \ - library/libobject.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/pp.cmx parsing/printer.cmx kernel/reduction.cmx \ - kernel/safe_typing.cmx kernel/sign.cmx pretyping/syntax_def.cmx \ - kernel/term.cmx pretyping/termops.cmx lib/util.cmx parsing/prettyp.cmi + library/libnames.cmx library/libobject.cmx library/nameops.cmx \ + kernel/names.cmx library/nametab.cmx lib/pp.cmx parsing/printer.cmx \ + kernel/reduction.cmx kernel/safe_typing.cmx kernel/sign.cmx \ + pretyping/syntax_def.cmx kernel/term.cmx pretyping/termops.cmx \ + lib/util.cmx parsing/prettyp.cmi parsing/printer.cmo: parsing/ast.cmi parsing/coqast.cmi library/declare.cmi \ kernel/environ.cmi parsing/extend.cmi library/global.cmi \ - library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/options.cmi \ - pretyping/pattern.cmi lib/pp.cmi parsing/ppconstr.cmi kernel/sign.cmi \ - kernel/term.cmi parsing/termast.cmi pretyping/termops.cmi lib/util.cmi \ - parsing/printer.cmi + library/libnames.cmi library/nameops.cmi kernel/names.cmi \ + library/nametab.cmi lib/options.cmi pretyping/pattern.cmi lib/pp.cmi \ + parsing/ppconstr.cmi kernel/sign.cmi kernel/term.cmi parsing/termast.cmi \ + pretyping/termops.cmi lib/util.cmi parsing/printer.cmi parsing/printer.cmx: parsing/ast.cmx parsing/coqast.cmx library/declare.cmx \ kernel/environ.cmx parsing/extend.cmx library/global.cmx \ - library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/options.cmx \ - pretyping/pattern.cmx lib/pp.cmx parsing/ppconstr.cmx kernel/sign.cmx \ - kernel/term.cmx parsing/termast.cmx pretyping/termops.cmx lib/util.cmx \ - parsing/printer.cmi -parsing/q_coqast.cmo: parsing/coqast.cmi parsing/genarg.cmi kernel/names.cmi \ - library/nametab.cmi parsing/pcoq.cmi parsing/q_util.cmi \ + library/libnames.cmx library/nameops.cmx kernel/names.cmx \ + library/nametab.cmx lib/options.cmx pretyping/pattern.cmx lib/pp.cmx \ + parsing/ppconstr.cmx kernel/sign.cmx kernel/term.cmx parsing/termast.cmx \ + pretyping/termops.cmx lib/util.cmx parsing/printer.cmi +parsing/q_coqast.cmo: parsing/coqast.cmi parsing/genarg.cmi \ + library/libnames.cmi kernel/names.cmi parsing/pcoq.cmi parsing/q_util.cmi \ pretyping/rawterm.cmi proofs/tacexpr.cmo -parsing/q_coqast.cmx: parsing/coqast.cmx parsing/genarg.cmx kernel/names.cmx \ - library/nametab.cmx parsing/pcoq.cmx parsing/q_util.cmx \ +parsing/q_coqast.cmx: parsing/coqast.cmx parsing/genarg.cmx \ + library/libnames.cmx kernel/names.cmx parsing/pcoq.cmx parsing/q_util.cmx \ pretyping/rawterm.cmx proofs/tacexpr.cmx parsing/q_util.cmo: parsing/q_util.cmi parsing/q_util.cmx: parsing/q_util.cmi parsing/search.cmo: parsing/astterm.cmi parsing/coqast.cmi parsing/coqlib.cmi \ - kernel/declarations.cmi library/declare.cmi kernel/environ.cmi \ - pretyping/evd.cmi library/global.cmi library/libobject.cmi \ - library/library.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi pretyping/pattern.cmi lib/pp.cmi \ - pretyping/pretyping.cmi parsing/printer.cmi pretyping/rawterm.cmi \ - pretyping/retyping.cmi kernel/term.cmi pretyping/typing.cmi lib/util.cmi \ - parsing/search.cmi + kernel/declarations.cmi library/declare.cmi library/declaremods.cmi \ + kernel/environ.cmi pretyping/evd.cmi library/global.cmi \ + library/libnames.cmi library/libobject.cmi library/library.cmi \ + library/nameops.cmi kernel/names.cmi library/nametab.cmi \ + pretyping/pattern.cmi lib/pp.cmi pretyping/pretyping.cmi \ + parsing/printer.cmi pretyping/rawterm.cmi pretyping/retyping.cmi \ + kernel/term.cmi pretyping/typing.cmi lib/util.cmi parsing/search.cmi parsing/search.cmx: parsing/astterm.cmx parsing/coqast.cmx parsing/coqlib.cmx \ - kernel/declarations.cmx library/declare.cmx kernel/environ.cmx \ - pretyping/evd.cmx library/global.cmx library/libobject.cmx \ - library/library.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx pretyping/pattern.cmx lib/pp.cmx \ - pretyping/pretyping.cmx parsing/printer.cmx pretyping/rawterm.cmx \ - pretyping/retyping.cmx kernel/term.cmx pretyping/typing.cmx lib/util.cmx \ - parsing/search.cmi + kernel/declarations.cmx library/declare.cmx library/declaremods.cmx \ + kernel/environ.cmx pretyping/evd.cmx library/global.cmx \ + library/libnames.cmx library/libobject.cmx library/library.cmx \ + library/nameops.cmx kernel/names.cmx library/nametab.cmx \ + pretyping/pattern.cmx lib/pp.cmx pretyping/pretyping.cmx \ + parsing/printer.cmx pretyping/rawterm.cmx pretyping/retyping.cmx \ + kernel/term.cmx pretyping/typing.cmx lib/util.cmx parsing/search.cmi parsing/tacextend.cmo: parsing/ast.cmi parsing/genarg.cmi parsing/pcoq.cmi \ lib/pp.cmi lib/pp_control.cmi parsing/q_coqast.cmo parsing/q_util.cmi \ lib/util.cmi toplevel/vernacexpr.cmo @@ -795,17 +893,19 @@ parsing/tacextend.cmx: parsing/ast.cmx parsing/genarg.cmx parsing/pcoq.cmx \ parsing/termast.cmo: parsing/ast.cmi pretyping/classops.cmi \ parsing/coqast.cmi library/declare.cmi pretyping/detyping.cmi \ kernel/environ.cmi library/impargs.cmi kernel/inductive.cmi \ - library/nameops.cmi kernel/names.cmi library/nametab.cmi \ - pretyping/pattern.cmi lib/pp.cmi pretyping/rawterm.cmi \ - pretyping/reductionops.cmi kernel/sign.cmi kernel/term.cmi \ - pretyping/termops.cmi kernel/univ.cmi lib/util.cmi parsing/termast.cmi + library/libnames.cmi library/nameops.cmi kernel/names.cmi \ + library/nametab.cmi pretyping/pattern.cmi lib/pp.cmi \ + pretyping/rawterm.cmi pretyping/reductionops.cmi kernel/sign.cmi \ + kernel/term.cmi pretyping/termops.cmi kernel/univ.cmi lib/util.cmi \ + parsing/termast.cmi parsing/termast.cmx: parsing/ast.cmx pretyping/classops.cmx \ parsing/coqast.cmx library/declare.cmx pretyping/detyping.cmx \ kernel/environ.cmx library/impargs.cmx kernel/inductive.cmx \ - library/nameops.cmx kernel/names.cmx library/nametab.cmx \ - pretyping/pattern.cmx lib/pp.cmx pretyping/rawterm.cmx \ - pretyping/reductionops.cmx kernel/sign.cmx kernel/term.cmx \ - pretyping/termops.cmx kernel/univ.cmx lib/util.cmx parsing/termast.cmi + library/libnames.cmx library/nameops.cmx kernel/names.cmx \ + library/nametab.cmx pretyping/pattern.cmx lib/pp.cmx \ + pretyping/rawterm.cmx pretyping/reductionops.cmx kernel/sign.cmx \ + kernel/term.cmx pretyping/termops.cmx kernel/univ.cmx lib/util.cmx \ + parsing/termast.cmi parsing/vernacextend.cmo: parsing/ast.cmi parsing/genarg.cmi parsing/pcoq.cmi \ lib/pp.cmi lib/pp_control.cmi parsing/q_coqast.cmo parsing/q_util.cmi \ lib/util.cmi toplevel/vernacexpr.cmo @@ -836,16 +936,18 @@ pretyping/cbv.cmx: kernel/closure.cmx kernel/environ.cmx kernel/esubst.cmx \ kernel/term.cmx kernel/univ.cmx lib/util.cmx pretyping/cbv.cmi pretyping/classops.cmo: library/declare.cmi kernel/environ.cmi \ pretyping/evd.cmi library/global.cmi library/goptions.cmi library/lib.cmi \ - library/libobject.cmi library/library.cmi kernel/names.cmi \ - library/nametab.cmi lib/options.cmi lib/pp.cmi pretyping/rawterm.cmi \ - pretyping/reductionops.cmi library/summary.cmi pretyping/tacred.cmi \ - kernel/term.cmi pretyping/termops.cmi lib/util.cmi pretyping/classops.cmi + library/libnames.cmi library/libobject.cmi library/library.cmi \ + kernel/names.cmi library/nametab.cmi lib/options.cmi lib/pp.cmi \ + pretyping/rawterm.cmi pretyping/reductionops.cmi library/summary.cmi \ + pretyping/tacred.cmi kernel/term.cmi pretyping/termops.cmi lib/util.cmi \ + pretyping/classops.cmi pretyping/classops.cmx: library/declare.cmx kernel/environ.cmx \ pretyping/evd.cmx library/global.cmx library/goptions.cmx library/lib.cmx \ - library/libobject.cmx library/library.cmx kernel/names.cmx \ - library/nametab.cmx lib/options.cmx lib/pp.cmx pretyping/rawterm.cmx \ - pretyping/reductionops.cmx library/summary.cmx pretyping/tacred.cmx \ - kernel/term.cmx pretyping/termops.cmx lib/util.cmx pretyping/classops.cmi + library/libnames.cmx library/libobject.cmx library/library.cmx \ + kernel/names.cmx library/nametab.cmx lib/options.cmx lib/pp.cmx \ + pretyping/rawterm.cmx pretyping/reductionops.cmx library/summary.cmx \ + pretyping/tacred.cmx kernel/term.cmx pretyping/termops.cmx lib/util.cmx \ + pretyping/classops.cmi pretyping/coercion.cmo: pretyping/classops.cmi kernel/environ.cmi \ pretyping/evarconv.cmi pretyping/evarutil.cmi pretyping/evd.cmi \ kernel/names.cmi pretyping/pretype_errors.cmi pretyping/rawterm.cmi \ @@ -859,15 +961,17 @@ pretyping/coercion.cmx: pretyping/classops.cmx kernel/environ.cmx \ pretyping/detyping.cmo: kernel/declarations.cmi library/declare.cmi \ kernel/environ.cmi library/global.cmi library/goptions.cmi \ library/impargs.cmi kernel/inductive.cmi pretyping/inductiveops.cmi \ - library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/pp.cmi \ - pretyping/rawterm.cmi kernel/sign.cmi kernel/term.cmi \ - pretyping/termops.cmi kernel/univ.cmi lib/util.cmi pretyping/detyping.cmi + library/libnames.cmi library/nameops.cmi kernel/names.cmi \ + library/nametab.cmi lib/pp.cmi pretyping/rawterm.cmi kernel/sign.cmi \ + kernel/term.cmi pretyping/termops.cmi kernel/univ.cmi lib/util.cmi \ + pretyping/detyping.cmi pretyping/detyping.cmx: kernel/declarations.cmx library/declare.cmx \ kernel/environ.cmx library/global.cmx library/goptions.cmx \ library/impargs.cmx kernel/inductive.cmx pretyping/inductiveops.cmx \ - library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/pp.cmx \ - pretyping/rawterm.cmx kernel/sign.cmx kernel/term.cmx \ - pretyping/termops.cmx kernel/univ.cmx lib/util.cmx pretyping/detyping.cmi + library/libnames.cmx library/nameops.cmx kernel/names.cmx \ + library/nametab.cmx lib/pp.cmx pretyping/rawterm.cmx kernel/sign.cmx \ + kernel/term.cmx pretyping/termops.cmx kernel/univ.cmx lib/util.cmx \ + pretyping/detyping.cmi pretyping/evarconv.cmo: pretyping/classops.cmi kernel/closure.cmi \ library/declare.cmi kernel/environ.cmi pretyping/evarutil.cmi \ pretyping/evd.cmi pretyping/instantiate.cmi kernel/names.cmi \ @@ -897,21 +1001,23 @@ pretyping/evd.cmo: kernel/names.cmi kernel/sign.cmi kernel/term.cmi \ pretyping/evd.cmx: kernel/names.cmx kernel/sign.cmx kernel/term.cmx \ lib/util.cmx pretyping/evd.cmi pretyping/indrec.cmo: kernel/declarations.cmi library/declare.cmi \ - kernel/environ.cmi pretyping/evd.cmi library/global.cmi \ - kernel/indtypes.cmi kernel/inductive.cmi pretyping/inductiveops.cmi \ - pretyping/instantiate.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/options.cmi lib/pp.cmi kernel/reduction.cmi \ - pretyping/reductionops.cmi kernel/safe_typing.cmi kernel/term.cmi \ - pretyping/termops.cmi kernel/type_errors.cmi kernel/typeops.cmi \ - lib/util.cmi pretyping/indrec.cmi + kernel/entries.cmi kernel/environ.cmi pretyping/evd.cmi \ + library/global.cmi kernel/indtypes.cmi kernel/inductive.cmi \ + pretyping/inductiveops.cmi pretyping/instantiate.cmi library/libnames.cmi \ + library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/options.cmi \ + lib/pp.cmi kernel/reduction.cmi pretyping/reductionops.cmi \ + kernel/safe_typing.cmi kernel/term.cmi pretyping/termops.cmi \ + kernel/type_errors.cmi kernel/typeops.cmi lib/util.cmi \ + pretyping/indrec.cmi pretyping/indrec.cmx: kernel/declarations.cmx library/declare.cmx \ - kernel/environ.cmx pretyping/evd.cmx library/global.cmx \ - kernel/indtypes.cmx kernel/inductive.cmx pretyping/inductiveops.cmx \ - pretyping/instantiate.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/options.cmx lib/pp.cmx kernel/reduction.cmx \ - pretyping/reductionops.cmx kernel/safe_typing.cmx kernel/term.cmx \ - pretyping/termops.cmx kernel/type_errors.cmx kernel/typeops.cmx \ - lib/util.cmx pretyping/indrec.cmi + kernel/entries.cmx kernel/environ.cmx pretyping/evd.cmx \ + library/global.cmx kernel/indtypes.cmx kernel/inductive.cmx \ + pretyping/inductiveops.cmx pretyping/instantiate.cmx library/libnames.cmx \ + library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/options.cmx \ + lib/pp.cmx kernel/reduction.cmx pretyping/reductionops.cmx \ + kernel/safe_typing.cmx kernel/term.cmx pretyping/termops.cmx \ + kernel/type_errors.cmx kernel/typeops.cmx lib/util.cmx \ + pretyping/indrec.cmi pretyping/inductiveops.cmo: kernel/declarations.cmi kernel/environ.cmi \ pretyping/evd.cmi library/global.cmi kernel/inductive.cmi \ kernel/names.cmi pretyping/reductionops.cmi kernel/sign.cmi \ @@ -929,13 +1035,13 @@ pretyping/instantiate.cmx: kernel/declarations.cmx kernel/environ.cmx \ pretyping/evd.cmx kernel/names.cmx lib/pp.cmx kernel/sign.cmx \ kernel/term.cmx lib/util.cmx pretyping/instantiate.cmi pretyping/pattern.cmo: library/declare.cmi kernel/environ.cmi \ - library/nameops.cmi kernel/names.cmi library/nametab.cmi \ - pretyping/rawterm.cmi pretyping/reductionops.cmi kernel/term.cmi \ - pretyping/termops.cmi lib/util.cmi pretyping/pattern.cmi + library/libnames.cmi library/nameops.cmi kernel/names.cmi \ + library/nametab.cmi pretyping/rawterm.cmi pretyping/reductionops.cmi \ + kernel/term.cmi pretyping/termops.cmi lib/util.cmi pretyping/pattern.cmi pretyping/pattern.cmx: library/declare.cmx kernel/environ.cmx \ - library/nameops.cmx kernel/names.cmx library/nametab.cmx \ - pretyping/rawterm.cmx pretyping/reductionops.cmx kernel/term.cmx \ - pretyping/termops.cmx lib/util.cmx pretyping/pattern.cmi + library/libnames.cmx library/nameops.cmx kernel/names.cmx \ + library/nametab.cmx pretyping/rawterm.cmx pretyping/reductionops.cmx \ + kernel/term.cmx pretyping/termops.cmx lib/util.cmx pretyping/pattern.cmi pretyping/pretype_errors.cmo: kernel/environ.cmi pretyping/evd.cmi \ pretyping/inductiveops.cmi kernel/names.cmi pretyping/rawterm.cmi \ kernel/reduction.cmi pretyping/reductionops.cmi kernel/sign.cmi \ @@ -949,40 +1055,40 @@ pretyping/pretype_errors.cmx: kernel/environ.cmx pretyping/evd.cmx \ pretyping/pretyping.cmo: pretyping/cases.cmi pretyping/classops.cmi \ pretyping/coercion.cmi kernel/declarations.cmi library/declare.cmi \ lib/dyn.cmi kernel/environ.cmi pretyping/evarconv.cmi \ - pretyping/evarutil.cmi pretyping/evd.cmi library/global.cmi \ - pretyping/indrec.cmi kernel/inductive.cmi pretyping/inductiveops.cmi \ - pretyping/instantiate.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/pp.cmi pretyping/pretype_errors.cmi \ - pretyping/rawterm.cmi pretyping/recordops.cmi pretyping/reductionops.cmi \ - pretyping/retyping.cmi kernel/sign.cmi kernel/term.cmi \ - pretyping/termops.cmi kernel/type_errors.cmi kernel/typeops.cmi \ - lib/util.cmi pretyping/pretyping.cmi + pretyping/evarutil.cmi pretyping/evd.cmi pretyping/indrec.cmi \ + kernel/inductive.cmi pretyping/inductiveops.cmi pretyping/instantiate.cmi \ + library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/pp.cmi \ + pretyping/pretype_errors.cmi pretyping/rawterm.cmi \ + pretyping/recordops.cmi pretyping/reductionops.cmi pretyping/retyping.cmi \ + kernel/sign.cmi kernel/term.cmi pretyping/termops.cmi \ + kernel/type_errors.cmi kernel/typeops.cmi lib/util.cmi \ + pretyping/pretyping.cmi pretyping/pretyping.cmx: pretyping/cases.cmx pretyping/classops.cmx \ pretyping/coercion.cmx kernel/declarations.cmx library/declare.cmx \ lib/dyn.cmx kernel/environ.cmx pretyping/evarconv.cmx \ - pretyping/evarutil.cmx pretyping/evd.cmx library/global.cmx \ - pretyping/indrec.cmx kernel/inductive.cmx pretyping/inductiveops.cmx \ - pretyping/instantiate.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/pp.cmx pretyping/pretype_errors.cmx \ - pretyping/rawterm.cmx pretyping/recordops.cmx pretyping/reductionops.cmx \ - pretyping/retyping.cmx kernel/sign.cmx kernel/term.cmx \ - pretyping/termops.cmx kernel/type_errors.cmx kernel/typeops.cmx \ - lib/util.cmx pretyping/pretyping.cmi -pretyping/rawterm.cmo: lib/dyn.cmi kernel/names.cmi library/nametab.cmi \ - kernel/sign.cmi kernel/term.cmi kernel/univ.cmi lib/util.cmi \ - pretyping/rawterm.cmi -pretyping/rawterm.cmx: lib/dyn.cmx kernel/names.cmx library/nametab.cmx \ - kernel/sign.cmx kernel/term.cmx kernel/univ.cmx lib/util.cmx \ - pretyping/rawterm.cmi + pretyping/evarutil.cmx pretyping/evd.cmx pretyping/indrec.cmx \ + kernel/inductive.cmx pretyping/inductiveops.cmx pretyping/instantiate.cmx \ + library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/pp.cmx \ + pretyping/pretype_errors.cmx pretyping/rawterm.cmx \ + pretyping/recordops.cmx pretyping/reductionops.cmx pretyping/retyping.cmx \ + kernel/sign.cmx kernel/term.cmx pretyping/termops.cmx \ + kernel/type_errors.cmx kernel/typeops.cmx lib/util.cmx \ + pretyping/pretyping.cmi +pretyping/rawterm.cmo: lib/dyn.cmi library/libnames.cmi kernel/names.cmi \ + library/nametab.cmi kernel/sign.cmi kernel/term.cmi kernel/univ.cmi \ + lib/util.cmi pretyping/rawterm.cmi +pretyping/rawterm.cmx: lib/dyn.cmx library/libnames.cmx kernel/names.cmx \ + library/nametab.cmx kernel/sign.cmx kernel/term.cmx kernel/univ.cmx \ + lib/util.cmx pretyping/rawterm.cmi pretyping/recordops.cmo: pretyping/classops.cmi library/lib.cmi \ - library/libobject.cmi library/library.cmi kernel/names.cmi \ - library/nametab.cmi lib/pp.cmi library/summary.cmi kernel/term.cmi \ - pretyping/termops.cmi kernel/typeops.cmi lib/util.cmi \ + library/libnames.cmi library/libobject.cmi library/library.cmi \ + kernel/names.cmi library/nametab.cmi lib/pp.cmi library/summary.cmi \ + kernel/term.cmi pretyping/termops.cmi kernel/typeops.cmi lib/util.cmi \ pretyping/recordops.cmi pretyping/recordops.cmx: pretyping/classops.cmx library/lib.cmx \ - library/libobject.cmx library/library.cmx kernel/names.cmx \ - library/nametab.cmx lib/pp.cmx library/summary.cmx kernel/term.cmx \ - pretyping/termops.cmx kernel/typeops.cmx lib/util.cmx \ + library/libnames.cmx library/libobject.cmx library/library.cmx \ + kernel/names.cmx library/nametab.cmx lib/pp.cmx library/summary.cmx \ + kernel/term.cmx pretyping/termops.cmx kernel/typeops.cmx lib/util.cmx \ pretyping/recordops.cmi pretyping/reductionops.cmo: kernel/closure.cmi kernel/declarations.cmi \ kernel/environ.cmi kernel/esubst.cmi pretyping/evd.cmi \ @@ -1004,34 +1110,36 @@ pretyping/retyping.cmx: kernel/declarations.cmx kernel/environ.cmx \ kernel/inductive.cmx pretyping/instantiate.cmx kernel/names.cmx \ pretyping/reductionops.cmx kernel/term.cmx kernel/typeops.cmx \ kernel/univ.cmx lib/util.cmx pretyping/retyping.cmi -pretyping/syntax_def.cmo: library/lib.cmi library/libobject.cmi \ - library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/pp.cmi \ - pretyping/rawterm.cmi library/summary.cmi lib/util.cmi \ - pretyping/syntax_def.cmi -pretyping/syntax_def.cmx: library/lib.cmx library/libobject.cmx \ - library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/pp.cmx \ - pretyping/rawterm.cmx library/summary.cmx lib/util.cmx \ - pretyping/syntax_def.cmi +pretyping/syntax_def.cmo: library/lib.cmi library/libnames.cmi \ + library/libobject.cmi library/nameops.cmi kernel/names.cmi \ + library/nametab.cmi lib/pp.cmi pretyping/rawterm.cmi library/summary.cmi \ + lib/util.cmi pretyping/syntax_def.cmi +pretyping/syntax_def.cmx: library/lib.cmx library/libnames.cmx \ + library/libobject.cmx library/nameops.cmx kernel/names.cmx \ + library/nametab.cmx lib/pp.cmx pretyping/rawterm.cmx library/summary.cmx \ + lib/util.cmx pretyping/syntax_def.cmi pretyping/tacred.cmo: pretyping/cbv.cmi kernel/closure.cmi \ kernel/conv_oracle.cmi kernel/declarations.cmi kernel/environ.cmi \ pretyping/evd.cmi library/global.cmi kernel/inductive.cmi \ - pretyping/instantiate.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/pp.cmi pretyping/rawterm.cmi \ + pretyping/instantiate.cmi library/libnames.cmi library/nameops.cmi \ + kernel/names.cmi library/nametab.cmi lib/pp.cmi pretyping/rawterm.cmi \ pretyping/reductionops.cmi pretyping/retyping.cmi library/summary.cmi \ kernel/term.cmi pretyping/termops.cmi lib/util.cmi pretyping/tacred.cmi pretyping/tacred.cmx: pretyping/cbv.cmx kernel/closure.cmx \ kernel/conv_oracle.cmx kernel/declarations.cmx kernel/environ.cmx \ pretyping/evd.cmx library/global.cmx kernel/inductive.cmx \ - pretyping/instantiate.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/pp.cmx pretyping/rawterm.cmx \ + pretyping/instantiate.cmx library/libnames.cmx library/nameops.cmx \ + kernel/names.cmx library/nametab.cmx lib/pp.cmx pretyping/rawterm.cmx \ pretyping/reductionops.cmx pretyping/retyping.cmx library/summary.cmx \ kernel/term.cmx pretyping/termops.cmx lib/util.cmx pretyping/tacred.cmi -pretyping/termops.cmo: kernel/environ.cmi library/nameops.cmi \ - kernel/names.cmi library/nametab.cmi lib/pp.cmi kernel/sign.cmi \ - kernel/term.cmi kernel/univ.cmi lib/util.cmi pretyping/termops.cmi -pretyping/termops.cmx: kernel/environ.cmx library/nameops.cmx \ - kernel/names.cmx library/nametab.cmx lib/pp.cmx kernel/sign.cmx \ - kernel/term.cmx kernel/univ.cmx lib/util.cmx pretyping/termops.cmi +pretyping/termops.cmo: kernel/environ.cmi library/lib.cmi \ + library/libnames.cmi library/nameops.cmi kernel/names.cmi \ + library/nametab.cmi lib/pp.cmi kernel/sign.cmi kernel/term.cmi \ + kernel/univ.cmi lib/util.cmi pretyping/termops.cmi +pretyping/termops.cmx: kernel/environ.cmx library/lib.cmx \ + library/libnames.cmx library/nameops.cmx kernel/names.cmx \ + library/nametab.cmx lib/pp.cmx kernel/sign.cmx kernel/term.cmx \ + kernel/univ.cmx lib/util.cmx pretyping/termops.cmi pretyping/typing.cmo: kernel/environ.cmi kernel/inductive.cmi \ pretyping/instantiate.cmi kernel/names.cmi pretyping/pretype_errors.cmi \ pretyping/reductionops.cmi kernel/term.cmi kernel/type_errors.cmi \ @@ -1087,41 +1195,41 @@ proofs/logic.cmx: parsing/coqast.cmx library/declare.cmx kernel/environ.cmx \ pretyping/termops.cmx kernel/type_errors.cmx kernel/typeops.cmx \ pretyping/typing.cmx lib/util.cmx proofs/logic.cmi proofs/pfedit.cmo: parsing/astterm.cmi kernel/declarations.cmi \ - library/declare.cmi lib/edit.cmi kernel/environ.cmi \ + library/declare.cmi lib/edit.cmi kernel/entries.cmi kernel/environ.cmi \ proofs/evar_refiner.cmi pretyping/evd.cmi library/lib.cmi \ - library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/pp.cmi \ + library/libnames.cmi library/nameops.cmi kernel/names.cmi lib/pp.cmi \ proofs/proof_trees.cmi proofs/proof_type.cmi kernel/safe_typing.cmi \ kernel/sign.cmi proofs/tacmach.cmi kernel/term.cmi pretyping/typing.cmi \ lib/util.cmi proofs/pfedit.cmi proofs/pfedit.cmx: parsing/astterm.cmx kernel/declarations.cmx \ - library/declare.cmx lib/edit.cmx kernel/environ.cmx \ + library/declare.cmx lib/edit.cmx kernel/entries.cmx kernel/environ.cmx \ proofs/evar_refiner.cmx pretyping/evd.cmx library/lib.cmx \ - library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/pp.cmx \ + library/libnames.cmx library/nameops.cmx kernel/names.cmx lib/pp.cmx \ proofs/proof_trees.cmx proofs/proof_type.cmx kernel/safe_typing.cmx \ kernel/sign.cmx proofs/tacmach.cmx kernel/term.cmx pretyping/typing.cmx \ lib/util.cmx proofs/pfedit.cmi proofs/proof_trees.cmo: kernel/closure.cmi pretyping/detyping.cmi \ kernel/environ.cmi pretyping/evarutil.cmi pretyping/evd.cmi \ - library/global.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/pp.cmi parsing/printer.cmi proofs/proof_type.cmi \ - kernel/sign.cmi pretyping/tacred.cmi kernel/term.cmi \ - pretyping/termops.cmi pretyping/typing.cmi lib/util.cmi \ + library/global.cmi library/libnames.cmi library/nameops.cmi \ + kernel/names.cmi library/nametab.cmi lib/pp.cmi parsing/printer.cmi \ + proofs/proof_type.cmi kernel/sign.cmi pretyping/tacred.cmi \ + kernel/term.cmi pretyping/termops.cmi pretyping/typing.cmi lib/util.cmi \ proofs/proof_trees.cmi proofs/proof_trees.cmx: kernel/closure.cmx pretyping/detyping.cmx \ kernel/environ.cmx pretyping/evarutil.cmx pretyping/evd.cmx \ - library/global.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/pp.cmx parsing/printer.cmx proofs/proof_type.cmx \ - kernel/sign.cmx pretyping/tacred.cmx kernel/term.cmx \ - pretyping/termops.cmx pretyping/typing.cmx lib/util.cmx \ + library/global.cmx library/libnames.cmx library/nameops.cmx \ + kernel/names.cmx library/nametab.cmx lib/pp.cmx parsing/printer.cmx \ + proofs/proof_type.cmx kernel/sign.cmx pretyping/tacred.cmx \ + kernel/term.cmx pretyping/termops.cmx pretyping/typing.cmx lib/util.cmx \ proofs/proof_trees.cmi proofs/proof_type.cmo: kernel/closure.cmi kernel/environ.cmi \ - pretyping/evd.cmi parsing/genarg.cmi kernel/names.cmi library/nametab.cmi \ - pretyping/rawterm.cmi proofs/tacexpr.cmo kernel/term.cmi lib/util.cmi \ - proofs/proof_type.cmi + pretyping/evd.cmi parsing/genarg.cmi library/libnames.cmi \ + kernel/names.cmi pretyping/rawterm.cmi proofs/tacexpr.cmo kernel/term.cmi \ + lib/util.cmi proofs/proof_type.cmi proofs/proof_type.cmx: kernel/closure.cmx kernel/environ.cmx \ - pretyping/evd.cmx parsing/genarg.cmx kernel/names.cmx library/nametab.cmx \ - pretyping/rawterm.cmx proofs/tacexpr.cmx kernel/term.cmx lib/util.cmx \ - proofs/proof_type.cmi + pretyping/evd.cmx parsing/genarg.cmx library/libnames.cmx \ + kernel/names.cmx pretyping/rawterm.cmx proofs/tacexpr.cmx kernel/term.cmx \ + lib/util.cmx proofs/proof_type.cmi proofs/refiner.cmo: kernel/environ.cmi pretyping/evarutil.cmi \ pretyping/evd.cmi library/global.cmi pretyping/instantiate.cmi \ proofs/logic.cmi lib/pp.cmi parsing/pptactic.cmi parsing/printer.cmi \ @@ -1135,9 +1243,11 @@ proofs/refiner.cmx: kernel/environ.cmx pretyping/evarutil.cmx \ kernel/sign.cmx proofs/tacexpr.cmx kernel/term.cmx pretyping/termops.cmx \ kernel/type_errors.cmx lib/util.cmx proofs/refiner.cmi proofs/tacexpr.cmo: parsing/coqast.cmi lib/dyn.cmi parsing/genarg.cmi \ - kernel/names.cmi library/nametab.cmi pretyping/rawterm.cmi lib/util.cmi + library/libnames.cmi kernel/names.cmi library/nametab.cmi \ + pretyping/rawterm.cmi lib/util.cmi proofs/tacexpr.cmx: parsing/coqast.cmx lib/dyn.cmx parsing/genarg.cmx \ - kernel/names.cmx library/nametab.cmx pretyping/rawterm.cmx lib/util.cmx + library/libnames.cmx kernel/names.cmx library/nametab.cmx \ + pretyping/rawterm.cmx lib/util.cmx proofs/tacmach.cmo: parsing/astterm.cmi library/declare.cmi \ kernel/environ.cmi pretyping/evarutil.cmi pretyping/evd.cmi \ library/global.cmi pretyping/instantiate.cmi proofs/logic.cmi \ @@ -1170,28 +1280,30 @@ tactics/auto.cmo: parsing/astterm.cmi tactics/btermdn.cmi proofs/clenv.cmi \ kernel/declarations.cmi library/declare.cmi tactics/dhyp.cmi \ proofs/evar_refiner.cmi pretyping/evd.cmi library/global.cmi \ tactics/hiddentac.cmi tactics/hipattern.cmi kernel/inductive.cmi \ - library/lib.cmi library/libobject.cmi library/library.cmi \ - proofs/logic.cmi library/nameops.cmi kernel/names.cmi library/nametab.cmi \ - lib/options.cmi pretyping/pattern.cmi proofs/pfedit.cmi lib/pp.cmi \ - parsing/pptactic.cmi parsing/printer.cmi proofs/proof_type.cmi \ - pretyping/rawterm.cmi kernel/reduction.cmi proofs/refiner.cmi \ - kernel/sign.cmi library/summary.cmi proofs/tacexpr.cmo proofs/tacmach.cmi \ - pretyping/tacred.cmi tactics/tacticals.cmi tactics/tactics.cmi \ - kernel/term.cmi pretyping/termops.cmi pretyping/typing.cmi lib/util.cmi \ - toplevel/vernacexpr.cmo tactics/auto.cmi + library/lib.cmi library/libnames.cmi library/libobject.cmi \ + library/library.cmi proofs/logic.cmi library/nameops.cmi kernel/names.cmi \ + library/nametab.cmi lib/options.cmi pretyping/pattern.cmi \ + proofs/pfedit.cmi lib/pp.cmi parsing/pptactic.cmi parsing/printer.cmi \ + proofs/proof_type.cmi pretyping/rawterm.cmi kernel/reduction.cmi \ + proofs/refiner.cmi kernel/sign.cmi library/summary.cmi proofs/tacexpr.cmo \ + proofs/tacmach.cmi pretyping/tacred.cmi tactics/tacticals.cmi \ + tactics/tactics.cmi kernel/term.cmi pretyping/termops.cmi \ + pretyping/typing.cmi lib/util.cmi toplevel/vernacexpr.cmo \ + tactics/auto.cmi tactics/auto.cmx: parsing/astterm.cmx tactics/btermdn.cmx proofs/clenv.cmx \ kernel/declarations.cmx library/declare.cmx tactics/dhyp.cmx \ proofs/evar_refiner.cmx pretyping/evd.cmx library/global.cmx \ tactics/hiddentac.cmx tactics/hipattern.cmx kernel/inductive.cmx \ - library/lib.cmx library/libobject.cmx library/library.cmx \ - proofs/logic.cmx library/nameops.cmx kernel/names.cmx library/nametab.cmx \ - lib/options.cmx pretyping/pattern.cmx proofs/pfedit.cmx lib/pp.cmx \ - parsing/pptactic.cmx parsing/printer.cmx proofs/proof_type.cmx \ - pretyping/rawterm.cmx kernel/reduction.cmx proofs/refiner.cmx \ - kernel/sign.cmx library/summary.cmx proofs/tacexpr.cmx proofs/tacmach.cmx \ - pretyping/tacred.cmx tactics/tacticals.cmx tactics/tactics.cmx \ - kernel/term.cmx pretyping/termops.cmx pretyping/typing.cmx lib/util.cmx \ - toplevel/vernacexpr.cmx tactics/auto.cmi + library/lib.cmx library/libnames.cmx library/libobject.cmx \ + library/library.cmx proofs/logic.cmx library/nameops.cmx kernel/names.cmx \ + library/nametab.cmx lib/options.cmx pretyping/pattern.cmx \ + proofs/pfedit.cmx lib/pp.cmx parsing/pptactic.cmx parsing/printer.cmx \ + proofs/proof_type.cmx pretyping/rawterm.cmx kernel/reduction.cmx \ + proofs/refiner.cmx kernel/sign.cmx library/summary.cmx proofs/tacexpr.cmx \ + proofs/tacmach.cmx pretyping/tacred.cmx tactics/tacticals.cmx \ + tactics/tactics.cmx kernel/term.cmx pretyping/termops.cmx \ + pretyping/typing.cmx lib/util.cmx toplevel/vernacexpr.cmx \ + tactics/auto.cmi tactics/autorewrite.cmo: parsing/ast.cmi parsing/coqast.cmi \ tactics/equality.cmi tactics/hipattern.cmi library/lib.cmi \ library/libobject.cmi kernel/names.cmi lib/pp.cmi proofs/proof_type.cmi \ @@ -1253,15 +1365,15 @@ tactics/eauto.cmx: tactics/auto.cmx toplevel/cerrors.cmx proofs/clenv.cmx \ tactics/tacinterp.cmx proofs/tacmach.cmx tactics/tacticals.cmx \ tactics/tactics.cmx kernel/term.cmx pretyping/termops.cmx lib/util.cmx tactics/elim.cmo: proofs/clenv.cmi kernel/environ.cmi tactics/hiddentac.cmi \ - tactics/hipattern.cmi pretyping/inductiveops.cmi kernel/names.cmi \ - library/nametab.cmi lib/pp.cmi parsing/printer.cmi proofs/proof_type.cmi \ + tactics/hipattern.cmi pretyping/inductiveops.cmi library/libnames.cmi \ + kernel/names.cmi lib/pp.cmi parsing/printer.cmi proofs/proof_type.cmi \ pretyping/rawterm.cmi kernel/reduction.cmi proofs/refiner.cmi \ proofs/tacexpr.cmo proofs/tacmach.cmi tactics/tacticals.cmi \ tactics/tactics.cmi kernel/term.cmi pretyping/termops.cmi lib/util.cmi \ tactics/elim.cmi tactics/elim.cmx: proofs/clenv.cmx kernel/environ.cmx tactics/hiddentac.cmx \ - tactics/hipattern.cmx pretyping/inductiveops.cmx kernel/names.cmx \ - library/nametab.cmx lib/pp.cmx parsing/printer.cmx proofs/proof_type.cmx \ + tactics/hipattern.cmx pretyping/inductiveops.cmx library/libnames.cmx \ + kernel/names.cmx lib/pp.cmx parsing/printer.cmx proofs/proof_type.cmx \ pretyping/rawterm.cmx kernel/reduction.cmx proofs/refiner.cmx \ proofs/tacexpr.cmx proofs/tacmach.cmx tactics/tacticals.cmx \ tactics/tactics.cmx kernel/term.cmx pretyping/termops.cmx lib/util.cmx \ @@ -1375,24 +1487,24 @@ tactics/inv.cmx: proofs/clenv.cmx parsing/coqlib.cmx tactics/elim.cmx \ tactics/tactics.cmx kernel/term.cmx pretyping/termops.cmx \ pretyping/typing.cmx lib/util.cmx tactics/wcclausenv.cmx tactics/inv.cmi tactics/leminv.cmo: parsing/astterm.cmi proofs/clenv.cmi \ - kernel/declarations.cmi library/declare.cmi kernel/environ.cmi \ - proofs/evar_refiner.cmi pretyping/evd.cmi library/global.cmi \ - pretyping/inductiveops.cmi tactics/inv.cmi library/nameops.cmi \ - kernel/names.cmi library/nametab.cmi proofs/pfedit.cmi lib/pp.cmi \ - parsing/printer.cmi proofs/proof_trees.cmi proofs/proof_type.cmi \ - pretyping/reductionops.cmi kernel/safe_typing.cmi kernel/sign.cmi \ - proofs/tacmach.cmi tactics/tacticals.cmi tactics/tactics.cmi \ - kernel/term.cmi pretyping/termops.cmi lib/util.cmi \ + kernel/declarations.cmi library/declare.cmi kernel/entries.cmi \ + kernel/environ.cmi proofs/evar_refiner.cmi pretyping/evd.cmi \ + library/global.cmi pretyping/inductiveops.cmi tactics/inv.cmi \ + library/libnames.cmi library/nameops.cmi kernel/names.cmi \ + proofs/pfedit.cmi lib/pp.cmi parsing/printer.cmi proofs/proof_trees.cmi \ + proofs/proof_type.cmi pretyping/reductionops.cmi kernel/safe_typing.cmi \ + kernel/sign.cmi proofs/tacmach.cmi tactics/tacticals.cmi \ + tactics/tactics.cmi kernel/term.cmi pretyping/termops.cmi lib/util.cmi \ toplevel/vernacexpr.cmo tactics/wcclausenv.cmi tactics/leminv.cmi tactics/leminv.cmx: parsing/astterm.cmx proofs/clenv.cmx \ - kernel/declarations.cmx library/declare.cmx kernel/environ.cmx \ - proofs/evar_refiner.cmx pretyping/evd.cmx library/global.cmx \ - pretyping/inductiveops.cmx tactics/inv.cmx library/nameops.cmx \ - kernel/names.cmx library/nametab.cmx proofs/pfedit.cmx lib/pp.cmx \ - parsing/printer.cmx proofs/proof_trees.cmx proofs/proof_type.cmx \ - pretyping/reductionops.cmx kernel/safe_typing.cmx kernel/sign.cmx \ - proofs/tacmach.cmx tactics/tacticals.cmx tactics/tactics.cmx \ - kernel/term.cmx pretyping/termops.cmx lib/util.cmx \ + kernel/declarations.cmx library/declare.cmx kernel/entries.cmx \ + kernel/environ.cmx proofs/evar_refiner.cmx pretyping/evd.cmx \ + library/global.cmx pretyping/inductiveops.cmx tactics/inv.cmx \ + library/libnames.cmx library/nameops.cmx kernel/names.cmx \ + proofs/pfedit.cmx lib/pp.cmx parsing/printer.cmx proofs/proof_trees.cmx \ + proofs/proof_type.cmx pretyping/reductionops.cmx kernel/safe_typing.cmx \ + kernel/sign.cmx proofs/tacmach.cmx tactics/tacticals.cmx \ + tactics/tactics.cmx kernel/term.cmx pretyping/termops.cmx lib/util.cmx \ toplevel/vernacexpr.cmx tactics/wcclausenv.cmx tactics/leminv.cmi tactics/nbtermdn.cmo: tactics/btermdn.cmi lib/gmap.cmi library/libobject.cmi \ library/library.cmi kernel/names.cmi pretyping/pattern.cmi \ @@ -1413,35 +1525,37 @@ tactics/refine.cmx: proofs/clenv.cmx kernel/environ.cmx pretyping/evd.cmx \ pretyping/termops.cmx pretyping/typing.cmx lib/util.cmx \ tactics/refine.cmi tactics/setoid_replace.cmo: parsing/astterm.cmi tactics/auto.cmi \ - toplevel/command.cmi library/declare.cmi kernel/environ.cmi \ - pretyping/evd.cmi library/global.cmi lib/gmap.cmi library/lib.cmi \ - library/libobject.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/options.cmi lib/pp.cmi parsing/printer.cmi \ - proofs/proof_type.cmi pretyping/reductionops.cmi kernel/safe_typing.cmi \ - library/summary.cmi proofs/tacmach.cmi tactics/tacticals.cmi \ - tactics/tactics.cmi kernel/term.cmi parsing/termast.cmi \ - pretyping/termops.cmi pretyping/typing.cmi lib/util.cmi \ - toplevel/vernacentries.cmi toplevel/vernacexpr.cmo \ - toplevel/vernacinterp.cmi tactics/setoid_replace.cmi + toplevel/command.cmi library/declare.cmi kernel/entries.cmi \ + kernel/environ.cmi pretyping/evd.cmi library/global.cmi lib/gmap.cmi \ + library/lib.cmi library/libnames.cmi library/libobject.cmi \ + library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/options.cmi \ + lib/pp.cmi parsing/printer.cmi proofs/proof_type.cmi \ + pretyping/reductionops.cmi kernel/safe_typing.cmi library/summary.cmi \ + proofs/tacmach.cmi tactics/tacticals.cmi tactics/tactics.cmi \ + kernel/term.cmi parsing/termast.cmi pretyping/termops.cmi \ + pretyping/typing.cmi lib/util.cmi toplevel/vernacentries.cmi \ + toplevel/vernacexpr.cmo toplevel/vernacinterp.cmi \ + tactics/setoid_replace.cmi tactics/setoid_replace.cmx: parsing/astterm.cmx tactics/auto.cmx \ - toplevel/command.cmx library/declare.cmx kernel/environ.cmx \ - pretyping/evd.cmx library/global.cmx lib/gmap.cmx library/lib.cmx \ - library/libobject.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/options.cmx lib/pp.cmx parsing/printer.cmx \ - proofs/proof_type.cmx pretyping/reductionops.cmx kernel/safe_typing.cmx \ - library/summary.cmx proofs/tacmach.cmx tactics/tacticals.cmx \ - tactics/tactics.cmx kernel/term.cmx parsing/termast.cmx \ - pretyping/termops.cmx pretyping/typing.cmx lib/util.cmx \ - toplevel/vernacentries.cmx toplevel/vernacexpr.cmx \ - toplevel/vernacinterp.cmx tactics/setoid_replace.cmi + toplevel/command.cmx library/declare.cmx kernel/entries.cmx \ + kernel/environ.cmx pretyping/evd.cmx library/global.cmx lib/gmap.cmx \ + library/lib.cmx library/libnames.cmx library/libobject.cmx \ + library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/options.cmx \ + lib/pp.cmx parsing/printer.cmx proofs/proof_type.cmx \ + pretyping/reductionops.cmx kernel/safe_typing.cmx library/summary.cmx \ + proofs/tacmach.cmx tactics/tacticals.cmx tactics/tactics.cmx \ + kernel/term.cmx parsing/termast.cmx pretyping/termops.cmx \ + pretyping/typing.cmx lib/util.cmx toplevel/vernacentries.cmx \ + toplevel/vernacexpr.cmx toplevel/vernacinterp.cmx \ + tactics/setoid_replace.cmi tactics/tacinterp.cmo: parsing/ast.cmi parsing/astterm.cmi tactics/auto.cmi \ kernel/closure.cmi parsing/coqast.cmi kernel/declarations.cmi \ library/declare.cmi tactics/dhyp.cmi lib/dyn.cmi tactics/elim.cmi \ - kernel/environ.cmi pretyping/evd.cmi parsing/genarg.cmi \ - library/global.cmi lib/gmap.cmi tactics/hiddentac.cmi library/lib.cmi \ - library/libobject.cmi proofs/logic.cmi library/nameops.cmi \ - kernel/names.cmi library/nametab.cmi lib/options.cmi \ - pretyping/pattern.cmi proofs/pfedit.cmi lib/pp.cmi \ + kernel/entries.cmi kernel/environ.cmi pretyping/evd.cmi \ + parsing/genarg.cmi library/global.cmi lib/gmap.cmi tactics/hiddentac.cmi \ + library/lib.cmi library/libnames.cmi library/libobject.cmi \ + proofs/logic.cmi library/nameops.cmi kernel/names.cmi library/nametab.cmi \ + lib/options.cmi pretyping/pattern.cmi proofs/pfedit.cmi lib/pp.cmi \ pretyping/pretype_errors.cmi pretyping/pretyping.cmi parsing/printer.cmi \ proofs/proof_type.cmi pretyping/rawterm.cmi pretyping/reductionops.cmi \ proofs/refiner.cmi kernel/safe_typing.cmi kernel/sign.cmi \ @@ -1452,11 +1566,11 @@ tactics/tacinterp.cmo: parsing/ast.cmi parsing/astterm.cmi tactics/auto.cmi \ tactics/tacinterp.cmx: parsing/ast.cmx parsing/astterm.cmx tactics/auto.cmx \ kernel/closure.cmx parsing/coqast.cmx kernel/declarations.cmx \ library/declare.cmx tactics/dhyp.cmx lib/dyn.cmx tactics/elim.cmx \ - kernel/environ.cmx pretyping/evd.cmx parsing/genarg.cmx \ - library/global.cmx lib/gmap.cmx tactics/hiddentac.cmx library/lib.cmx \ - library/libobject.cmx proofs/logic.cmx library/nameops.cmx \ - kernel/names.cmx library/nametab.cmx lib/options.cmx \ - pretyping/pattern.cmx proofs/pfedit.cmx lib/pp.cmx \ + kernel/entries.cmx kernel/environ.cmx pretyping/evd.cmx \ + parsing/genarg.cmx library/global.cmx lib/gmap.cmx tactics/hiddentac.cmx \ + library/lib.cmx library/libnames.cmx library/libobject.cmx \ + proofs/logic.cmx library/nameops.cmx kernel/names.cmx library/nametab.cmx \ + lib/options.cmx pretyping/pattern.cmx proofs/pfedit.cmx lib/pp.cmx \ pretyping/pretype_errors.cmx pretyping/pretyping.cmx parsing/printer.cmx \ proofs/proof_type.cmx pretyping/rawterm.cmx pretyping/reductionops.cmx \ proofs/refiner.cmx kernel/safe_typing.cmx kernel/sign.cmx \ @@ -1467,57 +1581,59 @@ tactics/tacinterp.cmx: parsing/ast.cmx parsing/astterm.cmx tactics/auto.cmx \ tactics/tacticals.cmo: proofs/clenv.cmi kernel/declarations.cmi \ library/declare.cmi kernel/environ.cmi proofs/evar_refiner.cmi \ library/global.cmi pretyping/indrec.cmi kernel/inductive.cmi \ - kernel/names.cmi pretyping/pattern.cmi lib/pp.cmi kernel/reduction.cmi \ - proofs/refiner.cmi kernel/sign.cmi proofs/tacmach.cmi kernel/term.cmi \ - pretyping/termops.cmi lib/util.cmi tactics/wcclausenv.cmi \ - tactics/tacticals.cmi + library/libnames.cmi kernel/names.cmi pretyping/pattern.cmi lib/pp.cmi \ + kernel/reduction.cmi proofs/refiner.cmi kernel/sign.cmi \ + proofs/tacmach.cmi kernel/term.cmi pretyping/termops.cmi lib/util.cmi \ + tactics/wcclausenv.cmi tactics/tacticals.cmi tactics/tacticals.cmx: proofs/clenv.cmx kernel/declarations.cmx \ library/declare.cmx kernel/environ.cmx proofs/evar_refiner.cmx \ library/global.cmx pretyping/indrec.cmx kernel/inductive.cmx \ - kernel/names.cmx pretyping/pattern.cmx lib/pp.cmx kernel/reduction.cmx \ - proofs/refiner.cmx kernel/sign.cmx proofs/tacmach.cmx kernel/term.cmx \ - pretyping/termops.cmx lib/util.cmx tactics/wcclausenv.cmx \ - tactics/tacticals.cmi + library/libnames.cmx kernel/names.cmx pretyping/pattern.cmx lib/pp.cmx \ + kernel/reduction.cmx proofs/refiner.cmx kernel/sign.cmx \ + proofs/tacmach.cmx kernel/term.cmx pretyping/termops.cmx lib/util.cmx \ + tactics/wcclausenv.cmx tactics/tacticals.cmi tactics/tactics.cmo: parsing/astterm.cmi proofs/clenv.cmi kernel/closure.cmi \ parsing/coqlib.cmi kernel/declarations.cmi library/declare.cmi \ - kernel/environ.cmi proofs/evar_refiner.cmi pretyping/evd.cmi \ - library/global.cmi tactics/hipattern.cmi pretyping/indrec.cmi \ - kernel/inductive.cmi pretyping/inductiveops.cmi proofs/logic.cmi \ - library/nameops.cmi kernel/names.cmi library/nametab.cmi \ - proofs/pfedit.cmi lib/pp.cmi proofs/proof_trees.cmi proofs/proof_type.cmi \ - pretyping/rawterm.cmi pretyping/reductionops.cmi proofs/refiner.cmi \ - kernel/safe_typing.cmi kernel/sign.cmi proofs/tacexpr.cmo \ - proofs/tacmach.cmi pretyping/tacred.cmi tactics/tacticals.cmi \ - kernel/term.cmi pretyping/termops.cmi lib/util.cmi tactics/tactics.cmi + kernel/entries.cmi kernel/environ.cmi proofs/evar_refiner.cmi \ + pretyping/evd.cmi library/global.cmi tactics/hipattern.cmi \ + pretyping/indrec.cmi kernel/inductive.cmi pretyping/inductiveops.cmi \ + library/libnames.cmi proofs/logic.cmi library/nameops.cmi \ + kernel/names.cmi library/nametab.cmi proofs/pfedit.cmi lib/pp.cmi \ + proofs/proof_trees.cmi proofs/proof_type.cmi pretyping/rawterm.cmi \ + pretyping/reductionops.cmi proofs/refiner.cmi kernel/sign.cmi \ + proofs/tacexpr.cmo proofs/tacmach.cmi pretyping/tacred.cmi \ + tactics/tacticals.cmi kernel/term.cmi pretyping/termops.cmi lib/util.cmi \ + tactics/tactics.cmi tactics/tactics.cmx: parsing/astterm.cmx proofs/clenv.cmx kernel/closure.cmx \ parsing/coqlib.cmx kernel/declarations.cmx library/declare.cmx \ - kernel/environ.cmx proofs/evar_refiner.cmx pretyping/evd.cmx \ - library/global.cmx tactics/hipattern.cmx pretyping/indrec.cmx \ - kernel/inductive.cmx pretyping/inductiveops.cmx proofs/logic.cmx \ - library/nameops.cmx kernel/names.cmx library/nametab.cmx \ - proofs/pfedit.cmx lib/pp.cmx proofs/proof_trees.cmx proofs/proof_type.cmx \ - pretyping/rawterm.cmx pretyping/reductionops.cmx proofs/refiner.cmx \ - kernel/safe_typing.cmx kernel/sign.cmx proofs/tacexpr.cmx \ - proofs/tacmach.cmx pretyping/tacred.cmx tactics/tacticals.cmx \ - kernel/term.cmx pretyping/termops.cmx lib/util.cmx tactics/tactics.cmi + kernel/entries.cmx kernel/environ.cmx proofs/evar_refiner.cmx \ + pretyping/evd.cmx library/global.cmx tactics/hipattern.cmx \ + pretyping/indrec.cmx kernel/inductive.cmx pretyping/inductiveops.cmx \ + library/libnames.cmx proofs/logic.cmx library/nameops.cmx \ + kernel/names.cmx library/nametab.cmx proofs/pfedit.cmx lib/pp.cmx \ + proofs/proof_trees.cmx proofs/proof_type.cmx pretyping/rawterm.cmx \ + pretyping/reductionops.cmx proofs/refiner.cmx kernel/sign.cmx \ + proofs/tacexpr.cmx proofs/tacmach.cmx pretyping/tacred.cmx \ + tactics/tacticals.cmx kernel/term.cmx pretyping/termops.cmx lib/util.cmx \ + tactics/tactics.cmi tactics/tauto.cmo: parsing/ast.cmi toplevel/cerrors.cmi parsing/coqast.cmi \ parsing/egrammar.cmi parsing/genarg.cmi tactics/hipattern.cmi \ - kernel/names.cmi library/nametab.cmi parsing/pcoq.cmi lib/pp.cmi \ + library/libnames.cmi kernel/names.cmi parsing/pcoq.cmi lib/pp.cmi \ parsing/pptactic.cmi proofs/proof_type.cmi pretyping/rawterm.cmi \ proofs/refiner.cmi proofs/tacexpr.cmo tactics/tacinterp.cmi \ tactics/tacticals.cmi tactics/tactics.cmi lib/util.cmi tactics/tauto.cmx: parsing/ast.cmx toplevel/cerrors.cmx parsing/coqast.cmx \ parsing/egrammar.cmx parsing/genarg.cmx tactics/hipattern.cmx \ - kernel/names.cmx library/nametab.cmx parsing/pcoq.cmx lib/pp.cmx \ + library/libnames.cmx kernel/names.cmx parsing/pcoq.cmx lib/pp.cmx \ parsing/pptactic.cmx proofs/proof_type.cmx pretyping/rawterm.cmx \ proofs/refiner.cmx proofs/tacexpr.cmx tactics/tacinterp.cmx \ tactics/tacticals.cmx tactics/tactics.cmx lib/util.cmx -tactics/termdn.cmo: tactics/dn.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi pretyping/pattern.cmi pretyping/rawterm.cmi \ - kernel/term.cmi lib/util.cmi tactics/termdn.cmi -tactics/termdn.cmx: tactics/dn.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx pretyping/pattern.cmx pretyping/rawterm.cmx \ - kernel/term.cmx lib/util.cmx tactics/termdn.cmi +tactics/termdn.cmo: tactics/dn.cmi library/libnames.cmi library/nameops.cmi \ + kernel/names.cmi library/nametab.cmi pretyping/pattern.cmi \ + pretyping/rawterm.cmi kernel/term.cmi lib/util.cmi tactics/termdn.cmi +tactics/termdn.cmx: tactics/dn.cmx library/libnames.cmx library/nameops.cmx \ + kernel/names.cmx library/nametab.cmx pretyping/pattern.cmx \ + pretyping/rawterm.cmx kernel/term.cmx lib/util.cmx tactics/termdn.cmi tactics/wcclausenv.cmo: proofs/clenv.cmi kernel/environ.cmi \ proofs/evar_refiner.cmi pretyping/evd.cmi library/global.cmi \ proofs/logic.cmi library/nameops.cmi kernel/names.cmi lib/pp.cmi \ @@ -1530,6 +1646,14 @@ tactics/wcclausenv.cmx: proofs/clenv.cmx kernel/environ.cmx \ proofs/proof_trees.cmx pretyping/reductionops.cmx proofs/refiner.cmx \ kernel/sign.cmx proofs/tacmach.cmx kernel/term.cmx pretyping/termops.cmx \ lib/util.cmx tactics/wcclausenv.cmi +tmp/probj.cmo: library/declaremods.cmi library/libobject.cmi +tmp/probj.cmx: library/declaremods.cmx library/libobject.cmx +tmp/vernacentries.cmo: parsing/ast.cmi toplevel/class.cmi \ + toplevel/command.cmi library/declare.cmi library/libnames.cmi \ + tactics/tactics.cmi +tmp/vernacentries.cmx: parsing/ast.cmx toplevel/class.cmx \ + toplevel/command.cmx library/declare.cmx library/libnames.cmx \ + tactics/tactics.cmx tools/coqdep_lexer.cmo: config/coq_config.cmi tools/coqdep_lexer.cmx: config/coq_config.cmx tools/coqdep.cmo: config/coq_config.cmi tools/coqdep_lexer.cmo @@ -1539,57 +1663,57 @@ tools/coq_vo2xml.cmx: config/coq_config.cmx toplevel/usage.cmx tools/gallina.cmo: tools/gallina_lexer.cmo tools/gallina.cmx: tools/gallina_lexer.cmx toplevel/cerrors.cmo: parsing/ast.cmi pretyping/cases.cmi toplevel/himsg.cmi \ - kernel/indtypes.cmi parsing/lexer.cmi proofs/logic.cmi \ - library/nametab.cmi lib/options.cmi lib/pp.cmi \ + kernel/indtypes.cmi parsing/lexer.cmi library/libnames.cmi \ + proofs/logic.cmi library/nametab.cmi lib/options.cmi lib/pp.cmi \ pretyping/pretype_errors.cmi proofs/refiner.cmi kernel/type_errors.cmi \ kernel/univ.cmi lib/util.cmi toplevel/cerrors.cmi toplevel/cerrors.cmx: parsing/ast.cmx pretyping/cases.cmx toplevel/himsg.cmx \ - kernel/indtypes.cmx parsing/lexer.cmx proofs/logic.cmx \ - library/nametab.cmx lib/options.cmx lib/pp.cmx \ + kernel/indtypes.cmx parsing/lexer.cmx library/libnames.cmx \ + proofs/logic.cmx library/nametab.cmx lib/options.cmx lib/pp.cmx \ pretyping/pretype_errors.cmx proofs/refiner.cmx kernel/type_errors.cmx \ kernel/univ.cmx lib/util.cmx toplevel/cerrors.cmi toplevel/class.cmo: pretyping/classops.cmi kernel/declarations.cmi \ - library/declare.cmi kernel/environ.cmi pretyping/evd.cmi \ - library/global.cmi kernel/inductive.cmi library/lib.cmi \ - library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/options.cmi \ - lib/pp.cmi parsing/printer.cmi pretyping/reductionops.cmi \ - pretyping/retyping.cmi kernel/safe_typing.cmi kernel/sign.cmi \ - kernel/term.cmi pretyping/termops.cmi pretyping/typing.cmi lib/util.cmi \ - toplevel/class.cmi + library/declare.cmi kernel/entries.cmi kernel/environ.cmi \ + pretyping/evd.cmi library/global.cmi kernel/inductive.cmi library/lib.cmi \ + library/libnames.cmi library/nameops.cmi kernel/names.cmi \ + library/nametab.cmi lib/options.cmi lib/pp.cmi parsing/printer.cmi \ + pretyping/reductionops.cmi pretyping/retyping.cmi kernel/safe_typing.cmi \ + kernel/sign.cmi kernel/term.cmi pretyping/termops.cmi \ + pretyping/typing.cmi lib/util.cmi toplevel/class.cmi toplevel/class.cmx: pretyping/classops.cmx kernel/declarations.cmx \ - library/declare.cmx kernel/environ.cmx pretyping/evd.cmx \ - library/global.cmx kernel/inductive.cmx library/lib.cmx \ - library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/options.cmx \ - lib/pp.cmx parsing/printer.cmx pretyping/reductionops.cmx \ - pretyping/retyping.cmx kernel/safe_typing.cmx kernel/sign.cmx \ - kernel/term.cmx pretyping/termops.cmx pretyping/typing.cmx lib/util.cmx \ - toplevel/class.cmi + library/declare.cmx kernel/entries.cmx kernel/environ.cmx \ + pretyping/evd.cmx library/global.cmx kernel/inductive.cmx library/lib.cmx \ + library/libnames.cmx library/nameops.cmx kernel/names.cmx \ + library/nametab.cmx lib/options.cmx lib/pp.cmx parsing/printer.cmx \ + pretyping/reductionops.cmx pretyping/retyping.cmx kernel/safe_typing.cmx \ + kernel/sign.cmx kernel/term.cmx pretyping/termops.cmx \ + pretyping/typing.cmx lib/util.cmx toplevel/class.cmi toplevel/command.cmo: parsing/ast.cmi parsing/astterm.cmi toplevel/class.cmi \ parsing/coqast.cmi kernel/declarations.cmi library/declare.cmi \ - kernel/environ.cmi pretyping/evd.cmi library/global.cmi \ - tactics/hiddentac.cmi library/impargs.cmi pretyping/indrec.cmi \ - kernel/indtypes.cmi kernel/inductive.cmi library/lib.cmi \ - library/libobject.cmi library/library.cmi proofs/logic.cmi \ - library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/options.cmi \ - proofs/pfedit.cmi lib/pp.cmi parsing/printer.cmi proofs/proof_type.cmi \ - kernel/reduction.cmi proofs/refiner.cmi kernel/safe_typing.cmi \ - library/states.cmi pretyping/syntax_def.cmi proofs/tacmach.cmi \ - pretyping/tacred.cmi kernel/term.cmi pretyping/termops.cmi \ - kernel/typeops.cmi lib/util.cmi toplevel/vernacexpr.cmo \ - toplevel/command.cmi + kernel/entries.cmi kernel/environ.cmi pretyping/evd.cmi \ + library/global.cmi tactics/hiddentac.cmi library/impargs.cmi \ + pretyping/indrec.cmi kernel/indtypes.cmi kernel/inductive.cmi \ + library/lib.cmi library/libnames.cmi library/libobject.cmi \ + library/library.cmi proofs/logic.cmi library/nameops.cmi kernel/names.cmi \ + library/nametab.cmi lib/options.cmi proofs/pfedit.cmi lib/pp.cmi \ + parsing/printer.cmi proofs/proof_type.cmi kernel/reduction.cmi \ + proofs/refiner.cmi kernel/safe_typing.cmi library/states.cmi \ + pretyping/syntax_def.cmi proofs/tacmach.cmi pretyping/tacred.cmi \ + kernel/term.cmi pretyping/termops.cmi kernel/typeops.cmi lib/util.cmi \ + toplevel/vernacexpr.cmo toplevel/command.cmi toplevel/command.cmx: parsing/ast.cmx parsing/astterm.cmx toplevel/class.cmx \ parsing/coqast.cmx kernel/declarations.cmx library/declare.cmx \ - kernel/environ.cmx pretyping/evd.cmx library/global.cmx \ - tactics/hiddentac.cmx library/impargs.cmx pretyping/indrec.cmx \ - kernel/indtypes.cmx kernel/inductive.cmx library/lib.cmx \ - library/libobject.cmx library/library.cmx proofs/logic.cmx \ - library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/options.cmx \ - proofs/pfedit.cmx lib/pp.cmx parsing/printer.cmx proofs/proof_type.cmx \ - kernel/reduction.cmx proofs/refiner.cmx kernel/safe_typing.cmx \ - library/states.cmx pretyping/syntax_def.cmx proofs/tacmach.cmx \ - pretyping/tacred.cmx kernel/term.cmx pretyping/termops.cmx \ - kernel/typeops.cmx lib/util.cmx toplevel/vernacexpr.cmx \ - toplevel/command.cmi + kernel/entries.cmx kernel/environ.cmx pretyping/evd.cmx \ + library/global.cmx tactics/hiddentac.cmx library/impargs.cmx \ + pretyping/indrec.cmx kernel/indtypes.cmx kernel/inductive.cmx \ + library/lib.cmx library/libnames.cmx library/libobject.cmx \ + library/library.cmx proofs/logic.cmx library/nameops.cmx kernel/names.cmx \ + library/nametab.cmx lib/options.cmx proofs/pfedit.cmx lib/pp.cmx \ + parsing/printer.cmx proofs/proof_type.cmx kernel/reduction.cmx \ + proofs/refiner.cmx kernel/safe_typing.cmx library/states.cmx \ + pretyping/syntax_def.cmx proofs/tacmach.cmx pretyping/tacred.cmx \ + kernel/term.cmx pretyping/termops.cmx kernel/typeops.cmx lib/util.cmx \ + toplevel/vernacexpr.cmx toplevel/command.cmi toplevel/coqinit.cmo: config/coq_config.cmi toplevel/mltop.cmi \ library/nameops.cmi kernel/names.cmi lib/options.cmi lib/pp.cmi \ lib/system.cmi toplevel/toplevel.cmi toplevel/vernac.cmi \ @@ -1599,35 +1723,37 @@ toplevel/coqinit.cmx: config/coq_config.cmx toplevel/mltop.cmx \ lib/system.cmx toplevel/toplevel.cmx toplevel/vernac.cmx \ toplevel/coqinit.cmi toplevel/coqtop.cmo: toplevel/cerrors.cmi config/coq_config.cmi \ - toplevel/coqinit.cmi library/lib.cmi library/library.cmi \ - toplevel/mltop.cmi library/nameops.cmi kernel/names.cmi lib/options.cmi \ - lib/pp.cmi lib/profile.cmi library/states.cmi lib/system.cmi \ - toplevel/toplevel.cmi toplevel/usage.cmi lib/util.cmi toplevel/vernac.cmi \ - toplevel/coqtop.cmi + toplevel/coqinit.cmi library/lib.cmi library/libnames.cmi \ + library/library.cmi toplevel/mltop.cmi library/nameops.cmi \ + kernel/names.cmi lib/options.cmi lib/pp.cmi lib/profile.cmi \ + library/states.cmi lib/system.cmi toplevel/toplevel.cmi \ + toplevel/usage.cmi lib/util.cmi toplevel/vernac.cmi toplevel/coqtop.cmi toplevel/coqtop.cmx: toplevel/cerrors.cmx config/coq_config.cmx \ - toplevel/coqinit.cmx library/lib.cmx library/library.cmx \ - toplevel/mltop.cmx library/nameops.cmx kernel/names.cmx lib/options.cmx \ - lib/pp.cmx lib/profile.cmx library/states.cmx lib/system.cmx \ - toplevel/toplevel.cmx toplevel/usage.cmx lib/util.cmx toplevel/vernac.cmx \ - toplevel/coqtop.cmi + toplevel/coqinit.cmx library/lib.cmx library/libnames.cmx \ + library/library.cmx toplevel/mltop.cmx library/nameops.cmx \ + kernel/names.cmx lib/options.cmx lib/pp.cmx lib/profile.cmx \ + library/states.cmx lib/system.cmx toplevel/toplevel.cmx \ + toplevel/usage.cmx lib/util.cmx toplevel/vernac.cmx toplevel/coqtop.cmi toplevel/discharge.cmo: toplevel/class.cmi pretyping/classops.cmi \ kernel/cooking.cmi kernel/declarations.cmi library/declare.cmi \ - kernel/environ.cmi library/global.cmi library/impargs.cmi \ - kernel/indtypes.cmi kernel/inductive.cmi pretyping/instantiate.cmi \ - library/lib.cmi library/libobject.cmi library/library.cmi \ - library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/options.cmi \ - lib/pp.cmi toplevel/recordobj.cmi pretyping/recordops.cmi \ - kernel/reduction.cmi kernel/sign.cmi library/summary.cmi kernel/term.cmi \ - kernel/typeops.cmi kernel/univ.cmi lib/util.cmi toplevel/discharge.cmi + kernel/entries.cmi kernel/environ.cmi library/global.cmi \ + library/impargs.cmi kernel/indtypes.cmi kernel/inductive.cmi \ + pretyping/instantiate.cmi library/lib.cmi library/libnames.cmi \ + library/libobject.cmi library/library.cmi library/nameops.cmi \ + kernel/names.cmi library/nametab.cmi lib/options.cmi lib/pp.cmi \ + toplevel/recordobj.cmi pretyping/recordops.cmi kernel/reduction.cmi \ + kernel/sign.cmi library/summary.cmi kernel/term.cmi kernel/typeops.cmi \ + kernel/univ.cmi lib/util.cmi toplevel/discharge.cmi toplevel/discharge.cmx: toplevel/class.cmx pretyping/classops.cmx \ kernel/cooking.cmx kernel/declarations.cmx library/declare.cmx \ - kernel/environ.cmx library/global.cmx library/impargs.cmx \ - kernel/indtypes.cmx kernel/inductive.cmx pretyping/instantiate.cmx \ - library/lib.cmx library/libobject.cmx library/library.cmx \ - library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/options.cmx \ - lib/pp.cmx toplevel/recordobj.cmx pretyping/recordops.cmx \ - kernel/reduction.cmx kernel/sign.cmx library/summary.cmx kernel/term.cmx \ - kernel/typeops.cmx kernel/univ.cmx lib/util.cmx toplevel/discharge.cmi + kernel/entries.cmx kernel/environ.cmx library/global.cmx \ + library/impargs.cmx kernel/indtypes.cmx kernel/inductive.cmx \ + pretyping/instantiate.cmx library/lib.cmx library/libnames.cmx \ + library/libobject.cmx library/library.cmx library/nameops.cmx \ + kernel/names.cmx library/nametab.cmx lib/options.cmx lib/pp.cmx \ + toplevel/recordobj.cmx pretyping/recordops.cmx kernel/reduction.cmx \ + kernel/sign.cmx library/summary.cmx kernel/term.cmx kernel/typeops.cmx \ + kernel/univ.cmx lib/util.cmx toplevel/discharge.cmi toplevel/fhimsg.cmo: kernel/environ.cmi parsing/g_minicoq.cmi \ kernel/names.cmi lib/pp.cmi kernel/reduction.cmi kernel/sign.cmi \ kernel/term.cmi kernel/type_errors.cmi lib/util.cmi toplevel/fhimsg.cmi @@ -1653,15 +1779,15 @@ toplevel/line_oriented_parser.cmx: toplevel/line_oriented_parser.cmi toplevel/metasyntax.cmo: parsing/ast.cmi parsing/astterm.cmi \ parsing/coqast.cmi parsing/egrammar.cmi parsing/esyntax.cmi \ parsing/extend.cmi parsing/genarg.cmi library/global.cmi library/lib.cmi \ - library/libobject.cmi library/library.cmi library/nametab.cmi \ - parsing/pcoq.cmi lib/pp.cmi library/summary.cmi lib/util.cmi \ - toplevel/vernacexpr.cmo toplevel/metasyntax.cmi + library/libnames.cmi library/libobject.cmi library/library.cmi \ + library/nametab.cmi parsing/pcoq.cmi lib/pp.cmi library/summary.cmi \ + lib/util.cmi toplevel/vernacexpr.cmo toplevel/metasyntax.cmi toplevel/metasyntax.cmx: parsing/ast.cmx parsing/astterm.cmx \ parsing/coqast.cmx parsing/egrammar.cmx parsing/esyntax.cmx \ parsing/extend.cmx parsing/genarg.cmx library/global.cmx library/lib.cmx \ - library/libobject.cmx library/library.cmx library/nametab.cmx \ - parsing/pcoq.cmx lib/pp.cmx library/summary.cmx lib/util.cmx \ - toplevel/vernacexpr.cmx toplevel/metasyntax.cmi + library/libnames.cmx library/libobject.cmx library/library.cmx \ + library/nametab.cmx parsing/pcoq.cmx lib/pp.cmx library/summary.cmx \ + lib/util.cmx toplevel/vernacexpr.cmx toplevel/metasyntax.cmi toplevel/minicoq.cmo: kernel/declarations.cmi toplevel/fhimsg.cmi \ parsing/g_minicoq.cmi kernel/inductive.cmi kernel/names.cmi lib/pp.cmi \ kernel/safe_typing.cmi kernel/sign.cmi kernel/term.cmi \ @@ -1686,31 +1812,33 @@ toplevel/protectedtoplevel.cmx: toplevel/cerrors.cmx \ toplevel/protectedtoplevel.cmi toplevel/record.cmo: parsing/ast.cmi parsing/astterm.cmi toplevel/class.cmi \ toplevel/command.cmi parsing/coqast.cmi kernel/declarations.cmi \ - library/declare.cmi kernel/environ.cmi pretyping/evd.cmi \ - library/global.cmi kernel/indtypes.cmi kernel/inductive.cmi \ - pretyping/inductiveops.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/options.cmi lib/pp.cmi parsing/printer.cmi \ - pretyping/recordops.cmi kernel/safe_typing.cmi kernel/term.cmi \ - pretyping/termops.cmi kernel/type_errors.cmi kernel/typeops.cmi \ - lib/util.cmi toplevel/vernacexpr.cmo toplevel/record.cmi + library/declare.cmi kernel/entries.cmi kernel/environ.cmi \ + pretyping/evd.cmi library/global.cmi kernel/indtypes.cmi \ + kernel/inductive.cmi pretyping/inductiveops.cmi library/libnames.cmi \ + library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/options.cmi \ + lib/pp.cmi parsing/printer.cmi pretyping/recordops.cmi \ + kernel/safe_typing.cmi kernel/term.cmi pretyping/termops.cmi \ + kernel/type_errors.cmi lib/util.cmi toplevel/vernacexpr.cmo \ + toplevel/record.cmi toplevel/record.cmx: parsing/ast.cmx parsing/astterm.cmx toplevel/class.cmx \ toplevel/command.cmx parsing/coqast.cmx kernel/declarations.cmx \ - library/declare.cmx kernel/environ.cmx pretyping/evd.cmx \ - library/global.cmx kernel/indtypes.cmx kernel/inductive.cmx \ - pretyping/inductiveops.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/options.cmx lib/pp.cmx parsing/printer.cmx \ - pretyping/recordops.cmx kernel/safe_typing.cmx kernel/term.cmx \ - pretyping/termops.cmx kernel/type_errors.cmx kernel/typeops.cmx \ - lib/util.cmx toplevel/vernacexpr.cmx toplevel/record.cmi + library/declare.cmx kernel/entries.cmx kernel/environ.cmx \ + pretyping/evd.cmx library/global.cmx kernel/indtypes.cmx \ + kernel/inductive.cmx pretyping/inductiveops.cmx library/libnames.cmx \ + library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/options.cmx \ + lib/pp.cmx parsing/printer.cmx pretyping/recordops.cmx \ + kernel/safe_typing.cmx kernel/term.cmx pretyping/termops.cmx \ + kernel/type_errors.cmx lib/util.cmx toplevel/vernacexpr.cmx \ + toplevel/record.cmi toplevel/recordobj.cmo: pretyping/classops.cmi library/declare.cmi \ kernel/environ.cmi library/global.cmi pretyping/instantiate.cmi \ - library/lib.cmi library/nameops.cmi kernel/names.cmi library/nametab.cmi \ - lib/pp.cmi pretyping/recordops.cmi kernel/term.cmi pretyping/termops.cmi \ + library/lib.cmi library/libnames.cmi library/nameops.cmi kernel/names.cmi \ + library/nametab.cmi lib/pp.cmi pretyping/recordops.cmi kernel/term.cmi \ lib/util.cmi toplevel/recordobj.cmi toplevel/recordobj.cmx: pretyping/classops.cmx library/declare.cmx \ kernel/environ.cmx library/global.cmx pretyping/instantiate.cmx \ - library/lib.cmx library/nameops.cmx kernel/names.cmx library/nametab.cmx \ - lib/pp.cmx pretyping/recordops.cmx kernel/term.cmx pretyping/termops.cmx \ + library/lib.cmx library/libnames.cmx library/nameops.cmx kernel/names.cmx \ + library/nametab.cmx lib/pp.cmx pretyping/recordops.cmx kernel/term.cmx \ lib/util.cmx toplevel/recordobj.cmi toplevel/toplevel.cmo: parsing/ast.cmi toplevel/cerrors.cmi library/lib.cmi \ toplevel/mltop.cmi kernel/names.cmi lib/options.cmi parsing/pcoq.cmi \ @@ -1722,68 +1850,72 @@ toplevel/toplevel.cmx: parsing/ast.cmx toplevel/cerrors.cmx library/lib.cmx \ toplevel/vernac.cmx toplevel/vernacexpr.cmx toplevel/toplevel.cmi toplevel/usage.cmo: config/coq_config.cmi toplevel/usage.cmi toplevel/usage.cmx: config/coq_config.cmx toplevel/usage.cmi -toplevel/vernacentries.cmo: parsing/ast.cmi parsing/astterm.cmi \ - tactics/auto.cmi toplevel/class.cmi pretyping/classops.cmi \ - toplevel/command.cmi parsing/coqast.cmi library/declare.cmi \ - tactics/dhyp.cmi toplevel/discharge.cmi kernel/environ.cmi \ +toplevel/vernacentries.cmo: parsing/ast.cmi parsing/astmod.cmi \ + parsing/astterm.cmi tactics/auto.cmi toplevel/class.cmi \ + pretyping/classops.cmi toplevel/command.cmi parsing/coqast.cmi \ + library/declare.cmi library/declaremods.cmi tactics/dhyp.cmi \ + toplevel/discharge.cmi kernel/entries.cmi kernel/environ.cmi \ pretyping/evarutil.cmi pretyping/evd.cmi library/global.cmi \ library/goptions.cmi library/impargs.cmi pretyping/inductiveops.cmi \ - library/lib.cmi library/library.cmi toplevel/metasyntax.cmi \ - toplevel/mltop.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/options.cmi proofs/pfedit.cmi lib/pp.cmi \ - lib/pp_control.cmi parsing/prettyp.cmi parsing/printer.cmi \ + library/lib.cmi library/libnames.cmi library/library.cmi \ + toplevel/metasyntax.cmi toplevel/mltop.cmi library/nameops.cmi \ + kernel/names.cmi library/nametab.cmi lib/options.cmi proofs/pfedit.cmi \ + lib/pp.cmi lib/pp_control.cmi parsing/prettyp.cmi parsing/printer.cmi \ proofs/proof_trees.cmi proofs/proof_type.cmi toplevel/record.cmi \ toplevel/recordobj.cmi proofs/refiner.cmi kernel/safe_typing.cmi \ - parsing/search.cmi library/states.cmi pretyping/syntax_def.cmi \ - lib/system.cmi tactics/tacinterp.cmi proofs/tacmach.cmi \ - pretyping/tacred.cmi proofs/tactic_debug.cmi tactics/tactics.cmi \ - kernel/term.cmi parsing/termast.cmi kernel/typeops.cmi kernel/univ.cmi \ - lib/util.cmi toplevel/vernacexpr.cmo toplevel/vernacinterp.cmi \ + parsing/search.cmi library/states.cmi lib/system.cmi \ + tactics/tacinterp.cmi proofs/tacmach.cmi pretyping/tacred.cmi \ + proofs/tactic_debug.cmi tactics/tactics.cmi kernel/term.cmi \ + parsing/termast.cmi kernel/typeops.cmi kernel/univ.cmi lib/util.cmi \ + toplevel/vernacexpr.cmo toplevel/vernacinterp.cmi \ toplevel/vernacentries.cmi -toplevel/vernacentries.cmx: parsing/ast.cmx parsing/astterm.cmx \ - tactics/auto.cmx toplevel/class.cmx pretyping/classops.cmx \ - toplevel/command.cmx parsing/coqast.cmx library/declare.cmx \ - tactics/dhyp.cmx toplevel/discharge.cmx kernel/environ.cmx \ +toplevel/vernacentries.cmx: parsing/ast.cmx parsing/astmod.cmx \ + parsing/astterm.cmx tactics/auto.cmx toplevel/class.cmx \ + pretyping/classops.cmx toplevel/command.cmx parsing/coqast.cmx \ + library/declare.cmx library/declaremods.cmx tactics/dhyp.cmx \ + toplevel/discharge.cmx kernel/entries.cmx kernel/environ.cmx \ pretyping/evarutil.cmx pretyping/evd.cmx library/global.cmx \ library/goptions.cmx library/impargs.cmx pretyping/inductiveops.cmx \ - library/lib.cmx library/library.cmx toplevel/metasyntax.cmx \ - toplevel/mltop.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/options.cmx proofs/pfedit.cmx lib/pp.cmx \ - lib/pp_control.cmx parsing/prettyp.cmx parsing/printer.cmx \ + library/lib.cmx library/libnames.cmx library/library.cmx \ + toplevel/metasyntax.cmx toplevel/mltop.cmx library/nameops.cmx \ + kernel/names.cmx library/nametab.cmx lib/options.cmx proofs/pfedit.cmx \ + lib/pp.cmx lib/pp_control.cmx parsing/prettyp.cmx parsing/printer.cmx \ proofs/proof_trees.cmx proofs/proof_type.cmx toplevel/record.cmx \ toplevel/recordobj.cmx proofs/refiner.cmx kernel/safe_typing.cmx \ - parsing/search.cmx library/states.cmx pretyping/syntax_def.cmx \ - lib/system.cmx tactics/tacinterp.cmx proofs/tacmach.cmx \ - pretyping/tacred.cmx proofs/tactic_debug.cmx tactics/tactics.cmx \ - kernel/term.cmx parsing/termast.cmx kernel/typeops.cmx kernel/univ.cmx \ - lib/util.cmx toplevel/vernacexpr.cmx toplevel/vernacinterp.cmx \ + parsing/search.cmx library/states.cmx lib/system.cmx \ + tactics/tacinterp.cmx proofs/tacmach.cmx pretyping/tacred.cmx \ + proofs/tactic_debug.cmx tactics/tactics.cmx kernel/term.cmx \ + parsing/termast.cmx kernel/typeops.cmx kernel/univ.cmx lib/util.cmx \ + toplevel/vernacexpr.cmx toplevel/vernacinterp.cmx \ toplevel/vernacentries.cmi toplevel/vernacexpr.cmo: parsing/ast.cmi parsing/coqast.cmi \ parsing/extend.cmi parsing/genarg.cmi library/goptions.cmi \ - kernel/names.cmi library/nametab.cmi proofs/proof_type.cmi \ - proofs/tacexpr.cmo lib/util.cmi + library/libnames.cmi kernel/names.cmi library/nametab.cmi \ + proofs/proof_type.cmi proofs/tacexpr.cmo lib/util.cmi toplevel/vernacexpr.cmx: parsing/ast.cmx parsing/coqast.cmx \ parsing/extend.cmx parsing/genarg.cmx library/goptions.cmx \ - kernel/names.cmx library/nametab.cmx proofs/proof_type.cmx \ - proofs/tacexpr.cmx lib/util.cmx + library/libnames.cmx kernel/names.cmx library/nametab.cmx \ + proofs/proof_type.cmx proofs/tacexpr.cmx lib/util.cmx toplevel/vernacinterp.cmo: parsing/ast.cmi parsing/coqast.cmi \ - parsing/extend.cmi toplevel/himsg.cmi kernel/names.cmi lib/options.cmi \ - lib/pp.cmi proofs/proof_type.cmi proofs/tacexpr.cmo tactics/tacinterp.cmi \ - lib/util.cmi toplevel/vernacexpr.cmo toplevel/vernacinterp.cmi + parsing/extend.cmi toplevel/himsg.cmi library/libnames.cmi \ + kernel/names.cmi lib/options.cmi lib/pp.cmi proofs/proof_type.cmi \ + proofs/tacexpr.cmo tactics/tacinterp.cmi lib/util.cmi \ + toplevel/vernacexpr.cmo toplevel/vernacinterp.cmi toplevel/vernacinterp.cmx: parsing/ast.cmx parsing/coqast.cmx \ - parsing/extend.cmx toplevel/himsg.cmx kernel/names.cmx lib/options.cmx \ - lib/pp.cmx proofs/proof_type.cmx proofs/tacexpr.cmx tactics/tacinterp.cmx \ - lib/util.cmx toplevel/vernacexpr.cmx toplevel/vernacinterp.cmi + parsing/extend.cmx toplevel/himsg.cmx library/libnames.cmx \ + kernel/names.cmx lib/options.cmx lib/pp.cmx proofs/proof_type.cmx \ + proofs/tacexpr.cmx tactics/tacinterp.cmx lib/util.cmx \ + toplevel/vernacexpr.cmx toplevel/vernacinterp.cmi toplevel/vernac.cmo: parsing/ast.cmi parsing/coqast.cmi library/lib.cmi \ - library/library.cmi library/nameops.cmi kernel/names.cmi lib/options.cmi \ - parsing/pcoq.cmi lib/pp.cmi library/states.cmi lib/system.cmi \ - pretyping/termops.cmi lib/util.cmi toplevel/vernacentries.cmi \ - toplevel/vernacexpr.cmo toplevel/vernacinterp.cmi toplevel/vernac.cmi + library/library.cmi kernel/names.cmi lib/options.cmi parsing/pcoq.cmi \ + lib/pp.cmi library/states.cmi lib/system.cmi lib/util.cmi \ + toplevel/vernacentries.cmi toplevel/vernacexpr.cmo \ + toplevel/vernacinterp.cmi toplevel/vernac.cmi toplevel/vernac.cmx: parsing/ast.cmx parsing/coqast.cmx library/lib.cmx \ - library/library.cmx library/nameops.cmx kernel/names.cmx lib/options.cmx \ - parsing/pcoq.cmx lib/pp.cmx library/states.cmx lib/system.cmx \ - pretyping/termops.cmx lib/util.cmx toplevel/vernacentries.cmx \ - toplevel/vernacexpr.cmx toplevel/vernacinterp.cmx toplevel/vernac.cmi + library/library.cmx kernel/names.cmx lib/options.cmx parsing/pcoq.cmx \ + lib/pp.cmx library/states.cmx lib/system.cmx lib/util.cmx \ + toplevel/vernacentries.cmx toplevel/vernacexpr.cmx \ + toplevel/vernacinterp.cmx toplevel/vernac.cmi contrib/correctness/pcicenv.cmo: library/global.cmi kernel/names.cmi \ contrib/correctness/past.cmi contrib/correctness/penv.cmi \ contrib/correctness/pmisc.cmi contrib/correctness/pmonad.cmi \ @@ -2029,31 +2161,31 @@ contrib/correctness/pwp.cmx: kernel/environ.cmx library/global.cmx \ pretyping/reductionops.cmx kernel/term.cmx pretyping/termops.cmx \ lib/util.cmx contrib/correctness/pwp.cmi contrib/extraction/common.cmo: library/declare.cmi \ - contrib/extraction/extraction.cmi library/global.cmi \ - contrib/extraction/haskell.cmi contrib/extraction/miniml.cmi \ + contrib/extraction/extraction.cmi contrib/extraction/haskell.cmi \ + library/libnames.cmi contrib/extraction/miniml.cmi \ contrib/extraction/mlutil.cmi library/nameops.cmi kernel/names.cmi \ library/nametab.cmi contrib/extraction/ocaml.cmi lib/pp.cmi \ lib/pp_control.cmi parsing/printer.cmi contrib/extraction/scheme.cmi \ - contrib/extraction/table.cmi pretyping/termops.cmi lib/util.cmi \ - contrib/extraction/common.cmi + contrib/extraction/table.cmi lib/util.cmi contrib/extraction/common.cmi contrib/extraction/common.cmx: library/declare.cmx \ - contrib/extraction/extraction.cmx library/global.cmx \ - contrib/extraction/haskell.cmx contrib/extraction/miniml.cmi \ + contrib/extraction/extraction.cmx contrib/extraction/haskell.cmx \ + library/libnames.cmx contrib/extraction/miniml.cmi \ contrib/extraction/mlutil.cmx library/nameops.cmx kernel/names.cmx \ library/nametab.cmx contrib/extraction/ocaml.cmx lib/pp.cmx \ lib/pp_control.cmx parsing/printer.cmx contrib/extraction/scheme.cmx \ - contrib/extraction/table.cmx pretyping/termops.cmx lib/util.cmx \ - contrib/extraction/common.cmi + contrib/extraction/table.cmx lib/util.cmx contrib/extraction/common.cmi contrib/extraction/extract_env.cmo: contrib/extraction/common.cmi \ - library/declare.cmi contrib/extraction/extraction.cmi library/lib.cmi \ - library/libobject.cmi library/library.cmi contrib/extraction/miniml.cmi \ + library/declare.cmi library/declaremods.cmi \ + contrib/extraction/extraction.cmi library/lib.cmi library/libnames.cmi \ + library/libobject.cmi contrib/extraction/miniml.cmi \ contrib/extraction/mlutil.cmi library/nameops.cmi kernel/names.cmi \ library/nametab.cmi lib/pp.cmi contrib/extraction/table.cmi \ kernel/term.cmi lib/util.cmi toplevel/vernacinterp.cmi \ contrib/extraction/extract_env.cmi contrib/extraction/extract_env.cmx: contrib/extraction/common.cmx \ - library/declare.cmx contrib/extraction/extraction.cmx library/lib.cmx \ - library/libobject.cmx library/library.cmx contrib/extraction/miniml.cmi \ + library/declare.cmx library/declaremods.cmx \ + contrib/extraction/extraction.cmx library/lib.cmx library/libnames.cmx \ + library/libobject.cmx contrib/extraction/miniml.cmi \ contrib/extraction/mlutil.cmx library/nameops.cmx kernel/names.cmx \ library/nametab.cmx lib/pp.cmx contrib/extraction/table.cmx \ kernel/term.cmx lib/util.cmx toplevel/vernacinterp.cmx \ @@ -2061,19 +2193,21 @@ contrib/extraction/extract_env.cmx: contrib/extraction/common.cmx \ contrib/extraction/extraction.cmo: kernel/closure.cmi kernel/declarations.cmi \ kernel/environ.cmi pretyping/evd.cmi library/global.cmi lib/gmap.cmi \ kernel/inductive.cmi pretyping/inductiveops.cmi pretyping/instantiate.cmi \ - contrib/extraction/miniml.cmi contrib/extraction/mlutil.cmi \ - library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/options.cmi \ - lib/pp.cmi pretyping/reductionops.cmi pretyping/retyping.cmi \ - library/summary.cmi contrib/extraction/table.cmi kernel/term.cmi \ - pretyping/termops.cmi lib/util.cmi contrib/extraction/extraction.cmi + library/libnames.cmi contrib/extraction/miniml.cmi \ + contrib/extraction/mlutil.cmi library/nameops.cmi kernel/names.cmi \ + library/nametab.cmi lib/options.cmi lib/pp.cmi pretyping/reductionops.cmi \ + pretyping/retyping.cmi library/summary.cmi contrib/extraction/table.cmi \ + kernel/term.cmi pretyping/termops.cmi lib/util.cmi \ + contrib/extraction/extraction.cmi contrib/extraction/extraction.cmx: kernel/closure.cmx kernel/declarations.cmx \ kernel/environ.cmx pretyping/evd.cmx library/global.cmx lib/gmap.cmx \ kernel/inductive.cmx pretyping/inductiveops.cmx pretyping/instantiate.cmx \ - contrib/extraction/miniml.cmi contrib/extraction/mlutil.cmx \ - library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/options.cmx \ - lib/pp.cmx pretyping/reductionops.cmx pretyping/retyping.cmx \ - library/summary.cmx contrib/extraction/table.cmx kernel/term.cmx \ - pretyping/termops.cmx lib/util.cmx contrib/extraction/extraction.cmi + library/libnames.cmx contrib/extraction/miniml.cmi \ + contrib/extraction/mlutil.cmx library/nameops.cmx kernel/names.cmx \ + library/nametab.cmx lib/options.cmx lib/pp.cmx pretyping/reductionops.cmx \ + pretyping/retyping.cmx library/summary.cmx contrib/extraction/table.cmx \ + kernel/term.cmx pretyping/termops.cmx lib/util.cmx \ + contrib/extraction/extraction.cmi contrib/extraction/g_extraction.cmo: toplevel/cerrors.cmi \ parsing/egrammar.cmi parsing/extend.cmi \ contrib/extraction/extract_env.cmi parsing/genarg.cmi parsing/pcoq.cmi \ @@ -2092,16 +2226,14 @@ contrib/extraction/haskell.cmx: contrib/extraction/miniml.cmi \ contrib/extraction/mlutil.cmx library/nameops.cmx kernel/names.cmx \ library/nametab.cmx contrib/extraction/ocaml.cmx lib/options.cmx \ lib/pp.cmx kernel/term.cmx lib/util.cmx contrib/extraction/haskell.cmi -contrib/extraction/mlutil.cmo: kernel/declarations.cmi \ - contrib/extraction/miniml.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/options.cmi lib/pp.cmi \ - contrib/extraction/table.cmi kernel/term.cmi lib/util.cmi \ - contrib/extraction/mlutil.cmi -contrib/extraction/mlutil.cmx: kernel/declarations.cmx \ - contrib/extraction/miniml.cmi library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/options.cmx lib/pp.cmx \ - contrib/extraction/table.cmx kernel/term.cmx lib/util.cmx \ - contrib/extraction/mlutil.cmi +contrib/extraction/mlutil.cmo: kernel/declarations.cmi library/libnames.cmi \ + contrib/extraction/miniml.cmi kernel/names.cmi library/nametab.cmi \ + lib/options.cmi lib/pp.cmi contrib/extraction/table.cmi kernel/term.cmi \ + lib/util.cmi contrib/extraction/mlutil.cmi +contrib/extraction/mlutil.cmx: kernel/declarations.cmx library/libnames.cmx \ + contrib/extraction/miniml.cmi kernel/names.cmx library/nametab.cmx \ + lib/options.cmx lib/pp.cmx contrib/extraction/table.cmx kernel/term.cmx \ + lib/util.cmx contrib/extraction/mlutil.cmi contrib/extraction/ocaml.cmo: contrib/extraction/miniml.cmi \ contrib/extraction/mlutil.cmi library/nameops.cmi kernel/names.cmi \ library/nametab.cmi lib/options.cmi lib/pp.cmi \ @@ -2124,22 +2256,22 @@ contrib/extraction/scheme.cmx: contrib/extraction/miniml.cmi \ contrib/extraction/scheme.cmi contrib/extraction/table.cmo: kernel/declarations.cmi kernel/environ.cmi \ parsing/extend.cmi library/global.cmi library/goptions.cmi \ - library/lib.cmi library/libobject.cmi kernel/names.cmi \ - library/nametab.cmi lib/pp.cmi parsing/printer.cmi kernel/reduction.cmi \ - library/summary.cmi kernel/term.cmi lib/util.cmi \ + library/lib.cmi library/libnames.cmi library/libobject.cmi \ + kernel/names.cmi library/nametab.cmi lib/pp.cmi parsing/printer.cmi \ + kernel/reduction.cmi library/summary.cmi kernel/term.cmi lib/util.cmi \ toplevel/vernacinterp.cmi contrib/extraction/table.cmi contrib/extraction/table.cmx: kernel/declarations.cmx kernel/environ.cmx \ parsing/extend.cmx library/global.cmx library/goptions.cmx \ - library/lib.cmx library/libobject.cmx kernel/names.cmx \ - library/nametab.cmx lib/pp.cmx parsing/printer.cmx kernel/reduction.cmx \ - library/summary.cmx kernel/term.cmx lib/util.cmx \ + library/lib.cmx library/libnames.cmx library/libobject.cmx \ + kernel/names.cmx library/nametab.cmx lib/pp.cmx parsing/printer.cmx \ + kernel/reduction.cmx library/summary.cmx kernel/term.cmx lib/util.cmx \ toplevel/vernacinterp.cmx contrib/extraction/table.cmi contrib/field/field.cmo: parsing/astterm.cmi toplevel/cerrors.cmi \ parsing/coqast.cmi library/declare.cmi parsing/egrammar.cmi \ pretyping/evd.cmi parsing/extend.cmi parsing/genarg.cmi \ - library/global.cmi library/lib.cmi library/libobject.cmi \ - library/library.cmi kernel/names.cmi library/nametab.cmi parsing/pcoq.cmi \ - lib/pp.cmi parsing/pptactic.cmi proofs/proof_type.cmi \ + library/global.cmi lib/gmap.cmi library/lib.cmi library/libnames.cmi \ + library/libobject.cmi library/library.cmi kernel/names.cmi \ + parsing/pcoq.cmi lib/pp.cmi parsing/pptactic.cmi proofs/proof_type.cmi \ contrib/ring/quote.cmo pretyping/rawterm.cmi pretyping/reductionops.cmi \ proofs/refiner.cmi contrib/ring/ring.cmo library/summary.cmi \ proofs/tacexpr.cmo tactics/tacinterp.cmi proofs/tacmach.cmi \ @@ -2148,26 +2280,28 @@ contrib/field/field.cmo: parsing/astterm.cmi toplevel/cerrors.cmi \ contrib/field/field.cmx: parsing/astterm.cmx toplevel/cerrors.cmx \ parsing/coqast.cmx library/declare.cmx parsing/egrammar.cmx \ pretyping/evd.cmx parsing/extend.cmx parsing/genarg.cmx \ - library/global.cmx library/lib.cmx library/libobject.cmx \ - library/library.cmx kernel/names.cmx library/nametab.cmx parsing/pcoq.cmx \ - lib/pp.cmx parsing/pptactic.cmx proofs/proof_type.cmx \ + library/global.cmx lib/gmap.cmx library/lib.cmx library/libnames.cmx \ + library/libobject.cmx library/library.cmx kernel/names.cmx \ + parsing/pcoq.cmx lib/pp.cmx parsing/pptactic.cmx proofs/proof_type.cmx \ contrib/ring/quote.cmx pretyping/rawterm.cmx pretyping/reductionops.cmx \ proofs/refiner.cmx contrib/ring/ring.cmx library/summary.cmx \ proofs/tacexpr.cmx tactics/tacinterp.cmx proofs/tacmach.cmx \ tactics/tacticals.cmx kernel/term.cmx pretyping/typing.cmx lib/util.cmx \ toplevel/vernacexpr.cmx toplevel/vernacinterp.cmx contrib/fourier/fourierR.cmo: parsing/astterm.cmi proofs/clenv.cmi \ - tactics/contradiction.cmi tactics/equality.cmi pretyping/evd.cmi \ - contrib/fourier/fourier.cmo library/global.cmi library/library.cmi \ - kernel/names.cmi parsing/pcoq.cmi contrib/ring/ring.cmo \ - proofs/tacmach.cmi tactics/tacticals.cmi tactics/tactics.cmi \ - kernel/term.cmi toplevel/vernacexpr.cmo + tactics/contradiction.cmi parsing/coqlib.cmi tactics/equality.cmi \ + pretyping/evd.cmi contrib/fourier/fourier.cmo library/global.cmi \ + library/libnames.cmi library/library.cmi kernel/names.cmi \ + parsing/pcoq.cmi contrib/ring/ring.cmo proofs/tacmach.cmi \ + tactics/tacticals.cmi tactics/tactics.cmi kernel/term.cmi \ + toplevel/vernacexpr.cmo contrib/fourier/fourierR.cmx: parsing/astterm.cmx proofs/clenv.cmx \ - tactics/contradiction.cmx tactics/equality.cmx pretyping/evd.cmx \ - contrib/fourier/fourier.cmx library/global.cmx library/library.cmx \ - kernel/names.cmx parsing/pcoq.cmx contrib/ring/ring.cmx \ - proofs/tacmach.cmx tactics/tacticals.cmx tactics/tactics.cmx \ - kernel/term.cmx toplevel/vernacexpr.cmx + tactics/contradiction.cmx parsing/coqlib.cmx tactics/equality.cmx \ + pretyping/evd.cmx contrib/fourier/fourier.cmx library/global.cmx \ + library/libnames.cmx library/library.cmx kernel/names.cmx \ + parsing/pcoq.cmx contrib/ring/ring.cmx proofs/tacmach.cmx \ + tactics/tacticals.cmx tactics/tactics.cmx kernel/term.cmx \ + toplevel/vernacexpr.cmx contrib/fourier/g_fourier.cmo: toplevel/cerrors.cmi parsing/egrammar.cmi \ contrib/fourier/fourierR.cmo parsing/pcoq.cmi lib/pp.cmi \ parsing/pptactic.cmi proofs/refiner.cmi @@ -2212,7 +2346,7 @@ contrib/interface/centaur.cmo: contrib/interface/ascent.cmi parsing/ast.cmi \ library/declare.cmi parsing/egrammar.cmi kernel/environ.cmi \ pretyping/evd.cmi parsing/extend.cmi parsing/genarg.cmi \ library/global.cmi contrib/interface/history.cmi library/lib.cmi \ - library/libobject.cmi library/library.cmi \ + library/libnames.cmi library/libobject.cmi library/library.cmi \ toplevel/line_oriented_parser.cmi toplevel/mltop.cmi \ contrib/interface/name_to_ast.cmi library/nameops.cmi kernel/names.cmi \ library/nametab.cmi contrib/interface/pbp.cmi parsing/pcoq.cmi \ @@ -2234,7 +2368,7 @@ contrib/interface/centaur.cmx: contrib/interface/ascent.cmi parsing/ast.cmx \ library/declare.cmx parsing/egrammar.cmx kernel/environ.cmx \ pretyping/evd.cmx parsing/extend.cmx parsing/genarg.cmx \ library/global.cmx contrib/interface/history.cmx library/lib.cmx \ - library/libobject.cmx library/library.cmx \ + library/libnames.cmx library/libobject.cmx library/library.cmx \ toplevel/line_oriented_parser.cmx toplevel/mltop.cmx \ contrib/interface/name_to_ast.cmx library/nameops.cmx kernel/names.cmx \ library/nametab.cmx contrib/interface/pbp.cmx parsing/pcoq.cmx \ @@ -2249,29 +2383,29 @@ contrib/interface/centaur.cmx: contrib/interface/ascent.cmi parsing/ast.cmx \ toplevel/vernacinterp.cmx contrib/interface/vtp.cmx \ contrib/interface/xlate.cmx contrib/interface/ctast.cmo: parsing/ast.cmi parsing/coqast.cmi lib/dyn.cmi \ - kernel/names.cmi + library/libnames.cmi kernel/names.cmi contrib/interface/ctast.cmx: parsing/ast.cmx parsing/coqast.cmx lib/dyn.cmx \ - kernel/names.cmx + library/libnames.cmx kernel/names.cmx contrib/interface/dad.cmo: parsing/astterm.cmi parsing/coqast.cmi \ contrib/interface/ctast.cmo kernel/environ.cmi pretyping/evd.cmi \ - parsing/genarg.cmi library/global.cmi kernel/names.cmi \ - library/nametab.cmi contrib/interface/paths.cmi pretyping/pattern.cmi \ - lib/pp.cmi proofs/proof_trees.cmi proofs/proof_type.cmi \ - pretyping/rawterm.cmi kernel/reduction.cmi proofs/tacexpr.cmo \ - proofs/tacmach.cmi tactics/tacticals.cmi tactics/tactics.cmi \ - kernel/term.cmi parsing/termast.cmi pretyping/typing.cmi lib/util.cmi \ - toplevel/vernacexpr.cmo toplevel/vernacinterp.cmi \ - contrib/interface/dad.cmi + parsing/genarg.cmi library/global.cmi library/libnames.cmi \ + kernel/names.cmi library/nametab.cmi contrib/interface/paths.cmi \ + pretyping/pattern.cmi lib/pp.cmi proofs/proof_trees.cmi \ + proofs/proof_type.cmi pretyping/rawterm.cmi kernel/reduction.cmi \ + proofs/tacexpr.cmo proofs/tacmach.cmi tactics/tacticals.cmi \ + tactics/tactics.cmi kernel/term.cmi parsing/termast.cmi \ + pretyping/typing.cmi lib/util.cmi toplevel/vernacexpr.cmo \ + toplevel/vernacinterp.cmi contrib/interface/dad.cmi contrib/interface/dad.cmx: parsing/astterm.cmx parsing/coqast.cmx \ contrib/interface/ctast.cmx kernel/environ.cmx pretyping/evd.cmx \ - parsing/genarg.cmx library/global.cmx kernel/names.cmx \ - library/nametab.cmx contrib/interface/paths.cmx pretyping/pattern.cmx \ - lib/pp.cmx proofs/proof_trees.cmx proofs/proof_type.cmx \ - pretyping/rawterm.cmx kernel/reduction.cmx proofs/tacexpr.cmx \ - proofs/tacmach.cmx tactics/tacticals.cmx tactics/tactics.cmx \ - kernel/term.cmx parsing/termast.cmx pretyping/typing.cmx lib/util.cmx \ - toplevel/vernacexpr.cmx toplevel/vernacinterp.cmx \ - contrib/interface/dad.cmi + parsing/genarg.cmx library/global.cmx library/libnames.cmx \ + kernel/names.cmx library/nametab.cmx contrib/interface/paths.cmx \ + pretyping/pattern.cmx lib/pp.cmx proofs/proof_trees.cmx \ + proofs/proof_type.cmx pretyping/rawterm.cmx kernel/reduction.cmx \ + proofs/tacexpr.cmx proofs/tacmach.cmx tactics/tacticals.cmx \ + tactics/tactics.cmx kernel/term.cmx parsing/termast.cmx \ + pretyping/typing.cmx lib/util.cmx toplevel/vernacexpr.cmx \ + toplevel/vernacinterp.cmx contrib/interface/dad.cmi contrib/interface/debug_tac.cmo: parsing/ast.cmi toplevel/cerrors.cmi \ parsing/coqast.cmi parsing/genarg.cmi lib/pp.cmi parsing/pptactic.cmi \ proofs/proof_trees.cmi proofs/proof_type.cmi proofs/refiner.cmi \ @@ -2291,29 +2425,31 @@ contrib/interface/line_parser.cmx: contrib/interface/line_parser.cmi contrib/interface/name_to_ast.cmo: parsing/ast.cmi pretyping/classops.cmi \ parsing/coqast.cmi kernel/declarations.cmi library/declare.cmi \ kernel/environ.cmi library/global.cmi library/impargs.cmi \ - kernel/inductive.cmi library/lib.cmi library/libobject.cmi \ - library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/pp.cmi \ - parsing/prettyp.cmi kernel/reduction.cmi kernel/sign.cmi \ - pretyping/syntax_def.cmi kernel/term.cmi parsing/termast.cmi lib/util.cmi \ + kernel/inductive.cmi library/lib.cmi library/libnames.cmi \ + library/libobject.cmi library/nameops.cmi kernel/names.cmi \ + library/nametab.cmi lib/pp.cmi parsing/prettyp.cmi kernel/reduction.cmi \ + kernel/sign.cmi kernel/term.cmi parsing/termast.cmi lib/util.cmi \ toplevel/vernacexpr.cmo contrib/interface/name_to_ast.cmi contrib/interface/name_to_ast.cmx: parsing/ast.cmx pretyping/classops.cmx \ parsing/coqast.cmx kernel/declarations.cmx library/declare.cmx \ kernel/environ.cmx library/global.cmx library/impargs.cmx \ - kernel/inductive.cmx library/lib.cmx library/libobject.cmx \ - library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/pp.cmx \ - parsing/prettyp.cmx kernel/reduction.cmx kernel/sign.cmx \ - pretyping/syntax_def.cmx kernel/term.cmx parsing/termast.cmx lib/util.cmx \ + kernel/inductive.cmx library/lib.cmx library/libnames.cmx \ + library/libobject.cmx library/nameops.cmx kernel/names.cmx \ + library/nametab.cmx lib/pp.cmx parsing/prettyp.cmx kernel/reduction.cmx \ + kernel/sign.cmx kernel/term.cmx parsing/termast.cmx lib/util.cmx \ toplevel/vernacexpr.cmx contrib/interface/name_to_ast.cmi contrib/interface/parse.cmo: contrib/interface/ascent.cmi \ toplevel/cerrors.cmi config/coq_config.cmi contrib/interface/ctast.cmo \ - parsing/esyntax.cmi library/libobject.cmi library/library.cmi \ + library/declaremods.cmi parsing/esyntax.cmi library/libnames.cmi \ + library/libobject.cmi library/library.cmi \ contrib/interface/line_parser.cmi library/nameops.cmi kernel/names.cmi \ library/nametab.cmi parsing/pcoq.cmi lib/pp.cmi lib/system.cmi \ lib/util.cmi toplevel/vernacentries.cmi toplevel/vernacexpr.cmo \ contrib/interface/vtp.cmi contrib/interface/xlate.cmi contrib/interface/parse.cmx: contrib/interface/ascent.cmi \ toplevel/cerrors.cmx config/coq_config.cmx contrib/interface/ctast.cmx \ - parsing/esyntax.cmx library/libobject.cmx library/library.cmx \ + library/declaremods.cmx parsing/esyntax.cmx library/libnames.cmx \ + library/libobject.cmx library/library.cmx \ contrib/interface/line_parser.cmx library/nameops.cmx kernel/names.cmx \ library/nametab.cmx parsing/pcoq.cmx lib/pp.cmx lib/system.cmx \ lib/util.cmx toplevel/vernacentries.cmx toplevel/vernacexpr.cmx \ @@ -2323,8 +2459,8 @@ contrib/interface/paths.cmx: contrib/interface/paths.cmi contrib/interface/pbp.cmo: parsing/astterm.cmi parsing/coqast.cmi \ parsing/coqlib.cmi contrib/interface/ctast.cmo library/declare.cmi \ kernel/environ.cmi pretyping/evd.cmi library/global.cmi \ - tactics/hipattern.cmi proofs/logic.cmi kernel/names.cmi \ - library/nametab.cmi pretyping/pattern.cmi lib/pp.cmi \ + tactics/hipattern.cmi library/libnames.cmi proofs/logic.cmi \ + kernel/names.cmi library/nametab.cmi pretyping/pattern.cmi lib/pp.cmi \ pretyping/pretyping.cmi proofs/proof_trees.cmi proofs/proof_type.cmi \ pretyping/rawterm.cmi kernel/reduction.cmi proofs/tacexpr.cmo \ tactics/tacinterp.cmi proofs/tacmach.cmi tactics/tacticals.cmi \ @@ -2333,8 +2469,8 @@ contrib/interface/pbp.cmo: parsing/astterm.cmi parsing/coqast.cmi \ contrib/interface/pbp.cmx: parsing/astterm.cmx parsing/coqast.cmx \ parsing/coqlib.cmx contrib/interface/ctast.cmx library/declare.cmx \ kernel/environ.cmx pretyping/evd.cmx library/global.cmx \ - tactics/hipattern.cmx proofs/logic.cmx kernel/names.cmx \ - library/nametab.cmx pretyping/pattern.cmx lib/pp.cmx \ + tactics/hipattern.cmx library/libnames.cmx proofs/logic.cmx \ + kernel/names.cmx library/nametab.cmx pretyping/pattern.cmx lib/pp.cmx \ pretyping/pretyping.cmx proofs/proof_trees.cmx proofs/proof_type.cmx \ pretyping/rawterm.cmx kernel/reduction.cmx proofs/tacexpr.cmx \ tactics/tacinterp.cmx proofs/tacmach.cmx tactics/tacticals.cmx \ @@ -2352,9 +2488,9 @@ contrib/interface/showproof.cmo: parsing/ast.cmi parsing/astterm.cmi \ proofs/clenv.cmi parsing/coqast.cmi kernel/declarations.cmi \ kernel/environ.cmi pretyping/evd.cmi parsing/genarg.cmi \ library/global.cmi kernel/inductive.cmi pretyping/inductiveops.cmi \ - library/nameops.cmi kernel/names.cmi proofs/pfedit.cmi lib/pp.cmi \ - parsing/printer.cmi proofs/proof_trees.cmi proofs/proof_type.cmi \ - pretyping/rawterm.cmi pretyping/reductionops.cmi \ + library/libnames.cmi library/nameops.cmi kernel/names.cmi \ + proofs/pfedit.cmi lib/pp.cmi parsing/printer.cmi proofs/proof_trees.cmi \ + proofs/proof_type.cmi pretyping/rawterm.cmi pretyping/reductionops.cmi \ contrib/interface/showproof_ct.cmo kernel/sign.cmi proofs/tacexpr.cmo \ proofs/tacmach.cmi kernel/term.cmi parsing/termast.cmi \ pretyping/termops.cmi contrib/interface/translate.cmi \ @@ -2364,9 +2500,9 @@ contrib/interface/showproof.cmx: parsing/ast.cmx parsing/astterm.cmx \ proofs/clenv.cmx parsing/coqast.cmx kernel/declarations.cmx \ kernel/environ.cmx pretyping/evd.cmx parsing/genarg.cmx \ library/global.cmx kernel/inductive.cmx pretyping/inductiveops.cmx \ - library/nameops.cmx kernel/names.cmx proofs/pfedit.cmx lib/pp.cmx \ - parsing/printer.cmx proofs/proof_trees.cmx proofs/proof_type.cmx \ - pretyping/rawterm.cmx pretyping/reductionops.cmx \ + library/libnames.cmx library/nameops.cmx kernel/names.cmx \ + proofs/pfedit.cmx lib/pp.cmx parsing/printer.cmx proofs/proof_trees.cmx \ + proofs/proof_type.cmx pretyping/rawterm.cmx pretyping/reductionops.cmx \ contrib/interface/showproof_ct.cmx kernel/sign.cmx proofs/tacexpr.cmx \ proofs/tacmach.cmx kernel/term.cmx parsing/termast.cmx \ pretyping/termops.cmx contrib/interface/translate.cmx \ @@ -2392,16 +2528,14 @@ contrib/interface/vtp.cmx: contrib/interface/ascent.cmi \ contrib/interface/vtp.cmi contrib/interface/xlate.cmo: contrib/interface/ascent.cmi parsing/ast.cmi \ parsing/astterm.cmi parsing/coqast.cmi contrib/interface/ctast.cmo \ - tactics/extraargs.cmi parsing/genarg.cmi library/nameops.cmi \ - kernel/names.cmi library/nametab.cmi pretyping/rawterm.cmi \ - proofs/tacexpr.cmo lib/util.cmi toplevel/vernacexpr.cmo \ - contrib/interface/xlate.cmi + tactics/extraargs.cmi parsing/genarg.cmi library/libnames.cmi \ + kernel/names.cmi pretyping/rawterm.cmi proofs/tacexpr.cmo lib/util.cmi \ + toplevel/vernacexpr.cmo contrib/interface/xlate.cmi contrib/interface/xlate.cmx: contrib/interface/ascent.cmi parsing/ast.cmx \ parsing/astterm.cmx parsing/coqast.cmx contrib/interface/ctast.cmx \ - tactics/extraargs.cmx parsing/genarg.cmx library/nameops.cmx \ - kernel/names.cmx library/nametab.cmx pretyping/rawterm.cmx \ - proofs/tacexpr.cmx lib/util.cmx toplevel/vernacexpr.cmx \ - contrib/interface/xlate.cmi + tactics/extraargs.cmx parsing/genarg.cmx library/libnames.cmx \ + kernel/names.cmx pretyping/rawterm.cmx proofs/tacexpr.cmx lib/util.cmx \ + toplevel/vernacexpr.cmx contrib/interface/xlate.cmi contrib/jprover/jall.cmo: contrib/jprover/jlogic.cmi \ contrib/jprover/jterm.cmi contrib/jprover/jtunify.cmi \ contrib/jprover/opname.cmi contrib/jprover/jall.cmi @@ -2441,8 +2575,8 @@ contrib/jprover/opname.cmx: contrib/jprover/opname.cmi contrib/omega/coq_omega.cmo: parsing/ast.cmi proofs/clenv.cmi \ kernel/closure.cmi tactics/contradiction.cmi parsing/coqlib.cmi \ kernel/declarations.cmi library/declare.cmi kernel/environ.cmi \ - tactics/equality.cmi proofs/evar_refiner.cmi library/global.cmi \ - library/goptions.cmi kernel/inductive.cmi library/library.cmi \ + tactics/equality.cmi proofs/evar_refiner.cmi library/goptions.cmi \ + kernel/inductive.cmi library/libnames.cmi library/library.cmi \ proofs/logic.cmi library/nameops.cmi kernel/names.cmi library/nametab.cmi \ contrib/omega/omega.cmo lib/pp.cmi parsing/printer.cmi \ proofs/proof_type.cmi pretyping/rawterm.cmi kernel/reduction.cmi \ @@ -2451,8 +2585,8 @@ contrib/omega/coq_omega.cmo: parsing/ast.cmi proofs/clenv.cmi \ contrib/omega/coq_omega.cmx: parsing/ast.cmx proofs/clenv.cmx \ kernel/closure.cmx tactics/contradiction.cmx parsing/coqlib.cmx \ kernel/declarations.cmx library/declare.cmx kernel/environ.cmx \ - tactics/equality.cmx proofs/evar_refiner.cmx library/global.cmx \ - library/goptions.cmx kernel/inductive.cmx library/library.cmx \ + tactics/equality.cmx proofs/evar_refiner.cmx library/goptions.cmx \ + kernel/inductive.cmx library/libnames.cmx library/library.cmx \ proofs/logic.cmx library/nameops.cmx kernel/names.cmx library/nametab.cmx \ contrib/omega/omega.cmx lib/pp.cmx parsing/printer.cmx \ proofs/proof_type.cmx pretyping/rawterm.cmx kernel/reduction.cmx \ @@ -2481,45 +2615,43 @@ contrib/ring/g_ring.cmx: toplevel/cerrors.cmx parsing/egrammar.cmx \ contrib/ring/quote.cmx proofs/refiner.cmx contrib/ring/ring.cmx \ toplevel/vernacinterp.cmx contrib/ring/quote.cmo: library/declare.cmi kernel/environ.cmi \ - library/global.cmi pretyping/instantiate.cmi library/library.cmi \ - kernel/names.cmi library/nametab.cmi pretyping/pattern.cmi lib/pp.cmi \ + library/global.cmi pretyping/instantiate.cmi library/libnames.cmi \ + library/library.cmi kernel/names.cmi pretyping/pattern.cmi lib/pp.cmi \ proofs/proof_trees.cmi proofs/tacexpr.cmo proofs/tacmach.cmi \ tactics/tactics.cmi kernel/term.cmi pretyping/termops.cmi lib/util.cmi contrib/ring/quote.cmx: library/declare.cmx kernel/environ.cmx \ - library/global.cmx pretyping/instantiate.cmx library/library.cmx \ - kernel/names.cmx library/nametab.cmx pretyping/pattern.cmx lib/pp.cmx \ + library/global.cmx pretyping/instantiate.cmx library/libnames.cmx \ + library/library.cmx kernel/names.cmx pretyping/pattern.cmx lib/pp.cmx \ proofs/proof_trees.cmx proofs/tacexpr.cmx proofs/tacmach.cmx \ tactics/tactics.cmx kernel/term.cmx pretyping/termops.cmx lib/util.cmx contrib/ring/ring.cmo: parsing/astterm.cmi kernel/closure.cmi \ parsing/coqlib.cmi library/declare.cmi tactics/equality.cmi \ pretyping/evd.cmi library/global.cmi tactics/hiddentac.cmi \ - tactics/hipattern.cmi library/lib.cmi library/libobject.cmi \ - library/library.cmi library/nameops.cmi kernel/names.cmi \ - library/nametab.cmi lib/options.cmi pretyping/pattern.cmi lib/pp.cmi \ - parsing/printer.cmi proofs/proof_trees.cmi contrib/ring/quote.cmo \ - pretyping/reductionops.cmi tactics/setoid_replace.cmi library/summary.cmi \ - proofs/tacexpr.cmo proofs/tacmach.cmi pretyping/tacred.cmi \ - tactics/tacticals.cmi tactics/tactics.cmi kernel/term.cmi \ - pretyping/typing.cmi lib/util.cmi toplevel/vernacexpr.cmo \ - toplevel/vernacinterp.cmi + tactics/hipattern.cmi library/lib.cmi library/libnames.cmi \ + library/libobject.cmi library/library.cmi library/nameops.cmi \ + kernel/names.cmi library/nametab.cmi lib/options.cmi \ + pretyping/pattern.cmi lib/pp.cmi parsing/printer.cmi \ + proofs/proof_trees.cmi contrib/ring/quote.cmo pretyping/reductionops.cmi \ + tactics/setoid_replace.cmi library/summary.cmi proofs/tacexpr.cmo \ + proofs/tacmach.cmi pretyping/tacred.cmi tactics/tacticals.cmi \ + tactics/tactics.cmi kernel/term.cmi pretyping/typing.cmi lib/util.cmi \ + toplevel/vernacexpr.cmo toplevel/vernacinterp.cmi contrib/ring/ring.cmx: parsing/astterm.cmx kernel/closure.cmx \ parsing/coqlib.cmx library/declare.cmx tactics/equality.cmx \ pretyping/evd.cmx library/global.cmx tactics/hiddentac.cmx \ - tactics/hipattern.cmx library/lib.cmx library/libobject.cmx \ - library/library.cmx library/nameops.cmx kernel/names.cmx \ - library/nametab.cmx lib/options.cmx pretyping/pattern.cmx lib/pp.cmx \ - parsing/printer.cmx proofs/proof_trees.cmx contrib/ring/quote.cmx \ - pretyping/reductionops.cmx tactics/setoid_replace.cmx library/summary.cmx \ - proofs/tacexpr.cmx proofs/tacmach.cmx pretyping/tacred.cmx \ - tactics/tacticals.cmx tactics/tactics.cmx kernel/term.cmx \ - pretyping/typing.cmx lib/util.cmx toplevel/vernacexpr.cmx \ - toplevel/vernacinterp.cmx -contrib/romega/const_omega.cmo: library/declare.cmi library/global.cmi \ - kernel/names.cmi library/nametab.cmi kernel/term.cmi \ - pretyping/termops.cmi lib/util.cmi -contrib/romega/const_omega.cmx: library/declare.cmx library/global.cmx \ - kernel/names.cmx library/nametab.cmx kernel/term.cmx \ - pretyping/termops.cmx lib/util.cmx + tactics/hipattern.cmx library/lib.cmx library/libnames.cmx \ + library/libobject.cmx library/library.cmx library/nameops.cmx \ + kernel/names.cmx library/nametab.cmx lib/options.cmx \ + pretyping/pattern.cmx lib/pp.cmx parsing/printer.cmx \ + proofs/proof_trees.cmx contrib/ring/quote.cmx pretyping/reductionops.cmx \ + tactics/setoid_replace.cmx library/summary.cmx proofs/tacexpr.cmx \ + proofs/tacmach.cmx pretyping/tacred.cmx tactics/tacticals.cmx \ + tactics/tactics.cmx kernel/term.cmx pretyping/typing.cmx lib/util.cmx \ + toplevel/vernacexpr.cmx toplevel/vernacinterp.cmx +contrib/romega/const_omega.cmo: library/declare.cmi library/libnames.cmi \ + kernel/names.cmi library/nametab.cmi kernel/term.cmi lib/util.cmi +contrib/romega/const_omega.cmx: library/declare.cmx library/libnames.cmx \ + kernel/names.cmx library/nametab.cmx kernel/term.cmx lib/util.cmx contrib/romega/g_romega.cmo: toplevel/cerrors.cmi parsing/egrammar.cmi \ parsing/pcoq.cmi lib/pp.cmi parsing/pptactic.cmi proofs/refiner.cmi \ contrib/romega/refl_omega.cmo @@ -2543,7 +2675,8 @@ contrib/romega/refl_omega.cmx: parsing/ast.cmx tactics/auto.cmx \ proofs/tacmach.cmx tactics/tacticals.cmx tactics/tactics.cmx \ kernel/term.cmx lib/util.cmx contrib/xml/xmlcommand.cmo: kernel/declarations.cmi library/declare.cmi \ - kernel/environ.cmi pretyping/evd.cmi library/global.cmi library/lib.cmi \ + library/declaremods.cmi kernel/environ.cmi pretyping/evd.cmi \ + library/global.cmi library/lib.cmi library/libnames.cmi \ library/libobject.cmi library/library.cmi library/nameops.cmi \ kernel/names.cmi library/nametab.cmi proofs/pfedit.cmi \ proofs/proof_trees.cmi kernel/reduction.cmi pretyping/retyping.cmi \ @@ -2551,7 +2684,8 @@ contrib/xml/xmlcommand.cmo: kernel/declarations.cmi library/declare.cmi \ kernel/term.cmi lib/util.cmi contrib/xml/xml.cmi \ contrib/xml/xmlcommand.cmi contrib/xml/xmlcommand.cmx: kernel/declarations.cmx library/declare.cmx \ - kernel/environ.cmx pretyping/evd.cmx library/global.cmx library/lib.cmx \ + library/declaremods.cmx kernel/environ.cmx pretyping/evd.cmx \ + library/global.cmx library/lib.cmx library/libnames.cmx \ library/libobject.cmx library/library.cmx library/nameops.cmx \ kernel/names.cmx library/nametab.cmx proofs/pfedit.cmx \ proofs/proof_trees.cmx kernel/reduction.cmx pretyping/retyping.cmx \ @@ -2568,6 +2702,26 @@ contrib/xml/xmlentries.cmx: toplevel/cerrors.cmx parsing/egrammar.cmx \ toplevel/vernacinterp.cmx contrib/xml/xmlcommand.cmx contrib/xml/xml.cmo: contrib/xml/xml.cmi contrib/xml/xml.cmx: contrib/xml/xml.cmi +tmp/Discharge/declare.cmo: kernel/declarations.cmi kernel/entries.cmi \ + kernel/environ.cmi library/global.cmi library/impargs.cmi \ + kernel/indtypes.cmi kernel/inductive.cmi library/lib.cmi \ + library/libnames.cmi library/libobject.cmi library/library.cmi \ + library/nameops.cmi kernel/names.cmi library/nametab.cmi lib/pp.cmi \ + kernel/reduction.cmi kernel/safe_typing.cmi kernel/sign.cmi \ + library/summary.cmi kernel/term.cmi kernel/type_errors.cmi \ + kernel/typeops.cmi kernel/univ.cmi lib/util.cmi +tmp/Discharge/declare.cmx: kernel/declarations.cmx kernel/entries.cmx \ + kernel/environ.cmx library/global.cmx library/impargs.cmx \ + kernel/indtypes.cmx kernel/inductive.cmx library/lib.cmx \ + library/libnames.cmx library/libobject.cmx library/library.cmx \ + library/nameops.cmx kernel/names.cmx library/nametab.cmx lib/pp.cmx \ + kernel/reduction.cmx kernel/safe_typing.cmx kernel/sign.cmx \ + library/summary.cmx kernel/term.cmx kernel/type_errors.cmx \ + kernel/typeops.cmx kernel/univ.cmx lib/util.cmx +tmp/Discharge/libnames.cmo: kernel/names.cmi lib/pp.cmi lib/predicate.cmi \ + lib/util.cmi tmp/Discharge/libnames.cmi +tmp/Discharge/libnames.cmx: kernel/names.cmx lib/pp.cmx lib/predicate.cmx \ + lib/util.cmx tmp/Discharge/libnames.cmi tactics/tauto.cmo: parsing/grammar.cma kernel/names.cmo parsing/ast.cmo parsing/g_tactic.cmo parsing/g_ltac.cmo parsing/g_constr.cmo tactics/tauto.cmx: parsing/grammar.cma kernel/names.cmo parsing/ast.cmo parsing/g_tactic.cmo parsing/g_ltac.cmo parsing/g_constr.cmo tactics/eqdecide.cmo: parsing/grammar.cma @@ -2624,6 +2778,8 @@ parsing/g_cases.cmo: parsing/grammar.cma parsing/g_cases.cmx: parsing/grammar.cma parsing/g_constr.cmo: parsing/grammar.cma parsing/g_constr.cmx: parsing/grammar.cma +parsing/g_module.cmo: parsing/grammar.cma +parsing/g_module.cmx: parsing/grammar.cma parsing/g_tactic.cmo: parsing/grammar.cma parsing/g_tactic.cmx: parsing/grammar.cma parsing/g_ltac.cmo: parsing/grammar.cma diff --git a/.depend.camlp4 b/.depend.camlp4 index fcfc52d1a..91bd4ef06 100644 --- a/.depend.camlp4 +++ b/.depend.camlp4 @@ -26,6 +26,7 @@ parsing/g_vernac.ml: parsing/grammar.cma parsing/g_proofs.ml: parsing/grammar.cma parsing/g_cases.ml: parsing/grammar.cma parsing/g_constr.ml: parsing/grammar.cma +parsing/g_module.ml: parsing/grammar.cma parsing/g_tactic.ml: parsing/grammar.cma parsing/g_ltac.ml: parsing/grammar.cma parsing/argextend.ml: diff --git a/.depend.coq b/.depend.coq index 0f4703012..f8c411800 100644 --- a/.depend.coq +++ b/.depend.coq @@ -33,13 +33,15 @@ theories/Setoids/Setoid.vo: theories/Setoids/Setoid.v theories/Reals/Reals.vo: theories/Reals/Reals.v theories/Reals/Rdefinitions.vo theories/Reals/TypeSyntax.vo theories/Reals/Raxioms.vo theories/Reals/Rbase.vo theories/Reals/R_Ifp.vo theories/Reals/Rbasic_fun.vo theories/Reals/R_sqr.vo theories/Reals/R_sqrt.vo theories/Reals/Rsigma.vo theories/Reals/Rlimit.vo theories/Reals/Rfunctions.vo theories/Reals/Rtrigo.vo theories/Reals/Rtrigo_calc.vo theories/Reals/Rgeom.vo theories/Reals/Rderiv.vo theories/Reals/Rseries.vo theories/Reals/Binome.vo theories/Reals/Rcomplet.vo theories/Reals/Alembert_compl.vo theories/Reals/Rtrigo_fun.vo theories/Reals/Ranalysis.vo theories/Reals/DiscrR.vo theories/Reals/SplitRmult.vo theories/Reals/SplitAbsolu.vo theories/Reals/Rgeom.vo: theories/Reals/Rgeom.v theories/Reals/Rbase.vo theories/Reals/R_sqr.vo theories/Reals/Rtrigo.vo theories/Reals/R_sqrt.vo theories/Reals/Ranalysis.vo: theories/Reals/Ranalysis.v theories/Reals/Ranalysis1.vo theories/Reals/Ranalysis2.vo theories/Reals/Ranalysis3.vo theories/Reals/Ranalysis4.vo -theories/Reals/Ranalysis4.vo: theories/Reals/Ranalysis4.v theories/Reals/Rbase.vo theories/Reals/Rbasic_fun.vo theories/Reals/R_sqr.vo theories/Reals/Rlimit.vo theories/Reals/Rderiv.vo theories/Reals/DiscrR.vo theories/Reals/Rtrigo.vo theories/Reals/Ranalysis1.vo theories/Reals/R_sqrt.vo theories/Reals/Ranalysis2.vo theories/Reals/Ranalysis3.vo theories/Reals/Sqrt_reg.vo +theories/Reals/Ranalysis4.vo: theories/Reals/Ranalysis4.v theories/Reals/Rbase.vo theories/Reals/Rbasic_fun.vo theories/Reals/R_sqr.vo theories/Reals/Rlimit.vo theories/Reals/Rderiv.vo theories/Reals/DiscrR.vo theories/Reals/Rtrigo.vo theories/Reals/Ranalysis1.vo theories/Reals/R_sqrt.vo theories/Reals/Ranalysis2.vo theories/Reals/Ranalysis3.vo theories/Reals/Rtrigo_reg.vo theories/Reals/Sqrt_reg.vo theories/Reals/Sqrt_reg.vo: theories/Reals/Sqrt_reg.v theories/Reals/Rbase.vo theories/Reals/Rbasic_fun.vo theories/Reals/Rderiv.vo theories/Reals/R_sqr.vo theories/Reals/DiscrR.vo theories/Reals/Rtrigo.vo theories/Reals/Ranalysis1.vo theories/Reals/R_sqrt.vo theories/Reals/Ranalysis3.vo: theories/Reals/Ranalysis3.v theories/Reals/Rbase.vo theories/Reals/Rbasic_fun.vo theories/Reals/R_sqr.vo theories/Reals/Rlimit.vo theories/Reals/Rderiv.vo theories/Reals/DiscrR.vo theories/Reals/Rtrigo.vo theories/Reals/Ranalysis1.vo theories/Reals/Ranalysis2.vo theories/Reals/Ranalysis2.vo: theories/Reals/Ranalysis2.v theories/Reals/Rbase.vo theories/Reals/Rbasic_fun.vo theories/Reals/R_sqr.vo theories/Reals/Rlimit.vo theories/Reals/Rderiv.vo theories/Reals/DiscrR.vo theories/Reals/Rtrigo.vo theories/Reals/Ranalysis1.vo contrib/omega/Omega.vo theories/Reals/Rtrigo_calc.vo: theories/Reals/Rtrigo_calc.v theories/Reals/Rbase.vo theories/Reals/DiscrR.vo theories/Reals/Rseries.vo theories/Reals/R_sqr.vo theories/Reals/Rlimit.vo theories/Reals/Rtrigo.vo theories/Reals/R_sqrt.vo theories/Reals/R_sqrt.vo: theories/Reals/R_sqrt.v theories/Reals/Rbase.vo theories/Reals/DiscrR.vo theories/Reals/Rbasic_fun.vo theories/Reals/R_sqr.vo theories/Reals/Rsqrt_def.vo theories/Reals/Rsqrt_def.vo: theories/Reals/Rsqrt_def.v theories/Bool/Sumbool.vo theories/Reals/Rbase.vo theories/Reals/DiscrR.vo theories/Reals/Rseries.vo theories/Reals/Alembert.vo theories/Reals/Rcomplet.vo theories/Reals/AltSeries.vo theories/Reals/Rtrigo_alt.vo theories/Reals/Cv_prop.vo theories/Reals/Ranalysis1.vo +theories/Reals/Rtrigo_reg.vo: theories/Reals/Rtrigo_reg.v theories/Reals/Rbase.vo theories/Reals/DiscrR.vo theories/Reals/Rfunctions.vo theories/Reals/Rseries.vo theories/Reals/Alembert.vo theories/Reals/Binome.vo theories/Reals/Rtrigo.vo theories/Reals/Ranalysis1.vo theories/Reals/PSeries_reg.vo +theories/Reals/PSeries_reg.vo: theories/Reals/PSeries_reg.v theories/Reals/Rbase.vo theories/Reals/DiscrR.vo theories/Reals/Rfunctions.vo theories/Reals/Rseries.vo theories/Reals/Rsigma.vo theories/Reals/Alembert.vo theories/Reals/Alembert_compl.vo theories/Reals/Binome.vo theories/Reals/Cv_prop.vo theories/Reals/Rcomplet.vo theories/Reals/Rtrigo_alt.vo theories/Reals/Cos_plus.vo theories/Reals/Ranalysis1.vo theories/Arith/Max.vo theories/Arith/Even.vo theories/Reals/Ranalysis1.vo: theories/Reals/Ranalysis1.v theories/Reals/Rbase.vo theories/Reals/Rbasic_fun.vo theories/Reals/R_sqr.vo theories/Reals/Rlimit.vo theories/Reals/Rderiv.vo theories/Reals/DiscrR.vo theories/Reals/Rtrigo.vo theories/Reals/Rtrigo.vo: theories/Reals/Rtrigo.v theories/ZArith/ZArith_base.vo theories/ZArith/Zcomplements.vo theories/Logic/Classical_Prop.vo theories/Reals/DiscrR.vo theories/Reals/Rbase.vo theories/Reals/R_sqr.vo theories/Reals/Rfunctions.vo theories/Reals/Rsigma.vo theories/Reals/Rlimit.vo theories/Reals/Binome.vo theories/Reals/Rtrigo_def.vo theories/Reals/Rtrigo_alt.vo theories/Reals/Cos_plus.vo theories/Reals/Cos_plus.vo: theories/Reals/Cos_plus.v theories/Arith/Max.vo theories/Reals/Rbase.vo theories/Reals/DiscrR.vo theories/Reals/Rseries.vo theories/Reals/Binome.vo theories/Reals/Rtrigo_def.vo theories/Reals/Rtrigo_alt.vo theories/Reals/Rprod.vo theories/Reals/Cv_prop.vo theories/Reals/Cos_rel.vo diff --git a/Makefile b/Makefile index d85e1cc2b..81d305af2 100644 --- a/Makefile +++ b/Makefile @@ -81,14 +81,17 @@ LIBREP=lib/pp_control.cmo lib/pp.cmo lib/util.cmo \ KERNEL=kernel/names.cmo kernel/univ.cmo \ kernel/esubst.cmo kernel/term.cmo kernel/sign.cmo \ kernel/declarations.cmo kernel/environ.cmo kernel/closure.cmo \ - kernel/conv_oracle.cmo kernel/reduction.cmo \ + kernel/conv_oracle.cmo kernel/reduction.cmo kernel/entries.cmo \ + kernel/modops.cmo \ kernel/type_errors.cmo kernel/inductive.cmo kernel/typeops.cmo \ - kernel/indtypes.cmo kernel/cooking.cmo kernel/safe_typing.cmo + kernel/indtypes.cmo kernel/cooking.cmo kernel/term_typing.cmo \ + kernel/subtyping.cmo kernel/mod_typing.cmo kernel/safe_typing.cmo -LIBRARY=library/nameops.cmo library/libobject.cmo library/summary.cmo \ - library/nametab.cmo library/lib.cmo library/global.cmo \ - library/goptions.cmo library/library.cmo library/states.cmo \ - library/impargs.cmo library/declare.cmo +LIBRARY=library/libnames.cmo library/nameops.cmo library/libobject.cmo \ + library/summary.cmo \ + library/nametab.cmo library/global.cmo library/lib.cmo \ + library/declaremods.cmo library/library.cmo library/states.cmo \ + library/impargs.cmo library/declare.cmo library/goptions.cmo PRETYPING=pretyping/termops.cmo \ pretyping/evd.cmo pretyping/instantiate.cmo \ @@ -103,7 +106,7 @@ PRETYPING=pretyping/termops.cmo \ PARSING=parsing/lexer.cmo parsing/coqast.cmo \ parsing/genarg.cmo proofs/tacexpr.cmo parsing/ast.cmo \ - parsing/termast.cmo parsing/astterm.cmo \ + parsing/termast.cmo parsing/astterm.cmo parsing/astmod.cmo \ parsing/extend.cmo parsing/esyntax.cmo \ parsing/ppconstr.cmo parsing/printer.cmo parsing/pptactic.cmo \ parsing/coqlib.cmo parsing/prettyp.cmo parsing/search.cmo @@ -111,6 +114,7 @@ PARSING=parsing/lexer.cmo parsing/coqast.cmo \ HIGHPARSING= parsing/g_prim.cmo parsing/g_basevernac.cmo \ parsing/g_vernac.cmo parsing/g_proofs.cmo parsing/g_tactic.cmo \ parsing/g_ltac.cmo parsing/g_constr.cmo parsing/g_cases.cmo \ + parsing/g_module.cmo \ parsing/g_natsyntax.cmo parsing/g_zsyntax.cmo parsing/g_rsyntax.cmo ARITHSYNTAX=parsing/g_natsyntax.cmo parsing/g_zsyntax.cmo parsing/g_rsyntax.cmo @@ -181,10 +185,14 @@ PARSERREQUIRES=config/coq_config.cmo lib/pp_control.cmo lib/pp.cmo \ kernel/term.cmo kernel/sign.cmo kernel/declarations.cmo \ kernel/environ.cmo \ kernel/closure.cmo kernel/conv_oracle.cmo kernel/reduction.cmo \ + kernel/modops.cmo \ kernel/type_errors.cmo kernel/inductive.cmo kernel/typeops.cmo \ - kernel/indtypes.cmo kernel/cooking.cmo kernel/safe_typing.cmo \ + kernel/indtypes.cmo kernel/cooking.cmo kernel/term_typing.cmo \ + kernel/subtyping.cmo kernel/mod_typing.cmo kernel/safe_typing.cmo \ + library/libnames.cmo \ library/nameops.cmo library/libobject.cmo library/summary.cmo \ library/nametab.cmo library/lib.cmo library/global.cmo \ + library/declaremods.cmo \ library/library.cmo lib/options.cmo library/impargs.cmo \ library/goptions.cmo \ pretyping/evd.cmo pretyping/instantiate.cmo \ @@ -203,7 +211,7 @@ PARSERREQUIRES=config/coq_config.cmo lib/pp_control.cmo lib/pp.cmo \ parsing/g_prim.cmo parsing/g_basevernac.cmo \ parsing/extend.cmo \ parsing/coqlib.cmo pretyping/detyping.cmo \ - parsing/termast.cmo parsing/astterm.cmo \ + parsing/termast.cmo parsing/astterm.cmo parsing/astmod.cmo \ parsing/egrammar.cmo parsing/esyntax.cmo toplevel/metasyntax.cmo \ parsing/ppconstr.cmo parsing/printer.cmo parsing/pptactic.cmo \ lib/stamps.cmo pretyping/typing.cmo \ @@ -272,9 +280,11 @@ JPROVERCMO=contrib/jprover/opname.cmo \ ML4FILES += contrib/jprover/jprover.ml4 -CONTRIB=$(OMEGACMO) $(ROMEGACMO) $(RINGCMO) $(FIELDCMO) $(XMLCMO) \ - $(FOURIERCMO) \ - $(EXTRACTIONCMO) $(CORRECTNESSCMO) $(JPROVERCMO) +CONTRIB=$(OMEGACMO) $(ROMEGACMO) $(RINGCMO) $(FIELDCMO) \ + $(FOURIERCMO) $(EXTRACTIONCMO) $(JPROVERCMO) $(XMLCMO) + +#later $(CORRECTNESSCMO) +#later :) CMA=$(CLIBS) $(CAMLP4OBJS) CMXA=$(CMA:.cma=.cmxa) @@ -368,7 +378,7 @@ bin/coq-interface$(EXE): $(COQMKTOP) $(CMO) $(USERTACCMO) $(INTERFACE) $(COQMKTOP) -top $(BYTEFLAGS) -o $@ $(INTERFACE) bin/parser$(EXE): contrib/interface/ctast.cmo contrib/interface/parse.cmo contrib/interface/line_parser.cmo $(PARSERREQUIRES) contrib/interface/xlate.cmo contrib/interface/vtp.cmo - $(OCAMLC) -cclib -lunix -custom $(MLINCLUDES) -o $@ $(CMA) \ + $(OCAMLC) -cclib -lunix -custom $(BYTEFLAGS) -o $@ $(CMA) \ $(PARSERREQUIRES) \ ctast.cmo line_parser.cmo vtp.cmo xlate.cmo parse.cmo @@ -394,6 +404,10 @@ SYNTAXPP=syntax/PPConstr.v syntax/PPCases.v states/barestate.coq: $(SYNTAXPP) $(BESTCOQTOP) $(BESTCOQTOP) -boot -batch -silent -nois -I syntax -load-vernac-source syntax/MakeBare.v -outputstate states/barestate.coq +# theories/Init/DatatypesHints.vo theories/Init/PeanoHints.vo \ +# theories/Init/LogicHints.vo theories/Init/SpecifHints.vo \ +# theories/Init/Logic_TypeHints.vo \ + INITVO=theories/Init/Datatypes.vo theories/Init/Peano.vo \ theories/Init/DatatypesSyntax.vo theories/Init/Prelude.vo \ theories/Init/Logic.vo theories/Init/Specif.vo \ @@ -784,7 +798,7 @@ clean:: # NB: the -maxdepth 3 is for excluding files from contrib/extraction/test tags: - find . -maxdepth 3 -name "*.ml*" | sort -r | xargs \ + find . -maxdepth 3 -regex ".*\.ml[i4]?" | sort -r | xargs \ etags --language=none\ "--regex=/let[ \t]+\([^ \t]+\)/\1/" \ "--regex=/let[ \t]+rec[ \t]+\([^ \t]+\)/\1/" \ @@ -822,10 +836,13 @@ CAMLP4EXTENSIONS= parsing/argextend.cmo parsing/tacextend.cmo \ GRAMMARCMO=lib/pp_control.cmo lib/pp.cmo lib/util.cmo \ lib/dyn.cmo lib/options.cmo \ - lib/hashcons.cmo lib/predicate.cmo lib/rtree.cmo $(KERNEL) \ - library/summary.cmo library/nameops.cmo library/nametab.cmo \ - library/libobject.cmo library/lib.cmo library/global.cmo \ - library/goptions.cmo pretyping/rawterm.cmo pretyping/evd.cmo \ + lib/hashcons.cmo lib/predicate.cmo lib/rtree.cmo \ + $(KERNEL) \ + library/libnames.cmo \ + library/summary.cmo library/nameops.cmo \ + library/nametab.cmo \ + library/libobject.cmo library/lib.cmo library/goptions.cmo \ + pretyping/rawterm.cmo pretyping/evd.cmo \ parsing/coqast.cmo parsing/genarg.cmo \ proofs/tacexpr.cmo proofs/proof_type.cmo parsing/ast.cmo \ parsing/lexer.cmo parsing/q_util.cmo parsing/extend.cmo \ @@ -839,9 +856,12 @@ clean:: rm -f parsing/grammar.cma ML4FILES +=parsing/g_basevernac.ml4 parsing/g_minicoq.ml4 \ - parsing/g_vernac.ml4 parsing/g_proofs.ml4 parsing/g_cases.ml4 \ - parsing/g_constr.ml4 parsing/g_tactic.ml4 parsing/g_ltac.ml4 \ - parsing/argextend.ml4 parsing/tacextend.ml4 parsing/vernacextend.ml4 + parsing/g_vernac.ml4 parsing/g_proofs.ml4 \ + parsing/g_cases.ml4 \ + parsing/g_constr.ml4 parsing/g_module.ml4 \ + parsing/g_tactic.ml4 parsing/g_ltac.ml4 \ + parsing/argextend.ml4 parsing/tacextend.ml4 \ + parsing/vernacextend.ml4 # beforedepend:: $(GRAMMARCMO) @@ -1002,6 +1022,9 @@ depend: beforedepend done # 5. Finally, we erase the generated .ml files rm -f $(ML4FILESML) +# 6. Since .depend contains correct dependencies .depend.devel can be deleted +# (see dev/Makefile.dir for details about this file) + if [ -e makefile ]; then >.depend.devel; else rm -f .depend.devel; fi ml4clean:: rm -f $(ML4FILESML) @@ -1011,3 +1034,10 @@ clean:: include .depend include .depend.coq + + +# this sets up developper supporting stuff +devel: + touch .depend.devel + $(MAKE) -f dev/Makefile.devel setup-devel + $(MAKE) dev/top_printers.cmo diff --git a/contrib/correctness/ptactic.ml b/contrib/correctness/ptactic.ml index dbae1fce5..1a0d4dc41 100644 --- a/contrib/correctness/ptactic.ml +++ b/contrib/correctness/ptactic.ml @@ -229,7 +229,7 @@ let correctness_hook _ ref = register pf_id None let correctness s p opttac = - Library.check_required_module ["Coq";"correctness";"Correctness"]; + Library.check_required_library ["Coq";"correctness";"Correctness"]; Pmisc.reset_names(); let p,oc,cty,v = coqast_of_prog p in let env = Global.env () in diff --git a/contrib/extraction/common.ml b/contrib/extraction/common.ml index 139f849c8..c7f0a97d9 100644 --- a/contrib/extraction/common.ml +++ b/contrib/extraction/common.ml @@ -16,9 +16,10 @@ open Table open Mlutil open Extraction open Ocaml -open Nametab +open Libnames open Util open Declare +open Nametab (*s Modules considerations *) @@ -34,7 +35,7 @@ let qualid_of_dirpath d = let is_long_module d r = let dir = repr_dirpath d - and dir' = repr_dirpath (dirpath (sp_of_r r)) in + and dir' = repr_dirpath (fst (decode_kn (kn_of_r r))) in let l = List.length dir and l' = List.length dir' in if l' < l then false @@ -106,7 +107,7 @@ let cache r f = module ToplevelParams = struct let globals () = Idset.empty - let rename_global r _ = Termops.id_of_global (Global.env()) r + let rename_global r _ = id_of_global None r let pp_global r _ _ = Printer.pr_global r end @@ -124,7 +125,7 @@ module MonoParams = struct let rename_global r upper = cache r (fun r -> - let id = Termops.id_of_global (Global.env()) r in + let id = id_of_global None r in rename_global_id (if upper || (is_construct r) then uppercase_id id else lowercase_id id)) @@ -143,7 +144,7 @@ module ModularParams = struct let clash r id = try - let _ = locate (make_qualid (dirpath (sp_of_r r)) id) + let _ = locate (make_qualid (fst (decode_kn (kn_of_r r))) id) in true with _ -> false @@ -160,7 +161,7 @@ module ModularParams = struct let rename_global r upper = cache r (fun r -> - let id = Termops.id_of_global (Global.env()) r in + let id = id_of_global None r in if upper || (is_construct r) then rename_global_id r id (uppercase_id id) "Coq_" else rename_global_id r id (lowercase_id id) "coq_") diff --git a/contrib/extraction/common.mli b/contrib/extraction/common.mli index 299ed508c..9ebb11069 100644 --- a/contrib/extraction/common.mli +++ b/contrib/extraction/common.mli @@ -12,7 +12,7 @@ open Pp open Miniml open Mlutil open Names -open Nametab +open Libnames val is_long_module : dir_path -> global_reference -> bool diff --git a/contrib/extraction/extract_env.ml b/contrib/extraction/extract_env.ml index 206de8a28..d04a65fde 100644 --- a/contrib/extraction/extract_env.ml +++ b/contrib/extraction/extract_env.ml @@ -18,6 +18,7 @@ open Extraction open Miniml open Table open Mlutil +open Libnames open Nametab open Vernacinterp open Common @@ -83,12 +84,12 @@ let check_modules m = We just keep constants and inductives. *) let extract_module m = - let seg = Library.module_segment (Some m) in + let seg = Declaremods.module_objects (MPfile m) in let get_reference = function - | sp, Leaf o -> + | (_,kn), Leaf o -> (match Libobject.object_tag o with - | "CONSTANT" | "PARAMETER" -> ConstRef sp - | "INDUCTIVE" -> IndRef (sp,0) + | "CONSTANT" | "PARAMETER" -> ConstRef kn + | "INDUCTIVE" -> IndRef (kn,0) | _ -> failwith "caught") | _ -> failwith "caught" in @@ -204,7 +205,7 @@ let print_user_extract r = let decl_in_r r0 = function | Dterm (r,_) -> r = r0 | Dtype (r,_,_) -> r = r0 - | Dind ((_,r,_)::_, _) -> sp_of_r r = sp_of_r r0 + | Dind ((_,r,_)::_, _) -> kn_of_r r = kn_of_r r0 | Dind ([],_) -> false | DdummyType r -> r = r0 | DcustomTerm (r,_) -> r = r0 diff --git a/contrib/extraction/extract_env.mli b/contrib/extraction/extract_env.mli index e019df342..215161898 100644 --- a/contrib/extraction/extract_env.mli +++ b/contrib/extraction/extract_env.mli @@ -12,7 +12,7 @@ open Util open Names -open Nametab +open Libnames val extraction : qualid located -> unit val extraction_rec : qualid located list -> unit diff --git a/contrib/extraction/extraction.ml b/contrib/extraction/extraction.ml index 8b842c5b5..f574cecae 100644 --- a/contrib/extraction/extraction.ml +++ b/contrib/extraction/extraction.ml @@ -26,6 +26,7 @@ open Table open Mlutil open Closure open Summary +open Libnames open Nametab (*i*) @@ -91,13 +92,13 @@ let add_constructor c e = constructor_table := Gmap.add c e !constructor_table let lookup_constructor c = Gmap.find c !constructor_table let constant_table = - ref (Gmap.empty : (section_path, ml_decl) Gmap.t) -let add_constant sp d = constant_table := Gmap.add sp d !constant_table -let lookup_constant sp = Gmap.find sp !constant_table + ref (Gmap.empty : (kernel_name, ml_decl) Gmap.t) +let add_constant kn d = constant_table := Gmap.add kn d !constant_table +let lookup_constant kn = Gmap.find kn !constant_table -let signature_table = ref (Gmap.empty : (section_path, signature) Gmap.t) -let add_signature sp s = signature_table := Gmap.add sp s !signature_table -let lookup_signature sp = Gmap.find sp !signature_table +let signature_table = ref (Gmap.empty : (kernel_name, signature) Gmap.t) +let add_signature kn s = signature_table := Gmap.add kn s !signature_table +let lookup_signature kn = Gmap.find kn !signature_table (* Tables synchronization. *) @@ -116,15 +117,15 @@ let _ = declare_summary "Extraction tables" (*S Warning and Error messages. *) -let axiom_error_message sp = +let axiom_error_message kn = errorlabstrm "axiom_message" (str "You must specify an extraction for axiom" ++ spc () ++ - pr_sp sp ++ spc () ++ str "first.") + pr_kn kn ++ spc () ++ str "first.") -let axiom_warning_message sp = +let axiom_warning_message kn = Options.if_verbose warn (str "This extraction depends on logical axiom" ++ spc () ++ - pr_sp sp ++ str "." ++ spc() ++ + pr_kn kn ++ str "." ++ spc() ++ str "Having false logical axiom in the environment when extracting" ++ spc () ++ str "may lead to incorrect or non-terminating ML terms.") @@ -140,7 +141,7 @@ let type_of env c = Retyping.get_type_of env none (strip_outer_cast c) let sort_of env c = Retyping.get_sort_family_of env none (strip_outer_cast c) -let is_axiom sp = (Global.lookup_constant sp).const_body = None +let is_axiom kn = (Global.lookup_constant kn).const_body = None (*s [flag_of_type] transforms a type [t] into a [flag]. Really important function. *) @@ -191,12 +192,12 @@ let rec type_sign_vl env c = (*s Function recording signatures of section paths. *) -let signature_of_sp sp = - try lookup_signature sp +let signature_of_kn kn = + try lookup_signature kn with Not_found -> let env = Global.env () in - let s = term_sign env (constant_type env sp) - in add_signature sp s; s + let s = term_sign env (constant_type env kn) + in add_signature kn s; s (*S Management of type variable contexts. *) @@ -261,16 +262,16 @@ let rec extract_type env db c args = | _ -> let n' = List.nth db (n-1) in if n' = 0 then Tunknown else Tvar n') - | Const sp -> - let t = constant_type env sp in + | Const kn -> + let t = constant_type env kn in (match flag_of_type env t with | (Info,Arity) -> - extract_type_app env db (ConstRef sp, type_sign env t) args + extract_type_app env db (ConstRef kn, type_sign env t) args | (Info,_) -> Tunknown | (Logic,_) -> Tdummy) - | Ind spi -> - (match extract_inductive spi with - | Iml (si,_) -> extract_type_app env db (IndRef spi,si) args + | Ind kni -> + (match extract_inductive kni with + | Iml (si,_) -> extract_type_app env db (IndRef kni,si) args | Iprop -> Tdummy) | Sort _ -> Tdummy | Case _ | Fix _ | CoFix _ -> Tunknown @@ -341,8 +342,8 @@ and extract_constructor (((sp,_),_) as c) = extract_mib sp; lookup_constructor c -and extract_mib sp = - let ind = (sp,0) in +and extract_mib kn = + let ind = (kn,0) in if not (Gmap.mem ind !inductive_table) then begin let (mib,mip) = Global.lookup_inductive ind in let env = Global.env () in @@ -353,7 +354,7 @@ and extract_mib sp = (* First pass: we store inductive signatures together with *) (* their type var list. *) for i = 0 to mib.mind_ntypes - 1 do - let ip = (sp,i) in + let ip = (kn,i) in let (mib,mip) = Global.lookup_inductive ip in if mip.mind_sort = (Prop Null) then add_inductive ip Iprop @@ -364,7 +365,7 @@ and extract_mib sp = done; (* Second pass: we extract constructors *) for i = 0 to mib.mind_ntypes - 1 do - let ip = (sp,i) in + let ip = (kn,i) in let (mib,mip) = Global.lookup_inductive ip in match lookup_inductive ip with | Iprop -> @@ -400,11 +401,11 @@ let is_singleton_inductive ind = (mib.mind_ntypes = 1) && (Array.length mip.mind_consnames = 1) && match extract_constructor (ind,1) with - | Cml ([mlt],_,_)-> not (type_mem_sp (fst ind) mlt) + | Cml ([mlt],_,_)-> not (type_mem_kn (fst ind) mlt) | _ -> false -let is_singleton_constructor ((sp,i),_) = - is_singleton_inductive (sp,i) +let is_singleton_constructor ((kn,i),_) = + is_singleton_inductive (kn,i) (*S Modification of the signature of terms. *) @@ -465,9 +466,9 @@ let rec extract_real_args env args s = (*s Abstraction of an constant. *) -and apply_constant env sp args = - let head = MLglob (ConstRef sp) in - let s = signature_of_sp sp in +and apply_constant env kn args = + let head = MLglob (ConstRef kn) in + let s = signature_of_kn kn in let ls = List.length s in let la = Array.length args in if ls = 0 then begin @@ -552,15 +553,15 @@ and extract_term env c = | App (f,a) -> (match kind_of_term (strip_outer_cast f) with | App _ -> assert false - | Const sp -> apply_constant env sp a + | Const kn -> apply_constant env kn a | Construct cp -> apply_constructor env cp a | _ -> let mlargs = Array.fold_right (fun a l -> (extract_constr_to_term env a) :: l) a [] in MLapp (extract_term env f, mlargs)) - | Const sp -> - apply_constant env sp [||] + | Const kn -> + apply_constant env kn [||] | Construct cp -> apply_constructor env cp [||] | Case ({ci_ind=ip},_,c,br) -> @@ -652,17 +653,17 @@ and extract_constr_to_term_wt env c t = (*s From a constant to a ML declaration. *) -let extract_constant sp r = +let extract_constant kn r = let env = Global.env() in - let cb = Global.lookup_constant sp in + let cb = Global.lookup_constant kn in let typ = cb.const_type in match cb.const_body with | None -> (* A logical axiom is risky, an informative one is fatal. *) (match flag_of_type env typ with - | (Info,_) -> axiom_error_message sp - | (Logic,Arity) -> axiom_warning_message sp; + | (Info,_) -> axiom_error_message kn + | (Logic,Arity) -> axiom_warning_message kn; DdummyType r - | (Logic,_) -> axiom_warning_message sp; + | (Logic,_) -> axiom_warning_message kn; Dterm (r, MLdummy')) | Some body -> (match flag_of_type env typ with @@ -676,20 +677,20 @@ let extract_constant sp r = | (Info, _) -> let a = extract_term env body in if a <> MLdummy' then - Dterm (r, kill_prop_lams_eta a (signature_of_sp sp)) + Dterm (r, kill_prop_lams_eta a (signature_of_kn kn)) else Dterm (r, a)) -let extract_constant_cache sp r = - try lookup_constant sp +let extract_constant_cache kn r = + try lookup_constant kn with Not_found -> - let d = extract_constant sp r - in add_constant sp d; d + let d = extract_constant kn r + in add_constant kn d; d (*s From an inductive to a ML declaration. *) -let extract_inductive_declaration sp = - extract_mib sp; - let ip = (sp,0) in +let extract_inductive_declaration kn = + extract_mib kn; + let ip = (kn,0) in if is_singleton_inductive ip then let t = match lookup_constructor (ip,1) with | Cml ([t],_,_)-> t @@ -701,7 +702,7 @@ let extract_inductive_declaration sp = in Dtype (IndRef ip,vl,t) else - let mib = Global.lookup_mind sp in + let mib = Global.lookup_mind kn in let one_ind ip n = iterate_for (-n) (-1) (fun j l -> @@ -713,7 +714,7 @@ let extract_inductive_declaration sp = let l = iterate_for (1 - mib.mind_ntypes) 0 (fun i acc -> - let ip = (sp,-i) in + let ip = (kn,-i) in let nc = Array.length mib.mind_packets.(-i).mind_consnames in match lookup_inductive ip with | Iprop -> acc @@ -725,9 +726,9 @@ let extract_inductive_declaration sp = (*s From a global reference to a ML declaration. *) let extract_declaration r = match r with - | ConstRef sp -> extract_constant sp r - | IndRef (sp,_) -> extract_inductive_declaration sp - | ConstructRef ((sp,_),_) -> extract_inductive_declaration sp + | ConstRef kn -> extract_constant kn r + | IndRef (kn,_) -> extract_inductive_declaration kn + | ConstructRef ((kn,_),_) -> extract_inductive_declaration kn | VarRef _ -> assert false (*s Check if a global reference corresponds to a logical inductive. *) @@ -743,3 +744,4 @@ let decl_is_logical_ind = function let decl_is_singleton = function | ConstructRef cp -> is_singleton_constructor cp | _ -> false + diff --git a/contrib/extraction/extraction.mli b/contrib/extraction/extraction.mli index fe57be427..0a273e752 100644 --- a/contrib/extraction/extraction.mli +++ b/contrib/extraction/extraction.mli @@ -11,7 +11,8 @@ (*s Extraction from Coq terms to Miniml. *) open Miniml -open Nametab +open Environ +open Libnames (*s ML declaration corresponding to a Coq reference. *) diff --git a/contrib/extraction/miniml.mli b/contrib/extraction/miniml.mli index f87f11ed2..eb82e4752 100644 --- a/contrib/extraction/miniml.mli +++ b/contrib/extraction/miniml.mli @@ -13,7 +13,7 @@ open Pp open Names open Term -open Nametab +open Libnames (*s ML type expressions. *) diff --git a/contrib/extraction/mlutil.ml b/contrib/extraction/mlutil.ml index 544d8af6e..5abd599ed 100644 --- a/contrib/extraction/mlutil.ml +++ b/contrib/extraction/mlutil.ml @@ -18,7 +18,7 @@ open Miniml open Nametab open Table open Options -open Nameops +open Libnames (*i*) (*s Exceptions. *) @@ -40,15 +40,15 @@ let id_of_name = function (*s Does a section path occur in a ML type ? *) -let sp_of_r r = match r with - | ConstRef sp -> sp - | IndRef (sp,_) -> sp - | ConstructRef ((sp,_),_) -> sp +let kn_of_r r = match r with + | ConstRef kn -> kn + | IndRef (kn,_) -> kn + | ConstructRef ((kn,_),_) -> kn | _ -> assert false -let rec type_mem_sp sp = function - | Tglob (r,l) -> (sp_of_r r) = sp || List.exists (type_mem_sp sp) l - | Tarr (a,b) -> (type_mem_sp sp a) || (type_mem_sp sp b) +let rec type_mem_kn kn = function + | Tglob (r,l) -> (kn_of_r r) = kn || List.exists (type_mem_kn kn) l + | Tarr (a,b) -> (type_mem_kn kn a) || (type_mem_kn kn b) | _ -> false (*S Generic functions over ML ast terms. *) @@ -650,9 +650,9 @@ let is_ind = function | _ -> false let is_rec_principle = function - | ConstRef sp -> - let d,i = repr_path sp in - let s = string_of_id i in + | ConstRef c -> + let m,d,l = repr_kn c in + let s = string_of_label l in if Filename.check_suffix s "_rec" then let i' = id_of_string (Filename.chop_suffix s "_rec") in (try is_ind (locate (make_qualid d i')) @@ -752,12 +752,13 @@ let inline_test t = not (is_fix t) && (is_constr t || (ml_size t < 12 && is_not_strict t)) let manual_inline_list = - List.map (fun s -> path_of_string ("Coq.Init."^s)) - [ (* "Wf.Acc_rec" ; "Wf.Acc_rect" ; *) - "Wf.well_founded_induction" ; "Wf.well_founded_induction_type" ] + let dir = dirpath_of_string "Coq.Init.Wf" + in List.map (fun s -> encode_kn dir (id_of_string s)) + [ "Acc_rec" ; "Acc_rect" ; + "well_founded_induction" ; "well_founded_induction_type" ] let manual_inline = function - | ConstRef sp -> List.mem sp manual_inline_list + | ConstRef c -> List.mem c manual_inline_list | _ -> false (* If the user doesn't say he wants to keep [t], we inline in two cases: @@ -838,7 +839,7 @@ and optimize_Dfix prm r t b l = else optimize prm l else let v = try - let d = dirpath (sp_of_r r) in + let d,_ = decode_kn (kn_of_r r) in Array.map (fun id -> locate (make_qualid d id)) f with Not_found -> raise Impossible in diff --git a/contrib/extraction/mlutil.mli b/contrib/extraction/mlutil.mli index 778c7ee51..854e3b5e4 100644 --- a/contrib/extraction/mlutil.mli +++ b/contrib/extraction/mlutil.mli @@ -10,7 +10,7 @@ open Names open Term -open Nametab +open Libnames open Miniml @@ -41,9 +41,9 @@ val named_lams : identifier list -> ml_ast -> ml_ast (*s Utility functions over ML types. [update_args sp vl t] puts [vl] as arguments behind every inductive types [(sp,_)]. *) -val sp_of_r : global_reference -> section_path +val kn_of_r : global_reference -> kernel_name -val type_mem_sp : section_path -> ml_type -> bool +val type_mem_kn : kernel_name -> ml_type -> bool (*s Utility functions over ML terms. [occurs n t] checks whether [Rel n] occurs (freely) in [t]. [ml_lift] is de Bruijn @@ -55,6 +55,8 @@ val occurs : int -> ml_ast -> bool val ml_lift : int -> ml_ast -> ml_ast +val ml_subst : ml_ast -> ml_ast -> ml_ast + val ml_pop : ml_ast -> ml_ast (*s Some transformations of ML terms. [optimize] simplify diff --git a/contrib/extraction/table.ml b/contrib/extraction/table.ml index 34b57a45c..af1a9c226 100644 --- a/contrib/extraction/table.ml +++ b/contrib/extraction/table.ml @@ -19,7 +19,7 @@ open Util open Pp open Term open Declarations -open Nametab +open Libnames open Reduction (*s AutoInline parameter *) @@ -77,11 +77,11 @@ let lang_ref = ref Ocaml let lang () = !lang_ref let (extr_lang,_) = - declare_object ("Extraction Lang", - {cache_function = (fun (_,l) -> lang_ref := l); - load_function = (fun (_,l) -> lang_ref := l); - open_function = (fun _ -> ()); - export_function = (fun x -> Some x) }) + declare_object + {(default_object "Extraction Lang") with + cache_function = (fun (_,l) -> lang_ref := l); + load_function = (fun _ (_,l) -> lang_ref := l); + export_function = (fun x -> Some x)} let _ = declare_summary "Extraction Lang" { freeze_function = (fun () -> !lang_ref); @@ -111,11 +111,11 @@ let add_inline_entries b l = (*s Registration of operations for rollback. *) let (inline_extraction,_) = - declare_object ("Extraction Inline", - { cache_function = (fun (_,(b,l)) -> add_inline_entries b l); - load_function = (fun (_,(b,l)) -> add_inline_entries b l); - open_function = (fun _ -> ()); - export_function = (fun x -> Some x) }) + declare_object + {(default_object "Extraction Inline") with + cache_function = (fun (_,(b,l)) -> add_inline_entries b l); + load_function = (fun _ (_,(b,l)) -> add_inline_entries b l); + export_function = (fun x -> Some x)} let _ = declare_summary "Extraction Inline" { freeze_function = (fun () -> !inline_table); @@ -148,11 +148,10 @@ let print_extraction_inline () = let (reset_inline,_) = declare_object - ("Reset Extraction Inline", - { cache_function = (fun (_,_)-> inline_table := empty_inline_table); - load_function = (fun (_,_)-> inline_table := empty_inline_table); - open_function = (fun _ -> ()); - export_function = (fun x -> Some x) }) + {(default_object "Reset Extraction Inline") with + cache_function = (fun (_,_)-> inline_table := empty_inline_table); + load_function = (fun _ (_,_)-> inline_table := empty_inline_table); + export_function = (fun x -> Some x)} let reset_extraction_inline () = add_anonymous_leaf (reset_inline ()) @@ -199,12 +198,11 @@ let find_ml_extraction r = snd (Refmap.find r (fst !extractions)) let (in_ml_extraction,_) = declare_object - ("ML extractions", - { cache_function = (fun (_,(r,k,s)) -> add_ml_extraction r k s); - load_function = (fun (_,(r,k,s)) -> add_ml_extraction r k s); - open_function = (fun _ -> ()); - export_function = (fun x -> Some x) }) - + {(default_object "ML extractions") with + cache_function = (fun (_,(r,k,s)) -> add_ml_extraction r k s); + load_function = (fun _ (_,(r,k,s)) -> add_ml_extraction r k s); + export_function = (fun x -> Some x)} + let _ = declare_summary "ML extractions" { freeze_function = (fun () -> !extractions); unfreeze_function = ((:=) extractions); diff --git a/contrib/extraction/table.mli b/contrib/extraction/table.mli index 1e21e494b..063c18a3c 100644 --- a/contrib/extraction/table.mli +++ b/contrib/extraction/table.mli @@ -10,7 +10,7 @@ open Vernacinterp open Names -open Nametab +open Libnames (*s AutoInline parameter *) diff --git a/contrib/field/field.ml4 b/contrib/field/field.ml4 index 7dc6d961f..d5c50f9d3 100644 --- a/contrib/field/field.ml4 +++ b/contrib/field/field.ml4 @@ -34,7 +34,7 @@ let constant dir s = Declare.global_reference_in_absolute_module dir id with Not_found -> anomaly ("Field: cannot find "^ - (Nametab.string_of_qualid (Nametab.make_qualid dir id))) + (Libnames.string_of_qualid (Libnames.make_qualid dir id))) (* To deal with the optional arguments *) let constr_of_opt a opt = @@ -44,12 +44,12 @@ let constr_of_opt a opt = | Some f -> mkApp ((constant ["Field_Compl"] "Some"),[|ac;constr_of f|]) (* Table of theories *) -let th_tab = ref ((Hashtbl.create 53) : (constr,constr) Hashtbl.t) +let th_tab = ref (Gmap.empty : (constr,constr) Gmap.t) -let lookup typ = Hashtbl.find !th_tab typ +let lookup typ = Gmap.find typ !th_tab let _ = - let init () = th_tab := (Hashtbl.create 53) in + let init () = th_tab := Gmap.empty in let freeze () = !th_tab in let unfreeze fs = th_tab := fs in Summary.declare_summary "field" @@ -59,17 +59,22 @@ let _ = Summary.survive_section = false } let load_addfield _ = () -let cache_addfield (_,(typ,th)) = Hashtbl.add !th_tab typ th +let cache_addfield (_,(typ,th)) = th_tab := Gmap.add typ th !th_tab +let subst_addfield (_,subst,(typ,th as obj)) = + let typ' = subst_mps subst typ in + let th' = subst_mps subst th in + if typ' == typ && th' == th then obj else + (typ',th') let export_addfield x = Some x (* Declaration of the Add Field library object *) let (in_addfield,out_addfield)= - Libobject.declare_object - ("ADD_FIELD", - { Libobject.load_function = load_addfield; - Libobject.open_function = cache_addfield; + Libobject.declare_object {(Libobject.default_object "ADD_FIELD") with + Libobject.open_function = (fun i o -> if i=1 then cache_addfield o); Libobject.cache_function = cache_addfield; - Libobject.export_function = export_addfield }) + Libobject.subst_function = subst_addfield; + Libobject.classify_function = (fun (_,a) -> Libobject.Substitute a); + Libobject.export_function = export_addfield } (* Adds a theory to the table *) let add_field a aplus amult aone azero aopp aeq ainv aminus_o adiv_o rth @@ -128,7 +133,7 @@ END (* Guesses the type and calls Field_Gen with the right theory *) let field g = - Library.check_required_module ["Coq";"field";"Field"]; + Library.check_required_library ["Coq";"field";"Field"]; let evc = project g and env = pf_env g in let ist = { evc=evc; env=env; lfun=[]; lmatch=[]; @@ -162,7 +167,7 @@ let guess_theory env evc = function (* Guesses the type and calls Field_Term with the right theory *) let field_term l g = - Library.check_required_module ["Coq";"field";"Field"]; + Library.check_required_library ["Coq";"field";"Field"]; let env = (pf_env g) and evc = (project g) in let th = valueIn (VConstr (guess_theory env evc l)) diff --git a/contrib/fourier/fourierR.ml b/contrib/fourier/fourierR.ml index 2cdb929ae..aac632de9 100644 --- a/contrib/fourier/fourierR.ml +++ b/contrib/fourier/fourierR.ml @@ -18,6 +18,7 @@ open Term open Tactics open Clenv open Names +open Libnames open Tacticals open Tacmach open Fourier @@ -76,43 +77,46 @@ let parse s = Astterm.interp_constr Evd.empty (Global.env()) (parse_ast s);; let pf_parse_constr gl s = Astterm.interp_constr Evd.empty (pf_env gl) (parse_ast s);; -let rec string_of_constr c = +let string_of_R_constant kn = + match Names.repr_kn kn with + | MPfile dir, sec_dir, id when + sec_dir = empty_dirpath && + string_of_dirpath dir = "Coq.Reals.Rdefinitions" + -> string_of_label id + | _ -> "constant_not_of_R" + +let rec string_of_R_constr c = match kind_of_term c with - Cast (c,t) -> string_of_constr c - |Const c -> string_of_path c - |Var(c) -> string_of_id c + Cast (c,t) -> string_of_R_constr c + |Const c -> string_of_R_constant c | _ -> "not_of_constant" -;; let rec rational_of_constr c = match kind_of_term c with | Cast (c,t) -> (rational_of_constr c) | App (c,args) -> - (match kind_of_term c with - Const c -> - (match (string_of_path c) with - "Coq.Reals.Rdefinitions.Ropp" -> - rop (rational_of_constr args.(0)) - |"Coq.Reals.Rdefinitions.Rinv" -> - rinv (rational_of_constr args.(0)) - |"Coq.Reals.Rdefinitions.Rmult" -> - rmult (rational_of_constr args.(0)) - (rational_of_constr args.(1)) - |"Coq.Reals.Rdefinitions.Rdiv" -> - rdiv (rational_of_constr args.(0)) - (rational_of_constr args.(1)) - |"Coq.Reals.Rdefinitions.Rplus" -> - rplus (rational_of_constr args.(0)) - (rational_of_constr args.(1)) - |"Coq.Reals.Rdefinitions.Rminus" -> - rminus (rational_of_constr args.(0)) - (rational_of_constr args.(1)) - | _ -> failwith "not a rational") - | _ -> failwith "not a rational") - | Const c -> - (match (string_of_path c) with - "Coq.Reals.Rdefinitions.R1" -> r1 - |"Coq.Reals.Rdefinitions.R0" -> r0 + (match (string_of_R_constr c) with + | "Ropp" -> + rop (rational_of_constr args.(0)) + | "Rinv" -> + rinv (rational_of_constr args.(0)) + | "Rmult" -> + rmult (rational_of_constr args.(0)) + (rational_of_constr args.(1)) + | "Rdiv" -> + rdiv (rational_of_constr args.(0)) + (rational_of_constr args.(1)) + | "Rplus" -> + rplus (rational_of_constr args.(0)) + (rational_of_constr args.(1)) + | "Rminus" -> + rminus (rational_of_constr args.(0)) + (rational_of_constr args.(1)) + | _ -> failwith "not a rational") + | Const kn -> + (match (string_of_R_constant kn) with + "R1" -> r1 + |"R0" -> r0 | _ -> failwith "not a rational") | _ -> failwith "not a rational" ;; @@ -122,43 +126,40 @@ let rec flin_of_constr c = match kind_of_term c with | Cast (c,t) -> (flin_of_constr c) | App (c,args) -> - (match kind_of_term c with - Const c -> - (match (string_of_path c) with - "Coq.Reals.Rdefinitions.Ropp" -> - flin_emult (rop r1) (flin_of_constr args.(0)) - |"Coq.Reals.Rdefinitions.Rplus"-> - flin_plus (flin_of_constr args.(0)) - (flin_of_constr args.(1)) - |"Coq.Reals.Rdefinitions.Rminus"-> - flin_minus (flin_of_constr args.(0)) - (flin_of_constr args.(1)) - |"Coq.Reals.Rdefinitions.Rmult"-> + (match (string_of_R_constr c) with + "Ropp" -> + flin_emult (rop r1) (flin_of_constr args.(0)) + | "Rplus"-> + flin_plus (flin_of_constr args.(0)) + (flin_of_constr args.(1)) + | "Rminus"-> + flin_minus (flin_of_constr args.(0)) + (flin_of_constr args.(1)) + | "Rmult"-> (try (let a=(rational_of_constr args.(0)) in - try (let b = (rational_of_constr args.(1)) in + try (let b = (rational_of_constr args.(1)) in (flin_add_cste (flin_zero()) (rmult a b))) - with _-> (flin_add (flin_zero()) - args.(1) - a)) + with _-> (flin_add (flin_zero()) + args.(1) + a)) with _-> (flin_add (flin_zero()) - args.(0) - (rational_of_constr args.(1)))) - |"Coq.Reals.Rdefinitions.Rinv"-> - let a=(rational_of_constr args.(0)) in + args.(0) + (rational_of_constr args.(1)))) + | "Rinv"-> + let a=(rational_of_constr args.(0)) in flin_add_cste (flin_zero()) (rinv a) - |"Coq.Reals.Rdefinitions.Rdiv"-> - (let b=(rational_of_constr args.(1)) in - try (let a = (rational_of_constr args.(0)) in + | "Rdiv"-> + (let b=(rational_of_constr args.(1)) in + try (let a = (rational_of_constr args.(0)) in (flin_add_cste (flin_zero()) (rdiv a b))) - with _-> (flin_add (flin_zero()) - args.(0) - (rinv b))) - |_->assert false) - |_ -> assert false) + with _-> (flin_add (flin_zero()) + args.(0) + (rinv b))) + |_->assert false) | Const c -> - (match (string_of_path c) with - "Coq.Reals.Rdefinitions.R1" -> flin_one () - |"Coq.Reals.Rdefinitions.R0" -> flin_zero () + (match (string_of_R_constant c) with + "R1" -> flin_one () + |"R0" -> flin_zero () |_-> assert false) |_-> assert false) with _ -> flin_add (flin_zero()) @@ -191,29 +192,29 @@ let ineq1_of_constr (h,t) = let t2= args.(1) in (match kind_of_term f with Const c -> - (match (string_of_path c) with - "Coq.Reals.Rdefinitions.Rlt" -> [{hname=h; + (match (string_of_R_constant c) with + "Rlt" -> [{hname=h; htype="Rlt"; hleft=t1; hright=t2; hflin= flin_minus (flin_of_constr t1) (flin_of_constr t2); hstrict=true}] - |"Coq.Reals.Rdefinitions.Rgt" -> [{hname=h; + |"Rgt" -> [{hname=h; htype="Rgt"; hleft=t2; hright=t1; hflin= flin_minus (flin_of_constr t2) (flin_of_constr t1); hstrict=true}] - |"Coq.Reals.Rdefinitions.Rle" -> [{hname=h; + |"Rle" -> [{hname=h; htype="Rle"; hleft=t1; hright=t2; hflin= flin_minus (flin_of_constr t1) (flin_of_constr t2); hstrict=false}] - |"Coq.Reals.Rdefinitions.Rge" -> [{hname=h; + |"Rge" -> [{hname=h; htype="Rge"; hleft=t2; hright=t1; @@ -221,15 +222,15 @@ let ineq1_of_constr (h,t) = (flin_of_constr t1); hstrict=false}] |_->assert false) - | Ind (sp,i) -> - (match (string_of_path sp) with - "Coq.Init.Logic_Type.eqT" -> let t0= args.(0) in + | Ind (kn,i) -> + if IndRef(kn,i) = Coqlib.glob_eqT then + let t0= args.(0) in let t1= args.(1) in let t2= args.(2) in (match (kind_of_term t0) with Const c -> - (match (string_of_path c) with - "Coq.Reals.Rdefinitions.R"-> + (match (string_of_R_constant c) with + "R"-> [{hname=h; htype="eqTLR"; hleft=t1; @@ -246,7 +247,8 @@ let ineq1_of_constr (h,t) = hstrict=false}] |_-> assert false) |_-> assert false) - |_-> assert false) + else + assert false |_-> assert false) |_-> assert false ;; @@ -373,17 +375,17 @@ let tac_use h = match h.htype with let is_ineq (h,t) = match (kind_of_term t) with - App (f,args) -> - (match (string_of_constr f) with - "Coq.Reals.Rdefinitions.Rlt" -> true - |"Coq.Reals.Rdefinitions.Rgt" -> true - |"Coq.Reals.Rdefinitions.Rle" -> true - |"Coq.Reals.Rdefinitions.Rge" -> true - |"Coq.Init.Logic_Type.eqT" -> (match (string_of_constr args.(0)) with - "Coq.Reals.Rdefinitions.R"->true - |_->false) - |_->false) - |_->false + App (f,args) -> + (match (string_of_R_constr f) with + "Rlt" -> true + | "Rgt" -> true + | "Rle" -> true + | "Rge" -> true + | "eqT" -> (match (string_of_R_constr args.(0)) with + "R" -> true + | _ -> false) + | _ ->false) + |_->false ;; let list_of_sign s = List.map (fun (x,_,z)->(x,z)) s;; @@ -395,7 +397,7 @@ let mkAppL a = (* Résolution d'inéquations linéaires dans R *) let rec fourier gl= - Library.check_required_module ["Coq";"fourier";"Fourier"]; + Library.check_required_library ["Coq";"fourier";"Fourier"]; let parse = pf_parse_constr gl in let goal = strip_outer_cast (pf_concl gl) in let fhyp=id_of_string "new_hyp_for_fourier" in @@ -404,23 +406,23 @@ let rec fourier gl= try (let tac = match (kind_of_term goal) with App (f,args) -> - (match (string_of_constr f) with - "Coq.Reals.Rdefinitions.Rlt" -> + (match (string_of_R_constr f) with + "Rlt" -> (tclTHEN (tclTHEN (apply (parse "Rfourier_not_ge_lt")) (intro_using fhyp)) fourier) - |"Coq.Reals.Rdefinitions.Rle" -> + |"Rle" -> (tclTHEN (tclTHEN (apply (parse "Rfourier_not_gt_le")) (intro_using fhyp)) fourier) - |"Coq.Reals.Rdefinitions.Rgt" -> + |"Rgt" -> (tclTHEN (tclTHEN (apply (parse "Rfourier_not_le_gt")) (intro_using fhyp)) fourier) - |"Coq.Reals.Rdefinitions.Rge" -> + |"Rge" -> (tclTHEN (tclTHEN (apply (parse "Rfourier_not_lt_ge")) (intro_using fhyp)) diff --git a/contrib/interface/centaur.ml4 b/contrib/interface/centaur.ml4 index 92f6f285f..a0620d4ba 100644 --- a/contrib/interface/centaur.ml4 +++ b/contrib/interface/centaur.ml4 @@ -7,6 +7,7 @@ open Util;; open Ast;; open Term;; open Pp;; +open Libnames;; open Libobject;; open Library;; open Vernacinterp;; @@ -258,9 +259,8 @@ let filter_by_module_from_varg_list l = let add_search (global_reference:global_reference) assumptions cstr = try - let env = Global.env() in let id_string = - string_of_qualid (Nametab.shortest_qualid_of_global env + string_of_qualid (Nametab.shortest_qualid_of_global None global_reference) in let ast = try @@ -321,15 +321,13 @@ and ntyp = nf_betaiota typ in (* The following function is copied from globpr in env/printer.ml *) let globcv x = - let env = Global.env() in match x with | Node(_,"MUTIND", (Path(_,sp))::(Num(_,tyi))::_) -> - let env = Global.env() in convert_qualid - (Nametab.shortest_qualid_of_global env (IndRef(sp,tyi))) + (Nametab.shortest_qualid_of_global None (IndRef(sp,tyi))) | Node(_,"MUTCONSTRUCT",(Path(_,sp))::(Num(_,tyi))::(Num(_,i))::_) -> convert_qualid - (Nametab.shortest_qualid_of_global env + (Nametab.shortest_qualid_of_global None (ConstructRef ((sp, tyi), i))) | _ -> failwith "globcv : unexpected value";; @@ -412,20 +410,20 @@ let inspect n = (fun a -> try (match a with - sp, Lib.Leaf lobj -> - (match sp, object_tag lobj with - _, "VARIABLE" -> + oname, Lib.Leaf lobj -> + (match oname, object_tag lobj with + (sp,_), "VARIABLE" -> let ((_, _, v), _) = get_variable (basename sp) in add_search2 (Nametab.locate (qualid_of_sp sp)) v - | sp, ("CONSTANT"|"PARAMETER") -> - let {const_type=typ} = Global.lookup_constant sp in + | (sp,kn), ("CONSTANT"|"PARAMETER") -> + let {const_type=typ} = Global.lookup_constant kn in add_search2 (Nametab.locate (qualid_of_sp sp)) typ - | sp, "MUTUALINDUCTIVE" -> + | (sp,kn), "MUTUALINDUCTIVE" -> add_search2 (Nametab.locate (qualid_of_sp sp)) (Pretyping.understand Evd.empty (Global.env()) - (RRef(dummy_loc, IndRef(sp,0)))) + (RRef(dummy_loc, IndRef(kn,0)))) | _ -> failwith ("unexpected value 1 for "^ - (string_of_id (basename sp)))) + (string_of_id (basename (fst oname))))) | _ -> failwith "unexpected value") with e -> ()) l; diff --git a/contrib/interface/ctast.ml b/contrib/interface/ctast.ml index c5611f985..2345ff471 100644 --- a/contrib/interface/ctast.ml +++ b/contrib/interface/ctast.ml @@ -1,6 +1,7 @@ (* A copy of pre V7 ast *) open Names +open Libnames type loc = int * int @@ -17,7 +18,7 @@ type t = let section_path sl = match List.rev sl with | s::pa -> - make_path + Libnames.encode_kn (make_dirpath (List.map id_of_string pa)) (id_of_string s) | [] -> invalid_arg "section_path" @@ -55,7 +56,7 @@ let rec ast_to_ct = function | Coqast.Id (loc,a) -> Id (loc,a) | Coqast.Str (loc,a) -> Str (loc,a) | Coqast.Path (loc,a) -> - let (sl,bn) = repr_path a in + let (sl,bn) = Libnames.decode_kn a in Path(loc, (List.map string_of_id (List.rev (repr_dirpath sl))) @ [string_of_id bn]) | Coqast.Dynamic (loc,a) -> Dynamic (loc,a) diff --git a/contrib/interface/dad.ml b/contrib/interface/dad.ml index 9d90782e4..3be5d8a36 100644 --- a/contrib/interface/dad.ml +++ b/contrib/interface/dad.ml @@ -15,6 +15,7 @@ open Ctast;; open Termast;; open Astterm;; open Vernacinterp;; +open Libnames;; open Nametab open Proof_type;; diff --git a/contrib/interface/name_to_ast.ml b/contrib/interface/name_to_ast.ml index 8c6293ae2..79375cf2b 100644 --- a/contrib/interface/name_to_ast.ml +++ b/contrib/interface/name_to_ast.ml @@ -8,6 +8,7 @@ open Termast;; open Term;; open Impargs;; open Reduction;; +open Libnames;; open Libobject;; open Environ;; open Declarations;; @@ -83,7 +84,7 @@ let implicit_args_to_ast_list sp mipv = | _ -> [VernacComments (List.rev implicit_args_descriptions)];; let convert_qualid qid = - let d, id = Nametab.repr_qualid qid in + let d, id = Libnames.repr_qualid qid in match repr_dirpath d with [] -> nvar id | d -> ope("QUALID", List.fold_left (fun l s -> (nvar s)::l) @@ -108,7 +109,7 @@ let convert_one_inductive sp tyi = let (ref, params, arity, cstrnames, cstrtypes) = build_inductive sp tyi in let env = Global.env () in let envpar = push_rel_context params env in - let sp = sp_of_global env (IndRef (sp, tyi)) in + let sp = sp_of_global None (IndRef (sp, tyi)) in (basename sp, convert_env(List.rev params), (ast_of_constr true envpar arity), @@ -167,16 +168,16 @@ let make_definition_ast name c typ implicits = ::(implicits_to_ast_list implicits);; (* This function is inspired by print_constant *) -let constant_to_ast_list sp = - let cb = Global.lookup_constant sp in +let constant_to_ast_list kn = + let cb = Global.lookup_constant kn in let c = cb.const_body in let typ = cb.const_type in - let l = constant_implicits_list sp in + let l = constant_implicits_list kn in (match c with None -> - make_variable_ast (basename sp) typ l + make_variable_ast (id_of_label (label kn)) typ l | Some c1 -> - make_definition_ast (basename sp) c1 typ l) + make_definition_ast (id_of_label (label kn)) c1 typ l) let variable_to_ast_list sp = let ((id, c, v), _) = get_variable sp in @@ -195,13 +196,13 @@ let inductive_to_ast_list sp = (* this function is inspired by print_leaf_entry from pretty.ml *) -let leaf_entry_to_ast_list (sp,lobj) = +let leaf_entry_to_ast_list ((sp,kn),lobj) = let tag = object_tag lobj in - match (sp,tag) with - | (_, "VARIABLE") -> variable_to_ast_list (basename sp) - | (_, ("CONSTANT"|"PARAMETER")) -> constant_to_ast_list sp - | (_, "INDUCTIVE") -> inductive_to_ast_list sp - | (_, s) -> + match tag with + | "VARIABLE" -> variable_to_ast_list (basename sp) + | "CONSTANT"|"PARAMETER" -> constant_to_ast_list kn + | "INDUCTIVE" -> inductive_to_ast_list kn + | s -> errorlabstrm "print" (str ("printing of unrecognized object " ^ s ^ " has been required"));; @@ -210,13 +211,13 @@ let leaf_entry_to_ast_list (sp,lobj) = (* this function is inspired by print_name *) -let name_to_ast (qid:Nametab.qualid) = +let name_to_ast qid = let l = try let sp = Nametab.locate_obj qid in let (sp,lobj) = let (sp,entry) = - List.find (fun en -> (fst en) = sp) (Lib.contents_after None) + List.find (fun en -> (fst (fst en)) = sp) (Lib.contents_after None) in match entry with | Lib.Leaf obj -> (sp,obj) @@ -232,7 +233,7 @@ let name_to_ast (qid:Nametab.qualid) = | VarRef sp -> variable_to_ast_list sp with Not_found -> try (* Var locale de but, pas var de section... donc pas d'implicits *) - let dir,name = Nametab.repr_qualid qid in + let dir,name = repr_qualid qid in if (repr_dirpath dir) <> [] then raise Not_found; let (_,c,typ) = Global.lookup_named name in (match c with @@ -240,12 +241,12 @@ let name_to_ast (qid:Nametab.qualid) = | Some c1 -> make_definition_ast name c1 typ []) with Not_found -> try - let sp = Syntax_def.locate_syntactic_definition qid in + let sp = Nametab.locate_syntactic_definition qid in errorlabstrm "print" (str "printing of syntax definitions not implemented") with Not_found -> errorlabstrm "print" - (Nametab.pr_qualid qid ++ + (pr_qualid qid ++ spc () ++ str "not a defined object") in VernacList (List.map (fun x -> (dummy_loc,x)) l) diff --git a/contrib/interface/name_to_ast.mli b/contrib/interface/name_to_ast.mli index 4a68e0134..600ec5f91 100644 --- a/contrib/interface/name_to_ast.mli +++ b/contrib/interface/name_to_ast.mli @@ -1,2 +1,2 @@ -val name_to_ast : Nametab.qualid -> Vernacexpr.vernac_expr;; -val convert_qualid : Nametab.qualid -> Coqast.t;; +val name_to_ast : Libnames.qualid -> Vernacexpr.vernac_expr;; +val convert_qualid : Libnames.qualid -> Coqast.t;; diff --git a/contrib/interface/parse.ml b/contrib/interface/parse.ml index 5bce60f22..61fd06072 100644 --- a/contrib/interface/parse.ml +++ b/contrib/interface/parse.ml @@ -6,6 +6,8 @@ open Ctast;; open Pp;; +open Libnames;; + open Library;; open Ascent;; @@ -67,7 +69,7 @@ let try_require_module import specif names = else Some (specif = "SPECIFICATION")) (List.map (fun name -> - (dummy_loc,Nametab.make_short_qualid (Names.id_of_string name))) + (dummy_loc,Libnames.make_short_qualid (Names.id_of_string name))) names) (import = "IMPORT") with @@ -110,7 +112,7 @@ let execute_when_necessary v = (try Vernacentries.interp v with _ -> - let l=prlist_with_sep spc (fun (_,qid) -> Nametab.pr_qualid qid) l in + let l=prlist_with_sep spc (fun (_,qid) -> pr_qualid qid) l in msgnl (str "Reinterning of " ++ l ++ str " failed")) | VernacRequireFrom (_,_,name,_) -> (try @@ -427,10 +429,10 @@ let print_version_action () = let load_syntax_action reqid module_name = msg (str "loading " ++ str module_name ++ str "... "); try - (let qid = Nametab.make_short_qualid (Names.id_of_string module_name) in - read_module (dummy_loc,qid); + (let qid = Libnames.make_short_qualid (Names.id_of_string module_name) in + read_library (dummy_loc,qid); msg (str "opening... "); - import_module false (dummy_loc,qid); + Declaremods.import_module (Nametab.locate_module qid); msgnl (str "done" ++ fnl ()); ()) with diff --git a/contrib/interface/pbp.ml b/contrib/interface/pbp.ml index 5dd9d071b..55aae90be 100644 --- a/contrib/interface/pbp.ml +++ b/contrib/interface/pbp.ml @@ -29,7 +29,7 @@ let get_hyp_by_name g name = ("hyp",judgment.uj_type)) (* je sais, c'est pas beau, mais je ne sais pas trop me servir de look_up... Loïc *) - with _ -> (let ast = Termast.ast_of_qualid (Nametab.make_short_qualid name)in + with _ -> (let ast = Termast.ast_of_qualid (Libnames.make_short_qualid name)in let c = Astterm.interp_constr evd env ast in ("cste",type_of (Global.env()) Evd.empty c)) ;; @@ -252,7 +252,7 @@ let reference dir s = Nametab.locate_in_absolute_module dir id with Not_found -> anomaly ("Coqlib: cannot find "^ - (Nametab.string_of_qualid (Nametab.make_qualid dir id))) + (Libnames.string_of_qualid (Libnames.make_qualid dir id))) let constant dir s = Declare.constr_of_reference (reference dir s);; diff --git a/contrib/interface/showproof.ml b/contrib/interface/showproof.ml index ae876e129..c7e6be131 100644 --- a/contrib/interface/showproof.ml +++ b/contrib/interface/showproof.ml @@ -6,6 +6,7 @@ open Environ open Evd open Names open Nameops +open Libnames open Term open Termops open Util diff --git a/contrib/interface/xlate.ml b/contrib/interface/xlate.ml index d1af58df7..a2620c67f 100644 --- a/contrib/interface/xlate.ml +++ b/contrib/interface/xlate.ml @@ -344,9 +344,9 @@ let qualid_to_ct_ID = | Node(_, "QUALIDMETA",[Num(_,n)]) -> Some(CT_metac (CT_int n)) | _ -> None;; -let tac_qualid_to_ct_ID qid = CT_ident (Nametab.string_of_qualid qid) +let tac_qualid_to_ct_ID qid = CT_ident (Libnames.string_of_qualid qid) -let loc_qualid_to_ct_ID (_,qid) = CT_ident (Nametab.string_of_qualid qid) +let loc_qualid_to_ct_ID (_,qid) = CT_ident (Libnames.string_of_qualid qid) let qualid_or_meta_to_ct_ID = function | AN (_,qid) -> tac_qualid_to_ct_ID qid @@ -360,7 +360,7 @@ let xlate_qualid_list l = CT_id_list (List.map loc_qualid_to_ct_ID l) let reference_to_ct_ID = function | RIdent (_,id) -> CT_ident (Names.string_of_id id) - | RQualid (_,qid) -> CT_ident (Nametab.string_of_qualid qid) + | RQualid (_,qid) -> CT_ident (Libnames.string_of_qualid qid) let xlate_class = function | FunClass -> CT_ident "FUNCLASS" @@ -437,10 +437,10 @@ let xlate_op the_node opn a b = (match a, b with | ((Path (_, sl)) :: []), [] -> CT_coerce_ID_to_FORMULA(CT_ident - (Names.string_of_id (Nameops.basename (section_path sl)))) + (Names.string_of_label (Names.label (section_path sl)))) | ((Path (_, sl)) :: []), tl -> CT_coerce_ID_to_FORMULA(CT_ident - (Names.string_of_id(Nameops.basename (section_path sl)))) + (Names.string_of_label(Names.label (section_path sl)))) | _, _ -> xlate_error "xlate_op : CONST") | (** string_of_path needs to be investigated. *) @@ -454,8 +454,8 @@ let xlate_op the_node opn a b = | _ -> assert false else CT_coerce_ID_to_FORMULA( - CT_ident(Names.string_of_id - (Nameops.basename (section_path sl)))) + CT_ident(Names.string_of_label + (Names.label (section_path sl)))) | _, _ -> xlate_error "xlate_op : MUTIND") | "CASE" | "MATCH" -> @@ -478,7 +478,10 @@ let xlate_op the_node opn a b = | Some(Rform x) -> x | _ -> assert false else - let name = Names.string_of_path (section_path sl) in + let name = + let dir,id = Libnames.decode_kn (section_path sl) in + Names.string_of_dirpath (Libnames.extend_dirpath dir id) + in (* This is rather a patch to cope with the fact that identifier names have disappeared from the vo files for grammar rules *) let type_desc = (try Some (Hashtbl.find type_table name) with @@ -2864,7 +2867,7 @@ let xlate_vernac = (* | "EndSection", ((Varg_ident id) :: []) -> *) - | VernacEndSection id -> CT_section_end (xlate_ident id) + | VernacEndSegment id -> CT_section_end (xlate_ident id) (* | "StartProof", ((Varg_string (CT_string kind)) :: ((Varg_ident s) :: (c :: []))) -> @@ -3249,9 +3252,9 @@ let xlate_vernac = let local_opt = match s with (* Cannot decide whether it is a global or a Local but at toplevel *) - | Nametab.NeverDischarge -> CT_coerce_NONE_to_LOCAL_OPT CT_none - | Nametab.DischargeAt _ -> CT_local - | Nametab.NotDeclare -> assert false in + | Libnames.NeverDischarge -> CT_coerce_NONE_to_LOCAL_OPT CT_none + | Libnames.DischargeAt _ -> CT_local + | Libnames.NotDeclare -> assert false in CT_coercion (local_opt, id_opt, loc_qualid_to_ct_ID id1, xlate_class id2, xlate_class id3) @@ -3260,9 +3263,9 @@ let xlate_vernac = let local_opt = match s with (* Cannot decide whether it is a global or a Local but at toplevel *) - | Nametab.NeverDischarge -> CT_coerce_NONE_to_LOCAL_OPT CT_none - | Nametab.DischargeAt _ -> CT_local - | Nametab.NotDeclare -> assert false in + | Libnames.NeverDischarge -> CT_coerce_NONE_to_LOCAL_OPT CT_none + | Libnames.DischargeAt _ -> CT_local + | Libnames.NotDeclare -> assert false in CT_coercion (local_opt, id_opt, xlate_ident id1, xlate_class id2, xlate_class id3) (* Not supported diff --git a/contrib/omega/coq_omega.ml b/contrib/omega/coq_omega.ml index 1eb22424b..47d5c5f4d 100644 --- a/contrib/omega/coq_omega.ml +++ b/contrib/omega/coq_omega.ml @@ -34,6 +34,7 @@ open Evar_refiner open Tactics open Clenv open Logic +open Libnames open Nametab open Omega open Contradiction @@ -151,7 +152,7 @@ let dest_const_apply t = | Ind isp -> IndRef isp | _ -> raise Destruct in - id_of_global (Global.env()) ref, args + id_of_global None ref, args type result = | Kvar of string @@ -161,14 +162,14 @@ type result = let destructurate t = let c, args = get_applist t in - let env = Global.env() in +(* let env = Global.env() in*) match kind_of_term c, args with | Const sp, args -> - Kapp (string_of_id (id_of_global env (ConstRef sp)),args) + Kapp (string_of_id (id_of_global None (ConstRef sp)),args) | Construct csp , args -> - Kapp (string_of_id (id_of_global env(ConstructRef csp)), args) + Kapp (string_of_id (id_of_global None (ConstructRef csp)), args) | Ind isp, args -> - Kapp (string_of_id (id_of_global env (IndRef isp)),args) + Kapp (string_of_id (id_of_global None (IndRef isp)),args) | Var id,[] -> Kvar(string_of_id id) | Prod (Anonymous,typ,body), [] -> Kimp(typ,body) | Prod (Name _,_,_),[] -> error "Omega: Not a quantifier-free goal" @@ -200,7 +201,7 @@ let constant dir s = Declare.global_reference_in_absolute_module dir id with Not_found -> anomaly ("Coq_omega: cannot find "^ - (Nametab.string_of_qualid (Nametab.make_qualid dir id))) + (Libnames.string_of_qualid (Libnames.make_qualid dir id))) let arith_constant dir = constant ("Arith"::dir) let zarith_constant dir = constant ("ZArith"::dir) @@ -364,12 +365,12 @@ let make_coq_path dir s = try Nametab.locate_in_absolute_module dir id with Not_found -> anomaly("Coq_omega: cannot find "^ - (Nametab.string_of_qualid(Nametab.make_qualid dir id))) + (Libnames.string_of_qualid(Libnames.make_qualid dir id))) in match ref with | ConstRef sp -> EvalConstRef sp | _ -> anomaly ("Coq_omega: "^ - (Nametab.string_of_qualid (Nametab.make_qualid dir id))^ + (Libnames.string_of_qualid (Libnames.make_qualid dir id))^ " is not a constant") let sp_Zs = lazy (make_coq_path ["ZArith";"zarith_aux"] "Zs") @@ -1861,7 +1862,7 @@ let destructure_goal gl = let destructure_goal = all_time (destructure_goal) let omega_solver gl = - Library.check_required_module ["Coq";"omega";"Omega"]; + Library.check_required_library ["Coq";"omega";"Omega"]; let result = destructure_goal gl in (* if !display_time_flag then begin text_time (); flush Pervasives.stdout end; *) diff --git a/contrib/ring/quote.ml b/contrib/ring/quote.ml index c150a4bfb..a2798d7a7 100644 --- a/contrib/ring/quote.ml +++ b/contrib/ring/quote.ml @@ -127,7 +127,7 @@ let constant dir s = Declare.global_reference_in_absolute_module dir id with Not_found -> anomaly ("Quote: cannot find "^ - (Nametab.string_of_qualid (Nametab.make_qualid dir id))) + (Libnames.string_of_qualid (Libnames.make_qualid dir id))) let coq_Empty_vm = lazy (constant ["Quote"] "Empty_vm") let coq_Node_vm = lazy (constant ["Quote"] "Node_vm") @@ -389,7 +389,7 @@ let rec sort_subterm gl l = [gl: goal sigma]\\ *) let quote_terms ivs lc gl = - Library.check_required_module ["Coq";"ring";"Quote"]; + Library.check_required_library ["Coq";"ring";"Quote"]; let varhash = (Hashtbl.create 17 : (constr, constr) Hashtbl.t) in let varlist = ref ([] : constr list) in (* list of variables *) let counter = ref 1 in (* number of variables created + 1 *) diff --git a/contrib/ring/ring.ml b/contrib/ring/ring.ml index 786a74674..db9b00c38 100644 --- a/contrib/ring/ring.ml +++ b/contrib/ring/ring.ml @@ -15,6 +15,7 @@ open Util open Options open Term open Names +open Libnames open Nameops open Reductionops open Tacticals @@ -44,7 +45,7 @@ let constant dir s = Declare.global_reference_in_absolute_module dir id with Not_found -> anomaly ("Ring: cannot find "^ - (Nametab.string_of_qualid (Nametab.make_qualid dir id))) + (Libnames.string_of_qualid (Libnames.make_qualid dir id))) (* Ring theory *) let coq_Ring_Theory = lazy (constant ["ring";"Ring_theory"] "Ring_Theory") @@ -199,14 +200,93 @@ let _ = The functions theory_to_obj and obj_to_theory do the conversions between theories and environement objects. *) + +let subst_morph subst morph = + let plusm' = subst_mps subst morph.plusm in + let multm' = subst_mps subst morph.multm in + let oppm' = option_smartmap (subst_mps subst) morph.oppm in + if plusm' == morph.plusm + && multm' == morph.multm + && oppm' == morph.oppm then + morph + else + { plusm = plusm' ; + multm = multm' ; + oppm = oppm' ; + } + +let subst_set subst cset = + let same = ref true in + let copy_subst c newset = + let c' = subst_mps subst c in + if not (c' == c) then same := false; + ConstrSet.add c' newset + in + let cset' = ConstrSet.fold copy_subst cset ConstrSet.empty in + if !same then cset else cset' + +let subst_theory subst th = + let th_equiv' = option_smartmap (subst_mps subst) th.th_equiv in + let th_setoid_th' = option_smartmap (subst_mps subst) th.th_setoid_th in + let th_morph' = option_smartmap (subst_morph subst) th.th_morph in + let th_a' = subst_mps subst th.th_a in + let th_plus' = subst_mps subst th.th_plus in + let th_mult' = subst_mps subst th.th_mult in + let th_one' = subst_mps subst th.th_one in + let th_zero' = subst_mps subst th.th_zero in + let th_opp' = option_smartmap (subst_mps subst) th.th_opp in + let th_eq' = subst_mps subst th.th_eq in + let th_t' = subst_mps subst th.th_t in + let th_closed' = subst_set subst th.th_closed in + if th_equiv' == th.th_equiv + && th_setoid_th' == th.th_setoid_th + && th_morph' == th.th_morph + && th_a' == th.th_a + && th_plus' == th.th_plus + && th_mult' == th.th_mult + && th_one' == th.th_one + && th_zero' == th.th_zero + && th_opp' == th.th_opp + && th_eq' == th.th_eq + && th_t' == th.th_t + && th_closed' == th.th_closed + then + th + else + { th_ring = th.th_ring ; + th_abstract = th.th_abstract ; + th_setoid = th.th_setoid ; + th_equiv = th_equiv' ; + th_setoid_th = th_setoid_th' ; + th_morph = th_morph' ; + th_a = th_a' ; + th_plus = th_plus' ; + th_mult = th_mult' ; + th_one = th_one' ; + th_zero = th_zero' ; + th_opp = th_opp' ; + th_eq = th_eq' ; + th_t = th_t' ; + th_closed = th_closed' ; + } + + +let subst_th (_,subst,(c,th as obj)) = + let c' = subst_mps subst c in + let th' = subst_theory subst th in + if c' == c && th' == th then obj else + (c',th') + + let (theory_to_obj, obj_to_theory) = let cache_th (_,(c, th)) = theories_map_add (c,th) and export_th x = Some x in - declare_object ("tactic-ring-theory", - { load_function = (fun _ -> ()); - open_function = cache_th; + declare_object {(default_object "tactic-ring-theory") with + open_function = (fun i o -> if i=1 then cache_th o); cache_function = cache_th; - export_function = export_th }) + subst_function = subst_th; + classify_function = (fun (_,x) -> Substitute x); + export_function = export_th } (* from the set A, guess the associated theory *) (* With this simple solution, the theory to use is automatically guessed *) @@ -647,24 +727,30 @@ module SectionPathSet = SectionPathSet; peut-être faudra-t-il la déplacer dans Closure *) let constants_to_unfold = (* List.fold_right SectionPathSet.add *) - [ path_of_string "Coq.ring.Ring_normalize.interp_cs"; - path_of_string "Coq.ring.Ring_normalize.interp_var"; - path_of_string "Coq.ring.Ring_normalize.interp_vl"; - path_of_string "Coq.ring.Ring_abstract.interp_acs"; - path_of_string "Coq.ring.Ring_abstract.interp_sacs"; - path_of_string "Coq.ring.Quote.varmap_find"; + let transform s = + let sp = path_of_string s in + let dir, id = repr_path sp in + Libnames.encode_kn dir id + in + List.map transform + [ "Coq.ring.Ring_normalize.interp_cs"; + "Coq.ring.Ring_normalize.interp_var"; + "Coq.ring.Ring_normalize.interp_vl"; + "Coq.ring.Ring_abstract.interp_acs"; + "Coq.ring.Ring_abstract.interp_sacs"; + "Coq.ring.Quote.varmap_find"; (* anciennement des Local devenus Definition *) - path_of_string "Coq.ring.Ring_normalize.ics_aux"; - path_of_string "Coq.ring.Ring_normalize.ivl_aux"; - path_of_string "Coq.ring.Ring_normalize.interp_m"; - path_of_string "Coq.ring.Ring_abstract.iacs_aux"; - path_of_string "Coq.ring.Ring_abstract.isacs_aux"; - path_of_string "Coq.ring.Setoid_ring_normalize.interp_cs"; - path_of_string "Coq.ring.Setoid_ring_normalize.interp_var"; - path_of_string "Coq.ring.Setoid_ring_normalize.interp_vl"; - path_of_string "Coq.ring.Setoid_ring_normalize.ics_aux"; - path_of_string "Coq.ring.Setoid_ring_normalize.ivl_aux"; - path_of_string "Coq.ring.Setoid_ring_normalize.interp_m"; + "Coq.ring.Ring_normalize.ics_aux"; + "Coq.ring.Ring_normalize.ivl_aux"; + "Coq.ring.Ring_normalize.interp_m"; + "Coq.ring.Ring_abstract.iacs_aux"; + "Coq.ring.Ring_abstract.isacs_aux"; + "Coq.ring.Setoid_ring_normalize.interp_cs"; + "Coq.ring.Setoid_ring_normalize.interp_var"; + "Coq.ring.Setoid_ring_normalize.interp_vl"; + "Coq.ring.Setoid_ring_normalize.ics_aux"; + "Coq.ring.Setoid_ring_normalize.ivl_aux"; + "Coq.ring.Setoid_ring_normalize.interp_m"; ] (* SectionPathSet.empty *) @@ -782,7 +868,7 @@ let match_with_equiv c = match (kind_of_term c) with | _ -> None let polynom lc gl = - Library.check_required_module ["Coq";"ring";"Ring"]; + Library.check_required_library ["Coq";"ring";"Ring"]; match lc with (* If no argument is given, try to recognize either an equality or a declared relation with arguments c1 ... cn, diff --git a/contrib/romega/const_omega.ml b/contrib/romega/const_omega.ml index 76a6bdf52..42b61a150 100644 --- a/contrib/romega/const_omega.ml +++ b/contrib/romega/const_omega.ml @@ -17,19 +17,19 @@ type result = let destructurate t = let c, args = Term.decompose_app t in - let env = Global.env() in +(* let env = Global.env() in*) match Term.kind_of_term c, args with | Term.Const sp, args -> Kapp (Names.string_of_id - (Termops.id_of_global env (Nametab.ConstRef sp)), + (Nametab.id_of_global None (Libnames.ConstRef sp)), args) | Term.Construct csp , args -> Kapp (Names.string_of_id - (Termops.id_of_global env (Nametab.ConstructRef csp)), + (Nametab.id_of_global None (Libnames.ConstructRef csp)), args) | Term.Ind isp, args -> Kapp (Names.string_of_id - (Termops.id_of_global env (Nametab.IndRef isp)),args) + (Nametab.id_of_global None (Libnames.IndRef isp)),args) | Term.Var id,[] -> Kvar(Names.string_of_id id) | Term.Prod (Names.Anonymous,typ,body), [] -> Kimp(typ,body) | Term.Prod (Names.Name _,_,_),[] -> @@ -42,12 +42,12 @@ let dest_const_apply t = let f,args = Term.decompose_app t in let ref = match Term.kind_of_term f with - | Term.Const sp -> Nametab.ConstRef sp - | Term.Construct csp -> Nametab.ConstructRef csp - | Term.Ind isp -> Nametab.IndRef isp + | Term.Const sp -> Libnames.ConstRef sp + | Term.Construct csp -> Libnames.ConstructRef csp + | Term.Ind isp -> Libnames.IndRef isp | _ -> raise Destruct in - Termops.id_of_global (Global.env()) ref, args + Nametab.id_of_global None ref, args let recognize_number t = let rec loop t = @@ -65,7 +65,7 @@ let recognize_number t = let constant dir s = try Declare.global_absolute_reference - (Names.make_path + (Libnames.make_path (Names.make_dirpath (List.map Names.id_of_string (List.rev dir))) (Names.id_of_string s)) with e -> print_endline (String.concat "." dir); print_endline s; diff --git a/contrib/romega/refl_omega.ml b/contrib/romega/refl_omega.ml index f87f3db96..e6e7074aa 100644 --- a/contrib/romega/refl_omega.ml +++ b/contrib/romega/refl_omega.ml @@ -818,7 +818,7 @@ let destructure_hyps gl = loop (pf_ids_of_hyps gl) (pf_hyps gl) gl let omega_solver gl = - Library.check_required_module ["Coq";"romega";"ROmega"]; + Library.check_required_library ["Coq";"romega";"ROmega"]; let concl = pf_concl gl in let rec loop t = match destructurate t with diff --git a/contrib/xml/xmlcommand.ml4 b/contrib/xml/xmlcommand.ml4 index 3c06c00da..a2baceade 100644 --- a/contrib/xml/xmlcommand.ml4 +++ b/contrib/xml/xmlcommand.ml4 @@ -78,16 +78,16 @@ let extract_nparams pack = (* than that could exists in cooked form with the same name in a super *) (* section of the actual section *) let could_have_namesakes o sp = (* namesake = omonimo in italian *) - let module N = Nametab in + let module L = Libnames in let module D = Declare in let tag = Libobject.object_tag o in print_if_verbose ("Object tag: " ^ tag ^ "\n") ; match tag with "CONSTANT" -> (match D.constant_strength sp with - | N.DischargeAt _ -> false (* a local definition *) - | N.NotDeclare -> false (* not a definition *) - | N.NeverDischarge -> true (* a non-local one *) + | L.DischargeAt _ -> false (* a local definition *) + | L.NotDeclare -> false (* not a definition *) + | L.NeverDischarge -> true (* a non-local one *) ) | "PARAMETER" (* axioms and *) | "INDUCTIVE" -> true (* mutual inductive types are never local *) @@ -100,9 +100,9 @@ let could_have_namesakes o sp = (* namesake = omonimo in italian *) (* section path is sp *) let uri_of_path sp tag = let module N = Names in - let module No = Nameops in + let module L = Libnames in let ext_of_sp sp = ext_of_tag tag in - let dir0 = No.extend_dirpath (No.dirpath sp) (No.basename sp) in + let dir0 = L.extend_dirpath (L.dirpath sp) (L.basename sp) in let dir = List.map N.string_of_id (List.rev (N.repr_dirpath dir0)) in "cic:/" ^ (String.concat "/" dir) ^ "." ^ (ext_of_sp sp) ;; @@ -260,6 +260,8 @@ let print_term inner_types l env csr = let module N = Names in let module E = Environ in let module T = Term in + let module Nt = Nametab in + let module L = Libnames in let module X = Xml in let module R = Retyping in let rec names_to_ids = @@ -419,23 +421,27 @@ let print_term inner_types l env csr = (fun x i -> [< (term_display idradix false l env x); i >]) t [<>]) >] ) - | T.Const sp -> - X.xml_empty "CONST" + | T.Const kn -> + let sp = Nt.sp_of_global None (L.ConstRef kn) in + X.xml_empty "CONST" (add_sort_attribute false ["uri",(uri_of_path sp Constant) ; "id", next_id]) - | T.Ind (sp,i) -> + | T.Ind (kn,i) -> + let sp = Nt.sp_of_global None (L.IndRef(kn,0)) in X.xml_empty "MUTIND" ["uri",(uri_of_path sp Inductive) ; "noType",(string_of_int i) ; "id", next_id] - | T.Construct ((sp,i),j) -> + | T.Construct ((kn,i),j) -> + let sp = Nt.sp_of_global None (L.IndRef(kn,0)) in X.xml_empty "MUTCONSTRUCT" (add_sort_attribute false ["uri",(uri_of_path sp Inductive) ; "noType",(string_of_int i) ; "noConstr",(string_of_int j) ; "id", next_id]) - | T.Case ({T.ci_ind=(sp,i)},ty,term,a) -> + | T.Case ({T.ci_ind=(kn,i)},ty,term,a) -> + let sp = Nt.sp_of_global None (L.IndRef(kn,0)) in let (uri, typeno) = (uri_of_path sp Inductive),i in X.xml_nempty "MUTCASE" (add_sort_attribute true @@ -709,22 +715,24 @@ let print (_,qid as locqid) fn = let module G = Global in let module N = Names in let module Nt = Nametab in + let module Ln = Libnames in let module T = Term in let module X = Xml in - let (_,id) = Nt.repr_qualid qid in + let (_,id) = Ln.repr_qualid qid in let glob_ref = Nametab.global locqid in - let env = (Safe_typing.env_of_safe_env (G.safe_env ())) in + let env = G.env () in reset_ids () ; let inner_types = ref [] in let sp,tag,pp_cmds = match glob_ref with - Nt.VarRef id -> + Ln.VarRef id -> let sp = Declare.find_section_variable id in let (_,body,typ) = G.lookup_named id in sp,Variable,print_variable id body (T.body_of_type typ) env inner_types - | Nt.ConstRef sp -> + | Ln.ConstRef kn -> + let sp = Nt.sp_of_global None glob_ref in let {D.const_body=val0 ; D.const_type = typ ; D.const_hyps = hyps} = - G.lookup_constant sp in + G.lookup_constant kn in let hyps = string_list_of_named_context_list hyps in let typ = T.body_of_type typ in sp,Constant, @@ -733,14 +741,15 @@ let print (_,qid as locqid) fn = None -> print_axiom id typ [] hyps env inner_types | Some c -> print_definition id c typ [] hyps env inner_types end - | Nt.IndRef (sp,_) -> + | Ln.IndRef (kn,_) -> + let sp = Nt.sp_of_global None (Ln.IndRef(kn,0)) in let {D.mind_packets=packs ; D.mind_hyps=hyps; - D.mind_finite=finite} = G.lookup_mind sp in + D.mind_finite=finite} = G.lookup_mind kn in let hyps = string_list_of_named_context_list hyps in sp,Inductive, print_mutual_inductive finite packs [] hyps env inner_types - | Nt.ConstructRef _ -> + | Ln.ConstructRef _ -> Util.anomaly ("print: this should not happen") in Xml.pp pp_cmds fn ; @@ -802,11 +811,11 @@ let mkfilename dn sp ext = let module L = Library in let module S = System in let module N = Names in - let module No = Nameops in + let module Ln = Libnames in match dn with None -> None | Some basedir -> - let dir0 = No.extend_dirpath (No.dirpath sp) (No.basename sp) in + let dir0 = Ln.extend_dirpath (Ln.dirpath sp) (Ln.basename sp) in let dir = List.map N.string_of_id (List.rev (N.repr_dirpath dir0)) in Some (basedir ^ join_dirs basedir dir ^ "." ^ ext) ;; @@ -823,7 +832,7 @@ let mkfilename dn sp ext = (* Note: it is printed only the uncooked available form of the object plus *) (* the list of parameters of the object deduced from it's most cooked *) (* form *) -let print_object lobj id sp dn fv env = +let print_object lobj id (sp,kn) dn fv env = let module D = Declarations in let module E = Environ in let module G = Global in @@ -840,7 +849,7 @@ let print_object lobj id sp dn fv env = "CONSTANT" (* = Definition, Theorem *) | "PARAMETER" (* = Axiom *) -> let {D.const_body=val0 ; D.const_type = typ ; D.const_hyps = hyps} = - G.lookup_constant sp + G.lookup_constant kn in let hyps = string_list_of_named_context_list hyps in let typ = T.body_of_type typ in @@ -854,7 +863,7 @@ let print_object lobj id sp dn fv env = {D.mind_packets=packs ; D.mind_hyps = hyps; D.mind_finite = finite - } = G.lookup_mind sp + } = G.lookup_mind kn in let hyps = string_list_of_named_context_list hyps in print_mutual_inductive finite packs fv hyps env inner_types @@ -892,14 +901,14 @@ let print_object lobj id sp dn fv env = (* print via print_node all the nodes (leafs excluded if bprintleaf is false) *)(* in state *) let rec print_library_segment state bprintleaf dn = List.iter - (function (sp, node) -> - print_if_verbose ("Print_library_segment: " ^ Names.string_of_path sp ^ "\n") ; - print_node node (Nameops.basename sp) sp bprintleaf dn ; + (function ((sp,_ as oname), node) -> + print_if_verbose ("Print_library_segment: " ^ Libnames.string_of_path sp ^ "\n") ; + print_node node (Libnames.basename sp) oname bprintleaf dn ; print_if_verbose "\n" ) (List.rev state) (* print_node node id section_path bprintleaf directory_name *) (* prints a single node (and all it's subnodes via print_library_segment *) -and print_node n id sp bprintleaf dn = +and print_node n id (sp,kn as oname) bprintleaf dn = let module L = Lib in match n with L.Leaf o -> @@ -914,14 +923,14 @@ try begin (* this is an uncooked term *) print_if_verbose ("OK, stampo solo questa volta " ^ Names.string_of_id id ^ "\n") ; - print_object o id sp dn !pvars !cumenv ; + print_object o id oname dn !pvars !cumenv ; printed := id::!printed end else begin (* this is a local term *) print_if_verbose ("OK, stampo " ^ Names.string_of_id id ^ "\n") ; - print_object o id sp dn !pvars !cumenv + print_object o id oname dn !pvars !cumenv end with _ -> print_if_verbose ("EXCEPTION RAISED!!!\n"); end @@ -931,11 +940,11 @@ with _ -> print_if_verbose ("EXCEPTION RAISED!!!\n"); print_if_verbose ("Suppongo gia' stampato " ^ Names.string_of_id id ^ "\n") end end - | L.OpenedSection (dir,_) -> - let id = snd (Nameops.split_dirpath dir) in + | L.OpenedSection ((dir,_),_) -> + let id = snd (Libnames.split_dirpath dir) in print_if_verbose ("OpenDir " ^ Names.string_of_id id ^ "\n") | L.ClosedSection (_,dir,state) -> - let id = snd (Nameops.split_dirpath dir) in + let id = snd (Libnames.split_dirpath dir) in print_if_verbose("ClosedDir " ^ Names.string_of_id id ^ "\n") ; if bprintleaf then begin @@ -948,8 +957,12 @@ with _ -> print_if_verbose ("EXCEPTION RAISED!!!\n"); | he::tl -> pvars := tl end ; print_if_verbose "/ClosedDir\n" - | L.Module s -> - print_if_verbose ("Module " ^ (Names.string_of_dirpath s) ^ "\n") + | L.CompilingModule (s,_) -> + print_if_verbose ("Library " ^ (Names.string_of_dirpath s) ^ "\n") + | L.OpenedModtype ((s,_), _) -> + print_if_verbose ("Open Module Type " ^ (Names.string_of_dirpath s) ^ "\n") + | L.OpenedModule ((s,_), _) -> + print_if_verbose ("Open Module " ^ (Names.string_of_dirpath s) ^ "\n") | L.FrozenState _ -> print_if_verbose ("FrozenState\n") ;; @@ -977,20 +990,23 @@ let print_closed_section s ls dn = (* and terms of the module d *) (* Note: the terms are printed in their uncooked form plus the informations *) (* on the parameters of their most cooked form *) -let printModule (loc,qid) dn = +let printLibrary (loc,qid) dn = + let module N = Names in let module L = Library in - let module N = Nametab in + let module Ln = Libnames in + let module Dm = Declaremods in + let module Nt = Nametab in let module X = Xml in let (_,dir_path,_) = L.locate_qualified_library qid in - let str = N.string_of_qualid qid in - let ls = L.module_segment (Some dir_path) in - print_if_verbose ("MODULE_BEGIN " ^ str ^ " " ^ - (L.module_full_filename dir_path) ^ "\n") ; + let str = Ln.string_of_qualid qid in + let ls = Dm.module_objects (N.MPfile dir_path) in + print_if_verbose ("LIBRARY_BEGIN " ^ str ^ " " ^ + (L.library_full_filename dir_path) ^ "\n") ; print_closed_section str (List.rev ls) dn ; - print_if_verbose ("MODULE_END " ^ str ^ " " ^ - (L.module_full_filename dir_path) ^ "\n") + print_if_verbose ("LIBRARY_END " ^ str ^ " " ^ + (L.library_full_filename dir_path) ^ "\n") ;; (* printSection identifier directory_name *) @@ -1003,30 +1019,30 @@ let printModule (loc,qid) dn = let printSection id dn = let module L = Library in let module N = Names in - let module No = Nameops in + let module Ln = Libnames in let module X = Xml in let sp = Lib.make_path id in let ls = let rec find_closed_section = function [] -> raise Not_found - | (_,Lib.ClosedSection (_,dir,ls))::_ when snd (No.split_dirpath dir) = id + | (_,Lib.ClosedSection (_,dir,ls))::_ when snd (Ln.split_dirpath dir) = id -> ls | _::t -> find_closed_section t in - print_string ("Searching " ^ Names.string_of_path sp ^ "\n") ; + print_string ("Searching " ^ Ln.string_of_path sp ^ "\n") ; find_closed_section (Lib.contents_after None) in let str = N.string_of_id id in - print_if_verbose ("SECTION_BEGIN " ^ str ^ " " ^ N.string_of_path sp ^ "\n"); + print_if_verbose ("SECTION_BEGIN " ^ str ^ " " ^ Ln.string_of_path sp ^ "\n"); print_closed_section str ls dn ; - print_if_verbose ("SECTION_END " ^ str ^ " " ^ N.string_of_path sp ^ "\n") + print_if_verbose ("SECTION_END " ^ str ^ " " ^ Ln.string_of_path sp ^ "\n") ;; (* print All () prints what is the structure of the current environment of *) (* Coq. No terms are printed. Useful only for debugging *) let printAll () = - let state = Library.module_segment None in + let state = Lib.contents_after None in let oldverbose = !verbose in verbose := true ; print_library_segment state false None ; diff --git a/contrib/xml/xmlcommand.mli b/contrib/xml/xmlcommand.mli index 07335ea6a..485bb93c2 100644 --- a/contrib/xml/xmlcommand.mli +++ b/contrib/xml/xmlcommand.mli @@ -28,7 +28,7 @@ (* Note: it is printed only (and directly) the most cooked available *) (* form of the definition (all the parameters are *) (* lambda-abstracted, but the object can still refer to variables) *) -val print : Nametab.qualid Util.located -> string option -> unit +val print : Libnames.qualid Util.located -> string option -> unit (* show dest *) (* where dest is either None (for stdout) or (Some filename) *) @@ -39,14 +39,14 @@ val show : string option -> unit (* Coq. No terms are printed. Useful only for debugging *) val printAll : unit -> unit -(* printModule identifier directory_name *) -(* where identifier is the qualified name of a module d *) +(* printLibrary identifier directory_name *) +(* where identifier is the qualified name of a library d *) (* and directory_name is the directory in which to root all the xml files *) (* prints all the xml files and directories corresponding to the subsections *) -(* and terms of the module d *) +(* and terms of the library d *) (* Note: the terms are printed in their uncooked form plus the informations *) (* on the parameters of their most cooked form *) -val printModule : Nametab.qualid Util.located -> string option -> unit +val printLibrary : Libnames.qualid Util.located -> string option -> unit (* printSection identifier directory_name *) (* where identifier is the name of a closed section d *) diff --git a/contrib/xml/xmlentries.ml4 b/contrib/xml/xmlentries.ml4 index d81269702..876da02e2 100644 --- a/contrib/xml/xmlentries.ml4 +++ b/contrib/xml/xmlentries.ml4 @@ -88,7 +88,7 @@ VERNAC COMMAND EXTEND Xml | [ "Print" "XML" "All" ] -> [ Xmlcommand.printAll () ] | [ "Print" "XML" "Module" diskname(dn) qualid(id) ] -> - [ Xmlcommand.printModule id dn ] + [ Xmlcommand.printLibrary id dn ] | [ "Print" "XML" "Section" diskname(dn) ident(id) ] -> [ Xmlcommand.printSection id dn ] diff --git a/dev/.cvsignore b/dev/.cvsignore index 6ae156b95..e494556aa 100644 --- a/dev/.cvsignore +++ b/dev/.cvsignore @@ -1 +1,3 @@ +other ocamldebug-v7 +debug_* \ No newline at end of file diff --git a/dev/Makefile.common b/dev/Makefile.common new file mode 100644 index 000000000..1ff5cf799 --- /dev/null +++ b/dev/Makefile.common @@ -0,0 +1,52 @@ +# this Makefile contains goals common for directory and main devel makefiles + +ifndef TOPDIR +TOPDIR=.. +endif + +ifndef BASEDIR +BASEDIR= +endif + +# the following entries are used to make synchronize two source trees +# (on big computer and on a laptop for example) + +OTHER_FILE=$(TOPDIR)/dev/other +OTHER=$(shell cat $(OTHER_FILE)) + +# this is a directory of useful temporary things +WORKDIR=tmp + +ifneq (,$(findstring n,$(MAKEFLAGS))) +NFLAG=-n +else +NFLAG= +endif + +check_other: + +@(if [ "$(OTHER)" = "" ] ; then \ + echo You must put the ssh path to the other Coq source in $(OTHER_FILE) ; \ + echo For example: chrzaszc@ruta:coq/V7 ; \ + exit 1; \ + fi) + +get: check_other + +rsync -Cauvz $(NFLAG) $(OTHER)/ $(TOPDIR)/ + +@(if [ -d $(TOPDIR)/$(WORKDIR) ]; then \ + rsync -auvz $(NFLAG) $(OTHER)/tmp/ $(TOPDIR)/tmp/ ; \ + fi) + +put: check_other + +rsync -Cauvz $(NFLAG) $(TOPDIR)/ $(OTHER)/ + +@(if [ -d $(TOPDIR)/$(WORKDIR) ]; then \ + rsync -auvz $(NFLAG) $(TOPDIR)/tmp/ $(OTHER)/tmp/ ; \ + fi) + +sync: get put + + +conflicts: + cvs status | grep File | grep conflicts | less + +confl: conflicts + diff --git a/dev/Makefile.devel b/dev/Makefile.devel new file mode 100644 index 000000000..78318d59b --- /dev/null +++ b/dev/Makefile.devel @@ -0,0 +1,46 @@ +# to be linked to makefile (lowercase - takes precedence over Makefile) +# in main directory + +TOPDIR=. +BASEDIR= + +SOURCEDIRS=lib kernel library pretyping parsing proofs tactics toplevel + +default: noargument + +# set the devel makefile +setup-devel: + @ln -sfv dev/Makefile.devel makefile + @(for i in $(SOURCEDIRS); do \ + (cd $(TOPDIR)/$$i; ln -sfv ../dev/Makefile.dir Makefile) \ + done) + +clean-devel: + echo rm -f makefile .depend.devel + echo rm -f $(foreach dir,$(SOURCEDIRS), $(TOPDIR)/$(dir)/Makefile) + +coqtop: bin/coqtop.byte + ln -sf bin/coqtop.byte coqtop + +quick: + $(MAKE) states BEST=byte + +include Makefile + +include $(TOPDIR)/dev/Makefile.common + +# this file is better described in dev/Makefile.dir +include .depend.devel + +#if dev/Makefile.local exists, it is included +ifneq ($(wildcard $(TOPDIR)/dev/Makefile.local),) +include $(TOPDIR)/dev/Makefile.local +endif + +#runs coqtop with all theories required +total: + ledit ./bin/coqtop.byte $(foreach th,$(THEORIESVO),-require $(notdir $(basename $(th)))) + +#runs coqtop storing using the history file +run: $(TOPDIR)/coqtop + ledit -h $(TOPDIR)/dev/debug_history -x $(TOPDIR)/coqtop diff --git a/dev/Makefile.dir b/dev/Makefile.dir new file mode 100644 index 000000000..df205c1f6 --- /dev/null +++ b/dev/Makefile.dir @@ -0,0 +1,102 @@ +# make a link to this file if you are working hard in one directory of Coq +# ln -s ../dev/Makefile.dir Makefile +# if you are working in a sub/dir/ make a link to dev/Makefile.subdir instead +# this Makefile provides many useful facilities to develop Coq +# it is not completely compatible with .ml4 files unfortunately + +ifndef TOPDIR +TOPDIR=.. +endif + +# this complicated thing should work for subsubdirs as well +BASEDIR=$(shell (dir=`pwd`; cd $(TOPDIR); top=`pwd`; echo $$dir | sed -e "s|$$top/||")) + +noargs: dir + +test-dir: + @echo TOPDIR=$(TOPDIR) + @echo BASEDIR=$(BASEDIR) + +include $(TOPDIR)/dev/Makefile.common + +# make this directory +dir: + $(MAKE) -C $(TOPDIR) $(notdir $(BASEDIR)) + +# make all cmo's in this directory. Useful in case the main Makefile is not +# up-to-date +all: + @(for i in *.ml{,4}; do \ + echo -n $(BASEDIR)/`basename $$i .ml`.cmo "" ; \ + done \ + | xargs $(MAKE) -C $(TOPDIR) ) + +# lists all files that should be compiled in this directory +list: + @(for i in *.mli; do \ + ls -l `basename $$i .mli`.cmi; \ + done) + @(for i in *.ml{,4}; do \ + ls -l `basename $$i .ml`.cmo; \ + done) + + +clean:: + rm -f *.cmi *.cmo *.cmx *.o + +# if grammar.cmo files cannot be compiled and main .depend cannot be +# rebuilt, this is quite useful +depend: + (cd $(TOPDIR); ocamldep -I $(BASEDIR) $(BASEDIR)/*.ml $(BASEDIR)/*.mli > .depend.devel) + +# displays the dependency graph of the current directory (vertically, +# unlike in doc/) +graph: + (ocamldep *.ml *.mli | ocamldot | dot -Tps | gv -) & + + +# the pretty entry draws a dependency graph marking red those nodes +# which do not have their .cmo files + +.INTERMEDIATE: depend.dot depend.2.dot +.PHONY: depend.ps + +depend.dot: + ocamldep *.ml *.mli | ocamldot > $@ + +depend.2.dot: depend.dot + (i=`cat $< | wc -l`; i=`expr $$i - 1`; head -n $$i $<) > $@ + (for ml in *.ml; do \ + base=`basename $$ml .ml`; \ + fst=`echo $$base | cut -c1 | tr [:lower:] [:upper:]`; \ + rest=`echo $$base | cut -c2-`; \ + name=`echo $$fst $$rest | tr -d " "`; \ + cmo=$$base.cmo; \ + if [ ! -e $$cmo ]; then \ + echo \"$$name\" [color=red]\; >> $@;\ + fi;\ + done;\ + echo } >> $@) + +depend.ps: depend.2.dot + dot -Tps $< > $@ + +clean:: + rm -f depend.ps + +pretty: depend.ps + (gv -spartan $<; rm $<) & +# gv -spartan $< & + + +# this is not perfect bot mostly WORKS! It just calls the main makefile + +%.cmi: *.mli + $(MAKE) -C $(TOPDIR) $(BASEDIR)/$@ + +SOURCES=$(wildcard *.mli) $(wildcard *.ml) $(wildcard *.ml4) +%.cmo: $(SOURCES) + $(MAKE) -C $(TOPDIR) $(BASEDIR)/$@ + +coqtop: + $(MAKE) -C $(TOPDIR) bin/coqtop.byte diff --git a/dev/Makefile.subdir b/dev/Makefile.subdir new file mode 100644 index 000000000..45358c426 --- /dev/null +++ b/dev/Makefile.subdir @@ -0,0 +1,7 @@ +# if you work in a sub/sub-rectory of Coq +# you should make a link to that makefile +# ln -s ../../dev/Makefile.subdir Makefile +# in order to have all the facilities of dev/Makefile.dir + +TOPDIR=../.. +include $(TOPDIR)/dev/Makefile.dir diff --git a/dev/base_include b/dev/base_include index a9d8d7812..83d967ce4 100644 --- a/dev/base_include +++ b/dev/base_include @@ -6,10 +6,20 @@ #use "top_printers.ml";; #install_printer (* identifier *) prid;; +#install_printer (* label *) prlab;; +#install_printer prmsid;; +#install_printer prmbid;; +#install_printer prdir;; +#install_printer prmp;; #install_printer (* section_path *) prsp;; #install_printer (* qualid *) prqualid;; +#install_printer (* kernel_name *) prkn;; #install_printer (* constr *) print_pure_constr;; +(* parsing of names *) + +let qid = Libnames.qualid_of_string;; + (* parsing of terms *) let parse_com = Pcoq.parse_string Pcoq.Constr.constr;; @@ -38,7 +48,7 @@ let constr_of_string s = open Declarations;; let constbody_of_string s = - let b = Global.lookup_constant (path_of_string s) in + let b = Global.lookup_constant (Nametab.locate_constant (qualid_of_sp (path_of_string s))) in Util.out_some b.const_body;; (* Get the current goal *) diff --git a/dev/changements.txt b/dev/changements.txt index c75774ee5..d1df2a810 100644 --- a/dev/changements.txt +++ b/dev/changements.txt @@ -1,3 +1,188 @@ +CHANGES DUE TO INTRODUCTION OF MODULES +====================================== + +1.Kernel +-------- + + The module level has no effect on constr except for the structure of +section_path. The type of unique names for constructions (what +section_path served) is now called a kernel name and is defined by + +type uniq_ident = int * string * dir_path (* int may be enough *) +type module_path = + | MPfile of dir_path (* reference to physical module, e.g. file *) + | MPbound of uniq_ident (* reference to a module parameter in a functor *) + | MPself of uniq_ident (* reference to one of the containing module *) + | MPdot of module_path * label +type label = identifier +type kernel_name = module_path * dir_path * label + ^^^^^^^^^^^ ^^^^^^^^ ^^^^^ + | | \ + | | the base name + | \ + / the (true) section path + example: (non empty only inside open sections) + L = (* i.e. some file of logical name L *) + struct + module A = struct Def a = ... end + end + M = (* i.e. some file of logical name M *) + struct + Def t = ... + N = functor (X : sig module T = struct Def b = ... end end) -> struct + module O = struct + Def u = ... + end + Def x := ... .t ... .O.u ... X.T.b ... L.A.a + + and are self-references, X is a bound reference and L is a +reference to a physical module. + + Notice that functor application is not part of a path: it must be +named by a "module M = F(A)" declaration to be used in a kernel +name. + + Notice that Jacek chose a practical approach, making directories not +modules. Another approach could have been to replace the constructor +MPfile by a constant constructor MProot representing the root of the +world. + + Other relevant informations are in kernel/entries.ml (type +module_expr) and kernel/declarations.ml (type module_body and +module_type_body). + +2. Library +---------- + +i) tables +[Summaries] - the only change is the special treatment of the +global environmet. + +ii) objects +[Libobject] declares persistent objects, given with methods: + + * cache_function specifying how to add the object in the current + scope; + * load_function, specifying what to do when the module + containing the object is loaded; + * open_function, specifying what to do when the module + containing the object is opened (imported); + * classify_function, specyfying what to do with the object, + when the current module (containing the object) is ended. + * subst_function + * export_function, to signal end_section survival + +(Almost) Each of these methods is called with a parameter of type +object_name = section_path * kernel_name +where section_path is the full user name of the object (such as +Coq.Init.Datatypes.Fst) and kernel_name is its substitutive internal +version such as (MPself,[],"Fst") (see above) + + +What happens at the end of an interactive module ? +================================================== +(or when a file is stored and reloaded from disk) + +All summaries (except Global environment) are reverted to the state +from before the beginning of the module, and: + +a) the objects (again, since last Declaremods.start_module or + Library.start_library) are classified using the classify_function. + To simplify consider only those who returned Substitute _ or Keep _. + +b) If the module is not a functor, the subst_function for each object of + the first group is called with the substitution + [MPself "" |-> MPfile "Coq.Init.Datatypes"]. + Then the load_function is called for substituted objects and the + "keep" object. + (If the module is a library the substitution is done at reloading). + +c) The objects which returned substitute are stored in the modtab + together with the self ident of the module, and functor argument + names if the module was a functor. + + They will be used (substituted and loaded) when a command like + Module M := F(N) or + Module Z := N + is evaluated + + +The difference between "substitute" and "keep" objects +======================================================== +i) The "keep" objects can _only_ reference other objects by section_paths +and qualids. They do not need the substitution function. + +They will work after end_module (or reloading a compiled library), +because these operations do not change section_path's + +They will obviously not work after Module Z:=N. + +These would typically be grammar rules, pretty printing rules etc. + + + +ii) The "substitute" objects can _only_ reference objects by +kernel_names. They must have a valid subst_function. + +They will work after end_module _and_ after Module Z:=N or +Module Z:=F(M). + + + +Other kinds of objects: +iii) "Dispose" - objects which do not survive end_module + As a consequence, objects which reference other objects sometimes + by kernel_names and sometimes by section_path must be of this kind... + +iv) "Anticipate" - objects which must be treated individually by + end_module (typically "REQUIRE" objects) + + + +Writing subst_thing functions +============================= +The subst_thing shoud not copy the thing if it hasn't actually +changed. There are some cool emacs macros in dev/objects.el +to help writing subst functions this way quickly and without errors. +Also there are *_smartmap functions in Util. + +The subst_thing functions are already written for many types, +including constr (Term.subst_mps), +global_reference (Libnames.subst_global), +rawconstr (Rawterm.subst_raw) etc + +They are all (apart from constr, for now) written in the non-copying +way. + + +Nametab +======= + +Nametab has been made more uniform. For every kind of thing there is +only one "push" function and one "locate" function. + + +Lib +=== + +library_segment is now a list of object_name * library_item, where +object_name = section_path * kernel_name (see above) + +New items have been added for open modules and module types + + +Declaremods +========== +Functions to declare interactive and noninteractive modules and module +types. + + +Library +======= +Uses Declaremods to actually communicate with Global and to register +objects. + + MAIN CHANGES FROM COQ V7.3 ========================== diff --git a/dev/db b/dev/db index f9a7f8d84..5440ff8f1 100644 --- a/dev/db +++ b/dev/db @@ -1,6 +1,12 @@ load_printer "gramlib.cma" load_printer "top_printers.cmo" install_printer Top_printers.prid +install_printer Top_printers.prlab +install_printer Top_printers.prmsid +install_printer Top_printers.prmbid +install_printer Top_printers.prdir +install_printer Top_printers.prmp +install_printer Top_printers.prkn install_printer Top_printers.prsp install_printer Top_printers.prqualid install_printer Top_printers.prast @@ -24,4 +30,5 @@ install_printer Top_printers.prevc install_printer Top_printers.prwc install_printer Top_printers.prclenv +install_printer Top_printers.pr_obj diff --git a/dev/include b/dev/include index 165853221..84f45fb2a 100644 --- a/dev/include +++ b/dev/include @@ -27,3 +27,4 @@ #install_printer (* env *) ppenv;; #install_printer (* tactic *) pptac;; +#install_printer (* object *) pr_obj;; diff --git a/dev/objects.el b/dev/objects.el new file mode 100644 index 000000000..4e531fd7b --- /dev/null +++ b/dev/objects.el @@ -0,0 +1,137 @@ +; functions to change old style object declaration to new style + +(defun repl-open nil + (interactive) + (query-replace-regexp + "open_function\\([ ]*\\)=\\([ ]*\\)cache_\\([a-zA-Z0-9'_]*\\)\\( *\\);" + "open_function\\1=\\2(fun i o -> if i=1 then cache_\\3 o)\\4;") + ) + +(global-set-key [f6] 'repl-open) + +(defun repl-load nil + (interactive) + (query-replace-regexp + "load_function\\([ ]*\\)=\\([ ]*\\)cache_\\([a-zA-Z0-9'_]*\\)\\( *\\);" + "load_function\\1=\\2(fun _ -> cache_\\3)\\4;") + ) + +(global-set-key [f7] 'repl-load) + +(defun repl-decl nil + (interactive) + (query-replace-regexp + "\\(Libobject\.\\)?declare_object[ + ]*([ ]*\\(.*\\)[ + ]*,[ ]* +\\([ ]*\\){\\([ ]*\\)\\([^ ][^}]*\\)}[ ]*)" + + "\\1declare_object {(\\1default_object \\2) with +\\3 \\4\\5}") +; "|$1=\\1|$2=\\2|$3=\\3|$4=\\4|") +) + +(global-set-key [f9] 'repl-decl) + +; eval the above and try f9 f6 f7 on the following: + +let (inThing,outThing) = + declare_object + ("THING", + { load_function = cache_thing; + cache_function = cache_thing; + open_function = cache_thing; + export_function = (function x -> Some x) + }) + + +;%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +;%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +; functions helping writing non-copying substitutions + +(defun make-subst (name) + (interactive "s") + (defun f (l) + (save-excursion + (query-replace-regexp + (concat "\\([a-zA-z_0-9]*\\)[ ]*:[ ]*" + (car l) + "\\([ ]*;\\|[ +]*\}\\)") + (concat "let \\1\' = " (cdr l) " " name "\\1 in")) + ) + ) + (mapcar 'f '(("constr"."subst_mps subst") + ("Coqast.t"."subst_ast subst") + ("Coqast.t list"."list_smartmap (subst_ast subst)") + ("'pat"."subst_pat subst") + ("'pat unparsing_hunk"."subst_hunk subst_pat subst") + ("'pat unparsing_hunk list"."list_smartmap (subst_hunk subst_pat subst)") + ("'pat syntax_entry"."subst_syntax_entry subst_pat subst") + ("'pat syntax_entry list"."list_smartmap (subst_syntax_entry subst_pat subst)") + ("constr option"."option_smartmap (subst_mps subst)") + ("constr list"."list_smartmap (subst_mps subst)") + ("constr array"."array_smartmap (subst_mps subst)") + ("global_reference"."subst_global subst") + ("extended_global_reference"."subst_ext subst") + ("obj_typ"."subst_obj subst") + ) + ) + ) + + +(global-set-key [f2] 'make-subst) + +(defun make-if (name) + (interactive "s") + (save-excursion + (query-replace-regexp + "\\([a-zA-z_0-9]*\\)[ ]*:[ ]*['a-zA-z_. ]*\\(;\\|[ +]*\}\\)" + (concat "&& \\1\' == " name "\\1") + ) + ) + ) + +(global-set-key [f3] 'make-if) + +(defun make-record nil + (interactive) + (save-excursion + (query-replace-regexp + "\\([a-zA-z_0-9]*\\)[ ]*:[ ]*['a-zA-z_. ]*\\(;\\|[ +]*\}\\)" + (concat "\\1 = \\1\' ;") + ) + ) + ) + +(global-set-key [f4] 'make-record) + +(defun make-prim nil + (interactive) + (save-excursion (query-replace-regexp "\\<[a-zA-Z'_0-9]*\\>" "\\&'")) + ) + +(global-set-key [f5] 'make-prim) + + +; eval the above, yank the text below and do +; paste f2 morph. +; paste f3 morph. +; paste f4 + + lem : constr; + profil : bool list; + arg_types : constr list; + lem2 : constr option } + + +; and you almost get Setoid_replace.subst_morph :) + +; and now f5 on this: + + (ref,(c1,c2)) + + + diff --git a/dev/ocamldebug-v7.template b/dev/ocamldebug-v7.template index 115cba72f..cbb72ce27 100644 --- a/dev/ocamldebug-v7.template +++ b/dev/ocamldebug-v7.template @@ -5,9 +5,9 @@ export COQTOP=COQTOPDIRECTORY export COQLIB=COQLIBDIRECTORY export COQTH=$COQLIB/theories -export CAMLP4LIB=CAMLP4LIBDIRECTORY CAMLBIN=CAMLBINDIRECTORY OCAMLDEBUG=$CAMLBIN/ocamldebug +export CAMLP4LIB=`$CAMLBIN/camlp4 -where` args="" coqdebug="no" diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 417bbfce2..0339a8fdd 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -12,6 +12,7 @@ open System open Pp open Ast open Names +open Libnames open Nameops open Sign open Univ @@ -39,9 +40,22 @@ let pppattern = (fun x -> pp(pr_pattern x)) let pptype = (fun x -> pp(prtype x)) let prid id = pp (pr_id id) +let prlab l = pp (pr_lab l) + +let prmsid msid = pp (str (string_of_msid msid)) +let prmbid mbid = pp (str (string_of_mbid mbid)) + +let prdir dir = pp (pr_dirpath dir) + +let prmp mp = pp(str (string_of_mp mp)) +let prkn kn = pp(pr_kn kn) + +let prsp sp = pp(pr_sp sp) + +let prqualid qid = pp(pr_qualid qid) let prconst (sp,j) = - pp (str"#" ++ pr_sp sp ++ str"=" ++ prterm j.uj_val) + pp (str"#" ++ pr_kn sp ++ str"=" ++ prterm j.uj_val) let prvar ((id,a)) = pp (str"#" ++ pr_id id ++ str":" ++ prterm a) @@ -51,10 +65,6 @@ let genprj f j = let prj j = pp (genprj prjudge j) -let prsp sp = pp(pr_sp sp) - -let prqualid qid = pp(Nametab.pr_qualid qid) - let prgoal g = pp(prgl g) let prsigmagoal g = pp(prgl (sig_it g)) @@ -81,6 +91,8 @@ let ppenv e = pp (pr_rel_context e (rel_context e)) let pptac = (fun x -> pp(Pptactic.pr_raw_tactic x)) +let pr_obj obj = Format.print_string (Libobject.object_tag obj) + let cnt = ref 0 let constr_display csr = @@ -99,11 +111,11 @@ let constr_display csr = ^(term_display t)^","^(term_display c)^")" | App (c,l) -> "App("^(term_display c)^","^(array_display l)^")\n" | Evar (e,l) -> "Evar("^(string_of_int e)^","^(array_display l)^")" - | Const c -> "Const("^(string_of_path c)^")" + | Const c -> "Const("^(string_of_kn c)^")" | Ind (sp,i) -> - "MutInd("^(string_of_path sp)^","^(string_of_int i)^")" + "MutInd("^(string_of_kn sp)^","^(string_of_int i)^")" | Construct ((sp,i),j) -> - "MutConstruct(("^(string_of_path sp)^","^(string_of_int i)^")," + "MutConstruct(("^(string_of_kn sp)^","^(string_of_int i)^")," ^(string_of_int j)^")" | Case (ci,p,c,bl) -> "MutCase(,"^(term_display p)^","^(term_display c)^"," @@ -243,13 +255,15 @@ let print_pure_constr csr = | Name id -> print_string (string_of_id id) | Anonymous -> print_string "_" (* Remove the top names for library and Scratch to avoid long names *) - and sp_display sp = let ls = - match List.rev (List.map string_of_id (repr_dirpath (dirpath sp))) with - ("Scratch"::l)-> l - | ("Coq"::_::l) -> l - | l -> l - in List.iter (fun x -> print_string x; print_string ".") ls; - print_string (string_of_id (basename sp)) + and sp_display sp = + let dir,l = decode_kn sp in + let ls = + match List.rev (List.map string_of_id (repr_dirpath dir)) with + ("Top"::l)-> l + | ("Coq"::_::l) -> l + | l -> l + in List.iter (fun x -> print_string x; print_string ".") ls; + print_string (string_of_id l) in box_display csr; print_flush() @@ -272,3 +286,4 @@ let _ = print_pure_constr (Astterm.interp_constr evmap sign c)) | _ -> bad_vernac_args "PrintPureConstr") *) + diff --git a/doc/library.dep.ps b/doc/library.dep.ps index 4f9297273..fc4b9f01e 100644 --- a/doc/library.dep.ps +++ b/doc/library.dep.ps @@ -1,9 +1,9 @@ %!PS-Adobe-2.0 %%Creator: dot version gviz 1.7.3b (Mon Feb 26 14:10:22 EST 2001) -%%For: (herbelin) Hugo Herbelin +%%For: (jacek) Jacek Chrzaszcz %%Title: G %%Pages: (atend) -%%BoundingBox: 36 36 543 181 +%%BoundingBox: 36 36 576 142 %%EndComments %%BeginProlog save @@ -150,257 +150,293 @@ def %%EndResource %%EndSetup %%Page: 1 1 -%%PageBoundingBox: 36 36 543 181 +%%PageBoundingBox: 36 36 576 142 %%PageOrientation: Portrait gsave -35 35 508 146 boxprim clip newpath +35 35 541 107 boxprim clip newpath 36 36 translate 0 0 1 beginpage +0.7317 set_scale 0 0 translate 0 rotate -[ /CropBox [36 36 543 181] /PAGES pdfmark +[ /CropBox [36 36 576 142] /PAGES pdfmark 0.000 0.000 0.000 graphcolor 14.00 /Times-Roman set_font -% States +% Summary gsave 10 dict begin -40 126 27 18 ellipse_path +580 126 37 18 ellipse_path stroke gsave 10 dict begin -40 127 moveto (States) 33 14.00 -0.50 alignedtext +580 127 moveto (Summary) 54 14.00 -0.50 alignedtext end grestore end grestore -% Library +% Libnames gsave 10 dict begin -144 126 31 18 ellipse_path +692 76 38 18 ellipse_path stroke gsave 10 dict begin -144 127 moveto (Library) 41 14.00 -0.50 alignedtext +692 77 moveto (Libnames) 55 14.00 -0.50 alignedtext end grestore end grestore -% States -> Library -newpath 67 126 moveto -78 126 91 126 103 126 curveto +% Summary -> Libnames +newpath 608 114 moveto +622 107 640 99 655 92 curveto stroke -newpath 103 124 moveto -113 126 lineto -103 129 lineto +newpath 654 90 moveto +664 88 lineto +656 94 lineto closepath gsave 0 setgray stroke grestore fill -% Global +% States gsave 10 dict begin -246 99 29 18 ellipse_path +40 72 27 18 ellipse_path stroke gsave 10 dict begin -246 100 moveto (Global) 38 14.00 -0.50 alignedtext +40 73 moveto (States) 33 14.00 -0.50 alignedtext end grestore end grestore -% Library -> Global -newpath 172 118 moveto -184 115 197 112 210 108 curveto +% Library +gsave 10 dict begin +139 72 31 18 ellipse_path stroke -newpath 209 106 moveto -219 106 lineto -210 111 lineto +gsave 10 dict begin +139 73 moveto (Library) 41 14.00 -0.50 alignedtext +end grestore +end grestore + +% States -> Library +newpath 67 72 moveto +77 72 88 72 98 72 curveto +stroke +newpath 98 70 moveto +108 72 lineto +98 75 lineto closepath gsave 0 setgray stroke grestore fill -% Lib +% Declaremods gsave 10 dict begin -246 45 27 18 ellipse_path +253 72 47 18 ellipse_path stroke gsave 10 dict begin -246 46 moveto (Lib) 19 14.00 -0.50 alignedtext +253 73 moveto (Declaremods) 73 14.00 -0.50 alignedtext end grestore end grestore -% Library -> Lib -newpath 165 112 moveto -171 109 176 105 180 102 curveto -193 91 204 79 216 69 curveto -218 67 220 65 223 63 curveto +% Library -> Declaremods +newpath 170 72 moveto +178 72 187 72 196 72 curveto stroke -newpath 219 62 moveto -229 59 lineto -222 67 lineto +newpath 196 70 moveto +206 72 lineto +196 75 lineto closepath gsave 0 setgray stroke grestore fill % Nametab gsave 10 dict begin -349 99 35 18 ellipse_path +469 80 35 18 ellipse_path stroke gsave 10 dict begin -349 100 moveto (Nametab) 50 14.00 -0.50 alignedtext +469 81 moveto (Nametab) 50 14.00 -0.50 alignedtext end grestore end grestore +% Nametab -> Summary +newpath 497 91 moveto +511 97 528 104 543 110 curveto +stroke +newpath 544 108 moveto +552 114 lineto +542 112 lineto +closepath +gsave 0 setgray stroke grestore fill + % Nameops gsave 10 dict begin -460 126 36 18 ellipse_path +580 72 36 18 ellipse_path stroke gsave 10 dict begin -460 127 moveto (Nameops) 52 14.00 -0.50 alignedtext +580 73 moveto (Nameops) 52 14.00 -0.50 alignedtext end grestore end grestore % Nametab -> Nameops -newpath 381 107 moveto -392 110 405 113 417 116 curveto +newpath 504 77 moveto +514 76 524 75 534 75 curveto stroke -newpath 418 114 moveto -427 118 lineto -417 118 lineto +newpath 534 73 moveto +544 74 lineto +534 77 lineto closepath gsave 0 setgray stroke grestore fill -% Summary +% Lib gsave 10 dict begin -460 72 37 18 ellipse_path +366 72 27 18 ellipse_path stroke gsave 10 dict begin -460 73 moveto (Summary) 54 14.00 -0.50 alignedtext +366 73 moveto (Lib) 19 14.00 -0.50 alignedtext end grestore end grestore -% Nametab -> Summary -newpath 381 91 moveto -392 89 405 85 417 82 curveto +% Declaremods -> Lib +newpath 300 72 moveto +310 72 320 72 329 72 curveto stroke -newpath 416 80 moveto -426 80 lineto -417 85 lineto +newpath 329 70 moveto +339 72 lineto +329 75 lineto closepath gsave 0 setgray stroke grestore fill -% Global -> Nametab -newpath 276 99 moveto -285 99 294 99 303 99 curveto +% Global +gsave 10 dict begin +366 126 29 18 ellipse_path stroke -newpath 303 97 moveto -313 99 lineto -303 102 lineto -closepath -gsave 0 setgray stroke grestore fill +gsave 10 dict begin +366 127 moveto (Global) 38 14.00 -0.50 alignedtext +end grestore +end grestore -% Lib -> Nametab -newpath 267 56 moveto -281 64 300 73 316 82 curveto +% Declaremods -> Global +newpath 282 86 moveto +298 94 318 103 334 111 curveto stroke -newpath 316 79 moveto -324 86 lineto -314 84 lineto +newpath 334 108 moveto +342 115 lineto +332 113 lineto closepath gsave 0 setgray stroke grestore fill % Libobject gsave 10 dict begin -349 45 37 18 ellipse_path +469 26 37 18 ellipse_path stroke gsave 10 dict begin -349 46 moveto (Libobject) 53 14.00 -0.50 alignedtext +469 27 moveto (Libobject) 53 14.00 -0.50 alignedtext end grestore end grestore +% Libobject -> Libnames +newpath 506 29 moveto +540 31 589 36 618 42 curveto +629 44 645 51 658 57 curveto +stroke +newpath 659 55 moveto +667 62 lineto +657 59 lineto +closepath +gsave 0 setgray stroke grestore fill + +% Lib -> Nametab +newpath 393 74 moveto +404 75 415 76 427 77 curveto +stroke +newpath 424 75 moveto +434 77 lineto +424 80 lineto +closepath +gsave 0 setgray stroke grestore fill + % Lib -> Libobject -newpath 273 45 moveto -282 45 292 45 302 45 curveto +newpath 389 62 moveto +401 56 418 49 432 42 curveto stroke -newpath 302 43 moveto -312 45 lineto -302 48 lineto +newpath 431 40 moveto +441 38 lineto +433 44 lineto closepath gsave 0 setgray stroke grestore fill % Impargs gsave 10 dict begin -144 72 33 18 ellipse_path +253 126 33 18 ellipse_path stroke gsave 10 dict begin -144 73 moveto (Impargs) 45 14.00 -0.50 alignedtext +253 127 moveto (Impargs) 45 14.00 -0.50 alignedtext end grestore end grestore +% Impargs -> Lib +newpath 278 114 moveto +295 106 318 95 336 86 curveto +stroke +newpath 334 84 moveto +344 82 lineto +336 89 lineto +closepath +gsave 0 setgray stroke grestore fill + % Impargs -> Global -newpath 174 80 moveto -185 83 198 86 210 89 curveto +newpath 286 126 moveto +299 126 314 126 327 126 curveto stroke -newpath 210 86 moveto -219 92 lineto -209 91 lineto +newpath 326 124 moveto +336 126 lineto +326 129 lineto closepath gsave 0 setgray stroke grestore fill -% Impargs -> Lib -newpath 174 64 moveto -186 61 199 57 212 54 curveto +% Global -> Summary +newpath 396 126 moveto +431 126 492 126 533 126 curveto stroke -newpath 211 52 moveto -221 52 lineto -212 57 lineto +newpath 532 124 moveto +542 126 lineto +532 129 lineto closepath gsave 0 setgray stroke grestore fill % Goptions gsave 10 dict begin -144 18 36 18 ellipse_path +253 18 36 18 ellipse_path stroke gsave 10 dict begin -144 19 moveto (Goptions) 51 14.00 -0.50 alignedtext +253 19 moveto (Goptions) 51 14.00 -0.50 alignedtext end grestore end grestore -% Goptions -> Global -newpath 166 32 moveto -171 36 176 39 180 42 curveto -193 53 204 65 216 75 curveto -218 77 220 78 223 80 curveto -stroke -newpath 223 76 moveto -228 85 lineto -219 80 lineto -closepath -gsave 0 setgray stroke grestore fill - % Goptions -> Lib -newpath 176 26 moveto -188 29 200 32 212 35 curveto +newpath 279 31 moveto +296 39 318 49 336 57 curveto stroke -newpath 212 32 moveto -221 38 lineto -211 37 lineto +newpath 336 54 moveto +344 61 lineto +334 59 lineto closepath gsave 0 setgray stroke grestore fill % Declare gsave 10 dict begin -40 72 32 18 ellipse_path +40 126 32 18 ellipse_path stroke gsave 10 dict begin -40 73 moveto (Declare) 43 14.00 -0.50 alignedtext +40 127 moveto (Declare) 43 14.00 -0.50 alignedtext end grestore end grestore % Declare -> Library -newpath 64 84 moveto -78 92 97 101 113 110 curveto +newpath 63 113 moveto +76 105 93 97 108 89 curveto stroke -newpath 113 107 moveto -121 114 lineto -111 112 lineto +newpath 106 87 moveto +116 84 lineto +109 91 lineto closepath gsave 0 setgray stroke grestore fill % Declare -> Impargs -newpath 72 72 moveto -81 72 92 72 101 72 curveto +newpath 72 126 moveto +109 126 171 126 212 126 curveto stroke -newpath 101 70 moveto -111 72 lineto -101 75 lineto +newpath 210 124 moveto +220 126 lineto +210 129 lineto closepath gsave 0 setgray stroke grestore fill endpage diff --git a/kernel/closure.ml b/kernel/closure.ml index ace0ed39d..463f6f21d 100644 --- a/kernel/closure.ml +++ b/kernel/closure.ml @@ -55,10 +55,10 @@ type evaluable_global_reference = | EvalVarRef of identifier | EvalConstRef of constant -type transparent_state = Idpred.t * Sppred.t +type transparent_state = Idpred.t * KNpred.t -let all_opaque = (Idpred.empty, Sppred.empty) -let all_transparent = (Idpred.full, Sppred.full) +let all_opaque = (Idpred.empty, KNpred.empty) +let all_transparent = (Idpred.full, KNpred.full) module type RedFlagsSig = sig type reds @@ -98,7 +98,7 @@ module RedFlags = (struct let fDELTA = DELTA let fIOTA = IOTA let fZETA = ZETA - let fCONST sp = CONST sp + let fCONST kn = CONST kn let fVAR id = VAR id let no_red = { r_beta = false; @@ -111,9 +111,9 @@ module RedFlags = (struct let red_add red = function | BETA -> { red with r_beta = true } | DELTA -> { red with r_delta = true; r_const = all_transparent } - | CONST sp -> + | CONST kn -> let (l1,l2) = red.r_const in - { red with r_const = l1, Sppred.add sp l2 } + { red with r_const = l1, KNpred.add kn l2 } | IOTA -> { red with r_iota = true } | ZETA -> { red with r_zeta = true } | VAR id -> @@ -123,9 +123,9 @@ module RedFlags = (struct let red_sub red = function | BETA -> { red with r_beta = false } | DELTA -> { red with r_delta = false } - | CONST sp -> + | CONST kn -> let (l1,l2) = red.r_const in - { red with r_const = l1, Sppred.remove sp l2 } + { red with r_const = l1, KNpred.remove kn l2 } | IOTA -> { red with r_iota = false } | ZETA -> { red with r_zeta = false } | VAR id -> @@ -139,11 +139,11 @@ module RedFlags = (struct let red_set red = function | BETA -> incr_cnt red.r_beta beta - | CONST sp -> + | CONST kn -> let (_,l) = red.r_const in - let c = Sppred.mem sp l in + let c = KNpred.mem kn l in incr_cnt c delta - | VAR id -> (* En attendant d'avoir des sp pour les Var *) + | VAR id -> (* En attendant d'avoir des kn pour les Var *) let (l,_) = red.r_const in let c = Idpred.mem id l in incr_cnt c delta @@ -155,7 +155,7 @@ module RedFlags = (struct let red_get_const red = let p1,p2 = red.r_const in let (b1,l1) = Idpred.elements p1 in - let (b2,l2) = Sppred.elements p2 in + let (b2,l2) = KNpred.elements p2 in if b1=b2 then let l1' = List.map (fun x -> EvalVarRef x) l1 in let l2' = List.map (fun x -> EvalConstRef x) l2 in @@ -171,10 +171,10 @@ let betadeltaiotanolet = mkflags [fBETA;fDELTA;fIOTA] let betaiota = mkflags [fBETA;fIOTA] let beta = mkflags [fBETA] let betaiotazeta = mkflags [fBETA;fIOTA;fZETA] -let unfold_red sp = - let flag = match sp with +let unfold_red kn = + let flag = match kn with | EvalVarRef id -> fVAR id - | EvalConstRef sp -> fCONST sp + | EvalConstRef kn -> fCONST kn in (* Remove fZETA for finer behaviour ? *) mkflags [fBETA;flag;fIOTA;fZETA] @@ -223,10 +223,10 @@ let betaiotazeta_red = { r_evar = false; r_iota = true } -let unfold_red sp = - let c = match sp with +let unfold_red kn = + let c = match kn with | EvalVarRef id -> false,[],[id] - | EvalConstRef sp -> false,[sp],[] + | EvalConstRef kn -> false,[kn],[] in { r_beta = true; r_const = c; @@ -236,7 +236,7 @@ let unfold_red sp = (* Sets of reduction kinds. Main rule: delta implies all consts (both global (= by - section_path) and local (= by Rel or Var)), all evars, and zeta (= letin's). + kernel_name) and local (= by Rel or Var)), all evars, and zeta (= letin's). Rem: reduction of a Rel/Var bound to a term is Delta, but reduction of a LetIn expression is Letin reduction *) @@ -284,11 +284,11 @@ let red_local_const = red_delta_set (* to know if a redex is allowed, only a subset of red_kind is used ... *) let red_set red = function | BETA -> incr_cnt red.r_beta beta - | CONST [sp] -> + | CONST [kn] -> let (b,l,_) = red.r_const in - let c = List.mem sp l in + let c = List.mem kn l in incr_cnt ((b & not c) or (c & not b)) delta - | VAR id -> (* En attendant d'avoir des sp pour les Var *) + | VAR id -> (* En attendant d'avoir des kn pour les Var *) let (b,_,l) = red.r_const in let c = List.mem id l in incr_cnt ((b & not c) or (c & not b)) delta @@ -584,8 +584,8 @@ let rec mk_clos e t = | Rel i -> clos_rel e i | Var x -> { norm = Red; term = FFlex (VarKey x) } | Meta _ | Sort _ -> { norm = Norm; term = FAtom t } - | Ind sp -> { norm = Norm; term = FInd sp } - | Construct sp -> { norm = Cstr; term = FConstruct sp } + | Ind kn -> { norm = Norm; term = FInd kn } + | Construct kn -> { norm = Cstr; term = FConstruct kn } | Evar (ev,args) -> { norm = Cstr; term = FEvar (ev,Array.map (mk_clos e) args) } | (Fix _|CoFix _|Lambda _|Prod _) -> @@ -609,9 +609,9 @@ let mk_clos_deep clos_fun env t = | App (f,v) -> { norm = Red; term = FApp (clos_fun env f, Array.map (clos_fun env) v) } - | Const sp -> + | Const kn -> { norm = Red; - term = FFlex (ConstKey sp) } + term = FFlex (ConstKey kn) } | Case (ci,p,c,v) -> { norm = Red; term = FCases (ci, clos_fun env p, clos_fun env c, @@ -903,8 +903,8 @@ let rec knr info m stk = (match get_arg m stk with (Some(depth,arg),s) -> knit info (subs_shift_cons(depth,e,arg)) f s | (None,s) -> (m,s)) - | FFlex(ConstKey sp) when red_set info.i_flags (fCONST sp) -> - (match ref_value_cache info (ConstKey sp) with + | FFlex(ConstKey kn) when red_set info.i_flags (fCONST kn) -> + (match ref_value_cache info (ConstKey kn) with Some v -> kni info v stk | None -> (set_norm m; (m,stk))) | FFlex(VarKey id) when red_set info.i_flags (fVAR id) -> diff --git a/kernel/closure.mli b/kernel/closure.mli index 1360862c9..826d58fba 100644 --- a/kernel/closure.mli +++ b/kernel/closure.mli @@ -24,14 +24,14 @@ val with_stats: 'a Lazy.t -> 'a type evaluable_global_reference = | EvalVarRef of identifier - | EvalConstRef of section_path + | EvalConstRef of kernel_name (*s Delta implies all consts (both global (= by - [section_path]) and local (= by [Rel] or [Var])), all evars, and letin's. + [kernel_name]) and local (= by [Rel] or [Var])), all evars, and letin's. Rem: reduction of a Rel/Var bound to a term is Delta, but reduction of a LetIn expression is Letin reduction *) -type transparent_state = Idpred.t * Sppred.t +type transparent_state = Idpred.t * KNpred.t val all_opaque : transparent_state val all_transparent : transparent_state diff --git a/kernel/conv_oracle.ml b/kernel/conv_oracle.ml index 350e1a5a0..ffa7735e7 100644 --- a/kernel/conv_oracle.ml +++ b/kernel/conv_oracle.ml @@ -12,36 +12,36 @@ open Names open Closure (* Opaque constants *) -let cst_transp = ref Sppred.full +let cst_transp = ref KNpred.full -let set_opaque_const sp = cst_transp := Sppred.remove sp !cst_transp -let set_transparent_const sp = cst_transp := Sppred.add sp !cst_transp +let set_opaque_const kn = cst_transp := KNpred.remove kn !cst_transp +let set_transparent_const kn = cst_transp := KNpred.add kn !cst_transp -let is_opaque_cst sp = not (Sppred.mem sp !cst_transp) +let is_opaque_cst kn = not (KNpred.mem kn !cst_transp) (* Unfold the first only if it is not opaque and the second is opaque *) -let const_order sp1 sp2 = is_opaque_cst sp2 & not (is_opaque_cst sp1) +let const_order kn1 kn2 = is_opaque_cst kn2 & not (is_opaque_cst kn1) (* Opaque variables *) let var_transp = ref Idpred.full -let set_opaque_var sp = var_transp := Idpred.remove sp !var_transp -let set_transparent_var sp = var_transp := Idpred.add sp !var_transp +let set_opaque_var kn = var_transp := Idpred.remove kn !var_transp +let set_transparent_var kn = var_transp := Idpred.add kn !var_transp -let is_opaque_var sp = not (Idpred.mem sp !var_transp) +let is_opaque_var kn = not (Idpred.mem kn !var_transp) let var_order id1 id2 = is_opaque_var id2 & not (is_opaque_var id1) (* *) let oracle_order k1 k2 = match (k1,k2) with - (ConstKey sp1, ConstKey sp2) -> const_order sp1 sp2 + (ConstKey kn1, ConstKey kn2) -> const_order kn1 kn2 | (VarKey id1, VarKey id2) -> var_order id1 id2 | _ -> false (* summary operations *) -let init() = (cst_transp := Sppred.full; var_transp := Idpred.full) +let init() = (cst_transp := KNpred.full; var_transp := Idpred.full) let freeze () = (!var_transp, !cst_transp) let unfreeze (vo,co) = (cst_transp := co; var_transp := vo) diff --git a/kernel/conv_oracle.mli b/kernel/conv_oracle.mli index 20ec5239e..da2bbfa23 100644 --- a/kernel/conv_oracle.mli +++ b/kernel/conv_oracle.mli @@ -12,14 +12,14 @@ open Names open Closure (* Order on section paths for unfolding. - If [oracle_order sp1 sp2] is true, then unfold sp1 first. + If [oracle_order kn1 kn2] is true, then unfold kn1 first. Note: the oracle does not introduce incompleteness, it only tries to postpone unfolding of "opaque" constants. *) val oracle_order : table_key -> table_key -> bool (* Changing the oracle *) -val set_opaque_const : section_path -> unit -val set_transparent_const : section_path -> unit +val set_opaque_const : kernel_name -> unit +val set_transparent_const : kernel_name -> unit val set_opaque_var : identifier -> unit val set_transparent_var : identifier -> unit diff --git a/kernel/cooking.ml b/kernel/cooking.ml index 0313d4d46..616be45f1 100644 --- a/kernel/cooking.ml +++ b/kernel/cooking.ml @@ -81,13 +81,13 @@ let modify_opers replfun (constl,indl,cstrl) = with | Not_found -> c') - | Const sp -> + | Const kn -> (try - (match List.assoc sp constl with + (match List.assoc kn constl with | NOT_OCCUR -> failure () | DO_ABSTRACT (oper',abs_vars) -> mkApp (mkConst oper', abs_vars) - | DO_REPLACE cb -> substrec (replfun (sp,cb))) + | DO_REPLACE cb -> substrec (replfun (kn,cb))) with | Not_found -> c') @@ -98,11 +98,11 @@ let modify_opers replfun (constl,indl,cstrl) = let expmod_constr modlist c = let simpfun = if modlist = ([],[],[]) then fun x -> x else nf_betaiota in - let expfun (sp,cb) = + let expfun (kn,cb) = if cb.const_opaque then errorlabstrm "expmod_constr" (str"Cannot unfold the value of " ++ - str(string_of_path sp) ++ spc () ++ + str(string_of_kn kn) ++ spc () ++ str"You cannot declare local lemmas as being opaque" ++ spc () ++ str"and then require that theorems which use them" ++ spc () ++ str"be transparent"); diff --git a/kernel/declarations.ml b/kernel/declarations.ml index fca3e0a50..67c0cb998 100644 --- a/kernel/declarations.ml +++ b/kernel/declarations.ml @@ -9,6 +9,7 @@ (*i $Id$ i*) (*i*) +open Util open Names open Univ open Term @@ -35,6 +36,11 @@ type recarg = | Mrec of int | Imbr of inductive +let subst_recarg sub r = match r with + | Norec | Mrec _ -> r + | Imbr (kn,i) -> let kn' = subst_kn sub kn in + if kn==kn' then r else Imbr (kn',i) + type wf_paths = recarg Rtree.t let mk_norec = Rtree.mk_node Norec [||] @@ -49,6 +55,8 @@ let dest_subterms p = let (_,cstrs) = Rtree.dest_node p in Array.map (fun t -> Array.to_list (snd (Rtree.dest_node t))) cstrs +let subst_wf_paths sub p = Rtree.smartmap (subst_recarg sub) p + (* [mind_typename] is the name of the inductive; [mind_arity] is the arity generalized over global parameters; [mind_lc] is the list of types of constructors generalized over global parameters and @@ -77,3 +85,78 @@ type mutual_inductive_body = { mind_packets : one_inductive_body array; mind_constraints : constraints; mind_singl : constr option } + +(* TODO: should be changed to non-coping after Term.subst_mps *) +let subst_const_body sub cb = + { const_body = option_app (Term.subst_mps sub) cb.const_body; + const_type = type_app (Term.subst_mps sub) cb.const_type; + const_hyps = (assert (cb.const_hyps=[]); []); + const_constraints = cb.const_constraints; + const_opaque = cb.const_opaque} + +let subst_mind_packet sub mbp = + { mind_consnames = mbp.mind_consnames; + mind_typename = mbp.mind_typename; + mind_nf_lc = + array_smartmap (type_app (Term.subst_mps sub)) mbp.mind_nf_lc; + mind_nf_arity = type_app (Term.subst_mps sub) mbp.mind_nf_arity; + mind_user_lc = + array_smartmap (type_app (Term.subst_mps sub)) mbp.mind_user_lc; + mind_user_arity = type_app (Term.subst_mps sub) mbp.mind_user_arity; + mind_sort = mbp.mind_sort; + mind_nrealargs = mbp.mind_nrealargs; + mind_kelim = mbp.mind_kelim; + mind_nparams = mbp.mind_nparams; + mind_params_ctxt = + map_rel_context (Term.subst_mps sub) mbp.mind_params_ctxt; + mind_recargs = subst_wf_paths sub mbp.mind_recargs (*wf_paths*); +} + +let subst_mind sub mib = + { mind_finite = mib.mind_finite ; + mind_ntypes = mib.mind_ntypes ; + mind_hyps = (assert (mib.mind_hyps=[]); []) ; + mind_packets = array_smartmap (subst_mind_packet sub) mib.mind_packets ; + mind_constraints = mib.mind_constraints ; + mind_singl = option_app (Term.subst_mps sub) mib.mind_singl; +} + + +(*s Modules: signature component specifications, module types, and + module declarations *) + +type specification_body = + | SPBconst of constant_body + | SPBmind of mutual_inductive_body + | SPBmodule of module_specification_body + | SPBmodtype of module_type_body + +and module_signature_body = (label * specification_body) list + +and module_type_body = + | MTBident of kernel_name + | MTBfunsig of mod_bound_id * module_type_body * module_type_body + | MTBsig of mod_self_id * module_signature_body + +and module_expr_body = + | MEBident of module_path + | MEBfunctor of mod_bound_id * module_type_body * module_expr_body + | MEBstruct of mod_self_id * module_structure_body + | MEBapply of module_expr_body * module_expr_body + * constraints + +and module_specification_body = module_type_body * module_path option + +and structure_elem_body = + | SEBconst of constant_body + | SEBmind of mutual_inductive_body + | SEBmodule of module_body + | SEBmodtype of module_type_body + +and module_structure_body = (label * structure_elem_body) list + +and module_body = + { mod_expr : module_expr_body option; + mod_user_type : (module_type_body * constraints) option; + mod_type : module_type_body; + mod_equiv : module_path option } diff --git a/kernel/declarations.mli b/kernel/declarations.mli index 8a03e1cda..9ea7e20f8 100644 --- a/kernel/declarations.mli +++ b/kernel/declarations.mli @@ -15,10 +15,11 @@ open Term open Sign (*i*) -(* This module defines the types of global declarations. This includes - global constants/axioms and mutual inductive definitions *) +(* This module defines the internal representation of global + declarations. This includes global constants/axioms, mutual + inductive definitions, modules and module types *) -(*s Constants (internal representation) (Definition/Axiom) *) +(*s Constants (Definition/Axiom) *) type constant_body = { const_hyps : section_context; (* New: younger hyp at top *) @@ -27,6 +28,8 @@ type constant_body = { const_constraints : constraints; const_opaque : bool } +val subst_const_body : substitution -> constant_body -> constant_body + (*s Inductive types (internal representation with redundant information). *) @@ -35,6 +38,8 @@ type recarg = | Mrec of int | Imbr of inductive +val subst_recarg : substitution -> recarg -> recarg + type wf_paths = recarg Rtree.t val mk_norec : wf_paths @@ -42,6 +47,8 @@ val mk_paths : recarg -> wf_paths list array -> wf_paths val dest_recarg : wf_paths -> recarg val dest_subterms : wf_paths -> wf_paths list array +val subst_wf_paths : substitution -> wf_paths -> wf_paths + (* [mind_typename] is the name of the inductive; [mind_arity] is the arity generalized over global parameters; [mind_lc] is the list of types of constructors generalized over global parameters and @@ -70,3 +77,51 @@ type mutual_inductive_body = { mind_packets : one_inductive_body array; mind_constraints : constraints; mind_singl : constr option } + + +val subst_mind : substitution -> mutual_inductive_body -> mutual_inductive_body + + +(*s Modules: signature component specifications, module types, and + module declarations *) + +type specification_body = + | SPBconst of constant_body + | SPBmind of mutual_inductive_body + | SPBmodule of module_specification_body + | SPBmodtype of module_type_body + +and module_signature_body = (label * specification_body) list + +and module_type_body = + | MTBident of kernel_name + | MTBfunsig of mod_bound_id * module_type_body * module_type_body + | MTBsig of mod_self_id * module_signature_body + +and module_expr_body = + | MEBident of module_path + | MEBfunctor of mod_bound_id * module_type_body * module_expr_body + | MEBstruct of mod_self_id * module_structure_body + | MEBapply of module_expr_body * module_expr_body + * constraints + +and module_specification_body = module_type_body * module_path option + +and structure_elem_body = + | SEBconst of constant_body + | SEBmind of mutual_inductive_body + | SEBmodule of module_body + | SEBmodtype of module_type_body + +and module_structure_body = (label * structure_elem_body) list + +and module_body = + { mod_expr : module_expr_body option; + mod_user_type : (module_type_body * constraints) option; + mod_type : module_type_body; + mod_equiv : module_path option } + + + + + diff --git a/kernel/entries.ml b/kernel/entries.ml new file mode 100644 index 000000000..4d86d6d2c --- /dev/null +++ b/kernel/entries.ml @@ -0,0 +1,98 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* body @@ -156,16 +154,14 @@ let evaluable_constant cst env = with Not_found | NotEvaluableConst _ -> false (* Mutual Inductives *) -let lookup_mind sp env = - Spmap.find sp env.env_globals.env_inductives +let lookup_mind kn env = + KNmap.find kn env.env_globals.env_inductives -let add_mind sp mib env = - let new_inds = Spmap.add sp mib env.env_globals.env_inductives in - let new_locals = (Inductive,sp)::env.env_globals.env_locals in +let add_mind kn mib env = + let new_inds = KNmap.add kn mib env.env_globals.env_inductives in let new_globals = { env.env_globals with - env_inductives = new_inds; - env_locals = new_locals } in + env_inductives = new_inds } in { env with env_globals = new_globals } (* Universe constraints *) @@ -183,8 +179,8 @@ let lookup_constant_variables c env = let cmap = lookup_constant c env in Sign.instance_from_named_context cmap.const_hyps -let lookup_inductive_variables (sp,i) env = - let mis = lookup_mind sp env in +let lookup_inductive_variables (kn,i) env = + let mis = lookup_mind kn env in Sign.instance_from_named_context mis.mind_hyps let lookup_constructor_variables (ind,_) env = @@ -195,9 +191,9 @@ let lookup_constructor_variables (ind,_) env = let vars_of_global env constr = match kind_of_term constr with Var id -> [id] - | Const sp -> + | Const kn -> List.map destVar - (Array.to_list (lookup_constant_variables sp env)) + (Array.to_list (lookup_constant_variables kn env)) | Ind ind -> List.map destVar (Array.to_list (lookup_inductive_variables ind env)) @@ -241,67 +237,27 @@ let keep_hyps env needed = ~init:empty_named_context -(* Constants *) - -(*s Modules (i.e. compiled environments). *) - -type compiled_env = { - cenv_stamped_id : compilation_unit_name; - cenv_needed : compilation_unit_name list; - cenv_constants : (section_path * constant_body) list; - cenv_inductives : (section_path * mutual_inductive_body) list } - -let exported_objects env = - let gl = env.env_globals in - let separate (cst,ind) = function - | (Constant,sp) -> (sp,Spmap.find sp gl.env_constants)::cst,ind - | (Inductive,sp) -> cst,(sp,Spmap.find sp gl.env_inductives)::ind - in - List.fold_left separate ([],[]) gl.env_locals - -let export env id = - let (cst,ind) = exported_objects env in - { cenv_stamped_id = (id,0); - cenv_needed = env.env_globals.env_imports; - cenv_constants = cst; - cenv_inductives = ind } - -let check_imports env needed = - let imports = env.env_globals.env_imports in - let check (id,stamp) = - try - let actual_stamp = List.assoc id imports in - if stamp <> actual_stamp then - error ("Inconsistent assumptions over module " ^(string_of_dirpath id)) - with Not_found -> - error ("Reference to unknown module " ^ (string_of_dirpath id)) - in - List.iter check needed - -let import_constraints g sp cst = - try - merge_constraints cst g - with UniverseInconsistency -> - error "import_constraints: Universe Inconsistency during import" - -let import cenv env = - check_imports env cenv.cenv_needed; - let add_list t = List.fold_left (fun t (sp,x) -> Spmap.add sp x t) t in - let gl = env.env_globals in +(* Modules *) + +let add_modtype ln mtb env = + let new_modtypes = KNmap.add ln mtb env.env_globals.env_modtypes in + let new_globals = + { env.env_globals with + env_modtypes = new_modtypes } in + { env with env_globals = new_globals } + +let shallow_add_module mp mb env = + let new_mods = MPmap.add mp mb env.env_globals.env_modules in let new_globals = - { env_constants = add_list gl.env_constants cenv.cenv_constants; - env_inductives = add_list gl.env_inductives cenv.cenv_inductives; - env_locals = gl.env_locals; - env_imports = cenv.cenv_stamped_id :: gl.env_imports } - in - let g = universes env in - let g = List.fold_left - (fun g (sp,cb) -> import_constraints g sp cb.const_constraints) - g cenv.cenv_constants in - let g = List.fold_left - (fun g (sp,mib) -> import_constraints g sp mib.mind_constraints) - g cenv.cenv_inductives in - { env with env_globals = new_globals; env_universes = g } + { env.env_globals with + env_modules = new_mods } in + { env with env_globals = new_globals } + +let lookup_module mp env = + MPmap.find mp env.env_globals.env_modules + +let lookup_modtype ln env = + KNmap.find ln env.env_globals.env_modtypes (*s Judgments. *) diff --git a/kernel/environ.mli b/kernel/environ.mli index 98f6380c0..2b2be6b6b 100644 --- a/kernel/environ.mli +++ b/kernel/environ.mli @@ -100,6 +100,16 @@ val add_mind : mutual_inductive -> mutual_inductive_body -> env -> env (* raises [Not_found] if the required path is not found *) val lookup_mind : mutual_inductive -> env -> mutual_inductive_body +(***********************************************************************) +(*s Modules *) +val add_modtype : kernel_name -> module_type_body -> env -> env + +(* [shallow_add_module] does not add module components *) +val shallow_add_module : module_path -> module_body -> env -> env + +val lookup_module : module_path -> env -> module_body +val lookup_modtype : kernel_name -> env -> module_type_body + (***********************************************************************) (*s Universe constraints *) val set_universes : Univ.universes -> env -> env @@ -115,14 +125,6 @@ val vars_of_global : env -> constr -> identifier list val keep_hyps : env -> Idset.t -> section_context -(***********************************************************************) -(*s Modules. *) - -type compiled_env - -val export : env -> dir_path -> compiled_env -val import : compiled_env -> env -> env - (***********************************************************************) (*s Unsafe judgments. We introduce here the pre-type of judgments, which is actually only a datatype to store a term with its type and the type of its diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml index 96c2eaf98..036bce60b 100644 --- a/kernel/indtypes.ml +++ b/kernel/indtypes.ml @@ -18,25 +18,13 @@ open Sign open Environ open Reduction open Typeops +open Entries (* [check_constructors_names id s cl] checks that all the constructors names appearing in [l] are not present in the set [s], and returns the new set of names. The name [id] is the name of the current inductive type, used when reporting the error. *) -(*s Declaration. *) - -type one_inductive_entry = { - mind_entry_params : (identifier * local_entry) list; - mind_entry_typename : identifier; - mind_entry_arity : constr; - mind_entry_consnames : identifier list; - mind_entry_lc : constr list } - -type mutual_inductive_entry = { - mind_entry_finite : bool; - mind_entry_inds : one_inductive_entry list } - (***********************************************************************) (* Various well-formedness check for inductive declarations *) @@ -48,6 +36,7 @@ type inductive_error = | NonPar of env * constr * int * constr * constr | SameNamesTypes of identifier | SameNamesConstructors of identifier * identifier + | SameNamesOverlap of identifier list | NotAnArity of identifier | BadEntry (* These are errors related to recursors building in Indrec *) @@ -356,11 +345,11 @@ let check_positivity_one (env, _,ntypes,_ as ienv) hyps i indlc = if not (List.for_all (noccur_between n ntypes) largs) then raise (IllFormedInd (LocalNonPos n))); rarg - | Ind ind_sp -> + | Ind ind_kn -> (* If the inductive type being defined appears in a parameter, then we have an imbricated type *) if List.for_all (noccur_between n ntypes) largs then mk_norec - else check_positive_imbr ienv (ind_sp, largs) + else check_positive_imbr ienv (ind_kn, largs) | err -> if noccur_between n ntypes x && List.for_all (noccur_between n ntypes) largs @@ -548,3 +537,4 @@ let check_inductive env mie = let recargs = check_positivity env_arities inds in (* Build the inductive packets *) build_inductive env env_arities mie.mind_entry_finite inds recargs cst + diff --git a/kernel/indtypes.mli b/kernel/indtypes.mli index 532471ebc..4e0f9012c 100644 --- a/kernel/indtypes.mli +++ b/kernel/indtypes.mli @@ -14,6 +14,7 @@ open Univ open Term open Declarations open Environ +open Entries open Typeops (*i*) @@ -29,6 +30,7 @@ type inductive_error = | NonPar of env * constr * int * constr * constr | SameNamesTypes of identifier | SameNamesConstructors of identifier * identifier + | SameNamesOverlap of identifier list | NotAnArity of identifier | BadEntry (* These are errors related to recursors building in Indrec *) @@ -38,30 +40,6 @@ type inductive_error = exception InductiveError of inductive_error -(*s Declaration of inductive types. *) - -(* Assume the following definition in concrete syntax: -\begin{verbatim} -Inductive I1 [x1:X1;...;xn:Xn] : A1 := c11 : T11 | ... | c1n1 : T1n1 -... -with Ip [x1:X1;...;xn:Xn] : Ap := cp1 : Tp1 | ... | cpnp : Tpnp. -\end{verbatim} -then, in $i^{th}$ block, [mind_entry_params] is [[xn:Xn;...;x1:X1]]; -[mind_entry_arity] is [Ai], defined in context [[[x1:X1;...;xn:Xn]]; -[mind_entry_lc] is [Ti1;...;Tini], defined in context [[A'1;...;A'p;x1:X1;...;xn:Xn]] where [A'i] is [Ai] generalized over [[x1:X1;...;xn:Xn]]. -*) - -type one_inductive_entry = { - mind_entry_params : (identifier * local_entry) list; - mind_entry_typename : identifier; - mind_entry_arity : constr; - mind_entry_consnames : identifier list; - mind_entry_lc : constr list } - -type mutual_inductive_entry = { - mind_entry_finite : bool; - mind_entry_inds : one_inductive_entry list } - (*s The following function does checks on inductive declarations. *) val check_inductive : diff --git a/kernel/inductive.ml b/kernel/inductive.ml index 8fe4fa032..cc2c37790 100644 --- a/kernel/inductive.ml +++ b/kernel/inductive.ml @@ -19,8 +19,8 @@ open Reduction open Type_errors (* raise Not_found if not an inductive type *) -let lookup_mind_specif env (sp,tyi) = - let mib = Environ.lookup_mind sp env in +let lookup_mind_specif env (kn,tyi) = + let mib = Environ.lookup_mind kn env in if tyi >= Array.length mib.mind_packets then error "Inductive.lookup_mind_specif: invalid inductive index"; (mib, mib.mind_packets.(tyi)) @@ -109,10 +109,12 @@ let type_of_constructor env cstr = if i > nconstr then error "Not enough constructors in the type"; constructor_instantiate (fst ind) mib specif.(i-1) -let arities_of_constructors env ind = - let (mib,mip) = lookup_mind_specif env ind in +let arities_of_specif kn (mib,mip) = let specif = mip.mind_nf_lc in - Array.map (constructor_instantiate (fst ind) mib) specif + Array.map (constructor_instantiate kn mib) specif + +let arities_of_constructors env ind = + arities_of_specif (fst ind) (lookup_mind_specif env ind) @@ -337,8 +339,8 @@ type guard_env = genv : subterm_spec list; } -let make_renv env minds recarg (sp,tyi) = - let mib = Environ.lookup_mind sp env in +let make_renv env minds recarg (kn,tyi) = + let mib = Environ.lookup_mind kn env in let mind_recvec = Array.map (fun mip -> mip.mind_recargs) mib.mind_packets in { env = env; @@ -586,12 +588,12 @@ let check_one_fix renv recpos def = bodies in array_for_all (fun b -> b) ok_vect - | Const sp as c -> + | Const kn as c -> (try List.for_all (check_rec_call renv) l with (FixGuardError _ ) as e -> - if evaluable_constant sp renv.env then + if evaluable_constant kn renv.env then check_rec_call renv - (applist(constant_value renv.env sp, l)) + (applist(constant_value renv.env kn, l)) else raise e) (* The cases below simply check recursively the condition on the @@ -734,9 +736,9 @@ let check_one_cofix env nbfix def deftype = else error "check_one_cofix: ???" (* ??? *) - | Construct (_,i as cstr_sp) -> + | Construct (_,i as cstr_kn) -> let lra =vlra.(i-1) in - let mI = inductive_of_constructor cstr_sp in + let mI = inductive_of_constructor cstr_kn in let (mib,mip) = lookup_mind_specif env mI in let _,realargs = list_chop mip.mind_nparams args in let rec process_args_of_constr = function diff --git a/kernel/inductive.mli b/kernel/inductive.mli index 79900a84e..f279eb79a 100644 --- a/kernel/inductive.mli +++ b/kernel/inductive.mli @@ -43,6 +43,10 @@ val type_of_constructor : env -> constructor -> types (* Return constructor types in normal form *) val arities_of_constructors : env -> inductive -> types array +(* Transforms inductive specification into types (in nf) *) +val arities_of_specif : + kernel_name -> mutual_inductive_body * one_inductive_body -> types array + (* [type_case_branches env (I,args) (p:A) c] computes useful types about the following Cases expression:

Cases (c :: (I args)) of b1..bn end diff --git a/kernel/mod_typing.ml b/kernel/mod_typing.ml new file mode 100644 index 000000000..5c4b00d94 --- /dev/null +++ b/kernel/mod_typing.ml @@ -0,0 +1,206 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* mb + | _ -> raise Not_path + + +let rec list_fold_map2 f e = function + | [] -> (e,[],[]) + | h::t -> + let e',h1',h2' = f e h in + let e'',t1',t2' = list_fold_map2 f e' t in + e'',h1'::t1',h2'::t2' + + +let rec translate_modtype env mte = + match mte with + | MTEident ln -> MTBident ln + | MTEfunsig (arg_id,arg_e,body_e) -> + let arg_b = translate_modtype env arg_e in + let env' = + add_module (MPbound arg_id) (module_body_of_type arg_b) env in + let body_b = translate_modtype env' body_e in + MTBfunsig (arg_id,arg_b,body_b) + | MTEsig (msid,sig_e) -> + let str_b,sig_b = translate_entry_list env msid false sig_e in + MTBsig (msid,sig_b) + +and translate_entry_list env msid is_definition sig_e = + let mp = MPself msid in + let do_entry env (l,e) = + let kn = make_kn mp empty_dirpath l in + match e with + | SPEconst ce -> + let cb = translate_constant env ce in + add_constant kn cb env, (l, SEBconst cb), (l, SPBconst cb) + | SPEmind mie -> + let mib = translate_mind env mie in + add_mind kn mib env, (l, SEBmind mib), (l, SPBmind mib) + | SPEmodule me -> + let mb = translate_module env is_definition me in + let mspec = mb.mod_type, mb.mod_equiv in + let mp' = MPdot (mp,l) in + add_module mp' mb env, (l, SEBmodule mb), (l, SPBmodule mspec) + | SPEmodtype mte -> + let mtb = translate_modtype env mte in + add_modtype kn mtb env, (l, SEBmodtype mtb), (l, SPBmodtype mtb) + in + let _,str_b,sig_b = list_fold_map2 do_entry env sig_e + in + str_b,sig_b + +(* if [is_definition=true], [mod_entry_expr] may be any expression. + Otherwise it must be a path *) + +and translate_module env is_definition me = + match me.mod_entry_expr, me.mod_entry_type with + | None, None -> + anomaly "Mod_typing.translate_module: empty type and expr in module entry" + | None, Some mte -> + let mtb = translate_modtype env mte in + { mod_expr = None; + mod_user_type = Some (mtb, Constraint.empty); + mod_type = mtb; + mod_equiv = None } + | Some mexpr, _ -> + let meq_o = (* do we have a transparent module ? *) + try (* TODO: transparent field in module_entry *) + Some (path_of_mexpr mexpr) + with + | Not_path -> None + in + let meb,mtb1 = + if is_definition then + translate_mexpr env mexpr + else + let mp = + try + path_of_mexpr mexpr + with + | Not_path -> error_declaration_not_path mexpr + in + MEBident mp, (lookup_module mp env).mod_type + in + let mtb,mod_user_type = + match me.mod_entry_type with + | None -> mtb1, None + | Some mte -> + let mtb2 = translate_modtype env mte in + mtb2, Some (mtb2, check_subtypes env mtb1 mtb2) + in + { mod_type = mtb; + mod_user_type = mod_user_type; + mod_expr = Some meb; + mod_equiv = meq_o } + +(* translate_mexpr : env -> module_expr -> module_expr_body * module_type_body *) +and translate_mexpr env mexpr = match mexpr with + | MEident mp -> + MEBident mp, + (lookup_module mp env).mod_type + | MEfunctor (arg_id, arg_e, body_expr) -> + let arg_b = translate_modtype env arg_e in + let env' = add_module (MPbound arg_id) (module_body_of_type arg_b) env in + let (body_b,body_tb) = translate_mexpr env' body_expr in + MEBfunctor (arg_id, arg_b, body_b), + MTBfunsig (arg_id, arg_b, body_tb) + | MEapply (fexpr,mexpr) -> + let feb,ftb = translate_mexpr env fexpr in + let ftb = scrape_modtype env ftb in + let farg_id, farg_b, fbody_b = destr_functor ftb in + let meb,mtb = translate_mexpr env mexpr in + let cst = check_subtypes env mtb farg_b in + let mp = + try + path_of_mexpr mexpr + with + | Not_path -> error_application_to_not_path mexpr + (* place for nondep_supertype *) + in + MEBapply(feb,meb,cst), + subst_modtype (map_mbid farg_id mp) fbody_b + | MEstruct (msid,structure) -> + let structure,signature = translate_entry_list env msid true structure in + MEBstruct (msid,structure), + MTBsig (msid,signature) + + +(* is_definition is true - me.mod_entry_expr may be any expression *) +let translate_module env me = translate_module env true me + +let rec add_module_expr_constraints env = function + | MEBident _ -> env + + | MEBfunctor (_,mtb,meb) -> + add_module_expr_constraints (add_modtype_constraints env mtb) meb + + | MEBstruct (_,mod_struct_body) -> + List.fold_left + (fun env (l,item) -> add_struct_elem_constraints env item) + env + mod_struct_body + + | MEBapply (meb1,meb2,cst) -> + Environ.add_constraints cst + (add_module_expr_constraints + (add_module_expr_constraints env meb1) + meb2) + +and add_struct_elem_constraints env = function + | SEBconst cb -> Environ.add_constraints cb.const_constraints env + | SEBmind mib -> Environ.add_constraints mib.mind_constraints env + | SEBmodule mb -> add_module_constraints env mb + | SEBmodtype mtb -> add_modtype_constraints env mtb + +and add_module_constraints env mb = + let env = match mb.mod_expr with + | None -> env + | Some meb -> add_module_expr_constraints env meb + in + let env = match mb.mod_user_type with + | None -> env + | Some (mtb,cst) -> + Environ.add_constraints cst (add_modtype_constraints env mtb) + in + env + +and add_modtype_constraints env = function + | MTBident _ -> env + | MTBfunsig (_,mtb1,mtb2) -> + add_modtype_constraints + (add_modtype_constraints env mtb1) + mtb2 + | MTBsig (_,mod_sig_body) -> + List.fold_left + (fun env (l,item) -> add_sig_elem_constraints env item) + env + mod_sig_body + +and add_sig_elem_constraints env = function + | SPBconst cb -> Environ.add_constraints cb.const_constraints env + | SPBmind mib -> Environ.add_constraints mib.mind_constraints env + | SPBmodule (mtb,_) -> add_modtype_constraints env mtb + | SPBmodtype mtb -> add_modtype_constraints env mtb + + diff --git a/kernel/mod_typing.mli b/kernel/mod_typing.mli new file mode 100644 index 000000000..7d17d4823 --- /dev/null +++ b/kernel/mod_typing.mli @@ -0,0 +1,25 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* module_type_entry -> module_type_body + +val translate_module : env -> module_entry -> module_body + +val add_modtype_constraints : env -> module_type_body -> env + +val add_module_constraints : env -> module_body -> env + diff --git a/kernel/modops.ml b/kernel/modops.ml new file mode 100644 index 000000000..1b0ff8ace --- /dev/null +++ b/kernel/modops.ml @@ -0,0 +1,151 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* "^string_of_label l'^" !") + +let error_result_must_be_signature mtb = + error "The result module type must be a signature" + +let error_no_module_to_end _ = + error "No open module to end" + +let error_no_modtype_to_end _ = + error "No open module type to end" + +let error_not_a_modtype s = + error ("\""^s^"\" is not a module type") + +let error_not_a_module s = + error ("\""^s^"\" is not a module") + + +let rec scrape_modtype env = function + | MTBident ln -> scrape_modtype env (lookup_modtype ln env) + | mtb -> mtb + + +let module_body_of_spec spec = + { mod_type = fst spec; + mod_equiv = snd spec; + mod_expr = None; + mod_user_type = None} + +let module_body_of_type mtb = + { mod_type = mtb; + mod_equiv = None; + mod_expr = None; + mod_user_type = None} + +let module_spec_of_body mb = mb.mod_type, mb.mod_equiv + +let destr_functor = function + | MTBfunsig (arg_id,arg_t,body_t) -> (arg_id,arg_t,body_t) + | mtb -> error_not_a_functor mtb + + +let rec check_modpath_equiv env mp1 mp2 = + if mp1=mp2 then (); + let mb1 = lookup_module mp1 env in + match mb1.mod_equiv with + | None -> + let mb2 = lookup_module mp2 env in + (match mb2.mod_equiv with + | None -> error_not_equal mp1 mp2 + | Some mp2' -> check_modpath_equiv env mp2' mp1) + | Some mp1' -> check_modpath_equiv env mp2 mp1' + + +let rec subst_modtype sub = function + | MTBident ln -> MTBident (subst_kn sub ln) + | MTBfunsig (arg_id, arg_b, body_b) -> + assert (not (occur_mbid arg_id sub)); + MTBfunsig (arg_id, + subst_modtype sub arg_b, + subst_modtype sub body_b) + | MTBsig (sid1, msb) -> + assert (not (occur_msid sid1 sub)); + MTBsig (sid1, subst_signature sub msb) + +and subst_signature sub sign = + let subst_body = function + SPBconst cb -> + SPBconst (subst_const_body sub cb) + | SPBmind mib -> + SPBmind (subst_mind sub mib) + | SPBmodule mb -> + SPBmodule (subst_module sub mb) + | SPBmodtype mtb -> + SPBmodtype (subst_modtype sub mtb) + in + List.map (fun (l,b) -> (l,subst_body b)) sign + +and subst_module sub (mtb,mpo as mb) = + let mtb' = subst_modtype sub mtb in + let mpo' = option_smartmap (subst_mp sub) mpo in + if mtb'==mtb && mpo'==mpo then mb else + (mtb',mpo') + +let subst_signature_msid msid mp = + subst_signature (map_msid msid mp) + +(* we assume that the substitution of "mp" into "msid" is already done +(or unnecessary) *) +let rec add_signature mp sign env = + let add_one env (l,elem) = + let kn = make_kn mp empty_dirpath l in + match elem with + | SPBconst cb -> Environ.add_constant kn cb env + | SPBmind mib -> Environ.add_mind kn mib env + | SPBmodule mb -> + add_module (MPdot (mp,l)) (module_body_of_spec mb) env + (* adds components as well *) + | SPBmodtype mtb -> Environ.add_modtype kn mtb env + in + List.fold_left add_one env sign + + +and add_module mp mb env = + let env = Environ.shallow_add_module mp mb env in + match scrape_modtype env mb.mod_type with + | MTBident _ -> anomaly "scrape_modtype does not work!" + | MTBsig (msid,sign) -> + add_signature mp (subst_signature_msid msid mp sign) env + + | MTBfunsig _ -> env + + diff --git a/kernel/modops.mli b/kernel/modops.mli new file mode 100644 index 000000000..240144785 --- /dev/null +++ b/kernel/modops.mli @@ -0,0 +1,79 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* module_type_body -> module_type_body + +(* make the environment entry out of type *) +val module_body_of_type : module_type_body -> module_body + +val module_body_of_spec : module_specification_body -> module_body + +val module_spec_of_body : module_body -> module_specification_body + + +val destr_functor : + module_type_body -> mod_bound_id * module_type_body * module_type_body + + +val subst_modtype : substitution -> module_type_body -> module_type_body + +val subst_signature_msid : + mod_self_id -> module_path -> + module_signature_body -> module_signature_body + +(* [add_signature mp sign env] assumes that the substitution [msid] + \mapsto [mp] has already been performed (or is not necessary, like + when [mp = MPself msid]) *) +val add_signature : + module_path -> module_signature_body -> env -> env + +(* adds a module and its components, but not the constraints *) +val add_module : + module_path -> module_body -> env -> env + +val check_modpath_equiv : env -> module_path -> module_path -> unit + +val error_existing_label : label -> 'a + +val error_declaration_not_path : module_expr -> 'a + +val error_application_to_not_path : module_expr -> 'a + +val error_not_a_functor : module_expr -> 'a + +val error_incompatible_modtypes : + module_type_body -> module_type_body -> 'a + +val error_not_match : label -> specification_body -> 'a + +val error_incompatible_labels : label -> label -> 'a + +val error_no_such_label : label -> 'a + +val error_result_must_be_signature : module_type_body -> 'a + +val error_no_module_to_end : unit -> 'a + +val error_no_modtype_to_end : unit -> 'a + +val error_not_a_modtype : string -> 'a + +val error_not_a_module : string -> 'a diff --git a/kernel/names.ml b/kernel/names.ml index a92adc607..0209d1c33 100644 --- a/kernel/names.ml +++ b/kernel/names.ml @@ -61,55 +61,203 @@ module ModIdmap = Map.Make(ModIdOrdered) let make_dirpath x = x let repr_dirpath x = x +let empty_dirpath = [] + let string_of_dirpath = function | [] -> "" | sl -> String.concat "." (List.map string_of_id (List.rev sl)) -(*s Section paths are absolute names *) -type section_path = { - dirpath : dir_path ; - basename : identifier } +let u_number = ref 0 +type uniq_ident = int * string * dir_path +let make_uid dir s = incr u_number;(!u_number,s,dir) +let string_of_uid (i,s,p) = + "<"(*^string_of_dirpath p ^"#"^*) ^ s ^"#"^ string_of_int i^">" -let make_path pa id = { dirpath = pa; basename = id } -let repr_path { dirpath = pa; basename = id } = (pa,id) +module Umap = Map.Make(struct + type t = uniq_ident + let compare = Pervasives.compare + end) -(* parsing and printing of section paths *) -let string_of_path sp = - let (sl,id) = repr_path sp in - if sl = [] then string_of_id id - else (string_of_dirpath sl) ^ "." ^ (string_of_id id) -let sp_ord sp1 sp2 = - let (p1,id1) = repr_path sp1 - and (p2,id2) = repr_path sp2 in - let p_bit = compare p1 p2 in - if p_bit = 0 then id_ord id1 id2 else p_bit +type mod_self_id = uniq_ident +let make_msid = make_uid +let string_of_msid = string_of_uid + +type mod_bound_id = uniq_ident +let make_mbid = make_uid +let string_of_mbid = string_of_uid + +type label = string +let mk_label l = l +let string_of_label l = l + +let id_of_label l = l +let label_of_id id = id + +module Labset = Idset +module Labmap = Idmap + +type module_path = + | MPfile of dir_path + | MPbound of mod_bound_id + | MPself of mod_self_id + | MPdot of module_path * label + +let rec string_of_mp = function + | MPfile sl -> string_of_dirpath sl + | MPbound uid -> string_of_uid uid + | MPself uid -> string_of_uid uid + | MPdot (mp,l) -> string_of_mp mp ^ "." ^ string_of_label l + +(* we compare labels first if two MPdot's *) +let rec mp_ord mp1 mp2 = match (mp1,mp2) with + MPdot(mp1,l1), MPdot(mp2,l2) -> + let c = Pervasives.compare l1 l2 in + if c<>0 then + c + else + mp_ord mp1 mp2 + | _,_ -> Pervasives.compare mp1 mp2 + +module MPord = struct + type t = module_path + let compare = mp_ord +end + +module MPmap = Map.Make(MPord) + + +(* this is correct under the condition that bound and struct + identifiers can never be identical (i.e. get the same stamp)! *) + +type substitution = module_path Umap.t + +let empty_subst = Umap.empty + +let add_msid = Umap.add +let add_mbid = Umap.add + +let map_msid msid mp = add_msid msid mp empty_subst +let map_mbid mbid mp = add_msid mbid mp empty_subst + +let join subst = + Umap.fold Umap.add subst + +let list_contents sub = + let one_pair uid mp l = + (string_of_uid uid, string_of_mp mp)::l + in + Umap.fold one_pair sub [] + +let debug_string_of_subst sub = + let l = List.map (fun (s1,s2) -> s1^"|->"^s2) (list_contents sub) in + "{" ^ String.concat "; " l ^ "}" + +let debug_pr_subst sub = + let l = list_contents sub in + let f (s1,s2) = hov 2 (str s1 ++ spc () ++ str "|-> " ++ str s2) + in + str "{" ++ hov 2 (prlist_with_sep pr_coma f l) ++ str "}" + +let rec subst_mp sub mp = (* 's like subst *) + match mp with + | MPself sid -> + (try Umap.find sid sub with Not_found -> mp) + | MPbound bid -> + (try Umap.find bid sub with Not_found -> mp) + | MPdot (mp1,l) -> + let mp1' = subst_mp sub mp1 in + if mp1==mp1' then + mp + else + MPdot (mp1',l) + | _ -> mp + + +let rec occur_in_path uid = function + | MPself sid -> sid = uid + | MPbound bid -> bid = uid + | MPdot (mp1,_) -> occur_in_path uid mp1 + | _ -> false + +let occur_uid uid sub = + let check_one uid' mp = + if uid = uid' || occur_in_path uid mp then raise Exit + in + try + Umap.iter check_one sub; + false + with Exit -> true + +let occur_msid = occur_uid +let occur_mbid = occur_uid + + + +(* Kernel names *) + +type kernel_name = module_path * dir_path * label + +let make_kn mp dir l = (mp,dir,l) +let repr_kn kn = kn + +let modpath kn = + let mp,_,_ = repr_kn kn in mp + +let label kn = + let _,_,l = repr_kn kn in l + +let string_of_kn (mp,dir,l) = + string_of_mp mp ^ "#" ^ string_of_dirpath dir ^ "#" ^ string_of_label l + +let pr_kn kn = str (string_of_kn kn) + + +let subst_kn sub (mp,dir,l as kn) = + let mp' = subst_mp sub mp in + if mp==mp' then kn else (mp',dir,l) + + +let kn_ord kn1 kn2 = + let mp1,dir1,l1 = kn1 in + let mp2,dir2,l2 = kn2 in + let c = Pervasives.compare l1 l2 in + if c <> 0 then + c + else + let c = Pervasives.compare dir1 dir2 in + if c<>0 then + c + else + MPord.compare mp1 mp2 + + +module KNord = struct + type t = kernel_name + let compare =kn_ord +end + +module KNmap = Map.Make(KNord) +module KNpred = Predicate.Make(KNord) +module KNset = Set.Make(KNord) -module SpOrdered = - struct - type t = section_path - let compare = sp_ord - end -module Spset = Set.Make(SpOrdered) -module Sppred = Predicate.Make(SpOrdered) -module Spmap = Map.Make(SpOrdered) +let initial_msid = (make_msid [] "TOP") +let initial_path = MPself initial_msid -(*s********************************************************************) -(* type of global reference *) type variable = identifier -type constant = section_path -type inductive = section_path * int +type constant = kernel_name +type mutual_inductive = kernel_name +type inductive = mutual_inductive * int type constructor = inductive * int -type mutual_inductive = section_path -let ith_mutual_inductive (sp,_) i = (sp,i) -let ith_constructor_of_inductive ind_sp i = (ind_sp,i) -let inductive_of_constructor (ind_sp,i) = ind_sp -let index_of_constructor (ind_sp,i) = i +let ith_mutual_inductive (kn,_) i = (kn,i) +let ith_constructor_of_inductive ind i = (ind,i) +let inductive_of_constructor (ind,i) = ind +let index_of_constructor (ind,i) = i (* Hash-consing of name objects *) module Hname = Hashcons.Make( @@ -136,16 +284,12 @@ module Hdir = Hashcons.Make( let hash = Hashtbl.hash end) -module Hsp = Hashcons.Make( +module Hkn = Hashcons.Make( struct - type t = section_path + type t = kernel_name type u = identifier -> identifier - let hash_sub hident sp = - { dirpath = List.map hident sp.dirpath; - basename = hident sp.basename } - let equal sp1 sp2 = - (List.length sp1.dirpath = List.length sp2.dirpath) && - (List.for_all2 (==) sp1.dirpath sp2.dirpath) + let hash_sub hident kn = Util.todo "hash_cons"; kn + let equal kn1 kn2 = kn1==kn2 let hash = Hashtbl.hash end) @@ -154,5 +298,5 @@ let hcons_names () = let hident = Hashcons.simple_hcons Hident.f hstring in let hname = Hashcons.simple_hcons Hname.f hident in let hdir = Hashcons.simple_hcons Hdir.f hident in - let hspcci = Hashcons.simple_hcons Hsp.f hident in - (hspcci,hdir,hname,hident,hstring) + let hkn = Hashcons.simple_hcons Hkn.f hident in + (hkn,hdir,hname,hident,hstring) diff --git a/kernel/names.mli b/kernel/names.mli index 08d2db357..b1bcc0b83 100644 --- a/kernel/names.mli +++ b/kernel/names.mli @@ -16,6 +16,8 @@ type name = Name of identifier | Anonymous val string_of_id : identifier -> string val id_of_string : string -> identifier +val id_ord : identifier -> identifier -> int + (* Identifiers sets and maps *) module Idset : Set.S with type elt = identifier module Idpred : Predicate.S with type elt = identifier @@ -32,44 +34,129 @@ type dir_path val make_dirpath : module_ident list -> dir_path val repr_dirpath : dir_path -> module_ident list +val empty_dirpath : dir_path + (* Printing of directory paths as ["coq_root.module.submodule"] *) val string_of_dirpath : dir_path -> string -(*s Section paths are {\em absolute} names *) -type section_path +(*s Unique identifier to be used as "self" in structures and + signatures - invisible for users *) + +type mod_self_id + +(* The first argument is a file name - to prevent conflict between + different files *) +val make_msid : dir_path -> string -> mod_self_id +val string_of_msid : mod_self_id -> string + +(*s Unique names for bound modules *) +type mod_bound_id + +val make_mbid : dir_path -> string -> mod_bound_id +val string_of_mbid : mod_bound_id -> string + +(*s Names of structure elements *) +type label +val mk_label : string -> label +val string_of_label : label -> string + +val label_of_id : identifier -> label +val id_of_label : label -> identifier + +module Labset : Set.S with type elt = label +module Labmap : Map.S with type key = label + +(*s The module part of the kernel name *) +type module_path = + | MPfile of dir_path + | MPbound of mod_bound_id + | MPself of mod_self_id + | MPdot of module_path * label +(*i | MPapply of module_path * module_path in the future (maybe) i*) + + +val string_of_mp : module_path -> string + +module MPmap : Map.S with type key = module_path + + +(*s Substitutions *) + +type substitution -(* Constructors of [section_path] *) -val make_path : dir_path -> identifier -> section_path +val empty_subst : substitution -(* Destructors of [section_path] *) -val repr_path : section_path -> dir_path * identifier +val add_msid : + mod_self_id -> module_path -> substitution -> substitution +val add_mbid : + mod_bound_id -> module_path -> substitution -> substitution -(* Parsing and printing of section path as ["coq_root.module.id"] *) -val string_of_path : section_path -> string +val map_msid : mod_self_id -> module_path -> substitution +val map_mbid : mod_bound_id -> module_path -> substitution -module Spset : Set.S with type elt = section_path -module Sppred : Predicate.S with type elt = section_path -module Spmap : Map.S with type key = section_path +val join : substitution -> substitution -> substitution -(*s********************************************************************) -(* type of global reference *) +(*i debugging *) +val debug_string_of_subst : substitution -> string +val debug_pr_subst : substitution -> Pp.std_ppcmds +(*i*) + +(* [subst_mp sub mp] guarantees that whenever the result of the + substitution is structutally equal [mp], it is equal by pointers + as well [==] *) + +val subst_mp : + substitution -> module_path -> module_path + +(* [occur_*id id sub] returns true iff [id] occurs in [sub] + on either side *) + +val occur_msid : mod_self_id -> substitution -> bool +val occur_mbid : mod_bound_id -> substitution -> bool + + +(* Name of the toplevel structure *) +val initial_msid : mod_self_id +val initial_path : module_path (* [= MPself initial_msid] *) + +(*s The absolute names of objects seen by kernel *) + +type kernel_name + +(* Constructor and destructor *) +val make_kn : module_path -> dir_path -> label -> kernel_name +val repr_kn : kernel_name -> module_path * dir_path * label + +val modpath : kernel_name -> module_path +val label : kernel_name -> label + +val string_of_kn : kernel_name -> string +val pr_kn : kernel_name -> Pp.std_ppcmds +val subst_kn : substitution -> kernel_name -> kernel_name + + +module KNset : Set.S with type elt = kernel_name +module KNpred : Predicate.S with type elt = kernel_name +module KNmap : Map.S with type key = kernel_name + + +(*s Specific paths for declarations *) type variable = identifier -type constant = section_path +type constant = kernel_name +type mutual_inductive = kernel_name (* Beware: first inductive has index 0 *) -type inductive = section_path * int +type inductive = mutual_inductive * int (* Beware: first constructor has index 1 *) type constructor = inductive * int -type mutual_inductive = section_path val ith_mutual_inductive : inductive -> int -> inductive - val ith_constructor_of_inductive : inductive -> int -> constructor val inductive_of_constructor : constructor -> inductive val index_of_constructor : constructor -> int (* Hash-consing *) val hcons_names : unit -> - (section_path -> section_path) * (dir_path -> dir_path) * + (kernel_name -> kernel_name) * (dir_path -> dir_path) * (name -> name) * (identifier -> identifier) * (string -> string) diff --git a/kernel/reduction.ml b/kernel/reduction.ml index a5b773c24..4512546ae 100644 --- a/kernel/reduction.ml +++ b/kernel/reduction.ml @@ -143,6 +143,11 @@ let sort_cmp pb s0 s1 cuniv = | (_, _) -> raise NotConvertible +let conv_sort env s0 s1 = sort_cmp CONV s0 s1 Constraint.empty + +let conv_sort_leq env s0 s1 = sort_cmp CUMUL s0 s1 Constraint.empty + + (* Conversion between [lft1]term1 and [lft2]term2 *) let rec ccnv cv_pb infos lft1 lft2 term1 term2 cuniv = eqappr cv_pb infos diff --git a/kernel/reduction.mli b/kernel/reduction.mli index 50371e85f..63949ceb2 100644 --- a/kernel/reduction.mli +++ b/kernel/reduction.mli @@ -29,6 +29,9 @@ exception NotConvertible exception NotConvertibleVect of int type 'a conversion_function = env -> 'a -> 'a -> Univ.constraints +val conv_sort : sorts conversion_function +val conv_sort_leq : sorts conversion_function + val conv : types conversion_function val conv_leq : types conversion_function val conv_leq_vecti : types array conversion_function diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml index 74e33cdfb..909b12704 100644 --- a/kernel/safe_typing.ml +++ b/kernel/safe_typing.ml @@ -17,43 +17,75 @@ open Sign open Declarations open Inductive open Environ +open Entries +open Typeops open Type_errors open Indtypes +open Term_typing +open Modops +open Subtyping +open Mod_typing -type judgment = unsafe_judgment -let j_val = j_val -let j_type = j_type +type modvariant = + | NONE + | SIG of (* funsig params *) (mod_bound_id * module_type_body) list + | STRUCT of (* functor params *) (mod_bound_id * module_type_body) list + * (* optional result type *) module_type_body option + | LIBRARY of dir_path + +type module_info = + { msid : mod_self_id; + modpath : module_path; + label : label; + variant : modvariant} -(* Exported machines. *) +let check_label l labset = + if Labset.mem l labset then error_existing_label l -let safe_infer = Typeops.infer -let safe_infer_type = Typeops.infer_type +type library_info = dir_path * Digest.t -(*s Safe environments. *) +type safe_environment = + { old : safe_environment; + env : env; + modinfo : module_info; + labset : Labset.t; + revsign : module_signature_body; + revstruct : module_structure_body; + imports : library_info list; + loads : (module_path * module_body) list } -type safe_environment = env +(* + { old = senv.old; + env = ; + modinfo = senv.modinfo; + labset = ; + revsign = ; + imports = senv.imports ; + loads = senv.loads } +*) -let empty_environment = empty_env -(* Insertion of variables (named and de Bruijn'ed). They are now typed before - being added to the environment. *) +(* a small hack to avoid variants and an unused case in all functions *) +let rec empty_environment = + { old = empty_environment; + env = empty_env; + modinfo = { + msid = initial_msid; + modpath = initial_path; + label = mk_label "_"; + variant = NONE}; + labset = Labset.empty; + revsign = []; + revstruct = []; + imports = []; + loads = [] } -let constrain_type env j cst1 = function - | None -> j.uj_type, cst1 - | Some t -> - let (tj,cst2) = safe_infer_type env t in - let cst3 = - try conv_leq env j.uj_type tj.utj_val - with NotConvertible -> error_actual_type env j tj.utj_val in - tj.utj_val, Constraint.union (Constraint.union cst1 cst2) cst3 +let env_of_safe_env senv = senv.env +let env_of_senv = env_of_safe_env -let push_rel_or_named_def push (id,b,topt) env = - let (j,cst) = safe_infer env b in - let (typ,cst) = constrain_type env j cst topt in - let env' = add_constraints cst env in - let env'' = push (id,Some j.uj_val,typ) env' in - (cst,env'') +(* Insertion of section variables. They are now typed before being + added to the environment. *) (* Same as push_named, but check that the variable is not already there. Should *not* be done in Environ because tactics add temporary @@ -65,90 +97,401 @@ let safe_push_named (id,_,_ as d) env = let _ = lookup_named id env in error ("identifier "^string_of_id id^" already defined") with Not_found -> () in - push_named d env + Environ.push_named d env -let push_named_def = push_rel_or_named_def safe_push_named -let push_rel_def = push_rel_or_named_def push_rel +let push_named_def (id,b,topt) senv = + let (c,typ,cst) = translate_local_def senv.env (b,topt) in + let env' = add_constraints cst senv.env in + let env'' = safe_push_named (id,Some c,typ) env' in + (cst, {senv with env=env''}) -let push_rel_or_named_assum push (id,t) env = - let (j,cst) = safe_infer env t in - let t = Typeops.assumption_of_judgment env j in - let env' = add_constraints cst env in - let env'' = push (id,None,t) env' in - (cst,env'') +let push_named_assum (id,t) senv = + let (t,cst) = translate_local_assum senv.env t in + let env' = add_constraints cst senv.env in + let env'' = safe_push_named (id,None,t) env' in + (cst, {senv with env=env''}) -let push_named_assum = push_rel_or_named_assum push_named -let push_rel_assum d env = snd (push_rel_or_named_assum push_rel d env) - -let push_rels_with_univ vars env = - List.fold_left (fun env nvar -> push_rel_assum nvar env) env vars (* Insertion of constants and parameters in environment. *) -type constant_entry = { - const_entry_body : constr; - const_entry_type : types option; - const_entry_opaque : bool } type global_declaration = | ConstantEntry of constant_entry - | ParameterEntry of types | GlobalRecipe of Cooking.recipe -(* Definition always declared transparent *) -let safe_infer_declaration env dcl = - match dcl with - | ConstantEntry c -> - let (j,cst) = safe_infer env c.const_entry_body in - let (typ,cst) = constrain_type env j cst c.const_entry_type in - Some j.uj_val, typ, cst, c.const_entry_opaque - | ParameterEntry t -> - let (j,cst) = safe_infer env t in - None, Typeops.assumption_of_judgment env j, cst, false - | GlobalRecipe r -> - Cooking.cook_constant env r - -let add_global_declaration sp env (body,typ,cst,op) = - let env' = add_constraints cst env in - let ids = match body with - | None -> global_vars_set env typ - | Some b -> - Idset.union (global_vars_set env b) (global_vars_set env typ) in - let hyps = keep_hyps env ids in - let cb = { - const_body = body; - const_type = typ; - const_hyps = hyps; - const_constraints = cst; - const_opaque = op } in - Environ.add_constant sp cb env' - -(*s Global and local constant declaration. *) - -let add_constant sp ce env = - let _ = - try - let _ = lookup_constant sp env in - error ("constant "^string_of_path sp^" already declared") - with Not_found -> () in - add_global_declaration sp env (safe_infer_declaration env ce) +let add_constant dir l decl senv = + check_label l senv.labset; + let cb = match decl with + ConstantEntry ce -> translate_constant senv.env ce + | GlobalRecipe r -> translate_recipe senv.env r + in + let env' = Environ.add_constraints cb.const_constraints senv.env in + let kn = make_kn senv.modinfo.modpath dir l in + let env'' = Environ.add_constant kn cb env' in + kn, { old = senv.old; + env = env''; + modinfo = senv.modinfo; + labset = Labset.add l senv.labset; + revsign = (l,SPBconst cb)::senv.revsign; + revstruct = (l,SEBconst cb)::senv.revstruct; + imports = senv.imports; + loads = senv.loads } + (* Insertion of inductive types. *) -let add_mind sp mie env = - let _ = - try - let _ = lookup_mind sp env in - error ("inductive "^string_of_path sp^" already declared") - with Not_found -> () in - let mib = check_inductive env mie in - let cst = mib.mind_constraints in - Environ.add_mind sp mib (add_constraints cst env) +let add_mind dir l mie senv = + if mie.mind_entry_inds = [] then + anomaly "empty inductive types declaration"; + (* this test is repeated by translate_mind *) + let id = (List.nth mie.mind_entry_inds 0).mind_entry_typename in + if l <> label_of_id id then + anomaly ("the label of inductive packet and its first inductive"^ + " type do not match"); + check_label l senv.labset; + (* TODO: when we will allow reorderings we will have to verify + all labels *) + let mib = translate_mind senv.env mie in + let env' = Environ.add_constraints mib.mind_constraints senv.env in + let kn = make_kn senv.modinfo.modpath dir l in + let env'' = Environ.add_mind kn mib env' in + kn, { old = senv.old; + env = env''; + modinfo = senv.modinfo; + labset = Labset.add l senv.labset; (* TODO: the same as above *) + revsign = (l,SPBmind mib)::senv.revsign; + revstruct = (l,SEBmind mib)::senv.revstruct; + imports = senv.imports; + loads = senv.loads } + + +(* Insertion of module types *) + +let add_modtype l mte senv = + check_label l senv.labset; + let mtb = translate_modtype senv.env mte in + let env' = add_modtype_constraints senv.env mtb in + let kn = make_kn senv.modinfo.modpath empty_dirpath l in + let env'' = Environ.add_modtype kn mtb env' in + kn, { old = senv.old; + env = env''; + modinfo = senv.modinfo; + labset = Labset.add l senv.labset; + revsign = (l,SPBmodtype mtb)::senv.revsign; + revstruct = (l,SEBmodtype mtb)::senv.revstruct; + imports = senv.imports; + loads = senv.loads } + + + +(* full_add_module adds module with universes and constraints *) +let full_add_module mp mb env = + let env = add_module_constraints env mb in + let env = Modops.add_module mp mb env in + env + +(* Insertion of modules *) + +let add_module l me senv = + check_label l senv.labset; + let mb = translate_module senv.env me in + let mspec = module_spec_of_body mb in + let mp = MPdot(senv.modinfo.modpath, l) in + let env' = full_add_module mp mb senv.env in + mp, { old = senv.old; + env = env'; + modinfo = senv.modinfo; + labset = Labset.add l senv.labset; + revsign = (l,SPBmodule mspec)::senv.revsign; + revstruct = (l,SEBmodule mb)::senv.revstruct; + imports = senv.imports; + loads = senv.loads } + + +(* Interactive modules *) + +let start_module dir l params result senv = + check_label l senv.labset; + let rec trans_params env = function + | [] -> env,[] + | (mbid,mte)::rest -> + let mtb = translate_modtype env mte in + let env = + full_add_module (MPbound mbid) (module_body_of_type mtb) env + in + let env,transrest = trans_params env rest in + env, (mbid,mtb)::transrest + in + let env,params_body = trans_params senv.env params in + let check_sig mtb = match scrape_modtype env mtb with + | MTBsig _ -> () + | MTBfunsig _ -> error_result_must_be_signature mtb + | _ -> anomaly "start_module: modtype not scraped" + in + let result_body = option_app (translate_modtype env) result in + ignore (option_app check_sig result_body); + let msid = make_msid dir (string_of_label l) in + let mp = MPself msid in + let modinfo = { msid = msid; + modpath = mp; + label = l; + variant = STRUCT(params_body,result_body) } + in + mp, { old = senv; + env = env; + modinfo = modinfo; + labset = Labset.empty; + revsign = []; + revstruct = []; + imports = senv.imports; + loads = [] } + + + +let end_module l senv = + let oldsenv = senv.old in + let modinfo = senv.modinfo in + let params, restype = + match modinfo.variant with + | NONE | LIBRARY _ | SIG _ -> error_no_module_to_end () + | STRUCT(params,restype) -> (params,restype) + in + if l <> modinfo.label then error_incompatible_labels l modinfo.label; + let functorize_type = + List.fold_right + (fun (arg_id,arg_b) mtb -> MTBfunsig (arg_id,arg_b,mtb)) + params + in + let auto_tb = MTBsig (modinfo.msid, List.rev senv.revsign) in + let mtb, mod_user_type = + match restype with + | None -> functorize_type auto_tb, None + | Some res_tb -> + let cnstrs = check_subtypes senv.env auto_tb res_tb in + let mtb = functorize_type res_tb in + mtb, Some (mtb, cnstrs) + in + let mexpr = + List.fold_right + (fun (arg_id,arg_b) mtb -> MEBfunctor (arg_id,arg_b,mtb)) + params + (MEBstruct (modinfo.msid, List.rev senv.revstruct)) + in + let mb = + { mod_expr = Some mexpr; + mod_user_type = mod_user_type; + mod_type = mtb; + mod_equiv = None } + in + let mspec = mtb , None in + let mp = MPdot (oldsenv.modinfo.modpath, l) in + let newenv = oldsenv.env in + let newenv = + List.fold_left + (fun env (mp,mb) -> full_add_module mp mb env) + newenv + senv.loads + in + let newenv = + full_add_module mp mb newenv + in + mp, { old = oldsenv.old; + env = newenv; + modinfo = oldsenv.modinfo; + labset = Labset.add l oldsenv.labset; + revsign = (l,SPBmodule mspec)::oldsenv.revsign; + revstruct = (l,SEBmodule mb)::oldsenv.revstruct; + imports = senv.imports; + loads = senv.loads@oldsenv.loads } + + +(* Interactive module types *) + +let start_modtype dir l params senv = + check_label l senv.labset; + let rec trans_params env = function + | [] -> env,[] + | (mbid,mte)::rest -> + let mtb = translate_modtype env mte in + let env = + full_add_module (MPbound mbid) (module_body_of_type mtb) env + in + let env,transrest = trans_params env rest in + env, (mbid,mtb)::transrest + in + let env,params_body = trans_params senv.env params in + let msid = make_msid dir (string_of_label l) in + let mp = MPself msid in + let modinfo = { msid = msid; + modpath = mp; + label = l; + variant = SIG params_body } + in + mp, { old = senv; + env = env; + modinfo = modinfo; + labset = Labset.empty; + revsign = []; + revstruct = []; + imports = senv.imports; + loads = [] } -let add_constraints = Environ.add_constraints +let end_modtype l senv = + let oldsenv = senv.old in + let modinfo = senv.modinfo in + let params = + match modinfo.variant with + | LIBRARY _ | NONE | STRUCT _ -> error_no_modtype_to_end () + | SIG params -> params + in + if l <> modinfo.label then error_incompatible_labels l modinfo.label; + let res_tb = MTBsig (modinfo.msid, List.rev senv.revsign) in + let mtb = + List.fold_right + (fun (arg_id,arg_b) mtb -> MTBfunsig (arg_id,arg_b,mtb)) + params + res_tb + in + let kn = make_kn oldsenv.modinfo.modpath empty_dirpath l in + let newenv = oldsenv.env in + let newenv = + List.fold_left + (fun env (mp,mb) -> full_add_module mp mb env) + newenv + senv.loads + in + let newenv = + add_modtype_constraints newenv mtb + in + let newenv = + Environ.add_modtype kn mtb newenv + in + kn, { old = oldsenv.old; + env = newenv; + modinfo = oldsenv.modinfo; + labset = Labset.add l oldsenv.labset; + revsign = (l,SPBmodtype mtb)::oldsenv.revsign; + revstruct = (l,SEBmodtype mtb)::oldsenv.revstruct; + imports = senv.imports; + loads = senv.loads@oldsenv.loads } + -let export = export -let import = import +let current_modpath senv = senv.modinfo.modpath +let current_msid senv = senv.modinfo.msid + + +let add_constraints cst senv = + {senv with env = Environ.add_constraints cst senv.env} + + + +(* Libraries = Compiled modules *) + +type compiled_library = + dir_path * module_body * library_info list + + +(* We check that only initial state Require's were performed before + [start_library] was called *) + +let start_library dir senv = + if not (senv.revsign = [] && + senv.modinfo.msid = initial_msid && + senv.modinfo.variant = NONE) + then + anomaly "Safe_typing.start_library: environment should be empty"; + let dir_path,l = + match (repr_dirpath dir) with + [] -> anomaly "Empty dirpath in Safe_typing.start_library" + | hd::tl -> + make_dirpath tl, label_of_id hd + in + let msid = make_msid dir_path (string_of_label l) in + let mp = MPself msid in + let modinfo = { msid = msid; + modpath = mp; + label = l; + variant = LIBRARY dir } + in + mp, { old = senv; + env = senv.env; + modinfo = modinfo; + labset = Labset.empty; + revsign = []; + revstruct = []; + imports = senv.imports; + loads = [] } + + +let export senv dir = + let modinfo = senv.modinfo in + begin + match modinfo.variant with + | LIBRARY dp -> + if dir <> dp then + anomaly "We are not exporting the right library!" + | _ -> + anomaly "We are not exporting the library" + end; + (*if senv.modinfo.params <> [] || senv.modinfo.restype <> None then + (* error_export_simple *) (); *) + let mb = + { mod_expr = Some (MEBstruct (modinfo.msid, List.rev senv.revstruct)); + mod_type = MTBsig (modinfo.msid, List.rev senv.revsign); + mod_user_type = None; + mod_equiv = None } + in + modinfo.msid, (dir,mb,senv.imports) + + +let import_constraints g kn cst = + try + merge_constraints cst g + with UniverseInconsistency -> + error "import_constraints: Universe Inconsistency during import" + + +let check_imports senv needed = + let imports = senv.imports in + let check (id,stamp) = + try + let actual_stamp = List.assoc id imports in + if stamp <> actual_stamp then + error ("Inconsistent assumptions over module " ^(string_of_dirpath id)) + with Not_found -> + error ("Reference to unknown module " ^ (string_of_dirpath id)) + in + List.iter check needed + + +(* we have an inefficiency: Since loaded files are added to the +environment every time a module is closed, their components are +calculated many times. Thic could be avoided in several ways: + +1 - for each file create a dummy environment containing only this +file's components, merge this environment with the global +environment, and store for the future (instead of just its type) + +2 - create "persistent modules" environment table in Environ add put +loaded by side-effect once and for all (like it is done in OCaml). +Would this be correct with respect to undo's and stuff ? +*) + +let import (dp,mb,depends) digest senv = + check_imports senv depends; + let mp = MPfile dp in + let env = senv.env in + mp, { senv with + env = full_add_module mp mb env; + imports = (dp,digest)::senv.imports; + loads = (mp,mb)::senv.loads } + + + +type judgment = unsafe_judgment -let env_of_safe_env e = e +let j_val j = j.uj_val +let j_type j = j.uj_type -let typing = Typeops.typing +let safe_infer senv = infer (env_of_senv senv) + +let typing senv = Typeops.typing (env_of_senv senv) diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli index 10357f407..2d4d2403c 100644 --- a/kernel/safe_typing.mli +++ b/kernel/safe_typing.mli @@ -12,12 +12,16 @@ open Names open Term open Declarations +open Entries (*i*) -(*s Safe environments. Since we are now able to type terms, we can define an - abstract type of safe environments, where objects are typed before being - added. Internally, the datatype is still [env]. We re-export the - functions of [Environ] for the new type [environment]. *) +(*s Safe environments. Since we are now able to type terms, we can + define an abstract type of safe environments, where objects are + typed before being added. + + We also add [open_structure] and [close_section], [close_module] to + provide functionnality for sections and interactive modules +*) type safe_environment @@ -34,30 +38,69 @@ val push_named_def : Univ.constraints * safe_environment (* Adding global axioms or definitions *) -type constant_entry = { - const_entry_body : constr; - const_entry_type : types option; - const_entry_opaque : bool } - type global_declaration = | ConstantEntry of constant_entry - | ParameterEntry of types | GlobalRecipe of Cooking.recipe val add_constant : - constant -> global_declaration -> safe_environment -> safe_environment + dir_path -> label -> global_declaration -> safe_environment -> + kernel_name * safe_environment (* Adding an inductive type *) val add_mind : - mutual_inductive -> Indtypes.mutual_inductive_entry -> safe_environment -> - safe_environment + dir_path -> label -> mutual_inductive_entry -> safe_environment -> + mutual_inductive * safe_environment + +(* Adding a module *) +val add_module : + label -> module_entry -> safe_environment + -> module_path * safe_environment + +(* Adding a module type *) +val add_modtype : + label -> module_type_entry -> safe_environment + -> kernel_name * safe_environment (* Adding universe constraints *) -val add_constraints : Univ.constraints -> safe_environment -> safe_environment +val add_constraints : + Univ.constraints -> safe_environment -> safe_environment + + +(*s Interactive module functions *) +val start_module : + dir_path -> label -> (mod_bound_id * module_type_entry) list + -> module_type_entry option + -> safe_environment -> module_path * safe_environment + +val end_module : + label -> safe_environment -> module_path * safe_environment + + +val start_modtype : + dir_path -> label -> (mod_bound_id * module_type_entry) list + -> safe_environment -> module_path * safe_environment + +val end_modtype : + label -> safe_environment -> kernel_name * safe_environment + + +val current_modpath : safe_environment -> module_path +val current_msid : safe_environment -> mod_self_id + + +(* Loading and saving compilation units *) + +(* exporting and importing modules *) +type compiled_library + +val start_library : dir_path -> safe_environment + -> module_path * safe_environment + +val export : safe_environment -> dir_path + -> mod_self_id * compiled_library -(* Loading and saving a module *) -val export : safe_environment -> dir_path -> Environ.compiled_env -val import : Environ.compiled_env -> safe_environment -> safe_environment +val import : compiled_library -> Digest.t -> safe_environment + -> module_path * safe_environment (*s Typing judgments *) diff --git a/kernel/sign.ml b/kernel/sign.ml index f0c275fa3..29e7379e9 100644 --- a/kernel/sign.ml +++ b/kernel/sign.ml @@ -71,6 +71,15 @@ let rel_context_nhyps hyps = let fold_rel_context f l ~init:x = List.fold_right f l x let fold_rel_context_reverse f ~init:x l = List.fold_left f x l +let map_rel_context f l = + let map_decl (n, body_o, typ as decl) = + let body_o' = option_smartmap f body_o in + let typ' = f typ in + if body_o' == body_o && typ' == typ then decl else + (n, body_o', typ') + in + list_smartmap map_decl l + (* Push named declarations on top of a rel context *) (* Bizarre. Should be avoided. *) let push_named_to_rel_context hyps ctxt = diff --git a/kernel/sign.mli b/kernel/sign.mli index e5de5613d..d66421fbd 100644 --- a/kernel/sign.mli +++ b/kernel/sign.mli @@ -56,6 +56,9 @@ val fold_rel_context : val fold_rel_context_reverse : ('a -> rel_declaration -> 'a) -> init:'a -> rel_context -> 'a +(*s Map function of [rel_context] *) +val map_rel_context : (constr -> constr) -> rel_context -> rel_context + (*s Term constructors *) val it_mkLambda_or_LetIn : constr -> rel_context -> constr diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml new file mode 100644 index 000000000..5a7ce3cb3 --- /dev/null +++ b/kernel/subtyping.ml @@ -0,0 +1,241 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Idmap.add id (IndConstr ((ip,i), mib)) map) + oib.mind_consnames + map + in + Idmap.add oib.mind_typename (IndType (ip, mib)) map + in + array_fold_right_i add_nameobjects_of_one mib.mind_packets map + +(* creates namedobject map for the whole signature *) + +let make_label_map msid list = + let add_one (l,e) map = + let obj = + match e with + | SPBconst cb -> Constant cb + | SPBmind mib -> Mind mib + | SPBmodule mb -> Module mb + | SPBmodtype mtb -> Modtype mtb + in +(* let map = match obj with + | Mind mib -> + add_nameobjects_of_mib (make_ln (MPself msid) l) mib map + | _ -> map + in *) + Labmap.add l obj map + in + List.fold_right add_one list Labmap.empty + +let check_conv_error error cst f env a1 a2 = + try + Constraint.union cst (f env a1 a2) + with + NotConvertible -> error () + +(* for now we do not allow reorderings *) +let check_inductive cst env msid1 l info1 mib2 spec2 = + let kn = make_kn (MPself msid1) empty_dirpath l in + let error () = error_not_match l spec2 in + let check_conv cst f = check_conv_error error cst f in + let mib1 = + match info1 with + | Mind mib -> mib + (* | IndType (_,mib) -> mib we will enable this later*) + | _ -> error () + in + let check_packet cst p1 p2 = + let check f = if f p1 <> f p2 then error () in + check (fun p -> p.mind_consnames); + check (fun p -> p.mind_typename); + (* nf_lc later *) + (* nf_arity later *) + (* user_lc ignored *) + (* user_arity ignored *) + let cst = check_conv cst conv_sort env p1.mind_sort p2.mind_sort in + check (fun p -> p.mind_nrealargs); + (* kelim ignored *) + (* listrec ignored *) + (* finite done *) + (* nparams done *) + (* params_ctxt done *) + let cst = check_conv cst conv env p1.mind_nf_arity p2.mind_nf_arity in + cst + in + (* this function uses mis_something and the others do not. + Perhaps we should uniformize it at some point... *) + let check_cons_types i cst p1 p2 = + array_fold_left2 + (fun cst t1 t2 -> check_conv cst conv env t1 t2) + cst + (arities_of_specif kn (mib1,p1)) + (arities_of_specif kn (mib2,p2)) + in + let check f = if f mib1 <> f mib2 then error () in + check (fun mib -> mib.mind_finite); + check (fun mib -> mib.mind_ntypes); + assert (mib1.mind_hyps=[] && mib2.mind_hyps=[]); + assert (Array.length mib1.mind_packets >= 1 + && Array.length mib2.mind_packets >= 1); + check (fun mib -> mib.mind_packets.(0).mind_nparams); + check (fun mib -> mib.mind_packets.(0).mind_params_ctxt); + (* TODO: we should allow renaming of parameters at least ! *) + let cst = match mib1.mind_singl, mib2.mind_singl with + None, None -> cst + | Some c1, Some c2 -> check_conv cst conv env c1 c2 + | _ -> error () + in + (* we first check simple things *) + let cst = + array_fold_left2 check_packet cst mib1.mind_packets mib2.mind_packets + in + (* and constructor types in the end *) + let cst = + array_fold_left2_i check_cons_types cst mib1.mind_packets mib2.mind_packets + in + cst + +let check_constant cst env msid1 l info1 cb2 spec2 = + let error () = error_not_match l spec2 in + let check_conv cst f = check_conv_error error cst f in + let cb1 = + match info1 with + | Constant cb -> cb + | _ -> error () + in + assert (cb1.const_hyps=[] && cb2.const_hyps=[]) ; + let cst = check_conv cst conv_leq env cb1.const_type cb2.const_type in + match cb1.const_body, cb2.const_body with + | None, None -> cst + | Some _, None -> cst + | Some c1, Some c2 -> + check_conv cst conv env c1 c2 + | None, Some c2 -> + check_conv + cst + conv + env + (mkConst (make_kn (MPself msid1) empty_dirpath l)) + c2 + +let rec check_modules cst env msid1 l msb1 msb2 = + let cst = check_modtypes cst env (fst msb1) (fst msb2) false in + begin + match (snd msb1), (snd msb2) with + | _, None -> () + | None, Some mp2 -> + check_modpath_equiv env (MPdot(MPself msid1,l)) mp2 + | Some mp1, Some mp2 -> + check_modpath_equiv env mp1 mp2 + end; + cst + + +and check_signatures cst env' (msid1,sig1) (msid2,sig2') = + let mp1 = MPself msid1 in + let env = add_signature mp1 sig1 env' in + let sig2 = subst_signature_msid msid2 mp1 sig2' in + let map1 = make_label_map msid1 sig1 in + let check_one_body cst (l,spec2) = + let info1 = + try + Labmap.find l map1 + with + Not_found -> error_no_such_label l + in + match spec2 with + | SPBconst cb2 -> + check_constant cst env msid1 l info1 cb2 spec2 + | SPBmind mib2 -> + check_inductive cst env msid1 l info1 mib2 spec2 + | SPBmodule msb2 -> + let msb1 = + match info1 with + | Module msb -> msb + | _ -> error_not_match l spec2 + in + check_modules cst env msid1 l msb1 msb2 + | SPBmodtype mtb2 -> + let mtb1 = + match info1 with + | Modtype mtb -> mtb + | _ -> error_not_match l spec2 + in + check_modtypes cst env mtb1 mtb2 true + in + List.fold_left check_one_body cst sig2 + +and check_modtypes cst env mtb1 mtb2 equiv = + if mtb1==mtb2 then (); (* just in case :) *) + let mtb1' = scrape_modtype env mtb1 in + let mtb2' = scrape_modtype env mtb2 in + if mtb1'==mtb2' then (); + match mtb1', mtb2' with + | MTBsig (msid1,list1), + MTBsig (msid2,list2) -> + let cst = check_signatures cst env (msid1,list1) (msid2,list2) in + if equiv then + check_signatures cst env (msid2,list2) (msid1,list1) + else + cst + | MTBfunsig (arg_id1,arg_t1,body_t1), + MTBfunsig (arg_id2,arg_t2,body_t2) -> + let cst = check_modtypes cst env arg_t2 arg_t1 equiv in + (* contravariant *) + let env' = + add_module (MPbound arg_id2) (module_body_of_type arg_t2) env + in + let body_t1' = + subst_modtype + (map_mbid arg_id1 (MPbound arg_id2)) + body_t1 + in + check_modtypes cst env' body_t1' body_t2 equiv + | MTBident _ , _ -> anomaly "Subtyping: scrape failed" + | _ , MTBident _ -> anomaly "Subtyping: scrape failed" + | _ , _ -> error_incompatible_modtypes mtb1 mtb2 + +let check_subtypes env sup super = + check_modtypes Constraint.empty env sup super false + diff --git a/kernel/subtyping.mli b/kernel/subtyping.mli new file mode 100644 index 000000000..71c111339 --- /dev/null +++ b/kernel/subtyping.mli @@ -0,0 +1,19 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* module_type_body -> module_type_body -> constraints + + diff --git a/kernel/term.ml b/kernel/term.ml index 95968f6c8..0c5dc2a54 100644 --- a/kernel/term.ml +++ b/kernel/term.ml @@ -147,7 +147,7 @@ let comp_term t1 t2 = & array_for_all2 (==) bl1 bl2 | _ -> false -let hash_term (sh_rec,(sh_sort,sh_sp,sh_na,sh_id)) t = +let hash_term (sh_rec,(sh_sort,sh_kn,sh_na,sh_id)) t = match t with | Rel _ | Meta _ -> t | Var x -> Var (sh_id x) @@ -158,10 +158,10 @@ let hash_term (sh_rec,(sh_sort,sh_sp,sh_na,sh_id)) t = | LetIn (na,b,t,c) -> LetIn (sh_na na, sh_rec b, sh_rec t, sh_rec c) | App (c,l) -> App (sh_rec c, Array.map sh_rec l) | Evar (e,l) -> Evar (e, Array.map sh_rec l) - | Const c -> Const (sh_sp c) - | Ind (sp,i) -> Ind (sh_sp sp,i) - | Construct ((sp,i),j) -> Construct ((sh_sp sp,i),j) - | Case (ci,p,c,bl) -> (* TO DO: extract ind_sp *) + | Const c -> Const (sh_kn c) + | Ind (kn,i) -> Ind (sh_kn kn,i) + | Construct ((kn,i),j) -> Construct ((sh_kn kn,i),j) + | Case (ci,p,c,bl) -> (* TO DO: extract ind_kn *) Case (ci, sh_rec p, sh_rec c, Array.map sh_rec bl) | Fix (ln,(lna,tl,bl)) -> Fix (ln,(Array.map sh_na lna, @@ -177,15 +177,15 @@ module Hconstr = struct type t = constr type u = (constr -> constr) * - ((sorts -> sorts) * (section_path -> section_path) + ((sorts -> sorts) * (kernel_name -> kernel_name) * (name -> name) * (identifier -> identifier)) let hash_sub = hash_term let equal = comp_term let hash = Hashtbl.hash end) -let hcons_term (hsorts,hsp,hname,hident) = - Hashcons.recursive_hcons Hconstr.f (hsorts,hsp,hname,hident) +let hcons_term (hsorts,hkn,hname,hident) = + Hashcons.recursive_hcons Hconstr.f (hsorts,hkn,hname,hident) (* Constructs a DeBrujin index with number n *) let rels = @@ -236,12 +236,12 @@ let mkConst c = Const c (* Constructs an existential variable *) let mkEvar e = Evar e -(* Constructs the ith (co)inductive type of the block named sp *) +(* Constructs the ith (co)inductive type of the block named kn *) (* The array of terms correspond to the variables introduced in the section *) let mkInd m = Ind m (* Constructs the jth constructor of the ith (co)inductive type of the - block named sp. The array of terms correspond to the variables + block named kn. The array of terms correspond to the variables introduced in the section *) let mkConstruct c = Construct c @@ -380,28 +380,28 @@ let isApp c = match kind_of_term c with App _ -> true | _ -> false (* Destructs a constant *) let destConst c = match kind_of_term c with - | Const sp -> sp + | Const kn -> kn | _ -> invalid_arg "destConst" let isConst c = match kind_of_term c with Const _ -> true | _ -> false (* Destructs an existential variable *) let destEvar c = match kind_of_term c with - | Evar (sp, a as r) -> r + | Evar (kn, a as r) -> r | _ -> invalid_arg "destEvar" let num_of_evar c = match kind_of_term c with | Evar (n, _) -> n | _ -> anomaly "num_of_evar called with bad args" -(* Destructs a (co)inductive type named sp *) +(* Destructs a (co)inductive type named kn *) let destInd c = match kind_of_term c with - | Ind (sp, a as r) -> r + | Ind (kn, a as r) -> r | _ -> invalid_arg "destInd" (* Destructs a constructor *) let destConstruct c = match kind_of_term c with - | Construct (sp, a as r) -> r + | Construct (kn, a as r) -> r | _ -> invalid_arg "dest" let isConstruct c = match kind_of_term c with @@ -813,6 +813,32 @@ let substn_vars p vars = let subst_vars = substn_vars 1 +(* +map_kn : (kernel_name -> kernel_name) -> constr -> constr + +This should be rewritten to prevent duplication of constr's when not +necessary. +For now, it uses map_constr and is rather ineffective +*) + +let rec map_kn f c = + let func = map_kn f in + match kind_of_term c with + | Const kn -> + mkConst (f kn) + | Ind (kn,i) -> + mkInd (f kn,i) + | Construct ((kn,i),j) -> + mkConstruct ((f kn,i),j) + | Case (ci,p,c,l) -> + let ci' = { ci with ci_ind = let (kn,i) = ci.ci_ind in f kn, i } in + mkCase (ci', func p, func c, array_smartmap func l) + | _ -> map_constr func c + +let subst_mps sub = + map_kn (subst_kn sub) + + (*********************) (* Term constructors *) (*********************) @@ -910,12 +936,12 @@ let mkConst = mkConst (* Constructs an existential variable *) let mkEvar = mkEvar -(* Constructs the ith (co)inductive type of the block named sp *) +(* Constructs the ith (co)inductive type of the block named kn *) (* The array of terms correspond to the variables introduced in the section *) let mkInd = mkInd (* Constructs the jth constructor of the ith (co)inductive type of the - block named sp. The array of terms correspond to the variables + block named kn. The array of terms correspond to the variables introduced in the section *) let mkConstruct = mkConstruct @@ -1169,9 +1195,9 @@ module Hsorts = let hsort = Hsorts.f -let hcons_constr (hspcci,hdir,hname,hident,hstr) = +let hcons_constr (hkn,hdir,hname,hident,hstr) = let hsortscci = Hashcons.simple_hcons hsort hcons1_univ in - let hcci = hcons_term (hsortscci,hspcci,hname,hident) in + let hcci = hcons_term (hsortscci,hkn,hname,hident) in let htcci = Hashcons.simple_hcons Htype.f (hcci,hsortscci) in (hcci,htcci) diff --git a/kernel/term.mli b/kernel/term.mli index 643d06443..6da9d1f5f 100644 --- a/kernel/term.mli +++ b/kernel/term.mli @@ -127,12 +127,12 @@ val mkConst : constant -> constr (* Inductive types *) -(* Constructs the ith (co)inductive type of the block named sp *) +(* Constructs the ith (co)inductive type of the block named kn *) (* The array of terms correspond to the variables introduced in the section *) val mkInd : inductive -> constr (* Constructs the jth constructor of the ith (co)inductive type of the - block named sp. The array of terms correspond to the variables + block named kn. The array of terms correspond to the variables introduced in the section *) val mkConstruct : constructor -> constr @@ -439,6 +439,12 @@ val subst_vars : identifier list -> constr -> constr if two names are identical, the one of least indice is keeped *) val substn_vars : int -> identifier list -> constr -> constr + +(* [subst_mps sub c] performs the substitution [sub] on all kernel + names appearing in [c] *) +val subst_mps : substitution -> constr -> constr + + (*s Functionals working on the immediate subterm of a construction *) (* [fold_constr f acc c] folds [f] on the immediate subterms of [c] @@ -496,7 +502,7 @@ val compare_constr : (constr -> constr -> bool) -> constr -> constr -> bool (*********************************************************************) val hcons_constr: - (section_path -> section_path) * + (kernel_name -> kernel_name) * (dir_path -> dir_path) * (name -> name) * (identifier -> identifier) * diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml new file mode 100644 index 000000000..1abe65c20 --- /dev/null +++ b/kernel/term_typing.ml @@ -0,0 +1,111 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* j.uj_type, cst1 + | Some t -> + let (tj,cst2) = infer_type env t in + let cst3 = + try conv_leq env j.uj_type tj.utj_val + with NotConvertible -> error_actual_type env j tj.utj_val in + tj.utj_val, Constraint.union (Constraint.union cst1 cst2) cst3 + + +let translate_local_def env (b,topt) = + let (j,cst) = infer env b in + let (typ,cst) = constrain_type env j cst topt in + (j.uj_val,typ,cst) + +let translate_local_assum env t = + let (j,cst) = infer env t in + let t = Typeops.assumption_of_judgment env j in + (t,cst) + +(* + +(* Same as push_named, but check that the variable is not already + there. Should *not* be done in Environ because tactics add temporary + hypothesis many many times, and the check performed here would + cost too much. *) +let safe_push_named (id,_,_ as d) env = + let _ = + try + let _ = lookup_named id env in + error ("identifier "^string_of_id id^" already defined") + with Not_found -> () in + push_named d env + +let push_named_def = push_rel_or_named_def safe_push_named +let push_rel_def = push_rel_or_named_def push_rel + +let push_rel_or_named_assum push (id,t) env = + let (j,cst) = safe_infer env t in + let t = Typeops.assumption_of_judgment env j in + let env' = add_constraints cst env in + let env'' = push (id,None,t) env' in + (cst,env'') + +let push_named_assum = push_rel_or_named_assum push_named +let push_rel_assum d env = snd (push_rel_or_named_assum push_rel d env) + +let push_rels_with_univ vars env = + List.fold_left (fun env nvar -> push_rel_assum nvar env) env vars +*) + + +(* Insertion of constants and parameters in environment. *) + +let infer_declaration env dcl = + match dcl with + | DefinitionEntry c -> + let (j,cst) = infer env c.const_entry_body in + let (typ,cst) = constrain_type env j cst c.const_entry_type in + Some j.uj_val, typ, cst, c.const_entry_opaque + | ParameterEntry t -> + let (j,cst) = infer env t in + None, Typeops.assumption_of_judgment env j, cst, false + +let build_constant_declaration env (body,typ,cst,op) = + let ids = match body with + | None -> global_vars_set env typ + | Some b -> + Idset.union (global_vars_set env b) (global_vars_set env typ) in + let hyps = keep_hyps env ids in + { const_body = body; + const_type = typ; + const_hyps = hyps; + const_constraints = cst; + const_opaque = op } + +(*s Global and local constant declaration. *) + +let translate_constant env ce = + build_constant_declaration env (infer_declaration env ce) + +let translate_recipe env r = + build_constant_declaration env (Cooking.cook_constant env r) + +(* Insertion of inductive types. *) + +let translate_mind env mie = check_inductive env mie diff --git a/kernel/term_typing.mli b/kernel/term_typing.mli new file mode 100644 index 000000000..4fb104c02 --- /dev/null +++ b/kernel/term_typing.mli @@ -0,0 +1,34 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* constr * types option -> + constr * types * Univ.constraints + +val translate_local_assum : env -> types -> + types * Univ.constraints + +val translate_constant : env -> constant_entry -> constant_body + +val translate_mind : + env -> mutual_inductive_entry -> mutual_inductive_body + +val translate_recipe : + env -> Cooking.recipe -> constant_body diff --git a/kernel/typeops.ml b/kernel/typeops.ml index 6da58adbc..29b09dde1 100644 --- a/kernel/typeops.ml +++ b/kernel/typeops.ml @@ -15,6 +15,7 @@ open Term open Declarations open Sign open Environ +open Entries open Reduction open Inductive open Type_errors @@ -242,8 +243,8 @@ let judge_of_cast env cj tj = let judge_of_inductive env i = let constr = mkInd i in let _ = - let (sp,_) = i in - let mib = lookup_mind sp env in + let (kn,_) = i in + let mib = lookup_mind kn env in check_args env constr mib.mind_hyps in make_judge constr (type_of_inductive env i) @@ -258,8 +259,8 @@ let judge_of_inductive env i let judge_of_constructor env c = let constr = mkConstruct c in let _ = - let ((sp,_),_) = c in - let mib = lookup_mind sp env in + let ((kn,_),_) = c in + let mib = lookup_mind kn env in check_args env constr mib.mind_hyps in make_judge constr (type_of_constructor env c) @@ -457,10 +458,6 @@ let infer_v env cv = (* Typing of several terms. *) -type local_entry = - | LocalDef of constr - | LocalAssum of constr - let infer_local_decl env id = function | LocalDef c -> let (j,cst) = infer env c in diff --git a/kernel/typeops.mli b/kernel/typeops.mli index 94d4e0844..55adf86a2 100644 --- a/kernel/typeops.mli +++ b/kernel/typeops.mli @@ -13,6 +13,7 @@ open Names open Univ open Term open Environ +open Entries (*i*) (*s Typing functions (not yet tagged as safe) *) @@ -21,10 +22,6 @@ val infer : env -> constr -> unsafe_judgment * constraints val infer_v : env -> constr array -> unsafe_judgment array * constraints val infer_type : env -> types -> unsafe_type_judgment * constraints -type local_entry = - | LocalDef of constr - | LocalAssum of constr - val infer_local_decls : env -> (identifier * local_entry) list -> env * Sign.rel_context * constraints diff --git a/lib/rtree.ml b/lib/rtree.ml index 9361a2858..593f73479 100644 --- a/lib/rtree.ml +++ b/lib/rtree.ml @@ -61,6 +61,26 @@ let rec subst_rtree_rec depth sub = function let subst_rtree sub t = subst_rtree_rec 0 sub t +let rec map f t = match t with + Param i -> Param i + | Node (a,sons) -> Node (f a, Array.map (map f) sons) + | Rec(j,defs) -> Rec (j, Array.map (map f) defs) + +let rec smartmap f t = match t with + Param i -> t + | Node (a,sons) -> + let a'=f a and sons' = Util.array_smartmap (map f) sons in + if a'==a && sons'==sons then + t + else + Node (a',sons') + | Rec(j,defs) -> + let defs' = Util.array_smartmap (map f) defs in + if defs'==defs then + t + else + Rec(j,defs') + (* To avoid looping, we must check that every body introduces a node or a parameter *) let rec expand_rtree = function diff --git a/lib/rtree.mli b/lib/rtree.mli index 5a34c819e..cb8fe0303 100644 --- a/lib/rtree.mli +++ b/lib/rtree.mli @@ -23,6 +23,11 @@ val mk_rec : 'a t array -> 'a t array to avoid captures when a tree appears under mk_rec *) val lift : int -> 'a t -> 'a t +val map : ('a -> 'b) -> 'a t -> 'b t + +(* [(smartmap f t) == t] if [(f a) ==a ] for all nodes *) +val smartmap : ('a -> 'a) -> 'a t -> 'a t + (* Destructors (recursive calls are expanded) *) val dest_param : 'a t -> int val dest_node : 'a t -> 'a * 'a t array diff --git a/lib/util.ml b/lib/util.ml index 3354bba53..98defb6fd 100644 --- a/lib/util.ml +++ b/lib/util.ml @@ -20,6 +20,8 @@ exception UserError of string * std_ppcmds (* User errors *) let error string = raise (UserError(string, str string)) let errorlabstrm l pps = raise (UserError(l,pps)) +let todo s = prerr_string ("TODO: "^s^"\n") + (* raising located exceptions *) type loc = int * int type 'a located = loc * 'a @@ -123,6 +125,13 @@ let list_assign l n e = in assrec [] (l,n) +let rec list_smartmap f l = match l with + [] -> l + | h::tl -> + let h' = f h and tl' = list_smartmap f tl in + if h'==h && tl'==tl then l + else h'::tl' + let list_map_left f = (* ensures the order in case of side-effects *) let rec map_rec = function | [] -> [] @@ -302,6 +311,23 @@ let list_share_tails l1 l2 = let list_join_map f l = List.flatten (List.map f l) +let rec list_fold_map f e = function + | [] -> (e,[]) + | h::t -> + let e',h' = f e h in + let e'',t' = list_fold_map f e' t in + e'',h'::t' + +(* (* tail-recursive version of the above function *) +let list_fold_map f e l = + let g (e,b') h = + let (e',h') = f e h in + (e',h'::b') + in + let (e',lrev) = List.fold_left g (e,[]) l in + (e',List.rev lrev) +*) + (* Arrays *) let array_exists f v = @@ -410,6 +436,35 @@ let array_chop n v = if n > vlen then failwith "array_chop"; (Array.sub v 0 n, Array.sub v n (vlen-n)) +exception Local of int + +(* If none of the elements is changed by f we return ar itself. + The for loop looks for the first such an element. + If found it is temporarily stored in a ref and the new array is produced, + but f is not re-applied to elements that are already checked *) +let array_smartmap f ar = + let ar_size = Array.length ar in + let aux = ref None in + try + for i = 0 to ar_size-1 do + let a = ar.(i) in + let a' = f a in + if a != a' then (* pointer (in)equality *) begin + aux := Some a'; + raise (Local i) + end + done; + ar + with + Local i -> + let copy j = + if j a' | None -> failwith "Error" + else f (ar.(j)) + in + Array.init ar_size copy + let array_map2 f v1 v2 = if Array.length v1 <> Array.length v2 then invalid_arg "array_map2"; if Array.length v1 == 0 then @@ -508,6 +563,10 @@ let option_compare f a b = match (a,b) with | Some a', Some b' -> f a' b' | _ -> failwith "option_compare" +let option_smartmap f a = match a with + | None -> a + | Some x -> let x' = f x in if x'==x then a else Some x' + let map_succeed f = let rec map_f = function | [] -> [] diff --git a/lib/util.mli b/lib/util.mli index e9cd97a7c..80b841b2b 100644 --- a/lib/util.mli +++ b/lib/util.mli @@ -26,6 +26,12 @@ exception UserError of string * std_ppcmds val error : string -> 'a val errorlabstrm : string -> std_ppcmds -> 'a +(* [todo] is for running of an incomplete code its implementation is + "do nothing" (or print a message), but this function should not be + used in a released code *) + +val todo : string -> unit + type loc = int * int type 'a located = loc * 'a @@ -67,6 +73,10 @@ val list_tabulate : (int -> 'a) -> int -> 'a list val list_assign : 'a list -> int -> 'a -> 'a list val list_distinct : 'a list -> bool val list_filter2 : ('a -> 'b -> bool) -> 'a list * 'b list -> 'a list * 'b list + +(* [list_smartmap f [a1...an] = List.map f [a1...an]] but if for all i + [ f ai == ai], then [list_smartmap f l==l] *) +val list_smartmap : ('a -> 'a) -> 'a list -> 'a list val list_map_left : ('a -> 'b) -> 'a list -> 'b list val list_map_i : (int -> 'a -> 'b) -> int -> 'a list -> 'b list val list_map2_i : @@ -98,6 +108,10 @@ val list_map_append : ('a -> 'b list) -> 'a list -> 'b list val list_map_append2 : ('a -> 'b -> 'c list) -> 'a list -> 'b list -> 'c list val list_share_tails : 'a list -> 'a list -> 'a list * 'a list * 'a list val list_join_map : ('a -> 'b list) -> 'a list -> 'b list +(* [list_fold_map f e_0 [l_1...l_n] = e_n,[k_1...k_n]] + where [(e_i,k_i)=f e_{i-1} l_i] *) +val list_fold_map : + ('a -> 'b -> 'a * 'c) -> 'a -> 'b list -> 'a * 'c list (*s Arrays. *) @@ -123,6 +137,7 @@ val array_app_tl : 'a array -> 'a list -> 'a list val array_list_of_tl : 'a array -> 'a list val array_map_to_list : ('a -> 'b) -> 'a array -> 'b list val array_chop : int -> 'a array -> 'a array * 'a array +val array_smartmap : ('a -> 'a) -> 'a array -> 'a array val array_map2 : ('a -> 'b -> 'c) -> 'a array -> 'b array -> 'c array val array_map2_i : (int -> 'a -> 'b -> 'c) -> 'a array -> 'b array -> 'c array val array_map3 : @@ -159,6 +174,7 @@ val in_some : 'a -> 'a option val out_some : 'a option -> 'a val option_app : ('a -> 'b) -> 'a option -> 'b option val option_compare : ('a -> 'b -> bool) -> 'a option -> 'b option -> bool +val option_smartmap : ('a -> 'a) -> 'a option -> 'a option (* In [map_succeed f l] an element [a] is removed if [f a] raises *) (* [Failure _] otherwise behaves as [List.map f l] *) diff --git a/library/declare.ml b/library/declare.ml index fc80cfda9..31af6dbbb 100644 --- a/library/declare.ml +++ b/library/declare.ml @@ -11,10 +11,12 @@ open Pp open Util open Names +open Libnames open Nameops open Term open Sign open Declarations +open Entries open Inductive open Indtypes open Reduction @@ -91,7 +93,7 @@ let _ = Summary.declare_summary "VARIABLE" Summary.init_function = (fun () -> vartab := Idmap.empty); Summary.survive_section = false } -let cache_variable (sp,(id,(p,d,strength))) = +let cache_variable ((sp,_),(id,(p,d,strength))) = (* Constr raisonne sur les noms courts *) if Idmap.mem id !vartab then errorlabstrm "cache_variable" (pr_id id ++ str " already exists"); @@ -100,17 +102,13 @@ let cache_variable (sp,(id,(p,d,strength))) = | SectionLocalDef (c,t) -> Global.push_named_def (id,c,t) in let (_,bd,ty) = Global.lookup_named id in let vd = (bd,ty,cst) in - Nametab.push 0 (restrict_path 0 sp) (VarRef id); + Nametab.push (Nametab.Until 1) (restrict_path 0 sp) (VarRef id); vartab := Idmap.add id (p,vd,strength) !vartab let (in_variable, out_variable) = - let od = { + declare_object { (default_object "VARIABLE") with cache_function = cache_variable; - load_function = (fun _ -> ()); - open_function = (fun _ -> ()); - export_function = (fun x -> Some x) } - in - declare_object ("VARIABLE", od) + classify_function = (fun _ -> Dispose) } let declare_variable id obj = let sp = add_leaf id (in_variable (id,obj)) in @@ -119,7 +117,7 @@ let declare_variable id obj = (* Globals: constants and parameters *) -type constant_declaration = global_declaration * strength +type constant_declaration = constant_entry * strength let csttab = ref (Spmap.empty : strength Spmap.t) @@ -129,90 +127,101 @@ let _ = Summary.declare_summary "CONSTANT" Summary.init_function = (fun () -> csttab := Spmap.empty); Summary.survive_section = false } -let cache_constant (sp,(cdt,stre)) = +let cache_constant ((sp,kn),(cdt,stre)) = (if Idmap.mem (basename sp) !vartab then errorlabstrm "cache_constant" (pr_id (basename sp) ++ str " already exists")); (if Nametab.exists_cci sp then let (_,id) = repr_path sp in errorlabstrm "cache_constant" (pr_id id ++ str " already exists")); - Global.add_constant sp cdt; + let _,dir,_ = repr_kn kn in + let kn' = Global.add_constant dir (basename sp) cdt in + if kn' <> kn then + anomaly "Kernel and Library names do not match"; (match stre with | DischargeAt (dp,n) when not (is_dirpath_prefix_of dp (Lib.cwd ())) -> (* Only qualifications including the sections segment from the current section to the discharge section is available for Remark & Fact *) - Nametab.push (n-Lib.sections_depth()) sp (ConstRef sp) + Nametab.push (Nametab.Until ((*n-Lib.sections_depth()+*)1)) sp (ConstRef kn) | (NeverDischarge| DischargeAt _) -> (* All qualifications of Theorem, Lemma & Definition are visible *) - Nametab.push 0 sp (ConstRef sp) + Nametab.push (Nametab.Until 1) sp (ConstRef kn) | NotDeclare -> assert false); csttab := Spmap.add sp stre !csttab (* At load-time, the segment starting from the module name to the discharge *) (* section (if Remark or Fact) is needed to access a construction *) -let load_constant (sp,(ce,stre)) = +let load_constant i ((sp,kn),(ce,stre)) = (if Nametab.exists_cci sp then let (_,id) = repr_path sp in errorlabstrm "cache_constant" (pr_id id ++ str " already exists")); csttab := Spmap.add sp stre !csttab; - Nametab.push (depth_of_strength stre + 1) sp (ConstRef sp) + Nametab.push (Nametab.Until ((*depth_of_strength stre + *)i)) sp (ConstRef kn) (* Opening means making the name without its module qualification available *) -let open_constant (sp,(_,stre)) = +let open_constant i ((sp,kn),(_,stre)) = let n = depth_of_strength stre in - Nametab.push n (restrict_path n sp) (ConstRef sp) + Nametab.push (Nametab.Exactly (i(*+n*))) sp (ConstRef kn) (* Hack to reduce the size of .vo: we keep only what load/open needs *) -let dummy_constant_entry = ParameterEntry mkProp +let dummy_constant_entry = ConstantEntry (ParameterEntry mkProp) + +let dummy_constant (ce,stre) = dummy_constant_entry,stre -let export_constant (ce,stre) = Some (dummy_constant_entry,stre) +let export_constant cst = Some (dummy_constant cst) + +let classify_constant (_,cst) = Substitute (dummy_constant cst) let (in_constant, out_constant) = - let od = { + declare_object { (default_object "CONSTANT") with cache_function = cache_constant; load_function = load_constant; open_function = open_constant; + classify_function = classify_constant; + subst_function = ident_subst_function; export_function = export_constant } - in - declare_object ("CONSTANT", od) let hcons_constant_declaration = function - | (ConstantEntry ce, stre) -> - (ConstantEntry + | (DefinitionEntry ce, stre) -> + (DefinitionEntry { const_entry_body = hcons1_constr ce.const_entry_body; const_entry_type = option_app hcons1_constr ce.const_entry_type; const_entry_opaque = ce.const_entry_opaque }, stre) | cd -> cd -let declare_constant id cd = +let declare_constant id (cd,stre) = (* let cd = hcons_constant_declaration cd in *) - let sp = add_leaf id (in_constant cd) in - if is_implicit_args() then declare_constant_implicits sp; - sp + let oname = add_leaf id (in_constant (ConstantEntry cd,stre)) in + if is_implicit_args() then declare_constant_implicits (snd oname); + oname -let redeclare_constant sp (cd,stre) = - add_absolutely_named_leaf sp (in_constant (GlobalRecipe cd,stre)); - if is_implicit_args() then declare_constant_implicits sp +let redeclare_constant id (cd,stre) = + let _,kn = add_leaf id (in_constant (GlobalRecipe cd,stre)) in + if is_implicit_args() then declare_constant_implicits kn (* Inductives. *) -let inductive_names sp mie = +let inductive_names sp kn mie = let (dp,_) = repr_path sp in let names, _ = List.fold_left (fun (names, n) ind -> - let indsp = (sp,n) in + let ind_p = (kn,n) in let names, _ = List.fold_left - (fun (names, p) id -> - let sp = Names.make_path dp id in - ((sp, ConstructRef (indsp,p)) :: names, p+1)) + (fun (names, p) l -> + let sp = + Libnames.make_path dp l + in + ((sp, ConstructRef (ind_p,p)) :: names, p+1)) (names, 1) ind.mind_entry_consnames in - let sp = Names.make_path dp ind.mind_entry_typename in - ((sp, IndRef indsp) :: names, n+1)) + let sp = Libnames.make_path dp ind.mind_entry_typename + in + ((sp, IndRef ind_p) :: names, n+1)) ([], 0) mie.mind_entry_inds in names + let check_exists_inductive (sp,_) = (if Idmap.mem (basename sp) !vartab then errorlabstrm "cache_inductive" @@ -221,22 +230,27 @@ let check_exists_inductive (sp,_) = let (_,id) = repr_path sp in errorlabstrm "cache_inductive" (pr_id id ++ str " already exists") -let cache_inductive (sp,mie) = - let names = inductive_names sp mie in +let cache_inductive ((sp,kn),mie) = + let names = inductive_names sp kn mie in List.iter check_exists_inductive names; - Global.add_mind sp mie; + let _,dir,_ = repr_kn kn in + let kn' = Global.add_mind dir (basename sp) mie in + if kn' <> kn then + anomaly "Kernel and Library names do not match"; + List.iter - (fun (sp, ref) -> Nametab.push 0 sp ref) + (fun (sp, ref) -> Nametab.push (Nametab.Until 1) sp ref) names -let load_inductive (sp,mie) = - let names = inductive_names sp mie in +let load_inductive i ((sp,kn),mie) = + let names = inductive_names sp kn mie in List.iter check_exists_inductive names; - List.iter (fun (sp, ref) -> Nametab.push 1 sp ref) names + List.iter (fun (sp, ref) -> Nametab.push (Nametab.Until i) sp ref) names -let open_inductive (sp,mie) = - let names = inductive_names sp mie in - List.iter (fun (sp, ref) -> Nametab.push 0 (restrict_path 0 sp) ref) names +let open_inductive i ((sp,kn),mie) = + let names = inductive_names sp kn mie in +(* List.iter (fun (sp, ref) -> Nametab.push 0 (restrict_path 0 sp) ref) names*) + List.iter (fun (sp, ref) -> Nametab.push (Nametab.Exactly i) sp ref) names let dummy_one_inductive_entry mie = { mind_entry_params = []; @@ -254,28 +268,28 @@ let dummy_inductive_entry m = { let export_inductive x = Some (dummy_inductive_entry x) let (in_inductive, out_inductive) = - let od = { + declare_object {(default_object "INDUCTIVE") with cache_function = cache_inductive; load_function = load_inductive; open_function = open_inductive; + classify_function = (fun (_,a) -> Substitute (dummy_inductive_entry a)); + subst_function = ident_subst_function; export_function = export_inductive } - in - declare_object ("INDUCTIVE", od) let declare_mind mie = let id = match mie.mind_entry_inds with | ind::_ -> ind.mind_entry_typename | [] -> anomaly "cannot declare an empty list of inductives" in - let sp = add_leaf id (in_inductive mie) in - if is_implicit_args() then declare_mib_implicits sp; - sp + let oname = add_leaf id (in_inductive mie) in + if is_implicit_args() then declare_mib_implicits (snd oname); + oname (*s Test and access functions. *) let is_constant sp = - try let _ = Global.lookup_constant sp in true with Not_found -> false + try let _ = Spmap.find sp !csttab in true with Not_found -> false let constant_strength sp = Spmap.find sp !csttab @@ -291,7 +305,7 @@ let variable_strength id = let (_,_,str) = Idmap.find id !vartab in str let find_section_variable id = - let (p,_,_) = Idmap.find id !vartab in Names.make_path p id + let (p,_,_) = Idmap.find id !vartab in Libnames.make_path p id (* Global references. *) @@ -353,12 +367,14 @@ let construct_qualified_reference qid = let ref = Nametab.locate qid in constr_of_reference ref -let construct_reference env id = - try - mkVar (let _ = Environ.lookup_named id env in id) - with Not_found -> - let ref = Nametab.sp_of_id id in - constr_of_reference ref +let construct_reference ctx_opt id = + match ctx_opt with + | None -> construct_qualified_reference (make_short_qualid id) + | Some ctx -> + try + mkVar (let _ = Sign.lookup_named id ctx in id) + with Not_found -> + construct_qualified_reference (make_short_qualid id) let global_qualified_reference qid = construct_qualified_reference qid @@ -370,36 +386,42 @@ let global_reference_in_absolute_module dir id = constr_of_reference (Nametab.locate_in_absolute_module dir id) let global_reference id = - construct_reference (Global.env()) id - -let dirpath sp = let (p,_) = repr_path sp in p - -let dirpath_of_global = function - | VarRef id -> empty_dirpath - | ConstRef sp -> dirpath sp - | IndRef (sp,_) -> dirpath sp - | ConstructRef ((sp,_),_) -> dirpath sp + construct_qualified_reference (make_short_qualid id) let is_section_variable = function | VarRef _ -> true | _ -> false +(* TODO temporary hack!!! *) +let rec is_imported_modpath = function + | MPfile dp -> dp <> (Lib.module_dp ()) +(* | MPdot (mp,_) -> is_imported_modpath mp *) + | _ -> false + +let is_imported_ref = function + | VarRef _ -> false + | ConstRef kn + | IndRef (kn,_) + | ConstructRef ((kn,_),_) +(* | ModTypeRef ln *) -> + let (mp,_,_) = repr_kn kn in is_imported_modpath mp +(* | ModRef mp -> + is_imported_modpath mp +*) let is_global id = try - let osp = Nametab.locate (make_short_qualid id) in - (* Compatibilité V6.3: Les variables de section ne sont pas globales - not (is_section_variable osp) && *) - is_dirpath_prefix_of (dirpath_of_global osp) (Lib.cwd()) + let ref = Nametab.locate (make_short_qualid id) in + not (is_imported_ref ref) with Not_found -> false -let strength_of_global = function - | ConstRef sp -> constant_strength sp +let strength_of_global ref = match ref with + | ConstRef kn -> constant_strength (full_name ref) | VarRef id -> variable_strength id | IndRef _ | ConstructRef _ -> NeverDischarge let library_part ref = - let sp = Nametab.sp_of_global (Global.env ()) ref in + let sp = Nametab.sp_of_global None ref in let dir,_ = repr_path sp in match strength_of_global ref with | DischargeAt (dp,n) -> @@ -412,3 +434,4 @@ let library_part ref = (* Theorem/Lemma outside its outer section of definition *) dir | NotDeclare -> assert false + diff --git a/library/declare.mli b/library/declare.mli index 07b9d98b6..273a2b344 100644 --- a/library/declare.mli +++ b/library/declare.mli @@ -10,9 +10,11 @@ (*i*) open Names +open Libnames open Term open Sign open Declarations +open Entries open Indtypes open Safe_typing open Library @@ -34,16 +36,16 @@ type section_variable_entry = type variable_declaration = dir_path * section_variable_entry * strength -val declare_variable : variable -> variable_declaration -> section_path +val declare_variable : variable -> variable_declaration -> object_name -type constant_declaration = global_declaration * strength +type constant_declaration = constant_entry * strength (* [declare_constant id cd] declares a global declaration (constant/parameter) with name [id] in the current section; it returns the full path of the declaration *) -val declare_constant : identifier -> constant_declaration -> constant +val declare_constant : identifier -> constant_declaration -> object_name -val redeclare_constant : constant -> Cooking.recipe * strength -> unit +val redeclare_constant : identifier -> Cooking.recipe * strength -> unit (* val declare_parameter : identifier -> constr -> constant @@ -52,7 +54,7 @@ val declare_parameter : identifier -> constr -> constant (* [declare_mind me] declares a block of inductive types with their constructors in the current section; it returns the path of the whole block *) -val declare_mind : mutual_inductive_entry -> mutual_inductive +val declare_mind : mutual_inductive_entry -> object_name val out_inductive : Libobject.obj -> mutual_inductive_entry @@ -65,7 +67,7 @@ val strength_min : strength * strength -> strength (*s Corresponding test and access functions. *) val is_constant : section_path -> bool -val constant_strength : constant -> strength +val constant_strength : section_path -> strength val out_variable : Libobject.obj -> identifier * variable_declaration val get_variable : variable -> named_declaration * strength @@ -86,11 +88,11 @@ val constr_of_reference : global_reference -> constr raise [Not_found] if not a global *) val reference_of_constr : constr -> global_reference -val global_qualified_reference : Nametab.qualid -> constr +val global_qualified_reference : qualid -> constr val global_absolute_reference : section_path -> constr val global_reference_in_absolute_module : dir_path -> identifier -> constr -val construct_qualified_reference : Nametab.qualid -> constr +val construct_qualified_reference : qualid -> constr val construct_absolute_reference : section_path -> constr (* This should eventually disappear *) @@ -98,7 +100,7 @@ val construct_absolute_reference : section_path -> constr the name [id] in the global environment. It looks also for variables in a given environment instead of looking in the current global environment. *) val global_reference : identifier -> constr -val construct_reference : Environ.env -> identifier -> constr +val construct_reference : Sign.named_context option -> identifier -> constr val is_global : identifier -> bool diff --git a/library/declaremods.ml b/library/declaremods.ml new file mode 100644 index 000000000..76ed46619 --- /dev/null +++ b/library/declaremods.ml @@ -0,0 +1,630 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* + !modtab_substobjs, + !modtab_objects, + !openmod_info); + Summary.unfreeze_function = (fun (sobjs,objs,info) -> + modtab_substobjs := sobjs; + modtab_objects := objs; + openmod_info := info); + Summary.init_function = (fun () -> + modtab_substobjs := MPmap.empty; + modtab_objects := MPmap.empty; + openmod_info := ([],None)); + Summary.survive_section = false } + +(* auxiliary functions to transform section_path and kernel_name given + by Lib into module_path and dir_path needed for modules *) + +let mp_of_kn kn = + let mp,sec,l = repr_kn kn in + if sec=empty_dirpath then + MPdot (mp,l) + else + anomaly ("Non-empty section in module name!" ^ string_of_kn kn) + +let dir_of_sp sp = + let dir,id = repr_path sp in + extend_dirpath dir id + +let msid_of_mp = function + MPself msid -> msid + | _ -> anomaly "'Self' module path expected!" + +let msid_of_prefix (_,(mp,sec)) = + if sec=empty_dirpath then + msid_of_mp mp + else + anomaly ("Non-empty section in module name!" ^ + string_of_mp mp ^ "." ^ string_of_dirpath sec) + +(* This function registers the visibility of the module and iterates + through its components. It is called by plenty module functions *) + +let do_module exists what iter_objects i dir mp substobjs objects = + let prefix = (dir,(mp,empty_dirpath)) in + let dirinfo = DirModule (dir,(mp,empty_dirpath)) in + let vis = + if exists then + if + try Nametab.locate_dir (qualid_of_dirpath dir) = dirinfo + with Not_found -> false + then + Nametab.Exactly i + else + errorlabstrm (what^"_module") + (pr_dirpath dir ++ str " should already exist!") + else + if Nametab.exists_dir dir then + errorlabstrm (what^"_module") (pr_dirpath dir ++ str " already exists") + else + Nametab.Until i + in + Nametab.push_dir vis dir dirinfo; + modtab_substobjs := MPmap.add mp substobjs !modtab_substobjs; + match objects with + Some seg -> + modtab_objects := MPmap.add mp (prefix,seg) !modtab_objects; + iter_objects (i+1) prefix seg + | None -> () + + +let conv_names_do_module exists what iter_objects i (sp,kn) substobjs substituted = + let dir,mp = dir_of_sp sp, mp_of_kn kn in + do_module exists what iter_objects i dir mp substobjs substituted + +(* Interactive modules and module types cannot be recached! cache_mod* + functions can be called only once (and "end_mod*" set the flag to + false then) +*) + +let cache_module ((sp,kn as oname),(entry,substobjs,substituted)) = + let _ = match entry with + | None -> + anomaly "You must not recache interactive modules!" + | Some me -> + let mp = Global.add_module (basename sp) me in + if mp <> mp_of_kn kn then + anomaly "Kernel and Library names do not match" + in + conv_names_do_module false "cache" load_objects 1 oname substobjs substituted + + +(* TODO: This check is not essential *) +let check_empty s = function + | None -> () + | Some _ -> + anomaly ("We should never have full info in " ^ s^"!") + + +(* When this function is called the module itself is already in the + environment. This function loads its objects only *) + +let load_module i (oname,(entry,substobjs,substituted)) = + (* TODO: This check is not essential *) + check_empty "load_module" entry; + conv_names_do_module false "load" load_objects i oname substobjs substituted + + +let open_module i (oname,(entry,substobjs,substituted)) = + (* TODO: This check is not essential *) + check_empty "open_module" entry; + conv_names_do_module true "open" open_objects i oname substobjs substituted + + +let subst_substobjs dir mp (subst,mbids,msid,objs) = + match mbids with + | [] -> + let prefix = dir,(mp,empty_dirpath) in + Some (subst_objects prefix (add_msid msid mp subst) objs) + | _ -> None + + +let subst_module ((sp,kn),subst,(entry,substobjs,_)) = + check_empty "subst_module" entry; + let dir,mp = dir_of_sp sp, mp_of_kn kn in + let (sub,mbids,msid,objs) = substobjs in + let subst' = join sub subst in + (* substitutive_objects get the new substitution *) + let substobjs = (subst',mbids,msid,objs) in + (* if we are not a functor - calculate substitued. + We add "msid |-> mp" to the substitution *) + let substituted = subst_substobjs dir mp substobjs + in + (None,substobjs,substituted) + + +let classify_module (_,(_,substobjs,_)) = + Substitute (None,substobjs,None) + +let (in_module,out_module) = + declare_object {(default_object "MODULE") with + cache_function = cache_module; + load_function = load_module; + open_function = open_module; + subst_function = subst_module; + classify_function = classify_module; + export_function = (fun _ -> anomaly "No modules in sections!") } + + +let cache_keep _ = anomaly "This module should not be cached!" + +let load_keep i ((sp,kn),seg) = + let mp = mp_of_kn kn in + let prefix = dir_of_sp sp, (mp,empty_dirpath) in + begin + try + let prefix',objects = MPmap.find mp !modtab_objects in + if prefix' <> prefix then + anomaly "Two different modules with the same path!"; + modtab_objects := MPmap.add mp (prefix,objects@seg) !modtab_objects; + with + Not_found -> anomaly "Keep objects before substitutive" + end; + load_objects i prefix seg + +let open_keep i ((sp,kn),seg) = + let dirpath,mp = dir_of_sp sp, mp_of_kn kn in + open_objects i (dirpath,(mp,empty_dirpath)) seg + +let (in_modkeep,out_modkeep) = + declare_object {(default_object "MODULE KEEP OBJECTS") with + cache_function = cache_keep; + load_function = load_keep; + open_function = open_keep; + export_function = (fun _ -> anomaly "No modules in sections!") } + +(* we remember objects for a module type. In case of a declaration: + Module M:SIG:=... + The module M gets its objects from SIG +*) +let modtypetab = + ref (KNmap.empty : substitutive_objects KNmap.t) + +(* currently started interactive module type. We remember its arguments + if it is a functor type *) +let openmodtype_info = + ref ([] : mod_bound_id list) + +let _ = Summary.declare_summary "MODTYPE-INFO" + { Summary.freeze_function = (fun () -> + !modtypetab,!openmodtype_info); + Summary.unfreeze_function = (fun ft -> + modtypetab := fst ft; + openmodtype_info := snd ft); + Summary.init_function = (fun () -> + modtypetab := KNmap.empty; + openmodtype_info := []); + Summary.survive_section = true } + + + + +let cache_modtype ((sp,kn),(entry,modtypeobjs)) = + let _ = + match entry with + | None -> + anomaly "You must not recache interactive module types!" + | Some mte -> + let kn' = Global.add_modtype (basename sp) mte in + if kn' <> kn then + anomaly "Kernel and Library names do not match" + in + + if Nametab.exists_modtype sp then + errorlabstrm "cache_modtype" + (pr_sp sp ++ str " already exists") ; + + Nametab.push_modtype (Nametab.Until 1) sp kn; + + modtypetab := KNmap.add kn modtypeobjs !modtypetab + + +let load_modtype i ((sp,kn),(entry,modtypeobjs)) = + check_empty "load_modtype" entry; + + if Nametab.exists_modtype sp then + errorlabstrm "cache_modtype" + (pr_sp sp ++ str " already exists") ; + + Nametab.push_modtype (Nametab.Until i) sp kn; + + modtypetab := KNmap.add kn modtypeobjs !modtypetab + + +let open_modtype i ((sp,kn),(entry,_)) = + check_empty "open_modtype" entry; + + if + try Nametab.locate_modtype (qualid_of_sp sp) <> kn + with Not_found -> true + then + errorlabstrm ("open_modtype") + (pr_sp sp ++ str " should already exist!"); + + Nametab.push_modtype (Nametab.Exactly i) sp kn + + +let subst_modtype (_,subst,(entry,(subs,mbids,msid,objs))) = + check_empty "subst_modtype" entry; + (entry,(join subs subst,mbids,msid,objs)) + + +let classify_modtype (_,(_,substobjs)) = + Substitute (None,substobjs) + + +let (in_modtype,out_modtype) = + declare_object {(default_object "MODULE TYPE") with + cache_function = cache_modtype; + load_function = load_modtype; + subst_function = subst_modtype; + classify_function = classify_modtype; + export_function = in_some } + + + +let abstract_substobjs mbids1 (subst, mbids2, msid, lib_stack) = + (subst, mbids1@mbids2, msid, lib_stack) + + +let rec get_modtype_substobjs = function + MTEident ln -> KNmap.find ln !modtypetab + | MTEfunsig (mbid,_,mte) -> + let (subst, mbids, msid, objs) = get_modtype_substobjs mte in + (subst, mbid::mbids, msid, objs) + | MTEsig (msid,_) -> (empty_subst, [], msid, []) + (* this is plainly wrong, but it is hard to do otherwise... *) + +(* push names of bound modules (and their components) to Nametab *) +(* add objects associated to them *) +let process_module_bindings argids args = + let process_arg id (mbid,mty) = + let dir = make_dirpath [id] in + let mp = MPbound mbid in + let substobjs = get_modtype_substobjs mty in + let substituted = subst_substobjs dir mp substobjs in + do_module false "begin" load_objects 1 dir mp substobjs substituted + in + List.iter2 process_arg argids args + +(* +(* this function removes keep objects from submodules *) +let rec kill_keep objs = + let kill = function + | (oname,Leaf obj) as node -> + if object_tag obj = "MODULE" then + let (entry,substobjs,substitute) = out_module obj in + match substitute,keep with + | [],[] -> node + | _ -> oname, Leaf (in_module (entry,(substobjs,[],[]))) + else + node + | _ -> anomaly "kill_keep expects Leafs only!" + in + match objs with + | [] -> objs + | h::tl -> let h'=kill h and tl'=kill_keep tl in + if h==h' && tl==tl' then + objs + else + h'::tl' +*) + +let start_module id argids args res_o = + let mp = Global.start_module (Lib.module_dp()) id args res_o in + let mbids = List.map fst args in + let fs = Summary.freeze_summaries () in + process_module_bindings argids args; + openmod_info:=(mbids,res_o); + Lib.start_module id mp fs; + Lib.add_frozen_state () + + +let end_module id = + + let oldoname,oldprefix,fs,lib_stack = Lib.end_module id in + let mp = Global.end_module id in + let substitute, keep, special = Lib.classify_objects lib_stack in + + let dir = fst oldprefix in + let msid = msid_of_prefix oldprefix in + let mbids, res_o = !openmod_info in + + Summary.unfreeze_other_summaries fs; + + let substobjs = match res_o with + | None -> + empty_subst, mbids, msid, substitute + | Some (MTEident ln) -> + abstract_substobjs mbids (KNmap.find ln (!modtypetab)) + | Some (MTEsig (msid,_)) -> + todo "Anonymous signatures not supported"; + empty_subst, mbids, msid, [] + | Some (MTEfunsig _) -> + anomaly "Funsig cannot be here..." + in + let substituted = subst_substobjs dir mp substobjs in + let node = in_module (None,substobjs,substituted) in + let objects = + if keep = [] then + special@[node] + else + special@[node;in_modkeep keep] + in + let newoname = Lib.add_leaves id objects in + + if (fst newoname) <> (fst oldoname) then + anomaly "Names generated on start_ and end_module do not match"; + if mp_of_kn (snd newoname) <> mp then + anomaly "Kernel and Library names do not match"; + + Lib.add_frozen_state () (* to prevent recaching *) + + + +let module_objects mp = + let prefix,objects = MPmap.find mp !modtab_objects in + segment_of_objects prefix objects + + +type library_name = dir_path + +(* The first two will form a substitutive_objects, the last one is keep *) +type library_objects = + mod_self_id * lib_objects * lib_objects + + +(* The library_cache here is needed to avoid recalculations of + substituted modules object during "reloading" of libraries *) +let library_cache = Hashtbl.create 17 + + +let register_library dir cenv objs digest = + let mp = MPfile dir in + let prefix = dir, (mp, empty_dirpath) in + let substobjs, objects = + try + ignore(Global.lookup_module mp); + (* if it's in the environment, the cached objects should be correct *) + Hashtbl.find library_cache dir + with + Not_found -> + if mp <> Global.import cenv digest then + anomaly "Unexpected disk module name"; + let msid,substitute,keep = objs in + let substobjs = empty_subst, [], msid, substitute in + let substituted = subst_substobjs dir mp substobjs in + let objects = option_app (fun seg -> seg@keep) substituted in + let modobjs = substobjs, objects in + Hashtbl.add library_cache dir modobjs; + modobjs + in + do_module false "register_compilation" load_objects 1 dir mp substobjs objects + + +let start_library dir = + let mp = Global.start_library dir in + openmod_info:=[],None; + Lib.start_compilation dir mp; + Lib.add_frozen_state () + + +let export_library dir = + let cenv = Global.export dir in + let prefix, lib_stack = Lib.end_compilation dir in + let msid = msid_of_prefix prefix in + let substitute, keep, _ = Lib.classify_objects lib_stack in + cenv,(msid,substitute,keep) + + + +let import_module mp = + let prefix,objects = MPmap.find mp !modtab_objects in + open_objects 1 prefix objects + + +let start_modtype id argids args = + let mp = Global.start_modtype (Lib.module_dp()) id args in + let fs= Summary.freeze_summaries () in + process_module_bindings argids args; + openmodtype_info := (List.map fst args); + Lib.start_modtype id mp fs; + Lib.add_frozen_state () + + +let end_modtype id = + + let oldoname,prefix,fs,lib_stack = Lib.end_modtype id in + let ln = Global.end_modtype id in + let substitute, _, special = Lib.classify_objects lib_stack in + + let msid = msid_of_prefix prefix in + let mbids = !openmodtype_info in + + Summary.unfreeze_other_summaries fs; + + let modtypeobjs = empty_subst, mbids, msid, substitute in + + let oname = Lib.add_leaves id (special@[in_modtype (None, modtypeobjs)]) in + if fst oname <> fst oldoname then + anomaly + "Section paths generated on start_ and end_modtype do not match"; + if snd oname <> ln then + anomaly + "Kernel and Library names do not match"; + + Lib.add_frozen_state ()(* to prevent recaching *) + + +let declare_modtype id mte = + let substobjs = get_modtype_substobjs mte in + ignore (add_leaf id (in_modtype (Some mte, substobjs))) + + + +let rec get_module_substobjs = function + MEident mp -> MPmap.find mp !modtab_substobjs + | MEfunctor (mbid,_,mexpr) -> + let (subst, mbids, msid, objs) = get_module_substobjs mexpr in + (subst, mbid::mbids, msid, objs) + | MEstruct (msid,_) -> + (empty_subst, [], msid, []) + | MEapply (mexpr, MEident mp) -> + let (subst, mbids, msid, objs) = get_module_substobjs mexpr in + (match mbids with + | mbid::mbids -> + (add_mbid mbid mp subst, mbids, msid, objs) + | [] -> anomaly "Too few arguments in functor") + | MEapply (_,_) -> + anomaly "The argument of a module application must be a path!" + + +let declare_module id me = + let substobjs = + match me with + | {mod_entry_type = Some mte} -> get_modtype_substobjs mte + | {mod_entry_expr = Some mexpr} -> get_module_substobjs mexpr + | _ -> anomaly "declare_module: No type, no body ..." + in + let dir,mp = dir_of_sp (Lib.make_path id), mp_of_kn (Lib.make_kn id) in + let substituted = subst_substobjs dir mp substobjs in + ignore (add_leaf + id + (in_module (Some me, substobjs, substituted))) + + +(*s Iterators. *) + +let fold_all_segments insec f x = + let acc' = + MPmap.fold + (fun _ (prefix,objects) acc -> + let apply_obj acc (id,obj) = f acc (make_oname prefix id) obj in + List.fold_left apply_obj acc objects) + !modtab_objects x + in + let rec apply_node acc = function + | sp, Leaf o -> f acc sp o + | _, ClosedSection (_,_,seg) -> + if insec then List.fold_left apply_node acc seg else acc + | _ -> acc + in + List.fold_left apply_node acc' (Lib.contents_after None) + +let iter_all_segments insec f = + let _ = + MPmap.iter + (fun _ (prefix,objects) -> + let apply_obj (id,obj) = f (make_oname prefix id) obj in + List.iter apply_obj objects) + !modtab_objects + in + let rec apply_node = function + | sp, Leaf o -> f sp o + | _, ClosedSection (_,_,seg) -> if insec then List.iter apply_node seg + | _ -> () + in + List.iter apply_node (Lib.contents_after None) + + + +let debug_print_modtab _ = + let pr_seg = function + | [] -> str "[]" + | l -> str ("[." ^ string_of_int (List.length l) ^ ".]") + in + let pr_modinfo mp (prefix,objects) s = + s ++ str (string_of_mp mp) ++ (spc ()) + ++ (pr_seg (segment_of_objects prefix objects)) + in + let modules = MPmap.fold pr_modinfo !modtab_objects (mt ()) in + hov 0 modules + + diff --git a/library/declaremods.mli b/library/declaremods.mli new file mode 100644 index 000000000..17db827e8 --- /dev/null +++ b/library/declaremods.mli @@ -0,0 +1,85 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* module_entry -> unit + +val start_module : + identifier -> identifier list -> (mod_bound_id * module_type_entry) list + -> module_type_entry option -> unit +val end_module : identifier -> unit + + + +(*s Module types *) + +val declare_modtype : identifier -> module_type_entry -> unit + +val start_modtype : + identifier -> identifier list -> (mod_bound_id * module_type_entry) list + -> unit +val end_modtype : identifier -> unit + + +(*s Objects of a module. They come in two lists: the substitutive ones + and the other *) + +val module_objects : module_path -> library_segment + + +(*s Libraries i.e. modules on disk *) + +type library_name = dir_path + +type library_objects + +val register_library : + library_name -> + Safe_typing.compiled_library -> library_objects -> Digest.t -> unit + +val start_library : library_name -> unit + +val export_library : + library_name -> Safe_typing.compiled_library * library_objects + + +(* [import_module mp] opens the module [mp] (in a Caml sense). + It modifies Nametab and performs the "open_object" function + for every object of the module. *) + +val import_module : module_path -> unit + + +(*s [fold_all_segments] and [iter_all_segments] iterate over all + segments, the modules' segments first and then the current + segment. Modules are presented in an arbitrary order. The given + function is applied to all leaves (together with their section + path). The boolean indicates if we must enter closed sections. *) + +val fold_all_segments : bool -> ('a -> object_name -> obj -> 'a) -> 'a -> 'a +val iter_all_segments : bool -> (object_name -> obj -> unit) -> unit + + +val debug_print_modtab : unit -> Pp.std_ppcmds + +(*val debug_print_modtypetab : unit -> Pp.std_ppcmds*) diff --git a/library/global.ml b/library/global.ml index db0e5f470..cdebc52f5 100644 --- a/library/global.ml +++ b/library/global.ml @@ -26,7 +26,7 @@ let safe_env () = !global_env let env () = env_of_safe_env !global_env let _ = - declare_summary "Global environment" + declare_global_environment { freeze_function = (fun () -> !global_env); unfreeze_function = (fun fr -> global_env := fr); init_function = (fun () -> global_env := empty_environment); @@ -46,17 +46,81 @@ let push_named_def d = global_env := env; cst -let add_constant sp ce = global_env := add_constant sp ce !global_env -let add_mind sp mie = global_env := add_mind sp mie !global_env +(*let add_thing add kn thing = + let _,dir,l = repr_kn kn in + let kn',newenv = add dir l thing !global_env in + if kn = kn' then + global_env := newenv + else + anomaly "Kernel names do not match." +*) + +let add_thing add dir id thing = + let kn, newenv = add dir (label_of_id id) thing !global_env in + global_env := newenv; + kn + +let add_constant = add_thing add_constant +let add_mind = add_thing add_mind +let add_modtype = add_thing (fun _ -> add_modtype) () +let add_module = add_thing (fun _ -> add_module) () + let add_constraints c = global_env := add_constraints c !global_env + + +let start_module dir id params mtyo = + let l = label_of_id id in + let mp,newenv = start_module dir l params mtyo !global_env in + global_env := newenv; + mp + +let end_module id = + let l = label_of_id id in + let mp,newenv = end_module l !global_env in + global_env := newenv; + mp + + +let start_modtype dir id params = + let l = label_of_id id in + let mp,newenv = start_modtype dir l params !global_env in + global_env := newenv; + mp + +let end_modtype id = + let l = label_of_id id in + let kn,newenv = end_modtype l !global_env in + global_env := newenv; + kn + + + + let lookup_named id = lookup_named id (env()) -let lookup_constant sp = lookup_constant sp (env()) +let lookup_constant kn = lookup_constant kn (env()) let lookup_inductive ind = Inductive.lookup_mind_specif (env()) ind -let lookup_mind sp = lookup_mind sp (env()) +let lookup_mind kn = lookup_mind kn (env()) + +let lookup_module mp = lookup_module mp (env()) +let lookup_modtype kn = lookup_modtype kn (env()) + + + + +let start_library dir = + let mp,newenv = start_library dir !global_env in + global_env := newenv; + mp + +let export s = snd (export !global_env s) + +let import cenv digest = + let mp,newenv = import cenv digest !global_env in + global_env := newenv; + mp + -let export s = export !global_env s -let import cenv = global_env := import cenv !global_env (*s Function to get an environment from the constants part of the global environment and a given context. *) @@ -64,7 +128,7 @@ let import cenv = global_env := import cenv !global_env let env_of_context hyps = reset_with_named_context hyps (env()) -open Nametab +open Libnames let type_of_reference env = function | VarRef id -> let (_,_,t) = Environ.lookup_named id env in t @@ -73,3 +137,8 @@ let type_of_reference env = function | ConstructRef cstr -> Inductive.type_of_constructor env cstr let type_of_global t = type_of_reference (env ()) t + + +(*let get_kn dp l = + make_kn (current_modpath !global_env) dp l +*) diff --git a/library/global.mli b/library/global.mli index 48e90d12e..4a215167a 100644 --- a/library/global.mli +++ b/library/global.mli @@ -13,13 +13,18 @@ open Names open Univ open Term open Declarations +open Entries open Indtypes open Safe_typing -(*i*) + (*i*) + +(* This module defines the global environment of Coq. The functions + below are exactly the same as the ones in [Safe_typing], operating on + that global environment. [add_*] functions perform name verification, + i.e. check that the name given as argument match those provided by + [Safe_typing]. *) + -(* This module defines the global environment of Coq. - The functions below are exactly the same as the ones in [Typing], - operating on that global environment. *) val safe_env : unit -> safe_environment val env : unit -> Environ.env @@ -27,26 +32,62 @@ val env : unit -> Environ.env val universes : unit -> universes val named_context : unit -> Sign.named_context -(* Extending env with variables, constants and inductives *) +(*s Extending env with variables and local definitions *) val push_named_assum : (identifier * types) -> Univ.constraints val push_named_def : (identifier * constr * types option) -> Univ.constraints -val add_constant : constant -> global_declaration -> unit -val add_mind : mutual_inductive -> mutual_inductive_entry -> unit +(*s Adding constants, inductives, modules and module types. All these + functions verify that given names match those generated by kernel *) + +val add_constant : + dir_path -> identifier -> global_declaration -> kernel_name +val add_mind : + dir_path -> identifier -> mutual_inductive_entry -> kernel_name + +val add_module : identifier -> module_entry -> module_path +val add_modtype : identifier -> module_type_entry -> kernel_name + val add_constraints : constraints -> unit +(*s Interactive modules and module types *) +(* Both [start_*] functions take the [dir_path] argument to create a + [mod_self_id]. This should be the name of the compilation unit. *) + +(* [start_*] functions return the [module_path] valid for components + of the started module / module type *) + +val start_module : + dir_path -> identifier -> (mod_bound_id * module_type_entry) list + -> module_type_entry option + -> module_path + +val end_module : + identifier -> module_path + +val start_modtype : + dir_path -> identifier -> (mod_bound_id * module_type_entry) list + -> module_path + +val end_modtype : + identifier -> kernel_name + + (* Queries *) val lookup_named : variable -> named_declaration val lookup_constant : constant -> constant_body val lookup_inductive : inductive -> mutual_inductive_body * one_inductive_body val lookup_mind : mutual_inductive -> mutual_inductive_body +val lookup_module : module_path -> module_body +val lookup_modtype : kernel_name -> module_type_body -(* Modules *) -val export : dir_path -> Environ.compiled_env -val import : Environ.compiled_env -> unit +(* Compiled modules *) +val start_library : dir_path -> module_path +val export : dir_path -> compiled_library +val import : compiled_library -> Digest.t -> module_path (*s Function to get an environment from the constants part of the global - environment and a given context. *) - -val type_of_global : Nametab.global_reference -> types + * environment and a given context. *) + +val type_of_global : Libnames.global_reference -> types val env_of_context : Sign.named_context -> Environ.env + diff --git a/library/goptions.ml b/library/goptions.ml index b4056472b..4d505b5aa 100644 --- a/library/goptions.ml +++ b/library/goptions.ml @@ -14,6 +14,7 @@ open Pp open Util open Libobject open Names +open Libnames open Term open Nametab @@ -55,6 +56,7 @@ module MakeTable = type key val table : (string * key table_of_A) list ref val encode : key -> t + val subst : substitution -> t -> t val printer : t -> std_ppcmds val key : option_name val title : string @@ -66,10 +68,10 @@ module MakeTable = | GOadd | GOrmv - let kn = nickname A.key + let nick = nickname A.key let _ = - if List.mem_assoc kn !A.table then + if List.mem_assoc nick !A.table then error "Sorry, this table name is already used" module MyType = struct type t = A.t let compare = Pervasives.compare end @@ -82,7 +84,7 @@ module MakeTable = let freeze () = !t in let unfreeze c = t := c in let init () = t := MySet.empty in - Summary.declare_summary kn + Summary.declare_summary nick { Summary.freeze_function = freeze; Summary.unfreeze_function = unfreeze; Summary.init_function = init; @@ -90,17 +92,25 @@ module MakeTable = let (add_option,remove_option) = if A.synchronous then - let load_options _ = () in let cache_options (_,(f,p)) = match f with | GOadd -> t := MySet.add p !t | GOrmv -> t := MySet.remove p !t in - let export_options fp = Some fp in + let load_options i o = if i=1 then cache_options o in + let subst_options (_,subst,(f,p as obj)) = + let p' = A.subst subst p in + if p' == p then obj else + (f,p') + in + let export_options fp = Some fp in let (inGo,outGo) = - Libobject.declare_object (kn, - { Libobject.load_function = load_options; - Libobject.open_function = cache_options; + Libobject.declare_object {(Libobject.default_object nick) with + Libobject.load_function = load_options; + Libobject.open_function = load_options; Libobject.cache_function = cache_options; - Libobject.export_function = export_options}) in + Libobject.subst_function = subst_options; + Libobject.classify_function = (fun (_,x) -> Substitute x); + Libobject.export_function = export_options} + in ((fun c -> Lib.add_anonymous_leaf (inGo (GOadd, c))), (fun c -> Lib.add_anonymous_leaf (inGo (GOrmv, c)))) else @@ -126,7 +136,7 @@ module MakeTable = method print = print_table A.title A.printer !t end - let _ = A.table := (kn,new table_of_A ())::!A.table + let _ = A.table := (nick,new table_of_A ())::!A.table let active c = MySet.mem c !t let elements () = MySet.elements !t end @@ -149,6 +159,7 @@ struct type key = string let table = string_table let encode x = x + let subst _ x = x let printer = str let key = A.key let title = A.title @@ -167,6 +178,7 @@ module type RefConvertArg = sig type t val encode : qualid located -> t + val subst : substitution -> t -> t val printer : t -> std_ppcmds val key : option_name val title : string @@ -180,6 +192,7 @@ struct type key = qualid located let table = ref_table let encode = A.encode + let subst = A.subst let printer = A.printer let key = A.key let title = A.title @@ -228,11 +241,9 @@ let declare_option cast uncast check_key key; let write = if sync then let (decl_obj,_) = - declare_object ((nickname key), - {load_function = (fun _ -> ()); - open_function = (fun _ -> ()); + declare_object {(default_object (nickname key)) with cache_function = (fun (_,v) -> write v); - export_function = (fun x -> None)}) + classify_function = (fun _ -> Dispose)} in let _ = declare_summary (nickname key) {freeze_function = read; @@ -298,7 +309,7 @@ let msg_option_value (name,v) = | IntValue (Some n) -> int n | IntValue None -> str "undefined" | StringValue s -> str s - | IdentValue r -> pr_global_env (Global.env()) r + | IdentValue r -> pr_global_env None r let print_option_value key = let (name,(_,read,_)) = get_option key in diff --git a/library/goptions.mli b/library/goptions.mli index d43a4283e..28da69ea6 100644 --- a/library/goptions.mli +++ b/library/goptions.mli @@ -56,6 +56,7 @@ open Pp open Util open Names +open Libnames open Term open Nametab (*i*) @@ -107,6 +108,7 @@ module MakeRefTable : (A : sig type t val encode : qualid located -> t + val subst : substitution -> t -> t val printer : t -> std_ppcmds val key : option_name val title : string diff --git a/library/impargs.ml b/library/impargs.ml index af6bfd6b7..14a93123b 100644 --- a/library/impargs.ml +++ b/library/impargs.ml @@ -10,6 +10,7 @@ open Util open Names +open Libnames open Term open Reduction open Declarations @@ -92,31 +93,35 @@ let list_of_implicits = function (*s Constants. *) -let constants_table = ref Spmap.empty +let constants_table = ref KNmap.empty -let compute_constant_implicits sp = +let compute_constant_implicits kn = let env = Global.env () in - let cb = lookup_constant sp env in + let cb = lookup_constant kn env in auto_implicits env (body_of_type cb.const_type) -let cache_constant_implicits (_,(sp,imps)) = - constants_table := Spmap.add sp imps !constants_table +let cache_constant_implicits (_,(kn,imps)) = + constants_table := KNmap.add kn imps !constants_table + +let subst_constant_implicits (_,subst,(kn,imps as obj)) = + let kn' = subst_kn subst kn in + if kn' == kn then obj else + kn',imps let (in_constant_implicits, _) = - let od = { + declare_object {(default_object "CONSTANT-IMPLICITS") with cache_function = cache_constant_implicits; - load_function = cache_constant_implicits; - open_function = (fun _ -> ()); + load_function = (fun _ -> cache_constant_implicits); + subst_function = subst_constant_implicits; + classify_function = (fun (_,x) -> Substitute x); export_function = (fun x -> Some x) } - in - declare_object ("CONSTANT-IMPLICITS", od) -let declare_constant_implicits sp = - let imps = compute_constant_implicits sp in - add_anonymous_leaf (in_constant_implicits (sp,imps)) +let declare_constant_implicits kn = + let imps = compute_constant_implicits kn in + add_anonymous_leaf (in_constant_implicits (kn,imps)) let constant_implicits sp = - try Spmap.find sp !constants_table with Not_found -> No_impl + try KNmap.find sp !constants_table with Not_found -> No_impl let constant_implicits_list sp = list_of_implicits (constant_implicits sp) @@ -151,30 +156,40 @@ let constructors_table = ref Constrmap.empty let cache_inductive_implicits (_,(indp,imps)) = inductives_table := Indmap.add indp imps !inductives_table +let subst_inductive_implicits (_,subst,((kn,i),imps as obj)) = + let kn' = subst_kn subst kn in + if kn' == kn then obj else + (kn',i),imps + let (in_inductive_implicits, _) = - let od = { + declare_object {(default_object "INDUCTIVE-IMPLICITS") with cache_function = cache_inductive_implicits; - load_function = cache_inductive_implicits; - open_function = (fun _ -> ()); - export_function = (fun x -> Some x) } - in - declare_object ("INDUCTIVE-IMPLICITS", od) + load_function = (fun _ -> cache_inductive_implicits); + subst_function = subst_inductive_implicits; + classify_function = (fun (_,x) -> Substitute x); + export_function = (fun x -> Some x) } let cache_constructor_implicits (_,(consp,imps)) = constructors_table := Constrmap.add consp imps !constructors_table +let subst_constructor_implicits (_,subst,(((kn,i),j),imps as obj)) = + let kn' = subst_kn subst kn in + if kn' == kn then obj else + ((kn',i),j),imps + + let (in_constructor_implicits, _) = - let od = { + declare_object {(default_object "CONSTRUCTOR-IMPLICITS") with cache_function = cache_constructor_implicits; - load_function = cache_constructor_implicits; - open_function = (fun _ -> ()); - export_function = (fun x -> Some x) } - in - declare_object ("CONSTRUCTOR-IMPLICITS", od) + load_function = (fun _ -> cache_constructor_implicits); + subst_function = subst_constructor_implicits; + classify_function = (fun (_,x) -> Substitute x); + export_function = (fun x -> Some x) } -let compute_mib_implicits sp = + +let compute_mib_implicits kn = let env = Global.env () in - let mib = lookup_mind sp env in + let mib = lookup_mind kn env in let ar = Array.to_list (Array.map (* No need to lift, arities contain no de Bruijn *) @@ -188,10 +203,10 @@ let compute_mib_implicits sp = in Array.map imps_one_inductive mib.mind_packets -let cache_mib_implicits (_,(sp,mibimps)) = +let cache_mib_implicits (_,(kn,mibimps)) = Array.iteri (fun i (imps,v) -> - let indp = (sp,i) in + let indp = (kn,i) in inductives_table := Indmap.add indp imps !inductives_table; Array.iteri (fun j imps -> @@ -199,18 +214,22 @@ let cache_mib_implicits (_,(sp,mibimps)) = Constrmap.add (indp,succ j) imps !constructors_table) v) mibimps +let subst_mib_implicits (_,subst,(kn,mibimps as obj)) = + let kn' = subst_kn subst kn in + if kn' == kn then obj else + kn',mibimps + let (in_mib_implicits, _) = - let od = { + declare_object {(default_object "MIB-IMPLICITS") with cache_function = cache_mib_implicits; - load_function = cache_mib_implicits; - open_function = (fun _ -> ()); - export_function = (fun x -> Some x) } - in - declare_object ("MIB-IMPLICITS", od) + load_function = (fun _ -> cache_mib_implicits); + subst_function = subst_mib_implicits; + classify_function = (fun (_,x) -> Substitute x); + export_function = (fun x -> Some x) } -let declare_mib_implicits sp = - let imps = compute_mib_implicits sp in - add_anonymous_leaf (in_mib_implicits (sp,imps)) +let declare_mib_implicits kn = + let imps = compute_mib_implicits kn in + add_anonymous_leaf (in_mib_implicits (kn,imps)) let inductive_implicits indp = try Indmap.find indp !inductives_table with Not_found -> No_impl @@ -237,13 +256,12 @@ let cache_var_implicits (_,(id,imps)) = var_table := Idmap.add id imps !var_table let (in_var_implicits, _) = - let od = { + declare_object {(default_object "VARIABLE-IMPLICITS") with cache_function = cache_var_implicits; - load_function = cache_var_implicits; - open_function = (fun _ -> ()); - export_function = (fun x -> Some x) } - in - declare_object ("VARIABLE-IMPLICITS", od) + load_function = (fun _ -> cache_var_implicits); + export_function = (fun x -> Some x) } + + let declare_var_implicits id = let imps = compute_var_implicits id in @@ -255,16 +273,16 @@ let implicits_of_var id = (*s Implicits of a global reference. *) let declare_implicits = function - | VarRef sp -> - declare_var_implicits sp - | ConstRef sp -> - declare_constant_implicits sp - | IndRef ((sp,i) as indp) -> - let mib_imps = compute_mib_implicits sp in + | VarRef id -> + declare_var_implicits id + | ConstRef kn -> + declare_constant_implicits kn + | IndRef ((kn,i) as indp) -> + let mib_imps = compute_mib_implicits kn in let imps = fst mib_imps.(i) in add_anonymous_leaf (in_inductive_implicits (indp, imps)) - | ConstructRef (((sp,i),j) as consp) -> - let mib_imps = compute_mib_implicits sp in + | ConstructRef (((kn,i),j) as consp) -> + let mib_imps = compute_mib_implicits kn in let imps = (snd mib_imps.(i)).(j-1) in add_anonymous_leaf (in_constructor_implicits (consp, imps)) @@ -296,7 +314,7 @@ let declare_manual_implicits r l = (*s Tests if declared implicit *) let is_implicit_constant sp = - try let _ = Spmap.find sp !constants_table in true with Not_found -> false + try let _ = KNmap.find sp !constants_table in true with Not_found -> false let is_implicit_inductive_definition indp = try let _ = Indmap.find indp !inductives_table in true @@ -313,13 +331,13 @@ let implicits_of_global = function (*s Registration as global tables and rollback. *) -type frozen_t = implicits Spmap.t +type frozen_t = implicits KNmap.t * implicits Indmap.t * implicits Constrmap.t * implicits Idmap.t let init () = - constants_table := Spmap.empty; + constants_table := KNmap.empty; inductives_table := Indmap.empty; constructors_table := Constrmap.empty; var_table := Idmap.empty diff --git a/library/impargs.mli b/library/impargs.mli index 022a38230..751d96844 100644 --- a/library/impargs.mli +++ b/library/impargs.mli @@ -10,6 +10,7 @@ (*i*) open Names +open Libnames open Term open Environ open Nametab diff --git a/library/lib.ml b/library/lib.ml index 8eaba772e..286133305 100644 --- a/library/lib.ml +++ b/library/lib.ml @@ -11,23 +11,102 @@ open Pp open Util open Names +open Libnames open Nameops open Libobject open Summary + type node = | Leaf of obj - | Module of dir_path - | OpenedSection of dir_path * Summary.frozen + | CompilingModule of object_prefix + | OpenedModule of object_prefix * Summary.frozen + | OpenedModtype of object_prefix * Summary.frozen + | OpenedSection of object_prefix * Summary.frozen (* bool is to tell if the section must be opened automatically *) | ClosedSection of bool * dir_path * library_segment | FrozenState of Summary.frozen -and library_entry = section_path * node +and library_entry = object_name * node and library_segment = library_entry list +type lib_objects = (identifier * obj) list + +let rec iter_leaves f i seg = + match seg with + | [] -> () + | (oname ,Leaf obj)::segtl -> + f i (oname,obj); + iter_leaves f i segtl + | _ -> anomaly "We should have leaves only here" + + +let open_segment = iter_leaves open_object + +let load_segment = iter_leaves load_object + +let change_kns mp seg = + let subst_one = function + | ((sp,kn),(Leaf obj as lobj)) -> + let kn' = make_kn mp empty_dirpath (label kn) in + ((sp,kn'),lobj) + | _ -> anomaly "We should have leaves only here" + in + List.map subst_one seg + +let subst_segment (dirpath,(mp,dir)) subst seg = + let subst_one = function + | ((sp,kn),Leaf obj) -> + let sp' = make_path dirpath (basename sp) in + let kn' = make_kn mp dir (label kn) in + let oname' = sp',kn' in + let obj' = subst_object (oname',subst,obj) in + (oname', Leaf obj') + | _ -> anomaly "We should have leaves only here" + in + List.map subst_one seg + + +let iter_objects f i prefix = + List.iter (fun (id,obj) -> f i (make_oname prefix id, obj)) + +let load_objects = iter_objects load_object +let open_objects = iter_objects open_object + +let subst_objects prefix subst seg = + let subst_one = fun (id,obj as node) -> + let obj' = subst_object (make_oname prefix id, subst, obj) in + if obj' == obj then node else + (id, obj') + in + list_smartmap subst_one seg + +let classify_objects seg = + let rec clean ((substl,keepl,anticipl) as acc) = function + | (_,CompilingModule _) :: _ | [] -> acc + | ((sp,kn as oname),Leaf o) as node :: stk -> + let id = id_of_label (label kn) in + (match classify_object (oname,o) with + | Dispose -> clean acc stk + | Keep o' -> + clean (substl, (id,o')::keepl, anticipl) stk + | Substitute o' -> + clean ((id,o')::substl, keepl, anticipl) stk + | Anticipate o' -> + clean (substl, keepl, o'::anticipl) stk) + | (oname,ClosedSection _ as item) :: stk -> clean acc stk + | (_,OpenedSection _) :: _ -> error "there are still opened sections" + | (_,OpenedModule _) :: _ -> error "there are still opened modules" + | (_,OpenedModtype _) :: _ -> error "there are still opened module types" + | (_,FrozenState _) :: stk -> clean acc stk + in + clean ([],[],[]) (List.rev seg) + + +let segment_of_objects prefix = + List.map (fun (id,obj) -> (make_oname prefix id, Leaf obj)) (* We keep trace of operations in the stack [lib_stk]. [path_prefix] is the current path of sections, where sections are stored in @@ -36,31 +115,44 @@ and library_segment = library_entry list sections, but on the contrary there are many constructions of section paths based on the library path. *) +let initial_prefix = default_module,(initial_path,empty_dirpath) + let lib_stk = ref ([] : library_segment) -let module_name = ref None -let path_prefix = ref (default_module : dir_path) +let comp_name = ref None +let path_prefix = ref initial_prefix -let module_sp () = - match !module_name with Some m -> m | None -> default_module +let module_dp () = + match !comp_name with Some m -> m | None -> default_module let recalc_path_prefix () = let rec recalc = function | (sp, OpenedSection (dir,_)) :: _ -> dir + | (sp, OpenedModule (dir,_)) :: _ -> dir + | (sp, OpenedModtype (dir,_)) :: _ -> dir + | (sp, CompilingModule dir) :: _ -> dir | _::l -> recalc l - | [] -> module_sp () + | [] -> initial_prefix in path_prefix := recalc !lib_stk -let pop_path_prefix () = path_prefix := fst (split_dirpath !path_prefix) +let pop_path_prefix () = + let dir,(mp,sec) = !path_prefix in + path_prefix := fst (split_dirpath dir), (mp, fst (split_dirpath sec)) + +let make_path id = Libnames.make_path (fst !path_prefix) id -let make_path id = Names.make_path !path_prefix id +let make_kn id = + let mp,dir = snd !path_prefix in + Names.make_kn mp dir (label_of_id id) + +let make_oname id = make_path id, make_kn id let sections_depth () = - List.length (repr_dirpath !path_prefix) - - List.length (repr_dirpath (module_sp ())) + List.length (repr_dirpath (snd (snd !path_prefix))) -let cwd () = !path_prefix +let cwd () = fst !path_prefix +let current_prefix () = snd !path_prefix let find_entry_p p = let rec find = function @@ -70,12 +162,17 @@ let find_entry_p p = find !lib_stk let split_lib sp = - let rec findrec after = function - | ((sp',obj) as hd)::before -> - if sp = sp' then (after,hd,before) else findrec (hd::after) before + let rec collect after equal = function + | ((sp',_) as hd)::before -> + if sp = sp' then collect after (hd::equal) before else after,equal,hd::before + | [] -> after,equal,[] + in + let rec findeq after = function + | ((sp',_) as hd)::before -> + if sp = sp' then collect after [hd] before else findeq (hd::after) before | [] -> error "no such entry" in - findrec [] !lib_stk + findeq [] !lib_stk (* Adding operations. *) @@ -87,119 +184,267 @@ let anonymous_id = fun () -> incr n; id_of_string ("_" ^ (string_of_int !n)) let add_anonymous_entry node = - let sp = make_path (anonymous_id()) in - add_entry sp node; - sp + let id = anonymous_id () in + let name = make_oname id in + add_entry name node; + name let add_absolutely_named_leaf sp obj = cache_object (sp,obj); add_entry sp (Leaf obj) let add_leaf id obj = - let sp = make_path id in - cache_object (sp,obj); - add_entry sp (Leaf obj); - sp + let oname = make_oname id in + cache_object (oname,obj); + add_entry oname (Leaf obj); + oname + +let add_leaves id objs = + let oname = make_oname id in + let add_obj obj = + add_entry oname (Leaf obj); + load_object 1 (oname,obj) + in + List.iter add_obj objs; + oname let add_anonymous_leaf obj = - let sp = make_path (anonymous_id()) in - cache_object (sp,obj); - add_entry sp (Leaf obj) + let id = anonymous_id () in + let oname = make_oname id in + cache_object (oname,obj); + add_entry oname (Leaf obj) let add_frozen_state () = let _ = add_anonymous_entry (FrozenState (freeze_summaries())) in () +(* Modules. *) + +let is_something_opened = function + (_,OpenedSection _) -> true + | (_,OpenedModule _) -> true + | (_,OpenedModtype _) -> true + | _ -> false + +let classify_segment seg = + let rec clean ((substl,keepl,anticipl) as acc) = function + | (_,CompilingModule _) :: _ | [] -> acc + | (oname,Leaf o) as node :: stk -> + (match classify_object (oname,o) with + | Dispose -> clean acc stk + | Keep o' -> + clean (substl, (oname,Leaf o')::keepl, anticipl) stk + | Substitute o' -> + clean ((oname,Leaf o')::substl, keepl, anticipl) stk + | Anticipate o' -> + clean (substl, keepl, (oname,Leaf o')::anticipl) stk) + | (oname,ClosedSection _ as item) :: stk -> clean acc stk + | (_,OpenedSection _) :: _ -> error "there are still opened sections" + | (_,OpenedModule _) :: _ -> error "there are still opened modules" + | (_,OpenedModtype _) :: _ -> error "there are still opened module types" + | (_,FrozenState _) :: stk -> clean acc stk + in + clean ([],[],[]) (List.rev seg) + +let export_segment seg = + let rec clean acc = function + | (_,CompilingModule _) :: _ | [] -> acc + | (oname,Leaf o) as node :: stk -> + (match export_object o with + | None -> clean acc stk + | Some o' -> clean ((oname,Leaf o') :: acc) stk) + | (oname,ClosedSection _ as item) :: stk -> clean (item :: acc) stk + | (_,OpenedSection _) :: _ -> error "there are still opened sections" + | (_,OpenedModule _) :: _ -> error "there are still opened modules" + | (_,OpenedModtype _) :: _ -> error "there are still opened module types" + | (_,FrozenState _) :: stk -> clean acc stk + in + clean [] seg + + +let start_module id mp nametab = + let dir = extend_dirpath (fst !path_prefix) id in + let prefix = dir,(mp,empty_dirpath) in + let oname = make_path id, make_kn id in + if Nametab.exists_module dir then + errorlabstrm "open_module" (pr_id id ++ str " already exists") ; + add_entry oname (OpenedModule (prefix,nametab)); + path_prefix := prefix +(* add_frozen_state () must be called in declaremods *) + +let end_module id = + let oname,nametab = + try match find_entry_p is_something_opened with + | oname,OpenedModule (_,nametab) -> + let sp = fst oname in + let id' = basename sp in + if id<>id' then error "this is not the last opened module"; + oname,nametab + | _,OpenedModtype _ -> + error "there are some open module types" + | _,OpenedSection _ -> + error "there are some open sections" + | _ -> assert false + with Not_found -> + error "no opened modules" + in + let (after,_,before) = split_lib oname in + lib_stk := before; + let prefix = !path_prefix in + recalc_path_prefix (); + (* add_frozen_state must be called after processing the module, + because we cannot recache interactive modules *) + (oname, prefix, nametab,after) + +let start_modtype id mp nametab = + let dir = extend_dirpath (fst !path_prefix) id in + let prefix = dir,(mp,empty_dirpath) in + let sp = make_path id in + let name = sp, make_kn id in + if Nametab.exists_cci sp then + errorlabstrm "open_modtype" (pr_id id ++ str " already exists") ; + add_entry name (OpenedModtype (prefix,nametab)); + path_prefix := prefix + +let end_modtype id = + let sp,nametab = + try match find_entry_p is_something_opened with + | sp,OpenedModtype (_,nametab) -> + let id' = basename (fst sp) in + if id<>id' then error "this is not the last opened module"; + sp,nametab + | _,OpenedModule _ -> + error "there are some open modules" + | _,OpenedSection _ -> + error "there are some open sections" + | _ -> assert false + with Not_found -> + error "no opened module types" + in + let (after,_,before) = split_lib sp in + lib_stk := before; + let dir = !path_prefix in + recalc_path_prefix (); + (* add_frozen_state must be called after processing the module type. + This is because we cannot recache interactive module types *) + (sp,dir,nametab,after) + + + let contents_after = function | None -> !lib_stk | Some sp -> let (after,_,_) = split_lib sp in after (* Modules. *) -let check_for_module () = +let check_for_comp_unit () = let is_decl = function (_,FrozenState _) -> false | _ -> true in try let _ = find_entry_p is_decl in - error "a module can not be started after some declarations" + error "a module cannot be started after some declarations" with Not_found -> () (* TODO: use check_for_module ? *) -let start_module s = - if !module_name <> None then - error "a module is already started"; - if !path_prefix <> default_module then +let start_compilation s mp = + if !comp_name <> None then + error "compilation unit is already started"; + if snd (snd (!path_prefix)) <> empty_dirpath then error "some sections are already opened"; - module_name := Some s; - let _ = add_anonymous_entry (Module s) in - path_prefix := s - -let end_module s = - match !module_name with - | None -> error "no module declared" - | Some m -> - let (_,bm) = split_dirpath m in - if bm <> s then - error ("The current open module has basename "^(string_of_id bm)); - m + let prefix = s, (mp, empty_dirpath) in + let _ = add_anonymous_entry (CompilingModule prefix) in + comp_name := Some s; + path_prefix := prefix + +let end_compilation dir = + let _ = + try match find_entry_p is_something_opened with + | _, OpenedSection _ -> error "There are some open sections" + | _, OpenedModule _ -> error "There are some open modules" + | _, OpenedModtype _ -> error "There are some open module types" + | _ -> assert false + with + Not_found -> () + in + let module_p = + function (_,CompilingModule _) -> true | x -> is_something_opened x + in + let oname = + try match find_entry_p module_p with + (oname, CompilingModule prefix) -> oname + | _ -> assert false + with + Not_found -> anomaly "No module declared" + in + let _ = match !comp_name with + | None -> anomaly "There should be a module name..." + | Some m -> + if m <> dir then anomaly + ("The current open module has name "^ (string_of_dirpath m) ^ + " and not " ^ (string_of_dirpath m)); + in + let (after,_,before) = split_lib oname in + !path_prefix,after + +(* Returns true if we are inside an opened module type *) +let is_specification () = + let opened_p = function + | _, OpenedModtype _ -> true + | _ -> false + in + try + let _ = find_entry_p opened_p in true + with + Not_found -> false + +(* Returns the most recent OpenedThing node *) +let what_is_opened () = find_entry_p is_something_opened (* Sections. *) let open_section id = - let dir = extend_dirpath !path_prefix id in - let sp = make_path id in + let olddir,(mp,oldsec) = !path_prefix in + let dir = extend_dirpath olddir id in + let prefix = dir, (mp, extend_dirpath oldsec id) in + let name = make_path id, make_kn id (* this makes little sense however *) in if Nametab.exists_section dir then errorlabstrm "open_section" (pr_id id ++ str " already exists"); let sum = freeze_summaries() in - add_entry sp (OpenedSection (dir, sum)); + add_entry name (OpenedSection (prefix, sum)); (*Pushed for the lifetime of the section: removed by unfrozing the summary*) - Nametab.push_section dir; - path_prefix := dir; - sp - -let is_opened_section = function (_,OpenedSection _) -> true | _ -> false + Nametab.push_dir (Nametab.Until 1) dir (DirOpenSection prefix); + path_prefix := prefix; + prefix let sections_are_opened () = - try let _ = find_entry_p is_opened_section in true + try + match find_entry_p is_something_opened with + | (_,OpenedSection _) -> true + | (_,OpenedModule _) -> false + | _ -> false with Not_found -> false -let export_segment seg = - let rec clean acc = function - | (_,Module _) :: _ | [] -> acc - | (sp,Leaf o) as node :: stk -> - (match export_object o with - | None -> clean acc stk - | Some o' -> clean ((sp,Leaf o') :: acc) stk) - | (sp,ClosedSection _ as item) :: stk -> clean (item :: acc) stk - | (_,OpenedSection _) :: _ -> error "there are still opened sections" - | (_,FrozenState _) :: stk -> clean acc stk - in - clean [] seg - (* Restore lib_stk and summaries as before the section opening, and add a ClosedSection object. *) let close_section ~export id = - let sp,dir,fs = - try match find_entry_p is_opened_section with - | sp,OpenedSection (dir,fs) -> - if id<>snd(split_dirpath dir) then + let oname,fs = + try match find_entry_p is_something_opened with + | oname,OpenedSection (_,fs) -> + if id <> basename (fst oname) then error "this is not the last opened section"; - (sp,dir,fs) - | _ -> assert false + (oname,fs) + | _ -> assert false with Not_found -> error "no opened section" in - let (after,_,before) = split_lib sp in + let (after,_,before) = split_lib oname in lib_stk := before; + let prefix = !path_prefix in pop_path_prefix (); - let closed_sec = ClosedSection (export, dir, export_segment after) in - add_entry (make_path id) closed_sec; - (dir,after,fs) - -(* The following function exports the whole library segment, that will be - saved as a module. Objects are presented in chronological order, and - frozen states are removed. *) - -let export_module s = - export_segment !lib_stk + let closed_sec = + ClosedSection (export, (fst prefix), export_segment after) + in + let name = make_path id, make_kn id in + add_entry name closed_sec; + (prefix, after, fs) (* Backtracking. *) @@ -207,6 +452,7 @@ let recache_decl = function | (sp, Leaf o) -> cache_object (sp,o) | _ -> () + let recache_context ctx = List.iter recache_decl ctx @@ -226,21 +472,16 @@ let reset_to sp = let reset_name id = let (sp,_) = try - find_entry_p (fun (sp,_) -> let (_,spi) = repr_path sp in id = spi) + find_entry_p (fun (sp,_) -> let (_,spi) = repr_path (fst sp) in id = spi) with Not_found -> error (string_of_id id ^ ": no such entry") in reset_to sp let point_obj = - let (f,_) = - declare_object - ("DOT", - {cache_function = (fun _ -> ()); - load_function = (fun _ -> ()); - open_function = (fun _ -> ()); - export_function = (fun _ -> None) }) in - f() + let (f,_) = declare_object {(default_object "DOT") with + classify_function = (fun _ -> Dispose)} in + f() let mark_end_of_command () = match !lib_stk with @@ -258,41 +499,41 @@ let back n = reset_to (back_stk n !lib_stk) (* [dir] is a section dir if [module] < [dir] <= [path_prefix] *) let is_section_p sp = - not (is_dirpath_prefix_of sp (module_sp ())) - & (is_dirpath_prefix_of sp !path_prefix) + not (is_dirpath_prefix_of sp (module_dp ())) + & (is_dirpath_prefix_of sp (fst !path_prefix)) (* State and initialization. *) type frozen = dir_path option * library_segment -let freeze () = (!module_name, !lib_stk) +let freeze () = (!comp_name, !lib_stk) let unfreeze (mn,stk) = - module_name := mn; + comp_name := mn; lib_stk := stk; recalc_path_prefix () let init () = lib_stk := []; add_frozen_state (); - module_name := None; - path_prefix := make_dirpath []; + comp_name := None; + path_prefix := initial_prefix; init_summaries() (* Initial state. *) -let initial_state = ref (None : section_path option) +let initial_state = ref None let declare_initial_state () = - let sp = add_anonymous_entry (FrozenState (freeze_summaries())) in - initial_state := Some sp + let name = add_anonymous_entry (FrozenState (freeze_summaries())) in + initial_state := Some name let reset_initial () = match !initial_state with | None -> init () | Some sp -> begin match split_lib sp with - | (_,(_,FrozenState fs as hd),before) -> + | (_,[_,FrozenState fs as hd],before) -> lib_stk := hd::before; recalc_path_prefix (); unfreeze_summaries fs diff --git a/library/lib.mli b/library/lib.mli index b86c08f72..90f111c4b 100644 --- a/library/lib.mli +++ b/library/lib.mli @@ -10,6 +10,7 @@ (*i*) open Names +open Libnames open Libobject open Summary (*i*) @@ -20,53 +21,127 @@ open Summary type node = | Leaf of obj - | Module of dir_path - | OpenedSection of dir_path * Summary.frozen + | CompilingModule of object_prefix + | OpenedModule of object_prefix * Summary.frozen + | OpenedModtype of object_prefix * Summary.frozen + | OpenedSection of object_prefix * Summary.frozen | ClosedSection of bool * dir_path * library_segment | FrozenState of Summary.frozen -and library_entry = section_path * node +and library_entry = object_name * node and library_segment = library_entry list +type lib_objects = (identifier * obj) list -(*s Adding operations (which calls the [cache] method, and getting the +(*s These functions iterate (or map) object functions. + + [subst_segment prefix subst seg] makes an assumption that all + objects in [seg] have the same prefix. This prefix is universally + changed to [prefix]. + + [classify_segment seg] divides segment into objects which responded + with [Substitute], [Keep], [Anticipate] respectively, to the + [classify_object] function. The order of each returned list is the + same as in the input list. + + [change_kns mp lib] only changes the prefix of the [kernel_name] + part of the [object_name] of every object to [(mp,empty_dirpath)]. + The [section_path] part of the [object_name] and the object itself + are unchanged. +*) + + +val open_segment : int -> library_segment -> unit +val load_segment : int -> library_segment -> unit +val subst_segment : + object_prefix -> substitution -> library_segment -> library_segment +val classify_segment : + library_segment -> library_segment * library_segment * library_segment +val change_kns : module_path -> library_segment -> library_segment + + + +val open_objects : int -> object_prefix -> lib_objects -> unit +val load_objects : int -> object_prefix -> lib_objects -> unit +val subst_objects : + object_prefix -> substitution -> lib_objects -> lib_objects +val classify_objects : + library_segment -> lib_objects * lib_objects * obj list + +val segment_of_objects : + object_prefix -> lib_objects -> library_segment + +(*s Adding operations (which call the [cache] method, and getting the current list of operations (most recent ones coming first). *) -val add_leaf : identifier -> obj -> section_path -val add_absolutely_named_leaf : section_path -> obj -> unit +val add_leaf : identifier -> obj -> object_name +val add_absolutely_named_leaf : object_name -> obj -> unit val add_anonymous_leaf : obj -> unit + +(* this operation adds all objects with the same name and calls load_object + for each of them *) +val add_leaves : identifier -> obj list -> object_name + val add_frozen_state : unit -> unit val mark_end_of_command : unit -> unit (*s The function [contents_after] returns the current library segment, - starting from a given section path. If not given, the entire segment - is returned. *) + starting from a given section path. If not given, the entire segment + is returned. *) -val contents_after : section_path option -> library_segment +val contents_after : object_name option -> library_segment +(* Returns true if we are inside an opened module type *) +val is_specification : unit -> bool + +(* Returns the most recent OpenedThing node *) +val what_is_opened : unit -> library_entry (*s Opening and closing a section. *) -val open_section : identifier -> section_path -val close_section : - export:bool -> identifier -> dir_path * library_segment * Summary.frozen +val open_section : identifier -> object_prefix + +val close_section : export:bool -> identifier -> + object_prefix * library_segment * Summary.frozen + val sections_are_opened : unit -> bool val make_path : identifier -> section_path +val make_kn : identifier -> kernel_name + val cwd : unit -> dir_path val sections_depth : unit -> int val is_section_p : dir_path -> bool -val start_module : dir_path -> unit -val end_module : module_ident -> dir_path -val export_module : dir_path -> library_segment +(*s Compilation units *) + +val start_compilation : dir_path -> module_path -> unit +val end_compilation : dir_path -> object_prefix * library_segment + +val module_dp : unit -> dir_path +(*s Modules and module types *) + +val start_module : module_ident -> module_path -> Summary.frozen -> unit +val end_module : module_ident + -> object_name * object_prefix * Summary.frozen * library_segment + +val start_modtype : module_ident -> module_path -> Summary.frozen -> unit +val end_modtype : module_ident + -> object_name * object_prefix * Summary.frozen * library_segment + +(* Lib.add_frozen_state must be called after all of the above functions *) + +val current_prefix : unit -> module_path * dir_path (*s Backtracking (undo). *) -val reset_to : section_path -> unit +val reset_to : object_name -> unit val reset_name : identifier -> unit + +(* [back n] resets to the place corresponding to the $n$-th call of + [mark_end_of_command] (counting backwards) *) val back : int -> unit (*s We can get and set the state of the operations (used in [States]). *) diff --git a/library/libnames.ml b/library/libnames.ml new file mode 100644 index 000000000..d5e9c8dc7 --- /dev/null +++ b/library/libnames.ml @@ -0,0 +1,212 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* ref + | ConstRef kn -> + let kn' = subst_kn subst kn in if kn==kn' then ref else + ConstRef kn' + | IndRef (kn,i) -> + let kn' = subst_kn subst kn in if kn==kn' then ref else + IndRef(kn',i) + | ConstructRef ((kn,i),j) -> + let kn' = subst_kn subst kn in if kn==kn' then ref else + ConstructRef ((kn',i),j) + + +(**********************************************) + +let pr_dirpath sl = (str (string_of_dirpath sl)) + +(*s Operations on dirpaths *) + +(* Pop the last n module idents *) +let extract_dirpath_prefix n dir = + let (_,dir') = list_chop n (repr_dirpath dir) in + make_dirpath dir' + +let dirpath_prefix p = match repr_dirpath p with + | [] -> anomaly "dirpath_prefix: empty dirpath" + | _::l -> make_dirpath l + +let is_dirpath_prefix_of d1 d2 = + list_prefix_of (List.rev (repr_dirpath d1)) (List.rev (repr_dirpath d2)) + +(* To know how qualified a name should be to be understood in the current env*) +let add_dirpath_prefix id d = make_dirpath (repr_dirpath d @ [id]) + +let split_dirpath d = + let l = repr_dirpath d in (make_dirpath (List.tl l), List.hd l) + +let extend_dirpath p id = make_dirpath (id :: repr_dirpath p) + + +(* +let path_of_constructor env ((sp,tyi),ind) = + let mib = Environ.lookup_mind sp env in + let mip = mib.mind_packets.(tyi) in + let (pa,_) = repr_path sp in + Names.make_path pa (mip.mind_consnames.(ind-1)) + +let path_of_inductive env (sp,tyi) = + if tyi = 0 then sp + else + let mib = Environ.lookup_mind sp env in + let mip = mib.mind_packets.(tyi) in + let (pa,_) = repr_path sp in + Names.make_path pa (mip.mind_typename) +*) +(* parsing *) +let parse_dir s = + let len = String.length s in + let rec decoupe_dirs dirs n = + if n>=len then dirs else + let pos = + try + String.index_from s n '.' + with Not_found -> len + in + let dir = String.sub s n (pos-n) in + decoupe_dirs ((id_of_string dir)::dirs) (pos+1) + in + decoupe_dirs [] 0 + +let dirpath_of_string s = + match parse_dir s with + [] -> invalid_arg "dirpath_of_string" + | dir -> make_dirpath dir + + +(*s Section paths are absolute names *) + +type section_path = { + dirpath : dir_path ; + basename : identifier } + +let make_path pa id = { dirpath = pa; basename = id } +let repr_path { dirpath = pa; basename = id } = (pa,id) + +(* parsing and printing of section paths *) +let string_of_path sp = + let (sl,id) = repr_path sp in + if repr_dirpath sl = [] then string_of_id id + else (string_of_dirpath sl) ^ "." ^ (string_of_id id) + +let sp_ord sp1 sp2 = + let (p1,id1) = repr_path sp1 + and (p2,id2) = repr_path sp2 in + let p_bit = compare p1 p2 in + if p_bit = 0 then id_ord id1 id2 else p_bit + +module SpOrdered = + struct + type t = section_path + let compare = sp_ord + end + +module Spset = Set.Make(SpOrdered) +module Sppred = Predicate.Make(SpOrdered) +module Spmap = Map.Make(SpOrdered) + +let dirpath sp = let (p,_) = repr_path sp in p +let basename sp = let (_,id) = repr_path sp in id + +let path_of_string s = + try + let dir, id = split_dirpath (dirpath_of_string s) in + make_path dir id + with + | Invalid_argument _ -> invalid_arg "path_of_string" + +let pr_sp sp = str (string_of_path sp) + +let restrict_path n sp = + let dir, s = repr_path sp in + let (dir',_) = list_chop n (repr_dirpath dir) in + make_path (make_dirpath dir') s + +type extended_global_reference = + | TrueGlobal of global_reference + | SyntacticDef of kernel_name + +let subst_ext subst glref = match glref with + | TrueGlobal ref -> + let ref' = subst_global subst ref in + if ref' == ref then glref else + TrueGlobal ref' + | SyntacticDef kn -> + let kn' = subst_kn subst kn in + if kn' == kn then glref else + SyntacticDef kn' + +let encode_kn dir id = make_kn (MPfile dir) empty_dirpath (label_of_id id) + +let decode_kn kn = + let mp,sec_dir,l = repr_kn kn in + match mp,(repr_dirpath sec_dir) with + MPfile dir,[] -> (dir,id_of_label l) + | _ , [] -> anomaly "MPfile expected!" + | _ -> anomaly "Section part should be empty!" + +(*s qualified names *) +type qualid = section_path + +let make_qualid = make_path +let repr_qualid = repr_path + +let string_of_qualid = string_of_path +let pr_qualid = pr_sp + +let qualid_of_string = path_of_string + +let qualid_of_sp sp = sp +let make_short_qualid id = make_qualid empty_dirpath id +let qualid_of_dirpath dir = + let (l,a) = split_dirpath dir in + make_qualid l a + +type object_name = section_path * kernel_name + +type object_prefix = dir_path * (module_path * dir_path) + +let make_oname (dirpath,(mp,dir)) id = + make_path dirpath id, make_kn mp dir (label_of_id id) + +(* to this type are mapped dir_path's in the nametab *) +type global_dir_reference = + | DirOpenModule of object_prefix + | DirOpenModtype of object_prefix + | DirOpenSection of object_prefix + | DirModule of object_prefix + | DirClosedSection of dir_path + (* this won't last long I hope! *) + +(* | ModRef mp -> + let mp' = subst_modpath subst mp in if mp==mp' then ref else + ModRef mp' + | ModTypeRef kn -> + let kn' = subst_kernel_name subst kn in if kn==kn' then ref else + ModTypeRef kn' +*) + +type strength = + | NotDeclare + | DischargeAt of dir_path * int + | NeverDischarge diff --git a/library/libnames.mli b/library/libnames.mli new file mode 100644 index 000000000..915cdf3d2 --- /dev/null +++ b/library/libnames.mli @@ -0,0 +1,118 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* global_reference -> global_reference + +(* dirpaths *) +val pr_dirpath : dir_path -> Pp.std_ppcmds + +val dirpath_of_string : string -> dir_path + +(* Give the immediate prefix of a [dir_path] *) +val dirpath_prefix : dir_path -> dir_path + +(* Give the immediate prefix and basename of a [dir_path] *) +val split_dirpath : dir_path -> dir_path * identifier + +val extend_dirpath : dir_path -> module_ident -> dir_path +val add_dirpath_prefix : module_ident -> dir_path -> dir_path + +val extract_dirpath_prefix : int -> dir_path -> dir_path +val is_dirpath_prefix_of : dir_path -> dir_path -> bool + +(*s Section paths are {\em absolute} names *) +type section_path + +(* Constructors of [section_path] *) +val make_path : dir_path -> identifier -> section_path + +(* Destructors of [section_path] *) +val repr_path : section_path -> dir_path * identifier +val dirpath : section_path -> dir_path +val basename : section_path -> identifier + +(* Parsing and printing of section path as ["coq_root.module.id"] *) +val path_of_string : string -> section_path +val string_of_path : section_path -> string +val pr_sp : section_path -> std_ppcmds + +module Sppred : Predicate.S with type elt = section_path +module Spmap : Map.S with type key = section_path + +val restrict_path : int -> section_path -> section_path + +type extended_global_reference = + | TrueGlobal of global_reference + | SyntacticDef of kernel_name + +val subst_ext : + substitution -> extended_global_reference -> extended_global_reference + +(*s Temporary function to brutally form kernel names from section paths *) + +val encode_kn : dir_path -> identifier -> kernel_name +val decode_kn : kernel_name -> dir_path * identifier + + +(*s A [qualid] is a partially qualified ident; it includes fully + qualified names (= absolute names) and all intermediate partial + qualifications of absolute names, including single identifiers *) +type qualid + +val make_qualid : dir_path -> identifier -> qualid +val repr_qualid : qualid -> dir_path * identifier + +val string_of_qualid : qualid -> string +val pr_qualid : qualid -> std_ppcmds + +val qualid_of_string : string -> qualid + +(* Turns an absolute name into a qualified name denoting the same name *) +val qualid_of_sp : section_path -> qualid + +val qualid_of_dirpath : dir_path -> qualid + +val make_short_qualid : identifier -> qualid + +(* Both names are passed to objects: a "semantic" kernel_name, which + can be substituted and a "syntactic" section_path which can be printed +*) + +type object_name = section_path * kernel_name + +type object_prefix = dir_path * (module_path * dir_path) + +val make_oname : object_prefix -> identifier -> object_name + +(* to this type are mapped dir_path's in the nametab *) +type global_dir_reference = + | DirOpenModule of object_prefix + | DirOpenModtype of object_prefix + | DirOpenSection of object_prefix + | DirModule of object_prefix + | DirClosedSection of dir_path + (* this won't last long I hope! *) + +type strength = + | NotDeclare + | DischargeAt of dir_path * int + | NeverDischarge diff --git a/library/libobject.ml b/library/libobject.ml index bbd8615be..2049e8e04 100644 --- a/library/libobject.ml +++ b/library/libobject.ml @@ -10,6 +10,7 @@ open Util open Names +open Libnames (* The relax flag is used to make it possible to load files while ignoring failures to incorporate some objects. This can be useful when one @@ -23,18 +24,53 @@ let relax_flag = ref false;; let relax b = relax_flag := b;; +type 'a substitutivity = + Dispose | Substitute of 'a | Keep of 'a | Anticipate of 'a + + type 'a object_declaration = { - cache_function : section_path * 'a -> unit; - load_function : section_path * 'a -> unit; - open_function : section_path * 'a -> unit; + object_name : string; + cache_function : object_name * 'a -> unit; + load_function : int -> object_name * 'a -> unit; + open_function : int -> object_name * 'a -> unit; + classify_function : object_name * 'a -> 'a substitutivity; + subst_function : object_name * substitution * 'a -> 'a; export_function : 'a -> 'a option } +let yell s = anomaly s + +let default_object s = { + object_name = s; + cache_function = (fun _ -> ()); + load_function = (fun _ _ -> ()); + open_function = (fun _ _ -> ()); + subst_function = (fun _ -> + yell ("The object "^s^" does not know how to substitute!")); + classify_function = (fun (_,obj) -> Keep obj); + export_function = (fun _ -> None)} + + +(* The suggested object declaration is the following: + + declare_object { (default_object "MY OBJECT") with + cache_function = fun (sp,a) -> Mytbl.add sp a} + + and the listed functions are only those which definitions accually + differ from the default. + + This helps introducing new functions in objects. +*) + +let ident_subst_function (_,_,a) = a + type obj = Dyn.t (* persistent dynamic objects *) type dynamic_object_declaration = { - dyn_cache_function : section_path * obj -> unit; - dyn_load_function : section_path * obj -> unit; - dyn_open_function : section_path * obj -> unit; + dyn_cache_function : object_name * obj -> unit; + dyn_load_function : int -> object_name * obj -> unit; + dyn_open_function : int -> object_name * obj -> unit; + dyn_subst_function : object_name * substitution * obj -> obj; + dyn_classify_function : object_name * obj -> obj substitutivity; dyn_export_function : obj -> obj option } let object_tag lobj = Dyn.tag lobj @@ -42,23 +78,43 @@ let object_tag lobj = Dyn.tag lobj let cache_tab = (Hashtbl.create 17 : (string,dynamic_object_declaration) Hashtbl.t) -let declare_object (na,odecl) = +let declare_object odecl = + let na = odecl.object_name in let (infun,outfun) = Dyn.create na in - let cacher (spopt,lobj) = - if Dyn.tag lobj = na then odecl.cache_function(spopt,outfun lobj) + let cacher (oname,lobj) = + if Dyn.tag lobj = na then odecl.cache_function (oname,outfun lobj) else anomaly "somehow we got the wrong dynamic object in the cachefun" - and loader (spopt,lobj) = - if Dyn.tag lobj = na then odecl.load_function (spopt,outfun lobj) + and loader i (oname,lobj) = + if Dyn.tag lobj = na then odecl.load_function i (oname,outfun lobj) else anomaly "somehow we got the wrong dynamic object in the loadfun" - and opener (spopt,lobj) = - if Dyn.tag lobj = na then odecl.open_function (spopt,outfun lobj) + and opener i (oname,lobj) = + if Dyn.tag lobj = na then odecl.open_function i (oname,outfun lobj) else anomaly "somehow we got the wrong dynamic object in the openfun" + and substituter (oname,sub,lobj) = + if Dyn.tag lobj = na then + infun (odecl.subst_function (oname,sub,outfun lobj)) + else anomaly "somehow we got the wrong dynamic object in the substfun" + and classifier (spopt,lobj) = + if Dyn.tag lobj = na then + match odecl.classify_function (spopt,outfun lobj) with + | Dispose -> Dispose + | Substitute obj -> Substitute (infun obj) + | Keep obj -> Keep (infun obj) + | Anticipate (obj) -> Anticipate (infun obj) + else + anomaly "somehow we got the wrong dynamic object in the classifyfun" and exporter lobj = - option_app infun (odecl.export_function (outfun lobj)) + if Dyn.tag lobj = na then + option_app infun (odecl.export_function (outfun lobj)) + else + anomaly "somehow we got the wrong dynamic object in the exportfun" + in Hashtbl.add cache_tab na { dyn_cache_function = cacher; dyn_load_function = loader; dyn_open_function = opener; + dyn_subst_function = substituter; + dyn_classify_function = classifier; dyn_export_function = exporter }; (infun,outfun) @@ -85,11 +141,17 @@ let apply_dyn_fun deflt f lobj = let cache_object ((_,lobj) as node) = apply_dyn_fun () (fun d -> d.dyn_cache_function node) lobj -let load_object ((_,lobj) as node) = - apply_dyn_fun () (fun d -> d.dyn_load_function node) lobj +let load_object i ((_,lobj) as node) = + apply_dyn_fun () (fun d -> d.dyn_load_function i node) lobj + +let open_object i ((_,lobj) as node) = + apply_dyn_fun () (fun d -> d.dyn_open_function i node) lobj + +let subst_object ((_,_,lobj) as node) = + apply_dyn_fun lobj (fun d -> d.dyn_subst_function node) lobj -let open_object ((_,lobj) as node) = - apply_dyn_fun () (fun d -> d.dyn_open_function node) lobj +let classify_object ((_,lobj) as node) = + apply_dyn_fun Dispose (fun d -> d.dyn_classify_function node) lobj let export_object lobj = apply_dyn_fun None (fun d -> d.dyn_export_function lobj) lobj diff --git a/library/libobject.mli b/library/libobject.mli index 017650e76..1c1209019 100644 --- a/library/libobject.mli +++ b/library/libobject.mli @@ -10,37 +10,96 @@ (*i*) open Names +open Libnames (*i*) -(* [Libobject] declares persistent objects, given with three methods: - a caching function specifying how to add the object in the current - scope of theory modules; - a loading function, specifying what to do when the object is loaded - from the disk; - an opening function, specifying what to do when the module containing - the object is opened; - a specification function, to extract its specification when writing - the specification of a module to the disk (.vi) *) +(* [Libobject] declares persistent objects, given with methods: + + * a caching function specifying how to add the object in the current + scope; + If the object wishes to register its visibility in the Nametab, + it should do so for all possible sufixes. + + * a loading function, specifying what to do when the module + containing the object is loaded; + If the object wishes to register its visibility in the Nametab, + it should do so for all sufixes no shorter then the "int" argument + + * an opening function, specifying what to do when the module + containing the object is opened (imported); + If the object wishes to register its visibility in the Nametab, + it should do so for the sufix of the length the "int" argument + + * a classification function, specyfying what to do with the object, + when the current module (containing the object) is ended; + The possibilities are: + Dispose - the object dies at the end of the module + Substitue - meaning the object is substitutive and + the module name must be updated + Keep - the object is not substitutive, but survives module + closing + Anticipate - this is for objects which have to be explicitely + managed by the end_module function (like Require + and Read markers) + + The classification function is also an occasion for a cleanup + (if this function returns Keep or Substitute of some object, the + cache method is never called for it) + + * a substitution function, performing the substitution; + this function should be declared for substitutive objects + only (see obove) + + * an export function, to enable optional writing of its contents + to disk (.vo). This function is also the oportunity to remove + redundant information in order to keep .vo size small + + The export function is a little obsolete and will be removed + in the near future... + +*) + +type 'a substitutivity = + Dispose | Substitute of 'a | Keep of 'a | Anticipate of 'a type 'a object_declaration = { - cache_function : section_path * 'a -> unit; - load_function : section_path * 'a -> unit; - open_function : section_path * 'a -> unit; + object_name : string; + cache_function : object_name * 'a -> unit; + load_function : int -> object_name * 'a -> unit; + open_function : int -> object_name * 'a -> unit; + classify_function : object_name * 'a -> 'a substitutivity; + subst_function : object_name * substitution * 'a -> 'a; export_function : 'a -> 'a option } -(*s Given an object name and a declaration, the function [declare_object] +(* The default object is a "Keep" object with empty methods. + Object creators are advised to use the construction + [{(default_object "MY_OBJECT") with + cache_function = ... + }] + and specify only these functions which are not empty/meaningless + +*) + +val default_object : string -> 'a object_declaration + +(* the identity substitution function *) +val ident_subst_function : object_name * substitution * 'a -> 'a + +(*s Given an object declaration, the function [declare_object] will hand back two functions, the "injection" and "projection" functions for dynamically typed library-objects. *) type obj val declare_object : - (string * 'a object_declaration) -> ('a -> obj) * (obj -> 'a) + 'a object_declaration -> ('a -> obj) * (obj -> 'a) val object_tag : obj -> string -val cache_object : section_path * obj -> unit -val load_object : section_path * obj -> unit -val open_object : section_path * obj -> unit +val cache_object : object_name * obj -> unit +val load_object : int -> object_name * obj -> unit +val open_object : int -> object_name * obj -> unit +val subst_object : object_name * substitution * obj -> obj +val classify_object : object_name * obj -> obj substitutivity val export_object : obj -> obj option val relax : bool -> unit diff --git a/library/library.ml b/library/library.ml index 93ab76371..8dd3c5432 100644 --- a/library/library.ml +++ b/library/library.ml @@ -12,13 +12,16 @@ open Pp open Util open Names +open Libnames open Nameops -open Environ +open Safe_typing open Libobject open Lib open Nametab +open Declaremods -(*s Load path. *) +(*************************************************************************) +(*s Load path. Mapping from physical to logical paths etc.*) type logical_path = dir_path @@ -54,6 +57,7 @@ let canonical_path_name p = (* We give up to find a canonical name and just simplify it... *) strip_path p + let find_logical_path phys_dir = let phys_dir = canonical_path_name phys_dir in match list_filter2 (fun p d -> p = phys_dir) !load_path with @@ -75,7 +79,7 @@ let add_load_path_entry (phys_path,coq_path) = && coq_path = Nameops.default_root_prefix) then begin - (* Assume the user is concerned by module naming *) + (* Assume the user is concerned by library naming *) if dir <> Nameops.default_root_prefix then (Options.if_verbose warning (phys_path^" was previously bound to " ^(string_of_dirpath dir) @@ -95,31 +99,32 @@ let load_path_of_logical_path dir = let get_full_load_path () = List.combine (fst !load_path) (snd !load_path) +(***********************************************************************) (*s Modules on disk contain the following informations (after the magic number, and before the digest). *) type compilation_unit_name = dir_path -type module_disk = { +type library_disk = { md_name : compilation_unit_name; - md_compiled_env : compiled_env; - md_declarations : library_segment; + md_compiled : compiled_library; + md_objects : library_objects; md_deps : (compilation_unit_name * Digest.t) list; md_imports : compilation_unit_name list } (*s Modules loaded in memory contain the following informations. They are - kept in the global table [modules_table]. *) - -type module_t = { - module_name : compilation_unit_name; - module_filename : System.physical_path; - module_compiled_env : compiled_env; - module_declarations : library_segment; - module_deps : (compilation_unit_name * Digest.t) list; - module_imports : compilation_unit_name list; - module_digest : Digest.t } - -module CompUnitOrdered = + kept in the global table [libraries_table]. *) + +type library_t = { + library_name : compilation_unit_name; + library_filename : System.physical_path; + library_compiled : compiled_library; + library_objects : library_objects; + library_deps : (compilation_unit_name * Digest.t) list; + library_imports : compilation_unit_name list; + library_digest : Digest.t } + +module CompilingModuleOrdered = struct type t = dir_path let compare d1 d2 = @@ -127,33 +132,33 @@ module CompUnitOrdered = (List.rev (repr_dirpath d1)) (List.rev (repr_dirpath d2)) end -module CompUnitmap = Map.Make(CompUnitOrdered) +module CompilingModulemap = Map.Make(CompilingModuleOrdered) -(* This is a map from names to modules *) -let modules_table = ref CompUnitmap.empty +(* This is a map from names to libraries *) +let libraries_table = ref CompilingModulemap.empty -(* These are the _ordered_ lists of loaded, imported and exported modules *) -let modules_loaded_list = ref [] -let modules_imports_list = ref [] -let modules_exports_list = ref [] +(* These are the _ordered_ lists of loaded, imported and exported libraries *) +let libraries_loaded_list = ref [] +let libraries_imports_list = ref [] +let libraries_exports_list = ref [] let freeze () = - !modules_table, - !modules_loaded_list, - !modules_imports_list, - !modules_exports_list + !libraries_table, + !libraries_loaded_list, + !libraries_imports_list, + !libraries_exports_list let unfreeze (mt,mo,mi,me) = - modules_table := mt; - modules_loaded_list := mo; - modules_imports_list := mi; - modules_exports_list := me + libraries_table := mt; + libraries_loaded_list := mo; + libraries_imports_list := mi; + libraries_exports_list := me let init () = - modules_table := CompUnitmap.empty; - modules_loaded_list := []; - modules_imports_list := []; - modules_exports_list := [] + libraries_table := CompilingModulemap.empty; + libraries_loaded_list := []; + libraries_imports_list := []; + libraries_exports_list := [] let _ = Summary.declare_summary "MODULES" @@ -162,40 +167,42 @@ let _ = Summary.init_function = init; Summary.survive_section = false } -let find_module s = +let find_library s = try - CompUnitmap.find s !modules_table + CompilingModulemap.find s !libraries_table with Not_found -> - error ("Unknown module " ^ (string_of_dirpath s)) + error ("Unknown library " ^ (string_of_dirpath s)) -let module_is_loaded dir = - try let _ = CompUnitmap.find dir !modules_table in true +let library_full_filename m = (find_library m).library_filename + +let library_is_loaded dir = + try let _ = CompilingModulemap.find dir !libraries_table in true with Not_found -> false -let module_is_opened dir = - List.exists (fun m -> m.module_name = dir) !modules_imports_list +let library_is_opened dir = + List.exists (fun m -> m.library_name = dir) !libraries_imports_list -let module_is_exported dir = - List.exists (fun m -> m.module_name = dir) !modules_exports_list +let library_is_exported dir = + List.exists (fun m -> m.library_name = dir) !libraries_exports_list -let loaded_modules () = - List.map (fun m -> m.module_name) !modules_loaded_list +let loaded_libraries () = + List.map (fun m -> m.library_name) !libraries_loaded_list -let opened_modules () = - List.map (fun m -> m.module_name) !modules_imports_list +let opened_libraries () = + List.map (fun m -> m.library_name) !libraries_imports_list - (* If a module is loaded several time, then the first occurrence must - be performed first, thus the modules_loaded_list ... *) + (* If a library is loaded several time, then the first occurrence must + be performed first, thus the libraries_loaded_list ... *) -let register_loaded_module m = +let register_loaded_library m = let rec aux = function | [] -> [m] - | m'::_ as l when m'.module_name = m.module_name -> l + | m'::_ as l when m'.library_name = m.library_name -> l | m'::l' -> m' :: aux l' in - modules_loaded_list := aux !modules_loaded_list; - modules_table := CompUnitmap.add m.module_name m !modules_table + libraries_loaded_list := aux !libraries_loaded_list; + libraries_table := CompilingModulemap.add m.library_name m !libraries_table - (* ... while if a module is imported/exported several time, then + (* ... while if a library is imported/exported several time, then only the last occurrence is really needed - though the imported list may differ from the exported list (consider the sequence Export A; Export B; Import A which results in A;B for exports but @@ -204,87 +211,89 @@ let register_loaded_module m = let rec remember_last_of_each l m = match l with | [] -> [m] - | m'::l' when m'.module_name = m.module_name -> remember_last_of_each l' m + | m'::l' when m'.library_name = m.library_name -> remember_last_of_each l' m | m'::l' -> m' :: remember_last_of_each l' m -let register_open_module export m = - modules_imports_list := remember_last_of_each !modules_imports_list m; +let register_open_library export m = + libraries_imports_list := remember_last_of_each !libraries_imports_list m; if export then - modules_exports_list := remember_last_of_each !modules_exports_list m - -let compunit_cache = ref CompUnitmap.empty - -let module_segment = function - | None -> contents_after None - | Some m -> (find_module m).module_declarations - -let module_full_filename m = (find_module m).module_filename - -let vo_magic_number = 0703 (* V7.3 *) + libraries_exports_list := remember_last_of_each !libraries_exports_list m -let (raw_extern_module, raw_intern_module) = - System.raw_extern_intern vo_magic_number ".vo" +(************************************************************************) +(*s Operations on library objects and opening *) -let segment_rec_iter f = +(*let segment_rec_iter f = let rec apply = function | sp,Leaf obj -> f (sp,obj) | _,OpenedSection _ -> assert false | _,ClosedSection (_,_,seg) -> iter seg - | _,(FrozenState _ | Module _) -> () + | _,(FrozenState _ | CompilingModule _ + | OpenedModule _ | OpenedModtype _) -> () and iter seg = List.iter apply seg in iter -let segment_iter f = +let segment_iter i v f = let rec apply = function - | sp,Leaf obj -> f (sp,obj) + | sp,Leaf obj -> f (i,sp,obj) | _,OpenedSection _ -> assert false | sp,ClosedSection (export,dir,seg) -> - push_section dir; + push_dir v dir (DirClosedSection dir); if export then iter seg - | _,(FrozenState _ | Module _) -> () + | _,(FrozenState _ | CompilingModule _ + | OpenedModule _ | OpenedModtype _) -> () and iter seg = List.iter apply seg in iter +*) -(*s [open_module s] opens a module. The module [s] and all modules needed by +(*s [open_library s] opens a library. The library [s] and all libraries needed by [s] are assumed to be already loaded. When opening [s] we recursively open - all the modules needed by [s] and tagged [exported]. *) + all the libraries needed by [s] and tagged [exported]. *) -let open_objects decls = - segment_iter open_object decls +(*let open_objects i decls = + segment_iter i (Exactly i) open_object decls*) -let open_module export m = - if not (module_is_opened m.module_name) then - (register_open_module export m; - open_objects m.module_declarations) +let open_library export m = + if not (library_is_opened m.library_name) then begin + register_open_library export m; + Declaremods.import_module (MPfile m.library_name) + end else if export then - modules_exports_list := remember_last_of_each !modules_exports_list m + libraries_exports_list := remember_last_of_each !libraries_exports_list m -let open_modules export modl = +let open_libraries export modl = let to_open_list = List.fold_left (fun l m -> let subimport = List.fold_left - (fun l m -> remember_last_of_each l (find_module m)) - [] m.module_imports + (fun l m -> remember_last_of_each l (find_library m)) + [] m.library_imports in remember_last_of_each subimport m) [] modl in - List.iter (open_module export) to_open_list + List.iter (open_library export) to_open_list -(*s [load_module s] loads the module [s] from the disk, and [find_module s] - returns the module of name [s], loading it if necessary. - The boolean [doexp] specifies if we open the modules which are declared + +(************************************************************************) +(*s Loading from disk to cache (preparation phase) *) + +let compunit_cache = ref CompilingModulemap.empty + +let vo_magic_number = 0703 (* V7.3 *) + +let (raw_extern_library, raw_intern_library) = + System.raw_extern_intern vo_magic_number ".vo" + +(*s [load_library s] loads the library [s] from the disk, and [find_library s] + returns the library of name [s], loading it if necessary. + The boolean [doexp] specifies if we open the libraries which are declared exported in the dependencies (it is [true] at the highest level; then same value as for caller is reused in recursive loadings). *) -let load_objects decls = - segment_iter load_object decls - exception LibUnmappedDir exception LibNotFound type library_location = LibLoaded | LibInPath @@ -292,8 +301,8 @@ type library_location = LibLoaded | LibInPath let locate_absolute_library dir = (* Look if loaded in current environment *) try - let m = CompUnitmap.find dir !modules_table in - (dir, m.module_filename) + let m = CompilingModulemap.find dir !libraries_table in + (dir, m.library_filename) with Not_found -> (* Look if in loadpath *) try @@ -313,112 +322,94 @@ let with_magic_number_check f a = spc () ++ str"It is corrupted" ++ spc () ++ str"or was compiled with another version of Coq.") -let rec load_module dir = - try - (* Look if loaded in current env (and consequently its dependencies) *) - CompUnitmap.find dir !modules_table - with Not_found -> - (* [dir] is an absolute name which matches [f] which must be in loadpath *) - let m = - try CompUnitmap.find dir !compunit_cache - with Not_found -> - anomaly ((string_of_dirpath dir)^" should be in cache") - in - List.iter (fun (dir,_) -> ignore (load_module dir)) m.module_deps; - Global.import m.module_compiled_env; - load_objects m.module_declarations; - register_loaded_module m; - Nametab.push_loaded_library m.module_name; - m - -let mk_module md f digest = { - module_name = md.md_name; - module_filename = f; - module_compiled_env = md.md_compiled_env; - module_declarations = md.md_declarations; - module_deps = md.md_deps; - module_imports = md.md_imports; - module_digest = digest } +let mk_library md f digest = { + library_name = md.md_name; + library_filename = f; + library_compiled = md.md_compiled; + library_objects = md.md_objects; + library_deps = md.md_deps; + library_imports = md.md_imports; + library_digest = digest } let intern_from_file f = - let ch = with_magic_number_check raw_intern_module f in + let ch = with_magic_number_check raw_intern_library f in let md = System.marshal_in ch in let digest = System.marshal_in ch in close_in ch; - mk_module md f digest + mk_library md f digest -let rec intern_module (dir, f) = +let rec intern_library (dir, f) = try (* Look if in the current logical environment *) - CompUnitmap.find dir !modules_table + CompilingModulemap.find dir !libraries_table with Not_found -> try (* Look if already loaded in cache and consequently its dependencies *) - CompUnitmap.find dir !compunit_cache + CompilingModulemap.find dir !compunit_cache with Not_found -> (* [dir] is an absolute name which matches [f] which must be in loadpath *) let m = intern_from_file f in - if dir <> m.module_name then - errorlabstrm "load_physical_module" - (str ("The file " ^ f ^ " contains module") ++ spc () ++ - pr_dirpath m.module_name ++ spc () ++ str "and not module" ++ + if dir <> m.library_name then + errorlabstrm "load_physical_library" + (str ("The file " ^ f ^ " contains library") ++ spc () ++ + pr_dirpath m.library_name ++ spc () ++ str "and not library" ++ spc() ++ pr_dirpath dir); - compunit_cache := CompUnitmap.add dir m !compunit_cache; - List.iter (intern_mandatory_module dir) m.module_deps; + compunit_cache := CompilingModulemap.add dir m !compunit_cache; + List.iter (intern_mandatory_library dir) m.library_deps; m -and intern_mandatory_module caller (dir,d) = - let m = intern_absolute_module_from dir in - if d <> m.module_digest then - error ("compiled module "^(string_of_dirpath caller)^ - " makes inconsistent assumptions over module " +and intern_mandatory_library caller (dir,d) = + let m = intern_absolute_library_from dir in + if d <> m.library_digest then + error ("compiled library "^(string_of_dirpath caller)^ + " makes inconsistent assumptions over library " ^(string_of_dirpath dir)) -and intern_absolute_module_from dir = +and intern_absolute_library_from dir = try - intern_module (locate_absolute_library dir) + intern_library (locate_absolute_library dir) with | LibUnmappedDir -> let prefix, dir = fst (split_dirpath dir), string_of_dirpath dir in - errorlabstrm "load_absolute_module_from" + errorlabstrm "load_absolute_library_from" (str ("Cannot load "^dir^":") ++ spc () ++ str "no physical path bound to" ++ spc () ++ pr_dirpath prefix ++ fnl ()) | LibNotFound -> - errorlabstrm "load_absolute_module_from" - (str"Cannot find module " ++ pr_dirpath dir ++ str" in loadpath") + errorlabstrm "load_absolute_library_from" + (str"Cannot find library " ++ pr_dirpath dir ++ str" in loadpath") | e -> raise e -let rec_intern_module mref = let _ = intern_module mref in () +let rec_intern_library mref = let _ = intern_library mref in () -let check_module_short_name f dir = function +let check_library_short_name f dir = function | Some id when id <> snd (split_dirpath dir) -> - errorlabstrm "check_module_short_name" - (str ("The file " ^ f ^ " contains module") ++ spc () ++ - pr_dirpath dir ++ spc () ++ str "and not module" ++ spc () ++ + errorlabstrm "check_library_short_name" + (str ("The file " ^ f ^ " contains library") ++ spc () ++ + pr_dirpath dir ++ spc () ++ str "and not library" ++ spc () ++ pr_id id) | _ -> () let rec_intern_by_filename_only id f = let m = intern_from_file f in (* Only the base name is expected to match *) - check_module_short_name f m.module_name id; - (* We check no other file containing same module is loaded *) + check_library_short_name f m.library_name id; + (* We check no other file containing same library is loaded *) try - let m' = CompUnitmap.find m.module_name !modules_table in + let m' = CompilingModulemap.find m.library_name !libraries_table in Options.if_verbose warning - ((string_of_dirpath m.module_name)^" is already loaded from file "^ - m'.module_filename); - m.module_name + ((string_of_dirpath m.library_name)^" is already loaded from file "^ + m'.library_filename); + m.library_name with Not_found -> - compunit_cache := CompUnitmap.add m.module_name m !compunit_cache; - List.iter (intern_mandatory_module m.module_name) m.module_deps; - m.module_name + compunit_cache := CompilingModulemap.add m.library_name m !compunit_cache; + List.iter (intern_mandatory_library m.library_name) m.library_deps; + m.library_name let locate_qualified_library qid = (* Look if loaded in current environment *) try let dir = Nametab.locate_loaded_library qid in - (LibLoaded, dir, module_full_filename dir) + (LibLoaded, dir, library_full_filename dir) with Not_found -> (* Look if in loadpath *) try @@ -438,7 +429,7 @@ let locate_qualified_library qid = let rec_intern_qualified_library (loc,qid) = try let (_,dir,f) = locate_qualified_library qid in - rec_intern_module (dir,f); + rec_intern_library (dir,f); dir with | LibUnmappedDir -> @@ -449,42 +440,79 @@ let rec_intern_qualified_library (loc,qid) = fnl ()) | LibNotFound -> user_err_loc (loc, "rec_intern_qualified_library", - str"Cannot find module " ++ pr_qualid qid ++ str" in loadpath") + str"Cannot find library " ++ pr_qualid qid ++ str" in loadpath") -let rec_intern_module_from_file idopt f = - (* A name is specified, we have to check it contains module id *) +let rec_intern_library_from_file idopt f = + (* A name is specified, we have to check it contains library id *) let _, f = System.find_file_in_path (get_load_path ()) (f^".vo") in rec_intern_by_filename_only idopt f -(*s [require_module] loads and opens a module. This is a synchronized - operation. *) +(**********************************************************************) +(*s [require_library] loads and opens a library. This is a synchronized + operation. It is performed as follows: + + preparation phase: (functions require_library* ) the library and its + dependencies are read from to disk to the compunit_cache + (using intern_* ) -type module_reference = dir_path list * bool option + execution phase: (through add_leaf and cache_require) + the library is loaded in the environment and Nametab, the objects are + registered etc, using functions from Declaremods (via load_library, + which recursively loads its dependencies) -let string_of_module (_,dir,_) = string_of_id (List.hd (repr_dirpath dir)) + + The [read_library] operation is very similar, but has does not "open" + the library +*) + +type library_reference = dir_path list * bool option + +let string_of_library (_,dir,_) = string_of_id (List.hd (repr_dirpath dir)) + +let rec load_library dir = + try + (* Look if loaded in current env (and consequently its dependencies) *) + CompilingModulemap.find dir !libraries_table + with Not_found -> + (* [dir] is an absolute name matching [f] which must be in loadpath *) + let m = + try CompilingModulemap.find dir !compunit_cache + with Not_found -> + anomaly ((string_of_dirpath dir)^" should be in cache") + in + List.iter (fun (dir,_) -> ignore (load_library dir)) m.library_deps; + Declaremods.register_library + m.library_name + m.library_compiled + m.library_objects + m.library_digest; + register_loaded_library m; + m let cache_require (_,(modl,export)) = - let ml = list_map_left load_module modl in + let ml = list_map_left load_library modl in match export with | None -> () - | Some export -> open_modules export ml + | Some export -> open_libraries export ml + +let load_require _ (_,(modl,_)) = + ignore(list_map_left load_library modl) (* keeps the require marker for closed section replay but removes OS dependent fields from .vo files for cross-platform compatibility *) let export_require (l,e) = Some (l,e) let (in_require, out_require) = - declare_object - ("REQUIRE", - { cache_function = cache_require; - load_function = (fun _ -> ()); - open_function = (fun _ -> ()); - export_function = export_require }) + declare_object {(default_object "REQUIRE") with + cache_function = cache_require; + load_function = load_require; + export_function = export_require; + classify_function = (fun (_,o) -> Anticipate o) } -let require_module spec qidl export = +let require_library spec qidl export = (* if sections_are_opened () && Options.verbose () then - warning ("Objets of "^(string_of_module modref)^ + warning ("Objets of "^(string_of_library modref)^ " not surviving sections (e.g. Grammar \nand Hints)\n"^ "will be removed at the end of the section"); *) @@ -492,53 +520,67 @@ let require_module spec qidl export = add_anonymous_leaf (in_require (modrefl,Some export)); add_frozen_state () -let require_module_from_file spec idopt file export = - let modref = rec_intern_module_from_file idopt file in +let require_library_from_file spec idopt file export = + let modref = rec_intern_library_from_file idopt file in add_anonymous_leaf (in_require ([modref],Some export)); add_frozen_state () -let import_module export (loc,qid) = - let modref = - try - Nametab.locate_loaded_library qid - with Not_found -> - user_err_loc - (loc,"import_module", - str ((Nametab.string_of_qualid qid)^" not loaded")) in - add_anonymous_leaf (in_require ([modref],Some export)) - -let read_module qid = +let export_library (loc,qid) = + try + match Nametab.locate_module qid with + MPfile dir -> + add_anonymous_leaf (in_require ([dir],Some true)) + | _ -> + raise Not_found + with + Not_found -> + user_err_loc + (loc,"export_library", + str ((string_of_qualid qid)^" is not a loaded library")) + +let read_library qid = let modref = rec_intern_qualified_library qid in add_anonymous_leaf (in_require ([modref],None)); add_frozen_state () -let read_module_from_file f = +let read_library_from_file f = let _, f = System.find_file_in_path (get_load_path ()) (f^".vo") in let modref = rec_intern_by_filename_only None f in add_anonymous_leaf (in_require ([modref],None)); add_frozen_state () -let reload_module (modrefl, export) = +let reload_library (modrefl, export) = add_anonymous_leaf (in_require (modrefl,export)); add_frozen_state () -(*s [save_module s] saves the module [m] to the disk. *) + +(***********************************************************************) +(*s [save_library s] saves the library [m] to the disk. *) + +let start_library f = + let _,longf = + System.find_file_in_path (get_load_path ()) (f^".v") in + let ldir0 = find_logical_path (Filename.dirname longf) in + let id = id_of_string (Filename.basename f) in + let ldir = extend_dirpath ldir0 id in + Declaremods.start_library ldir; + ldir,longf let current_deps () = - List.map (fun m -> (m.module_name, m.module_digest)) !modules_loaded_list + List.map (fun m -> (m.library_name, m.library_digest)) !libraries_loaded_list let current_reexports () = - List.map (fun m -> m.module_name) !modules_exports_list + List.map (fun m -> m.library_name) !libraries_exports_list -let save_module_to s f = - let seg = export_module s in +let save_library_to s f = + let cenv, seg = Declaremods.export_library s in let md = { md_name = s; - md_compiled_env = Global.export s; - md_declarations = seg; + md_compiled = cenv; + md_objects = seg; md_deps = current_deps (); md_imports = current_reexports () } in - let (f',ch) = raw_extern_module f in + let (f',ch) = raw_extern_library f in try System.marshal_out ch md; flush ch; @@ -547,62 +589,38 @@ let save_module_to s f = close_out ch with e -> (warning ("Removed file "^f');close_out ch; Sys.remove f'; raise e) -(*s Iterators. *) - -let fold_all_segments insec f x = - let rec apply acc = function - | sp, Leaf o -> f acc sp o - | _, ClosedSection (_,_,seg) -> - if insec then List.fold_left apply acc seg else acc - | _ -> acc - in - let acc' = - CompUnitmap.fold - (fun _ m acc -> List.fold_left apply acc m.module_declarations) - !modules_table x - in - List.fold_left apply acc' (Lib.contents_after None) - -let iter_all_segments insec f = - let rec apply = function - | sp, Leaf o -> f sp o - | _, ClosedSection (_,_,seg) -> if insec then List.iter apply seg - | _ -> () - in - CompUnitmap.iter - (fun _ m -> List.iter apply m.module_declarations) !modules_table; - List.iter apply (Lib.contents_after None) - -(*s Pretty-printing of modules state. *) - -let fmt_modules_state () = - let opened = opened_modules () - and loaded = loaded_modules () in +(*s Pretty-printing of libraries state. *) + +let fmt_libraries_state () = + let opened = opened_libraries () + and loaded = loaded_libraries () in (str "Imported (open) Modules: " ++ prlist_with_sep pr_spc pr_dirpath opened ++ fnl () ++ str "Loaded Modules: " ++ prlist_with_sep pr_spc pr_dirpath loaded ++ fnl ()) + (*s For tactics/commands requiring vernacular libraries *) -let check_required_module d = +let check_required_library d = let d' = List.map id_of_string d in let dir = make_dirpath (List.rev d') in - if not (module_is_loaded dir) then + if not (library_is_loaded dir) then (* Loading silently ... let m, prefix = list_sep_last d' in - read_module + read_library (dummy_loc,make_qualid (make_dirpath (List.rev prefix)) m) *) (* or failing ...*) - error ("Module "^(list_last d)^" has to be required first") + error ("Library "^(list_last d)^" has to be required first") + -(*s Display the memory use of a module. *) +(*s Display the memory use of a library. *) open Printf let mem s = - let m = find_module s in + let m = find_library s in h 0 (str (sprintf "%dk (cenv = %dk / seg = %dk)" - (size_kb m) (size_kb m.module_compiled_env) - (size_kb m.module_declarations))) + (size_kb m) (size_kb m.library_compiled) + (size_kb m.library_objects))) diff --git a/library/library.mli b/library/library.mli index dc34ec72e..43d73d241 100644 --- a/library/library.mli +++ b/library/library.mli @@ -11,61 +11,59 @@ (*i*) open Util open Names +open Libnames open Libobject (*i*) (*s This module is the heart of the library. It provides low level functions to - load, open and save modules. Modules are saved onto the disk with checksums + load, open and save libraries. Modules are saved onto the disk with checksums (obtained with the [Digest] module), which are checked at loading time to prevent inconsistencies between files written at various dates. It also provides a high level function [require] which corresponds to the vernacular command [Require]. *) -val read_module : Nametab.qualid located -> unit -val read_module_from_file : System.physical_path -> unit -val import_module : bool -> Nametab.qualid located -> unit +val read_library : qualid located -> unit -val module_is_loaded : dir_path -> bool -val module_is_opened : dir_path -> bool +val read_library_from_file : System.physical_path -> unit -val loaded_modules : unit -> dir_path list -val opened_modules : unit -> dir_path list +val export_library : qualid located -> unit -val fmt_modules_state : unit -> Pp.std_ppcmds +val library_is_loaded : dir_path -> bool +val library_is_opened : dir_path -> bool -(*s Require. The command [require_module spec m file export] loads and opens - a module [m]. [file] specifies the filename, if not [None]. [spec] +val loaded_libraries : unit -> dir_path list +val opened_libraries : unit -> dir_path list + +val fmt_libraries_state : unit -> Pp.std_ppcmds + +(*s Require. The command [require_library spec m file export] loads and opens + a library [m]. [file] specifies the filename, if not [None]. [spec] specifies to look for a specification ([true]) or for an implementation - ([false]), if not [None]. And [export] specifies if the module must be + ([false]), if not [None]. And [export] specifies if the library must be exported. *) -val require_module : - bool option -> Nametab.qualid located list -> bool -> unit +val require_library : + bool option -> qualid located list -> bool -> unit -val require_module_from_file : +val require_library_from_file : bool option -> identifier option -> string -> bool -> unit -(*s [save_module_to s f] saves the current environment as a module [s] +(*s [save_library_to s f] saves the current environment as a library [s] in the file [f]. *) -val save_module_to : dir_path -> string -> unit +val start_library : string -> dir_path * string +val save_library_to : dir_path -> string -> unit -(*s [module_segment m] returns the segment of the loaded module - [m]; if not given, the segment of the current module is returned +(*s [library_segment m] returns the segment of the loaded library + [m]; if not given, the segment of the current library is returned (which is then the same as [Lib.contents_after None]). - [module_full_filename] returns the full filename of a loaded module. *) +*) +(*val library_segment : dir_path option -> Lib.library_segment*) -val module_segment : dir_path option -> Lib.library_segment -val module_full_filename : dir_path -> string +(* [library_full_filename] returns the full filename of a loaded library. *) -(*s [fold_all_segments] and [iter_all_segments] iterate over all - segments, the modules' segments first and then the current - segment. Modules are presented in an arbitrary order. The given - function is applied to all leaves (together with their section - path). The boolean indicates if we must enter closed sections. *) +val library_full_filename : dir_path -> string -val fold_all_segments : bool -> ('a -> section_path -> obj -> 'a) -> 'a -> 'a -val iter_all_segments : bool -> (section_path -> obj -> unit) -> unit (*s Global load path *) type logical_path = dir_path @@ -82,16 +80,16 @@ exception LibNotFound type library_location = LibLoaded | LibInPath val locate_qualified_library : - Nametab.qualid -> library_location * dir_path * System.physical_path + qualid -> library_location * dir_path * System.physical_path -val check_required_module : string list -> unit -(*s Displays the memory use of a module. *) +val check_required_library : string list -> unit +(*s Displays the memory use of a library. *) val mem : dir_path -> Pp.std_ppcmds (* For discharge *) -type module_reference +type library_reference -val out_require : Libobject.obj -> module_reference -val reload_module : module_reference -> unit +val out_require : Libobject.obj -> library_reference +val reload_library : library_reference -> unit diff --git a/library/nameops.ml b/library/nameops.ml index ecbe07e77..0fd9ec0d1 100644 --- a/library/nameops.ml +++ b/library/nameops.ml @@ -143,12 +143,7 @@ let next_name_away name l = | Name str -> next_ident_away str l | Anonymous -> id_of_string "_" -(**********************************************) - -let pr_dirpath sl = (str (string_of_dirpath sl)) - -(* Operations on dirpaths *) -let empty_dirpath = make_dirpath [] +let pr_lab l = str (string_of_label l) let default_module_name = id_of_string "Top" let default_module = make_dirpath [default_module_name] @@ -157,81 +152,3 @@ let default_module = make_dirpath [default_module_name] let coq_root = id_of_string "Coq" let default_root_prefix = make_dirpath [] -let restrict_path n sp = - let dir, s = repr_path sp in - let (dir',_) = list_chop n (repr_dirpath dir) in - Names.make_path (make_dirpath dir') s - -(* Pop the last n module idents *) -let extract_dirpath_prefix n dir = - let (_,dir') = list_chop n (repr_dirpath dir) in - make_dirpath dir' - -let dirpath_prefix p = match repr_dirpath p with - | [] -> anomaly "dirpath_prefix: empty dirpath" - | _::l -> make_dirpath l - -let is_dirpath_prefix_of d1 d2 = - list_prefix_of (List.rev (repr_dirpath d1)) (List.rev (repr_dirpath d2)) - -(* To know how qualified a name should be to be understood in the current env*) -let add_dirpath_prefix id d = make_dirpath (repr_dirpath d @ [id]) - -let split_dirpath d = - let l = repr_dirpath d in (make_dirpath (List.tl l), List.hd l) - -let extend_dirpath p id = make_dirpath (id :: repr_dirpath p) - - -(* Section paths *) - -let dirpath sp = let (p,_) = repr_path sp in p -let basename sp = let (_,id) = repr_path sp in id - -let path_of_constructor env ((sp,tyi),ind) = - let mib = Environ.lookup_mind sp env in - let mip = mib.mind_packets.(tyi) in - let (pa,_) = repr_path sp in - Names.make_path pa (mip.mind_consnames.(ind-1)) - -let path_of_inductive env (sp,tyi) = - if tyi = 0 then sp - else - let mib = Environ.lookup_mind sp env in - let mip = mib.mind_packets.(tyi) in - let (pa,_) = repr_path sp in - Names.make_path pa (mip.mind_typename) - -(* parsing *) -let parse_sp s = - let len = String.length s in - let rec decoupe_dirs n = - try - let pos = String.index_from s n '.' in - let dir = String.sub s n (pos-n) in - let dirs,n' = decoupe_dirs (succ pos) in - (id_of_string dir)::dirs,n' - with - | Not_found -> [],n - in - if len = 0 then invalid_arg "parse_section_path"; - let dirs,n = decoupe_dirs 0 in - let id = String.sub s n (len-n) in - make_dirpath (List.rev dirs), (id_of_string id) - -let dirpath_of_string s = - try - let sl,s = parse_sp s in - extend_dirpath sl s - with - | Invalid_argument _ -> invalid_arg "dirpath_of_string" - -let path_of_string s = - try - let sl,s = parse_sp s in - make_path sl s - with - | Invalid_argument _ -> invalid_arg "path_of_string" - -(* Section paths *) -let pr_sp sp = (str (string_of_path sp)) diff --git a/library/nameops.mli b/library/nameops.mli index 35346b0a6..591e9030d 100644 --- a/library/nameops.mli +++ b/library/nameops.mli @@ -34,9 +34,11 @@ val next_name_away_with_default : val out_name : name -> identifier -(* Section and module mechanism: dealinng with dir paths *) -val pr_dirpath : dir_path -> Pp.std_ppcmds -val empty_dirpath : dir_path + +val pr_lab : label -> Pp.std_ppcmds + +(* some preset paths *) + val default_module : dir_path (* This is the root of the standard library of Coq *) @@ -45,31 +47,3 @@ val coq_root : module_ident (* This is the default root prefix for developments which doesn't mention a root *) val default_root_prefix : dir_path - - -val dirpath_of_string : string -> dir_path -val path_of_string : string -> section_path - -val path_of_constructor : env -> constructor -> section_path -val path_of_inductive : env -> inductive -> section_path - - -val dirpath : section_path -> dir_path -val basename : section_path -> identifier - -(* Give the immediate prefix of a [dir_path] *) -val dirpath_prefix : dir_path -> dir_path - -(* Give the immediate prefix and basename of a [dir_path] *) -val split_dirpath : dir_path -> dir_path * identifier - -val extend_dirpath : dir_path -> module_ident -> dir_path -val add_dirpath_prefix : module_ident -> dir_path -> dir_path - -val extract_dirpath_prefix : int -> dir_path -> dir_path -val is_dirpath_prefix_of : dir_path -> dir_path -> bool - -val restrict_path : int -> section_path -> section_path - -(* Section path *) -val pr_sp : section_path -> Pp.std_ppcmds diff --git a/library/nametab.ml b/library/nametab.ml index b28506f91..c80675aec 100755 --- a/library/nametab.ml +++ b/library/nametab.ml @@ -11,23 +11,10 @@ open Util open Pp open Names +open Libnames open Nameops open Declarations -(*s qualified names *) -type qualid = section_path - -let make_qualid = make_path -let repr_qualid = repr_path - -let string_of_qualid = string_of_path -let pr_qualid = pr_sp - -let qualid_of_sp sp = sp -let make_short_qualid id = make_qualid empty_dirpath id -let qualid_of_dirpath dir = - let (l,a) = split_dirpath dir in - make_qualid l a exception GlobalizationError of qualid exception GlobalizationConstantError of qualid @@ -40,118 +27,150 @@ let error_global_constant_not_found_loc loc q = let error_global_not_found q = raise (GlobalizationError q) -(* Constructions and syntactic definitions live in the same space *) -type global_reference = - | VarRef of variable - | ConstRef of constant - | IndRef of inductive - | ConstructRef of constructor - -type extended_global_reference = - | TrueGlobal of global_reference - | SyntacticDef of section_path - -let sp_of_global env = function - | VarRef id -> make_path empty_dirpath id - | ConstRef sp -> sp - | IndRef (sp,tyi) -> - (* Does not work with extracted inductive types when the first - inductive is logic : if tyi=0 then basename sp else *) - let mib = Environ.lookup_mind sp env in - assert (tyi < mib.mind_ntypes && tyi >= 0); - let mip = mib.mind_packets.(tyi) in - let (p,_) = repr_path sp in - make_path p mip.mind_typename - | ConstructRef ((sp,tyi),i) -> - let mib = Environ.lookup_mind sp env in - assert (tyi < mib.mind_ntypes && i >= 0); - let mip = mib.mind_packets.(tyi) in - assert (i <= Array.length mip.mind_consnames && i > 0); - let (p,_) = repr_path sp in - make_path p mip.mind_consnames.(i-1) - +type 'a path_status = Nothing | Relative of 'a | Absolute of 'a (* Dictionaries of short names *) -type 'a nametree = ('a option * 'a nametree ModIdmap.t) +type 'a nametree = ('a path_status * 'a nametree ModIdmap.t) type ccitab = extended_global_reference nametree Idmap.t +type kntab = kernel_name nametree Idmap.t type objtab = section_path nametree Idmap.t -type dirtab = dir_path nametree ModIdmap.t +type dirtab = global_dir_reference nametree ModIdmap.t let the_ccitab = ref (Idmap.empty : ccitab) -let the_libtab = ref (ModIdmap.empty : dirtab) -let the_sectab = ref (ModIdmap.empty : dirtab) +let the_dirtab = ref (ModIdmap.empty : dirtab) let the_objtab = ref (Idmap.empty : objtab) +let the_modtypetab = ref (Idmap.empty : kntab) + +(* This table translates extended_global_references back to section paths *) +module Globtab = Map.Make(struct type t=extended_global_reference + let compare = compare end) + +type globtab = section_path Globtab.t + +let the_globtab = ref (Globtab.empty : globtab) + + +let sp_of_global ctx_opt ref = + match (ctx_opt,ref) with + | Some ctx, VarRef id -> + let _ = Sign.lookup_named id ctx in + make_path empty_dirpath id + | _ -> Globtab.find (TrueGlobal ref) !the_globtab + + +let full_name = sp_of_global None + +let id_of_global ctx_opt ref = + let (_,id) = repr_path (sp_of_global ctx_opt ref) in + id + +let sp_of_syntactic_definition kn = + Globtab.find (SyntacticDef kn) !the_globtab + +(******************************************************************) +(* the_dirpath is the table matching dir_paths to toplevel + modules/files or sections + + If we have a closed module M having a submodule N, than N does not + have the entry here. + +*) + +type visibility = Until of int | Exactly of int -let dirpath_of_reference ref = - let sp = match ref with - | ConstRef sp -> sp - | VarRef id -> make_path empty_dirpath id - | ConstructRef ((sp,_),_) -> sp - | IndRef (sp,_) -> sp in - let (p,_) = repr_path sp in - p - -let dirpath_of_extended_ref = function - | TrueGlobal ref -> dirpath_of_reference ref - | SyntacticDef sp -> - let (p,_) = repr_path sp in p - -(* How [visibility] works: a value of [0] means all suffixes of [dir] are - allowed to access the object, a value of [1] means all suffixes, except the - base name, are available, [2] means all suffixes except the base name and - the name qualified by the module name *) - -(* Concretely, library roots and directory are - always open but modules/files are open only during their interactive - construction or on demand if a precompiled one: for a name - "Root.Rep.Lib.name", then "Lib.name", "Rep.Lib.name" and - "Root.Rep.Lib.name", but not "name" are pushed; which means, visibility is - always 1 *) - -(* We add a binding of [[modid1;...;modidn;id]] to [o] in the name tab *) -(* We proceed in the reverse way, looking first to [id] *) -let push_tree extract_dirpath tab visibility dir o = - let extract = option_app (fun c -> repr_dirpath (extract_dirpath c)) in +(* is the short name visible? tests for 1 *) +let is_short_visible v sp = match v with + Until 1 | Exactly 1 -> true + | _ -> false + +(* In general, directories and sections are always open and modules + (and files) are open on request only *) + +(* We add a binding of ["modid1.....modidn.id"] to [o] in the table *) +(* Using the fact that the reprezentation of a [dir_path] is already + reversed, we proceed in the reverse way, looking first for [id] *) + +(* [push_many] is used to register [Until vis] visibility and + [push_one] to [Exactly vis] and [push_tree] chooses the right one*) + +let push_many vis tab dir o = let rec push level (current,dirmap) = function | modid :: path as dir -> let mc = try ModIdmap.find modid dirmap - with Not_found -> (None, ModIdmap.empty) + with Not_found -> (Nothing, ModIdmap.empty) in let this = - if level >= visibility then - if extract current = Some dir then + if level >= vis then + match current with + | Absolute _ -> + (* This is an absolute name, we must keep it otherwise it may + become unaccessible forever *) + warning "Trying to zaslonic an absolute name!"; current + | Nothing + | Relative _ -> Relative o + else current + in + (this, ModIdmap.add modid (push (level+1) mc path) dirmap) + | [] -> + match current with + | Absolute _ -> (* This is an absolute name, we must keep it otherwise it may - become unaccessible forever *) - current - else - Some o - else current in - (this, ModIdmap.add modid (push (level+1) mc path) dirmap) - | [] -> (Some o,dirmap) in - push 0 tab (repr_dirpath dir) - -let push_idtree extract_dirpath tab n dir id o = + become unaccessible forever *) + (* But ours is also absolute! This is an error! *) + error "Trying to zaslonic an absolute name!" + | Nothing + | Relative _ -> Absolute o, dirmap + in + push 0 tab (repr_dirpath dir) + +let push_one vis tab dir o = + let rec push level (current,dirmap) = function + | modid :: path as dir -> + let mc = + try ModIdmap.find modid dirmap + with Not_found -> (Nothing, ModIdmap.empty) + in + if level = vis then + let this = + match current with + | Absolute _ -> + (* This is an absolute name, we must keep it otherwise it may + become unaccessible forever *) + error "Trying to zaslonic an absolute name!" + | Nothing + | Relative _ -> Relative o + in + (this, dirmap) + else + (current, ModIdmap.add modid (push (level+1) mc path) dirmap) + | [] -> anomaly "We should never come to this point" + in + push 0 tab (repr_dirpath dir) + + +let push_tree = function + | Until i -> push_many (i-1) + | Exactly i -> push_one (i-1) + + + +let push_idtree tab visibility sp o = + let dir,id = repr_path sp in let modtab = try Idmap.find id !tab - with Not_found -> (None, ModIdmap.empty) in - tab := Idmap.add id (push_tree extract_dirpath modtab n dir o) !tab + with Not_found -> (Nothing, ModIdmap.empty) in + tab := + Idmap.add id (push_tree visibility modtab dir o) !tab -let push_long_names_ccipath = push_idtree dirpath_of_extended_ref the_ccitab -let push_short_name_ccipath = push_idtree dirpath_of_extended_ref the_ccitab -let push_long_names_objpath = - push_idtree (fun sp -> let (p,_) = repr_path sp in p) the_objtab -let push_short_name_objpath = - push_idtree (fun sp -> let (p,_) = repr_path sp in p) the_objtab -let push_modidtree tab dir id o = +let push_modidtree tab visibility dirpath o = + let dir,id = split_dirpath dirpath in let modtab = try ModIdmap.find id !tab - with Not_found -> (None, ModIdmap.empty) in - tab := ModIdmap.add id (push_tree (fun x -> x) modtab 0 dir o) !tab + with Not_found -> (Nothing, ModIdmap.empty) in + tab := ModIdmap.add id (push_tree visibility modtab dir o) !tab -let push_long_names_secpath = push_modidtree the_sectab -let push_long_names_libpath = push_modidtree the_libtab (* These are entry points for new declarations at toplevel *) @@ -159,95 +178,122 @@ let push_long_names_libpath = push_modidtree the_libtab possibly limited visibility, i.e. Theorem, Lemma, Definition, Axiom, Parameter but also Remark and Fact) *) -let push_cci ~visibility:n sp ref = - let dir, s = repr_path sp in - (* We push partially qualified name (with at least one prefix) *) - push_long_names_ccipath n dir s (TrueGlobal ref) +let push_xref visibility sp xref = + push_idtree the_ccitab visibility sp xref; + match visibility with + | Until _ -> + the_globtab := Globtab.add xref sp !the_globtab + | _ -> () -let push = push_cci +let push_cci visibility sp ref = + push_xref visibility sp (TrueGlobal ref) (* This is for Syntactic Definitions *) +let push_syntactic_definition visibility sp kn = + push_xref visibility sp (SyntacticDef kn) -let push_syntactic_definition sp = - let dir, s = repr_path sp in - push_long_names_ccipath 0 dir s (SyntacticDef sp) +let push = push_cci + +let push_modtype = push_idtree the_modtypetab -let push_short_name_syntactic_definition sp = - let _, s = repr_qualid (qualid_of_sp sp) in - push_short_name_ccipath Pervasives.max_int empty_dirpath s (SyntacticDef sp) (* This is for dischargeable non-cci objects (removed at the end of the section -- i.e. Hints, Grammar ...) *) (* --> Unused *) -let push_short_name_object sp = - let _, s = repr_qualid (qualid_of_sp sp) in - push_short_name_objpath 0 empty_dirpath s sp +let push_object visibility sp = + push_idtree the_objtab visibility sp sp (* This is for tactic definition names *) -let push_tactic_path sp = - let dir, s = repr_qualid (qualid_of_sp sp) in - push_long_names_objpath 0 dir s sp +let push_tactic = push_object (* This is to remember absolute Section/Module names and to avoid redundancy *) -let push_section fulldir = - let dir, s = split_dirpath fulldir in - (* We push all partially qualified name *) - push_long_names_secpath dir s fulldir; - push_long_names_secpath empty_dirpath s fulldir - -(* These are entry points to locate names *) -let locate_in_tree tab dir = - let dir = repr_dirpath dir in +let push_dir = push_modidtree the_dirtab + + +(* As before we start with generic functions *) + +let find_in_tree tab dir = let rec search (current,modidtab) = function | modid :: path -> search (ModIdmap.find modid modidtab) path - | [] -> match current with Some o -> o | _ -> raise Not_found + | [] -> current in search tab dir -let locate_cci (qid:qualid) = +let find_in_idtree tab qid = + let (dir,id) = repr_qualid qid in + find_in_tree (Idmap.find id !tab) (repr_dirpath dir) + +let find_in_modidtree tab qid = let (dir,id) = repr_qualid qid in - locate_in_tree (Idmap.find id !the_ccitab) dir + find_in_tree (ModIdmap.find id !tab) (repr_dirpath dir) + +let get = function + Absolute o | Relative o -> o + | Nothing -> raise Not_found + +let absolute = function + Absolute _ -> true + | Relative _ + | Nothing -> false + + +(* These are entry points to locate different things *) +let find_cci = find_in_idtree the_ccitab + +let find_dir = find_in_modidtree the_dirtab + +let find_modtype = find_in_idtree the_modtypetab (* This should be used when syntactic definitions are allowed *) -let extended_locate = locate_cci +let extended_locate qid = get (find_cci qid) (* This should be used when no syntactic definitions is expected *) -let locate qid = match locate_cci qid with +let locate qid = match extended_locate qid with | TrueGlobal ref -> ref | SyntacticDef _ -> raise Not_found -let locate_obj qid = - let (dir,id) = repr_qualid qid in - locate_in_tree (Idmap.find id !the_objtab) dir +let locate_syntactic_definition qid = match extended_locate qid with + | TrueGlobal _ -> raise Not_found + | SyntacticDef kn -> kn -let locate_tactic qid = - let (dir,id) = repr_qualid qid in - locate_in_tree (Idmap.find id !the_objtab) dir +let locate_modtype qid = get (find_modtype qid) -(* Actually, this table has only two levels, since only basename and *) -(* fullname are registered *) -let push_loaded_library fulldir = - let dir, s = split_dirpath fulldir in - push_long_names_libpath dir s fulldir; - push_long_names_libpath empty_dirpath s fulldir +let locate_obj qid = get (find_in_idtree the_objtab qid) -let locate_loaded_library qid = - let (dir,id) = repr_qualid qid in - locate_in_tree (ModIdmap.find id !the_libtab) dir +let locate_tactic qid = get (find_in_idtree the_objtab qid) + +let locate_dir qid = get (find_dir qid) + +let locate_module qid = + match locate_dir qid with + | DirModule (_,(mp,_)) -> mp + | _ -> raise Not_found + +let locate_loaded_library qid = + match locate_dir qid with + | DirModule (dir,_) -> dir + | _ -> raise Not_found let locate_section qid = - let (dir,id) = repr_qualid qid in - locate_in_tree (ModIdmap.find id !the_sectab) dir + match locate_dir qid with + | DirOpenSection (dir, _) + | DirClosedSection dir -> dir + | _ -> raise Not_found (* Derived functions *) let locate_constant qid = - (* TODO: restrict to defined constants *) - match locate_cci qid with - | TrueGlobal (ConstRef sp) -> sp + match extended_locate qid with + | TrueGlobal (ConstRef kn) -> kn + | _ -> raise Not_found + +let locate_mind qid = + match extended_locate qid with + | TrueGlobal (IndRef (kn,0)) -> kn | _ -> raise Not_found +(* let sp_of_id id = match locate_cci (make_short_qualid id) with | TrueGlobal ref -> ref | SyntacticDef _ -> @@ -258,14 +304,11 @@ let constant_sp_of_id id = match locate_cci (make_short_qualid id) with | TrueGlobal (ConstRef sp) -> sp | _ -> raise Not_found +*) let absolute_reference sp = - let a = locate_cci sp in - let (p,_) = repr_path sp in - if not (dirpath_of_extended_ref a = p) then - anomaly ("Not an absolute path: "^(string_of_path sp)); - match a with - | TrueGlobal ref -> ref + match find_cci (qualid_of_sp sp) with + | Absolute (TrueGlobal ref) -> ref | _ -> raise Not_found let locate_in_absolute_module dir id = @@ -282,13 +325,23 @@ let global (loc,qid) = error_global_not_found_loc loc qid let exists_cci sp = - try let _ = locate_cci sp in true + try + absolute (find_cci (qualid_of_sp sp)) with Not_found -> false - -let exists_section dir = - try let _ = locate_section (qualid_of_dirpath dir) in true + +let exists_dir dir = + try + absolute (find_dir (qualid_of_dirpath dir)) with Not_found -> false +let exists_section = exists_dir + +let exists_module = exists_dir + +let exists_modtype sp = + try + absolute (find_modtype (qualid_of_sp sp)) + with Not_found -> false (* For a sp Coq.A.B.x, try to find the shortest among x, B.x, A.B.x and Coq.A.B.x is a qualid that denotes the same object. *) @@ -323,25 +376,25 @@ let global_inductive (loc,qid as locqid) = (********************************************************************) (* Registration of tables as a global table and rollback *) -type frozen = ccitab * dirtab * dirtab * objtab * identifier list +type frozen = ccitab * dirtab * objtab * globtab let init () = the_ccitab := Idmap.empty; - the_libtab := ModIdmap.empty; - the_sectab := ModIdmap.empty; - the_objtab := Idmap.empty + the_dirtab := ModIdmap.empty; + the_objtab := Idmap.empty; + the_globtab := Globtab.empty let freeze () = !the_ccitab, - !the_libtab, - !the_sectab, - !the_objtab + !the_dirtab, + !the_objtab, + !the_globtab -let unfreeze (mc,ml,ms,mo) = +let unfreeze (mc,md,mo,gt) = the_ccitab := mc; - the_libtab := ml; - the_sectab := ms; - the_objtab := mo + the_dirtab := md; + the_objtab := mo; + the_globtab := gt let _ = Summary.declare_summary "names" @@ -350,7 +403,3 @@ let _ = Summary.init_function = init; Summary.survive_section = false } -type strength = - | NotDeclare - | DischargeAt of dir_path * int - | NeverDischarge diff --git a/library/nametab.mli b/library/nametab.mli index d71e3e379..b64fe0d9d 100755 --- a/library/nametab.mli +++ b/library/nametab.mli @@ -12,44 +12,49 @@ open Util open Pp open Names +open Libnames (*i*) -(*s This module contains the table for globalization, which associates global - names (section paths) to qualified names. *) +(*s This module contains the tables for globalization, which + associates internal object references to qualified names (qualid). *) -type global_reference = - | VarRef of variable - | ConstRef of constant - | IndRef of inductive - | ConstructRef of constructor +(* Most functions in this module fall into one of the following categories: + \begin{itemize} + \item [push : visibility -> full_user_name -> object_reference -> unit] + + Registers the [object_reference] to be referred to by the + [full_user_name] (and its suffixes according to [visibility]) -(* Finds the real name of a global (e.g. fetch the constructor names - from the inductive name and constructor number) *) -val sp_of_global : Environ.env -> global_reference -> section_path - -type extended_global_reference = - | TrueGlobal of global_reference - | SyntacticDef of section_path + \item [exists : full_user_name -> bool] + + Is the [full_user_name] already atributed as an absolute user name + of some object? -(*s A [qualid] is a partially qualified ident; it includes fully - qualified names (= absolute names) and all intermediate partial - qualifications of absolute names, including single identifiers *) -type qualid + \item [locate : qualid -> object_reference] -val make_qualid : dir_path -> identifier -> qualid -val repr_qualid : qualid -> dir_path * identifier -val make_short_qualid : identifier -> qualid + Finds the object referred to by [qualid] or raises Not_found + + \item [name_of] : object_reference -> user_name -val string_of_qualid : qualid -> string -val pr_qualid : qualid -> std_ppcmds + The [user_name] can be for example the shortest non ambiguous [qualid] or + the [full_user_name] or [identifier]. Such a function can also have a + local context argument. +*) + + -val qualid_of_sp : section_path -> qualid +(* Finds the real name of a global (e.g. fetch the constructor names + from the inductive name and constructor number) *) +val sp_of_global : Sign.named_context option -> global_reference -> section_path +val sp_of_syntactic_definition : kernel_name -> section_path (* Turns an absolute name into a qualified name denoting the same name *) -val shortest_qualid_of_global : Environ.env -> global_reference -> qualid +val full_name : global_reference -> section_path +val shortest_qualid_of_global : Sign.named_context option -> global_reference -> qualid +val id_of_global : Sign.named_context option -> global_reference -> identifier (* Printing of global references using names as short as possible *) -val pr_global_env : Environ.env -> global_reference -> std_ppcmds +val pr_global_env : Sign.named_context option -> global_reference -> std_ppcmds exception GlobalizationError of qualid exception GlobalizationConstantError of qualid @@ -59,29 +64,29 @@ val error_global_not_found_loc : loc -> qualid -> 'a val error_global_not_found : qualid -> 'a val error_global_constant_not_found_loc : loc -> qualid -> 'a -(*s Register visibility of absolute paths by qualified names *) -val push : visibility:int -> section_path -> global_reference -> unit -val push_syntactic_definition : section_path -> unit +(*s Register visibility of things *) -(*s Register visibility of absolute paths by short names *) -val push_short_name_syntactic_definition : section_path -> unit -val push_short_name_object : section_path -> unit +(* The visibility can be registered either for all suffixes not + shorter then a given int - when the object is loaded inside a module, + or for a precise suffix, when the module containing (the module + containing ...) the object is open (imported) *) -(*s Register visibility of absolute paths by short names *) -val push_tactic_path : section_path -> unit -val locate_tactic : qualid -> section_path +type visibility = Until of int | Exactly of int -(*s Register visibility by all qualifications *) -val push_section : dir_path -> unit +(* is the short name visible? tests for 1 *) +val is_short_visible : visibility -> section_path -> bool -(* This should eventually disappear *) -val sp_of_id : identifier -> global_reference +val push : visibility -> section_path -> global_reference -> unit +val push_syntactic_definition : + visibility -> section_path -> kernel_name -> unit +val push_modtype : visibility -> section_path -> kernel_name -> unit +val push_dir : visibility -> dir_path -> global_dir_reference -> unit +val push_object : visibility -> section_path -> unit +val push_tactic : visibility -> section_path -> unit (*s The following functions perform globalization of qualified names *) (* This returns the section path of a constant or fails with [Not_found] *) -val constant_sp_of_id : identifier -> section_path - val locate : qualid -> global_reference (* This function is used to transform a qualified identifier into a @@ -97,11 +102,23 @@ val extended_locate : qualid -> extended_global_reference val locate_obj : qualid -> section_path val locate_constant : qualid -> constant +val locate_mind : qualid -> mutual_inductive val locate_section : qualid -> dir_path +val locate_modtype : qualid -> kernel_name +val locate_syntactic_definition : qualid -> kernel_name + +val locate_tactic : qualid -> section_path +val locate_dir : qualid -> global_dir_reference +val locate_module : qualid -> module_path (* [exists sp] tells if [sp] is already bound to a cci term *) val exists_cci : section_path -> bool -val exists_section : dir_path -> bool +val exists_modtype : section_path -> bool + +(* Those three functions are the same *) +val exists_dir : dir_path -> bool +val exists_section : dir_path -> bool (* deprecated *) +val exists_module : dir_path -> bool (* deprecated *) (*s Roots of the space of absolute names *) @@ -114,10 +131,10 @@ val absolute_reference : section_path -> global_reference one of its section/subsection *) val locate_in_absolute_module : dir_path -> identifier -> global_reference -val push_loaded_library : dir_path -> unit val locate_loaded_library : qualid -> dir_path -type strength = - | NotDeclare - | DischargeAt of dir_path * int - | NeverDischarge +type frozen + +val freeze : unit -> frozen +val unfreeze : frozen -> unit + diff --git a/library/summary.ml b/library/summary.ml index 210a1a81b..59560af22 100644 --- a/library/summary.ml +++ b/library/summary.ml @@ -10,7 +10,6 @@ open Pp open Util -open Names type 'a summary_declaration = { freeze_function : unit -> 'a; @@ -21,8 +20,8 @@ type 'a summary_declaration = { let summaries = (Hashtbl.create 17 : (string, Dyn.t summary_declaration) Hashtbl.t) -let declare_summary sumname sdecl = - let (infun,outfun) = Dyn.create (sumname^"-SUMMARY") in +let internal_declare_summary sumname sdecl = + let (infun,outfun) = Dyn.create sumname in let dyn_freeze () = infun (sdecl.freeze_function()) and dyn_unfreeze sum = sdecl.unfreeze_function (outfun sum) and dyn_init = sdecl.init_function in @@ -37,6 +36,14 @@ let declare_summary sumname sdecl = (str "Cannot declare a summary twice: " ++ str sumname); Hashtbl.add summaries sumname ddecl +let declare_summary sumname decl = + internal_declare_summary (sumname^"-SUMMARY") decl + +let envsummary = "Global environment SUMMARY" + +let declare_global_environment sdecl = + internal_declare_summary envsummary sdecl + type frozen = Dyn.t Stringmap.t let freeze_summaries () = @@ -62,5 +69,14 @@ let unfreeze_lost_summaries fs = with Not_found -> decl.init_function()) summaries +let unfreeze_other_summaries fs = + Hashtbl.iter + (fun id decl -> + try + if id <> envsummary then + decl.unfreeze_function (Stringmap.find id fs) + with Not_found -> decl.init_function()) + summaries + let init_summaries () = Hashtbl.iter (fun _ decl -> decl.init_function()) summaries diff --git a/library/summary.mli b/library/summary.mli index d381543b9..781b31144 100644 --- a/library/summary.mli +++ b/library/summary.mli @@ -10,6 +10,7 @@ (*i*) open Names +open Libnames (*i*) (* This module registers the declaration of global tables, which will be kept @@ -22,12 +23,14 @@ type 'a summary_declaration = { survive_section : bool } val declare_summary : string -> 'a summary_declaration -> unit +val declare_global_environment : 'a summary_declaration -> unit type frozen val freeze_summaries : unit -> frozen val unfreeze_summaries : frozen -> unit val unfreeze_lost_summaries : frozen -> unit +val unfreeze_other_summaries : frozen -> unit val init_summaries : unit -> unit diff --git a/parsing/.cvsignore b/parsing/.cvsignore index ffeaac2ef..a2e50565c 100644 --- a/parsing/.cvsignore +++ b/parsing/.cvsignore @@ -1,7 +1,6 @@ lexer.ml *.ppo pcoq.ml -extend.ml g_prim.ml q_coqast.ml g_basevernac.ml diff --git a/parsing/ast.ml b/parsing/ast.ml index 46fb041a7..a41e627e6 100755 --- a/parsing/ast.ml +++ b/parsing/ast.ml @@ -11,6 +11,7 @@ open Pp open Util open Names +open Libnames open Coqast open Tacexpr open Genarg @@ -111,15 +112,6 @@ type act = | ActCase of act * (pat * act) list | ActCaseList of act * (pat * act) list -(* -type act = - | Aast of typed_ast - | Aastlist of patlist - | Acase of act * (Coqast.t * act) list - | Atypedcase of act * (typed_ast * act) list - | Acaselist of act * (patlist * act) list -*) - (* values associated to variables *) type typed_ast = | AstListNode of Coqast.t list @@ -139,7 +131,7 @@ let rec print_ast ast = match ast with | Num(_,n) -> int n | Str(_,s) -> qs s - | Path(_,sl) -> str (string_of_path sl) + | Path(_,sl) -> str (string_of_kn sl) | Id (_,s) -> str "{" ++ str s ++ str "}" | Nvar(_,s) -> str (string_of_id s) | Nmeta(_,s) -> str s @@ -226,8 +218,8 @@ let coerce_to_id a = match coerce_to_var a with (loc ast,"Ast.coerce_to_id", str"This expression should be a simple identifier") -let coerce_qualid_to_id (loc,qid) = match Nametab.repr_qualid qid with - | dir, id when dir = Nameops.empty_dirpath -> id +let coerce_qualid_to_id (loc,qid) = match repr_qualid qid with + | dir, id when dir = empty_dirpath -> id | _ -> user_err_loc (loc, "Ast.coerce_qualid_to_id", str"This expression should be a simple identifier") @@ -749,3 +741,27 @@ and caselist vars etyp (pl,a) = (AstListPat apl,aa) let to_act_check_vars = act_of_ast + +let rec subst_astpat subst = function + | Pquote a -> Pquote (subst_ast subst a) + | Pmeta _ as p -> p + | Pnode (s,pl) -> Pnode (s,subst_astpatlist subst pl) + | Pslam (ido,p) -> Pslam (ido,subst_astpat subst p) + | Pmeta_slam (s,p) -> Pmeta_slam (s,subst_astpat subst p) + +and subst_astpatlist subst = function + | Pcons (p,pl) -> Pcons (subst_astpat subst p, subst_astpatlist subst pl) + | (Plmeta _ | Pnil) as pl -> pl + +let subst_pat subst = function + | AstListPat pl -> AstListPat (subst_astpatlist subst pl) + | PureAstPat p -> PureAstPat (subst_astpat subst p) + +let rec subst_act subst = function + | Act p -> Act (subst_pat subst p) + | ActCase (a,l) -> + ActCase (subst_act subst a, + List.map (fun (p,a) -> subst_pat subst p, subst_act subst a) l) + | ActCaseList (a,l) -> + ActCaseList (subst_act subst a, + List.map (fun (p,a) -> subst_pat subst p, subst_act subst a) l) diff --git a/parsing/ast.mli b/parsing/ast.mli index 4a9048256..07dbd06f2 100755 --- a/parsing/ast.mli +++ b/parsing/ast.mli @@ -11,6 +11,7 @@ (*i*) open Pp open Names +open Libnames open Coqast open Genarg (*i*) @@ -30,13 +31,13 @@ val nvar : identifier -> Coqast.t val ide : string -> Coqast.t val num : int -> Coqast.t val string : string -> Coqast.t -val path : section_path -> Coqast.t +val path : kernel_name -> Coqast.t val dynamic : Dyn.t -> Coqast.t val set_loc : Coqast.loc -> Coqast.t -> Coqast.t -val path_section : Coqast.loc -> section_path -> Coqast.t -val section_path : section_path -> section_path +val path_section : Coqast.loc -> kernel_name -> Coqast.t +val section_path : kernel_name -> kernel_name (* ast destructors *) val num_of_ast : Coqast.t -> int @@ -100,7 +101,7 @@ val coerce_to_var : Coqast.t -> Coqast.t val coerce_to_id : Coqast.t -> identifier -val coerce_qualid_to_id : Nametab.qualid Util.located -> identifier +val coerce_qualid_to_id : qualid Util.located -> identifier val coerce_reference_to_id : Tacexpr.reference_expr -> identifier @@ -151,3 +152,6 @@ val to_pat : entry_env -> Coqast.t -> (astpat * entry_env) val eval_act : Coqast.loc -> env -> act -> typed_ast val to_act_check_vars : entry_env -> ast_action_type -> grammar_action -> act + +val subst_astpat : Names.substitution -> astpat -> astpat +val subst_act : Names.substitution -> act -> act diff --git a/parsing/astmod.ml b/parsing/astmod.ml new file mode 100644 index 000000000..acf0de62e --- /dev/null +++ b/parsing/astmod.ml @@ -0,0 +1,147 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* mp + | h::tl -> make_mp (MPdot(mp, label_of_id h)) tl + +(* +(* Since module components are not put in the nametab we try to locate +the module prefix *) +exception BadRef + +let lookup_qualid (modtype:bool) qid = + let rec make_mp mp = function + [] -> mp + | h::tl -> make_mp (MPdot(mp, label_of_id h)) tl + in + let rec find_module_prefix dir n = + if n<0 then raise Not_found; + let dir',dir'' = list_chop n dir in + let id',dir''' = + match dir'' with + | hd::tl -> hd,tl + | _ -> anomaly "This list should not be empty!" + in + let qid' = make_qualid dir' id' in + try + match Nametab.locate qid' with + | ModRef mp -> mp,dir''' + | _ -> raise BadRef + with + Not_found -> find_module_prefix dir (pred n) + in + try Nametab.locate qid + with Not_found -> + let (dir,id) = repr_qualid qid in + let pref_mp,dir' = find_module_prefix dir (List.length dir - 1) in + let mp = + List.fold_left (fun mp id -> MPdot (mp, label_of_id id)) pref_mp dir' + in + if modtype then + ModTypeRef (make_ln mp (label_of_id id)) + else + ModRef (MPdot (mp,label_of_id id)) + +*) + +(* Search for the head of [qid] in [binders]. + If found, returns the module_path/kernel_name created from the dirpath + and the basename. Searches Nametab otherwise. +*) + +let lookup_module binders qid = + try + let dir,id = repr_qualid qid in + (* dirpath in natural order *) + let idl = List.rev (id::repr_dirpath dir) in + let top_mp = MPbound (fst (List.assoc (List.hd idl) binders)) in + make_mp top_mp (List.tl idl) + with + Not_found -> Nametab.locate_module qid + +let lookup_modtype binders qid = + try + let dir,id = repr_qualid qid in + (* dirpath in natural order *) + match List.rev (repr_dirpath dir) with + | hd::tl -> + let top_mp = MPbound (fst (List.assoc hd binders)) in + let mp = make_mp top_mp tl in + make_kn mp empty_dirpath (label_of_id id) + | [] -> raise Not_found + (* may-be a module, but not a module type!*) + with + Not_found -> Nametab.locate_modtype qid + +let transl_modtype binders = function + | Node(loc,"MODTYPEQID",qid_ast) -> begin match qid_ast with + | [Node (loc, "QUALID", astl)] -> + let qid = interp_qualid astl in begin + try + MTEident (lookup_modtype binders qid) + with + | Not_found -> + Modops.error_not_a_modtype (*loc*) (string_of_qualid qid) + end + | _ -> anomaly "QUALID expected" + end + | _ -> anomaly "TODO: transl_modtype: I can handle qualid module types only" + + +let transl_binder binders (idl,mty_ast) = + let mte = transl_modtype binders mty_ast in + let add_one binders id = + if List.mem_assoc id binders then + error "Two identical module binders..." + else + let mbid = make_mbid (Lib.module_dp()) (string_of_id id) in + (id,(mbid,mte))::binders + in + List.fold_left add_one binders idl + +let transl_binders = List.fold_left transl_binder + + +let rec transl_modexpr binders = function + | Node(loc,"MODEXPRQID",qid_ast) -> begin match qid_ast with + | [Node (loc, "QUALID", astl)] -> + let qid = interp_qualid astl in begin + try + MEident (lookup_module binders qid) + with + | Not_found -> + Modops.error_not_a_module (*loc*) (string_of_qualid qid) + end + | _ -> anomaly "QUALID expected" + end + | Node(_,"MODEXPRAPP",[ast1;ast2]) -> + let me1 = transl_modexpr binders ast1 in + let me2 = transl_modexpr binders ast2 in + MEapply(me1,me2) + | Node(_,"MODEXPRAPP",_) -> + anomaly "transl_modexpr: MODEXPRAPP must have two arguments" + | _ -> anomaly "transl_modexpr: I can handle MODEXPRQID or MODEXPRAPP only..." + + +let interp_module_decl evd env args_ast mty_ast_o mexpr_ast_o = + let binders = transl_binders [] args_ast in + let mty_o = option_app (transl_modtype binders) mty_ast_o in + let mexpr_o = option_app (transl_modexpr binders) mexpr_ast_o in + (List.rev binders, mty_o, mexpr_o) + diff --git a/parsing/astmod.mli b/parsing/astmod.mli new file mode 100644 index 000000000..158f40e89 --- /dev/null +++ b/parsing/astmod.mli @@ -0,0 +1,30 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* env -> + (identifier list * Coqast.t) list -> + Coqast.t option -> + Coqast.t option + -> + (identifier * (mod_bound_id * module_type_entry)) list * + module_type_entry option * + module_expr option + diff --git a/parsing/astterm.ml b/parsing/astterm.ml index 0985d309f..89025de4f 100644 --- a/parsing/astterm.ml +++ b/parsing/astterm.ml @@ -19,6 +19,7 @@ open Termops open Environ open Evd open Reductionops +open Libnames open Impargs open Declare open Rawterm @@ -122,7 +123,7 @@ let add_glob loc ref = in let s = string_of_dirpath (find_module dir) in i*) - let sp = Nametab.sp_of_global (Global.env ()) ref in + let sp = Nametab.sp_of_global None ref in let id = let _,id = repr_path sp in string_of_id id in let dp = string_of_dirpath (Declare.library_part ref) in dump_string (Printf.sprintf "R%d %s.%s\n" (fst loc) dp id) @@ -229,14 +230,15 @@ let maybe_constructor allow_var env = function if !dump then add_glob loc (ConstructRef c); ConstrPat (loc,c) - | Path(loc,sp) -> - (match absolute_reference sp with - | ConstructRef c as r -> - if !dump then add_glob loc (ConstructRef c); - ConstrPat (loc,c) - | _ -> - error ("Unknown absolute constructor name: "^(string_of_path sp))) - + | Path(loc,kn) -> + (let dir,id = decode_kn kn in + let sp = make_path dir id in + match absolute_reference sp with + | ConstructRef c as r -> + if !dump then add_glob loc (ConstructRef c); + ConstrPat (loc,c) + | _ -> + error ("Unknown absolute constructor name: "^(string_of_path sp))) | Node(loc,("CONST"|"SECVAR"|"EVAR"|"MUTIND"|"SYNCONST" as key), l) -> user_err_loc (loc,"ast_to_pattern", (str "Found a pattern involving global references which are not constructors" @@ -622,7 +624,7 @@ let ast_of_syndef loc sp = Node (loc, "SYNCONST", [path_section loc sp]) let ast_of_extended_ref_loc loc = function | TrueGlobal ref -> ast_of_ref_loc loc ref - | SyntacticDef sp -> ast_of_syndef loc sp + | SyntacticDef kn -> ast_of_syndef loc kn let ast_of_extended_ref = ast_of_extended_ref_loc dummy_loc diff --git a/parsing/astterm.mli b/parsing/astterm.mli index 3074ff665..3a871cd53 100644 --- a/parsing/astterm.mli +++ b/parsing/astterm.mli @@ -14,6 +14,7 @@ open Term open Sign open Evd open Environ +open Libnames open Rawterm open Pattern (*i*) @@ -25,7 +26,7 @@ val constrIn : constr -> Coqast.t val constrOut : Coqast.t -> constr (* Interprets global names, including syntactic defs and section variables *) -val interp_global_constr : env -> Nametab.qualid Util.located -> constr +val interp_global_constr : env -> qualid Util.located -> constr val interp_rawconstr : evar_map -> env -> Coqast.t -> rawconstr val interp_rawconstr_gen : @@ -81,13 +82,13 @@ val interp_constrpattern : bound idents in grammar or pretty-printing rules) *) val globalize_constr : Coqast.t -> Coqast.t val globalize_ast : Coqast.t -> Coqast.t -val globalize_qualid : Nametab.qualid Util.located -> Coqast.t +val globalize_qualid : qualid Util.located -> Coqast.t -val ast_of_extended_ref_loc : loc -> Nametab.extended_global_reference -> Coqast.t +val ast_of_extended_ref_loc : loc -> Libnames.extended_global_reference -> Coqast.t (* This transforms args of a qualid keyword into a qualified ident *) (* it does no relocation *) -val interp_qualid : Coqast.t list -> Nametab.qualid +val interp_qualid : Coqast.t list -> qualid (*i Translation rules from V6 to V7: diff --git a/parsing/coqast.ml b/parsing/coqast.ml index 9b65bdfb1..e9b9e3c14 100644 --- a/parsing/coqast.ml +++ b/parsing/coqast.ml @@ -10,6 +10,7 @@ (*i*) open Names +open Libnames (*i*) type loc = int * int @@ -23,7 +24,7 @@ type t = | Num of loc * int | Str of loc * string | Id of loc * string - | Path of loc * section_path + | Path of loc * kernel_name | Dynamic of loc * Dyn.t type the_coq_ast = t @@ -62,7 +63,7 @@ module Hast = Hashcons.Make( type u = (the_coq_ast -> the_coq_ast) * ((loc -> loc) * (string -> string) - * (identifier -> identifier) * (section_path -> section_path)) + * (identifier -> identifier) * (kernel_name -> kernel_name)) let hash_sub (hast,(hloc,hstr,hid,hsp)) = function | Node(l,s,al) -> Node(hloc l, hstr s, List.map hast al) | Nmeta(l,s) -> Nmeta(hloc l, hstr s) @@ -98,6 +99,32 @@ let hcons_ast (hstr,hid,hpath) = let hast = Hashcons.recursive_hcons Hast.f (hloc,hstr,hid,hpath) in (hast,hloc) +let rec subst_ast subst ast = match ast with + | Node (l,s,astl) -> + let astl' = Util.list_smartmap (subst_ast subst) astl in + if astl' == astl then ast else + Node (l,s,astl') + | Slam (l,ido,ast1) -> + let ast1' = subst_ast subst ast1 in + if ast1' == ast1 then ast else + Slam (l,ido,ast1') + | Smetalam (l,s,ast1) -> + let ast1' = subst_ast subst ast1 in + if ast1' == ast1 then ast else + Smetalam (l,s,ast1') + | Path (loc,kn) -> + let kn' = Names.subst_kn subst kn in + if kn' == kn then ast else + Path(loc,kn') + | Nmeta _ + | Nvar _ -> ast + | Num _ + | Str _ + | Id _ + | Dynamic _ -> ast + + + (* type 'vernac_ast raw_typed_ast = | PureAstNode of t diff --git a/parsing/coqast.mli b/parsing/coqast.mli index b5ad45c74..a304aa06b 100644 --- a/parsing/coqast.mli +++ b/parsing/coqast.mli @@ -10,6 +10,7 @@ (*i*) open Names +open Libnames (*i*) (* Abstract syntax trees. *) @@ -25,7 +26,7 @@ type t = | Num of loc * int | Str of loc * string | Id of loc * string - | Path of loc * section_path + | Path of loc * kernel_name | Dynamic of loc * Dyn.t (* returns the list of metas occuring in the ast *) @@ -38,9 +39,11 @@ val subst_meta : (int * t) list -> t -> t (* hash-consing function *) val hcons_ast: (string -> string) * (Names.identifier -> Names.identifier) - * (section_path -> section_path) + * (kernel_name -> kernel_name) -> (t -> t) * (loc -> loc) +val subst_ast: Names.substitution -> t -> t + (* val map_tactic_expr : (t -> t) -> (tactic_expr -> tactic_expr) -> tactic_expr -> tactic_expr val fold_tactic_expr : diff --git a/parsing/coqlib.ml b/parsing/coqlib.ml index 07e439af8..52311d188 100644 --- a/parsing/coqlib.ml +++ b/parsing/coqlib.ml @@ -11,6 +11,7 @@ open Util open Names open Term +open Libnames open Declare open Pattern open Nametab @@ -26,6 +27,9 @@ let logic_type_module = make_dir ["Coq";"Init";"Logic_Type"] let datatypes_module = make_dir ["Coq";"Init";"Datatypes"] let arith_module = make_dir ["Coq";"Arith";"Arith"] +(* TODO: temporary hack *) +let make_path dir id = Libnames.encode_kn dir id + let nat_path = make_path datatypes_module (id_of_string "nat") let glob_nat = IndRef (nat_path,0) diff --git a/parsing/coqlib.mli b/parsing/coqlib.mli index 0b5bb0ec5..8eedab050 100644 --- a/parsing/coqlib.mli +++ b/parsing/coqlib.mli @@ -10,6 +10,7 @@ (*i*) open Names +open Libnames open Nametab open Term open Pattern @@ -21,8 +22,8 @@ open Pattern (*s Global references *) (* Modules *) -val logic_module : Names.dir_path -val logic_type_module : Names.dir_path +val logic_module : dir_path +val logic_type_module : dir_path (* Natural numbers *) val glob_nat : global_reference diff --git a/parsing/egrammar.ml b/parsing/egrammar.ml index 97cfd77b1..d665ad575 100644 --- a/parsing/egrammar.ml +++ b/parsing/egrammar.ml @@ -24,6 +24,10 @@ type all_grammar_command = (string * (string * grammar_production list) * Tacexpr.raw_tactic_expr) list +let subst_all_grammar_command subst = function + | AstGrammar gc -> AstGrammar (subst_grammar_command subst gc) + | TacticGrammar g -> TacticGrammar g (* TODO ... *) + let (grammar_state : all_grammar_command list ref) = ref [] (* Interpretation of the right hand side of grammar rules *) diff --git a/parsing/egrammar.mli b/parsing/egrammar.mli index e7d2cde12..9f9e34846 100644 --- a/parsing/egrammar.mli +++ b/parsing/egrammar.mli @@ -38,3 +38,6 @@ val extend_tactic_grammar : val extend_vernac_command_grammar : string -> (string * grammar_tactic_production list) list -> unit + +val subst_all_grammar_command : + Names.substitution -> all_grammar_command -> all_grammar_command diff --git a/parsing/esyntax.ml b/parsing/esyntax.ml index 608868ca6..b3d027802 100644 --- a/parsing/esyntax.ml +++ b/parsing/esyntax.ml @@ -10,6 +10,8 @@ open Pp open Util +open Names +open Libnames open Coqast open Ast open Extend @@ -195,14 +197,14 @@ let make_hunks (lp,rp) s e1 e2 = let build_syntax (ref,e1,e2,assoc) = let sp = match ref with - | TrueGlobal r -> Nametab.sp_of_global (Global.env()) r - | SyntacticDef sp -> sp in + | TrueGlobal r -> Nametab.sp_of_global None r + | SyntacticDef kn -> Nametab.sp_of_syntactic_definition kn in let rec find_symbol = function | [] -> let s = match ref with | TrueGlobal r -> - string_of_qualid (shortest_qualid_of_global (Global.env()) r) - | SyntacticDef sp -> string_of_path sp in + string_of_qualid (shortest_qualid_of_global None r) + | SyntacticDef _ -> string_of_path sp in UNP_BOX (PpHOVB 0, [RO "("; RO s; UNP_BRK (1, 1); PH (meta_pattern e1, None, E); UNP_BRK (1, 1); PH (meta_pattern e2, None, E); RO ")"]) diff --git a/parsing/esyntax.mli b/parsing/esyntax.mli index 5ddadb5bc..8e3445ed7 100644 --- a/parsing/esyntax.mli +++ b/parsing/esyntax.mli @@ -42,7 +42,7 @@ module Ppprim : val add : string * t -> unit end -val declare_infix_symbol : Names.section_path -> string -> unit +val declare_infix_symbol : Libnames.section_path -> string -> unit (* Generic printing functions *) val token_printer: std_printer -> std_printer diff --git a/parsing/extend.ml b/parsing/extend.ml index bc4ad211b..7c4c400eb 100644 --- a/parsing/extend.ml +++ b/parsing/extend.ml @@ -50,6 +50,16 @@ type raw_grammar_rule = string * grammar_production list * grammar_action type raw_grammar_entry = string * ast_action_type * grammar_associativity * raw_grammar_rule list +let subst_grammar_rule subst gr = + { gr with gr_action = subst_act subst gr.gr_action } + +let subst_grammar_entry subst ge = + { ge with gl_rules = List.map (subst_grammar_rule subst) ge.gl_rules } + +let subst_grammar_command subst gc = + { gc with gc_entries = List.map (subst_grammar_entry subst) gc.gc_entries } + + (*s Terminal symbols interpretation *) let is_ident_not_keyword s = @@ -177,9 +187,28 @@ type 'pat unparsing_hunk = | UNP_TBRK of int * int | UNP_TAB | UNP_FNL - | UNP_INFIX of Nametab.extended_global_reference * string * string * + | UNP_INFIX of Libnames.extended_global_reference * string * string * (parenRelation * parenRelation) +let rec subst_hunk subst_pat subst hunk = match hunk with + | PH (pat,so,pr) -> + let pat' = subst_pat subst pat in + if pat'==pat then hunk else + PH (pat',so,pr) + | RO _ -> hunk + | UNP_BOX (ppbox, hunkl) -> + let hunkl' = list_smartmap (subst_hunk subst_pat subst) hunkl in + if hunkl' == hunkl then hunk else + UNP_BOX (ppbox, hunkl') + | UNP_BRK _ + | UNP_TBRK _ + | UNP_TAB + | UNP_FNL -> hunk + | UNP_INFIX (ref,s1,s2,prs) -> + let ref' = Libnames.subst_ext subst ref in + if ref' == ref then hunk else + UNP_INFIX (ref',s1,s2,prs) + (* Checks if the precedence of the parent printer (None means the highest precedence), and the child's one, follow the given relation. *) @@ -213,10 +242,29 @@ type 'pat syntax_entry = { syn_astpat : 'pat; syn_hunks : 'pat unparsing_hunk list } +let subst_syntax_entry subst_pat subst sentry = + let syn_astpat' = subst_pat subst sentry.syn_astpat in + let syn_hunks' = list_smartmap (subst_hunk subst_pat subst) sentry.syn_hunks + in + if syn_astpat' == sentry.syn_astpat + && syn_hunks' == sentry.syn_hunks then sentry + else + { sentry with + syn_astpat = syn_astpat' ; + syn_hunks = syn_hunks' ; + } + type 'pat syntax_command = { sc_univ : string; sc_entries : 'pat syntax_entry list } +let subst_syntax_command subst_pat subst scomm = + let sc_entries' = + list_smartmap (subst_syntax_entry subst_pat subst) scomm.sc_entries + in + if sc_entries' == scomm.sc_entries then scomm else + { scomm with sc_entries = sc_entries' } + type syntax_rule = string * Coqast.t * Coqast.t unparsing_hunk list type syntax_entry_ast = precedence * syntax_rule list diff --git a/parsing/extend.mli b/parsing/extend.mli index e80f011d3..8734e3452 100644 --- a/parsing/extend.mli +++ b/parsing/extend.mli @@ -1,3 +1,15 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* string * string val rename_command : string -> string +val subst_grammar_command : + Names.substitution -> grammar_command -> grammar_command + (*s Pretty-print. *) (* Dealing with precedences *) @@ -67,9 +82,14 @@ type 'pat unparsing_hunk = | UNP_TBRK of int * int | UNP_TAB | UNP_FNL - | UNP_INFIX of Nametab.extended_global_reference * string * string * + | UNP_INFIX of Libnames.extended_global_reference * string * string * (parenRelation * parenRelation) +(*val subst_unparsing_hunk : + Names.substitution -> (Names.substitution -> 'pat -> 'pat) -> + 'pat unparsing_hunk -> 'pat unparsing_hunk +*) + (* Checks if the precedence of the parent printer (None means the highest precedence), and the child's one, follow the given relation. *) @@ -86,10 +106,19 @@ type 'pat syntax_entry = { syn_astpat : 'pat; syn_hunks : 'pat unparsing_hunk list } +val subst_syntax_entry : + (Names.substitution -> 'pat -> 'pat) -> + Names.substitution -> 'pat syntax_entry -> 'pat syntax_entry + + type 'pat syntax_command = { sc_univ : string; sc_entries : 'pat syntax_entry list } +val subst_syntax_command : + (Names.substitution -> 'pat -> 'pat) -> + Names.substitution -> 'pat syntax_command -> 'pat syntax_command + type syntax_rule = string * Coqast.t * Coqast.t unparsing_hunk list type syntax_entry_ast = precedence * syntax_rule list diff --git a/parsing/g_module.ml4 b/parsing/g_module.ml4 new file mode 100644 index 000000000..01e7256ef --- /dev/null +++ b/parsing/g_module.ml4 @@ -0,0 +1,81 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* id ] ] + ; + + ident_comma_list_tail: + [ [ ","; idl = LIST0 ident SEP "," -> idl + | -> [] ] ] + ; + + qualid: + [ [ id = Prim.var; l = fields -> <:ast< (QUALID $id ($LIST $l)) >> + | id = Prim.var -> <:ast< (QUALID $id) >> + ] ] + ; + fields: + [ [ id = FIELD; l = fields -> <:ast< ($VAR $id) >> :: l + | id = FIELD -> [ <:ast< ($VAR $id) >> ] + ] ] + ; + + vardecls: (* This is interpreted by Pcoq.abstract_binder *) + [ [ id = ident; idl = ident_comma_list_tail; + ":"; mty = module_type -> + <:ast< (BINDER $mty $id ($LIST $idl)) >> ] ] + ; + + ne_vardecls_list: + [ [ id = vardecls; ";"; idl = ne_vardecls_list -> id :: idl + | id = vardecls -> [id] ] ] + ; + + rawbinders: + [ [ "["; bl = ne_vardecls_list; "]" -> bl ] ] + ; + + ne_binders_list: + [ [ bl = rawbinders; bll = ne_binders_list -> bl @ bll + | bl = rawbinders -> bl ] ] + ; + + module_expr: + [ [ qid = qualid -> <:ast< (MODEXPRQID $qid) >> + | me1 = module_expr; me2 = module_expr -> + <:ast< (MODEXPRAPP $me1 $me2) >> + | "("; me = module_expr; ")" -> + me +(* ... *) + ] ] + ; + + module_type: + [ [ qid = qualid -> <:ast< (MODTYPEQID $qid) >> +(* ... *) + ] ] + ; +END diff --git a/parsing/g_prim.ml4 b/parsing/g_prim.ml4 index 51696ad03..d279499ab 100644 --- a/parsing/g_prim.ml4 +++ b/parsing/g_prim.ml4 @@ -15,6 +15,7 @@ camlp4o pa_ifdef.cmo pa_extend.cmo pr_o.cmo pr_extend.cmo -quotify -DQuotify -im open Coqast open Pcoq open Names +open Libnames ifdef Quotify then open Qast @@ -26,7 +27,7 @@ open Prim ifdef Quotify then module Prelude = struct let local_id_of_string s = Apply ("Names","id_of_string", [s]) -let local_make_dirpath l = Qast.Apply ("Nametab","make_dirpath",[l]) +let local_make_dirpath l = Qast.Apply ("Names","make_dirpath",[l]) let local_make_posint s = s let local_make_negint s = Apply ("Pervasives","~-", [s]) let local_make_qualid l id' = @@ -34,7 +35,7 @@ let local_make_qualid l id' = let local_make_short_qualid id = Qast.Apply ("Nametab","make_short_qualid",[id]) let local_make_path l id' = - Qast.Apply ("Names","make_path", [local_make_dirpath l;id']) + Qast.Apply ("Libnames","encode_kn", [local_make_dirpath l;id']) let local_make_binding loc a b = match a with | Qast.Node ("Nvar", [_;id]) -> @@ -56,7 +57,7 @@ let local_make_qualid l id' = make_qualid (local_make_dirpath l) id' let local_make_short_qualid id = make_short_qualid id let local_make_posint = int_of_string let local_make_negint n = - int_of_string n -let local_make_path l a = make_path (local_make_dirpath l) a +let local_make_path l a = encode_kn (local_make_dirpath l) a let local_make_binding loc a b = match a with | Nvar (_,id) -> Slam(loc,Some id,b) diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index 243f72bf1..c9a0d69c7 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -297,10 +297,46 @@ ident_comma_list_tail: [ [ (* Sections *) IDENT "Section"; id = ident -> VernacBeginSection id - | IDENT "Chapter"; id = ident -> VernacBeginSection id - | IDENT "Module"; id = ident -> - warning "Module is currently unsupported"; VernacNop - | IDENT "End"; id = ident -> VernacEndSection id + | IDENT "Chapter"; id = ident -> VernacBeginSection id ] ] +(* | IDENT "Module"; id = ident -> + warning "Module is currently unsupported"; VernacNop *) + ; + + module_vardecls: (* This is interpreted by Pcoq.abstract_binder *) + [ [ id = ident; idl = ident_comma_list_tail; + ":"; mty = Module.module_type -> + (id::idl,mty) ] ] + ; + module_binders: + [ [ "["; bl = LIST1 module_vardecls SEP ";"; "]" -> bl ] ] + ; + module_binders_list: + [ [ bls = LIST0 module_binders -> List.flatten bls ] ] + ; + of_module_type: + [ [ ":"; mty = Module.module_type -> mty ] ] + ; + is_module_type: + [ [ ":="; mty = Module.module_type -> mty ] ] + ; + is_module_expr: + [ [ ":="; mexpr = Module.module_expr -> mexpr ] ] + ; + gallina_ext: + [ [ + (* Interactive module declaration *) + IDENT "Module"; id = ident; bl = module_binders_list; + mty_o = OPT of_module_type; mexpr_o = OPT is_module_expr -> + VernacDeclareModule id bl mty_o mexpr_o + + | IDENT "Module"; "Type"; id = ident; + bl = module_binders_list; mty_o = OPT is_module_type -> + VernacDeclareModuleType id bl mty_o + + (* This end a Section a Module or a Module Type *) + + | IDENT "End"; id = ident -> VernacEndSegment id + (* Transparent and Opaque *) | IDENT "Transparent"; l = LIST1 qualid -> VernacSetOpacity (false, l) @@ -333,13 +369,13 @@ ident_comma_list_tail: VernacIdentityCoercion (Declare.make_strength_0 (), f, s, t) | IDENT "Identity"; IDENT "Coercion"; f = Prim.ident; ":"; s = class_rawexpr; ">->"; t = class_rawexpr -> - VernacIdentityCoercion (Nametab.NeverDischarge, f, s, t) + VernacIdentityCoercion (Libnames.NeverDischarge, f, s, t) | IDENT "Coercion"; IDENT "Local"; qid = qualid; ":"; s = class_rawexpr; ">->"; t = class_rawexpr -> VernacCoercion (Declare.make_strength_0 (), qid, s, t) | IDENT "Coercion"; qid = qualid; ":"; s = class_rawexpr; ">->"; t = class_rawexpr -> - VernacCoercion (Nametab.NeverDischarge, qid, s, t) + VernacCoercion (Libnames.NeverDischarge, qid, s, t) | IDENT "Class"; IDENT "Local"; c = qualid -> Pp.warning "Class is obsolete"; VernacNop | IDENT "Class"; c = qualid -> diff --git a/parsing/genarg.mli b/parsing/genarg.mli index 70e67330e..2991d237a 100644 --- a/parsing/genarg.mli +++ b/parsing/genarg.mli @@ -11,7 +11,7 @@ open Util open Names open Term -open Nametab +open Libnames open Rawterm type 'a or_var = ArgArg of 'a | ArgVar of loc * identifier @@ -82,7 +82,7 @@ val rawwit_pre_ident : (string,'co,'ta) abstract_argument_type val wit_pre_ident : (string,'co,'ta) abstract_argument_type val rawwit_qualid : (qualid located,constr_ast,'ta) abstract_argument_type -val wit_qualid : (Nametab.global_reference,constr,'ta) abstract_argument_type +val wit_qualid : (global_reference,constr,'ta) abstract_argument_type val rawwit_quant_hyp : (quantified_hypothesis,'co,'ta) abstract_argument_type val wit_quant_hyp : (quantified_hypothesis,'co,'ta) abstract_argument_type @@ -99,7 +99,7 @@ val wit_casted_open_constr : (open_constr,constr,'ta) abstract_argument_type val rawwit_constr_with_bindings : (constr_ast with_bindings,constr_ast,'ta) abstract_argument_type val wit_constr_with_bindings : (constr with_bindings,constr,'ta) abstract_argument_type -val rawwit_red_expr : ((constr_ast,Nametab.qualid or_metanum) red_expr_gen,constr_ast,'ta) abstract_argument_type +val rawwit_red_expr : ((constr_ast,qualid or_metanum) red_expr_gen,constr_ast,'ta) abstract_argument_type val wit_red_expr : ((constr,Closure.evaluable_global_reference) red_expr_gen,constr,'ta) abstract_argument_type (* TODO: transformer tactic en extra arg *) diff --git a/parsing/pcoq.ml4 b/parsing/pcoq.ml4 index d26613b10..2d444fcbb 100644 --- a/parsing/pcoq.ml4 +++ b/parsing/pcoq.ml4 @@ -222,6 +222,7 @@ let create_univ s = let uprim = create_univ "prim" let uconstr = create_univ "constr" +let umodule = create_univ "module" let utactic = create_univ "tactic" let uvernac = create_univ "vernac" @@ -398,6 +399,18 @@ module Constr = end +module Module = + struct + let gec = make_entry umodule inPureAstType + let gec_list = make_entry umodule inAstListType + + let ident = gec "ident" + let qualid = gec_list "qualid" + let ne_binders_list = gec_list "ne_binders_list" + let module_expr = gec "module_expr" + let module_type = gec "module_type" + end + module Tactic = struct let gec = make_entry utactic inPureAstType diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli index 61e2f9771..628f4035e 100644 --- a/parsing/pcoq.mli +++ b/parsing/pcoq.mli @@ -102,7 +102,7 @@ module Prim : sig open Util open Names - open Nametab + open Libnames val preident : string Gram.Entry.e val ident : identifier Gram.Entry.e val rawident : identifier located Gram.Entry.e @@ -150,6 +150,13 @@ module Constr : val ne_binders_list : Coqast.t list Gram.Entry.e end +module Module : + sig + val ne_binders_list : Coqast.t list Gram.Entry.e + val module_expr : Coqast.t Gram.Entry.e + val module_type : Coqast.t Gram.Entry.e + end + module Tactic : sig open Rawterm @@ -168,7 +175,7 @@ module Tactic : module Vernac_ : sig open Util - open Nametab + open Libnames val thm_token : theorem_kind Gram.Entry.e val class_rawexpr : class_rawexpr Gram.Entry.e val gallina : vernac_expr Gram.Entry.e diff --git a/parsing/ppconstr.ml b/parsing/ppconstr.ml index ad120fe19..fbee5ff27 100644 --- a/parsing/ppconstr.ml +++ b/parsing/ppconstr.ml @@ -16,33 +16,34 @@ open Pp open Nametab open Names open Nameops +open Libnames open Coqast open Extend open Util let dfltpr ast = (str"#GENTERM " ++ print_ast ast);; -let pr_global ref = pr_global_env (Global.env()) ref +let pr_global ref = pr_global_env None ref -let global_const_name sp = - try pr_global (ConstRef sp) +let global_const_name kn = + try pr_global (ConstRef kn) with Not_found -> (* May happen in debug *) - (str ("CONST("^(string_of_path sp)^")")) + (str ("CONST("^(string_of_kn kn)^")")) let global_var_name id = try pr_global (VarRef id) with Not_found -> (* May happen in debug *) (str ("SECVAR("^(string_of_id id)^")")) -let global_ind_name (sp,tyi) = - try pr_global (IndRef (sp,tyi)) +let global_ind_name (kn,tyi) = + try pr_global (IndRef (kn,tyi)) with Not_found -> (* May happen in debug *) - (str ("IND("^(string_of_path sp)^","^(string_of_int tyi)^")")) + (str ("IND("^(string_of_kn kn)^","^(string_of_int tyi)^")")) -let global_constr_name ((sp,tyi),i) = - try pr_global (ConstructRef ((sp,tyi),i)) +let global_constr_name ((kn,tyi),i) = + try pr_global (ConstructRef ((kn,tyi),i)) with Not_found -> (* May happen in debug *) - (str ("CONSTRUCT("^(string_of_path sp)^","^(string_of_int tyi) + (str ("CONSTRUCT("^(string_of_kn kn)^","^(string_of_int tyi) ^","^(string_of_int i)^")")) let globpr gt = match gt with @@ -89,7 +90,7 @@ let pr_constr = gentermpr let pr_pattern = gentermpr -let pr_qualid qid = str (Nametab.string_of_qualid qid) +let pr_qualid qid = str (string_of_qualid qid) open Rawterm diff --git a/parsing/ppconstr.mli b/parsing/ppconstr.mli index d9b8b6ea5..04225ef7a 100644 --- a/parsing/ppconstr.mli +++ b/parsing/ppconstr.mli @@ -14,7 +14,7 @@ open Pp open Environ open Term -open Nametab +open Libnames open Pcoq open Rawterm open Extend diff --git a/parsing/pptactic.ml b/parsing/pptactic.ml index 80cdf953b..8fa2326c4 100644 --- a/parsing/pptactic.ml +++ b/parsing/pptactic.ml @@ -245,9 +245,9 @@ open Closure let pr_evaluable_reference = function | EvalVarRef id -> pr_id id - | EvalConstRef sp -> pr_global (Nametab.ConstRef sp) + | EvalConstRef sp -> pr_global (Libnames.ConstRef sp) -let pr_inductive ind = pr_global (Nametab.IndRef ind) +let pr_inductive ind = pr_global (Libnames.IndRef ind) let rec pr_generic prtac x = match Genarg.genarg_tag x with diff --git a/parsing/prettyp.ml b/parsing/prettyp.ml index c7da75c36..6554659f3 100644 --- a/parsing/prettyp.ml +++ b/parsing/prettyp.ml @@ -25,6 +25,7 @@ open Declare open Impargs open Libobject open Printer +open Libnames open Nametab let print_basename sp = pr_global (ConstRef sp) @@ -291,20 +292,20 @@ let print_constant with_values sep sp = let print_inductive sp = (print_mutual sp ++ fnl ()) -let print_syntactic_def sep sp = - let id = basename sp in - let c = Syntax_def.search_syntactic_definition sp in - (str" Syntactic Definition " ++ pr_id id ++ str sep ++ pr_rawterm c ++ fnl ()) +let print_syntactic_def sep kn = + let _,_,l = repr_kn kn in + let c = Syntax_def.search_syntactic_definition kn in + (str" Syntactic Definition " ++ pr_lab l ++ str sep ++ pr_rawterm c ++ fnl ()) -let print_leaf_entry with_values sep (sp,lobj) = +let print_leaf_entry with_values sep ((sp,kn as oname),lobj) = let tag = object_tag lobj in - match (sp,tag) with + match (oname,tag) with | (_,"VARIABLE") -> print_section_variable (basename sp) | (_,("CONSTANT"|"PARAMETER")) -> - print_constant with_values sep sp + print_constant with_values sep kn | (_,"INDUCTIVE") -> - print_inductive sp + print_inductive kn | (_,"AUTOHINT") -> (* (str" Hint Marker" ++ fnl ())*) (mt ()) @@ -312,7 +313,7 @@ let print_leaf_entry with_values sep (sp,lobj) = (* (str" Grammar Marker" ++ fnl ())*) (mt ()) | (_,"SYNTAXCONSTANT") -> - print_syntactic_def sep sp + print_syntactic_def sep kn | (_,"PPSYNTAX") -> (* (str" Syntax Marker" ++ fnl ())*) (mt ()) @@ -340,15 +341,20 @@ let print_leaf_entry with_values sep (sp,lobj) = let rec print_library_entry with_values ent = let sep = if with_values then " = " else " : " in + let pr_name (sp,_) = pr_id (basename sp) in match ent with - | (sp,Lib.Leaf lobj) -> - (print_leaf_entry with_values sep (sp,lobj)) - | (sp,Lib.OpenedSection (dir,_)) -> - (str " >>>>>>> Section " ++ pr_id (basename sp) ++ fnl ()) - | (sp,Lib.ClosedSection _) -> - (str " >>>>>>> Closed Section " ++ pr_id (basename sp) ++ fnl ()) - | (_,Lib.Module dir) -> - (str " >>>>>>> Module " ++ pr_dirpath dir ++ fnl ()) + | (oname,Lib.Leaf lobj) -> + (print_leaf_entry with_values sep (oname,lobj)) + | (oname,Lib.OpenedSection (dir,_)) -> + (str " >>>>>>> Section " ++ pr_name oname ++ fnl ()) + | (oname,Lib.ClosedSection _) -> + (str " >>>>>>> Closed Section " ++ pr_name oname ++ fnl ()) + | (_,Lib.CompilingModule (dir,_)) -> + (str " >>>>>>> Library " ++ pr_dirpath dir ++ fnl ()) + | (oname,Lib.OpenedModule _) -> + (str " >>>>>>> Module " ++ pr_name oname ++ fnl ()) + | (oname,Lib.OpenedModtype _) -> + (str " >>>>>>> Module Type " ++ pr_name oname ++ fnl ()) | (_,Lib.FrozenState _) -> (mt ()) @@ -385,9 +391,9 @@ let read_sec_context (loc,qid) = with Not_found -> user_err_loc (loc,"read_sec_context", str "Unknown section") in let rec get_cxt in_cxt = function - | ((sp,Lib.OpenedSection (dir',_)) as hd)::rest -> + | ((_,Lib.OpenedSection ((dir',_),_)) as hd)::rest -> if dir = dir' then (hd::in_cxt) else get_cxt (hd::in_cxt) rest - | ((sp,Lib.ClosedSection (_,_,ctxt)) as hd)::rest -> + | ((_,Lib.ClosedSection (_,_,ctxt)) as hd)::rest -> error "Cannot print the contents of a closed section" | [] -> [] | hd::rest -> get_cxt (hd::in_cxt) rest @@ -414,15 +420,15 @@ let print_eval red_fun env {uj_val=trm;uj_type=typ} = let print_name (loc,qid) = try let sp = Nametab.locate_obj qid in - let (sp,lobj) = - let (sp,entry) = - List.find (fun en -> (fst en) = sp) (Lib.contents_after None) + let (oname,lobj) = + let (oname,entry) = + List.find (fun en -> (fst (fst en)) = sp) (Lib.contents_after None) in match entry with - | Lib.Leaf obj -> (sp,obj) + | Lib.Leaf obj -> (oname,obj) | _ -> raise Not_found in - print_leaf_entry true " = " (sp,lobj) + print_leaf_entry true " = " (oname,lobj) with Not_found -> try match Nametab.locate qid with @@ -438,8 +444,8 @@ let print_name (loc,qid) = (print_named_decl (str,c,typ)) with Not_found -> try - let sp = Syntax_def.locate_syntactic_definition qid in - print_syntactic_def " = " sp + let kn = Nametab.locate_syntactic_definition qid in + print_syntactic_def " = " kn with Not_found -> user_err_loc (loc,"print_name",pr_qualid qid ++ spc () ++ str "not a defined object") @@ -468,9 +474,9 @@ let print_local_context () = let env = Lib.contents_after None in let rec print_var_rec = function | [] -> (mt ()) - | (sp,Lib.Leaf lobj)::rest -> + | (oname,Lib.Leaf lobj)::rest -> if "VARIABLE" = object_tag lobj then - let (d,_) = get_variable (basename sp) in + let (d,_) = get_variable (basename (fst oname)) in (print_var_rec rest ++ print_named_decl d) else @@ -478,16 +484,18 @@ let print_local_context () = | _::rest -> print_var_rec rest and print_last_const = function - | (sp,Lib.Leaf lobj)::rest -> + | (oname,Lib.Leaf lobj)::rest -> (match object_tag lobj with | "CONSTANT" | "PARAMETER" -> + let kn = snd oname in let {const_body=val_0;const_type=typ} = - Global.lookup_constant sp in - (print_last_const rest ++ - print_basename sp ++str" = " ++ + Global.lookup_constant kn in + (print_last_const rest ++ + print_basename kn ++str" = " ++ print_typed_body (val_0,typ)) | "INDUCTIVE" -> - (print_last_const rest ++print_mutual sp ++ fnl ()) + let kn = snd oname in + (print_last_const rest ++print_mutual kn ++ fnl ()) | "VARIABLE" -> (mt ()) | _ -> print_last_const rest) | _ -> (mt ()) diff --git a/parsing/prettyp.mli b/parsing/prettyp.mli index 4a65492ae..4f3c1269e 100644 --- a/parsing/prettyp.mli +++ b/parsing/prettyp.mli @@ -16,6 +16,7 @@ open Sign open Term open Environ open Reductionops +open Libnames open Nametab (*i*) @@ -26,21 +27,21 @@ val assumptions_for_print : name list -> Termops.names_context val print_closed_sections : bool ref val print_impl_args : int list -> std_ppcmds val print_context : bool -> Lib.library_segment -> std_ppcmds -val print_library_entry : bool -> (section_path * Lib.node) -> std_ppcmds +val print_library_entry : bool -> (object_name * Lib.node) -> std_ppcmds val print_full_context : unit -> std_ppcmds val print_full_context_typ : unit -> std_ppcmds -val print_sec_context : Nametab.qualid located -> std_ppcmds -val print_sec_context_typ : Nametab.qualid located -> std_ppcmds +val print_sec_context : qualid located -> std_ppcmds +val print_sec_context_typ : qualid located -> std_ppcmds val print_judgment : env -> unsafe_judgment -> std_ppcmds val print_safe_judgment : env -> Safe_typing.judgment -> std_ppcmds val print_eval : reduction_function -> env -> unsafe_judgment -> std_ppcmds (* This function is exported for the graphical user-interface pcoq *) -val build_inductive : section_path -> int -> +val build_inductive : mutual_inductive -> int -> global_reference * rel_context * types * identifier array * types array -val print_mutual : section_path -> std_ppcmds -val print_name : Nametab.qualid located -> std_ppcmds -val print_opaque_name : Nametab.qualid located -> std_ppcmds +val print_mutual : mutual_inductive -> std_ppcmds +val print_name : qualid located -> std_ppcmds +val print_opaque_name : qualid located -> std_ppcmds val print_local_context : unit -> std_ppcmds (*i diff --git a/parsing/printer.ml b/parsing/printer.ml index a3ce32047..5867d8143 100644 --- a/parsing/printer.ml +++ b/parsing/printer.ml @@ -21,6 +21,7 @@ open Declare open Coqast open Ast open Termast +open Libnames open Extend open Nametab open Ppconstr diff --git a/parsing/printer.mli b/parsing/printer.mli index 17b6341fd..48ee955cc 100644 --- a/parsing/printer.mli +++ b/parsing/printer.mli @@ -11,6 +11,7 @@ (*i*) open Pp open Names +open Libnames open Term open Sign open Environ diff --git a/parsing/q_coqast.ml4 b/parsing/q_coqast.ml4 index debe3ba93..6b7eb89a3 100644 --- a/parsing/q_coqast.ml4 +++ b/parsing/q_coqast.ml4 @@ -8,6 +8,8 @@ (* $Id$ *) +open Names +open Libnames open Q_util let dummy_loc = (0,0) @@ -71,13 +73,13 @@ let rec mlexpr_of_ast = function | Coqast.Num loc i -> <:expr< Coqast.Num loc $int:string_of_int i$ >> | Coqast.Id loc id -> <:expr< Coqast.Id loc $str:id$ >> | Coqast.Str loc str -> <:expr< Coqast.Str loc $str:str$ >> - | Coqast.Path loc qid -> - let l,a = Names.repr_path qid in + | Coqast.Path loc kn -> + let l,a = Libnames.decode_kn kn in let mlexpr_of_modid id = - <:expr< Names.id_of_string $str:Names.string_of_id id$ >> in - let e = List.map mlexpr_of_modid (Names.repr_dirpath l) in + <:expr< Names.id_of_string $str:string_of_id id$ >> in + let e = List.map mlexpr_of_modid (repr_dirpath l) in let e = expr_list_of_var_list e in - <:expr< Coqast.Path loc (Names.make_path (Names.make_dirpath $e$) + <:expr< Coqast.Path loc (Libnames.encode_kn (Names.make_dirpath $e$) (Names.id_of_string $str:Names.string_of_id a$)) >> | Coqast.Dynamic _ _ -> failwith "Q_Coqast: dynamic: not implemented" @@ -121,8 +123,8 @@ let mlexpr_of_dirpath dir = <:expr< Names.make_dirpath $mlexpr_of_list mlexpr_of_ident l$ >> let mlexpr_of_qualid qid = - let (dir, id) = Nametab.repr_qualid qid in - <:expr< Nametab.make_qualid $mlexpr_of_dirpath dir$ $mlexpr_of_ident id$ >> + let (dir, id) = repr_qualid qid in + <:expr< make_qualid $mlexpr_of_dirpath dir$ $mlexpr_of_ident id$ >> let mlexpr_of_reference = function | Tacexpr.RQualid (loc,qid) -> <:expr< Tacexpr.RQualid loc $mlexpr_of_qualid qid$ >> diff --git a/parsing/search.ml b/parsing/search.ml index 1d5619969..78e549e13 100644 --- a/parsing/search.ml +++ b/parsing/search.ml @@ -22,6 +22,7 @@ open Astterm open Environ open Pattern open Printer +open Libnames open Nametab (* The functions print_constructors and crible implement the behavior needed @@ -49,9 +50,9 @@ let rec head_const c = match kind_of_term c with let crible (fn : global_reference -> env -> constr -> unit) ref = let env = Global.env () in - let imported = Library.opened_modules() in + let imported = Library.opened_libraries() in let const = constr_of_reference ref in - let crible_rec sp lobj = + let crible_rec (sp,_) lobj = match object_tag lobj with | "VARIABLE" -> (try @@ -60,26 +61,28 @@ let crible (fn : global_reference -> env -> constr -> unit) ref = with Not_found -> (* we are in a section *) ()) | "CONSTANT" | "PARAMETER" -> - let {const_type=typ} = Global.lookup_constant sp in - if (head_const typ) = const then fn (ConstRef sp) env typ + let kn=locate_constant (qualid_of_sp sp) in + let {const_type=typ} = Global.lookup_constant kn in + if (head_const typ) = const then fn (ConstRef kn) env typ | "INDUCTIVE" -> - let mib = Global.lookup_mind sp in - let arities = + let kn=locate_mind (qualid_of_sp sp) in + let mib = Global.lookup_mind kn in +(* let arities = array_map_to_list (fun mip -> (Name mip.mind_typename, None, mip.mind_nf_arity)) - mib.mind_packets in + mib.mind_packets in*) (match kind_of_term const with - | Ind ((sp',tyi) as indsp) -> - if sp=sp' then - print_constructors indsp fn env mib.mind_packets.(tyi) + | Ind ((kn',tyi) as ind) -> + if kn=kn' then + print_constructors ind fn env mib.mind_packets.(tyi) | _ -> ()) | _ -> () in - try - Library.iter_all_segments false crible_rec - with Not_found -> - errorlabstrm "search" + try + Declaremods.iter_all_segments false crible_rec + with Not_found -> + errorlabstrm "search" (pr_global ref ++ spc () ++ str "not declared") (* Fine Search. By Yves Bertot. *) @@ -104,9 +107,9 @@ let plain_display ref a c = msg (hov 2 (pr ++ str":" ++ spc () ++ pc) ++ fnl ()) let filter_by_module (module_list:dir_path list) (accept:bool) - (ref:global_reference) (env:env) _ = + (ref:global_reference) _ _ = try - let sp = sp_of_global env ref in + let sp = sp_of_global None ref in let sl = dirpath sp in let rec filter_aux = function | m :: tl -> (not (is_dirpath_prefix_of m sl)) && (filter_aux tl) @@ -117,9 +120,9 @@ let filter_by_module (module_list:dir_path list) (accept:bool) false let gref_eq = - IndRef (make_path Coqlib.logic_module (id_of_string "eq"), 0) + IndRef (Libnames.encode_kn Coqlib.logic_module (id_of_string "eq"), 0) let gref_eqT = - IndRef (make_path Coqlib.logic_type_module (id_of_string "eqT"), 0) + IndRef (Libnames.encode_kn Coqlib.logic_type_module (id_of_string "eqT"), 0) let mk_rewrite_pattern1 eq pattern = PApp (PRef eq, [| PMeta None; pattern; PMeta None |]) diff --git a/parsing/search.mli b/parsing/search.mli index 111858733..d1f6bccd4 100644 --- a/parsing/search.mli +++ b/parsing/search.mli @@ -13,6 +13,7 @@ open Names open Term open Environ open Pattern +open Libnames open Nametab (*s Search facilities. *) diff --git a/parsing/termast.ml b/parsing/termast.ml index 95e4b5163..d30e03903 100644 --- a/parsing/termast.ml +++ b/parsing/termast.ml @@ -18,6 +18,7 @@ open Termops open Inductive open Sign open Environ +open Libnames open Declare open Impargs open Coqast diff --git a/parsing/termast.mli b/parsing/termast.mli index fb0762446..c6e94fe16 100644 --- a/parsing/termast.mli +++ b/parsing/termast.mli @@ -14,6 +14,7 @@ open Term open Termops open Sign open Environ +open Libnames open Nametab open Rawterm open Pattern @@ -36,7 +37,7 @@ val ast_of_existential : env -> existential -> Coqast.t val ast_of_constructor : env -> constructor -> Coqast.t val ast_of_inductive : env -> inductive -> Coqast.t val ast_of_ref : global_reference -> Coqast.t -val ast_of_qualid : Nametab.qualid -> Coqast.t +val ast_of_qualid : qualid -> Coqast.t (* For debugging *) val print_implicits : bool ref diff --git a/pretyping/classops.ml b/pretyping/classops.ml index fb8a6c8a4..bc3b1310a 100755 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -12,6 +12,7 @@ open Util open Pp open Options open Names +open Libnames open Nametab open Environ open Libobject @@ -166,6 +167,29 @@ let lookup_pattern_path_between (s,t) = | Construct sp -> (sp, coe.coe_param) | _ -> raise Not_found) l + +let subst_cl_typ subst ct = match ct with + CL_SORT + | CL_FUN + | CL_SECVAR _ -> ct + | CL_CONST kn -> + let kn' = subst_kn subst kn in + if kn' == kn then ct else + CL_CONST kn' + | CL_IND (kn,i) -> + let kn' = subst_kn subst kn in + if kn' == kn then ct else + CL_IND (kn',i) + +let subst_coe_typ = subst_global + +let subst_coe_info subst info = + let jud = info.coe_value in + let val' = subst_mps subst (j_val jud) in + let type' = subst_mps subst (j_type jud) in + if val' == j_val jud && type' == j_type jud then info else + {info with coe_value = make_judge val' type'} + (* library, summary *) (*val inClass : (cl_typ * cl_info_typ) -> Libobject.object = @@ -173,12 +197,18 @@ let lookup_pattern_path_between (s,t) = let cache_class (_,x) = add_new_class x +let subst_class (_,subst,(ct,ci as obj)) = + let ct' = subst_cl_typ subst ct in + if ct' == ct then obj else + (ct',ci) + let (inClass,outClass) = - declare_object ("CLASS", - { load_function = cache_class; - open_function = (fun _ -> ()); + declare_object {(default_object "CLASS") with + load_function = (fun _ o -> cache_class o); cache_function = cache_class; - export_function = (function x -> Some x) }) + subst_function = subst_class; + classify_function = (fun (_,x) -> Substitute x); + export_function = (function x -> Some x) } let declare_class (cl,stre,p) = Lib.add_anonymous_leaf (inClass ((cl,{ cl_strength = stre; cl_param = p }))) @@ -225,7 +255,7 @@ let inductive_class_of ind = fst (class_info (CL_IND ind)) let class_args_of c = snd (decompose_app c) let strength_of_cl = function - | CL_CONST sp -> constant_strength sp + | CL_CONST kn -> constant_strength (sp_of_global None (ConstRef kn)) | CL_SECVAR sp -> variable_strength sp | _ -> NeverDischarge @@ -233,11 +263,11 @@ let string_of_class = function | CL_FUN -> "FUNCLASS" | CL_SORT -> "SORTCLASS" | CL_CONST sp -> - string_of_qualid (shortest_qualid_of_global (Global.env()) (ConstRef sp)) + string_of_qualid (shortest_qualid_of_global None (ConstRef sp)) | CL_IND sp -> - string_of_qualid (shortest_qualid_of_global (Global.env()) (IndRef sp)) + string_of_qualid (shortest_qualid_of_global None (IndRef sp)) | CL_SECVAR sp -> - string_of_qualid (shortest_qualid_of_global (Global.env()) (VarRef sp)) + string_of_qualid (shortest_qualid_of_global None (VarRef sp)) (* coercion_value : coe_index -> unsafe_judgment * bool *) @@ -316,17 +346,27 @@ let cache_coercion (_,((coe,xf),cls,clt)) = let jf = add_new_coercion (coe,xf) in add_coercion_in_graph (jf,is,it) +let subst_coercion (_,subst,((coe,xf),cls,clt as obj)) = + let coe' = subst_coe_typ subst coe in + let xf' = subst_coe_info subst xf in + let cls' = subst_cl_typ subst cls in + let clt' = subst_cl_typ subst clt in + if coe' == coe && xf' == xf && cls' == cls & clt' == clt then obj else + ((coe',xf'),cls',clt') + + (* val inCoercion : (coe_typ * coe_info_typ) * cl_typ * cl_typ -> -> Libobject.object val outCoercion : Libobject.object -> (coe_typ * coe_info_typ) * cl_typ * cl_typ *) let (inCoercion,outCoercion) = - declare_object ("COERCION", - { load_function = cache_coercion; - open_function = (fun _ -> ()); + declare_object {(default_object "COERCION") with + load_function = (fun _ o -> cache_coercion o); cache_function = cache_coercion; - export_function = (function x -> Some x) }) + subst_function = subst_coercion; + classify_function = (fun (_,x) -> Substitute x); + export_function = (function x -> Some x) } let declare_coercion coef v stre ~isid ~src:cls ~target:clt ~params:ps = Lib.add_anonymous_leaf @@ -353,14 +393,15 @@ let coercion_of_qualid qid = let coe = coe_of_reference ref in if not (coercion_exists coe) then errorlabstrm "try_add_coercion" - (Nametab.pr_global_env (Global.env()) ref ++ str" is not a coercion"); + (Nametab.pr_global_env None ref ++ str" is not a coercion"); coe module CoercionPrinting = struct type t = coe_typ let encode = coercion_of_qualid - let printer x = pr_global_env (Global.env()) x + let subst = subst_coe_typ + let printer x = pr_global_env None x let key = Goptions.SecondaryTable ("Printing","Coercion") let title = "Explicitly printed coercions: " let member_message x b = diff --git a/pretyping/classops.mli b/pretyping/classops.mli index cd5f31db8..d37588d06 100644 --- a/pretyping/classops.mli +++ b/pretyping/classops.mli @@ -10,6 +10,7 @@ (*i*) open Names +open Libnames open Nametab open Term open Evd @@ -32,7 +33,7 @@ type cl_info_typ = { cl_param : int } (* This is the type of coercion kinds *) -type coe_typ = global_reference +type coe_typ = Libnames.global_reference (* This is the type of infos for declared coercions *) type coe_info_typ @@ -46,7 +47,7 @@ type coe_index (* This is the type of paths from a class to another *) type inheritance_path = coe_index list -val coe_of_reference : global_reference -> coe_typ +val coe_of_reference : Libnames.global_reference -> coe_typ (*s [declare_class] adds a class to the set of declared classes *) val declare_class : cl_typ * strength * int -> unit diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index faa5e9e46..748c72f4c 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -23,6 +23,7 @@ open Impargs open Rawterm open Nameops open Termops +open Libnames open Nametab (****************************************************************************) @@ -66,7 +67,11 @@ module PrintingCasesMake = struct type t = inductive * int array let encode = Test.encode - let printer (ind,_) = pr_global_env (Global.env()) (IndRef ind) + let subst subst ((kn,i), ints as obj) = + let kn' = subst_kn subst kn in + if kn' == kn then obj else + (kn',i), ints + let printer (ind,_) = pr_global_env None (IndRef ind) let key = Goptions.SecondaryTable ("Printing",Test.field) let title = Test.title let member_message x = Test.member_message (printer x) diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml index ef9db0c44..e0eecf703 100644 --- a/pretyping/indrec.ml +++ b/pretyping/indrec.ml @@ -11,10 +11,12 @@ open Pp open Util open Names +open Libnames open Nameops open Term open Termops open Declarations +open Entries open Inductive open Inductiveops open Instantiate @@ -500,13 +502,14 @@ let declare_one_elimination ind = let (mib,mip) = Global.lookup_inductive ind in let mindstr = string_of_id mip.mind_typename in let declare na c t = - let sp = Declare.declare_constant (id_of_string na) - (ConstantEntry + let kn = Declare.declare_constant (id_of_string na) + (DefinitionEntry { const_entry_body = c; const_entry_type = t; const_entry_opaque = false }, NeverDischarge) in - Options.if_verbose ppnl (str na ++ str " is defined"); sp + Options.if_verbose ppnl (str na ++ str " is defined"); + kn in let env = Global.env () in let sigma = Evd.empty in @@ -519,7 +522,7 @@ let declare_one_elimination ind = if List.mem InType kelim then let cte = declare (mindstr^type_suff) (make_elim (new_sort_in_family InType)) None - in let c = mkConst cte and t = constant_type (Global.env ()) cte + in let c = mkConst (snd cte) and t = constant_type (Global.env ()) (snd cte) in List.iter (fun (sort,suff) -> let (t',c') = instanciate_type_indrec_scheme (new_sort_in_family sort) npars c t @@ -548,8 +551,32 @@ let declare_eliminations sp = (* Look up function for the default elimination constant *) let lookup_eliminator ind_sp s = - let env = Global.env() in - let path = sp_of_global env (IndRef ind_sp) in + let kn,i = ind_sp in + let mp,dp,l = repr_kn kn in + let ind_id = (Global.lookup_mind kn).mind_packets.(i).mind_typename in + let id = add_suffix ind_id (elimination_suffix s) in + (* Try first to get an eliminator defined in the same section as the *) + (* inductive type *) + let ref = ConstRef (make_kn mp dp (label_of_id id)) in + try + let _ = full_name ref in + constr_of_reference ref + with Not_found -> + (* Then try to get a user-defined eliminator in some other places *) + (* using short name (e.g. for "eq_rec") *) + try construct_reference None id + with Not_found -> + errorlabstrm "default_elim" + (str "Cannot find the elimination combinator " ++ + pr_id id ++ spc () ++ + str "The elimination of the inductive definition " ++ + pr_id id ++ spc () ++ str "on sort " ++ + spc () ++ print_sort_family s ++ + str " is probably not allowed") + + +(* let env = Global.env() in + let path = sp_of_global None (IndRef ind_sp) in let dir, base = repr_path path in let id = add_suffix base (elimination_suffix s) in (* Try first to get an eliminator defined in the same section as the *) @@ -558,7 +585,7 @@ let lookup_eliminator ind_sp s = with Not_found -> (* Then try to get a user-defined eliminator in some other places *) (* using short name (e.g. for "eq_rec") *) - try construct_reference env id + try construct_reference None id with Not_found -> errorlabstrm "default_elim" (str "Cannot find the elimination combinator " ++ @@ -567,3 +594,4 @@ let lookup_eliminator ind_sp s = pr_id base ++ spc () ++ str "on sort " ++ spc () ++ print_sort_family s ++ str " is probably not allowed") +*) diff --git a/pretyping/pattern.ml b/pretyping/pattern.ml index 85d38ab4d..bde1f31bd 100644 --- a/pretyping/pattern.ml +++ b/pretyping/pattern.ml @@ -10,6 +10,7 @@ open Util open Names +open Libnames open Nameops open Term open Termops diff --git a/pretyping/pattern.mli b/pretyping/pattern.mli index cc36b260a..5ff667f90 100644 --- a/pretyping/pattern.mli +++ b/pretyping/pattern.mli @@ -13,6 +13,7 @@ open Names open Sign open Term open Environ +open Libnames open Nametab (*i*) diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index b7d834d6f..91e0acd65 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -143,7 +143,7 @@ let error_unsolvable_implicit (loc,kind) = str "Cannot infer a type of this anonymous abstraction" | ImplicitArg (c,n) -> str "Cannot infer the " ++ pr_ord n ++ - str " implicit argument of " ++ Nametab.pr_global_env (Global.env()) c + str " implicit argument of " ++ Nametab.pr_global_env None c | InternalHole -> str "Cannot infer a term for an internal placeholder" in diff --git a/pretyping/rawterm.ml b/pretyping/rawterm.ml index 9f39a1db9..0f376ae5e 100644 --- a/pretyping/rawterm.ml +++ b/pretyping/rawterm.ml @@ -13,6 +13,7 @@ open Util open Names open Sign open Term +open Libnames open Nametab (*i*) @@ -88,6 +89,91 @@ type rawconstr = - boolean in POldCase means it is recursive i*) +let rec subst_pat subst pat = + match pat with + | PatVar _ -> pat + | PatCstr (loc,((kn,i),j),cpl,n) -> + let kn' = subst_kn subst kn + and cpl' = list_smartmap (subst_pat subst) cpl in + if kn' == kn && cpl' == cpl then pat else + PatCstr (loc,((kn',i),j),cpl',n) + +let rec subst_raw subst raw = + match raw with + | RRef (loc,ref) -> + let ref' = subst_global subst ref in + if ref' == ref then raw else + RRef (loc,ref') + + | RVar _ -> raw + | REvar _ -> raw + | RMeta _ -> raw + + | RApp (loc,r,rl) -> + let r' = subst_raw subst r + and rl' = list_smartmap (subst_raw subst) rl in + if r' == r && rl' == rl then raw else + RApp(loc,r',rl') + + | RLambda (loc,n,r1,r2) -> + let r1' = subst_raw subst r1 and r2' = subst_raw subst r2 in + if r1' == r1 && r2' == r2 then raw else + RLambda (loc,n,r1',r2') + + | RProd (loc,n,r1,r2) -> + let r1' = subst_raw subst r1 and r2' = subst_raw subst r2 in + if r1' == r1 && r2' == r2 then raw else + RProd (loc,n,r1',r2') + + | RLetIn (loc,n,r1,r2) -> + let r1' = subst_raw subst r1 and r2' = subst_raw subst r2 in + if r1' == r1 && r2' == r2 then raw else + RLetIn (loc,n,r1',r2') + + | RCases (loc,cs,ro,rl,branches) -> + let ro' = option_smartmap (subst_raw subst) ro + and rl' = list_smartmap (subst_raw subst) rl + and branches' = list_smartmap + (fun (loc,idl,cpl,r as branch) -> + let cpl' = list_smartmap (subst_pat subst) cpl + and r' = subst_raw subst r in + if cpl' == cpl && r' == r then branch else + (loc,idl,cpl',r')) + branches + in + if ro' == ro && rl' == rl && branches' == branches then raw else + RCases (loc,cs,ro',rl',branches') + + | ROldCase (loc,b,ro,r,ra) -> + let ro' = option_smartmap (subst_raw subst) ro + and r' = subst_raw subst r + and ra' = array_smartmap (subst_raw subst) ra in + if ro' == ro && r' == r && ra' == ra then raw else + ROldCase (loc,b,ro',r',ra') + + | RRec (loc,fix,ida,ra1,ra2) -> + let ra1' = array_smartmap (subst_raw subst) ra1 + and ra2' = array_smartmap (subst_raw subst) ra2 in + if ra1' == ra1 && ra2' == ra2 then raw else + RRec (loc,fix,ida,ra1',ra2') + + | RSort _ -> raw + + | RHole (loc,ImplicitArg (ref,i)) -> + let ref' = subst_global subst ref in + if ref' == ref then raw else + RHole (loc,ImplicitArg (ref',i)) + | RHole (loc, + (AbstractionType _ | QuestionMark | CasesType | InternalHole)) -> + raw + + | RCast (loc,r1,r2) -> + let r1' = subst_raw subst r1 and r2' = subst_raw subst r2 in + if r1' == r1 && r2' == r2 then raw else + RCast (loc,r1',r2') + + | RDynamic _ -> raw + let dummy_loc = (0,0) let loc_of_rawconstr = function diff --git a/pretyping/rawterm.mli b/pretyping/rawterm.mli index d6543323e..e8b0726c9 100644 --- a/pretyping/rawterm.mli +++ b/pretyping/rawterm.mli @@ -12,6 +12,7 @@ open Names open Sign open Term +open Libnames open Nametab (*i*) @@ -57,7 +58,7 @@ type 'ctxt reference = | REVar of int * 'ctxt type rawconstr = - | RRef of loc * global_reference + | RRef of loc * Libnames.global_reference | RVar of loc * identifier | REvar of loc * existential_key | RMeta of loc * int @@ -93,6 +94,8 @@ val loc_of_rawconstr : rawconstr -> loc val set_loc_of_rawconstr : loc -> rawconstr -> rawconstr val join_loc : loc -> loc -> loc +val subst_raw : Names.substitution -> rawconstr -> rawconstr + type 'a raw_red_flag = { rBeta : bool; rIota : bool; diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml index 6617a7a9b..e2d4036f4 100755 --- a/pretyping/recordops.ml +++ b/pretyping/recordops.ml @@ -11,6 +11,7 @@ open Util open Pp open Names +open Libnames open Nametab open Term open Termops @@ -44,14 +45,22 @@ let structure_table = (ref [] : (inductive * struc_typ) list ref) let cache_structure (_,x) = structure_table := x :: (!structure_table) +let subst_structure (_,subst,((kn,i),struc as obj)) = + let kn' = subst_kn subst kn in + let proj' = list_smartmap + (option_smartmap (subst_kn subst)) + struc.s_PROJ + in + if proj' == struc.s_PROJ && kn' == kn then obj else + (kn',i),{struc with s_PROJ = proj'} + let (inStruc,outStruc) = - declare_object - ("STRUCTURE", { - load_function = (fun _ -> ()); - cache_function = cache_structure; - open_function = cache_structure; - export_function = (function x -> Some x) - }) + declare_object {(default_object "STRUCTURE") with + cache_function = cache_structure; + open_function = (fun i o -> if i=1 then cache_structure o); + subst_function = subst_structure; + classify_function = (fun (_,x) -> Substitute x); + export_function = (function x -> Some x) } let add_new_struc (s,c,n,l) = Lib.add_anonymous_leaf (inStruc (s,{s_CONST=c;s_PARAM=n;s_PROJ=l})) @@ -81,19 +90,42 @@ type obj_typ = { o_TPARAMS : constr list; (* dans l'ordre *) o_TCOMPS : constr list } (* dans l'ordre *) +let subst_obj subst obj = + let o_DEF' = subst_mps subst obj.o_DEF in + let o_TABS' = list_smartmap (subst_mps subst) obj.o_TABS in + let o_TPARAMS' = list_smartmap (subst_mps subst) obj.o_TPARAMS in + let o_TCOMPS' = list_smartmap (subst_mps subst) obj.o_TCOMPS in + if o_DEF' == obj.o_DEF + && o_TABS' == obj.o_TABS + && o_TPARAMS' == obj.o_TPARAMS + && o_TCOMPS' == obj.o_TCOMPS + then + obj + else + { o_DEF = o_DEF' ; + o_TABS = o_TABS' ; + o_TPARAMS = o_TPARAMS' ; + o_TCOMPS = o_TCOMPS' } + let object_table = (ref [] : ((global_reference * global_reference) * obj_typ) list ref) let cache_object (_,x) = object_table := x :: (!object_table) -let (inObjDef, outObjDef) = - declare_object - ("OBJDEF", { - load_function = (fun _ -> ()); - open_function = cache_object; - cache_function = cache_object; - export_function = (function x -> Some x) - }) +let subst_object (_,subst,((r1,r2),o as obj)) = + let r1' = subst_global subst r1 in + let r2' = subst_global subst r2 in + let o' = subst_obj subst o in + if r1' == r1 && r2' == r2 && o' == o then obj else + (r1',r2'),o' + +let (inObjDef,outObjDef) = + declare_object {(default_object "OBJDEF") with + open_function = (fun i o -> if i=1 then cache_object o); + cache_function = cache_object; + subst_function = subst_object; + classify_function = (fun (_,x) -> Substitute x); + export_function = (function x -> Some x) } let add_new_objdef (o,c,la,lp,l) = try @@ -104,14 +136,11 @@ let add_new_objdef (o,c,la,lp,l) = let cache_objdef1 (_,sp) = () -let (inObjDef1, outObjDef1) = - declare_object - ("OBJDEF1", { - load_function = (fun _ -> ()); - open_function = cache_objdef1; - cache_function = cache_objdef1; - export_function = (function x -> Some x) - }) +let (inObjDef1,outObjDef1) = + declare_object {(default_object "OBJDEF1") with + open_function = (fun i o -> if i=1 then cache_objdef1 o); + cache_function = cache_objdef1; + export_function = (function x -> Some x) } let objdef_info o = List.assoc o !object_table diff --git a/pretyping/recordops.mli b/pretyping/recordops.mli index a3dd2f2a3..e5ffe4fd2 100755 --- a/pretyping/recordops.mli +++ b/pretyping/recordops.mli @@ -12,6 +12,7 @@ open Names open Nametab open Term +open Libnames open Classops open Libobject open Library @@ -27,10 +28,10 @@ val compter : bool ref type struc_typ = { s_CONST : identifier; s_PARAM : int; - s_PROJ : section_path option list } + s_PROJ : constant option list } val add_new_struc : - inductive * identifier * int * section_path option list -> unit + inductive * identifier * int * constant option list -> unit (* [find_structure isp] returns the infos associated to inductive path [isp] if it corresponds to a structure, otherwise fails with [Not_found] *) @@ -50,5 +51,5 @@ val add_new_objdef : val inStruc : inductive * struc_typ -> obj val outStruc : obj -> inductive * struc_typ -val inObjDef1 : section_path -> obj -val outObjDef1 : obj -> section_path +val inObjDef1 : kernel_name -> obj +val outObjDef1 : obj -> kernel_name diff --git a/pretyping/syntax_def.ml b/pretyping/syntax_def.ml index f13f31de0..dc5824dc7 100644 --- a/pretyping/syntax_def.ml +++ b/pretyping/syntax_def.ml @@ -11,6 +11,7 @@ open Util open Pp open Names +open Libnames open Rawterm open Libobject open Lib @@ -18,51 +19,50 @@ open Nameops (* Syntactic definitions. *) -let syntax_table = ref (Spmap.empty : rawconstr Spmap.t) +let syntax_table = ref (KNmap.empty : rawconstr KNmap.t) let _ = Summary.declare_summary "SYNTAXCONSTANT" { Summary.freeze_function = (fun () -> !syntax_table); Summary.unfreeze_function = (fun ft -> syntax_table := ft); - Summary.init_function = (fun () -> syntax_table := Spmap.empty); + Summary.init_function = (fun () -> syntax_table := KNmap.empty); Summary.survive_section = false } -let add_syntax_constant sp c = - syntax_table := Spmap.add sp c !syntax_table +let add_syntax_constant kn c = + syntax_table := KNmap.add kn c !syntax_table -let cache_syntax_constant (sp,c) = +let cache_syntax_constant ((sp,kn),c) = if Nametab.exists_cci sp then errorlabstrm "cache_syntax_constant" (pr_id (basename sp) ++ str " already exists"); - add_syntax_constant sp c; - Nametab.push_syntactic_definition sp; - Nametab.push_short_name_syntactic_definition sp + add_syntax_constant kn c; + Nametab.push_syntactic_definition (Nametab.Until 1) sp kn -let load_syntax_constant (sp,c) = +let load_syntax_constant i ((sp,kn),c) = if Nametab.exists_cci sp then errorlabstrm "cache_syntax_constant" (pr_id (basename sp) ++ str " already exists"); - add_syntax_constant sp c; - Nametab.push_syntactic_definition sp + add_syntax_constant kn c; + Nametab.push_syntactic_definition (Nametab.Until i) sp kn -let open_syntax_constant (sp,c) = - Nametab.push_short_name_syntactic_definition sp +let open_syntax_constant i ((sp,kn),c) = + Nametab.push_syntactic_definition (Nametab.Exactly i) sp kn + +let subst_syntax_constant ((sp,kn),subst,c) = + subst_raw subst c + +let classify_syntax_constant (_,c) = Substitute c let (in_syntax_constant, out_syntax_constant) = - let od = { + declare_object {(default_object "SYNTAXCONSTANT") with cache_function = cache_syntax_constant; load_function = load_syntax_constant; open_function = open_syntax_constant; + subst_function = subst_syntax_constant; + classify_function = classify_syntax_constant; export_function = (fun x -> Some x) } - in - declare_object ("SYNTAXCONSTANT", od) let declare_syntactic_definition id c = let _ = add_leaf id (in_syntax_constant c) in () -let search_syntactic_definition sp = Spmap.find sp !syntax_table - -let locate_syntactic_definition qid = - match Nametab.extended_locate qid with - | Nametab.SyntacticDef sp -> sp - | _ -> raise Not_found +let search_syntactic_definition kn = KNmap.find kn !syntax_table diff --git a/pretyping/syntax_def.mli b/pretyping/syntax_def.mli index 1a884300c..a3e1cad25 100644 --- a/pretyping/syntax_def.mli +++ b/pretyping/syntax_def.mli @@ -10,6 +10,7 @@ (*i*) open Names +open Libnames open Rawterm (*i*) @@ -17,8 +18,6 @@ open Rawterm val declare_syntactic_definition : identifier -> rawconstr -> unit -val search_syntactic_definition : section_path -> rawconstr - -val locate_syntactic_definition : Nametab.qualid -> section_path +val search_syntactic_definition : kernel_name -> rawconstr diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index b42318a52..f2c185621 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -13,6 +13,7 @@ open Util open Names open Nameops open Term +open Libnames open Termops open Declarations open Inductive @@ -32,7 +33,7 @@ let set_transparent_const sp = if cb.const_body <> None & cb.const_opaque then errorlabstrm "set_transparent_const" (str "Cannot make" ++ spc () ++ - Nametab.pr_global_env (Global.env()) (Nametab.ConstRef sp) ++ + Nametab.pr_global_env None (ConstRef sp) ++ spc () ++ str "transparent because it was declared opaque."); Conv_oracle.set_transparent_const sp @@ -48,10 +49,10 @@ let _ = let is_evaluable env ref = match ref with - EvalConstRef sp -> - let (ids,sps) = Conv_oracle.freeze() in - Sppred.mem sp sps & - let cb = Environ.lookup_constant sp env in + EvalConstRef kn -> + let (ids,kns) = Conv_oracle.freeze() in + KNpred.mem kn kns & + let cb = Environ.lookup_constant kn env in cb.const_body <> None & not cb.const_opaque | EvalVarRef id -> let (ids,sps) = Conv_oracle.freeze() in @@ -195,8 +196,9 @@ let invert_name labs l na0 env sigma ref = function let refi = match ref with | EvalRel _ | EvalEvar _ -> None | EvalVar id' -> Some (EvalVar id) - | EvalConst sp -> - Some (EvalConst (make_path (dirpath sp) id)) in + | EvalConst kn -> + let (mp,dp,_) = repr_kn kn in + Some (EvalConst (make_kn mp dp (label_of_id id))) in match refi with | None -> None | Some ref -> @@ -365,7 +367,7 @@ let contract_cofix_use_function f (bodynum,(_,names,bodies as typedbodies)) = let subbodies = list_tabulate make_Fi nbodies in substl subbodies bodies.(bodynum) -let reduce_mind_case_use_function sp env mia = +let reduce_mind_case_use_function kn env mia = match kind_of_term mia.mconstr with | Construct(ind_sp,i as cstr_sp) -> let real_cargs = snd (list_chop mia.mci.ci_npar mia.mcargs) in @@ -374,10 +376,11 @@ let reduce_mind_case_use_function sp env mia = let build_fix_name i = match names.(i) with | Name id -> - let sp = make_path (dirpath sp) id in - (match constant_opt_value env sp with + let (mp,dp,_) = repr_kn kn in + let kn = make_kn mp dp (label_of_id id) in + (match constant_opt_value env kn with | None -> None - | Some _ -> Some (mkConst sp)) + | Some _ -> Some (mkConst kn)) | Anonymous -> None in let cofix_def = contract_cofix_use_function build_fix_name cofix in mkCase (mia.mci, mia.mP, applist(cofix_def,mia.mcargs), mia.mlf) @@ -595,14 +598,14 @@ let nf env sigma c = strong whd_nf env sigma c * ol is the occurence list to find. *) let rec substlin env name n ol c = match kind_of_term c with - | Const const when EvalConstRef const = name -> + | Const kn when EvalConstRef kn = name -> if List.hd ol = n then try - (n+1, List.tl ol, constant_value env const) + (n+1, List.tl ol, constant_value env kn) with NotEvaluableConst _ -> errorlabstrm "substlin" - (pr_sp const ++ str " is not a defined constant") + (pr_kn kn ++ str " is not a defined constant") else ((n+1), ol, c) @@ -692,7 +695,7 @@ let rec substlin env name n ol c = let string_of_evaluable_ref = function | EvalVarRef id -> string_of_id id - | EvalConstRef sp -> string_of_path sp + | EvalConstRef kn -> string_of_kn kn let unfold env sigma name = if is_evaluable env name then diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli index 1d801c5a4..d41161efb 100644 --- a/pretyping/tacred.mli +++ b/pretyping/tacred.mli @@ -79,8 +79,8 @@ val reduction_of_redexp : red_expr -> reduction_function val declare_red_expr : string -> reduction_function -> unit (* Opaque and Transparent commands. *) -val set_opaque_const : section_path -> unit -val set_transparent_const : section_path -> unit +val set_opaque_const : constant -> unit +val set_transparent_const : constant -> unit val set_opaque_var : identifier -> unit val set_transparent_var : identifier -> unit diff --git a/pretyping/termops.ml b/pretyping/termops.ml index 7f77bcdba..05411c68d 100644 --- a/pretyping/termops.ml +++ b/pretyping/termops.ml @@ -15,6 +15,7 @@ open Nameops open Term open Sign open Environ +open Libnames open Nametab let print_sort = function @@ -27,15 +28,15 @@ let print_sort_family = function | InProp -> (str "Prop") | InType -> (str "Type") -let current_module = ref empty_dirpath +(*let current_module = ref empty_dirpath -let set_module m = current_module := m +let set_module m = current_module := m*) let new_univ = let univ_gen = ref 0 in (fun sp -> incr univ_gen; - Univ.make_univ (!current_module,!univ_gen)) + Univ.make_univ (Lib.module_dp(),!univ_gen)) let new_sort_in_family = function | InProp -> mk_Prop @@ -510,7 +511,7 @@ let first_char id = let lowercase_first_char id = String.lowercase (first_char id) -let id_of_global env ref = basename (sp_of_global env ref) +let id_of_global env ref = Nametab.id_of_global (Some (named_context env)) ref let sort_hdchar = function | Prop(_) -> "P" @@ -524,12 +525,12 @@ let hdchar env c = | LetIn (_,_,_,c) -> hdrec (k+1) c | Cast (c,_) -> hdrec k c | App (f,l) -> hdrec k f - | Const sp -> - let c = lowercase_first_char (basename sp) in + | Const kn -> + let c = lowercase_first_char (id_of_label (label kn)) in if c = "?" then "y" else c - | Ind ((sp,i) as x) -> + | Ind ((kn,i) as x) -> if i=0 then - lowercase_first_char (basename sp) + lowercase_first_char (id_of_label (label kn)) else lowercase_first_char (id_of_global env (IndRef x)) | Construct ((sp,i) as x) -> @@ -644,12 +645,12 @@ let occur_rel p env id = let occur_id env nenv id0 c = let rec occur n c = match kind_of_term c with | Var id when id=id0 -> raise Occur - | Const sp when basename sp = id0 -> raise Occur + | Const kn when id_of_global env (ConstRef kn) = id0 -> raise Occur | Ind ind_sp - when basename (path_of_inductive env ind_sp) = id0 -> + when id_of_global env (IndRef ind_sp) = id0 -> raise Occur | Construct cstr_sp - when basename (path_of_constructor env cstr_sp) = id0 -> + when id_of_global env (ConstructRef cstr_sp) = id0 -> raise Occur | Rel p when p>n & occur_rel (p-n) nenv id0 -> raise Occur | _ -> iter_constr_with_binders succ occur n c diff --git a/pretyping/termops.mli b/pretyping/termops.mli index ebcd93a34..431e69e7f 100644 --- a/pretyping/termops.mli +++ b/pretyping/termops.mli @@ -16,7 +16,7 @@ open Sign open Environ (* Universes *) -val set_module : Names.dir_path -> unit +(*val set_module : Names.dir_path -> unit*) val new_univ : unit -> Univ.universe val new_sort_in_family : sorts_family -> sorts @@ -104,7 +104,7 @@ val eta_eq_constr : constr -> constr -> bool (* finding "intuitive" names to hypotheses *) val first_char : identifier -> string val lowercase_first_char : identifier -> string -val id_of_global : env -> Nametab.global_reference -> identifier +(*val id_of_global : env -> Libnames.global_reference -> identifier*) val sort_hdchar : sorts -> string val hdchar : env -> types -> string val id_of_name_using_hdchar : diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml index e2932a904..8b47ced02 100644 --- a/proofs/pfedit.ml +++ b/proofs/pfedit.ml @@ -15,6 +15,7 @@ open Nameops open Sign open Term open Declarations +open Entries open Environ open Evd open Declare @@ -34,7 +35,7 @@ open Safe_typing type proof_topstate = { top_hyps : named_context * named_context; top_goal : goal; - top_strength : bool * Nametab.strength; + top_strength : bool * Libnames.strength; top_hook : declaration_hook } let proof_edits = diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli index 78314eacb..beeb56959 100644 --- a/proofs/pfedit.mli +++ b/proofs/pfedit.mli @@ -67,7 +67,7 @@ val get_undo : unit -> int option declare the built constructions as a coercion or a setoid morphism) *) val start_proof : - identifier -> bool * strength -> named_context -> constr + identifier -> bool * Libnames.strength -> named_context -> constr -> declaration_hook -> unit (* [restart_proof ()] restarts the current focused proof from the beginning @@ -96,7 +96,7 @@ val suspend_proof : unit -> unit val cook_proof : unit -> identifier * - (Safe_typing.constant_entry * (bool * strength) * declaration_hook) + (Entries.definition_entry * (bool * Libnames.strength) * declaration_hook) (* To export completed proofs to xml *) val set_xml_cook_proof : (pftreestate -> unit) -> unit diff --git a/proofs/proof_trees.ml b/proofs/proof_trees.ml index 0c916495a..a05464b00 100644 --- a/proofs/proof_trees.ml +++ b/proofs/proof_trees.ml @@ -21,6 +21,7 @@ open Evarutil open Proof_type open Tacred open Typing +open Libnames open Nametab (* @@ -277,9 +278,9 @@ let last_of_cvt_flags red = List.map (function | EvalVarRef id -> nvar id - | EvalConstRef sp -> + | EvalConstRef kn -> ast_of_qualid - (shortest_qualid_of_global (Global.env()) (ConstRef sp))) + (shortest_qualid_of_global None (ConstRef kn))) lconst in if lqid = [] then [] else if n_unf then [ope("Delta",[]);ope("UnfBut",lqid)] @@ -303,9 +304,9 @@ let ast_of_cvt_redexp = function ope("Unfold",List.map (fun (locc,sp) -> ope("UNFOLD", [match sp with | EvalVarRef id -> nvar id - | EvalConstRef sp -> + | EvalConstRef kn -> ast_of_qualid - (shortest_qualid_of_global (Global.env()) (ConstRef sp))] + (shortest_qualid_of_global None (ConstRef kn))] @(List.map num locc))) l) | Fold l -> ope("Fold",List.map (fun c -> ope ("CONSTR", diff --git a/proofs/proof_type.ml b/proofs/proof_type.ml index 756370b80..3d599cc48 100644 --- a/proofs/proof_type.ml +++ b/proofs/proof_type.ml @@ -12,6 +12,7 @@ open Environ open Evd open Names +open Libnames open Term open Util open Tacexpr @@ -99,4 +100,4 @@ type closed_generic_argument = type 'a closed_abstract_argument_type = ('a,constr,raw_tactic_expr) abstract_argument_type -type declaration_hook = Nametab.strength -> Nametab.global_reference -> unit +type declaration_hook = Libnames.strength -> global_reference -> unit diff --git a/proofs/proof_type.mli b/proofs/proof_type.mli index 9c4bf5c47..6f86fbe3a 100644 --- a/proofs/proof_type.mli +++ b/proofs/proof_type.mli @@ -12,6 +12,7 @@ open Environ open Evd open Names +open Libnames open Term open Util open Tacexpr @@ -127,4 +128,4 @@ type closed_generic_argument = type 'a closed_abstract_argument_type = ('a,constr,raw_tactic_expr) abstract_argument_type -type declaration_hook = Nametab.strength -> Nametab.global_reference -> unit +type declaration_hook = Libnames.strength -> global_reference -> unit diff --git a/proofs/tacexpr.ml b/proofs/tacexpr.ml index 65c9cd009..2e8d581ae 100644 --- a/proofs/tacexpr.ml +++ b/proofs/tacexpr.ml @@ -10,6 +10,7 @@ open Names open Coqast +open Libnames open Nametab open Rawterm open Util diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 7777f3463..e0b802c3e 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -83,7 +83,7 @@ let pf_interp_type gls c = let evc = project gls in Astterm.interp_type evc (pf_env gls) c -let pf_global gls id = Declare.construct_reference (pf_env gls) id +let pf_global gls id = Declare.construct_reference (Some (pf_hyps gls)) id let pf_parse_const gls = compose (pf_global gls) id_of_string diff --git a/tactics/auto.ml b/tactics/auto.ml index 256914d4c..08953ffe0 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -30,10 +30,11 @@ open Tactics open Tacticals open Clenv open Hiddentac +open Libnames +open Nametab open Libobject open Library open Printer -open Nametab open Declarations open Tacexpr @@ -161,34 +162,7 @@ let _ = Summary.declare_summary "search" (**************************************************************************) -(* declaration of the AUTOHINT library object *) -(**************************************************************************) - -(* If the database does not exist, it is created *) -(* TODO: should a warning be printed in this case ?? *) -let add_hint dbname hintlist = - try - let db = searchtable_map dbname in - let db' = Hint_db.add_list hintlist db in - searchtable_add (dbname,db') - with Not_found -> - let db = Hint_db.add_list hintlist Hint_db.empty in - searchtable_add (dbname,db) - -let cache_autohint (_,(name,hintlist)) = - try add_hint name hintlist with _ -> anomaly "Auto.add_hint" - -let export_autohint x = Some x - -let (inAutoHint,outAutoHint) = - declare_object ("AUTOHINT", - { load_function = (fun _ -> ()); - cache_function = cache_autohint; - open_function = cache_autohint; - export_function = export_autohint }) - -(**************************************************************************) -(* The "Hint" vernacular command *) +(* Auxiliary functions to prepare AUTOHINT objects *) (**************************************************************************) let rec nb_hyp c = match kind_of_term c with @@ -261,22 +235,6 @@ let make_resolve_hyp env sigma (hname,_,htyp) = | Failure _ -> [] | e when Logic.catchable_exception e -> anomaly "make_resolve_hyp" -let add_resolves env sigma clist dbnames = - List.iter - (fun dbname -> - Lib.add_anonymous_leaf - (inAutoHint - (dbname, - (List.flatten - (List.map - (fun (name,c) -> - let ty = type_of env sigma c in - let verbose = Options.is_verbose() in - make_resolves env sigma name (true,verbose) (c,ty)) clist - )) - ))) - dbnames - (* REM : in most cases hintname = id *) let make_unfold (hintname, ref) = (Pattern.label_of_ref ref, @@ -285,12 +243,6 @@ let make_unfold (hintname, ref) = pat = None; code = Unfold_nth ref }) -let add_unfolds l dbnames = - List.iter - (fun dbname -> - Lib.add_anonymous_leaf (inAutoHint (dbname, List.map make_unfold l))) - dbnames - let make_extern name pri pat tacast = let hdconstr = try_head_pattern pat in (hdconstr, @@ -299,6 +251,126 @@ let make_extern name pri pat tacast = pat = Some pat; code= Extern tacast }) +let make_trivial env sigma (name,c) = + let t = hnf_constr env sigma (type_of env sigma c) in + let hd = head_of_constr_reference (List.hd (head_constr t)) in + let ce = mk_clenv_from () (c,t) in + (hd, { hname = name; + pri=1; + pat = Some (Pattern.pattern_of_constr (clenv_template_type ce).rebus); + code=Res_pf_THEN_trivial_fail(c,ce) }) + +open Vernacexpr + +(**************************************************************************) +(* declaration of the AUTOHINT library object *) +(**************************************************************************) + +let eager o = ref (Lazy.Value o) + +(* If the database does not exist, it is created *) +(* TODO: should a warning be printed in this case ?? *) +let add_hint dbname hintlist = + try + let db = searchtable_map dbname in + let db' = Hint_db.add_list hintlist db in + searchtable_add (dbname,db') + with Not_found -> + let db = Hint_db.add_list hintlist Hint_db.empty in + searchtable_add (dbname,db) + +let cache_autohint (_,(name,l_hintlist)) = + try add_hint name (Lazy.force l_hintlist) with _ -> anomaly "Auto.add_hint" + +let subst_autohint (_,subst,((name,l_hintlist) as obj)) = + let recalc_hints hintlist = + let env = Global.env() and sigma = Evd.empty in + let recalc_hint ((_,data) as hint) = + match data.code with + | Res_pf (c,_) -> + let c' = Term.subst_mps subst c in + if c==c' then hint else + make_apply_entry env sigma (false,false) + data.hname (c', type_of env sigma c') + | ERes_pf (c,_) -> + let c' = Term.subst_mps subst c in + if c==c' then hint else + make_apply_entry env sigma (true,false) + data.hname (c', type_of env sigma c') + | Give_exact c -> + let c' = Term.subst_mps subst c in + if c==c' then hint else + make_exact_entry data.hname (c',type_of env sigma c') + | Res_pf_THEN_trivial_fail (c,_) -> + let c' = Term.subst_mps subst c in + if c==c' then hint else + make_trivial env sigma (data.hname,c') + | Unfold_nth ref -> + let ref' = subst_global subst ref in + if ref==ref' then hint else + make_unfold (data.hname,ref') + | Extern _ -> + anomaly "Extern hints cannot be substituted!!!" + in + list_smartmap recalc_hint hintlist + in + let hintlist = Lazy.force l_hintlist in + try + let hintlist' = recalc_hints hintlist in + if hintlist'==hintlist then + obj + else + (name,eager hintlist') + with + _ -> (name,lazy (recalc_hints hintlist)) + +let classify_autohint (_,((name,l_hintlist) as obj)) = + match Lazy.force l_hintlist with + [] -> Dispose + | (_,{code=Extern _})::_ -> Keep obj (* TODO! should be changed *) + | _ -> Substitute obj + +let export_autohint x = Some x + +let (inAutoHint,outAutoHint) = + declare_object {(default_object "AUTOHINT") with + cache_function = cache_autohint; + load_function = (fun _ -> cache_autohint); + subst_function = subst_autohint; + classify_function = classify_autohint; + export_function = export_autohint } + + +(**************************************************************************) +(* The "Hint" vernacular command *) +(**************************************************************************) +let add_resolves env sigma clist dbnames = + List.iter + (fun dbname -> + Lib.add_anonymous_leaf + (inAutoHint + (dbname, + let l = + List.flatten + (List.map + (fun (name,c) -> + let ty = type_of env sigma c in + let verbose = Options.is_verbose() in + make_resolves env sigma name (true,verbose) (c,ty)) clist + ) + in + eager l + ))) + dbnames + + +let add_unfolds l dbnames = + List.iter + (fun dbname -> Lib.add_anonymous_leaf + (inAutoHint (dbname, eager (List.map make_unfold l)))) + dbnames + + let add_extern name pri (patmetas,pat) tacast dbname = (* We check that all metas that appear in tacast have at least one occurence in the left pattern pat *) @@ -312,28 +384,58 @@ let add_extern name pri (patmetas,pat) tacast dbname = (str "The meta-variable ?" ++ int i ++ str" is not bound") | [] -> Lib.add_anonymous_leaf - (inAutoHint(dbname, [make_extern name pri pat tacast])) + (inAutoHint(dbname, eager [make_extern name pri pat tacast])) let add_externs name pri pat tacast dbnames = List.iter (add_extern name pri pat tacast) dbnames -let make_trivial (name,c) = - let sigma = Evd.empty and env = Global.env() in - let t = hnf_constr env sigma (type_of env sigma c) in - let hd = head_of_constr_reference (List.hd (head_constr t)) in - let ce = mk_clenv_from () (c,t) in - (hd, { hname = name; - pri=1; - pat = Some (Pattern.pattern_of_constr (clenv_template_type ce).rebus); - code=Res_pf_THEN_trivial_fail(c,ce) }) -let add_trivials l dbnames = +let add_trivials env sigma l dbnames = List.iter (fun dbname -> - Lib.add_anonymous_leaf (inAutoHint(dbname, List.map make_trivial l))) + Lib.add_anonymous_leaf ( + inAutoHint(dbname, eager (List.map (make_trivial env sigma) l)))) dbnames -open Vernacexpr + +let add_hints dbnames h = + let dbnames = if dbnames = [] then ["core"] else dbnames in match h with + | HintsResolve lhints -> + let env = Global.env() and sigma = Evd.empty in + let f (n,c) = + let c = Astterm.interp_constr sigma env c in + let n = match n with + | None -> basename (sp_of_global None (Declare.reference_of_constr c)) + | Some n -> n in + (n,c) in + add_resolves env sigma (List.map f lhints) dbnames + | HintsImmediate lhints -> + let env = Global.env() and sigma = Evd.empty in + let f (n,c) = + let c = Astterm.interp_constr sigma env c in + let n = match n with + | None -> basename (sp_of_global None (Declare.reference_of_constr c)) + | Some n -> n in + (n,c) in + add_trivials env sigma (List.map f lhints) dbnames + | HintsUnfold lhints -> + let f (n,locqid) = + let r = Nametab.global locqid in + let n = match n with + | None -> basename (sp_of_global None r) + | Some n -> n in + (n,r) in + add_unfolds (List.map f lhints) dbnames + | HintsConstructors (hintname, qid) -> + let env = Global.env() and sigma = Evd.empty in + let isp = global_inductive qid in + let consnames = (snd (Global.lookup_inductive isp)).mind_consnames in + let lcons = list_tabulate (fun i -> mkConstruct (isp,i+1)) (Array.length consnames) in + let lcons = List.map2 (fun id c -> (id,c)) (Array.to_list consnames) lcons in + add_resolves env sigma lcons dbnames + | HintsExtern (hintname, pri, patcom, tacexp) -> + let pat = Astterm.interp_constrpattern Evd.empty (Global.env()) patcom in + add_externs hintname pri pat tacexp dbnames (**************************************************************************) (* Functions for printing the hints *) @@ -437,45 +539,6 @@ let print_searchtable () = print_hint_db db) !searchtable -let add_hints dbnames h = - let dbnames = if dbnames = [] then ["core"] else dbnames in match h with - | HintsResolve lhints -> - let env = Global.env() and sigma = Evd.empty in - let f (n,c) = - let c = Astterm.interp_constr sigma env c in - let n = match n with - | None -> basename (sp_of_global env (Declare.reference_of_constr c)) - | Some n -> n in - (n,c) in - add_resolves env sigma (List.map f lhints) dbnames - | HintsImmediate lhints -> - let env = Global.env() and sigma = Evd.empty in - let f (n,c) = - let c = Astterm.interp_constr sigma env c in - let n = match n with - | None -> basename (sp_of_global env (Declare.reference_of_constr c)) - | Some n -> n in - (n,c) in - add_trivials (List.map f lhints) dbnames - | HintsUnfold lhints -> - let f (n,locqid) = - let r = Nametab.global locqid in - let n = match n with - | None -> basename (sp_of_global (Global.env()) r) - | Some n -> n in - (n,r) in - add_unfolds (List.map f lhints) dbnames - | HintsConstructors (hintname, qid) -> - let env = Global.env() and sigma = Evd.empty in - let isp = global_inductive qid in - let consnames = (snd (Global.lookup_inductive isp)).mind_consnames in - let lcons = list_tabulate (fun i -> mkConstruct (isp,i+1)) (Array.length consnames) in - let lcons = List.map2 (fun id c -> (id,c)) (Array.to_list consnames) lcons in - add_resolves env sigma lcons dbnames - | HintsExtern (hintname, pri, patcom, tacexp) -> - let pat = Astterm.interp_constrpattern Evd.empty (Global.env()) patcom in - add_externs hintname pri pat tacexp dbnames - (**************************************************************************) (* Automatic tactics *) (**************************************************************************) @@ -811,7 +874,7 @@ let default_superauto g = superauto !default_search_depth [] [] g let interp_to_add gl locqid = let r = Nametab.global locqid in - let id = basename (sp_of_global (Global.env()) r) in + let id = basename (sp_of_global None r) in (next_ident_away id (pf_ids_of_hyps gl), Declare.constr_of_reference r) let gen_superauto nopt l a b gl = @@ -822,4 +885,3 @@ let gen_superauto nopt l a b gl = let h_superauto no l a b = Refiner.abstract_tactic (TacSuperAuto (no,l,a,b)) (gen_superauto no l a b) - diff --git a/tactics/auto.mli b/tactics/auto.mli index c5266bc58..c887c1bb4 100644 --- a/tactics/auto.mli +++ b/tactics/auto.mli @@ -19,7 +19,7 @@ open Clenv open Pattern open Environ open Evd -open Nametab +open Libnames (*i*) type auto_tactic = diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index 75c7509ac..ad4da1235 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -61,7 +61,6 @@ let autorewrite tac_main lbas = tclTHEN tac (one_base tac_main bas)) tclIDTAC lbas)) (* Functions necessary to the library object declaration *) -let load_hintrewrite _ = () let cache_hintrewrite (_,(rbase,lrl)) = List.iter (fun (c,b,t) -> Hashtbl.add !rewtab rbase (c,b,Tacinterp.interp t)) lrl @@ -69,12 +68,10 @@ let export_hintrewrite x = Some x (* Declaration of the Hint Rewrite library object *) let (in_hintrewrite,out_hintrewrite)= - Libobject.declare_object - ("HINT_REWRITE", - { Libobject.load_function = load_hintrewrite; - Libobject.open_function = cache_hintrewrite; + Libobject.declare_object {(Libobject.default_object "HINT_REWRITE") with + Libobject.open_function = (fun i o -> if i=1 then cache_hintrewrite o); Libobject.cache_function = cache_hintrewrite; - Libobject.export_function = export_hintrewrite }) + Libobject.export_function = export_hintrewrite } (* To add rewriting rules to a base *) let add_rew_rules base lrul = diff --git a/tactics/dhyp.ml b/tactics/dhyp.ml index d73677617..92db61f07 100644 --- a/tactics/dhyp.ml +++ b/tactics/dhyp.ml @@ -200,11 +200,10 @@ type destructor_data_object = identifier * destructor_data let ((inDD : destructor_data_object->obj), (outDD : obj->destructor_data_object)) = - declare_object ("DESTRUCT-HYP-CONCL-DATA", - { load_function = (fun _ -> ()); + declare_object {(default_object "DESTRUCT-HYP-CONCL-DATA") with cache_function = cache_dd; - open_function = cache_dd; - export_function = export_dd }) + open_function = (fun i o -> if i=1 then cache_dd o); + export_function = export_dd } let add_destructor_hint na loc pat pri code = begin match loc, code with diff --git a/tactics/elim.ml b/tactics/elim.ml index cbe18ba36..b4f718fbd 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -14,6 +14,7 @@ open Names open Term open Termops open Environ +open Libnames open Reduction open Inductiveops open Proof_type @@ -103,7 +104,7 @@ let head_in gls indl t = with Not_found -> false let inductive_of = function - | Nametab.IndRef ity -> ity + | IndRef ity -> ity | r -> errorlabstrm "Decompose" (Printer.pr_global r ++ str " is not an inductive type") diff --git a/tactics/equality.ml b/tactics/equality.ml index fdb579729..bec72973d 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -1231,18 +1231,15 @@ let add_list_rules rbase lrl = let rules_of_base rbase = List.rev (Gmapl.find rbase !rew_tab) (*Functions necessary to the library object declaration*) -let load_autorewrite_rule _ = () let cache_autorewrite_rule (_,(rbase,lrl)) = add_list_rules rbase lrl let export_autorewrite_rule x = Some x (*Declaration of the AUTOREWRITE_RULE library object*) let (in_autorewrite_rule,out_autorewrite_rule)= - Libobject.declare_object - ("AUTOREWRITE_RULE", - { Libobject.load_function = load_autorewrite_rule; - Libobject.open_function = cache_autorewrite_rule; + Libobject.declare_object {(Libobject.default_object "AUTOREWRITE_RULE") with + Libobject.open_function = (fun i o -> if i=1 then cache_autorewrite_rule o); Libobject.cache_function = cache_autorewrite_rule; - Libobject.export_function = export_autorewrite_rule }) + Libobject.export_function = export_autorewrite_rule } (****The tactic****) diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 1a83dbb5d..36e0fa23f 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -19,6 +19,7 @@ open Evd open Printer open Reductionops open Declarations +open Entries open Inductiveops open Environ open Tacmach @@ -244,10 +245,10 @@ let add_inversion_lemma name env sigma t sort dep inv_op = let invProof = inversion_scheme env sigma t sort dep inv_op in let _ = declare_constant name - (ConstantEntry { const_entry_body = invProof; + (DefinitionEntry { const_entry_body = invProof; const_entry_type = None; const_entry_opaque = false }, - Nametab.NeverDischarge) + Libnames.NeverDischarge) in () (* open Pfedit *) diff --git a/tactics/setoid_replace.ml b/tactics/setoid_replace.ml index a147997ba..1962318f9 100644 --- a/tactics/setoid_replace.ml +++ b/tactics/setoid_replace.ml @@ -15,6 +15,8 @@ open Reductionops open Term open Termops open Names +open Entries +open Libnames open Nameops open Util open Pp @@ -51,7 +53,7 @@ let constant dir s = try Declare.global_reference_in_absolute_module dir id with Not_found -> - anomaly ("Setoid: cannot find "^(Nametab.string_of_qualid (Nametab.make_qualid dir id))) + anomaly ("Setoid: cannot find "^(string_of_qualid (make_qualid dir id))) let global_constant dir s = let dir = make_dirpath @@ -60,7 +62,7 @@ let global_constant dir s = try Declare.global_reference_in_absolute_module dir id with Not_found -> - anomaly ("Setoid: cannot find "^(Nametab.string_of_qualid (Nametab.make_qualid dir id))) + anomaly ("Setoid: cannot find "^(string_of_qualid (make_qualid dir id))) let current_constant id = try @@ -97,6 +99,21 @@ let setoid_table_add (s,th) = setoid_table := Gmap.add s th !setoid_table let setoid_table_find s = Gmap.find s !setoid_table let setoid_table_mem s = Gmap.mem s !setoid_table +let subst_setoid subst setoid = + let set_a' = subst_mps subst setoid.set_a in + let set_aeq' = subst_mps subst setoid.set_aeq in + let set_th' = subst_mps subst setoid.set_th in + if set_a' == setoid.set_a + && set_aeq' == setoid.set_aeq + && set_th' == setoid.set_th + then + setoid + else + { set_a = set_a' ; + set_aeq = set_aeq' ; + set_th = set_th' ; + } + let equiv_list () = List.map (fun x -> x.set_aeq) (Gmap.rng !setoid_table) let _ = @@ -110,13 +127,19 @@ let _ = let (setoid_to_obj, obj_to_setoid)= let cache_set (_,(s, th)) = setoid_table_add (s,th) + and subst_set (_,subst,(s,th as obj)) = + let s' = subst_mps subst s in + let th' = subst_setoid subst th in + if s' == s && th' == th then obj else + (s',th') and export_set x = Some x in - declare_object ("setoid-theory", - { cache_function = cache_set; - load_function = (fun _ -> ()); - open_function = cache_set; - export_function = export_set}) + declare_object {(default_object "setoid-theory") with + cache_function = cache_set; + open_function = (fun i o -> if i=1 then cache_set o); + subst_function = subst_set; + classify_function = (fun (_,x) -> Substitute x); + export_function = export_set} (******************************* Table of declared morphisms ********************) @@ -128,6 +151,23 @@ let morphism_table_add (m,c) = morphism_table := Gmap.add m c !morphism_table let morphism_table_find m = Gmap.find m !morphism_table let morphism_table_mem m = Gmap.mem m !morphism_table +let subst_morph subst morph = + let lem' = subst_mps subst morph.lem in + let arg_types' = list_smartmap (subst_mps subst) morph.arg_types in + let lem2' = option_smartmap (subst_mps subst) morph.lem2 in + if lem' == morph.lem + && arg_types' == morph.arg_types + && lem2' == morph.lem2 + then + morph + else + { lem = lem' ; + profil = morph.profil ; + arg_types = arg_types' ; + lem2 = lem2' ; + } + + let _ = Summary.declare_summary "morphism-table" { Summary.freeze_function = (fun () -> !morphism_table); @@ -139,13 +179,19 @@ let _ = let (morphism_to_obj, obj_to_morphism)= let cache_set (_,(m, c)) = morphism_table_add (m, c) + and subst_set (_,subst,(m,c as obj)) = + let m' = subst_mps subst m in + let c' = subst_morph subst c in + if m' == m && c' == c then obj else + (m',c') and export_set x = Some x in - declare_object ("morphism-definition", - { cache_function = cache_set; - load_function = (fun _ -> ()); - open_function = cache_set; - export_function = export_set}) + declare_object {(default_object "morphism-definition") with + cache_function = cache_set; + open_function = (fun i o -> if i=1 then cache_set o); + subst_function = subst_set; + classify_function = (fun (_,x) -> Substitute x); + export_function = export_set} (************************** Adding a setoid to the database *********************) @@ -238,15 +284,15 @@ let add_setoid a aeq th = let eq_ext_name = gen_eq_lem_name () in let eq_ext_name2 = gen_eq_lem_name () in let _ = Declare.declare_constant eq_ext_name - ((ConstantEntry {const_entry_body = eq_morph; + ((DefinitionEntry {const_entry_body = eq_morph; const_entry_type = None; const_entry_opaque = true}), - Nametab.NeverDischarge) in + Libnames.NeverDischarge) in let _ = Declare.declare_constant eq_ext_name2 - ((ConstantEntry {const_entry_body = eq_morph2; + ((DefinitionEntry {const_entry_body = eq_morph2; const_entry_type = None; const_entry_opaque = true}), - Nametab.NeverDischarge) in + Libnames.NeverDischarge) in let eqmorph = (current_constant eq_ext_name) in let eqmorph2 = (current_constant eq_ext_name2) in (Lib.add_anonymous_leaf @@ -290,10 +336,10 @@ let check_is_dependent t n = let gen_lem_name m = match kind_of_term m with | Var id -> add_suffix id "_ext" - | Const sp -> add_suffix (basename sp) "_ext" - | Ind (sp, i) -> add_suffix (basename sp) ((string_of_int i)^"_ext") - | Construct ((sp,i),j) -> add_suffix - (basename sp) ((string_of_int i)^(string_of_int j)^"_ext") + | Const kn -> add_suffix (id_of_label (label kn)) "_ext" + | Ind (kn, i) -> add_suffix (id_of_label (label kn)) ((string_of_int i)^"_ext") + | Construct ((kn,i),j) -> add_suffix + (id_of_label (label kn)) ((string_of_int i)^(string_of_int j)^"_ext") | _ -> errorlabstrm "New Morphism" (str "The term " ++ prterm m ++ str "is not a known name") let gen_lemma_tail m lisset body n = @@ -359,7 +405,7 @@ let new_morphism m id hook = let lem = (gen_compat_lemma env m body args_t poss) in let lemast = (ast_of_constr true env lem) in new_edited id m poss; - start_proof_com (Some id) (false,Nametab.NeverDischarge) lemast hook; + start_proof_com (Some id) (false,Libnames.NeverDischarge) lemast hook; (Options.if_verbose Vernacentries.show_open_subgoals ())) let rec sub_bool l1 n = function @@ -451,10 +497,10 @@ let add_morphism lem_name (m,profil) = (let lem_2 = gen_lem_iff env m mext args_t poss in let lem2_name = add_suffix lem_name "2" in let _ = Declare.declare_constant lem2_name - ((ConstantEntry {const_entry_body = lem_2; + ((DefinitionEntry {const_entry_body = lem_2; const_entry_type = None; const_entry_opaque = true}), - Nametab.NeverDischarge) in + Libnames.NeverDischarge) in let lem2 = (current_constant lem2_name) in (Lib.add_anonymous_leaf (morphism_to_obj (m, @@ -472,7 +518,7 @@ let add_morphism lem_name (m,profil) = lem2 = None})))); Options.if_verbose ppnl (prterm m ++str " is registered as a morphism") let morphism_hook stre ref = - let pf_id = basename (sp_of_global (Global.env()) ref) in + let pf_id = basename (sp_of_global None ref) in if (is_edited pf_id) then (add_morphism pf_id (what_edited pf_id); no_more_edited pf_id) diff --git a/tactics/tacinterp.ml b/tactics/tacinterp.ml index d3721d015..27eb73d0a 100644 --- a/tactics/tacinterp.ml +++ b/tactics/tacinterp.ml @@ -12,6 +12,7 @@ open Astterm open Closure open RedFlags open Declarations +open Entries open Dyn open Libobject open Pattern @@ -22,6 +23,7 @@ open Tacred open Util open Names open Nameops +open Libnames open Nametab open Pfedit open Proof_type @@ -121,7 +123,7 @@ let make_hyps = List.map (fun (id,_,typ) -> (id,body_of_type typ)) (* Transforms an id into a constr if possible *) let constr_of_id ist id = match ist.goalopt with - | None -> construct_reference ist.env id + | None -> construct_reference (Some (Environ.named_context ist.env)) id | Some goal -> let hyps = pf_hyps goal in if mem_named_context id hyps then @@ -196,7 +198,7 @@ let look_for_interp = Hashtbl.find interp_tab let find_reference ist qid = (* We first look for a variable of the current proof *) - match Nametab.repr_qualid qid, ist.goalopt with + match repr_qualid qid, ist.goalopt with | (d,id),Some gl when repr_dirpath d = [] & List.mem id (pf_ids_of_hyps gl) -> VarRef id | _ -> Nametab.locate qid @@ -308,7 +310,7 @@ let glob_hyp_or_metanum ist = function | MetaNum (loc,n) -> MetaNum (loc,glob_metanum ist loc n) let glob_qualid_or_metanum ist = function - | AN (loc,qid) -> AN (loc,qualid_of_sp (sp_of_global (Global.env())(Nametab.global (loc,qid)))) + | AN (loc,qid) -> AN (loc,qualid_of_sp (sp_of_global None (Nametab.global (loc,qid)))) | MetaNum (loc,n) -> MetaNum (loc,glob_metanum ist loc n) let glob_reference ist (_,qid as locqid) = @@ -317,7 +319,7 @@ let glob_reference ist (_,qid as locqid) = if dir = empty_dirpath && List.mem id (fst ist) then qid else raise Not_found with Not_found -> - qualid_of_sp (sp_of_global (Global.env()) (Nametab.global locqid)) + qualid_of_sp (sp_of_global None (Nametab.global locqid)) let glob_ltac_qualid ist (loc,qid as locqid) = try qualid_of_sp (locate_tactic qid) @@ -824,7 +826,7 @@ let constr_to_id loc c = else invalid_arg_loc (loc, "Not an identifier") let constr_to_qid loc c = - try qualid_of_sp (sp_of_global (Global.env ()) (reference_of_constr c)) + try qualid_of_sp (sp_of_global None (reference_of_constr c)) with _ -> invalid_arg_loc (loc, "Not a global reference") (* Check for LetTac *) @@ -1706,15 +1708,14 @@ let register_tacdef (sp,td) = let cache_md (_,defs) = (* Needs a rollback if something goes wrong *) - List.iter (fun (sp,_) -> Nametab.push_tactic_path sp) defs; + List.iter (fun (sp,_) -> Nametab.push_tactic (Until 1) sp) defs; List.iter add (List.map register_tacdef defs) let (inMD,outMD) = - declare_object ("TAC-DEFINITION", - {cache_function = cache_md; - load_function = (fun _ -> ()); - open_function = cache_md; - export_function = (fun x -> Some x)}) + declare_object {(default_object "TAC-DEFINITION") with + cache_function = cache_md; + open_function = (fun i o -> if i=1 then cache_md o); + export_function = (fun x -> Some x)} (* Adds a definition for tactics in the table *) let make_absolute_name (loc,id) = diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index b6fea7dc2..08004e971 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -18,6 +18,7 @@ open Declarations open Inductive open Reduction open Environ +open Libnames open Declare open Refiner open Tacmach @@ -314,7 +315,7 @@ let general_elim_then_using | _ -> let name_elim = match kind_of_term elim with - | Const sp -> string_of_path sp + | Const kn -> string_of_kn kn | Var id -> string_of_id id | _ -> "\b" in diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 690f1f4f9..37da503fd 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -20,6 +20,7 @@ open Inductive open Inductiveops open Reductionops open Environ +open Libnames open Declare open Evd open Pfedit @@ -906,8 +907,8 @@ let find_eliminator c gl = let env = pf_env gl in let (ind,t) = reduce_to_quantified_ind env (project gl) (pf_type_of gl c) in let s = elimination_sort_of_goal gl in - try Indrec.lookup_eliminator ind s - with Not_found -> + Indrec.lookup_eliminator ind s +(* with Not_found -> let dir, base = repr_path (path_of_inductive env ind) in let id = Indrec.make_elimination_ident base s in errorlabstrm "default_elim" @@ -917,7 +918,7 @@ let find_eliminator c gl = pr_id base ++ spc () ++ str "on sort " ++ spc () ++ print_sort (new_sort_in_family s) ++ str " is probably not allowed") - +(* lookup_eliminator prints the message *) *) let default_elim (c,lbindc) gl = general_elim (c,lbindc) (find_eliminator c gl,NoBindings) gl @@ -926,6 +927,7 @@ let elim (c,lbindc) elim gl = | Some (elimc,lbindelimc) -> general_elim (c,lbindc) (elimc,lbindelimc) gl | None -> general_elim (c,lbindc) (find_eliminator c gl,NoBindings) gl + (* The simplest elimination tactic, with no substitutions at all. *) let simplest_elim c = default_elim (c,NoBindings) @@ -1001,7 +1003,7 @@ let general_elim_in id (c,lbindc) (elimc,lbindelimc) gl = (* We recompute recargs because we are not sure the elimination lemma comes from a canonically generated one *) - +(* dead code ? let rec is_rec_arg env sigma indpath t = try let (ind_sp,_) = find_mrectype env sigma t in @@ -1014,7 +1016,7 @@ let rec recargs indpath env sigma t = (is_rec_arg env sigma indpath t) ::(recargs indpath (push_rel_assum (na,t) env) sigma c2) | _ -> [] - +*) let induct_discharge old_style mind statuslists cname destopt avoid ra gl = let (lstatus,rstatus) = statuslists in let tophyp = ref None in @@ -1598,7 +1600,7 @@ let abstract_subproof name tac gls = in if occur_existential concl then error "Abstract cannot handle existentials"; let lemme = - start_proof na (false,Nametab.NeverDischarge) current_sign concl (fun _ _ -> ()); + start_proof na (false,NeverDischarge) current_sign concl (fun _ _ -> ()); let _,(const,(_,strength),_) = try by (tclCOMPLETE (tclTHEN (tclDO (List.length sign) intro) tac)); @@ -1607,10 +1609,10 @@ let abstract_subproof name tac gls = with e when catchable_exception e -> (delete_current_proof(); raise e) in (* Faudrait un peu fonctionnaliser cela *) - let cd = Safe_typing.ConstantEntry const in + let cd = Entries.DefinitionEntry const in let sp = Declare.declare_constant na (cd,strength) in let newenv = Global.env() in - Declare.constr_of_reference (ConstRef sp) + Declare.constr_of_reference (ConstRef (snd sp)) in exact_no_check (applist (lemme, diff --git a/tactics/tactics.mli b/tactics/tactics.mli index 14d479362..9d1f61616 100644 --- a/tactics/tactics.mli +++ b/tactics/tactics.mli @@ -21,6 +21,7 @@ open Evar_refiner open Clenv open Tacred open Tacticals +open Libnames open Tacexpr open Nametab open Rawterm diff --git a/tactics/tauto.ml4 b/tactics/tauto.ml4 index f4f8452fa..dc28eb48c 100644 --- a/tactics/tauto.ml4 +++ b/tactics/tauto.ml4 @@ -14,6 +14,7 @@ open Ast open Coqast open Hipattern open Names +open Libnames open Pp open Proof_type open Tacticals diff --git a/tactics/termdn.ml b/tactics/termdn.ml index 2d0f49f4e..d3acf6e68 100644 --- a/tactics/termdn.ml +++ b/tactics/termdn.ml @@ -14,6 +14,7 @@ open Nameops open Term open Pattern open Rawterm +open Libnames open Nametab (* Discrimination nets of terms. diff --git a/test-suite/modules/Nametab.v b/test-suite/modules/Nametab.v new file mode 100644 index 000000000..9d8c49ef7 --- /dev/null +++ b/test-suite/modules/Nametab.v @@ -0,0 +1,100 @@ +Module M. + Definition id:=[x:Set]x. + Module M. + Definition zero:(id nat):= O. + Definition id:=Set. + Definition nacik:id:=nat. + End M. +End M. + +Module M. +Module N. +Module K. +Definition id:=Set. +End K. +End N. +End M. + + +Locate id. +Locate K.id. +Locate N.K.id. +Locate M.N.K.id. +Locate Top.M.N.K.id. +Locate K. +Locate N.K. +Locate M.N.K. +Locate Scratch.M.N.K. +Locate N. +Locate M.N. +Locate Scratch.M.N. +Locate M. +Locate Scratch.M. + +(* +#trace Nametab.push;; +#trace Nametab.push_short_name;; +#trace Nametab.freeze;; +#trace Nametab.unfreeze;; +#trace Nametab.exists_cci;; +*) + +Module M. +Module M[X:SIG]. +Module M[X,Y:SIG]. +Module M[X:SIG;Y:SIG]. +Module M[X,Y:SIG;X,Z:SIG]. +Module M[X:SIG][Y:SIG]. +Module M[X,Y:SIG][X,Z:SIG]. +Module M:SIG. +Module M[X:SIG]:SIG. +Module M[X,Y:SIG]:SIG. +Module M[X:SIG;Y:SIG]:SIG. +Module M[X,Y:SIG;X,Z:SIG]:SIG. +Module M[X:SIG][Y:SIG]:SIG. +Module M[X,Y:SIG][X,Z:SIG]:SIG. +Module M:=(M N). +Module M[X:SIG]:=(M N). +Module M[X,Y:SIG]:=(M N). +Module M[X:SIG;Y:SIG]:=(M N). +Module M[X,Y:SIG;X,Z:SIG]:=(M N). +Module M[X:SIG][Y:SIG]:=(M N). +Module M[X,Y:SIG][X,Z:SIG]:=(M N). +Module M:SIG:=(M N). +Module M[X:SIG]:SIG:=(M N). +Module M[X,Y:SIG]:SIG:=(M N). +Module M[X:SIG;Y:SIG]:SIG:=(M N). +Module M[X,Y:SIG;X,Z:SIG]:SIG:=(M N). +Module M[X:SIG][Y:SIG]:SIG:=(M N). +Module M[X,Y:SIG][X,Z:SIG]:SIG:=(M N). + + +Moduletype SIG. + Parameter A:Set. + Parameter x:A. +EndT SIG. + +Module Nat. + Definition A:=nat. + Definition x:=O. +End Nat. + +Module List[X:SIG]. + Inductive list : Set := nil : list + | cons : X.A -> list -> list. + + Definition head := + [l:list]Cases l of + nil => X.x + | (cons x _) => x + end. + + Definition singl := [x:X.A] (cons x nil). + + Lemma head_singl : (x:X.A)(head (singl x))=x. + Auto. + Qed. + +End List. + +Module N:=(List Nat). diff --git a/test-suite/modules/Nat.v b/test-suite/modules/Nat.v new file mode 100644 index 000000000..d0ad6a5e2 --- /dev/null +++ b/test-suite/modules/Nat.v @@ -0,0 +1,68 @@ +Definition T:=nat. + +Definition le:=Peano.le. + +Hints Unfold le. + +Lemma le_refl:(n:nat)(le n n). + Auto. +Save. + +Lemma le_trans:(n,m,k:nat)(le n m) -> (le m k) -> (le n k). + Unfold le. + Intros. + Generalize H. + Clear H. + Elim H0. + Auto. + Auto. +Save. + +Lemma le_mono_S : (n,m:nat)(le n m) -> (le (S n) (S m)). + Unfold le. + NewInduction 1. + Auto. + Auto. +Save. + +Hints Resolve le_mono_S. + +Lemma le_pred_mono : (n,m:nat) (le n m) -> (le (pred n) (pred m)). + Unfold le. + Induction 1. + Auto. + Intro. + Case m0. + Auto. + Intro. + Simpl. + Auto. +Save. + +Hints Resolve le_pred_mono. + +Lemma le_S_mono : (m,n:nat)(le (S n) (S m)) -> (le n m). + Intros. + Change (le (pred (S n)) (pred (S m))). + Auto. +Save. + +Hints Resolve le_S_mono. + +Lemma le_antis:(n,m:nat)(le n m) -> (le m n) -> n=m. + NewInduction n. + Intros. + Inversion H0. + Reflexivity. + + Intros. + Inversion H. + Auto. + Rewrite (IHn m0). + Auto. + Rewrite <- H2 in H. + Auto. + Rewrite <- H2 in H0. + Auto. +Save. + diff --git a/test-suite/modules/PO.v b/test-suite/modules/PO.v new file mode 100644 index 000000000..da098a420 --- /dev/null +++ b/test-suite/modules/PO.v @@ -0,0 +1,89 @@ +Implicit Arguments On. + +Implicits fst. +Implicits snd. + +Module Type PO. + Parameter T:Set. + Parameter le:T->T->Prop. + + Axiom le_refl : (x:T)(le x x). + Axiom le_trans : (x,y,z:T)(le x y) -> (le y z) -> (le x z). + Axiom le_antis : (x,y:T)(le x y) -> (le y x) -> (x=y). + + Hints Resolve le_refl le_trans le_antis. +End PO. + + +Module Pair[X:PO][Y:PO]. + Definition T:=X.T*Y.T. + Definition le:=[p1,p2] + (X.le (fst p1) (fst p2)) /\ (Y.le (snd p1) (snd p2)). + + Hints Unfold le. + + Lemma le_refl : (p:T)(le p p). + Auto. + Save. + + Lemma le_trans : (p1,p2,p3:T)(le p1 p2) -> (le p2 p3) -> (le p1 p3). + Unfold le. + Intuition; EAuto. + Save. + + Lemma le_antis : (p1,p2:T)(le p1 p2) -> (le p2 p1) -> (p1=p2). + NewDestruct p1. + NewDestruct p2. + Unfold le. + Intuition. + Cut t=t1;Auto. + Cut t0=t2;Auto. + Intros. + Rewrite H0. + Rewrite H4. + Reflexivity. + Save. + + Hints Resolve le_refl le_trans le_antis. + +End Pair. + +Module Check_Pair [X:PO][Y:PO] : PO := (Pair X Y). + + +Module Type Fmono. + Module X:PO. + Module Y:PO. + + Parameter f : X.T -> Y.T. + + Axiom f_mono : (x1,x2:X.T)(X.le x1 x2) -> (Y.le (f x1) (f x2)). +End Fmono. + + +Read Module Nat. + + +Module PlusMono:Fmono. + Module Y:=Nat. + Module X:=Pair Nat Nat. + + Definition f:=[p] (plus (fst p) (snd p)). + + Lemma f_mono : (p1,p2:nat*nat)(X.le p1 p2) -> (le (f p1) (f p2)). + NewDestruct p1;NewDestruct p2. + Unfold X.le Nat.le f. + Simpl. + NewDestruct 1. + Induction H. + + Induction n. + Auto. + Simpl. + Apply Nat.le_mono_S. + Auto. + + Simpl. + Auto. + Save. +End PlusMono. diff --git a/test-suite/modules/Przyklad.v b/test-suite/modules/Przyklad.v new file mode 100644 index 000000000..ecf9a8c25 --- /dev/null +++ b/test-suite/modules/Przyklad.v @@ -0,0 +1,193 @@ +Definition ifte := [T:Set][A:Prop][B:Prop][s:(sumbool A B)][th:T][el:T] + if s then [_]th else [_]el. + +Implicits ifte. + +Lemma Reflexivity_provable : + (A:Set)(a:A)(s:{a=a}+{~a=a})(EXT x| s==(left ? ? x)). +Intros. +Elim s. +Intro x. +Split with x; Reflexivity. + +Intro. +Absurd a=a; Auto. + +Save. + +Lemma Disequality_provable : + (A:Set)(a,b:A)(~a=b)->(s:{a=b}+{~a=b})(EXT x| s==(right ? ? x)). +Intros. +Elim s. +Intro. +Absurd a=a; Auto. + +Intro. +Split with b0; Reflexivity. + +Save. + +Modtype ELEM. + Parameter T : Set. + Parameter eq_dec : (a,a':T){a=a'}+{~ a=a'}. +EndT ELEM. + +Modtype SET[Elt : ELEM]. + Parameter T : Set. + Parameter empty : T. + Parameter add : Elt.T -> T -> T. + Parameter find : Elt.T -> T -> bool. + + (* Axioms *) + + Axiom find_empty_false : + (e:Elt.T) (find e empty) = false. + + Axiom find_add_true : + (s:T) (e:Elt.T) (find e (add e s)) = true. + + Axiom find_add_false : + (s:T) (e:Elt.T) (e':Elt.T) ~(e=e') -> + (find e (add e' s))=(find e s). + +EndT SET. + +Mod FuncDict[E : ELEM]. + Definition T := E.T -> bool. + Definition empty := [e':E.T] false. + Definition find := [e':E.T] [s:T] (s e'). + Definition add := [e:E.T][s:T][e':E.T] + (ifte (E.eq_dec e e') true (find e' s)). + + Lemma find_empty_false : (e:E.T) (find e empty) = false. + Auto. + Qed. + + Lemma find_add_true : + (s:T) (e:E.T) (find e (add e s)) = true. + + Intros. + Unfold find add. + Elim (Reflexivity_provable ? ? (E.eq_dec e e)). + Intros. + Rewrite H. + Auto. + + Qed. + + Lemma find_add_false : + (s:T) (e:E.T) (e':E.T) ~(e=e') -> + (find e (add e' s))=(find e s). + Intros. + Unfold add find. + Cut (EXT x:? | (E.eq_dec e' e)==(right ? ? x)). + Intros. + Elim H0. + Intros. + Rewrite H1. + Unfold ifte. + Reflexivity. + + Apply Disequality_provable. + Auto. + + Qed. + +EndM FuncDict. + +Mod F : SET := FuncDict. + + +Mod Nat. + Definition T:=nat. + Lemma eq_dec : (a,a':T){a=a'}+{~ a=a'}. + Decide Equality. + Qed. +EndM Nat. + + +Mod SetNat:=F Nat. + + +Lemma no_zero_in_empty:(SetNat.find O SetNat.empty)=false. +Apply SetNat.find_empty_false. +Save. + +(***************************************************************************) +Mod Lemmas[G:SET][E:ELEM]. + + Mod ESet:=G E. + + Lemma commute : (S:ESet.T)(a1,a2:E.T) + let S1 = (ESet.add a1 (ESet.add a2 S)) in + let S2 = (ESet.add a2 (ESet.add a1 S)) in + (a:E.T)(ESet.find a S1)=(ESet.find a S2). + + Intros. + Unfold S1 S2. + Elim (E.eq_dec a a1); Elim (E.eq_dec a a2); Intros H1 H2; + Try Rewrite <- H1; Try Rewrite <- H2; + Repeat + (Try (Rewrite ESet.find_add_true; Auto); + Try (Rewrite ESet.find_add_false; Auto); + Auto). + Save. +EndM Lemmas. + + +Inductive list [A:Set] : Set := nil : (list A) + | cons : A -> (list A) -> (list A). + +Mod ListDict[E : ELEM]. + Definition T := (list E.T). + Definition elt := E.T. + + Definition empty := (nil elt). + Definition add := [e:elt][s:T] (cons elt e s). + Fixpoint find [e:elt; s:T] : bool := + Cases s of + nil => false + | (cons e' s') => (ifte (E.eq_dec e e') + true + (find e s')) + end. + + Definition find_empty_false := [e:elt] (refl_equal bool false). + + Lemma find_add_true : + (s:T) (e:E.T) (find e (add e s)) = true. + Intros. + Simpl. + Elim (Reflexivity_provable ? ? (E.eq_dec e e)). + Intros. + Rewrite H. + Auto. + + Qed. + + + Lemma find_add_false : + (s:T) (e:E.T) (e':E.T) ~(e=e') -> + (find e (add e' s))=(find e s). + Intros. + Simpl. + Elim (Disequality_provable ? ? ? H (E.eq_dec e e')). + Intros. + Rewrite H0. + Simpl. + Reflexivity. + Save. + + +EndM ListDict. + + +Mod L : SET := ListDict. + + + + + + + + diff --git a/test-suite/modules/fun_objects.v b/test-suite/modules/fun_objects.v new file mode 100644 index 000000000..b14c3f793 --- /dev/null +++ b/test-suite/modules/fun_objects.v @@ -0,0 +1,32 @@ +Implicit Arguments On. + +Modtype SIG. + Parameter id:(A:Set)A->A. +EndT SIG. + +Mod M[X:SIG]. + Definition idid := (X.id X.id). + Definition id := (idid X.id). +EndM M. + +Mod N:=M. + +Mod Nat. + Definition T := nat. + Definition x := O. + Definition id := [A:Set][x:A]x. +EndM Nat. + +Mod Z:=(N Nat). + +Check (Z.idid O). + +Mod P[Y:SIG] := N. + +Mod Y:=P Nat Z. + +Check (Y.id O). + + + + diff --git a/test-suite/modules/grammar.v b/test-suite/modules/grammar.v new file mode 100644 index 000000000..7751e9f38 --- /dev/null +++ b/test-suite/modules/grammar.v @@ -0,0 +1,15 @@ +Mod N. +Definition f:=plus. +Syntax constr level 7: plus [ (f $n $m)] -> [ $n:L "+" $m:E]. +Check (f O O). +EndM N. +Check (N.f O O). +Imp N. +Check (N.f O O). +Check (f O O). +Mod M:=N. +Check (f O O). +Check (N.f O O). +Imp M. +Check (f O O). +Check (N.f O O). diff --git a/test-suite/modules/modul.v b/test-suite/modules/modul.v new file mode 100644 index 000000000..ba333c525 --- /dev/null +++ b/test-suite/modules/modul.v @@ -0,0 +1,39 @@ +Mod M. + Parameter rel:nat -> nat -> Prop. + + Axiom w : (n:nat)(rel O (S n)). + + Hints Resolve w. + + Grammar constr constr8 := + not_eq [ constr7($a) "#" constr7($b) ] -> [ (rel $a $b) ]. + + Print Hint *. + + Lemma w1 : (O#(S O)). + Auto. + Save. + +EndM M. + +(*Lemma w1 : (M.rel O (S O)). +Auto. +*) + +Imp M. + +Print Hint *. +Lemma w1 : (O#(S O)). +Print Hint. +Print Hint *. + +Auto. +Save. + +Check (O#O). +Locate rel. + +Locate M. + +Mod N:=Scratch.M. + diff --git a/test-suite/modules/modulik.v b/test-suite/modules/modulik.v new file mode 100644 index 000000000..50d985e52 --- /dev/null +++ b/test-suite/modules/modulik.v @@ -0,0 +1,5 @@ +Definition toto := Set. + +Mod M. + Definition id:=[X:Set]X. +EndM M. diff --git a/test-suite/modules/obj.v b/test-suite/modules/obj.v new file mode 100644 index 000000000..b02bdf189 --- /dev/null +++ b/test-suite/modules/obj.v @@ -0,0 +1,26 @@ +Implicit Arguments On. + +Mod M. + Definition a:=[s:Set]s. + Print a. +EndM M. + +Print M.a. + +Mod K. + Definition app:=[A,B:Set; f:(A->B); x:A](f x). + Mod N. + Definition apap:=[A,B:Set](app (app 1!A 2!B)). + Print app. + Print apap. + EndM N. + Print N.apap. +EndM K. + +Print K.app. +Print K.N.apap. + +Mod W:=K.N. + +Print W.apap. + diff --git a/test-suite/modules/objects.v b/test-suite/modules/objects.v new file mode 100644 index 000000000..8183c67c6 --- /dev/null +++ b/test-suite/modules/objects.v @@ -0,0 +1,35 @@ +Reset Initial. + +Modtype SET. + Axiom T:Set. + Axiom x:T. +EndT SET. + +Implicit Arguments On. + +Mod M[X:SET]. + Definition T := nat. + Definition x := O. + Definition f := [A:Set][x:A]X.x. +EndM M. + +Mod N:=M. + +Mod Nat. + Definition T := nat. + Definition x := O. +EndM Nat. + +Mod Z:=(N Nat). + +Check (Z.f O). + +Mod P[Y:SET] := N. + +Mod Y:=P Z Nat. + +Check (Y.f O). + + + + diff --git a/test-suite/modules/pliczek.v b/test-suite/modules/pliczek.v new file mode 100644 index 000000000..6061ace34 --- /dev/null +++ b/test-suite/modules/pliczek.v @@ -0,0 +1,3 @@ +Require Export plik. + +Definition tutu := [X:Set](toto X). diff --git a/test-suite/modules/plik.v b/test-suite/modules/plik.v new file mode 100644 index 000000000..f1f59df0f --- /dev/null +++ b/test-suite/modules/plik.v @@ -0,0 +1,4 @@ +Definition toto:=[x:Set]x. + +Grammar constr constr8 := + toto [ "#" constr7($b) ] -> [ (toto $b) ]. diff --git a/test-suite/modules/sub_objects.v b/test-suite/modules/sub_objects.v new file mode 100644 index 000000000..f800dec52 --- /dev/null +++ b/test-suite/modules/sub_objects.v @@ -0,0 +1,35 @@ +Reset Initial. + +Set Implicit Arguments. + + +Mod M. + Definition id:=[A:Set][x:A]x. + + Modtype SIG. + Parameter idid:(A:Set)A->A. + EndT SIG. + + Mod N. + Definition idid:=[A:Set][x:A](id x). + Grammar constr constr8 := + not_eq [ "#" constr7($b) ] -> [ (idid $b) ]. + Syntactic Definition inc := (plus (S O)). + EndM N. + + Definition zero:=(N.idid O). + +EndM M. + +Definition zero := (M.N.idid O). +Definition jeden := (M.N.inc O). + +Mod Goly:=M.N. + +Definition Gole_zero := (Goly.idid O). +Definition Goly_jeden := (Goly.inc O). + +Mod Ubrany : M.SIG := M.N. + +Definition Ubrane_zero := (Ubrany.idid O). + diff --git a/tools/coq_makefile.ml4 b/tools/coq_makefile.ml4 index 1cd9d350c..891d41e6c 100644 --- a/tools/coq_makefile.ml4 +++ b/tools/coq_makefile.ml4 @@ -177,7 +177,10 @@ let variables l = -I $(COQTOP)/contrib/ring -I $(COQTOP)/contrib/xml \\ -I $(CAMLP4LIB)\n"; print "ZFLAGS=$(OCAMLLIBS) $(COQSRC)\n"; - print "OPT="; if !opt = "-byte" then print "-byte"; print "\n"; + if !opt = "-byte" then + print "override OPT=-byte\n" + else + print "OPT=\n"; print "COQFLAGS=-q $(OPT) $(COQLIBS)\n"; print "COQC=$(COQBIN)coqc\n"; print "GALLINA=gallina\n"; diff --git a/toplevel/cerrors.ml b/toplevel/cerrors.ml index 1fbc3a90e..9d00e63f4 100644 --- a/toplevel/cerrors.ml +++ b/toplevel/cerrors.ml @@ -78,12 +78,12 @@ let rec explain_exn_default = function hov 0 (str "Error:" ++ spc () ++ Himsg.explain_refiner_error e) | Nametab.GlobalizationError q -> hov 0 (str "Error:" ++ spc () ++ - str "The reference" ++ spc () ++ Nametab.pr_qualid q ++ + str "The reference" ++ spc () ++ Libnames.pr_qualid q ++ spc () ++ str "was not found" ++ spc () ++ str "in the current" ++ spc () ++ str "environment") | Nametab.GlobalizationConstantError q -> hov 0 (str "Error:" ++ spc () ++ - str "No constant of this name:" ++ spc () ++ Nametab.pr_qualid q) + str "No constant of this name:" ++ spc () ++ Libnames.pr_qualid q) | Refiner.FailError i -> hov 0 (str "Error: Fail tactic always fails (level " ++ int i ++ str").") diff --git a/toplevel/class.ml b/toplevel/class.ml index 8bac0812e..09013d173 100644 --- a/toplevel/class.ml +++ b/toplevel/class.ml @@ -16,11 +16,13 @@ open Term open Termops open Inductive open Declarations +open Entries open Environ open Inductive open Lib open Classops open Declare +open Libnames open Nametab open Safe_typing @@ -143,7 +145,7 @@ let uniform_cond nargs lt = let id_of_cl = function | CL_FUN -> id_of_string "FUNCLASS" | CL_SORT -> id_of_string "SORTCLASS" - | CL_CONST sp -> basename sp + | CL_CONST kn -> id_of_label (label kn) | CL_IND ind -> let (_,mip) = Global.lookup_inductive ind in mip.mind_typename @@ -259,12 +261,12 @@ let build_id_coercion idf_opt source = (string_of_class (fst (find_class_type t)))) in let constr_entry = (* Cast is necessary to express [val_f] is identity *) - ConstantEntry + DefinitionEntry { const_entry_body = mkCast (val_f, typ_f); const_entry_type = None; const_entry_opaque = false } in - let sp = declare_constant idf (constr_entry,NeverDischarge) in - ConstRef sp + let (_,kn) = declare_constant idf (constr_entry,NeverDischarge) in + ConstRef kn (* nom de la fonction coercion @@ -332,7 +334,7 @@ let try_add_new_coercion_with_source ref stre ~source = let add_coercion_hook stre ref = try_add_new_coercion ref stre; Options.if_verbose message - (string_of_qualid (shortest_qualid_of_global (Global.env()) ref) + (string_of_qualid (shortest_qualid_of_global None ref) ^ " is now a coercion") let add_subclass_hook stre ref = @@ -374,52 +376,49 @@ let count_extra_abstractions hyps ids_to_discard = (hyps,0) ids_to_discard in n -let defined_in_sec sp sec_sp = dirpath sp = sec_sp +let defined_in_sec kn olddir = + let _,dir,_ = repr_kn kn in + dir = olddir (* This moves the global path one step below *) -let process_global sec_sp = function +let process_global olddir = function | VarRef _ -> anomaly "process_global only processes global surviving the section" - | ConstRef sp as x -> - if defined_in_sec sp sec_sp then - let (_,spid) = repr_path sp in - let newsp = Lib.make_path spid in - ConstRef newsp + | ConstRef kn as x -> + if defined_in_sec kn olddir then + let newkn = Lib.make_kn (id_of_label (label kn)) in + ConstRef newkn else x - | IndRef (sp,i) as x -> - if defined_in_sec sp sec_sp then - let (_,spid) = repr_path sp in - let newsp = Lib.make_path spid in - IndRef (newsp,i) + | IndRef (kn,i) as x -> + if defined_in_sec kn olddir then + let newkn = Lib.make_kn (id_of_label (label kn)) in + IndRef (newkn,i) else x - | ConstructRef ((sp,i),j) as x -> - if defined_in_sec sp sec_sp then - let (_,spid) = repr_path sp in - let newsp = Lib.make_path spid in - ConstructRef ((newsp,i),j) + | ConstructRef ((kn,i),j) as x -> + if defined_in_sec kn olddir then + let newkn = Lib.make_kn (id_of_label (label kn)) in + ConstructRef ((newkn,i),j) else x -let process_class sec_sp ids_to_discard x = +let process_class olddir ids_to_discard x = let (cl,{cl_strength=stre; cl_param=p}) = x in (* let env = Global.env () in*) match cl with | CL_SECVAR _ -> x - | CL_CONST sp -> - if defined_in_sec sp sec_sp then - let (_,spid) = repr_path sp in - let newsp = Lib.make_path spid in - let hyps = (Global.lookup_constant sp).const_hyps in + | CL_CONST kn -> + if defined_in_sec kn olddir then + let newkn = Lib.make_kn (id_of_label (label kn)) in + let hyps = (Global.lookup_constant kn).const_hyps in + let n = count_extra_abstractions hyps ids_to_discard in + (CL_CONST newkn,{cl_strength=stre;cl_param=p+n}) + else + x + | CL_IND (kn,i) -> + if defined_in_sec kn olddir then + let newkn = Lib.make_kn (id_of_label (label kn)) in + let hyps = (Global.lookup_mind kn).mind_hyps in let n = count_extra_abstractions hyps ids_to_discard in - (CL_CONST newsp,{cl_strength=stre;cl_param=p+n}) - else - x - | CL_IND (sp,i) -> - if defined_in_sec sp sec_sp then - let (_,spid) = repr_path sp in - let newsp = Lib.make_path spid in - let hyps = (Global.lookup_mind sp).mind_hyps in - let n = count_extra_abstractions hyps ids_to_discard in - (CL_IND (newsp,i),{cl_strength=stre;cl_param=p+n}) + (CL_IND (newkn,i),{cl_strength=stre;cl_param=p+n}) else x | _ -> anomaly "process_class" @@ -427,28 +426,26 @@ let process_class sec_sp ids_to_discard x = let process_cl sec_sp cl = match cl with | CL_SECVAR id -> cl - | CL_CONST sp -> - if defined_in_sec sp sec_sp then - let (_,spid) = repr_path sp in - let newsp = Lib.make_path spid in - CL_CONST newsp + | CL_CONST kn -> + if defined_in_sec kn sec_sp then + let newkn = Lib.make_kn (id_of_label (label kn)) in + CL_CONST newkn else cl - | CL_IND (sp,i) -> - if defined_in_sec sp sec_sp then - let (_,spid) = repr_path sp in - let newsp = Lib.make_path spid in - CL_IND (newsp,i) + | CL_IND (kn,i) -> + if defined_in_sec kn sec_sp then + let newkn = Lib.make_kn (id_of_label (label kn)) in + CL_IND (newkn,i) else cl | _ -> cl -let process_coercion sec_sp ids_to_discard ((coe,coeinfo),cls,clt) = +let process_coercion olddir ids_to_discard ((coe,coeinfo),cls,clt) = let hyps = context_of_global_reference coe in let nargs = count_extra_abstractions hyps ids_to_discard in - (process_global sec_sp coe, + (process_global olddir coe, coercion_strength coeinfo, coercion_identity coeinfo, - process_cl sec_sp cls, - process_cl sec_sp clt, + process_cl olddir cls, + process_cl olddir clt, nargs + coercion_params coeinfo) diff --git a/toplevel/class.mli b/toplevel/class.mli index 388f664e4..8e9bcc3c3 100644 --- a/toplevel/class.mli +++ b/toplevel/class.mli @@ -13,6 +13,7 @@ open Names open Term open Classops open Declare +open Libnames open Nametab (*i*) diff --git a/toplevel/command.ml b/toplevel/command.ml index 34adc0587..c9becbd3a 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -14,6 +14,7 @@ open Options open Term open Termops open Declarations +open Entries open Inductive open Environ open Reduction @@ -21,6 +22,7 @@ open Tacred open Declare open Nametab open Names +open Libnames open Nameops open Coqast open Ast @@ -45,6 +47,13 @@ let rec abstract_rawconstr c = function List.fold_right (fun x b -> mkLambdaC(x,t,b)) idl (abstract_rawconstr c bl) +let rec prod_rawconstr c = function + | [] -> c + | LocalRawDef (x,b)::bl -> mkLetInC(x,b,prod_rawconstr c bl) + | LocalRawAssum (idl,t)::bl -> + List.fold_right (fun x b -> mkProdC(x,t,b)) idl + (prod_rawconstr c bl) + let rec destSubCast c = match kind_of_term c with | Lambda (x,t,c) -> let (b,u) = destSubCast c in mkLambda (x,t,b), mkProd (x,t,u) @@ -83,11 +92,11 @@ let red_constant_entry ce = function reduction_of_redexp red (Global.env()) Evd.empty body } let declare_global_definition ident ce n local = - let sp = declare_constant ident (ConstantEntry ce,n) in + let (_,kn) = declare_constant ident (DefinitionEntry ce,n) in if local then msg_warning (pr_id ident ++ str" is declared as a global definition"); if_verbose message ((string_of_id ident) ^ " is defined"); - ConstRef sp + ConstRef kn let declare_definition ident (local,n) bl red_option c typopt = let ce = constant_entry_of_com (bl,c,typopt,false) in @@ -122,13 +131,14 @@ let syntax_definition ident c = let assumption_message id = if_verbose message ((string_of_id id) ^ " is assumed") -let declare_assumption ident n c = +let declare_assumption ident n bl c = + let c = prod_rawconstr c bl in let c = interp_type Evd.empty (Global.env()) c in match n with | NeverDischarge -> - let r = declare_constant ident (ParameterEntry c, NeverDischarge) in + let (_,kn) = declare_constant ident (ParameterEntry c, NeverDischarge) in assumption_message ident; - ConstRef r + ConstRef kn | DischargeAt (disch_sp,_) -> if Lib.is_section_p disch_sp then begin let r = declare_variable ident (Lib.cwd(),SectionLocalAssum c,n) in @@ -139,12 +149,12 @@ let declare_assumption ident n c = VarRef ident end else - let r = declare_constant ident (ParameterEntry c, NeverDischarge) in + let (_,kn) = declare_constant ident (ParameterEntry c, NeverDischarge) in assumption_message ident; if_verbose msg_warning (pr_id ident ++ str" is declared as a parameter" ++ str" because it is at a global level"); - ConstRef r + ConstRef kn | NotDeclare -> anomalylabstrm "Command.hypothesis_def_var" (str "Strength NotDeclare not for Variable, only for Let") @@ -230,10 +240,10 @@ let interp_mutual lparams lnamearconstrs finite = let declare_mutual_with_eliminations mie = let lrecnames = List.map (fun e -> e.mind_entry_typename) mie.mind_entry_inds in - let sp = declare_mind mie in + let (_,kn) = declare_mind mie in if_verbose ppnl (minductive_message lrecnames); - Indrec.declare_eliminations sp; - sp + Indrec.declare_eliminations kn; + kn let eq_la (id,ast) (id',ast') = id = id' & alpha_eq(ast,ast') @@ -351,8 +361,8 @@ let build_recursive lnameargsardef = recvec)); const_entry_type = None; const_entry_opaque = false } in - let sp = declare_constant fi (ConstantEntry ce, n) in - (ConstRef sp) + let (_,kn) = declare_constant fi (DefinitionEntry ce, n) in + (ConstRef kn) in (* declare the recursive definitions *) let lrefrec = Array.mapi declare namerec in @@ -365,7 +375,7 @@ let build_recursive lnameargsardef = let ce = { const_entry_body = replace_vars subst def; const_entry_type = Some t; const_entry_opaque = false } in - let _ = declare_constant f (ConstantEntry ce,n) in + let _ = declare_constant f (DefinitionEntry ce,n) in warning ((string_of_id f)^" is non-recursively defined"); (var_subst f) :: subst) (List.map var_subst (Array.to_list namerec)) @@ -415,8 +425,8 @@ let build_corecursive lnameardef = const_entry_type = None; const_entry_opaque = false } in - let sp = declare_constant fi (ConstantEntry ce,n) in - (ConstRef sp) + let _,kn = declare_constant fi (DefinitionEntry ce,n) in + (ConstRef kn) in let lrefrec = Array.mapi declare namerec in if_verbose ppnl (corecursive_message lrefrec); @@ -427,7 +437,7 @@ let build_corecursive lnameardef = let ce = { const_entry_body = replace_vars subst def; const_entry_type = Some t; const_entry_opaque = false } in - let _ = declare_constant f (ConstantEntry ce,n) in + let _ = declare_constant f (DefinitionEntry ce,n) in warning ((string_of_id f)^" is non-recursively defined"); (var_subst f) :: subst) (List.map var_subst (Array.to_list namerec)) @@ -453,8 +463,8 @@ let build_scheme lnamedepindsort = let ce = { const_entry_body = decl; const_entry_type = None; const_entry_opaque = false } in - let sp = declare_constant fi (ConstantEntry ce,n) in - ConstRef sp :: lrecref + let _,kn = declare_constant fi (DefinitionEntry ce,n) in + ConstRef kn :: lrecref in let lrecref = List.fold_right2 declare listdecl lrecnames [] in if_verbose ppnl (recursive_message (Array.of_list lrecref)) @@ -484,17 +494,18 @@ let apply_tac_not_declare id pft = function Pfedit.delete_current_proof (); Pfedit.by (Refiner.tclTHENSV cutt [|introduction id;exat|]) -let save id const strength hook = +let save id const (local,strength) hook = let {const_entry_body = pft; const_entry_type = tpo; const_entry_opaque = opacity } = const in begin match strength with - | DischargeAt (disch_sp,_) when Lib.is_section_p disch_sp && not opacity -> + | DischargeAt (disch_sp,_) when Lib.is_section_p disch_sp && local -> let c = SectionLocalDef (pft, tpo) in let _ = declare_variable id (Lib.cwd(), c, strength) in hook strength (VarRef id) | NeverDischarge | DischargeAt _ -> - let ref = ConstRef (declare_constant id (ConstantEntry const,strength)) in + let _,kn = declare_constant id (DefinitionEntry const,strength) in + let ref = ConstRef kn in hook strength ref | NotDeclare -> apply_tac_not_declare id pft tpo end; @@ -505,9 +516,9 @@ let save id const strength hook = end let save_named opacity = - let id,(const,(local,strength),hook) = Pfedit.cook_proof () in + let id,(const,persistence,hook) = Pfedit.cook_proof () in let const = { const with const_entry_opaque = opacity } in - save id const strength hook + save id const persistence hook let check_anonymity id save_ident = if atompart_of_id id <> "Unnamed_thm" then @@ -517,16 +528,17 @@ let check_anonymity id save_ident = *) let save_anonymous opacity save_ident = - let id,(const,(local,strength),hook) = Pfedit.cook_proof () in + let id,(const,persistence,hook) = Pfedit.cook_proof () in let const = { const with const_entry_opaque = opacity } in check_anonymity id save_ident; - save save_ident const strength hook + save save_ident const persistence hook let save_anonymous_with_strength strength opacity save_ident = let id,(const,_,hook) = Pfedit.cook_proof () in let const = { const with const_entry_opaque = opacity } in check_anonymity id save_ident; - save save_ident const strength hook + (* we consider that non opaque behaves as local for discharge *) + save save_ident const (not opacity,strength) hook let get_current_context () = try Pfedit.get_current_goal_context () diff --git a/toplevel/command.mli b/toplevel/command.mli index 1b8cbe49a..23c0db5b4 100644 --- a/toplevel/command.mli +++ b/toplevel/command.mli @@ -12,9 +12,12 @@ open Util open Names open Term -open Nametab +open Declare open Library +open Libnames +open Nametab open Vernacexpr + (*i*) (*s Declaration functions. The following functions take ASTs, @@ -28,12 +31,13 @@ val declare_definition : identifier -> bool * strength -> val syntax_definition : identifier -> Coqast.t -> unit -val declare_assumption : identifier -> strength -> Coqast.t -> global_reference +val declare_assumption : identifier -> strength -> + local_binder list -> Coqast.t -> global_reference val build_mutual : Vernacexpr.inductive_expr list -> bool -> unit val declare_mutual_with_eliminations : - Indtypes.mutual_inductive_entry -> section_path + Entries.mutual_inductive_entry -> mutual_inductive val build_recursive : (identifier * ((identifier * Coqast.t) list) * Coqast.t * Coqast.t) list @@ -41,7 +45,7 @@ val build_recursive : val build_corecursive : (identifier * Coqast.t * Coqast.t) list -> unit -val build_scheme : (identifier * bool * Nametab.qualid located * Coqast.t) list -> unit +val build_scheme : (identifier * bool * qualid located * Coqast.t) list -> unit val start_proof_com : identifier option -> bool * strength -> Coqast.t -> Proof_type.declaration_hook -> unit diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml index a62c1fae4..bcee6d0ce 100644 --- a/toplevel/coqtop.ml +++ b/toplevel/coqtop.ml @@ -13,6 +13,7 @@ open Util open System open Options open Names +open Libnames open Nameops open States open Toplevel @@ -62,12 +63,12 @@ let load_vernacular () = let load_vernacular_obj = ref ([] : string list) let add_vernac_obj s = load_vernacular_obj := s :: !load_vernacular_obj let load_vernac_obj () = - List.iter Library.read_module_from_file (List.rev !load_vernacular_obj) + List.iter Library.read_library_from_file (List.rev !load_vernacular_obj) let require_list = ref ([] : string list) let add_require s = require_list := s :: !require_list let require () = - List.iter (fun s -> Library.require_module_from_file None None s false) + List.iter (fun s -> Library.require_library_from_file None None s false) (List.rev !require_list) let compile_list = ref ([] : (bool * string) list) diff --git a/toplevel/discharge.ml b/toplevel/discharge.ml index 14e4feff5..93277300b 100644 --- a/toplevel/discharge.ml +++ b/toplevel/discharge.ml @@ -15,11 +15,13 @@ open Nameops open Sign open Term open Declarations +open Entries open Inductive open Instantiate open Reduction open Cooking open Typeops +open Libnames open Libobject open Lib open Nametab @@ -33,7 +35,11 @@ open Indtypes open Nametab let recalc_sp dir sp = - let (_,spid) = repr_path sp in Names.make_path dir spid + let (_,spid) = repr_path sp in Libnames.make_path dir spid + +let recalc_kn dir kn = + let (mp,_,l) = Names.repr_kn kn in + Names.make_kn mp dir l let rec find_var id = function | [] -> false @@ -93,8 +99,8 @@ let abstract_inductive ids_to_abs hyps inds = let params' = List.map (function - | (Name id,None,p) -> id, LocalAssum p - | (Name id,Some p,_) -> id, LocalDef p + | (Name id,None,p) -> id, Entries.LocalAssum p + | (Name id,Some p,_) -> id, Entries.LocalDef p | (Anonymous,_,_) -> anomaly"Unnamed inductive local variable") params in { mind_entry_params = params'; @@ -170,20 +176,21 @@ type opacity = bool type discharge_operation = | Variable of identifier * section_variable_entry * strength * bool - | Constant of section_path * recipe * strength * bool + | Constant of identifier * recipe * strength * constant * bool | Inductive of mutual_inductive_entry * bool | Class of cl_typ * cl_info_typ | Struc of inductive * (unit -> struc_typ) | Objdef of constant | Coercion of coercion_entry - | Require of module_reference + | Require of library_reference | Constraints of Univ.constraints (* Main function to traverse the library segment and compute the various discharge operations. *) -let process_object oldenv dir sec_sp - (ops,ids_to_discard,(constl,indl,cstrl as work_alist)) (sp,lobj) = +let process_object oldenv olddir full_olddir newdir +(* {dir -> newdir} {sec_sp -> full_olddir, olddir} *) + (ops,ids_to_discard,(constl,indl,cstrl as work_alist)) ((sp,kn),lobj) = let tag = object_tag lobj in match tag with | "VARIABLE" -> @@ -222,60 +229,63 @@ let process_object oldenv dir sec_sp (ops, ids_to_discard, (constl,indl,cstrl)) else *) - let cb = Environ.lookup_constant sp oldenv in - let imp = is_implicit_constant sp in - let newsp = match stre with - | DischargeAt (d,_) when not (is_dirpath_prefix_of d dir) -> sp - | _ -> recalc_sp dir sp in + let kn = Nametab.locate_constant (qualid_of_sp sp) in + let lab = label kn in + let cb = Environ.lookup_constant kn oldenv in + let imp = is_implicit_constant kn in + let newkn = (*match stre with (* this did not work anyway...*) + | DischargeAt (d,_) when not (is_dirpath_prefix_of d dir) -> kn + | _ -> *)recalc_kn newdir kn in let mods = let abs_vars = build_abstract_list cb.const_hyps ids_to_discard in - [ (sp, DO_ABSTRACT(newsp,abs_vars)) ] + [ (kn, DO_ABSTRACT(newkn,abs_vars)) ] in let r = { d_from = cb; d_modlist = work_alist; d_abstract = ids_to_discard } in - let op = Constant (newsp,r,stre,imp) in + let op = Constant (id_of_label lab, r,stre,newkn,imp) in (op :: ops, ids_to_discard, (mods@constl, indl, cstrl)) | "INDUCTIVE" -> - let mib = Environ.lookup_mind sp oldenv in - let newsp = recalc_sp dir sp in + let kn = Nametab.locate_mind (qualid_of_sp sp) in + let mib = Environ.lookup_mind kn oldenv in + let newkn = recalc_kn newdir kn in let imp = is_implicit_args() (* CHANGE *) in let (mie,indmods,cstrmods) = - process_inductive sp newsp oldenv (ids_to_discard,work_alist) mib in + process_inductive kn newkn oldenv (ids_to_discard,work_alist) mib in ((Inductive(mie,imp)) :: ops, ids_to_discard, (constl,indmods@indl,cstrmods@cstrl)) | "CLASS" -> let ((cl,clinfo) as x) = outClass lobj in - if (match clinfo.cl_strength with DischargeAt (sp,_) -> sp = sec_sp | _ -> false) then + if (match clinfo.cl_strength with DischargeAt (dp,_) -> dp = full_olddir | _ -> false) then (ops,ids_to_discard,work_alist) else - let (y1,y2) = process_class sec_sp ids_to_discard x in + let (y1,y2) = process_class olddir ids_to_discard x in ((Class (y1,y2))::ops, ids_to_discard, work_alist) | "COERCION" -> let (((_,coeinfo),_,_)as x) = outCoercion lobj in - if (match coercion_strength coeinfo with DischargeAt (sp,_) -> sp = sec_sp | _ -> false) then + if (match coercion_strength coeinfo with DischargeAt (dp,_) -> dp = full_olddir | _ -> false) then (ops,ids_to_discard,work_alist) else - let y = process_coercion sec_sp ids_to_discard x in + let y = process_coercion olddir ids_to_discard x in ((Coercion y)::ops, ids_to_discard, work_alist) | "STRUCTURE" -> - let ((sp,i),info) = outStruc lobj in - let newsp = recalc_sp dir sp in + let ((kn,i),info) = outStruc lobj in + let newkn = recalc_kn newdir kn in let strobj () = - let mib = Environ.lookup_mind newsp (Global.env ()) in + let mib = Environ.lookup_mind newkn (Global.env ()) in { s_CONST = info.s_CONST; s_PARAM = mib.mind_packets.(0).mind_nparams; - s_PROJ = List.map (option_app (recalc_sp dir)) info.s_PROJ } in - ((Struc ((newsp,i),strobj))::ops, ids_to_discard, work_alist) + s_PROJ = List.map (option_app (fun kn -> recalc_kn newdir kn)) info.s_PROJ } in + ((Struc ((newkn,i),strobj))::ops, ids_to_discard, work_alist) | "OBJDEF1" -> - let sp = outObjDef1 lobj in - let new_sp = recalc_sp dir sp in - ((Objdef new_sp)::ops, ids_to_discard, work_alist) + let kn = outObjDef1 lobj in + let new_kn = recalc_kn newdir kn in + ((Objdef new_kn)::ops, ids_to_discard, work_alist) | "REQUIRE" -> let c = out_require lobj in @@ -283,8 +293,9 @@ let process_object oldenv dir sec_sp | _ -> (ops,ids_to_discard,work_alist) -let process_item oldenv dir sec_sp acc = function - | (sp,Leaf lobj) -> process_object oldenv dir sec_sp acc (sp,lobj) +let process_item oldenv olddir full_olddir newdir acc = function + | (sp,Leaf lobj) -> + process_object oldenv olddir full_olddir newdir acc (sp,lobj) | (_,_) -> acc let process_operation = function @@ -293,9 +304,9 @@ let process_operation = function let _ = with_implicits imp (declare_variable id) (Lib.cwd(),expmod_a,stre) in () - | Constant (sp,r,stre,imp) -> - with_implicits imp (redeclare_constant sp) (r,stre); - constant_message (basename sp) + | Constant (id,r,stre,kn,imp) -> + with_implicits imp (redeclare_constant id) (r,stre); + constant_message id | Inductive (mie,imp) -> let _ = with_implicits imp declare_mind mie in inductive_message mie.mind_entry_inds @@ -306,7 +317,7 @@ let process_operation = function | Objdef newsp -> begin try Recordobj.objdef_declare (ConstRef newsp) with _ -> () end | Coercion y -> add_new_coercion y - | Require y -> reload_module y + | Require y -> reload_library y | Constraints y -> Global.add_constraints y let catch_not_found f x = @@ -317,13 +328,14 @@ let catch_not_found f x = let close_section _ s = let oldenv = Global.env() in - let olddir,decls,fs = close_section false s in + let prefix,decls,fs = close_section false s in + let full_olddir, (_,olddir) = prefix in let newdir = fst (split_dirpath olddir) in let (ops,ids,_) = List.fold_left - (process_item oldenv newdir olddir) ([],[],([],[],[])) decls + (process_item oldenv olddir full_olddir newdir) ([],[],([],[],[])) decls in let ids = last_section_hyps olddir in Summary.unfreeze_lost_summaries fs; catch_not_found (List.iter process_operation) (List.rev ops); - Nametab.push_section olddir + Nametab.push_dir (Until 1) full_olddir (DirClosedSection full_olddir) diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index c2892a74b..82991495d 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -448,6 +448,11 @@ let error_same_names_constructors id cid = str "is used twice is the definition of type" ++ spc () ++ pr_id id +let error_same_names_overlap idl = + str "The following names" ++ spc () ++ + str "are used both as type names and constructor names:" ++ spc () ++ + prlist_with_sep pr_coma pr_id idl + let error_not_an_arity id = str "The type of" ++ spc () ++ pr_id id ++ spc () ++ str "is not an arity." @@ -477,6 +482,7 @@ let explain_inductive_error = function | NonPar (env,c,n,v1,v2) -> error_bad_ind_parameters env c n v1 v2 | SameNamesTypes id -> error_same_names_types id | SameNamesConstructors (id,cid) -> error_same_names_constructors id cid + | SameNamesOverlap idl -> error_same_names_overlap idl | NotAnArity id -> error_not_an_arity id | BadEntry -> error_bad_entry () (* These are errors related to recursors *) diff --git a/toplevel/metasyntax.ml b/toplevel/metasyntax.ml index cabdee95c..523b81b41 100644 --- a/toplevel/metasyntax.ml +++ b/toplevel/metasyntax.ml @@ -59,13 +59,16 @@ let _ = (* Pretty-printing objects = syntax_entry *) let cache_syntax (_,ppobj) = Esyntax.add_ppobject ppobj +let subst_syntax (_,subst,ppobj) = + Extend.subst_syntax_command Ast.subst_astpat subst ppobj + let (inPPSyntax,outPPSyntax) = - declare_object - ("PPSYNTAX", - { load_function = (fun _ -> ()); - open_function = cache_syntax; + declare_object {(default_object "PPSYNTAX") with + open_function = (fun i o -> if i=1 then cache_syntax o); cache_function = cache_syntax; - export_function = (fun x -> Some x) }) + subst_function = subst_syntax; + classify_function = (fun (_,o) -> Substitute o); + export_function = (fun x -> Some x) } (* Syntax extension functions (registered in the environnement) *) @@ -93,25 +96,27 @@ let _ = let cache_token (_,s) = Pcoq.lexer.Token.using ("", s) let (inToken, outToken) = - declare_object - ("TOKEN", - { load_function = (fun _ -> ()); - open_function = cache_token; + declare_object {(default_object "TOKEN") with + open_function = (fun i o -> if i=1 then cache_token o); cache_function = cache_token; - export_function = (fun x -> Some x)}) + subst_function = Libobject.ident_subst_function; + classify_function = (fun (_,o) -> Substitute o); + export_function = (fun x -> Some x)} let add_token_obj s = Lib.add_anonymous_leaf (inToken s) (* Grammar rules *) let cache_grammar (_,a) = Egrammar.extend_grammar a +let subst_grammar (_,subst,a) = Egrammar.subst_all_grammar_command subst a + let (inGrammar, outGrammar) = - declare_object - ("GRAMMAR", - { load_function = (fun _ -> ()); - open_function = cache_grammar; + declare_object {(default_object "GRAMMAR") with + open_function = (fun i o -> if i=1 then cache_grammar o); cache_function = cache_grammar; - export_function = (fun x -> Some x)}) + subst_function = subst_grammar; + classify_function = (fun (_,o) -> Substitute o); + export_function = (fun x -> Some x)} let gram_define_entry (u,_ as univ) ((ntl,nt),et,assoc,rl) = let etyp = match et with None -> entry_type_from_name u | Some e -> e in @@ -164,13 +169,17 @@ let cache_infix (_,(gr,se)) = Egrammar.extend_grammar gr; Esyntax.add_ppobject {sc_univ="constr";sc_entries=se} +let subst_infix (_,subst,(gr,se)) = + (Egrammar.subst_all_grammar_command subst gr, + list_smartmap (Extend.subst_syntax_entry Ast.subst_astpat subst) se) + let (inInfix, outInfix) = - declare_object - ("INFIX", - { load_function = (fun _ -> ()); - open_function = cache_infix; + declare_object {(default_object "INFIX") with + open_function = (fun i o -> if i=1 then cache_infix o); cache_function = cache_infix; - export_function = (fun x -> Some x)}) + subst_function = subst_infix; + classify_function = (fun (_,o) -> Substitute o); + export_function = (fun x -> Some x)} (* Build the syntax and grammar rules *) @@ -261,27 +270,31 @@ let make_infix_symbols assoc n sl = NonTerminal ((n,lp),"$a")::(List.map (fun s -> Terminal s) sl) @[NonTerminal ((n,rp),"$b")] +let subst_infix2 (_, subst, (ref,inf as node)) = + let ref' = Libnames.subst_ext subst ref in + if ref' == ref then node else + (ref', inf) let cache_infix2 (_,(ref,inf)) = let sp = match ref with - | Nametab.TrueGlobal r -> Nametab.sp_of_global (Global.env()) r - | Nametab.SyntacticDef sp -> sp in + | Libnames.TrueGlobal r -> Nametab.sp_of_global None r + | Libnames.SyntacticDef kn -> Nametab.sp_of_syntactic_definition kn in declare_infix_symbol sp inf let (inInfix2, outInfix2) = - declare_object - ("INFIX2", - { load_function = (fun _ -> ()); - open_function = cache_infix2; + declare_object {(default_object "INFIX2") with + open_function = (fun i o -> if i=1 then cache_infix2 o); cache_function = cache_infix2; - export_function = (fun x -> Some x)}) + subst_function = subst_infix2; + classify_function = (fun (_,o) -> Substitute o); + export_function = (fun x -> Some x)} let add_infix assoc n inf (loc,qid) = let ref = try Nametab.extended_locate qid with Not_found -> - error ("Unknown reference: "^(Nametab.string_of_qualid qid)) in + error ("Unknown reference: "^(Libnames.string_of_qualid qid)) in let pr = Astterm.ast_of_extended_ref_loc loc ref in (* check the precedence *) if n<1 or n>10 then diff --git a/toplevel/metasyntax.mli b/toplevel/metasyntax.mli index 1eeec61b0..839e0fdbe 100644 --- a/toplevel/metasyntax.mli +++ b/toplevel/metasyntax.mli @@ -22,7 +22,7 @@ val add_token_obj : string -> unit val add_tactic_grammar : (string * (string * grammar_production list) * Tacexpr.raw_tactic_expr) list -> unit val add_infix : - Gramext.g_assoc option -> int -> string -> Nametab.qualid Util.located -> unit + Gramext.g_assoc option -> int -> string -> Libnames.qualid Util.located -> unit val add_distfix : Gramext.g_assoc option -> int -> string -> Coqast.t -> unit diff --git a/toplevel/mltop.ml4 b/toplevel/mltop.ml4 index 0b86c7656..20d87306a 100644 --- a/toplevel/mltop.ml4 +++ b/toplevel/mltop.ml4 @@ -41,7 +41,7 @@ open Vernacinterp put all the needed names into the ML Module object.) *) (* This path is where we look for .cmo *) -let coq_mlpath_copy = ref [] +let coq_mlpath_copy = ref ["."] let keep_copy_mlpath s = let dir = glob s in coq_mlpath_copy := dir :: !coq_mlpath_copy @@ -283,11 +283,10 @@ let cache_ml_module_object (_,{mnames=mnames}) = let export_ml_module_object x = Some x let (inMLModule,outMLModule) = - declare_object ("ML-MODULE", - { load_function = cache_ml_module_object; + declare_object {(default_object "ML-MODULE") with + load_function = (fun _ -> cache_ml_module_object); cache_function = cache_ml_module_object; - open_function = (fun _ -> ()); - export_function = export_ml_module_object }) + export_function = export_ml_module_object } let declare_ml_modules l = Lib.add_anonymous_leaf (inMLModule {mnames=l}) diff --git a/toplevel/record.ml b/toplevel/record.ml index 10b9c42da..ab3482c91 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -11,11 +11,13 @@ open Pp open Util open Names +open Libnames open Nameops open Term open Termops open Environ open Declarations +open Entries open Declare open Nametab open Coqast @@ -49,7 +51,7 @@ let interp_decl sigma env = function (Name id,Some j.uj_val, j.uj_type) let build_decl_entry sigma env (id,t) = - (id,Typeops.LocalAssum (interp_type Evd.empty env t)) + (id,Entries.LocalAssum (interp_type Evd.empty env t)) let typecheck_params_and_fields ps fs = let env0 = Global.env () in @@ -70,8 +72,8 @@ let typecheck_params_and_fields ps fs = newps, newfs let degenerate_decl id = function - (_,None,t) -> (id,Typeops.LocalAssum t) - | (_,Some c,t) -> (id,Typeops.LocalDef c) + (_,None,t) -> (id,Entries.LocalAssum t) + | (_,Some c,t) -> (id,Entries.LocalDef c) type record_error = | MissingProj of identifier * identifier list @@ -184,7 +186,7 @@ let declare_projections indsp coers fields = const_entry_type = None; const_entry_opaque = false } in let sp = - declare_constant fid (ConstantEntry cie,NeverDischarge) + declare_constant fid (DefinitionEntry cie,NeverDischarge) in Some sp with Type_errors.TypeError (ctx,te) -> begin warning_or_error coe indsp (BadTypedProj (fid,ctx,te)); @@ -193,9 +195,9 @@ let declare_projections indsp coers fields = match name with | None -> (nfi-1, None::sp_projs, NoProjection fi::subst) - | Some sp -> - let refi = ConstRef sp in - let constr_fi = mkConst sp in + | Some (sp,kn) -> + let refi = ConstRef kn in + let constr_fi = mkConst kn in if coe then begin let cl = Class.class_of_ref (IndRef indsp) in Class.try_add_new_coercion_with_source @@ -203,7 +205,7 @@ let declare_projections indsp coers fields = end; let proj_args = (*Rel 1 refers to "x"*) paramargs@[mkRel 1] in let constr_fip = applist (constr_fi,proj_args) in - (nfi-1, name::sp_projs, Projection constr_fip::subst)) + (nfi-1, (Some kn)::sp_projs, Projection constr_fip::subst)) (List.length fields,[],[]) coers (List.rev fields) in sp_projs diff --git a/toplevel/recordobj.ml b/toplevel/recordobj.ml index 92c8e5524..58fbb9781 100755 --- a/toplevel/recordobj.ml +++ b/toplevel/recordobj.ml @@ -11,6 +11,7 @@ open Util open Pp open Names +open Libnames open Nameops open Term open Instantiate @@ -31,7 +32,7 @@ let typ_lams_of t = let objdef_err ref = errorlabstrm "object_declare" - (pr_id (Termops.id_of_global (Global.env()) ref) ++ + (pr_id (id_of_global None ref) ++ str" is not a structure object") let objdef_declare ref = diff --git a/toplevel/recordobj.mli b/toplevel/recordobj.mli index 590482653..71dbc6816 100755 --- a/toplevel/recordobj.mli +++ b/toplevel/recordobj.mli @@ -8,7 +8,5 @@ (* $Id$ *) -open Nametab - -val objdef_declare : global_reference -> unit +val objdef_declare : Libnames.global_reference -> unit val add_object_hook : Proof_type.declaration_hook diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index ed2c3c154..937d05a22 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -156,15 +156,21 @@ let load_vernac verb file = (* Compile a vernac file (f is assumed without .v suffix) *) let compile verbosely f = - let s = Filename.basename f in - let m = Names.id_of_string s in - let _,longf = find_file_in_path (Library.get_load_path ()) (f^".v") in - let ldir0 = Library.find_logical_path (Filename.dirname longf) in - let ldir = Nameops.extend_dirpath ldir0 m in - Termops.set_module ldir; (* Just for universe naming *) - Lib.start_module ldir; +(* + let s = Filename.basename f in + let m = Names.id_of_string s in + let _,longf = find_file_in_path (Library.get_load_path ()) (f^".v") in + let ldir0 = Library.find_logical_path (Filename.dirname longf) in + let ldir = Libnames.extend_dirpath ldir0 m in + Termops.set_module ldir; (* Just for universe naming *) + Lib.start_module ldir; + if !dump then dump_string ("F" ^ Names.string_of_dirpath ldir ^ "\n"); + let _ = load_vernac verbosely longf in + let mid = Lib.end_module m in + assert (mid = ldir); + Library.save_module_to ldir (longf^"o") +*) + let ldir,long_f_dot_v = Library.start_library f in if !dump then dump_string ("F" ^ Names.string_of_dirpath ldir ^ "\n"); - let _ = load_vernac verbosely longf in - let mid = Lib.end_module m in - assert (mid = ldir); - Library.save_module_to ldir (longf^"o") + let _ = load_vernac verbosely long_f_dot_v in + Library.save_library_to ldir (long_f_dot_v^"o") diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index 3852ba7fe..6e0ba7087 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -14,6 +14,7 @@ open Pp open Util open Options open Names +open Entries open Nameops open Term open Pfedit @@ -25,6 +26,8 @@ open Printer open Tacinterp open Command open Goptions +(*open Declare*) +open Libnames open Nametab open Vernacexpr @@ -161,8 +164,8 @@ let print_loadpath () = prlist_with_sep pr_fnl print_path_entry l)) let print_modules () = - let opened = Library.opened_modules () - and loaded = Library.loaded_modules () in + let opened = Library.opened_libraries () + and loaded = Library.loaded_libraries () in let loaded_opened = list_intersect loaded opened and only_loaded = list_subtract loaded opened in str"Loaded and imported modules: " ++ @@ -194,14 +197,26 @@ let locate_file f = let print_located_qualid (_,qid) = try - let sp = Nametab.sp_of_global (Global.env()) (Nametab.locate qid) in + let sp = Nametab.sp_of_global None (Nametab.locate qid) in msgnl (pr_sp sp) with Not_found -> try - msgnl (pr_sp (Syntax_def.locate_syntactic_definition qid) ++ fnl ()) + let kn = Nametab.locate_syntactic_definition qid in + let sp = Nametab.sp_of_syntactic_definition kn in + msgnl (pr_sp sp) with Not_found -> msgnl (pr_qualid qid ++ str " is not a defined object") +(*let print_path_entry (s,l) = + (str s ++ tbrk (0,2) ++ str (string_of_dirpath l)) + +let print_loadpath () = + let l = Library.get_full_load_path () in + msgnl (Pp.t (str "Physical path: " ++ + tab () ++ str "Logical Path:" ++ fnl () ++ + prlist_with_sep pr_fnl print_path_entry l)) +*) + let msg_found_library = function | Library.LibLoaded, fulldir, file -> msgnl(pr_dirpath fulldir ++ str " has been loaded from file" ++ fnl () ++ @@ -210,13 +225,13 @@ let msg_found_library = function msgnl(pr_dirpath fulldir ++ str " is bound to file " ++ str file) let msg_notfound_library loc qid = function | Library.LibUnmappedDir -> - let dir = fst (Nametab.repr_qualid qid) in + let dir = fst (repr_qualid qid) in user_err_loc (loc,"locate_library", str "Cannot find a physical path bound to logical path " ++ pr_dirpath dir ++ fnl ()) | Library.LibNotFound -> msgnl (hov 0 - (str"Unable to locate library" ++ spc () ++ Nametab.pr_qualid qid)) + (str"Unable to locate library" ++ spc () ++ pr_qualid qid)) | e -> assert false let print_located_library (loc,qid) = @@ -232,8 +247,8 @@ let vernac_syntax = Metasyntax.add_syntax_obj let vernac_grammar = Metasyntax.add_grammar_obj let vernac_infix assoc n inf qid = - let sp = sp_of_global (Global.env()) (global qid) in - let dir = repr_dirpath (Nameops.dirpath sp) in + let sp = sp_of_global None (global qid) in + let dir = repr_dirpath (dirpath sp) in (* if dir <> [] then begin let modname = List.hd dir in @@ -252,14 +267,14 @@ let vernac_distfix assoc n inf qid = let interp_assumption = function | (AssumptionHypothesis | AssumptionVariable) -> Declare.make_strength_0 () - | (AssumptionAxiom | AssumptionParameter) -> Nametab.NeverDischarge + | (AssumptionAxiom | AssumptionParameter) -> Libnames.NeverDischarge let interp_definition = function - | Definition -> (false, Nametab.NeverDischarge) + | Definition -> (false, Libnames.NeverDischarge) | LocalDefinition -> (true, Declare.make_strength_0 ()) let interp_theorem = function - | (Theorem | Lemma | Decl) -> Nametab.NeverDischarge + | (Theorem | Lemma | Decl) -> Libnames.NeverDischarge | Fact -> Declare.make_strength_1 () | Remark -> Declare.make_strength_0 () @@ -282,10 +297,14 @@ let rec generalize_rawconstr c = function let vernac_definition kind id def hook = let (local,stre as k) = interp_definition kind in match def with - | ProveBody (bl,t) -> - let hook _ _ = () in - let t = generalize_rawconstr t bl in - start_proof_and_print (Some id) k t hook + | ProveBody (bl,t) -> (* local binders, typ *) + if Lib.is_specification () then + let ref = declare_assumption id stre bl t in + hook stre ref + else + let hook _ _ = () in + let t = generalize_rawconstr t bl in + start_proof_and_print (Some id) k t hook | DefineBody (bl,red_option,c,typ_opt) -> let red_option = match red_option with | None -> None @@ -299,10 +318,16 @@ let vernac_start_proof kind sopt t lettop hook = if not(refining ()) then if lettop then errorlabstrm "Vernacentries.StartProof" - (str "Let declarations can only be used in proof editing mode") -(* else if s = None then - error "repeated Goal not permitted in refining mode"*); - start_proof_and_print sopt (false, interp_theorem kind) t hook + (str "Let declarations can only be used in proof editing mode"); + let stre = interp_theorem kind in + match Lib.is_specification (), sopt with + | true, Some id -> + let ref = declare_assumption id stre [] t in + hook stre ref + | _ -> + (* an explicit Goal command starts the refining mode + even in a specification *) + start_proof_and_print sopt (false, stre) t hook let vernac_end_proof is_opaque idopt = if_verbose show_script (); @@ -323,7 +348,7 @@ let vernac_assumption kind l = let stre = interp_assumption kind in List.iter (fun (is_coe,(id,c)) -> - let r = declare_assumption id stre c in + let r = declare_assumption id stre [] c in if is_coe then Class.try_add_new_coercion r stre) l let vernac_inductive f indl = build_mutual indl f @@ -334,6 +359,92 @@ let vernac_cofixpoint = build_corecursive let vernac_scheme = build_scheme +(**********************) +(* Modules *) + +let vernac_declare_module id bl mty_ast_o mexpr_ast_o = + let evd = Evd.empty in + let env = Global.env () in + let arglist,base_mty_o,base_mexpr_o = + Astmod.interp_module_decl evd env bl mty_ast_o mexpr_ast_o + in + let argids, args = List.split arglist + in (* We check the state of the system (in module, in module type) + and what parts are supplied *) + match Lib.is_specification (), base_mty_o, base_mexpr_o with + | _, None, None + | false, _, None -> + Declaremods.start_module id argids args base_mty_o; + if_verbose message + ("Interactive Module "^ string_of_id id ^" started") + + | true, Some _, None + | false, _, Some _ -> + let mexpr_o = + option_app + (List.fold_right + (fun (mbid,mte) me -> MEfunctor(mbid,mte,me)) + args) + base_mexpr_o + in + let mty_o = + option_app + (List.fold_right + (fun (arg_id,arg_t) mte -> MTEfunsig(arg_id,arg_t,mte)) + args) + base_mty_o + in + let mod_entry = + {mod_entry_type = mty_o; + mod_entry_expr = mexpr_o} + in + Declaremods.declare_module id mod_entry; + if_verbose message + ("Module "^ string_of_id id ^" is defined") + + | true, _, Some _ -> + error "Module definition not allowed in a Module Type" + + +let vernac_end_module id = + Declaremods.end_module id; + if_verbose message + ("Module "^ string_of_id id ^" is defined") + + + +let vernac_declare_module_type id bl mty_ast_o = + let evd = Evd.empty in + let env = Global.env () in + let arglist,base_mty_o,_ = + Astmod.interp_module_decl evd env bl mty_ast_o None + in + let argids, args = List.split arglist + in (* We check the state of the system (in module, in module type) + and what parts are supplied *) + match Lib.is_specification (), base_mty_o with + | _, None -> + Declaremods.start_modtype id argids args; + if_verbose message + ("Interactive Module "^ string_of_id id ^" started") + + | _, Some base_mty -> + let mty = + List.fold_right + (fun (arg_id,arg_t) mte -> MTEfunsig(arg_id,arg_t,mte)) + args + base_mty + in + Declaremods.declare_modtype id mty; + if_verbose message + ("Module Type "^ string_of_id id ^" is defined") + + +let vernac_end_modtype id = + Declaremods.end_modtype id; + if_verbose message + ("Module Type "^ string_of_id id ^" is defined") + (**********************) (* Gallina extensions *) @@ -350,9 +461,20 @@ let vernac_record struc binders sort nameopt cfs = let vernac_begin_section id = let _ = Lib.open_section id in () let vernac_end_section id = - check_no_pending_proofs (); Discharge.close_section (is_verbose ()) id + +let vernac_end_segment id = + check_no_pending_proofs (); + try + match Lib.what_is_opened () with + | _,Lib.OpenedModule _ -> vernac_end_module id + | _,Lib.OpenedModtype _ -> vernac_end_modtype id + | _,Lib.OpenedSection _ -> vernac_end_section id + | _ -> anomaly "No more opened things" + with + Not_found -> error "There is nothing to end." + let is_obsolete_module (_,qid) = match repr_qualid qid with | dir, id when dir = empty_dirpath -> @@ -367,8 +489,8 @@ let is_obsolete_module (_,qid) = let vernac_require import _ qidl = try match import with - | None -> List.iter Library.read_module qidl - | Some b -> Library.require_module None qidl b + | None -> List.iter Library.read_library qidl + | Some b -> Library.require_library None qidl b with e -> (* Compatibility message *) let qidl' = List.filter is_obsolete_module qidl in @@ -381,7 +503,20 @@ let vernac_require import _ qidl = raise e let vernac_import export qidl = - List.iter (Library.import_module export) qidl + if export then + List.iter Library.export_library qidl + else + let import (loc,qid) = + try + let mp = Nametab.locate_module qid in + Declaremods.import_module mp + with Not_found -> + user_err_loc + (loc,"vernac_import", + str ((string_of_qualid qid)^" is not a module")) + in + List.iter import qidl; + Lib.add_frozen_state () let vernac_canonical locqid = Recordobj.objdef_declare (Nametab.global locqid) @@ -396,7 +531,7 @@ let vernac_coercion stre (loc,qid as locqid) qids qidt = let source = cl_of_qualid qids in let ref = Nametab.global locqid in Class.try_add_new_coercion_with_target ref stre source target; - if_verbose message ((Nametab.string_of_qualid qid) ^ " is now a coercion") + if_verbose message ((string_of_qualid qid) ^ " is now a coercion") let vernac_identity_coercion stre id qids qidt = let target = cl_of_qualid qidt in @@ -431,7 +566,7 @@ let vernac_solve_existential = instantiate_nth_evar_com (* Auxiliary file management *) let vernac_require_from export spec id filename = - Library.require_module_from_file spec (Some id) filename export + Library.require_library_from_file spec (Some id) filename export let vernac_add_loadpath isrec pdir ldiropt = let alias = match ldiropt with @@ -595,9 +730,9 @@ let vernac_mem_option key lv = let vernac_print_option key = try (get_ref_table key)#print - with not_found -> + with Not_found -> try (get_string_table key)#print - with not_found -> + with Not_found -> try print_option_value key with Not_found -> error_undeclared_key key @@ -655,7 +790,6 @@ let vernac_print = function | PrintHintDbName s -> Auto.print_hint_db_by_name s | PrintHintDb -> Auto.print_searchtable () - let global_loaded_library (loc, qid) = try Nametab.locate_loaded_library qid with Not_found -> @@ -692,7 +826,7 @@ let vernac_goal = function | None -> () | Some c -> if not (refining()) then begin - start_proof_com None (false,Nametab.NeverDischarge) c (fun _ _ ->()); + start_proof_com None (false,Libnames.NeverDischarge) c (fun _ _ ->()); print_subgoals () end else error "repeated Goal not permitted in refining mode" @@ -844,7 +978,7 @@ let _ = let ref = Nametab.global (dummy_loc, qid) in Class.try_add_new_class ref stre; if_verbose message - ((Nametab.string_of_qualid qid) ^ " is now a class") + ((string_of_qualid qid) ^ " is now a class") | _ -> bad_vernac_args "CLASS") (* Meta-syntax commands *) @@ -893,10 +1027,18 @@ let interp c = match c with | VernacCoFixpoint l -> vernac_cofixpoint l | VernacScheme l -> vernac_scheme l + (* Modules *) + | VernacDeclareModule (id,bl,mtyo,mexpro) -> + vernac_declare_module id bl mtyo mexpro + | VernacDeclareModuleType (id,bl,mtyo) -> + vernac_declare_module_type id bl mtyo + (* Gallina extensions *) - | VernacRecord (id,bl,s,idopt,fs) -> vernac_record id bl s idopt fs | VernacBeginSection id -> vernac_begin_section id - | VernacEndSection id -> vernac_end_section id + + | VernacEndSegment id -> vernac_end_segment id + + | VernacRecord (id,bl,s,idopt,fs) -> vernac_record id bl s idopt fs | VernacRequire (export,spec,qidl) -> vernac_require export spec qidl | VernacImport (export,qidl) -> vernac_import export qidl | VernacCanonical qid -> vernac_canonical qid diff --git a/toplevel/vernacentries.mli b/toplevel/vernacentries.mli index 85daf4375..0eca1143f 100644 --- a/toplevel/vernacentries.mli +++ b/toplevel/vernacentries.mli @@ -52,7 +52,7 @@ type pcoq_hook = { solve : int -> unit; abort : string -> unit; search : searchable -> dir_path list * bool -> unit; - print_name : Nametab.qualid Util.located -> unit; + print_name : Libnames.qualid Util.located -> unit; print_check : Environ.unsafe_judgment -> unit; print_eval : (constr -> constr) -> Environ.env -> Coqast.t -> Environ.unsafe_judgment -> unit; show_goal : int option -> unit diff --git a/toplevel/vernacexpr.ml b/toplevel/vernacexpr.ml index ace11aee3..869760411 100644 --- a/toplevel/vernacexpr.ml +++ b/toplevel/vernacexpr.ml @@ -25,6 +25,7 @@ type def_kind = DEFINITION | LET | LOCAL | THEOREM | LETTOP | DECL | REMARK | COERCION | LCOERCION | OBJECT | LOBJECT | OBJCOERCION | LOBJCOERCION | SUBCLASS | LSUBCLASS +open Libnames open Nametab type class_rawexpr = FunClass | SortClass | RefClass of qualid located @@ -165,6 +166,9 @@ type grammar_entry_ast = (loc * string) * Ast.entry_type option * grammar_associativity * raw_grammar_rule list +type module_ast = Coqast.t +type module_binder = identifier list * module_ast + type vernac_expr = (* Control *) | VernacList of located_vernac_expr list @@ -197,7 +201,7 @@ type vernac_expr = | VernacRecord of identifier with_coercion * simple_binder list * sort_expr * identifier option * local_decl_expr with_coercion list | VernacBeginSection of identifier - | VernacEndSection of identifier + | VernacEndSegment of identifier | VernacRequire of export_flag option * specif_flag option * qualid located list | VernacImport of export_flag * qualid located list @@ -206,6 +210,12 @@ type vernac_expr = | VernacIdentityCoercion of strength * identifier * class_rawexpr * class_rawexpr + (* Modules and Module Types *) + | VernacDeclareModule of identifier * + module_binder list * module_ast option * module_ast option + | VernacDeclareModuleType of identifier * + module_binder list * module_ast option + (* Solving *) | VernacSolve of int * raw_tactic_expr | VernacSolveExistential of int * constr_ast diff --git a/toplevel/vernacinterp.ml b/toplevel/vernacinterp.ml index 8e4189e2e..da4c56c57 100644 --- a/toplevel/vernacinterp.ml +++ b/toplevel/vernacinterp.ml @@ -11,6 +11,7 @@ open Pp open Util open Names +open Libnames open Himsg open Proof_type open Tacinterp -- cgit v1.2.3