GithubHelp home page GithubHelp logo

metacoq / metacoq Goto Github PK

View Code? Open in Web Editor NEW
361.0 17.0 79.0 31.35 MB

Metaprogramming, verified meta-theory and implementation of Coq in Coq

Home Page: https://metacoq.github.io

License: MIT License

Makefile 0.19% OCaml 2.23% Coq 97.15% Shell 0.16% CSS 0.09% JavaScript 0.07% HTML 0.01% Nix 0.09%
coq metaprogramming coq-formalization

metacoq's Introduction

MetaCoq

MetaCoq

Build status MetaCoq Chat Open in Visual Studio Code

MetaCoq is a project formalizing Coq in Coq and providing tools for manipulating Coq terms and developing certified plugins (i.e. translations, compilers or tactics) in Coq.

Quick jump

Getting started

Installation instructions

See INSTALL.md

Documentation

See DOC.md

Overview of the project

At the center of this project is the Template-Coq quoting library for Coq. The project currently has a single repository extending Template-Coq with additional features. Each extension is in a dedicated folder. The dependency graph might be useful to navigate the project. Statistics: ~300kLoC of Coq, ~30kLoC of OCaml.

Template-Coq is a quoting library for Coq. It takes Coq terms and constructs a representation of their syntax tree as an inductive data type. The representation is based on the kernel's term representation.

After importing MetaCoq.Template.Loader there are commands MetaCoq Test Quote t., MetaCoq Quote Definition name := (t). and MetaCoq Quote Recursively Definition name := (t). as well as a tactic quote_term t k, where in all cases t is a term and k a continuation tactic.

In addition to this representation of terms, Template Coq includes:

  • Reification of the environment structures, for constant and inductive declarations along with their universe structures.

  • Denotation of terms and global declarations.

  • A monad for querying the environment, manipulating global declarations, calling the type checker, and inserting them in the global environment, in the style of MTac. Monadic programs p : TemplateMonad A can be run using MetaCoq Run p.

  • A formalization of the typing rules reflecting the ones of Coq, covering all of Coq except eta-expansion and template polymorphism.

PCUIC, the Polymorphic Cumulative Calculus of Inductive Constructions is a cleaned up version of the term language of Coq and its associated type system, shown equivalent to the one of Coq. This version of the calculus has proofs of standard metatheoretical results:

  • Weakening for global declarations, weakening and substitution for local contexts.

  • Confluence of reduction using a notion of parallel reduction

  • Context cumulativity / conversion and validity of typing.

  • Subject Reduction (case/cofix reduction excluded)

  • Principality: every typeable term has a smallest type.

  • Bidirectional presentation: an equivalent presentation of the system that enforces directionality of the typing rules. Strengthening follows from this presentation.

  • Elimination restrictions: the elimination restrictions ensure that singleton elimination (from Prop to Type) is only allowed on singleton inductives in Prop.

  • Canonicity: The weak head normal form of a term of inductive type is a constructor application.

  • Consistency under the assumption of strong normalization

  • Weak call-by-value standardization: Normal forms of terms of first-order inductive type can be found via weak call-by-value evaluation.

See the PCUIC README for a detailed view of the development.

Implementation of a fuel-free and verified reduction machine, conversion checker and type checker for PCUIC. This relies on a postulate of strong normalization of the reduction relation of PCUIC on well-typed terms. The checker is shown to be correct and complete w.r.t. the PCUIC specification. The extracted safe checker is available in Coq through a new vernacular command:

MetaCoq SafeCheck <term>

After importing MetaCoq.SafeChecker.Loader.

To roughly compare the time used to check a definition with Coq's vanilla type-checker, one can use:

MetaCoq CoqCheck <term>

This also includes a verified, efficient re-typing procedure (useful in tactics) in MetaCoq.SafeChecker.PCUICSafeRetyping.

See the SafeChecker README for a detailed view of the development.

An erasure procedure to untyped lambda-calculus accomplishing the same as the type and proof erasure phase of the Extraction plugin of Coq. The extracted safe erasure is available in Coq through a new vernacular command:

MetaCoq Erase <term>

After importing MetaCoq.Erasure.Loader.

The erasure pipeline includes verified optimizations to remove lets in constructors, remove cases on propositional terms, switch to an unguarded fixpoint reduction rule and transform the higher-order constructor applications to first-order blocks for easier translation to usual programming languages. See the erasure README for a detailed view of the development.

Examples of translations built on top of this:

The Quotation module is geared at providing functions □T → □□T for □T := Ast.term (currently implemented) and for □T := { t : Ast.term & Σ ;;; [] |- t : T } (still in the works).

