diff options
| author | Pierre-Marie Pédrot | 2019-02-05 17:28:05 +0100 |
|---|---|---|
| committer | Pierre-Marie Pédrot | 2019-02-11 15:48:43 +0100 |
| commit | fba77d134e27c086a982e61de5741038bf92bb8a (patch) | |
| tree | 89af4953ee85a55a8ab25ee73b30733d065c8c5d | |
| parent | ccfa2585dbe9a482040f2b2d405fdb4451d19e39 (diff) | |
Make Grammar truly functorial.
The old implementation was simply hiding the fact that the functor relied on
a generic data type. If we want to implement the grammar engine in a truly
type-safe way, we need to make the ancilliary datatypes depend on the type
parameters.
| -rw-r--r-- | gramlib/grammar.ml | 120 |
1 files changed, 62 insertions, 58 deletions
diff --git a/gramlib/grammar.ml b/gramlib/grammar.ml index 1562a275db..8710dc7bb0 100644 --- a/gramlib/grammar.ml +++ b/gramlib/grammar.ml @@ -5,6 +5,66 @@ open Gramext open Format +(* Functorial interface *) + +module type GLexerType = sig type te val lexer : te Plexing.lexer end + +module type S = + sig + type te + type parsable + val parsable : char Stream.t -> parsable + val tokens : string -> (string * int) list + module Entry : + sig + type 'a e + val create : string -> 'a e + val parse : 'a e -> parsable -> 'a + val name : 'a e -> string + val of_parser : string -> (te Stream.t -> 'a) -> 'a e + val parse_token_stream : 'a e -> te Stream.t -> 'a + val print : Format.formatter -> 'a e -> unit + end + type ('self, 'a) ty_symbol + type ('self, 'f, 'r) ty_rule + type 'a ty_production + val s_nterm : 'a Entry.e -> ('self, 'a) ty_symbol + val s_nterml : 'a Entry.e -> string -> ('self, 'a) ty_symbol + val s_list0 : ('self, 'a) ty_symbol -> ('self, 'a list) ty_symbol + val s_list0sep : + ('self, 'a) ty_symbol -> ('self, 'b) ty_symbol -> bool -> + ('self, 'a list) ty_symbol + val s_list1 : ('self, 'a) ty_symbol -> ('self, 'a list) ty_symbol + val s_list1sep : + ('self, 'a) ty_symbol -> ('self, 'b) ty_symbol -> bool -> + ('self, 'a list) ty_symbol + val s_opt : ('self, 'a) ty_symbol -> ('self, 'a option) ty_symbol + val s_self : ('self, 'self) ty_symbol + val s_next : ('self, 'self) ty_symbol + val s_token : Plexing.pattern -> ('self, string) ty_symbol + val s_rules : warning:(string -> unit) option -> 'a ty_production list -> ('self, 'a) ty_symbol + val r_stop : ('self, 'r, 'r) ty_rule + val r_next : + ('self, 'a, 'r) ty_rule -> ('self, 'b) ty_symbol -> + ('self, 'b -> 'a, 'r) ty_rule + val production : ('a, 'f, Loc.t -> 'a) ty_rule * 'f -> 'a ty_production + module Unsafe : + sig + val clear_entry : 'a Entry.e -> unit + end + val safe_extend : warning:(string -> unit) option -> + 'a Entry.e -> Gramext.position option -> + (string option * Gramext.g_assoc option * 'a ty_production list) + list -> + unit + val safe_delete_rule : 'a Entry.e -> ('a, 'r, 'f) ty_rule -> unit + end + +(* Implementation *) + +module GMake (L : GLexerType) = +struct + type 'a parser_t = 'a Stream.t -> Obj.t type 'te grammar = @@ -1188,63 +1248,6 @@ let clear_entry e = Dlevels _ -> e.edesc <- Dlevels [] | Dparser _ -> () -(* Functorial interface *) - -module type GLexerType = sig type te val lexer : te Plexing.lexer end - -module type S = - sig - type te - type parsable - val parsable : char Stream.t -> parsable - val tokens : string -> (string * int) list - module Entry : - sig - type 'a e - val create : string -> 'a e - val parse : 'a e -> parsable -> 'a - val name : 'a e -> string - val of_parser : string -> (te Stream.t -> 'a) -> 'a e - val parse_token_stream : 'a e -> te Stream.t -> 'a - val print : Format.formatter -> 'a e -> unit - end - type ('self, 'a) ty_symbol - type ('self, 'f, 'r) ty_rule - type 'a ty_production - val s_nterm : 'a Entry.e -> ('self, 'a) ty_symbol - val s_nterml : 'a Entry.e -> string -> ('self, 'a) ty_symbol - val s_list0 : ('self, 'a) ty_symbol -> ('self, 'a list) ty_symbol - val s_list0sep : - ('self, 'a) ty_symbol -> ('self, 'b) ty_symbol -> bool -> - ('self, 'a list) ty_symbol - val s_list1 : ('self, 'a) ty_symbol -> ('self, 'a list) ty_symbol - val s_list1sep : - ('self, 'a) ty_symbol -> ('self, 'b) ty_symbol -> bool -> - ('self, 'a list) ty_symbol - val s_opt : ('self, 'a) ty_symbol -> ('self, 'a option) ty_symbol - val s_self : ('self, 'self) ty_symbol - val s_next : ('self, 'self) ty_symbol - val s_token : Plexing.pattern -> ('self, string) ty_symbol - val s_rules : warning:(string -> unit) option -> 'a ty_production list -> ('self, 'a) ty_symbol - val r_stop : ('self, 'r, 'r) ty_rule - val r_next : - ('self, 'a, 'r) ty_rule -> ('self, 'b) ty_symbol -> - ('self, 'b -> 'a, 'r) ty_rule - val production : ('a, 'f, Loc.t -> 'a) ty_rule * 'f -> 'a ty_production - module Unsafe : - sig - val clear_entry : 'a Entry.e -> unit - end - val safe_extend : warning:(string -> unit) option -> - 'a Entry.e -> Gramext.position option -> - (string option * Gramext.g_assoc option * 'a ty_production list) - list -> - unit - val safe_delete_rule : 'a Entry.e -> ('a, 'r, 'f) ty_rule -> unit - end - -module GMake (L : GLexerType) = - struct type te = L.te type parsable = te gen_parsable let gram = gcreate L.lexer @@ -1303,4 +1306,5 @@ module GMake (L : GLexerType) = list) = extend_entry ~warning e pos (Obj.magic r) let safe_delete_rule e r = delete_rule (Entry.obj e) r - end + +end |
