Library MetaCoq.Template.TemplateMonad.Extractable

From MetaCoq.Template Require Import utils Ast AstUtils TemplateMonad.Common.

Local Set Universe Polymorphism.

The Extractable Template Monad

A monad for programming with Template Coq structures. Use Run TemplateProgram on a monad action to produce its side-effects.

Cumulative Inductive TM@{t} : Type@{t} Type :=

| tmReturn {A:Type@{t}}
  : A TM A
| tmBind {A B : Type@{t}}
  : TM A (A TM B) TM B


| tmPrint : Ast.term TM unit
| tmMsg : string TM unit
| tmFail : {A:Type@{t}}, string TM A
| tmEval (red : reductionStrategy) (tm : Ast.term)
  : TM Ast.term


| tmDefinition_ (opaque : bool)
               (nm : ident)
               (type : option Ast.term) (term : Ast.term)
  : TM kername
| tmAxiom (nm : ident)
          (type : Ast.term)
  : TM kername
| tmLemma (nm : ident)
          (type : Ast.term)
  : TM kername


| tmFreshName : ident TM ident

| tmLocate : qualid TM (list global_reference)
| tmCurrentModPath : TM modpath


| tmQuoteInductive (nm : kername)
  : TM mutual_inductive_body
| tmQuoteConstant (nm : kername) (bypass_opacity : bool)
  : TM constant_body
| tmQuoteUniverses : TM ConstraintSet.t
| tmQuoteModule : qualid TM (list global_reference)



| tmInductive : bool mutual_inductive_entry TM unit


| tmExistingInstance : global_reference TM unit
| tmInferInstance (type : Ast.term)
  : TM (option Ast.term)
.

Definition TypeInstance : Common.TMInstance :=
  {| Common.TemplateMonad := TM
   ; Common.tmReturn:=@tmReturn
   ; Common.tmBind:=@tmBind
   ; Common.tmFail:=@tmFail
   ; Common.tmFreshName:=@tmFreshName
   ; Common.tmLocate:=@tmLocate
   ; Common.tmCurrentModPath:=fun _ ⇒ @tmCurrentModPath
   ; Common.tmQuoteInductive:=@tmQuoteInductive
   ; Common.tmQuoteUniverses:=@tmQuoteUniverses
   ; Common.tmQuoteConstant:=@tmQuoteConstant
   ; Common.tmMkInductive:=@tmInductive
   ; Common.tmExistingInstance:=@tmExistingInstance
   |}.

Definition tmOpaqueDefinition (nm : ident)
               (type : option Ast.term) (term : Ast.term) :=
  tmDefinition_ true nm type term.

Definition tmDefinition (nm : ident)
               (type : option Ast.term) (term : Ast.term) :=
  tmDefinition_ false nm type term.

Definition tmInductive' (mind : mutual_inductive_body) : TM unit
  := tmInductive false (mind_body_to_entry mind).

Definition tmLemmaRed (i : ident) (rd : reductionStrategy)
           (ty : Ast.term) :=
  tmBind (tmEval rd ty) (fun tytmLemma i ty).

Definition tmAxiomRed (i : ident) (rd : reductionStrategy) (ty : Ast.term)
  :=
    tmBind (tmEval rd ty) (fun tytmAxiom i ty).

Definition tmDefinitionRed (opaque : bool) (i : ident) (rd : reductionStrategy)
           (ty : option Ast.term) (body : Ast.term) :=
  match ty with
  | NonetmDefinition_ opaque i None body
  | Some ty
    tmBind (tmEval rd ty) (fun tytmDefinition_ opaque i (Some ty) body)
  end.

Definition tmInferInstanceRed (rd : reductionStrategy) (type : Ast.term)
  : TM (option Ast.term) :=
  tmBind (tmEval rd type) (fun typetmInferInstance type).