Ultimately the goal of this development is to prove that is a lax monoidal semicomonad (a functor with cojoin : □T → □□T that codistributes over unit and ×), which is sufficient for proving Löb's theorem.

The public-facing interface of this development is provided in MetaCoq.Quotation.ToTemplate.All and MetaCoq.Quotation.ToPCUIC.All.

See the Quotation README for a more detailed view of the development.

Examples

Papers

Related Projects

  • The CertiCoq project develops a certified compiler from the output of verified erasure down to CompCert C-light. It provides in particular OCaml and fully foundationally verified plugins for the whole compilation pipeline from Gallina to Clight and the verified type-checker of MetaCoq.

  • The ConCert project develops certified or certifying compilers from Gallina to smart contract languages (Liquidity and CameLIGO), the functional language Elm, and a subset of the Rust programming languages. It uses the typed erasure variant to gather more type information about erased terms and perform optimizations based on this information. The project focuses in particular on the verification and safe extraction of smart contracts for blockchains.

Team & Credits

Abhishek Anand Danil Annenkov Simon Boulier
Cyril Cohen Yannick Forster Jason Gross
Meven Lennon-Bertrand Kenji Maillard Gregory Malecha
Jakob Botsch Nielsen Matthieu Sozeau Nicolas Tabareau
Théo Winterhalter

MetaCoq is developed by (left to right) Abhishek Anand, Danil Annenkov, Simon Boulier, Cyril Cohen, Yannick Forster, Jason Gross, Meven Lennon-Bertrand, Kenji Maillard, Gregory Malecha, Jakob Botsch Nielsen, Matthieu Sozeau, Nicolas Tabareau and Théo Winterhalter.

Copyright (c) 2014-2023 Gregory Malecha
Copyright (c) 2015-2023 Abhishek Anand, Matthieu Sozeau
Copyright (c) 2017-2023 Simon Boulier, Nicolas Tabareau, Cyril Cohen
Copyright (c) 2018-2023 Danil Annenkov, Yannick Forster, Théo Winterhalter
Copyright (c) 2020-2023 Jakob Botsch Nielsen, Meven Lennon-Bertrand
Copyright (c) 2022-2023 Kenji Maillard
Copyright (c) 2023      Jason Gross

This software is distributed under the terms of the MIT license. See LICENSE for details.

Bugs

Please report any bugs or feature requests on the github issue tracker.

metacoq's People

Contributors

aa755 avatar annenkov avatar dependabot[bot] avatar ejgallego avatar fakusb avatar flupe avatar gares avatar gasche avatar gmalecha avatar herbelin avatar jakobbotsch avatar jasongross avatar kyodralliam avatar liyishuai avatar lysxia avatar mattam82 avatar maximedenes avatar mevenbertrand avatar namin avatar neuralcoder3 avatar ppedrot avatar proux01 avatar simonboulier avatar skyskimmer avatar tabareau avatar theowinterhalter avatar vbgl avatar vzaliva avatar yannl35133 avatar yforster 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  avatar  avatar  avatar  avatar  avatar  avatar  avatar

metacoq's Issues

Unquoting doesn't work inside functors

This seems to involve names inside of functors.

Require Import Template.Loader.
Require Import String.
Require Import Template.Ast.

Definition x : nat := 3.
Quote Definition _t := (x = x).
Run TemplateProgram (tmBind (tmUnquote _t) (fun x => tmPrint x)). (* works! *)

Module Type M.
  Parameter x : nat.
End M.

Module V (x : M).
  Quote Definition _t := (x.x = x.x).

  Run TemplateProgram (tmBind (tmUnquote _t) (fun x => tmPrint x)).
  (* Toplevel input, characters 0-65: *)
  (* > Run TemplateProgram (tmBind (tmUnquote _t) (fun x => tmPrint x)). *)
  (* > ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ *)
  (* Error: Constant not found : Top.x.x *)

"not a sort" anomaly

Why the following example:

Require Import MetaCoq.Template.All.
Require Import List.
Import ListNotations MonadNotation.

Parameter A : Type.
Parameter F: Vector.t A 3 -> nat.

Definition a :=
  tProd (nNamed "a")
        (tApp
           (tInd {| inductive_mind := "Coq.Vectors.VectorDef.t"; inductive_ind := 0 |}
                 [])
           [tConst "A" [];
              tApp
                (tConstruct
                   {| inductive_mind := "Coq.Init.Datatypes.nat"; inductive_ind := 0 |} 1
                   [])
                [tApp
                   (tConstruct
                      {|
                        inductive_mind := "Coq.Init.Datatypes.nat";
                        inductive_ind := 0 |} 1 [])
                   [tApp
                      (tConstruct
                         {|
                           inductive_mind := "Coq.Init.Datatypes.nat";
                           inductive_ind := 0 |} 1 [])
                      [tConstruct
                         {|
                           inductive_mind := "Coq.Init.Datatypes.nat";
                           inductive_ind := 0 |} 0 []]]]])
        (tApp (tConst "F" []) [tRel 0]).

