Library MetaCoq.Template.Ast



AST of Coq kernel terms and kernel data structures

Basic data-types:

We reflect identifiers ident, sort families sort_family, names name, cast kinds cast_kind, inductives inductive and primitive projections projection and (co)-fixpoint blocks mfixpoint and def. These are defined in the BasicAst file.

Terms:

The AST is term : Set
Smart constructors mkApp, mkApps maintain the invariant of no nested or empty n-ary applications. List in fixpoints and cofixpoint should be non-empty.

Kernel interface: entries and declarations

Kernel input declarations for constants constant_entry and mutual inductives mutual_inductive_entry. Kernel safe declarations for constants constand_decl and inductives minductive_decl.

Environments of declarations

The global environment global_env_ext: a list of global_decl and a universe graph constraints.


Inductive term : Set :=
| tRel (n : nat)
| tVar (id : ident)
| tEvar (ev : nat) (args : list term)
| tSort (s : universe)
| tCast (t : term) (kind : cast_kind) (v : term)
| tProd (na : name) (ty : term) (body : term)
| tLambda (na : name) (ty : term) (body : term)
| tLetIn (na : name) (def : term) (def_ty : term) (body : term)
| tApp (f : term) (args : list term)
| tConst (c : kername) (u : universe_instance)
| tInd (ind : inductive) (u : universe_instance)
| tConstruct (ind : inductive) (idx : nat) (u : universe_instance)
| tCase (ind_and_nbparams: inductive×nat) (type_info:term)
        (discr:term) (branches : list (nat × term))
| tProj (proj : projection) (t : term)
| tFix (mfix : mfixpoint term) (idx : nat)
| tCoFix (mfix : mfixpoint term) (idx : nat).

Definition mkApps t us :=
  match us with
  | nilt
  | _match t with
        | tApp f argstApp f (args ++ us)
        | _tApp t us
        end
  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.

Well-formed terms: invariants which are not ensured by the OCaml type system

Inductive wf : term Prop :=
| wf_tRel n : wf (tRel n)
| wf_tVar id : wf (tVar id)
| wf_tEvar n l : Forall wf l wf (tEvar n l)
| wf_tSort u : wf (tSort u)
| wf_tCast t k t' : wf t wf t' wf (tCast t k t')
| wf_tProd na t b : wf t wf b wf (tProd na t b)
| wf_tLambda na t b : wf t wf b wf (tLambda na t b)
| wf_tLetIn na t b b' : wf t wf b wf b' wf (tLetIn na t b b')
| wf_tApp t u : isApp t = false u nil wf t Forall wf u wf (tApp t u)
| wf_tConst k u : wf (tConst k u)
| wf_tInd i u : wf (tInd i u)
| wf_tConstruct i k u : wf (tConstruct i k u)
| wf_tCase ci p c brs : wf p wf c Forall (Program.Basics.compose wf snd) brs wf (tCase ci p c brs)
| wf_tProj p t : wf t wf (tProj p t)
| wf_tFix mfix k : Forall (fun defwf def.(dtype) wf def.(dbody) isLambda def.(dbody) = true) mfix
                   wf (tFix mfix k)
| wf_tCoFix mfix k : Forall (fun defwf def.(dtype) wf def.(dbody)) mfix wf (tCoFix mfix k).

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 := mkdecl {
  decl_name : name ;
  decl_body : option term ;
  decl_type : term }.

Local (de Bruijn) variable binding

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

Local (de Bruijn) let-binding

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

Local (de Bruijn) context

Definition context := list context_decl.

Last declaration first

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

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

Environments

See one_inductive_body from declarations.ml.
Record one_inductive_body := {
  ind_name : ident;
  ind_type : term;
  ind_kelim : list sort_family;
  ind_ctors : list (ident × term
                    × nat );
  ind_projs : list (ident × term) }.

See mutual_inductive_body from declarations.ml.
See constant_body from declarations.ml

Programs

A set of declarations and a term, as produced by Quote Recursively.

Definition program : Type := global_env × term.