Giter Site home page Giter Site logo

lpcic / coq-elpi Goto Github PK

View Code? Open in Web Editor NEW
134.0 10.0 50.0 8.64 MB

Coq plugin embedding elpi

License: GNU Lesser General Public License v2.1

Makefile 0.07% OCaml 37.56% Coq 33.58% Prolog 27.63% Python 0.87% Nix 0.24% Perl 0.02% Shell 0.02%
coq metaprogramming extension-language lambda-prolog scripting

coq-elpi's Introduction

Docker CI Nix CI Nix CI DOC project chat Coq-Elpi logo

Coq-Elpi

Coq plugin embedding Elpi.

What is Elpi

Elpi provides an easy-to-embed implementation of a dialect of λProlog, a programming language well suited to manipulate abstract syntax trees containing binders and unification variables.

What is Coq-Elpi

Coq-Elpi provides a Coq plugin that lets one define new commands and tactics in Elpi. For that purpose it provides an embedding of Coq's terms into λProlog using the Higher-Order Abstract Syntax approach (HOAS). It also exports to Elpi a comprehensive set of Coq's primitives, so that one can print a message, access the environment of theorems and data types, define a new constant, declare implicit arguments, type classes instances, and so on. For convenience it also provides quotations and anti-quotations for Coq's syntax, so that one can write {{ nat -> lp:X }} in the middle of a λProlog program instead of the equivalent AST.

What is the purpose of all that

In the short term, provide an extension language for Coq well suited to manipulate terms containing binders. One can already use Elpi to implement commands and tactics. As ongoing research we are looking forward to express algorithms like higher order unification and type inference, and to provide an alternative elaborator for Coq.

Installation

The simplest way is to use OPAM and type

opam repo add coq-released https://coq.inria.fr/opam/released
opam install coq-elpi

Editor Setup

The recommended user interface is VSCoq. We provide an extension for vscode in the market place, just look for Coq Elpi. The extension provides syntax hilighting for both languages even when they are nested via quotations and antiquotations.

Other editors (click to expand)

At the time of writing Proof General does not handle quotations correctly, see ProofGeneral/PG#437. In particular Elpi Accumulate lp:{{ .... }}. is used in tutorials to mix Coq and Elpi code without escaping. Coq-Elpi also accepts Elpi Accumulate " .... ". but strings part of the Elpi code needs to be escaped. Finally, for non-tutorial material, one can always put the code in an external file declared with From some.load.path Extra Dependency "filename" as f. and use Elpi Accumulate File f..

CoqIDE does handle quotations. The installation process puts coq-elpi.lang in a place where CoqIDE can find it. Then you can select coq-elpi from the menu Edit -> Preferences -> Colors.

For Vim users, Coqtail provides syntax highlighting and handles quotations.

Development version (click to expand)

To install the development version one can type

opam pin add coq-elpi https://github.com/LPCIC/coq-elpi.git

One can also clone this repository and type make, but check you have all the dependencies installed first (see coq-elpi.opam).

We recommend to look at the CI setup for ocaml versions being tested. Also, we recommend to install dot-merlin-reader and ocaml-lsp-server (version 1.15).

Documentation

Tutorials

  • The Elpi programming language is an Elpi tutorial, there is nothing Coq specific in there even if the tutorial uses Coq to step trough the various examples. If you never heard of λProlog or HOAS based languages (like Twelf or Beluga) then you are strongly encouraged to read this tutorial and have a look at λProlog's home page for additional documentation. Even if you are familiar with λProlog or HOAS it may be worth reading the last sections since they focus on Elpi specific features. Last but not least it covers common pitfalls for people with a background in functional programming and the tracing mechanisms (useful for debugging).
  • HOAS of Coq terms focuses on how Coq terms are represented in Elpi, how to inspect them and call Coq APIs under a context of binders, and finally how holes ("evars" in Coq slang) are represented. It assumes the reader is familiar with Elpi.
  • Writing commands in Elpi focuses on how to write commands, in particular how to store a state across calls via so called DBs and how to handled command arguments. It assumes the reader is familiar with Elpi and the HOAS of Coq terms.
  • Writing tactics in Elpi describes how goals and tactics are represented, how to handle tactic arguments and finally how to define tactic notations. It assumes the reader is familiar with Elpi and the HOAS of Coq terms.
  • Coq-Elpi in 20 minutes video recording of a talk given at the Coq Users and Developers Workshop 2020.