Run TemplateProgram (tmUnquote a >>= tmPrint).

Gives me:

Error: Anomaly "Not a sort." Please report at http://coq.inria.fr/bugs/.

Environment:

  1. Coq 8.8.2
  2. MetCoq coq-8.8 branch from git
  3. OCaml 4.07.0

`eq_term` does not account for the difference eq/leq

Right now, eq_term is parametrised by a relation on universes, but it should actually have two, so that leq_term can use a different one when recursing on domains.

At the moment, it checks A ≤ A' and B ≤ B' to get A → B ≤ A' → B'.

Anomaly when unquoting polymorphic universes

I am quite puzzled by the following snippet

Quote Definition t0 := nat.
Fail Run TemplateProgram (tmUnquoteTyped Type t0).
(* The command has indeed failed with message: *)
(* Conversion test raised an anomaly: Anomaly "Universe Top.4964 undefined." *)
(* Please report at http://coq.inria.fr/bugs/. *)

Polymorphic Definition ty_poly : Type := Type.
Fail Run TemplateProgram (tmUnquoteTyped ty_poly t0).
(* same as above *)
Definition ty : Type := Type.
Run TemplateProgram (tmUnquoteTyped ty t0).

The only way to unquote the quoted nat in a typed way is to provide a non-polymorphic definition of Type.

Making compiling with coq 8.9 and coq master

Dear MetaCoq devs,
I created a branch that is compiling with coq 8.9:
https://github.com/MetaCoq/metacoq/tree/coq-8.9
Please make / rebase your PR on top of it.

Remain to do:

  • edit readme and make this branch the default branch after a bit of testing
  • rebase master on top of it (currently, all commits of master are to make compiling with coq master, no new feature with respect to coq-8.9)
  • merge pcuic extraction

Compatibility with Coq 8.10

Hi all,

Since 8.10+beta1 is out and SProp is included in this release, it's probably time to start porting MetaCoq to support Coq 8.10 (which adds SProp to MetaCoq as well).
I've added support of SProp to MetaCoq at some point https://github.com/annenkov/template-coq/tree/master but now it is out of sync with master branch of the main MetaCoq repo.
What do you think, how should we approach this?

template-coq plugin interacts badly with asynchronous proofs

Trying to use asynchronous proofs in coqide or coqc -schedule-vio2vo when Template.Loader has been loaded cause an anomaly:

$ cat bug.v
From Template Require Import Loader.

Lemma test: 0 = 0.
Proof.
  trivial.
Qed.
$ coqc -quick bug.v
$ coqc -schedule-vio2vo 1 bug
Warning: File ./bug.v: proof of test: chars 66-74
Anomaly "cannot find String in module Coq.Strings.String."
Please report at http://coq.inria.fr/bugs/.

On the other hand, if we include everything but the template-coq ML module, things work:

$ cat bug.v
From Template Require Import Ast TemplateMonad.

Lemma test: 0 = 0.
Proof.
  trivial.
Qed.
$ coqc -quick bug.v
$ coqc -schedule-vio2vo 1 bug

I am not sure if this should be reported here or on the Coq issue tracker; I reported here because it's specific to the template-coq plugin.

Versions:

  • coq 8.9.0 (from OPAM, OCaml 4.06.0)
  • metacoq - branch coq-8.9, commit 6f306ae.

make install from coq8.8 installs in Top instead of Template

When I do make install in branch coq 8.8, the script succeeds but then files are installed in user-contrib/Top instead of user-contrib/Template.
This is annoying because I cannot refer to it in a separate project (like I do for the coq 8.7 release).

This is probably related to the warnings we get during make:

COQC theories/Loader.v
Warning: /Users/theo/Documents/PhD/template-coq/checker/theories was
previously bound to Top.theories; it is remapped to TemplateChecker
[overriding-logical-loadpath,loadpath]
COQC demo.v
Warning: /Users/theo/Documents/PhD/template-coq/checker/theories was
previously bound to Top.theories; it is remapped to TemplateChecker
[overriding-logical-loadpath,loadpath]

Helper function for tCase

While playing around with MetaCoq, I found getting tCase right to be quite a hurdle, because of all the required annotations.

So I wrote a little wrapper that seems fairly general-purpose. Did I miss some existing utility, or does this tMatch function below seem to have a place in the library? It's very much in draft state, but if you have any comments to simplify this, I would welcome that too!

Roughly, given x, y, z : term and a function branches, tMatch x y z branches constructs the expression match x : y return z where ... end, assuming y is some inductive type, using the function branches to fill every branch given the name, type, and arity of the corresponding constructor.

https://gist.github.com/Lysxia/6a9025aeec5f89f36728da1426a3c9dd

Adapting to coq 8.8, extraction problem

With @annenkov, we are trying to adapt Template Coq to coq 8.8alpha:
https://github.com/Template-Coq/template-coq/commits/coq-8.8

We have a problem with extraction. The dependency of TemplateCoqCompiler.v over template_coq_plugin.mllib is not recorded.
This is, according to @ppedrot, due to the new dependency system where there is a global .coqdeps.d generated: it is generated by the first Makefile and not by the others. If I try to manually remove it (3310f81) so that it is regenerated, I have a linking error.

@ppedrot suggests splitting the three layers of plugin in three directory. @mattam82 Any suggestion?

intactic no longer used?

I was looking through the code in run_template_program_rec and it looks like the optional intactic argument is no longer used anywhere? Is this a bug, or is it no longer necessary?

First beta release

I'd like to make a beta-release for the current coq-8.7 branch so that we can put a package on opam to be used by others (in particular certicoq). Do you guys see other things we absolutely need to do before doing so, or expected changes in the "interfaces"/inductive definitions that we should rather implement before making a beta release (I'd hope Ast.v/univ.v will not move much anymore) ? @aa755 @CohenCyril @SimonBoulier @tabareau ?

  • Do a cleanup pass on univ/uGraph/Ast/Typing.v files
  • Produce coqdoc corresponding to that part
  • Remove program #18
  • Template Cast cleanups #20

