| author | wenzelm | 
| Tue, 26 Sep 2023 15:09:31 +0200 | |
| changeset 78723 | 3dc56a11d89e | 
| parent 78009 | f906f7f83dae | 
| child 79079 | 01f9128ec655 | 
| permissions | -rw-r--r-- | 
| 18 | 1 | (* Title: Pure/Syntax/parser.ML | 
| 67545 | 2 | Author: Carsten Clasohm, Sonia Mahjoub | 
| 3 | Author: Makarius | |
| 18 | 4 | |
| 67545 | 5 | General context-free parser for the inner syntax of terms and types. | 
| 18 | 6 | *) | 
| 7 | ||
| 8 | signature PARSER = | |
| 15752 | 9 | sig | 
| 1507 | 10 | type gram | 
| 11 | val empty_gram: gram | |
| 78009 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 12 | val make_gram: Syntax_Ext.xprod list -> Syntax_Ext.xprod list -> gram option -> gram | 
| 1507 | 13 | val pretty_gram: gram -> Pretty.T list | 
| 14 | datatype parsetree = | |
| 15 | Node of string * parsetree list | | |
| 16 | Tip of Lexicon.token | |
| 42374 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 17 | exception PARSETREE of parsetree | 
| 42205 
22f5cc06c419
direct pretty printing of parsetrees -- prevent diagnostic output from crashing due to undeclared entities;
 wenzelm parents: 
41378diff
changeset | 18 | val pretty_parsetree: parsetree -> Pretty.T | 
| 45641 | 19 | val parse: gram -> string -> Lexicon.token list -> parsetree list | 
| 41378 | 20 | val branching_level: int Config.T | 
| 15752 | 21 | end; | 
| 1507 | 22 | |
| 15752 | 23 | structure Parser: PARSER = | 
| 18 | 24 | struct | 
| 15752 | 25 | |
| 18 | 26 | (** datatype gram **) | 
| 27 | ||
| 67541 | 28 | (* nonterminals *) | 
| 29 | ||
| 67545 | 30 | (*production for the NTs are stored in a vector, indexed by the NT tag*) | 
| 67541 | 31 | type nt = int; | 
| 32 | ||
| 67545 | 33 | type tags = nt Symtab.table; | 
| 34 | val tags_empty: tags = Symtab.empty; | |
| 35 | fun tags_content (tags: tags) = sort_by #1 (Symtab.dest tags); | |
| 36 | fun tags_lookup (tags: tags) = Symtab.lookup tags; | |
| 37 | fun tags_insert tag (tags: tags) = Symtab.update_new tag tags; | |
| 78006 | 38 | fun tags_name (tags: tags) = | 
| 39 | the o Inttab.lookup (Inttab.build (Symtab.fold (Inttab.update_new o swap) tags)); | |
| 67545 | 40 | |
| 77723 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 wenzelm parents: 
71468diff
changeset | 41 | type nts = Intset.T; | 
| 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 wenzelm parents: 
71468diff
changeset | 42 | val nts_empty: nts = Intset.empty; | 
| 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 wenzelm parents: 
71468diff
changeset | 43 | val nts_merge: nts * nts -> nts = Intset.merge; | 
| 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 wenzelm parents: 
71468diff
changeset | 44 | fun nts_insert nt : nts -> nts = Intset.insert nt; | 
| 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 wenzelm parents: 
71468diff
changeset | 45 | fun nts_member (nts: nts) = Intset.member nts; | 
| 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 wenzelm parents: 
71468diff
changeset | 46 | fun nts_fold f (nts: nts) = Intset.fold f nts; | 
| 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 wenzelm parents: 
71468diff
changeset | 47 | fun nts_subset (nts1: nts, nts2: nts) = Intset.forall (nts_member nts2) nts1; | 
| 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 wenzelm parents: 
71468diff
changeset | 48 | fun nts_is_empty (nts: nts) = Intset.is_empty nts; | 
| 77741 
1951f6470792
discontinue somewhat pointless is_single, which also depends on details of internal data representation;
 wenzelm parents: 
77723diff
changeset | 49 | fun nts_is_unique (nts: nts) = Intset.size nts <= 1; | 
| 67541 | 50 | |
| 51 | ||
| 67539 | 52 | (* tokens *) | 
| 53 | ||
| 77823 | 54 | structure Tokens = Set(type key = Lexicon.token val ord = Lexicon.tokens_match_ord); | 
| 67539 | 55 | |
| 77823 | 56 | fun tokens_find P tokens = Tokens.get_first (fn tok => if P tok then SOME tok else NONE) tokens; | 
| 57 | fun tokens_add (nt: nt, tokens) = if Tokens.is_empty tokens then I else cons (nt, tokens); | |
| 67533 
f253e5eaf995
clarified types and operations: potentially more efficient add_prods;
 wenzelm parents: 
