File: genarg.mli

package info (click to toggle)
coq 9.1.0%2Bdfsg-3
  • links: PTS, VCS
  • area: main
  • in suites: experimental
  • size: 35,964 kB
  • sloc: ml: 239,908; sh: 4,355; python: 2,985; ansic: 2,644; makefile: 874; lisp: 171; javascript: 63; xml: 24; sed: 2
file content (286 lines) | stat: -rw-r--r-- 11,944 bytes parent folder | download
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
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
(************************************************************************)
(*         *      The Rocq Prover / The Rocq Development Team           *)
(*  v      *         Copyright INRIA, CNRS and contributors             *)
(* <O___,, * (see version control and CREDITS file for authors & dates) *)
(*   \VV/  **************************************************************)
(*    //   *    This file is distributed under the terms of the         *)
(*         *     GNU Lesser General Public License Version 2.1          *)
(*         *     (see LICENSE file for the text of the license)         *)
(************************************************************************)

(** Generic arguments used by the extension mechanisms of several Rocq ASTs. *)

(** Generic arguments must be registered according to their usage:

    (raw level printers are always useful for clearer [-time] output, for beautify,
    and some other debug prints)

    - extensible constr syntax beyond notations (eg [ltac:()], [ltac2:()] and ltac2 [$x]).
      Such genargs appear in glob_term GGenarg and constrexpr CGenarg.
      They must be registered with [Genintern.register_intern0]
      and [GlobEnv.register_constr_interp0].

      The glob level may be kept through notations and other operations like Ltac definitions
      (eg [Ltac foo := exact ltac2:(foo)]) in which case [Gensubst.register_subst0]
      and a glob level printer are useful.

      Other useful registrations are
      - [Genintern.register_intern_pat] and [Patternops.register_interp_pat]
        to be used in tactic patterns.
      - [Genintern.register_ntn_subst0] to be used in notations
        (eg [Notation "foo" := ltac2:(foo)]).

      NB: only the base [ExtraArg] is allowed here.

    - tactic arguments to commands defined without depending on ltac_plugin
      (VernacProof, HintsExtern, Hint Rewrite, etc).

      Must be registered with [Genintern.register_intern0] and
      [Genintern.register_interp0].

      The glob level can be kept (currently with Hint Extern and Hint
      Rewrite) so [Gensubst.register_subst0] is also needed.

      Currently AFAICT this is just [Tacarg.wit_ltac].

      NB: only the base [ExtraArg] is allowed here.

    - vernac arguments, used by vernac extend. Usually declared in mlg
      using VERNAC ARGUMENT EXTEND then used in VERNAC EXTEND.

      With VERNAC ARGUMENT EXTEND the raw level printer is registered
      by including PRINTED BY.

      Must be registered with [Procq.register_grammar] (handled by
      VERNAC ARGUMENT EXTEND when declared that way) as vernac extend
      only gets the genarg as argument so must get the grammar from
      the registration.

      Unless combined with some other use, the glob and top levels will be empty
      (as in [vernac_genarg_type]).

    - Ltac tactic_extend arguments. Usually declared in mlg using ARGUMENT EXTEND
      then used in TACTIC EXTEND.

      Must be registered with [Genintern.register_intern0],
      [Gensubst.register_subst0] and [Genintern.register_interp0].

      Must be registered with [Procq.register_grammar] as tactic extend
      only gets the genarg as argument so must get the grammar from
      the registration.

      They must be associated with a [Geninterp.Val.tag] using [Geninterp.register_val0]
      (which creates a fresh tag if passed [None]).
      Note: although [Genintern.register_interp0] registers a producer
      of arbitrary [Geninterp.Val.t], tactic_extend requires them to be of the tag
      registered by [Geninterp.register_val0] to work properly.

      They should also have all printer levels registered with [Genprint.register_print0].

      All registrations are handled by the arguments to ARGUMENT EXTEND when declared that way.

      All of them can also be used as vernac_extend arguments since
      vernac_extend uses a subset of the registrations needed for tactic_extend.

    - some hack in Tacentries.ml_val_tactic_extend and its variant in
      Tac2core_ltac1 for Ltac1.lambda.
*)

(** The route of a generic argument, from parsing to evaluation.
In the following diagram, "object" can be ltac_expr, constr, tactic_value, etc.

{% \begin{verbatim} %}
             parsing          in_raw                            out_raw
   char stream ---> raw_object ---> raw_object generic_argument -------+
                          encapsulation                          decaps|
                                                                       |
                                                                       V
                                                                   raw_object
                                                                       |
                                                         globalization |
                                                                       V
                                                                   glob_object
                                                                       |
                                                                encaps |
                                                               in_glob |
                                                                       V
                                                           glob_object generic_argument
                                                                       |
          out                          in                     out_glob |
  object <--- object generic_argument <--- object <--- glob_object <---+
    |   decaps                       encaps      interp           decaps
    |
    V
effective use
{% \end{verbatim} %}

To distinguish between the uninterpreted, globalized and
interpreted worlds, we annotate the type [generic_argument] by a
phantom argument.

*)

(** {5 Generic types} *)

module ArgT :
sig
  type ('a, 'b, 'c) tag
  val eq : ('a1, 'b1, 'c1) tag -> ('a2, 'b2, 'c2) tag -> ('a1 * 'b1 * 'c1, 'a2 * 'b2 * 'c2) CSig.eq option
  val repr : ('a, 'b, 'c) tag -> string
  type any = Any : ('a, 'b, 'c) tag -> any
  val name : string -> any option
  val dump : unit -> (int * string) list
end

(** Generic types. The first parameter is the OCaml lowest level, the second one
    is the globalized level, and third one the internalized level. *)
type (_, _, _) genarg_type =
| ExtraArg : ('a, 'b, 'c) ArgT.tag -> ('a, 'b, 'c) genarg_type
| ListArg : ('a, 'b, 'c) genarg_type -> ('a list, 'b list, 'c list) genarg_type
| OptArg : ('a, 'b, 'c) genarg_type -> ('a option, 'b option, 'c option) genarg_type
| PairArg : ('a1, 'b1, 'c1) genarg_type * ('a2, 'b2, 'c2) genarg_type ->
  ('a1 * 'a2, 'b1 * 'b2, 'c1 * 'c2) genarg_type

type 'a uniform_genarg_type = ('a, 'a, 'a) genarg_type
(** Alias for concision when the three types agree. *)

type 'a vernac_genarg_type = ('a, Util.Empty.t, Util.Empty.t) genarg_type
(** Produced  by VERNAC ARGUMENT EXTEND *)

val make0 : string -> ('raw, 'glob, 'top) genarg_type
(** Create a new generic type of argument: force to associate
    unique ML types at each of the three levels. *)

val create_arg : string -> ('raw, 'glob, 'top) genarg_type
(** Alias for [make0]. *)

(** {5 Specialized types} *)

(** All of [rlevel], [glevel] and [tlevel] must be non convertible
    to ensure the injectivity of the GADT type inference. *)

type rlevel = [ `rlevel ]
type glevel = [ `glevel ]
type tlevel = [ `tlevel ]

(** Generic types at a fixed level. The first parameter embeds the OCaml type
    and the second one the level. *)
type (_, _) abstract_argument_type =
| Rawwit : ('a, 'b, 'c) genarg_type -> ('a, rlevel) abstract_argument_type
| Glbwit : ('a, 'b, 'c) genarg_type -> ('b, glevel) abstract_argument_type
| Topwit : ('a, 'b, 'c) genarg_type -> ('c, tlevel) abstract_argument_type

type 'a raw_abstract_argument_type = ('a, rlevel) abstract_argument_type
(** Specialized type at raw level. *)

type 'a glob_abstract_argument_type = ('a, glevel) abstract_argument_type
(** Specialized type at globalized level. *)

type 'a typed_abstract_argument_type = ('a, tlevel) abstract_argument_type
(** Specialized type at internalized level. *)

(** {6 Projections} *)

val rawwit : ('a, 'b, 'c) genarg_type -> ('a, rlevel) abstract_argument_type
(** Projection on the raw type constructor. *)

val glbwit : ('a, 'b, 'c) genarg_type -> ('b, glevel) abstract_argument_type
(** Projection on the globalized type constructor. *)

val topwit : ('a, 'b, 'c) genarg_type -> ('c, tlevel) abstract_argument_type
(** Projection on the internalized type constructor. *)

(** {5 Generic arguments} *)

type 'l generic_argument = GenArg : ('a, 'l) abstract_argument_type * 'a -> 'l generic_argument
(** A inhabitant of ['level generic_argument] is a inhabitant of some type at
    level ['level], together with the representation of this type. *)

type raw_generic_argument = rlevel generic_argument
type glob_generic_argument = glevel generic_argument
type typed_generic_argument = tlevel generic_argument

(** {6 Constructors} *)

val in_gen : ('a, 'co) abstract_argument_type -> 'a -> 'co generic_argument
(** [in_gen t x] embeds an argument of type [t] into a generic argument. *)

val out_gen : ('a, 'co) abstract_argument_type -> 'co generic_argument -> 'a
(** [out_gen t x] recovers an argument of type [t] from a generic argument. It
    fails if [x] has not the right dynamic type. *)

val has_type : 'co generic_argument -> ('a, 'co) abstract_argument_type -> bool
(** [has_type v t] tells whether [v] has type [t]. If true, it ensures that
    [out_gen t v] will not raise a dynamic type exception. *)

(** {6 Type reification} *)

type argument_type = ArgumentType : ('a, 'b, 'c) genarg_type -> argument_type

(** {6 Equalities} *)

val argument_type_eq : argument_type -> argument_type -> bool
val genarg_type_eq :
  ('a1, 'b1, 'c1) genarg_type ->
  ('a2, 'b2, 'c2) genarg_type ->
  ('a1 * 'b1 * 'c1, 'a2 * 'b2 * 'c2) CSig.eq option
val abstract_argument_type_eq :
  ('a, 'l) abstract_argument_type -> ('b, 'l) abstract_argument_type ->
  ('a, 'b) CSig.eq option

val pr_argument_type : argument_type -> Pp.t
(** Print a human-readable representation for a given type. *)

val genarg_tag : 'a generic_argument -> argument_type

val unquote : ('a, 'co) abstract_argument_type -> argument_type

(** {6 Registering genarg-manipulating functions}

  This is boilerplate code used here and there in the code of Rocq. *)

val get_arg_tag : ('a, 'b, 'c) genarg_type -> ('a, 'b, 'c) ArgT.tag
(** Works only on base objects (ExtraArg), otherwise fails badly. *)

module type GenObj =
sig
  type ('raw, 'glb, 'top) obj
  (** An object manipulating generic arguments. *)

  val name : string
  (** A name for such kind of manipulation, e.g. [interp]. *)

  val default : ('raw, 'glb, 'top) ArgT.tag -> ('raw, 'glb, 'top) obj option
  (** A generic object when there is no registered object for this type. *)
end

module Register (M : GenObj) :
sig
  (** Warning: although the following APIs use [genarg_type] the
      values must always be [ExtraArg some_tag]. *)

  val register0 : ('raw, 'glb, 'top) genarg_type ->
    ('raw, 'glb, 'top) M.obj -> unit
  (** Register a ground type manipulation function. Must be [ExtraArg]. *)

  val obj : ('raw, 'glb, 'top) genarg_type -> ('raw, 'glb, 'top) M.obj
  (** Recover a manipulation function at a given type. Must be [ExtraArg]. *)

  val mem : _ genarg_type -> bool
  (** Is this type registered? (must be [ExtraArg]) *)

  val fold_keys : (ArgT.any -> 'acc -> 'acc) -> 'acc -> 'acc
  (** Fold over the registered keys. *)

end

(** {5 Compatibility layer}

The functions below are aliases for generic_type constructors.

*)

val wit_list : ('a, 'b, 'c) genarg_type -> ('a list, 'b list, 'c list) genarg_type
val wit_opt : ('a, 'b, 'c) genarg_type -> ('a option, 'b option, 'c option) genarg_type
val wit_pair : ('a1, 'b1, 'c1) genarg_type -> ('a2, 'b2, 'c2) genarg_type ->
  ('a1 * 'a2, 'b1 * 'b2, 'c1 * 'c2) genarg_type