Project name change

After discussion with a number of people, it seems that template-Coq is not the best name to use for the whole project. @betaziliani is ok with us reusing the name MetaCoq (which was used for a former MTac version) and people seem to like it. To me it sounds better as well, given the we provide much different functionality than just reifying / reflecting terms, with the template monad to script the vernacular and the typing/checker work to provide a kind of Coq-in-Coq implementation. In my opinion, Template-Coq / MetaCoq-Template could still refer to the core reification/reflection framework we have, while MetaCoq-Gallina (or some better name) could refer to the current template monad, MetaCoq-Checker, MetaCoq-Extraction and the various translations should also be separate projects under a larger MetaCoq umbrella. If you agree, I can create a new organization and work on separating the various parts into different projects, which should be easy as they are already in separate directories. So, what do you think?

Better quoting of universes

We should quote universes properly.
@mattam82 Why do you propose

Definition universe := list (level * bool).

(with true for +1)? In the kernel it seems to be list (level * int).

Also, do you know how I can access the level list of universes: the type universe seems not to be exposed and there are no such functions. Do we need to expose a new function in the code of Coq?

tmLemma treats implicit arguments wrong

Run TemplateProgram (tmLemma "test" (@nil nat = @nil nat)).

creates the error message Error: Cannot infer the implicit parameter A of nil whose type is "Type".

Interestingly, this also happens after I do Arguments nil _ : clear implicits.

Branch managing

Master is out-dated and it is the branch people see when they arrive on the web page ...
We should follow coq.8.[last] no?

Implement tmInstance

We have a use case for a tmInstance templateMonad command that would take a type and return (if proof search succeeds) the associated Coq proof term. Should be easily implementable as a call to resolve_one_typeclass. Maybe a more general interface could be found.

Plugin Extraction (more)

The things that remaining.

  • Write an Ocaml interpreter for the extracted version of Template.Monad.Extractable.TM which interprets it into PluginCore.tm.
  • Determine whether it is possible to directly translate Template.Ast.term to Constr.t using extraction directives, e.g. Extract Inductive. If this doesn't work then we need to write a function that converts back and forth.
  • Extract a library that can be exposed through a hand-written ml4 file.

Compatibility with HoTT

