Library MetaCoq.Erasure.EAst

From MetaCoq.Template Require Import utils BasicAst Universes.

Extracted terms

These are the terms produced by extraction: compared to kernel terms, all proofs and types are translated to tBox or erased (type annotations at lambda and let-ins, types of fix/cofixpoints), applications are in binary form and casts are removed.

Declare Scope erasure.
Local Open Scope erasure.

Record def (term : Set) := { dname : name; dbody : term; rarg : nat }.
Derive NoConfusion for def.
Arguments dname {term} d.
Arguments dbody {term} d.
Arguments rarg {term} d.

Definition map_def {term : Set} (f : term term) (d : def term) :=
  {| dname := d.(dname); dbody := f d.(dbody); rarg := d.(rarg) |}.

Definition test_def {term : Set} (f : term bool) (d : def term) :=
  f d.(dbody).

Definition mfixpoint (term : Set) := list (def term).

Inductive term : Set :=
| tBox : term
| tRel : nat term
| tVar : ident term
| tEvar : nat list term term
| tLambda : name term term
| tLetIn : name term term term
| tApp : term term term
| tConst : kername term
| tConstruct : inductive nat list term term
| tCase : (inductive × nat)
               term list (list name × term) term
| tProj : projection term term
| tFix : mfixpoint term nat term
| tCoFix : mfixpoint term nat term.

Derive NoConfusion for term.

Bind Scope erasure with term.

Fixpoint mkApps t us :=
  match us with
  | nilt
  | a :: argsmkApps (tApp t a) args
  end.

Definition mkApp t u := Eval cbn in mkApps t [u].

Definition isApp t :=
  match t with
  | tApp _ _true
  | _false
  end.

Definition isLambda t :=
  match t with
  | tLambda _ _true
  | _false
  end.

Entries

The kernel accepts these inputs and typechecks them to produce declarations. Reflects kernel/entries.mli.

Constant and axiom entries

Inductive entries

This is the representation of mutual inductives. nearly copied from kernel/entries.mli
Assume the following definition in concrete syntax:

  Inductive I1 (x1:X1) ... (xn:Xn) : A1 := c11 : T11 | ... | c1n1 : T1n1
  ...
  with Ip (x1:X1) ... (xn:Xn) : Ap := cp1 : Tp1 ... | cpnp : Tpnp.
then, in ith block, mind_entry_params is xn:Xn;...;x1:X1; mind_entry_arity is Ai, defined in context x1:X1;...;xn:Xn; mind_entry_lc is Ti1;...;Tini, defined in context A'1;...;A'p;x1:X1;...;xn:Xn where A'i is Ai generalized over x1:X1;...;xn:Xn.

Declarations

The context of De Bruijn indices


Record context_decl := {
  decl_name : name ;
  decl_body : option term }.

Local (de Bruijn) variable binding

Definition vass x := {| decl_name := x; decl_body := None |}.

Local (de Bruijn) let-binding

Definition vdef x t := {| decl_name := x; decl_body := Some t |}.

Local (de Bruijn) context

Definition context := list context_decl.

Bind Scope erasure with context.

Mapping over a declaration and context.

Definition map_decl f (d : context_decl) :=
  {| decl_name := d.(decl_name);
     decl_body := option_map f d.(decl_body) |}.

Definition map_context f c :=
  List.map (map_decl f) c.

Last declaration first

Definition snoc {A} (Γ : list A) (d : A) := d :: Γ.

Notation " Γ ,, d " := (snoc Γ d) (at level 20, d at next level) : erasure.

Environments


Record constructor_body :=
  mkConstructor {
    cstr_name : ident;
    cstr_nargs : nat
  }.
Derive NoConfusion for constructor_body.

Record projection_body :=
  mkProjection {
    proj_name : ident;
  }.
Derive NoConfusion for projection_body.

See one_inductive_body from declarations.ml.
See mutual_inductive_body from declarations.ml.
See constant_body from declarations.ml
A context of global declarations

Programs

A set of declarations and a term, as produced by extraction from template-coq programs.

Definition program : Type := global_declarations × term.