Small examples (proofs of concept)

  • reification is the typical use case for meta programs: reading the syntax of terms into an inductive representing a sub language on which some decision procedure can be implemented
  • data bases shows how Elpi programs can store data and reuse it across multiple runs
  • record expansion sketches a program to unpack records in a definition: it replaces an abstraction over a records with abstractions over all of its components
  • record to sigma sketches a program that de-sugars a record type to iterated sigma types
  • fuzzer sketches a program to alter an inductive type while preserving its well typedness. It makes nothing useful per se, but shows how to map a term and call the type checker deep inside it.
  • tactics show how to create simple tactics by using (proof) terms and the elaborator of Coq
  • generalize show how to abstract subterms out (one way to skin the cat, there are many)
  • abs_evars show how to close a term containing holes (evars) with binders
  • record import gives short names to record projections applied to the given record instance.
  • reduction surgery implements a tactic fine tuning cbv with a list of allowed unfoldings taken from a module.

Applications written in Coq-Elpi

  • Derive shows how to obtain proved equality tests and a few extra gadgets out of inductive type declarations. See the README for the list of derivations. It comes bundled with Coq-Elpi.
  • Locker lets one hide the computational contents of definitions via modules or opaque locks. It comes bundled with Coq-Elpi.
  • Hierarchy Builder is a Coq extension to declare hierarchies of algebraic structures.
  • Algebra Tactics is a port of the ring and field tactics to the Mathematical Components library.
  • Trakt is a generic goal preprocessing tool for proof automation tactics in Coq.
  • Namespace Emulation System implements most of the features of namespaces (on top of Coq's modules).
  • Dx uses elpi to generate an intermediate representation of Coq terms, to be later tranformed into C.
  • Coercion enable to program coercions in Elpi. It comes bundled with Coq-Elpi.

Quick Reference

In order to load Coq-Elpi use From elpi Require Import elpi.

Vernacular commands

(click to expand)
  • Elpi Command <qname> creates command named <qname> containing the preamble elpi-command.
  • Elpi Tactic <qname> creates a tactic <qname> containing the preamble elpi-tactic.
  • Elpi Db <dbname> <code> creates a Db (a program that is accumulated into other programs). <code> is the initial contents of the Db, including the type declaration of its constituting predicates. It understands the #[phase] attribute, see synterp-vs-interp.
  • Elpi Program <qname> <code> lower level primitive letting one crate a command/tactic with a custom preamble <code>.
  • From some.load.path Extra Dependency <filename> as <fname>.
  • Elpi Accumulate [<dbname>|<qname>] [<code>|File <fname>|Db <dbname>] adds code to the current program (or <dbname> or <qname> if specified). The code can be verbatim, from a file or a Db. File names <fname> must have been previously declared with the above command. It understands the #[skip="rex"] and #[only="rex"] which make the command a no op if the Coq version is matched (or not) by the given regular expression. It understands the #[phase] attribute, see synterp-vs-interp
  • Elpi Typecheck [<qname>] typechecks the current program (or <qname> if specified). It understands the #[phase] attribute, see synterp-vs-interp
  • Elpi Debug <string> sets the variable <string>, relevant for conditional clause compilation (the :if VARIABLE clause attribute). It understands the #[phase] attribute, see synterp-vs-interp
  • Elpi Trace [[<start> <stop>] <predicate-filter>*|Off] enable/disable tracing, eventually limiting it to a specific range of execution steps or predicate names. It understands the #[phase] attribute, see synterp-vs-interp
  • Elpi Trace Browser enable/disable tracing for Elpi's trace browser.
  • Elpi Bound Steps <number> limits the number of steps an Elpi program can make.
  • Elpi Print <qname> [<string> <filter>*] prints the program <qname> to an HTML file named <qname>.html and a text file called <qname>.txt (or <string> if provided) filtering out clauses whose file or clause-name matches <filter>. It understands the #[phase] attribute, see synterp-vs-interp

where:

  • <qname> is a qualified Coq name, e.g. derive.eq or my_program.
  • <dbname> is like <qname> but lives in a different namespace. By convention <dbname> ends in .db, e.g. derive.eq.db.
  • <code> is verbatim Elpi code, either lp:{{ ... }} or " ... " (in the latter case, strings delimiters need to be escaped following Coq rules, e.g. lp:{{ coq.say "hello!" }} becomes " coq.say ""hello!"" ").
  • <filename> is a string containing the path of an external file, e.g. "this_file.elpi".
  • <start> and <stop> are numbers, e.g. 17 24.
  • <predicate-filter> is a regexp against which the predicate name is matched, e.g. "derive.*".

Separation of parsing from execution of vernacular commands

(click to expand)

Since version 8.18 Coq has separate parsing and execution phases, respectively called synterp and interp.

Since Coq has an extensible grammar the parsing phase is not entirely performed by the parser: after parsing one sentence Coq evaluates its synterp action. The synterp actions of a command like Import A. are the subset of its effect which affect parsing, like enabling a notation. Later, during the execution phase Coq evaluates the its interp action, which includes effects like putting lemma names in scope or enables type class instances etc.

Being able to parse an entire document quickly, without actually executing any sentence, is important for developing reactive user interfaces, but requires some extra work when defining new commands, in particular to separate their synterp actions from their interp ones. Each command defined with Coq-Elpi is split into two programs, one running during the parsing phase and the other one during the execution phase.

Declaration of synterp actions

Each Elpi Command internally declares two programs with the same name. One to be run while the Coq document is parsed, the synterp-command, and the other one while it is executed, the interp command. Elpi Accumulate, by default, adds code to the interp-command. The #[phase] attribute can be used to accumulate code to the synterp-command or to both commands. Elpi Typecheck checks both commands.

Each Elpi Db internally declares one db, by default for the interp phase. The #[phase] attribute can be used crate a database for the synterp phase, or for both phases. Note that databases for the two phases are distinct, no data is shared among them. In particular the coq.elpi.accumulate* API exists in both phases and only acts on data bases for the current phase.

The alignment of phases

All synterp actions, i.e. calls to APIs dealing with modules and sections like begin/end-module or import/export, have to happen at both synterp and interp time and in the same order.

In order to do so, the synterp-command may need to communicate data to the corresponding interp-command. There are two ways for doing so.

The first one is to use, as the main entry points, the following ones:

pred main-synterp i:list argument, o:any.
pred main-interp i:list argument, i:any.

Unlike main the former outputs a datum while the latter receives it in input. During the synterp phase the API coq.synterp-actions lists the actions performed so far. An excerpt from the coq-builtin-synterp file:

% Action executed during the parsing phase (aka synterp)
kind synterp-action type.
type begin-module id -> synterp-action.
type end-module modpath -> synterp-action.

The synterp-command can output data of that type, but also any other data it wishes.

The second way to communicate data is implicit, but limited to synterp actions. Such synterp actions can be recorded into (nested) groups whose structure is declared using well-bracketed calls to predicates coq.begin-synterp-group and coq.end-synterp-group in the synterp phase. In the interp phase, one can then use predicate coq.replay-synterp-action-group to replay all the synterp actions of the group with the given name at once.

In the case where one wishes to interleave code between the actions of a given group, it is also possible to match the synterp group structure at interp, via coq.begin-synterp-group and coq.end-synterp-group. Individual actions that are contained in the group then need to be replayed individually.

One can use coq.replay-next-synterp-actions to replay all synterp actions until the next beginning/end of a synterp group. However, this is discouraged in favour of using groups explicitly, as this is more modular. Code that used to rely on the now-removed coq.replay-all-missing-synterp-actions predicate can rely on coq.replay-next-synterp-actions instead, but this is discouraged in favour of using groups explicitly)