It would be nice to make MetaCoq compatible with the HoTT library (https://github.com/HoTT/HoTT). I know that @andreaslyn would be interested in this kind of compatibility. I guess, making translations to work with HoTT would be already nice, the rest (like CIC meta-theory and extraction) is probably not needed to be compatible.
Of course, this requires at least to resolve #138 first.

Make kername fully qualified

Since name resolution is context dependent, we should make it an explicit effect in the monad rather than something that happens implicitly during denotation.

Template polymorphism is not implemented

Template polymorphism is not implemented.
For template polymorphism see: refman
The following compiles in Coq:

Definition T := {A : Type & {B : Type & A}}.

While it does not in Template Coq:

Require Import Template.TemplateCoqChecker.
Fail Template Check T.

I am not sure we will implement it because we can use full polymorphism instead but people should be aware of it.
@mattam82 Do you confirm?

Verified extraction to OCaml

To verify extraction to OCaml, one needs a semantics of OCaml in Coq (which may or may not be verified itself in some way). This is an attempt to document relevant resources to make this happen.

One early effort on formalizing OCaml semantics is OCaml Light by Scott Owens and his collaborators. The language and semantics is encoded in the Ott language and then exported to HOL4, Isabelle/HOL, and Coq. The soundness proof of the type system is only in HOL4, however. A paper describing the formalization and proofs was published in ESOP 2008. All files are available in the Ott repository on GitHub. The repository version of Ott can produce a definition of OCaml Light that compiles with recent Coq.

While there are some related efforts to formalize OCaml in different proof assistants, as discussed here, OCaml Light seems to me to still be the most realistic and relevant one for verifying extraction in MetaCoq.

CI broken

Travis CI is broken currently on all coq-8.8 builds, and I had the same compilation problem locally.

After a git clean -xf, I can now build perfectly. I think the problem appeared in commit df158c4, where Equations was introduced.

Is the CI doing incremental builds? If so, we probably need a git clean there as well.

Issue when chaining parametricity translation with definition

When compiling the translations directory (with make translations), I get the following error message: Error: ID not found.

I investigated the issue with @CohenCyril and we found out that chaining the use of tTranslate with a definition using the new translation fails. For instance:

Notation t :=
     (tInd {| inductive_mind := "Top.boolᵗ"; inductive_ind := 0 |} []).
Fail Run TemplateProgram (tTranslate emptyTC "bool";; tmMkDefinition "toto" t).

After these two commands, Coq answers

boolᵗ is defined
boolᵗ_rect is defined
boolᵗ_ind is defined
boolᵗ_rec is defined
"bool has been translated as boolᵗ"
The command has indeed failed with message:
         bool not found

This error message comes from the tTranslate function so there is some sort of backtracking to this function.

This issue does not happen if we split the program in two instead:

Notation t :=
     (tInd {| inductive_mind := "Top.boolᵗ"; inductive_ind := 0 |} []).
Run TemplateProgram (tTranslate emptyTC "bool").
Run TemplateProgram (tmMkDefinition "toto" t).

translations and universes

@SimonBoulier we have a problem: we can't port translations easily to your new representation of universes because there's no "empty" universe anymore, which was used to refresh universes before. I'm not sure what to do? Maybe we should make term polymorphic over the representation of universes, allowing unquoting of terms where some universes are None (which will get refreshed (heavy!)), or put the translations in the monad and add actions new_univ : TM universe_level, add_constraint : univ_constraint -> TM unit that would add the constraint. The interpretation of the monad is already up to an evar_map so these will be very easy to implement.

extraction mode: adding level to environment in unquote_universe

There seems to be a discrepancy in how unquote_level is instantiated in extraction mode and in template-coq (live-coq) mode.

The extraction mode implementation is here and looks like:

  let unquote_level (trm : Univ0.Level.t) : Univ.Level.t =
    match trm with
    | Univ0.Level.Coq_lProp -> Univ.Level.prop
    | Univ0.Level.Coq_lSet -> Univ.Level.set
    | Univ0.Level.Level s ->
      let s = unquote_string s in
      let comps = CString.split '.' s in
      let last, dp = CList.sep_last comps in
      let dp = DirPath.make (List.map Id.of_string comps) in
      let idx = int_of_string last in
      Univ.Level.make dp idx
    | Univ0.Level.Var n -> Univ.Level.var (unquote_int n)

Clearly from the types, it does not thread the environment (Evd. evar_map).
In contrast, the template-coq version, implemented here, threads the environment and makes a non-trivial change to it (the tLevel case calls get_level which can add a universe to the environment):

let get_level evm s =
  if CString.string_contains ~where:s ~what:"." then
    match List.rev (CString.split '.' s) with
    | [] -> CErrors.anomaly (str"Invalid universe name " ++ str s ++ str".")
    | n :: dp ->
       let num = int_of_string n in
       let dp = DirPath.make (List.map Id.of_string dp) in
       let l = Univ.Level.make dp num in
       try
         let evm = Evd.add_global_univ evm l in
         if !strict_unquote_universe_mode then
           CErrors.user_err ~hdr:"unquote_level" (str ("Level "^s^" is not a declared level and you are in Strict Unquote Universe Mode."))
         else (Feedback.msg_info (str"Fresh universe " ++ Level.pr l ++ str" was added to the context.");
               evm, l)
       with
       | UGraph.AlreadyDeclared -> evm, l
  else
    try
      evm, Evd.universe_of_name evm (Id.of_string s)
    with Not_found ->
         try
           let univ, k = Nametab.locate_universe (Libnames.qualid_of_string s) in
           evm, Univ.Level.make univ k
         with Not_found ->
           CErrors.user_err ~hdr:"unquote_level" (str ("Level "^s^" is not a declared level."))

let unquote_level evm trm (* of type level *) : Evd.evar_map * Univ.Level.t =
  let (h,args) = app_full trm [] in
  if Constr.equal h lProp then
    match args with
    | [] -> evm, Univ.Level.prop
    | _ -> bad_term_verb trm "unquote_level"
  else if Constr.equal h lSet then
    match args with
    | [] -> evm, Univ.Level.set
    | _ -> bad_term_verb trm "unquote_level"
  else if Constr.equal h tLevel then
    match args with
    | s :: [] -> debug (fun () -> str "Unquoting level " ++ pr_constr trm);
                 get_level evm (unquote_string s)
    | _ -> bad_term_verb trm "unquote_level"
  else if Constr.equal h tLevelVar then
    match args with
    | l :: [] -> evm, Univ.Level.var (unquote_nat l)
    | _ -> bad_term_verb trm "unquote_level"
  else
    not_supported_verb trm "unquote_level"

Should the extraction version be fixed to thread/update the context in a similar way?

Reorganize the available commands in the monad

Here is a list of possible changes, let me know what you think of them.

  • tmUnquote [term -> TemplateMonad {A : Type & A}] : todo

  • tmMkDefinition:
    replace by tmDefinition [ident -> forall {A:Type}, A -> TemplateMonad unit] (primitive without unquoting)
    and tmMkDefinition [ident -> term -> TemplateMonad unit] (derived, with unquoting)

  • the actual tmQuote could be derived from:

    • tmQuoteTerm [forall {A:Type}, A -> TemplateMonad term]

    • tmGetBody [ident -> TemplateMonad {A : Type & option A}]
      (None if it is an axiom (or an inductive/constructor/projection ?))
      (might be implemented as an "unfold" if the reduction strategies were richer)

    • tmQuoteInductive [ident -> TemplateMonad mutual_inductive_entry]

  • tmUnQReduceQ: derived

  • tmFreshName : I propose to either return a fresh name with the argument as prefix, or to change the name as "tmIsFreshName"

  • add a tmFreshUniv : generate a fresh universe

  • add a tmLemma [Type -> bool (* opacity *)-> TemplateMonad unit] : open a goal and make the definition

  • add tmAxiom [Type -> TemplateMonad unit]

  • tmLocate ? [ident -> TemplateMonad fullpath] (instead of using it in tmGetBody)

And for future : Fixpoint, Extraction, Let, Program Definition, Existing Instance, Existing Class, ... :-)

