Module Ltac2_plugin.Tac2quote

val constructor : ?⁠loc:Loc.t -> Tac2expr.ltac_constructor -> Tac2expr.raw_tacexpr list -> Tac2expr.raw_tacexpr
val thunk : Tac2expr.raw_tacexpr -> Tac2expr.raw_tacexpr
val of_anti : ('a -> Tac2expr.raw_tacexpr) -> 'a Tac2qexpr.or_anti -> Tac2expr.raw_tacexpr
val of_int : int CAst.t -> Tac2expr.raw_tacexpr
val of_pair : ('a -> Tac2expr.raw_tacexpr) -> ('b -> Tac2expr.raw_tacexpr) -> ('a * 'b) CAst.t -> Tac2expr.raw_tacexpr
val of_tuple : ?⁠loc:Loc.t -> Tac2expr.raw_tacexpr list -> Tac2expr.raw_tacexpr
val of_variable : Names.Id.t CAst.t -> Tac2expr.raw_tacexpr
val of_ident : Names.Id.t CAst.t -> Tac2expr.raw_tacexpr
val of_constr : ?⁠delimiters:Names.Id.t list -> Constrexpr.constr_expr -> Tac2expr.raw_tacexpr
val of_open_constr : Constrexpr.constr_expr -> Tac2expr.raw_tacexpr
val of_list : ?⁠loc:Loc.t -> ('a -> Tac2expr.raw_tacexpr) -> 'a list -> Tac2expr.raw_tacexpr
val of_bindings : Tac2qexpr.bindings -> Tac2expr.raw_tacexpr
val of_intro_pattern : Tac2qexpr.intro_pattern -> Tac2expr.raw_tacexpr
val of_intro_patterns : Tac2qexpr.intro_pattern list CAst.t -> Tac2expr.raw_tacexpr
val of_clause : Tac2qexpr.clause -> Tac2expr.raw_tacexpr
val of_destruction_arg : Tac2qexpr.destruction_arg -> Tac2expr.raw_tacexpr
val of_induction_clause : Tac2qexpr.induction_clause -> Tac2expr.raw_tacexpr
val of_conversion : Tac2qexpr.conversion -> Tac2expr.raw_tacexpr
val of_rewriting : Tac2qexpr.rewriting -> Tac2expr.raw_tacexpr
val of_occurrences : Tac2qexpr.occurrences -> Tac2expr.raw_tacexpr
val of_hintdb : Tac2qexpr.hintdb -> Tac2expr.raw_tacexpr
val of_move_location : Tac2qexpr.move_location -> Tac2expr.raw_tacexpr
val of_reference : Tac2qexpr.reference Tac2qexpr.or_anti -> Tac2expr.raw_tacexpr
val of_hyp : ?⁠loc:Loc.t -> Names.Id.t CAst.t -> Tac2expr.raw_tacexpr

id ↦ 'Control.hyp @id'

val of_exact_hyp : ?⁠loc:Loc.t -> Names.Id.t CAst.t -> Tac2expr.raw_tacexpr

id ↦ 'Control.refine (fun () => Control.hyp @id')

val of_exact_var : ?⁠loc:Loc.t -> Names.Id.t CAst.t -> Tac2expr.raw_tacexpr

id ↦ 'Control.refine (fun () => Control.hyp @id')

val of_dispatch : Tac2qexpr.dispatch -> Tac2expr.raw_tacexpr
val of_strategy_flag : Tac2qexpr.strategy_flag -> Tac2expr.raw_tacexpr
val of_pose : Tac2qexpr.pose -> Tac2expr.raw_tacexpr
val of_assertion : Tac2qexpr.assertion -> Tac2expr.raw_tacexpr
val of_constr_matching : Tac2qexpr.constr_matching -> Tac2expr.raw_tacexpr
val of_goal_matching : Tac2qexpr.goal_matching -> Tac2expr.raw_tacexpr
Generic arguments
val wit_pattern : (Constrexpr.constr_exprPattern.constr_pattern) Tac2dyn.Arg.tag
val wit_ident : (Names.Id.tNames.Id.t) Tac2dyn.Arg.tag
val wit_reference : (Tac2qexpr.referenceNames.GlobRef.t) Tac2dyn.Arg.tag
val wit_constr : (Constrexpr.constr_exprGlob_term.glob_constr) Tac2dyn.Arg.tag
val wit_open_constr : (Constrexpr.constr_exprGlob_term.glob_constr) Tac2dyn.Arg.tag
val wit_preterm : (Util.Empty.tGlob_term.glob_constr) Tac2dyn.Arg.tag
val wit_ltac1 : (Names.Id.t CAst.t list * Ltac_plugin.Tacexpr.raw_tactic_exprNames.Id.t list * Ltac_plugin.Tacexpr.glob_tactic_expr) Tac2dyn.Arg.tag

Ltac1 AST quotation, seen as a 'tactic'. Its type is unit in Ltac2.

val wit_ltac1val : (Names.Id.t CAst.t list * Ltac_plugin.Tacexpr.raw_tactic_exprNames.Id.t list * Ltac_plugin.Tacexpr.glob_tactic_expr) Tac2dyn.Arg.tag

Ltac1 AST quotation, seen as a value-returning expression, with type Ltac1.t.