aboutsummaryrefslogtreecommitdiffhomepage
path: root/parsing/egrammar.ml
blob: cec7e4458637497c8e83a5ad2f9e48f2ab7098b8 (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
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
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
(***********************************************************************)
(*  v      *   The Coq Proof Assistant  /  The Coq Development Team    *)
(* <O___,, *        INRIA-Rocquencourt  &  LRI-CNRS-Orsay              *)
(*   \VV/  *************************************************************)
(*    //   *      This file is distributed under the terms of the      *)
(*         *       GNU Lesser General Public License Version 2.1       *)
(***********************************************************************)

(* $Id$ *)

open Pp
open Util
open Extend
open Pcoq
open Topconstr
open Ast
open Genarg
open Libnames

(* State of the grammar extensions *)

type all_grammar_command =
  | Notation of (string * notation * prod_item list)
  | Delimiters of (scope_name * prod_item list * prod_item list)
  | Grammar of grammar_command
  | TacticGrammar of
      (string * (string * grammar_production list) * Tacexpr.raw_tactic_expr)
      list

let subst_all_grammar_command subst = function
  | Notation _ -> anomaly "Notation not in GRAMMAR summary"
  | Delimiters _ -> anomaly "Delimiters not in GRAMMAR summary"
  | Grammar gc -> Grammar (subst_grammar_command subst gc)
  | TacticGrammar g -> TacticGrammar g (* TODO ... *)

let (grammar_state : all_grammar_command list ref) = ref []

(* Interpretation of the right hand side of grammar rules *)

(* When reporting errors, we add the name of the grammar rule that failed *)
let specify_name name e =
  match e with
    | UserError(lab,strm) ->
        UserError(lab, (str"during interpretation of grammar rule " ++
                          str name ++ str"," ++ spc () ++ strm))
    | Anomaly(lab,strm) ->
        Anomaly(lab, (str"during interpretation of grammar rule " ++
                        str name ++ str"," ++ spc () ++ strm))
    | Failure s ->
        Failure("during interpretation of grammar rule "^name^", "^s)
    | e -> e

(* Translation of environments: a production
 *   [ nt1(x1) ... nti(xi) ] -> act(x1..xi)
 * is written (with camlp4 conventions):
 *   (fun vi -> .... (fun v1 -> act(v1 .. vi) )..)
 * where v1..vi are the values generated by non-terminals nt1..nti.
 * Since the actions are executed by substituting an environment,
 * make_act builds the following closure:
 *
 *      ((fun env ->
 *          (fun vi -> 
 *             (fun env -> ...
 *           
 *                  (fun v1 ->
 *                     (fun env -> gram_action .. env act)
 *                     ((x1,v1)::env))
 *                  ...)
 *             ((xi,vi)::env)))
 *         [])
 *)

open Names

let make_act f pil =
  let rec make env = function
    | [] ->
	Gramext.action (fun loc -> f loc env)
    | None :: tl -> (* parse a non-binding item *)
        Gramext.action (fun _ -> make env tl)
    | Some (p, ETConstr) :: tl -> (* non-terminal *)
        Gramext.action (fun (v:constr_expr) -> make ((p,v) :: env) tl)
    | Some (p, ETReference) :: tl -> (* non-terminal *)
        Gramext.action (fun (v:reference) -> make ((p,CRef v) :: env) tl)
    | Some (p, ETIdent) :: tl -> (* non-terminal *)
        Gramext.action (fun (v:identifier) ->
	  make ((p,CRef (Ident (dummy_loc,v))) :: env) tl) in
  make [] (List.rev pil)

let make_cases_pattern_act f pil =
  let rec make env = function
    | [] ->
	Gramext.action (fun loc -> f loc env)
    | None :: tl -> (* parse a non-binding item *)
        Gramext.action (fun _ -> make env tl)
    | Some (p, ETConstr) :: tl -> (* non-terminal *)
        Gramext.action (fun v -> make ((p,v) :: env) tl)
    | Some (p, ETReference) :: tl -> (* non-terminal *)
        Gramext.action (fun v -> make ((p,CPatAtom (dummy_loc,Some v)) :: env) tl)
    | Some (p, ETIdent) :: tl ->
	error "ident entry not admitted in patterns cases syntax extensions" in
  make [] (List.rev pil)

(* Grammar extension command. Rules are assumed correct.
 * Type-checking of grammar rules is done during the translation of
 * ast to the type grammar_command.  We only check that the existing
 * entries have the type assumed in the grammar command (these types
 * annotations are added when type-checking the command, function
 * Extend.of_ast) *)

let get_entry_type (u,n) =
  if u = "constr" & n = "pattern" then Gram.Entry.obj Constr.pattern
  else Gram.Entry.obj (object_of_typed_entry (get_entry (get_univ u) n))

let rec build_prod_item univ = function
  | ProdList0 s -> Gramext.Slist0 (build_prod_item univ s)
  | ProdList1 s -> Gramext.Slist1 (build_prod_item univ s)
  | ProdOpt s   -> Gramext.Sopt   (build_prod_item univ s)
  | ProdPrimitive nt -> 
      let eobj = get_entry_type nt in
      Gramext.Snterm eobj

let symbol_of_prod_item univ = function
  | Term tok -> (Gramext.Stoken tok, None)
  | NonTerm (nt, ovar) ->
      let eobj = build_prod_item univ nt in
      (eobj, ovar)

(*
let make_rule univ etyp rule =
  let pil = List.map (symbol_of_prod_item univ) rule.gr_production in
  let (symbs,ntl) = List.split pil in
  let act = make_act (rule.gr_name,etyp) rule.gr_action ntl in
  (symbs, act)
*)

let make_rule univ etyp rule =
  let pil = List.map (symbol_of_prod_item univ) rule.gr_production in
  let (symbs,ntl) = List.split pil in
  let f loc env = CGrammar (loc, rule.gr_action, env) in
  let act = make_act f ntl in
  (symbs, act)


(* Rules of a level are entered in reverse order, so that the first rules
   are applied before the last ones *)
let extend_entry univ (te, etyp, ass, rls) =
  let rules = List.rev (List.map (make_rule univ etyp) rls) in
  grammar_extend (object_of_typed_entry te) None [(None, ass, rules)]

(* Defines new entries. If the entry already exists, check its type *)
let define_entry univ {ge_name=n; ge_type=typ; gl_assoc=ass; gl_rules=rls} =
  let typ = entry_type_of_constr_entry_type typ in
  let e = force_entry_type univ n typ in
  (e,typ,ass,rls)

(* Add a bunch of grammar rules. Does not check if it is well formed *)
let extend_grammar_rules gram =
  let univ = get_univ gram.gc_univ in
  let tl = List.map (define_entry univ) gram.gc_entries in
  List.iter (extend_entry univ) tl

(* Add a grammar rules for tactics *)
type grammar_tactic_production =
  | TacTerm of string
  | TacNonTerm of loc * (Gram.te Gramext.g_symbol * argument_type) * string option

let make_prod_item = function
  | TacTerm s -> (Gramext.Stoken (Extend.terminal s), None)
  | TacNonTerm (_,(nont,t), po) ->
      (nont, option_app (fun p -> (p,t)) po)

let make_gen_act f pil =
  let rec make env = function
    | [] -> 
	Gramext.action (fun loc -> f loc env)
    | None :: tl -> (* parse a non-binding item *)
        Gramext.action (fun _ -> make env tl)
    | Some (p, t) :: tl -> (* non-terminal *)
        Gramext.action (fun v -> make ((p,in_generic t v) :: env) tl) in
  make [] (List.rev pil)

let extend_constr entry make_act pt =
  let univ = get_univ "constr" in
  let pil = List.map (symbol_of_prod_item univ) pt in
  let (symbs,ntl) = List.split pil in
  let act = make_act ntl in
  grammar_extend entry None [(None, None, [symbs, act])]

let constr_entry name =
  object_of_typed_entry (get_entry (get_univ "constr") name)

let extend_constr_notation (name,ntn,rule) =
  let mkact loc env = CNotation (loc,ntn,env) in
  extend_constr (constr_entry name) (make_act mkact) rule

let extend_constr_grammar (name,c,rule) =
  let mkact loc env = CGrammar (loc,c,env) in
  extend_constr (constr_entry name) (make_act mkact) rule

let extend_constr_delimiters (sc,rule,pat_rule) =
  let mkact loc env = CDelimiters (loc,sc,snd (List.hd env)) in
  extend_constr (constr_entry "constr8") (make_act mkact) rule;
  let mkact loc env = CPatDelimiters (loc,sc,snd (List.hd env)) in
  extend_constr Constr.pattern (make_cases_pattern_act mkact) pat_rule

(* These grammars are not a removable *)
let make_rule univ f g (s,pt) =
  let hd = Gramext.Stoken ("IDENT", s) in
  let pil = (hd,None) :: List.map g pt in
  let (symbs,ntl) = List.split pil in
  let act = make_gen_act f ntl in
  (symbs, act)

let extend_tactic_grammar s gl =
  let univ = get_univ "tactic" in
  let make_act loc l = Tacexpr.TacExtend (loc,s,List.map snd l) in
  let rules = List.map (make_rule univ make_act make_prod_item) gl in
  Gram.extend Tactic.simple_tactic None [(None, None, List.rev rules)]

let extend_vernac_command_grammar s gl =
  let univ = get_univ "vernac" in
  let make_act loc l = Vernacexpr.VernacExtend (s,List.map snd l) in
  let rules = List.map (make_rule univ make_act make_prod_item) gl in
  Gram.extend Vernac_.command None [(None, None, List.rev rules)]

let rec interp_entry_name u s =
  let l = String.length s in
  if l > 8 & String.sub s 0 3 = "ne_" & String.sub s (l-5) 5 = "_list" then
    let t, g = interp_entry_name u (String.sub s 3 (l-8)) in
    List1ArgType t, Gramext.Slist1 g
  else if l > 5 & String.sub s (l-5) 5 = "_list" then
    let t, g = interp_entry_name u (String.sub s 0 (l-5)) in
    List0ArgType t, Gramext.Slist0 g
  else if l > 4 & String.sub s (l-4) 4 = "_opt" then
    let t, g = interp_entry_name u (String.sub s 0 (l-4)) in
    OptArgType t, Gramext.Sopt g
  else
    let n = Extend.rename_command s in
    let e = get_entry (get_univ u) n in
    let o = object_of_typed_entry e in
    let t = type_of_typed_entry e in
    t, Gramext.Snterm (Pcoq.Gram.Entry.obj o)

let qualified_nterm current_univ = function
  | NtQual (univ, en) -> (rename_command univ, rename_command en)
  | NtShort en -> (current_univ, rename_command en)

let make_vprod_item univ = function
  | VTerm s -> (Gramext.Stoken (Extend.terminal s), None)
  | VNonTerm (loc, nt, po) ->
      let (u,nt) = qualified_nterm univ nt in
      let (etyp, e) = interp_entry_name u nt in
      e, option_app (fun p -> (p,etyp)) po

let add_tactic_entries gl =
  let univ = get_univ "tactic" in
  let make_act s tac loc l = Tacexpr.TacAlias (s,l,tac) in
  let f (s,l,tac) =
    make_rule univ (make_act s tac) (make_vprod_item "tactic") l in
  let rules = List.map f gl in
  grammar_extend Tactic.simple_tactic None [(None, None, List.rev rules)]

let extend_grammar gram =
  (match gram with
  | Notation a -> extend_constr_notation a
  | Delimiters a -> extend_constr_delimiters a
  | Grammar g -> extend_grammar_rules g
  | TacticGrammar l -> add_tactic_entries l);
  grammar_state := gram :: !grammar_state


(* Summary functions: the state of the lexer is included in that of the parser.
   Because the grammar affects the set of keywords when adding or removing
   grammar rules. *)
type frozen_t = all_grammar_command list * Lexer.frozen_t

let freeze () = (!grammar_state, Lexer.freeze ())

(* We compare the current state of the grammar and the state to unfreeze, 
   by computing the longest common suffixes *)
let factorize_grams l1 l2 =
  if l1 == l2 then ([], [], l1) else list_share_tails l1 l2

let number_of_entries gcl =
  List.fold_left
    (fun n -> function
      | Notation _ -> n + 1
      | Delimiters _ -> n + 2 (* One rule for constr, one for pattern *)
      | Grammar gc -> n + (List.length gc.gc_entries)
      | TacticGrammar l -> n + 1)
    0 gcl

let unfreeze (grams, lex) =
  let (undo, redo, common) = factorize_grams !grammar_state grams in
  remove_grammars (number_of_entries undo);
  grammar_state := common;
  Lexer.unfreeze lex;
  List.iter extend_grammar (List.rev redo)
    
let init_grammar () =
  remove_grammars (number_of_entries !grammar_state);
  grammar_state := []

let _ = Lexer.init ()

let init () =
  init_grammar ()