67532diff
changeset | 58 | |
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 59 | |
| 67545 | 60 | (* productions *) | 
| 61 | ||
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 62 | datatype symb = | 
| 67545 | 63 | Terminal of Lexicon.token | | 
| 64 | Nonterminal of nt * int; (*(tag, prio)*) | |
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 65 | |
| 77823 | 66 | structure Prods = Table(Tokens.Key); | 
| 67 | type prods = (symb list * string * int) list Prods.table; (*start_token ~> [(rhs, name, prio)]*) | |
| 67540 | 68 | |
| 77823 | 69 | val prods_empty: prods = Prods.empty; | 
| 70 | fun prods_lookup (prods: prods) = Prods.lookup_list prods; | |
| 71 | fun prods_update entry : prods -> prods = Prods.update entry; | |
| 72 | fun prods_content (prods: prods) = distinct (op =) (maps #2 (Prods.dest prods)); | |
| 67540 | 73 | |
| 77823 | 74 | type nt_gram = (nts * Tokens.T) * prods; (*dependent_nts, start_tokens, prods*) | 
| 67545 | 75 | (*depent_nts is a set of all NTs whose lookahead depends on this NT's lookahead*) | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 76 | |
| 77823 | 77 | val nt_gram_empty: nt_gram = ((nts_empty, Tokens.empty), prods_empty); | 
| 67533 
f253e5eaf995
clarified types and operations: potentially more efficient add_prods;
 wenzelm parents: 
67532diff
changeset | 78 | |
| 67531 | 79 | type chains = unit Int_Graph.T; | 
| 80 | fun chains_preds (chains: chains) = Int_Graph.immediate_preds chains; | |
| 81 | fun chains_all_preds (chains: chains) = Int_Graph.all_preds chains; | |
| 82 | fun chains_all_succs (chains: chains) = Int_Graph.all_succs chains; | |
| 83 | val chains_empty: chains = Int_Graph.empty; | |
| 84 | fun chains_member (chains: chains) = Int_Graph.is_edge chains; | |
| 85 | fun chains_declare nt : chains -> chains = Int_Graph.default_node (nt, ()); | |
| 86 | fun chains_insert (from, to) = | |
| 87 | chains_declare from #> chains_declare to #> Int_Graph.add_edge (from, to); | |
| 88 | ||
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 89 | datatype gram = | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 90 | Gram of | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 91 |    {nt_count: int,
 | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 92 | prod_count: int, | 
| 67531 | 93 | tags: tags, | 
| 94 | chains: chains, | |
| 67533 
f253e5eaf995
clarified types and operations: potentially more efficient add_prods;
 wenzelm parents: 
67532diff
changeset | 95 | lambdas: nts, | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 96 | prods: nt_gram Vector.vector}; | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 97 | (*"tags" is used to map NT names (i.e. strings) to tags; | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 98 | chain productions are not stored as normal productions | 
| 67517 
add9a9f6a290
explicit graph for chains, which contains all nts as nodes;
 wenzelm parents: 
67516diff
changeset | 99 | but instead as an entry in "chains": from -> to; | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 100 | lambda productions are stored as normal productions | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 101 | and also as an entry in "lambdas"*) | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 102 | |
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 103 | (*productions for which no starting token is | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 104 | known yet are associated with this token*) | 
| 38713 | 105 | val unknown_start = Lexicon.eof; | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 106 | |
| 67539 | 107 | fun get_start tks = | 
| 77823 | 108 | (case Tokens.min tks of | 
| 77818 | 109 | SOME tk => tk | 
| 67539 | 110 | | NONE => unknown_start); | 
| 67516 | 111 | |
| 78007 | 112 | fun add_production array_prods (lhs, new_prod as (rhs, _, pri)) (chains, lambdas) = | 
| 67543 | 113 | let | 
| 114 | (*store chain if it does not already exist*) | |
| 115 | val (chain, new_chain, chains') = | |
| 116 | (case (pri, rhs) of | |
| 117 | (~1, [Nonterminal (from, ~1)]) => | |
| 118 | if chains_member chains (from, lhs) | |
| 119 | then (SOME from, false, chains) | |
| 120 | else (SOME from, true, chains_insert (from, lhs) chains) | |
| 71468 
53fcbede7bf7
more robust (amending add9a9f6a290): proper syntax error instead of exception for grammar with unreachable nonterminals, e.g. nonterminal f1 syntax "_F" :: "f1 ⇒ 'b"  ("F _" 10);
 wenzelm parents: 
69575diff
changeset | 121 | | _ => | 
| 
53fcbede7bf7
more robust (amending add9a9f6a290): proper syntax error instead of exception for grammar with unreachable nonterminals, e.g. nonterminal f1 syntax "_F" :: "f1 ⇒ 'b"  ("F _" 10);
 wenzelm parents: 
69575diff
changeset | 122 | let | 
| 
53fcbede7bf7
more robust (amending add9a9f6a290): proper syntax error instead of exception for grammar with unreachable nonterminals, e.g. nonterminal f1 syntax "_F" :: "f1 ⇒ 'b"  ("F _" 10);
 wenzelm parents: 
69575diff
changeset | 123 | val chains' = chains | 
| 
53fcbede7bf7
more robust (amending add9a9f6a290): proper syntax error instead of exception for grammar with unreachable nonterminals, e.g. nonterminal f1 syntax "_F" :: "f1 ⇒ 'b"  ("F _" 10);
 wenzelm parents: 
69575diff
changeset | 124 | |> chains_declare lhs | 
| 
53fcbede7bf7
more robust (amending add9a9f6a290): proper syntax error instead of exception for grammar with unreachable nonterminals, e.g. nonterminal f1 syntax "_F" :: "f1 ⇒ 'b"  ("F _" 10);
 wenzelm parents: 
69575diff
changeset | 125 | |> fold (fn Nonterminal (nt, _) => chains_declare nt | _ => I) rhs; | 
| 
53fcbede7bf7
more robust (amending add9a9f6a290): proper syntax error instead of exception for grammar with unreachable nonterminals, e.g. nonterminal f1 syntax "_F" :: "f1 ⇒ 'b"  ("F _" 10);
 wenzelm parents: 
69575diff
changeset | 126 | in (NONE, false, chains') end); | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 127 | |
| 67543 | 128 | (*propagate new chain in lookahead and lambdas; | 
| 129 | added_starts is used later to associate existing | |
| 130 | productions with new starting tokens*) | |
| 131 | val (added_starts, lambdas') = | |
| 132 | if not new_chain then ([], lambdas) | |
| 133 | else | |
| 134 | let (*lookahead of chain's source*) | |
| 78007 | 135 | val ((_, from_tks), _) = Array.nth array_prods (the chain); | 
| 1175 
1b15a4b1a4f7
added comments; fixed a bug; reduced memory usage slightly
 clasohm parents: 
1147diff
changeset | 136 | |
| 67543 | 137 | (*copy from's lookahead to chain's destinations*) | 
| 138 | fun copy_lookahead to = | |
| 139 | let | |
| 78007 | 140 | val ((to_nts, to_tks), ps) = Array.nth array_prods to; | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 141 | |
| 77823 | 142 | val new_tks = Tokens.subtract to_tks from_tks; (*added lookahead tokens*) | 
| 143 | val to_tks' = Tokens.merge (to_tks, new_tks); | |
| 78007 | 144 | val _ = Array.upd array_prods to ((to_nts, to_tks'), ps); | 
| 67543 | 145 | in tokens_add (to, new_tks) end; | 
| 1175 
1b15a4b1a4f7
added comments; fixed a bug; reduced memory usage slightly
 clasohm parents: 
1147diff
changeset | 146 | |
| 67543 | 147 | val tos = chains_all_succs chains' [lhs]; | 
| 148 | in | |
| 149 | (fold copy_lookahead tos [], | |
| 150 | lambdas |> nts_member lambdas lhs ? fold nts_insert tos) | |
| 151 | end; | |
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 152 | |
| 67543 | 153 | (*test if new production can produce lambda | 
| 154 | (rhs must either be empty or only consist of lambda NTs)*) | |
| 155 | val new_lambdas = | |
| 156 | if forall | |
| 157 | (fn Nonterminal (id, _) => nts_member lambdas' id | |
| 158 | | Terminal _ => false) rhs | |
| 159 | then SOME (filter_out (nts_member lambdas') (chains_all_succs chains' [lhs])) | |
| 160 | else NONE; | |
| 161 | val lambdas'' = fold nts_insert (these new_lambdas) lambdas'; | |
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 162 | |
| 67543 | 163 | (*list optional terminal and all nonterminals on which the lookahead | 
| 164 | of a production depends*) | |
| 165 | fun lookahead_dependency _ [] nts = (NONE, nts) | |
| 166 | | lookahead_dependency _ (Terminal tk :: _) nts = (SOME tk, nts) | |
| 167 | | lookahead_dependency lambdas (Nonterminal (nt, _) :: symbs) nts = | |
| 168 | if nts_member lambdas nt then | |
| 169 | lookahead_dependency lambdas symbs (nts_insert nt nts) | |
| 170 | else (NONE, nts_insert nt nts); | |
| 330 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
258diff
changeset | 171 | |
| 67543 | 172 | (*get all known starting tokens for a nonterminal*) | 
| 78007 | 173 | fun starts_for_nt nt = snd (fst (Array.nth array_prods nt)); | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 174 | |
| 67543 | 175 | (*update prods, lookaheads, and lambdas according to new lambda NTs*) | 
| 176 | val (added_starts', lambdas') = | |
| 177 | let | |
| 178 | (*propagate added lambda NT*) | |
| 179 | fun propagate_lambda [] added_starts lambdas = (added_starts, lambdas) | |
| 180 | | propagate_lambda (l :: ls) added_starts lambdas = | |
| 181 | let | |
| 182 | (*get lookahead for lambda NT*) | |
| 78007 | 183 | val ((dependent, l_starts), _) = Array.nth array_prods l; | 
| 1175 
1b15a4b1a4f7
added comments; fixed a bug; reduced memory usage slightly
 clasohm parents: 
1147diff
changeset | 184 | |
| 67543 | 185 | (*check productions whose lookahead may depend on lambda NT*) | 
| 186 | fun examine_prods [] add_lambda nt_dependencies added_tks nt_prods = | |
| 187 | (add_lambda, nt_dependencies, added_tks, nt_prods) | |
| 188 | | examine_prods ((p as (rhs, _, _)) :: ps) add_lambda | |
| 189 | nt_dependencies added_tks nt_prods = | |
| 190 | let val (tk, nts) = lookahead_dependency lambdas rhs nts_empty in | |
| 191 | if nts_member nts l then (*update production's lookahead*) | |
| 192 | let | |
| 193 | val new_lambda = | |
| 194 | is_none tk andalso nts_subset (nts, lambdas); | |
| 18 | 195 | |
| 67543 | 196 | val new_tks = | 
| 77823 | 197 | Tokens.empty | 
| 198 | |> fold Tokens.insert (the_list tk) | |
| 199 | |> nts_fold (curry Tokens.merge o starts_for_nt) nts | |
| 200 | |> Tokens.subtract l_starts; | |
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 201 | |
| 77823 | 202 | val added_tks' = Tokens.merge (added_tks, new_tks); | 
| 330 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
258diff
changeset | 203 | |
| 67543 | 204 | val nt_dependencies' = nts_merge (nt_dependencies, nts); | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 205 | |
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 206 | (*associate production with new starting tokens*) | 
| 67539 | 207 | fun copy tk nt_prods = | 
| 67543 | 208 | prods_update (tk, p :: prods_lookup nt_prods tk) nt_prods; | 
| 209 | ||
| 210 | val nt_prods' = nt_prods | |
| 77823 | 211 | |> Tokens.fold copy new_tks | 
| 67550 | 212 | |> new_lambda ? copy Lexicon.dummy; | 
| 67543 | 213 | in | 
| 214 | examine_prods ps (add_lambda orelse new_lambda) | |
| 215 | nt_dependencies' added_tks' nt_prods' | |
| 216 | end | |
| 217 | else (*skip production*) | |
| 218 | examine_prods ps add_lambda nt_dependencies added_tks nt_prods | |
| 219 | end; | |
| 220 | ||
| 221 | (*check each NT whose lookahead depends on new lambda NT*) | |
| 222 | fun process_nts nt (added_lambdas, added_starts) = | |
| 223 | let | |
| 78007 | 224 | val ((old_nts, old_tks), nt_prods) = Array.nth array_prods nt; | 
| 67543 | 225 | |
| 226 | (*existing productions whose lookahead may depend on l*) | |
| 227 | val tk_prods = prods_lookup nt_prods (get_start l_starts); | |
| 228 | ||
| 229 | (*add_lambda is true if an existing production of the nt | |
| 230 | produces lambda due to the new lambda NT l*) | |
| 231 | val (add_lambda, nt_dependencies, added_tks, nt_prods') = | |
| 77823 | 232 | examine_prods tk_prods false nts_empty Tokens.empty nt_prods; | 
| 67543 | 233 | |
| 234 | val new_nts = nts_merge (old_nts, nt_dependencies); | |
| 77823 | 235 | val new_tks = Tokens.merge (old_tks, added_tks); | 
| 67543 | 236 | |
| 237 | val added_lambdas' = added_lambdas |> add_lambda ? cons nt; | |
| 78007 | 238 | val _ = Array.upd array_prods nt ((new_nts, new_tks), nt_prods'); | 
| 67543 | 239 | (*N.B. that because the tks component | 
| 240 | is used to access existing | |
| 241 | productions we have to add new | |
| 242 | tokens at the _end_ of the list*) | |
| 243 | val added_starts' = tokens_add (nt, added_tks) added_starts; | |
| 244 | in (added_lambdas', added_starts') end; | |
| 377 
ab8917806779
lookaheads are now computed faster (during the grammar is built)
 clasohm parents: 
373diff
changeset | 245 | |
| 67543 | 246 | val (added_lambdas, added_starts') = | 
| 247 | nts_fold process_nts dependent ([], added_starts); | |
| 248 | val added_lambdas' = filter_out (nts_member lambdas) added_lambdas; | |
| 249 | in | |
| 250 | propagate_lambda (ls @ added_lambdas') added_starts' | |
| 251 | (fold nts_insert added_lambdas' lambdas) | |
| 252 | end; | |
| 253 | in | |
| 254 | propagate_lambda | |
| 255 | (nts_fold (fn l => not (nts_member lambdas l) ? cons l) lambdas'' []) | |
| 256 | added_starts lambdas'' | |
| 257 | end; | |
| 258 | ||
| 259 | (*insert production into grammar*) | |
| 260 | val added_starts' = | |
| 261 | if is_some chain then added_starts' (*don't store chain production*) | |
| 262 | else | |
| 263 | let | |
| 264 | (*lookahead tokens of new production and on which | |
| 265 | NTs lookahead depends*) | |
| 266 | val (start_tk, start_nts) = lookahead_dependency lambdas' rhs nts_empty; | |
| 267 | ||
| 268 | val start_tks = | |
| 77823 | 269 | Tokens.empty | 
| 270 | |> fold Tokens.insert (the_list start_tk) | |
| 271 | |> nts_fold (curry Tokens.merge o starts_for_nt) start_nts; | |
| 67543 | 272 | |
| 273 | val start_tks' = | |
| 274 | start_tks | |
| 77823 | 275 | |> (if is_some new_lambdas then Tokens.insert Lexicon.dummy | 
| 276 | else if Tokens.is_empty start_tks then Tokens.insert unknown_start | |
| 67543 | 277 | else I); | 
| 278 | ||
| 279 | (*add lhs NT to list of dependent NTs in lookahead*) | |
| 280 | fun add_nts nt initial = | |
| 281 | (if initial then | |
| 78007 | 282 | let val ((old_nts, old_tks), ps) = Array.nth array_prods nt in | 
| 67543 | 283 | if nts_member old_nts lhs then () | 
| 78007 | 284 | else Array.upd array_prods nt ((nts_insert lhs old_nts, old_tks), ps) | 
| 67543 | 285 | end | 
| 286 | else (); false); | |
| 287 | ||
| 288 | (*add new start tokens to chained NTs' lookahead list; | |
| 289 | also store new production for lhs NT*) | |
| 290 | fun add_tks [] added = added | |
| 291 | | add_tks (nt :: nts) added = | |
| 292 | let | |
| 78007 | 293 | val ((old_nts, old_tks), nt_prods) = Array.nth array_prods nt; | 
| 67543 | 294 | |
| 77823 | 295 | val new_tks = Tokens.subtract old_tks start_tks; | 
| 330 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
258diff
changeset | 296 | |
| 67543 | 297 | (*store new production*) | 
| 298 | fun store tk (prods, _) = | |
| 299 | let | |
| 300 | val tk_prods = prods_lookup prods tk; | |
| 301 | val tk_prods' = new_prod :: tk_prods; | |
| 302 | val prods' = prods_update (tk, tk_prods') prods; | |
| 303 | in (prods', true) end; | |
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 304 | |
| 67543 | 305 | val (nt_prods', changed) = (nt_prods, false) | 
| 77823 | 306 | |> nt = lhs ? Tokens.fold store start_tks'; | 
| 67543 | 307 | val _ = | 
| 77823 | 308 | if not changed andalso Tokens.is_empty new_tks then () | 
| 78007 | 309 | else Array.upd array_prods nt ((old_nts, Tokens.merge (old_tks, new_tks)), nt_prods'); | 
| 67543 | 310 | in add_tks nts (tokens_add (nt, new_tks) added) end; | 
| 311 | val _ = nts_fold add_nts start_nts true; | |
| 312 | in add_tks (chains_all_succs chains' [lhs]) [] end; | |
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 313 | |
| 67543 | 314 | (*associate productions with new lookaheads*) | 
| 315 | val _ = | |
| 316 | let | |
| 317 | (*propagate added start tokens*) | |
| 318 | fun add_starts [] = () | |
| 319 | | add_starts ((changed_nt, new_tks) :: starts) = | |
| 320 | let | |
| 321 | (*token under which old productions which | |
| 322 | depend on changed_nt could be stored*) | |
| 323 | val key = | |
| 77823 | 324 | tokens_find (not o Tokens.member new_tks) (starts_for_nt changed_nt) | 
| 67543 | 325 | |> the_default unknown_start; | 
| 326 | ||
| 327 | (*copy productions whose lookahead depends on changed_nt; | |
| 328 | if key = SOME unknown_start then tk_prods is used to hold | |
| 329 | the productions not copied*) | |
| 330 | fun update_prods [] result = result | |
| 331 | | update_prods ((p as (rhs, _: string, _: nt)) :: ps) | |
| 332 | (tk_prods, nt_prods) = | |
| 333 | let | |
| 334 | (*lookahead dependency for production*) | |
| 335 | val (tk, depends) = lookahead_dependency lambdas' rhs nts_empty; | |
| 336 | ||
| 337 | (*test if this production has to be copied*) | |
| 338 | val update = nts_member depends changed_nt; | |
| 18 | 339 | |
| 67543 | 340 | (*test if production could already be associated with | 
| 341 | a member of new_tks*) | |
| 342 | val lambda = | |
| 343 | not (nts_is_unique depends) orelse | |
| 344 | not (nts_is_empty depends) andalso is_some tk | |
| 77823 | 345 | andalso Tokens.member new_tks (the tk); | 
| 67543 | 346 | |
| 347 | (*associate production with new starting tokens*) | |
| 348 | fun copy tk nt_prods = | |
| 349 | let | |
| 350 | val tk_prods = prods_lookup nt_prods tk; | |
| 351 | val tk_prods' = | |
| 352 | if not lambda then p :: tk_prods | |
| 353 | else insert (op =) p tk_prods; | |
| 354 | (*if production depends on lambda NT we | |
| 355 | have to look for duplicates*) | |
| 356 | in prods_update (tk, tk_prods') nt_prods end; | |
| 357 | val result = | |
| 77823 | 358 | if update then (tk_prods, Tokens.fold copy new_tks nt_prods) | 
| 67543 | 359 | else if key = unknown_start then (p :: tk_prods, nt_prods) | 
| 360 | else (tk_prods, nt_prods); | |
| 361 | in update_prods ps result end; | |
| 18 | 362 | |
| 67543 | 363 | (*copy existing productions for new starting tokens*) | 
| 364 | fun process_nts nt = | |
| 365 | let | |
| 78007 | 366 | val ((nts, tks), nt_prods) = Array.nth array_prods nt; | 
| 67543 | 367 | |
| 368 | val tk_prods = prods_lookup nt_prods key; | |
| 369 | ||
| 370 | (*associate productions with new lookahead tokens*) | |
| 371 | val (tk_prods', nt_prods') = update_prods tk_prods ([], nt_prods); | |
| 372 | ||
| 373 | val nt_prods'' = | |
| 374 | if key = unknown_start then | |
| 375 | prods_update (key, tk_prods') nt_prods' | |
| 376 | else nt_prods'; | |
| 377 | ||
| 77823 | 378 | val added_tks = Tokens.subtract tks new_tks; | 
| 379 | val tks' = Tokens.merge (tks, added_tks); | |
| 78007 | 380 | val _ = Array.upd array_prods nt ((nts, tks'), nt_prods''); | 
| 67543 | 381 | in tokens_add (nt, added_tks) end; | 
| 382 | ||
| 78007 | 383 | val ((dependent, _), _) = Array.nth array_prods changed_nt; | 
| 67543 | 384 | in add_starts (starts @ nts_fold process_nts dependent []) end; | 
| 385 | in add_starts added_starts' end; | |
| 386 | in (chains', lambdas') end; | |
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 387 | |
| 18 | 388 | |
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 389 | (* pretty_gram *) | 
| 18 | 390 | |
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 391 | fun pretty_gram (Gram {tags, prods, chains, ...}) =
 | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 392 | let | 
| 67531 | 393 | val print_nt = tags_name tags; | 
| 67518 | 394 |     fun print_pri p = if p < 0 then "" else Symbol.make_sup ("(" ^ string_of_int p ^ ")");
 | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 395 | |
| 67552 | 396 | fun pretty_symb (Terminal tok) = | 
| 397 | if Lexicon.is_literal tok | |
| 398 | then Pretty.quote (Pretty.keyword1 (Lexicon.str_of_token tok)) | |
| 399 | else Pretty.str (Lexicon.str_of_token tok) | |
| 67513 | 400 | | pretty_symb (Nonterminal (tag, p)) = Pretty.str (print_nt tag ^ print_pri p); | 
| 18 | 401 | |
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 402 | fun pretty_const "" = [] | 
| 67513 | 403 |       | pretty_const c = [Pretty.str ("\<^bold>\<Rightarrow> " ^ quote c)];
 | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 404 | |
| 67513 | 405 | fun prod_of_chain from = ([Nonterminal (from, ~1)], "", ~1); | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 406 | |
| 67513 | 407 | fun pretty_prod (name, tag) = | 
| 77846 
5ba68d3bd741
more operations, following Isabelle/ML conventions;
 wenzelm parents: 
77823diff
changeset | 408 | (prods_content (#2 (Vector.nth prods tag)) @ map prod_of_chain (chains_preds chains tag)) | 
| 67513 | 409 | |> map (fn (symbs, const, p) => | 
| 410 | Pretty.block (Pretty.breaks | |
| 411 | (Pretty.str (name ^ print_pri p ^ " =") :: map pretty_symb symbs @ pretty_const const))); | |
| 67531 | 412 | in maps pretty_prod (tags_content tags) end; | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 413 | |
| 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 414 | |
| 42217 | 415 | |
| 67545 | 416 | (** operations on grammars **) | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 417 | |
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 418 | val empty_gram = | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 419 | Gram | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 420 |    {nt_count = 0,
 | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 421 | prod_count = 0, | 
| 67531 | 422 | tags = tags_empty, | 
| 423 | chains = chains_empty, | |
| 67533 
f253e5eaf995
clarified types and operations: potentially more efficient add_prods;
 wenzelm parents: 
67532diff
changeset | 424 | lambdas = nts_empty, | 
| 
f253e5eaf995
clarified types and operations: potentially more efficient add_prods;
 wenzelm parents: 
67532diff
changeset | 425 | prods = Vector.fromList [nt_gram_empty]}; | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 426 | |
| 78009 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 427 | fun extend_gram xprods (Gram {nt_count, prod_count, tags, chains, lambdas, prods}) =
 | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 428 | let | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 429 | (*Get tag for existing nonterminal or create a new one*) | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 430 | fun get_tag nt_count tags nt = | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 431 | (case tags_lookup tags nt of | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 432 | SOME tag => (nt_count, tags, tag) | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 433 | | NONE => (nt_count + 1, tags_insert (nt, nt_count) tags, nt_count)); | 
| 1438 | 434 | |
| 78009 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 435 | (*Convert symbols to the form used by the parser; | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 436 | delimiters and predefined terms are stored as terminals, | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 437 | nonterminals are converted to integer tags*) | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 438 | fun symb_of [] nt_count tags result = (nt_count, tags, rev result) | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 439 | | symb_of (Syntax_Ext.Delim s :: ss) nt_count tags result = | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 440 | symb_of ss nt_count tags (Terminal (Lexicon.literal s) :: result) | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 441 | | symb_of (Syntax_Ext.Argument (s, p) :: ss) nt_count tags result = | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 442 | let | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 443 | val (nt_count', tags', new_symb) = | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 444 | (case Lexicon.get_terminal s of | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 445 | NONE => | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 446 | let val (nt_count', tags', s_tag) = get_tag nt_count tags s; | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 447 | in (nt_count', tags', Nonterminal (s_tag, p)) end | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 448 | | SOME tk => (nt_count, tags, Terminal tk)); | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 449 | in symb_of ss nt_count' tags' (new_symb :: result) end | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 450 | | symb_of (_ :: ss) nt_count tags result = symb_of ss nt_count tags result; | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 451 | |
| 78009 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 452 | (*Convert list of productions by invoking symb_of for each of them*) | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 453 | fun prod_of [] nt_count prod_count tags result = | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 454 | (nt_count, prod_count, tags, result) | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 455 | | prod_of (Syntax_Ext.XProd (lhs, xsymbs, const, pri) :: ps) | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 456 | nt_count prod_count tags result = | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 457 | let | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 458 | val (nt_count', tags', lhs_tag) = get_tag nt_count tags lhs; | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 459 | val (nt_count'', tags'', prods) = symb_of xsymbs nt_count' tags' []; | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 460 | in | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 461 | prod_of ps nt_count'' (prod_count + 1) tags'' | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 462 | ((lhs_tag, (prods, const, pri)) :: result) | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 463 | end; | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 464 | |
| 78009 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 465 | val (nt_count', prod_count', tags', xprods') = | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 466 | prod_of xprods nt_count prod_count tags []; | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 467 | |
| 78009 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 468 | val array_prods' = | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 469 | Array.tabulate (nt_count', fn i => | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 470 | if i < nt_count then Vector.nth prods i | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 471 | else nt_gram_empty); | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 472 | |
| 78009 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 473 | val (chains', lambdas') = | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 474 | (chains, lambdas) |> fold (add_production array_prods') xprods'; | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 475 | in | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 476 | Gram | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 477 |      {nt_count = nt_count',
 | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 478 | prod_count = prod_count', | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 479 | tags = tags', | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 480 | chains = chains', | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 481 | lambdas = lambdas', | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 482 | prods = Array.vector array_prods'} | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 483 | end; | 
| 18 | 484 | |
| 78009 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 485 | fun make_gram [] _ (SOME gram) = gram | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 486 | | make_gram new_xprods _ (SOME gram) = extend_gram new_xprods gram | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 487 | | make_gram [] [] NONE = empty_gram | 
| 
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
 wenzelm parents: 
78007diff
changeset | 488 | | make_gram new_xprods old_xprods NONE = extend_gram (new_xprods @ old_xprods) empty_gram; | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 489 | |
| 18 | 490 | |
| 42217 | 491 | |
| 42374 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 492 | (** parser **) | 
| 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 493 | |
| 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 494 | (* parsetree *) | 
| 18 | 495 | |
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 496 | datatype parsetree = | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 497 | Node of string * parsetree list | | 
| 37683 | 498 | Tip of Lexicon.token; | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 499 | |
| 42374 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 500 | exception PARSETREE of parsetree; | 
| 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 501 | |
| 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 502 | fun pretty_parsetree parsetree = | 
| 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 503 | let | 
| 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 504 | fun pretty (Node (c, pts)) = | 
| 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 505 |           [Pretty.enclose "(" ")" (Pretty.breaks (Pretty.quote (Pretty.str c) :: maps pretty pts))]
 | 
| 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 506 | | pretty (Tip tok) = | 
| 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 507 | if Lexicon.valued_token tok then [Pretty.str (Lexicon.str_of_token tok)] else []; | 
| 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 508 | in (case pretty parsetree of [prt] => prt | _ => raise PARSETREE parsetree) end; | 
| 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 509 | |
| 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 510 | |
| 
b9a6b465da25
clarified pretty_parsetree: suppress literal tokens;
 wenzelm parents: 
42288diff
changeset | 511 | (* parser state *) | 
| 42205 
22f5cc06c419
direct pretty printing of parsetrees -- prevent diagnostic output from crashing due to undeclared entities;
 wenzelm parents: 
41378diff
changeset | 512 | |
| 18 | 513 | type state = | 
| 67533 
f253e5eaf995
clarified types and operations: potentially more efficient add_prods;
 wenzelm parents: 
67532diff
changeset | 514 | nt * int * (*identification and production precedence*) | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 515 | parsetree list * (*already parsed nonterminals on rhs*) | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 516 | symb list * (*rest of rhs*) | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 517 | string * (*name of production*) | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 518 | int; (*index for previous state list*) | 
| 18 | 519 | |
| 520 | ||
| 38713 | 521 | (*Get all rhss with precedence >= min_prec*) | 
| 42217 | 522 | fun get_RHS min_prec = filter (fn (_, _, prec: int) => prec >= min_prec); | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 523 | |
| 38713 | 524 | (*Get all rhss with precedence >= min_prec and < max_prec*) | 
| 525 | fun get_RHS' min_prec max_prec = | |
| 42217 | 526 | filter (fn (_, _, prec: int) => prec >= min_prec andalso prec < max_prec); | 
| 18 | 527 | |
| 330 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
258diff
changeset | 528 | (*Make states using a list of rhss*) | 
| 62669 | 529 | fun mk_states i lhs_ID rhss = | 
| 38713 | 530 | let fun mk_state (rhs, id, prod_prec) = (lhs_ID, prod_prec, [], rhs, id, i); | 
| 531 | in map mk_state rhss end; | |
| 697 
40f72ab196f8
changed warning for extremely ambiguous expressions
 clasohm parents: 
682diff
changeset | 532 | |
| 15752 | 533 | (*Add parse tree to list and eliminate duplicates | 
| 330 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
258diff
changeset | 534 | saving the maximum precedence*) | 
| 42217 | 535 | fun conc (t: parsetree list, prec: int) [] = (NONE, [(t, prec)]) | 
| 330 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
258diff
changeset | 536 | | conc (t, prec) ((t', prec') :: ts) = | 
| 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
258diff
changeset | 537 | if t = t' then | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 538 | (SOME prec', | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 539 | if prec' >= prec then (t', prec') :: ts | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 540 | else (t, prec) :: ts) | 
| 330 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
258diff
changeset | 541 | else | 
| 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
258diff
changeset | 542 | let val (n, ts') = conc (t, prec) ts | 
| 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
258diff
changeset | 543 | in (n, (t', prec') :: ts') end; | 
| 18 | 544 | |
| 330 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
258diff
changeset | 545 | (*Update entry in used*) | 
| 42226 
cb650789f2f0
use standard tables with standard argument order;
 wenzelm parents: 
42222diff
changeset | 546 | fun update_trees (A, t) used = | 
| 
cb650789f2f0
use standard tables with standard argument order;
 wenzelm parents: 
42222diff
changeset | 547 | let | 
| 
cb650789f2f0
use standard tables with standard argument order;
 wenzelm parents: 
42222diff
changeset | 548 | val (i, ts) = the (Inttab.lookup used A); | 
| 
cb650789f2f0
use standard tables with standard argument order;
 wenzelm parents: 
42222diff
changeset | 549 | val (n, ts') = conc t ts; | 
| 
cb650789f2f0
use standard tables with standard argument order;
 wenzelm parents: 
42222diff
changeset | 550 | in (n, Inttab.update (A, (i, ts')) used) end; | 
| 18 | 551 | |
| 330 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
258diff
changeset | 552 | (*Replace entry in used*) | 
| 42226 
cb650789f2f0
use standard tables with standard argument order;
 wenzelm parents: 
42222diff
changeset | 553 | fun update_prec (A, prec) = | 
| 
cb650789f2f0
use standard tables with standard argument order;
 wenzelm parents: 
42222diff
changeset | 554 | Inttab.map_entry A (fn (_, ts) => (prec, ts)); | 
| 18 | 555 | |
| 42220 | 556 | fun getS A max_prec NONE Si = | 
| 557 | filter | |
| 558 | (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => A = B andalso prec <= max_prec | |
| 559 | | _ => false) Si | |
| 560 | | getS A max_prec (SOME min_prec) Si = | |
| 561 | filter | |
| 562 | (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => | |
| 563 | A = B andalso prec > min_prec andalso prec <= max_prec | |
| 564 | | _ => false) Si; | |
| 18 | 565 | |
| 62669 | 566 | fun get_states Estate j A max_prec = | 
| 33317 | 567 | filter | 
| 38713 | 568 | (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => A = B andalso prec <= max_prec | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 569 | | _ => false) | 
| 77888 | 570 | (Array.nth Estate j); | 
| 18 | 571 | |
| 572 | ||
| 42219 | 573 | fun movedot_term c (A, j, ts, Terminal a :: sa, id, i) = | 
| 42282 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42226diff
changeset | 574 | if Lexicon.valued_token c orelse id <> "" | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42226diff
changeset | 575 | then (A, j, Tip c :: ts, sa, id, i) | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 576 | else (A, j, ts, sa, id, i); | 
| 18 | 577 | |
| 42219 | 578 | fun movedot_nonterm tt (A, j, ts, Nonterminal _ :: sa, id, i) = | 
| 42282 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42226diff
changeset | 579 | (A, j, tt @ ts, sa, id, i); | 
| 18 | 580 | |
| 42219 | 581 | fun movedot_lambda [] _ = [] | 
| 42221 | 582 | | movedot_lambda ((t, ki) :: ts) (state as (B, j, tss, Nonterminal (A, k) :: sa, id, i)) = | 
| 42282 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42226diff
changeset | 583 | if k <= ki then (B, j, t @ tss, sa, id, i) :: movedot_lambda ts state | 
| 42221 | 584 | else movedot_lambda ts state; | 
| 18 | 585 | |
| 586 | ||
| 41378 | 587 | (*trigger value for warnings*) | 
| 69575 | 588 | val branching_level = Config.declare_int ("syntax_branching_level", \<^here>) (K 600);
 | 
| 18 | 589 | |
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 590 | (*get all productions of a NT and NTs chained to it which can | 
| 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 591 | be started by specified token*) | 
| 67540 | 592 | fun prods_for (Gram {prods, chains, ...}) tk nt =
 | 
| 37683 | 593 | let | 
| 67540 | 594 | fun token_prods prods = | 
| 595 | fold cons (prods_lookup prods tk) #> | |
| 67550 | 596 | fold cons (prods_lookup prods Lexicon.dummy); | 
| 77846 
5ba68d3bd741
more operations, following Isabelle/ML conventions;
 wenzelm parents: 
77823diff
changeset | 597 | fun nt_prods nt = #2 (Vector.nth prods nt); | 
| 67540 | 598 | in fold (token_prods o nt_prods) (chains_all_preds chains [nt]) [] end; | 
| 1147 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 599 | |
| 
57b5f55bf879
removed 'raw' productions from gram datatype; replaced mk_gram by add_prods;
 clasohm parents: 
890diff
changeset | 600 | |
| 67515 | 601 | fun PROCESSS gram Estate i c states = | 
| 37683 | 602 | let | 
| 62669 | 603 | fun processS _ [] (Si, Sii) = (Si, Sii) | 
| 37683 | 604 | | processS used (S :: States) (Si, Sii) = | 
| 605 | (case S of | |
| 38713 | 606 | (_, _, _, Nonterminal (nt, min_prec) :: _, _, _) => | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 607 | let (*predictor operation*) | 
| 37683 | 608 | val (used', new_states) = | 
| 42226 
cb650789f2f0
use standard tables with standard argument order;
 wenzelm parents: 
42222diff
changeset | 609 | (case Inttab.lookup used nt of | 
| 38713 | 610 | SOME (used_prec, l) => (*nonterminal has been processed*) | 
| 611 | if used_prec <= min_prec then | |
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 612 | (*wanted precedence has been processed*) | 
| 42219 | 613 | (used, movedot_lambda l S) | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 614 | else (*wanted precedence hasn't been parsed yet*) | 
| 37683 | 615 | let | 
| 67540 | 616 | val tk_prods = prods_for gram c nt; | 
| 42217 | 617 | val States' = | 
| 62669 | 618 | mk_states i nt (get_RHS' min_prec used_prec tk_prods); | 
| 42219 | 619 | in (update_prec (nt, min_prec) used, movedot_lambda l S @ States') end | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 620 | | NONE => (*nonterminal is parsed for the first time*) | 
| 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 621 | let | 
| 67540 | 622 | val tk_prods = prods_for gram c nt; | 
| 62669 | 623 | val States' = mk_states i nt (get_RHS min_prec tk_prods); | 
| 42226 
cb650789f2f0
use standard tables with standard argument order;
 wenzelm parents: 
42222diff
changeset | 624 | in (Inttab.update (nt, (min_prec, [])) used, States') end); | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 625 | in | 
| 37683 | 626 | processS used' (new_states @ States) (S :: Si, Sii) | 
| 15752 | 627 | end | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 628 | | (_, _, _, Terminal a :: _, _, _) => (*scanner operation*) | 
| 37683 | 629 | processS used States | 
| 67549 | 630 | (S :: Si, | 
| 631 | if Lexicon.tokens_match_ord (a, c) = EQUAL then movedot_term c S :: Sii else Sii) | |
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 632 | | (A, prec, ts, [], id, j) => (*completer operation*) | 
| 42222 
ff50c436b199
accumulate parsetrees in canonical reverse order;
 wenzelm parents: 
42221diff
changeset | 633 | let val tt = if id = "" then ts else [Node (id, rev ts)] in | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 634 | if j = i then (*lambda production?*) | 
| 37683 | 635 | let | 
| 42226 
cb650789f2f0
use standard tables with standard argument order;
 wenzelm parents: 
42222diff
changeset | 636 | val (prec', used') = update_trees (A, (tt, prec)) used; | 
| 42220 | 637 | val Slist = getS A prec prec' Si; | 
| 638 | val States' = map (movedot_nonterm tt) Slist; | |
| 639 | in processS used' (States' @ States) (S :: Si, Sii) end | |
| 37683 | 640 | else | 
| 62669 | 641 | let val Slist = get_states Estate j A prec | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 642 | in processS used (map (movedot_nonterm tt) Slist @ States) (S :: Si, Sii) end | 
| 37683 | 643 | end) | 
| 42226 
cb650789f2f0
use standard tables with standard argument order;
 wenzelm parents: 
42222diff
changeset | 644 | in processS Inttab.empty states ([], []) end; | 
| 18 | 645 | |
| 646 | ||
| 67515 | 647 | fun produce gram stateset i indata prev_token = | 
| 77888 | 648 | (case Array.nth stateset i of | 
| 25986 
26f1e4c172c3
syntax error: reduced spam -- print expected nonterminals instead of terminals;
 wenzelm parents: 
24245diff
changeset | 649 | [] => | 
| 
26f1e4c172c3
syntax error: reduced spam -- print expected nonterminals instead of terminals;
 wenzelm parents: 
24245diff
changeset | 650 | let | 
| 37683 | 651 | val toks = if Lexicon.is_eof prev_token then indata else prev_token :: indata; | 
| 48992 | 652 | val pos = Position.here (Lexicon.pos_of_token prev_token); | 
| 27801 
0d0adaf0228d
datatype token: maintain range, tuned representation;
 wenzelm parents: 
26678diff
changeset | 653 | in | 
| 55624 | 654 | if null toks then | 
| 655 |           error ("Inner syntax error: unexpected end of input" ^ pos)
 | |
| 656 | else | |
| 657 |           error ("Inner syntax error" ^ pos ^
 | |
| 658 | Markup.markup Markup.no_report | |
| 659 |               ("\n" ^ Pretty.string_of
 | |
| 660 | (Pretty.block [ | |
| 661 | Pretty.str "at", Pretty.brk 1, | |
| 662 | Pretty.block | |
| 663 | (Pretty.str "\"" :: | |
| 664 | Pretty.breaks (map (Pretty.str o Lexicon.str_of_token) (#1 (split_last toks))) @ | |
| 665 | [Pretty.str "\""])]))) | |
| 27801 
0d0adaf0228d
datatype token: maintain range, tuned representation;
 wenzelm parents: 
26678diff
changeset | 666 | end | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 667 | | s => | 
| 38712 
f7688fd819a8
some attempts to recover Isabelle/ML style from the depths of time;
 wenzelm parents: 
38711diff
changeset | 668 | (case indata of | 
| 42218 
8e0a4d500e5b
streamlined token list operations, assuming that the order of union does not matter;
 wenzelm parents: 
42217diff
changeset | 669 | [] => s | 
| 42217 | 670 | | c :: cs => | 
| 671 | let | |
| 67515 | 672 | val (si, sii) = PROCESSS gram stateset i c s; | 
| 77896 | 673 | val _ = Array.upd stateset i si; | 
| 674 | val _ = Array.upd stateset (i + 1) sii; | |
| 67515 | 675 | in produce gram stateset (i + 1) cs c end)); | 
| 18 | 676 | |
| 677 | ||
| 42217 | 678 | fun get_trees states = map_filter (fn (_, _, [pt], _, _, _) => SOME pt | _ => NONE) states; | 
| 18 | 679 | |
| 67515 | 680 | fun earley (gram as Gram {tags, ...}) startsymbol indata =
 | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 681 | let | 
| 37683 | 682 | val start_tag = | 
| 67531 | 683 | (case tags_lookup tags startsymbol of | 
| 37683 | 684 | SOME tag => tag | 
| 40959 | 685 |       | NONE => error ("Inner syntax: bad grammar root symbol " ^ quote startsymbol));
 | 
| 37683 | 686 | val S0 = [(~1, 0, [], [Nonterminal (start_tag, 0), Terminal Lexicon.eof], "", 0)]; | 
| 330 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
258diff
changeset | 687 | val s = length indata + 1; | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 688 | val Estate = Array.array (s, []); | 
| 77896 | 689 | val _ = Array.upd Estate 0 S0; | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 690 | in | 
| 67515 | 691 | get_trees (produce gram Estate 0 indata Lexicon.eof) | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
46diff
changeset | 692 | end; | 
| 18 | 693 | |
| 694 | ||
| 67515 | 695 | fun parse gram start toks = | 
| 27801 
0d0adaf0228d
datatype token: maintain range, tuned representation;
 wenzelm parents: 
26678diff
changeset | 696 | let | 
| 
0d0adaf0228d
datatype token: maintain range, tuned representation;
 wenzelm parents: 
26678diff
changeset | 697 | val end_pos = | 
| 
0d0adaf0228d
datatype token: maintain range, tuned representation;
 wenzelm parents: 
26678diff
changeset | 698 | (case try List.last toks of | 
| 
0d0adaf0228d
datatype token: maintain range, tuned representation;
 wenzelm parents: 
26678diff
changeset | 699 | NONE => Position.none | 
| 67551 | 700 | | SOME tok => Lexicon.end_pos_of_token tok); | 
| 27801 
0d0adaf0228d
datatype token: maintain range, tuned representation;
 wenzelm parents: 
26678diff
changeset | 701 | val r = | 
| 67515 | 702 | (case earley gram start (toks @ [Lexicon.mk_eof end_pos]) of | 
| 37852 
a902f158b4fc
eliminated old-style sys_error/SYS_ERROR in favour of exception Fail -- after careful checking that there is no overlap with existing handling of that;
 wenzelm parents: 
37684diff
changeset | 703 | [] => raise Fail "Inner syntax: no parse trees" | 
| 27801 
0d0adaf0228d
datatype token: maintain range, tuned representation;
 wenzelm parents: 
26678diff
changeset | 704 | | pts => pts); | 
| 
0d0adaf0228d
datatype token: maintain range, tuned representation;
 wenzelm parents: 
26678diff
changeset | 705 | in r end; | 
| 26678 | 706 | |
| 18 | 707 | end; |