Syntax of the #[phase] attribute
  • #[phase="ph"] where "ph" can be "parsing", "execution" or "both"
  • #[synterp] is a shorthand for #[phase="parsing"]
  • #[interp] is a shorthand for #[phase="execution]

Invocation of Elpi code

(click to expand)
  • Elpi <qname> <argument>*. invokes the main predicate of the <qname> program passing a possible empty list of arguments. This is how you invoke a command.

  • elpi <qname> <argument>*. invokes the solve predicate of the <qname> program passing a possible empty list of arguments and the current goal. This is how you invoke a tactic.

  • Elpi Export <qname> [As <other-qname>] makes it possible to invoke command <qname> (or <other-qname> if given) without the Elpi prefix or invoke tactic <qname> in the middle of a term just writing <qname> args instead of ltac:(elpi <qname> args). Note that in the case of tactics, all arguments are considered to be terms. Moreover, remember that one can use Tactic Notation to give the tactic a better syntax and a shorter name when used in the middle of a proof script.

where <argument> can be:

  • a number, e.g. 3, represented in Elpi as (int 3)
  • a string, e.g. "foo" or bar.baz, represented in Elpi as (str "foo") and (str "bar.baz"). Coq keywords and symbols are recognized as strings, eg => requires no quotes. Quotes are necessary if the string contains a space or a character that is not accepted for qualified identifiers or if the string is Definition, Axiom, Record, Structure, Inductive, CoInductive, Variant or Context.
  • a term, e.g. (3) or (f x), represented in Elpi as (trm ...). Note that terms always require parentheses, that is 3 is a number while (3) is a Coq term and depending on the context could be a natural number (i.e. S (S (S O))) or a Z or ... See also the section Terms as arguments down below, and the syntax for Ltac variables down below.