Deprecate `tmMkDefinition`?

Looking at the current implementation, it looks like tmMkDefinition is s special case of tmDefinition. Does it make sense to drop tmMkDefinition?

tmUnquote unquotes invalid terms

I stumbled across an invalid, untypable term (T below) accepted by tmUnquote. This might be intended behaviour, but it's for sure not intuitive. I would have expected tmUnquote to tell me that the term does not have a type.

I was able to boil down the definition to the example below. Unquoting T and printing the typed_term works, but copy-pasting the printed output into a definition does not work (rightly so).

Even more problematic: Unquoting T and passing the result to the tmDefinition operation raises an anomaly.

Require Import Template.All.
Require Export String List.

Import ListNotations.
Import MonadNotation.

Inductive test (X : Type) := test_T : test X -> test X.

Definition T := 
tFix
  [mkdef term (nNamed "f") (tProd (nNamed "x") (tApp (tInd (mkInd "Top.test" 0) []) [tInd (mkInd "Coq.Init.Datatypes.unit" 0) []]) (tInd (mkInd "Coq.Init.Datatypes.unit" 0) []))
     (tLambda (nNamed "x") (tApp (tInd (mkInd "Top.test" 0) []) [tRel 0])
        (tCase (mkInd "Top.test" 0, 1)
           (tLambda (nNamed "x") (tApp (tInd (mkInd "Top.test" 0) []) [tInd (mkInd "Coq.Init.Datatypes.unit" 0) []]) (tInd (mkInd "Coq.Init.Datatypes.unit" 0) [])) 
           (tRel 0)
           [(1, tLambda (nNamed "x0") (tApp (tInd (mkInd "Top.test" 0) []) [tInd (mkInd "Coq.Init.Datatypes.unit" 0) []]) (tApp (tRel 2) [tRel 0]))]))
     0] 0.
