blob: 273cf02c9bab2221e16c047db9253f0f107bf586 (
plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
|
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Names
open Libnames
open Tac2expr
(** {5 Toplevel definitions} *)
val register_ltac : ?local:bool -> ?mut:bool -> rec_flag ->
(Names.lname * raw_tacexpr) list -> unit
val register_type : ?local:bool -> rec_flag ->
(qualid * redef_flag * raw_quant_typedef) list -> unit
val register_primitive : ?local:bool ->
Names.lident -> raw_typexpr -> ml_tactic_name -> unit
val register_struct
: ?local:bool
-> strexpr
-> unit
val register_notation : ?local:bool -> sexpr list -> int option ->
raw_tacexpr -> unit
val perform_eval : pstate:Proof_global.pstate option -> raw_tacexpr -> unit
(** {5 Notations} *)
type scope_rule =
| ScopeRule : (raw_tacexpr, _, 'a) Extend.symbol * ('a -> raw_tacexpr) -> scope_rule
type scope_interpretation = sexpr list -> scope_rule
val register_scope : Id.t -> scope_interpretation -> unit
(** Create a new scope with the provided name *)
val parse_scope : sexpr -> scope_rule
(** Use this to interpret the subscopes for interpretation functions *)
(** {5 Inspecting} *)
val print_ltac : Libnames.qualid -> unit
(** {5 Eval loop} *)
(** Evaluate a tactic expression in the current environment *)
val call : pstate:Proof_global.pstate -> default:bool -> raw_tacexpr -> Proof_global.pstate
(** {5 Toplevel exceptions} *)
val backtrace : backtrace Exninfo.t
(** {5 Parsing entries} *)
module Pltac :
sig
val tac2expr : raw_tacexpr Pcoq.Entry.t
(** Quoted entries. To be used for complex notations. *)
open Tac2qexpr
val q_ident : Id.t CAst.t or_anti Pcoq.Entry.t
val q_bindings : bindings Pcoq.Entry.t
val q_with_bindings : bindings Pcoq.Entry.t
val q_intropattern : intro_pattern Pcoq.Entry.t
val q_intropatterns : intro_pattern list CAst.t Pcoq.Entry.t
val q_destruction_arg : destruction_arg Pcoq.Entry.t
val q_induction_clause : induction_clause Pcoq.Entry.t
val q_conversion : conversion Pcoq.Entry.t
val q_rewriting : rewriting Pcoq.Entry.t
val q_clause : clause Pcoq.Entry.t
val q_dispatch : dispatch Pcoq.Entry.t
val q_occurrences : occurrences Pcoq.Entry.t
val q_reference : reference or_anti Pcoq.Entry.t
val q_strategy_flag : strategy_flag Pcoq.Entry.t
val q_constr_matching : constr_matching Pcoq.Entry.t
val q_goal_matching : goal_matching Pcoq.Entry.t
val q_hintdb : hintdb Pcoq.Entry.t
val q_move_location : move_location Pcoq.Entry.t
val q_pose : pose Pcoq.Entry.t
val q_assert : assertion Pcoq.Entry.t
end
(** {5 Hooks} *)
val register_constr_quotations : (unit -> unit) Hook.t
|