File: cLexer.mli

package info (click to toggle)
coq 8.20.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: sid, trixie
  • size: 44,116 kB
  • sloc: ml: 234,160; sh: 4,301; python: 3,270; ansic: 2,644; makefile: 882; lisp: 172; javascript: 63; xml: 24; sed: 2
file content (99 lines) | stat: -rw-r--r-- 3,467 bytes parent folder | download | duplicates (2)
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
(************************************************************************)
(*         *   The Coq Proof Assistant / The Coq 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)         *)
(************************************************************************)

(** When one registers a keyword she can declare it starts a quotation.
  In particular using QUOTATION("name:") in a grammar rule
  declares "name:" as a keyword and the token QUOTATION is
  matched whenever the keyword is followed by an identifier or a
  parenthesized text. Eg

    constr:x
    string:[....]
    ltac:(....)
    ltac:{....}

  The delimiter is made of 1 or more occurrences of the same parenthesis,
  eg ((.....)) or [[[[....]]]]. The idea being that if the text happens to
  contain the closing delimiter, one can make the delimiter longer and avoid
  confusion (no escaping). Eg

    string:[[ .. ']' .. ]]


  Nesting the delimiter is allowed, eg ((..((...))..)) is OK.

  Keywords don't need to end in ':' *)
type starts_quotation = NoQuotation | Quotation

type keyword_state

val empty_keyword_state : keyword_state

val add_keyword : ?quotation:starts_quotation -> keyword_state -> string -> keyword_state

val is_keyword : keyword_state -> string -> bool
val keywords : keyword_state -> CString.Set.t

val check_ident : string -> unit
val is_ident : string -> bool
val check_keyword : string -> unit

val add_keyword_tok : keyword_state -> 'c Tok.p -> keyword_state

(** When string is not an ident, returns a keyword. *)
val terminal : keyword_state -> string -> string Tok.p

(** Precondition: the input is a number (c.f. [NumTok.t]) *)
val terminal_number : string -> NumTok.Unsigned.t Tok.p

(** [after loc] Will advance a lexing location as the lexer does; this
    can be used to implement parsing resumption from a given position:
{[
  let loc = Pcoq.Parsable.loc pa |> after in
  let str = Gramlib.Stream.of_string text in
  (* Stream.count being correct is critical for Coq's lexer *)
  Gramlib.Stream.njunk loc.ep str;
  let pa = Pcoq.Parsable.make ~loc str in
  (* ready to resume parsing *)
]}
*)
val after : Loc.t -> Loc.t

(** The lexer of Coq: *)

module Lexer :
  Gramlib.Plexing.S
  with type keyword_state = keyword_state
   and type te = Tok.t
   and type 'c pattern = 'c Tok.p


module Error : sig
  type t
  exception E of t
  val to_string : t -> string
end

(** LexerDiff ensures that, ignoring white space, the concatenated
    tokens equal the input string. Specifically:
    - for strings, return the enclosing quotes as tokens and treat the
      quoted value as if it was unquoted, possibly becoming multiple
      tokens.
    - for comments, return the "(\*" (\ to be kind to syntax
      highlighters) as a token and treat the contents of the comment as
      if it was not in a comment, possibly becoming multiple tokens.
    - return any unrecognized Ascii or UTF-8 character as a string.
*)

module LexerDiff :
  Gramlib.Plexing.S
  with type keyword_state = keyword_state
   and type te = Tok.t
   and type 'c pattern = 'c Tok.p