Run TemplateProgram (tmUnquote T >>= tmPrint).

Fail Let bla := (existT_typed_term (test unit -> unit) (fix f (x : test f) : unit := match x with
                                                                              | test_T _ x0 => f x0
                                                                              end)).
Run TemplateProgram (tmUnquote T >>= tmDefinition "fails").

The last command raises

Error: Anomaly "Uncaught exception Reduction.NotArity."
Please report at http://coq.inria.fr/bugs/.

That's a mix between a Coq Bug and a Template Coq bug, but I felt that writing it down here is the first step. The problem is not really bothering me (I shouldn't unquote garbage anyways), but eventually we should probably fix that.

Constr vs Econstr

Everywhere in the code we convert our Constr.t to some EConstr.t and back.
I consider moving everything to EConstr (and continue carrying the evarmap).
Do you agree it is the right thing to do?

Defining lambda from Ast

The following code fails with "Universe Var(0) undefined":

Definition f@{i j k} := fun (E:Type@{i}) => TypeVal@{i k} E (type@{i j} E) (TypeErr@{i j} E).
Quote Definition qf := Eval cbv in f. (* To expands definition *)
(*
qf = 
tLambda (nNamed "E") (tSort [(Level.Level "Top.257", false)])
  (tApp
     (tConstruct {| inductive_mind := "Top.type"; inductive_ind := 0 |} 0
        [Level.Level "Top.257"; Level.Level "Top.259"])
     [tRel 0;
     tApp
       (tInd {| inductive_mind := "Top.type"; inductive_ind := 0 |}
          [Level.Level "Top.257"; Level.Level "Top.258"]) [tRel 0];
     tApp
       (tConstruct {| inductive_mind := "Top.type"; inductive_ind := 0 |} 1
          [Level.Level "Top.257"; Level.Level "Top.258"]) [tRel 0]])
     : term
*)
Make Definition uqf := Eval cbv in qf.

It would be nice if the following is accepted:


Make Definition uqf := 
tLambda (nNamed "E") (tSort [(Level.Var 0, false)])
  (tApp
     (tConstruct {| inductive_mind := "Top.type"; inductive_ind := 0 |} 0
        [Level.Var 0; Level.Var 2])
     [tRel 0;
     tApp
       (tInd {| inductive_mind := "Top.type"; inductive_ind := 0 |}
          [Level.Var 0; Level.Var 1]) [tRel 0];
     tApp
       (tConstruct {| inductive_mind := "Top.type"; inductive_ind := 0 |} 1
          [Level.Var 0; Level.Var 1]) [tRel 0]])

Declaring new universes for every Var.

The inductive type is polymorphic, so to my understanding I have to give the universe instance. Therefore, it should be possible te generate fresh universes variables.

TemplateMonad Semantics Questions

I have two questions about the semantics of TemplateMonad:

  1. What are the semantics of tmLemma? It seems to take a type, but not a value of that type to be used as the body. Where does the body of the lemma (i.e. the proof) come from?
  2. Is there a reason to have both tmDefinitionRed and tmMkDefinition as primitives? It seems like it is possible to define tmDefinitionRed in terms of tmMkDefinition and tmQuote (and tmEval) or to define tmMkDefinition in terms of tmDefinitionRed and tmUnquote.

Unfolding fully qualified names

tmEval (unfold "name") does not allow fully qualified names like "a.b.c".
The error is: Invalid character '.' in identifier "MathClasses.interfaces.canonical_names.plus".

make plugin in template-coq alqays recompiles

When I run make and then make immediately again, there are recompilations happening during the second make. It's happening in template-coq on the plugin subgoal - this always recompiles things, no matter what.

Is that by design? I'm currently using MetaCoq as submodule in another project, and my make all first does make all on all submodules. Previously (at least on commit d5088ac) there was nothing done in template-coq if nothing changed.

@gmalecha @aa755 do you think it's easy to fix this? It's not an issue for people installing via opam and I can probably work around it, but maybe it's easy to resolve.

Variant types

I can create Inductive types. How can I create Variant types in cases where I do not need and induction scheme generated.

Quoting inside sections

Hi all,

I have a use case for quoting inside sections. For now, I just commented out the check_inside_section (); in the definition of Run TemplateProgram, which is working for most examples. However, I now have a case where I get the error message Error: Invalid character '#' in identifier "Top#Fix_X#type_unit_termT". where Fix_X is the section name and type_unit_term is a self-chosen identifier passed to tmLemma.

