src/Pure/Syntax/parser.ML
changeset 38713 29d0298439be
parent 38712 f7688fd819a8
child 38831 4933a3dfd745
equal deleted inserted replaced
38712:f7688fd819a8 38713:29d0298439be
    54      and also as an entry in "lambdas"*)
    54      and also as an entry in "lambdas"*)
    55 
    55 
    56 
    56 
    57 (*productions for which no starting token is
    57 (*productions for which no starting token is
    58   known yet are associated with this token*)
    58   known yet are associated with this token*)
    59 val UnknownStart = Lexicon.eof;
    59 val unknown_start = Lexicon.eof;
    60 
    60 
    61 (*get all NTs that are connected with a list of NTs*)
    61 (*get all NTs that are connected with a list of NTs*)
    62 fun connected_with _ ([]: nt_tag list) relatives = relatives
    62 fun connected_with _ ([]: nt_tag list) relatives = relatives
    63   | connected_with chains (root :: roots) relatives =
    63   | connected_with chains (root :: roots) relatives =
    64       let val branches = subtract (op =) relatives (these (AList.lookup (op =) chains root));
    64       let val branches = subtract (op =) relatives (these (AList.lookup (op =) chains root));
   197 
   197 
   198                             (*existing productions whose lookahead may depend on l*)
   198                             (*existing productions whose lookahead may depend on l*)
   199                             val tk_prods =
   199                             val tk_prods =
   200                               these
   200                               these
   201                                 (AList.lookup (op =) nt_prods
   201                                 (AList.lookup (op =) nt_prods
   202                                   (SOME (hd l_starts handle Empty => UnknownStart)));
   202                                   (SOME (hd l_starts handle Empty => unknown_start)));
   203 
   203 
   204                             (*add_lambda is true if an existing production of the nt
   204                             (*add_lambda is true if an existing production of the nt
   205                               produces lambda due to the new lambda NT l*)
   205                               produces lambda due to the new lambda NT l*)
   206                             val (add_lambda, nt_dependencies, added_tks, nt_prods') =
   206                             val (add_lambda, nt_dependencies, added_tks, nt_prods') =
   207                               examine_prods tk_prods false [] [] nt_prods;
   207                               examine_prods tk_prods false [] [] nt_prods;
   247                   (if is_some start_tk then [the start_tk] else [],
   247                   (if is_some start_tk then [the start_tk] else [],
   248                     map starts_for_nt start_nts);
   248                     map starts_for_nt start_nts);
   249 
   249 
   250               val opt_starts =
   250               val opt_starts =
   251                (if new_lambda then [NONE]
   251                (if new_lambda then [NONE]
   252                 else if null start_tks then [SOME UnknownStart]
   252                 else if null start_tks then [SOME unknown_start]
   253                 else []) @ map SOME start_tks;
   253                 else []) @ map SOME start_tks;
   254 
   254 
   255               (*add lhs NT to list of dependent NTs in lookahead*)
   255               (*add lhs NT to list of dependent NTs in lookahead*)
   256               fun add_nts [] = ()
   256               fun add_nts [] = ()
   257                 | add_nts (nt :: nts) =
   257                 | add_nts (nt :: nts) =
   316                   let
   316                   let
   317                     (*token under which old productions which
   317                     (*token under which old productions which
   318                       depend on changed_nt could be stored*)
   318                       depend on changed_nt could be stored*)
   319                     val key =
   319                     val key =
   320                       (case find_first (not o member (op =) new_tks) (starts_for_nt changed_nt) of
   320                       (case find_first (not o member (op =) new_tks) (starts_for_nt changed_nt) of
   321                         NONE => SOME UnknownStart
   321                         NONE => SOME unknown_start
   322                       | t => t);
   322                       | t => t);
   323 
   323 
   324                     (*copy productions whose lookahead depends on changed_nt;
   324                     (*copy productions whose lookahead depends on changed_nt;
   325                       if key = SOME UnknownToken then tk_prods is used to hold
   325                       if key = SOME unknown_start then tk_prods is used to hold
   326                       the productions not copied*)
   326                       the productions not copied*)
   327                     fun update_prods [] result = result
   327                     fun update_prods [] result = result
   328                       | update_prods ((p as (rhs, _: string, _: nt_tag)) :: ps)
   328                       | update_prods ((p as (rhs, _: string, _: nt_tag)) :: ps)
   329                             (tk_prods, nt_prods) =
   329                             (tk_prods, nt_prods) =
   330                           let
   330                           let
   357                                    |> AList.update (op =) (SOME tk, tk_prods')
   357                                    |> AList.update (op =) (SOME tk, tk_prods')
   358                                    |> copy tks
   358                                    |> copy tks
   359                                  end;
   359                                  end;
   360                             val result =
   360                             val result =
   361                               if update then (tk_prods, copy new_tks nt_prods)
   361                               if update then (tk_prods, copy new_tks nt_prods)
   362                               else if key = SOME UnknownStart then (p :: tk_prods, nt_prods)
   362                               else if key = SOME unknown_start then (p :: tk_prods, nt_prods)
   363                               else (tk_prods, nt_prods);
   363                               else (tk_prods, nt_prods);
   364                           in update_prods ps result end;
   364                           in update_prods ps result end;
   365 
   365 
   366                     (*copy existing productions for new starting tokens*)
   366                     (*copy existing productions for new starting tokens*)
   367                     fun process_nts [] added = added
   367                     fun process_nts [] added = added
   374                             (*associate productions with new lookahead tokens*)
   374                             (*associate productions with new lookahead tokens*)
   375                             val (tk_prods', nt_prods') = update_prods tk_prods ([], nt_prods);
   375                             val (tk_prods', nt_prods') = update_prods tk_prods ([], nt_prods);
   376 
   376 
   377                             val nt_prods' =
   377                             val nt_prods' =
   378                               nt_prods'
   378                               nt_prods'
   379                               |> (key = SOME UnknownStart) ? AList.update (op =) (key, tk_prods');
   379                               |> (key = SOME unknown_start) ? AList.update (op =) (key, tk_prods');
   380 
   380 
   381                             val added_tks = subtract Lexicon.matching_tokens old_tks new_tks;
   381                             val added_tks = subtract Lexicon.matching_tokens old_tks new_tks;
   382                           in
   382                           in
   383                             if null added_tks then
   383                             if null added_tks then
   384                               (Array.update (prods, nt, (lookahead, nt_prods'));
   384                               (Array.update (prods, nt, (lookahead, nt_prods'));
   638   symb list *       (*rest of rhs*)
   638   symb list *       (*rest of rhs*)
   639   string *          (*name of production*)
   639   string *          (*name of production*)
   640   int;              (*index for previous state list*)
   640   int;              (*index for previous state list*)
   641 
   641 
   642 
   642 
   643 (*Get all rhss with precedence >= minPrec*)
   643 (*Get all rhss with precedence >= min_prec*)
   644 fun getRHS minPrec = filter (fn (_, _, prec:int) => prec >= minPrec);
   644 fun get_RHS min_prec = filter (fn (_, _, prec:int) => prec >= min_prec);
   645 
   645 
   646 (*Get all rhss with precedence >= minPrec and < maxPrec*)
   646 (*Get all rhss with precedence >= min_prec and < max_prec*)
   647 fun getRHS' minPrec maxPrec =
   647 fun get_RHS' min_prec max_prec =
   648   filter (fn (_, _, prec:int) => prec >= minPrec andalso prec < maxPrec);
   648   filter (fn (_, _, prec:int) => prec >= min_prec andalso prec < max_prec);
   649 
   649 
   650 (*Make states using a list of rhss*)
   650 (*Make states using a list of rhss*)
   651 fun mkStates i minPrec lhsID rhss =
   651 fun mk_states i min_prec lhs_ID rhss =
   652   let fun mkState (rhs, id, prodPrec) = (lhsID, prodPrec, [], rhs, id, i);
   652   let fun mk_state (rhs, id, prod_prec) = (lhs_ID, prod_prec, [], rhs, id, i);
   653   in map mkState rhss end;
   653   in map mk_state rhss end;
   654 
   654 
   655 (*Add parse tree to list and eliminate duplicates
   655 (*Add parse tree to list and eliminate duplicates
   656   saving the maximum precedence*)
   656   saving the maximum precedence*)
   657 fun conc (t: parsetree list, prec:int) [] = (NONE, [(t, prec)])
   657 fun conc (t: parsetree list, prec:int) [] = (NONE, [(t, prec)])
   658   | conc (t, prec) ((t', prec') :: ts) =
   658   | conc (t, prec) ((t', prec') :: ts) =
   680       if A = B
   680       if A = B
   681       then used' @ ((A, (prec, ts)) :: used)
   681       then used' @ ((A, (prec, ts)) :: used)
   682       else update (used, hd :: used')
   682       else update (used, hd :: used')
   683   in update (used, []) end;
   683   in update (used, []) end;
   684 
   684 
   685 fun getS A maxPrec Si =
   685 fun getS A max_prec Si =
   686   filter
   686   filter
   687     (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => A = B andalso prec <= maxPrec
   687     (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => A = B andalso prec <= max_prec
   688       | _ => false) Si;
   688       | _ => false) Si;
   689 
   689 
   690 fun getS' A maxPrec minPrec Si =
   690 fun getS' A max_prec min_prec Si =
   691   filter
   691   filter
   692     (fn (_, _, _, Nonterminal (B, prec) :: _, _, _)
   692     (fn (_, _, _, Nonterminal (B, prec) :: _, _, _)
   693           => A = B andalso prec > minPrec andalso prec <= maxPrec
   693           => A = B andalso prec > min_prec andalso prec <= max_prec
   694       | _ => false) Si;
   694       | _ => false) Si;
   695 
   695 
   696 fun getStates Estate i ii A maxPrec =
   696 fun get_states Estate i ii A max_prec =
   697   filter
   697   filter
   698     (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => A = B andalso prec <= maxPrec
   698     (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => A = B andalso prec <= max_prec
   699       | _ => false)
   699       | _ => false)
   700     (Array.sub (Estate, ii));
   700     (Array.sub (Estate, ii));
   701 
   701 
   702 
   702 
   703 fun movedot_term (A, j, ts, Terminal a :: sa, id, i) c =
   703 fun movedot_term (A, j, ts, Terminal a :: sa, id, i) c =
   743     fun all_prods_for nt = prods_for prods chains true c [nt];
   743     fun all_prods_for nt = prods_for prods chains true c [nt];
   744 
   744 
   745     fun processS used [] (Si, Sii) = (Si, Sii)
   745     fun processS used [] (Si, Sii) = (Si, Sii)
   746       | processS used (S :: States) (Si, Sii) =
   746       | processS used (S :: States) (Si, Sii) =
   747           (case S of
   747           (case S of
   748             (_, _, _, Nonterminal (nt, minPrec) :: _, _, _) =>
   748             (_, _, _, Nonterminal (nt, min_prec) :: _, _, _) =>
   749               let (*predictor operation*)
   749               let (*predictor operation*)
   750                 val (used', new_states) =
   750                 val (used', new_states) =
   751                   (case AList.lookup (op =) used nt of
   751                   (case AList.lookup (op =) used nt of
   752                     SOME (usedPrec, l) => (*nonterminal has been processed*)
   752                     SOME (used_prec, l) => (*nonterminal has been processed*)
   753                       if usedPrec <= minPrec then
   753                       if used_prec <= min_prec then
   754                         (*wanted precedence has been processed*)
   754                         (*wanted precedence has been processed*)
   755                         (used, movedot_lambda S l)
   755                         (used, movedot_lambda S l)
   756                       else (*wanted precedence hasn't been parsed yet*)
   756                       else (*wanted precedence hasn't been parsed yet*)
   757                         let
   757                         let
   758                           val tk_prods = all_prods_for nt;
   758                           val tk_prods = all_prods_for nt;
   759 
   759 
   760                           val States' = mkStates i minPrec nt
   760                           val States' = mk_states i min_prec nt
   761                                           (getRHS' minPrec usedPrec tk_prods);
   761                                           (get_RHS' min_prec used_prec tk_prods);
   762                         in (update_prec (nt, minPrec) used,
   762                         in (update_prec (nt, min_prec) used,
   763                             movedot_lambda S l @ States')
   763                             movedot_lambda S l @ States')
   764                         end
   764                         end
   765 
   765 
   766                   | NONE => (*nonterminal is parsed for the first time*)
   766                   | NONE => (*nonterminal is parsed for the first time*)
   767                       let
   767                       let
   768                         val tk_prods = all_prods_for nt;
   768                         val tk_prods = all_prods_for nt;
   769                         val States' = mkStates i minPrec nt (getRHS minPrec tk_prods);
   769                         val States' = mk_states i min_prec nt (get_RHS min_prec tk_prods);
   770                       in ((nt, (minPrec, [])) :: used, States') end);
   770                       in ((nt, (min_prec, [])) :: used, States') end);
   771 
   771 
   772                 val dummy =
   772                 val dummy =
   773                   if not (! warned) andalso
   773                   if not (! warned) andalso
   774                      length (new_states @ States) > ! branching_level then
   774                      length (new_states @ States) > ! branching_level then
   775                     (warning "Currently parsed expression could be extremely ambiguous.";
   775                     (warning "Currently parsed expression could be extremely ambiguous.";
   801                             val Slist = getS' A prec n Si;
   801                             val Slist = getS' A prec n Si;
   802                             val States' = map (movedot_nonterm tt) Slist;
   802                             val States' = map (movedot_nonterm tt) Slist;
   803                           in processS used' (States' @ States) (S :: Si, Sii) end)
   803                           in processS used' (States' @ States) (S :: Si, Sii) end)
   804                   end
   804                   end
   805                 else
   805                 else
   806                   let val Slist = getStates Estate i j A prec
   806                   let val Slist = get_states Estate i j A prec
   807                   in processS used (map (movedot_nonterm tt) Slist @ States) (S :: Si, Sii) end
   807                   in processS used (map (movedot_nonterm tt) Slist @ States) (S :: Si, Sii) end
   808               end)
   808               end)
   809   in processS [] states ([], []) end;
   809   in processS [] states ([], []) end;
   810 
   810 
   811 
   811