Commands also accept the following arguments (the syntax is as close as possible to the Coq one: [...] means optional, * means 0 or more). See the argument data type in coq-builtin.elpi for their HOAS encoding. See also the section Terms as arguments down below.

  • Definition name binder* [: term] := term
  • Axiom name : term
  • [ Record | Structure ] name binder* [: sort] := [name] { name : term ; * }
  • [ Inductive | CoInductive | Variant ] name binder* [| binder] [: term] := | name binder : term *
  • Context binder*
Ltac Variables

Tactics also accept Ltac variables as follows:

  • ltac_string:(v) (for v of type string or ident)
  • ltac_int:(v) (for v of type int or integer)
  • ltac_term:(v) (for v of type constr or open_constr or uconstr or hyp)
  • ltac_(string|int|term)_list:(v) (for v of type list of ...)
  • ltac_tactic:(t) (for t of type tactic_expr)
  • ltac_attributes:(v) (for v of type attributes) For example:
Tactic Notation "tac" string(X) ident(Y) int(Z) hyp(T) constr_list(L) simple_intropattern_list(P) :=
  elpi tac ltac_string:(X) ltac_string:(Y) ltac_int:(Z) ltac_term:(T) ltac_term_list:(L) ltac_tactic:(intros P).

lets one write tac "a" b 3 H t1 t2 t3 [|m] in any Ltac context. Arguments are first interpreted by Ltac according to the types declared in the tactic notation and then injected in the corresponding Elpi argument. For example H must be an existing hypothesis, since it is typed with the hyp Ltac type, but in Elpi it will appear as a term, eg trm c0. Similarly t1, t2 and t3 are checked to be well typed and to contain no unresolved implicit arguments, since this is what the constr Ltac type means If they were typed as open_constr or uconstr, the last or both checks would be respectively skipped. In any case they are passed to the Elpi code as trm .... Both "a" and b are passed to Elpi as str .... Finally, ltac_term:(T) and (T) are not synonyms: but the former must be used when defining tactic notations, the latter when invoking elpi tactics directly.

Attributes

