package coq

  1. Overview
  2. Docs

doc/ltac_plugin/Ltac_plugin/Tacentries/index.html

Module Ltac_plugin.TacentriesSource

Ltac toplevel command entries.

Tactic Definitions
Sourceval register_ltac : Vernacexpr.locality_flag -> ?deprecation:Deprecation.t -> Tacexpr.tacdef_body list -> unit

Adds new Ltac definitions to the environment.

Tactic Notations
Sourcetype 'a grammar_tactic_prod_item_expr = 'a Pptactic.grammar_tactic_prod_item_expr =
  1. | TacTerm of string
  2. | TacNonTerm of ('a * Names.Id.t option) Loc.located
Sourcetype raw_argument = string * string option

An argument type as provided in Tactic notations, i.e. a string like "ne_foo_list_opt" together with a separator that only makes sense in the "_sep" cases.

A fully resolved argument type given as an AST with generic arguments on the leaves.

Sourceval add_tactic_notation : Vernacexpr.locality_flag -> int -> ?deprecation:Deprecation.t -> raw_argument grammar_tactic_prod_item_expr list -> Tacexpr.raw_tactic_expr -> unit

add_tactic_notation local level prods expr adds a tactic notation in the environment at level level with locality local made of the grammar productions prods and returning the body expr

Sourceval register_tactic_notation_entry : string -> ('a, 'b, 'c) Genarg.genarg_type -> unit

Register an argument under a given entry name for tactic notations. When translating raw_argument into argument, atomic names will be first looked up according to names registered through this function and fallback to finding an argument by name (as in Genarg) if there is none matching.

Sourceval add_ml_tactic_notation : Tacexpr.ml_tactic_name -> level:int -> ?deprecation:Deprecation.t -> argument grammar_tactic_prod_item_expr list list -> unit

A low-level variant of add_tactic_notation used by the TACTIC EXTEND ML-side macro.

Tactic Quotations
Sourceval create_ltac_quotation : string -> ('grm Loc.located -> Tacexpr.raw_tactic_arg) -> ('grm Pcoq.Entry.t * int option) -> unit

create_ltac_quotation name f e adds a quotation rule to Ltac, that is, Ltac grammar now accepts arguments of the form "name" ":" "(" <e> ")", and generates an argument using f on the entry parsed by e.

Queries
Sourceval print_ltacs : unit -> unit

Display the list of ltac definitions currently available.

Sourceval print_located_tactic : Libnames.qualid -> unit

Display the absolute name of a tactic.

Sourceval print_ltac : Libnames.qualid -> Pp.t

Display the definition of a tactic.

Low-level registering of tactics
Sourcetype (_, 'a) ml_ty_sig =
  1. | MLTyNil : ('a, 'a) ml_ty_sig
  2. | MLTyArg : ('r, 'a) ml_ty_sig -> (Geninterp.Val.t -> 'r, 'a) ml_ty_sig
Sourceval ml_tactic_extend : plugin:string -> name:string -> local:Vernacexpr.locality_flag -> ?deprecation:Deprecation.t -> ('r, unit Proofview.tactic) ml_ty_sig -> 'r -> unit

Helper function to define directly an Ltac function in OCaml without any associated parsing rule nor further shenanigans. The Ltac function will be defined as name in the Coq file that loads the ML plugin where this function is called. It will have the arity given by the ml_ty_sig argument.

Sourceval ml_val_tactic_extend : plugin:string -> name:string -> local:Vernacexpr.locality_flag -> ?deprecation:Deprecation.t -> ('r, Geninterp.Val.t Ftactic.t) ml_ty_sig -> 'r -> unit

Same as ml_tactic_extend but the function can return an argument instead.

TACTIC EXTEND
Sourcetype _ ty_sig =
  1. | TyNil : (Geninterp.interp_sign -> unit Proofview.tactic) ty_sig
  2. | TyIdent : string * 'r ty_sig -> 'r ty_sig
  3. | TyArg : ('a, 'b, 'c) Extend.ty_user_symbol * 'r ty_sig -> ('c -> 'r) ty_sig
Sourcetype ty_ml =
  1. | TyML : 'r ty_sig * 'r -> ty_ml
Sourceval tactic_extend : string -> string -> level:Int.t -> ?deprecation:Deprecation.t -> ty_ml list -> unit
Sourceval eval_of_ty_ml : ty_ml -> Geninterp.Val.t list -> Geninterp.interp_sign -> unit Proofview.tactic

grammar rule for add_tactic_notation

ARGUMENT EXTEND

This is the main entry point for the ARGUMENT EXTEND macro that allows to easily create user-made Ltac arguments.

Each argument has three type parameters. See Genarg for more details. There are two kinds of Ltac arguments, uniform and non-uniform. The former have the same type at each level (raw, glob, top) while the latter may vary.

When declaring an argument one must provide the following data:

  • Internalization : raw -> glob
  • Substitution : glob -> glob
  • Interpretation : glob -> Ltac dynamic value
  • Printing for every level
  • An optional toplevel tag of type top (with the proviso that the interpretation function only produces values with this tag)

This data can be either given explicitly with the Fun constructors, or it can be inherited from another argument with the Wit constructors.

Sourcetype ('a, 'b) argument_intern =
  1. | ArgInternFun : ('a, 'b) Genintern.intern_fun -> ('a, 'b) argument_intern
  2. | ArgInternWit : ('a, 'b, 'c) Genarg.genarg_type -> ('a, 'b) argument_intern
Sourcetype 'b argument_subst =
  1. | ArgSubstFun : 'b Genintern.subst_fun -> 'b argument_subst
  2. | ArgSubstWit : ('a, 'b, 'c) Genarg.genarg_type -> 'b argument_subst
Sourcetype ('b, 'c) argument_interp =
  1. | ArgInterpRet : ('c, 'c) argument_interp
  2. | ArgInterpFun : ('b, Geninterp.Val.t) Geninterp.interp_fun -> ('b, 'c) argument_interp
  3. | ArgInterpWit : ('a, 'b, 'r) Genarg.genarg_type -> ('b, 'c) argument_interp
  4. | ArgInterpSimple : (Geninterp.interp_sign -> Environ.env -> Evd.evar_map -> 'b -> 'c) -> ('b, 'c) argument_interp
Sourcetype ('a, 'b, 'c) tactic_argument = {
  1. arg_parsing : 'a Vernacextend.argument_rule;
  2. arg_tag : 'c Geninterp.Val.tag option;
  3. arg_intern : ('a, 'b) argument_intern;
  4. arg_subst : 'b argument_subst;
  5. arg_interp : ('b, 'c) argument_interp;
  6. arg_printer : ('a, 'b, 'c) argument_printer;
}
Sourceval argument_extend : name:string -> ('a, 'b, 'c) tactic_argument -> ('a, 'b, 'c) Genarg.genarg_type * 'a Pcoq.Entry.t