diff options
| author | Théo Zimmermann | 2019-05-07 13:41:04 +0200 |
|---|---|---|
| committer | Théo Zimmermann | 2019-05-07 13:41:04 +0200 |
| commit | 7602c2cb547fe6664f7a065d17717baf12b9da88 (patch) | |
| tree | 1b034b3646090d35a8d730cbec6a5cf1c91da804 | |
| parent | 8aa64e7c4661549fef63a1c9c2e4e5284db911d8 (diff) | |
| parent | 9779c0bf4945693bfd37b141e2c9f0fea200ba4d (diff) | |
Merge PR #10002: Integrate ltac2
Ack-by: JasonGross
Reviewed-by: gares
Reviewed-by: ppedrot
Reviewed-by: jfehrle
Ack-by: SkySkimmer
Reviewed-by: Zimmi48
Reviewed-by: ejgallego
73 files changed, 12642 insertions, 31 deletions
diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 3c24ec28c4..2bfb91f27f 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -169,9 +169,7 @@ before_script: - not-a-real-job script: - cd _install_ci - - find lib/coq/ -name '*.vo' -print0 > vofiles - - for regexp in 's/.vo//' 's:lib/coq/plugins:Coq:' 's:lib/coq/theories:Coq:' 's:/:.:g'; do sed -z -i "$regexp" vofiles; done - - xargs -0 --arg-file=vofiles bin/coqchk -silent -o -m -coqlib lib/coq/ + - find lib/coq/ -name '*.vo' -print0 | xargs -0 bin/coqchk -silent -o -m -coqlib lib/coq/ .ci-template: stage: test @@ -66,7 +66,7 @@ FIND_SKIP_DIRS:='(' \ ')' -prune -o define find - $(shell find . $(FIND_SKIP_DIRS) '(' -name $(1) ')' -print | sed 's|^\./||') + $(shell find . user-contrib/Ltac2 $(FIND_SKIP_DIRS) '(' -name $(1) ')' -print | sed 's|^\./||') endef define findindir diff --git a/Makefile.build b/Makefile.build index 2a071fd820..034c9ea03c 100644 --- a/Makefile.build +++ b/Makefile.build @@ -158,11 +158,14 @@ endif VDFILE := .vfiles MLDFILE := .mlfiles PLUGMLDFILE := plugins/.mlfiles +USERCONTRIBMLDFILE := user-contrib/.mlfiles MLLIBDFILE := .mllibfiles PLUGMLLIBDFILE := plugins/.mllibfiles +USERCONTRIBMLLIBDFILE := user-contrib/.mllibfiles DEPENDENCIES := \ - $(addsuffix .d, $(MLDFILE) $(MLLIBDFILE) $(PLUGMLDFILE) $(PLUGMLLIBDFILE) $(CFILES) $(VDFILE)) + $(addsuffix .d, $(MLDFILE) $(MLLIBDFILE) $(PLUGMLDFILE) $(PLUGMLLIBDFILE) \ + $(USERCONTRIBMLDFILE) $(USERCONTRIBMLLIBDFILE) $(CFILES) $(VDFILE)) -include $(DEPENDENCIES) @@ -209,12 +212,14 @@ BOOTCOQC=$(TIMER) $(COQC) -coqlib . -q $(COQOPTS) LOCALINCLUDES=$(addprefix -I ,$(SRCDIRS)) MLINCLUDES=$(LOCALINCLUDES) +USERCONTRIBINCLUDES=$(addprefix -I user-contrib/,$(USERCONTRIBDIRS)) + OCAMLC := $(OCAMLFIND) ocamlc $(CAMLFLAGS) OCAMLOPT := $(OCAMLFIND) opt $(CAMLFLAGS) BYTEFLAGS=$(CAMLDEBUG) $(USERFLAGS) OPTFLAGS=$(CAMLDEBUGOPT) $(CAMLTIMEPROF) $(USERFLAGS) $(FLAMBDA_FLAGS) -DEPFLAGS=$(LOCALINCLUDES) -map gramlib/.pack/gramlib.ml $(if $(filter plugins/%,$@),, -I ide -I ide/protocol) +DEPFLAGS=$(LOCALINCLUDES) -map gramlib/.pack/gramlib.ml $(if $(filter plugins/% user-contrib/%,$@),, -I ide -I ide/protocol) # On MacOS, the binaries are signed, except our private ones ifeq ($(shell which codesign > /dev/null 2>&1 && echo $(ARCH)),Darwin) @@ -442,11 +447,11 @@ tools/coqdep_boot.cmx : tools/coqdep_common.cmx $(COQDEPBOOT): $(call bestobj, $(COQDEPBOOTSRC)) $(SHOW)'OCAMLBEST -o $@' - $(HIDE)$(call bestocaml, -I tools -package unix) + $(HIDE)$(call bestocaml, -I tools -package unix -package str) $(COQDEPBOOTBYTE): $(COQDEPBOOTSRC) $(SHOW)'OCAMLC -o $@' - $(HIDE)$(call ocamlbyte, -I tools -package unix) + $(HIDE)$(call ocamlbyte, -I tools -package unix -package str) $(OCAMLLIBDEP): $(call bestobj, tools/ocamllibdep.cmo) $(SHOW)'OCAMLBEST -o $@' @@ -567,7 +572,7 @@ VALIDOPTS=$(if $(VERBOSE),,-silent) -o -m -coqlib . validate: $(CHICKEN) | $(ALLVO:.$(VO)=.vo) $(SHOW)'COQCHK <theories & plugins>' - $(HIDE)$(CHICKEN) $(VALIDOPTS) $(ALLMODS) + $(HIDE)$(CHICKEN) $(VALIDOPTS) $(ALLVO) $(ALLSTDLIB).v: $(SHOW)'MAKE $(notdir $@)' @@ -743,6 +748,10 @@ plugins/%.cmx: plugins/%.ml $(SHOW)'OCAMLOPT $<' $(HIDE)$(OCAMLOPT) $(COND_OPTFLAGS) $(HACKMLI) $($(@:.cmx=_FORPACK)) -c $< +user-contrib/%.cmx: user-contrib/%.ml + $(SHOW)'OCAMLOPT $<' + $(HIDE)$(OCAMLOPT) $(COND_OPTFLAGS) $(HACKMLI) $($(@:.cmx=_FORPACK)) -c $< + kernel/%.cmx: COND_OPTFLAGS+=-w +a-4-44-50 %.cmx: %.ml @@ -776,8 +785,8 @@ kernel/%.cmx: COND_OPTFLAGS+=-w +a-4-44-50 # Ocamldep is now used directly again (thanks to -ml-synonym in OCaml >= 3.12) OCAMLDEP = $(OCAMLFIND) ocamldep -slash -ml-synonym .mlpack -MAINMLFILES := $(filter-out gramlib/.pack/% checker/% plugins/%, $(MLFILES) $(MLIFILES)) -MAINMLLIBFILES := $(filter-out gramlib/.pack/% checker/% plugins/%, $(MLLIBFILES) $(MLPACKFILES)) +MAINMLFILES := $(filter-out gramlib/.pack/% checker/% plugins/% user-contrib/%, $(MLFILES) $(MLIFILES)) +MAINMLLIBFILES := $(filter-out gramlib/.pack/% checker/% plugins/% user-contrib/%, $(MLLIBFILES) $(MLPACKFILES)) $(MLDFILE).d: $(D_DEPEND_BEFORE_SRC) $(MAINMLFILES) $(D_DEPEND_AFTER_SRC) $(GENFILES) $(GENGRAMFILES) $(SHOW)'OCAMLDEP MLFILES MLIFILES' @@ -796,6 +805,14 @@ $(PLUGMLLIBDFILE).d: $(D_DEPEND_BEFORE_SRC) $(filter plugins/%, $(MLLIBFILES) $( $(SHOW)'OCAMLLIBDEP plugins/MLLIBFILES plugins/MLPACKFILES' $(HIDE)$(OCAMLLIBDEP) $(DEPFLAGS) $(filter plugins/%, $(MLLIBFILES) $(MLPACKFILES)) $(TOTARGET) +$(USERCONTRIBMLDFILE).d: $(D_DEPEND_BEFORE_SRC) $(filter user-contrib/%, $(MLFILES) $(MLIFILES)) $(D_DEPEND_AFTER_SRC) $(GENFILES) + $(SHOW)'OCAMLDEP user-contrib/MLFILES user-contrib/MLIFILES' + $(HIDE)$(OCAMLDEP) $(DEPFLAGS) $(filter user-contrib/%, $(MLFILES) $(MLIFILES)) $(TOTARGET) + +$(USERCONTRIBMLLIBDFILE).d: $(D_DEPEND_BEFORE_SRC) $(filter user-contrib/%, $(MLLIBFILES) $(MLPACKFILES)) $(D_DEPEND_AFTER_SRC) $(OCAMLLIBDEP) $(GENFILES) + $(SHOW)'OCAMLLIBDEP user-contrib/MLLIBFILES user-contrib/MLPACKFILES' + $(HIDE)$(OCAMLLIBDEP) $(DEPFLAGS) $(filter user-contrib/%, $(MLLIBFILES) $(MLPACKFILES)) $(TOTARGET) + ########################################################################### # Compilation of .v files ########################################################################### @@ -861,7 +878,7 @@ endif $(VDFILE).d: $(D_DEPEND_BEFORE_SRC) $(VFILES) $(D_DEPEND_AFTER_SRC) $(COQDEPBOOT) $(SHOW)'COQDEP VFILES' - $(HIDE)$(COQDEPBOOT) -boot $(DYNDEP) $(VFILES) $(TOTARGET) + $(HIDE)$(COQDEPBOOT) -boot $(DYNDEP) -Q user-contrib "" $(USERCONTRIBINCLUDES) $(VFILES) $(TOTARGET) ########################################################################### diff --git a/Makefile.common b/Makefile.common index bd0e19cd00..ee3bfb43c5 100644 --- a/Makefile.common +++ b/Makefile.common @@ -104,10 +104,14 @@ PLUGINDIRS:=\ rtauto nsatz syntax btauto \ ssrmatching ltac ssr +USERCONTRIBDIRS:=\ + Ltac2 + SRCDIRS:=\ $(CORESRCDIRS) \ tools tools/coqdoc \ - $(addprefix plugins/, $(PLUGINDIRS)) + $(addprefix plugins/, $(PLUGINDIRS)) \ + $(addprefix user-contrib/, $(USERCONTRIBDIRS)) COQRUN := coqrun LIBCOQRUN:=kernel/byterun/lib$(COQRUN).a @@ -149,13 +153,14 @@ DERIVECMO:=plugins/derive/derive_plugin.cmo LTACCMO:=plugins/ltac/ltac_plugin.cmo plugins/ltac/tauto_plugin.cmo SSRMATCHINGCMO:=plugins/ssrmatching/ssrmatching_plugin.cmo SSRCMO:=plugins/ssr/ssreflect_plugin.cmo +LTAC2CMO:=user-contrib/Ltac2/ltac2_plugin.cmo PLUGINSCMO:=$(LTACCMO) $(OMEGACMO) $(MICROMEGACMO) \ $(RINGCMO) \ $(EXTRACTIONCMO) \ $(CCCMO) $(FOCMO) $(RTAUTOCMO) $(BTAUTOCMO) \ $(FUNINDCMO) $(NSATZCMO) $(SYNTAXCMO) \ - $(DERIVECMO) $(SSRMATCHINGCMO) $(SSRCMO) + $(DERIVECMO) $(SSRMATCHINGCMO) $(SSRCMO) $(LTAC2CMO) ifeq ($(HASNATDYNLINK)-$(BEST),false-opt) STATICPLUGINS:=$(PLUGINSCMO) diff --git a/Makefile.vofiles b/Makefile.vofiles index a71d68e565..e05822c889 100644 --- a/Makefile.vofiles +++ b/Makefile.vofiles @@ -13,7 +13,7 @@ endif ########################################################################### THEORIESVO := $(patsubst %.v,%.$(VO),$(shell find theories -type f -name "*.v")) -PLUGINSVO := $(patsubst %.v,%.$(VO),$(shell find plugins -type f -name "*.v")) +PLUGINSVO := $(patsubst %.v,%.$(VO),$(shell find plugins user-contrib -type f -name "*.v")) ALLVO := $(THEORIESVO) $(PLUGINSVO) VFILES := $(ALLVO:.$(VO)=.v) @@ -24,16 +24,16 @@ THEORIESLIGHTVO:= \ # convert a (stdlib) filename into a module name: # remove .vo, replace theories and plugins by Coq, and replace slashes by dots -vo_to_mod = $(subst /,.,$(patsubst theories/%,Coq.%,$(patsubst plugins/%,Coq.%,$(1:.vo=)))) +vo_to_mod = $(subst /,.,$(patsubst user-contrib/%,%,$(patsubst theories/%,Coq.%,$(patsubst plugins/%,Coq.%,$(1:.vo=))))) ALLMODS:=$(call vo_to_mod,$(ALLVO:.$(VO)=.vo)) # Converting a stdlib filename into native compiler filenames # Used for install targets -vo_to_cm = $(foreach vo,$(1),$(dir $(vo)).coq-native/$(subst /,_,$(patsubst theories/%,NCoq_%,$(patsubst plugins/%,NCoq_%,$(vo:.$(VO)=.cm*))))) +vo_to_cm = $(foreach vo,$(1),$(dir $(vo)).coq-native/$(subst /,_,$(patsubst user-contrib/%, N%, $(patsubst theories/%,NCoq_%,$(patsubst plugins/%,NCoq_%,$(vo:.$(VO)=.cm*)))))) -vo_to_obj = $(foreach vo,$(1),$(dir $(vo)).coq-native/$(subst /,_,$(patsubst theories/%,NCoq_%,$(patsubst plugins/%,NCoq_%,$(vo:.$(VO)=.o))))) +vo_to_obj = $(foreach vo,$(1),$(dir $(vo)).coq-native/$(subst /,_,$(patsubst user-contrib/%, N%, $(patsubst theories/%,NCoq_%,$(patsubst plugins/%,NCoq_%,$(vo:.$(VO)=.o)))))) ifdef QUICK GLOBFILES:= diff --git a/doc/sphinx/biblio.bib b/doc/sphinx/biblio.bib index 0467852b19..85b02013d8 100644 --- a/doc/sphinx/biblio.bib +++ b/doc/sphinx/biblio.bib @@ -551,3 +551,20 @@ the Calculus of Inductive Constructions}}, biburl = {http://dblp.uni-trier.de/rec/bib/conf/cpp/BoespflugDG11}, bibsource = {dblp computer science bibliography, http://dblp.org} } + +@inproceedings{MilnerPrincipalTypeSchemes, + author = {Damas, Luis and Milner, Robin}, + title = {Principal Type-schemes for Functional Programs}, + booktitle = {Proceedings of the 9th ACM SIGPLAN-SIGACT Symposium on Principles of Programming Languages}, + series = {POPL '82}, + year = {1982}, + isbn = {0-89791-065-6}, + location = {Albuquerque, New Mexico}, + pages = {207--212}, + numpages = {6}, + url = {http://doi.acm.org/10.1145/582153.582176}, + doi = {10.1145/582153.582176}, + acmid = {582176}, + publisher = {ACM}, + address = {New York, NY, USA}, +} diff --git a/doc/sphinx/index.html.rst b/doc/sphinx/index.html.rst index a91c6a9c5f..0a20d1c47b 100644 --- a/doc/sphinx/index.html.rst +++ b/doc/sphinx/index.html.rst @@ -42,6 +42,7 @@ Contents proof-engine/proof-handling proof-engine/tactics proof-engine/ltac + proof-engine/ltac2 proof-engine/detailed-tactic-examples proof-engine/ssreflect-proof-language diff --git a/doc/sphinx/index.latex.rst b/doc/sphinx/index.latex.rst index 708820fff7..5562736997 100644 --- a/doc/sphinx/index.latex.rst +++ b/doc/sphinx/index.latex.rst @@ -41,6 +41,7 @@ The proof engine proof-engine/proof-handling proof-engine/tactics proof-engine/ltac + proof-engine/ltac2 proof-engine/detailed-tactic-examples proof-engine/ssreflect-proof-language diff --git a/doc/sphinx/proof-engine/ltac.rst b/doc/sphinx/proof-engine/ltac.rst index 0322b43694..d3562b52c5 100644 --- a/doc/sphinx/proof-engine/ltac.rst +++ b/doc/sphinx/proof-engine/ltac.rst @@ -1,7 +1,7 @@ .. _ltac: -The tactic language -=================== +Ltac +==== This chapter gives a compact documentation of |Ltac|, the tactic language available in |Coq|. We start by giving the syntax, and next, we present the diff --git a/doc/sphinx/proof-engine/ltac2.rst b/doc/sphinx/proof-engine/ltac2.rst new file mode 100644 index 0000000000..6e33862b39 --- /dev/null +++ b/doc/sphinx/proof-engine/ltac2.rst @@ -0,0 +1,992 @@ +.. _ltac2: + +.. coqtop:: none + + From Ltac2 Require Import Ltac2. + +Ltac2 +===== + +The Ltac tactic language is probably one of the ingredients of the success of +Coq, yet it is at the same time its Achilles' heel. Indeed, Ltac: + +- has often unclear semantics +- is very non-uniform due to organic growth +- lacks expressivity (data structures, combinators, types, ...) +- is slow +- is error-prone and fragile +- has an intricate implementation + +Following the need of users that start developing huge projects relying +critically on Ltac, we believe that we should offer a proper modern language +that features at least the following: + +- at least informal, predictable semantics +- a typing system +- standard programming facilities (i.e. datatypes) + +This new language, called Ltac2, is described in this chapter. It is still +experimental but we encourage nonetheless users to start testing it, +especially wherever an advanced tactic language is needed. The previous +implementation of Ltac, described in the previous chapter, will be referred to +as Ltac1. + +.. _ltac2_design: + +General design +-------------- + +There are various alternatives to Ltac1, such that Mtac or Rtac for instance. +While those alternatives can be quite distinct from Ltac1, we designed +Ltac2 to be closest as reasonably possible to Ltac1, while fixing the +aforementioned defects. + +In particular, Ltac2 is: + +- a member of the ML family of languages, i.e. + + * a call-by-value functional language + * with effects + * together with Hindley-Milner type system + +- a language featuring meta-programming facilities for the manipulation of + Coq-side terms +- a language featuring notation facilities to help writing palatable scripts + +We describe more in details each point in the remainder of this document. + +ML component +------------ + +Overview +~~~~~~~~ + +Ltac2 is a member of the ML family of languages, in the sense that it is an +effectful call-by-value functional language, with static typing à la +Hindley-Milner (see :cite:`MilnerPrincipalTypeSchemes`). It is commonly accepted +that ML constitutes a sweet spot in PL design, as it is relatively expressive +while not being either too lax (unlike dynamic typing) nor too strict +(unlike, say, dependent types). + +The main goal of Ltac2 is to serve as a meta-language for Coq. As such, it +naturally fits in the ML lineage, just as the historical ML was designed as +the tactic language for the LCF prover. It can also be seen as a general-purpose +language, by simply forgetting about the Coq-specific features. + +Sticking to a standard ML type system can be considered somewhat weak for a +meta-language designed to manipulate Coq terms. In particular, there is no +way to statically guarantee that a Coq term resulting from an Ltac2 +computation will be well-typed. This is actually a design choice, motivated +by retro-compatibility with Ltac1. Instead, well-typedness is deferred to +dynamic checks, allowing many primitive functions to fail whenever they are +provided with an ill-typed term. + +The language is naturally effectful as it manipulates the global state of the +proof engine. This allows to think of proof-modifying primitives as effects +in a straightforward way. Semantically, proof manipulation lives in a monad, +which allows to ensure that Ltac2 satisfies the same equations as a generic ML +with unspecified effects would do, e.g. function reduction is substitution +by a value. + +Type Syntax +~~~~~~~~~~~ + +At the level of terms, we simply elaborate on Ltac1 syntax, which is quite +close to e.g. the one of OCaml. Types follow the simply-typed syntax of OCaml. + +The non-terminal :production:`lident` designates identifiers starting with a +lowercase. + +.. productionlist:: coq + ltac2_type : ( `ltac2_type`, ... , `ltac2_type` ) `ltac2_typeconst` + : ( `ltac2_type` * ... * `ltac2_type` ) + : `ltac2_type` -> `ltac2_type` + : `ltac2_typevar` + ltac2_typeconst : ( `modpath` . )* `lident` + ltac2_typevar : '`lident` + ltac2_typeparams : ( `ltac2_typevar`, ... , `ltac2_typevar` ) + +The set of base types can be extended thanks to the usual ML type +declarations such as algebraic datatypes and records. + +Built-in types include: + +- ``int``, machine integers (size not specified, in practice inherited from OCaml) +- ``string``, mutable strings +- ``'a array``, mutable arrays +- ``exn``, exceptions +- ``constr``, kernel-side terms +- ``pattern``, term patterns +- ``ident``, well-formed identifiers + +Type declarations +~~~~~~~~~~~~~~~~~ + +One can define new types by the following commands. + +.. cmd:: Ltac2 Type @ltac2_typeparams @lident + :name: Ltac2 Type + + This command defines an abstract type. It has no use for the end user and + is dedicated to types representing data coming from the OCaml world. + +.. cmdv:: Ltac2 Type {? rec} @ltac2_typeparams @lident := @ltac2_typedef + + This command defines a type with a manifest. There are four possible + kinds of such definitions: alias, variant, record and open variant types. + + .. productionlist:: coq + ltac2_typedef : `ltac2_type` + : [ `ltac2_constructordef` | ... | `ltac2_constructordef` ] + : { `ltac2_fielddef` ; ... ; `ltac2_fielddef` } + : [ .. ] + ltac2_constructordef : `uident` [ ( `ltac2_type` , ... , `ltac2_type` ) ] + ltac2_fielddef : [ mutable ] `ident` : `ltac2_type` + + Aliases are just a name for a given type expression and are transparently + unfoldable to it. They cannot be recursive. The non-terminal + :production:`uident` designates identifiers starting with an uppercase. + + Variants are sum types defined by constructors and eliminated by + pattern-matching. They can be recursive, but the `rec` flag must be + explicitly set. Pattern-maching must be exhaustive. + + Records are product types with named fields and eliminated by projection. + Likewise they can be recursive if the `rec` flag is set. + + .. cmdv:: Ltac2 Type @ltac2_typeparams @ltac2_qualid := [ @ltac2_constructordef ] + + Open variants are a special kind of variant types whose constructors are not + statically defined, but can instead be extended dynamically. A typical example + is the standard `exn` type. Pattern-matching must always include a catch-all + clause. They can be extended by this command. + +Term Syntax +~~~~~~~~~~~ + +The syntax of the functional fragment is very close to the one of Ltac1, except +that it adds a true pattern-matching feature, as well as a few standard +constructions from ML. + +.. productionlist:: coq + ltac2_var : `lident` + ltac2_qualid : ( `modpath` . )* `lident` + ltac2_constructor: `uident` + ltac2_term : `ltac2_qualid` + : `ltac2_constructor` + : `ltac2_term` `ltac2_term` ... `ltac2_term` + : fun `ltac2_var` => `ltac2_term` + : let `ltac2_var` := `ltac2_term` in `ltac2_term` + : let rec `ltac2_var` := `ltac2_term` in `ltac2_term` + : match `ltac2_term` with `ltac2_branch` ... `ltac2_branch` end + : `int` + : `string` + : `ltac2_term` ; `ltac2_term` + : [| `ltac2_term` ; ... ; `ltac2_term` |] + : ( `ltac2_term` , ... , `ltac2_term` ) + : { `ltac2_field` `ltac2_field` ... `ltac2_field` } + : `ltac2_term` . ( `ltac2_qualid` ) + : `ltac2_term` . ( `ltac2_qualid` ) := `ltac2_term` + : [; `ltac2_term` ; ... ; `ltac2_term` ] + : `ltac2_term` :: `ltac2_term` + : ... + ltac2_branch : `ltac2_pattern` => `ltac2_term` + ltac2_pattern : `ltac2_var` + : _ + : ( `ltac2_pattern` , ... , `ltac2_pattern` ) + : `ltac2_constructor` `ltac2_pattern` ... `ltac2_pattern` + : [ ] + : `ltac2_pattern` :: `ltac2_pattern` + ltac2_field : `ltac2_qualid` := `ltac2_term` + +In practice, there is some additional syntactic sugar that allows e.g. to +bind a variable and match on it at the same time, in the usual ML style. + +There is a dedicated syntax for list and array literals. + +.. note:: + + For now, deep pattern matching is not implemented. + +Ltac Definitions +~~~~~~~~~~~~~~~~ + +.. cmd:: Ltac2 {? mutable} {? rec} @lident := @ltac2_term + :name: Ltac2 + + This command defines a new global Ltac2 value. + + For semantic reasons, the body of the Ltac2 definition must be a syntactical + value, i.e. a function, a constant or a pure constructor recursively applied to + values. + + If ``rec`` is set, the tactic is expanded into a recursive binding. + + If ``mutable`` is set, the definition can be redefined at a later stage (see below). + +.. cmd:: Ltac2 Set @qualid := @ltac2_term + :name: Ltac2 Set + + This command redefines a previous ``mutable`` definition. + Mutable definitions act like dynamic binding, i.e. at runtime, the last defined + value for this entry is chosen. This is useful for global flags and the like. + +Reduction +~~~~~~~~~ + +We use the usual ML call-by-value reduction, with an otherwise unspecified +evaluation order. This is a design choice making it compatible with OCaml, +if ever we implement native compilation. The expected equations are as follows:: + + (fun x => t) V ≡ t{x := V} (βv) + + let x := V in t ≡ t{x := V} (let) + + match C V₀ ... Vₙ with ... | C x₀ ... xₙ => t | ... end ≡ t {xᵢ := Vᵢ} (ι) + + (t any term, V values, C constructor) + +Note that call-by-value reduction is already a departure from Ltac1 which uses +heuristics to decide when evaluating an expression. For instance, the following +expressions do not evaluate the same way in Ltac1. + +:n:`foo (idtac; let x := 0 in bar)` + +:n:`foo (let x := 0 in bar)` + +Instead of relying on the :n:`idtac` idiom, we would now require an explicit thunk +not to compute the argument, and :n:`foo` would have e.g. type +:n:`(unit -> unit) -> unit`. + +:n:`foo (fun () => let x := 0 in bar)` + +Typing +~~~~~~ + +Typing is strict and follows Hindley-Milner system. Unlike Ltac1, there +are no type casts at runtime, and one has to resort to conversion +functions. See notations though to make things more palatable. + +In this setting, all usual argument-free tactics have type :n:`unit -> unit`, but +one can return as well a value of type :n:`t` thanks to terms of type :n:`unit -> t`, +or take additional arguments. + +Effects +~~~~~~~ + +Effects in Ltac2 are straightforward, except that instead of using the +standard IO monad as the ambient effectful world, Ltac2 is going to use the +tactic monad. + +Note that the order of evaluation of application is *not* specified and is +implementation-dependent, as in OCaml. + +We recall that the `Proofview.tactic` monad is essentially a IO monad together +with backtracking state representing the proof state. + +Intuitively a thunk of type :n:`unit -> 'a` can do the following: + +- It can perform non-backtracking IO like printing and setting mutable variables +- It can fail in a non-recoverable way +- It can use first-class backtrack. The proper way to figure that is that we + morally have the following isomorphism: + :n:`(unit -> 'a) ~ (unit -> exn + ('a * (exn -> 'a)))` + i.e. thunks can produce a lazy list of results where each + tail is waiting for a continuation exception. +- It can access a backtracking proof state, made out amongst other things of + the current evar assignation and the list of goals under focus. + +We describe more thoroughly the various effects existing in Ltac2 hereafter. + +Standard IO ++++++++++++ + +The Ltac2 language features non-backtracking IO, notably mutable data and +printing operations. + +Mutable fields of records can be modified using the set syntax. Likewise, +built-in types like `string` and `array` feature imperative assignment. See +modules `String` and `Array` respectively. + +A few printing primitives are provided in the `Message` module, allowing to +display information to the user. + +Fatal errors +++++++++++++ + +The Ltac2 language provides non-backtracking exceptions, also known as *panics*, +through the following primitive in module `Control`.:: + + val throw : exn -> 'a + +Unlike backtracking exceptions from the next section, this kind of error +is never caught by backtracking primitives, that is, throwing an exception +destroys the stack. This is materialized by the following equation, where `E` +is an evaluation context.:: + + E[throw e] ≡ throw e + + (e value) + +There is currently no way to catch such an exception and it is a design choice. +There might be at some future point a way to catch it in a brutal way, +destroying all backtrack and return values. + +Backtrack ++++++++++ + +In Ltac2, we have the following backtracking primitives, defined in the +`Control` module.:: + + Ltac2 Type 'a result := [ Val ('a) | Err (exn) ]. + + val zero : exn -> 'a + val plus : (unit -> 'a) -> (exn -> 'a) -> 'a + val case : (unit -> 'a) -> ('a * (exn -> 'a)) result + +If one sees thunks as lazy lists, then `zero` is the empty list and `plus` is +list concatenation, while `case` is pattern-matching. + +The backtracking is first-class, i.e. one can write +:n:`plus (fun () => "x") (fun _ => "y") : string` producing a backtracking string. + +These operations are expected to satisfy a few equations, most notably that they +form a monoid compatible with sequentialization.:: + + plus t zero ≡ t () + plus (fun () => zero e) f ≡ f e + plus (plus t f) g ≡ plus t (fun e => plus (f e) g) + + case (fun () => zero e) ≡ Err e + case (fun () => plus (fun () => t) f) ≡ Val (t,f) + + let x := zero e in u ≡ zero e + let x := plus t f in u ≡ plus (fun () => let x := t in u) (fun e => let x := f e in u) + + (t, u, f, g, e values) + +Goals ++++++ + +A goal is given by the data of its conclusion and hypotheses, i.e. it can be +represented as `[Γ ⊢ A]`. + +The tactic monad naturally operates over the whole proofview, which may +represent several goals, including none. Thus, there is no such thing as +*the current goal*. Goals are naturally ordered, though. + +It is natural to do the same in Ltac2, but we must provide a way to get access +to a given goal. This is the role of the `enter` primitive, that applies a +tactic to each currently focused goal in turn.:: + + val enter : (unit -> unit) -> unit + +It is guaranteed that when evaluating `enter f`, `f` is called with exactly one +goal under focus. Note that `f` may be called several times, or never, depending +on the number of goals under focus before the call to `enter`. + +Accessing the goal data is then implicit in the Ltac2 primitives, and may panic +if the invariants are not respected. The two essential functions for observing +goals are given below.:: + + val hyp : ident -> constr + val goal : unit -> constr + +The two above functions panic if there is not exactly one goal under focus. +In addition, `hyp` may also fail if there is no hypothesis with the +corresponding name. + +Meta-programming +---------------- + +Overview +~~~~~~~~ + +One of the major implementation issues of Ltac1 is the fact that it is +never clear whether an object refers to the object world or the meta-world. +This is an incredible source of slowness, as the interpretation must be +aware of bound variables and must use heuristics to decide whether a variable +is a proper one or referring to something in the Ltac context. + +Likewise, in Ltac1, constr parsing is implicit, so that ``foo 0`` is +not ``foo`` applied to the Ltac integer expression ``0`` (Ltac does have a +notion of integers, though it is not first-class), but rather the Coq term +:g:`Datatypes.O`. + +The implicit parsing is confusing to users and often gives unexpected results. +Ltac2 makes these explicit using quoting and unquoting notation, although there +are notations to do it in a short and elegant way so as not to be too cumbersome +to the user. + +Generic Syntax for Quotations +~~~~~~~~~~~~~~~~~~~~~~~~~~~~~ + +In general, quotations can be introduced in terms using the following syntax, where +:production:`quotentry` is some parsing entry. + +.. prodn:: + ltac2_term += @ident : ( @quotentry ) + +Built-in quotations ++++++++++++++++++++ + +The current implementation recognizes the following built-in quotations: + +- ``ident``, which parses identifiers (type ``Init.ident``). +- ``constr``, which parses Coq terms and produces an-evar free term at runtime + (type ``Init.constr``). +- ``open_constr``, which parses Coq terms and produces a term potentially with + holes at runtime (type ``Init.constr`` as well). +- ``pattern``, which parses Coq patterns and produces a pattern used for term + matching (type ``Init.pattern``). +- ``reference``, which parses either a :n:`@qualid` or :n:`& @ident`. Qualified names + are globalized at internalization into the corresponding global reference, + while ``&id`` is turned into ``Std.VarRef id``. This produces at runtime a + ``Std.reference``. + +The following syntactic sugar is provided for two common cases. + +- ``@id`` is the same as ``ident:(id)`` +- ``'t`` is the same as ``open_constr:(t)`` + +Strict vs. non-strict mode +++++++++++++++++++++++++++ + +Depending on the context, quotations producing terms (i.e. ``constr`` or +``open_constr``) are not internalized in the same way. There are two possible +modes, respectively called the *strict* and the *non-strict* mode. + +- In strict mode, all simple identifiers appearing in a term quotation are + required to be resolvable statically. That is, they must be the short name of + a declaration which is defined globally, excluding section variables and + hypotheses. If this doesn't hold, internalization will fail. To work around + this error, one has to specifically use the ``&`` notation. +- In non-strict mode, any simple identifier appearing in a term quotation which + is not bound in the global context is turned into a dynamic reference to a + hypothesis. That is to say, internalization will succeed, but the evaluation + of the term at runtime will fail if there is no such variable in the dynamic + context. + +Strict mode is enforced by default, e.g. for all Ltac2 definitions. Non-strict +mode is only set when evaluating Ltac2 snippets in interactive proof mode. The +rationale is that it is cumbersome to explicitly add ``&`` interactively, while it +is expected that global tactics enforce more invariants on their code. + +Term Antiquotations +~~~~~~~~~~~~~~~~~~~ + +Syntax +++++++ + +One can also insert Ltac2 code into Coq terms, similarly to what is possible in +Ltac1. + +.. prodn:: + term += ltac2:( @ltac2_term ) + +Antiquoted terms are expected to have type ``unit``, as they are only evaluated +for their side-effects. + +Semantics ++++++++++ + +Interpretation of a quoted Coq term is done in two phases, internalization and +evaluation. + +- Internalization is part of the static semantics, i.e. it is done at Ltac2 + typing time. +- Evaluation is part of the dynamic semantics, i.e. it is done when + a term gets effectively computed by Ltac2. + +Note that typing of Coq terms is a *dynamic* process occurring at Ltac2 +evaluation time, and not at Ltac2 typing time. + +Static semantics +**************** + +During internalization, Coq variables are resolved and antiquotations are +type-checked as Ltac2 terms, effectively producing a ``glob_constr`` in Coq +implementation terminology. Note that although it went through the +type-checking of **Ltac2**, the resulting term has not been fully computed and +is potentially ill-typed as a runtime **Coq** term. + +.. example:: + + The following term is valid (with type `unit -> constr`), but will fail at runtime: + + .. coqtop:: in + + Ltac2 myconstr () := constr:(nat -> 0). + +Term antiquotations are type-checked in the enclosing Ltac2 typing context +of the corresponding term expression. + +.. example:: + + The following will type-check, with type `constr`. + + .. coqdoc:: + + let x := '0 in constr:(1 + ltac2:(exact x)) + +Beware that the typing environment of antiquotations is **not** +expanded by the Coq binders from the term. + + .. example:: + + The following Ltac2 expression will **not** type-check:: + + `constr:(fun x : nat => ltac2:(exact x))` + `(* Error: Unbound variable 'x' *)` + +There is a simple reason for that, which is that the following expression would +not make sense in general. + +`constr:(fun x : nat => ltac2:(clear @x; exact x))` + +Indeed, a hypothesis can suddenly disappear from the runtime context if some +other tactic pulls the rug from under you. + +Rather, the tactic writer has to resort to the **dynamic** goal environment, +and must write instead explicitly that she is accessing a hypothesis, typically +as follows. + +`constr:(fun x : nat => ltac2:(exact (hyp @x)))` + +This pattern is so common that we provide dedicated Ltac2 and Coq term notations +for it. + +- `&x` as an Ltac2 expression expands to `hyp @x`. +- `&x` as a Coq constr expression expands to + `ltac2:(Control.refine (fun () => hyp @x))`. + +Dynamic semantics +***************** + +During evaluation, a quoted term is fully evaluated to a kernel term, and is +in particular type-checked in the current environment. + +Evaluation of a quoted term goes as follows. + +- The quoted term is first evaluated by the pretyper. +- Antiquotations are then evaluated in a context where there is exactly one goal + under focus, with the hypotheses coming from the current environment extended + with the bound variables of the term, and the resulting term is fed into the + quoted term. + +Relative orders of evaluation of antiquotations and quoted term are not +specified. + +For instance, in the following example, `tac` will be evaluated in a context +with exactly one goal under focus, whose last hypothesis is `H : nat`. The +whole expression will thus evaluate to the term :g:`fun H : nat => H`. + +`let tac () := hyp @H in constr:(fun H : nat => ltac2:(tac ()))` + +Many standard tactics perform type-checking of their argument before going +further. It is your duty to ensure that terms are well-typed when calling +such tactics. Failure to do so will result in non-recoverable exceptions. + +**Trivial Term Antiquotations** + +It is possible to refer to a variable of type `constr` in the Ltac2 environment +through a specific syntax consistent with the antiquotations presented in +the notation section. + +.. prodn:: term += $@lident + +In a Coq term, writing :g:`$x` is semantically equivalent to +:g:`ltac2:(Control.refine (fun () => x))`, up to re-typechecking. It allows to +insert in a concise way an Ltac2 variable of type :n:`constr` into a Coq term. + +Match over terms +~~~~~~~~~~~~~~~~ + +Ltac2 features a construction similar to Ltac1 :n:`match` over terms, although +in a less hard-wired way. + +.. productionlist:: coq + ltac2_term : match! `ltac2_term` with `constrmatching` .. `constrmatching` end + : lazy_match! `ltac2_term` with `constrmatching` .. `constrmatching` end + : multi_match! `ltac2_term` with `constrmatching` .. `constrmatching` end + constrmatching : | `constrpattern` => `ltac2_term` + constrpattern : `term` + : context [ `term` ] + : context `lident` [ `term` ] + +This construction is not primitive and is desugared at parsing time into +calls to term matching functions from the `Pattern` module. Internally, it is +implemented thanks to a specific scope accepting the :n:`@constrmatching` syntax. + +Variables from the :n:`@constrpattern` are statically bound in the body of the branch, to +values of type `constr` for the variables from the :n:`@constr` pattern and to a +value of type `Pattern.context` for the variable :n:`@lident`. + +Note that unlike Ltac, only lowercase identifiers are valid as Ltac2 +bindings, so that there will be a syntax error if one of the bound variables +starts with an uppercase character. + +The semantics of this construction is otherwise the same as the corresponding +one from Ltac1, except that it requires the goal to be focused. + +Match over goals +~~~~~~~~~~~~~~~~ + +Similarly, there is a way to match over goals in an elegant way, which is +just a notation desugared at parsing time. + +.. productionlist:: coq + ltac2_term : match! [ reverse ] goal with `goalmatching` ... `goalmatching` end + : lazy_match! [ reverse ] goal with `goalmatching` ... `goalmatching` end + : multi_match! [ reverse ] goal with `goalmatching` ... `goalmatching` end + goalmatching : | [ `hypmatching` ... `hypmatching` |- `constrpattern` ] => `ltac2_term` + hypmatching : `lident` : `constrpattern` + : _ : `constrpattern` + +Variables from :n:`@hypmatching` and :n:`@constrpattern` are bound in the body of the +branch. Their types are: + +- ``constr`` for pattern variables appearing in a :n:`@term` +- ``Pattern.context`` for variables binding a context +- ``ident`` for variables binding a hypothesis name. + +The same identifier caveat as in the case of matching over constr applies, and +this features has the same semantics as in Ltac1. In particular, a ``reverse`` +flag can be specified to match hypotheses from the more recently introduced to +the least recently introduced one. + +Notations +--------- + +Notations are the crux of the usability of Ltac1. We should be able to recover +a feeling similar to the old implementation by using and abusing notations. + +Scopes +~~~~~~ + +A scope is a name given to a grammar entry used to produce some Ltac2 expression +at parsing time. Scopes are described using a form of S-expression. + +.. prodn:: + ltac2_scope ::= @string %| @integer %| @lident ({+, @ltac2_scope}) + +A few scopes contain antiquotation features. For sake of uniformity, all +antiquotations are introduced by the syntax :n:`$@lident`. + +The following scopes are built-in. + +- :n:`constr`: + + + parses :n:`c = @term` and produces :n:`constr:(c)` + +- :n:`ident`: + + + parses :n:`id = @ident` and produces :n:`ident:(id)` + + parses :n:`$(x = @ident)` and produces the variable :n:`x` + +- :n:`list0(@ltac2_scope)`: + + + if :n:`@ltac2_scope` parses :production:`entry`, parses :n:`(@entry__0, ..., @entry__n)` and produces + :n:`[@entry__0; ...; @entry__n]`. + +- :n:`list0(@ltac2_scope, sep = @string__sep)`: + + + if :n:`@ltac2_scope` parses :n:`@entry`, parses :n:`(@entry__0 @string__sep ... @string__sep @entry__n)` + and produces :n:`[@entry__0; ...; @entry__n]`. + +- :n:`list1`: same as :n:`list0` (with or without separator) but parses :n:`{+ @entry}` instead + of :n:`{* @entry}`. + +- :n:`opt(@ltac2_scope)` + + + if :n:`@ltac2_scope` parses :n:`@entry`, parses :n:`{? @entry}` and produces either :n:`None` or + :n:`Some x` where :n:`x` is the parsed expression. + +- :n:`self`: + + + parses a Ltac2 expression at the current level and return it as is. + +- :n:`next`: + + + parses a Ltac2 expression at the next level and return it as is. + +- :n:`tactic(n = @integer)`: + + + parses a Ltac2 expression at the provided level :n:`n` and return it as is. + +- :n:`thunk(@ltac2_scope)`: + + + parses the same as :n:`scope`, and if :n:`e` is the parsed expression, returns + :n:`fun () => e`. + +- :n:`STRING`: + + + parses the corresponding string as an identifier and returns :n:`()`. + +- :n:`keyword(s = @string)`: + + + parses the string :n:`s` as a keyword and returns `()`. + +- :n:`terminal(s = @string)`: + + + parses the string :n:`s` as a keyword, if it is already a + keyword, otherwise as an :n:`@ident`. Returns `()`. + +- :n:`seq(@ltac2_scope__1, ..., @ltac2_scope__2)`: + + + parses :n:`scope__1`, ..., :n:`scope__n` in this order, and produces a tuple made + out of the parsed values in the same order. As an optimization, all + subscopes of the form :n:`STRING` are left out of the returned tuple, instead + of returning a useless unit value. It is forbidden for the various + subscopes to refer to the global entry using self or next. + +A few other specific scopes exist to handle Ltac1-like syntax, but their use is +discouraged and they are thus not documented. + +For now there is no way to declare new scopes from Ltac2 side, but this is +planned. + +Notations +~~~~~~~~~ + +The Ltac2 parser can be extended by syntactic notations. + +.. cmd:: Ltac2 Notation {+ @lident (@ltac2_scope) %| @string } {? : @integer} := @ltac2_term + :name: Ltac2 Notation + + A Ltac2 notation adds a parsing rule to the Ltac2 grammar, which is expanded + to the provided body where every token from the notation is let-bound to the + corresponding generated expression. + + .. example:: + + Assume we perform: + + .. coqdoc:: + + Ltac2 Notation "foo" c(thunk(constr)) ids(list0(ident)) := Bar.f c ids. + + Then the following expression + + `let y := @X in foo (nat -> nat) x $y` + + will expand at parsing time to + + `let y := @X in` + `let c := fun () => constr:(nat -> nat) with ids := [@x; y] in Bar.f c ids` + + Beware that the order of evaluation of multiple let-bindings is not specified, + so that you may have to resort to thunking to ensure that side-effects are + performed at the right time. + +Abbreviations +~~~~~~~~~~~~~ + +.. cmdv:: Ltac2 Notation @lident := @ltac2_term + + This command introduces a special kind of notations, called abbreviations, + that is designed so that it does not add any parsing rules. It is similar in + spirit to Coq abbreviations, insofar as its main purpose is to give an + absolute name to a piece of pure syntax, which can be transparently referred + by this name as if it were a proper definition. + + The abbreviation can then be manipulated just as a normal Ltac2 definition, + except that it is expanded at internalization time into the given expression. + Furthermore, in order to make this kind of construction useful in practice in + an effectful language such as Ltac2, any syntactic argument to an abbreviation + is thunked on-the-fly during its expansion. + +For instance, suppose that we define the following. + +:n:`Ltac2 Notation foo := fun x => x ().` + +Then we have the following expansion at internalization time. + +:n:`foo 0 ↦ (fun x => x ()) (fun _ => 0)` + +Note that abbreviations are not typechecked at all, and may result in typing +errors after expansion. + +Evaluation +---------- + +Ltac2 features a toplevel loop that can be used to evaluate expressions. + +.. cmd:: Ltac2 Eval @ltac2_term + :name: Ltac2 Eval + + This command evaluates the term in the current proof if there is one, or in the + global environment otherwise, and displays the resulting value to the user + together with its type. This command is pure in the sense that it does not + modify the state of the proof, and in particular all side-effects are discarded. + +Debug +----- + +.. opt:: Ltac2 Backtrace + + When this option is set, toplevel failures will be printed with a backtrace. + +Compatibility layer with Ltac1 +------------------------------ + +Ltac1 from Ltac2 +~~~~~~~~~~~~~~~~ + +Simple API +++++++++++ + +One can call Ltac1 code from Ltac2 by using the :n:`ltac1` quotation. It parses +a Ltac1 expression, and semantics of this quotation is the evaluation of the +corresponding code for its side effects. In particular, it cannot return values, +and the quotation has type :n:`unit`. + +Beware, Ltac1 **cannot** access variables from the Ltac2 scope. One is limited +to the use of standalone function calls. + +Low-level API ++++++++++++++ + +There exists a lower-level FFI into Ltac1 that is not recommended for daily use, +which is available in the `Ltac2.Ltac1` module. This API allows to directly +manipulate dynamically-typed Ltac1 values, either through the function calls, +or using the `ltac1val` quotation. The latter parses the same as `ltac1`, but +has type `Ltac2.Ltac1.t` instead of `unit`, and dynamically behaves as an Ltac1 +thunk, i.e. `ltac1val:(foo)` corresponds to the tactic closure that Ltac1 +would generate from `idtac; foo`. + +Due to intricate dynamic semantics, understanding when Ltac1 value quotations +focus is very hard. This is why some functions return a continuation-passing +style value, as it can dispatch dynamically between focused and unfocused +behaviour. + +Ltac2 from Ltac1 +~~~~~~~~~~~~~~~~ + +Same as above by switching Ltac1 by Ltac2 and using the `ltac2` quotation +instead. + +Note that the tactic expression is evaluated eagerly, if one wants to use it as +an argument to a Ltac1 function, she has to resort to the good old +:n:`idtac; ltac2:(foo)` trick. For instance, the code below will fail immediately +and won't print anything. + +.. coqtop:: in + + From Ltac2 Require Import Ltac2. + Set Default Proof Mode "Classic". + +.. coqtop:: all + + Ltac mytac tac := idtac "wow"; tac. + + Goal True. + Proof. + Fail mytac ltac2:(fail). + +Transition from Ltac1 +--------------------- + +Owing to the use of a lot of notations, the transition should not be too +difficult. In particular, it should be possible to do it incrementally. That +said, we do *not* guarantee you it is going to be a blissful walk either. +Hopefully, owing to the fact Ltac2 is typed, the interactive dialogue with Coq +will help you. + +We list the major changes and the transition strategies hereafter. + +Syntax changes +~~~~~~~~~~~~~~ + +Due to conflicts, a few syntactic rules have changed. + +- The dispatch tactical :n:`tac; [foo|bar]` is now written :n:`tac > [foo|bar]`. +- Levels of a few operators have been revised. Some tacticals now parse as if + they were a normal function, i.e. one has to put parentheses around the + argument when it is complex, e.g an abstraction. List of affected tacticals: + :n:`try`, :n:`repeat`, :n:`do`, :n:`once`, :n:`progress`, :n:`time`, :n:`abstract`. +- :n:`idtac` is no more. Either use :n:`()` if you expect nothing to happen, + :n:`(fun () => ())` if you want a thunk (see next section), or use printing + primitives from the :n:`Message` module if you want to display something. + +Tactic delay +~~~~~~~~~~~~ + +Tactics are not magically delayed anymore, neither as functions nor as +arguments. It is your responsibility to thunk them beforehand and apply them +at the call site. + +A typical example of a delayed function: + +:n:`Ltac foo := blah.` + +becomes + +:n:`Ltac2 foo () := blah.` + +All subsequent calls to `foo` must be applied to perform the same effect as +before. + +Likewise, for arguments: + +:n:`Ltac bar tac := tac; tac; tac.` + +becomes + +:n:`Ltac2 bar tac := tac (); tac (); tac ().` + +We recommend the use of syntactic notations to ease the transition. For +instance, the first example can alternatively be written as: + +:n:`Ltac2 foo0 () := blah.` +:n:`Ltac2 Notation foo := foo0 ().` + +This allows to keep the subsequent calls to the tactic as-is, as the +expression `foo` will be implicitly expanded everywhere into `foo0 ()`. Such +a trick also works for arguments, as arguments of syntactic notations are +implicitly thunked. The second example could thus be written as follows. + +:n:`Ltac2 bar0 tac := tac (); tac (); tac ().` +:n:`Ltac2 Notation bar := bar0.` + +Variable binding +~~~~~~~~~~~~~~~~ + +Ltac1 relies on complex dynamic trickery to be able to tell apart bound +variables from terms, hypotheses, etc. There is no such thing in Ltac2, +as variables are recognized statically and other constructions do not live in +the same syntactic world. Due to the abuse of quotations, it can sometimes be +complicated to know what a mere identifier represents in a tactic expression. We +recommend tracking the context and letting the compiler print typing errors to +understand what is going on. + +We list below the typical changes one has to perform depending on the static +errors produced by the typechecker. + +In Ltac expressions ++++++++++++++++++++ + +.. exn:: Unbound ( value | constructor ) X + + * if `X` is meant to be a term from the current stactic environment, replace + the problematic use by `'X`. + * if `X` is meant to be a hypothesis from the goal context, replace the + problematic use by `&X`. + +In quotations ++++++++++++++ + +.. exn:: The reference X was not found in the current environment + + * if `X` is meant to be a tactic expression bound by a Ltac2 let or function, + replace the problematic use by `$X`. + * if `X` is meant to be a hypothesis from the goal context, replace the + problematic use by `&X`. + +Exception catching +~~~~~~~~~~~~~~~~~~ + +Ltac2 features a proper exception-catching mechanism. For this reason, the +Ltac1 mechanism relying on `fail` taking integers, and tacticals decreasing it, +has been removed. Now exceptions are preserved by all tacticals, and it is +your duty to catch them and reraise them depending on your use. @@ -18,8 +18,9 @@ (targets .vfiles.d) (deps (source_tree theories) - (source_tree plugins)) - (action (with-stdout-to .vfiles.d (bash "%{bin:coqdep} -dyndep both -noglob -boot `find theories plugins -type f -name *.v`")))) + (source_tree plugins) + (source_tree user-contrib)) + (action (with-stdout-to .vfiles.d (bash "%{bin:coqdep} -dyndep both -noglob -boot `find theories plugins user-contrib -type f -name *.v`")))) (alias (name vodeps) diff --git a/test-suite/Makefile b/test-suite/Makefile index ba591ede20..94011447d7 100644 --- a/test-suite/Makefile +++ b/test-suite/Makefile @@ -99,7 +99,7 @@ INTERACTIVE := interactive UNIT_TESTS := unit-tests VSUBSYSTEMS := prerequisite success failure $(BUGS) output \ output-modulo-time $(INTERACTIVE) micromega $(COMPLEXITY) modules stm \ - coqdoc ssr arithmetic + coqdoc ssr arithmetic ltac2 # All subsystems SUBSYSTEMS := $(VSUBSYSTEMS) misc bugs ide vio coqchk coqwc coq-makefile tools $(UNIT_TESTS) @@ -181,6 +181,7 @@ summary: $(call summary_dir, "tools/ tests", tools); \ $(call summary_dir, "Unit tests", unit-tests); \ $(call summary_dir, "Machine arithmetic tests", arithmetic); \ + $(call summary_dir, "Ltac2 tests", ltac2); \ nb_success=`find . -name '*.log' -exec tail -n2 '{}' \; | grep -e $(log_success) | wc -l`; \ nb_failure=`find . -name '*.log' -exec tail -n2 '{}' \; | grep -e $(log_failure) | wc -l`; \ nb_tests=`expr $$nb_success + $$nb_failure`; \ @@ -319,7 +320,7 @@ $(addsuffix .log,$(wildcard prerequisite/*.v)): %.v.log: %.v } > "$@" ssr: $(wildcard ssr/*.v:%.v=%.v.log) -$(addsuffix .log,$(wildcard ssr/*.v success/*.v micromega/*.v modules/*.v arithmetic/*.v)): %.v.log: %.v $(PREREQUISITELOG) +$(addsuffix .log,$(wildcard ssr/*.v success/*.v micromega/*.v modules/*.v arithmetic/*.v ltac2/*.v)): %.v.log: %.v $(PREREQUISITELOG) @echo "TEST $< $(call get_coq_prog_args_in_parens,"$<")" $(HIDE){ \ opts="$(if $(findstring modules/,$<),-R modules Mods)"; \ diff --git a/test-suite/ltac2/compat.v b/test-suite/ltac2/compat.v new file mode 100644 index 0000000000..489fa638e4 --- /dev/null +++ b/test-suite/ltac2/compat.v @@ -0,0 +1,58 @@ +Require Import Ltac2.Ltac2. + +Import Ltac2.Notations. + +(** Test calls to Ltac1 from Ltac2 *) + +Ltac2 foo () := ltac1:(discriminate). + +Goal true = false -> False. +Proof. +foo (). +Qed. + +Goal true = false -> false = true. +Proof. +intros H; ltac1:(match goal with [ H : ?P |- _ ] => rewrite H end); reflexivity. +Qed. + +Goal true = false -> false = true. +Proof. +intros H; ltac1:(rewrite H); reflexivity. +Abort. + +(** Variables do not cross the compatibility layer boundary. *) +Fail Ltac2 bar nay := ltac1:(discriminate nay). + +Fail Ltac2 pose1 (v : constr) := + ltac1:(pose $v). + +(** Test calls to Ltac2 from Ltac1 *) + +Set Default Proof Mode "Classic". + +Ltac foo := ltac2:(foo ()). + +Goal true = false -> False. +Proof. +ltac2:(foo ()). +Qed. + +Goal true = false -> False. +Proof. +foo. +Qed. + +(** Variables do not cross the compatibility layer boundary. *) +Fail Ltac bar x := ltac2:(foo x). + +Ltac mytac tac := idtac "wow". + +Goal True. +Proof. +(** Fails because quotation is evaluated eagerly *) +Fail mytac ltac2:(fail). +(** One has to thunk thanks to the idtac trick *) +let t := idtac; ltac2:(fail) in mytac t. +constructor. +Qed. diff --git a/test-suite/ltac2/errors.v b/test-suite/ltac2/errors.v new file mode 100644 index 0000000000..c677f6af5d --- /dev/null +++ b/test-suite/ltac2/errors.v @@ -0,0 +1,12 @@ +Require Import Ltac2.Ltac2. + +Goal True. +Proof. +let x := Control.plus + (fun () => let _ := constr:(nat -> 0) in 0) + (fun e => match e with Not_found => 1 | _ => 2 end) in +match Int.equal x 2 with +| true => () +| false => Control.throw (Tactic_failure None) +end. +Abort. diff --git a/test-suite/ltac2/example1.v b/test-suite/ltac2/example1.v new file mode 100644 index 0000000000..023791050f --- /dev/null +++ b/test-suite/ltac2/example1.v @@ -0,0 +1,27 @@ +Require Import Ltac2.Ltac2. + +Import Ltac2.Control. + +(** Alternative implementation of the hyp primitive *) +Ltac2 get_hyp_by_name x := + let h := hyps () in + let rec find x l := match l with + | [] => zero Not_found + | p :: l => + match p with + | (id, _, t) => + match Ident.equal x id with + | true => t + | false => find x l + end + end + end in + find x h. + +Print Ltac2 get_hyp_by_name. + +Goal forall n m, n + m = 0 -> n = 0. +Proof. +refine (fun () => '(fun n m H => _)). +let t := get_hyp_by_name @H in Message.print (Message.of_constr t). +Abort. diff --git a/test-suite/ltac2/example2.v b/test-suite/ltac2/example2.v new file mode 100644 index 0000000000..c953d25061 --- /dev/null +++ b/test-suite/ltac2/example2.v @@ -0,0 +1,281 @@ +Require Import Ltac2.Ltac2. + +Import Ltac2.Notations. + +Set Default Goal Selector "all". + +Goal exists n, n = 0. +Proof. +split with (x := 0). +reflexivity. +Qed. + +Goal exists n, n = 0. +Proof. +split with 0. +split. +Qed. + +Goal exists n, n = 0. +Proof. +let myvar := Std.NamedHyp @x in split with ($myvar := 0). +split. +Qed. + +Goal (forall n : nat, n = 0 -> False) -> True. +Proof. +intros H. +eelim &H. +split. +Qed. + +Goal (forall n : nat, n = 0 -> False) -> True. +Proof. +intros H. +elim &H with 0. +split. +Qed. + +Goal forall (P : nat -> Prop), (forall n m, n = m -> P n) -> P 0. +Proof. +intros P H. +Fail apply &H. +apply &H with (m := 0). +split. +Qed. + +Goal forall (P : nat -> Prop), (forall n m, n = m -> P n) -> (0 = 1) -> P 0. +Proof. +intros P H e. +apply &H with (m := 1) in e. +exact e. +Qed. + +Goal forall (P : nat -> Prop), (forall n m, n = m -> P n) -> P 0. +Proof. +intros P H. +eapply &H. +split. +Qed. + +Goal exists n, n = 0. +Proof. +Fail constructor 1. +constructor 1 with (x := 0). +split. +Qed. + +Goal exists n, n = 0. +Proof. +econstructor 1. +split. +Qed. + +Goal forall n, 0 + n = n. +Proof. +intros n. +induction &n as [|n] using nat_rect; split. +Qed. + +Goal forall n, 0 + n = n. +Proof. +intros n. +let n := @X in +let q := Std.NamedHyp @P in +induction &n as [|$n] using nat_rect with ($q := fun m => 0 + m = m); split. +Qed. + +Goal forall n, 0 + n = n. +Proof. +intros n. +destruct &n as [|n] using nat_rect; split. +Qed. + +Goal forall n, 0 + n = n. +Proof. +intros n. +let n := @X in +let q := Std.NamedHyp @P in +destruct &n as [|$n] using nat_rect with ($q := fun m => 0 + m = m); split. +Qed. + +Goal forall b1 b2, andb b1 b2 = andb b2 b1. +Proof. +intros b1 b2. +destruct &b1 as [|], &b2 as [|]; split. +Qed. + +Goal forall n m, n = 0 -> n + m = m. +Proof. +intros n m Hn. +rewrite &Hn; split. +Qed. + +Goal forall n m p, n = m -> p = m -> 0 = n -> p = 0. +Proof. +intros n m p He He' Hn. +rewrite &He, <- &He' in Hn. +rewrite &Hn. +split. +Qed. + +Goal forall n m, (m = n -> n = m) -> m = n -> n = 0 -> m = 0. +Proof. +intros n m He He' He''. +rewrite <- &He by assumption. +Control.refine (fun () => &He''). +Qed. + +Goal forall n (r := if true then n else 0), r = n. +Proof. +intros n r. +hnf in r. +split. +Qed. + +Goal 1 = 0 -> 0 = 0. +Proof. +intros H. +pattern 0 at 1. +let occ := 2 in pattern 1 at 1, 0 at $occ in H. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +vm_compute. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +native_compute. +reflexivity. +Qed. + +Goal 1 + 1 = 2 - 0 -> True. +Proof. +intros H. +vm_compute plus in H. +reflexivity. +Qed. + +Goal 1 = 0 -> True /\ True. +Proof. +intros H. +split; fold (1 + 0) (1 + 0) in H. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +cbv [ Nat.add ]. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +let x := reference:(Nat.add) in +cbn beta iota delta [ $x ]. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +simpl beta. +reflexivity. +Qed. + +Goal 1 + 1 = 2. +Proof. +lazy. +reflexivity. +Qed. + +Goal let x := 1 + 1 - 1 in x = x. +Proof. +intros x. +unfold &x at 1. +let x := reference:(Nat.sub) in unfold Nat.add, $x in x. +reflexivity. +Qed. + +Goal exists x y : nat, x = y. +Proof. +exists 0, 0; reflexivity. +Qed. + +Goal exists x y : nat, x = y. +Proof. +eexists _, 0; reflexivity. +Qed. + +Goal exists x y : nat, x = y. +Proof. +refine '(let x := 0 in _). +eexists; exists &x; reflexivity. +Qed. + +Goal True. +Proof. +pose (X := True). +constructor. +Qed. + +Goal True. +Proof. +pose True as X. +constructor. +Qed. + +Goal True. +Proof. +let x := @foo in +set ($x := True) in * |-. +constructor. +Qed. + +Goal 0 = 0. +Proof. +remember 0 as n eqn: foo at 1. +rewrite foo. +reflexivity. +Qed. + +Goal True. +Proof. +assert (H := 0 + 0). +constructor. +Qed. + +Goal True. +Proof. +assert (exists n, n = 0) as [n Hn]. ++ exists 0; reflexivity. ++ exact I. +Qed. + +Goal True -> True. +Proof. +assert (H : 0 + 0 = 0) by reflexivity. +intros x; exact x. +Qed. + +Goal 1 + 1 = 2. +Proof. +change (?a + 1 = 2) with (2 = $a + 1). +reflexivity. +Qed. + +Goal (forall n, n = 0 -> False) -> False. +Proof. +intros H. +specialize (H 0 eq_refl). +destruct H. +Qed. + +Goal (forall n, n = 0 -> False) -> False. +Proof. +intros H. +specialize (H 0 eq_refl) as []. +Qed. diff --git a/test-suite/ltac2/matching.v b/test-suite/ltac2/matching.v new file mode 100644 index 0000000000..4338cbd32f --- /dev/null +++ b/test-suite/ltac2/matching.v @@ -0,0 +1,71 @@ +Require Import Ltac2.Ltac2 Ltac2.Notations. + +Ltac2 Type exn ::= [ Nope ]. + +Ltac2 check_id id id' := match Ident.equal id id' with +| true => () +| false => Control.throw Nope +end. + +Goal True -> False. +Proof. +Fail +let b := { contents := true } in +let f c := + match b.(contents) with + | true => Message.print (Message.of_constr c); b.(contents) := false; fail + | false => () + end +in +(** This fails because the matching is not allowed to backtrack once + it commits to a branch*) +lazy_match! '(nat -> bool) with context [?a] => f a end. +lazy_match! Control.goal () with ?a -> ?b => Message.print (Message.of_constr b) end. + +(** This one works by taking the second match context, i.e. ?a := nat *) +let b := { contents := true } in +let f c := + match b.(contents) with + | true => b.(contents) := false; fail + | false => Message.print (Message.of_constr c) + end +in +match! '(nat -> bool) with context [?a] => f a end. +Abort. + +Goal forall (i j : unit) (x y : nat) (b : bool), True. +Proof. +Fail match! goal with +| [ h : ?t, h' : ?t |- _ ] => () +end. +intros i j x y b. +match! goal with +| [ h : ?t, h' : ?t |- _ ] => + check_id h @x; + check_id h' @y +end. +match! reverse goal with +| [ h : ?t, h' : ?t |- _ ] => + check_id h @j; + check_id h' @i +end. +Abort. + +(* Check #79 *) +Goal 2 = 3. + Control.plus + (fun () + => lazy_match! goal with + | [ |- 2 = 3 ] => Control.zero (Tactic_failure None) + | [ |- 2 = _ ] => Control.zero (Tactic_failure (Some (Message.of_string "should not be printed"))) + end) + (fun e + => match e with + | Tactic_failure c + => match c with + | None => () + | _ => Control.zero e + end + | e => Control.zero e + end). +Abort. diff --git a/test-suite/ltac2/quot.v b/test-suite/ltac2/quot.v new file mode 100644 index 0000000000..624c4ad0c1 --- /dev/null +++ b/test-suite/ltac2/quot.v @@ -0,0 +1,26 @@ +Require Import Ltac2.Ltac2. + +(** Test for quotations *) + +Ltac2 ref0 () := reference:(&x). +Ltac2 ref1 () := reference:(nat). +Ltac2 ref2 () := reference:(Datatypes.nat). +Fail Ltac2 ref () := reference:(i_certainly_dont_exist). +Fail Ltac2 ref () := reference:(And.Me.neither). + +Goal True. +Proof. +let x := constr:(I) in +let y := constr:((fun z => z) $x) in +Control.refine (fun _ => y). +Qed. + +Goal True. +Proof. +(** Here, Ltac2 should not put its variables in the same environment as + Ltac1 otherwise the second binding fails as x is bound but not an + ident. *) +let x := constr:(I) in +let y := constr:((fun x => x) $x) in +Control.refine (fun _ => y). +Qed. diff --git a/test-suite/ltac2/rebind.v b/test-suite/ltac2/rebind.v new file mode 100644 index 0000000000..e1c20a2059 --- /dev/null +++ b/test-suite/ltac2/rebind.v @@ -0,0 +1,34 @@ +Require Import Ltac2.Ltac2 Ltac2.Notations. + +Ltac2 mutable foo () := constructor. + +Goal True. +Proof. +foo (). +Qed. + +Ltac2 Set foo := fun _ => fail. + +Goal True. +Proof. +Fail foo (). +constructor. +Qed. + +(** Not the right type *) +Fail Ltac2 Set foo := 0. + +Ltac2 bar () := (). + +(** Cannot redefine non-mutable tactics *) +Fail Ltac2 Set bar := fun _ => (). + +(** Subtype check *) + +Ltac2 mutable rec f x := f x. + +Fail Ltac2 Set f := fun x => x. + +Ltac2 mutable g x := x. + +Ltac2 Set g := f. diff --git a/test-suite/ltac2/stuff/ltac2.v b/test-suite/ltac2/stuff/ltac2.v new file mode 100644 index 0000000000..370bc70d15 --- /dev/null +++ b/test-suite/ltac2/stuff/ltac2.v @@ -0,0 +1,143 @@ +Require Import Ltac2.Ltac2. + +Ltac2 foo (_ : int) := + let f (x : int) := x in + let _ := f 0 in + f 1. + +Print Ltac2 foo. + +Import Control. + +Ltac2 exact x := refine (fun () => x). + +Print Ltac2 refine. +Print Ltac2 exact. + +Ltac2 foo' () := ident:(bla). + +Print Ltac2 foo'. + +Ltac2 bar x h := match x with +| None => constr:(fun H => ltac2:(exact (hyp ident:(H))) -> nat) +| Some x => x +end. + +Print Ltac2 bar. + +Ltac2 qux := Some 0. + +Print Ltac2 qux. + +Ltac2 Type foo := [ Foo (int) ]. + +Fail Ltac2 qux0 := Foo None. + +Ltac2 Type 'a ref := { mutable contents : 'a }. + +Fail Ltac2 qux0 := { contents := None }. +Ltac2 foo0 () := { contents := None }. + +Print Ltac2 foo0. + +Ltac2 qux0 x := x.(contents). +Ltac2 qux1 x := x.(contents) := x.(contents). + +Ltac2 qux2 := ([1;2], true). + +Print Ltac2 qux0. +Print Ltac2 qux1. +Print Ltac2 qux2. + +Import Control. + +Ltac2 qux3 x := constr:(nat -> ltac2:(refine (fun () => hyp x))). + +Print Ltac2 qux3. + +Ltac2 Type rec nat := [ O | S (nat) ]. + +Ltac2 message_of_nat n := +let rec aux n := +match n with +| O => Message.of_string "O" +| S n => Message.concat (Message.of_string "S") (aux n) +end in aux n. + +Print Ltac2 message_of_nat. + +Ltac2 numgoals () := + let r := { contents := O } in + enter (fun () => r.(contents) := S (r.(contents))); + r.(contents). + +Print Ltac2 numgoals. + +Goal True /\ False. +Proof. +let n := numgoals () in Message.print (message_of_nat n). +refine (fun () => open_constr:((fun x => conj _ _) 0)); (). +let n := numgoals () in Message.print (message_of_nat n). + +Fail (hyp ident:(x)). +Fail (enter (fun () => hyp ident:(There_is_no_spoon); ())). + +enter (fun () => Message.print (Message.of_string "foo")). + +enter (fun () => Message.print (Message.of_constr (goal ()))). +Fail enter (fun () => Message.print (Message.of_constr (qux3 ident:(x)))). +enter (fun () => plus (fun () => constr:(_); ()) (fun _ => ())). +plus + (fun () => enter (fun () => let x := ident:(foo) in let _ := hyp x in ())) (fun _ => Message.print (Message.of_string "failed")). +let x := { contents := 0 } in +let x := x.(contents) := x.(contents) in x. +Abort. + +Ltac2 Type exn ::= [ Foo ]. + +Goal True. +Proof. +plus (fun () => zero Foo) (fun _ => ()). +Abort. + +Ltac2 Type exn ::= [ Bar (string) ]. + +Goal True. +Proof. +Fail zero (Bar "lol"). +Abort. + +Ltac2 Notation "refine!" c(thunk(constr)) := refine c. + +Goal True. +Proof. +refine! I. +Abort. + +Goal True. +Proof. +let x () := plus (fun () => 0) (fun _ => 1) in +match case x with +| Val x => + match x with + | (x, k) => Message.print (Message.of_int (k Not_found)) + end +| Err x => Message.print (Message.of_string "Err") +end. +Abort. + +Goal (forall n : nat, n = 0 -> False) -> True. +Proof. +refine (fun () => '(fun H => _)). +Std.case true (hyp @H, Std.ExplicitBindings [Std.NamedHyp @n, '0]). +refine (fun () => 'eq_refl). +Qed. + +Goal forall x, 1 + x = x + 1. +Proof. +refine (fun () => '(fun x => _)). +Std.cbv { + Std.rBeta := true; Std.rMatch := true; Std.rFix := true; Std.rCofix := true; + Std.rZeta := true; Std.rDelta := true; Std.rConst := []; +} { Std.on_hyps := None; Std.on_concl := Std.AllOccurrences }. +Abort. diff --git a/test-suite/ltac2/tacticals.v b/test-suite/ltac2/tacticals.v new file mode 100644 index 0000000000..1a2fbcbb37 --- /dev/null +++ b/test-suite/ltac2/tacticals.v @@ -0,0 +1,34 @@ +Require Import Ltac2.Ltac2. + +Import Ltac2.Notations. + +Goal True. +Proof. +Fail fail. +Fail solve [ () ]. +try fail. +repeat fail. +repeat (). +solve [ constructor ]. +Qed. + +Goal True. +Proof. +first [ + Message.print (Message.of_string "Yay"); fail +| constructor +| Message.print (Message.of_string "I won't be printed") +]. +Qed. + +Goal True /\ True. +Proof. +Fail split > [ split | |]. +split > [split | split]. +Qed. + +Goal True /\ (True -> True) /\ True. +Proof. +split > [ | split] > [split | .. | split]. +intros H; refine &H. +Qed. diff --git a/test-suite/ltac2/typing.v b/test-suite/ltac2/typing.v new file mode 100644 index 0000000000..9f18292716 --- /dev/null +++ b/test-suite/ltac2/typing.v @@ -0,0 +1,72 @@ +Require Import Ltac2.Ltac2. + +(** Ltac2 is typed à la ML. *) + +Ltac2 test0 n := Int.add n 1. + +Print Ltac2 test0. + +Ltac2 test1 () := test0 0. + +Print Ltac2 test1. + +Fail Ltac2 test2 () := test0 true. + +Fail Ltac2 test2 () := test0 0 0. + +Ltac2 test3 f x := x, (f x, x). + +Print Ltac2 test3. + +(** Polymorphism *) + +Ltac2 rec list_length l := +match l with +| [] => 0 +| x :: l => Int.add 1 (list_length l) +end. + +Print Ltac2 list_length. + +(** Pattern-matching *) + +Ltac2 ifb b f g := match b with +| true => f () +| false => g () +end. + +Print Ltac2 ifb. + +Ltac2 if_not_found e f g := match e with +| Not_found => f () +| _ => g () +end. + +Fail Ltac2 ifb' b f g := match b with +| true => f () +end. + +Fail Ltac2 if_not_found' e f g := match e with +| Not_found => f () +end. + +(** Reimplementing 'do'. Return value of the function useless. *) + +Ltac2 rec do n tac := match Int.equal n 0 with +| true => () +| false => tac (); do (Int.sub n 1) tac +end. + +Print Ltac2 do. + +(** Non-function pure values are OK. *) + +Ltac2 tuple0 := ([1; 2], true, (fun () => "yay")). + +Print Ltac2 tuple0. + +(** Impure values are not. *) + +Fail Ltac2 not_a_value := { contents := 0 }. +Fail Ltac2 not_a_value := "nope". +Fail Ltac2 not_a_value := list_length []. diff --git a/tools/coq_dune.ml b/tools/coq_dune.ml index fa8b771a74..6ddc503542 100644 --- a/tools/coq_dune.ml +++ b/tools/coq_dune.ml @@ -214,7 +214,7 @@ let record_dune d ff = if Sys.file_exists sd && Sys.is_directory sd then let out = open_out (bpath [sd;"dune"]) in let fmt = formatter_of_out_channel out in - if List.nth d 0 = "plugins" then + if List.nth d 0 = "plugins" || List.nth d 0 = "user-contrib" then fprintf fmt "(include plugin_base.dune)@\n"; out_install fmt d ff; List.iter (pp_dep d fmt) ff; @@ -224,17 +224,20 @@ let record_dune d ff = eprintf "error in coq_dune, a directory disappeared: %s@\n%!" sd (* File Scanning *) -let scan_mlg m d = - let dir = ["plugins"; d] in +let scan_mlg ~root m d = + let dir = [root; d] in let m = DirMap.add dir [] m in let mlg = Sys.(List.filter (fun f -> Filename.(check_suffix f ".mlg")) Array.(to_list @@ readdir (bpath dir))) in - List.fold_left (fun m f -> add_map_list ["plugins"; d] (MLG f) m) m mlg + List.fold_left (fun m f -> add_map_list [root; d] (MLG f) m) m mlg -let scan_plugins m = +let scan_dir ~root m = let is_plugin_directory dir = Sys.(is_directory dir && file_exists (bpath [dir;"plugin_base.dune"])) in - let dirs = Sys.(List.filter (fun f -> is_plugin_directory @@ bpath ["plugins";f]) Array.(to_list @@ readdir "plugins")) in - List.fold_left scan_mlg m dirs + let dirs = Sys.(List.filter (fun f -> is_plugin_directory @@ bpath [root;f]) Array.(to_list @@ readdir root)) in + List.fold_left (scan_mlg ~root) m dirs + +let scan_plugins m = scan_dir ~root:"plugins" m +let scan_usercontrib m = scan_dir ~root:"user-contrib" m (* This will be removed when we drop support for Make *) let fix_cmo_cma file = @@ -291,5 +294,6 @@ let exec_ifile f = let _ = exec_ifile (fun ic -> let map = scan_plugins DirMap.empty in + let map = scan_usercontrib map in let map = read_vfiles ic map in out_map map) diff --git a/tools/coqdep.ml b/tools/coqdep.ml index 7114965a11..8823206252 100644 --- a/tools/coqdep.ml +++ b/tools/coqdep.ml @@ -529,6 +529,11 @@ let coqdep () = add_rec_dir_import add_known "plugins" ["Coq"]; add_rec_dir_import (fun _ -> add_caml_known) "theories" ["Coq"]; add_rec_dir_import (fun _ -> add_caml_known) "plugins" ["Coq"]; + let user = "user-contrib" in + if Sys.file_exists user then begin + add_rec_dir_no_import add_known user []; + add_rec_dir_no_import (fun _ -> add_caml_known) user []; + end; end else begin (* option_boot is actually always false in this branch *) Envars.set_coqlib ~fail:(fun msg -> raise (CoqlibError msg)); diff --git a/tools/coqdep_boot.ml b/tools/coqdep_boot.ml index aa023e6986..a638906c11 100644 --- a/tools/coqdep_boot.ml +++ b/tools/coqdep_boot.ml @@ -17,6 +17,9 @@ open Coqdep_common options (see for instance [option_natdynlk] below). *) +let split_period = Str.split (Str.regexp (Str.quote ".")) +let add_q_include path l = add_rec_dir_no_import add_known path (split_period l) + let rec parse = function | "-dyndep" :: "no" :: ll -> option_dynlink := No; parse ll | "-dyndep" :: "opt" :: ll -> option_dynlink := Opt; parse ll @@ -33,6 +36,7 @@ let rec parse = function add_caml_dir r; norec_dirs := StrSet.add r !norec_dirs; parse ll + | "-Q" :: r :: ln :: ll -> add_q_include r ln; parse ll | f :: ll -> treat_file None f; parse ll | [] -> () diff --git a/user-contrib/Ltac2/Array.v b/user-contrib/Ltac2/Array.v new file mode 100644 index 0000000000..11b64e3515 --- /dev/null +++ b/user-contrib/Ltac2/Array.v @@ -0,0 +1,14 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Require Import Ltac2.Init. + +Ltac2 @external make : int -> 'a -> 'a array := "ltac2" "array_make". +Ltac2 @external length : 'a array -> int := "ltac2" "array_length". +Ltac2 @external get : 'a array -> int -> 'a := "ltac2" "array_get". +Ltac2 @external set : 'a array -> int -> 'a -> unit := "ltac2" "array_set". diff --git a/user-contrib/Ltac2/Char.v b/user-contrib/Ltac2/Char.v new file mode 100644 index 0000000000..29fef60f2c --- /dev/null +++ b/user-contrib/Ltac2/Char.v @@ -0,0 +1,12 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Require Import Ltac2.Init. + +Ltac2 @external of_int : int -> char := "ltac2" "char_of_int". +Ltac2 @external to_int : char -> int := "ltac2" "char_to_int". diff --git a/user-contrib/Ltac2/Constr.v b/user-contrib/Ltac2/Constr.v new file mode 100644 index 0000000000..d8d222730e --- /dev/null +++ b/user-contrib/Ltac2/Constr.v @@ -0,0 +1,72 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Require Import Ltac2.Init. + +Ltac2 @ external type : constr -> constr := "ltac2" "constr_type". +(** Return the type of a term *) + +Ltac2 @ external equal : constr -> constr -> bool := "ltac2" "constr_equal". +(** Strict syntactic equality: only up to α-conversion and evar expansion *) + +Module Unsafe. + +(** Low-level access to kernel terms. Use with care! *) + +Ltac2 Type case. + +Ltac2 Type kind := [ +| Rel (int) +| Var (ident) +| Meta (meta) +| Evar (evar, constr array) +| Sort (sort) +| Cast (constr, cast, constr) +| Prod (ident option, constr, constr) +| Lambda (ident option, constr, constr) +| LetIn (ident option, constr, constr, constr) +| App (constr, constr array) +| Constant (constant, instance) +| Ind (inductive, instance) +| Constructor (constructor, instance) +| Case (case, constr, constr, constr array) +| Fix (int array, int, ident option array, constr array, constr array) +| CoFix (int, ident option array, constr array, constr array) +| Proj (projection, constr) +]. + +Ltac2 @ external kind : constr -> kind := "ltac2" "constr_kind". + +Ltac2 @ external make : kind -> constr := "ltac2" "constr_make". + +Ltac2 @ external check : constr -> constr result := "ltac2" "constr_check". +(** Checks that a constr generated by unsafe means is indeed safe in the + current environment, and returns it, or the error otherwise. Panics if + not focussed. *) + +Ltac2 @ external substnl : constr list -> int -> constr -> constr := "ltac2" "constr_substnl". +(** [substnl [r₁;...;rₙ] k c] substitutes in parallel [Rel(k+1); ...; Rel(k+n)] with + [r₁;...;rₙ] in [c]. *) + +Ltac2 @ external closenl : ident list -> int -> constr -> constr := "ltac2" "constr_closenl". +(** [closenl [x₁;...;xₙ] k c] abstracts over variables [x₁;...;xₙ] and replaces them with + [Rel(k); ...; Rel(k+n-1)] in [c]. If two names are identical, the one of least index is kept. *) + +Ltac2 @ external case : inductive -> case := "ltac2" "constr_case". +(** Generate the case information for a given inductive type. *) + +Ltac2 @ external constructor : inductive -> int -> constructor := "ltac2" "constr_constructor". +(** Generate the i-th constructor for a given inductive type. Indexing starts + at 0. Panics if there is no such constructor. *) + +End Unsafe. + +Ltac2 @ external in_context : ident -> constr -> (unit -> unit) -> constr := "ltac2" "constr_in_context". +(** On a focussed goal [Γ ⊢ A], [in_context id c tac] evaluates [tac] in a + focussed goal [Γ, id : c ⊢ ?X] and returns [fun (id : c) => t] where [t] is + the proof built by the tactic. *) diff --git a/user-contrib/Ltac2/Control.v b/user-contrib/Ltac2/Control.v new file mode 100644 index 0000000000..071c2ea8ce --- /dev/null +++ b/user-contrib/Ltac2/Control.v @@ -0,0 +1,76 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Require Import Ltac2.Init. + +(** Panic *) + +Ltac2 @ external throw : exn -> 'a := "ltac2" "throw". +(** Fatal exception throwing. This does not induce backtracking. *) + +(** Generic backtracking control *) + +Ltac2 @ external zero : exn -> 'a := "ltac2" "zero". +Ltac2 @ external plus : (unit -> 'a) -> (exn -> 'a) -> 'a := "ltac2" "plus". +Ltac2 @ external once : (unit -> 'a) -> 'a := "ltac2" "once". +Ltac2 @ external dispatch : (unit -> unit) list -> unit := "ltac2" "dispatch". +Ltac2 @ external extend : (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit := "ltac2" "extend". +Ltac2 @ external enter : (unit -> unit) -> unit := "ltac2" "enter". +Ltac2 @ external case : (unit -> 'a) -> ('a * (exn -> 'a)) result := "ltac2" "case". + +(** Proof state manipulation *) + +Ltac2 @ external focus : int -> int -> (unit -> 'a) -> 'a := "ltac2" "focus". +Ltac2 @ external shelve : unit -> unit := "ltac2" "shelve". +Ltac2 @ external shelve_unifiable : unit -> unit := "ltac2" "shelve_unifiable". + +Ltac2 @ external new_goal : evar -> unit := "ltac2" "new_goal". +(** Adds the given evar to the list of goals as the last one. If it is + already defined in the current state, don't do anything. Panics if the + evar is not in the current state. *) + +Ltac2 @ external progress : (unit -> 'a) -> 'a := "ltac2" "progress". + +(** Goal inspection *) + +Ltac2 @ external goal : unit -> constr := "ltac2" "goal". +(** Panics if there is not exactly one goal under focus. Otherwise returns + the conclusion of this goal. *) + +Ltac2 @ external hyp : ident -> constr := "ltac2" "hyp". +(** Panics if there is more than one goal under focus. If there is no + goal under focus, looks for the section variable with the given name. + If there is one, looks for the hypothesis with the given name. *) + +Ltac2 @ external hyps : unit -> (ident * constr option * constr) list := "ltac2" "hyps". +(** Panics if there is more than one goal under focus. If there is no + goal under focus, returns the list of section variables. + If there is one, returns the list of hypotheses. In both cases, the + list is ordered with rightmost values being last introduced. *) + +(** Refinement *) + +Ltac2 @ external refine : (unit -> constr) -> unit := "ltac2" "refine". + +(** Evars *) + +Ltac2 @ external with_holes : (unit -> 'a) -> ('a -> 'b) -> 'b := "ltac2" "with_holes". +(** [with_holes x f] evaluates [x], then apply [f] to the result, and fails if + all evars generated by the call to [x] have not been solved when [f] + returns. *) + +(** Misc *) + +Ltac2 @ external time : string option -> (unit -> 'a) -> 'a := "ltac2" "time". +(** Displays the time taken by a tactic to evaluate. *) + +Ltac2 @ external abstract : ident option -> (unit -> unit) -> unit := "ltac2" "abstract". +(** Abstract a subgoal. *) + +Ltac2 @ external check_interrupt : unit -> unit := "ltac2" "check_interrupt". +(** For internal use. *) diff --git a/user-contrib/Ltac2/Env.v b/user-contrib/Ltac2/Env.v new file mode 100644 index 0000000000..4aa1718c9a --- /dev/null +++ b/user-contrib/Ltac2/Env.v @@ -0,0 +1,26 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +From Ltac2 Require Import Init Std. + +Ltac2 @ external get : ident list -> Std.reference option := "ltac2" "env_get". +(** Returns the global reference corresponding to the absolute name given as + argument if it exists. *) + +Ltac2 @ external expand : ident list -> Std.reference list := "ltac2" "env_expand". +(** Returns the list of all global references whose absolute name contains + the argument list as a prefix. *) + +Ltac2 @ external path : Std.reference -> ident list := "ltac2" "env_path". +(** Returns the absolute name of the given reference. Panics if the reference + does not exist. *) + +Ltac2 @ external instantiate : Std.reference -> constr := "ltac2" "env_instantiate". +(** Returns a fresh instance of the corresponding reference, in particular + generating fresh universe variables and constraints when this reference is + universe-polymorphic. *) diff --git a/user-contrib/Ltac2/Fresh.v b/user-contrib/Ltac2/Fresh.v new file mode 100644 index 0000000000..5e876bb077 --- /dev/null +++ b/user-contrib/Ltac2/Fresh.v @@ -0,0 +1,26 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Require Import Ltac2.Init. + +Module Free. + +Ltac2 Type t. +(** Type of sets of free variables *) + +Ltac2 @ external union : t -> t -> t := "ltac2" "fresh_free_union". + +Ltac2 @ external of_ids : ident list -> t := "ltac2" "fresh_free_of_ids". + +Ltac2 @ external of_constr : constr -> t := "ltac2" "fresh_free_of_constr". + +End Free. + +Ltac2 @ external fresh : Free.t -> ident -> ident := "ltac2" "fresh_fresh". +(** Generate a fresh identifier with the given base name which is not a + member of the provided set of free variables. *) diff --git a/user-contrib/Ltac2/Ident.v b/user-contrib/Ltac2/Ident.v new file mode 100644 index 0000000000..55456afbe2 --- /dev/null +++ b/user-contrib/Ltac2/Ident.v @@ -0,0 +1,17 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Require Import Ltac2.Init. + +Ltac2 Type t := ident. + +Ltac2 @ external equal : t -> t -> bool := "ltac2" "ident_equal". + +Ltac2 @ external of_string : string -> t option := "ltac2" "ident_of_string". + +Ltac2 @ external to_string : t -> string := "ltac2" "ident_to_string". diff --git a/user-contrib/Ltac2/Init.v b/user-contrib/Ltac2/Init.v new file mode 100644 index 0000000000..16e7d7a6f9 --- /dev/null +++ b/user-contrib/Ltac2/Init.v @@ -0,0 +1,69 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Declare ML Module "ltac2_plugin". + +(** Primitive types *) + +Ltac2 Type int. +Ltac2 Type string. +Ltac2 Type char. +Ltac2 Type ident. + +(** Constr-specific built-in types *) +Ltac2 Type meta. +Ltac2 Type evar. +Ltac2 Type sort. +Ltac2 Type cast. +Ltac2 Type instance. +Ltac2 Type constant. +Ltac2 Type inductive. +Ltac2 Type constructor. +Ltac2 Type projection. +Ltac2 Type pattern. +Ltac2 Type constr. + +Ltac2 Type message. +Ltac2 Type exn := [ .. ]. +Ltac2 Type 'a array. + +(** Pervasive types *) + +Ltac2 Type 'a option := [ None | Some ('a) ]. + +Ltac2 Type 'a ref := { mutable contents : 'a }. + +Ltac2 Type bool := [ true | false ]. + +Ltac2 Type 'a result := [ Val ('a) | Err (exn) ]. + +(** Pervasive exceptions *) + +Ltac2 Type err. +(** Coq internal errors. Cannot be constructed, merely passed around. *) + +Ltac2 Type exn ::= [ Internal (err) ]. +(** Wrapper around the errors raised by Coq implementation. *) + +Ltac2 Type exn ::= [ Out_of_bounds ]. +(** Used for bound checking, e.g. with String and Array. *) + +Ltac2 Type exn ::= [ Not_focussed ]. +(** In Ltac2, the notion of "current environment" only makes sense when there is + at most one goal under focus. Contrarily to Ltac1, instead of dynamically + focussing when we need it, we raise this non-backtracking error when it does + not make sense. *) + +Ltac2 Type exn ::= [ Not_found ]. +(** Used when something is missing. *) + +Ltac2 Type exn ::= [ Match_failure ]. +(** Used to signal a pattern didn't match a term. *) + +Ltac2 Type exn ::= [ Tactic_failure (message option) ]. +(** Generic error for tactic failure. *) diff --git a/user-contrib/Ltac2/Int.v b/user-contrib/Ltac2/Int.v new file mode 100644 index 0000000000..0a90d757b6 --- /dev/null +++ b/user-contrib/Ltac2/Int.v @@ -0,0 +1,18 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Require Import Ltac2.Init. + +Ltac2 Type exn ::= [ Division_by_zero ]. + +Ltac2 @ external equal : int -> int -> bool := "ltac2" "int_equal". +Ltac2 @ external compare : int -> int -> int := "ltac2" "int_compare". +Ltac2 @ external add : int -> int -> int := "ltac2" "int_add". +Ltac2 @ external sub : int -> int -> int := "ltac2" "int_sub". +Ltac2 @ external mul : int -> int -> int := "ltac2" "int_mul". +Ltac2 @ external neg : int -> int := "ltac2" "int_neg". diff --git a/user-contrib/Ltac2/Ltac1.v b/user-contrib/Ltac2/Ltac1.v new file mode 100644 index 0000000000..c4e0b606d0 --- /dev/null +++ b/user-contrib/Ltac2/Ltac1.v @@ -0,0 +1,36 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +(** This module defines the Ltac2 FFI to Ltac1 code. Due to intricate semantics + of the latter, the functions described here are voluntarily under-specified. + Not for the casual user, handle with care and expect undefined behaviours + otherwise. **) + +Require Import Ltac2.Init. + +Ltac2 Type t. +(** Dynamically-typed Ltac1 values. *) + +Ltac2 @ external ref : ident list -> t := "ltac2" "ltac1_ref". +(** Returns the Ltac1 definition with the given absolute name. *) + +Ltac2 @ external run : t -> unit := "ltac2" "ltac1_run". +(** Runs an Ltac1 value, assuming it is a 'tactic', i.e. not returning + anything. *) + +Ltac2 @ external apply : t -> t list -> (t -> unit) -> unit := "ltac2" "ltac1_apply". +(** Applies an Ltac1 value to a list of arguments, and provides the result in + CPS style. It does **not** run the returned value. *) + +(** Conversion functions *) + +Ltac2 @ external of_constr : constr -> t := "ltac2" "ltac1_of_constr". +Ltac2 @ external to_constr : t -> constr option := "ltac2" "ltac1_to_constr". + +Ltac2 @ external of_list : t list -> t := "ltac2" "ltac1_of_list". +Ltac2 @ external to_list : t -> t list option := "ltac2" "ltac1_to_list". diff --git a/user-contrib/Ltac2/Ltac2.v b/user-contrib/Ltac2/Ltac2.v new file mode 100644 index 0000000000..ac90f63560 --- /dev/null +++ b/user-contrib/Ltac2/Ltac2.v @@ -0,0 +1,24 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Require Export Ltac2.Init. + +Require Ltac2.Int. +Require Ltac2.Char. +Require Ltac2.String. +Require Ltac2.Ident. +Require Ltac2.Array. +Require Ltac2.Message. +Require Ltac2.Constr. +Require Ltac2.Control. +Require Ltac2.Fresh. +Require Ltac2.Pattern. +Require Ltac2.Std. +Require Ltac2.Env. +Require Ltac2.Ltac1. +Require Export Ltac2.Notations. diff --git a/user-contrib/Ltac2/Message.v b/user-contrib/Ltac2/Message.v new file mode 100644 index 0000000000..7bffe0746b --- /dev/null +++ b/user-contrib/Ltac2/Message.v @@ -0,0 +1,25 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Require Import Ltac2.Init. + +Ltac2 @ external print : message -> unit := "ltac2" "print". + +Ltac2 @ external of_string : string -> message := "ltac2" "message_of_string". + +Ltac2 @ external of_int : int -> message := "ltac2" "message_of_int". + +Ltac2 @ external of_ident : ident -> message := "ltac2" "message_of_ident". + +Ltac2 @ external of_constr : constr -> message := "ltac2" "message_of_constr". +(** Panics if there is more than one goal under focus. *) + +Ltac2 @ external of_exn : exn -> message := "ltac2" "message_of_exn". +(** Panics if there is more than one goal under focus. *) + +Ltac2 @ external concat : message -> message -> message := "ltac2" "message_concat". diff --git a/user-contrib/Ltac2/Notations.v b/user-contrib/Ltac2/Notations.v new file mode 100644 index 0000000000..0eab36df82 --- /dev/null +++ b/user-contrib/Ltac2/Notations.v @@ -0,0 +1,556 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Require Import Ltac2.Init. +Require Ltac2.Control Ltac2.Pattern Ltac2.Array Ltac2.Int Ltac2.Std. + +(** Constr matching *) + +Ltac2 Notation "lazy_match!" t(tactic(6)) "with" m(constr_matching) "end" := + Pattern.lazy_match0 t m. + +Ltac2 Notation "multi_match!" t(tactic(6)) "with" m(constr_matching) "end" := + Pattern.multi_match0 t m. + +Ltac2 Notation "match!" t(tactic(6)) "with" m(constr_matching) "end" := + Pattern.one_match0 t m. + +(** Goal matching *) + +Ltac2 Notation "lazy_match!" "goal" "with" m(goal_matching) "end" := + Pattern.lazy_goal_match0 false m. + +Ltac2 Notation "multi_match!" "goal" "with" m(goal_matching) "end" := + Pattern.multi_goal_match0 false m. + +Ltac2 Notation "match!" "goal" "with" m(goal_matching) "end" := + Pattern.one_goal_match0 false m. + +Ltac2 Notation "lazy_match!" "reverse" "goal" "with" m(goal_matching) "end" := + Pattern.lazy_goal_match0 true m. + +Ltac2 Notation "multi_match!" "reverse" "goal" "with" m(goal_matching) "end" := + Pattern.multi_goal_match0 true m. + +Ltac2 Notation "match!" "reverse" "goal" "with" m(goal_matching) "end" := + Pattern.one_goal_match0 true m. + +(** Tacticals *) + +Ltac2 orelse t f := +match Control.case t with +| Err e => f e +| Val ans => + let (x, k) := ans in + Control.plus (fun _ => x) k +end. + +Ltac2 ifcatch t s f := +match Control.case t with +| Err e => f e +| Val ans => + let (x, k) := ans in + Control.plus (fun _ => s x) (fun e => s (k e)) +end. + +Ltac2 fail0 (_ : unit) := Control.enter (fun _ => Control.zero (Tactic_failure None)). + +Ltac2 Notation fail := fail0 (). + +Ltac2 try0 t := Control.enter (fun _ => orelse t (fun _ => ())). + +Ltac2 Notation try := try0. + +Ltac2 rec repeat0 (t : unit -> unit) := + Control.enter (fun () => + ifcatch (fun _ => Control.progress t) + (fun _ => Control.check_interrupt (); repeat0 t) (fun _ => ())). + +Ltac2 Notation repeat := repeat0. + +Ltac2 dispatch0 t (head, tail) := + match tail with + | None => Control.enter (fun _ => t (); Control.dispatch head) + | Some tacs => + let (def, rem) := tacs in + Control.enter (fun _ => t (); Control.extend head def rem) + end. + +Ltac2 Notation t(thunk(self)) ">" "[" l(dispatch) "]" : 4 := dispatch0 t l. + +Ltac2 do0 n t := + let rec aux n t := match Int.equal n 0 with + | true => () + | false => t (); aux (Int.sub n 1) t + end in + aux (n ()) t. + +Ltac2 Notation do := do0. + +Ltac2 Notation once := Control.once. + +Ltac2 progress0 tac := Control.enter (fun _ => Control.progress tac). + +Ltac2 Notation progress := progress0. + +Ltac2 rec first0 tacs := +match tacs with +| [] => Control.zero (Tactic_failure None) +| tac :: tacs => Control.enter (fun _ => orelse tac (fun _ => first0 tacs)) +end. + +Ltac2 Notation "first" "[" tacs(list0(thunk(tactic(6)), "|")) "]" := first0 tacs. + +Ltac2 complete tac := + let ans := tac () in + Control.enter (fun () => Control.zero (Tactic_failure None)); + ans. + +Ltac2 rec solve0 tacs := +match tacs with +| [] => Control.zero (Tactic_failure None) +| tac :: tacs => + Control.enter (fun _ => orelse (fun _ => complete tac) (fun _ => solve0 tacs)) +end. + +Ltac2 Notation "solve" "[" tacs(list0(thunk(tactic(6)), "|")) "]" := solve0 tacs. + +Ltac2 time0 tac := Control.time None tac. + +Ltac2 Notation time := time0. + +Ltac2 abstract0 tac := Control.abstract None tac. + +Ltac2 Notation abstract := abstract0. + +(** Base tactics *) + +(** Note that we redeclare notations that can be parsed as mere identifiers + as abbreviations, so that it allows to parse them as function arguments + without having to write them within parentheses. *) + +(** Enter and check evar resolution *) +Ltac2 enter_h ev f arg := +match ev with +| true => Control.enter (fun () => f ev (arg ())) +| false => + Control.enter (fun () => + Control.with_holes arg (fun x => f ev x)) +end. + +Ltac2 intros0 ev p := + Control.enter (fun () => Std.intros false p). + +Ltac2 Notation "intros" p(intropatterns) := intros0 false p. +Ltac2 Notation intros := intros. + +Ltac2 Notation "eintros" p(intropatterns) := intros0 true p. +Ltac2 Notation eintros := eintros. + +Ltac2 split0 ev bnd := + enter_h ev Std.split bnd. + +Ltac2 Notation "split" bnd(thunk(with_bindings)) := split0 false bnd. +Ltac2 Notation split := split. + +Ltac2 Notation "esplit" bnd(thunk(with_bindings)) := split0 true bnd. +Ltac2 Notation esplit := esplit. + +Ltac2 exists0 ev bnds := match bnds with +| [] => split0 ev (fun () => Std.NoBindings) +| _ => + let rec aux bnds := match bnds with + | [] => () + | bnd :: bnds => split0 ev bnd; aux bnds + end in + aux bnds +end. + +Ltac2 Notation "exists" bnd(list0(thunk(bindings), ",")) := exists0 false bnd. +(* Ltac2 Notation exists := exists. *) + +Ltac2 Notation "eexists" bnd(list0(thunk(bindings), ",")) := exists0 true bnd. +Ltac2 Notation eexists := eexists. + +Ltac2 left0 ev bnd := enter_h ev Std.left bnd. + +Ltac2 Notation "left" bnd(thunk(with_bindings)) := left0 false bnd. +Ltac2 Notation left := left. + +Ltac2 Notation "eleft" bnd(thunk(with_bindings)) := left0 true bnd. +Ltac2 Notation eleft := eleft. + +Ltac2 right0 ev bnd := enter_h ev Std.right bnd. + +Ltac2 Notation "right" bnd(thunk(with_bindings)) := right0 false bnd. +Ltac2 Notation right := right. + +Ltac2 Notation "eright" bnd(thunk(with_bindings)) := right0 true bnd. +Ltac2 Notation eright := eright. + +Ltac2 constructor0 ev n bnd := + enter_h ev (fun ev bnd => Std.constructor_n ev n bnd) bnd. + +Ltac2 Notation "constructor" := Control.enter (fun () => Std.constructor false). +Ltac2 Notation constructor := constructor. +Ltac2 Notation "constructor" n(tactic) bnd(thunk(with_bindings)) := constructor0 false n bnd. + +Ltac2 Notation "econstructor" := Control.enter (fun () => Std.constructor true). +Ltac2 Notation econstructor := econstructor. +Ltac2 Notation "econstructor" n(tactic) bnd(thunk(with_bindings)) := constructor0 true n bnd. + +Ltac2 specialize0 c pat := + enter_h false (fun _ c => Std.specialize c pat) c. + +Ltac2 Notation "specialize" c(thunk(seq(constr, with_bindings))) ipat(opt(seq("as", intropattern))) := + specialize0 c ipat. + +Ltac2 elim0 ev c bnd use := + let f ev (c, bnd, use) := Std.elim ev (c, bnd) use in + enter_h ev f (fun () => c (), bnd (), use ()). + +Ltac2 Notation "elim" c(thunk(constr)) bnd(thunk(with_bindings)) + use(thunk(opt(seq("using", constr, with_bindings)))) := + elim0 false c bnd use. + +Ltac2 Notation "eelim" c(thunk(constr)) bnd(thunk(with_bindings)) + use(thunk(opt(seq("using", constr, with_bindings)))) := + elim0 true c bnd use. + +Ltac2 apply0 adv ev cb cl := + Std.apply adv ev cb cl. + +Ltac2 Notation "eapply" + cb(list1(thunk(seq(constr, with_bindings)), ",")) + cl(opt(seq("in", ident, opt(seq("as", intropattern))))) := + apply0 true true cb cl. + +Ltac2 Notation "apply" + cb(list1(thunk(seq(constr, with_bindings)), ",")) + cl(opt(seq("in", ident, opt(seq("as", intropattern))))) := + apply0 true false cb cl. + +Ltac2 default_on_concl cl := +match cl with +| None => { Std.on_hyps := Some []; Std.on_concl := Std.AllOccurrences } +| Some cl => cl +end. + +Ltac2 pose0 ev p := + enter_h ev (fun ev (na, p) => Std.pose na p) p. + +Ltac2 Notation "pose" p(thunk(pose)) := + pose0 false p. + +Ltac2 Notation "epose" p(thunk(pose)) := + pose0 true p. + +Ltac2 Notation "set" p(thunk(pose)) cl(opt(clause)) := + Std.set false p (default_on_concl cl). + +Ltac2 Notation "eset" p(thunk(pose)) cl(opt(clause)) := + Std.set true p (default_on_concl cl). + +Ltac2 assert0 ev ast := + enter_h ev (fun _ ast => Std.assert ast) ast. + +Ltac2 Notation "assert" ast(thunk(assert)) := assert0 false ast. + +Ltac2 Notation "eassert" ast(thunk(assert)) := assert0 true ast. + +Ltac2 default_everywhere cl := +match cl with +| None => { Std.on_hyps := None; Std.on_concl := Std.AllOccurrences } +| Some cl => cl +end. + +Ltac2 Notation "remember" + c(thunk(open_constr)) + na(opt(seq("as", ident))) + pat(opt(seq("eqn", ":", intropattern))) + cl(opt(clause)) := + Std.remember false na c pat (default_everywhere cl). + +Ltac2 Notation "eremember" + c(thunk(open_constr)) + na(opt(seq("as", ident))) + pat(opt(seq("eqn", ":", intropattern))) + cl(opt(clause)) := + Std.remember true na c pat (default_everywhere cl). + +Ltac2 induction0 ev ic use := + let f ev use := Std.induction ev ic use in + enter_h ev f use. + +Ltac2 Notation "induction" + ic(list1(induction_clause, ",")) + use(thunk(opt(seq("using", constr, with_bindings)))) := + induction0 false ic use. + +Ltac2 Notation "einduction" + ic(list1(induction_clause, ",")) + use(thunk(opt(seq("using", constr, with_bindings)))) := + induction0 true ic use. + +Ltac2 generalize0 gen := + enter_h false (fun _ gen => Std.generalize gen) gen. + +Ltac2 Notation "generalize" + gen(thunk(list1(seq (open_constr, occurrences, opt(seq("as", ident))), ","))) := + generalize0 gen. + +Ltac2 destruct0 ev ic use := + let f ev use := Std.destruct ev ic use in + enter_h ev f use. + +Ltac2 Notation "destruct" + ic(list1(induction_clause, ",")) + use(thunk(opt(seq("using", constr, with_bindings)))) := + destruct0 false ic use. + +Ltac2 Notation "edestruct" + ic(list1(induction_clause, ",")) + use(thunk(opt(seq("using", constr, with_bindings)))) := + destruct0 true ic use. + +Ltac2 Notation "simple" "inversion" + arg(destruction_arg) + pat(opt(seq("as", intropattern))) + ids(opt(seq("in", list1(ident)))) := + Std.inversion Std.SimpleInversion arg pat ids. + +Ltac2 Notation "inversion" + arg(destruction_arg) + pat(opt(seq("as", intropattern))) + ids(opt(seq("in", list1(ident)))) := + Std.inversion Std.FullInversion arg pat ids. + +Ltac2 Notation "inversion_clear" + arg(destruction_arg) + pat(opt(seq("as", intropattern))) + ids(opt(seq("in", list1(ident)))) := + Std.inversion Std.FullInversionClear arg pat ids. + +Ltac2 Notation "red" cl(opt(clause)) := + Std.red (default_on_concl cl). +Ltac2 Notation red := red. + +Ltac2 Notation "hnf" cl(opt(clause)) := + Std.hnf (default_on_concl cl). +Ltac2 Notation hnf := hnf. + +Ltac2 Notation "simpl" s(strategy) pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := + Std.simpl s pl (default_on_concl cl). +Ltac2 Notation simpl := simpl. + +Ltac2 Notation "cbv" s(strategy) cl(opt(clause)) := + Std.cbv s (default_on_concl cl). +Ltac2 Notation cbv := cbv. + +Ltac2 Notation "cbn" s(strategy) cl(opt(clause)) := + Std.cbn s (default_on_concl cl). +Ltac2 Notation cbn := cbn. + +Ltac2 Notation "lazy" s(strategy) cl(opt(clause)) := + Std.lazy s (default_on_concl cl). +Ltac2 Notation lazy := lazy. + +Ltac2 Notation "unfold" pl(list1(seq(reference, occurrences), ",")) cl(opt(clause)) := + Std.unfold pl (default_on_concl cl). + +Ltac2 fold0 pl cl := + let cl := default_on_concl cl in + Control.enter (fun () => Control.with_holes pl (fun pl => Std.fold pl cl)). + +Ltac2 Notation "fold" pl(thunk(list1(open_constr))) cl(opt(clause)) := + fold0 pl cl. + +Ltac2 Notation "pattern" pl(list1(seq(constr, occurrences), ",")) cl(opt(clause)) := + Std.pattern pl (default_on_concl cl). + +Ltac2 Notation "vm_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := + Std.vm pl (default_on_concl cl). +Ltac2 Notation vm_compute := vm_compute. + +Ltac2 Notation "native_compute" pl(opt(seq(pattern, occurrences))) cl(opt(clause)) := + Std.native pl (default_on_concl cl). +Ltac2 Notation native_compute := native_compute. + +Ltac2 change0 p cl := + let (pat, c) := p in + Std.change pat c (default_on_concl cl). + +Ltac2 Notation "change" c(conversion) cl(opt(clause)) := change0 c cl. + +Ltac2 rewrite0 ev rw cl tac := + let cl := default_on_concl cl in + Std.rewrite ev rw cl tac. + +Ltac2 Notation "rewrite" + rw(list1(rewriting, ",")) + cl(opt(clause)) + tac(opt(seq("by", thunk(tactic)))) := + rewrite0 false rw cl tac. + +Ltac2 Notation "erewrite" + rw(list1(rewriting, ",")) + cl(opt(clause)) + tac(opt(seq("by", thunk(tactic)))) := + rewrite0 true rw cl tac. + +(** coretactics *) + +Ltac2 exact0 ev c := + Control.enter (fun _ => + match ev with + | true => + let c := c () in + Control.refine (fun _ => c) + | false => + Control.with_holes c (fun c => Control.refine (fun _ => c)) + end + ). + +Ltac2 Notation "exact" c(thunk(open_constr)) := exact0 false c. +Ltac2 Notation "eexact" c(thunk(open_constr)) := exact0 true c. + +Ltac2 Notation "intro" id(opt(ident)) mv(opt(move_location)) := Std.intro id mv. +Ltac2 Notation intro := intro. + +Ltac2 Notation "move" id(ident) mv(move_location) := Std.move id mv. + +Ltac2 Notation reflexivity := Std.reflexivity (). + +Ltac2 symmetry0 cl := + Std.symmetry (default_on_concl cl). + +Ltac2 Notation "symmetry" cl(opt(clause)) := symmetry0 cl. +Ltac2 Notation symmetry := symmetry. + +Ltac2 Notation "revert" ids(list1(ident)) := Std.revert ids. + +Ltac2 Notation assumption := Std.assumption (). + +Ltac2 Notation etransitivity := Std.etransitivity (). + +Ltac2 Notation admit := Std.admit (). + +Ltac2 clear0 ids := match ids with +| [] => Std.keep [] +| _ => Std.clear ids +end. + +Ltac2 Notation "clear" ids(list0(ident)) := clear0 ids. +Ltac2 Notation "clear" "-" ids(list1(ident)) := Std.keep ids. +Ltac2 Notation clear := clear. + +Ltac2 Notation refine := Control.refine. + +(** extratactics *) + +Ltac2 absurd0 c := Control.enter (fun _ => Std.absurd (c ())). + +Ltac2 Notation "absurd" c(thunk(open_constr)) := absurd0 c. + +Ltac2 subst0 ids := match ids with +| [] => Std.subst_all () +| _ => Std.subst ids +end. + +Ltac2 Notation "subst" ids(list0(ident)) := subst0 ids. +Ltac2 Notation subst := subst. + +Ltac2 Notation "discriminate" arg(opt(destruction_arg)) := + Std.discriminate false arg. +Ltac2 Notation discriminate := discriminate. + +Ltac2 Notation "ediscriminate" arg(opt(destruction_arg)) := + Std.discriminate true arg. +Ltac2 Notation ediscriminate := ediscriminate. + +Ltac2 Notation "injection" arg(opt(destruction_arg)) ipat(opt(seq("as", intropatterns))):= + Std.injection false ipat arg. + +Ltac2 Notation "einjection" arg(opt(destruction_arg)) ipat(opt(seq("as", intropatterns))):= + Std.injection true ipat arg. + +(** Auto *) + +Ltac2 default_db dbs := match dbs with +| None => Some [] +| Some dbs => + match dbs with + | None => None + | Some l => Some l + end +end. + +Ltac2 default_list use := match use with +| None => [] +| Some use => use +end. + +Ltac2 trivial0 use dbs := + let dbs := default_db dbs in + let use := default_list use in + Std.trivial Std.Off use dbs. + +Ltac2 Notation "trivial" + use(opt(seq("using", list1(thunk(constr), ",")))) + dbs(opt(seq("with", hintdb))) := trivial0 use dbs. + +Ltac2 Notation trivial := trivial. + +Ltac2 auto0 n use dbs := + let dbs := default_db dbs in + let use := default_list use in + Std.auto Std.Off n use dbs. + +Ltac2 Notation "auto" n(opt(tactic(0))) + use(opt(seq("using", list1(thunk(constr), ",")))) + dbs(opt(seq("with", hintdb))) := auto0 n use dbs. + +Ltac2 Notation auto := auto. + +Ltac2 new_eauto0 n use dbs := + let dbs := default_db dbs in + let use := default_list use in + Std.new_auto Std.Off n use dbs. + +Ltac2 Notation "new" "auto" n(opt(tactic(0))) + use(opt(seq("using", list1(thunk(constr), ",")))) + dbs(opt(seq("with", hintdb))) := new_eauto0 n use dbs. + +Ltac2 eauto0 n p use dbs := + let dbs := default_db dbs in + let use := default_list use in + Std.eauto Std.Off n p use dbs. + +Ltac2 Notation "eauto" n(opt(tactic(0))) p(opt(tactic(0))) + use(opt(seq("using", list1(thunk(constr), ",")))) + dbs(opt(seq("with", hintdb))) := eauto0 n p use dbs. + +Ltac2 Notation eauto := eauto. + +Ltac2 Notation "typeclasses_eauto" n(opt(tactic(0))) + dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto None n dbs. + +Ltac2 Notation "typeclasses_eauto" "bfs" n(opt(tactic(0))) + dbs(opt(seq("with", list1(ident)))) := Std.typeclasses_eauto (Some Std.BFS) n dbs. + +Ltac2 Notation typeclasses_eauto := typeclasses_eauto. + +(** Congruence *) + +Ltac2 f_equal0 () := ltac1:(f_equal). +Ltac2 Notation f_equal := f_equal0 (). + +(** now *) + +Ltac2 now0 t := t (); ltac1:(easy). +Ltac2 Notation "now" t(thunk(self)) := now0 t. diff --git a/user-contrib/Ltac2/Pattern.v b/user-contrib/Ltac2/Pattern.v new file mode 100644 index 0000000000..8d1fb0cd8a --- /dev/null +++ b/user-contrib/Ltac2/Pattern.v @@ -0,0 +1,145 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Require Import Ltac2.Init. +Require Ltac2.Control. + +Ltac2 Type t := pattern. + +Ltac2 Type context. + +Ltac2 Type match_kind := [ +| MatchPattern +| MatchContext +]. + +Ltac2 @ external empty_context : unit -> context := + "ltac2" "pattern_empty_context". +(** A trivial context only made of the hole. *) + +Ltac2 @ external matches : t -> constr -> (ident * constr) list := + "ltac2" "pattern_matches". +(** If the term matches the pattern, returns the bound variables. If it doesn't, + fail with [Match_failure]. Panics if not focussed. *) + +Ltac2 @ external matches_subterm : t -> constr -> context * ((ident * constr) list) := + "ltac2" "pattern_matches_subterm". +(** Returns a stream of results corresponding to all of the subterms of the term + that matches the pattern as in [matches]. The stream is encoded as a + backtracking value whose last exception is [Match_failure]. The additional + value compared to [matches] is the context of the match, to be filled with + the instantiate function. *) + +Ltac2 @ external matches_vect : t -> constr -> constr array := + "ltac2" "pattern_matches_vect". +(** Internal version of [matches] that does not return the identifiers. *) + +Ltac2 @ external matches_subterm_vect : t -> constr -> context * constr array := + "ltac2" "pattern_matches_subterm_vect". +(** Internal version of [matches_subterms] that does not return the identifiers. *) + +Ltac2 @ external matches_goal : bool -> (match_kind * t) list -> (match_kind * t) -> + ident array * context array * constr array * context := + "ltac2" "pattern_matches_goal". +(** Given a list of patterns [hpats] for hypotheses and one pattern [cpat] for the + conclusion, [matches_goal rev hpats cpat] produces (a stream of) tuples of: + - An array of idents, whose size is the length of [hpats], corresponding to the + name of matched hypotheses. + - An array of contexts, whose size is the length of [hpats], corresponding to + the contexts matched for every hypothesis pattern. In case the match kind of + a hypothesis was [MatchPattern], the corresponding context is ensured to be empty. + - An array of terms, whose size is the total number of pattern variables without + duplicates. Terms are ordered by identifier order, e.g. ?a comes before ?b. + - A context corresponding to the conclusion, which is ensured to be empty if + the kind of [cpat] was [MatchPattern]. + This produces a backtracking stream of results containing all the possible + result combinations. The order of considered hypotheses is reversed if [rev] + is true. +*) + +Ltac2 @ external instantiate : context -> constr -> constr := + "ltac2" "pattern_instantiate". +(** Fill the hole of a context with the given term. *) + +(** Implementation of Ltac matching over terms and goals *) + +Ltac2 lazy_match0 t pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + let next _ := interp m in + let (knd, pat, f) := p in + let p := match knd with + | MatchPattern => + (fun _ => + let context := empty_context () in + let bind := matches_vect pat t in + fun _ => f context bind) + | MatchContext => + (fun _ => + let (context, bind) := matches_subterm_vect pat t in + fun _ => f context bind) + end in + Control.plus p next + end in + Control.once (fun () => interp pats) (). + +Ltac2 multi_match0 t pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + let next _ := interp m in + let (knd, pat, f) := p in + let p := match knd with + | MatchPattern => + (fun _ => + let context := empty_context () in + let bind := matches_vect pat t in + f context bind) + | MatchContext => + (fun _ => + let (context, bind) := matches_subterm_vect pat t in + f context bind) + end in + Control.plus p next + end in + interp pats. + +Ltac2 one_match0 t m := Control.once (fun _ => multi_match0 t m). + +Ltac2 lazy_goal_match0 rev pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + let next _ := interp m in + let (pat, f) := p in + let (phyps, pconcl) := pat in + let cur _ := + let (hids, hctx, subst, cctx) := matches_goal rev phyps pconcl in + fun _ => f hids hctx subst cctx + in + Control.plus cur next + end in + Control.once (fun () => interp pats) (). + +Ltac2 multi_goal_match0 rev pats := + let rec interp m := match m with + | [] => Control.zero Match_failure + | p :: m => + let next _ := interp m in + let (pat, f) := p in + let (phyps, pconcl) := pat in + let cur _ := + let (hids, hctx, subst, cctx) := matches_goal rev phyps pconcl in + f hids hctx subst cctx + in + Control.plus cur next + end in + interp pats. + +Ltac2 one_goal_match0 rev pats := Control.once (fun _ => multi_goal_match0 rev pats). diff --git a/user-contrib/Ltac2/Std.v b/user-contrib/Ltac2/Std.v new file mode 100644 index 0000000000..6c3f465f33 --- /dev/null +++ b/user-contrib/Ltac2/Std.v @@ -0,0 +1,259 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Require Import Ltac2.Init. + +(** ML-facing types *) + +Ltac2 Type hypothesis := [ AnonHyp (int) | NamedHyp (ident) ]. + +Ltac2 Type bindings := [ +| NoBindings +| ImplicitBindings (constr list) +| ExplicitBindings ((hypothesis * constr) list) +]. + +Ltac2 Type constr_with_bindings := constr * bindings. + +Ltac2 Type occurrences := [ +| AllOccurrences +| AllOccurrencesBut (int list) +| NoOccurrences +| OnlyOccurrences (int list) +]. + +Ltac2 Type hyp_location_flag := [ InHyp | InHypTypeOnly | InHypValueOnly ]. + +Ltac2 Type clause := { + on_hyps : (ident * occurrences * hyp_location_flag) list option; + on_concl : occurrences; +}. + +Ltac2 Type reference := [ +| VarRef (ident) +| ConstRef (constant) +| IndRef (inductive) +| ConstructRef (constructor) +]. + +Ltac2 Type red_flags := { + rBeta : bool; + rMatch : bool; + rFix : bool; + rCofix : bool; + rZeta : bool; + rDelta : bool; (** true = delta all but rConst; false = delta only on rConst*) + rConst : reference list +}. + +Ltac2 Type 'a not_implemented. + +Ltac2 Type rec intro_pattern := [ +| IntroForthcoming (bool) +| IntroNaming (intro_pattern_naming) +| IntroAction (intro_pattern_action) +] +with intro_pattern_naming := [ +| IntroIdentifier (ident) +| IntroFresh (ident) +| IntroAnonymous +] +with intro_pattern_action := [ +| IntroWildcard +| IntroOrAndPattern (or_and_intro_pattern) +| IntroInjection (intro_pattern list) +| IntroApplyOn ((unit -> constr), intro_pattern) +| IntroRewrite (bool) +] +with or_and_intro_pattern := [ +| IntroOrPattern (intro_pattern list list) +| IntroAndPattern (intro_pattern list) +]. + +Ltac2 Type destruction_arg := [ +| ElimOnConstr (unit -> constr_with_bindings) +| ElimOnIdent (ident) +| ElimOnAnonHyp (int) +]. + +Ltac2 Type induction_clause := { + indcl_arg : destruction_arg; + indcl_eqn : intro_pattern_naming option; + indcl_as : or_and_intro_pattern option; + indcl_in : clause option; +}. + +Ltac2 Type assertion := [ +| AssertType (intro_pattern option, constr, (unit -> unit) option) +| AssertValue (ident, constr) +]. + +Ltac2 Type repeat := [ +| Precisely (int) +| UpTo (int) +| RepeatStar +| RepeatPlus +]. + +Ltac2 Type orientation := [ LTR | RTL ]. + +Ltac2 Type rewriting := { + rew_orient : orientation option; + rew_repeat : repeat; + rew_equatn : (unit -> constr_with_bindings); +}. + +Ltac2 Type evar_flag := bool. +Ltac2 Type advanced_flag := bool. + +Ltac2 Type move_location := [ +| MoveAfter (ident) +| MoveBefore (ident) +| MoveFirst +| MoveLast +]. + +Ltac2 Type inversion_kind := [ +| SimpleInversion +| FullInversion +| FullInversionClear +]. + +(** Standard, built-in tactics. See Ltac1 for documentation. *) + +Ltac2 @ external intros : evar_flag -> intro_pattern list -> unit := "ltac2" "tac_intros". + +Ltac2 @ external apply : advanced_flag -> evar_flag -> + (unit -> constr_with_bindings) list -> (ident * (intro_pattern option)) option -> unit := "ltac2" "tac_apply". + +Ltac2 @ external elim : evar_flag -> constr_with_bindings -> constr_with_bindings option -> unit := "ltac2" "tac_elim". +Ltac2 @ external case : evar_flag -> constr_with_bindings -> unit := "ltac2" "tac_case". + +Ltac2 @ external generalize : (constr * occurrences * ident option) list -> unit := "ltac2" "tac_generalize". + +Ltac2 @ external assert : assertion -> unit := "ltac2" "tac_assert". +Ltac2 @ external enough : constr -> (unit -> unit) option option -> intro_pattern option -> unit := "ltac2" "tac_enough". + +Ltac2 @ external pose : ident option -> constr -> unit := "ltac2" "tac_pose". +Ltac2 @ external set : evar_flag -> (unit -> ident option * constr) -> clause -> unit := "ltac2" "tac_set". + +Ltac2 @ external remember : evar_flag -> ident option -> (unit -> constr) -> intro_pattern option -> clause -> unit := "ltac2" "tac_remember". + +Ltac2 @ external destruct : evar_flag -> induction_clause list -> + constr_with_bindings option -> unit := "ltac2" "tac_induction". + +Ltac2 @ external induction : evar_flag -> induction_clause list -> + constr_with_bindings option -> unit := "ltac2" "tac_induction". + +Ltac2 @ external red : clause -> unit := "ltac2" "tac_red". +Ltac2 @ external hnf : clause -> unit := "ltac2" "tac_hnf". +Ltac2 @ external simpl : red_flags -> (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_simpl". +Ltac2 @ external cbv : red_flags -> clause -> unit := "ltac2" "tac_cbv". +Ltac2 @ external cbn : red_flags -> clause -> unit := "ltac2" "tac_cbn". +Ltac2 @ external lazy : red_flags -> clause -> unit := "ltac2" "tac_lazy". +Ltac2 @ external unfold : (reference * occurrences) list -> clause -> unit := "ltac2" "tac_unfold". +Ltac2 @ external fold : constr list -> clause -> unit := "ltac2" "tac_fold". +Ltac2 @ external pattern : (constr * occurrences) list -> clause -> unit := "ltac2" "tac_pattern". +Ltac2 @ external vm : (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_vm". +Ltac2 @ external native : (pattern * occurrences) option -> clause -> unit := "ltac2" "tac_native". + +Ltac2 @ external eval_red : constr -> constr := "ltac2" "eval_red". +Ltac2 @ external eval_hnf : constr -> constr := "ltac2" "eval_hnf". +Ltac2 @ external eval_red : constr -> constr := "ltac2" "eval_red". +Ltac2 @ external eval_simpl : red_flags -> (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_simpl". +Ltac2 @ external eval_cbv : red_flags -> constr -> constr := "ltac2" "eval_cbv". +Ltac2 @ external eval_cbn : red_flags -> constr -> constr := "ltac2" "eval_cbn". +Ltac2 @ external eval_lazy : red_flags -> constr -> constr := "ltac2" "eval_lazy". +Ltac2 @ external eval_unfold : (reference * occurrences) list -> constr -> constr := "ltac2" "eval_unfold". +Ltac2 @ external eval_fold : constr list -> constr -> constr := "ltac2" "eval_fold". +Ltac2 @ external eval_pattern : (constr * occurrences) list -> constr -> constr := "ltac2" "eval_pattern". +Ltac2 @ external eval_vm : (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_vm". +Ltac2 @ external eval_native : (pattern * occurrences) option -> constr -> constr := "ltac2" "eval_native". + +Ltac2 @ external change : pattern option -> (constr array -> constr) -> clause -> unit := "ltac2" "tac_change". + +Ltac2 @ external rewrite : evar_flag -> rewriting list -> clause -> (unit -> unit) option -> unit := "ltac2" "tac_rewrite". + +Ltac2 @ external reflexivity : unit -> unit := "ltac2" "tac_reflexivity". + +Ltac2 @ external assumption : unit -> unit := "ltac2" "tac_assumption". + +Ltac2 @ external transitivity : constr -> unit := "ltac2" "tac_transitivity". + +Ltac2 @ external etransitivity : unit -> unit := "ltac2" "tac_etransitivity". + +Ltac2 @ external cut : constr -> unit := "ltac2" "tac_cut". + +Ltac2 @ external left : evar_flag -> bindings -> unit := "ltac2" "tac_left". +Ltac2 @ external right : evar_flag -> bindings -> unit := "ltac2" "tac_right". + +Ltac2 @ external constructor : evar_flag -> unit := "ltac2" "tac_constructor". +Ltac2 @ external split : evar_flag -> bindings -> unit := "ltac2" "tac_split". + +Ltac2 @ external constructor_n : evar_flag -> int -> bindings -> unit := "ltac2" "tac_constructorn". + +Ltac2 @ external intros_until : hypothesis -> unit := "ltac2" "tac_introsuntil". + +Ltac2 @ external symmetry : clause -> unit := "ltac2" "tac_symmetry". + +Ltac2 @ external rename : (ident * ident) list -> unit := "ltac2" "tac_rename". + +Ltac2 @ external revert : ident list -> unit := "ltac2" "tac_revert". + +Ltac2 @ external admit : unit -> unit := "ltac2" "tac_admit". + +Ltac2 @ external fix_ : ident option -> int -> unit := "ltac2" "tac_fix". +Ltac2 @ external cofix_ : ident option -> unit := "ltac2" "tac_cofix". + +Ltac2 @ external clear : ident list -> unit := "ltac2" "tac_clear". +Ltac2 @ external keep : ident list -> unit := "ltac2" "tac_keep". + +Ltac2 @ external clearbody : ident list -> unit := "ltac2" "tac_clearbody". + +Ltac2 @ external exact_no_check : constr -> unit := "ltac2" "tac_exactnocheck". +Ltac2 @ external vm_cast_no_check : constr -> unit := "ltac2" "tac_vmcastnocheck". +Ltac2 @ external native_cast_no_check : constr -> unit := "ltac2" "tac_nativecastnocheck". + +Ltac2 @ external inversion : inversion_kind -> destruction_arg -> intro_pattern option -> ident list option -> unit := "ltac2" "tac_inversion". + +(** coretactics *) + +Ltac2 @ external move : ident -> move_location -> unit := "ltac2" "tac_move". + +Ltac2 @ external intro : ident option -> move_location option -> unit := "ltac2" "tac_intro". + +Ltac2 @ external specialize : constr_with_bindings -> intro_pattern option -> unit := "ltac2" "tac_specialize". + +(** extratactics *) + +Ltac2 @ external discriminate : evar_flag -> destruction_arg option -> unit := "ltac2" "tac_discriminate". +Ltac2 @ external injection : evar_flag -> intro_pattern list option -> destruction_arg option -> unit := "ltac2" "tac_injection". + +Ltac2 @ external absurd : constr -> unit := "ltac2" "tac_absurd". +Ltac2 @ external contradiction : constr_with_bindings option -> unit := "ltac2" "tac_contradiction". + +Ltac2 @ external autorewrite : bool -> (unit -> unit) option -> ident list -> clause -> unit := "ltac2" "tac_autorewrite". + +Ltac2 @ external subst : ident list -> unit := "ltac2" "tac_subst". +Ltac2 @ external subst_all : unit -> unit := "ltac2" "tac_substall". + +(** auto *) + +Ltac2 Type debug := [ Off | Info | Debug ]. + +Ltac2 Type strategy := [ BFS | DFS ]. + +Ltac2 @ external trivial : debug -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_trivial". + +Ltac2 @ external auto : debug -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_auto". + +Ltac2 @ external new_auto : debug -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_newauto". + +Ltac2 @ external eauto : debug -> int option -> int option -> (unit -> constr) list -> ident list option -> unit := "ltac2" "tac_eauto". + +Ltac2 @ external typeclasses_eauto : strategy option -> int option -> ident list option -> unit := "ltac2" "tac_typeclasses_eauto". diff --git a/user-contrib/Ltac2/String.v b/user-contrib/Ltac2/String.v new file mode 100644 index 0000000000..99e1dab76b --- /dev/null +++ b/user-contrib/Ltac2/String.v @@ -0,0 +1,14 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +Require Import Ltac2.Init. + +Ltac2 @external make : int -> char -> string := "ltac2" "string_make". +Ltac2 @external length : string -> int := "ltac2" "string_length". +Ltac2 @external get : string -> int -> char := "ltac2" "string_get". +Ltac2 @external set : string -> int -> char -> unit := "ltac2" "string_set". diff --git a/user-contrib/Ltac2/g_ltac2.mlg b/user-contrib/Ltac2/g_ltac2.mlg new file mode 100644 index 0000000000..890ed76d52 --- /dev/null +++ b/user-contrib/Ltac2/g_ltac2.mlg @@ -0,0 +1,933 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +{ + +open Pp +open Util +open Names +open Tok +open Pcoq +open Attributes +open Constrexpr +open Tac2expr +open Tac2qexpr +open Ltac_plugin + +let err () = raise Stream.Failure + +type lookahead = int -> Tok.t Stream.t -> int option + +let entry_of_lookahead s (lk : lookahead) = + let run strm = match lk 0 strm with None -> err () | Some _ -> () in + Pcoq.Entry.of_parser s run + +let (>>) (lk1 : lookahead) lk2 n strm = match lk1 n strm with +| None -> None +| Some n -> lk2 n strm + +let (<+>) (lk1 : lookahead) lk2 n strm = match lk1 n strm with +| None -> lk2 n strm +| Some n -> Some n + +let lk_kw kw n strm = match stream_nth n strm with +| KEYWORD kw' | IDENT kw' -> if String.equal kw kw' then Some (n + 1) else None +| _ -> None + +let lk_ident n strm = match stream_nth n strm with +| IDENT _ -> Some (n + 1) +| _ -> None + +let lk_int n strm = match stream_nth n strm with +| NUMERAL { NumTok.int = _; frac = ""; exp = "" } -> Some (n + 1) +| _ -> None + +let lk_ident_or_anti = lk_ident <+> (lk_kw "$" >> lk_ident) + +(* lookahead for (x:=t), (?x:=t) and (1:=t) *) +let test_lpar_idnum_coloneq = + entry_of_lookahead "test_lpar_idnum_coloneq" begin + lk_kw "(" >> (lk_ident_or_anti <+> lk_int) >> lk_kw ":=" + end + +(* lookahead for (x:t), (?x:t) *) +let test_lpar_id_colon = + entry_of_lookahead "test_lpar_id_colon" begin + lk_kw "(" >> lk_ident_or_anti >> lk_kw ":" + end + +(* Hack to recognize "(x := t)" and "($x := t)" *) +let test_lpar_id_coloneq = + entry_of_lookahead "test_lpar_id_coloneq" begin + lk_kw "(" >> lk_ident_or_anti >> lk_kw ":=" + end + +(* Hack to recognize "(x)" *) +let test_lpar_id_rpar = + entry_of_lookahead "test_lpar_id_rpar" begin + lk_kw "(" >> lk_ident >> lk_kw ")" + end + +let test_ampersand_ident = + entry_of_lookahead "test_ampersand_ident" begin + lk_kw "&" >> lk_ident + end + +let test_dollar_ident = + entry_of_lookahead "test_dollar_ident" begin + lk_kw "$" >> lk_ident + end + +let tac2expr = Tac2entries.Pltac.tac2expr +let tac2type = Entry.create "tactic:tac2type" +let tac2def_val = Entry.create "tactic:tac2def_val" +let tac2def_typ = Entry.create "tactic:tac2def_typ" +let tac2def_ext = Entry.create "tactic:tac2def_ext" +let tac2def_syn = Entry.create "tactic:tac2def_syn" +let tac2def_mut = Entry.create "tactic:tac2def_mut" +let tac2def_run = Entry.create "tactic:tac2def_run" +let tac2mode = Entry.create "vernac:ltac2_command" + +let ltac1_expr = Pltac.tactic_expr + +let inj_wit wit loc x = CAst.make ~loc @@ CTacExt (wit, x) +let inj_open_constr loc c = inj_wit Tac2quote.wit_open_constr loc c +let inj_pattern loc c = inj_wit Tac2quote.wit_pattern loc c +let inj_reference loc c = inj_wit Tac2quote.wit_reference loc c +let inj_ltac1 loc e = inj_wit Tac2quote.wit_ltac1 loc e +let inj_ltac1val loc e = inj_wit Tac2quote.wit_ltac1val loc e + +let pattern_of_qualid qid = + if Tac2env.is_constructor qid then CAst.make ?loc:qid.CAst.loc @@ CPatRef (RelId qid, []) + else + let open Libnames in + if qualid_is_ident qid then CAst.make ?loc:qid.CAst.loc @@ CPatVar (Name (qualid_basename qid)) + else + CErrors.user_err ?loc:qid.CAst.loc (Pp.str "Syntax error") + +} + +GRAMMAR EXTEND Gram + GLOBAL: tac2expr tac2type tac2def_val tac2def_typ tac2def_ext tac2def_syn + tac2def_mut tac2def_run; + tac2pat: + [ "1" LEFTA + [ qid = Prim.qualid; pl = LIST1 tac2pat LEVEL "0" -> { + if Tac2env.is_constructor qid then + CAst.make ~loc @@ CPatRef (RelId qid, pl) + else + CErrors.user_err ~loc (Pp.str "Syntax error") } + | qid = Prim.qualid -> { pattern_of_qualid qid } + | "["; "]" -> { CAst.make ~loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_nil), []) } + | p1 = tac2pat; "::"; p2 = tac2pat -> + { CAst.make ~loc @@ CPatRef (AbsKn (Other Tac2core.Core.c_cons), [p1; p2])} + ] + | "0" + [ "_" -> { CAst.make ~loc @@ CPatVar Anonymous } + | "()" -> { CAst.make ~loc @@ CPatRef (AbsKn (Tuple 0), []) } + | qid = Prim.qualid -> { pattern_of_qualid qid } + | "("; p = atomic_tac2pat; ")" -> { p } + ] ] + ; + atomic_tac2pat: + [ [ -> + { CAst.make ~loc @@ CPatRef (AbsKn (Tuple 0), []) } + | p = tac2pat; ":"; t = tac2type -> + { CAst.make ~loc @@ CPatCnv (p, t) } + | p = tac2pat; ","; pl = LIST0 tac2pat SEP "," -> + { let pl = p :: pl in + CAst.make ~loc @@ CPatRef (AbsKn (Tuple (List.length pl)), pl) } + | p = tac2pat -> { p } + ] ] + ; + tac2expr: + [ "6" RIGHTA + [ e1 = SELF; ";"; e2 = SELF -> { CAst.make ~loc @@ CTacSeq (e1, e2) } ] + | "5" + [ "fun"; it = LIST1 input_fun ; "=>"; body = tac2expr LEVEL "6" -> + { CAst.make ~loc @@ CTacFun (it, body) } + | "let"; isrec = rec_flag; + lc = LIST1 let_clause SEP "with"; "in"; + e = tac2expr LEVEL "6" -> + { CAst.make ~loc @@ CTacLet (isrec, lc, e) } + | "match"; e = tac2expr LEVEL "5"; "with"; bl = branches; "end" -> + { CAst.make ~loc @@ CTacCse (e, bl) } + ] + | "4" LEFTA [ ] + | "::" RIGHTA + [ e1 = tac2expr; "::"; e2 = tac2expr -> + { CAst.make ~loc @@ CTacApp (CAst.make ~loc @@ CTacCst (AbsKn (Other Tac2core.Core.c_cons)), [e1; e2]) } + ] + | [ e0 = SELF; ","; el = LIST1 NEXT SEP "," -> + { let el = e0 :: el in + CAst.make ~loc @@ CTacApp (CAst.make ~loc @@ CTacCst (AbsKn (Tuple (List.length el))), el) } ] + | "1" LEFTA + [ e = tac2expr; el = LIST1 tac2expr LEVEL "0" -> + { CAst.make ~loc @@ CTacApp (e, el) } + | e = SELF; ".("; qid = Prim.qualid; ")" -> + { CAst.make ~loc @@ CTacPrj (e, RelId qid) } + | e = SELF; ".("; qid = Prim.qualid; ")"; ":="; r = tac2expr LEVEL "5" -> + { CAst.make ~loc @@ CTacSet (e, RelId qid, r) } ] + | "0" + [ "("; a = SELF; ")" -> { a } + | "("; a = SELF; ":"; t = tac2type; ")" -> + { CAst.make ~loc @@ CTacCnv (a, t) } + | "()" -> + { CAst.make ~loc @@ CTacCst (AbsKn (Tuple 0)) } + | "("; ")" -> + { CAst.make ~loc @@ CTacCst (AbsKn (Tuple 0)) } + | "["; a = LIST0 tac2expr LEVEL "5" SEP ";"; "]" -> + { Tac2quote.of_list ~loc (fun x -> x) a } + | "{"; a = tac2rec_fieldexprs; "}" -> + { CAst.make ~loc @@ CTacRec a } + | a = tactic_atom -> { a } ] + ] + ; + branches: + [ [ -> { [] } + | "|"; bl = LIST1 branch SEP "|" -> { bl } + | bl = LIST1 branch SEP "|" -> { bl } ] + ] + ; + branch: + [ [ pat = tac2pat LEVEL "1"; "=>"; e = tac2expr LEVEL "6" -> { (pat, e) } ] ] + ; + rec_flag: + [ [ IDENT "rec" -> { true } + | -> { false } ] ] + ; + mut_flag: + [ [ IDENT "mutable" -> { true } + | -> { false } ] ] + ; + typ_param: + [ [ "'"; id = Prim.ident -> { id } ] ] + ; + tactic_atom: + [ [ n = Prim.integer -> { CAst.make ~loc @@ CTacAtm (AtmInt n) } + | s = Prim.string -> { CAst.make ~loc @@ CTacAtm (AtmStr s) } + | qid = Prim.qualid -> + { if Tac2env.is_constructor qid then + CAst.make ~loc @@ CTacCst (RelId qid) + else + CAst.make ~loc @@ CTacRef (RelId qid) } + | "@"; id = Prim.ident -> { Tac2quote.of_ident (CAst.make ~loc id) } + | "&"; id = lident -> { Tac2quote.of_hyp ~loc id } + | "'"; c = Constr.constr -> { inj_open_constr loc c } + | IDENT "constr"; ":"; "("; c = Constr.lconstr; ")" -> { Tac2quote.of_constr c } + | IDENT "open_constr"; ":"; "("; c = Constr.lconstr; ")" -> { Tac2quote.of_open_constr c } + | IDENT "ident"; ":"; "("; c = lident; ")" -> { Tac2quote.of_ident c } + | IDENT "pattern"; ":"; "("; c = Constr.lconstr_pattern; ")" -> { inj_pattern loc c } + | IDENT "reference"; ":"; "("; c = globref; ")" -> { inj_reference loc c } + | IDENT "ltac1"; ":"; "("; qid = ltac1_expr; ")" -> { inj_ltac1 loc qid } + | IDENT "ltac1val"; ":"; "("; qid = ltac1_expr; ")" -> { inj_ltac1val loc qid } + ] ] + ; + let_clause: + [ [ binder = let_binder; ":="; te = tac2expr -> + { let (pat, fn) = binder in + let te = match fn with + | None -> te + | Some args -> CAst.make ~loc @@ CTacFun (args, te) + in + (pat, te) } + ] ] + ; + let_binder: + [ [ pats = LIST1 input_fun -> + { match pats with + | [{CAst.v=CPatVar _} as pat] -> (pat, None) + | ({CAst.v=CPatVar (Name id)} as pat) :: args -> (pat, Some args) + | [pat] -> (pat, None) + | _ -> CErrors.user_err ~loc (str "Invalid pattern") } + ] ] + ; + tac2type: + [ "5" RIGHTA + [ t1 = tac2type; "->"; t2 = tac2type -> { CAst.make ~loc @@ CTypArrow (t1, t2) } ] + | "2" + [ t = tac2type; "*"; tl = LIST1 tac2type LEVEL "1" SEP "*" -> + { let tl = t :: tl in + CAst.make ~loc @@ CTypRef (AbsKn (Tuple (List.length tl)), tl) } ] + | "1" LEFTA + [ t = SELF; qid = Prim.qualid -> { CAst.make ~loc @@ CTypRef (RelId qid, [t]) } ] + | "0" + [ "("; t = tac2type LEVEL "5"; ")" -> { t } + | id = typ_param -> { CAst.make ~loc @@ CTypVar (Name id) } + | "_" -> { CAst.make ~loc @@ CTypVar Anonymous } + | qid = Prim.qualid -> { CAst.make ~loc @@ CTypRef (RelId qid, []) } + | "("; p = LIST1 tac2type LEVEL "5" SEP ","; ")"; qid = Prim.qualid -> + { CAst.make ~loc @@ CTypRef (RelId qid, p) } ] + ]; + locident: + [ [ id = Prim.ident -> { CAst.make ~loc id } ] ] + ; + binder: + [ [ "_" -> { CAst.make ~loc Anonymous } + | l = Prim.ident -> { CAst.make ~loc (Name l) } ] ] + ; + input_fun: + [ [ b = tac2pat LEVEL "0" -> { b } ] ] + ; + tac2def_body: + [ [ name = binder; it = LIST0 input_fun; ":="; e = tac2expr -> + { let e = if List.is_empty it then e else CAst.make ~loc @@ CTacFun (it, e) in + (name, e) } + ] ] + ; + tac2def_val: + [ [ mut = mut_flag; isrec = rec_flag; l = LIST1 tac2def_body SEP "with" -> + { StrVal (mut, isrec, l) } + ] ] + ; + tac2def_mut: + [ [ "Set"; qid = Prim.qualid; ":="; e = tac2expr -> { StrMut (qid, e) } ] ] + ; + tac2def_run: + [ [ "Eval"; e = tac2expr -> { StrRun e } ] ] + ; + tac2typ_knd: + [ [ t = tac2type -> { CTydDef (Some t) } + | "["; ".."; "]" -> { CTydOpn } + | "["; t = tac2alg_constructors; "]" -> { CTydAlg t } + | "{"; t = tac2rec_fields; "}"-> { CTydRec t } ] ] + ; + tac2alg_constructors: + [ [ "|"; cs = LIST1 tac2alg_constructor SEP "|" -> { cs } + | cs = LIST0 tac2alg_constructor SEP "|" -> { cs } ] ] + ; + tac2alg_constructor: + [ [ c = Prim.ident -> { (c, []) } + | c = Prim.ident; "("; args = LIST0 tac2type SEP ","; ")"-> { (c, args) } ] ] + ; + tac2rec_fields: + [ [ f = tac2rec_field; ";"; l = tac2rec_fields -> { f :: l } + | f = tac2rec_field; ";" -> { [f] } + | f = tac2rec_field -> { [f] } + | -> { [] } ] ] + ; + tac2rec_field: + [ [ mut = mut_flag; id = Prim.ident; ":"; t = tac2type -> { (id, mut, t) } ] ] + ; + tac2rec_fieldexprs: + [ [ f = tac2rec_fieldexpr; ";"; l = tac2rec_fieldexprs -> { f :: l } + | f = tac2rec_fieldexpr; ";" -> { [f] } + | f = tac2rec_fieldexpr-> { [f] } + | -> { [] } ] ] + ; + tac2rec_fieldexpr: + [ [ qid = Prim.qualid; ":="; e = tac2expr LEVEL "1" -> { RelId qid, e } ] ] + ; + tac2typ_prm: + [ [ -> { [] } + | id = typ_param -> { [CAst.make ~loc id] } + | "("; ids = LIST1 [ id = typ_param -> { CAst.make ~loc id } ] SEP "," ;")" -> { ids } + ] ] + ; + tac2typ_def: + [ [ prm = tac2typ_prm; id = Prim.qualid; b = tac2type_body -> { let (r, e) = b in (id, r, (prm, e)) } ] ] + ; + tac2type_body: + [ [ -> { false, CTydDef None } + | ":="; e = tac2typ_knd -> { false, e } + | "::="; e = tac2typ_knd -> { true, e } + ] ] + ; + tac2def_typ: + [ [ "Type"; isrec = rec_flag; l = LIST1 tac2typ_def SEP "with" -> + { StrTyp (isrec, l) } + ] ] + ; + tac2def_ext: + [ [ "@"; IDENT "external"; id = locident; ":"; t = tac2type LEVEL "5"; ":="; + plugin = Prim.string; name = Prim.string -> + { let ml = { mltac_plugin = plugin; mltac_tactic = name } in + StrPrm (id, t, ml) } + ] ] + ; + syn_node: + [ [ "_" -> { CAst.make ~loc None } + | id = Prim.ident -> { CAst.make ~loc (Some id) } + ] ] + ; + sexpr: + [ [ s = Prim.string -> { SexprStr (CAst.make ~loc s) } + | n = Prim.integer -> { SexprInt (CAst.make ~loc n) } + | id = syn_node -> { SexprRec (loc, id, []) } + | id = syn_node; "("; tok = LIST1 sexpr SEP "," ; ")" -> + { SexprRec (loc, id, tok) } + ] ] + ; + syn_level: + [ [ -> { None } + | ":"; n = Prim.integer -> { Some n } + ] ] + ; + tac2def_syn: + [ [ "Notation"; toks = LIST1 sexpr; n = syn_level; ":="; + e = tac2expr -> + { StrSyn (toks, n, e) } + ] ] + ; + lident: + [ [ id = Prim.ident -> { CAst.make ~loc id } ] ] + ; + globref: + [ [ "&"; id = Prim.ident -> { CAst.make ~loc (QHypothesis id) } + | qid = Prim.qualid -> { CAst.make ~loc @@ QReference qid } + ] ] + ; +END + +(* Quotation scopes used by notations *) + +{ + +open Tac2entries.Pltac + +let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) + +} + +GRAMMAR EXTEND Gram + GLOBAL: q_ident q_bindings q_intropattern q_intropatterns q_induction_clause + q_conversion q_rewriting q_clause q_dispatch q_occurrences q_strategy_flag + q_destruction_arg q_reference q_with_bindings q_constr_matching + q_goal_matching q_hintdb q_move_location q_pose q_assert; + anti: + [ [ "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } ] ] + ; + ident_or_anti: + [ [ id = lident -> { QExpr id } + | "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } + ] ] + ; + lident: + [ [ id = Prim.ident -> { CAst.make ~loc id } ] ] + ; + lnatural: + [ [ n = Prim.natural -> { CAst.make ~loc n } ] ] + ; + q_ident: + [ [ id = ident_or_anti -> { id } ] ] + ; + qhyp: + [ [ x = anti -> { x } + | n = lnatural -> { QExpr (CAst.make ~loc @@ QAnonHyp n) } + | id = lident -> { QExpr (CAst.make ~loc @@ QNamedHyp id) } + ] ] + ; + simple_binding: + [ [ "("; h = qhyp; ":="; c = Constr.lconstr; ")" -> + { CAst.make ~loc (h, c) } + ] ] + ; + bindings: + [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> + { CAst.make ~loc @@ QExplicitBindings bl } + | bl = LIST1 Constr.constr -> + { CAst.make ~loc @@ QImplicitBindings bl } + ] ] + ; + q_bindings: + [ [ bl = bindings -> { bl } ] ] + ; + q_with_bindings: + [ [ bl = with_bindings -> { bl } ] ] + ; + intropatterns: + [ [ l = LIST0 nonsimple_intropattern -> { CAst.make ~loc l } ] ] + ; +(* ne_intropatterns: *) +(* [ [ l = LIST1 nonsimple_intropattern -> l ]] *) +(* ; *) + or_and_intropattern: + [ [ "["; tc = LIST1 intropatterns SEP "|"; "]" -> { CAst.make ~loc @@ QIntroOrPattern tc } + | "()" -> { CAst.make ~loc @@ QIntroAndPattern (CAst.make ~loc []) } + | "("; si = simple_intropattern; ")" -> { CAst.make ~loc @@ QIntroAndPattern (CAst.make ~loc [si]) } + | "("; si = simple_intropattern; ","; + tc = LIST1 simple_intropattern SEP "," ; ")" -> + { CAst.make ~loc @@ QIntroAndPattern (CAst.make ~loc (si::tc)) } + | "("; si = simple_intropattern; "&"; + tc = LIST1 simple_intropattern SEP "&" ; ")" -> + (* (A & B & C) is translated into (A,(B,C)) *) + { let rec pairify = function + | ([]|[_]|[_;_]) as l -> CAst.make ~loc l + | t::q -> + let q = + CAst.make ~loc @@ + QIntroAction (CAst.make ~loc @@ + QIntroOrAndPattern (CAst.make ~loc @@ + QIntroAndPattern (pairify q))) + in + CAst.make ~loc [t; q] + in CAst.make ~loc @@ QIntroAndPattern (pairify (si::tc)) } ] ] + ; + equality_intropattern: + [ [ "->" -> { CAst.make ~loc @@ QIntroRewrite true } + | "<-" -> { CAst.make ~loc @@ QIntroRewrite false } + | "[="; tc = intropatterns; "]" -> { CAst.make ~loc @@ QIntroInjection tc } ] ] + ; + naming_intropattern: + [ [ LEFTQMARK; id = lident -> + { CAst.make ~loc @@ QIntroFresh (QExpr id) } + | "?$"; id = lident -> + { CAst.make ~loc @@ QIntroFresh (QAnti id) } + | "?" -> + { CAst.make ~loc @@ QIntroAnonymous } + | id = ident_or_anti -> + { CAst.make ~loc @@ QIntroIdentifier id } + ] ] + ; + nonsimple_intropattern: + [ [ l = simple_intropattern -> { l } + | "*" -> { CAst.make ~loc @@ QIntroForthcoming true } + | "**" -> { CAst.make ~loc @@ QIntroForthcoming false } ] ] + ; + simple_intropattern: + [ [ pat = simple_intropattern_closed -> +(* l = LIST0 ["%"; c = operconstr LEVEL "0" -> c] -> *) + (** TODO: handle %pat *) + { pat } + ] ] + ; + simple_intropattern_closed: + [ [ pat = or_and_intropattern -> + { CAst.make ~loc @@ QIntroAction (CAst.make ~loc @@ QIntroOrAndPattern pat) } + | pat = equality_intropattern -> + { CAst.make ~loc @@ QIntroAction pat } + | "_" -> + { CAst.make ~loc @@ QIntroAction (CAst.make ~loc @@ QIntroWildcard) } + | pat = naming_intropattern -> + { CAst.make ~loc @@ QIntroNaming pat } + ] ] + ; + q_intropatterns: + [ [ ipat = intropatterns -> { ipat } ] ] + ; + q_intropattern: + [ [ ipat = simple_intropattern -> { ipat } ] ] + ; + nat_or_anti: + [ [ n = lnatural -> { QExpr n } + | "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } + ] ] + ; + eqn_ipat: + [ [ IDENT "eqn"; ":"; pat = naming_intropattern -> { Some pat } + | -> { None } + ] ] + ; + with_bindings: + [ [ "with"; bl = bindings -> { bl } | -> { CAst.make ~loc @@ QNoBindings } ] ] + ; + constr_with_bindings: + [ [ c = Constr.constr; l = with_bindings -> { CAst.make ~loc @@ (c, l) } ] ] + ; + destruction_arg: + [ [ n = lnatural -> { CAst.make ~loc @@ QElimOnAnonHyp n } + | id = lident -> { CAst.make ~loc @@ QElimOnIdent id } + | c = constr_with_bindings -> { CAst.make ~loc @@ QElimOnConstr c } + ] ] + ; + q_destruction_arg: + [ [ arg = destruction_arg -> { arg } ] ] + ; + as_or_and_ipat: + [ [ "as"; ipat = or_and_intropattern -> { Some ipat } + | -> { None } + ] ] + ; + occs_nums: + [ [ nl = LIST1 nat_or_anti -> { CAst.make ~loc @@ QOnlyOccurrences nl } + | "-"; n = nat_or_anti; nl = LIST0 nat_or_anti -> + { CAst.make ~loc @@ QAllOccurrencesBut (n::nl) } + ] ] + ; + occs: + [ [ "at"; occs = occs_nums -> { occs } | -> { CAst.make ~loc QAllOccurrences } ] ] + ; + hypident: + [ [ id = ident_or_anti -> + { id,Locus.InHyp } + | "("; IDENT "type"; IDENT "of"; id = ident_or_anti; ")" -> + { id,Locus.InHypTypeOnly } + | "("; IDENT "value"; IDENT "of"; id = ident_or_anti; ")" -> + { id,Locus.InHypValueOnly } + ] ] + ; + hypident_occ: + [ [ h=hypident; occs=occs -> { let (id,l) = h in ((occs,id),l) } ] ] + ; + in_clause: + [ [ "*"; occs=occs -> + { { q_onhyps = None; q_concl_occs = occs } } + | "*"; "|-"; occs = concl_occ -> + { { q_onhyps = None; q_concl_occs = occs } } + | hl = LIST0 hypident_occ SEP ","; "|-"; occs = concl_occ -> + { { q_onhyps = Some hl; q_concl_occs = occs } } + | hl = LIST0 hypident_occ SEP "," -> + { { q_onhyps = Some hl; q_concl_occs = CAst.make ~loc QNoOccurrences } } + ] ] + ; + clause: + [ [ "in"; cl = in_clause -> { CAst.make ~loc @@ cl } + | "at"; occs = occs_nums -> + { CAst.make ~loc @@ { q_onhyps = Some []; q_concl_occs = occs } } + ] ] + ; + q_clause: + [ [ cl = clause -> { cl } ] ] + ; + concl_occ: + [ [ "*"; occs = occs -> { occs } + | -> { CAst.make ~loc QNoOccurrences } + ] ] + ; + induction_clause: + [ [ c = destruction_arg; pat = as_or_and_ipat; eq = eqn_ipat; + cl = OPT clause -> + { CAst.make ~loc @@ { + indcl_arg = c; + indcl_eqn = eq; + indcl_as = pat; + indcl_in = cl; + } } + ] ] + ; + q_induction_clause: + [ [ cl = induction_clause -> { cl } ] ] + ; + conversion: + [ [ c = Constr.constr -> + { CAst.make ~loc @@ QConvert c } + | c1 = Constr.constr; "with"; c2 = Constr.constr -> + { CAst.make ~loc @@ QConvertWith (c1, c2) } + ] ] + ; + q_conversion: + [ [ c = conversion -> { c } ] ] + ; + orient: + [ [ "->" -> { CAst.make ~loc (Some true) } + | "<-" -> { CAst.make ~loc (Some false) } + | -> { CAst.make ~loc None } + ]] + ; + rewriter: + [ [ "!"; c = constr_with_bindings -> + { (CAst.make ~loc @@ QRepeatPlus,c) } + | [ "?" -> { () } | LEFTQMARK -> { () } ]; c = constr_with_bindings -> + { (CAst.make ~loc @@ QRepeatStar,c) } + | n = lnatural; "!"; c = constr_with_bindings -> + { (CAst.make ~loc @@ QPrecisely n,c) } + | n = lnatural; ["?" -> { () } | LEFTQMARK -> { () } ]; c = constr_with_bindings -> + { (CAst.make ~loc @@ QUpTo n,c) } + | n = lnatural; c = constr_with_bindings -> + { (CAst.make ~loc @@ QPrecisely n,c) } + | c = constr_with_bindings -> + { (CAst.make ~loc @@ QPrecisely (CAst.make 1), c) } + ] ] + ; + oriented_rewriter: + [ [ b = orient; r = rewriter -> + { let (m, c) = r in + CAst.make ~loc @@ { + rew_orient = b; + rew_repeat = m; + rew_equatn = c; + } } + ] ] + ; + q_rewriting: + [ [ r = oriented_rewriter -> { r } ] ] + ; + tactic_then_last: + [ [ "|"; lta = LIST0 (OPT tac2expr LEVEL "6") SEP "|" -> { lta } + | -> { [] } + ] ] + ; + tactic_then_gen: + [ [ ta = tac2expr; "|"; tg = tactic_then_gen -> { let (first,last) = tg in (Some ta :: first, last) } + | ta = tac2expr; ".."; l = tactic_then_last -> { ([], Some (Some ta, l)) } + | ".."; l = tactic_then_last -> { ([], Some (None, l)) } + | ta = tac2expr -> { ([Some ta], None) } + | "|"; tg = tactic_then_gen -> { let (first,last) = tg in (None :: first, last) } + | -> { ([None], None) } + ] ] + ; + q_dispatch: + [ [ d = tactic_then_gen -> { CAst.make ~loc d } ] ] + ; + q_occurrences: + [ [ occs = occs -> { occs } ] ] + ; + red_flag: + [ [ IDENT "beta" -> { CAst.make ~loc @@ QBeta } + | IDENT "iota" -> { CAst.make ~loc @@ QIota } + | IDENT "match" -> { CAst.make ~loc @@ QMatch } + | IDENT "fix" -> { CAst.make ~loc @@ QFix } + | IDENT "cofix" -> { CAst.make ~loc @@ QCofix } + | IDENT "zeta" -> { CAst.make ~loc @@ QZeta } + | IDENT "delta"; d = delta_flag -> { d } + ] ] + ; + refglobal: + [ [ "&"; id = Prim.ident -> { QExpr (CAst.make ~loc @@ QHypothesis id) } + | qid = Prim.qualid -> { QExpr (CAst.make ~loc @@ QReference qid) } + | "$"; id = Prim.ident -> { QAnti (CAst.make ~loc id) } + ] ] + ; + q_reference: + [ [ r = refglobal -> { r } ] ] + ; + refglobals: + [ [ gl = LIST1 refglobal -> { CAst.make ~loc gl } ] ] + ; + delta_flag: + [ [ "-"; "["; idl = refglobals; "]" -> { CAst.make ~loc @@ QDeltaBut idl } + | "["; idl = refglobals; "]" -> { CAst.make ~loc @@ QConst idl } + | -> { CAst.make ~loc @@ QDeltaBut (CAst.make ~loc []) } + ] ] + ; + strategy_flag: + [ [ s = LIST1 red_flag -> { CAst.make ~loc s } + | d = delta_flag -> + { CAst.make ~loc + [CAst.make ~loc QBeta; CAst.make ~loc QIota; CAst.make ~loc QZeta; d] } + ] ] + ; + q_strategy_flag: + [ [ flag = strategy_flag -> { flag } ] ] + ; + hintdb: + [ [ "*" -> { CAst.make ~loc @@ QHintAll } + | l = LIST1 ident_or_anti -> { CAst.make ~loc @@ QHintDbs l } + ] ] + ; + q_hintdb: + [ [ db = hintdb -> { db } ] ] + ; + match_pattern: + [ [ IDENT "context"; id = OPT Prim.ident; + "["; pat = Constr.lconstr_pattern; "]" -> { CAst.make ~loc @@ QConstrMatchContext (id, pat) } + | pat = Constr.lconstr_pattern -> { CAst.make ~loc @@ QConstrMatchPattern pat } ] ] + ; + match_rule: + [ [ mp = match_pattern; "=>"; tac = tac2expr -> + { CAst.make ~loc @@ (mp, tac) } + ] ] + ; + match_list: + [ [ mrl = LIST1 match_rule SEP "|" -> { CAst.make ~loc @@ mrl } + | "|"; mrl = LIST1 match_rule SEP "|" -> { CAst.make ~loc @@ mrl } ] ] + ; + q_constr_matching: + [ [ m = match_list -> { m } ] ] + ; + gmatch_hyp_pattern: + [ [ na = Prim.name; ":"; pat = match_pattern -> { (na, pat) } ] ] + ; + gmatch_pattern: + [ [ "["; hl = LIST0 gmatch_hyp_pattern SEP ","; "|-"; p = match_pattern; "]" -> + { CAst.make ~loc @@ { + q_goal_match_concl = p; + q_goal_match_hyps = hl; + } } + ] ] + ; + gmatch_rule: + [ [ mp = gmatch_pattern; "=>"; tac = tac2expr -> + { CAst.make ~loc @@ (mp, tac) } + ] ] + ; + gmatch_list: + [ [ mrl = LIST1 gmatch_rule SEP "|" -> { CAst.make ~loc @@ mrl } + | "|"; mrl = LIST1 gmatch_rule SEP "|" -> { CAst.make ~loc @@ mrl } ] ] + ; + q_goal_matching: + [ [ m = gmatch_list -> { m } ] ] + ; + move_location: + [ [ "at"; IDENT "top" -> { CAst.make ~loc @@ QMoveFirst } + | "at"; IDENT "bottom" -> { CAst.make ~loc @@ QMoveLast } + | IDENT "after"; id = ident_or_anti -> { CAst.make ~loc @@ QMoveAfter id } + | IDENT "before"; id = ident_or_anti -> { CAst.make ~loc @@ QMoveBefore id } + ] ] + ; + q_move_location: + [ [ mv = move_location -> { mv } ] ] + ; + as_name: + [ [ -> { None } + | "as"; id = ident_or_anti -> { Some id } + ] ] + ; + pose: + [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> + { CAst.make ~loc (Some id, c) } + | c = Constr.constr; na = as_name -> { CAst.make ~loc (na, c) } + ] ] + ; + q_pose: + [ [ p = pose -> { p } ] ] + ; + as_ipat: + [ [ "as"; ipat = simple_intropattern -> { Some ipat } + | -> { None } + ] ] + ; + by_tactic: + [ [ "by"; tac = tac2expr -> { Some tac } + | -> { None } + ] ] + ; + assertion: + [ [ test_lpar_id_coloneq; "("; id = ident_or_anti; ":="; c = Constr.lconstr; ")" -> + { CAst.make ~loc (QAssertValue (id, c)) } + | test_lpar_id_colon; "("; id = ident_or_anti; ":"; c = Constr.lconstr; ")"; tac = by_tactic -> + { let ipat = CAst.make ~loc @@ QIntroNaming (CAst.make ~loc @@ QIntroIdentifier id) in + CAst.make ~loc (QAssertType (Some ipat, c, tac)) } + | c = Constr.constr; ipat = as_ipat; tac = by_tactic -> + { CAst.make ~loc (QAssertType (ipat, c, tac)) } + ] ] + ; + q_assert: + [ [ a = assertion -> { a } ] ] + ; +END + +(** Extension of constr syntax *) + +(* +GRAMMAR EXTEND Gram + Pcoq.Constr.operconstr: LEVEL "0" + [ [ IDENT "ltac2"; ":"; "("; tac = tac2expr; ")" -> + { let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) } + | test_ampersand_ident; "&"; id = Prim.ident -> + { let tac = Tac2quote.of_exact_hyp ~loc (CAst.make ~loc id) in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) } + | test_dollar_ident; "$"; id = Prim.ident -> + { let id = Loc.tag ~loc id in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2_quotation) id in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) } + ] ] + ; +END +*) +{ + +let () = + +let open Extend in +let open Tok in +let (++) r s = Next (r, s) in +let rules = [ + Rule ( + Stop ++ Aentry test_dollar_ident ++ Atoken (PKEYWORD "$") ++ Aentry Prim.ident, + begin fun id _ _ loc -> + let id = Loc.tag ~loc id in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2_quotation) id in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) + end + ); + + Rule ( + Stop ++ Aentry test_ampersand_ident ++ Atoken (PKEYWORD "&") ++ Aentry Prim.ident, + begin fun id _ _ loc -> + let tac = Tac2quote.of_exact_hyp ~loc (CAst.make ~loc id) in + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) + end + ); + + Rule ( + Stop ++ Atoken (PIDENT (Some "ltac2")) ++ Atoken (PKEYWORD ":") ++ + Atoken (PKEYWORD "(") ++ Aentry tac2expr ++ Atoken (PKEYWORD ")"), + begin fun _ tac _ _ _ loc -> + let arg = Genarg.in_gen (Genarg.rawwit Tac2env.wit_ltac2) tac in + CAst.make ~loc (CHole (None, Namegen.IntroAnonymous, Some arg)) + end + ) +] in + +Hook.set Tac2entries.register_constr_quotations begin fun () -> + Pcoq.grammar_extend Pcoq.Constr.operconstr None (Some (Gramlib.Gramext.Level "0"), [(None, None, rules)]) +end + +} + +{ + +let pr_ltac2entry _ = mt () (* FIXME *) +let pr_ltac2expr _ = mt () (* FIXME *) + +} + +VERNAC ARGUMENT EXTEND ltac2_entry +PRINTED BY { pr_ltac2entry } +| [ tac2def_val(v) ] -> { v } +| [ tac2def_typ(t) ] -> { t } +| [ tac2def_ext(e) ] -> { e } +| [ tac2def_syn(e) ] -> { e } +| [ tac2def_mut(e) ] -> { e } +| [ tac2def_run(e) ] -> { e } +END + +{ + +let classify_ltac2 = function +| StrSyn _ -> Vernacextend.(VtSideff [], VtNow) +| StrMut _ | StrVal _ | StrPrm _ | StrTyp _ | StrRun _ -> Vernacextend.classify_as_sideeff + +} + +VERNAC COMMAND EXTEND VernacDeclareTactic2Definition +| #[ local = locality ] ![proof] [ "Ltac2" ltac2_entry(e) ] => { classify_ltac2 e } -> { + fun ~pstate -> Tac2entries.register_struct ?local ~pstate e; pstate + } +END + +{ + +let _ = Pvernac.register_proof_mode "Ltac2" tac2mode + +} + +VERNAC ARGUMENT EXTEND ltac2_expr +PRINTED BY { pr_ltac2expr } +| [ tac2expr(e) ] -> { e } +END + +{ + +open G_ltac +open Vernacextend + +} + +VERNAC { tac2mode } EXTEND VernacLtac2 +| ![proof] [ ltac2_expr(t) ltac_use_default(default) ] => + { classify_as_proofstep } -> { +(* let g = Option.default (Proof_global.get_default_goal_selector ()) g in *) + fun ~pstate -> + Option.map (fun pstate -> Tac2entries.call ~pstate ~default t) pstate + } +END + +{ + +open Stdarg + +} + +VERNAC COMMAND EXTEND Ltac2Print CLASSIFIED AS SIDEFF +| [ "Print" "Ltac2" reference(tac) ] -> { Tac2entries.print_ltac tac } +END diff --git a/user-contrib/Ltac2/ltac2_plugin.mlpack b/user-contrib/Ltac2/ltac2_plugin.mlpack new file mode 100644 index 0000000000..2a25e825cb --- /dev/null +++ b/user-contrib/Ltac2/ltac2_plugin.mlpack @@ -0,0 +1,14 @@ +Tac2dyn +Tac2ffi +Tac2env +Tac2print +Tac2intern +Tac2interp +Tac2entries +Tac2quote +Tac2match +Tac2core +Tac2extffi +Tac2tactics +Tac2stdlib +G_ltac2 diff --git a/user-contrib/Ltac2/plugin_base.dune b/user-contrib/Ltac2/plugin_base.dune new file mode 100644 index 0000000000..711e9b95d3 --- /dev/null +++ b/user-contrib/Ltac2/plugin_base.dune @@ -0,0 +1,6 @@ +(library + (name ltac2_plugin) + (public_name coq.plugins.ltac2) + (synopsis "Coq's Ltac2 plugin") + (modules_without_implementation tac2expr tac2qexpr tac2types) + (libraries coq.plugins.ltac)) diff --git a/user-contrib/Ltac2/tac2core.ml b/user-contrib/Ltac2/tac2core.ml new file mode 100644 index 0000000000..d7e7b91ee6 --- /dev/null +++ b/user-contrib/Ltac2/tac2core.ml @@ -0,0 +1,1446 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Util +open Pp +open Names +open Genarg +open Tac2env +open Tac2expr +open Tac2entries.Pltac +open Proofview.Notations + +(** Standard values *) + +module Value = Tac2ffi +open Value + +let core_prefix path n = KerName.make path (Label.of_id (Id.of_string_soft n)) + +let std_core n = core_prefix Tac2env.std_prefix n +let coq_core n = core_prefix Tac2env.coq_prefix n +let ltac1_core n = core_prefix Tac2env.ltac1_prefix n + +module Core = +struct + +let t_int = coq_core "int" +let t_string = coq_core "string" +let t_array = coq_core "array" +let t_unit = coq_core "unit" +let t_list = coq_core "list" +let t_constr = coq_core "constr" +let t_pattern = coq_core "pattern" +let t_ident = coq_core "ident" +let t_option = coq_core "option" +let t_exn = coq_core "exn" +let t_reference = std_core "reference" +let t_ltac1 = ltac1_core "t" + +let c_nil = coq_core "[]" +let c_cons = coq_core "::" + +let c_none = coq_core "None" +let c_some = coq_core "Some" + +let c_true = coq_core "true" +let c_false = coq_core "false" + +end + +open Core + +let v_unit = Value.of_unit () +let v_blk = Valexpr.make_block + +let of_name c = match c with +| Anonymous -> Value.of_option Value.of_ident None +| Name id -> Value.of_option Value.of_ident (Some id) + +let to_name c = match Value.to_option Value.to_ident c with +| None -> Anonymous +| Some id -> Name id + +let of_relevance = function + | Sorts.Relevant -> ValInt 0 + | Sorts.Irrelevant -> ValInt 1 + +let to_relevance = function + | ValInt 0 -> Sorts.Relevant + | ValInt 1 -> Sorts.Irrelevant + | _ -> assert false + +let of_annot f Context.{binder_name;binder_relevance} = + of_tuple [|(f binder_name); of_relevance binder_relevance|] + +let to_annot f x = + match to_tuple x with + | [|x;y|] -> + let x = f x in + let y = to_relevance y in + Context.make_annot x y + | _ -> assert false + +let of_instance u = + let u = Univ.Instance.to_array (EConstr.Unsafe.to_instance u) in + Value.of_array (fun v -> Value.of_ext Value.val_univ v) u + +let to_instance u = + let u = Value.to_array (fun v -> Value.to_ext Value.val_univ v) u in + EConstr.EInstance.make (Univ.Instance.of_array u) + +let of_rec_declaration (nas, ts, cs) = + (Value.of_array (of_annot of_name) nas, + Value.of_array Value.of_constr ts, + Value.of_array Value.of_constr cs) + +let to_rec_declaration (nas, ts, cs) = + (Value.to_array (to_annot to_name) nas, + Value.to_array Value.to_constr ts, + Value.to_array Value.to_constr cs) + +let of_result f = function +| Inl c -> v_blk 0 [|f c|] +| Inr e -> v_blk 1 [|Value.of_exn e|] + +(** Stdlib exceptions *) + +let err_notfocussed = + Tac2interp.LtacError (coq_core "Not_focussed", [||]) + +let err_outofbounds = + Tac2interp.LtacError (coq_core "Out_of_bounds", [||]) + +let err_notfound = + Tac2interp.LtacError (coq_core "Not_found", [||]) + +let err_matchfailure = + Tac2interp.LtacError (coq_core "Match_failure", [||]) + +(** Helper functions *) + +let thaw f = Tac2ffi.apply f [v_unit] + +let fatal_flag : unit Exninfo.t = Exninfo.make () + +let set_bt info = + if !Tac2interp.print_ltac2_backtrace then + Tac2interp.get_backtrace >>= fun bt -> + Proofview.tclUNIT (Exninfo.add info Tac2entries.backtrace bt) + else Proofview.tclUNIT info + +let throw ?(info = Exninfo.null) e = + set_bt info >>= fun info -> + let info = Exninfo.add info fatal_flag () in + Proofview.tclLIFT (Proofview.NonLogical.raise ~info e) + +let fail ?(info = Exninfo.null) e = + set_bt info >>= fun info -> + Proofview.tclZERO ~info e + +let return x = Proofview.tclUNIT x +let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } + +let wrap f = + return () >>= fun () -> return (f ()) + +let wrap_unit f = + return () >>= fun () -> f (); return v_unit + +let assert_focussed = + Proofview.Goal.goals >>= fun gls -> + match gls with + | [_] -> Proofview.tclUNIT () + | [] | _ :: _ :: _ -> throw err_notfocussed + +let pf_apply f = + Proofview.Goal.goals >>= function + | [] -> + Proofview.tclENV >>= fun env -> + Proofview.tclEVARMAP >>= fun sigma -> + f env sigma + | [gl] -> + gl >>= fun gl -> + f (Proofview.Goal.env gl) (Tacmach.New.project gl) + | _ :: _ :: _ -> + throw err_notfocussed + +(** Primitives *) + +let define_primitive name arity f = + Tac2env.define_primitive (pname name) (mk_closure arity f) + +let define0 name f = define_primitive name arity_one (fun _ -> f) + +let define1 name r0 f = define_primitive name arity_one begin fun x -> + f (Value.repr_to r0 x) +end + +let define2 name r0 r1 f = define_primitive name (arity_suc arity_one) begin fun x y -> + f (Value.repr_to r0 x) (Value.repr_to r1 y) +end + +let define3 name r0 r1 r2 f = define_primitive name (arity_suc (arity_suc arity_one)) begin fun x y z -> + f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z) +end + +(** Printing *) + +let () = define1 "print" pp begin fun pp -> + wrap_unit (fun () -> Feedback.msg_notice pp) +end + +let () = define1 "message_of_int" int begin fun n -> + return (Value.of_pp (Pp.int n)) +end + +let () = define1 "message_of_string" string begin fun s -> + return (Value.of_pp (str (Bytes.to_string s))) +end + +let () = define1 "message_of_constr" constr begin fun c -> + pf_apply begin fun env sigma -> + let pp = Printer.pr_econstr_env env sigma c in + return (Value.of_pp pp) + end +end + +let () = define1 "message_of_ident" ident begin fun c -> + let pp = Id.print c in + return (Value.of_pp pp) +end + +let () = define1 "message_of_exn" valexpr begin fun v -> + Proofview.tclENV >>= fun env -> + Proofview.tclEVARMAP >>= fun sigma -> + let pp = Tac2print.pr_valexpr env sigma v (GTypRef (Other Core.t_exn, [])) in + return (Value.of_pp pp) +end + + +let () = define2 "message_concat" pp pp begin fun m1 m2 -> + return (Value.of_pp (Pp.app m1 m2)) +end + +(** Array *) + +let () = define2 "array_make" int valexpr begin fun n x -> + if n < 0 || n > Sys.max_array_length then throw err_outofbounds + else wrap (fun () -> v_blk 0 (Array.make n x)) +end + +let () = define1 "array_length" block begin fun (_, v) -> + return (Value.of_int (Array.length v)) +end + +let () = define3 "array_set" block int valexpr begin fun (_, v) n x -> + if n < 0 || n >= Array.length v then throw err_outofbounds + else wrap_unit (fun () -> v.(n) <- x) +end + +let () = define2 "array_get" block int begin fun (_, v) n -> + if n < 0 || n >= Array.length v then throw err_outofbounds + else wrap (fun () -> v.(n)) +end + +(** Ident *) + +let () = define2 "ident_equal" ident ident begin fun id1 id2 -> + return (Value.of_bool (Id.equal id1 id2)) +end + +let () = define1 "ident_to_string" ident begin fun id -> + return (Value.of_string (Bytes.of_string (Id.to_string id))) +end + +let () = define1 "ident_of_string" string begin fun s -> + let id = try Some (Id.of_string (Bytes.to_string s)) with _ -> None in + return (Value.of_option Value.of_ident id) +end + +(** Int *) + +let () = define2 "int_equal" int int begin fun m n -> + return (Value.of_bool (m == n)) +end + +let binop n f = define2 n int int begin fun m n -> + return (Value.of_int (f m n)) +end + +let () = binop "int_compare" Int.compare +let () = binop "int_add" (+) +let () = binop "int_sub" (-) +let () = binop "int_mul" ( * ) + +let () = define1 "int_neg" int begin fun m -> + return (Value.of_int (~- m)) +end + +(** Char *) + +let () = define1 "char_of_int" int begin fun n -> + wrap (fun () -> Value.of_char (Char.chr n)) +end + +let () = define1 "char_to_int" char begin fun n -> + wrap (fun () -> Value.of_int (Char.code n)) +end + +(** String *) + +let () = define2 "string_make" int char begin fun n c -> + if n < 0 || n > Sys.max_string_length then throw err_outofbounds + else wrap (fun () -> Value.of_string (Bytes.make n c)) +end + +let () = define1 "string_length" string begin fun s -> + return (Value.of_int (Bytes.length s)) +end + +let () = define3 "string_set" string int char begin fun s n c -> + if n < 0 || n >= Bytes.length s then throw err_outofbounds + else wrap_unit (fun () -> Bytes.set s n c) +end + +let () = define2 "string_get" string int begin fun s n -> + if n < 0 || n >= Bytes.length s then throw err_outofbounds + else wrap (fun () -> Value.of_char (Bytes.get s n)) +end + +(** Terms *) + +(** constr -> constr *) +let () = define1 "constr_type" constr begin fun c -> + let get_type env sigma = + Proofview.V82.wrap_exceptions begin fun () -> + let (sigma, t) = Typing.type_of env sigma c in + let t = Value.of_constr t in + Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclUNIT t + end in + pf_apply get_type +end + +(** constr -> constr *) +let () = define2 "constr_equal" constr constr begin fun c1 c2 -> + Proofview.tclEVARMAP >>= fun sigma -> + let b = EConstr.eq_constr sigma c1 c2 in + Proofview.tclUNIT (Value.of_bool b) +end + +let () = define1 "constr_kind" constr begin fun c -> + let open Constr in + Proofview.tclEVARMAP >>= fun sigma -> + return begin match EConstr.kind sigma c with + | Rel n -> + v_blk 0 [|Value.of_int n|] + | Var id -> + v_blk 1 [|Value.of_ident id|] + | Meta n -> + v_blk 2 [|Value.of_int n|] + | Evar (evk, args) -> + v_blk 3 [| + Value.of_int (Evar.repr evk); + Value.of_array Value.of_constr args; + |] + | Sort s -> + v_blk 4 [|Value.of_ext Value.val_sort s|] + | Cast (c, k, t) -> + v_blk 5 [| + Value.of_constr c; + Value.of_ext Value.val_cast k; + Value.of_constr t; + |] + | Prod (na, t, u) -> + v_blk 6 [| + of_annot of_name na; + Value.of_constr t; + Value.of_constr u; + |] + | Lambda (na, t, c) -> + v_blk 7 [| + of_annot of_name na; + Value.of_constr t; + Value.of_constr c; + |] + | LetIn (na, b, t, c) -> + v_blk 8 [| + of_annot of_name na; + Value.of_constr b; + Value.of_constr t; + Value.of_constr c; + |] + | App (c, cl) -> + v_blk 9 [| + Value.of_constr c; + Value.of_array Value.of_constr cl; + |] + | Const (cst, u) -> + v_blk 10 [| + Value.of_constant cst; + of_instance u; + |] + | Ind (ind, u) -> + v_blk 11 [| + Value.of_ext Value.val_inductive ind; + of_instance u; + |] + | Construct (cstr, u) -> + v_blk 12 [| + Value.of_ext Value.val_constructor cstr; + of_instance u; + |] + | Case (ci, c, t, bl) -> + v_blk 13 [| + Value.of_ext Value.val_case ci; + Value.of_constr c; + Value.of_constr t; + Value.of_array Value.of_constr bl; + |] + | Fix ((recs, i), def) -> + let (nas, ts, cs) = of_rec_declaration def in + v_blk 14 [| + Value.of_array Value.of_int recs; + Value.of_int i; + nas; + ts; + cs; + |] + | CoFix (i, def) -> + let (nas, ts, cs) = of_rec_declaration def in + v_blk 15 [| + Value.of_int i; + nas; + ts; + cs; + |] + | Proj (p, c) -> + v_blk 16 [| + Value.of_ext Value.val_projection p; + Value.of_constr c; + |] + | Int _ -> + assert false + end +end + +let () = define1 "constr_make" valexpr begin fun knd -> + let c = match Tac2ffi.to_block knd with + | (0, [|n|]) -> + let n = Value.to_int n in + EConstr.mkRel n + | (1, [|id|]) -> + let id = Value.to_ident id in + EConstr.mkVar id + | (2, [|n|]) -> + let n = Value.to_int n in + EConstr.mkMeta n + | (3, [|evk; args|]) -> + let evk = Evar.unsafe_of_int (Value.to_int evk) in + let args = Value.to_array Value.to_constr args in + EConstr.mkEvar (evk, args) + | (4, [|s|]) -> + let s = Value.to_ext Value.val_sort s in + EConstr.mkSort (EConstr.Unsafe.to_sorts s) + | (5, [|c; k; t|]) -> + let c = Value.to_constr c in + let k = Value.to_ext Value.val_cast k in + let t = Value.to_constr t in + EConstr.mkCast (c, k, t) + | (6, [|na; t; u|]) -> + let na = to_annot to_name na in + let t = Value.to_constr t in + let u = Value.to_constr u in + EConstr.mkProd (na, t, u) + | (7, [|na; t; c|]) -> + let na = to_annot to_name na in + let t = Value.to_constr t in + let u = Value.to_constr c in + EConstr.mkLambda (na, t, u) + | (8, [|na; b; t; c|]) -> + let na = to_annot to_name na in + let b = Value.to_constr b in + let t = Value.to_constr t in + let c = Value.to_constr c in + EConstr.mkLetIn (na, b, t, c) + | (9, [|c; cl|]) -> + let c = Value.to_constr c in + let cl = Value.to_array Value.to_constr cl in + EConstr.mkApp (c, cl) + | (10, [|cst; u|]) -> + let cst = Value.to_constant cst in + let u = to_instance u in + EConstr.mkConstU (cst, u) + | (11, [|ind; u|]) -> + let ind = Value.to_ext Value.val_inductive ind in + let u = to_instance u in + EConstr.mkIndU (ind, u) + | (12, [|cstr; u|]) -> + let cstr = Value.to_ext Value.val_constructor cstr in + let u = to_instance u in + EConstr.mkConstructU (cstr, u) + | (13, [|ci; c; t; bl|]) -> + let ci = Value.to_ext Value.val_case ci in + let c = Value.to_constr c in + let t = Value.to_constr t in + let bl = Value.to_array Value.to_constr bl in + EConstr.mkCase (ci, c, t, bl) + | (14, [|recs; i; nas; ts; cs|]) -> + let recs = Value.to_array Value.to_int recs in + let i = Value.to_int i in + let def = to_rec_declaration (nas, ts, cs) in + EConstr.mkFix ((recs, i), def) + | (15, [|i; nas; ts; cs|]) -> + let i = Value.to_int i in + let def = to_rec_declaration (nas, ts, cs) in + EConstr.mkCoFix (i, def) + | (16, [|p; c|]) -> + let p = Value.to_ext Value.val_projection p in + let c = Value.to_constr c in + EConstr.mkProj (p, c) + | _ -> assert false + in + return (Value.of_constr c) +end + +let () = define1 "constr_check" constr begin fun c -> + pf_apply begin fun env sigma -> + try + let (sigma, _) = Typing.type_of env sigma c in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + return (of_result Value.of_constr (Inl c)) + with e when CErrors.noncritical e -> + let e = CErrors.push e in + return (of_result Value.of_constr (Inr e)) + end +end + +let () = define3 "constr_substnl" (list constr) int constr begin fun subst k c -> + let ans = EConstr.Vars.substnl subst k c in + return (Value.of_constr ans) +end + +let () = define3 "constr_closenl" (list ident) int constr begin fun ids k c -> + let ans = EConstr.Vars.substn_vars k ids c in + return (Value.of_constr ans) +end + +let () = define1 "constr_case" (repr_ext val_inductive) begin fun ind -> + Proofview.tclENV >>= fun env -> + try + let ans = Inductiveops.make_case_info env ind Sorts.Relevant Constr.RegularStyle in + return (Value.of_ext Value.val_case ans) + with e when CErrors.noncritical e -> + throw err_notfound +end + +let () = define2 "constr_constructor" (repr_ext val_inductive) int begin fun (ind, i) k -> + Proofview.tclENV >>= fun env -> + try + let open Declarations in + let ans = Environ.lookup_mind ind env in + let _ = ans.mind_packets.(i).mind_consnames.(k) in + return (Value.of_ext val_constructor ((ind, i), (k + 1))) + with e when CErrors.noncritical e -> + throw err_notfound +end + +let () = define3 "constr_in_context" ident constr closure begin fun id t c -> + Proofview.Goal.goals >>= function + | [gl] -> + gl >>= fun gl -> + let env = Proofview.Goal.env gl in + let sigma = Proofview.Goal.sigma gl in + let has_var = + try + let _ = Environ.lookup_named_val id env in + true + with Not_found -> false + in + if has_var then + Tacticals.New.tclZEROMSG (str "Variable already exists") + else + let open Context.Named.Declaration in + let nenv = EConstr.push_named (LocalAssum (Context.make_annot id Sorts.Relevant, t)) env in + let (sigma, (evt, _)) = Evarutil.new_type_evar nenv sigma Evd.univ_flexible in + let (sigma, evk) = Evarutil.new_pure_evar (Environ.named_context_val nenv) sigma evt in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + Proofview.Unsafe.tclSETGOALS [Proofview.with_empty_state evk] >>= fun () -> + thaw c >>= fun _ -> + Proofview.Unsafe.tclSETGOALS [Proofview.with_empty_state (Proofview.Goal.goal gl)] >>= fun () -> + let args = List.map (fun d -> EConstr.mkVar (get_id d)) (EConstr.named_context env) in + let args = Array.of_list (EConstr.mkRel 1 :: args) in + let ans = EConstr.mkEvar (evk, args) in + let ans = EConstr.mkLambda (Context.make_annot (Name id) Sorts.Relevant, t, ans) in + return (Value.of_constr ans) + | _ -> + throw err_notfocussed +end + +(** Patterns *) + +let empty_context = EConstr.mkMeta Constr_matching.special_meta + +let () = define0 "pattern_empty_context" begin + return (Value.of_constr empty_context) +end + +let () = define2 "pattern_matches" pattern constr begin fun pat c -> + pf_apply begin fun env sigma -> + let ans = + try Some (Constr_matching.matches env sigma pat c) + with Constr_matching.PatternMatchingFailure -> None + in + begin match ans with + | None -> fail err_matchfailure + | Some ans -> + let ans = Id.Map.bindings ans in + let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in + return (Value.of_list of_pair ans) + end + end +end + +let () = define2 "pattern_matches_subterm" pattern constr begin fun pat c -> + let open Constr_matching in + let rec of_ans s = match IStream.peek s with + | IStream.Nil -> fail err_matchfailure + | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> + let ans = Id.Map.bindings sub in + let of_pair (id, c) = Value.of_tuple [| Value.of_ident id; Value.of_constr c |] in + let ans = Value.of_tuple [| Value.of_constr (Lazy.force m_ctx); Value.of_list of_pair ans |] in + Proofview.tclOR (return ans) (fun _ -> of_ans s) + in + pf_apply begin fun env sigma -> + let ans = Constr_matching.match_subterm env sigma (Id.Set.empty,pat) c in + of_ans ans + end +end + +let () = define2 "pattern_matches_vect" pattern constr begin fun pat c -> + pf_apply begin fun env sigma -> + let ans = + try Some (Constr_matching.matches env sigma pat c) + with Constr_matching.PatternMatchingFailure -> None + in + begin match ans with + | None -> fail err_matchfailure + | Some ans -> + let ans = Id.Map.bindings ans in + let ans = Array.map_of_list snd ans in + return (Value.of_array Value.of_constr ans) + end + end +end + +let () = define2 "pattern_matches_subterm_vect" pattern constr begin fun pat c -> + let open Constr_matching in + let rec of_ans s = match IStream.peek s with + | IStream.Nil -> fail err_matchfailure + | IStream.Cons ({ m_sub = (_, sub); m_ctx }, s) -> + let ans = Id.Map.bindings sub in + let ans = Array.map_of_list snd ans in + let ans = Value.of_tuple [| Value.of_constr (Lazy.force m_ctx); Value.of_array Value.of_constr ans |] in + Proofview.tclOR (return ans) (fun _ -> of_ans s) + in + pf_apply begin fun env sigma -> + let ans = Constr_matching.match_subterm env sigma (Id.Set.empty,pat) c in + of_ans ans + end +end + +let () = define3 "pattern_matches_goal" bool (list (pair bool pattern)) (pair bool pattern) begin fun rev hp cp -> + assert_focussed >>= fun () -> + Proofview.Goal.enter_one begin fun gl -> + let env = Proofview.Goal.env gl in + let sigma = Proofview.Goal.sigma gl in + let concl = Proofview.Goal.concl gl in + let mk_pattern (b, pat) = if b then Tac2match.MatchPattern pat else Tac2match.MatchContext pat in + let r = (List.map mk_pattern hp, mk_pattern cp) in + Tac2match.match_goal env sigma concl ~rev r >>= fun (hyps, ctx, subst) -> + let of_ctxopt ctx = Value.of_constr (Option.default empty_context ctx) in + let hids = Value.of_array Value.of_ident (Array.map_of_list fst hyps) in + let hctx = Value.of_array of_ctxopt (Array.map_of_list snd hyps) in + let subs = Value.of_array Value.of_constr (Array.map_of_list snd (Id.Map.bindings subst)) in + let cctx = of_ctxopt ctx in + let ans = Value.of_tuple [| hids; hctx; subs; cctx |] in + Proofview.tclUNIT ans + end +end + +let () = define2 "pattern_instantiate" constr constr begin fun ctx c -> + let ctx = EConstr.Unsafe.to_constr ctx in + let c = EConstr.Unsafe.to_constr c in + let ans = Termops.subst_meta [Constr_matching.special_meta, c] ctx in + return (Value.of_constr (EConstr.of_constr ans)) +end + +(** Error *) + +let () = define1 "throw" exn begin fun (e, info) -> + throw ~info e +end + +(** Control *) + +(** exn -> 'a *) +let () = define1 "zero" exn begin fun (e, info) -> + fail ~info e +end + +(** (unit -> 'a) -> (exn -> 'a) -> 'a *) +let () = define2 "plus" closure closure begin fun x k -> + Proofview.tclOR (thaw x) (fun e -> Tac2ffi.apply k [Value.of_exn e]) +end + +(** (unit -> 'a) -> 'a *) +let () = define1 "once" closure begin fun f -> + Proofview.tclONCE (thaw f) +end + +(** (unit -> unit) list -> unit *) +let () = define1 "dispatch" (list closure) begin fun l -> + let l = List.map (fun f -> Proofview.tclIGNORE (thaw f)) l in + Proofview.tclDISPATCH l >>= fun () -> return v_unit +end + +(** (unit -> unit) list -> (unit -> unit) -> (unit -> unit) list -> unit *) +let () = define3 "extend" (list closure) closure (list closure) begin fun lft tac rgt -> + let lft = List.map (fun f -> Proofview.tclIGNORE (thaw f)) lft in + let tac = Proofview.tclIGNORE (thaw tac) in + let rgt = List.map (fun f -> Proofview.tclIGNORE (thaw f)) rgt in + Proofview.tclEXTEND lft tac rgt >>= fun () -> return v_unit +end + +(** (unit -> unit) -> unit *) +let () = define1 "enter" closure begin fun f -> + let f = Proofview.tclIGNORE (thaw f) in + Proofview.tclINDEPENDENT f >>= fun () -> return v_unit +end + +(** (unit -> 'a) -> ('a * ('exn -> 'a)) result *) +let () = define1 "case" closure begin fun f -> + Proofview.tclCASE (thaw f) >>= begin function + | Proofview.Next (x, k) -> + let k = Tac2ffi.mk_closure arity_one begin fun e -> + let (e, info) = Value.to_exn e in + set_bt info >>= fun info -> + k (e, info) + end in + return (v_blk 0 [| Value.of_tuple [| x; Value.of_closure k |] |]) + | Proofview.Fail e -> return (v_blk 1 [| Value.of_exn e |]) + end +end + +(** int -> int -> (unit -> 'a) -> 'a *) +let () = define3 "focus" int int closure begin fun i j tac -> + Proofview.tclFOCUS i j (thaw tac) +end + +(** unit -> unit *) +let () = define0 "shelve" begin + Proofview.shelve >>= fun () -> return v_unit +end + +(** unit -> unit *) +let () = define0 "shelve_unifiable" begin + Proofview.shelve_unifiable >>= fun () -> return v_unit +end + +let () = define1 "new_goal" int begin fun ev -> + let ev = Evar.unsafe_of_int ev in + Proofview.tclEVARMAP >>= fun sigma -> + if Evd.mem sigma ev then + Proofview.Unsafe.tclNEWGOALS [Proofview.with_empty_state ev] <*> Proofview.tclUNIT v_unit + else throw err_notfound +end + +(** unit -> constr *) +let () = define0 "goal" begin + assert_focussed >>= fun () -> + Proofview.Goal.enter_one begin fun gl -> + let concl = Tacmach.New.pf_nf_concl gl in + return (Value.of_constr concl) + end +end + +(** ident -> constr *) +let () = define1 "hyp" ident begin fun id -> + pf_apply begin fun env _ -> + let mem = try ignore (Environ.lookup_named id env); true with Not_found -> false in + if mem then return (Value.of_constr (EConstr.mkVar id)) + else Tacticals.New.tclZEROMSG + (str "Hypothesis " ++ quote (Id.print id) ++ str " not found") (* FIXME: Do something more sensible *) + end +end + +let () = define0 "hyps" begin + pf_apply begin fun env _ -> + let open Context in + let open Named.Declaration in + let hyps = List.rev (Environ.named_context env) in + let map = function + | LocalAssum (id, t) -> + let t = EConstr.of_constr t in + Value.of_tuple [|Value.of_ident id.binder_name; Value.of_option Value.of_constr None; Value.of_constr t|] + | LocalDef (id, c, t) -> + let c = EConstr.of_constr c in + let t = EConstr.of_constr t in + Value.of_tuple [|Value.of_ident id.binder_name; Value.of_option Value.of_constr (Some c); Value.of_constr t|] + in + return (Value.of_list map hyps) + end +end + +(** (unit -> constr) -> unit *) +let () = define1 "refine" closure begin fun c -> + let c = thaw c >>= fun c -> Proofview.tclUNIT ((), Value.to_constr c) in + Proofview.Goal.enter begin fun gl -> + Refine.generic_refine ~typecheck:true c gl + end >>= fun () -> return v_unit +end + +let () = define2 "with_holes" closure closure begin fun x f -> + Proofview.tclEVARMAP >>= fun sigma0 -> + thaw x >>= fun ans -> + Proofview.tclEVARMAP >>= fun sigma -> + Proofview.Unsafe.tclEVARS sigma0 >>= fun () -> + Tacticals.New.tclWITHHOLES false (Tac2ffi.apply f [ans]) sigma +end + +let () = define1 "progress" closure begin fun f -> + Proofview.tclPROGRESS (thaw f) +end + +let () = define2 "abstract" (option ident) closure begin fun id f -> + Abstract.tclABSTRACT id (Proofview.tclIGNORE (thaw f)) >>= fun () -> + return v_unit +end + +let () = define2 "time" (option string) closure begin fun s f -> + let s = Option.map Bytes.to_string s in + Proofview.tclTIME s (thaw f) +end + +let () = define0 "check_interrupt" begin + Proofview.tclCHECKINTERRUPT >>= fun () -> return v_unit +end + +(** Fresh *) + +let () = define2 "fresh_free_union" (repr_ext val_free) (repr_ext val_free) begin fun set1 set2 -> + let ans = Id.Set.union set1 set2 in + return (Value.of_ext Value.val_free ans) +end + +let () = define1 "fresh_free_of_ids" (list ident) begin fun ids -> + let free = List.fold_right Id.Set.add ids Id.Set.empty in + return (Value.of_ext Value.val_free free) +end + +let () = define1 "fresh_free_of_constr" constr begin fun c -> + Proofview.tclEVARMAP >>= fun sigma -> + let rec fold accu c = match EConstr.kind sigma c with + | Constr.Var id -> Id.Set.add id accu + | _ -> EConstr.fold sigma fold accu c + in + let ans = fold Id.Set.empty c in + return (Value.of_ext Value.val_free ans) +end + +let () = define2 "fresh_fresh" (repr_ext val_free) ident begin fun avoid id -> + let nid = Namegen.next_ident_away_from id (fun id -> Id.Set.mem id avoid) in + return (Value.of_ident nid) +end + +(** Env *) + +let () = define1 "env_get" (list ident) begin fun ids -> + let r = match ids with + | [] -> None + | _ :: _ as ids -> + let (id, path) = List.sep_last ids in + let path = DirPath.make (List.rev path) in + let fp = Libnames.make_path path id in + try Some (Nametab.global_of_path fp) with Not_found -> None + in + return (Value.of_option Value.of_reference r) +end + +let () = define1 "env_expand" (list ident) begin fun ids -> + let r = match ids with + | [] -> [] + | _ :: _ as ids -> + let (id, path) = List.sep_last ids in + let path = DirPath.make (List.rev path) in + let qid = Libnames.make_qualid path id in + Nametab.locate_all qid + in + return (Value.of_list Value.of_reference r) +end + +let () = define1 "env_path" reference begin fun r -> + match Nametab.path_of_global r with + | fp -> + let (path, id) = Libnames.repr_path fp in + let path = DirPath.repr path in + return (Value.of_list Value.of_ident (List.rev_append path [id])) + | exception Not_found -> + throw err_notfound +end + +let () = define1 "env_instantiate" reference begin fun r -> + Proofview.tclENV >>= fun env -> + Proofview.tclEVARMAP >>= fun sigma -> + let (sigma, c) = Evd.fresh_global env sigma r in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + return (Value.of_constr c) +end + +(** Ltac1 in Ltac2 *) + +let ltac1 = Tac2ffi.repr_ext Value.val_ltac1 +let of_ltac1 v = Value.of_ext Value.val_ltac1 v + +let () = define1 "ltac1_ref" (list ident) begin fun ids -> + let open Ltac_plugin in + let r = match ids with + | [] -> raise Not_found + | _ :: _ as ids -> + let (id, path) = List.sep_last ids in + let path = DirPath.make (List.rev path) in + let fp = Libnames.make_path path id in + if Tacenv.exists_tactic fp then + List.hd (Tacenv.locate_extended_all_tactic (Libnames.qualid_of_path fp)) + else raise Not_found + in + let tac = Tacinterp.Value.of_closure (Tacinterp.default_ist ()) (Tacenv.interp_ltac r) in + return (Value.of_ext val_ltac1 tac) +end + +let () = define1 "ltac1_run" ltac1 begin fun v -> + let open Ltac_plugin in + Tacinterp.tactic_of_value (Tacinterp.default_ist ()) v >>= fun () -> + return v_unit +end + +let () = define3 "ltac1_apply" ltac1 (list ltac1) closure begin fun f args k -> + let open Ltac_plugin in + let open Tacexpr in + let open Locus in + let k ret = + Proofview.tclIGNORE (Tac2ffi.apply k [Value.of_ext val_ltac1 ret]) + in + let fold arg (i, vars, lfun) = + let id = Id.of_string ("x" ^ string_of_int i) in + let x = Reference (ArgVar CAst.(make id)) in + (succ i, x :: vars, Id.Map.add id arg lfun) + in + let (_, args, lfun) = List.fold_right fold args (0, [], Id.Map.empty) in + let lfun = Id.Map.add (Id.of_string "F") f lfun in + let ist = { (Tacinterp.default_ist ()) with Tacinterp.lfun = lfun; } in + let tac = TacArg(CAst.make @@ TacCall (CAst.make (ArgVar CAst.(make @@ Id.of_string "F"),args))) in + Tacinterp.val_interp ist tac k >>= fun () -> + return v_unit +end + +let () = define1 "ltac1_of_constr" constr begin fun c -> + let open Ltac_plugin in + return (Value.of_ext val_ltac1 (Tacinterp.Value.of_constr c)) +end + +let () = define1 "ltac1_to_constr" ltac1 begin fun v -> + let open Ltac_plugin in + return (Value.of_option Value.of_constr (Tacinterp.Value.to_constr v)) +end + +let () = define1 "ltac1_of_list" (list ltac1) begin fun l -> + let open Geninterp.Val in + return (Value.of_ext val_ltac1 (inject (Base typ_list) l)) +end + +let () = define1 "ltac1_to_list" ltac1 begin fun v -> + let open Ltac_plugin in + return (Value.of_option (Value.of_list of_ltac1) (Tacinterp.Value.to_list v)) +end + +(** ML types *) + +let constr_flags () = + let open Pretyping in + { + use_typeclasses = true; + solve_unification_constraints = true; + fail_evar = true; + expand_evars = true; + program_mode = false; + polymorphic = false; + } + +let open_constr_no_classes_flags () = + let open Pretyping in + { + use_typeclasses = false; + solve_unification_constraints = true; + fail_evar = false; + expand_evars = true; + program_mode = false; + polymorphic = false; + } + +(** Embed all Ltac2 data into Values *) +let to_lvar ist = + let open Glob_ops in + let lfun = Tac2interp.set_env ist Id.Map.empty in + { empty_lvar with Ltac_pretype.ltac_genargs = lfun } + +let gtypref kn = GTypRef (Other kn, []) + +let intern_constr self ist c = + let (_, (c, _)) = Genintern.intern Stdarg.wit_constr ist c in + (GlbVal c, gtypref t_constr) + +let catchable_exception = function + | Logic_monad.Exception _ -> false + | e -> CErrors.noncritical e + +let interp_constr flags ist c = + let open Pretyping in + let ist = to_lvar ist in + pf_apply begin fun env sigma -> + try + let (sigma, c) = understand_ltac flags env sigma ist WithoutTypeConstraint c in + let c = Value.of_constr c in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + Proofview.tclUNIT c + with e when catchable_exception e -> + let (e, info) = CErrors.push e in + set_bt info >>= fun info -> + match Exninfo.get info fatal_flag with + | None -> Proofview.tclZERO ~info e + | Some () -> throw ~info e + end + +let () = + let intern = intern_constr in + let interp ist c = interp_constr (constr_flags ()) ist c in + let print env c = str "constr:(" ++ Printer.pr_lglob_constr_env env c ++ str ")" in + let subst subst c = Detyping.subst_glob_constr (Global.env()) subst c in + let obj = { + ml_intern = intern; + ml_subst = subst; + ml_interp = interp; + ml_print = print; + } in + define_ml_object Tac2quote.wit_constr obj + +let () = + let intern = intern_constr in + let interp ist c = interp_constr (open_constr_no_classes_flags ()) ist c in + let print env c = str "open_constr:(" ++ Printer.pr_lglob_constr_env env c ++ str ")" in + let subst subst c = Detyping.subst_glob_constr (Global.env()) subst c in + let obj = { + ml_intern = intern; + ml_subst = subst; + ml_interp = interp; + ml_print = print; + } in + define_ml_object Tac2quote.wit_open_constr obj + +let () = + let interp _ id = return (Value.of_ident id) in + let print _ id = str "ident:(" ++ Id.print id ++ str ")" in + let obj = { + ml_intern = (fun _ _ id -> GlbVal id, gtypref t_ident); + ml_interp = interp; + ml_subst = (fun _ id -> id); + ml_print = print; + } in + define_ml_object Tac2quote.wit_ident obj + +let () = + let intern self ist c = + let env = ist.Genintern.genv in + let sigma = Evd.from_env env in + let warn = if !Ltac_plugin.Tacintern.strict_check then fun x -> x else Constrintern.for_grammar in + let _, pat = warn (fun () ->Constrintern.intern_constr_pattern env sigma ~as_type:false c) () in + GlbVal pat, gtypref t_pattern + in + let subst subst c = + let env = Global.env () in + let sigma = Evd.from_env env in + Patternops.subst_pattern env sigma subst c + in + let print env pat = str "pattern:(" ++ Printer.pr_lconstr_pattern_env env Evd.empty pat ++ str ")" in + let interp _ c = return (Value.of_pattern c) in + let obj = { + ml_intern = intern; + ml_interp = interp; + ml_subst = subst; + ml_print = print; + } in + define_ml_object Tac2quote.wit_pattern obj + +let () = + let intern self ist ref = match ref.CAst.v with + | Tac2qexpr.QHypothesis id -> + GlbVal (Globnames.VarRef id), gtypref t_reference + | Tac2qexpr.QReference qid -> + let gr = + try Nametab.locate qid + with Not_found -> + Nametab.error_global_not_found qid + in + GlbVal gr, gtypref t_reference + in + let subst s c = Globnames.subst_global_reference s c in + let interp _ gr = return (Value.of_reference gr) in + let print _ = function + | Globnames.VarRef id -> str "reference:(" ++ str "&" ++ Id.print id ++ str ")" + | r -> str "reference:(" ++ Printer.pr_global r ++ str ")" + in + let obj = { + ml_intern = intern; + ml_subst = subst; + ml_interp = interp; + ml_print = print; + } in + define_ml_object Tac2quote.wit_reference obj + +let () = + let intern self ist tac = + (* Prevent inner calls to Ltac2 values *) + let extra = Tac2intern.drop_ltac2_env ist.Genintern.extra in + let ist = { ist with Genintern.extra } in + let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in + GlbVal tac, gtypref t_unit + in + let interp ist tac = + let ist = { env_ist = Id.Map.empty } in + let lfun = Tac2interp.set_env ist Id.Map.empty in + let ist = Ltac_plugin.Tacinterp.default_ist () in + let ist = { ist with Geninterp.lfun = lfun } in + let tac = (Ltac_plugin.Tacinterp.eval_tactic_ist ist tac : unit Proofview.tactic) in + let wrap (e, info) = set_bt info >>= fun info -> Proofview.tclZERO ~info e in + Proofview.tclOR tac wrap >>= fun () -> + return v_unit + in + let subst s tac = Genintern.substitute Ltac_plugin.Tacarg.wit_tactic s tac in + let print env tac = + str "ltac1:(" ++ Ltac_plugin.Pptactic.pr_glob_tactic env tac ++ str ")" + in + let obj = { + ml_intern = intern; + ml_subst = subst; + ml_interp = interp; + ml_print = print; + } in + define_ml_object Tac2quote.wit_ltac1 obj + +let () = + let open Ltac_plugin in + let intern self ist tac = + (* Prevent inner calls to Ltac2 values *) + let extra = Tac2intern.drop_ltac2_env ist.Genintern.extra in + let ist = { ist with Genintern.extra } in + let _, tac = Genintern.intern Ltac_plugin.Tacarg.wit_tactic ist tac in + GlbVal tac, gtypref t_ltac1 + in + let interp ist tac = + let ist = { env_ist = Id.Map.empty } in + let lfun = Tac2interp.set_env ist Id.Map.empty in + let ist = Ltac_plugin.Tacinterp.default_ist () in + let ist = { ist with Geninterp.lfun = lfun } in + return (Value.of_ext val_ltac1 (Tacinterp.Value.of_closure ist tac)) + in + let subst s tac = Genintern.substitute Tacarg.wit_tactic s tac in + let print env tac = + str "ltac1val:(" ++ Ltac_plugin.Pptactic.pr_glob_tactic env tac ++ str ")" + in + let obj = { + ml_intern = intern; + ml_subst = subst; + ml_interp = interp; + ml_print = print; + } in + define_ml_object Tac2quote.wit_ltac1val obj + +(** Ltac2 in terms *) + +let () = + let interp ist poly env sigma concl tac = + let ist = Tac2interp.get_env ist in + let tac = Proofview.tclIGNORE (Tac2interp.interp ist tac) in + let name, poly = Id.of_string "ltac2", poly in + let c, sigma = Pfedit.refine_by_tactic ~name ~poly env sigma concl tac in + (EConstr.of_constr c, sigma) + in + GlobEnv.register_constr_interp0 wit_ltac2 interp + +let () = + let interp ist poly env sigma concl id = + let ist = Tac2interp.get_env ist in + let c = Id.Map.find id ist.env_ist in + let c = Value.to_constr c in + let sigma = Typing.check env sigma c concl in + (c, sigma) + in + GlobEnv.register_constr_interp0 wit_ltac2_quotation interp + +let () = + let pr_raw id = Genprint.PrinterBasic (fun _env _sigma -> mt ()) in + let pr_glb id = Genprint.PrinterBasic (fun _env _sigma -> str "$" ++ Id.print id) in + let pr_top _ = Genprint.TopPrinterBasic mt in + Genprint.register_print0 wit_ltac2_quotation pr_raw pr_glb pr_top + +(** Ltac2 in Ltac1 *) + +let () = + let e = Tac2entries.Pltac.tac2expr in + let inject (loc, v) = Ltac_plugin.Tacexpr.TacGeneric (in_gen (rawwit wit_ltac2) v) in + Ltac_plugin.Tacentries.create_ltac_quotation "ltac2" inject (e, None) + +let () = + let open Ltac_plugin in + let open Tacinterp in + let idtac = Value.of_closure (default_ist ()) (Tacexpr.TacId []) in + let interp ist tac = +(* let ist = Tac2interp.get_env ist.Geninterp.lfun in *) + let ist = { env_ist = Id.Map.empty } in + Tac2interp.interp ist tac >>= fun _ -> + Ftactic.return idtac + in + Geninterp.register_interp0 wit_ltac2 interp + +let () = + let pr_raw _ = Genprint.PrinterBasic (fun _env _sigma -> mt ()) in + let pr_glb e = Genprint.PrinterBasic (fun _env _sigma -> Tac2print.pr_glbexpr e) in + let pr_top _ = Genprint.TopPrinterBasic mt in + Genprint.register_print0 wit_ltac2 pr_raw pr_glb pr_top + +(** Built-in notation scopes *) + +let add_scope s f = + Tac2entries.register_scope (Id.of_string s) f + +let rec pr_scope = let open CAst in function +| SexprStr {v=s} -> qstring s +| SexprInt {v=n} -> Pp.int n +| SexprRec (_, {v=na}, args) -> + let na = match na with + | None -> str "_" + | Some id -> Id.print id + in + na ++ str "(" ++ prlist_with_sep (fun () -> str ", ") pr_scope args ++ str ")" + +let scope_fail s args = + let args = str "(" ++ prlist_with_sep (fun () -> str ", ") pr_scope args ++ str ")" in + CErrors.user_err (str "Invalid arguments " ++ args ++ str " in scope " ++ str s) + +let q_unit = CAst.make @@ CTacCst (AbsKn (Tuple 0)) + +let add_generic_scope s entry arg = + let parse = function + | [] -> + let scope = Extend.Aentry entry in + let act x = CAst.make @@ CTacExt (arg, x) in + Tac2entries.ScopeRule (scope, act) + | arg -> scope_fail s arg + in + add_scope s parse + +open CAst + +let () = add_scope "keyword" begin function +| [SexprStr {loc;v=s}] -> + let scope = Extend.Atoken (Tok.PKEYWORD s) in + Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) +| arg -> scope_fail "keyword" arg +end + +let () = add_scope "terminal" begin function +| [SexprStr {loc;v=s}] -> + let scope = Extend.Atoken (CLexer.terminal s) in + Tac2entries.ScopeRule (scope, (fun _ -> q_unit)) +| arg -> scope_fail "terminal" arg +end + +let () = add_scope "list0" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let scope = Extend.Alist0 scope in + let act l = Tac2quote.of_list act l in + Tac2entries.ScopeRule (scope, act) +| [tok; SexprStr {v=str}] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let sep = Extend.Atoken (CLexer.terminal str) in + let scope = Extend.Alist0sep (scope, sep) in + let act l = Tac2quote.of_list act l in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "list0" arg +end + +let () = add_scope "list1" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let scope = Extend.Alist1 scope in + let act l = Tac2quote.of_list act l in + Tac2entries.ScopeRule (scope, act) +| [tok; SexprStr {v=str}] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let sep = Extend.Atoken (CLexer.terminal str) in + let scope = Extend.Alist1sep (scope, sep) in + let act l = Tac2quote.of_list act l in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "list1" arg +end + +let () = add_scope "opt" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let scope = Extend.Aopt scope in + let act opt = match opt with + | None -> + CAst.make @@ CTacCst (AbsKn (Other Core.c_none)) + | Some x -> + CAst.make @@ CTacApp (CAst.make @@ CTacCst (AbsKn (Other Core.c_some)), [act x]) + in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "opt" arg +end + +let () = add_scope "self" begin function +| [] -> + let scope = Extend.Aself in + let act tac = tac in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "self" arg +end + +let () = add_scope "next" begin function +| [] -> + let scope = Extend.Anext in + let act tac = tac in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "next" arg +end + +let () = add_scope "tactic" begin function +| [] -> + (* Default to level 5 parsing *) + let scope = Extend.Aentryl (tac2expr, "5") in + let act tac = tac in + Tac2entries.ScopeRule (scope, act) +| [SexprInt {loc;v=n}] as arg -> + let () = if n < 0 || n > 6 then scope_fail "tactic" arg in + let scope = Extend.Aentryl (tac2expr, string_of_int n) in + let act tac = tac in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "tactic" arg +end + +let () = add_scope "thunk" begin function +| [tok] -> + let Tac2entries.ScopeRule (scope, act) = Tac2entries.parse_scope tok in + let act e = Tac2quote.thunk (act e) in + Tac2entries.ScopeRule (scope, act) +| arg -> scope_fail "thunk" arg +end + +let add_expr_scope name entry f = + add_scope name begin function + | [] -> Tac2entries.ScopeRule (Extend.Aentry entry, f) + | arg -> scope_fail name arg + end + +let () = add_expr_scope "ident" q_ident (fun id -> Tac2quote.of_anti Tac2quote.of_ident id) +let () = add_expr_scope "bindings" q_bindings Tac2quote.of_bindings +let () = add_expr_scope "with_bindings" q_with_bindings Tac2quote.of_bindings +let () = add_expr_scope "intropattern" q_intropattern Tac2quote.of_intro_pattern +let () = add_expr_scope "intropatterns" q_intropatterns Tac2quote.of_intro_patterns +let () = add_expr_scope "destruction_arg" q_destruction_arg Tac2quote.of_destruction_arg +let () = add_expr_scope "induction_clause" q_induction_clause Tac2quote.of_induction_clause +let () = add_expr_scope "conversion" q_conversion Tac2quote.of_conversion +let () = add_expr_scope "rewriting" q_rewriting Tac2quote.of_rewriting +let () = add_expr_scope "clause" q_clause Tac2quote.of_clause +let () = add_expr_scope "hintdb" q_hintdb Tac2quote.of_hintdb +let () = add_expr_scope "occurrences" q_occurrences Tac2quote.of_occurrences +let () = add_expr_scope "dispatch" q_dispatch Tac2quote.of_dispatch +let () = add_expr_scope "strategy" q_strategy_flag Tac2quote.of_strategy_flag +let () = add_expr_scope "reference" q_reference Tac2quote.of_reference +let () = add_expr_scope "move_location" q_move_location Tac2quote.of_move_location +let () = add_expr_scope "pose" q_pose Tac2quote.of_pose +let () = add_expr_scope "assert" q_assert Tac2quote.of_assertion +let () = add_expr_scope "constr_matching" q_constr_matching Tac2quote.of_constr_matching +let () = add_expr_scope "goal_matching" q_goal_matching Tac2quote.of_goal_matching + +let () = add_generic_scope "constr" Pcoq.Constr.constr Tac2quote.wit_constr +let () = add_generic_scope "open_constr" Pcoq.Constr.constr Tac2quote.wit_open_constr +let () = add_generic_scope "pattern" Pcoq.Constr.constr Tac2quote.wit_pattern + +(** seq scope, a bit hairy *) + +open Extend +exception SelfSymbol + +let rec generalize_symbol : + type a tr s. (s, tr, a) Extend.symbol -> (s, Extend.norec, a) Extend.symbol = function +| Atoken tok -> Atoken tok +| Alist1 e -> Alist1 (generalize_symbol e) +| Alist1sep (e, sep) -> + let e = generalize_symbol e in + let sep = generalize_symbol sep in + Alist1sep (e, sep) +| Alist0 e -> Alist0 (generalize_symbol e) +| Alist0sep (e, sep) -> + let e = generalize_symbol e in + let sep = generalize_symbol sep in + Alist0sep (e, sep) +| Aopt e -> Aopt (generalize_symbol e) +| Aself -> raise SelfSymbol +| Anext -> raise SelfSymbol +| Aentry e -> Aentry e +| Aentryl (e, l) -> Aentryl (e, l) +| Arules r -> Arules r + +type _ converter = +| CvNil : (Loc.t -> raw_tacexpr) converter +| CvCns : 'act converter * ('a -> raw_tacexpr) option -> ('a -> 'act) converter + +let rec apply : type a. a converter -> raw_tacexpr list -> a = function +| CvNil -> fun accu loc -> Tac2quote.of_tuple ~loc accu +| CvCns (c, None) -> fun accu x -> apply c accu +| CvCns (c, Some f) -> fun accu x -> apply c (f x :: accu) + +type seqrule = +| Seqrule : (Tac2expr.raw_tacexpr, Extend.norec, 'act, Loc.t -> raw_tacexpr) rule * 'act converter -> seqrule + +let rec make_seq_rule = function +| [] -> + Seqrule (Stop, CvNil) +| tok :: rem -> + let Tac2entries.ScopeRule (scope, f) = Tac2entries.parse_scope tok in + let scope = generalize_symbol scope in + let Seqrule (r, c) = make_seq_rule rem in + let r = NextNoRec (r, scope) in + let f = match tok with + | SexprStr _ -> None (* Leave out mere strings *) + | _ -> Some f + in + Seqrule (r, CvCns (c, f)) + +let () = add_scope "seq" begin fun toks -> + let scope = + try + let Seqrule (r, c) = make_seq_rule (List.rev toks) in + Arules [Rules (r, apply c [])] + with SelfSymbol -> + CErrors.user_err (str "Recursive symbols (self / next) are not allowed in local rules") + in + Tac2entries.ScopeRule (scope, (fun e -> e)) +end diff --git a/user-contrib/Ltac2/tac2core.mli b/user-contrib/Ltac2/tac2core.mli new file mode 100644 index 0000000000..9fae65bb3e --- /dev/null +++ b/user-contrib/Ltac2/tac2core.mli @@ -0,0 +1,30 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Tac2expr + +(** {5 Hardwired data} *) + +module Core : +sig + +val t_list : type_constant +val c_nil : ltac_constructor +val c_cons : ltac_constructor + +val t_int : type_constant +val t_option : type_constant +val t_string : type_constant +val t_array : type_constant + +val c_true : ltac_constructor +val c_false : ltac_constructor + +end + +val pf_apply : (Environ.env -> Evd.evar_map -> 'a Proofview.tactic) -> 'a Proofview.tactic diff --git a/user-contrib/Ltac2/tac2dyn.ml b/user-contrib/Ltac2/tac2dyn.ml new file mode 100644 index 0000000000..896676f08b --- /dev/null +++ b/user-contrib/Ltac2/tac2dyn.ml @@ -0,0 +1,27 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +module Arg = +struct + module DYN = Dyn.Make(struct end) + module Map = DYN.Map + type ('a, 'b) tag = ('a * 'b) DYN.tag + let eq = DYN.eq + let repr = DYN.repr + let create = DYN.create +end + +module type Param = sig type ('raw, 'glb) t end + +module ArgMap (M : Param) = +struct + type _ pack = Pack : ('raw, 'glb) M.t -> ('raw * 'glb) pack + include Arg.Map(struct type 'a t = 'a pack end) +end + +module Val = Dyn.Make(struct end) diff --git a/user-contrib/Ltac2/tac2dyn.mli b/user-contrib/Ltac2/tac2dyn.mli new file mode 100644 index 0000000000..e995296840 --- /dev/null +++ b/user-contrib/Ltac2/tac2dyn.mli @@ -0,0 +1,34 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +(** Dynamic arguments for Ltac2. *) + +module Arg : +sig + type ('a, 'b) tag + val create : string -> ('a, 'b) tag + val eq : ('a1, 'b1) tag -> ('a2, 'b2) tag -> ('a1 * 'b1, 'a2 * 'b2) CSig.eq option + val repr : ('a, 'b) tag -> string +end +(** Arguments that are part of an AST. *) + +module type Param = sig type ('raw, 'glb) t end + +module ArgMap (M : Param) : +sig + type _ pack = Pack : ('raw, 'glb) M.t -> ('raw * 'glb) pack + type t + val empty : t + val add : ('a, 'b) Arg.tag -> ('a * 'b) pack -> t -> t + val remove : ('a, 'b) Arg.tag -> t -> t + val find : ('a, 'b) Arg.tag -> t -> ('a * 'b) pack + val mem : ('a, 'b) Arg.tag -> t -> bool +end + +module Val : Dyn.S +(** Toplevel values *) diff --git a/user-contrib/Ltac2/tac2entries.ml b/user-contrib/Ltac2/tac2entries.ml new file mode 100644 index 0000000000..9fd01426de --- /dev/null +++ b/user-contrib/Ltac2/tac2entries.ml @@ -0,0 +1,938 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Pp +open Util +open CAst +open CErrors +open Names +open Libnames +open Libobject +open Nametab +open Tac2expr +open Tac2print +open Tac2intern + +(** Grammar entries *) + +module Pltac = +struct +let tac2expr = Pcoq.Entry.create "tactic:tac2expr" + +let q_ident = Pcoq.Entry.create "tactic:q_ident" +let q_bindings = Pcoq.Entry.create "tactic:q_bindings" +let q_with_bindings = Pcoq.Entry.create "tactic:q_with_bindings" +let q_intropattern = Pcoq.Entry.create "tactic:q_intropattern" +let q_intropatterns = Pcoq.Entry.create "tactic:q_intropatterns" +let q_destruction_arg = Pcoq.Entry.create "tactic:q_destruction_arg" +let q_induction_clause = Pcoq.Entry.create "tactic:q_induction_clause" +let q_conversion = Pcoq.Entry.create "tactic:q_conversion" +let q_rewriting = Pcoq.Entry.create "tactic:q_rewriting" +let q_clause = Pcoq.Entry.create "tactic:q_clause" +let q_dispatch = Pcoq.Entry.create "tactic:q_dispatch" +let q_occurrences = Pcoq.Entry.create "tactic:q_occurrences" +let q_reference = Pcoq.Entry.create "tactic:q_reference" +let q_strategy_flag = Pcoq.Entry.create "tactic:q_strategy_flag" +let q_constr_matching = Pcoq.Entry.create "tactic:q_constr_matching" +let q_goal_matching = Pcoq.Entry.create "tactic:q_goal_matching" +let q_hintdb = Pcoq.Entry.create "tactic:q_hintdb" +let q_move_location = Pcoq.Entry.create "tactic:q_move_location" +let q_pose = Pcoq.Entry.create "tactic:q_pose" +let q_assert = Pcoq.Entry.create "tactic:q_assert" +end + +(** Tactic definition *) + +type tacdef = { + tacdef_local : bool; + tacdef_mutable : bool; + tacdef_expr : glb_tacexpr; + tacdef_type : type_scheme; +} + +let perform_tacdef visibility ((sp, kn), def) = + let () = if not def.tacdef_local then Tac2env.push_ltac visibility sp (TacConstant kn) in + let data = { + Tac2env.gdata_expr = def.tacdef_expr; + gdata_type = def.tacdef_type; + gdata_mutable = def.tacdef_mutable; + } in + Tac2env.define_global kn data + +let load_tacdef i obj = perform_tacdef (Until i) obj +let open_tacdef i obj = perform_tacdef (Exactly i) obj + +let cache_tacdef ((sp, kn), def) = + let () = Tac2env.push_ltac (Until 1) sp (TacConstant kn) in + let data = { + Tac2env.gdata_expr = def.tacdef_expr; + gdata_type = def.tacdef_type; + gdata_mutable = def.tacdef_mutable; + } in + Tac2env.define_global kn data + +let subst_tacdef (subst, def) = + let expr' = subst_expr subst def.tacdef_expr in + let type' = subst_type_scheme subst def.tacdef_type in + if expr' == def.tacdef_expr && type' == def.tacdef_type then def + else { def with tacdef_expr = expr'; tacdef_type = type' } + +let classify_tacdef o = Substitute o + +let inTacDef : tacdef -> obj = + declare_object {(default_object "TAC2-DEFINITION") with + cache_function = cache_tacdef; + load_function = load_tacdef; + open_function = open_tacdef; + subst_function = subst_tacdef; + classify_function = classify_tacdef} + +(** Type definition *) + +type typdef = { + typdef_local : bool; + typdef_expr : glb_quant_typedef; +} + +let change_kn_label kn id = + let mp = KerName.modpath kn in + KerName.make mp (Label.of_id id) + +let change_sp_label sp id = + let (dp, _) = Libnames.repr_path sp in + Libnames.make_path dp id + +let push_typedef visibility sp kn (_, def) = match def with +| GTydDef _ -> + Tac2env.push_type visibility sp kn +| GTydAlg { galg_constructors = cstrs } -> + (* Register constructors *) + let iter (c, _) = + let spc = change_sp_label sp c in + let knc = change_kn_label kn c in + Tac2env.push_constructor visibility spc knc + in + Tac2env.push_type visibility sp kn; + List.iter iter cstrs +| GTydRec fields -> + (* Register fields *) + let iter (c, _, _) = + let spc = change_sp_label sp c in + let knc = change_kn_label kn c in + Tac2env.push_projection visibility spc knc + in + Tac2env.push_type visibility sp kn; + List.iter iter fields +| GTydOpn -> + Tac2env.push_type visibility sp kn + +let next i = + let ans = !i in + let () = incr i in + ans + +let define_typedef kn (params, def as qdef) = match def with +| GTydDef _ -> + Tac2env.define_type kn qdef +| GTydAlg { galg_constructors = cstrs } -> + (* Define constructors *) + let constant = ref 0 in + let nonconstant = ref 0 in + let iter (c, args) = + let knc = change_kn_label kn c in + let tag = if List.is_empty args then next constant else next nonconstant in + let data = { + Tac2env.cdata_prms = params; + cdata_type = kn; + cdata_args = args; + cdata_indx = Some tag; + } in + Tac2env.define_constructor knc data + in + Tac2env.define_type kn qdef; + List.iter iter cstrs +| GTydRec fs -> + (* Define projections *) + let iter i (id, mut, t) = + let knp = change_kn_label kn id in + let proj = { + Tac2env.pdata_prms = params; + pdata_type = kn; + pdata_ptyp = t; + pdata_mutb = mut; + pdata_indx = i; + } in + Tac2env.define_projection knp proj + in + Tac2env.define_type kn qdef; + List.iteri iter fs +| GTydOpn -> + Tac2env.define_type kn qdef + +let perform_typdef vs ((sp, kn), def) = + let () = if not def.typdef_local then push_typedef vs sp kn def.typdef_expr in + define_typedef kn def.typdef_expr + +let load_typdef i obj = perform_typdef (Until i) obj +let open_typdef i obj = perform_typdef (Exactly i) obj + +let cache_typdef ((sp, kn), def) = + let () = push_typedef (Until 1) sp kn def.typdef_expr in + define_typedef kn def.typdef_expr + +let subst_typdef (subst, def) = + let expr' = subst_quant_typedef subst def.typdef_expr in + if expr' == def.typdef_expr then def else { def with typdef_expr = expr' } + +let classify_typdef o = Substitute o + +let inTypDef : typdef -> obj = + declare_object {(default_object "TAC2-TYPE-DEFINITION") with + cache_function = cache_typdef; + load_function = load_typdef; + open_function = open_typdef; + subst_function = subst_typdef; + classify_function = classify_typdef} + +(** Type extension *) + +type extension_data = { + edata_name : Id.t; + edata_args : int glb_typexpr list; +} + +type typext = { + typext_local : bool; + typext_prms : int; + typext_type : type_constant; + typext_expr : extension_data list; +} + +let push_typext vis sp kn def = + let iter data = + let spc = change_sp_label sp data.edata_name in + let knc = change_kn_label kn data.edata_name in + Tac2env.push_constructor vis spc knc + in + List.iter iter def.typext_expr + +let define_typext kn def = + let iter data = + let knc = change_kn_label kn data.edata_name in + let cdata = { + Tac2env.cdata_prms = def.typext_prms; + cdata_type = def.typext_type; + cdata_args = data.edata_args; + cdata_indx = None; + } in + Tac2env.define_constructor knc cdata + in + List.iter iter def.typext_expr + +let cache_typext ((sp, kn), def) = + let () = define_typext kn def in + push_typext (Until 1) sp kn def + +let perform_typext vs ((sp, kn), def) = + let () = if not def.typext_local then push_typext vs sp kn def in + define_typext kn def + +let load_typext i obj = perform_typext (Until i) obj +let open_typext i obj = perform_typext (Exactly i) obj + +let subst_typext (subst, e) = + let open Mod_subst in + let subst_data data = + let edata_args = List.Smart.map (fun e -> subst_type subst e) data.edata_args in + if edata_args == data.edata_args then data + else { data with edata_args } + in + let typext_type = subst_kn subst e.typext_type in + let typext_expr = List.Smart.map subst_data e.typext_expr in + if typext_type == e.typext_type && typext_expr == e.typext_expr then + e + else + { e with typext_type; typext_expr } + +let classify_typext o = Substitute o + +let inTypExt : typext -> obj = + declare_object {(default_object "TAC2-TYPE-EXTENSION") with + cache_function = cache_typext; + load_function = load_typext; + open_function = open_typext; + subst_function = subst_typext; + classify_function = classify_typext} + +(** Toplevel entries *) + +let fresh_var avoid x = + let bad id = + Id.Set.mem id avoid || + (try ignore (Tac2env.locate_ltac (qualid_of_ident id)); true with Not_found -> false) + in + Namegen.next_ident_away_from (Id.of_string x) bad + +let extract_pattern_type ({loc;v=p} as pat) = match p with +| CPatCnv (pat, ty) -> pat, Some ty +| CPatVar _ | CPatRef _ -> pat, None + +(** Mangle recursive tactics *) +let inline_rec_tactic tactics = + let avoid = List.fold_left (fun accu ({v=id}, _) -> Id.Set.add id accu) Id.Set.empty tactics in + let map (id, e) = match e.v with + | CTacFun (pat, _) -> (id, List.map extract_pattern_type pat, e) + | _ -> + user_err ?loc:id.loc (str "Recursive tactic definitions must be functions") + in + let tactics = List.map map tactics in + let map (id, pat, e) = + let fold_var (avoid, ans) (pat, _) = + let id = fresh_var avoid "x" in + let loc = pat.loc in + (Id.Set.add id avoid, CAst.make ?loc id :: ans) + in + (* Fresh variables to abstract over the function patterns *) + let _, vars = List.fold_left fold_var (avoid, []) pat in + let map_body ({loc;v=id}, _, e) = CAst.(make ?loc @@ CPatVar (Name id)), e in + let bnd = List.map map_body tactics in + let pat_of_id {loc;v=id} = CAst.make ?loc @@ CPatVar (Name id) in + let var_of_id {loc;v=id} = + let qid = qualid_of_ident ?loc id in + CAst.make ?loc @@ CTacRef (RelId qid) + in + let loc0 = e.loc in + let vpat = List.map pat_of_id vars in + let varg = List.map var_of_id vars in + let e = CAst.make ?loc:loc0 @@ CTacLet (true, bnd, CAst.make ?loc:loc0 @@ CTacApp (var_of_id id, varg)) in + (id, CAst.make ?loc:loc0 @@ CTacFun (vpat, e)) + in + List.map map tactics + +let check_lowercase {loc;v=id} = + if Tac2env.is_constructor (Libnames.qualid_of_ident id) then + user_err ?loc (str "The identifier " ++ Id.print id ++ str " must be lowercase") + +let register_ltac ?(local = false) ?(mut = false) isrec tactics = + let map ({loc;v=na}, e) = + let id = match na with + | Anonymous -> + user_err ?loc (str "Tactic definition must have a name") + | Name id -> id + in + let () = check_lowercase CAst.(make ?loc id) in + (CAst.(make ?loc id), e) + in + let tactics = List.map map tactics in + let tactics = + if isrec then inline_rec_tactic tactics else tactics + in + let map ({loc;v=id}, e) = + let (e, t) = intern ~strict:true e in + let () = + if not (is_value e) then + user_err ?loc (str "Tactic definition must be a syntactical value") + in + let kn = Lib.make_kn id in + let exists = + try let _ = Tac2env.interp_global kn in true with Not_found -> false + in + let () = + if exists then + user_err ?loc (str "Tactic " ++ Names.Id.print id ++ str " already exists") + in + (id, e, t) + in + let defs = List.map map tactics in + let iter (id, e, t) = + let def = { + tacdef_local = local; + tacdef_mutable = mut; + tacdef_expr = e; + tacdef_type = t; + } in + ignore (Lib.add_leaf id (inTacDef def)) + in + List.iter iter defs + +let qualid_to_ident qid = + if qualid_is_ident qid then CAst.make ?loc:qid.CAst.loc @@ qualid_basename qid + else user_err ?loc:qid.CAst.loc (str "Identifier expected") + +let register_typedef ?(local = false) isrec types = + let same_name ({v=id1}, _) ({v=id2}, _) = Id.equal id1 id2 in + let () = match List.duplicates same_name types with + | [] -> () + | ({loc;v=id}, _) :: _ -> + user_err ?loc (str "Multiple definition of the type name " ++ Id.print id) + in + let check ({loc;v=id}, (params, def)) = + let same_name {v=id1} {v=id2} = Id.equal id1 id2 in + let () = match List.duplicates same_name params with + | [] -> () + | {loc;v=id} :: _ -> + user_err ?loc (str "The type parameter " ++ Id.print id ++ + str " occurs several times") + in + match def with + | CTydDef _ -> + if isrec then + user_err ?loc (str "The type abbreviation " ++ Id.print id ++ + str " cannot be recursive") + | CTydAlg cs -> + let same_name (id1, _) (id2, _) = Id.equal id1 id2 in + let () = match List.duplicates same_name cs with + | [] -> () + | (id, _) :: _ -> + user_err (str "Multiple definitions of the constructor " ++ Id.print id) + in + () + | CTydRec ps -> + let same_name (id1, _, _) (id2, _, _) = Id.equal id1 id2 in + let () = match List.duplicates same_name ps with + | [] -> () + | (id, _, _) :: _ -> + user_err (str "Multiple definitions of the projection " ++ Id.print id) + in + () + | CTydOpn -> + if isrec then + user_err ?loc (str "The open type declaration " ++ Id.print id ++ + str " cannot be recursive") + in + let () = List.iter check types in + let self = + if isrec then + let fold accu ({v=id}, (params, _)) = + Id.Map.add id (Lib.make_kn id, List.length params) accu + in + List.fold_left fold Id.Map.empty types + else Id.Map.empty + in + let map ({v=id}, def) = + let typdef = { + typdef_local = local; + typdef_expr = intern_typedef self def; + } in + (id, typdef) + in + let types = List.map map types in + let iter (id, def) = ignore (Lib.add_leaf id (inTypDef def)) in + List.iter iter types + +let register_primitive ?(local = false) {loc;v=id} t ml = + let t = intern_open_type t in + let rec count_arrow = function + | GTypArrow (_, t) -> 1 + count_arrow t + | _ -> 0 + in + let arrows = count_arrow (snd t) in + let () = if Int.equal arrows 0 then + user_err ?loc (str "External tactic must have at least one argument") in + let () = + try let _ = Tac2env.interp_primitive ml in () with Not_found -> + user_err ?loc (str "Unregistered primitive " ++ + quote (str ml.mltac_plugin) ++ spc () ++ quote (str ml.mltac_tactic)) + in + let init i = Id.of_string (Printf.sprintf "x%i" i) in + let names = List.init arrows init in + let bnd = List.map (fun id -> Name id) names in + let arg = List.map (fun id -> GTacVar id) names in + let e = GTacFun (bnd, GTacPrm (ml, arg)) in + let def = { + tacdef_local = local; + tacdef_mutable = false; + tacdef_expr = e; + tacdef_type = t; + } in + ignore (Lib.add_leaf id (inTacDef def)) + +let register_open ?(local = false) qid (params, def) = + let kn = + try Tac2env.locate_type qid + with Not_found -> + user_err ?loc:qid.CAst.loc (str "Unbound type " ++ pr_qualid qid) + in + let (tparams, t) = Tac2env.interp_type kn in + let () = match t with + | GTydOpn -> () + | GTydAlg _ | GTydRec _ | GTydDef _ -> + user_err ?loc:qid.CAst.loc (str "Type " ++ pr_qualid qid ++ str " is not an open type") + in + let () = + if not (Int.equal (List.length params) tparams) then + Tac2intern.error_nparams_mismatch ?loc:qid.CAst.loc (List.length params) tparams + in + match def with + | CTydOpn -> () + | CTydAlg def -> + let intern_type t = + let tpe = CTydDef (Some t) in + let (_, ans) = intern_typedef Id.Map.empty (params, tpe) in + match ans with + | GTydDef (Some t) -> t + | _ -> assert false + in + let map (id, tpe) = + let tpe = List.map intern_type tpe in + { edata_name = id; edata_args = tpe } + in + let def = List.map map def in + let def = { + typext_local = local; + typext_type = kn; + typext_prms = tparams; + typext_expr = def; + } in + Lib.add_anonymous_leaf (inTypExt def) + | CTydRec _ | CTydDef _ -> + user_err ?loc:qid.CAst.loc (str "Extensions only accept inductive constructors") + +let register_type ?local isrec types = match types with +| [qid, true, def] -> + let () = if isrec then user_err ?loc:qid.CAst.loc (str "Extensions cannot be recursive") in + register_open ?local qid def +| _ -> + let map (qid, redef, def) = + let () = if redef then + user_err ?loc:qid.loc (str "Types can only be extended one by one") + in + (qualid_to_ident qid, def) + in + let types = List.map map types in + register_typedef ?local isrec types + +(** Parsing *) + +type 'a token = +| TacTerm of string +| TacNonTerm of Name.t * 'a + +type scope_rule = +| ScopeRule : (raw_tacexpr, _, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule + +type scope_interpretation = sexpr list -> scope_rule + +let scope_table : scope_interpretation Id.Map.t ref = ref Id.Map.empty + +let register_scope id s = + scope_table := Id.Map.add id s !scope_table + +module ParseToken = +struct + +let loc_of_token = function +| SexprStr {loc} -> loc +| SexprInt {loc} -> loc +| SexprRec (loc, _, _) -> Some loc + +let parse_scope = function +| SexprRec (_, {loc;v=Some id}, toks) -> + if Id.Map.mem id !scope_table then + Id.Map.find id !scope_table toks + else + CErrors.user_err ?loc (str "Unknown scope" ++ spc () ++ Names.Id.print id) +| SexprStr {v=str} -> + let v_unit = CAst.make @@ CTacCst (AbsKn (Tuple 0)) in + ScopeRule (Extend.Atoken (Tok.PIDENT (Some str)), (fun _ -> v_unit)) +| tok -> + let loc = loc_of_token tok in + CErrors.user_err ?loc (str "Invalid parsing token") + +let parse_token = function +| SexprStr {v=s} -> TacTerm s +| SexprRec (_, {v=na}, [tok]) -> + let na = match na with None -> Anonymous | Some id -> Name id in + let scope = parse_scope tok in + TacNonTerm (na, scope) +| tok -> + let loc = loc_of_token tok in + CErrors.user_err ?loc (str "Invalid parsing token") + +end + +let parse_scope = ParseToken.parse_scope + +type synext = { + synext_tok : sexpr list; + synext_exp : raw_tacexpr; + synext_lev : int option; + synext_loc : bool; +} + +type krule = +| KRule : + (raw_tacexpr, _, 'act, Loc.t -> raw_tacexpr) Extend.rule * + ((Loc.t -> (Name.t * raw_tacexpr) list -> raw_tacexpr) -> 'act) -> krule + +let rec get_rule (tok : scope_rule token list) : krule = match tok with +| [] -> KRule (Extend.Stop, fun k loc -> k loc []) +| TacNonTerm (na, ScopeRule (scope, inj)) :: tok -> + let KRule (rule, act) = get_rule tok in + let rule = Extend.Next (rule, scope) in + let act k e = act (fun loc acc -> k loc ((na, inj e) :: acc)) in + KRule (rule, act) +| TacTerm t :: tok -> + let KRule (rule, act) = get_rule tok in + let rule = Extend.Next (rule, Extend.Atoken (CLexer.terminal t)) in + let act k _ = act k in + KRule (rule, act) + +let perform_notation syn st = + let tok = List.rev_map ParseToken.parse_token syn.synext_tok in + let KRule (rule, act) = get_rule tok in + let mk loc args = + let map (na, e) = + ((CAst.make ?loc:e.loc @@ CPatVar na), e) + in + let bnd = List.map map args in + CAst.make ~loc @@ CTacLet (false, bnd, syn.synext_exp) + in + let rule = Extend.Rule (rule, act mk) in + let lev = match syn.synext_lev with + | None -> None + | Some lev -> Some (string_of_int lev) + in + let rule = (lev, None, [rule]) in + ([Pcoq.ExtendRule (Pltac.tac2expr, None, (None, [rule]))], st) + +let ltac2_notation = + Pcoq.create_grammar_command "ltac2-notation" perform_notation + +let cache_synext (_, syn) = + Pcoq.extend_grammar_command ltac2_notation syn + +let open_synext i (_, syn) = + if Int.equal i 1 then Pcoq.extend_grammar_command ltac2_notation syn + +let subst_synext (subst, syn) = + let e = Tac2intern.subst_rawexpr subst syn.synext_exp in + if e == syn.synext_exp then syn else { syn with synext_exp = e } + +let classify_synext o = + if o.synext_loc then Dispose else Substitute o + +let inTac2Notation : synext -> obj = + declare_object {(default_object "TAC2-NOTATION") with + cache_function = cache_synext; + open_function = open_synext; + subst_function = subst_synext; + classify_function = classify_synext} + +type abbreviation = { + abbr_body : raw_tacexpr; +} + +let perform_abbreviation visibility ((sp, kn), abbr) = + let () = Tac2env.push_ltac visibility sp (TacAlias kn) in + Tac2env.define_alias kn abbr.abbr_body + +let load_abbreviation i obj = perform_abbreviation (Until i) obj +let open_abbreviation i obj = perform_abbreviation (Exactly i) obj + +let cache_abbreviation ((sp, kn), abbr) = + let () = Tac2env.push_ltac (Until 1) sp (TacAlias kn) in + Tac2env.define_alias kn abbr.abbr_body + +let subst_abbreviation (subst, abbr) = + let body' = subst_rawexpr subst abbr.abbr_body in + if body' == abbr.abbr_body then abbr + else { abbr_body = body' } + +let classify_abbreviation o = Substitute o + +let inTac2Abbreviation : abbreviation -> obj = + declare_object {(default_object "TAC2-ABBREVIATION") with + cache_function = cache_abbreviation; + load_function = load_abbreviation; + open_function = open_abbreviation; + subst_function = subst_abbreviation; + classify_function = classify_abbreviation} + +let register_notation ?(local = false) tkn lev body = match tkn, lev with +| [SexprRec (_, {loc;v=Some id}, [])], None -> + (* Tactic abbreviation *) + let () = check_lowercase CAst.(make ?loc id) in + let body = Tac2intern.globalize Id.Set.empty body in + let abbr = { abbr_body = body } in + ignore (Lib.add_leaf id (inTac2Abbreviation abbr)) +| _ -> + (* Check that the tokens make sense *) + let entries = List.map ParseToken.parse_token tkn in + let fold accu tok = match tok with + | TacTerm _ -> accu + | TacNonTerm (Name id, _) -> Id.Set.add id accu + | TacNonTerm (Anonymous, _) -> accu + in + let ids = List.fold_left fold Id.Set.empty entries in + (* Globalize so that names are absolute *) + let body = Tac2intern.globalize ids body in + let lev = match lev with Some _ -> lev | None -> Some 5 in + let ext = { + synext_tok = tkn; + synext_exp = body; + synext_lev = lev; + synext_loc = local; + } in + Lib.add_anonymous_leaf (inTac2Notation ext) + +type redefinition = { + redef_kn : ltac_constant; + redef_body : glb_tacexpr; +} + +let perform_redefinition (_, redef) = + let kn = redef.redef_kn in + let data = Tac2env.interp_global kn in + let data = { data with Tac2env.gdata_expr = redef.redef_body } in + Tac2env.define_global kn data + +let subst_redefinition (subst, redef) = + let kn = Mod_subst.subst_kn subst redef.redef_kn in + let body = Tac2intern.subst_expr subst redef.redef_body in + if kn == redef.redef_kn && body == redef.redef_body then redef + else { redef_kn = kn; redef_body = body } + +let classify_redefinition o = Substitute o + +let inTac2Redefinition : redefinition -> obj = + declare_object {(default_object "TAC2-REDEFINITION") with + cache_function = perform_redefinition; + open_function = (fun _ -> perform_redefinition); + subst_function = subst_redefinition; + classify_function = classify_redefinition } + +let register_redefinition ?(local = false) qid e = + let kn = + try Tac2env.locate_ltac qid + with Not_found -> user_err ?loc:qid.CAst.loc (str "Unknown tactic " ++ pr_qualid qid) + in + let kn = match kn with + | TacConstant kn -> kn + | TacAlias _ -> + user_err ?loc:qid.CAst.loc (str "Cannot redefine syntactic abbreviations") + in + let data = Tac2env.interp_global kn in + let () = + if not (data.Tac2env.gdata_mutable) then + user_err ?loc:qid.CAst.loc (str "The tactic " ++ pr_qualid qid ++ str " is not declared as mutable") + in + let (e, t) = intern ~strict:true e in + let () = + if not (is_value e) then + user_err ?loc:qid.CAst.loc (str "Tactic definition must be a syntactical value") + in + let () = + if not (Tac2intern.check_subtype t data.Tac2env.gdata_type) then + let name = int_name () in + user_err ?loc:qid.CAst.loc (str "Type " ++ pr_glbtype name (snd t) ++ + str " is not a subtype of " ++ pr_glbtype name (snd data.Tac2env.gdata_type)) + in + let def = { + redef_kn = kn; + redef_body = e; + } in + Lib.add_anonymous_leaf (inTac2Redefinition def) + +let perform_eval ~pstate e = + let open Proofview.Notations in + let env = Global.env () in + let (e, ty) = Tac2intern.intern ~strict:false e in + let v = Tac2interp.interp Tac2interp.empty_environment e in + let selector, proof = + match pstate with + | None -> + let sigma = Evd.from_env env in + let name, poly = Id.of_string "ltac2", false in + Goal_select.SelectAll, Proof.start ~name ~poly sigma [] + | Some pstate -> + Goal_select.get_default_goal_selector (), + Proof_global.give_me_the_proof pstate + in + let v = match selector with + | Goal_select.SelectNth i -> Proofview.tclFOCUS i i v + | Goal_select.SelectList l -> Proofview.tclFOCUSLIST l v + | Goal_select.SelectId id -> Proofview.tclFOCUSID id v + | Goal_select.SelectAll -> v + | Goal_select.SelectAlreadyFocused -> assert false (* TODO **) + in + (* HACK: the API doesn't allow to return a value *) + let ans = ref None in + let tac = (v >>= fun r -> ans := Some r; Proofview.tclUNIT ()) in + let (proof, _) = Proof.run_tactic (Global.env ()) tac proof in + let sigma = Proof.in_proof proof (fun sigma -> sigma) in + let ans = match !ans with None -> assert false | Some r -> r in + let name = int_name () in + Feedback.msg_notice (str "- : " ++ pr_glbtype name (snd ty) + ++ spc () ++ str "=" ++ spc () ++ + Tac2print.pr_valexpr env sigma ans (snd ty)) + +(** Toplevel entries *) + +let register_struct ?local ~pstate str = match str with +| StrVal (mut, isrec, e) -> register_ltac ?local ~mut isrec e +| StrTyp (isrec, t) -> register_type ?local isrec t +| StrPrm (id, t, ml) -> register_primitive ?local id t ml +| StrSyn (tok, lev, e) -> register_notation ?local tok lev e +| StrMut (qid, e) -> register_redefinition ?local qid e +| StrRun e -> perform_eval ~pstate e + +(** Toplevel exception *) + +let _ = Goptions.declare_bool_option { + Goptions.optdepr = false; + Goptions.optname = "print Ltac2 backtrace"; + Goptions.optkey = ["Ltac2"; "Backtrace"]; + Goptions.optread = (fun () -> !Tac2interp.print_ltac2_backtrace); + Goptions.optwrite = (fun b -> Tac2interp.print_ltac2_backtrace := b); +} + +let backtrace : backtrace Exninfo.t = Exninfo.make () + +let pr_frame = function +| FrAnon e -> str "Call {" ++ pr_glbexpr e ++ str "}" +| FrLtac kn -> + str "Call " ++ Libnames.pr_qualid (Tac2env.shortest_qualid_of_ltac (TacConstant kn)) +| FrPrim ml -> + str "Prim <" ++ str ml.mltac_plugin ++ str ":" ++ str ml.mltac_tactic ++ str ">" +| FrExtn (tag, arg) -> + let obj = Tac2env.interp_ml_object tag in + str "Extn " ++ str (Tac2dyn.Arg.repr tag) ++ str ":" ++ spc () ++ + obj.Tac2env.ml_print (Global.env ()) arg + +let () = register_handler begin function +| Tac2interp.LtacError (kn, args) -> + let t_exn = KerName.make Tac2env.coq_prefix (Label.make "exn") in + let v = Tac2ffi.of_open (kn, args) in + let t = GTypRef (Other t_exn, []) in + let c = Tac2print.pr_valexpr (Global.env ()) Evd.empty v t in + hov 0 (str "Uncaught Ltac2 exception:" ++ spc () ++ hov 0 c) +| _ -> raise Unhandled +end + +let () = ExplainErr.register_additional_error_info begin fun (e, info) -> + if !Tac2interp.print_ltac2_backtrace then + let bt = Exninfo.get info backtrace in + let bt = match bt with + | Some bt -> bt + | None -> raise Exit + in + let bt = + str "Backtrace:" ++ fnl () ++ prlist_with_sep fnl pr_frame bt ++ fnl () + in + Some (Loc.tag @@ Some bt) + else raise Exit +end + +(** Printing *) + +let print_ltac qid = + if Tac2env.is_constructor qid then + let kn = + try Tac2env.locate_constructor qid + with Not_found -> user_err ?loc:qid.CAst.loc (str "Unknown constructor " ++ pr_qualid qid) + in + let _ = Tac2env.interp_constructor kn in + Feedback.msg_notice (hov 2 (str "Constructor" ++ spc () ++ str ":" ++ spc () ++ pr_qualid qid)) + else + let kn = + try Tac2env.locate_ltac qid + with Not_found -> user_err ?loc:qid.CAst.loc (str "Unknown tactic " ++ pr_qualid qid) + in + match kn with + | TacConstant kn -> + let data = Tac2env.interp_global kn in + let e = data.Tac2env.gdata_expr in + let (_, t) = data.Tac2env.gdata_type in + let name = int_name () in + Feedback.msg_notice ( + hov 0 ( + hov 2 (pr_qualid qid ++ spc () ++ str ":" ++ spc () ++ pr_glbtype name t) ++ fnl () ++ + hov 2 (pr_qualid qid ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr e) + ) + ) + | TacAlias kn -> + Feedback.msg_notice (str "Alias to ...") + +(** Calling tactics *) + +let solve ~pstate default tac = + let pstate, status = Proof_global.with_current_proof begin fun etac p -> + let with_end_tac = if default then Some etac else None in + let g = Goal_select.get_default_goal_selector () in + let (p, status) = Pfedit.solve g None tac ?with_end_tac p in + (* in case a strict subtree was completed, + go back to the top of the prooftree *) + let p = Proof.maximal_unfocus Vernacentries.command_focus p in + p, status + end pstate in + if not status then Feedback.feedback Feedback.AddedAxiom; + pstate + +let call ~pstate ~default e = + let loc = e.loc in + let (e, t) = intern ~strict:false e in + let () = check_unit ?loc t in + let tac = Tac2interp.interp Tac2interp.empty_environment e in + solve ~pstate default (Proofview.tclIGNORE tac) + +(** Primitive algebraic types than can't be defined Coq-side *) + +let register_prim_alg name params def = + let id = Id.of_string name in + let def = List.map (fun (cstr, tpe) -> (Id.of_string_soft cstr, tpe)) def in + let getn (const, nonconst) (c, args) = match args with + | [] -> (succ const, nonconst) + | _ :: _ -> (const, succ nonconst) + in + let nconst, nnonconst = List.fold_left getn (0, 0) def in + let alg = { + galg_constructors = def; + galg_nconst = nconst; + galg_nnonconst = nnonconst; + } in + let def = (params, GTydAlg alg) in + let def = { typdef_local = false; typdef_expr = def } in + ignore (Lib.add_leaf id (inTypDef def)) + +let coq_def n = KerName.make Tac2env.coq_prefix (Label.make n) + +let def_unit = { + typdef_local = false; + typdef_expr = 0, GTydDef (Some (GTypRef (Tuple 0, []))); +} + +let t_list = coq_def "list" + +let (f_register_constr_quotations, register_constr_quotations) = Hook.make () + +let cache_ltac2_init (_, ()) = + Hook.get f_register_constr_quotations () + +let load_ltac2_init _ (_, ()) = + Hook.get f_register_constr_quotations () + +let open_ltac2_init _ (_, ()) = + Goptions.set_string_option_value_gen ["Default"; "Proof"; "Mode"] "Ltac2" + +(** Dummy object that register global rules when Require is called *) +let inTac2Init : unit -> obj = + declare_object {(default_object "TAC2-INIT") with + cache_function = cache_ltac2_init; + load_function = load_ltac2_init; + open_function = open_ltac2_init; + } + +let _ = Mltop.declare_cache_obj begin fun () -> + ignore (Lib.add_leaf (Id.of_string "unit") (inTypDef def_unit)); + register_prim_alg "list" 1 [ + ("[]", []); + ("::", [GTypVar 0; GTypRef (Other t_list, [GTypVar 0])]); + ]; + Lib.add_anonymous_leaf (inTac2Init ()); +end "ltac2_plugin" diff --git a/user-contrib/Ltac2/tac2entries.mli b/user-contrib/Ltac2/tac2entries.mli new file mode 100644 index 0000000000..d493192bb3 --- /dev/null +++ b/user-contrib/Ltac2/tac2entries.mli @@ -0,0 +1,93 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open Libnames +open Tac2expr + +(** {5 Toplevel definitions} *) + +val register_ltac : ?local:bool -> ?mut:bool -> rec_flag -> + (Names.lname * raw_tacexpr) list -> unit + +val register_type : ?local:bool -> rec_flag -> + (qualid * redef_flag * raw_quant_typedef) list -> unit + +val register_primitive : ?local:bool -> + Names.lident -> raw_typexpr -> ml_tactic_name -> unit + +val register_struct + : ?local:bool + -> pstate:Proof_global.t option + -> strexpr + -> unit + +val register_notation : ?local:bool -> sexpr list -> int option -> + raw_tacexpr -> unit + +(** {5 Notations} *) + +type scope_rule = +| ScopeRule : (raw_tacexpr, _, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule + +type scope_interpretation = sexpr list -> scope_rule + +val register_scope : Id.t -> scope_interpretation -> unit +(** Create a new scope with the provided name *) + +val parse_scope : sexpr -> scope_rule +(** Use this to interpret the subscopes for interpretation functions *) + +(** {5 Inspecting} *) + +val print_ltac : Libnames.qualid -> unit + +(** {5 Eval loop} *) + +(** Evaluate a tactic expression in the current environment *) +val call : pstate:Proof_global.t -> default:bool -> raw_tacexpr -> Proof_global.t + +(** {5 Toplevel exceptions} *) + +val backtrace : backtrace Exninfo.t + +(** {5 Parsing entries} *) + +module Pltac : +sig +val tac2expr : raw_tacexpr Pcoq.Entry.t + +(** Quoted entries. To be used for complex notations. *) + +open Tac2qexpr + +val q_ident : Id.t CAst.t or_anti Pcoq.Entry.t +val q_bindings : bindings Pcoq.Entry.t +val q_with_bindings : bindings Pcoq.Entry.t +val q_intropattern : intro_pattern Pcoq.Entry.t +val q_intropatterns : intro_pattern list CAst.t Pcoq.Entry.t +val q_destruction_arg : destruction_arg Pcoq.Entry.t +val q_induction_clause : induction_clause Pcoq.Entry.t +val q_conversion : conversion Pcoq.Entry.t +val q_rewriting : rewriting Pcoq.Entry.t +val q_clause : clause Pcoq.Entry.t +val q_dispatch : dispatch Pcoq.Entry.t +val q_occurrences : occurrences Pcoq.Entry.t +val q_reference : reference or_anti Pcoq.Entry.t +val q_strategy_flag : strategy_flag Pcoq.Entry.t +val q_constr_matching : constr_matching Pcoq.Entry.t +val q_goal_matching : goal_matching Pcoq.Entry.t +val q_hintdb : hintdb Pcoq.Entry.t +val q_move_location : move_location Pcoq.Entry.t +val q_pose : pose Pcoq.Entry.t +val q_assert : assertion Pcoq.Entry.t +end + +(** {5 Hooks} *) + +val register_constr_quotations : (unit -> unit) Hook.t diff --git a/user-contrib/Ltac2/tac2env.ml b/user-contrib/Ltac2/tac2env.ml new file mode 100644 index 0000000000..93ad57e97e --- /dev/null +++ b/user-contrib/Ltac2/tac2env.ml @@ -0,0 +1,298 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Util +open Names +open Libnames +open Tac2expr +open Tac2ffi + +type global_data = { + gdata_expr : glb_tacexpr; + gdata_type : type_scheme; + gdata_mutable : bool; +} + +type constructor_data = { + cdata_prms : int; + cdata_type : type_constant; + cdata_args : int glb_typexpr list; + cdata_indx : int option; +} + +type projection_data = { + pdata_prms : int; + pdata_type : type_constant; + pdata_ptyp : int glb_typexpr; + pdata_mutb : bool; + pdata_indx : int; +} + +type ltac_state = { + ltac_tactics : global_data KNmap.t; + ltac_constructors : constructor_data KNmap.t; + ltac_projections : projection_data KNmap.t; + ltac_types : glb_quant_typedef KNmap.t; + ltac_aliases : raw_tacexpr KNmap.t; +} + +let empty_state = { + ltac_tactics = KNmap.empty; + ltac_constructors = KNmap.empty; + ltac_projections = KNmap.empty; + ltac_types = KNmap.empty; + ltac_aliases = KNmap.empty; +} + +let ltac_state = Summary.ref empty_state ~name:"ltac2-state" + +let define_global kn e = + let state = !ltac_state in + ltac_state := { state with ltac_tactics = KNmap.add kn e state.ltac_tactics } + +let interp_global kn = + let data = KNmap.find kn ltac_state.contents.ltac_tactics in + data + +let define_constructor kn t = + let state = !ltac_state in + ltac_state := { state with ltac_constructors = KNmap.add kn t state.ltac_constructors } + +let interp_constructor kn = KNmap.find kn ltac_state.contents.ltac_constructors + +let define_projection kn t = + let state = !ltac_state in + ltac_state := { state with ltac_projections = KNmap.add kn t state.ltac_projections } + +let interp_projection kn = KNmap.find kn ltac_state.contents.ltac_projections + +let define_type kn e = + let state = !ltac_state in + ltac_state := { state with ltac_types = KNmap.add kn e state.ltac_types } + +let interp_type kn = KNmap.find kn ltac_state.contents.ltac_types + +let define_alias kn tac = + let state = !ltac_state in + ltac_state := { state with ltac_aliases = KNmap.add kn tac state.ltac_aliases } + +let interp_alias kn = KNmap.find kn ltac_state.contents.ltac_aliases + +module ML = +struct + type t = ml_tactic_name + let compare n1 n2 = + let c = String.compare n1.mltac_plugin n2.mltac_plugin in + if Int.equal c 0 then String.compare n1.mltac_tactic n2.mltac_tactic + else c +end + +module MLMap = Map.Make(ML) + +let primitive_map = ref MLMap.empty + +let define_primitive name f = primitive_map := MLMap.add name f !primitive_map +let interp_primitive name = MLMap.find name !primitive_map + +(** Name management *) + +module FullPath = +struct + type t = full_path + let equal = eq_full_path + let to_string = string_of_path + let repr sp = + let dir,id = repr_path sp in + id, (DirPath.repr dir) +end + +type tacref = Tac2expr.tacref = +| TacConstant of ltac_constant +| TacAlias of ltac_alias + +module TacRef = +struct +type t = tacref +let compare r1 r2 = match r1, r2 with +| TacConstant c1, TacConstant c2 -> KerName.compare c1 c2 +| TacAlias c1, TacAlias c2 -> KerName.compare c1 c2 +| TacConstant _, TacAlias _ -> -1 +| TacAlias _, TacConstant _ -> 1 + +let equal r1 r2 = compare r1 r2 == 0 + +end + +module KnTab = Nametab.Make(FullPath)(KerName) +module RfTab = Nametab.Make(FullPath)(TacRef) +module RfMap = Map.Make(TacRef) + +type nametab = { + tab_ltac : RfTab.t; + tab_ltac_rev : full_path RfMap.t; + tab_cstr : KnTab.t; + tab_cstr_rev : full_path KNmap.t; + tab_type : KnTab.t; + tab_type_rev : full_path KNmap.t; + tab_proj : KnTab.t; + tab_proj_rev : full_path KNmap.t; +} + +let empty_nametab = { + tab_ltac = RfTab.empty; + tab_ltac_rev = RfMap.empty; + tab_cstr = KnTab.empty; + tab_cstr_rev = KNmap.empty; + tab_type = KnTab.empty; + tab_type_rev = KNmap.empty; + tab_proj = KnTab.empty; + tab_proj_rev = KNmap.empty; +} + +let nametab = Summary.ref empty_nametab ~name:"ltac2-nametab" + +let push_ltac vis sp kn = + let tab = !nametab in + let tab_ltac = RfTab.push vis sp kn tab.tab_ltac in + let tab_ltac_rev = RfMap.add kn sp tab.tab_ltac_rev in + nametab := { tab with tab_ltac; tab_ltac_rev } + +let locate_ltac qid = + let tab = !nametab in + RfTab.locate qid tab.tab_ltac + +let locate_extended_all_ltac qid = + let tab = !nametab in + RfTab.find_prefixes qid tab.tab_ltac + +let shortest_qualid_of_ltac kn = + let tab = !nametab in + let sp = RfMap.find kn tab.tab_ltac_rev in + RfTab.shortest_qualid Id.Set.empty sp tab.tab_ltac + +let push_constructor vis sp kn = + let tab = !nametab in + let tab_cstr = KnTab.push vis sp kn tab.tab_cstr in + let tab_cstr_rev = KNmap.add kn sp tab.tab_cstr_rev in + nametab := { tab with tab_cstr; tab_cstr_rev } + +let locate_constructor qid = + let tab = !nametab in + KnTab.locate qid tab.tab_cstr + +let locate_extended_all_constructor qid = + let tab = !nametab in + KnTab.find_prefixes qid tab.tab_cstr + +let shortest_qualid_of_constructor kn = + let tab = !nametab in + let sp = KNmap.find kn tab.tab_cstr_rev in + KnTab.shortest_qualid Id.Set.empty sp tab.tab_cstr + +let push_type vis sp kn = + let tab = !nametab in + let tab_type = KnTab.push vis sp kn tab.tab_type in + let tab_type_rev = KNmap.add kn sp tab.tab_type_rev in + nametab := { tab with tab_type; tab_type_rev } + +let locate_type qid = + let tab = !nametab in + KnTab.locate qid tab.tab_type + +let locate_extended_all_type qid = + let tab = !nametab in + KnTab.find_prefixes qid tab.tab_type + +let shortest_qualid_of_type ?loc kn = + let tab = !nametab in + let sp = KNmap.find kn tab.tab_type_rev in + KnTab.shortest_qualid ?loc Id.Set.empty sp tab.tab_type + +let push_projection vis sp kn = + let tab = !nametab in + let tab_proj = KnTab.push vis sp kn tab.tab_proj in + let tab_proj_rev = KNmap.add kn sp tab.tab_proj_rev in + nametab := { tab with tab_proj; tab_proj_rev } + +let locate_projection qid = + let tab = !nametab in + KnTab.locate qid tab.tab_proj + +let locate_extended_all_projection qid = + let tab = !nametab in + KnTab.find_prefixes qid tab.tab_proj + +let shortest_qualid_of_projection kn = + let tab = !nametab in + let sp = KNmap.find kn tab.tab_proj_rev in + KnTab.shortest_qualid Id.Set.empty sp tab.tab_proj + +type 'a or_glb_tacexpr = +| GlbVal of 'a +| GlbTacexpr of glb_tacexpr + +type environment = { + env_ist : valexpr Id.Map.t; +} + +type ('a, 'b, 'r) intern_fun = Genintern.glob_sign -> 'a -> 'b * 'r glb_typexpr + +type ('a, 'b) ml_object = { + ml_intern : 'r. (raw_tacexpr, glb_tacexpr, 'r) intern_fun -> ('a, 'b or_glb_tacexpr, 'r) intern_fun; + ml_subst : Mod_subst.substitution -> 'b -> 'b; + ml_interp : environment -> 'b -> valexpr Proofview.tactic; + ml_print : Environ.env -> 'b -> Pp.t; +} + +module MLTypeObj = +struct + type ('a, 'b) t = ('a, 'b) ml_object +end + +module MLType = Tac2dyn.ArgMap(MLTypeObj) + +let ml_object_table = ref MLType.empty + +let define_ml_object t tpe = + ml_object_table := MLType.add t (MLType.Pack tpe) !ml_object_table + +let interp_ml_object t = + try + let MLType.Pack ans = MLType.find t !ml_object_table in + ans + with Not_found -> + CErrors.anomaly Pp.(str "Unknown object type " ++ str (Tac2dyn.Arg.repr t)) + +(** Absolute paths *) + +let coq_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Init"; "Ltac2"])) + +let std_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Std"; "Ltac2"])) + +let ltac1_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Ltac1"; "Ltac2"])) + +(** Generic arguments *) + +let wit_ltac2 = Genarg.make0 "ltac2:value" +let wit_ltac2_quotation = Genarg.make0 "ltac2:quotation" +let () = Geninterp.register_val0 wit_ltac2 None +let () = Geninterp.register_val0 wit_ltac2_quotation None + +let is_constructor qid = + let (_, id) = repr_qualid qid in + let id = Id.to_string id in + assert (String.length id > 0); + match id with + | "true" | "false" -> true (* built-in constructors *) + | _ -> + match id.[0] with + | 'A'..'Z' -> true + | _ -> false diff --git a/user-contrib/Ltac2/tac2env.mli b/user-contrib/Ltac2/tac2env.mli new file mode 100644 index 0000000000..c7e87c5432 --- /dev/null +++ b/user-contrib/Ltac2/tac2env.mli @@ -0,0 +1,146 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Genarg +open Names +open Libnames +open Nametab +open Tac2expr +open Tac2ffi + +(** Ltac2 global environment *) + +(** {5 Toplevel definition of values} *) + +type global_data = { + gdata_expr : glb_tacexpr; + gdata_type : type_scheme; + gdata_mutable : bool; +} + +val define_global : ltac_constant -> global_data -> unit +val interp_global : ltac_constant -> global_data + +(** {5 Toplevel definition of types} *) + +val define_type : type_constant -> glb_quant_typedef -> unit +val interp_type : type_constant -> glb_quant_typedef + +(** {5 Toplevel definition of algebraic constructors} *) + +type constructor_data = { + cdata_prms : int; + (** Type parameters *) + cdata_type : type_constant; + (** Inductive definition to which the constructor pertains *) + cdata_args : int glb_typexpr list; + (** Types of the constructor arguments *) + cdata_indx : int option; + (** Index of the constructor in the ADT. Numbering is duplicated between + argumentless and argument-using constructors, e.g. in type ['a option] + [None] and [Some] have both index 0. This field is empty whenever the + constructor is a member of an open type. *) +} + +val define_constructor : ltac_constructor -> constructor_data -> unit +val interp_constructor : ltac_constructor -> constructor_data + +(** {5 Toplevel definition of projections} *) + +type projection_data = { + pdata_prms : int; + (** Type parameters *) + pdata_type : type_constant; + (** Record definition to which the projection pertains *) + pdata_ptyp : int glb_typexpr; + (** Type of the projection *) + pdata_mutb : bool; + (** Whether the field is mutable *) + pdata_indx : int; + (** Index of the projection *) +} + +val define_projection : ltac_projection -> projection_data -> unit +val interp_projection : ltac_projection -> projection_data + +(** {5 Toplevel definition of aliases} *) + +val define_alias : ltac_constant -> raw_tacexpr -> unit +val interp_alias : ltac_constant -> raw_tacexpr + +(** {5 Name management} *) + +val push_ltac : visibility -> full_path -> tacref -> unit +val locate_ltac : qualid -> tacref +val locate_extended_all_ltac : qualid -> tacref list +val shortest_qualid_of_ltac : tacref -> qualid + +val push_constructor : visibility -> full_path -> ltac_constructor -> unit +val locate_constructor : qualid -> ltac_constructor +val locate_extended_all_constructor : qualid -> ltac_constructor list +val shortest_qualid_of_constructor : ltac_constructor -> qualid + +val push_type : visibility -> full_path -> type_constant -> unit +val locate_type : qualid -> type_constant +val locate_extended_all_type : qualid -> type_constant list +val shortest_qualid_of_type : ?loc:Loc.t -> type_constant -> qualid + +val push_projection : visibility -> full_path -> ltac_projection -> unit +val locate_projection : qualid -> ltac_projection +val locate_extended_all_projection : qualid -> ltac_projection list +val shortest_qualid_of_projection : ltac_projection -> qualid + +(** {5 Toplevel definitions of ML tactics} *) + +(** This state is not part of the summary, contrarily to the ones above. It is + intended to be used from ML plugins to register ML-side functions. *) + +val define_primitive : ml_tactic_name -> closure -> unit +val interp_primitive : ml_tactic_name -> closure + +(** {5 ML primitive types} *) + +type 'a or_glb_tacexpr = +| GlbVal of 'a +| GlbTacexpr of glb_tacexpr + +type ('a, 'b, 'r) intern_fun = Genintern.glob_sign -> 'a -> 'b * 'r glb_typexpr + +type environment = { + env_ist : valexpr Id.Map.t; +} + +type ('a, 'b) ml_object = { + ml_intern : 'r. (raw_tacexpr, glb_tacexpr, 'r) intern_fun -> ('a, 'b or_glb_tacexpr, 'r) intern_fun; + ml_subst : Mod_subst.substitution -> 'b -> 'b; + ml_interp : environment -> 'b -> valexpr Proofview.tactic; + ml_print : Environ.env -> 'b -> Pp.t; +} + +val define_ml_object : ('a, 'b) Tac2dyn.Arg.tag -> ('a, 'b) ml_object -> unit +val interp_ml_object : ('a, 'b) Tac2dyn.Arg.tag -> ('a, 'b) ml_object + +(** {5 Absolute paths} *) + +val coq_prefix : ModPath.t +(** Path where primitive datatypes are defined in Ltac2 plugin. *) + +val std_prefix : ModPath.t +(** Path where Ltac-specific datatypes are defined in Ltac2 plugin. *) + +val ltac1_prefix : ModPath.t +(** Path where the Ltac1 legacy FFI is defined. *) + +(** {5 Generic arguments} *) + +val wit_ltac2 : (raw_tacexpr, glb_tacexpr, Util.Empty.t) genarg_type +val wit_ltac2_quotation : (Id.t Loc.located, Id.t, Util.Empty.t) genarg_type + +(** {5 Helper functions} *) + +val is_constructor : qualid -> bool diff --git a/user-contrib/Ltac2/tac2expr.mli b/user-contrib/Ltac2/tac2expr.mli new file mode 100644 index 0000000000..1069d0bfa3 --- /dev/null +++ b/user-contrib/Ltac2/tac2expr.mli @@ -0,0 +1,190 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open Libnames + +type mutable_flag = bool +type rec_flag = bool +type redef_flag = bool +type lid = Id.t +type uid = Id.t + +type ltac_constant = KerName.t +type ltac_alias = KerName.t +type ltac_constructor = KerName.t +type ltac_projection = KerName.t +type type_constant = KerName.t + +type tacref = +| TacConstant of ltac_constant +| TacAlias of ltac_alias + +type 'a or_relid = +| RelId of qualid +| AbsKn of 'a + +(** {5 Misc} *) + +type ml_tactic_name = { + mltac_plugin : string; + mltac_tactic : string; +} + +type 'a or_tuple = +| Tuple of int +| Other of 'a + +(** {5 Type syntax} *) + +type raw_typexpr_r = +| CTypVar of Name.t +| CTypArrow of raw_typexpr * raw_typexpr +| CTypRef of type_constant or_tuple or_relid * raw_typexpr list + +and raw_typexpr = raw_typexpr_r CAst.t + +type raw_typedef = +| CTydDef of raw_typexpr option +| CTydAlg of (uid * raw_typexpr list) list +| CTydRec of (lid * mutable_flag * raw_typexpr) list +| CTydOpn + +type 'a glb_typexpr = +| GTypVar of 'a +| GTypArrow of 'a glb_typexpr * 'a glb_typexpr +| GTypRef of type_constant or_tuple * 'a glb_typexpr list + +type glb_alg_type = { + galg_constructors : (uid * int glb_typexpr list) list; + (** Constructors of the algebraic type *) + galg_nconst : int; + (** Number of constant constructors *) + galg_nnonconst : int; + (** Number of non-constant constructors *) +} + +type glb_typedef = +| GTydDef of int glb_typexpr option +| GTydAlg of glb_alg_type +| GTydRec of (lid * mutable_flag * int glb_typexpr) list +| GTydOpn + +type type_scheme = int * int glb_typexpr + +type raw_quant_typedef = Names.lident list * raw_typedef +type glb_quant_typedef = int * glb_typedef + +(** {5 Term syntax} *) + +type atom = +| AtmInt of int +| AtmStr of string + +(** Tactic expressions *) +type raw_patexpr_r = +| CPatVar of Name.t +| CPatRef of ltac_constructor or_tuple or_relid * raw_patexpr list +| CPatCnv of raw_patexpr * raw_typexpr + +and raw_patexpr = raw_patexpr_r CAst.t + +type raw_tacexpr_r = +| CTacAtm of atom +| CTacRef of tacref or_relid +| CTacCst of ltac_constructor or_tuple or_relid +| CTacFun of raw_patexpr list * raw_tacexpr +| CTacApp of raw_tacexpr * raw_tacexpr list +| CTacLet of rec_flag * (raw_patexpr * raw_tacexpr) list * raw_tacexpr +| CTacCnv of raw_tacexpr * raw_typexpr +| CTacSeq of raw_tacexpr * raw_tacexpr +| CTacCse of raw_tacexpr * raw_taccase list +| CTacRec of raw_recexpr +| CTacPrj of raw_tacexpr * ltac_projection or_relid +| CTacSet of raw_tacexpr * ltac_projection or_relid * raw_tacexpr +| CTacExt : ('a, _) Tac2dyn.Arg.tag * 'a -> raw_tacexpr_r + +and raw_tacexpr = raw_tacexpr_r CAst.t + +and raw_taccase = raw_patexpr * raw_tacexpr + +and raw_recexpr = (ltac_projection or_relid * raw_tacexpr) list + +type case_info = type_constant or_tuple + +type 'a open_match = { + opn_match : 'a; + opn_branch : (Name.t * Name.t array * 'a) KNmap.t; + (** Invariant: should not be empty *) + opn_default : Name.t * 'a; +} + +type glb_tacexpr = +| GTacAtm of atom +| GTacVar of Id.t +| GTacRef of ltac_constant +| GTacFun of Name.t list * glb_tacexpr +| GTacApp of glb_tacexpr * glb_tacexpr list +| GTacLet of rec_flag * (Name.t * glb_tacexpr) list * glb_tacexpr +| GTacCst of case_info * int * glb_tacexpr list +| GTacCse of glb_tacexpr * case_info * glb_tacexpr array * (Name.t array * glb_tacexpr) array +| GTacPrj of type_constant * glb_tacexpr * int +| GTacSet of type_constant * glb_tacexpr * int * glb_tacexpr +| GTacOpn of ltac_constructor * glb_tacexpr list +| GTacWth of glb_tacexpr open_match +| GTacExt : (_, 'a) Tac2dyn.Arg.tag * 'a -> glb_tacexpr +| GTacPrm of ml_tactic_name * glb_tacexpr list + +(** {5 Parsing & Printing} *) + +type exp_level = +| E5 +| E4 +| E3 +| E2 +| E1 +| E0 + +type sexpr = +| SexprStr of string CAst.t +| SexprInt of int CAst.t +| SexprRec of Loc.t * Id.t option CAst.t * sexpr list + +(** {5 Toplevel statements} *) + +type strexpr = +| StrVal of mutable_flag * rec_flag * (Names.lname * raw_tacexpr) list + (** Term definition *) +| StrTyp of rec_flag * (qualid * redef_flag * raw_quant_typedef) list + (** Type definition *) +| StrPrm of Names.lident * raw_typexpr * ml_tactic_name + (** External definition *) +| StrSyn of sexpr list * int option * raw_tacexpr + (** Syntactic extensions *) +| StrMut of qualid * raw_tacexpr + (** Redefinition of mutable globals *) +| StrRun of raw_tacexpr + (** Toplevel evaluation of an expression *) + +(** {5 Dynamic semantics} *) + +(** Values are represented in a way similar to OCaml, i.e. they constrast + immediate integers (integers, constructors without arguments) and structured + blocks (tuples, arrays, constructors with arguments), as well as a few other + base cases, namely closures, strings, named constructors, and dynamic type + coming from the Coq implementation. *) + +type tag = int + +type frame = +| FrLtac of ltac_constant +| FrAnon of glb_tacexpr +| FrPrim of ml_tactic_name +| FrExtn : ('a, 'b) Tac2dyn.Arg.tag * 'b -> frame + +type backtrace = frame list diff --git a/user-contrib/Ltac2/tac2extffi.ml b/user-contrib/Ltac2/tac2extffi.ml new file mode 100644 index 0000000000..315c970f9e --- /dev/null +++ b/user-contrib/Ltac2/tac2extffi.ml @@ -0,0 +1,40 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Tac2ffi +open Tac2types + +module Value = Tac2ffi + +(** Make a representation with a dummy from function *) +let make_to_repr f = Tac2ffi.make_repr (fun _ -> assert false) f + +(** More ML representations *) + +let to_qhyp v = match Value.to_block v with +| (0, [| i |]) -> AnonHyp (Value.to_int i) +| (1, [| id |]) -> NamedHyp (Value.to_ident id) +| _ -> assert false + +let qhyp = make_to_repr to_qhyp + +let to_bindings = function +| ValInt 0 -> NoBindings +| ValBlk (0, [| vl |]) -> + ImplicitBindings (Value.to_list Value.to_constr vl) +| ValBlk (1, [| vl |]) -> + ExplicitBindings ((Value.to_list (fun p -> to_pair to_qhyp Value.to_constr p) vl)) +| _ -> assert false + +let bindings = make_to_repr to_bindings + +let to_constr_with_bindings v = match Value.to_tuple v with +| [| c; bnd |] -> (Value.to_constr c, to_bindings bnd) +| _ -> assert false + +let constr_with_bindings = make_to_repr to_constr_with_bindings diff --git a/user-contrib/Ltac2/tac2extffi.mli b/user-contrib/Ltac2/tac2extffi.mli new file mode 100644 index 0000000000..f5251c3d0d --- /dev/null +++ b/user-contrib/Ltac2/tac2extffi.mli @@ -0,0 +1,16 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Tac2ffi +open Tac2types + +val qhyp : quantified_hypothesis repr + +val bindings : bindings repr + +val constr_with_bindings : constr_with_bindings repr diff --git a/user-contrib/Ltac2/tac2ffi.ml b/user-contrib/Ltac2/tac2ffi.ml new file mode 100644 index 0000000000..e3127ab9df --- /dev/null +++ b/user-contrib/Ltac2/tac2ffi.ml @@ -0,0 +1,382 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Util +open Names +open Globnames +open Tac2dyn +open Tac2expr +open Proofview.Notations + +type ('a, _) arity0 = +| OneAty : ('a, 'a -> 'a Proofview.tactic) arity0 +| AddAty : ('a, 'b) arity0 -> ('a, 'a -> 'b) arity0 + +type valexpr = +| ValInt of int + (** Immediate integers *) +| ValBlk of tag * valexpr array + (** Structured blocks *) +| ValStr of Bytes.t + (** Strings *) +| ValCls of closure + (** Closures *) +| ValOpn of KerName.t * valexpr array + (** Open constructors *) +| ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr + (** Arbitrary data *) + +and closure = MLTactic : (valexpr, 'v) arity0 * 'v -> closure + +let arity_one = OneAty +let arity_suc a = AddAty a + +type 'a arity = (valexpr, 'a) arity0 + +let mk_closure arity f = MLTactic (arity, f) + +module Valexpr = +struct + +type t = valexpr + +let is_int = function +| ValInt _ -> true +| ValBlk _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> false + +let tag v = match v with +| ValBlk (n, _) -> n +| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> + CErrors.anomaly (Pp.str "Unexpected value shape") + +let field v n = match v with +| ValBlk (_, v) -> v.(n) +| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> + CErrors.anomaly (Pp.str "Unexpected value shape") + +let set_field v n w = match v with +| ValBlk (_, v) -> v.(n) <- w +| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ -> + CErrors.anomaly (Pp.str "Unexpected value shape") + +let make_block tag v = ValBlk (tag, v) +let make_int n = ValInt n + +end + +type 'a repr = { + r_of : 'a -> valexpr; + r_to : valexpr -> 'a; + r_id : bool; +} + +let repr_of r x = r.r_of x +let repr_to r x = r.r_to x + +let make_repr r_of r_to = { r_of; r_to; r_id = false; } + +(** Dynamic tags *) + +let val_exn = Val.create "exn" +let val_constr = Val.create "constr" +let val_ident = Val.create "ident" +let val_pattern = Val.create "pattern" +let val_pp = Val.create "pp" +let val_sort = Val.create "sort" +let val_cast = Val.create "cast" +let val_inductive = Val.create "inductive" +let val_constant = Val.create "constant" +let val_constructor = Val.create "constructor" +let val_projection = Val.create "projection" +let val_case = Val.create "case" +let val_univ = Val.create "universe" +let val_free : Names.Id.Set.t Val.tag = Val.create "free" +let val_ltac1 : Geninterp.Val.t Val.tag = Val.create "ltac1" + +let extract_val (type a) (type b) (tag : a Val.tag) (tag' : b Val.tag) (v : b) : a = +match Val.eq tag tag' with +| None -> assert false +| Some Refl -> v + +(** Exception *) + +exception LtacError of KerName.t * valexpr array + +(** Conversion functions *) + +let valexpr = { + r_of = (fun obj -> obj); + r_to = (fun obj -> obj); + r_id = true; +} + +let of_unit () = ValInt 0 + +let to_unit = function +| ValInt 0 -> () +| _ -> assert false + +let unit = { + r_of = of_unit; + r_to = to_unit; + r_id = false; +} + +let of_int n = ValInt n +let to_int = function +| ValInt n -> n +| _ -> assert false + +let int = { + r_of = of_int; + r_to = to_int; + r_id = false; +} + +let of_bool b = if b then ValInt 0 else ValInt 1 + +let to_bool = function +| ValInt 0 -> true +| ValInt 1 -> false +| _ -> assert false + +let bool = { + r_of = of_bool; + r_to = to_bool; + r_id = false; +} + +let of_char n = ValInt (Char.code n) +let to_char = function +| ValInt n -> Char.chr n +| _ -> assert false + +let char = { + r_of = of_char; + r_to = to_char; + r_id = false; +} + +let of_string s = ValStr s +let to_string = function +| ValStr s -> s +| _ -> assert false + +let string = { + r_of = of_string; + r_to = to_string; + r_id = false; +} + +let rec of_list f = function +| [] -> ValInt 0 +| x :: l -> ValBlk (0, [| f x; of_list f l |]) + +let rec to_list f = function +| ValInt 0 -> [] +| ValBlk (0, [|v; vl|]) -> f v :: to_list f vl +| _ -> assert false + +let list r = { + r_of = (fun l -> of_list r.r_of l); + r_to = (fun l -> to_list r.r_to l); + r_id = false; +} + +let of_closure cls = ValCls cls + +let to_closure = function +| ValCls cls -> cls +| ValExt _ | ValInt _ | ValBlk _ | ValStr _ | ValOpn _ -> assert false + +let closure = { + r_of = of_closure; + r_to = to_closure; + r_id = false; +} + +let of_ext tag c = + ValExt (tag, c) + +let to_ext tag = function +| ValExt (tag', e) -> extract_val tag tag' e +| _ -> assert false + +let repr_ext tag = { + r_of = (fun e -> of_ext tag e); + r_to = (fun e -> to_ext tag e); + r_id = false; +} + +let of_constr c = of_ext val_constr c +let to_constr c = to_ext val_constr c +let constr = repr_ext val_constr + +let of_ident c = of_ext val_ident c +let to_ident c = to_ext val_ident c +let ident = repr_ext val_ident + +let of_pattern c = of_ext val_pattern c +let to_pattern c = to_ext val_pattern c +let pattern = repr_ext val_pattern + +let internal_err = + let open Names in + let coq_prefix = + MPfile (DirPath.make (List.map Id.of_string ["Init"; "Ltac2"])) + in + KerName.make coq_prefix (Label.of_id (Id.of_string "Internal")) + +(** FIXME: handle backtrace in Ltac2 exceptions *) +let of_exn c = match fst c with +| LtacError (kn, c) -> ValOpn (kn, c) +| _ -> ValOpn (internal_err, [|of_ext val_exn c|]) + +let to_exn c = match c with +| ValOpn (kn, c) -> + if Names.KerName.equal kn internal_err then + to_ext val_exn c.(0) + else + (LtacError (kn, c), Exninfo.null) +| _ -> assert false + +let exn = { + r_of = of_exn; + r_to = to_exn; + r_id = false; +} + +let of_option f = function +| None -> ValInt 0 +| Some c -> ValBlk (0, [|f c|]) + +let to_option f = function +| ValInt 0 -> None +| ValBlk (0, [|c|]) -> Some (f c) +| _ -> assert false + +let option r = { + r_of = (fun l -> of_option r.r_of l); + r_to = (fun l -> to_option r.r_to l); + r_id = false; +} + +let of_pp c = of_ext val_pp c +let to_pp c = to_ext val_pp c +let pp = repr_ext val_pp + +let of_tuple cl = ValBlk (0, cl) +let to_tuple = function +| ValBlk (0, cl) -> cl +| _ -> assert false + +let of_pair f g (x, y) = ValBlk (0, [|f x; g y|]) +let to_pair f g = function +| ValBlk (0, [|x; y|]) -> (f x, g y) +| _ -> assert false +let pair r0 r1 = { + r_of = (fun p -> of_pair r0.r_of r1.r_of p); + r_to = (fun p -> to_pair r0.r_to r1.r_to p); + r_id = false; +} + +let of_array f vl = ValBlk (0, Array.map f vl) +let to_array f = function +| ValBlk (0, vl) -> Array.map f vl +| _ -> assert false +let array r = { + r_of = (fun l -> of_array r.r_of l); + r_to = (fun l -> to_array r.r_to l); + r_id = false; +} + +let of_block (n, args) = ValBlk (n, args) +let to_block = function +| ValBlk (n, args) -> (n, args) +| _ -> assert false + +let block = { + r_of = of_block; + r_to = to_block; + r_id = false; +} + +let of_open (kn, args) = ValOpn (kn, args) + +let to_open = function +| ValOpn (kn, args) -> (kn, args) +| _ -> assert false + +let open_ = { + r_of = of_open; + r_to = to_open; + r_id = false; +} + +let of_constant c = of_ext val_constant c +let to_constant c = to_ext val_constant c +let constant = repr_ext val_constant + +let of_reference = function +| VarRef id -> ValBlk (0, [| of_ident id |]) +| ConstRef cst -> ValBlk (1, [| of_constant cst |]) +| IndRef ind -> ValBlk (2, [| of_ext val_inductive ind |]) +| ConstructRef cstr -> ValBlk (3, [| of_ext val_constructor cstr |]) + +let to_reference = function +| ValBlk (0, [| id |]) -> VarRef (to_ident id) +| ValBlk (1, [| cst |]) -> ConstRef (to_constant cst) +| ValBlk (2, [| ind |]) -> IndRef (to_ext val_inductive ind) +| ValBlk (3, [| cstr |]) -> ConstructRef (to_ext val_constructor cstr) +| _ -> assert false + +let reference = { + r_of = of_reference; + r_to = to_reference; + r_id = false; +} + +type ('a, 'b) fun1 = closure + +let fun1 (r0 : 'a repr) (r1 : 'b repr) : ('a, 'b) fun1 repr = closure +let to_fun1 r0 r1 f = to_closure f + +let rec apply : type a. a arity -> a -> valexpr list -> valexpr Proofview.tactic = + fun arity f args -> match args, arity with + | [], arity -> Proofview.tclUNIT (ValCls (MLTactic (arity, f))) + (* A few hardcoded cases for efficiency *) + | [a0], OneAty -> f a0 + | [a0; a1], AddAty OneAty -> f a0 a1 + | [a0; a1; a2], AddAty (AddAty OneAty) -> f a0 a1 a2 + | [a0; a1; a2; a3], AddAty (AddAty (AddAty OneAty)) -> f a0 a1 a2 a3 + (* Generic cases *) + | a :: args, OneAty -> + f a >>= fun f -> + let MLTactic (arity, f) = to_closure f in + apply arity f args + | a :: args, AddAty arity -> + apply arity (f a) args + +let apply (MLTactic (arity, f)) args = apply arity f args + +type n_closure = +| NClosure : 'a arity * (valexpr list -> 'a) -> n_closure + +let rec abstract n f = + if Int.equal n 1 then NClosure (OneAty, fun accu v -> f (List.rev (v :: accu))) + else + let NClosure (arity, fe) = abstract (n - 1) f in + NClosure (AddAty arity, fun accu v -> fe (v :: accu)) + +let abstract n f = + let () = assert (n > 0) in + let NClosure (arity, f) = abstract n f in + MLTactic (arity, f []) + +let app_fun1 cls r0 r1 x = + apply cls [r0.r_of x] >>= fun v -> Proofview.tclUNIT (r1.r_to v) diff --git a/user-contrib/Ltac2/tac2ffi.mli b/user-contrib/Ltac2/tac2ffi.mli new file mode 100644 index 0000000000..bfc93d99e6 --- /dev/null +++ b/user-contrib/Ltac2/tac2ffi.mli @@ -0,0 +1,189 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open EConstr +open Tac2dyn +open Tac2expr + +(** {5 Toplevel values} *) + +type closure + +type valexpr = +| ValInt of int + (** Immediate integers *) +| ValBlk of tag * valexpr array + (** Structured blocks *) +| ValStr of Bytes.t + (** Strings *) +| ValCls of closure + (** Closures *) +| ValOpn of KerName.t * valexpr array + (** Open constructors *) +| ValExt : 'a Tac2dyn.Val.tag * 'a -> valexpr + (** Arbitrary data *) + +type 'a arity + +val arity_one : (valexpr -> valexpr Proofview.tactic) arity +val arity_suc : 'a arity -> (valexpr -> 'a) arity + +val mk_closure : 'v arity -> 'v -> closure + +module Valexpr : +sig + type t = valexpr + val is_int : t -> bool + val tag : t -> int + val field : t -> int -> t + val set_field : t -> int -> t -> unit + val make_block : int -> t array -> t + val make_int : int -> t +end + +(** {5 Ltac2 FFI} *) + +type 'a repr + +val repr_of : 'a repr -> 'a -> valexpr +val repr_to : 'a repr -> valexpr -> 'a + +val make_repr : ('a -> valexpr) -> (valexpr -> 'a) -> 'a repr + +(** These functions allow to convert back and forth between OCaml and Ltac2 + data representation. The [to_*] functions raise an anomaly whenever the data + has not expected shape. *) + +val of_unit : unit -> valexpr +val to_unit : valexpr -> unit +val unit : unit repr + +val of_int : int -> valexpr +val to_int : valexpr -> int +val int : int repr + +val of_bool : bool -> valexpr +val to_bool : valexpr -> bool +val bool : bool repr + +val of_char : char -> valexpr +val to_char : valexpr -> char +val char : char repr + +val of_string : Bytes.t -> valexpr +val to_string : valexpr -> Bytes.t +val string : Bytes.t repr + +val of_list : ('a -> valexpr) -> 'a list -> valexpr +val to_list : (valexpr -> 'a) -> valexpr -> 'a list +val list : 'a repr -> 'a list repr + +val of_constr : EConstr.t -> valexpr +val to_constr : valexpr -> EConstr.t +val constr : EConstr.t repr + +val of_exn : Exninfo.iexn -> valexpr +val to_exn : valexpr -> Exninfo.iexn +val exn : Exninfo.iexn repr + +val of_ident : Id.t -> valexpr +val to_ident : valexpr -> Id.t +val ident : Id.t repr + +val of_closure : closure -> valexpr +val to_closure : valexpr -> closure +val closure : closure repr + +val of_block : (int * valexpr array) -> valexpr +val to_block : valexpr -> (int * valexpr array) +val block : (int * valexpr array) repr + +val of_array : ('a -> valexpr) -> 'a array -> valexpr +val to_array : (valexpr -> 'a) -> valexpr -> 'a array +val array : 'a repr -> 'a array repr + +val of_tuple : valexpr array -> valexpr +val to_tuple : valexpr -> valexpr array + +val of_pair : ('a -> valexpr) -> ('b -> valexpr) -> 'a * 'b -> valexpr +val to_pair : (valexpr -> 'a) -> (valexpr -> 'b) -> valexpr -> 'a * 'b +val pair : 'a repr -> 'b repr -> ('a * 'b) repr + +val of_option : ('a -> valexpr) -> 'a option -> valexpr +val to_option : (valexpr -> 'a) -> valexpr -> 'a option +val option : 'a repr -> 'a option repr + +val of_pattern : Pattern.constr_pattern -> valexpr +val to_pattern : valexpr -> Pattern.constr_pattern +val pattern : Pattern.constr_pattern repr + +val of_pp : Pp.t -> valexpr +val to_pp : valexpr -> Pp.t +val pp : Pp.t repr + +val of_constant : Constant.t -> valexpr +val to_constant : valexpr -> Constant.t +val constant : Constant.t repr + +val of_reference : GlobRef.t -> valexpr +val to_reference : valexpr -> GlobRef.t +val reference : GlobRef.t repr + +val of_ext : 'a Val.tag -> 'a -> valexpr +val to_ext : 'a Val.tag -> valexpr -> 'a +val repr_ext : 'a Val.tag -> 'a repr + +val of_open : KerName.t * valexpr array -> valexpr +val to_open : valexpr -> KerName.t * valexpr array +val open_ : (KerName.t * valexpr array) repr + +type ('a, 'b) fun1 + +val app_fun1 : ('a, 'b) fun1 -> 'a repr -> 'b repr -> 'a -> 'b Proofview.tactic + +val to_fun1 : 'a repr -> 'b repr -> valexpr -> ('a, 'b) fun1 +val fun1 : 'a repr -> 'b repr -> ('a, 'b) fun1 repr + +val valexpr : valexpr repr + +(** {5 Dynamic tags} *) + +val val_constr : EConstr.t Val.tag +val val_ident : Id.t Val.tag +val val_pattern : Pattern.constr_pattern Val.tag +val val_pp : Pp.t Val.tag +val val_sort : ESorts.t Val.tag +val val_cast : Constr.cast_kind Val.tag +val val_inductive : inductive Val.tag +val val_constant : Constant.t Val.tag +val val_constructor : constructor Val.tag +val val_projection : Projection.t Val.tag +val val_case : Constr.case_info Val.tag +val val_univ : Univ.Level.t Val.tag +val val_free : Id.Set.t Val.tag +val val_ltac1 : Geninterp.Val.t Val.tag + +val val_exn : Exninfo.iexn Tac2dyn.Val.tag +(** Toplevel representation of OCaml exceptions. Invariant: no [LtacError] + should be put into a value with tag [val_exn]. *) + +(** Closures *) + +val apply : closure -> valexpr list -> valexpr Proofview.tactic +(** Given a closure, apply it to some arguments. Handling of argument mismatches + is done automatically, i.e. in case of over or under-application. *) + +val abstract : int -> (valexpr list -> valexpr Proofview.tactic) -> closure +(** Turn a fixed-arity function into a closure. The inner function is guaranteed + to be applied to a list whose size is the integer argument. *) + +(** Exception *) + +exception LtacError of KerName.t * valexpr array +(** Ltac2-defined exceptions seen from OCaml side *) diff --git a/user-contrib/Ltac2/tac2intern.ml b/user-contrib/Ltac2/tac2intern.ml new file mode 100644 index 0000000000..de99fb167f --- /dev/null +++ b/user-contrib/Ltac2/tac2intern.ml @@ -0,0 +1,1545 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Pp +open Util +open CAst +open CErrors +open Names +open Libnames +open Locus +open Tac2env +open Tac2print +open Tac2expr + +(** Hardwired types and constants *) + +let coq_type n = KerName.make Tac2env.coq_prefix (Label.make n) + +let t_int = coq_type "int" +let t_string = coq_type "string" +let t_constr = coq_type "constr" + +(** Union find *) + +module UF : +sig +type elt +type 'a t +val equal : elt -> elt -> bool +val create : unit -> 'a t +val fresh : 'a t -> elt +val find : elt -> 'a t -> (elt * 'a option) +val union : elt -> elt -> 'a t -> unit +val set : elt -> 'a -> 'a t -> unit +module Map : +sig + type key = elt + type +'a t + val empty : 'a t + val add : key -> 'a -> 'a t -> 'a t + val mem : key -> 'a t -> bool + val find : key -> 'a t -> 'a + val exists : (key -> 'a -> bool) -> 'a t -> bool +end +end += +struct +type elt = int +let equal = Int.equal +module Map = Int.Map + +type 'a node = +| Canon of int * 'a option +| Equiv of elt + +type 'a t = { + mutable uf_data : 'a node array; + mutable uf_size : int; +} + +let resize p = + if Int.equal (Array.length p.uf_data) p.uf_size then begin + let nsize = 2 * p.uf_size + 1 in + let v = Array.make nsize (Equiv 0) in + Array.blit p.uf_data 0 v 0 (Array.length p.uf_data); + p.uf_data <- v; + end + +let create () = { uf_data = [||]; uf_size = 0 } + +let fresh p = + resize p; + let n = p.uf_size in + p.uf_data.(n) <- (Canon (1, None)); + p.uf_size <- n + 1; + n + +let rec lookup n p = + let node = Array.get p.uf_data n in + match node with + | Canon (size, v) -> n, size, v + | Equiv y -> + let ((z, _, _) as res) = lookup y p in + if not (Int.equal z y) then Array.set p.uf_data n (Equiv z); + res + +let find n p = + let (x, _, v) = lookup n p in (x, v) + +let union x y p = + let ((x, size1, _) as xcan) = lookup x p in + let ((y, size2, _) as ycan) = lookup y p in + let xcan, ycan = if size1 < size2 then xcan, ycan else ycan, xcan in + let x, _, xnode = xcan in + let y, _, ynode = ycan in + assert (Option.is_empty xnode); + assert (Option.is_empty ynode); + p.uf_data.(x) <- Equiv y; + p.uf_data.(y) <- Canon (size1 + size2, None) + +let set x v p = + let (x, s, v') = lookup x p in + assert (Option.is_empty v'); + p.uf_data.(x) <- Canon (s, Some v) + +end + +type mix_var = +| GVar of UF.elt +| LVar of int + +type mix_type_scheme = int * mix_var glb_typexpr + +type environment = { + env_var : mix_type_scheme Id.Map.t; + (** Type schemes of bound variables *) + env_cst : UF.elt glb_typexpr UF.t; + (** Unification state *) + env_als : UF.elt Id.Map.t ref; + (** Map user-facing type variables to unification variables *) + env_opn : bool; + (** Accept unbound type variables *) + env_rec : (KerName.t * int) Id.Map.t; + (** Recursive type definitions *) + env_str : bool; + (** True iff in strict mode *) +} + +let empty_env () = { + env_var = Id.Map.empty; + env_cst = UF.create (); + env_als = ref Id.Map.empty; + env_opn = true; + env_rec = Id.Map.empty; + env_str = true; +} + +let env_name env = + (* Generate names according to a provided environment *) + let mk num = + let base = num mod 26 in + let rem = num / 26 in + let name = String.make 1 (Char.chr (97 + base)) in + let suff = if Int.equal rem 0 then "" else string_of_int rem in + let name = name ^ suff in + name + in + let fold id elt acc = UF.Map.add elt (Id.to_string id) acc in + let vars = Id.Map.fold fold env.env_als.contents UF.Map.empty in + let vars = ref vars in + let rec fresh n = + let name = mk n in + if UF.Map.exists (fun _ name' -> String.equal name name') !vars then fresh (succ n) + else name + in + fun n -> + if UF.Map.mem n !vars then UF.Map.find n !vars + else + let ans = fresh 0 in + let () = vars := UF.Map.add n ans !vars in + ans + +let ltac2_env : environment Genintern.Store.field = + Genintern.Store.field () + +let drop_ltac2_env store = + Genintern.Store.remove store ltac2_env + +let fresh_id env = UF.fresh env.env_cst + +let get_alias {loc;v=id} env = + try Id.Map.find id env.env_als.contents + with Not_found -> + if env.env_opn then + let n = fresh_id env in + let () = env.env_als := Id.Map.add id n env.env_als.contents in + n + else user_err ?loc (str "Unbound type parameter " ++ Id.print id) + +let push_name id t env = match id with +| Anonymous -> env +| Name id -> { env with env_var = Id.Map.add id t env.env_var } + +let error_nargs_mismatch ?loc kn nargs nfound = + let cstr = Tac2env.shortest_qualid_of_constructor kn in + user_err ?loc (str "Constructor " ++ pr_qualid cstr ++ str " expects " ++ + int nargs ++ str " arguments, but is applied to " ++ int nfound ++ + str " arguments") + +let error_nparams_mismatch ?loc nargs nfound = + user_err ?loc (str "Type expects " ++ int nargs ++ + str " arguments, but is applied to " ++ int nfound ++ + str " arguments") + +let rec subst_type subst (t : 'a glb_typexpr) = match t with +| GTypVar id -> subst id +| GTypArrow (t1, t2) -> GTypArrow (subst_type subst t1, subst_type subst t2) +| GTypRef (qid, args) -> + GTypRef (qid, List.map (fun t -> subst_type subst t) args) + +let rec intern_type env ({loc;v=t} : raw_typexpr) : UF.elt glb_typexpr = match t with +| CTypVar (Name id) -> GTypVar (get_alias (CAst.make ?loc id) env) +| CTypVar Anonymous -> GTypVar (fresh_id env) +| CTypRef (rel, args) -> + let (kn, nparams) = match rel with + | RelId qid -> + let id = qualid_basename qid in + if qualid_is_ident qid && Id.Map.mem id env.env_rec then + let (kn, n) = Id.Map.find id env.env_rec in + (Other kn, n) + else + let kn = + try Tac2env.locate_type qid + with Not_found -> + user_err ?loc (str "Unbound type constructor " ++ pr_qualid qid) + in + let (nparams, _) = Tac2env.interp_type kn in + (Other kn, nparams) + | AbsKn (Other kn) -> + let (nparams, _) = Tac2env.interp_type kn in + (Other kn, nparams) + | AbsKn (Tuple n) -> + (Tuple n, n) + in + let nargs = List.length args in + let () = + if not (Int.equal nparams nargs) then + let qid = match rel with + | RelId lid -> lid + | AbsKn (Other kn) -> shortest_qualid_of_type ?loc kn + | AbsKn (Tuple _) -> assert false + in + user_err ?loc (strbrk "The type constructor " ++ pr_qualid qid ++ + strbrk " expects " ++ int nparams ++ strbrk " argument(s), but is here \ + applied to " ++ int nargs ++ strbrk "argument(s)") + in + GTypRef (kn, List.map (fun t -> intern_type env t) args) +| CTypArrow (t1, t2) -> GTypArrow (intern_type env t1, intern_type env t2) + +let fresh_type_scheme env (t : type_scheme) : UF.elt glb_typexpr = + let (n, t) = t in + let subst = Array.init n (fun _ -> fresh_id env) in + let substf i = GTypVar subst.(i) in + subst_type substf t + +let fresh_mix_type_scheme env (t : mix_type_scheme) : UF.elt glb_typexpr = + let (n, t) = t in + let subst = Array.init n (fun _ -> fresh_id env) in + let substf = function + | LVar i -> GTypVar subst.(i) + | GVar n -> GTypVar n + in + subst_type substf t + +let fresh_reftype env (kn : KerName.t or_tuple) = + let n = match kn with + | Other kn -> fst (Tac2env.interp_type kn) + | Tuple n -> n + in + let subst = Array.init n (fun _ -> fresh_id env) in + let t = GTypRef (kn, Array.map_to_list (fun i -> GTypVar i) subst) in + (subst, t) + +(** First-order unification algorithm *) +let is_unfoldable kn = match snd (Tac2env.interp_type kn) with +| GTydDef (Some _) -> true +| GTydDef None | GTydAlg _ | GTydRec _ | GTydOpn -> false + +let unfold env kn args = + let (nparams, def) = Tac2env.interp_type kn in + let def = match def with + | GTydDef (Some t) -> t + | _ -> assert false + in + let args = Array.of_list args in + let subst n = args.(n) in + subst_type subst def + +(** View function, allows to ensure head normal forms *) +let rec kind env t = match t with +| GTypVar id -> + let (id, v) = UF.find id env.env_cst in + begin match v with + | None -> GTypVar id + | Some t -> kind env t + end +| GTypRef (Other kn, tl) -> + if is_unfoldable kn then kind env (unfold env kn tl) else t +| GTypArrow _ | GTypRef (Tuple _, _) -> t + +(** Normalize unification variables without unfolding type aliases *) +let rec nf env t = match t with +| GTypVar id -> + let (id, v) = UF.find id env.env_cst in + begin match v with + | None -> GTypVar id + | Some t -> nf env t + end +| GTypRef (kn, tl) -> + let tl = List.map (fun t -> nf env t) tl in + GTypRef (kn, tl) +| GTypArrow (t, u) -> + let t = nf env t in + let u = nf env u in + GTypArrow (t, u) + +let pr_glbtype env t = + let t = nf env t in + let name = env_name env in + pr_glbtype name t + +exception Occur + +let rec occur_check env id t = match kind env t with +| GTypVar id' -> if UF.equal id id' then raise Occur +| GTypArrow (t1, t2) -> + let () = occur_check env id t1 in + occur_check env id t2 +| GTypRef (kn, tl) -> + List.iter (fun t -> occur_check env id t) tl + +exception CannotUnify of UF.elt glb_typexpr * UF.elt glb_typexpr + +let unify_var env id t = match kind env t with +| GTypVar id' -> + if not (UF.equal id id') then UF.union id id' env.env_cst +| GTypArrow _ | GTypRef _ -> + try + let () = occur_check env id t in + UF.set id t env.env_cst + with Occur -> raise (CannotUnify (GTypVar id, t)) + +let eq_or_tuple eq t1 t2 = match t1, t2 with +| Tuple n1, Tuple n2 -> Int.equal n1 n2 +| Other o1, Other o2 -> eq o1 o2 +| _ -> false + +let rec unify0 env t1 t2 = match kind env t1, kind env t2 with +| GTypVar id, t | t, GTypVar id -> + unify_var env id t +| GTypArrow (t1, u1), GTypArrow (t2, u2) -> + let () = unify0 env t1 t2 in + unify0 env u1 u2 +| GTypRef (kn1, tl1), GTypRef (kn2, tl2) -> + if eq_or_tuple KerName.equal kn1 kn2 then + List.iter2 (fun t1 t2 -> unify0 env t1 t2) tl1 tl2 + else raise (CannotUnify (t1, t2)) +| _ -> raise (CannotUnify (t1, t2)) + +let unify ?loc env t1 t2 = + try unify0 env t1 t2 + with CannotUnify (u1, u2) -> + user_err ?loc (str "This expression has type" ++ spc () ++ pr_glbtype env t1 ++ + spc () ++ str "but an expression was expected of type" ++ spc () ++ pr_glbtype env t2) + +let unify_arrow ?loc env ft args = + let ft0 = ft in + let rec iter ft args is_fun = match kind env ft, args with + | t, [] -> t + | GTypArrow (t1, ft), (loc, t2) :: args -> + let () = unify ?loc env t2 t1 in + iter ft args true + | GTypVar id, (_, t) :: args -> + let ft = GTypVar (fresh_id env) in + let () = unify_var env id (GTypArrow (t, ft)) in + iter ft args true + | GTypRef _, _ :: _ -> + if is_fun then + user_err ?loc (str "This function has type" ++ spc () ++ pr_glbtype env ft0 ++ + spc () ++ str "and is applied to too many arguments") + else + user_err ?loc (str "This expression has type" ++ spc () ++ pr_glbtype env ft0 ++ + spc () ++ str "and is not a function") + in + iter ft args false + +(** Term typing *) + +let is_pure_constructor kn = + match snd (Tac2env.interp_type kn) with + | GTydAlg _ | GTydOpn -> true + | GTydRec fields -> + let is_pure (_, mut, _) = not mut in + List.for_all is_pure fields + | GTydDef _ -> assert false (** Type definitions have no constructors *) + +let rec is_value = function +| GTacAtm (AtmInt _) | GTacVar _ | GTacRef _ | GTacFun _ -> true +| GTacAtm (AtmStr _) | GTacApp _ | GTacLet _ -> false +| GTacCst (Tuple _, _, el) -> List.for_all is_value el +| GTacCst (_, _, []) -> true +| GTacOpn (_, el) -> List.for_all is_value el +| GTacCst (Other kn, _, el) -> is_pure_constructor kn && List.for_all is_value el +| GTacCse _ | GTacPrj _ | GTacSet _ | GTacExt _ | GTacPrm _ +| GTacWth _ -> false + +let is_rec_rhs = function +| GTacFun _ -> true +| GTacAtm _ | GTacVar _ | GTacRef _ | GTacApp _ | GTacLet _ | GTacPrj _ +| GTacSet _ | GTacExt _ | GTacPrm _ | GTacCst _ +| GTacCse _ | GTacOpn _ | GTacWth _ -> false + +let rec fv_type f t accu = match t with +| GTypVar id -> f id accu +| GTypArrow (t1, t2) -> fv_type f t1 (fv_type f t2 accu) +| GTypRef (kn, tl) -> List.fold_left (fun accu t -> fv_type f t accu) accu tl + +let fv_env env = + let rec f id accu = match UF.find id env.env_cst with + | id, None -> UF.Map.add id () accu + | _, Some t -> fv_type f t accu + in + let fold_var id (_, t) accu = + let fmix id accu = match id with + | LVar _ -> accu + | GVar id -> f id accu + in + fv_type fmix t accu + in + let fv_var = Id.Map.fold fold_var env.env_var UF.Map.empty in + let fold_als _ id accu = f id accu in + Id.Map.fold fold_als !(env.env_als) fv_var + +let abstract_var env (t : UF.elt glb_typexpr) : mix_type_scheme = + let fv = fv_env env in + let count = ref 0 in + let vars = ref UF.Map.empty in + let rec subst id = + let (id, t) = UF.find id env.env_cst in + match t with + | None -> + if UF.Map.mem id fv then GTypVar (GVar id) + else + begin try UF.Map.find id !vars + with Not_found -> + let n = !count in + let var = GTypVar (LVar n) in + let () = incr count in + let () = vars := UF.Map.add id var !vars in + var + end + | Some t -> subst_type subst t + in + let t = subst_type subst t in + (!count, t) + +let monomorphic (t : UF.elt glb_typexpr) : mix_type_scheme = + let subst id = GTypVar (GVar id) in + (0, subst_type subst t) + +let warn_not_unit = + CWarnings.create ~name:"not-unit" ~category:"ltac" + (fun () -> strbrk "The following expression should have type unit.") + +let warn_redundant_clause = + CWarnings.create ~name:"redundant-clause" ~category:"ltac" + (fun () -> strbrk "The following clause is redundant.") + +let check_elt_unit loc env t = + let maybe_unit = match kind env t with + | GTypVar _ -> true + | GTypArrow _ -> false + | GTypRef (Tuple 0, []) -> true + | GTypRef _ -> false + in + if not maybe_unit then warn_not_unit ?loc () + +let check_elt_empty loc env t = match kind env t with +| GTypVar _ -> + user_err ?loc (str "Cannot infer an empty type for this expression") +| GTypArrow _ | GTypRef (Tuple _, _) -> + user_err ?loc (str "Type" ++ spc () ++ pr_glbtype env t ++ spc () ++ str "is not an empty type") +| GTypRef (Other kn, _) -> + let def = Tac2env.interp_type kn in + match def with + | _, GTydAlg { galg_constructors = [] } -> kn + | _ -> + user_err ?loc (str "Type" ++ spc () ++ pr_glbtype env t ++ spc () ++ str "is not an empty type") + +let check_unit ?loc t = + let env = empty_env () in + (* Should not matter, t should be closed. *) + let t = fresh_type_scheme env t in + let maybe_unit = match kind env t with + | GTypVar _ -> true + | GTypArrow _ -> false + | GTypRef (Tuple 0, []) -> true + | GTypRef _ -> false + in + if not maybe_unit then warn_not_unit ?loc () + +let check_redundant_clause = function +| [] -> () +| (p, _) :: _ -> warn_redundant_clause ?loc:p.loc () + +let get_variable0 mem var = match var with +| RelId qid -> + let id = qualid_basename qid in + if qualid_is_ident qid && mem id then ArgVar CAst.(make ?loc:qid.CAst.loc id) + else + let kn = + try Tac2env.locate_ltac qid + with Not_found -> + CErrors.user_err ?loc:qid.CAst.loc (str "Unbound value " ++ pr_qualid qid) + in + ArgArg kn +| AbsKn kn -> ArgArg kn + +let get_variable env var = + let mem id = Id.Map.mem id env.env_var in + get_variable0 mem var + +let get_constructor env var = match var with +| RelId qid -> + let c = try Some (Tac2env.locate_constructor qid) with Not_found -> None in + begin match c with + | Some knc -> Other knc + | None -> + CErrors.user_err ?loc:qid.CAst.loc (str "Unbound constructor " ++ pr_qualid qid) + end +| AbsKn knc -> knc + +let get_projection var = match var with +| RelId qid -> + let kn = try Tac2env.locate_projection qid with Not_found -> + user_err ?loc:qid.CAst.loc (pr_qualid qid ++ str " is not a projection") + in + Tac2env.interp_projection kn +| AbsKn kn -> + Tac2env.interp_projection kn + +let intern_atm env = function +| AtmInt n -> (GTacAtm (AtmInt n), GTypRef (Other t_int, [])) +| AtmStr s -> (GTacAtm (AtmStr s), GTypRef (Other t_string, [])) + +let invalid_pattern ?loc kn kn' = + let pr t = match t with + | Other kn' -> str "type " ++ pr_typref kn' + | Tuple n -> str "tuple of size " ++ int n + in + user_err ?loc (str "Invalid pattern, expected a pattern for " ++ + pr kn ++ str ", found a pattern for " ++ pr kn') (** FIXME *) + +(** Pattern view *) + +type glb_patexpr = +| GPatVar of Name.t +| GPatRef of ltac_constructor or_tuple * glb_patexpr list + +let rec intern_patexpr env {loc;v=pat} = match pat with +| CPatVar na -> GPatVar na +| CPatRef (qid, pl) -> + let kn = get_constructor env qid in + GPatRef (kn, List.map (fun p -> intern_patexpr env p) pl) +| CPatCnv (pat, ty) -> + user_err ?loc (str "Pattern not handled yet") + +type pattern_kind = +| PKind_empty +| PKind_variant of type_constant or_tuple +| PKind_open of type_constant +| PKind_any + +let get_pattern_kind env pl = match pl with +| [] -> PKind_empty +| p :: pl -> + let rec get_kind (p, _) pl = match intern_patexpr env p with + | GPatVar _ -> + begin match pl with + | [] -> PKind_any + | p :: pl -> get_kind p pl + end + | GPatRef (Other kn, pl) -> + let data = Tac2env.interp_constructor kn in + if Option.is_empty data.cdata_indx then PKind_open data.cdata_type + else PKind_variant (Other data.cdata_type) + | GPatRef (Tuple _, tp) -> PKind_variant (Tuple (List.length tp)) + in + get_kind p pl + +(** Internalization *) + +(** Used to generate a fresh tactic variable for pattern-expansion *) +let fresh_var avoid = + let bad id = + Id.Set.mem id avoid || + (try ignore (locate_ltac (qualid_of_ident id)); true with Not_found -> false) + in + Namegen.next_ident_away_from (Id.of_string "p") bad + +let add_name accu = function +| Name id -> Id.Set.add id accu +| Anonymous -> accu + +let rec ids_of_pattern accu {v=pat} = match pat with +| CPatVar Anonymous -> accu +| CPatVar (Name id) -> Id.Set.add id accu +| CPatRef (_, pl) -> + List.fold_left ids_of_pattern accu pl +| CPatCnv (pat, _) -> ids_of_pattern accu pat + +let loc_of_relid = function +| RelId {loc} -> loc +| AbsKn _ -> None + +let extract_pattern_type ({loc;v=p} as pat) = match p with +| CPatCnv (pat, ty) -> pat, Some ty +| CPatVar _ | CPatRef _ -> pat, None + +(** Expand pattern: [p => t] becomes [x => match x with p => t end] *) +let expand_pattern avoid bnd = + let fold (avoid, bnd) (pat, t) = + let na, expand = match pat.v with + | CPatVar na -> + (* Don't expand variable patterns *) + na, None + | _ -> + let id = fresh_var avoid in + let qid = RelId (qualid_of_ident ?loc:pat.loc id) in + Name id, Some qid + in + let avoid = ids_of_pattern avoid pat in + let avoid = add_name avoid na in + (avoid, (na, pat, expand) :: bnd) + in + let (_, bnd) = List.fold_left fold (avoid, []) bnd in + let fold e (na, pat, expand) = match expand with + | None -> e + | Some qid -> + let loc = loc_of_relid qid in + CAst.make ?loc @@ CTacCse (CAst.make ?loc @@ CTacRef qid, [pat, e]) + in + let expand e = List.fold_left fold e bnd in + let nas = List.rev_map (fun (na, _, _) -> na) bnd in + (nas, expand) + +let is_alias env qid = match get_variable env qid with +| ArgArg (TacAlias _) -> true +| ArgVar _ | (ArgArg (TacConstant _)) -> false + +let rec intern_rec env {loc;v=e} = match e with +| CTacAtm atm -> intern_atm env atm +| CTacRef qid -> + begin match get_variable env qid with + | ArgVar {CAst.v=id} -> + let sch = Id.Map.find id env.env_var in + (GTacVar id, fresh_mix_type_scheme env sch) + | ArgArg (TacConstant kn) -> + let { Tac2env.gdata_type = sch } = + try Tac2env.interp_global kn + with Not_found -> + CErrors.anomaly (str "Missing hardwired primitive " ++ KerName.print kn) + in + (GTacRef kn, fresh_type_scheme env sch) + | ArgArg (TacAlias kn) -> + let e = + try Tac2env.interp_alias kn + with Not_found -> + CErrors.anomaly (str "Missing hardwired alias " ++ KerName.print kn) + in + intern_rec env e + end +| CTacCst qid -> + let kn = get_constructor env qid in + intern_constructor env loc kn [] +| CTacFun (bnd, e) -> + let bnd = List.map extract_pattern_type bnd in + let map (_, t) = match t with + | None -> GTypVar (fresh_id env) + | Some t -> intern_type env t + in + let tl = List.map map bnd in + let (nas, exp) = expand_pattern (Id.Map.domain env.env_var) bnd in + let env = List.fold_left2 (fun env na t -> push_name na (monomorphic t) env) env nas tl in + let (e, t) = intern_rec env (exp e) in + let t = List.fold_right (fun t accu -> GTypArrow (t, accu)) tl t in + (GTacFun (nas, e), t) +| CTacApp ({loc;v=CTacCst qid}, args) -> + let kn = get_constructor env qid in + intern_constructor env loc kn args +| CTacApp ({v=CTacRef qid}, args) when is_alias env qid -> + let kn = match get_variable env qid with + | ArgArg (TacAlias kn) -> kn + | ArgVar _ | (ArgArg (TacConstant _)) -> assert false + in + let e = Tac2env.interp_alias kn in + let map arg = + (* Thunk alias arguments *) + let loc = arg.loc in + let t_unit = CAst.make ?loc @@ CTypRef (AbsKn (Tuple 0), []) in + let var = CAst.make ?loc @@ CPatCnv (CAst.make ?loc @@ CPatVar Anonymous, t_unit) in + CAst.make ?loc @@ CTacFun ([var], arg) + in + let args = List.map map args in + intern_rec env (CAst.make ?loc @@ CTacApp (e, args)) +| CTacApp (f, args) -> + let loc = f.loc in + let (f, ft) = intern_rec env f in + let fold arg (args, t) = + let loc = arg.loc in + let (arg, argt) = intern_rec env arg in + (arg :: args, (loc, argt) :: t) + in + let (args, t) = List.fold_right fold args ([], []) in + let ret = unify_arrow ?loc env ft t in + (GTacApp (f, args), ret) +| CTacLet (is_rec, el, e) -> + let map (pat, e) = + let (pat, ty) = extract_pattern_type pat in + (pat, ty, e) + in + let el = List.map map el in + let fold accu (pat, _, e) = + let ids = ids_of_pattern Id.Set.empty pat in + let common = Id.Set.inter ids accu in + if Id.Set.is_empty common then Id.Set.union ids accu + else + let id = Id.Set.choose common in + user_err ?loc:pat.loc (str "Variable " ++ Id.print id ++ str " is bound several \ + times in this matching") + in + let ids = List.fold_left fold Id.Set.empty el in + if is_rec then intern_let_rec env loc ids el e + else intern_let env loc ids el e +| CTacCnv (e, tc) -> + let (e, t) = intern_rec env e in + let tc = intern_type env tc in + let () = unify ?loc env t tc in + (e, tc) +| CTacSeq (e1, e2) -> + let loc1 = e1.loc in + let (e1, t1) = intern_rec env e1 in + let (e2, t2) = intern_rec env e2 in + let () = check_elt_unit loc1 env t1 in + (GTacLet (false, [Anonymous, e1], e2), t2) +| CTacCse (e, pl) -> + intern_case env loc e pl +| CTacRec fs -> + intern_record env loc fs +| CTacPrj (e, proj) -> + let pinfo = get_projection proj in + let loc = e.loc in + let (e, t) = intern_rec env e in + let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in + let params = Array.map_to_list (fun i -> GTypVar i) subst in + let exp = GTypRef (Other pinfo.pdata_type, params) in + let () = unify ?loc env t exp in + let substf i = GTypVar subst.(i) in + let ret = subst_type substf pinfo.pdata_ptyp in + (GTacPrj (pinfo.pdata_type, e, pinfo.pdata_indx), ret) +| CTacSet (e, proj, r) -> + let pinfo = get_projection proj in + let () = + if not pinfo.pdata_mutb then + let loc = match proj with + | RelId {CAst.loc} -> loc + | AbsKn _ -> None + in + user_err ?loc (str "Field is not mutable") + in + let subst = Array.init pinfo.pdata_prms (fun _ -> fresh_id env) in + let params = Array.map_to_list (fun i -> GTypVar i) subst in + let exp = GTypRef (Other pinfo.pdata_type, params) in + let e = intern_rec_with_constraint env e exp in + let substf i = GTypVar subst.(i) in + let ret = subst_type substf pinfo.pdata_ptyp in + let r = intern_rec_with_constraint env r ret in + (GTacSet (pinfo.pdata_type, e, pinfo.pdata_indx, r), GTypRef (Tuple 0, [])) +| CTacExt (tag, arg) -> + let open Genintern in + let self ist e = + let env = match Store.get ist.extra ltac2_env with + | None -> empty_env () + | Some env -> env + in + intern_rec env e + in + let obj = interp_ml_object tag in + (* External objects do not have access to the named context because this is + not stable by dynamic semantics. *) + let genv = Global.env_of_context Environ.empty_named_context_val in + let ist = empty_glob_sign genv in + let ist = { ist with extra = Store.set ist.extra ltac2_env env } in + let arg, tpe = + if env.env_str then + let arg () = obj.ml_intern self ist arg in + Flags.with_option Ltac_plugin.Tacintern.strict_check arg () + else + obj.ml_intern self ist arg + in + let e = match arg with + | GlbVal arg -> GTacExt (tag, arg) + | GlbTacexpr e -> e + in + (e, tpe) + +and intern_rec_with_constraint env e exp = + let (er, t) = intern_rec env e in + let () = unify ?loc:e.loc env t exp in + er + +and intern_let env loc ids el e = + let avoid = Id.Set.union ids (Id.Map.domain env.env_var) in + let fold (pat, t, e) (avoid, accu) = + let nas, exp = expand_pattern avoid [pat, t] in + let na = match nas with [x] -> x | _ -> assert false in + let avoid = List.fold_left add_name avoid nas in + (avoid, (na, exp, t, e) :: accu) + in + let (_, el) = List.fold_right fold el (avoid, []) in + let fold (na, exp, tc, e) (body, el, p) = + let (e, t) = match tc with + | None -> intern_rec env e + | Some tc -> + let tc = intern_type env tc in + (intern_rec_with_constraint env e tc, tc) + in + let t = if is_value e then abstract_var env t else monomorphic t in + (exp body, (na, e) :: el, (na, t) :: p) + in + let (e, el, p) = List.fold_right fold el (e, [], []) in + let env = List.fold_left (fun accu (na, t) -> push_name na t accu) env p in + let (e, t) = intern_rec env e in + (GTacLet (false, el, e), t) + +and intern_let_rec env loc ids el e = + let map env (pat, t, e) = + let na = match pat.v with + | CPatVar na -> na + | CPatRef _ | CPatCnv _ -> + user_err ?loc:pat.loc (str "This kind of pattern is forbidden in let-rec bindings") + in + let id = fresh_id env in + let env = push_name na (monomorphic (GTypVar id)) env in + (env, (loc, na, t, e, id)) + in + let (env, el) = List.fold_left_map map env el in + let fold (loc, na, tc, e, id) (el, tl) = + let loc_e = e.loc in + let (e, t) = intern_rec env e in + let () = + if not (is_rec_rhs e) then + user_err ?loc:loc_e (str "This kind of expression is not allowed as \ + right-hand side of a recursive binding") + in + let () = unify ?loc env t (GTypVar id) in + let () = match tc with + | None -> () + | Some tc -> + let tc = intern_type env tc in + unify ?loc env t tc + in + ((na, e) :: el, t :: tl) + in + let (el, tl) = List.fold_right fold el ([], []) in + let (e, t) = intern_rec env e in + (GTacLet (true, el, e), t) + +(** For now, patterns recognized by the pattern-matching compiling are limited + to depth-one where leaves are either variables or catch-all *) +and intern_case env loc e pl = + let (e', t) = intern_rec env e in + let todo ?loc () = user_err ?loc (str "Pattern not handled yet") in + match get_pattern_kind env pl with + | PKind_any -> + let (pat, b) = List.hd pl in + let na = match intern_patexpr env pat with + | GPatVar na -> na + | _ -> assert false + in + let () = check_redundant_clause (List.tl pl) in + let env = push_name na (monomorphic t) env in + let (b, tb) = intern_rec env b in + (GTacLet (false, [na, e'], b), tb) + | PKind_empty -> + let kn = check_elt_empty loc env t in + let r = fresh_id env in + (GTacCse (e', Other kn, [||], [||]), GTypVar r) + | PKind_variant kn -> + let subst, tc = fresh_reftype env kn in + let () = unify ?loc:e.loc env t tc in + let (nconst, nnonconst, arities) = match kn with + | Tuple 0 -> 1, 0, [0] + | Tuple n -> 0, 1, [n] + | Other kn -> + let (_, def) = Tac2env.interp_type kn in + let galg = match def with | GTydAlg c -> c | _ -> assert false in + let arities = List.map (fun (_, args) -> List.length args) galg.galg_constructors in + galg.galg_nconst, galg.galg_nnonconst, arities + in + let const = Array.make nconst None in + let nonconst = Array.make nnonconst None in + let ret = GTypVar (fresh_id env) in + let rec intern_branch = function + | [] -> () + | (pat, br) :: rem -> + let tbr = match pat.v with + | CPatVar (Name _) -> + let loc = pat.loc in + todo ?loc () + | CPatVar Anonymous -> + let () = check_redundant_clause rem in + let (br', brT) = intern_rec env br in + (* Fill all remaining branches *) + let fill (ncst, narg) arity = + if Int.equal arity 0 then + let () = + if Option.is_empty const.(ncst) then const.(ncst) <- Some br' + in + (succ ncst, narg) + else + let () = + if Option.is_empty nonconst.(narg) then + let ids = Array.make arity Anonymous in + nonconst.(narg) <- Some (ids, br') + in + (ncst, succ narg) + in + let _ = List.fold_left fill (0, 0) arities in + brT + | CPatRef (qid, args) -> + let loc = pat.loc in + let knc = get_constructor env qid in + let kn', index, arity = match knc with + | Tuple n -> Tuple n, 0, List.init n (fun i -> GTypVar i) + | Other knc -> + let data = Tac2env.interp_constructor knc in + let index = Option.get data.cdata_indx in + Other data.cdata_type, index, data.cdata_args + in + let () = + if not (eq_or_tuple KerName.equal kn kn') then + invalid_pattern ?loc kn kn' + in + let get_id pat = match pat with + | {v=CPatVar na} -> na + | {loc} -> todo ?loc () + in + let ids = List.map get_id args in + let nids = List.length ids in + let nargs = List.length arity in + let () = match knc with + | Tuple n -> assert (n == nids) + | Other knc -> + if not (Int.equal nids nargs) then error_nargs_mismatch ?loc knc nargs nids + in + let fold env id tpe = + (* Instantiate all arguments *) + let subst n = GTypVar subst.(n) in + let tpe = subst_type subst tpe in + push_name id (monomorphic tpe) env + in + let nenv = List.fold_left2 fold env ids arity in + let (br', brT) = intern_rec nenv br in + let () = + if List.is_empty args then + if Option.is_empty const.(index) then const.(index) <- Some br' + else warn_redundant_clause ?loc () + else + let ids = Array.of_list ids in + if Option.is_empty nonconst.(index) then nonconst.(index) <- Some (ids, br') + else warn_redundant_clause ?loc () + in + brT + | CPatCnv _ -> + user_err ?loc (str "Pattern not handled yet") + in + let () = unify ?loc:br.loc env tbr ret in + intern_branch rem + in + let () = intern_branch pl in + let map n is_const = function + | None -> + let kn = match kn with Other kn -> kn | _ -> assert false in + let cstr = pr_internal_constructor kn n is_const in + user_err ?loc (str "Unhandled match case for constructor " ++ cstr) + | Some x -> x + in + let const = Array.mapi (fun i o -> map i true o) const in + let nonconst = Array.mapi (fun i o -> map i false o) nonconst in + let ce = GTacCse (e', kn, const, nonconst) in + (ce, ret) + | PKind_open kn -> + let subst, tc = fresh_reftype env (Other kn) in + let () = unify ?loc:e.loc env t tc in + let ret = GTypVar (fresh_id env) in + let rec intern_branch map = function + | [] -> + user_err ?loc (str "Missing default case") + | (pat, br) :: rem -> + match intern_patexpr env pat with + | GPatVar na -> + let () = check_redundant_clause rem in + let nenv = push_name na (monomorphic tc) env in + let br' = intern_rec_with_constraint nenv br ret in + let def = (na, br') in + (map, def) + | GPatRef (knc, args) -> + let get = function + | GPatVar na -> na + | GPatRef _ -> + user_err ?loc (str "TODO: Unhandled match case") (* FIXME *) + in + let loc = pat.loc in + let knc = match knc with + | Other knc -> knc + | Tuple n -> invalid_pattern ?loc (Other kn) (Tuple n) + in + let ids = List.map get args in + let data = Tac2env.interp_constructor knc in + let () = + if not (KerName.equal kn data.cdata_type) then + invalid_pattern ?loc (Other kn) (Other data.cdata_type) + in + let nids = List.length ids in + let nargs = List.length data.cdata_args in + let () = + if not (Int.equal nids nargs) then error_nargs_mismatch ?loc knc nargs nids + in + let fold env id tpe = + (* Instantiate all arguments *) + let subst n = GTypVar subst.(n) in + let tpe = subst_type subst tpe in + push_name id (monomorphic tpe) env + in + let nenv = List.fold_left2 fold env ids data.cdata_args in + let br' = intern_rec_with_constraint nenv br ret in + let map = + if KNmap.mem knc map then + let () = warn_redundant_clause ?loc () in + map + else + KNmap.add knc (Anonymous, Array.of_list ids, br') map + in + intern_branch map rem + in + let (map, def) = intern_branch KNmap.empty pl in + (GTacWth { opn_match = e'; opn_branch = map; opn_default = def }, ret) + +and intern_constructor env loc kn args = match kn with +| Other kn -> + let cstr = interp_constructor kn in + let nargs = List.length cstr.cdata_args in + if Int.equal nargs (List.length args) then + let subst = Array.init cstr.cdata_prms (fun _ -> fresh_id env) in + let substf i = GTypVar subst.(i) in + let types = List.map (fun t -> subst_type substf t) cstr.cdata_args in + let targs = List.init cstr.cdata_prms (fun i -> GTypVar subst.(i)) in + let ans = GTypRef (Other cstr.cdata_type, targs) in + let map arg tpe = intern_rec_with_constraint env arg tpe in + let args = List.map2 map args types in + match cstr.cdata_indx with + | Some idx -> + (GTacCst (Other cstr.cdata_type, idx, args), ans) + | None -> + (GTacOpn (kn, args), ans) + else + error_nargs_mismatch ?loc kn nargs (List.length args) +| Tuple n -> + assert (Int.equal n (List.length args)); + let types = List.init n (fun i -> GTypVar (fresh_id env)) in + let map arg tpe = intern_rec_with_constraint env arg tpe in + let args = List.map2 map args types in + let ans = GTypRef (Tuple n, types) in + GTacCst (Tuple n, 0, args), ans + +and intern_record env loc fs = + let map (proj, e) = + let loc = match proj with + | RelId {CAst.loc} -> loc + | AbsKn _ -> None + in + let proj = get_projection proj in + (loc, proj, e) + in + let fs = List.map map fs in + let kn = match fs with + | [] -> user_err ?loc (str "Cannot infer the corresponding record type") + | (_, proj, _) :: _ -> proj.pdata_type + in + let params, typdef = match Tac2env.interp_type kn with + | n, GTydRec def -> n, def + | _ -> assert false + in + let subst = Array.init params (fun _ -> fresh_id env) in + (* Set the answer [args] imperatively *) + let args = Array.make (List.length typdef) None in + let iter (loc, pinfo, e) = + if KerName.equal kn pinfo.pdata_type then + let index = pinfo.pdata_indx in + match args.(index) with + | None -> + let exp = subst_type (fun i -> GTypVar subst.(i)) pinfo.pdata_ptyp in + let e = intern_rec_with_constraint env e exp in + args.(index) <- Some e + | Some _ -> + let (name, _, _) = List.nth typdef pinfo.pdata_indx in + user_err ?loc (str "Field " ++ Id.print name ++ str " is defined \ + several times") + else + user_err ?loc (str "Field " ++ (*KerName.print knp ++*) str " does not \ + pertain to record definition " ++ pr_typref pinfo.pdata_type) + in + let () = List.iter iter fs in + let () = match Array.findi (fun _ o -> Option.is_empty o) args with + | None -> () + | Some i -> + let (field, _, _) = List.nth typdef i in + user_err ?loc (str "Field " ++ Id.print field ++ str " is undefined") + in + let args = Array.map_to_list Option.get args in + let tparam = List.init params (fun i -> GTypVar subst.(i)) in + (GTacCst (Other kn, 0, args), GTypRef (Other kn, tparam)) + +let normalize env (count, vars) (t : UF.elt glb_typexpr) = + let get_var id = + try UF.Map.find id !vars + with Not_found -> + let () = assert env.env_opn in + let n = GTypVar !count in + let () = incr count in + let () = vars := UF.Map.add id n !vars in + n + in + let rec subst id = match UF.find id env.env_cst with + | id, None -> get_var id + | _, Some t -> subst_type subst t + in + subst_type subst t + +let intern ~strict e = + let env = empty_env () in + let env = if strict then env else { env with env_str = false } in + let (e, t) = intern_rec env e in + let count = ref 0 in + let vars = ref UF.Map.empty in + let t = normalize env (count, vars) t in + (e, (!count, t)) + +let intern_typedef self (ids, t) : glb_quant_typedef = + let env = { (empty_env ()) with env_rec = self } in + (* Initialize type parameters *) + let map id = get_alias id env in + let ids = List.map map ids in + let count = ref (List.length ids) in + let vars = ref UF.Map.empty in + let iter n id = vars := UF.Map.add id (GTypVar n) !vars in + let () = List.iteri iter ids in + (* Do not accept unbound type variables *) + let env = { env with env_opn = false } in + let intern t = + let t = intern_type env t in + normalize env (count, vars) t + in + let count = !count in + match t with + | CTydDef None -> (count, GTydDef None) + | CTydDef (Some t) -> (count, GTydDef (Some (intern t))) + | CTydAlg constrs -> + let map (c, t) = (c, List.map intern t) in + let constrs = List.map map constrs in + let getn (const, nonconst) (c, args) = match args with + | [] -> (succ const, nonconst) + | _ :: _ -> (const, succ nonconst) + in + let nconst, nnonconst = List.fold_left getn (0, 0) constrs in + let galg = { + galg_constructors = constrs; + galg_nconst = nconst; + galg_nnonconst = nnonconst; + } in + (count, GTydAlg galg) + | CTydRec fields -> + let map (c, mut, t) = (c, mut, intern t) in + let fields = List.map map fields in + (count, GTydRec fields) + | CTydOpn -> (count, GTydOpn) + +let intern_open_type t = + let env = empty_env () in + let t = intern_type env t in + let count = ref 0 in + let vars = ref UF.Map.empty in + let t = normalize env (count, vars) t in + (!count, t) + +(** Subtyping *) + +let check_subtype t1 t2 = + let env = empty_env () in + let t1 = fresh_type_scheme env t1 in + (* We build a substitution mimicking rigid variable by using dummy tuples *) + let rigid i = GTypRef (Tuple (i + 1), []) in + let (n, t2) = t2 in + let subst = Array.init n rigid in + let substf i = subst.(i) in + let t2 = subst_type substf t2 in + try unify0 env t1 t2; true with CannotUnify _ -> false + +(** Globalization *) + +let get_projection0 var = match var with +| RelId qid -> + let kn = try Tac2env.locate_projection qid with Not_found -> + user_err ?loc:qid.CAst.loc (pr_qualid qid ++ str " is not a projection") + in + kn +| AbsKn kn -> kn + +let rec globalize ids ({loc;v=er} as e) = match er with +| CTacAtm _ -> e +| CTacRef ref -> + let mem id = Id.Set.mem id ids in + begin match get_variable0 mem ref with + | ArgVar _ -> e + | ArgArg kn -> CAst.make ?loc @@ CTacRef (AbsKn kn) + end +| CTacCst qid -> + let knc = get_constructor () qid in + CAst.make ?loc @@ CTacCst (AbsKn knc) +| CTacFun (bnd, e) -> + let fold (pats, accu) pat = + let accu = ids_of_pattern accu pat in + let pat = globalize_pattern ids pat in + (pat :: pats, accu) + in + let bnd, ids = List.fold_left fold ([], ids) bnd in + let bnd = List.rev bnd in + let e = globalize ids e in + CAst.make ?loc @@ CTacFun (bnd, e) +| CTacApp (e, el) -> + let e = globalize ids e in + let el = List.map (fun e -> globalize ids e) el in + CAst.make ?loc @@ CTacApp (e, el) +| CTacLet (isrec, bnd, e) -> + let fold accu (pat, _) = ids_of_pattern accu pat in + let ext = List.fold_left fold Id.Set.empty bnd in + let eids = Id.Set.union ext ids in + let e = globalize eids e in + let map (qid, e) = + let ids = if isrec then eids else ids in + let qid = globalize_pattern ids qid in + (qid, globalize ids e) + in + let bnd = List.map map bnd in + CAst.make ?loc @@ CTacLet (isrec, bnd, e) +| CTacCnv (e, t) -> + let e = globalize ids e in + CAst.make ?loc @@ CTacCnv (e, t) +| CTacSeq (e1, e2) -> + let e1 = globalize ids e1 in + let e2 = globalize ids e2 in + CAst.make ?loc @@ CTacSeq (e1, e2) +| CTacCse (e, bl) -> + let e = globalize ids e in + let bl = List.map (fun b -> globalize_case ids b) bl in + CAst.make ?loc @@ CTacCse (e, bl) +| CTacRec r -> + let map (p, e) = + let p = get_projection0 p in + let e = globalize ids e in + (AbsKn p, e) + in + CAst.make ?loc @@ CTacRec (List.map map r) +| CTacPrj (e, p) -> + let e = globalize ids e in + let p = get_projection0 p in + CAst.make ?loc @@ CTacPrj (e, AbsKn p) +| CTacSet (e, p, e') -> + let e = globalize ids e in + let p = get_projection0 p in + let e' = globalize ids e' in + CAst.make ?loc @@ CTacSet (e, AbsKn p, e') +| CTacExt (tag, arg) -> + let arg = str (Tac2dyn.Arg.repr tag) in + CErrors.user_err ?loc (str "Cannot globalize generic arguments of type" ++ spc () ++ arg) + +and globalize_case ids (p, e) = + (globalize_pattern ids p, globalize ids e) + +and globalize_pattern ids ({loc;v=pr} as p) = match pr with +| CPatVar _ -> p +| CPatRef (cst, pl) -> + let knc = get_constructor () cst in + let cst = AbsKn knc in + let pl = List.map (fun p -> globalize_pattern ids p) pl in + CAst.make ?loc @@ CPatRef (cst, pl) +| CPatCnv (pat, ty) -> + let pat = globalize_pattern ids pat in + CAst.make ?loc @@ CPatCnv (pat, ty) + +(** Kernel substitution *) + +open Mod_subst + +let subst_or_tuple f subst o = match o with +| Tuple _ -> o +| Other v -> + let v' = f subst v in + if v' == v then o else Other v' + +let rec subst_type subst t = match t with +| GTypVar _ -> t +| GTypArrow (t1, t2) -> + let t1' = subst_type subst t1 in + let t2' = subst_type subst t2 in + if t1' == t1 && t2' == t2 then t + else GTypArrow (t1', t2') +| GTypRef (kn, tl) -> + let kn' = subst_or_tuple subst_kn subst kn in + let tl' = List.Smart.map (fun t -> subst_type subst t) tl in + if kn' == kn && tl' == tl then t else GTypRef (kn', tl') + +let rec subst_expr subst e = match e with +| GTacAtm _ | GTacVar _ | GTacPrm _ -> e +| GTacRef kn -> GTacRef (subst_kn subst kn) +| GTacFun (ids, e) -> GTacFun (ids, subst_expr subst e) +| GTacApp (f, args) -> + GTacApp (subst_expr subst f, List.map (fun e -> subst_expr subst e) args) +| GTacLet (r, bs, e) -> + let bs = List.map (fun (na, e) -> (na, subst_expr subst e)) bs in + GTacLet (r, bs, subst_expr subst e) +| GTacCst (t, n, el) as e0 -> + let t' = subst_or_tuple subst_kn subst t in + let el' = List.Smart.map (fun e -> subst_expr subst e) el in + if t' == t && el' == el then e0 else GTacCst (t', n, el') +| GTacCse (e, ci, cse0, cse1) -> + let cse0' = Array.map (fun e -> subst_expr subst e) cse0 in + let cse1' = Array.map (fun (ids, e) -> (ids, subst_expr subst e)) cse1 in + let ci' = subst_or_tuple subst_kn subst ci in + GTacCse (subst_expr subst e, ci', cse0', cse1') +| GTacWth { opn_match = e; opn_branch = br; opn_default = (na, def) } as e0 -> + let e' = subst_expr subst e in + let def' = subst_expr subst def in + let fold kn (self, vars, p) accu = + let kn' = subst_kn subst kn in + let p' = subst_expr subst p in + if kn' == kn && p' == p then accu + else KNmap.add kn' (self, vars, p') (KNmap.remove kn accu) + in + let br' = KNmap.fold fold br br in + if e' == e && br' == br && def' == def then e0 + else GTacWth { opn_match = e'; opn_default = (na, def'); opn_branch = br' } +| GTacPrj (kn, e, p) as e0 -> + let kn' = subst_kn subst kn in + let e' = subst_expr subst e in + if kn' == kn && e' == e then e0 else GTacPrj (kn', e', p) +| GTacSet (kn, e, p, r) as e0 -> + let kn' = subst_kn subst kn in + let e' = subst_expr subst e in + let r' = subst_expr subst r in + if kn' == kn && e' == e && r' == r then e0 else GTacSet (kn', e', p, r') +| GTacExt (tag, arg) -> + let tpe = interp_ml_object tag in + let arg' = tpe.ml_subst subst arg in + if arg' == arg then e else GTacExt (tag, arg') +| GTacOpn (kn, el) as e0 -> + let kn' = subst_kn subst kn in + let el' = List.Smart.map (fun e -> subst_expr subst e) el in + if kn' == kn && el' == el then e0 else GTacOpn (kn', el') + +let subst_typedef subst e = match e with +| GTydDef t -> + let t' = Option.Smart.map (fun t -> subst_type subst t) t in + if t' == t then e else GTydDef t' +| GTydAlg galg -> + let map (c, tl as p) = + let tl' = List.Smart.map (fun t -> subst_type subst t) tl in + if tl' == tl then p else (c, tl') + in + let constrs' = List.Smart.map map galg.galg_constructors in + if constrs' == galg.galg_constructors then e + else GTydAlg { galg with galg_constructors = constrs' } +| GTydRec fields -> + let map (c, mut, t as p) = + let t' = subst_type subst t in + if t' == t then p else (c, mut, t') + in + let fields' = List.Smart.map map fields in + if fields' == fields then e else GTydRec fields' +| GTydOpn -> GTydOpn + +let subst_quant_typedef subst (prm, def as qdef) = + let def' = subst_typedef subst def in + if def' == def then qdef else (prm, def') + +let subst_type_scheme subst (prm, t as sch) = + let t' = subst_type subst t in + if t' == t then sch else (prm, t') + +let subst_or_relid subst ref = match ref with +| RelId _ -> ref +| AbsKn kn -> + let kn' = subst_or_tuple subst_kn subst kn in + if kn' == kn then ref else AbsKn kn' + +let rec subst_rawtype subst ({loc;v=tr} as t) = match tr with +| CTypVar _ -> t +| CTypArrow (t1, t2) -> + let t1' = subst_rawtype subst t1 in + let t2' = subst_rawtype subst t2 in + if t1' == t1 && t2' == t2 then t else CAst.make ?loc @@ CTypArrow (t1', t2') +| CTypRef (ref, tl) -> + let ref' = subst_or_relid subst ref in + let tl' = List.Smart.map (fun t -> subst_rawtype subst t) tl in + if ref' == ref && tl' == tl then t else CAst.make ?loc @@ CTypRef (ref', tl') + +let subst_tacref subst ref = match ref with +| RelId _ -> ref +| AbsKn (TacConstant kn) -> + let kn' = subst_kn subst kn in + if kn' == kn then ref else AbsKn (TacConstant kn') +| AbsKn (TacAlias kn) -> + let kn' = subst_kn subst kn in + if kn' == kn then ref else AbsKn (TacAlias kn') + +let subst_projection subst prj = match prj with +| RelId _ -> prj +| AbsKn kn -> + let kn' = subst_kn subst kn in + if kn' == kn then prj else AbsKn kn' + +let rec subst_rawpattern subst ({loc;v=pr} as p) = match pr with +| CPatVar _ -> p +| CPatRef (c, pl) -> + let pl' = List.Smart.map (fun p -> subst_rawpattern subst p) pl in + let c' = subst_or_relid subst c in + if pl' == pl && c' == c then p else CAst.make ?loc @@ CPatRef (c', pl') +| CPatCnv (pat, ty) -> + let pat' = subst_rawpattern subst pat in + let ty' = subst_rawtype subst ty in + if pat' == pat && ty' == ty then p else CAst.make ?loc @@ CPatCnv (pat', ty') + +(** Used for notations *) +let rec subst_rawexpr subst ({loc;v=tr} as t) = match tr with +| CTacAtm _ -> t +| CTacRef ref -> + let ref' = subst_tacref subst ref in + if ref' == ref then t else CAst.make ?loc @@ CTacRef ref' +| CTacCst ref -> + let ref' = subst_or_relid subst ref in + if ref' == ref then t else CAst.make ?loc @@ CTacCst ref' +| CTacFun (bnd, e) -> + let map pat = subst_rawpattern subst pat in + let bnd' = List.Smart.map map bnd in + let e' = subst_rawexpr subst e in + if bnd' == bnd && e' == e then t else CAst.make ?loc @@ CTacFun (bnd', e') +| CTacApp (e, el) -> + let e' = subst_rawexpr subst e in + let el' = List.Smart.map (fun e -> subst_rawexpr subst e) el in + if e' == e && el' == el then t else CAst.make ?loc @@ CTacApp (e', el') +| CTacLet (isrec, bnd, e) -> + let map (na, e as p) = + let na' = subst_rawpattern subst na in + let e' = subst_rawexpr subst e in + if na' == na && e' == e then p else (na', e') + in + let bnd' = List.Smart.map map bnd in + let e' = subst_rawexpr subst e in + if bnd' == bnd && e' == e then t else CAst.make ?loc @@ CTacLet (isrec, bnd', e') +| CTacCnv (e, c) -> + let e' = subst_rawexpr subst e in + let c' = subst_rawtype subst c in + if c' == c && e' == e then t else CAst.make ?loc @@ CTacCnv (e', c') +| CTacSeq (e1, e2) -> + let e1' = subst_rawexpr subst e1 in + let e2' = subst_rawexpr subst e2 in + if e1' == e1 && e2' == e2 then t else CAst.make ?loc @@ CTacSeq (e1', e2') +| CTacCse (e, bl) -> + let map (p, e as x) = + let p' = subst_rawpattern subst p in + let e' = subst_rawexpr subst e in + if p' == p && e' == e then x else (p', e') + in + let e' = subst_rawexpr subst e in + let bl' = List.Smart.map map bl in + if e' == e && bl' == bl then t else CAst.make ?loc @@ CTacCse (e', bl') +| CTacRec el -> + let map (prj, e as p) = + let prj' = subst_projection subst prj in + let e' = subst_rawexpr subst e in + if prj' == prj && e' == e then p else (prj', e') + in + let el' = List.Smart.map map el in + if el' == el then t else CAst.make ?loc @@ CTacRec el' +| CTacPrj (e, prj) -> + let prj' = subst_projection subst prj in + let e' = subst_rawexpr subst e in + if prj' == prj && e' == e then t else CAst.make ?loc @@ CTacPrj (e', prj') +| CTacSet (e, prj, r) -> + let prj' = subst_projection subst prj in + let e' = subst_rawexpr subst e in + let r' = subst_rawexpr subst r in + if prj' == prj && e' == e && r' == r then t else CAst.make ?loc @@ CTacSet (e', prj', r') +| CTacExt _ -> assert false (** Should not be generated by globalization *) + +(** Registering *) + +let () = + let open Genintern in + let intern ist tac = + let env = match Genintern.Store.get ist.extra ltac2_env with + | None -> + (* Only happens when Ltac2 is called from a constr or ltac1 quotation *) + let env = empty_env () in + if !Ltac_plugin.Tacintern.strict_check then env + else { env with env_str = false } + | Some env -> env + in + let loc = tac.loc in + let (tac, t) = intern_rec env tac in + let () = check_elt_unit loc env t in + (ist, tac) + in + Genintern.register_intern0 wit_ltac2 intern +let () = Genintern.register_subst0 wit_ltac2 subst_expr + +let () = + let open Genintern in + let intern ist (loc, id) = + let env = match Genintern.Store.get ist.extra ltac2_env with + | None -> + (* Only happens when Ltac2 is called from a constr or ltac1 quotation *) + let env = empty_env () in + if !Ltac_plugin.Tacintern.strict_check then env + else { env with env_str = false } + | Some env -> env + in + let t = + try Id.Map.find id env.env_var + with Not_found -> + CErrors.user_err ?loc (str "Unbound value " ++ Id.print id) + in + let t = fresh_mix_type_scheme env t in + let () = unify ?loc env t (GTypRef (Other t_constr, [])) in + (ist, id) + in + Genintern.register_intern0 wit_ltac2_quotation intern + +let () = Genintern.register_subst0 wit_ltac2_quotation (fun _ id -> id) diff --git a/user-contrib/Ltac2/tac2intern.mli b/user-contrib/Ltac2/tac2intern.mli new file mode 100644 index 0000000000..d646b5cda5 --- /dev/null +++ b/user-contrib/Ltac2/tac2intern.mli @@ -0,0 +1,46 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open Mod_subst +open Tac2expr + +val intern : strict:bool -> raw_tacexpr -> glb_tacexpr * type_scheme +val intern_typedef : (KerName.t * int) Id.Map.t -> raw_quant_typedef -> glb_quant_typedef +val intern_open_type : raw_typexpr -> type_scheme + +(** Check that a term is a value. Only values are safe to marshall between + processes. *) +val is_value : glb_tacexpr -> bool +val check_unit : ?loc:Loc.t -> type_scheme -> unit + +val check_subtype : type_scheme -> type_scheme -> bool +(** [check_subtype t1 t2] returns [true] iff all values of intances of type [t1] + also have type [t2]. *) + +val subst_type : substitution -> 'a glb_typexpr -> 'a glb_typexpr +val subst_expr : substitution -> glb_tacexpr -> glb_tacexpr +val subst_quant_typedef : substitution -> glb_quant_typedef -> glb_quant_typedef +val subst_type_scheme : substitution -> type_scheme -> type_scheme + +val subst_rawexpr : substitution -> raw_tacexpr -> raw_tacexpr + +(** {5 Notations} *) + +val globalize : Id.Set.t -> raw_tacexpr -> raw_tacexpr +(** Replaces all qualified identifiers by their corresponding kernel name. The + set represents bound variables in the context. *) + +(** Errors *) + +val error_nargs_mismatch : ?loc:Loc.t -> ltac_constructor -> int -> int -> 'a +val error_nparams_mismatch : ?loc:Loc.t -> int -> int -> 'a + +(** Misc *) + +val drop_ltac2_env : Genintern.Store.t -> Genintern.Store.t diff --git a/user-contrib/Ltac2/tac2interp.ml b/user-contrib/Ltac2/tac2interp.ml new file mode 100644 index 0000000000..db779db471 --- /dev/null +++ b/user-contrib/Ltac2/tac2interp.ml @@ -0,0 +1,227 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Util +open Pp +open CErrors +open Names +open Proofview.Notations +open Tac2expr +open Tac2ffi + +exception LtacError = Tac2ffi.LtacError + +let backtrace : backtrace Evd.Store.field = Evd.Store.field () + +let print_ltac2_backtrace = ref false + +let get_backtrace = + Proofview.tclEVARMAP >>= fun sigma -> + match Evd.Store.get (Evd.get_extra_data sigma) backtrace with + | None -> Proofview.tclUNIT [] + | Some bt -> Proofview.tclUNIT bt + +let set_backtrace bt = + Proofview.tclEVARMAP >>= fun sigma -> + let store = Evd.get_extra_data sigma in + let store = Evd.Store.set store backtrace bt in + let sigma = Evd.set_extra_data store sigma in + Proofview.Unsafe.tclEVARS sigma + +let with_frame frame tac = + if !print_ltac2_backtrace then + get_backtrace >>= fun bt -> + set_backtrace (frame :: bt) >>= fun () -> + tac >>= fun ans -> + set_backtrace bt >>= fun () -> + Proofview.tclUNIT ans + else tac + +type environment = Tac2env.environment = { + env_ist : valexpr Id.Map.t; +} + +let empty_environment = { + env_ist = Id.Map.empty; +} + +type closure = { + mutable clos_env : valexpr Id.Map.t; + (** Mutable so that we can implement recursive functions imperatively *) + clos_var : Name.t list; + (** Bound variables *) + clos_exp : glb_tacexpr; + (** Body *) + clos_ref : ltac_constant option; + (** Global constant from which the closure originates *) +} + +let push_name ist id v = match id with +| Anonymous -> ist +| Name id -> { env_ist = Id.Map.add id v ist.env_ist } + +let get_var ist id = + try Id.Map.find id ist.env_ist with Not_found -> + anomaly (str "Unbound variable " ++ Id.print id) + +let get_ref ist kn = + try + let data = Tac2env.interp_global kn in + data.Tac2env.gdata_expr + with Not_found -> + anomaly (str "Unbound reference" ++ KerName.print kn) + +let return = Proofview.tclUNIT + +let rec interp (ist : environment) = function +| GTacAtm (AtmInt n) -> return (Tac2ffi.of_int n) +| GTacAtm (AtmStr s) -> return (Tac2ffi.of_string (Bytes.of_string s)) +| GTacVar id -> return (get_var ist id) +| GTacRef kn -> + let data = get_ref ist kn in + return (eval_pure (Some kn) data) +| GTacFun (ids, e) -> + let cls = { clos_ref = None; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in + let f = interp_app cls in + return (Tac2ffi.of_closure f) +| GTacApp (f, args) -> + interp ist f >>= fun f -> + Proofview.Monad.List.map (fun e -> interp ist e) args >>= fun args -> + Tac2ffi.apply (Tac2ffi.to_closure f) args +| GTacLet (false, el, e) -> + let fold accu (na, e) = + interp ist e >>= fun e -> + return (push_name accu na e) + in + Proofview.Monad.List.fold_left fold ist el >>= fun ist -> + interp ist e +| GTacLet (true, el, e) -> + let map (na, e) = match e with + | GTacFun (ids, e) -> + let cls = { clos_ref = None; clos_env = ist.env_ist; clos_var = ids; clos_exp = e } in + let f = Tac2ffi.of_closure (interp_app cls) in + na, cls, f + | _ -> anomaly (str "Ill-formed recursive function") + in + let fixs = List.map map el in + let fold accu (na, _, cls) = match na with + | Anonymous -> accu + | Name id -> { env_ist = Id.Map.add id cls accu.env_ist } + in + let ist = List.fold_left fold ist fixs in + (* Hack to make a cycle imperatively in the environment *) + let iter (_, e, _) = e.clos_env <- ist.env_ist in + let () = List.iter iter fixs in + interp ist e +| GTacCst (_, n, []) -> return (Valexpr.make_int n) +| GTacCst (_, n, el) -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + return (Valexpr.make_block n (Array.of_list el)) +| GTacCse (e, _, cse0, cse1) -> + interp ist e >>= fun e -> interp_case ist e cse0 cse1 +| GTacWth { opn_match = e; opn_branch = cse; opn_default = def } -> + interp ist e >>= fun e -> interp_with ist e cse def +| GTacPrj (_, e, p) -> + interp ist e >>= fun e -> interp_proj ist e p +| GTacSet (_, e, p, r) -> + interp ist e >>= fun e -> + interp ist r >>= fun r -> + interp_set ist e p r +| GTacOpn (kn, el) -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + return (Tac2ffi.of_open (kn, Array.of_list el)) +| GTacPrm (ml, el) -> + Proofview.Monad.List.map (fun e -> interp ist e) el >>= fun el -> + with_frame (FrPrim ml) (Tac2ffi.apply (Tac2env.interp_primitive ml) el) +| GTacExt (tag, e) -> + let tpe = Tac2env.interp_ml_object tag in + with_frame (FrExtn (tag, e)) (tpe.Tac2env.ml_interp ist e) + +and interp_app f = + let ans = fun args -> + let { clos_env = ist; clos_var = ids; clos_exp = e; clos_ref = kn } = f in + let frame = match kn with + | None -> FrAnon e + | Some kn -> FrLtac kn + in + let ist = { env_ist = ist } in + let ist = List.fold_left2 push_name ist ids args in + with_frame frame (interp ist e) + in + Tac2ffi.abstract (List.length f.clos_var) ans + +and interp_case ist e cse0 cse1 = + if Valexpr.is_int e then + interp ist cse0.(Tac2ffi.to_int e) + else + let (n, args) = Tac2ffi.to_block e in + let (ids, e) = cse1.(n) in + let ist = CArray.fold_left2 push_name ist ids args in + interp ist e + +and interp_with ist e cse def = + let (kn, args) = Tac2ffi.to_open e in + let br = try Some (KNmap.find kn cse) with Not_found -> None in + begin match br with + | None -> + let (self, def) = def in + let ist = push_name ist self e in + interp ist def + | Some (self, ids, p) -> + let ist = push_name ist self e in + let ist = CArray.fold_left2 push_name ist ids args in + interp ist p + end + +and interp_proj ist e p = + return (Valexpr.field e p) + +and interp_set ist e p r = + let () = Valexpr.set_field e p r in + return (Valexpr.make_int 0) + +and eval_pure kn = function +| GTacAtm (AtmInt n) -> Valexpr.make_int n +| GTacRef kn -> + let { Tac2env.gdata_expr = e } = + try Tac2env.interp_global kn + with Not_found -> assert false + in + eval_pure (Some kn) e +| GTacFun (na, e) -> + let cls = { clos_ref = kn; clos_env = Id.Map.empty; clos_var = na; clos_exp = e } in + let f = interp_app cls in + Tac2ffi.of_closure f +| GTacCst (_, n, []) -> Valexpr.make_int n +| GTacCst (_, n, el) -> Valexpr.make_block n (Array.map_of_list eval_unnamed el) +| GTacOpn (kn, el) -> Tac2ffi.of_open (kn, Array.map_of_list eval_unnamed el) +| GTacAtm (AtmStr _) | GTacLet _ | GTacVar _ | GTacSet _ +| GTacApp _ | GTacCse _ | GTacPrj _ | GTacPrm _ | GTacExt _ | GTacWth _ -> + anomaly (Pp.str "Term is not a syntactical value") + +and eval_unnamed e = eval_pure None e + + +(** Cross-boundary hacks. *) + +open Geninterp + +let val_env : environment Val.typ = Val.create "ltac2:env" +let env_ref = Id.of_string_soft "@@ltac2_env@@" + +let extract_env (Val.Dyn (tag, v)) : environment = +match Val.eq tag val_env with +| None -> assert false +| Some Refl -> v + +let get_env ist = + try extract_env (Id.Map.find env_ref ist) + with Not_found -> empty_environment + +let set_env env ist = + Id.Map.add env_ref (Val.Dyn (val_env, env)) ist diff --git a/user-contrib/Ltac2/tac2interp.mli b/user-contrib/Ltac2/tac2interp.mli new file mode 100644 index 0000000000..21fdcd03af --- /dev/null +++ b/user-contrib/Ltac2/tac2interp.mli @@ -0,0 +1,37 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open Tac2expr +open Tac2ffi + +type environment = Tac2env.environment + +val empty_environment : environment + +val interp : environment -> glb_tacexpr -> valexpr Proofview.tactic + +(* val interp_app : closure -> ml_tactic *) + +(** {5 Cross-boundary encodings} *) + +val get_env : Ltac_pretype.unbound_ltac_var_map -> environment +val set_env : environment -> Ltac_pretype.unbound_ltac_var_map -> Ltac_pretype.unbound_ltac_var_map + +(** {5 Exceptions} *) + +exception LtacError of KerName.t * valexpr array +(** Ltac2-defined exceptions seen from OCaml side *) + +(** {5 Backtrace} *) + +val get_backtrace : backtrace Proofview.tactic + +val with_frame : frame -> 'a Proofview.tactic -> 'a Proofview.tactic + +val print_ltac2_backtrace : bool ref diff --git a/user-contrib/Ltac2/tac2match.ml b/user-contrib/Ltac2/tac2match.ml new file mode 100644 index 0000000000..058d02adde --- /dev/null +++ b/user-contrib/Ltac2/tac2match.ml @@ -0,0 +1,232 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open Context.Named.Declaration + +module NamedDecl = Context.Named.Declaration + +type context = EConstr.t + +type result = { + subst : Ltac_pretype.patvar_map ; +} + +type match_pattern = +| MatchPattern of Pattern.constr_pattern +| MatchContext of Pattern.constr_pattern + +(** TODO: handle definitions *) +type match_context_hyps = match_pattern + +type match_rule = match_context_hyps list * match_pattern + +(** {6 Utilities} *) + +(** Tests whether the substitution [s] is empty. *) +let is_empty_subst = Id.Map.is_empty + +(** {6 Non-linear patterns} *) + + +(** The patterns of Ltac are not necessarily linear. Non-linear + pattern are partially handled by the {!Matching} module, however + goal patterns are not primitive to {!Matching}, hence we must deal + with non-linearity between hypotheses and conclusion. Subterms are + considered equal up to the equality implemented in + [equal_instances]. *) +(* spiwack: it doesn't seem to be quite the same rule for non-linear + term patterns and non-linearity between hypotheses and/or + conclusion. Indeed, in [Matching], matching is made modulo + syntactic equality, and here we merge modulo conversion. It may be + a good idea to have an entry point of [Matching] with a partial + substitution as argument instead of merging substitution here. That + would ensure consistency. *) +let equal_instances env sigma c1 c2 = + (* How to compare instances? Do we want the terms to be convertible? + unifiable? Do we want the universe levels to be relevant? + (historically, conv_x is used) *) + Reductionops.is_conv env sigma c1 c2 + +(** Merges two substitutions. Raises [Not_coherent_metas] when + encountering two instances of the same metavariable which are not + equal according to {!equal_instances}. *) +exception Not_coherent_metas +let verify_metas_coherence env sigma s1 s2 = + let merge id oc1 oc2 = match oc1, oc2 with + | None, None -> None + | None, Some c | Some c, None -> Some c + | Some c1, Some c2 -> + if equal_instances env sigma c1 c2 then Some c1 + else raise Not_coherent_metas + in + Id.Map.merge merge s1 s2 + +let matching_error = + CErrors.UserError (Some "tactic matching" , Pp.str "No matching clauses for match.") + +let imatching_error = (matching_error, Exninfo.null) + +(** A functor is introduced to share the environment and the + evar_map. They do not change and it would be a pity to introduce + closures everywhere just for the occasional calls to + {!equal_instances}. *) +module type StaticEnvironment = sig + val env : Environ.env + val sigma : Evd.evar_map +end +module PatternMatching (E:StaticEnvironment) = struct + + + (** {6 The pattern-matching monad } *) + + + (** To focus on the algorithmic portion of pattern-matching, the + bookkeeping is relegated to a monad: the composition of the + bactracking monad of {!IStream.t} with a "writer" effect. *) + (* spiwack: as we don't benefit from the various stream optimisations + of Haskell, it may be costly to give the monad in direct style such as + here. We may want to use some continuation passing style. *) + type 'a tac = 'a Proofview.tactic + type 'a m = { stream : 'r. ('a -> result -> 'r tac) -> result -> 'r tac } + + (** The empty substitution. *) + let empty_subst = Id.Map.empty + + (** Composes two substitutions using {!verify_metas_coherence}. It + must be a monoid with neutral element {!empty_subst}. Raises + [Not_coherent_metas] when composition cannot be achieved. *) + let subst_prod s1 s2 = + if is_empty_subst s1 then s2 + else if is_empty_subst s2 then s1 + else verify_metas_coherence E.env E.sigma s1 s2 + + (** Merge two writers (and ignore the first value component). *) + let merge m1 m2 = + try Some { + subst = subst_prod m1.subst m2.subst; + } + with Not_coherent_metas -> None + + (** Monadic [return]: returns a single success with empty substitutions. *) + let return (type a) (lhs:a) : a m = + { stream = fun k ctx -> k lhs ctx } + + (** Monadic bind: each success of [x] is replaced by the successes + of [f x]. The substitutions of [x] and [f x] are composed, + dropping the apparent successes when the substitutions are not + coherent. *) + let (>>=) (type a) (type b) (m:a m) (f:a -> b m) : b m = + { stream = fun k ctx -> m.stream (fun x ctx -> (f x).stream k ctx) ctx } + + (** A variant of [(>>=)] when the first argument returns [unit]. *) + let (<*>) (type a) (m:unit m) (y:a m) : a m = + { stream = fun k ctx -> m.stream (fun () ctx -> y.stream k ctx) ctx } + + (** Failure of the pattern-matching monad: no success. *) + let fail (type a) : a m = { stream = fun _ _ -> Proofview.tclZERO matching_error } + + let run (m : 'a m) = + let ctx = { + subst = empty_subst ; + } in + let eval x ctx = Proofview.tclUNIT (x, ctx) in + m.stream eval ctx + + (** Chooses in a list, in the same order as the list *) + let rec pick (l:'a list) (e, info) : 'a m = match l with + | [] -> { stream = fun _ _ -> Proofview.tclZERO ~info e } + | x :: l -> + { stream = fun k ctx -> Proofview.tclOR (k x ctx) (fun e -> (pick l e).stream k ctx) } + + let pick l = pick l imatching_error + + let put_subst subst : unit m = + let s = { subst } in + { stream = fun k ctx -> match merge s ctx with None -> Proofview.tclZERO matching_error | Some s -> k () s } + + (** {6 Pattern-matching} *) + + let pattern_match_term pat term = + match pat with + | MatchPattern p -> + begin + try + put_subst (Constr_matching.matches E.env E.sigma p term) <*> + return None + with Constr_matching.PatternMatchingFailure -> fail + end + | MatchContext p -> + + let rec map s (e, info) = + { stream = fun k ctx -> match IStream.peek s with + | IStream.Nil -> Proofview.tclZERO ~info e + | IStream.Cons ({ Constr_matching.m_sub = (_, subst); m_ctx }, s) -> + let nctx = { subst } in + match merge ctx nctx with + | None -> (map s (e, info)).stream k ctx + | Some nctx -> Proofview.tclOR (k (Some (Lazy.force m_ctx)) nctx) (fun e -> (map s e).stream k ctx) + } + in + map (Constr_matching.match_subterm E.env E.sigma (Id.Set.empty,p) term) imatching_error + + let hyp_match_type pat hyps = + pick hyps >>= fun decl -> + let id = NamedDecl.get_id decl in + pattern_match_term pat (NamedDecl.get_type decl) >>= fun ctx -> + return (id, ctx) + + let _hyp_match_body_and_type bodypat typepat hyps = + pick hyps >>= function + | LocalDef (id,body,hyp) -> + pattern_match_term bodypat body >>= fun ctx_body -> + pattern_match_term typepat hyp >>= fun ctx_typ -> + return (id, ctx_body, ctx_typ) + | LocalAssum (id,hyp) -> fail + + let hyp_match pat hyps = + match pat with + | typepat -> + hyp_match_type typepat hyps +(* | Def ((_,hypname),bodypat,typepat) -> *) +(* hyp_match_body_and_type hypname bodypat typepat hyps *) + + (** [hyp_pattern_list_match pats hyps lhs], matches the list of + patterns [pats] against the hypotheses in [hyps], and eventually + returns [lhs]. *) + let rec hyp_pattern_list_match pats hyps accu = + match pats with + | pat::pats -> + hyp_match pat hyps >>= fun (matched_hyp, hyp_ctx) -> + let select_matched_hyp decl = Id.equal (NamedDecl.get_id decl) matched_hyp in + let hyps = CList.remove_first select_matched_hyp hyps in + hyp_pattern_list_match pats hyps ((matched_hyp, hyp_ctx) :: accu) + | [] -> return accu + + let rule_match_goal hyps concl = function + | (hyppats,conclpat) -> + (* the rules are applied from the topmost one (in the concrete + syntax) to the bottommost. *) + let hyppats = List.rev hyppats in + pattern_match_term conclpat concl >>= fun ctx_concl -> + hyp_pattern_list_match hyppats hyps [] >>= fun hyps -> + return (hyps, ctx_concl) + +end + +let match_goal env sigma concl ~rev rule = + let open Proofview.Notations in + let hyps = EConstr.named_context env in + let hyps = if rev then List.rev hyps else hyps in + let module E = struct + let env = env + let sigma = sigma + end in + let module M = PatternMatching(E) in + M.run (M.rule_match_goal hyps concl rule) >>= fun ((hyps, ctx_concl), subst) -> + Proofview.tclUNIT (hyps, ctx_concl, subst.subst) diff --git a/user-contrib/Ltac2/tac2match.mli b/user-contrib/Ltac2/tac2match.mli new file mode 100644 index 0000000000..c82c40d238 --- /dev/null +++ b/user-contrib/Ltac2/tac2match.mli @@ -0,0 +1,33 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open EConstr + +(** This file extends Matching with the main logic for Ltac2 match goal. *) + +type context = EConstr.t + +type match_pattern = +| MatchPattern of Pattern.constr_pattern +| MatchContext of Pattern.constr_pattern + +(** TODO: handle definitions *) +type match_context_hyps = match_pattern + +type match_rule = match_context_hyps list * match_pattern + +val match_goal: + Environ.env -> + Evd.evar_map -> + constr -> + rev:bool -> + match_rule -> + ((Id.t * context option) list * (* List of hypotheses matching: name + context *) + context option * (* Context for conclusion *) + Ltac_pretype.patvar_map (* Pattern variable substitution *)) Proofview.tactic diff --git a/user-contrib/Ltac2/tac2print.ml b/user-contrib/Ltac2/tac2print.ml new file mode 100644 index 0000000000..f4cb290265 --- /dev/null +++ b/user-contrib/Ltac2/tac2print.ml @@ -0,0 +1,488 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Util +open Pp +open Names +open Tac2expr +open Tac2env +open Tac2ffi + +(** Utils *) + +let change_kn_label kn id = + let mp = KerName.modpath kn in + KerName.make mp (Label.of_id id) + +let paren p = hov 2 (str "(" ++ p ++ str ")") + +let t_list = + KerName.make Tac2env.coq_prefix (Label.of_id (Id.of_string "list")) + + +(** Type printing *) + +type typ_level = +| T5_l +| T5_r +| T2 +| T1 +| T0 + +let t_unit = + KerName.make Tac2env.coq_prefix (Label.of_id (Id.of_string "unit")) + +let pr_typref kn = + Libnames.pr_qualid (Tac2env.shortest_qualid_of_type kn) + +let pr_glbtype_gen pr lvl c = + let rec pr_glbtype lvl = function + | GTypVar n -> str "'" ++ str (pr n) + | GTypRef (Other kn, []) -> pr_typref kn + | GTypRef (Other kn, [t]) -> + let paren = match lvl with + | T5_r | T5_l | T2 | T1 -> fun x -> x + | T0 -> paren + in + paren (pr_glbtype T1 t ++ spc () ++ pr_typref kn) + | GTypRef (Other kn, tl) -> + let paren = match lvl with + | T5_r | T5_l | T2 | T1 -> fun x -> x + | T0 -> paren + in + paren (str "(" ++ prlist_with_sep (fun () -> str ", ") (pr_glbtype lvl) tl ++ str ")" ++ spc () ++ pr_typref kn) + | GTypArrow (t1, t2) -> + let paren = match lvl with + | T5_r -> fun x -> x + | T5_l | T2 | T1 | T0 -> paren + in + paren (pr_glbtype T5_l t1 ++ spc () ++ str "->" ++ spc () ++ pr_glbtype T5_r t2) + | GTypRef (Tuple 0, []) -> + Libnames.pr_qualid (Tac2env.shortest_qualid_of_type t_unit) + | GTypRef (Tuple _, tl) -> + let paren = match lvl with + | T5_r | T5_l -> fun x -> x + | T2 | T1 | T0 -> paren + in + paren (prlist_with_sep (fun () -> str " * ") (pr_glbtype T2) tl) + in + hov 0 (pr_glbtype lvl c) + +let pr_glbtype pr c = pr_glbtype_gen pr T5_r c + +let int_name () = + let vars = ref Int.Map.empty in + fun n -> + if Int.Map.mem n !vars then Int.Map.find n !vars + else + let num = Int.Map.cardinal !vars in + let base = num mod 26 in + let rem = num / 26 in + let name = String.make 1 (Char.chr (97 + base)) in + let suff = if Int.equal rem 0 then "" else string_of_int rem in + let name = name ^ suff in + let () = vars := Int.Map.add n name !vars in + name + +(** Term printing *) + +let pr_constructor kn = + Libnames.pr_qualid (Tac2env.shortest_qualid_of_constructor kn) + +let pr_projection kn = + Libnames.pr_qualid (Tac2env.shortest_qualid_of_projection kn) + +type exp_level = Tac2expr.exp_level = +| E5 +| E4 +| E3 +| E2 +| E1 +| E0 + +let pr_atom = function +| AtmInt n -> Pp.int n +| AtmStr s -> qstring s + +let pr_name = function +| Name id -> Id.print id +| Anonymous -> str "_" + +let find_constructor n empty def = + let rec find n = function + | [] -> assert false + | (id, []) as ans :: rem -> + if empty then + if Int.equal n 0 then ans + else find (pred n) rem + else find n rem + | (id, _ :: _) as ans :: rem -> + if not empty then + if Int.equal n 0 then ans + else find (pred n) rem + else find n rem + in + find n def + +let pr_internal_constructor tpe n is_const = + let data = match Tac2env.interp_type tpe with + | (_, GTydAlg data) -> data + | _ -> assert false + in + let (id, _) = find_constructor n is_const data.galg_constructors in + let kn = change_kn_label tpe id in + pr_constructor kn + +let order_branches cbr nbr def = + let rec order cidx nidx def = match def with + | [] -> [] + | (id, []) :: rem -> + let ans = order (succ cidx) nidx rem in + (id, [], cbr.(cidx)) :: ans + | (id, _ :: _) :: rem -> + let ans = order cidx (succ nidx) rem in + let (vars, e) = nbr.(nidx) in + (id, Array.to_list vars, e) :: ans + in + order 0 0 def + +let pr_glbexpr_gen lvl c = + let rec pr_glbexpr lvl = function + | GTacAtm atm -> pr_atom atm + | GTacVar id -> Id.print id + | GTacRef gr -> + let qid = shortest_qualid_of_ltac (TacConstant gr) in + Libnames.pr_qualid qid + | GTacFun (nas, c) -> + let nas = pr_sequence pr_name nas in + let paren = match lvl with + | E0 | E1 | E2 | E3 | E4 -> paren + | E5 -> fun x -> x + in + paren (hov 0 (hov 2 (str "fun" ++ spc () ++ nas) ++ spc () ++ str "=>" ++ spc () ++ + pr_glbexpr E5 c)) + | GTacApp (c, cl) -> + let paren = match lvl with + | E0 -> paren + | E1 | E2 | E3 | E4 | E5 -> fun x -> x + in + paren (hov 2 (pr_glbexpr E1 c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl))) + | GTacLet (mut, bnd, e) -> + let paren = match lvl with + | E0 | E1 | E2 | E3 | E4 -> paren + | E5 -> fun x -> x + in + let mut = if mut then str "rec" ++ spc () else mt () in + let pr_bnd (na, e) = + pr_name na ++ spc () ++ str ":=" ++ spc () ++ hov 2 (pr_glbexpr E5 e) ++ spc () + in + let bnd = prlist_with_sep (fun () -> str "with" ++ spc ()) pr_bnd bnd in + paren (hv 0 (hov 2 (str "let" ++ spc () ++ mut ++ bnd ++ str "in") ++ spc () ++ pr_glbexpr E5 e)) + | GTacCst (Tuple 0, _, _) -> str "()" + | GTacCst (Tuple _, _, cl) -> + let paren = match lvl with + | E0 | E1 -> paren + | E2 | E3 | E4 | E5 -> fun x -> x + in + paren (prlist_with_sep (fun () -> str "," ++ spc ()) (pr_glbexpr E1) cl) + | GTacCst (Other tpe, n, cl) -> + pr_applied_constructor lvl tpe n cl + | GTacCse (e, info, cst_br, ncst_br) -> + let e = pr_glbexpr E5 e in + let br = match info with + | Other kn -> + let def = match Tac2env.interp_type kn with + | _, GTydAlg { galg_constructors = def } -> def + | _, GTydDef _ | _, GTydRec _ | _, GTydOpn -> assert false + in + let br = order_branches cst_br ncst_br def in + let pr_branch (cstr, vars, p) = + let cstr = change_kn_label kn cstr in + let cstr = pr_constructor cstr in + let vars = match vars with + | [] -> mt () + | _ -> spc () ++ pr_sequence pr_name vars + in + hov 4 (str "|" ++ spc () ++ hov 0 (cstr ++ vars ++ spc () ++ str "=>") ++ spc () ++ + hov 2 (pr_glbexpr E5 p)) ++ spc () + in + prlist pr_branch br + | Tuple n -> + let (vars, p) = if Int.equal n 0 then ([||], cst_br.(0)) else ncst_br.(0) in + let p = pr_glbexpr E5 p in + let vars = prvect_with_sep (fun () -> str "," ++ spc ()) pr_name vars in + hov 4 (str "|" ++ spc () ++ hov 0 (paren vars ++ spc () ++ str "=>") ++ spc () ++ p) + in + v 0 (hv 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ br ++ spc () ++ str "end") + | GTacWth wth -> + let e = pr_glbexpr E5 wth.opn_match in + let pr_pattern c self vars p = + let self = match self with + | Anonymous -> mt () + | Name id -> spc () ++ str "as" ++ spc () ++ Id.print id + in + hov 4 (str "|" ++ spc () ++ hov 0 (c ++ vars ++ self ++ spc () ++ str "=>") ++ spc () ++ + hov 2 (pr_glbexpr E5 p)) ++ spc () + in + let pr_branch (cstr, (self, vars, p)) = + let cstr = pr_constructor cstr in + let vars = match Array.to_list vars with + | [] -> mt () + | vars -> spc () ++ pr_sequence pr_name vars + in + pr_pattern cstr self vars p + in + let br = prlist pr_branch (KNmap.bindings wth.opn_branch) in + let (def_as, def_p) = wth.opn_default in + let def = pr_pattern (str "_") def_as (mt ()) def_p in + let br = br ++ def in + v 0 (hv 0 (str "match" ++ spc () ++ e ++ spc () ++ str "with") ++ spc () ++ br ++ str "end") + | GTacPrj (kn, e, n) -> + let def = match Tac2env.interp_type kn with + | _, GTydRec def -> def + | _, GTydDef _ | _, GTydAlg _ | _, GTydOpn -> assert false + in + let (proj, _, _) = List.nth def n in + let proj = change_kn_label kn proj in + let proj = pr_projection proj in + let e = pr_glbexpr E0 e in + hov 0 (e ++ str "." ++ paren proj) + | GTacSet (kn, e, n, r) -> + let def = match Tac2env.interp_type kn with + | _, GTydRec def -> def + | _, GTydDef _ | _, GTydAlg _ | _, GTydOpn -> assert false + in + let (proj, _, _) = List.nth def n in + let proj = change_kn_label kn proj in + let proj = pr_projection proj in + let e = pr_glbexpr E0 e in + let r = pr_glbexpr E1 r in + hov 0 (e ++ str "." ++ paren proj ++ spc () ++ str ":=" ++ spc () ++ r) + | GTacOpn (kn, cl) -> + let paren = match lvl with + | E0 -> paren + | E1 | E2 | E3 | E4 | E5 -> fun x -> x + in + let c = pr_constructor kn in + paren (hov 0 (c ++ spc () ++ (pr_sequence (pr_glbexpr E0) cl))) + | GTacExt (tag, arg) -> + let tpe = interp_ml_object tag in + hov 0 (tpe.ml_print (Global.env ()) arg) (* FIXME *) + | GTacPrm (prm, args) -> + let args = match args with + | [] -> mt () + | _ -> spc () ++ pr_sequence (pr_glbexpr E0) args + in + hov 0 (str "@external" ++ spc () ++ qstring prm.mltac_plugin ++ spc () ++ + qstring prm.mltac_tactic ++ args) + and pr_applied_constructor lvl tpe n cl = + let _, data = Tac2env.interp_type tpe in + if KerName.equal tpe t_list then + let rec factorize accu = function + | GTacCst (_, 0, []) -> accu, None + | GTacCst (_, 0, [e; l]) -> factorize (e :: accu) l + | e -> accu, Some e + in + let l, e = factorize [] (GTacCst (Other tpe, n, cl)) in + match e with + | None -> + let pr e = pr_glbexpr E4 e in + hov 2 (str "[" ++ prlist_with_sep pr_semicolon pr (List.rev l) ++ str "]") + | Some e -> + let paren = match lvl with + | E0 | E1 | E2 -> paren + | E3 | E4 | E5 -> fun x -> x + in + let pr e = pr_glbexpr E1 e in + let pr_cons () = spc () ++ str "::" ++ spc () in + paren (hov 2 (prlist_with_sep pr_cons pr (List.rev (e :: l)))) + else match data with + | GTydAlg def -> + let paren = match lvl with + | E0 -> + if List.is_empty cl then fun x -> x else paren + | E1 | E2 | E3 | E4 | E5 -> fun x -> x + in + let cstr = pr_internal_constructor tpe n (List.is_empty cl) in + let cl = match cl with + | [] -> mt () + | _ -> spc () ++ pr_sequence (pr_glbexpr E0) cl + in + paren (hov 2 (cstr ++ cl)) + | GTydRec def -> + let args = List.combine def cl in + let pr_arg ((id, _, _), arg) = + let kn = change_kn_label tpe id in + pr_projection kn ++ spc () ++ str ":=" ++ spc () ++ pr_glbexpr E1 arg + in + let args = prlist_with_sep pr_semicolon pr_arg args in + hv 0 (str "{" ++ spc () ++ args ++ spc () ++ str "}") + | (GTydDef _ | GTydOpn) -> assert false + in + hov 0 (pr_glbexpr lvl c) + + + +let pr_glbexpr c = + pr_glbexpr_gen E5 c + +(** Toplevel printers *) + +let rec subst_type subst (t : 'a glb_typexpr) = match t with +| GTypVar id -> subst.(id) +| GTypArrow (t1, t2) -> GTypArrow (subst_type subst t1, subst_type subst t2) +| GTypRef (qid, args) -> + GTypRef (qid, List.map (fun t -> subst_type subst t) args) + +let unfold kn args = + let (nparams, def) = Tac2env.interp_type kn in + match def with + | GTydDef (Some def) -> + let args = Array.of_list args in + Some (subst_type args def) + | _ -> None + +let rec kind t = match t with +| GTypVar id -> GTypVar id +| GTypRef (Other kn, tl) -> + begin match unfold kn tl with + | None -> t + | Some t -> kind t + end +| GTypArrow _ | GTypRef (Tuple _, _) -> t + +type val_printer = + { val_printer : 'a. Environ.env -> Evd.evar_map -> valexpr -> 'a glb_typexpr list -> Pp.t } + +let printers = ref KNmap.empty + +let register_val_printer kn pr = + printers := KNmap.add kn pr !printers + +open Tac2ffi + +let rec pr_valexpr env sigma v t = match kind t with +| GTypVar _ -> str "<poly>" +| GTypRef (Other kn, params) -> + let pr = try Some (KNmap.find kn !printers) with Not_found -> None in + begin match pr with + | Some pr -> pr.val_printer env sigma v params + | None -> + let n, repr = Tac2env.interp_type kn in + if KerName.equal kn t_list then + pr_val_list env sigma (to_list (fun v -> repr_to valexpr v) v) (List.hd params) + else match repr with + | GTydDef None -> str "<abstr>" + | GTydDef (Some _) -> + (* Shouldn't happen thanks to kind *) + assert false + | GTydAlg alg -> + if Valexpr.is_int v then + pr_internal_constructor kn (Tac2ffi.to_int v) true + else + let (n, args) = Tac2ffi.to_block v in + let (id, tpe) = find_constructor n false alg.galg_constructors in + let knc = change_kn_label kn id in + let args = pr_constrargs env sigma params args tpe in + hv 2 (pr_constructor knc ++ spc () ++ str "(" ++ args ++ str ")") + | GTydRec rcd -> + let (_, args) = Tac2ffi.to_block v in + pr_record env sigma params args rcd + | GTydOpn -> + begin match Tac2ffi.to_open v with + | (knc, [||]) -> pr_constructor knc + | (knc, args) -> + let data = Tac2env.interp_constructor knc in + let args = pr_constrargs env sigma params args data.Tac2env.cdata_args in + hv 2 (pr_constructor knc ++ spc () ++ str "(" ++ args ++ str ")") + end + end +| GTypArrow _ -> str "<fun>" +| GTypRef (Tuple 0, []) -> str "()" +| GTypRef (Tuple _, tl) -> + let blk = Array.to_list (snd (to_block v)) in + if List.length blk == List.length tl then + let prs = List.map2 (fun v t -> pr_valexpr env sigma v t) blk tl in + hv 2 (str "(" ++ prlist_with_sep pr_comma (fun p -> p) prs ++ str ")") + else + str "<unknown>" + +and pr_constrargs env sigma params args tpe = + let subst = Array.of_list params in + let tpe = List.map (fun t -> subst_type subst t) tpe in + let args = Array.to_list args in + let args = List.combine args tpe in + prlist_with_sep pr_comma (fun (v, t) -> pr_valexpr env sigma v t) args + +and pr_record env sigma params args rcd = + let subst = Array.of_list params in + let map (id, _, tpe) = (id, subst_type subst tpe) in + let rcd = List.map map rcd in + let args = Array.to_list args in + let fields = List.combine rcd args in + let pr_field ((id, t), arg) = + Id.print id ++ spc () ++ str ":=" ++ spc () ++ pr_valexpr env sigma arg t + in + str "{" ++ spc () ++ prlist_with_sep pr_semicolon pr_field fields ++ spc () ++ str "}" + +and pr_val_list env sigma args tpe = + let pr v = pr_valexpr env sigma v tpe in + str "[" ++ prlist_with_sep pr_semicolon pr args ++ str "]" + +let register_init n f = + let kn = KerName.make Tac2env.coq_prefix (Label.make n) in + register_val_printer kn { val_printer = fun env sigma v _ -> f env sigma v } + +let () = register_init "int" begin fun _ _ n -> + let n = to_int n in + Pp.int n +end + +let () = register_init "string" begin fun _ _ s -> + let s = to_string s in + Pp.quote (str (Bytes.to_string s)) +end + +let () = register_init "ident" begin fun _ _ id -> + let id = to_ident id in + str "@" ++ Id.print id +end + +let () = register_init "constr" begin fun env sigma c -> + let c = to_constr c in + let c = try Printer.pr_leconstr_env env sigma c with _ -> str "..." in + str "constr:(" ++ c ++ str ")" +end + +let () = register_init "pattern" begin fun env sigma c -> + let c = to_pattern c in + let c = try Printer.pr_lconstr_pattern_env env sigma c with _ -> str "..." in + str "pattern:(" ++ c ++ str ")" +end + +let () = register_init "message" begin fun _ _ pp -> + str "message:(" ++ to_pp pp ++ str ")" +end + +let () = register_init "err" begin fun _ _ e -> + let e = to_ext val_exn e in + let (e, _) = ExplainErr.process_vernac_interp_error ~allow_uncaught:true e in + str "err:(" ++ CErrors.print_no_report e ++ str ")" +end + +let () = + let kn = KerName.make Tac2env.coq_prefix (Label.make "array") in + let val_printer env sigma v arg = match arg with + | [arg] -> + let (_, v) = to_block v in + str "[|" ++ spc () ++ + prvect_with_sep pr_semicolon (fun a -> pr_valexpr env sigma a arg) v ++ + spc () ++ str "|]" + | _ -> assert false + in + register_val_printer kn { val_printer } diff --git a/user-contrib/Ltac2/tac2print.mli b/user-contrib/Ltac2/tac2print.mli new file mode 100644 index 0000000000..9b9db2937d --- /dev/null +++ b/user-contrib/Ltac2/tac2print.mli @@ -0,0 +1,46 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Tac2expr +open Tac2ffi + +(** {5 Printing types} *) + +type typ_level = +| T5_l +| T5_r +| T2 +| T1 +| T0 + +val pr_typref : type_constant -> Pp.t +val pr_glbtype_gen : ('a -> string) -> typ_level -> 'a glb_typexpr -> Pp.t +val pr_glbtype : ('a -> string) -> 'a glb_typexpr -> Pp.t + +(** {5 Printing expressions} *) + +val pr_constructor : ltac_constructor -> Pp.t +val pr_internal_constructor : type_constant -> int -> bool -> Pp.t +val pr_projection : ltac_projection -> Pp.t +val pr_glbexpr_gen : exp_level -> glb_tacexpr -> Pp.t +val pr_glbexpr : glb_tacexpr -> Pp.t + +(** {5 Printing values}*) + +type val_printer = + { val_printer : 'a. Environ.env -> Evd.evar_map -> valexpr -> 'a glb_typexpr list -> Pp.t } + +val register_val_printer : type_constant -> val_printer -> unit + +val pr_valexpr : Environ.env -> Evd.evar_map -> valexpr -> 'a glb_typexpr -> Pp.t + +(** {5 Utilities} *) + +val int_name : unit -> (int -> string) +(** Create a function that give names to integers. The names are generated on + the fly, in the order they are encountered. *) diff --git a/user-contrib/Ltac2/tac2qexpr.mli b/user-contrib/Ltac2/tac2qexpr.mli new file mode 100644 index 0000000000..400ab1a092 --- /dev/null +++ b/user-contrib/Ltac2/tac2qexpr.mli @@ -0,0 +1,173 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open Tac2expr + +(** Quoted variants of Ltac syntactic categories. Contrarily to the former, they + sometimes allow anti-quotations. Used for notation scopes. *) + +type 'a or_anti = +| QExpr of 'a +| QAnti of Id.t CAst.t + +type reference_r = +| QReference of Libnames.qualid +| QHypothesis of Id.t + +type reference = reference_r CAst.t + +type quantified_hypothesis = +| QAnonHyp of int CAst.t +| QNamedHyp of Id.t CAst.t + +type bindings_r = +| QImplicitBindings of Constrexpr.constr_expr list +| QExplicitBindings of (quantified_hypothesis CAst.t or_anti * Constrexpr.constr_expr) CAst.t list +| QNoBindings + +type bindings = bindings_r CAst.t + +type intro_pattern_r = +| QIntroForthcoming of bool +| QIntroNaming of intro_pattern_naming +| QIntroAction of intro_pattern_action +and intro_pattern_naming_r = +| QIntroIdentifier of Id.t CAst.t or_anti +| QIntroFresh of Id.t CAst.t or_anti +| QIntroAnonymous +and intro_pattern_action_r = +| QIntroWildcard +| QIntroOrAndPattern of or_and_intro_pattern +| QIntroInjection of intro_pattern list CAst.t +(* | QIntroApplyOn of Empty.t (** Not implemented yet *) *) +| QIntroRewrite of bool +and or_and_intro_pattern_r = +| QIntroOrPattern of intro_pattern list CAst.t list +| QIntroAndPattern of intro_pattern list CAst.t + +and intro_pattern = intro_pattern_r CAst.t +and intro_pattern_naming = intro_pattern_naming_r CAst.t +and intro_pattern_action = intro_pattern_action_r CAst.t +and or_and_intro_pattern = or_and_intro_pattern_r CAst.t + +type occurrences_r = +| QAllOccurrences +| QAllOccurrencesBut of int CAst.t or_anti list +| QNoOccurrences +| QOnlyOccurrences of int CAst.t or_anti list + +type occurrences = occurrences_r CAst.t + +type hyp_location = (occurrences * Id.t CAst.t or_anti) * Locus.hyp_location_flag + +type clause_r = + { q_onhyps : hyp_location list option; q_concl_occs : occurrences; } + +type clause = clause_r CAst.t + +type constr_with_bindings = (Constrexpr.constr_expr * bindings) CAst.t + +type destruction_arg_r = +| QElimOnConstr of constr_with_bindings +| QElimOnIdent of Id.t CAst.t +| QElimOnAnonHyp of int CAst.t + +type destruction_arg = destruction_arg_r CAst.t + +type induction_clause_r = { + indcl_arg : destruction_arg; + indcl_eqn : intro_pattern_naming option; + indcl_as : or_and_intro_pattern option; + indcl_in : clause option; +} + +type induction_clause = induction_clause_r CAst.t + +type conversion_r = +| QConvert of Constrexpr.constr_expr +| QConvertWith of Constrexpr.constr_expr * Constrexpr.constr_expr + +type conversion = conversion_r CAst.t + +type multi_r = +| QPrecisely of int CAst.t +| QUpTo of int CAst.t +| QRepeatStar +| QRepeatPlus + +type multi = multi_r CAst.t + +type rewriting_r = { + rew_orient : bool option CAst.t; + rew_repeat : multi; + rew_equatn : constr_with_bindings; +} + +type rewriting = rewriting_r CAst.t + +type dispatch_r = raw_tacexpr option list * (raw_tacexpr option * raw_tacexpr option list) option + +type dispatch = dispatch_r CAst.t + +type red_flag_r = +| QBeta +| QIota +| QMatch +| QFix +| QCofix +| QZeta +| QConst of reference or_anti list CAst.t +| QDeltaBut of reference or_anti list CAst.t + +type red_flag = red_flag_r CAst.t + +type strategy_flag = red_flag list CAst.t + +type constr_match_pattern_r = +| QConstrMatchPattern of Constrexpr.constr_expr +| QConstrMatchContext of Id.t option * Constrexpr.constr_expr + +type constr_match_pattern = constr_match_pattern_r CAst.t + +type constr_match_branch = (constr_match_pattern * raw_tacexpr) CAst.t + +type constr_matching = constr_match_branch list CAst.t + +type goal_match_pattern_r = { + q_goal_match_concl : constr_match_pattern; + q_goal_match_hyps : (Names.lname * constr_match_pattern) list; +} + +type goal_match_pattern = goal_match_pattern_r CAst.t + +type goal_match_branch = (goal_match_pattern * raw_tacexpr) CAst.t + +type goal_matching = goal_match_branch list CAst.t + +type hintdb_r = +| QHintAll +| QHintDbs of Id.t CAst.t or_anti list + +type hintdb = hintdb_r CAst.t + +type move_location_r = +| QMoveAfter of Id.t CAst.t or_anti +| QMoveBefore of Id.t CAst.t or_anti +| QMoveFirst +| QMoveLast + +type move_location = move_location_r CAst.t + +type pose = (Id.t CAst.t or_anti option * Constrexpr.constr_expr) CAst.t + +type assertion_r = +| QAssertType of intro_pattern option * Constrexpr.constr_expr * raw_tacexpr option +| QAssertValue of Id.t CAst.t or_anti * Constrexpr.constr_expr + +type assertion = assertion_r CAst.t diff --git a/user-contrib/Ltac2/tac2quote.ml b/user-contrib/Ltac2/tac2quote.ml new file mode 100644 index 0000000000..a98264745e --- /dev/null +++ b/user-contrib/Ltac2/tac2quote.ml @@ -0,0 +1,465 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Pp +open Names +open Util +open CAst +open Tac2dyn +open Tac2expr +open Tac2qexpr + +(** Generic arguments *) + +let wit_pattern = Arg.create "pattern" +let wit_reference = Arg.create "reference" +let wit_ident = Arg.create "ident" +let wit_constr = Arg.create "constr" +let wit_open_constr = Arg.create "open_constr" +let wit_ltac1 = Arg.create "ltac1" +let wit_ltac1val = Arg.create "ltac1val" + +(** Syntactic quoting of expressions. *) + +let prefix_gen n = + MPfile (DirPath.make (List.map Id.of_string [n; "Ltac2"])) + +let control_prefix = prefix_gen "Control" +let pattern_prefix = prefix_gen "Pattern" +let array_prefix = prefix_gen "Array" + +let kername prefix n = KerName.make prefix (Label.of_id (Id.of_string_soft n)) +let std_core n = kername Tac2env.std_prefix n +let coq_core n = kername Tac2env.coq_prefix n +let control_core n = kername control_prefix n +let pattern_core n = kername pattern_prefix n + +let global_ref ?loc kn = + CAst.make ?loc @@ CTacRef (AbsKn (TacConstant kn)) + +let constructor ?loc kn args = + let cst = CAst.make ?loc @@ CTacCst (AbsKn (Other kn)) in + if List.is_empty args then cst + else CAst.make ?loc @@ CTacApp (cst, args) + +let std_constructor ?loc name args = + constructor ?loc (std_core name) args + +let std_proj ?loc name = + AbsKn (std_core name) + +let thunk e = + let t_unit = coq_core "unit" in + let loc = e.loc in + let ty = CAst.make?loc @@ CTypRef (AbsKn (Other t_unit), []) in + let pat = CAst.make ?loc @@ CPatVar (Anonymous) in + let pat = CAst.make ?loc @@ CPatCnv (pat, ty) in + CAst.make ?loc @@ CTacFun ([pat], e) + +let of_pair f g {loc;v=(e1, e2)} = + CAst.make ?loc @@ CTacApp (CAst.make ?loc @@ CTacCst (AbsKn (Tuple 2)), [f e1; g e2]) + +let of_tuple ?loc el = match el with +| [] -> + CAst.make ?loc @@ CTacCst (AbsKn (Tuple 0)) +| [e] -> e +| el -> + let len = List.length el in + CAst.make ?loc @@ CTacApp (CAst.make ?loc @@ CTacCst (AbsKn (Tuple len)), el) + +let of_int {loc;v=n} = + CAst.make ?loc @@ CTacAtm (AtmInt n) + +let of_option ?loc f opt = match opt with +| None -> constructor ?loc (coq_core "None") [] +| Some e -> constructor ?loc (coq_core "Some") [f e] + +let inj_wit ?loc wit x = + CAst.make ?loc @@ CTacExt (wit, x) + +let of_variable {loc;v=id} = + let qid = Libnames.qualid_of_ident ?loc id in + if Tac2env.is_constructor qid then + CErrors.user_err ?loc (str "Invalid identifier") + else CAst.make ?loc @@ CTacRef (RelId qid) + +let of_anti f = function +| QExpr x -> f x +| QAnti id -> of_variable id + +let of_ident {loc;v=id} = inj_wit ?loc wit_ident id + +let of_constr c = + let loc = Constrexpr_ops.constr_loc c in + inj_wit ?loc wit_constr c + +let of_open_constr c = + let loc = Constrexpr_ops.constr_loc c in + inj_wit ?loc wit_open_constr c + +let of_bool ?loc b = + let c = if b then coq_core "true" else coq_core "false" in + constructor ?loc c [] + +let rec of_list ?loc f = function +| [] -> constructor (coq_core "[]") [] +| e :: l -> + constructor ?loc (coq_core "::") [f e; of_list ?loc f l] + +let of_qhyp {loc;v=h} = match h with +| QAnonHyp n -> std_constructor ?loc "AnonHyp" [of_int n] +| QNamedHyp id -> std_constructor ?loc "NamedHyp" [of_ident id] + +let of_bindings {loc;v=b} = match b with +| QNoBindings -> + std_constructor ?loc "NoBindings" [] +| QImplicitBindings tl -> + std_constructor ?loc "ImplicitBindings" [of_list ?loc of_open_constr tl] +| QExplicitBindings tl -> + let map e = of_pair (fun q -> of_anti of_qhyp q) of_open_constr e in + std_constructor ?loc "ExplicitBindings" [of_list ?loc map tl] + +let of_constr_with_bindings c = of_pair of_open_constr of_bindings c + +let rec of_intro_pattern {loc;v=pat} = match pat with +| QIntroForthcoming b -> + std_constructor ?loc "IntroForthcoming" [of_bool b] +| QIntroNaming iname -> + std_constructor ?loc "IntroNaming" [of_intro_pattern_naming iname] +| QIntroAction iact -> + std_constructor ?loc "IntroAction" [of_intro_pattern_action iact] + +and of_intro_pattern_naming {loc;v=pat} = match pat with +| QIntroIdentifier id -> + std_constructor ?loc "IntroIdentifier" [of_anti of_ident id] +| QIntroFresh id -> + std_constructor ?loc "IntroFresh" [of_anti of_ident id] +| QIntroAnonymous -> + std_constructor ?loc "IntroAnonymous" [] + +and of_intro_pattern_action {loc;v=pat} = match pat with +| QIntroWildcard -> + std_constructor ?loc "IntroWildcard" [] +| QIntroOrAndPattern pat -> + std_constructor ?loc "IntroOrAndPattern" [of_or_and_intro_pattern pat] +| QIntroInjection il -> + std_constructor ?loc "IntroInjection" [of_intro_patterns il] +| QIntroRewrite b -> + std_constructor ?loc "IntroRewrite" [of_bool ?loc b] + +and of_or_and_intro_pattern {loc;v=pat} = match pat with +| QIntroOrPattern ill -> + std_constructor ?loc "IntroOrPattern" [of_list ?loc of_intro_patterns ill] +| QIntroAndPattern il -> + std_constructor ?loc "IntroAndPattern" [of_intro_patterns il] + +and of_intro_patterns {loc;v=l} = + of_list ?loc of_intro_pattern l + +let of_hyp_location_flag ?loc = function +| Locus.InHyp -> std_constructor ?loc "InHyp" [] +| Locus.InHypTypeOnly -> std_constructor ?loc "InHypTypeOnly" [] +| Locus.InHypValueOnly -> std_constructor ?loc "InHypValueOnly" [] + +let of_occurrences {loc;v=occ} = match occ with +| QAllOccurrences -> std_constructor ?loc "AllOccurrences" [] +| QAllOccurrencesBut occs -> + let map occ = of_anti of_int occ in + let occs = of_list ?loc map occs in + std_constructor ?loc "AllOccurrencesBut" [occs] +| QNoOccurrences -> std_constructor ?loc "NoOccurrences" [] +| QOnlyOccurrences occs -> + let map occ = of_anti of_int occ in + let occs = of_list ?loc map occs in + std_constructor ?loc "OnlyOccurrences" [occs] + +let of_hyp_location ?loc ((occs, id), flag) = + of_tuple ?loc [ + of_anti of_ident id; + of_occurrences occs; + of_hyp_location_flag ?loc flag; + ] + +let of_clause {loc;v=cl} = + let hyps = of_option ?loc (fun l -> of_list ?loc of_hyp_location l) cl.q_onhyps in + let concl = of_occurrences cl.q_concl_occs in + CAst.make ?loc @@ CTacRec ([ + std_proj "on_hyps", hyps; + std_proj "on_concl", concl; + ]) + +let of_destruction_arg {loc;v=arg} = match arg with +| QElimOnConstr c -> + let arg = thunk (of_constr_with_bindings c) in + std_constructor ?loc "ElimOnConstr" [arg] +| QElimOnIdent id -> std_constructor ?loc "ElimOnIdent" [of_ident id] +| QElimOnAnonHyp n -> std_constructor ?loc "ElimOnAnonHyp" [of_int n] + +let of_induction_clause {loc;v=cl} = + let arg = of_destruction_arg cl.indcl_arg in + let eqn = of_option ?loc of_intro_pattern_naming cl.indcl_eqn in + let as_ = of_option ?loc of_or_and_intro_pattern cl.indcl_as in + let in_ = of_option ?loc of_clause cl.indcl_in in + CAst.make ?loc @@ CTacRec ([ + std_proj "indcl_arg", arg; + std_proj "indcl_eqn", eqn; + std_proj "indcl_as", as_; + std_proj "indcl_in", in_; + ]) + +let check_pattern_id ?loc id = + if Tac2env.is_constructor (Libnames.qualid_of_ident id) then + CErrors.user_err ?loc (str "Invalid pattern binding name " ++ Id.print id) + +let pattern_vars pat = + let rec aux () accu pat = match pat.CAst.v with + | Constrexpr.CPatVar id + | Constrexpr.CEvar (id, []) -> + let () = check_pattern_id ?loc:pat.CAst.loc id in + Id.Set.add id accu + | _ -> + Constrexpr_ops.fold_constr_expr_with_binders (fun _ () -> ()) aux () accu pat + in + aux () Id.Set.empty pat + +let abstract_vars loc vars tac = + let get_name = function Name id -> Some id | Anonymous -> None in + let def = try Some (List.find_map get_name vars) with Not_found -> None in + let na, tac = match def with + | None -> (Anonymous, tac) + | Some id0 -> + (* Trick: in order not to shadow a variable nor to choose an arbitrary + name, we reuse one which is going to be shadowed by the matched + variables anyways. *) + let build_bindings (n, accu) na = match na with + | Anonymous -> (n + 1, accu) + | Name _ -> + let get = global_ref ?loc (kername array_prefix "get") in + let args = [of_variable CAst.(make ?loc id0); of_int CAst.(make ?loc n)] in + let e = CAst.make ?loc @@ CTacApp (get, args) in + let accu = (CAst.make ?loc @@ CPatVar na, e) :: accu in + (n + 1, accu) + in + let (_, bnd) = List.fold_left build_bindings (0, []) vars in + let tac = CAst.make ?loc @@ CTacLet (false, bnd, tac) in + (Name id0, tac) + in + CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar na], tac) + +let of_pattern p = + inj_wit ?loc:p.CAst.loc wit_pattern p + +let of_conversion {loc;v=c} = match c with +| QConvert c -> + let pat = of_option ?loc of_pattern None in + let c = CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar Anonymous], of_constr c) in + of_tuple ?loc [pat; c] +| QConvertWith (pat, c) -> + let vars = pattern_vars pat in + let pat = of_option ?loc of_pattern (Some pat) in + let c = of_constr c in + (* Order is critical here *) + let vars = List.map (fun id -> Name id) (Id.Set.elements vars) in + let c = abstract_vars loc vars c in + of_tuple [pat; c] + +let of_repeat {loc;v=r} = match r with +| QPrecisely n -> std_constructor ?loc "Precisely" [of_int n] +| QUpTo n -> std_constructor ?loc "UpTo" [of_int n] +| QRepeatStar -> std_constructor ?loc "RepeatStar" [] +| QRepeatPlus -> std_constructor ?loc "RepeatPlus" [] + +let of_orient loc b = + if b then std_constructor ?loc "LTR" [] + else std_constructor ?loc "RTL" [] + +let of_rewriting {loc;v=rew} = + let orient = + let {loc;v=orient} = rew.rew_orient in + of_option ?loc (fun b -> of_orient loc b) orient + in + let repeat = of_repeat rew.rew_repeat in + let equatn = thunk (of_constr_with_bindings rew.rew_equatn) in + CAst.make ?loc @@ CTacRec ([ + std_proj "rew_orient", orient; + std_proj "rew_repeat", repeat; + std_proj "rew_equatn", equatn; + ]) + +let of_hyp ?loc id = + let hyp = global_ref ?loc (control_core "hyp") in + CAst.make ?loc @@ CTacApp (hyp, [of_ident id]) + +let of_exact_hyp ?loc id = + let refine = global_ref ?loc (control_core "refine") in + CAst.make ?loc @@ CTacApp (refine, [thunk (of_hyp ?loc id)]) + +let of_exact_var ?loc id = + let refine = global_ref ?loc (control_core "refine") in + CAst.make ?loc @@ CTacApp (refine, [thunk (of_variable id)]) + +let of_dispatch tacs = + let loc = tacs.loc in + let default = function + | Some e -> thunk e + | None -> thunk (CAst.make ?loc @@ CTacCst (AbsKn (Tuple 0))) + in + let map e = of_pair default (fun l -> of_list ?loc default l) (CAst.make ?loc e) in + of_pair (fun l -> of_list ?loc default l) (fun r -> of_option ?loc map r) tacs + +let make_red_flag l = + let open Genredexpr in + let rec add_flag red = function + | [] -> red + | {v=flag} :: lf -> + let red = match flag with + | QBeta -> { red with rBeta = true } + | QMatch -> { red with rMatch = true } + | QFix -> { red with rFix = true } + | QCofix -> { red with rCofix = true } + | QZeta -> { red with rZeta = true } + | QConst {loc;v=l} -> + if red.rDelta then + CErrors.user_err ?loc Pp.(str + "Cannot set both constants to unfold and constants not to unfold"); + { red with rConst = red.rConst @ l } + | QDeltaBut {loc;v=l} -> + if red.rConst <> [] && not red.rDelta then + CErrors.user_err ?loc Pp.(str + "Cannot set both constants to unfold and constants not to unfold"); + { red with rConst = red.rConst @ l; rDelta = true } + | QIota -> + { red with rMatch = true; rFix = true; rCofix = true } + in + add_flag red lf + in + add_flag + {rBeta = false; rMatch = false; rFix = false; rCofix = false; + rZeta = false; rDelta = false; rConst = []} + l + +let of_reference r = + let of_ref ref = + inj_wit ?loc:ref.loc wit_reference ref + in + of_anti of_ref r + +let of_strategy_flag {loc;v=flag} = + let open Genredexpr in + let flag = make_red_flag flag in + CAst.make ?loc @@ CTacRec ([ + std_proj "rBeta", of_bool ?loc flag.rBeta; + std_proj "rMatch", of_bool ?loc flag.rMatch; + std_proj "rFix", of_bool ?loc flag.rFix; + std_proj "rCofix", of_bool ?loc flag.rCofix; + std_proj "rZeta", of_bool ?loc flag.rZeta; + std_proj "rDelta", of_bool ?loc flag.rDelta; + std_proj "rConst", of_list ?loc of_reference flag.rConst; + ]) + +let of_hintdb {loc;v=hdb} = match hdb with +| QHintAll -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) None +| QHintDbs ids -> of_option ?loc (fun l -> of_list (fun id -> of_anti of_ident id) l) (Some ids) + +let extract_name ?loc oid = match oid with +| None -> Anonymous +| Some id -> + let () = check_pattern_id ?loc id in + Name id + +(** For every branch in the matching, generate a corresponding term of type + [(match_kind * pattern * (context -> constr array -> 'a))] + where the function binds the names from the pattern to the contents of the + constr array. *) +let of_constr_matching {loc;v=m} = + let map {loc;v=({loc=ploc;v=pat}, tac)} = + let (knd, pat, na) = match pat with + | QConstrMatchPattern pat -> + let knd = constructor ?loc (pattern_core "MatchPattern") [] in + (knd, pat, Anonymous) + | QConstrMatchContext (id, pat) -> + let na = extract_name ?loc id in + let knd = constructor ?loc (pattern_core "MatchContext") [] in + (knd, pat, na) + in + let vars = pattern_vars pat in + (* Order of elements is crucial here! *) + let vars = Id.Set.elements vars in + let vars = List.map (fun id -> Name id) vars in + let e = abstract_vars loc vars tac in + let e = CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar na], e) in + let pat = inj_wit ?loc:ploc wit_pattern pat in + of_tuple [knd; pat; e] + in + of_list ?loc map m + +(** From the patterns and the body of the branch, generate: + - a goal pattern: (constr_match list * constr_match) + - a branch function (ident array -> context array -> constr array -> context -> 'a) +*) +let of_goal_matching {loc;v=gm} = + let mk_pat {loc;v=p} = match p with + | QConstrMatchPattern pat -> + let knd = constructor ?loc (pattern_core "MatchPattern") [] in + (Anonymous, pat, knd) + | QConstrMatchContext (id, pat) -> + let na = extract_name ?loc id in + let knd = constructor ?loc (pattern_core "MatchContext") [] in + (na, pat, knd) + in + let mk_gpat {loc;v=p} = + let concl_pat = p.q_goal_match_concl in + let hyps_pats = p.q_goal_match_hyps in + let (concl_ctx, concl_pat, concl_knd) = mk_pat concl_pat in + let vars = pattern_vars concl_pat in + let map accu (na, pat) = + let (ctx, pat, knd) = mk_pat pat in + let vars = pattern_vars pat in + (Id.Set.union vars accu, (na, ctx, pat, knd)) + in + let (vars, hyps_pats) = List.fold_left_map map vars hyps_pats in + let map (_, _, pat, knd) = of_tuple [knd; of_pattern pat] in + let concl = of_tuple [concl_knd; of_pattern concl_pat] in + let r = of_tuple [of_list ?loc map hyps_pats; concl] in + let hyps = List.map (fun ({CAst.v=na}, _, _, _) -> na) hyps_pats in + let map (_, na, _, _) = na in + let hctx = List.map map hyps_pats in + (* Order of elements is crucial here! *) + let vars = Id.Set.elements vars in + let subst = List.map (fun id -> Name id) vars in + (r, hyps, hctx, subst, concl_ctx) + in + let map {loc;v=(pat, tac)} = + let (pat, hyps, hctx, subst, cctx) = mk_gpat pat in + let tac = CAst.make ?loc @@ CTacFun ([CAst.make ?loc @@ CPatVar cctx], tac) in + let tac = abstract_vars loc subst tac in + let tac = abstract_vars loc hctx tac in + let tac = abstract_vars loc hyps tac in + of_tuple ?loc [pat; tac] + in + of_list ?loc map gm + +let of_move_location {loc;v=mv} = match mv with +| QMoveAfter id -> std_constructor ?loc "MoveAfter" [of_anti of_ident id] +| QMoveBefore id -> std_constructor ?loc "MoveBefore" [of_anti of_ident id] +| QMoveFirst -> std_constructor ?loc "MoveFirst" [] +| QMoveLast -> std_constructor ?loc "MoveLast" [] + +let of_pose p = + of_pair (fun id -> of_option (fun id -> of_anti of_ident id) id) of_open_constr p + +let of_assertion {loc;v=ast} = match ast with +| QAssertType (ipat, c, tac) -> + let ipat = of_option of_intro_pattern ipat in + let c = of_constr c in + let tac = of_option thunk tac in + std_constructor ?loc "AssertType" [ipat; c; tac] +| QAssertValue (id, c) -> + let id = of_anti of_ident id in + let c = of_constr c in + std_constructor ?loc "AssertValue" [id; c] diff --git a/user-contrib/Ltac2/tac2quote.mli b/user-contrib/Ltac2/tac2quote.mli new file mode 100644 index 0000000000..1b03dad8ec --- /dev/null +++ b/user-contrib/Ltac2/tac2quote.mli @@ -0,0 +1,102 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open Tac2dyn +open Tac2qexpr +open Tac2expr + +(** Syntactic quoting of expressions. *) + +(** Contrarily to Tac2ffi, which lives on the semantic level, this module + manipulates pure syntax of Ltac2. Its main purpose is to write notations. *) + +val constructor : ?loc:Loc.t -> ltac_constructor -> raw_tacexpr list -> raw_tacexpr + +val thunk : raw_tacexpr -> raw_tacexpr + +val of_anti : ('a -> raw_tacexpr) -> 'a or_anti -> raw_tacexpr + +val of_int : int CAst.t -> raw_tacexpr + +val of_pair : ('a -> raw_tacexpr) -> ('b -> raw_tacexpr) -> ('a * 'b) CAst.t -> raw_tacexpr + +val of_tuple : ?loc:Loc.t -> raw_tacexpr list -> raw_tacexpr + +val of_variable : Id.t CAst.t -> raw_tacexpr + +val of_ident : Id.t CAst.t -> raw_tacexpr + +val of_constr : Constrexpr.constr_expr -> raw_tacexpr + +val of_open_constr : Constrexpr.constr_expr -> raw_tacexpr + +val of_list : ?loc:Loc.t -> ('a -> raw_tacexpr) -> 'a list -> raw_tacexpr + +val of_bindings : bindings -> raw_tacexpr + +val of_intro_pattern : intro_pattern -> raw_tacexpr + +val of_intro_patterns : intro_pattern list CAst.t -> raw_tacexpr + +val of_clause : clause -> raw_tacexpr + +val of_destruction_arg : destruction_arg -> raw_tacexpr + +val of_induction_clause : induction_clause -> raw_tacexpr + +val of_conversion : conversion -> raw_tacexpr + +val of_rewriting : rewriting -> raw_tacexpr + +val of_occurrences : occurrences -> raw_tacexpr + +val of_hintdb : hintdb -> raw_tacexpr + +val of_move_location : move_location -> raw_tacexpr + +val of_reference : reference or_anti -> raw_tacexpr + +val of_hyp : ?loc:Loc.t -> Id.t CAst.t -> raw_tacexpr +(** id ↦ 'Control.hyp @id' *) + +val of_exact_hyp : ?loc:Loc.t -> Id.t CAst.t -> raw_tacexpr +(** id ↦ 'Control.refine (fun () => Control.hyp @id') *) + +val of_exact_var : ?loc:Loc.t -> Id.t CAst.t -> raw_tacexpr +(** id ↦ 'Control.refine (fun () => Control.hyp @id') *) + +val of_dispatch : dispatch -> raw_tacexpr + +val of_strategy_flag : strategy_flag -> raw_tacexpr + +val of_pose : pose -> raw_tacexpr + +val of_assertion : assertion -> raw_tacexpr + +val of_constr_matching : constr_matching -> raw_tacexpr + +val of_goal_matching : goal_matching -> raw_tacexpr + +(** {5 Generic arguments} *) + +val wit_pattern : (Constrexpr.constr_expr, Pattern.constr_pattern) Arg.tag + +val wit_ident : (Id.t, Id.t) Arg.tag + +val wit_reference : (reference, GlobRef.t) Arg.tag + +val wit_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag + +val wit_open_constr : (Constrexpr.constr_expr, Glob_term.glob_constr) Arg.tag + +val wit_ltac1 : (Ltac_plugin.Tacexpr.raw_tactic_expr, Ltac_plugin.Tacexpr.glob_tactic_expr) Arg.tag +(** Ltac1 AST quotation, seen as a 'tactic'. Its type is unit in Ltac2. *) + +val wit_ltac1val : (Ltac_plugin.Tacexpr.raw_tactic_expr, Ltac_plugin.Tacexpr.glob_tactic_expr) Arg.tag +(** Ltac1 AST quotation, seen as a value-returning expression, with type Ltac1.t. *) diff --git a/user-contrib/Ltac2/tac2stdlib.ml b/user-contrib/Ltac2/tac2stdlib.ml new file mode 100644 index 0000000000..fb51fc965b --- /dev/null +++ b/user-contrib/Ltac2/tac2stdlib.ml @@ -0,0 +1,572 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open Genredexpr +open Tac2expr +open Tac2ffi +open Tac2types +open Tac2extffi +open Proofview.Notations + +module Value = Tac2ffi + +(** Make a representation with a dummy from function *) +let make_to_repr f = Tac2ffi.make_repr (fun _ -> assert false) f + +let return x = Proofview.tclUNIT x +let v_unit = Value.of_unit () +let thaw r f = Tac2ffi.app_fun1 f unit r () +let uthaw r f = Tac2ffi.app_fun1 (to_fun1 unit r f) unit r () +let thunk r = fun1 unit r + +let to_name c = match Value.to_option Value.to_ident c with +| None -> Anonymous +| Some id -> Name id + +let name = make_to_repr to_name + +let to_occurrences = function +| ValInt 0 -> AllOccurrences +| ValBlk (0, [| vl |]) -> AllOccurrencesBut (Value.to_list Value.to_int vl) +| ValInt 1 -> NoOccurrences +| ValBlk (1, [| vl |]) -> OnlyOccurrences (Value.to_list Value.to_int vl) +| _ -> assert false + +let occurrences = make_to_repr to_occurrences + +let to_hyp_location_flag v = match Value.to_int v with +| 0 -> InHyp +| 1 -> InHypTypeOnly +| 2 -> InHypValueOnly +| _ -> assert false + +let to_clause v = match Value.to_tuple v with +| [| hyps; concl |] -> + let cast v = match Value.to_tuple v with + | [| hyp; occ; flag |] -> + (Value.to_ident hyp, to_occurrences occ, to_hyp_location_flag flag) + | _ -> assert false + in + let hyps = Value.to_option (fun h -> Value.to_list cast h) hyps in + { onhyps = hyps; concl_occs = to_occurrences concl; } +| _ -> assert false + +let clause = make_to_repr to_clause + +let to_red_flag v = match Value.to_tuple v with +| [| beta; iota; fix; cofix; zeta; delta; const |] -> + { + rBeta = Value.to_bool beta; + rMatch = Value.to_bool iota; + rFix = Value.to_bool fix; + rCofix = Value.to_bool cofix; + rZeta = Value.to_bool zeta; + rDelta = Value.to_bool delta; + rConst = Value.to_list Value.to_reference const; + } +| _ -> assert false + +let red_flags = make_to_repr to_red_flag + +let pattern_with_occs = pair pattern occurrences + +let constr_with_occs = pair constr occurrences + +let reference_with_occs = pair reference occurrences + +let rec to_intro_pattern v = match Value.to_block v with +| (0, [| b |]) -> IntroForthcoming (Value.to_bool b) +| (1, [| pat |]) -> IntroNaming (to_intro_pattern_naming pat) +| (2, [| act |]) -> IntroAction (to_intro_pattern_action act) +| _ -> assert false + +and to_intro_pattern_naming = function +| ValBlk (0, [| id |]) -> IntroIdentifier (Value.to_ident id) +| ValBlk (1, [| id |]) -> IntroFresh (Value.to_ident id) +| ValInt 0 -> IntroAnonymous +| _ -> assert false + +and to_intro_pattern_action = function +| ValInt 0 -> IntroWildcard +| ValBlk (0, [| op |]) -> IntroOrAndPattern (to_or_and_intro_pattern op) +| ValBlk (1, [| inj |]) -> + let map ipat = to_intro_pattern ipat in + IntroInjection (Value.to_list map inj) +| ValBlk (2, [| c; ipat |]) -> + let c = Value.to_fun1 Value.unit Value.constr c in + IntroApplyOn (c, to_intro_pattern ipat) +| ValBlk (3, [| b |]) -> IntroRewrite (Value.to_bool b) +| _ -> assert false + +and to_or_and_intro_pattern v = match Value.to_block v with +| (0, [| ill |]) -> + IntroOrPattern (Value.to_list to_intro_patterns ill) +| (1, [| il |]) -> + IntroAndPattern (to_intro_patterns il) +| _ -> assert false + +and to_intro_patterns il = + Value.to_list to_intro_pattern il + +let intro_pattern = make_to_repr to_intro_pattern + +let intro_patterns = make_to_repr to_intro_patterns + +let to_destruction_arg v = match Value.to_block v with +| (0, [| c |]) -> + let c = uthaw constr_with_bindings c in + ElimOnConstr c +| (1, [| id |]) -> ElimOnIdent (Value.to_ident id) +| (2, [| n |]) -> ElimOnAnonHyp (Value.to_int n) +| _ -> assert false + +let destruction_arg = make_to_repr to_destruction_arg + +let to_induction_clause v = match Value.to_tuple v with +| [| arg; eqn; as_; in_ |] -> + let arg = to_destruction_arg arg in + let eqn = Value.to_option to_intro_pattern_naming eqn in + let as_ = Value.to_option to_or_and_intro_pattern as_ in + let in_ = Value.to_option to_clause in_ in + (arg, eqn, as_, in_) +| _ -> + assert false + +let induction_clause = make_to_repr to_induction_clause + +let to_assertion v = match Value.to_block v with +| (0, [| ipat; t; tac |]) -> + let to_tac t = Value.to_fun1 Value.unit Value.unit t in + let ipat = Value.to_option to_intro_pattern ipat in + let t = Value.to_constr t in + let tac = Value.to_option to_tac tac in + AssertType (ipat, t, tac) +| (1, [| id; c |]) -> + AssertValue (Value.to_ident id, Value.to_constr c) +| _ -> assert false + +let assertion = make_to_repr to_assertion + +let to_multi = function +| ValBlk (0, [| n |]) -> Precisely (Value.to_int n) +| ValBlk (1, [| n |]) -> UpTo (Value.to_int n) +| ValInt 0 -> RepeatStar +| ValInt 1 -> RepeatPlus +| _ -> assert false + +let to_rewriting v = match Value.to_tuple v with +| [| orient; repeat; c |] -> + let orient = Value.to_option Value.to_bool orient in + let repeat = to_multi repeat in + let c = uthaw constr_with_bindings c in + (orient, repeat, c) +| _ -> assert false + +let rewriting = make_to_repr to_rewriting + +let to_debug v = match Value.to_int v with +| 0 -> Hints.Off +| 1 -> Hints.Info +| 2 -> Hints.Debug +| _ -> assert false + +let debug = make_to_repr to_debug + +let to_strategy v = match Value.to_int v with +| 0 -> Class_tactics.Bfs +| 1 -> Class_tactics.Dfs +| _ -> assert false + +let strategy = make_to_repr to_strategy + +let to_inversion_kind v = match Value.to_int v with +| 0 -> Inv.SimpleInversion +| 1 -> Inv.FullInversion +| 2 -> Inv.FullInversionClear +| _ -> assert false + +let inversion_kind = make_to_repr to_inversion_kind + +let to_move_location = function +| ValInt 0 -> Logic.MoveFirst +| ValInt 1 -> Logic.MoveLast +| ValBlk (0, [|id|]) -> Logic.MoveAfter (Value.to_ident id) +| ValBlk (1, [|id|]) -> Logic.MoveBefore (Value.to_ident id) +| _ -> assert false + +let move_location = make_to_repr to_move_location + +let to_generalize_arg v = match Value.to_tuple v with +| [| c; occs; na |] -> + (Value.to_constr c, to_occurrences occs, to_name na) +| _ -> assert false + +let generalize_arg = make_to_repr to_generalize_arg + +(** Standard tactics sharing their implementation with Ltac1 *) + +let pname s = { mltac_plugin = "ltac2"; mltac_tactic = s } + +let lift tac = tac <*> return v_unit + +let define_prim0 name tac = + let tac _ = lift tac in + Tac2env.define_primitive (pname name) (mk_closure arity_one tac) + +let define_prim1 name r0 f = + let tac x = lift (f (Value.repr_to r0 x)) in + Tac2env.define_primitive (pname name) (mk_closure arity_one tac) + +let define_prim2 name r0 r1 f = + let tac x y = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y)) in + Tac2env.define_primitive (pname name) (mk_closure (arity_suc arity_one) tac) + +let define_prim3 name r0 r1 r2 f = + let tac x y z = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z)) in + Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc arity_one)) tac) + +let define_prim4 name r0 r1 r2 r3 f = + let tac x y z u = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z) (Value.repr_to r3 u)) in + Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc (arity_suc arity_one))) tac) + +let define_prim5 name r0 r1 r2 r3 r4 f = + let tac x y z u v = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z) (Value.repr_to r3 u) (Value.repr_to r4 v)) in + Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc (arity_suc (arity_suc arity_one)))) tac) + +(** Tactics from Tacexpr *) + +let () = define_prim2 "tac_intros" bool intro_patterns begin fun ev ipat -> + Tac2tactics.intros_patterns ev ipat +end + +let () = define_prim4 "tac_apply" bool bool (list (thunk constr_with_bindings)) (option (pair ident (option intro_pattern))) begin fun adv ev cb ipat -> + Tac2tactics.apply adv ev cb ipat +end + +let () = define_prim3 "tac_elim" bool constr_with_bindings (option constr_with_bindings) begin fun ev c copt -> + Tac2tactics.elim ev c copt +end + +let () = define_prim2 "tac_case" bool constr_with_bindings begin fun ev c -> + Tac2tactics.general_case_analysis ev c +end + +let () = define_prim1 "tac_generalize" (list generalize_arg) begin fun cl -> + Tac2tactics.generalize cl +end + +let () = define_prim1 "tac_assert" assertion begin fun ast -> + Tac2tactics.assert_ ast +end + +let () = define_prim3 "tac_enough" constr (option (option (thunk unit))) (option intro_pattern) begin fun c tac ipat -> + let tac = Option.map (fun o -> Option.map (fun f -> thaw unit f) o) tac in + Tac2tactics.forward false tac ipat c +end + +let () = define_prim2 "tac_pose" name constr begin fun na c -> + Tactics.letin_tac None na c None Locusops.nowhere +end + +let () = define_prim3 "tac_set" bool (thunk (pair name constr)) clause begin fun ev p cl -> + Proofview.tclEVARMAP >>= fun sigma -> + thaw (pair name constr) p >>= fun (na, c) -> + Tac2tactics.letin_pat_tac ev None na (sigma, c) cl +end + +let () = define_prim5 "tac_remember" bool name (thunk constr) (option intro_pattern) clause begin fun ev na c eqpat cl -> + let eqpat = Option.default (IntroNaming IntroAnonymous) eqpat in + match eqpat with + | IntroNaming eqpat -> + Proofview.tclEVARMAP >>= fun sigma -> + thaw constr c >>= fun c -> + Tac2tactics.letin_pat_tac ev (Some (true, eqpat)) na (sigma, c) cl + | _ -> + Tacticals.New.tclZEROMSG (Pp.str "Invalid pattern for remember") +end + +let () = define_prim3 "tac_destruct" bool (list induction_clause) (option constr_with_bindings) begin fun ev ic using -> + Tac2tactics.induction_destruct false ev ic using +end + +let () = define_prim3 "tac_induction" bool (list induction_clause) (option constr_with_bindings) begin fun ev ic using -> + Tac2tactics.induction_destruct true ev ic using +end + +let () = define_prim1 "tac_red" clause begin fun cl -> + Tac2tactics.reduce (Red false) cl +end + +let () = define_prim1 "tac_hnf" clause begin fun cl -> + Tac2tactics.reduce Hnf cl +end + +let () = define_prim3 "tac_simpl" red_flags (option pattern_with_occs) clause begin fun flags where cl -> + Tac2tactics.simpl flags where cl +end + +let () = define_prim2 "tac_cbv" red_flags clause begin fun flags cl -> + Tac2tactics.cbv flags cl +end + +let () = define_prim2 "tac_cbn" red_flags clause begin fun flags cl -> + Tac2tactics.cbn flags cl +end + +let () = define_prim2 "tac_lazy" red_flags clause begin fun flags cl -> + Tac2tactics.lazy_ flags cl +end + +let () = define_prim2 "tac_unfold" (list reference_with_occs) clause begin fun refs cl -> + Tac2tactics.unfold refs cl +end + +let () = define_prim2 "tac_fold" (list constr) clause begin fun args cl -> + Tac2tactics.reduce (Fold args) cl +end + +let () = define_prim2 "tac_pattern" (list constr_with_occs) clause begin fun where cl -> + Tac2tactics.pattern where cl +end + +let () = define_prim2 "tac_vm" (option pattern_with_occs) clause begin fun where cl -> + Tac2tactics.vm where cl +end + +let () = define_prim2 "tac_native" (option pattern_with_occs) clause begin fun where cl -> + Tac2tactics.native where cl +end + +(** Reduction functions *) + +let lift tac = tac >>= fun c -> Proofview.tclUNIT (Value.of_constr c) + +let define_red1 name r0 f = + let tac x = lift (f (Value.repr_to r0 x)) in + Tac2env.define_primitive (pname name) (mk_closure arity_one tac) + +let define_red2 name r0 r1 f = + let tac x y = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y)) in + Tac2env.define_primitive (pname name) (mk_closure (arity_suc arity_one) tac) + +let define_red3 name r0 r1 r2 f = + let tac x y z = lift (f (Value.repr_to r0 x) (Value.repr_to r1 y) (Value.repr_to r2 z)) in + Tac2env.define_primitive (pname name) (mk_closure (arity_suc (arity_suc arity_one)) tac) + +let () = define_red1 "eval_red" constr begin fun c -> + Tac2tactics.eval_red c +end + +let () = define_red1 "eval_hnf" constr begin fun c -> + Tac2tactics.eval_hnf c +end + +let () = define_red3 "eval_simpl" red_flags (option pattern_with_occs) constr begin fun flags where c -> + Tac2tactics.eval_simpl flags where c +end + +let () = define_red2 "eval_cbv" red_flags constr begin fun flags c -> + Tac2tactics.eval_cbv flags c +end + +let () = define_red2 "eval_cbn" red_flags constr begin fun flags c -> + Tac2tactics.eval_cbn flags c +end + +let () = define_red2 "eval_lazy" red_flags constr begin fun flags c -> + Tac2tactics.eval_lazy flags c +end + +let () = define_red2 "eval_unfold" (list reference_with_occs) constr begin fun refs c -> + Tac2tactics.eval_unfold refs c +end + +let () = define_red2 "eval_fold" (list constr) constr begin fun args c -> + Tac2tactics.eval_fold args c +end + +let () = define_red2 "eval_pattern" (list constr_with_occs) constr begin fun where c -> + Tac2tactics.eval_pattern where c +end + +let () = define_red2 "eval_vm" (option pattern_with_occs) constr begin fun where c -> + Tac2tactics.eval_vm where c +end + +let () = define_red2 "eval_native" (option pattern_with_occs) constr begin fun where c -> + Tac2tactics.eval_native where c +end + +let () = define_prim3 "tac_change" (option pattern) (fun1 (array constr) constr) clause begin fun pat c cl -> + Tac2tactics.change pat c cl +end + +let () = define_prim4 "tac_rewrite" bool (list rewriting) clause (option (thunk unit)) begin fun ev rw cl by -> + Tac2tactics.rewrite ev rw cl by +end + +let () = define_prim4 "tac_inversion" inversion_kind destruction_arg (option intro_pattern) (option (list ident)) begin fun knd arg pat ids -> + Tac2tactics.inversion knd arg pat ids +end + +(** Tactics from coretactics *) + +let () = define_prim0 "tac_reflexivity" Tactics.intros_reflexivity + +let () = define_prim2 "tac_move" ident move_location begin fun id mv -> + Tactics.move_hyp id mv +end + +let () = define_prim2 "tac_intro" (option ident) (option move_location) begin fun id mv -> + let mv = Option.default Logic.MoveLast mv in + Tactics.intro_move id mv +end + +(* + +TACTIC EXTEND exact + [ "exact" casted_constr(c) ] -> [ Tactics.exact_no_check c ] +END + +*) + +let () = define_prim0 "tac_assumption" Tactics.assumption + +let () = define_prim1 "tac_transitivity" constr begin fun c -> + Tactics.intros_transitivity (Some c) +end + +let () = define_prim0 "tac_etransitivity" (Tactics.intros_transitivity None) + +let () = define_prim1 "tac_cut" constr begin fun c -> + Tactics.cut c +end + +let () = define_prim2 "tac_left" bool bindings begin fun ev bnd -> + Tac2tactics.left_with_bindings ev bnd +end +let () = define_prim2 "tac_right" bool bindings begin fun ev bnd -> + Tac2tactics.right_with_bindings ev bnd +end + +let () = define_prim1 "tac_introsuntil" qhyp begin fun h -> + Tactics.intros_until h +end + +let () = define_prim1 "tac_exactnocheck" constr begin fun c -> + Tactics.exact_no_check c +end + +let () = define_prim1 "tac_vmcastnocheck" constr begin fun c -> + Tactics.vm_cast_no_check c +end + +let () = define_prim1 "tac_nativecastnocheck" constr begin fun c -> + Tactics.native_cast_no_check c +end + +let () = define_prim1 "tac_constructor" bool begin fun ev -> + Tactics.any_constructor ev None +end + +let () = define_prim3 "tac_constructorn" bool int bindings begin fun ev n bnd -> + Tac2tactics.constructor_tac ev None n bnd +end + +let () = define_prim2 "tac_specialize" constr_with_bindings (option intro_pattern) begin fun c ipat -> + Tac2tactics.specialize c ipat +end + +let () = define_prim1 "tac_symmetry" clause begin fun cl -> + Tac2tactics.symmetry cl +end + +let () = define_prim2 "tac_split" bool bindings begin fun ev bnd -> + Tac2tactics.split_with_bindings ev bnd +end + +let () = define_prim1 "tac_rename" (list (pair ident ident)) begin fun ids -> + Tactics.rename_hyp ids +end + +let () = define_prim1 "tac_revert" (list ident) begin fun ids -> + Tactics.revert ids +end + +let () = define_prim0 "tac_admit" Proofview.give_up + +let () = define_prim2 "tac_fix" ident int begin fun ident n -> + Tactics.fix ident n +end + +let () = define_prim1 "tac_cofix" ident begin fun ident -> + Tactics.cofix ident +end + +let () = define_prim1 "tac_clear" (list ident) begin fun ids -> + Tactics.clear ids +end + +let () = define_prim1 "tac_keep" (list ident) begin fun ids -> + Tactics.keep ids +end + +let () = define_prim1 "tac_clearbody" (list ident) begin fun ids -> + Tactics.clear_body ids +end + +(** Tactics from extratactics *) + +let () = define_prim2 "tac_discriminate" bool (option destruction_arg) begin fun ev arg -> + Tac2tactics.discriminate ev arg +end + +let () = define_prim3 "tac_injection" bool (option intro_patterns) (option destruction_arg) begin fun ev ipat arg -> + Tac2tactics.injection ev ipat arg +end + +let () = define_prim1 "tac_absurd" constr begin fun c -> + Contradiction.absurd c +end + +let () = define_prim1 "tac_contradiction" (option constr_with_bindings) begin fun c -> + Tac2tactics.contradiction c +end + +let () = define_prim4 "tac_autorewrite" bool (option (thunk unit)) (list ident) clause begin fun all by ids cl -> + Tac2tactics.autorewrite ~all by ids cl +end + +let () = define_prim1 "tac_subst" (list ident) begin fun ids -> + Equality.subst ids +end + +let () = define_prim0 "tac_substall" (return () >>= fun () -> Equality.subst_all ()) + +(** Auto *) + +let () = define_prim3 "tac_trivial" debug (list (thunk constr)) (option (list ident)) begin fun dbg lems dbs -> + Tac2tactics.trivial dbg lems dbs +end + +let () = define_prim5 "tac_eauto" debug (option int) (option int) (list (thunk constr)) (option (list ident)) begin fun dbg n p lems dbs -> + Tac2tactics.eauto dbg n p lems dbs +end + +let () = define_prim4 "tac_auto" debug (option int) (list (thunk constr)) (option (list ident)) begin fun dbg n lems dbs -> + Tac2tactics.auto dbg n lems dbs +end + +let () = define_prim4 "tac_newauto" debug (option int) (list (thunk constr)) (option (list ident)) begin fun dbg n lems dbs -> + Tac2tactics.new_auto dbg n lems dbs +end + +let () = define_prim3 "tac_typeclasses_eauto" (option strategy) (option int) (option (list ident)) begin fun str n dbs -> + Tac2tactics.typeclasses_eauto str n dbs +end diff --git a/user-contrib/Ltac2/tac2stdlib.mli b/user-contrib/Ltac2/tac2stdlib.mli new file mode 100644 index 0000000000..927b57074d --- /dev/null +++ b/user-contrib/Ltac2/tac2stdlib.mli @@ -0,0 +1,9 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +(** Standard tactics sharing their implementation with Ltac1 *) diff --git a/user-contrib/Ltac2/tac2tactics.ml b/user-contrib/Ltac2/tac2tactics.ml new file mode 100644 index 0000000000..603e00c815 --- /dev/null +++ b/user-contrib/Ltac2/tac2tactics.ml @@ -0,0 +1,447 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Pp +open Util +open Names +open Globnames +open Tac2types +open Tac2extffi +open Genredexpr +open Proofview.Notations + +let return = Proofview.tclUNIT +let thaw r f = Tac2ffi.app_fun1 f Tac2ffi.unit r () + +let tactic_infer_flags with_evar = { + Pretyping.use_typeclasses = true; + Pretyping.solve_unification_constraints = true; + Pretyping.fail_evar = not with_evar; + Pretyping.expand_evars = true; + Pretyping.program_mode = false; + Pretyping.polymorphic = false; +} + +(** FIXME: export a better interface in Tactics *) +let delayed_of_tactic tac env sigma = + let _, pv = Proofview.init sigma [] in + let name, poly = Id.of_string "ltac2_delayed", false in + let c, pv, _, _ = Proofview.apply ~name ~poly env tac pv in + (sigma, c) + +let delayed_of_thunk r tac env sigma = + delayed_of_tactic (thaw r tac) env sigma + +let mk_bindings = function +| ImplicitBindings l -> Tactypes.ImplicitBindings l +| ExplicitBindings l -> + let l = List.map CAst.make l in + Tactypes.ExplicitBindings l +| NoBindings -> Tactypes.NoBindings + +let mk_with_bindings (x, b) = (x, mk_bindings b) + +let rec mk_intro_pattern = function +| IntroForthcoming b -> CAst.make @@ Tactypes.IntroForthcoming b +| IntroNaming ipat -> CAst.make @@ Tactypes.IntroNaming (mk_intro_pattern_naming ipat) +| IntroAction ipat -> CAst.make @@ Tactypes.IntroAction (mk_intro_pattern_action ipat) + +and mk_intro_pattern_naming = function +| IntroIdentifier id -> Namegen.IntroIdentifier id +| IntroFresh id -> Namegen.IntroFresh id +| IntroAnonymous -> Namegen.IntroAnonymous + +and mk_intro_pattern_action = function +| IntroWildcard -> Tactypes.IntroWildcard +| IntroOrAndPattern ipat -> Tactypes.IntroOrAndPattern (mk_or_and_intro_pattern ipat) +| IntroInjection ipats -> Tactypes.IntroInjection (List.map mk_intro_pattern ipats) +| IntroApplyOn (c, ipat) -> + let c = CAst.make @@ delayed_of_thunk Tac2ffi.constr c in + Tactypes.IntroApplyOn (c, mk_intro_pattern ipat) +| IntroRewrite b -> Tactypes.IntroRewrite b + +and mk_or_and_intro_pattern = function +| IntroOrPattern ipatss -> + Tactypes.IntroOrPattern (List.map (fun ipat -> List.map mk_intro_pattern ipat) ipatss) +| IntroAndPattern ipats -> + Tactypes.IntroAndPattern (List.map mk_intro_pattern ipats) + +let mk_intro_patterns ipat = List.map mk_intro_pattern ipat + +let mk_occurrences f = function +| AllOccurrences -> Locus.AllOccurrences +| AllOccurrencesBut l -> Locus.AllOccurrencesBut (List.map f l) +| NoOccurrences -> Locus.NoOccurrences +| OnlyOccurrences l -> Locus.OnlyOccurrences (List.map f l) + +let mk_occurrences_expr occ = + mk_occurrences (fun i -> Locus.ArgArg i) occ + +let mk_hyp_location (id, occs, h) = + ((mk_occurrences_expr occs, id), h) + +let mk_clause cl = { + Locus.onhyps = Option.map (fun l -> List.map mk_hyp_location l) cl.onhyps; + Locus.concl_occs = mk_occurrences_expr cl.concl_occs; +} + +let intros_patterns ev ipat = + let ipat = mk_intro_patterns ipat in + Tactics.intros_patterns ev ipat + +let apply adv ev cb cl = + let map c = + let c = thaw constr_with_bindings c >>= fun p -> return (mk_with_bindings p) in + None, CAst.make (delayed_of_tactic c) + in + let cb = List.map map cb in + match cl with + | None -> Tactics.apply_with_delayed_bindings_gen adv ev cb + | Some (id, cl) -> + let cl = Option.map mk_intro_pattern cl in + Tactics.apply_delayed_in adv ev id cb cl + +let mk_destruction_arg = function +| ElimOnConstr c -> + let c = c >>= fun c -> return (mk_with_bindings c) in + Tactics.ElimOnConstr (delayed_of_tactic c) +| ElimOnIdent id -> Tactics.ElimOnIdent CAst.(make id) +| ElimOnAnonHyp n -> Tactics.ElimOnAnonHyp n + +let mk_induction_clause (arg, eqn, as_, occ) = + let eqn = Option.map (fun ipat -> CAst.make @@ mk_intro_pattern_naming ipat) eqn in + let as_ = Option.map (fun ipat -> CAst.make @@ mk_or_and_intro_pattern ipat) as_ in + let occ = Option.map mk_clause occ in + ((None, mk_destruction_arg arg), (eqn, as_), occ) + +let induction_destruct isrec ev (ic : induction_clause list) using = + let ic = List.map mk_induction_clause ic in + let using = Option.map mk_with_bindings using in + Tactics.induction_destruct isrec ev (ic, using) + +let elim ev c copt = + let c = mk_with_bindings c in + let copt = Option.map mk_with_bindings copt in + Tactics.elim ev None c copt + +let generalize pl = + let mk_occ occs = mk_occurrences (fun i -> i) occs in + let pl = List.map (fun (c, occs, na) -> (mk_occ occs, c), na) pl in + Tactics.new_generalize_gen pl + +let general_case_analysis ev c = + let c = mk_with_bindings c in + Tactics.general_case_analysis ev None c + +let constructor_tac ev n i bnd = + let bnd = mk_bindings bnd in + Tactics.constructor_tac ev n i bnd + +let left_with_bindings ev bnd = + let bnd = mk_bindings bnd in + Tactics.left_with_bindings ev bnd + +let right_with_bindings ev bnd = + let bnd = mk_bindings bnd in + Tactics.right_with_bindings ev bnd + +let split_with_bindings ev bnd = + let bnd = mk_bindings bnd in + Tactics.split_with_bindings ev [bnd] + +let specialize c pat = + let c = mk_with_bindings c in + let pat = Option.map mk_intro_pattern pat in + Tactics.specialize c pat + +let change pat c cl = + let open Tac2ffi in + Proofview.Goal.enter begin fun gl -> + let c subst env sigma = + let subst = Array.map_of_list snd (Id.Map.bindings subst) in + delayed_of_tactic (Tac2ffi.app_fun1 c (array constr) constr subst) env sigma + in + let cl = mk_clause cl in + Tactics.change pat c cl + end + +let rewrite ev rw cl by = + let map_rw (orient, repeat, c) = + let c = c >>= fun c -> return (mk_with_bindings c) in + (Option.default true orient, repeat, None, delayed_of_tactic c) + in + let rw = List.map map_rw rw in + let cl = mk_clause cl in + let by = Option.map (fun tac -> Tacticals.New.tclCOMPLETE (thaw Tac2ffi.unit tac), Equality.Naive) by in + Equality.general_multi_rewrite ev rw cl by + +let symmetry cl = + let cl = mk_clause cl in + Tactics.intros_symmetry cl + +let forward fst tac ipat c = + let ipat = Option.map mk_intro_pattern ipat in + Tactics.forward fst tac ipat c + +let assert_ = function +| AssertValue (id, c) -> + let ipat = CAst.make @@ Tactypes.IntroNaming (Namegen.IntroIdentifier id) in + Tactics.forward true None (Some ipat) c +| AssertType (ipat, c, tac) -> + let ipat = Option.map mk_intro_pattern ipat in + let tac = Option.map (fun tac -> thaw Tac2ffi.unit tac) tac in + Tactics.forward true (Some tac) ipat c + +let letin_pat_tac ev ipat na c cl = + let ipat = Option.map (fun (b, ipat) -> (b, CAst.make @@ mk_intro_pattern_naming ipat)) ipat in + let cl = mk_clause cl in + Tactics.letin_pat_tac ev ipat na c cl + +(** Ltac interface treats differently global references than other term + arguments in reduction expressions. In Ltac1, this is done at parsing time. + Instead, we parse indifferently any pattern and dispatch when the tactic is + called. *) +let map_pattern_with_occs (pat, occ) = match pat with +| Pattern.PRef (ConstRef cst) -> (mk_occurrences_expr occ, Inl (EvalConstRef cst)) +| Pattern.PRef (VarRef id) -> (mk_occurrences_expr occ, Inl (EvalVarRef id)) +| _ -> (mk_occurrences_expr occ, Inr pat) + +let get_evaluable_reference = function +| VarRef id -> Proofview.tclUNIT (EvalVarRef id) +| ConstRef cst -> Proofview.tclUNIT (EvalConstRef cst) +| r -> + Tacticals.New.tclZEROMSG (str "Cannot coerce" ++ spc () ++ + Nametab.pr_global_env Id.Set.empty r ++ spc () ++ + str "to an evaluable reference.") + +let reduce r cl = + let cl = mk_clause cl in + Tactics.reduce r cl + +let simpl flags where cl = + let where = Option.map map_pattern_with_occs where in + let cl = mk_clause cl in + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + Tactics.reduce (Simpl (flags, where)) cl + +let cbv flags cl = + let cl = mk_clause cl in + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + Tactics.reduce (Cbv flags) cl + +let cbn flags cl = + let cl = mk_clause cl in + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + Tactics.reduce (Cbn flags) cl + +let lazy_ flags cl = + let cl = mk_clause cl in + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + Tactics.reduce (Lazy flags) cl + +let unfold occs cl = + let cl = mk_clause cl in + let map (gr, occ) = + let occ = mk_occurrences_expr occ in + get_evaluable_reference gr >>= fun gr -> Proofview.tclUNIT (occ, gr) + in + Proofview.Monad.List.map map occs >>= fun occs -> + Tactics.reduce (Unfold occs) cl + +let pattern where cl = + let where = List.map (fun (c, occ) -> (mk_occurrences_expr occ, c)) where in + let cl = mk_clause cl in + Tactics.reduce (Pattern where) cl + +let vm where cl = + let where = Option.map map_pattern_with_occs where in + let cl = mk_clause cl in + Tactics.reduce (CbvVm where) cl + +let native where cl = + let where = Option.map map_pattern_with_occs where in + let cl = mk_clause cl in + Tactics.reduce (CbvNative where) cl + +let eval_fun red c = + Tac2core.pf_apply begin fun env sigma -> + let (redfun, _) = Redexpr.reduction_of_red_expr env red in + let (sigma, ans) = redfun env sigma c in + Proofview.Unsafe.tclEVARS sigma >>= fun () -> + Proofview.tclUNIT ans + end + +let eval_red c = + eval_fun (Red false) c + +let eval_hnf c = + eval_fun Hnf c + +let eval_simpl flags where c = + let where = Option.map map_pattern_with_occs where in + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + eval_fun (Simpl (flags, where)) c + +let eval_cbv flags c = + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + eval_fun (Cbv flags) c + +let eval_cbn flags c = + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + eval_fun (Cbn flags) c + +let eval_lazy flags c = + Proofview.Monad.List.map get_evaluable_reference flags.rConst >>= fun rConst -> + let flags = { flags with rConst } in + eval_fun (Lazy flags) c + +let eval_unfold occs c = + let map (gr, occ) = + let occ = mk_occurrences_expr occ in + get_evaluable_reference gr >>= fun gr -> Proofview.tclUNIT (occ, gr) + in + Proofview.Monad.List.map map occs >>= fun occs -> + eval_fun (Unfold occs) c + +let eval_fold cl c = + eval_fun (Fold cl) c + +let eval_pattern where c = + let where = List.map (fun (pat, occ) -> (mk_occurrences_expr occ, pat)) where in + eval_fun (Pattern where) c + +let eval_vm where c = + let where = Option.map map_pattern_with_occs where in + eval_fun (CbvVm where) c + +let eval_native where c = + let where = Option.map map_pattern_with_occs where in + eval_fun (CbvNative where) c + +let on_destruction_arg tac ev arg = + Proofview.Goal.enter begin fun gl -> + match arg with + | None -> tac ev None + | Some (clear, arg) -> + let arg = match arg with + | ElimOnConstr c -> + let env = Proofview.Goal.env gl in + Proofview.tclEVARMAP >>= fun sigma -> + c >>= fun (c, lbind) -> + let lbind = mk_bindings lbind in + Proofview.tclEVARMAP >>= fun sigma' -> + let flags = tactic_infer_flags ev in + let (sigma', c) = Unification.finish_evar_resolution ~flags env sigma' (sigma, c) in + Proofview.tclUNIT (Some sigma', Tactics.ElimOnConstr (c, lbind)) + | ElimOnIdent id -> Proofview.tclUNIT (None, Tactics.ElimOnIdent CAst.(make id)) + | ElimOnAnonHyp n -> Proofview.tclUNIT (None, Tactics.ElimOnAnonHyp n) + in + arg >>= fun (sigma', arg) -> + let arg = Some (clear, arg) in + match sigma' with + | None -> tac ev arg + | Some sigma' -> + Tacticals.New.tclWITHHOLES ev (tac ev arg) sigma' + end + +let discriminate ev arg = + let arg = Option.map (fun arg -> None, arg) arg in + on_destruction_arg Equality.discr_tac ev arg + +let injection ev ipat arg = + let arg = Option.map (fun arg -> None, arg) arg in + let ipat = Option.map mk_intro_patterns ipat in + let tac ev arg = Equality.injClause None ipat ev arg in + on_destruction_arg tac ev arg + +let autorewrite ~all by ids cl = + let conds = if all then Some Equality.AllMatches else None in + let ids = List.map Id.to_string ids in + let cl = mk_clause cl in + match by with + | None -> Autorewrite.auto_multi_rewrite ?conds ids cl + | Some by -> + let by = thaw Tac2ffi.unit by in + Autorewrite.auto_multi_rewrite_with ?conds by ids cl + +(** Auto *) + +let trivial debug lems dbs = + let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in + let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in + Auto.h_trivial ~debug lems dbs + +let auto debug n lems dbs = + let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in + let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in + Auto.h_auto ~debug n lems dbs + +let new_auto debug n lems dbs = + let make_depth n = snd (Eauto.make_dimension n None) in + let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in + match dbs with + | None -> Auto.new_full_auto ~debug (make_depth n) lems + | Some dbs -> + let dbs = List.map Id.to_string dbs in + Auto.new_auto ~debug (make_depth n) lems dbs + +let eauto debug n p lems dbs = + let lems = List.map (fun c -> delayed_of_thunk Tac2ffi.constr c) lems in + let dbs = Option.map (fun l -> List.map Id.to_string l) dbs in + Eauto.gen_eauto (Eauto.make_dimension n p) lems dbs + +let typeclasses_eauto strategy depth dbs = + let only_classes, dbs = match dbs with + | None -> + true, [Class_tactics.typeclasses_db] + | Some dbs -> + let dbs = List.map Id.to_string dbs in + false, dbs + in + Class_tactics.typeclasses_eauto ~only_classes ?strategy ~depth dbs + +(** Inversion *) + +let inversion knd arg pat ids = + let ids = match ids with + | None -> [] + | Some l -> l + in + begin match pat with + | None -> Proofview.tclUNIT None + | Some (IntroAction (IntroOrAndPattern p)) -> + Proofview.tclUNIT (Some (CAst.make @@ mk_or_and_intro_pattern p)) + | Some _ -> + Tacticals.New.tclZEROMSG (str "Inversion only accept disjunctive patterns") + end >>= fun pat -> + let inversion _ arg = + begin match arg with + | None -> assert false + | Some (_, Tactics.ElimOnAnonHyp n) -> + Inv.inv_clause knd pat ids (AnonHyp n) + | Some (_, Tactics.ElimOnIdent {CAst.v=id}) -> + Inv.inv_clause knd pat ids (NamedHyp id) + | Some (_, Tactics.ElimOnConstr c) -> + let open Tactypes in + let anon = CAst.make @@ IntroNaming Namegen.IntroAnonymous in + Tactics.specialize c (Some anon) >>= fun () -> + Tacticals.New.onLastHypId (fun id -> Inv.inv_clause knd pat ids (NamedHyp id)) + end + in + on_destruction_arg inversion true (Some (None, arg)) + +let contradiction c = + let c = Option.map mk_with_bindings c in + Contradiction.contradiction c diff --git a/user-contrib/Ltac2/tac2tactics.mli b/user-contrib/Ltac2/tac2tactics.mli new file mode 100644 index 0000000000..e56544cd68 --- /dev/null +++ b/user-contrib/Ltac2/tac2tactics.mli @@ -0,0 +1,122 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open Tac2expr +open EConstr +open Genredexpr +open Tac2types +open Proofview + +(** Local reimplementations of tactics variants from Coq *) + +val intros_patterns : evars_flag -> intro_pattern list -> unit tactic + +val apply : advanced_flag -> evars_flag -> + constr_with_bindings thunk list -> + (Id.t * intro_pattern option) option -> unit tactic + +val induction_destruct : rec_flag -> evars_flag -> + induction_clause list -> constr_with_bindings option -> unit tactic + +val elim : evars_flag -> constr_with_bindings -> constr_with_bindings option -> + unit tactic + +val general_case_analysis : evars_flag -> constr_with_bindings -> unit tactic + +val generalize : (constr * occurrences * Name.t) list -> unit tactic + +val constructor_tac : evars_flag -> int option -> int -> bindings -> unit tactic + +val left_with_bindings : evars_flag -> bindings -> unit tactic +val right_with_bindings : evars_flag -> bindings -> unit tactic +val split_with_bindings : evars_flag -> bindings -> unit tactic + +val specialize : constr_with_bindings -> intro_pattern option -> unit tactic + +val change : Pattern.constr_pattern option -> (constr array, constr) Tac2ffi.fun1 -> clause -> unit tactic + +val rewrite : + evars_flag -> rewriting list -> clause -> unit thunk option -> unit tactic + +val symmetry : clause -> unit tactic + +val forward : bool -> unit tactic option option -> + intro_pattern option -> constr -> unit tactic + +val assert_ : assertion -> unit tactic + +val letin_pat_tac : evars_flag -> (bool * intro_pattern_naming) option -> + Name.t -> (Evd.evar_map * constr) -> clause -> unit tactic + +val reduce : Redexpr.red_expr -> clause -> unit tactic + +val simpl : GlobRef.t glob_red_flag -> + (Pattern.constr_pattern * occurrences) option -> clause -> unit tactic + +val cbv : GlobRef.t glob_red_flag -> clause -> unit tactic + +val cbn : GlobRef.t glob_red_flag -> clause -> unit tactic + +val lazy_ : GlobRef.t glob_red_flag -> clause -> unit tactic + +val unfold : (GlobRef.t * occurrences) list -> clause -> unit tactic + +val pattern : (constr * occurrences) list -> clause -> unit tactic + +val vm : (Pattern.constr_pattern * occurrences) option -> clause -> unit tactic + +val native : (Pattern.constr_pattern * occurrences) option -> clause -> unit tactic + +val eval_red : constr -> constr tactic + +val eval_hnf : constr -> constr tactic + +val eval_simpl : GlobRef.t glob_red_flag -> + (Pattern.constr_pattern * occurrences) option -> constr -> constr tactic + +val eval_cbv : GlobRef.t glob_red_flag -> constr -> constr tactic + +val eval_cbn : GlobRef.t glob_red_flag -> constr -> constr tactic + +val eval_lazy : GlobRef.t glob_red_flag -> constr -> constr tactic + +val eval_unfold : (GlobRef.t * occurrences) list -> constr -> constr tactic + +val eval_fold : constr list -> constr -> constr tactic + +val eval_pattern : (EConstr.t * occurrences) list -> constr -> constr tactic + +val eval_vm : (Pattern.constr_pattern * occurrences) option -> constr -> constr tactic + +val eval_native : (Pattern.constr_pattern * occurrences) option -> constr -> constr tactic + +val discriminate : evars_flag -> destruction_arg option -> unit tactic + +val injection : evars_flag -> intro_pattern list option -> destruction_arg option -> unit tactic + +val autorewrite : all:bool -> unit thunk option -> Id.t list -> clause -> unit tactic + +val trivial : Hints.debug -> constr thunk list -> Id.t list option -> + unit Proofview.tactic + +val auto : Hints.debug -> int option -> constr thunk list -> + Id.t list option -> unit Proofview.tactic + +val new_auto : Hints.debug -> int option -> constr thunk list -> + Id.t list option -> unit Proofview.tactic + +val eauto : Hints.debug -> int option -> int option -> constr thunk list -> + Id.t list option -> unit Proofview.tactic + +val typeclasses_eauto : Class_tactics.search_strategy option -> int option -> + Id.t list option -> unit Proofview.tactic + +val inversion : Inv.inversion_kind -> destruction_arg -> intro_pattern option -> Id.t list option -> unit tactic + +val contradiction : constr_with_bindings option -> unit tactic diff --git a/user-contrib/Ltac2/tac2types.mli b/user-contrib/Ltac2/tac2types.mli new file mode 100644 index 0000000000..fa31153a27 --- /dev/null +++ b/user-contrib/Ltac2/tac2types.mli @@ -0,0 +1,92 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open EConstr +open Proofview + +(** Redefinition of Ltac1 data structures because of impedance mismatch *) + +type evars_flag = bool +type advanced_flag = bool + +type 'a thunk = (unit, 'a) Tac2ffi.fun1 + +type quantified_hypothesis = Tactypes.quantified_hypothesis = +| AnonHyp of int +| NamedHyp of Id.t + +type explicit_bindings = (quantified_hypothesis * EConstr.t) list + +type bindings = +| ImplicitBindings of EConstr.t list +| ExplicitBindings of explicit_bindings +| NoBindings + +type constr_with_bindings = EConstr.constr * bindings + +type core_destruction_arg = +| ElimOnConstr of constr_with_bindings tactic +| ElimOnIdent of Id.t +| ElimOnAnonHyp of int + +type destruction_arg = core_destruction_arg + +type intro_pattern = +| IntroForthcoming of bool +| IntroNaming of intro_pattern_naming +| IntroAction of intro_pattern_action +and intro_pattern_naming = +| IntroIdentifier of Id.t +| IntroFresh of Id.t +| IntroAnonymous +and intro_pattern_action = +| IntroWildcard +| IntroOrAndPattern of or_and_intro_pattern +| IntroInjection of intro_pattern list +| IntroApplyOn of EConstr.t thunk * intro_pattern +| IntroRewrite of bool +and or_and_intro_pattern = +| IntroOrPattern of intro_pattern list list +| IntroAndPattern of intro_pattern list + +type occurrences = +| AllOccurrences +| AllOccurrencesBut of int list +| NoOccurrences +| OnlyOccurrences of int list + +type hyp_location_flag = Locus.hyp_location_flag = +| InHyp | InHypTypeOnly | InHypValueOnly + +type hyp_location = Id.t * occurrences * hyp_location_flag + +type clause = + { onhyps : hyp_location list option; + concl_occs : occurrences } + +type induction_clause = + destruction_arg * + intro_pattern_naming option * + or_and_intro_pattern option * + clause option + +type multi = Equality.multi = +| Precisely of int +| UpTo of int +| RepeatStar +| RepeatPlus + +type rewriting = + bool option * + multi * + constr_with_bindings tactic + +type assertion = +| AssertType of intro_pattern option * constr * unit thunk option +| AssertValue of Id.t * constr |