Attributes are supported in both commands and tactics. Examples:

  • #[ att ] Elpi cmd
  • #[ att ] cmd for a command cmd exported via Elpi Export cmd
  • #[ att ] elpi tac
  • Tactic Notation ... attributes(A) ... := ltac_attributes:(A) elpi tac. Due to a parsing conflict in Coq grammar, at the time of writing this code:
    Tactic Notation "#[" attributes(A) "]" "tac" :=
      ltac_attributes:(A) elpi tac.
    has the following limitation:
    • #[ att ] tac. does not parse
    • (#[ att ] tac). works
    • idtac; #[ att ] tac. works
Terms as arguments

Since version 1.15, terms passed to Elpi commands code via (term) or via a declaration (like Record, Inductive ...) are in elaborated format by default. This means that all Coq notational facilities are available, like deep pattern matching, or tactics in terms. One can use the attribute #[arguments(raw)] to declare a command which instead takes arguments in raw format. In that case, notations are unfolded, implicit arguments are expanded (holes _ are added) and lexical analysis is performed (global names and bound names are identified, holes are applied to bound names in scope), but deep pattern matching or tactics in terms are not supported, and in particular type checking/inference is not performed. Once can use the coq.typecheck or coq.elaborate-skeleton APIs to fill in implicit arguments and insert coercions on raw terms.

Terms passed to Elpi tactics via tactic notations can be forced to be elaborated beforehand by declaring the parameters to be of type constr or open_constr. Arguments of type uconstr are passed raw.

Testing/debugging:
  • Elpi Query [<qname>] <code> runs <code> in the current program (or in <qname> if specified).
  • Elpi Query [<qname>] <synterp-code> <interp-code> runs <synterp-code> in the current (synterp) program (or in <qname> if specified) and <interp-code> in the current program (or <qname>).
  • elpi query [<qname>] <string> <argument>* runs the <string> predicate (that must have the same signature of the default predicate solve).

Supported features of Gallina (core calculus of Coq)

(click to expand)
  • functional core (fun, forall, match, application, let-in, sorts)
  • evars (unification variables)
  • single Inductive and CoInductive types (including parameters, non-uniform parameters, indexes)
  • mutual Inductive and CoInductive types
  • fixpoints
  • mutual fixpoints
  • cofixpoints
  • primitive records
  • primitive projections
  • primitive integers
  • primitive floats
  • primitive arrays
  • universe polymorphism
  • modules
  • module types
  • functor application
  • functor definition

Supported features of Gallina's extensions (extra logical features, APIs)

(click to expand)

Checked boxes are available, unchecked boxes are planned, missing items are not planned. This is a high level list, for the details see coq-builtin.

  • i/o: messages, warnings, errors, Coq version
  • logical environment: read, write, locate
    • dependencies between objects
  • type classes database: read, write
    • take over resolution
  • canonical structures database: read, write
    • take over resolution
  • coercions database: read, write
  • sections: open, close
  • scope management: import, export
  • hints: mode, opaque, resolve, strategy
  • arguments: implicit, name, scope, simpl
  • abbreviations: read, write, locate
  • typing and elaboration
  • unification
  • reduction: lazy, cbv, vm, native
    • flags for lazy and cbv
  • ltac1: bridge to call ltac1 code, mono and multi-goal tactics
  • option system: get, set, add
  • pretty printer: boxes, printing width
  • attributes: read

Relevant files

Organization of the repository

The code of the Coq plugin is at the root of the repository in the src, elpi and theories directories.

The apps directory contains client applications written in Coq-Elpi.

coq-elpi's People

Contributors

artagnon avatar bgregoir avatar blaisorblade avatar cohencyril avatar dwarfmaster avatar ecrancemerce avatar ejgallego avatar eponier avatar fajb avatar fissored avatar gares avatar herbelin avatar jashug avatar jfehrle avatar maximedenes avatar palmskog avatar pedrotst avatar phikal avatar pi8027 avatar ppedrot avatar proux01 avatar rlepigre avatar simonboulier avatar skyskimmer avatar tragicus avatar trilis avatar vbgl avatar whonore avatar ybertot avatar yoichi-at-bedrock avatar

Stargazers

 avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar

Watchers

 avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar

coq-elpi's Issues

coq.CS.db fails in the presence of default CS

From elpi Require Import elpi.

Record test := Test {sort :> Type}.
Canonical can_test T := Test T.

Elpi Command test.
Elpi Query lp:{{
   coq.CS.db DB
}}.

gives

Error: Pattern matching failure embedding: cs-pattern: Default_cs

coq.notation.abbreviation

This failed on HB generating an illtyped term:

(pi t\ coq.notation.abbreviation A [t] (X t),
print X % X = c\ a term

separate compilation

As of today every time you invoke an elpi command it is recompiled on the stop.
This is a waste of time, but more importantly makes {{ quotation }} see the state of Coq (eg open scopes).

Fixing this requires a proper separate compilation on the side of elpi (compile once in the Coq state the user expects), which has been on its way for some time, but is not there yet.

[Coq 8.9+] Please fix deprecation warnings

Please fix deprecation warnings for 8.9 so Elpi will be able to compile with 8.10, thanks!

File "src/coq_elpi_builtins.ml", line 571, characters 19-38:
Warning 3: deprecated: Check
Use the constructor in module [Declaremods]
File "src/coq_elpi_builtins.ml", line 575, characters 23-63:
Warning 3: deprecated: Enforce
Use the constructor in module [Declaremods]
File "src/coq_elpi_builtins.ml", line 575, characters 49-62:
Warning 3: deprecated: DefaultInline
Use the constructor in module [Declaremods]
File "src/coq_elpi_builtins.ml", line 625, characters 30-54:
Warning 3: deprecated: DefaultInline
Use the constructor in module [Declaremods]
File "src/coq_elpi_builtins.ml", line 638, characters 30-54:
Warning 3: deprecated: DefaultInline
Use the constructor in module [Declaremods]
CAMLC -c src/coq_elpi_vernacular.ml
File "src/coq_elpi_vernacular.ml", line 21, characters 0-50:
Warning 39: unused rec flag.
File "src/coq_elpi_builtins.ml", line 571, characters 19-38:
Warning 3: deprecated: Check
Use the constructor in module [Declaremods]
File "src/coq_elpi_builtins.ml", line 575, characters 23-63:
Warning 3: deprecated: Enforce
Use the constructor in module [Declaremods]
File "src/coq_elpi_builtins.ml", line 575, characters 49-62:
Warning 3: deprecated: DefaultInline
Use the constructor in module [Declaremods]
File "src/coq_elpi_builtins.ml", line 625, characters 30-54:
Warning 3: deprecated: DefaultInline
Use the constructor in module [Declaremods]
File "src/coq_elpi_builtins.ml", line 638, characters 30-54:
Warning 3: deprecated: DefaultInline
Use the constructor in module [Declaremods]
CAMLOPT -c -for-pack Elpi_plugin src/coq_elpi_vernacular.ml
File "src/coq_elpi_vernacular.ml", line 21, characters 0-50:
Warning 39: unused rec flag.

`coq.gr->string` duplicates the identifier

Elpi Command debug.
Elpi Query debug lp:{{
  coq.locate "bool" GR,
  coq.gr->string GR ID.
}}.

outputs

Query assignments:
  GR = indt «bool»
  ID = Coq.Init.Datatypes.bool.bool

Broken in 4.06.0 [-safe-string]

coq-elpi fails in newer OCaml version with

OCAMLOPT elpi_custom.cmx            -c
File "elpi_custom.ml", line 187, characters 32-33:
Error: This expression has type string but an expression was expected of type
         bytes

infixr leaks from databases

An infixr declaration contained inside a Database seems to be applied to all subsequent Elpi tactics, regardless of whether that database was actually accumulated. This potentially makes typechecks successful even though the type was not correctly specified.

error compiling coq-elpi master (dbf82c0447c9d15e679e5de5c79892a3b2ad24df)

I have the following error message:

Using coq found in /nix/store/5hgy5hknybi927nillzb646ygwhw929i-coq-8.10+beta1/bin, from COQBIN or PATH
COQDEP VFILES
COQPP src/coq_elpi_vernacular_syntax.mlg
CAMLDEP src/coq_elpi_builtins.mli
CAMLDEP src/coq_elpi_goal_HOAS.mli
CAMLDEP src/coq_elpi_glob_quotation.mli
CAMLDEP src/coq_elpi_HOAS.mli
CAMLDEP src/coq_elpi_utils.mli
CAMLDEP src/coq_elpi_vernacular.mli
COQDEP src/elpi_plugin.mlpack
CAMLDEP src/coq_elpi_config.ml
CAMLDEP src/coq_elpi_builtins.ml
CAMLDEP src/coq_elpi_goal_HOAS.ml
CAMLDEP src/coq_elpi_glob_quotation.ml
CAMLDEP src/coq_elpi_name_quotation.ml
CAMLDEP src/coq_elpi_HOAS.ml
CAMLDEP src/coq_elpi_utils.ml
CAMLDEP src/coq_elpi_vernacular.ml
CAMLDEP src/coq_elpi_vernacular_syntax.ml
CAMLOPT -c -for-pack Elpi_plugin src/coq_elpi_config.ml
CAMLC -c src/coq_elpi_utils.mli
CAMLOPT -c -for-pack Elpi_plugin src/coq_elpi_utils.ml
CAMLC -c src/coq_elpi_HOAS.mli
CAMLOPT -c -for-pack Elpi_plugin src/coq_elpi_HOAS.ml
File "src/coq_elpi_HOAS.ml", line 222, characters 2-58:
Error (warning 32): unused value pp_elpi_evar.
File "src/coq_elpi_HOAS.ml", line 227, characters 4-17:
Error (warning 32): unused value empty.
File "src/coq_elpi_HOAS.ml", line 234, characters 4-42:
Error (warning 32): unused value pp.
File "src/coq_elpi_HOAS.ml", line 251, characters 2-60:
Error (warning 32): unused value pp_coq_engine.
File "src/coq_elpi_HOAS.ml", line 1670, characters 4-16:
Error (warning 32): unused value cc_constr2lp.
make[3]: *** [Makefile.coq:607: src/coq_elpi_HOAS.cmx] Error 2
make[2]: *** [Makefile.coq:321: all] Error 2
make[1]: *** [Makefile.coq:746: opt] Error 2
make: *** [Makefile:26: all] Error 2

using

$ coqtop -v
The Coq Proof Assistant, version 8.10+beta1 (May 2019)
compiled on May 20 2019 10:24:51 with OCaml 4.06.1

and I do not know how to print elpi version...

missing API

coq.module.import M
coq.module.export M.
coq.section.begin S.
coq.section.end.

Binding symbol arguments

Can we add parsing for arbitrary symbols so that
Elpi foo x := a.
has arguments [str ":=", str "a"] ?

ltac.elpi

All combinators should use a solved? predicate to skip goals solved by side effect.

Binding `(x : A)`

Everyone can dream,... would it be possible one day to have

Elpi Command def.
Elpi Accumulate def lp:{{
main (arg-cons F (arg-decl X A Decl) :-
  pi x\ Decl x = arg-cons (str ":=") (arg-cons (trm (T x)) arg-nil)),
  Body = (fun X A x \ T),
  coq.env.typecheck Body _Ty,
  coq.env.add-const F Body _Ty ff ff _.
}}.
Elpi def f (x : bool) := x.
Check f.
(* f
	 : bool -> bool
*)

with

kind arguments type.
type arg-nil arguments.
type arg-cons argument -> arguments -> arguments.
type arg-binder name -> term -> (term -> arguments) -> arguments.

?

coq-elaborator and no-ops

something is fishy when X = {{ let x : T := _ in x }} is run, in particular it seems no new evar constraint is generated for the body. It might be the code that detects progress in combination with some eager iota/beta reduction (also {{ (fun x : T => x) _ }} is problematic, while {{ f _ }} is OK).

antiquotation in binder scope

Idea

{{ forall x (pp:X) : lp:T , lp:(B x) }}

to use X for the pretty printing hint of x.
Today you are forced to use prod X T x\B x to chose the pretty printing hint X

Document how to compile against a local Coq version

Context: I want to prepare a patch for coq-elpi in view of a change on Coq's side.

I have a local Coq tree and I want to compile coq-elpi against it. The corresponding coqc is in the path, but when I run make, coq-elpi starts cloning Coq without asking me anything.

About installation, the README mentions only:

``
[blabla OPAM blabla]

You can also clone this repository and type make (in this case the plugin is compiled against the Coq version in the coq/ submodule directory).
``

Could you document how to build coq-elpi against a local tree? I had a brief look at the Makefile, but got lost in the hacks. It seems you rely on the location of the file produced by Coq's configure, for example. What is the point of all this? Superficially, it looks a bit like you are reimplementing some of the management of coq-elpi's dependencies in the Makefile...

examples

In addition to the tutorial it would be nice to provide examples such as:

  • reification (example already in theories/tutorials/ by YB)
  • DB population/retrieval
  • ad-hoc proof search, eg the example in Beta's paper about CS
  • clean up and document a couple of bricks in derive, eg "eq" and "param1"
  • tactic to transfer the goal, eg bool-Prop, or maybe Zify
  • quickchick generation of counter example

clarify nuparams in `Inductive` argument

In Coq there is no syntax to distinguish between Uniform and Non-Uniform parameters to Inductives. Coq infers it by their usage in the types of the constructor. This is not orthogonal to implicits, since an implicit status may make a parameter automatically inferred (as Uniform?).

This is too complex for Elpi's "inductive" argument, IMO, and also does not reflect how inductives are represented in Elpi: Uniform parameters do not occur as arguments to the inductivetype in the type of the constructors (they are fixed, why should you repeat them?), while Non-Uniform (and indexes) must appear (they vary, you must tell Elpi how). Coq-Elpi 1.3 only supports Record arguments, but I'd like to support also Inductive in 1.4 (to be released in few weeks).

I'm leaning toward imposing this very discipline for inductives used as arguments to Elpi commands.
Eg:

Elpi mycommand
  Inductive ty UP NuP : idx -> Type :=
  | K1 : ty (NuP * NuP) i -> ty NuP j
  | K2.

This also makes type checking of constructor with no given type for the conclusion (K2), it is
sufficient to unify they missing type with ty NuP _ (in this case the index is not inferrable, while the exact argument NuP is exactly what makes NuP a parameter and not an index, but that is the idea, I hope it is understandable).

This is not ideal, since one cannot just prefix any Coq Inductive declaration with, say, Elpi derive since he may had implicits on, or use @ and pass all arguments. But at the same time this design is very simple. The alternative is to have Elpi call the implicit argument machinery during the interning phase (note that Coq does intern + pretype inductive in a mixed way, so Elpi is already running its own code for inductives, since it has to stop after interning.. but so far this code is small, calling impargs does not not really thrill me).

Before committing to this design/compromise I'd love to hear from @maximedenes / @SkySkimmer since IIRC they did clean up a little inductive + implicits recently.

collect-goals

It is too simplistic

  • may return duplicates
  • does not shelve

Clean up ctypes for names

We currently use these 3 data types in the API

macro @gref :- ctype "Globnames.global_reference". % name for a global constant
macro @name :- ctype "Name.t".                     % name hint (compare as =)
string

I think we could add

macro @id :- ctype "Id.t"

and have all API adding things to the environment to require an @id, have conversion functions from the others to @id.

This new type could also be used in the module-related API, since @name is broken wrt comparison.

remove hole

As in hol light, use UnifVar in place of hole.
Beware coq.elab won't instantiate them, maybe...

wish: syntax

Elpi foobar [ 33 | "ident" | ident | (term) | Record <record-syntax> ] .

Recommend Projects

  • React photo React

    A declarative, efficient, and flexible JavaScript library for building user interfaces.

  • Vue.js photo Vue.js

    🖖 Vue.js is a progressive, incrementally-adoptable JavaScript framework for building UI on the web.

  • Typescript photo Typescript

    TypeScript is a superset of JavaScript that compiles to clean JavaScript output.

  • TensorFlow photo TensorFlow

    An Open Source Machine Learning Framework for Everyone

  • Django photo Django

    The Web framework for perfectionists with deadlines.

  • D3 photo D3

    Bring data to life with SVG, Canvas and HTML. 📊📈🎉

Recommend Topics

  • javascript

    JavaScript (JS) is a lightweight interpreted programming language with first-class functions.

  • web

    Some thing interesting about web. New door for the world.

  • server

    A server is a program made to process requests and deliver data to clients.

  • Machine learning

    Machine learning is a way of modeling and interpreting data that allows a piece of software to respond intelligently.

  • Game

    Some thing interesting about game, make everyone happy.

Recommend Org

  • Facebook photo Facebook

    We are working to build community through open source technology. NB: members must have two-factor auth.

  • Microsoft photo Microsoft

    Open source projects and samples from Microsoft.

  • Google photo Google

    Google ❤️ Open Source for everyone.

  • D3 photo D3

    Data-Driven Documents codes.