I suppose errors like this are the reason why quoting inside sections is disabled. But I wonder what the technical reason is. Using Program inside section seems to be generally fine.

Any ideas what the reason is? And maybe even how much work it would be to make Template Coq work in sections consistently?

PCUIC vs Ast.term

I just recently started looking into PCUIC and it looks really great. Would it be possible to replace the Template.Ast.term definition with this one? Are there features that exist in one but not in the other?

Cannot compile v2.1-beta3

I followed the compilation instructions but could not get the latest version to compile; there seems to be a problem with the extraction subdirectory (the compilation output is here).

The previous tag, v2.1-beta2, compiles just fine. I am running Coq 8.8.1 and OCaml 4.04.1.

template-coq branch coq-8.8 fails to build with Coq 8.8.0 with template_coq_plugin.cmxs: undefined symbol: camlReify__fun_8253

https://travis-ci.org/mit-plv/reification-by-parametricity/jobs/374189145#L3250

CAMLOPT -a -o src/template_coq_plugin.cmxa
findlib: [WARNING] Interface logic.cmi occurs in several directories: /usr/lib/coq/proofs, src
CAMLOPT -shared -o src/template_coq_plugin.cmxs
findlib: [WARNING] Interface logic.cmi occurs in several directories: /usr/lib/coq/proofs, src
COQC theories/TemplateCoqCompiler.v
File "./theories/TemplateCoqCompiler.v", line 3, characters 0-40:
Error:
while loading
/home/travis/build/mit-plv/reification-by-parametricity/template-coq/src/template_coq_plugin.cmxs:
error loading shared library: /home/travis/build/mit-plv/reification-by-parametricity/template-coq/src/template_coq_plugin.cmxs: undefined symbol: camlReify__fun_8253

template-coq does not build with Coq master nor 8.7

CAMLDEP src/reify.mli
COQDEP src/template_plugin.mllib
CAMLDEP -pp src/reify.ml4
COQDEP theories/Template.v
COQDEP theories/Ast.v
*** Warning: template_plugin.mllib already found in src (discarding src/template_plugin.mllib)
COQC theories/Ast.v
CAMLC -c src/reify.mli
File "src/reify.mli", line 21, characters 17-28:
Warning 3: deprecated: API.Names.kernel_name
alias of API.Names.KerName.t
CAMLOPT -pp -c  src/reify.ml4
File "src/reify.ml4", line 8, characters 2-18:
Error: Unbound record field Goptions.optsync
Makefile.coq:582: recipe for target 'src/reify.cmx' failed
make[2]: *** [src/reify.cmx] Error 2
make[2]: *** Waiting for unfinished jobs....
Makefile.coq:315: recipe for target 'all' failed
make[1]: *** [all] Error 2
make[1]: Leaving directory '/home/jgross/Downloads/coq/template-coq-master'
Makefile:4: recipe for target 'coq' failed
make: *** [coq] Error 2

Make Template Coq work if called from a tactic

I would like to use Template Coq from within Ltac.

The main problem seems to be the Program usage of tmLemma. In tactics, Program is probably not needed, because a template program can just generate subgoals.

I plan to work on this myself this week eventually, just wanted to document the feature request somewhere.

Universes are slow

Here's an idea: to check universe constraints in the checker, we first build the transitive closure of the graph, then we should get constant time comparisons. We could also use a kind of union-find like in the kernel to handle equal universe levels.

Remove tMeta

Can I remove tMeta from Ast.term?
It is polluting for nothing ...

About universes

Currently, the typing rule for sorts is:

               Γ ⊢
--------------------------------------
Γ ⊢ Universe.make l : Universe.super l

Hence, algebraic universes (i.e. universes which are not a single level) should not appear at the left of a colon.
Do you agree that type inference should fail on terms containing algebraic universes?

Wellformedness is then:

Γ ⊢ t : A  ->  Γ ⊢ A : s  or  A is a sort

Is it correct?

If so, should we change the conversion rule from:

Γ ⊢ t : A      Γ ⊢ B : s  
------------------------------------  (A ≡ B)
            Γ ⊢ t : B

to

Γ ⊢ t : A      Γ ⊢ B : s or B a sort
------------------------------------  (A ≡ B)
            Γ ⊢ t : B

???

_CoqProject files broken

As noticed by @TheoWinterhalter, the _CoqProject files on coq-8.8 don't work in Proof General at the moment, because they don't have the necessary -I/-R statements. Using make it works, because the DEPS variables is passed.

I don't think there is an issue yet, so I'm opening this one to start a discussion.

@mattam82 mentioned that we can automate this. Maybe by adding a make proofgeneral target in the toplevel Makefile?

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.