aboutsummaryrefslogtreecommitdiffstats
path: root/cparser
diff options
context:
space:
mode:
authorBernhard Schommer <bschommer@users.noreply.github.com>2015-11-12 17:35:41 +0100
committerBernhard Schommer <bschommer@users.noreply.github.com>2015-11-12 17:35:41 +0100
commitd90ba4443294b80bd940daedfdcdc3d4334fdc7c (patch)
treed647ca2216c342c433c35783d32e1abe8545a72c /cparser
parent9054efbd25eedd5627b9e6e62bf1204e5fa0ae94 (diff)
parent0ebefc1d145f82783829174bad1f41bb319742b4 (diff)
downloadcompcert-kvx-d90ba4443294b80bd940daedfdcdc3d4334fdc7c.tar.gz
compcert-kvx-d90ba4443294b80bd940daedfdcdc3d4334fdc7c.zip
Merge pull request #69 from jhjourdan/parser_fix
Parser : duplicate identifier tokens, fix K&R definition parsing
Diffstat (limited to 'cparser')
-rw-r--r--cparser/Cabs.v4
-rw-r--r--cparser/Cabshelper.ml10
-rw-r--r--cparser/Elab.ml188
-rw-r--r--cparser/Lexer.mll79
-rw-r--r--cparser/Parser.vy146
-rw-r--r--cparser/deLexer.ml3
-rw-r--r--cparser/handcrafted.messages3260
-rw-r--r--cparser/pre_parser.mly478
-rw-r--r--cparser/pre_parser_aux.ml19
9 files changed, 1836 insertions, 2351 deletions
diff --git a/cparser/Cabs.v b/cparser/Cabs.v
index e15f8694..f5cab15a 100644
--- a/cparser/Cabs.v
+++ b/cparser/Cabs.v
@@ -81,6 +81,7 @@ with decl_type :=
| PTR : list cvspec -> decl_type -> decl_type
(* The bool is true for variable length parameters. *)
| PROTO : decl_type -> list parameter * bool -> decl_type
+ | PROTO_OLD : decl_type -> list string -> decl_type
with parameter :=
| PARAM : list spec_elem -> option string -> decl_type -> list attribute -> cabsloc -> parameter
@@ -190,8 +191,7 @@ Definition asm_flag := (bool * list char_code)%type.
** Declaration definition (at toplevel)
*)
Inductive definition :=
- | FUNDEF : list spec_elem -> name -> statement -> cabsloc -> definition
- | KRFUNDEF : list spec_elem -> name -> list string -> list definition -> statement -> cabsloc -> definition
+ | FUNDEF : list spec_elem -> name -> list definition -> statement -> cabsloc -> definition
| DECDEF : init_name_group -> cabsloc -> definition (* global variable(s), or function prototype *)
| PRAGMA : string -> cabsloc -> definition
diff --git a/cparser/Cabshelper.ml b/cparser/Cabshelper.ml
index 890679b4..b4e6a082 100644
--- a/cparser/Cabshelper.ml
+++ b/cparser/Cabshelper.ml
@@ -46,8 +46,7 @@ let rec isTypedef = function
let get_definitionloc (d : definition) : cabsloc =
match d with
- | FUNDEF(_, _, _, l) -> l
- | KRFUNDEF(_, _, _, _, _, l) -> l
+ | FUNDEF(_, _, _, _, l) -> l
| DECDEF(_, l) -> l
| PRAGMA(_, l) -> l
@@ -78,10 +77,3 @@ let string_of_cabsloc l =
let format_cabsloc pp l =
Format.fprintf pp "%s:%d" l.filename l.lineno
-
-let rec append_decltype dt1 dt2 =
- match dt1 with
- | JUSTBASE -> dt2
- | ARRAY(dt, attr, sz) -> ARRAY(append_decltype dt dt2, attr, sz)
- | PTR(attr, dt) -> PTR(attr, append_decltype dt dt2)
- | PROTO(dt, params) -> PROTO(append_decltype dt dt2, params)
diff --git a/cparser/Elab.ml b/cparser/Elab.ml
index 0e445b9d..27b650c0 100644
--- a/cparser/Elab.ml
+++ b/cparser/Elab.ml
@@ -518,9 +518,9 @@ and elab_cvspecs env cv_specs =
(* Elaboration of a type declarator. C99 section 6.7.5. *)
-and elab_type_declarator loc env ty = function
+and elab_type_declarator loc env ty kr_ok = function
| Cabs.JUSTBASE ->
- (ty, env)
+ ((ty, None), env)
| Cabs.ARRAY(d, cv_specs, sz) ->
let a = elab_cvspecs env cv_specs in
let sz' =
@@ -536,32 +536,41 @@ and elab_type_declarator loc env ty = function
| None ->
error loc "array size is not a compile-time constant";
Some 1L in (* produces better error messages later *)
- elab_type_declarator loc env (TArray(ty, sz', a)) d
+ elab_type_declarator loc env (TArray(ty, sz', a)) kr_ok d
| Cabs.PTR(cv_specs, d) ->
let a = elab_cvspecs env cv_specs in
- elab_type_declarator loc env (TPtr(ty, a)) d
+ elab_type_declarator loc env (TPtr(ty, a)) kr_ok d
| Cabs.PROTO(d, (params, vararg)) ->
- begin match unroll env ty with
- | TArray _ | TFun _ ->
- error loc "illegal function return type@ %a" Cprint.typ ty
- | _ -> ()
- end;
- let params' = elab_parameters env params in
- elab_type_declarator loc env (TFun(ty, params', vararg, [])) d
+ begin match unroll env ty with
+ | TArray _ | TFun _ ->
+ error loc "Illegal function return type@ %a" Cprint.typ ty
+ | _ -> ()
+ end;
+ let params' = elab_parameters env params in
+ elab_type_declarator loc env (TFun(ty, Some params', vararg, [])) kr_ok d
+ | Cabs.PROTO_OLD(d, params) ->
+ begin match unroll env ty with
+ | TArray _ | TFun _ ->
+ error loc "Illegal function return type@ %a" Cprint.typ ty
+ | _ -> ()
+ end;
+ match params with
+ | [] ->
+ elab_type_declarator loc env (TFun(ty, None, false, [])) kr_ok d
+ | _ ->
+ if not kr_ok || d <> Cabs.JUSTBASE then
+ fatal_error loc "Illegal old-style K&R function definition";
+ ((TFun(ty, None, false, []), Some params), env)
(* Elaboration of parameters in a prototype *)
and elab_parameters env params =
- match params with
- | [] -> (* old-style K&R prototype *)
- None
- | _ ->
- (* Prototype introduces a new scope *)
- let (vars, _) = mmap elab_parameter (Env.new_scope env) params in
- (* Catch special case f(t) where t is void or a typedef to void *)
- match vars with
- | [ ( {name=""}, t) ] when is_void_type env t -> Some []
- | _ -> Some vars
+ (* Prototype introduces a new scope *)
+ let (vars, _) = mmap elab_parameter (Env.new_scope env) params in
+ (* Catch special case f(t) where t is void or a typedef to void *)
+ match vars with
+ | [ ( {name=""}, t) ] when is_void_type env t -> []
+ | _ -> vars
(* Elaboration of a function parameter *)
@@ -569,7 +578,7 @@ and elab_parameter env (PARAM (spec, id, decl, attr, loc)) =
let (sto, inl, tydef, bty, env1) = elab_specifier loc env spec in
if tydef then
error loc "'typedef' used in function parameter";
- let (ty, env2) = elab_type_declarator loc env1 bty decl in
+ let ((ty, _), env2) = elab_type_declarator loc env1 bty false decl in
let ty = add_attributes_type (elab_attributes env attr) ty in
if sto <> Storage_default && sto <> Storage_register then
error loc
@@ -586,13 +595,13 @@ and elab_parameter env (PARAM (spec, id, decl, attr, loc)) =
(* Elaboration of a (specifier, Cabs "name") pair *)
-and elab_name env spec (Name (id, decl, attr, loc)) =
+and elab_fundef_name env spec (Name (id, decl, attr, loc)) =
let (sto, inl, tydef, bty, env') = elab_specifier loc env spec in
if tydef then
error loc "'typedef' is forbidden here";
- let (ty, env'') = elab_type_declarator loc env' bty decl in
+ let ((ty, kr_params), env'') = elab_type_declarator loc env' bty true decl in
let a = elab_attributes env attr in
- (id, sto, inl, add_attributes_type a ty, env'')
+ (id, sto, inl, add_attributes_type a ty, kr_params, env'')
(* Elaboration of a name group. C99 section 6.7.6 *)
@@ -604,8 +613,8 @@ and elab_name_group loc env (spec, namelist) =
if inl then
error loc "'inline' is forbidden here";
let elab_one_name env (Name (id, decl, attr, loc)) =
- let (ty, env1) =
- elab_type_declarator loc env bty decl in
+ let ((ty, _), env1) =
+ elab_type_declarator loc env bty false decl in
let a = elab_attributes env attr in
((id, add_attributes_type a ty), env1) in
(mmap elab_one_name env' namelist, sto)
@@ -616,8 +625,8 @@ and elab_init_name_group loc env (spec, namelist) =
let (sto, inl, tydef, bty, env') =
elab_specifier ~only:(namelist=[]) loc env spec in
let elab_one_name env (Init_name (Name (id, decl, attr, loc), init)) =
- let (ty, env1) =
- elab_type_declarator loc env bty decl in
+ let ((ty, _), env1) =
+ elab_type_declarator loc env bty false decl in
let a = elab_attributes env attr in
if inl && not (is_function_type env ty) then
error loc "'inline' can only appear on functions";
@@ -818,7 +827,7 @@ and elab_enum only loc tag optmembers attrs env =
let elab_type loc env spec decl =
let (sto, inl, tydef, bty, env') = elab_specifier loc env spec in
- let (ty, env'') = elab_type_declarator loc env' bty decl in
+ let ((ty, _), env'') = elab_type_declarator loc env' bty false decl in
if sto <> Storage_default || inl || tydef then
error loc "'typedef', 'extern', 'static', 'register' and 'inline' are meaningless in cast";
(ty, env'')
@@ -1881,16 +1890,66 @@ let enter_decdefs local loc env sto dl =
let (decls, env') = List.fold_left enter_decdef ([], env) dl in
(List.rev decls, env')
-let elab_fundef env spec name body loc =
- let (s, sto, inline, ty, env1) = elab_name env spec name in
+let elab_fundef env spec name defs body loc =
+ let (s, sto, inline, ty, kr_params, env1) = elab_fundef_name env spec name in
if sto = Storage_register then
- fatal_error loc "a function definition cannot have 'register' storage class";
- (* Fix up the type. We can have params = None but only for an
- old-style parameterless function "int f() {...}" *)
- let ty =
- match ty with
- | TFun(ty_ret, None, vararg, attr) -> TFun(ty_ret, Some [], vararg, attr)
- | _ -> ty in
+ fatal_error loc "A function definition cannot have 'register' storage class";
+ begin match kr_params, defs with
+ | None, d::_ ->
+ error (get_definitionloc d)
+ "Old-style parameter declaration in a new-style function definition"
+ | _ -> ()
+ end;
+ let (ty, env1) =
+ match ty, kr_params with
+ | TFun(ty_ret, None, vararg, attr), None ->
+ (TFun(ty_ret, Some [], vararg, attr), env1)
+ | ty, None ->
+ (ty, env1)
+ | TFun(ty_ret, None, false, attr), Some params ->
+ warning loc "Non-prototype, pre-standard function definition.@ Converting to prototype form";
+ (* Check that the parameters have unique names *)
+ List.iter (fun id ->
+ if List.length (List.filter ((=) id) params) > 1 then
+ fatal_error loc "Parameter '%s' appears more than once in function declaration" id)
+ params;
+ (* Check that the declarations only declare parameters *)
+ let extract_name (Init_name(Name(s, dty, attrs, loc') as name, ie)) =
+ if not (List.mem s params) then
+ error loc' "Declaration of '%s' which is not a function parameter" s;
+ if ie <> NO_INIT then
+ error loc' "Illegal initialization of function parameter '%s'" s;
+ name
+ in
+ (* Convert old-style K&R function definition to modern prototyped form *)
+ let elab_kr_param_def env = function
+ | DECDEF((spec', name_init_list), loc') ->
+ let name_list = List.map extract_name name_init_list in
+ let (paramsenv, sto) = elab_name_group loc' env (spec', name_list) in
+ if sto <> Storage_default && sto <> Storage_register then
+ error loc' "'extern' or 'static' storage not supported for function parameter";
+ paramsenv
+ | d ->
+ (* Should never be produced by the parser *)
+ fatal_error (get_definitionloc d)
+ "Illegal declaration of function parameter" in
+ let (kr_params_defs, env1) = mmap elab_kr_param_def env1 defs in
+ let kr_params_defs = List.concat kr_params_defs in
+ let rec search_param_type param =
+ match List.filter (fun (p, _) -> p = param) kr_params_defs with
+ | [] ->
+ (* Parameter is not declared, defaults to "int" in ISO C90,
+ is an error in ISO C99. Just emit a warning. *)
+ warning loc "Type of '%s' defaults to 'int'" param;
+ (Env.fresh_ident param, TInt (IInt, []))
+ | (_,ty)::q ->
+ if q <> [] then error loc "Parameter '%s' defined more than once" param;
+ (Env.fresh_ident param, argument_conversion env1 ty)
+ in
+ let params' = List.map search_param_type params in
+ (TFun(ty_ret, Some params', false, attr), env1)
+ | _, Some params -> assert false
+ in
(* Extract info from type *)
let (ty_ret, params, vararg, attr) =
match ty with
@@ -1938,61 +1997,14 @@ let elab_fundef env spec name body loc =
emit_elab env1 loc (Gfundef fn);
env1
-let elab_kr_fundef env spec name params defs body loc =
- warning loc "Non-prototype, pre-standard function definition.@ Converting to prototype form";
- (* Check that the declarations only declare parameters *)
- let check_one_decl (Init_name(Name(s, dty, attrs, loc'), ie)) =
- if not (List.mem s params) then
- error loc' "Declaration of '%s' which is not a function parameter" s;
- if ie <> NO_INIT then
- error loc' "Illegal initialization of function parameter '%s'" s in
- let check_decl = function
- | DECDEF((spec', name_init_list), loc') ->
- List.iter check_one_decl name_init_list
- | d ->
- (* Should never be produced by the parser *)
- fatal_error (get_definitionloc d)
- "Illegal declaration of function parameter" in
- List.iter check_decl defs;
- (* Convert old-style K&R function definition to modern prototyped form *)
- let rec convert_param param = function
- | [] ->
- (* Parameter is not declared, defaults to "int" in ISO C90,
- is an error in ISO C99. Just emit a warning. *)
- warning loc "Type of '%s' defaults to 'int'" param;
- PARAM([SpecType Tint], Some param, JUSTBASE, [], loc)
- | DECDEF((spec', name_init_list), loc') :: defs ->
- let rec convert = function
- | [] -> convert_param param defs
- | Init_name(Name(s, dty, attrs, loc''), ie) :: l ->
- if s = param
- then PARAM(spec', Some param, dty, attrs, loc'')
- else convert l
- in convert name_init_list
- | _ ->
- assert false (* checked earlier *) in
- let params' =
- List.map (fun p -> convert_param p defs) params in
- let name' =
- let (Name(s, dty, attr, loc')) = name in
- Name(s, append_decltype dty (PROTO(JUSTBASE, (params', false))),
- attr, loc') in
- (* Elaborate the prototyped form *)
- elab_fundef env spec name' body loc
-
let rec elab_definition (local: bool) (env: Env.t) (def: Cabs.definition)
: decl list * Env.t =
match def with
(* "int f(int x) { ... }" *)
- | FUNDEF(spec, name, body, loc) ->
- if local then error loc "local definition of a function";
- let env1 = elab_fundef env spec name body loc in
- ([], env1)
-
(* "int f(x, y) double y; { ... }" *)
- | KRFUNDEF(spec, name, params, defs, body, loc) ->
+ | FUNDEF(spec, name, defs, body, loc) ->
if local then error loc "local definition of a function";
- let env1 = elab_kr_fundef env spec name params defs body loc in
+ let env1 = elab_fundef env spec name defs body loc in
([], env1)
(* "int x = 12, y[10], *z" *)
diff --git a/cparser/Lexer.mll b/cparser/Lexer.mll
index 17c4528c..6fac15e8 100644
--- a/cparser/Lexer.mll
+++ b/cparser/Lexer.mll
@@ -20,13 +20,12 @@ open Pre_parser_aux
open Cabshelper
open Camlcoq
-module SMap = Map.Make(String)
+module SSet = Set.Make(String)
-let contexts_stk : (Cabs.cabsloc -> token) SMap.t list ref = ref []
+let lexicon : (string, Cabs.cabsloc -> token) Hashtbl.t = Hashtbl.create 17
-let init_ctx =
- List.fold_left (fun ctx (key, builder) -> SMap.add key builder ctx)
- SMap.empty
+let () =
+ List.iter (fun (key, builder) -> Hashtbl.add lexicon key builder)
[ ("_Alignas", fun loc -> ALIGNAS loc);
("_Alignof", fun loc -> ALIGNOF loc);
("_Bool", fun loc -> UNDERSCORE_BOOL loc);
@@ -82,38 +81,24 @@ let init_ctx =
("unsigned", fun loc -> UNSIGNED loc);
("void", fun loc -> VOID loc);
("volatile", fun loc -> VOLATILE loc);
- ("while", fun loc -> WHILE loc);
- (let id = "__builtin_va_list" in
- id, fun loc -> TYPEDEF_NAME (id, ref TypedefId, loc))]
+ ("while", fun loc -> WHILE loc)]
+
+let init_ctx = SSet.singleton "__builtin_va_list"
+let types_context : SSet.t ref = ref init_ctx
let _ =
(* See comments in pre_parser_aux.ml *)
- open_context := begin fun () ->
- contexts_stk := List.hd !contexts_stk::!contexts_stk
- end;
-
- close_context := begin fun () ->
- contexts_stk := List.tl !contexts_stk
- end;
-
- save_contexts_stk := begin fun () ->
- let save = !contexts_stk in
- fun () -> contexts_stk := save
+ save_context := begin fun () ->
+ let save = !types_context in
+ fun () -> types_context := save
end;
declare_varname := begin fun id ->
- match !contexts_stk with
- (* This is the default, so there is no need to have an entry in this case. *)
- | ctx::stk -> contexts_stk := SMap.remove id ctx::stk
- | [] -> assert false
+ types_context := SSet.remove id !types_context
end;
declare_typename := begin fun id ->
- match !contexts_stk with
- | ctx::stk ->
- contexts_stk :=
- SMap.add id (fun loc -> TYPEDEF_NAME (id, ref TypedefId, loc)) ctx::stk
- | [] -> assert false
+ types_context := SSet.add id !types_context
end
let init filename channel : Lexing.lexbuf =
@@ -340,8 +325,8 @@ rule initial = parse
| "," { COMMA(currentLoc lexbuf) }
| "." { DOT(currentLoc lexbuf) }
| identifier as id {
- try SMap.find id (List.hd !contexts_stk) (currentLoc lexbuf)
- with Not_found -> VAR_NAME (id, ref VarId, currentLoc lexbuf) }
+ try Hashtbl.find lexicon id (currentLoc lexbuf)
+ with Not_found -> PRE_NAME id }
| eof { EOF }
| _ as c { fatal_error lexbuf "invalid symbol %C" c }
@@ -452,16 +437,33 @@ and singleline_comment = parse
(* [lexer tokens buffer] is a new lexer, which wraps [lexer], and also: 1-
records the token stream into the FIFO queue [tokens] and 2- records the
start and end positions of the last two tokens in the two-place buffer
- [buffer]. *)
+ [buffer] and 3- duplicates identifier tokens into PRE_NAME and
+ VAR/TYPE_NAME. *)
let lexer tokens buffer : lexbuf -> Pre_parser.token =
+ let curr_id = ref None in
+ types_context := init_ctx;
fun lexbuf ->
- let token = lexer lexbuf in
- Queue.push token tokens;
- let startp = lexbuf.lex_start_p
- and endp = lexbuf.lex_curr_p in
- buffer := ErrorReports.update !buffer (startp, endp);
- token
+ match !curr_id with
+ | Some id ->
+ curr_id := None;
+ let loc = currentLoc lexbuf in
+ let token =
+ if SSet.mem id !types_context then TYPEDEF_NAME (id, ref TypedefId, loc)
+ else VAR_NAME (id, ref VarId, loc)
+ in
+ Queue.push token tokens;
+ token
+ | None ->
+ let token = lexer lexbuf in
+ begin match token with
+ | PRE_NAME id -> curr_id := Some id
+ | _ -> Queue.push token tokens
+ end;
+ let startp = lexbuf.lex_start_p
+ and endp = lexbuf.lex_curr_p in
+ buffer := ErrorReports.update !buffer (startp, endp);
+ token
(* [invoke_pre_parser] is in charge of calling the pre_parser. It uses
the incremental API, which allows us to do our own error handling. *)
@@ -482,11 +484,9 @@ and singleline_comment = parse
of (appropriately classified) tokens. *)
let tokens_stream filename text : token coq_Stream =
- contexts_stk := [init_ctx];
let tokens = Queue.create () in
let buffer = ref ErrorReports.Zero in
invoke_pre_parser filename text (lexer tokens buffer) buffer;
- assert (List.length !contexts_stk = 1);
let rec compute_token_stream () =
let loop t v =
Cons (Coq_existT (t, Obj.magic v), Lazy.from_fun compute_token_stream)
@@ -606,6 +606,7 @@ and singleline_comment = parse
| ATTRIBUTE loc -> loop ATTRIBUTE't loc
| ASM loc -> loop ASM't loc
| PRAGMA (s, loc) -> loop PRAGMA't (s, loc)
+ | PRE_NAME _ -> assert false
in
Lazy.from_fun compute_token_stream
diff --git a/cparser/Parser.vy b/cparser/Parser.vy
index 7c0bfb55..16f6a0ef 100644
--- a/cparser/Parser.vy
+++ b/cparser/Parser.vy
@@ -52,6 +52,7 @@ Require Import List.
%type<list expression (* Reverse order *)> argument_expression_list
%type<definition> declaration
%type<list spec_elem * cabsloc> declaration_specifiers
+%type<list spec_elem> declaration_specifiers_typespec_opt
%type<list init_name (* Reverse order *)> init_declarator_list
%type<init_name> init_declarator
%type<storage * cabsloc> storage_class_specifier
@@ -65,9 +66,9 @@ Require Import List.
%type<list (string * option expression * cabsloc) (* Reverse order *)> enumerator_list
%type<string * option expression * cabsloc> enumerator
%type<string * cabsloc> enumeration_constant
-%type<cvspec * cabsloc> type_qualifier
+%type<cvspec * cabsloc> type_qualifier type_qualifier_noattr
%type<cabsloc> function_specifier
-%type<name> declarator direct_declarator
+%type<name> declarator declarator_noattrend direct_declarator
%type<(decl_type -> decl_type) * cabsloc> pointer
%type<list cvspec (* Reverse order *)> type_qualifier_list
%type<list parameter * bool> parameter_type_list
@@ -95,7 +96,6 @@ Require Import List.
%type<gcc_attribute> gcc_attribute
%type<list gcc_attribute> gcc_attribute_list
%type<gcc_attribute_word> gcc_attribute_word
-%type<name * list string> old_function_declarator direct_old_function_declarator
%type<list string (* Reverse order *)> identifier_list
%type<list asm_flag> asm_flags
%type<option string> asm_op_name
@@ -337,23 +337,34 @@ declaration:
| decspec = declaration_specifiers SEMICOLON
{ DECDEF (fst decspec, []) (snd decspec) }
+declaration_specifiers_typespec_opt:
+| storage = storage_class_specifier rest = declaration_specifiers_typespec_opt
+ { SpecStorage (fst storage)::rest }
+| typ = type_specifier rest = declaration_specifiers_typespec_opt
+ { SpecType (fst typ)::rest }
+| qual = type_qualifier rest = declaration_specifiers_typespec_opt
+ { SpecCV (fst qual)::rest }
+| loc = function_specifier rest = declaration_specifiers_typespec_opt
+ { SpecInline::rest }
+| /* empty */
+ { [] }
+
+
+(* We impose a lighter constraint on declaration specifiers than in the
+ pre_parser: declaration specifiers should have at least one type
+ specifier. *)
declaration_specifiers:
| storage = storage_class_specifier rest = declaration_specifiers
{ (SpecStorage (fst storage)::fst rest, snd storage) }
-| storage = storage_class_specifier
- { ([SpecStorage (fst storage)], snd storage) }
-| typ = type_specifier rest = declaration_specifiers
- { (SpecType (fst typ)::fst rest, snd typ) }
-| typ = type_specifier
- { ([SpecType (fst typ)], snd typ) }
-| qual = type_qualifier rest = declaration_specifiers
+| typ = type_specifier rest = declaration_specifiers_typespec_opt
+ { (SpecType (fst typ)::rest, snd typ) }
+(* We have to inline type_qualifier in order to avoid a conflict. *)
+| qual = type_qualifier_noattr rest = declaration_specifiers
{ (SpecCV (fst qual)::fst rest, snd qual) }
-| qual = type_qualifier
- { ([SpecCV (fst qual)], snd qual) }
+| attr = attribute_specifier rest = declaration_specifiers
+ { (SpecCV (CV_ATTR (fst attr))::fst rest, snd attr) }
| loc = function_specifier rest = declaration_specifiers
{ (SpecInline::fst rest, loc) }
-| loc = function_specifier
- { ([SpecInline], loc) }
init_declarator_list:
| init = init_declarator
@@ -411,12 +422,17 @@ type_specifier:
(* 6.7.2.1 *)
struct_or_union_specifier:
-| str_uni = struct_or_union attrs = attribute_specifier_list id = OTHER_NAME LBRACE decls = struct_declaration_list RBRACE
- { (Tstruct_union (fst str_uni) (Some (fst id)) (Some (rev' decls)) (rev' attrs), snd str_uni) }
-| str_uni = struct_or_union attrs = attribute_specifier_list LBRACE decls = struct_declaration_list RBRACE
- { (Tstruct_union (fst str_uni) None (Some (rev' decls)) (rev' attrs), snd str_uni) }
| str_uni = struct_or_union attrs = attribute_specifier_list id = OTHER_NAME
- { (Tstruct_union (fst str_uni) (Some (fst id)) None (rev' attrs), snd str_uni) }
+ LBRACE decls = struct_declaration_list RBRACE
+ { (Tstruct_union (fst str_uni) (Some (fst id)) (Some (rev' decls)) attrs,
+ snd str_uni) }
+| str_uni = struct_or_union attrs = attribute_specifier_list
+ LBRACE decls = struct_declaration_list RBRACE
+ { (Tstruct_union (fst str_uni) None (Some (rev' decls)) attrs,
+ snd str_uni) }
+| str_uni = struct_or_union attrs = attribute_specifier_list id = OTHER_NAME
+ { (Tstruct_union (fst str_uni) (Some (fst id)) None attrs,
+ snd str_uni) }
struct_or_union:
| loc = STRUCT
@@ -463,16 +479,20 @@ struct_declarator:
(* 6.7.2.2 *)
enum_specifier:
-| loc = ENUM attrs = attribute_specifier_list name = OTHER_NAME LBRACE enum_list = enumerator_list RBRACE
- { (Tenum (Some (fst name)) (Some (rev' enum_list)) (rev' attrs), loc) }
-| loc = ENUM attrs = attribute_specifier_list LBRACE enum_list = enumerator_list RBRACE
- { (Tenum None (Some (rev' enum_list)) (rev' attrs), loc) }
-| loc = ENUM attrs = attribute_specifier_list name = OTHER_NAME LBRACE enum_list = enumerator_list COMMA RBRACE
- { (Tenum (Some (fst name)) (Some (rev' enum_list)) (rev' attrs), loc) }
-| loc = ENUM attrs = attribute_specifier_list LBRACE enum_list = enumerator_list COMMA RBRACE
- { (Tenum None (Some (rev' enum_list)) (rev' attrs), loc) }
| loc = ENUM attrs = attribute_specifier_list name = OTHER_NAME
- { (Tenum (Some (fst name)) None (rev' attrs), loc) }
+ LBRACE enum_list = enumerator_list RBRACE
+ { (Tenum (Some (fst name)) (Some (rev' enum_list)) attrs, loc) }
+| loc = ENUM attrs = attribute_specifier_list
+ LBRACE enum_list = enumerator_list RBRACE
+ { (Tenum None (Some (rev' enum_list)) attrs, loc) }
+| loc = ENUM attrs = attribute_specifier_list name = OTHER_NAME
+ LBRACE enum_list = enumerator_list COMMA RBRACE
+ { (Tenum (Some (fst name)) (Some (rev' enum_list)) attrs, loc) }
+| loc = ENUM attrs = attribute_specifier_list
+ LBRACE enum_list = enumerator_list COMMA RBRACE
+ { (Tenum None (Some (rev' enum_list)) attrs, loc) }
+| loc = ENUM attrs = attribute_specifier_list name = OTHER_NAME
+ { (Tenum (Some (fst name)) None attrs, loc) }
enumerator_list:
| enum = enumerator
@@ -491,13 +511,17 @@ enumeration_constant:
{ cst }
(* 6.7.3 *)
-type_qualifier:
+type_qualifier_noattr:
| loc = CONST
{ (CV_CONST, loc) }
| loc = RESTRICT
{ (CV_RESTRICT, loc) }
| loc = VOLATILE
{ (CV_VOLATILE, loc) }
+
+type_qualifier:
+| qual = type_qualifier_noattr
+ { qual }
(* Non-standard *)
| attr = attribute_specifier
{ (CV_ATTR (fst attr), snd attr) }
@@ -507,17 +531,14 @@ type_qualifier:
attribute_specifier_list:
| /* empty */
{ [] }
-| attrs = attribute_specifier_list attr = attribute_specifier
- { fst attr :: attrs }
+| attr = attribute_specifier attrs = attribute_specifier_list
+ { fst attr :: attrs }
attribute_specifier:
| loc = ATTRIBUTE LPAREN LPAREN attr = gcc_attribute_list RPAREN RPAREN
{ (GCC_ATTR (rev' attr) loc, loc) }
| loc = PACKED LPAREN args = argument_expression_list RPAREN
{ (PACKED_ATTR (rev' args) loc, loc) }
-/* TODO: slove conflict */
-/*| loc = PACKED
- { (PACKED_ATTR [] loc, loc) }*/
| loc = ALIGNAS LPAREN args = argument_expression_list RPAREN
{ (ALIGNAS_ATTR (rev' args) loc, loc) }
| loc = ALIGNAS LPAREN typ = type_name RPAREN
@@ -554,12 +575,16 @@ function_specifier:
(* 6.7.5 *)
declarator:
-| decl = direct_declarator attrs = attribute_specifier_list
+| decl = declarator_noattrend attrs = attribute_specifier_list
{ match decl with Name name typ attr loc =>
Name name typ (List.app attr attrs) loc end }
-| pt = pointer decl = direct_declarator attrs = attribute_specifier_list
+
+declarator_noattrend:
+| decl = direct_declarator
+ { decl }
+| pt = pointer decl = direct_declarator
{ match decl with Name name typ attr _ =>
- Name name ((fst pt) typ) (List.app attr attrs) (snd pt) end }
+ Name name ((fst pt) typ) attr (snd pt) end }
direct_declarator:
| id = VAR_NAME
@@ -585,7 +610,10 @@ direct_declarator:
Name name (PROTO typ params) attr loc end }
| decl = direct_declarator LPAREN RPAREN
{ match decl with Name name typ attr loc =>
- Name name (PROTO typ ([], false)) attr loc end }
+ Name name (PROTO_OLD typ []) attr loc end }
+| decl = direct_declarator LPAREN params = identifier_list RPAREN
+ { match decl with Name name typ attr loc =>
+ Name name (PROTO_OLD typ (rev' params)) attr loc end }
pointer:
| loc = STAR
@@ -624,6 +652,12 @@ parameter_declaration:
| specs = declaration_specifiers
{ PARAM (fst specs) None JUSTBASE [] (snd specs) }
+identifier_list:
+| id = VAR_NAME
+ { [fst id] }
+| idl = identifier_list COMMA id = VAR_NAME
+ { fst id :: idl }
+
(* 6.7.6 *)
type_name:
| specqual = specifier_qualifier_list
@@ -902,40 +936,18 @@ external_declaration:
(* 6.9.1 *)
function_definition:
| specs = declaration_specifiers
- decl = declarator
+ decl = declarator_noattrend
+ dlist = declaration_list
stmt = compound_statement
- { FUNDEF (fst specs) decl stmt (snd specs) }
+ { FUNDEF (fst specs) decl (List.rev' dlist) stmt (snd specs) }
| specs = declaration_specifiers
- decl = old_function_declarator
- dlist = declaration_list
+ decl = declarator
stmt = compound_statement
- { KRFUNDEF (fst specs) (fst decl) (snd decl) (List.rev' dlist) stmt (snd specs) }
-
-(* Hack to parse old-style, unprototyped function definitions
- without causing ambiguities between a regular declarator
- (which can end with attributes) and a declaration list
- (which can start with attributes). *)
-
-old_function_declarator:
-| decl = direct_old_function_declarator
- { decl }
-| pt = pointer decl = direct_old_function_declarator
- { match decl with (Name name typ attr _, params) =>
- (Name name ((fst pt) typ) attr (snd pt), params) end }
-
-direct_old_function_declarator:
-| decl = direct_declarator LPAREN params = identifier_list RPAREN
- { (decl, List.rev' params) }
-
-identifier_list:
-| id = VAR_NAME
- { [fst id] }
-| idl = identifier_list COMMA id = VAR_NAME
- { fst id :: idl }
+ { FUNDEF (fst specs) decl [] stmt (snd specs) }
declaration_list:
-| /*empty*/
- { [] }
+| d = declaration
+ { [d] }
| dl = declaration_list d = declaration
{ d :: dl }
diff --git a/cparser/deLexer.ml b/cparser/deLexer.ml
index 00308e4b..3aa168da 100644
--- a/cparser/deLexer.ml
+++ b/cparser/deLexer.ml
@@ -65,6 +65,7 @@ let delex (symbol : string) : string =
| "WHILE" -> "while"
| "TYPEDEF_NAME" -> "t" (* this should be a type name *)
| "VAR_NAME" -> "x" (* this should be a variable name *)
+ | "PRE_NAME" -> ""
| "CONSTANT" -> "42"
| "STRING_LITERAL" -> "\"\""
| "ELLIPSIS" -> "..."
@@ -121,6 +122,8 @@ let delex (symbol : string) : string =
let delex sentence =
let symbols = Str.split (Str.regexp " ") sentence in
+ if List.nth symbols (List.length symbols - 1) = "PRE_NAME" then
+ failwith "token sequence terminating with PRE_NAME";
let symbols = List.map delex symbols in
List.iter (fun symbol ->
Printf.printf "%s " symbol
diff --git a/cparser/handcrafted.messages b/cparser/handcrafted.messages
index 01416449..2528a42e 100644
--- a/cparser/handcrafted.messages
+++ b/cparser/handcrafted.messages
@@ -3,6 +3,7 @@
# The Compcert verified compiler #
# #
# François Pottier, INRIA Paris-Rocquencourt #
+# Jacques-Henri Jourdan, INRIA Paris-Rocquencourt #
# #
# Copyright Institut National de Recherche en Informatique et en #
# Automatique. All rights reserved. This file is distributed #
@@ -117,8 +118,8 @@
# a VAR_NAME which could be either the beginning of an expression or the
# beginning of a labeled statement:
-# general_identifier -> VAR_NAME . [ COLON ]
-# primary_expression -> VAR_NAME . [ <lots of tokens> ]
+# general_identifier -> PRE_NAME VAR_NAME . [ COLON ]
+# primary_expression -> PRE_NAME VAR_NAME . [ <lots of tokens> ]
# In the strict interpretation, no extra reduction is permitted here, because
# two reductions are enabled. In the lax interpretation, because we have
@@ -175,11 +176,11 @@
# ------------------------------------------------------------------------------
-translation_unit_file: ALIGNAS LPAREN TYPEDEF_NAME INT
+translation_unit_file: ALIGNAS LPAREN INT XOR_ASSIGN
##
-## Ends in an error in state: 177.
+## Ends in an error in state: 299.
##
-## attribute_specifier -> ALIGNAS LPAREN type_name . RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> ALIGNAS LPAREN type_name . RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ALIGNAS LPAREN type_name
@@ -188,9 +189,9 @@ translation_unit_file: ALIGNAS LPAREN TYPEDEF_NAME INT
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 185, spurious reduction of production specifier_qualifier_list(type_name) -> option(type_qualifier_list) TYPEDEF_NAME option(type_qualifier_list)
-## In state 148, spurious reduction of production option(abstract_declarator(type_name)) ->
-## In state 324, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+## In state 152, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 301, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 307, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
##
# Maybe the type name was not complete, but we have reduced anyway
@@ -202,15 +203,15 @@ translation_unit_file: ALIGNAS LPAREN TYPEDEF_NAME INT
Ill-formed _Alignas qualifier.
Up to this point, a type name has been recognized:
- $2 $1 $0
+ $0
If this type name is complete,
then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ ALIGNOF LPAREN VOID XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ ALIGNOF LPAREN VOID XOR_ASSIGN
##
-## Ends in an error in state: 122.
+## Ends in an error in state: 311.
##
## postfix_expression -> LPAREN type_name . RPAREN LBRACE initializer_list option(COMMA) RBRACE [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
## unary_expression -> ALIGNOF LPAREN type_name . RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LEQ LEFT_ASSIGN LEFT HAT GT GEQ EQEQ EQ DIV_ASSIGN COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
@@ -222,13 +223,13 @@ translation_unit_file: INT VAR_NAME EQ ALIGNOF LPAREN VOID XOR_ASSIGN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 214, spurious reduction of production specifier_qualifier_list(type_name) -> option(type_qualifier_list) type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
-## In state 148, spurious reduction of production option(abstract_declarator(type_name)) ->
-## In state 324, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+## In state 152, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 301, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 307, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
##
-translation_unit_file: INT VAR_NAME EQ SIZEOF LPAREN VOID XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ SIZEOF LPAREN VOID XOR_ASSIGN
##
-## Ends in an error in state: 340.
+## Ends in an error in state: 378.
##
## postfix_expression -> LPAREN type_name . RPAREN LBRACE initializer_list option(COMMA) RBRACE [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
## unary_expression -> SIZEOF LPAREN type_name . RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LEQ LEFT_ASSIGN LEFT HAT GT GEQ EQEQ EQ DIV_ASSIGN COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
@@ -240,9 +241,9 @@ translation_unit_file: INT VAR_NAME EQ SIZEOF LPAREN VOID XOR_ASSIGN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 214, spurious reduction of production specifier_qualifier_list(type_name) -> option(type_qualifier_list) type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
-## In state 148, spurious reduction of production option(abstract_declarator(type_name)) ->
-## In state 324, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+## In state 152, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 301, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 307, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
##
Ill-formed use of $2.
@@ -253,9 +254,9 @@ then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ BUILTIN_VA_ARG LPAREN VAR_NAME COMMA VOID XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ BUILTIN_VA_ARG LPAREN PRE_NAME VAR_NAME COMMA VOID XOR_ASSIGN
##
-## Ends in an error in state: 331.
+## Ends in an error in state: 342.
##
## postfix_expression -> BUILTIN_VA_ARG LPAREN assignment_expression COMMA type_name . RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
@@ -266,22 +267,22 @@ translation_unit_file: INT VAR_NAME EQ BUILTIN_VA_ARG LPAREN VAR_NAME COMMA VOID
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 214, spurious reduction of production specifier_qualifier_list(type_name) -> option(type_qualifier_list) type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
-## In state 148, spurious reduction of production option(abstract_declarator(type_name)) ->
-## In state 324, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+## In state 152, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 301, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 307, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
##
Ill-formed use of __builtin_va_arg.
Up to this point, a type name has been recognized:
- $4 $3 $2 $1 $0
+ $0
If this type name is complete,
then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ INC LPAREN VOID XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ INC LPAREN VOID XOR_ASSIGN
##
-## Ends in an error in state: 334.
+## Ends in an error in state: 372.
##
## postfix_expression -> LPAREN type_name . RPAREN LBRACE initializer_list option(COMMA) RBRACE [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
@@ -292,9 +293,9 @@ translation_unit_file: INT VAR_NAME EQ INC LPAREN VOID XOR_ASSIGN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 214, spurious reduction of production specifier_qualifier_list(type_name) -> option(type_qualifier_list) type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
-## In state 148, spurious reduction of production option(abstract_declarator(type_name)) ->
-## In state 324, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+## In state 152, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 301, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 307, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
##
# gcc simply says it expects a closing parenthesis,
@@ -302,15 +303,15 @@ translation_unit_file: INT VAR_NAME EQ INC LPAREN VOID XOR_ASSIGN
Ill-formed compound literal.
Up to this point, a type name has been recognized:
- $1 $0
+ $0
If this type name is complete,
then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LPAREN VOID XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LPAREN VOID XOR_ASSIGN
##
-## Ends in an error in state: 337.
+## Ends in an error in state: 375.
##
## cast_expression -> LPAREN type_name . RPAREN cast_expression [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LEQ LEFT_ASSIGN LEFT HAT GT GEQ EQEQ EQ DIV_ASSIGN COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
## postfix_expression -> LPAREN type_name . RPAREN LBRACE initializer_list option(COMMA) RBRACE [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
@@ -322,26 +323,26 @@ translation_unit_file: INT VAR_NAME EQ LPAREN VOID XOR_ASSIGN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 214, spurious reduction of production specifier_qualifier_list(type_name) -> option(type_qualifier_list) type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
-## In state 148, spurious reduction of production option(abstract_declarator(type_name)) ->
-## In state 324, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+## In state 152, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 301, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 307, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
##
# gcc and clang say they expect a closing parenthesis.
Up to this point, a type name has been recognized:
- $1 $0
+ $0
If this type name is complete,
then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: ALIGNAS LPAREN VAR_NAME SEMICOLON
+translation_unit_file: ALIGNAS LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 318.
+## Ends in an error in state: 309.
##
## argument_expression_list -> argument_expression_list . COMMA assignment_expression [ RPAREN COMMA ]
-## attribute_specifier -> ALIGNAS LPAREN argument_expression_list . RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> ALIGNAS LPAREN argument_expression_list . RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ALIGNAS LPAREN argument_expression_list
@@ -350,21 +351,21 @@ translation_unit_file: ALIGNAS LPAREN VAR_NAME SEMICOLON
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 109, spurious reduction of production argument_expression_list -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 133, spurious reduction of production argument_expression_list -> assignment_expression
##
# We are trying to recognize an alignas specifier.
@@ -388,20 +389,22 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: ALIGNAS LPAREN INT LBRACK RPAREN
##
-## Ends in an error in state: 152.
+## Ends in an error in state: 235.
##
## direct_abstract_declarator -> option(direct_abstract_declarator) LBRACK option(type_qualifier_list) . optional(assignment_expression,RBRACK) [ RPAREN LPAREN LBRACK COMMA ]
-## type_qualifier_list -> option(type_qualifier_list) . type_qualifier [ VOLATILE VAR_NAME TILDE STRING_LITERAL STAR SIZEOF RESTRICT RBRACK PLUS PACKED MINUS LPAREN INC DEC CONSTANT CONST BUILTIN_VA_ARG BANG ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## type_qualifier_list -> option(type_qualifier_list) . type_qualifier_noattr [ VOLATILE TILDE STRING_LITERAL STAR SIZEOF RESTRICT RBRACK PRE_NAME PLUS PACKED MINUS LPAREN INC DEC CONSTANT CONST BUILTIN_VA_ARG BANG ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## type_qualifier_list -> option(type_qualifier_list) . attribute_specifier [ VOLATILE TILDE STRING_LITERAL STAR SIZEOF RESTRICT RBRACK PRE_NAME PLUS PACKED MINUS LPAREN INC DEC CONSTANT CONST BUILTIN_VA_ARG BANG ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## option(direct_abstract_declarator) LBRACK option(type_qualifier_list)
##
-translation_unit_file: INT VAR_NAME LBRACK RPAREN
+translation_unit_file: INT PRE_NAME VAR_NAME LBRACK RPAREN
##
-## Ends in an error in state: 277.
+## Ends in an error in state: 252.
##
-## direct_declarator -> direct_declarator LBRACK option(type_qualifier_list) . optional(assignment_expression,RBRACK) [ SEMICOLON RPAREN PACKED LPAREN LBRACK LBRACE EQ COMMA COLON ATTRIBUTE ALIGNAS ]
-## type_qualifier_list -> option(type_qualifier_list) . type_qualifier [ VOLATILE VAR_NAME TILDE STRING_LITERAL STAR SIZEOF RESTRICT RBRACK PLUS PACKED MINUS LPAREN INC DEC CONSTANT CONST BUILTIN_VA_ARG BANG ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## direct_declarator -> direct_declarator LBRACK option(type_qualifier_list) . optional(assignment_expression,RBRACK) [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
+## type_qualifier_list -> option(type_qualifier_list) . type_qualifier_noattr [ VOLATILE TILDE STRING_LITERAL STAR SIZEOF RESTRICT RBRACK PRE_NAME PLUS PACKED MINUS LPAREN INC DEC CONSTANT CONST BUILTIN_VA_ARG BANG ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## type_qualifier_list -> option(type_qualifier_list) . attribute_specifier [ VOLATILE TILDE STRING_LITERAL STAR SIZEOF RESTRICT RBRACK PRE_NAME PLUS PACKED MINUS LPAREN INC DEC CONSTANT CONST BUILTIN_VA_ARG BANG ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## direct_declarator LBRACK option(type_qualifier_list)
@@ -429,32 +432,32 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: ALIGNAS LPAREN INT LPAREN RPAREN LPAREN INT COMMA ELLIPSIS XOR_ASSIGN
+translation_unit_file: ALIGNAS LPAREN INT LPAREN INT COMMA ELLIPSIS XOR_ASSIGN
##
-## Ends in an error in state: 271.
+## Ends in an error in state: 263.
##
-## direct_abstract_declarator -> direct_abstract_declarator LPAREN in_context(option(parameter_type_list)) . RPAREN [ RPAREN LPAREN LBRACK COMMA ]
+## direct_abstract_declarator -> LPAREN option(context_parameter_type_list) . RPAREN [ RPAREN LPAREN LBRACK COMMA ]
##
## The known suffix of the stack is as follows:
-## direct_abstract_declarator LPAREN in_context(option(parameter_type_list))
+## LPAREN option(context_parameter_type_list)
##
-translation_unit_file: ALIGNAS LPAREN INT LPAREN INT COMMA ELLIPSIS XOR_ASSIGN
+translation_unit_file: ALIGNAS LPAREN INT LBRACK RBRACK LPAREN INT COMMA ELLIPSIS XOR_ASSIGN
##
-## Ends in an error in state: 273.
+## Ends in an error in state: 246.
##
-## direct_abstract_declarator -> LPAREN in_context(option(parameter_type_list)) . RPAREN [ RPAREN LPAREN LBRACK COMMA ]
+## direct_abstract_declarator -> direct_abstract_declarator LPAREN option(context_parameter_type_list) . RPAREN [ RPAREN LPAREN LBRACK COMMA ]
##
## The known suffix of the stack is as follows:
-## LPAREN in_context(option(parameter_type_list))
+## direct_abstract_declarator LPAREN option(context_parameter_type_list)
##
-translation_unit_file: INT VAR_NAME LPAREN INT COMMA ELLIPSIS XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT COMMA ELLIPSIS XOR_ASSIGN
##
-## Ends in an error in state: 298.
+## Ends in an error in state: 280.
##
-## direct_declarator -> direct_declarator LPAREN open_context option(parameter_type_list) save_contexts_stk close_context . RPAREN [ SEMICOLON RPAREN PACKED LPAREN LBRACK LBRACE EQ COMMA COLON ATTRIBUTE ALIGNAS ]
+## direct_declarator -> direct_declarator LPAREN context_parameter_type_list . RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## direct_declarator LPAREN open_context option(parameter_type_list) save_contexts_stk close_context
+## direct_declarator LPAREN context_parameter_type_list
##
# Unlikely error, since only the ELLIPSIS allows us to tell that
@@ -466,18 +469,18 @@ At this point, a closing parenthesis ')' is expected.
translation_unit_file: ALIGNAS LPAREN INT LPAREN LPAREN RPAREN COMMA
##
-## Ends in an error in state: 285.
+## Ends in an error in state: 261.
##
-## direct_abstract_declarator -> LPAREN abstract_declarator(type_name) . RPAREN [ RPAREN LPAREN LBRACK COMMA ]
+## direct_abstract_declarator -> LPAREN save_context abstract_declarator(type_name) . RPAREN [ RPAREN LPAREN LBRACK COMMA ]
##
## The known suffix of the stack is as follows:
-## LPAREN abstract_declarator(type_name)
+## LPAREN save_context abstract_declarator(type_name)
##
## WARNING: This example involves spurious reductions.
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 282, spurious reduction of production abstract_declarator(type_name) -> direct_abstract_declarator
+## In state 254, spurious reduction of production abstract_declarator(type_name) -> direct_abstract_declarator
##
#
# The first LPAREN in this example must be the beginning of an abstract_declarator.
@@ -501,17 +504,17 @@ translation_unit_file: ALIGNAS LPAREN INT LPAREN LPAREN RPAREN COMMA
# gcc and clang simply request a closing parenthesis.
Up to this point, an abstract declarator has been recognized:
- $1 $0
+ $0
At this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
translation_unit_file: ALIGNAS LPAREN INT LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 149.
+## Ends in an error in state: 302.
##
-## direct_abstract_declarator -> LPAREN . abstract_declarator(type_name) RPAREN [ RPAREN LPAREN LBRACK ]
-## direct_abstract_declarator -> LPAREN . in_context(option(parameter_type_list)) RPAREN [ RPAREN LPAREN LBRACK ]
+## direct_abstract_declarator -> LPAREN . save_context abstract_declarator(type_name) RPAREN [ RPAREN LPAREN LBRACK ]
+## direct_abstract_declarator -> LPAREN . option(context_parameter_type_list) RPAREN [ RPAREN LPAREN LBRACK ]
##
## The known suffix of the stack is as follows:
## LPAREN
@@ -529,13 +532,13 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT LPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 261.
+## Ends in an error in state: 229.
##
-## direct_abstract_declarator -> LPAREN . abstract_declarator(type_name) RPAREN [ RPAREN LPAREN LBRACK COMMA ]
-## direct_abstract_declarator -> LPAREN . in_context(option(parameter_type_list)) RPAREN [ RPAREN LPAREN LBRACK COMMA ]
-## direct_declarator -> LPAREN . declarator RPAREN [ RPAREN PACKED LPAREN LBRACK COMMA ATTRIBUTE ALIGNAS ]
+## direct_abstract_declarator -> LPAREN . save_context abstract_declarator(type_name) RPAREN [ RPAREN LPAREN LBRACK COMMA ]
+## direct_abstract_declarator -> LPAREN . option(context_parameter_type_list) RPAREN [ RPAREN LPAREN LBRACK COMMA ]
+## direct_declarator -> LPAREN . save_context declarator RPAREN [ RPAREN PACKED LPAREN LBRACK COMMA ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
## LPAREN
@@ -554,14 +557,14 @@ At this point, one of the following is expected:
translation_unit_file: ALIGNAS LPAREN VOLATILE ADD_ASSIGN
##
-## Ends in an error in state: 179.
+## Ends in an error in state: 294.
##
-## specifier_qualifier_list(type_name) -> option(type_qualifier_list) . TYPEDEF_NAME option(type_qualifier_list) [ STAR RPAREN LPAREN LBRACK ]
-## specifier_qualifier_list(type_name) -> option(type_qualifier_list) . type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name) [ STAR RPAREN LPAREN LBRACK ]
-## type_qualifier_list -> option(type_qualifier_list) . type_qualifier [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME STRUCT SIGNED SHORT RESTRICT PACKED LONG INT FLOAT ENUM DOUBLE CONST CHAR ATTRIBUTE ALIGNAS ]
+## option(type_qualifier_list) -> type_qualifier_list . [ VOLATILE RESTRICT PACKED CONST ATTRIBUTE ALIGNAS ]
+## specifier_qualifier_list(type_name) -> type_qualifier_list . typedef_name option(type_qualifier_list) [ STAR RPAREN LPAREN LBRACK ]
+## specifier_qualifier_list(type_name) -> type_qualifier_list . type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name) [ STAR RPAREN LPAREN LBRACK ]
##
## The known suffix of the stack is as follows:
-## option(type_qualifier_list)
+## type_qualifier_list
##
# We are trying to recognize a specifier-qualifier-list, and have not yet seen
@@ -579,10 +582,10 @@ At this point, one of the following is expected:
translation_unit_file: ALIGNAS LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 176.
+## Ends in an error in state: 150.
##
-## attribute_specifier -> ALIGNAS LPAREN . argument_expression_list RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
-## attribute_specifier -> ALIGNAS LPAREN . type_name RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> ALIGNAS LPAREN . argument_expression_list RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> ALIGNAS LPAREN . type_name RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ALIGNAS LPAREN
@@ -601,10 +604,10 @@ At this point, one of the following is expected:
translation_unit_file: ALIGNAS XOR_ASSIGN
##
-## Ends in an error in state: 175.
+## Ends in an error in state: 149.
##
-## attribute_specifier -> ALIGNAS . LPAREN argument_expression_list RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
-## attribute_specifier -> ALIGNAS . LPAREN type_name RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> ALIGNAS . LPAREN argument_expression_list RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> ALIGNAS . LPAREN type_name RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ALIGNAS
@@ -619,7 +622,7 @@ At this point, an opening parenthesis '(' is expected.
translation_unit_file: ATTRIBUTE LPAREN LPAREN COMMA XOR_ASSIGN
##
-## Ends in an error in state: 172.
+## Ends in an error in state: 354.
##
## gcc_attribute_list -> gcc_attribute_list COMMA . gcc_attribute [ RPAREN COMMA ]
##
@@ -641,9 +644,9 @@ At this point, a gcc attribute is expected.
translation_unit_file: ATTRIBUTE LPAREN LPAREN RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 170.
+## Ends in an error in state: 352.
##
-## attribute_specifier -> ATTRIBUTE LPAREN LPAREN gcc_attribute_list RPAREN . RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> ATTRIBUTE LPAREN LPAREN gcc_attribute_list RPAREN . RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ATTRIBUTE LPAREN LPAREN gcc_attribute_list RPAREN
@@ -654,11 +657,11 @@ At this point, a second closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: ATTRIBUTE LPAREN LPAREN VAR_NAME LPAREN RPAREN XOR_ASSIGN
+translation_unit_file: ATTRIBUTE LPAREN LPAREN PRE_NAME VAR_NAME LPAREN RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 169.
+## Ends in an error in state: 351.
##
-## attribute_specifier -> ATTRIBUTE LPAREN LPAREN gcc_attribute_list . RPAREN RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> ATTRIBUTE LPAREN LPAREN gcc_attribute_list . RPAREN RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
## gcc_attribute_list -> gcc_attribute_list . COMMA gcc_attribute [ RPAREN COMMA ]
##
## The known suffix of the stack is as follows:
@@ -677,35 +680,35 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: ATTRIBUTE LPAREN LPAREN VAR_NAME LPAREN TYPEDEF_NAME COMMA VAR_NAME SEMICOLON
+translation_unit_file: ATTRIBUTE LPAREN LPAREN PRE_NAME VAR_NAME LPAREN PRE_NAME TYPEDEF_NAME COMMA PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 165.
+## Ends in an error in state: 347.
##
## argument_expression_list -> argument_expression_list . COMMA assignment_expression [ RPAREN COMMA ]
-## gcc_attribute -> gcc_attribute_word LPAREN TYPEDEF_NAME COMMA argument_expression_list . RPAREN [ RPAREN COMMA ]
+## gcc_attribute -> gcc_attribute_word LPAREN typedef_name COMMA argument_expression_list . RPAREN [ RPAREN COMMA ]
##
## The known suffix of the stack is as follows:
-## gcc_attribute_word LPAREN TYPEDEF_NAME COMMA argument_expression_list
+## gcc_attribute_word LPAREN typedef_name COMMA argument_expression_list
##
## WARNING: This example involves spurious reductions.
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 109, spurious reduction of production argument_expression_list -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 133, spurious reduction of production argument_expression_list -> assignment_expression
##
# We know for sure that we are parsing a gcc attribute.
@@ -721,14 +724,14 @@ then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: ATTRIBUTE LPAREN LPAREN VAR_NAME LPAREN TYPEDEF_NAME COMMA XOR_ASSIGN
+translation_unit_file: ATTRIBUTE LPAREN LPAREN PRE_NAME VAR_NAME LPAREN PRE_NAME TYPEDEF_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 164.
+## Ends in an error in state: 346.
##
-## gcc_attribute -> gcc_attribute_word LPAREN TYPEDEF_NAME COMMA . argument_expression_list RPAREN [ RPAREN COMMA ]
+## gcc_attribute -> gcc_attribute_word LPAREN typedef_name COMMA . argument_expression_list RPAREN [ RPAREN COMMA ]
##
## The known suffix of the stack is as follows:
-## gcc_attribute_word LPAREN TYPEDEF_NAME COMMA
+## gcc_attribute_word LPAREN typedef_name COMMA
##
# gcc/clang agree.
@@ -738,14 +741,14 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: ATTRIBUTE LPAREN LPAREN VAR_NAME LPAREN TYPEDEF_NAME XOR_ASSIGN
+translation_unit_file: ATTRIBUTE LPAREN LPAREN PRE_NAME VAR_NAME LPAREN PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 163.
+## Ends in an error in state: 345.
##
-## gcc_attribute -> gcc_attribute_word LPAREN TYPEDEF_NAME . COMMA argument_expression_list RPAREN [ RPAREN COMMA ]
+## gcc_attribute -> gcc_attribute_word LPAREN typedef_name . COMMA argument_expression_list RPAREN [ RPAREN COMMA ]
##
## The known suffix of the stack is as follows:
-## gcc_attribute_word LPAREN TYPEDEF_NAME
+## gcc_attribute_word LPAREN typedef_name
##
# gcc and clang complain about the TYPEDEF_NAME, not sure why.
@@ -755,12 +758,12 @@ At this point, a comma ',' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: ATTRIBUTE LPAREN LPAREN VAR_NAME LPAREN XOR_ASSIGN
+translation_unit_file: ATTRIBUTE LPAREN LPAREN PRE_NAME VAR_NAME LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 162.
+## Ends in an error in state: 45.
##
## gcc_attribute -> gcc_attribute_word LPAREN . option(argument_expression_list) RPAREN [ RPAREN COMMA ]
-## gcc_attribute -> gcc_attribute_word LPAREN . TYPEDEF_NAME COMMA argument_expression_list RPAREN [ RPAREN COMMA ]
+## gcc_attribute -> gcc_attribute_word LPAREN . typedef_name COMMA argument_expression_list RPAREN [ RPAREN COMMA ]
##
## The known suffix of the stack is as follows:
## gcc_attribute_word LPAREN
@@ -775,13 +778,13 @@ At this point, a list of expressions is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: ATTRIBUTE LPAREN LPAREN VAR_NAME XOR_ASSIGN
+translation_unit_file: ATTRIBUTE LPAREN LPAREN PRE_NAME VAR_NAME XOR_ASSIGN
##
-## Ends in an error in state: 161.
+## Ends in an error in state: 44.
##
## gcc_attribute -> gcc_attribute_word . [ RPAREN COMMA ]
## gcc_attribute -> gcc_attribute_word . LPAREN option(argument_expression_list) RPAREN [ RPAREN COMMA ]
-## gcc_attribute -> gcc_attribute_word . LPAREN TYPEDEF_NAME COMMA argument_expression_list RPAREN [ RPAREN COMMA ]
+## gcc_attribute -> gcc_attribute_word . LPAREN typedef_name COMMA argument_expression_list RPAREN [ RPAREN COMMA ]
##
## The known suffix of the stack is as follows:
## gcc_attribute_word
@@ -803,9 +806,9 @@ At this point, one of the following is expected:
translation_unit_file: ATTRIBUTE LPAREN LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 157.
+## Ends in an error in state: 37.
##
-## attribute_specifier -> ATTRIBUTE LPAREN LPAREN . gcc_attribute_list RPAREN RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> ATTRIBUTE LPAREN LPAREN . gcc_attribute_list RPAREN RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ATTRIBUTE LPAREN LPAREN
@@ -825,9 +828,9 @@ At this point, a gcc attribute is expected.
translation_unit_file: ATTRIBUTE LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 156.
+## Ends in an error in state: 36.
##
-## attribute_specifier -> ATTRIBUTE LPAREN . LPAREN gcc_attribute_list RPAREN RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> ATTRIBUTE LPAREN . LPAREN gcc_attribute_list RPAREN RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ATTRIBUTE LPAREN
@@ -840,9 +843,9 @@ At this point, a second opening parenthesis '(' is expected.
translation_unit_file: ATTRIBUTE XOR_ASSIGN
##
-## Ends in an error in state: 155.
+## Ends in an error in state: 35.
##
-## attribute_specifier -> ATTRIBUTE . LPAREN LPAREN gcc_attribute_list RPAREN RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> ATTRIBUTE . LPAREN LPAREN gcc_attribute_list RPAREN RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ATTRIBUTE
@@ -853,9 +856,9 @@ At this point, two opening parentheses '((' are expected.
# ------------------------------------------------------------------------------
-translation_unit_file: ENUM LBRACE VAR_NAME COMMA XOR_ASSIGN
+translation_unit_file: ENUM LBRACE PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 200.
+## Ends in an error in state: 362.
##
## enumerator_list -> enumerator_list COMMA . declare_varname(enumerator) [ RBRACE COMMA ]
## option(COMMA) -> COMMA . [ RBRACE ]
@@ -874,11 +877,11 @@ At this point, an enumerator is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: ENUM LBRACE VAR_NAME EQ CONSTANT SEMICOLON
+translation_unit_file: ENUM LBRACE PRE_NAME VAR_NAME EQ CONSTANT SEMICOLON
##
-## Ends in an error in state: 199.
+## Ends in an error in state: 361.
##
-## enum_specifier -> ENUM attribute_specifier_list option(other_identifier) LBRACE enumerator_list . option(COMMA) RBRACE [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
+## enum_specifier -> ENUM attribute_specifier_list option(other_identifier) LBRACE enumerator_list . option(COMMA) RBRACE [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
## enumerator_list -> enumerator_list . COMMA declare_varname(enumerator) [ RBRACE COMMA ]
##
## The known suffix of the stack is as follows:
@@ -888,22 +891,22 @@ translation_unit_file: ENUM LBRACE VAR_NAME EQ CONSTANT SEMICOLON
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 31, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 204, spurious reduction of production enumerator -> enumeration_constant EQ conditional_expression
-## In state 201, spurious reduction of production declare_varname(enumerator) -> enumerator
-## In state 208, spurious reduction of production enumerator_list -> declare_varname(enumerator)
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 57, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 366, spurious reduction of production enumerator -> enumeration_constant EQ conditional_expression
+## In state 363, spurious reduction of production declare_varname(enumerator) -> enumerator
+## In state 370, spurious reduction of production enumerator_list -> declare_varname(enumerator)
##
#
# At first sight, it seems that the last enumerator that we have recognized
@@ -932,9 +935,9 @@ then at this point, a closing brace '}' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: ENUM LBRACE VAR_NAME EQ XOR_ASSIGN
+translation_unit_file: ENUM LBRACE PRE_NAME VAR_NAME EQ XOR_ASSIGN
##
-## Ends in an error in state: 203.
+## Ends in an error in state: 365.
##
## enumerator -> enumeration_constant EQ . conditional_expression [ RBRACE COMMA ]
##
@@ -947,9 +950,9 @@ At this point, a constant expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: ENUM LBRACE VAR_NAME XOR_ASSIGN
+translation_unit_file: ENUM LBRACE PRE_NAME VAR_NAME XOR_ASSIGN
##
-## Ends in an error in state: 202.
+## Ends in an error in state: 364.
##
## enumerator -> enumeration_constant . [ RBRACE COMMA ]
## enumerator -> enumeration_constant . EQ conditional_expression [ RBRACE COMMA ]
@@ -970,9 +973,9 @@ At this point, one of the following is expected:
translation_unit_file: ENUM LBRACE XOR_ASSIGN
##
-## Ends in an error in state: 197.
+## Ends in an error in state: 359.
##
-## enum_specifier -> ENUM attribute_specifier_list option(other_identifier) LBRACE . enumerator_list option(COMMA) RBRACE [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
+## enum_specifier -> ENUM attribute_specifier_list option(other_identifier) LBRACE . enumerator_list option(COMMA) RBRACE [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ENUM attribute_specifier_list option(other_identifier) LBRACE
@@ -988,15 +991,20 @@ At this point, an enumerator is expected.
translation_unit_file: ENUM XOR_ASSIGN
##
-## Ends in an error in state: 195.
+## Ends in an error in state: 357.
##
-## attribute_specifier_list -> attribute_specifier_list . attribute_specifier [ VAR_NAME TYPEDEF_NAME PACKED LBRACE ATTRIBUTE ALIGNAS ]
-## enum_specifier -> ENUM attribute_specifier_list . option(other_identifier) LBRACE enumerator_list option(COMMA) RBRACE [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
-## enum_specifier -> ENUM attribute_specifier_list . general_identifier [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
+## enum_specifier -> ENUM attribute_specifier_list . option(other_identifier) LBRACE enumerator_list option(COMMA) RBRACE [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
+## enum_specifier -> ENUM attribute_specifier_list . general_identifier [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ENUM attribute_specifier_list
##
+## WARNING: This example involves spurious reductions.
+## This implies that, although the LR(1) items shown above provide an
+## accurate view of the past (what has been recognized so far), they
+## may provide an INCOMPLETE view of the future (what was expected next).
+## In state 34, spurious reduction of production attribute_specifier_list ->
+##
# Here, both clang and gcc give an incomplete diagnostic message.
@@ -1008,9 +1016,9 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ ALIGNOF LPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ ALIGNOF LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 29.
+## Ends in an error in state: 53.
##
## postfix_expression -> LPAREN . type_name RPAREN LBRACE initializer_list option(COMMA) RBRACE [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
## primary_expression -> LPAREN . expression RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
@@ -1019,9 +1027,9 @@ translation_unit_file: INT VAR_NAME EQ ALIGNOF LPAREN XOR_ASSIGN
## The known suffix of the stack is as follows:
## ALIGNOF LPAREN
##
-translation_unit_file: INT VAR_NAME EQ SIZEOF LPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ SIZEOF LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 18.
+## Ends in an error in state: 26.
##
## postfix_expression -> LPAREN . type_name RPAREN LBRACE initializer_list option(COMMA) RBRACE [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
## primary_expression -> LPAREN . expression RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
@@ -1043,9 +1051,9 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ ALIGNOF XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ ALIGNOF XOR_ASSIGN
##
-## Ends in an error in state: 28.
+## Ends in an error in state: 52.
##
## unary_expression -> ALIGNOF . unary_expression [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LEQ LEFT_ASSIGN LEFT HAT GT GEQ EQEQ EQ DIV_ASSIGN COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
## unary_expression -> ALIGNOF . LPAREN type_name RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LEQ LEFT_ASSIGN LEFT HAT GT GEQ EQEQ EQ DIV_ASSIGN COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
@@ -1053,9 +1061,9 @@ translation_unit_file: INT VAR_NAME EQ ALIGNOF XOR_ASSIGN
## The known suffix of the stack is as follows:
## ALIGNOF
##
-translation_unit_file: INT VAR_NAME EQ SIZEOF XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ SIZEOF XOR_ASSIGN
##
-## Ends in an error in state: 15.
+## Ends in an error in state: 21.
##
## unary_expression -> SIZEOF . unary_expression [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LEQ LEFT_ASSIGN LEFT HAT GT GEQ EQEQ EQ DIV_ASSIGN COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
## unary_expression -> SIZEOF . LPAREN type_name RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LEQ LEFT_ASSIGN LEFT HAT GT GEQ EQEQ EQ DIV_ASSIGN COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
@@ -1075,9 +1083,9 @@ followed with an expression or a type name.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ BUILTIN_VA_ARG LPAREN VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME EQ BUILTIN_VA_ARG LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 329.
+## Ends in an error in state: 340.
##
## postfix_expression -> BUILTIN_VA_ARG LPAREN assignment_expression . COMMA type_name RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
@@ -1088,20 +1096,20 @@ translation_unit_file: INT VAR_NAME EQ BUILTIN_VA_ARG LPAREN VAR_NAME SEMICOLON
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
##
Ill-formed use of $2.
@@ -1112,9 +1120,24 @@ then at this point, a comma ',' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ BUILTIN_VA_ARG LPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ BUILTIN_VA_ARG LPAREN PRE_NAME VAR_NAME COMMA XOR_ASSIGN
+##
+## Ends in an error in state: 341.
+##
+## postfix_expression -> BUILTIN_VA_ARG LPAREN assignment_expression COMMA . type_name RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
+##
+## The known suffix of the stack is as follows:
+## BUILTIN_VA_ARG LPAREN assignment_expression COMMA
+##
+
+Ill-formed use of $3.
+At this point, a type name is expected.
+
+# ------------------------------------------------------------------------------
+
+translation_unit_file: INT PRE_NAME VAR_NAME EQ BUILTIN_VA_ARG LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 25.
+## Ends in an error in state: 49.
##
## postfix_expression -> BUILTIN_VA_ARG LPAREN . assignment_expression COMMA type_name RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
@@ -1127,9 +1150,9 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ BUILTIN_VA_ARG XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ BUILTIN_VA_ARG XOR_ASSIGN
##
-## Ends in an error in state: 24.
+## Ends in an error in state: 48.
##
## postfix_expression -> BUILTIN_VA_ARG . LPAREN assignment_expression COMMA type_name RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
@@ -1142,18 +1165,18 @@ At this point, an opening parenthesis '(' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ DEC XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ DEC XOR_ASSIGN
##
-## Ends in an error in state: 22.
+## Ends in an error in state: 46.
##
## unary_expression -> DEC . unary_expression [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LEQ LEFT_ASSIGN LEFT HAT GT GEQ EQEQ EQ DIV_ASSIGN COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
## The known suffix of the stack is as follows:
## DEC
##
-translation_unit_file: INT VAR_NAME EQ INC XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ INC XOR_ASSIGN
##
-## Ends in an error in state: 20.
+## Ends in an error in state: 31.
##
## unary_expression -> INC . unary_expression [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LEQ LEFT_ASSIGN LEFT HAT GT GEQ EQEQ EQ DIV_ASSIGN COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
@@ -1166,9 +1189,9 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ INC LPAREN INT RPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ INC LPAREN INT RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 335.
+## Ends in an error in state: 373.
##
## postfix_expression -> LPAREN type_name RPAREN . LBRACE initializer_list option(COMMA) RBRACE [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
@@ -1197,9 +1220,9 @@ If this is intended to be the beginning of a cast expression,
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ INC LPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ INC LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 21.
+## Ends in an error in state: 32.
##
## postfix_expression -> LPAREN . type_name RPAREN LBRACE initializer_list option(COMMA) RBRACE [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
## primary_expression -> LPAREN . expression RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
@@ -1218,9 +1241,9 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LPAREN VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 326.
+## Ends in an error in state: 337.
##
## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
## primary_expression -> LPAREN expression . RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
@@ -1232,21 +1255,21 @@ translation_unit_file: INT VAR_NAME EQ LPAREN VAR_NAME SEMICOLON
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 127, spurious reduction of production expression -> assignment_expression
##
# Since we are saying "if this expression is complete",
@@ -1254,16 +1277,16 @@ translation_unit_file: INT VAR_NAME EQ LPAREN VAR_NAME SEMICOLON
# can be continued with a comma and another expression.
# So, let's just say a closing parenthesis is expected.
-Up to this point, a parenthesis '(' and an expression have been recognized:
- $1 $0
+Up to this point, an expression have been recognized:
+ $0
If this expression is complete,
then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LPAREN INT RPAREN LBRACE VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LPAREN INT RPAREN LBRACE PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 145.
+## Ends in an error in state: 334.
##
## initializer_list -> initializer_list . COMMA option(designation) c_initializer [ RBRACE COMMA ]
## postfix_expression -> LPAREN type_name RPAREN LBRACE initializer_list . option(COMMA) RBRACE [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
@@ -1275,22 +1298,22 @@ translation_unit_file: INT VAR_NAME EQ LPAREN INT RPAREN LBRACE VAR_NAME SEMICOL
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 138, spurious reduction of production c_initializer -> assignment_expression
-## In state 144, spurious reduction of production initializer_list -> option(designation) c_initializer
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 327, spurious reduction of production c_initializer -> assignment_expression
+## In state 333, spurious reduction of production initializer_list -> option(designation) c_initializer
##
# Let's ignore the fact that a comma can precede a closing brace.
@@ -1303,9 +1326,9 @@ then at this point, a closing brace '}' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LPAREN INT RPAREN LBRACE XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LPAREN INT RPAREN LBRACE XOR_ASSIGN
##
-## Ends in an error in state: 124.
+## Ends in an error in state: 313.
##
## postfix_expression -> LPAREN type_name RPAREN LBRACE . initializer_list option(COMMA) RBRACE [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
@@ -1320,9 +1343,9 @@ At this point, an initializer is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LPAREN INT RPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LPAREN INT RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 338.
+## Ends in an error in state: 376.
##
## cast_expression -> LPAREN type_name RPAREN . cast_expression [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LEQ LEFT_ASSIGN LEFT HAT GT GEQ EQEQ EQ DIV_ASSIGN COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
## postfix_expression -> LPAREN type_name RPAREN . LBRACE initializer_list option(COMMA) RBRACE [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
@@ -1342,9 +1365,9 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 19.
+## Ends in an error in state: 28.
##
## cast_expression -> LPAREN . type_name RPAREN cast_expression [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LEQ LEFT_ASSIGN LEFT HAT GT GEQ EQEQ EQ DIV_ASSIGN COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
## postfix_expression -> LPAREN . type_name RPAREN LBRACE initializer_list option(COMMA) RBRACE [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
@@ -1364,9 +1387,9 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ TILDE XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ TILDE XOR_ASSIGN
##
-## Ends in an error in state: 30.
+## Ends in an error in state: 56.
##
## unary_expression -> unary_operator . cast_expression [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LEQ LEFT_ASSIGN LEFT HAT GT GEQ EQEQ EQ DIV_ASSIGN COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
@@ -1381,90 +1404,90 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ VAR_NAME AND XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME AND XOR_ASSIGN
##
-## Ends in an error in state: 92.
+## Ends in an error in state: 116.
##
## and_expression -> and_expression AND . equality_expression [ SEMICOLON RPAREN RBRACK RBRACE QUESTION HAT COMMA COLON BARBAR BAR ANDAND AND ]
##
## The known suffix of the stack is as follows:
## and_expression AND
##
-translation_unit_file: INT VAR_NAME EQ VAR_NAME ANDAND XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME ANDAND XOR_ASSIGN
##
-## Ends in an error in state: 81.
+## Ends in an error in state: 105.
##
## logical_and_expression -> logical_and_expression ANDAND . inclusive_or_expression [ SEMICOLON RPAREN RBRACK RBRACE QUESTION COMMA COLON BARBAR ANDAND ]
##
## The known suffix of the stack is as follows:
## logical_and_expression ANDAND
##
-translation_unit_file: INT VAR_NAME EQ VAR_NAME BAR XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME BAR XOR_ASSIGN
##
-## Ends in an error in state: 83.
+## Ends in an error in state: 107.
##
## inclusive_or_expression -> inclusive_or_expression BAR . exclusive_or_expression [ SEMICOLON RPAREN RBRACK RBRACE QUESTION COMMA COLON BARBAR BAR ANDAND ]
##
## The known suffix of the stack is as follows:
## inclusive_or_expression BAR
##
-translation_unit_file: INT VAR_NAME EQ VAR_NAME BARBAR XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME BARBAR XOR_ASSIGN
##
-## Ends in an error in state: 104.
+## Ends in an error in state: 128.
##
## logical_or_expression -> logical_or_expression BARBAR . logical_and_expression [ SEMICOLON RPAREN RBRACK RBRACE QUESTION COMMA COLON BARBAR ]
##
## The known suffix of the stack is as follows:
## logical_or_expression BARBAR
##
-translation_unit_file: INT VAR_NAME EQ VAR_NAME HAT XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME HAT XOR_ASSIGN
##
-## Ends in an error in state: 85.
+## Ends in an error in state: 109.
##
## exclusive_or_expression -> exclusive_or_expression HAT . and_expression [ SEMICOLON RPAREN RBRACK RBRACE QUESTION HAT COMMA COLON BARBAR BAR ANDAND ]
##
## The known suffix of the stack is as follows:
## exclusive_or_expression HAT
##
-translation_unit_file: INT VAR_NAME EQ VAR_NAME LT XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME LT XOR_ASSIGN
##
-## Ends in an error in state: 75.
+## Ends in an error in state: 99.
##
## relational_expression -> relational_expression relational_operator . shift_expression [ SEMICOLON RPAREN RBRACK RBRACE QUESTION NEQ LT LEQ HAT GT GEQ EQEQ COMMA COLON BARBAR BAR ANDAND AND ]
##
## The known suffix of the stack is as follows:
## relational_expression relational_operator
##
-translation_unit_file: INT VAR_NAME EQ VAR_NAME NEQ XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME NEQ XOR_ASSIGN
##
-## Ends in an error in state: 89.
+## Ends in an error in state: 113.
##
## equality_expression -> equality_expression equality_operator . relational_expression [ SEMICOLON RPAREN RBRACK RBRACE QUESTION NEQ HAT EQEQ COMMA COLON BARBAR BAR ANDAND AND ]
##
## The known suffix of the stack is as follows:
## equality_expression equality_operator
##
-translation_unit_file: INT VAR_NAME EQ VAR_NAME PLUS XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME PLUS XOR_ASSIGN
##
-## Ends in an error in state: 68.
+## Ends in an error in state: 92.
##
## additive_expression -> additive_expression additive_operator . multiplicative_expression [ SEMICOLON RPAREN RIGHT RBRACK RBRACE QUESTION PLUS NEQ MINUS LT LEQ LEFT HAT GT GEQ EQEQ COMMA COLON BARBAR BAR ANDAND AND ]
##
## The known suffix of the stack is as follows:
## additive_expression additive_operator
##
-translation_unit_file: INT VAR_NAME EQ VAR_NAME RIGHT XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME RIGHT XOR_ASSIGN
##
-## Ends in an error in state: 57.
+## Ends in an error in state: 81.
##
## shift_expression -> shift_expression shift_operator . additive_expression [ SEMICOLON RPAREN RIGHT RBRACK RBRACE QUESTION NEQ LT LEQ LEFT HAT GT GEQ EQEQ COMMA COLON BARBAR BAR ANDAND AND ]
##
## The known suffix of the stack is as follows:
## shift_expression shift_operator
##
-translation_unit_file: INT VAR_NAME EQ VAR_NAME STAR XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME STAR XOR_ASSIGN
##
-## Ends in an error in state: 62.
+## Ends in an error in state: 86.
##
## multiplicative_expression -> multiplicative_expression multiplicative_operator . cast_expression [ STAR SLASH SEMICOLON RPAREN RIGHT RBRACK RBRACE QUESTION PLUS PERCENT NEQ MINUS LT LEQ LEFT HAT GT GEQ EQEQ COMMA COLON BARBAR BAR ANDAND AND ]
##
@@ -1479,9 +1502,9 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ VAR_NAME XOR_ASSIGN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME XOR_ASSIGN XOR_ASSIGN
##
-## Ends in an error in state: 53.
+## Ends in an error in state: 77.
##
## assignment_expression -> unary_expression assignment_operator . assignment_expression [ SEMICOLON RPAREN RBRACK RBRACE COMMA COLON ]
##
@@ -1496,9 +1519,9 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ VAR_NAME LPAREN VAR_NAME COMMA XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 111.
+## Ends in an error in state: 135.
##
## argument_expression_list -> argument_expression_list COMMA . assignment_expression [ RPAREN COMMA ]
##
@@ -1517,18 +1540,18 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ VAR_NAME DOT XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME DOT XOR_ASSIGN
##
-## Ends in an error in state: 117.
+## Ends in an error in state: 141.
##
## postfix_expression -> postfix_expression DOT . general_identifier [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
## The known suffix of the stack is as follows:
## postfix_expression DOT
##
-translation_unit_file: INT VAR_NAME EQ VAR_NAME PTR XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME PTR XOR_ASSIGN
##
-## Ends in an error in state: 36.
+## Ends in an error in state: 62.
##
## postfix_expression -> postfix_expression PTR . general_identifier [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
@@ -1543,9 +1566,9 @@ At this point, the name of a struct or union member is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ VAR_NAME LBRACK VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME LBRACK PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 114.
+## Ends in an error in state: 138.
##
## expression -> expression . COMMA assignment_expression [ RBRACK COMMA ]
## postfix_expression -> postfix_expression LBRACK expression . RBRACK [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
@@ -1557,21 +1580,21 @@ translation_unit_file: INT VAR_NAME EQ VAR_NAME LBRACK VAR_NAME SEMICOLON
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 127, spurious reduction of production expression -> assignment_expression
##
# We know for sure that an array subscript expression has begun, and
@@ -1588,9 +1611,9 @@ then at this point, a closing bracket ']' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ VAR_NAME LBRACK XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME LBRACK XOR_ASSIGN
##
-## Ends in an error in state: 113.
+## Ends in an error in state: 137.
##
## postfix_expression -> postfix_expression LBRACK . expression RBRACK [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
@@ -1603,9 +1626,9 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ VAR_NAME LPAREN VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 110.
+## Ends in an error in state: 134.
##
## argument_expression_list -> argument_expression_list . COMMA assignment_expression [ RPAREN COMMA ]
## option(argument_expression_list) -> argument_expression_list . [ RPAREN ]
@@ -1617,21 +1640,21 @@ translation_unit_file: INT VAR_NAME EQ VAR_NAME LPAREN VAR_NAME SEMICOLON
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 109, spurious reduction of production argument_expression_list -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 133, spurious reduction of production argument_expression_list -> assignment_expression
##
Up to this point, a list of expressions has been recognized:
@@ -1641,9 +1664,9 @@ then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ VAR_NAME LPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 40.
+## Ends in an error in state: 64.
##
## postfix_expression -> postfix_expression LPAREN . option(argument_expression_list) RPAREN [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
##
@@ -1659,18 +1682,18 @@ followed with a closing parenthesis ')', is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ VAR_NAME QUESTION VAR_NAME COLON XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME QUESTION PRE_NAME VAR_NAME COLON XOR_ASSIGN
##
-## Ends in an error in state: 101.
+## Ends in an error in state: 125.
##
## conditional_expression -> logical_or_expression QUESTION expression COLON . conditional_expression [ SEMICOLON RPAREN RBRACK RBRACE COMMA COLON ]
##
## The known suffix of the stack is as follows:
## logical_or_expression QUESTION expression COLON
##
-translation_unit_file: INT VAR_NAME EQ VAR_NAME QUESTION XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME QUESTION XOR_ASSIGN
##
-## Ends in an error in state: 79.
+## Ends in an error in state: 103.
##
## conditional_expression -> logical_or_expression QUESTION . expression COLON conditional_expression [ SEMICOLON RPAREN RBRACK RBRACE COMMA COLON ]
##
@@ -1683,9 +1706,9 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ VAR_NAME QUESTION VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME QUESTION PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 97.
+## Ends in an error in state: 121.
##
## conditional_expression -> logical_or_expression QUESTION expression . COLON conditional_expression [ SEMICOLON RPAREN RBRACK RBRACE COMMA COLON ]
## expression -> expression . COMMA assignment_expression [ COMMA COLON ]
@@ -1697,21 +1720,21 @@ translation_unit_file: INT VAR_NAME EQ VAR_NAME QUESTION VAR_NAME SEMICOLON
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 127, spurious reduction of production expression -> assignment_expression
##
# gcc and clang simply expect a colon.
@@ -1726,12 +1749,12 @@ then at this point, a colon ':' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: PACKED LPAREN VAR_NAME SEMICOLON
+translation_unit_file: PACKED LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 343.
+## Ends in an error in state: 381.
##
## argument_expression_list -> argument_expression_list . COMMA assignment_expression [ RPAREN COMMA ]
-## attribute_specifier -> PACKED LPAREN argument_expression_list . RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> PACKED LPAREN argument_expression_list . RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## PACKED LPAREN argument_expression_list
@@ -1740,21 +1763,21 @@ translation_unit_file: PACKED LPAREN VAR_NAME SEMICOLON
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 109, spurious reduction of production argument_expression_list -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 133, spurious reduction of production argument_expression_list -> assignment_expression
##
Ill-formed $2 attribute.
@@ -1767,9 +1790,9 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: PACKED LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 10.
+## Ends in an error in state: 17.
##
-## attribute_specifier -> PACKED LPAREN . argument_expression_list RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> PACKED LPAREN . argument_expression_list RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## PACKED LPAREN
@@ -1784,11 +1807,27 @@ At this point, a list of expressions is expected.
# ------------------------------------------------------------------------------
+translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME TYPEDEF_NAME
+##
+## Ends in an error in state: 22.
+##
+## primary_expression -> PRE_NAME . VAR_NAME [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RBRACK RBRACE QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA COLON BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
+##
+## The known suffix of the stack is as follows:
+## PRE_NAME
+##
+
+Ill-formed expression.
+The following identifier is used as a variable, but has been defined as a type:
+ $0
+
+# ------------------------------------------------------------------------------
+
translation_unit_file: PACKED XOR_ASSIGN
##
-## Ends in an error in state: 9.
+## Ends in an error in state: 16.
##
-## attribute_specifier -> PACKED . LPAREN argument_expression_list RPAREN [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## attribute_specifier -> PACKED . LPAREN argument_expression_list RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER RBRACK PRE_NAME PLUS PACKED MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## PACKED
@@ -1806,10 +1845,10 @@ is expected.
translation_unit_file: SEMICOLON XOR_ASSIGN
##
-## Ends in an error in state: 349.
+## Ends in an error in state: 399.
##
-## translation_unit -> translation_unit . external_declaration [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC SIGNED SHORT SEMICOLON RESTRICT REGISTER PRAGMA PACKED LONG INT INLINE FLOAT EXTERN EOF ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
-## translation_unit -> translation_unit . SEMICOLON [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC SIGNED SHORT SEMICOLON RESTRICT REGISTER PRAGMA PACKED LONG INT INLINE FLOAT EXTERN EOF ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## translation_unit -> translation_unit . external_declaration [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC SIGNED SHORT SEMICOLON RESTRICT REGISTER PRE_NAME PRAGMA PACKED LONG INT INLINE FLOAT EXTERN EOF ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## translation_unit -> translation_unit . SEMICOLON [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC SIGNED SHORT SEMICOLON RESTRICT REGISTER PRE_NAME PRAGMA PACKED LONG INT INLINE FLOAT EXTERN EOF ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
## translation_unit_file -> translation_unit . EOF [ # ]
##
## The known suffix of the stack is as follows:
@@ -1838,55 +1877,97 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: TYPEDEF TYPEDEF_NAME XOR_ASSIGN
+translation_unit_file: TYPEDEF PRE_NAME TYPEDEF_NAME XOR_ASSIGN
+##
+## Ends in an error in state: 384.
+##
+## declaration_specifiers_typedef -> TYPEDEF list(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## TYPEDEF list(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type)
+##
+translation_unit_file: PRE_NAME TYPEDEF_NAME TYPEDEF XOR_ASSIGN
##
-## Ends in an error in state: 356.
+## Ends in an error in state: 393.
##
-## declaration_specifiers_typedef -> list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) TYPEDEF_NAME list(declaration_specifier_no_type) . [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . declaration_specifier_no_type [ VOLATILE VAR_NAME TYPEDEF_NAME STATIC STAR SEMICOLON RESTRICT REGISTER PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers_typedef -> typedef_name list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) TYPEDEF_NAME list(declaration_specifier_no_type)
+## typedef_name list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type)
##
-translation_unit_file: TYPEDEF_NAME TYPEDEF XOR_ASSIGN
+translation_unit_file: VOLATILE TYPEDEF PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 5.
+## Ends in an error in state: 406.
##
-## declaration_specifiers_typedef -> TYPEDEF_NAME list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) . [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . declaration_specifier_no_type [ VOLATILE VAR_NAME TYPEDEF_NAME STATIC STAR SEMICOLON RESTRICT REGISTER PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## TYPEDEF_NAME list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type)
+## rlist(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type)
##
-translation_unit_file: VOLATILE TYPEDEF_NAME TYPEDEF XOR_ASSIGN
+translation_unit_file: VOLATILE PRE_NAME TYPEDEF_NAME TYPEDEF XOR_ASSIGN
##
-## Ends in an error in state: 367.
+## Ends in an error in state: 412.
##
-## declaration_specifiers_typedef -> list(declaration_specifier_no_type) declaration_specifier_no_type TYPEDEF_NAME list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) . [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . declaration_specifier_no_type [ VOLATILE VAR_NAME TYPEDEF_NAME STATIC STAR SEMICOLON RESTRICT REGISTER PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) declaration_specifier_no_type TYPEDEF_NAME list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type)
+## rlist(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type)
##
translation_unit_file: TYPEDEF INT XOR_ASSIGN
##
-## Ends in an error in state: 358.
+## Ends in an error in state: 386.
##
-## declaration_specifiers_typedef -> list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers_typedef -> TYPEDEF list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
+## TYPEDEF list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
##
translation_unit_file: INT TYPEDEF XOR_ASSIGN
##
-## Ends in an error in state: 362.
+## Ends in an error in state: 397.
+##
+## declaration_specifiers_typedef -> type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) TYPEDEF list(declaration_specifier_no_typedef_name) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) TYPEDEF list(declaration_specifier_no_typedef_name)
+##
+translation_unit_file: VOLATILE TYPEDEF INT XOR_ASSIGN
+##
+## Ends in an error in state: 408.
##
-## declaration_specifiers_typedef -> list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) TYPEDEF list(declaration_specifier_no_typedef_name) . [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) TYPEDEF list(declaration_specifier_no_typedef_name)
+## rlist(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
+##
+translation_unit_file: VOLATILE INT TYPEDEF XOR_ASSIGN
+##
+## Ends in an error in state: 416.
+##
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) TYPEDEF list(declaration_specifier_no_typedef_name) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## rlist(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) TYPEDEF list(declaration_specifier_no_typedef_name)
##
# We have begun a type definition (a.k.a. declaration_specifiers_typedef).
@@ -1919,14 +2000,31 @@ At this point, one of the following is expected:
translation_unit_file: TYPEDEF XOR_ASSIGN
##
-## Ends in an error in state: 354.
+## Ends in an error in state: 7.
##
-## declaration_specifiers_typedef -> list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) . TYPEDEF_NAME list(declaration_specifier_no_type) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## declaration_specifiers_typedef -> list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) . type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . declaration_specifier_no_type [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PACKED LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers_typedef -> TYPEDEF list(declaration_specifier_no_type) . typedef_name list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> TYPEDEF list(declaration_specifier_no_type) . type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PRE_NAME PACKED LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PRE_NAME PACKED LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PRE_NAME PACKED LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PRE_NAME PACKED LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type)
+## TYPEDEF list(declaration_specifier_no_type)
+##
+translation_unit_file: VOLATILE TYPEDEF XOR_ASSIGN
+##
+## Ends in an error in state: 404.
+##
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) . typedef_name list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) . type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PRE_NAME PACKED LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PRE_NAME PACKED LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PRE_NAME PACKED LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PRE_NAME PACKED LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## rlist(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type)
##
# We have seen the TYPEDEF keyword, and possibly some declaration_specifiers_no_type.
@@ -1948,15 +2046,21 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN VOLATILE XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN VOLATILE XOR_ASSIGN
##
-## Ends in an error in state: 257.
+## Ends in an error in state: 217.
##
-## declaration_specifiers(parameter_declaration) -> list(declaration_specifier_no_type) declaration_specifier_no_type . TYPEDEF_NAME list(declaration_specifier_no_type) [ VAR_NAME TYPEDEF_NAME STAR RPAREN LPAREN LBRACK COMMA ]
-## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) declaration_specifier_no_type . [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PACKED LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers(parameter_declaration) -> rlist(declaration_specifier_no_type) . typedef_name list(declaration_specifier_no_type) [ STAR RPAREN PRE_NAME LPAREN LBRACK COMMA ]
+## declaration_specifiers(parameter_declaration) -> rlist(declaration_specifier_no_type) . type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) [ STAR RPAREN PRE_NAME LPAREN LBRACK COMMA ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) declaration_specifier_no_type
+## rlist(declaration_specifier_no_type)
+##
+## WARNING: This example involves spurious reductions.
+## This implies that, although the LR(1) items shown above provide an
+## accurate view of the past (what has been recognized so far), they
+## may provide an INCOMPLETE view of the future (what was expected next).
+## In state 209, spurious reduction of production rlist(declaration_specifier_no_type) -> type_qualifier_noattr
##
# Analogous to the above, except we are in the context of a parameter declaration,
@@ -1978,38 +2082,55 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: TYPEDEF_NAME XOR_ASSIGN
+translation_unit_file: PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 2.
+## Ends in an error in state: 391.
##
-## declaration_specifiers(declaration(external_declaration)) -> TYPEDEF_NAME list(declaration_specifier_no_type) . [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## declaration_specifiers_typedef -> TYPEDEF_NAME list(declaration_specifier_no_type) . TYPEDEF list(declaration_specifier_no_type) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . declaration_specifier_no_type [ VOLATILE VAR_NAME TYPEDEF_NAME TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers(declaration(external_declaration)) -> typedef_name list(declaration_specifier_no_type) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> typedef_name list(declaration_specifier_no_type) . TYPEDEF list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## TYPEDEF_NAME list(declaration_specifier_no_type)
+## typedef_name list(declaration_specifier_no_type)
##
-translation_unit_file: VOLATILE TYPEDEF_NAME XOR_ASSIGN
+translation_unit_file: VOLATILE PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 365.
+## Ends in an error in state: 410.
##
-## declaration_specifiers(declaration(external_declaration)) -> list(declaration_specifier_no_type) declaration_specifier_no_type TYPEDEF_NAME list(declaration_specifier_no_type) . [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## declaration_specifiers_typedef -> list(declaration_specifier_no_type) declaration_specifier_no_type TYPEDEF_NAME list(declaration_specifier_no_type) . TYPEDEF list(declaration_specifier_no_type) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . declaration_specifier_no_type [ VOLATILE VAR_NAME TYPEDEF_NAME TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers(declaration(external_declaration)) -> rlist(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type) . TYPEDEF list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) declaration_specifier_no_type TYPEDEF_NAME list(declaration_specifier_no_type)
+## rlist(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type)
##
translation_unit_file: INT XOR_ASSIGN
##
-## Ends in an error in state: 360.
+## Ends in an error in state: 395.
##
-## declaration_specifiers(declaration(external_declaration)) -> list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## declaration_specifiers_typedef -> list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . TYPEDEF list(declaration_specifier_no_typedef_name) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers(declaration(external_declaration)) -> type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . TYPEDEF list(declaration_specifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
+## type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
+##
+translation_unit_file: VOLATILE INT XOR_ASSIGN
+##
+## Ends in an error in state: 414.
+##
+## declaration_specifiers(declaration(external_declaration)) -> rlist(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . TYPEDEF list(declaration_specifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## rlist(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
##
# We have seen a TYPEDEF_NAME or a primitive type specifier,
@@ -2039,7 +2160,7 @@ translation_unit_file: INT XOR_ASSIGN
# If this declaration_specifiers is complete,
# then we expect:
# init_declarator_list? SEMICOLON if this is a declaration;
-# ioption(pointer) x=direct_declarator if this is a function definition.
+# declarator if this is a function definition.
# (followed with a function body)
# gcc and clang expect identifier or '(', which is very low-level.
@@ -2055,42 +2176,57 @@ At this point, one of the following is expected:
a type qualifier; or
an init declarator,
if this is a declaration; or
- a possibly empty list of pointers,
- followed with a direct declarator,
+ a declarator,
followed with a function body,
if this is a function definition.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN TYPEDEF_NAME XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 235.
+## Ends in an error in state: 196.
##
-## declaration_specifiers(parameter_declaration) -> TYPEDEF_NAME list(declaration_specifier_no_type) . [ VAR_NAME TYPEDEF_NAME STAR RPAREN LPAREN LBRACK COMMA ]
-## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . declaration_specifier_no_type [ VOLATILE VAR_NAME TYPEDEF_NAME STATIC STAR RPAREN RESTRICT REGISTER PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers(parameter_declaration) -> typedef_name list(declaration_specifier_no_type) . [ STAR RPAREN PRE_NAME LPAREN LBRACK COMMA ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE STATIC STAR RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE STATIC STAR RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE STATIC STAR RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE STATIC STAR RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## TYPEDEF_NAME list(declaration_specifier_no_type)
+## typedef_name list(declaration_specifier_no_type)
##
-translation_unit_file: INT VAR_NAME LPAREN VOLATILE TYPEDEF_NAME XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN VOLATILE PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 259.
+## Ends in an error in state: 219.
##
-## declaration_specifiers(parameter_declaration) -> list(declaration_specifier_no_type) declaration_specifier_no_type TYPEDEF_NAME list(declaration_specifier_no_type) . [ VAR_NAME TYPEDEF_NAME STAR RPAREN LPAREN LBRACK COMMA ]
-## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . declaration_specifier_no_type [ VOLATILE VAR_NAME TYPEDEF_NAME STATIC STAR RPAREN RESTRICT REGISTER PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers(parameter_declaration) -> rlist(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type) . [ STAR RPAREN PRE_NAME LPAREN LBRACK COMMA ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE STATIC STAR RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE STATIC STAR RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE STATIC STAR RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE STATIC STAR RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) declaration_specifier_no_type TYPEDEF_NAME list(declaration_specifier_no_type)
+## rlist(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type)
##
-translation_unit_file: INT VAR_NAME LPAREN INT XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT XOR_ASSIGN
##
-## Ends in an error in state: 250.
+## Ends in an error in state: 202.
##
-## declaration_specifiers(parameter_declaration) -> list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ VAR_NAME TYPEDEF_NAME STAR RPAREN LPAREN LBRACK COMMA ]
-## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME STRUCT STATIC STAR SIGNED SHORT RPAREN RESTRICT REGISTER PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA CHAR AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers(parameter_declaration) -> type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ STAR RPAREN PRE_NAME LPAREN LBRACK COMMA ]
+## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC STAR SIGNED SHORT RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
+## type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
+##
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN VOLATILE INT XOR_ASSIGN
+##
+## Ends in an error in state: 221.
+##
+## declaration_specifiers(parameter_declaration) -> rlist(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ STAR RPAREN PRE_NAME LPAREN LBRACK COMMA ]
+## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC STAR SIGNED SHORT RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA CHAR AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## rlist(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
##
# Analogous to the above situation, except this time, we are in the
@@ -2131,16 +2267,25 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: VOLATILE VAR_NAME
+translation_unit_file: VOLATILE XOR_ASSIGN
##
-## Ends in an error in state: 363.
+## Ends in an error in state: 402.
##
-## declaration_specifiers(declaration(external_declaration)) -> list(declaration_specifier_no_type) declaration_specifier_no_type . TYPEDEF_NAME list(declaration_specifier_no_type) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## declaration_specifiers_typedef -> list(declaration_specifier_no_type) declaration_specifier_no_type . TYPEDEF_NAME list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) declaration_specifier_no_type . [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PACKED LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers(declaration(external_declaration)) -> rlist(declaration_specifier_no_type) . typedef_name list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers(declaration(external_declaration)) -> rlist(declaration_specifier_no_type) . type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) . TYPEDEF list(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) . typedef_name list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) . TYPEDEF list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) . type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) TYPEDEF list(declaration_specifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) declaration_specifier_no_type
+## rlist(declaration_specifier_no_type)
+##
+## WARNING: This example involves spurious reductions.
+## This implies that, although the LR(1) items shown above provide an
+## accurate view of the past (what has been recognized so far), they
+## may provide an INCOMPLETE view of the future (what was expected next).
+## In state 209, spurious reduction of production rlist(declaration_specifier_no_type) -> type_qualifier_noattr
##
# We have seen some specifiers or qualifiers. We have probably seen at least
@@ -2167,18 +2312,26 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE VOLATILE VAR_NAME
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE VOLATILE XOR_ASSIGN
##
-## Ends in an error in state: 424.
+## Ends in an error in state: 511.
##
-## declaration_specifiers(declaration(block_item)) -> list(declaration_specifier_no_type) declaration_specifier_no_type . TYPEDEF_NAME list(declaration_specifier_no_type) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## declaration_specifiers_typedef -> list(declaration_specifier_no_type) declaration_specifier_no_type . TYPEDEF_NAME list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) declaration_specifier_no_type . [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PACKED LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers(declaration(block_item)) -> rlist(declaration_specifier_no_type) . typedef_name list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers(declaration(block_item)) -> rlist(declaration_specifier_no_type) . type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) . TYPEDEF list(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) . typedef_name list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) . TYPEDEF list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) . type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) TYPEDEF list(declaration_specifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) declaration_specifier_no_type
+## rlist(declaration_specifier_no_type)
+##
+## WARNING: This example involves spurious reductions.
+## This implies that, although the LR(1) items shown above provide an
+## accurate view of the past (what has been recognized so far), they
+## may provide an INCOMPLETE view of the future (what was expected next).
+## In state 209, spurious reduction of production rlist(declaration_specifier_no_type) -> type_qualifier_noattr
##
-
# Identical to the previous one, except we are not at the top level,
# so we know this cannot be the beginning of a function definition.
@@ -2190,38 +2343,55 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE TYPEDEF_NAME VOLATILE XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE PRE_NAME TYPEDEF_NAME VOLATILE XOR_ASSIGN
##
-## Ends in an error in state: 393.
+## Ends in an error in state: 508.
##
-## declaration_specifiers(declaration(block_item)) -> TYPEDEF_NAME list(declaration_specifier_no_type) . [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## declaration_specifiers_typedef -> TYPEDEF_NAME list(declaration_specifier_no_type) . TYPEDEF list(declaration_specifier_no_type) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . declaration_specifier_no_type [ VOLATILE VAR_NAME TYPEDEF_NAME TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers(declaration(block_item)) -> typedef_name list(declaration_specifier_no_type) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> typedef_name list(declaration_specifier_no_type) . TYPEDEF list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## TYPEDEF_NAME list(declaration_specifier_no_type)
+## typedef_name list(declaration_specifier_no_type)
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE VOLATILE TYPEDEF_NAME XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE VOLATILE PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 426.
+## Ends in an error in state: 513.
##
-## declaration_specifiers(declaration(block_item)) -> list(declaration_specifier_no_type) declaration_specifier_no_type TYPEDEF_NAME list(declaration_specifier_no_type) . [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## declaration_specifiers_typedef -> list(declaration_specifier_no_type) declaration_specifier_no_type TYPEDEF_NAME list(declaration_specifier_no_type) . TYPEDEF list(declaration_specifier_no_type) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . declaration_specifier_no_type [ VOLATILE VAR_NAME TYPEDEF_NAME TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers(declaration(block_item)) -> rlist(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type) . TYPEDEF list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) declaration_specifier_no_type TYPEDEF_NAME list(declaration_specifier_no_type)
+## rlist(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type)
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE INT XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE INT XOR_ASSIGN
##
-## Ends in an error in state: 423.
+## Ends in an error in state: 510.
##
-## declaration_specifiers(declaration(block_item)) -> list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## declaration_specifiers_typedef -> list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . TYPEDEF list(declaration_specifier_no_typedef_name) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## declaration_specifiers(declaration(block_item)) -> type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . TYPEDEF list(declaration_specifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## list(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
+## type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
+##
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE VOLATILE INT XOR_ASSIGN
+##
+## Ends in an error in state: 515.
+##
+## declaration_specifiers(declaration(block_item)) -> rlist(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> rlist(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . TYPEDEF list(declaration_specifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## rlist(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
##
# This is analogous to the error sentence TYPEDEF_NAME VOLATILE XOR_ASSIGN,
@@ -2252,11 +2422,11 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: UNION LBRACE TYPEDEF_NAME XOR_ASSIGN
+translation_unit_file: UNION LBRACE PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 225.
+## Ends in an error in state: 177.
##
-## struct_declaration -> specifier_qualifier_list(struct_declaration) . option(struct_declarator_list) SEMICOLON [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME STRUCT SIGNED SHORT RESTRICT RBRACE PACKED LONG INT FLOAT ENUM DOUBLE CONST CHAR ATTRIBUTE ALIGNAS ]
+## struct_declaration -> specifier_qualifier_list(struct_declaration) . option(struct_declarator_list) SEMICOLON [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT SIGNED SHORT RESTRICT RBRACE PRE_NAME PACKED LONG INT FLOAT ENUM DOUBLE CONST CHAR ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
## specifier_qualifier_list(struct_declaration)
@@ -2265,7 +2435,7 @@ translation_unit_file: UNION LBRACE TYPEDEF_NAME XOR_ASSIGN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 313, spurious reduction of production specifier_qualifier_list(struct_declaration) -> option(type_qualifier_list) TYPEDEF_NAME option(type_qualifier_list)
+## In state 163, spurious reduction of production specifier_qualifier_list(struct_declaration) -> typedef_name option(type_qualifier_list)
##
# We have (spuriously) recognized a specifier_qualifier_list,
@@ -2299,7 +2469,7 @@ at this point, one of the following is expected:
translation_unit_file: UNION LBRACE LONG COLON CONSTANT RPAREN
##
-## Ends in an error in state: 301.
+## Ends in an error in state: 282.
##
## option(struct_declarator_list) -> struct_declarator_list . [ SEMICOLON ]
## struct_declarator_list -> struct_declarator_list . COMMA struct_declarator [ SEMICOLON COMMA ]
@@ -2311,21 +2481,21 @@ translation_unit_file: UNION LBRACE LONG COLON CONSTANT RPAREN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 31, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 306, spurious reduction of production struct_declarator -> option(declarator) COLON conditional_expression
-## In state 308, spurious reduction of production struct_declarator_list -> struct_declarator
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 57, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 287, spurious reduction of production struct_declarator -> option(declarator) COLON conditional_expression
+## In state 289, spurious reduction of production struct_declarator_list -> struct_declarator
##
# We have seen a non-empty struct_declarator_list.
@@ -2343,7 +2513,7 @@ then at this point, a semicolon ';' is expected.
translation_unit_file: UNION LBRACE INT COLON XOR_ASSIGN
##
-## Ends in an error in state: 305.
+## Ends in an error in state: 286.
##
## struct_declarator -> option(declarator) COLON . conditional_expression [ SEMICOLON COMMA ]
##
@@ -2356,9 +2526,9 @@ At this point, a constant expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: UNION LBRACE INT VAR_NAME COMMA XOR_ASSIGN
+translation_unit_file: UNION LBRACE INT PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 302.
+## Ends in an error in state: 283.
##
## struct_declarator_list -> struct_declarator_list COMMA . struct_declarator [ SEMICOLON COMMA ]
##
@@ -2371,9 +2541,9 @@ At this point, a struct declarator is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: UNION LBRACE INT VAR_NAME RPAREN
+translation_unit_file: UNION LBRACE INT PRE_NAME VAR_NAME RPAREN
##
-## Ends in an error in state: 307.
+## Ends in an error in state: 288.
##
## option(declarator) -> declarator . [ COLON ]
## struct_declarator -> declarator . [ SEMICOLON COMMA ]
@@ -2385,8 +2555,9 @@ translation_unit_file: UNION LBRACE INT VAR_NAME RPAREN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 275, spurious reduction of production attribute_specifier_list ->
-## In state 281, spurious reduction of production declarator -> direct_declarator attribute_specifier_list
+## In state 250, spurious reduction of production declarator_noattrend -> direct_declarator
+## In state 255, spurious reduction of production attribute_specifier_list ->
+## In state 256, spurious reduction of production declarator -> declarator_noattrend attribute_specifier_list
##
# Assuming the declarator so far is complete, we expect
@@ -2409,14 +2580,14 @@ then at this point, one of the following is expected:
translation_unit_file: UNION LBRACE VOLATILE ADD_ASSIGN
##
-## Ends in an error in state: 311.
+## Ends in an error in state: 171.
##
-## specifier_qualifier_list(struct_declaration) -> option(type_qualifier_list) . TYPEDEF_NAME option(type_qualifier_list) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN COLON ]
-## specifier_qualifier_list(struct_declaration) -> option(type_qualifier_list) . type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN COLON ]
-## type_qualifier_list -> option(type_qualifier_list) . type_qualifier [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME STRUCT SIGNED SHORT RESTRICT PACKED LONG INT FLOAT ENUM DOUBLE CONST CHAR ATTRIBUTE ALIGNAS ]
+## option(type_qualifier_list) -> type_qualifier_list . [ VOLATILE RESTRICT PACKED CONST ATTRIBUTE ALIGNAS ]
+## specifier_qualifier_list(struct_declaration) -> type_qualifier_list . typedef_name option(type_qualifier_list) [ STAR SEMICOLON PRE_NAME LPAREN COLON ]
+## specifier_qualifier_list(struct_declaration) -> type_qualifier_list . type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN COLON ]
##
## The known suffix of the stack is as follows:
-## option(type_qualifier_list)
+## type_qualifier_list
##
# A list of qualifiers has been read.
@@ -2432,10 +2603,10 @@ At this point, one of the following is expected:
translation_unit_file: UNION LBRACE XOR_ASSIGN
##
-## Ends in an error in state: 222.
+## Ends in an error in state: 160.
##
-## struct_declaration_list -> struct_declaration_list . struct_declaration [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME STRUCT SIGNED SHORT RESTRICT RBRACE PACKED LONG INT FLOAT ENUM DOUBLE CONST CHAR ATTRIBUTE ALIGNAS ]
-## struct_or_union_specifier -> struct_or_union attribute_specifier_list option(other_identifier) LBRACE struct_declaration_list . RBRACE [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
+## struct_declaration_list -> struct_declaration_list . struct_declaration [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT SIGNED SHORT RESTRICT RBRACE PRE_NAME PACKED LONG INT FLOAT ENUM DOUBLE CONST CHAR ATTRIBUTE ALIGNAS ]
+## struct_or_union_specifier -> struct_or_union attribute_specifier_list option(other_identifier) LBRACE struct_declaration_list . RBRACE [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
## struct_or_union attribute_specifier_list option(other_identifier) LBRACE struct_declaration_list
@@ -2452,15 +2623,20 @@ At this point, one of the following is expected:
translation_unit_file: UNION XOR_ASSIGN
##
-## Ends in an error in state: 219.
+## Ends in an error in state: 157.
##
-## attribute_specifier_list -> attribute_specifier_list . attribute_specifier [ VAR_NAME TYPEDEF_NAME PACKED LBRACE ATTRIBUTE ALIGNAS ]
-## struct_or_union_specifier -> struct_or_union attribute_specifier_list . option(other_identifier) LBRACE struct_declaration_list RBRACE [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
-## struct_or_union_specifier -> struct_or_union attribute_specifier_list . general_identifier [ VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
+## struct_or_union_specifier -> struct_or_union attribute_specifier_list . option(other_identifier) LBRACE struct_declaration_list RBRACE [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
+## struct_or_union_specifier -> struct_or_union attribute_specifier_list . general_identifier [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF STRUCT STATIC STAR SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
## struct_or_union attribute_specifier_list
##
+## WARNING: This example involves spurious reductions.
+## This implies that, although the LR(1) items shown above provide an
+## accurate view of the past (what has been recognized so far), they
+## may provide an INCOMPLETE view of the future (what was expected next).
+## In state 156, spurious reduction of production attribute_specifier_list ->
+##
# gcc expects '{'.
# clang gives a mysterious message: "declaration of anonymous union must be a definition".
@@ -2474,25 +2650,26 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT LPAREN VAR_NAME SEMICOLON
+translation_unit_file: INT LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 283.
+## Ends in an error in state: 259.
##
-## direct_declarator -> LPAREN declarator . RPAREN [ SEMICOLON RPAREN PACKED LPAREN LBRACK LBRACE EQ COMMA COLON ATTRIBUTE ALIGNAS ]
+## direct_declarator -> LPAREN save_context declarator . RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## LPAREN declarator
+## LPAREN save_context declarator
##
## WARNING: This example involves spurious reductions.
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 275, spurious reduction of production attribute_specifier_list ->
-## In state 281, spurious reduction of production declarator -> direct_declarator attribute_specifier_list
+## In state 250, spurious reduction of production declarator_noattrend -> direct_declarator
+## In state 255, spurious reduction of production attribute_specifier_list ->
+## In state 256, spurious reduction of production declarator -> declarator_noattrend attribute_specifier_list
##
-Up to this point, an opening parenthesis and a declarator have been recognized:
- $1 $0
+Up to this point, a declarator have been recognized:
+ $0
If this declarator is complete,
then at this point, a closing parenthesis ')' is expected.
@@ -2500,12 +2677,12 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: INT LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 226.
+## Ends in an error in state: 183.
##
-## direct_declarator -> LPAREN . declarator RPAREN [ SEMICOLON RPAREN PACKED LPAREN LBRACK LBRACE EQ COMMA COLON ATTRIBUTE ALIGNAS ]
+## direct_declarator -> LPAREN save_context . declarator RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## LPAREN
+## LPAREN save_context
##
# clang and gcc expect identifier or '(', as usual.
@@ -2515,53 +2692,15 @@ At this point, a declarator is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT STAR RPAREN
-##
-## Ends in an error in state: 658.
-##
-## declarator -> list(pointer1) STAR option(type_qualifier_list) . direct_declarator attribute_specifier_list [ SEMICOLON EQ COMMA ]
-## function_definition_begin -> declaration_specifiers(declaration(external_declaration)) list(pointer1) STAR option(type_qualifier_list) . direct_declarator [ LBRACE ]
-## function_definition_begin -> declaration_specifiers(declaration(external_declaration)) list(pointer1) STAR option(type_qualifier_list) . direct_declarator LPAREN identifier_list RPAREN open_context declaration_list [ LBRACE ]
-## list(pointer1) -> list(pointer1) STAR option(type_qualifier_list) . [ STAR ]
-## type_qualifier_list -> option(type_qualifier_list) . type_qualifier [ VOLATILE VAR_NAME TYPEDEF_NAME STAR RESTRICT PACKED LPAREN CONST ATTRIBUTE ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## declaration_specifiers(declaration(external_declaration)) list(pointer1) STAR option(type_qualifier_list)
-##
-
-# We are definitely at the beginning of a function definition, and have seen
-# a return type (declaration_specifiers) and a pointer1 (i.e. at least one STAR).
-# Here, either this pointer could be continued, or we expect a direct_declarator.
-
-# clang and gcc expect identifier or '(', as usual.
-
-Ill-formed function definition.
-At this point, one of the following is expected:
- a type qualifier; or
- a star '*', possibly followed with type qualifiers; or
- a direct declarator.
-
-# ------------------------------------------------------------------------------
-
-translation_unit_file: INT STAR VAR_NAME LPAREN VAR_NAME XOR_ASSIGN
-##
-## Ends in an error in state: 662.
-##
-## function_definition_begin -> declaration_specifiers(declaration(external_declaration)) list(pointer1) STAR option(type_qualifier_list) direct_declarator LPAREN identifier_list . RPAREN open_context declaration_list [ LBRACE ]
-## identifier_list -> identifier_list . COMMA VAR_NAME [ RPAREN COMMA ]
-##
-## The known suffix of the stack is as follows:
-## declaration_specifiers(declaration(external_declaration)) list(pointer1) STAR option(type_qualifier_list) direct_declarator LPAREN identifier_list
-##
-translation_unit_file: INT VAR_NAME LPAREN VAR_NAME XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME XOR_ASSIGN
##
-## Ends in an error in state: 671.
+## Ends in an error in state: 276.
##
-## function_definition_begin -> declaration_specifiers(declaration(external_declaration)) direct_declarator LPAREN identifier_list . RPAREN open_context declaration_list [ LBRACE ]
-## identifier_list -> identifier_list . COMMA VAR_NAME [ RPAREN COMMA ]
+## identifier_list -> identifier_list . COMMA PRE_NAME VAR_NAME [ RPAREN COMMA ]
+## option(identifier_list) -> identifier_list . [ RPAREN ]
##
## The known suffix of the stack is as follows:
-## declaration_specifiers(declaration(external_declaration)) direct_declarator LPAREN identifier_list
+## identifier_list
##
Ill-formed K&R function definition.
@@ -2572,28 +2711,19 @@ then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT STAR VAR_NAME LPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 660.
+## Ends in an error in state: 189.
##
-## direct_declarator -> direct_declarator LPAREN . open_context option(parameter_type_list) save_contexts_stk close_context RPAREN [ SEMICOLON PACKED LPAREN LBRACK LBRACE EQ COMMA ATTRIBUTE ALIGNAS ]
-## function_definition_begin -> declaration_specifiers(declaration(external_declaration)) list(pointer1) STAR option(type_qualifier_list) direct_declarator LPAREN . identifier_list RPAREN open_context declaration_list [ LBRACE ]
+## context_parameter_type_list -> save_context . parameter_type_list save_context [ RPAREN ]
+## direct_declarator -> direct_declarator LPAREN save_context . option(identifier_list) RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## declaration_specifiers(declaration(external_declaration)) list(pointer1) STAR option(type_qualifier_list) direct_declarator LPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 670.
-##
-## direct_declarator -> direct_declarator LPAREN . open_context option(parameter_type_list) save_contexts_stk close_context RPAREN [ SEMICOLON PACKED LPAREN LBRACK LBRACE EQ COMMA ATTRIBUTE ALIGNAS ]
-## function_definition_begin -> declaration_specifiers(declaration(external_declaration)) direct_declarator LPAREN . identifier_list RPAREN open_context declaration_list [ LBRACE ]
-##
-## The known suffix of the stack is as follows:
-## declaration_specifiers(declaration(external_declaration)) direct_declarator LPAREN
+## direct_declarator LPAREN save_context
##
# Ignore K&R syntax, just request ANSI syntax.
+
# Ignore the distinction between parameter-type-list and parameter-list.
# clang expects a parameter declarator (which is incomplete).
@@ -2605,13 +2735,14 @@ followed with a closing parenthesis ')', is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: TYPEDEF INT STAR RPAREN
+translation_unit_file: INT STAR RPAREN
##
-## Ends in an error in state: 229.
+## Ends in an error in state: 186.
##
-## declarator -> list(pointer1) STAR option(type_qualifier_list) . direct_declarator attribute_specifier_list [ SEMICOLON RPAREN EQ COMMA COLON ]
+## declarator_noattrend -> list(pointer1) STAR option(type_qualifier_list) . direct_declarator [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LONG LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
## list(pointer1) -> list(pointer1) STAR option(type_qualifier_list) . [ STAR ]
-## type_qualifier_list -> option(type_qualifier_list) . type_qualifier [ VOLATILE VAR_NAME TYPEDEF_NAME STAR RESTRICT PACKED LPAREN CONST ATTRIBUTE ALIGNAS ]
+## type_qualifier_list -> option(type_qualifier_list) . type_qualifier_noattr [ VOLATILE STAR RESTRICT PRE_NAME PACKED LPAREN CONST ATTRIBUTE ALIGNAS ]
+## type_qualifier_list -> option(type_qualifier_list) . attribute_specifier [ VOLATILE STAR RESTRICT PRE_NAME PACKED LPAREN CONST ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
## list(pointer1) STAR option(type_qualifier_list)
@@ -2637,9 +2768,9 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: TYPEDEF INT VAR_NAME XOR_ASSIGN
+translation_unit_file: TYPEDEF INT PRE_NAME VAR_NAME XOR_ASSIGN
##
-## Ends in an error in state: 515.
+## Ends in an error in state: 527.
##
## option(typedef_declarator_list) -> typedef_declarator_list . [ SEMICOLON ]
## typedef_declarator_list -> typedef_declarator_list . COMMA typedef_declarator [ SEMICOLON COMMA ]
@@ -2651,14 +2782,15 @@ translation_unit_file: TYPEDEF INT VAR_NAME XOR_ASSIGN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 275, spurious reduction of production attribute_specifier_list ->
-## In state 281, spurious reduction of production declarator -> direct_declarator attribute_specifier_list
-## In state 519, spurious reduction of production declare_typename(fst(declarator)) -> declarator
-## In state 518, spurious reduction of production typedef_declarator -> declare_typename(fst(declarator))
-## In state 520, spurious reduction of production typedef_declarator_list -> typedef_declarator
+## In state 250, spurious reduction of production declarator_noattrend -> direct_declarator
+## In state 255, spurious reduction of production attribute_specifier_list ->
+## In state 256, spurious reduction of production declarator -> declarator_noattrend attribute_specifier_list
+## In state 531, spurious reduction of production declare_typename(declarator) -> declarator
+## In state 530, spurious reduction of production typedef_declarator -> declare_typename(declarator)
+## In state 532, spurious reduction of production typedef_declarator_list -> typedef_declarator
##
-# Because attribute_specifier_list and declarator have been marked
+# Because attribute_specifier_list, declarator and declarator_noattrend have been marked
# %on_error_reduce, we perform several spurious reductions and end up here.
# Which is good, because the context is clear.
@@ -2679,9 +2811,9 @@ then at this point, a semicolon ';' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: TYPEDEF INT VAR_NAME COMMA XOR_ASSIGN
+translation_unit_file: TYPEDEF INT PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 516.
+## Ends in an error in state: 528.
##
## typedef_declarator_list -> typedef_declarator_list COMMA . typedef_declarator [ SEMICOLON COMMA ]
##
@@ -2693,23 +2825,14 @@ At this point, a declarator is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: TYPEDEF INT VAR_NAME LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 233.
-##
-## direct_declarator -> direct_declarator LPAREN open_context . option(parameter_type_list) save_contexts_stk close_context RPAREN [ SEMICOLON RPAREN PACKED LPAREN LBRACK LBRACE EQ COMMA COLON ATTRIBUTE ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## direct_declarator LPAREN open_context
-##
translation_unit_file: ALIGNAS LPAREN INT LPAREN RPAREN LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 262.
+## Ends in an error in state: 244.
##
-## in_context(option(parameter_type_list)) -> open_context . option(parameter_type_list) close_context [ RPAREN ]
+## direct_abstract_declarator -> direct_abstract_declarator LPAREN . option(context_parameter_type_list) RPAREN [ RPAREN LPAREN LBRACK COMMA ]
##
## The known suffix of the stack is as follows:
-## open_context
+## direct_abstract_declarator LPAREN
##
At this point, a list of parameter declarations,
@@ -2717,38 +2840,29 @@ followed with a closing parenthesis ')', is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM CONST XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM CONST XOR_ASSIGN
##
-## Ends in an error in state: 453.
+## Ends in an error in state: 443.
##
## asm_attributes -> CONST . asm_attributes [ LPAREN ]
##
## The known suffix of the stack is as follows:
## CONST
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM VOLATILE XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM VOLATILE XOR_ASSIGN
##
-## Ends in an error in state: 452.
+## Ends in an error in state: 442.
##
## asm_attributes -> VOLATILE . asm_attributes [ LPAREN ]
##
## The known suffix of the stack is as follows:
## VOLATILE
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM XOR_ASSIGN
##
-## Ends in an error in state: 613.
-##
-## asm_statement(nop) -> ASM . asm_attributes LPAREN string_literals_list asm_arguments RPAREN SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## ASM
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO ASM XOR_ASSIGN
-##
-## Ends in an error in state: 451.
+## Ends in an error in state: 441.
##
-## asm_statement(close_context) -> ASM . asm_attributes LPAREN string_literals_list asm_arguments RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## asm_statement -> ASM . asm_attributes LPAREN string_literals_list asm_arguments RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ASM
@@ -2761,9 +2875,9 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON COLON COLON STRING_LITERAL COMMA XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON COLON COLON STRING_LITERAL COMMA XOR_ASSIGN
##
-## Ends in an error in state: 477.
+## Ends in an error in state: 467.
##
## asm_flags -> asm_flags COMMA . string_literals_list [ RPAREN COMMA ]
##
@@ -2773,9 +2887,9 @@ translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN
# We are in the clobber list.
# We have seen a comma, so we expect a string literal.
# first(asm_flags) = STRING_LITERAL
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON COLON COLON XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON COLON COLON XOR_ASSIGN
##
-## Ends in an error in state: 474.
+## Ends in an error in state: 464.
##
## asm_arguments -> COLON asm_operands COLON asm_operands COLON . asm_flags [ RPAREN ]
##
@@ -2793,9 +2907,9 @@ Examples of clobbered resources:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON COLON COLON STRING_LITERAL XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON COLON COLON STRING_LITERAL XOR_ASSIGN
##
-## Ends in an error in state: 476.
+## Ends in an error in state: 466.
##
## asm_arguments -> COLON asm_operands COLON asm_operands COLON asm_flags . [ RPAREN ]
## asm_flags -> asm_flags . COMMA string_literals_list [ RPAREN COMMA ]
@@ -2807,7 +2921,7 @@ translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 475, spurious reduction of production asm_flags -> string_literals_list
+## In state 465, spurious reduction of production asm_flags -> string_literals_list
##
# Let's ignore the possibility of concatenating string literals.
@@ -2822,9 +2936,9 @@ then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON STRING_LITERAL LPAREN CONSTANT RPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON STRING_LITERAL LPAREN CONSTANT RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 471.
+## Ends in an error in state: 461.
##
## asm_arguments -> COLON asm_operands . [ RPAREN ]
## asm_arguments -> COLON asm_operands . COLON asm_operands [ RPAREN ]
@@ -2837,7 +2951,7 @@ translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 463, spurious reduction of production asm_operands -> asm_operands_ne
+## In state 453, spurious reduction of production asm_operands -> asm_operands_ne
##
# We have seen one COLON, hence the outputs. (The list of outputs may be empty.)
@@ -2852,9 +2966,9 @@ then at this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON COLON XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON COLON XOR_ASSIGN
##
-## Ends in an error in state: 473.
+## Ends in an error in state: 463.
##
## asm_arguments -> COLON asm_operands COLON asm_operands . [ RPAREN ]
## asm_arguments -> COLON asm_operands COLON asm_operands . COLON asm_flags [ RPAREN ]
@@ -2866,7 +2980,7 @@ translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 472, spurious reduction of production asm_operands ->
+## In state 462, spurious reduction of production asm_operands ->
##
# We have seen two COLONs, hence the outputs and inputs. (The list of inputs may be empty.)
@@ -2884,9 +2998,9 @@ then at this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON LBRACK VAR_NAME RBRACK XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON LBRACK PRE_NAME VAR_NAME RBRACK XOR_ASSIGN
##
-## Ends in an error in state: 466.
+## Ends in an error in state: 456.
##
## asm_operand -> asm_op_name . string_literals_list LPAREN expression RPAREN [ RPAREN COMMA COLON ]
##
@@ -2903,9 +3017,9 @@ At this point, a string literal, representing a constraint, is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON LBRACK VAR_NAME XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON LBRACK PRE_NAME VAR_NAME XOR_ASSIGN
##
-## Ends in an error in state: 461.
+## Ends in an error in state: 451.
##
## asm_op_name -> LBRACK general_identifier . RBRACK [ STRING_LITERAL ]
##
@@ -2918,9 +3032,9 @@ At this point, a closing bracket ']' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON LBRACK XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON LBRACK XOR_ASSIGN
##
-## Ends in an error in state: 460.
+## Ends in an error in state: 450.
##
## asm_op_name -> LBRACK . general_identifier RBRACK [ STRING_LITERAL ]
##
@@ -2933,9 +3047,9 @@ At this point, an identifier is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON STRING_LITERAL LPAREN CONSTANT RPAREN COMMA XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON STRING_LITERAL LPAREN CONSTANT RPAREN COMMA XOR_ASSIGN
##
-## Ends in an error in state: 464.
+## Ends in an error in state: 454.
##
## asm_operands_ne -> asm_operands_ne COMMA . asm_operand [ RPAREN COMMA COLON ]
##
@@ -2950,9 +3064,9 @@ At this point, an assembly operand is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON STRING_LITERAL LPAREN VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON STRING_LITERAL LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 469.
+## Ends in an error in state: 459.
##
## asm_operand -> asm_op_name string_literals_list LPAREN expression . RPAREN [ RPAREN COMMA COLON ]
## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
@@ -2964,21 +3078,21 @@ translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 127, spurious reduction of production expression -> assignment_expression
##
Ill-formed assembly operand.
@@ -2989,9 +3103,9 @@ then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON STRING_LITERAL LPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON STRING_LITERAL LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 468.
+## Ends in an error in state: 458.
##
## asm_operand -> asm_op_name string_literals_list LPAREN . expression RPAREN [ RPAREN COMMA COLON ]
##
@@ -3004,9 +3118,9 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON STRING_LITERAL XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL COLON STRING_LITERAL XOR_ASSIGN
##
-## Ends in an error in state: 467.
+## Ends in an error in state: 457.
##
## asm_operand -> asm_op_name string_literals_list . LPAREN expression RPAREN [ RPAREN COMMA COLON ]
## string_literals_list -> string_literals_list . STRING_LITERAL [ STRING_LITERAL LPAREN ]
@@ -3024,27 +3138,16 @@ followed with an expression and a closing parenthesis ')', is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL XOR_ASSIGN
-##
-## Ends in an error in state: 616.
-##
-## asm_statement(nop) -> ASM asm_attributes LPAREN string_literals_list . asm_arguments RPAREN SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## string_literals_list -> string_literals_list . STRING_LITERAL [ STRING_LITERAL RPAREN COLON ]
-##
-## The known suffix of the stack is as follows:
-## ASM asm_attributes LPAREN string_literals_list
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO ASM LPAREN STRING_LITERAL XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL XOR_ASSIGN
##
-## Ends in an error in state: 458.
+## Ends in an error in state: 448.
##
-## asm_statement(close_context) -> ASM asm_attributes LPAREN string_literals_list . asm_arguments RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## asm_statement -> ASM asm_attributes LPAREN string_literals_list . asm_arguments RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
## string_literals_list -> string_literals_list . STRING_LITERAL [ STRING_LITERAL RPAREN COLON ]
##
## The known suffix of the stack is as follows:
## ASM asm_attributes LPAREN string_literals_list
##
-
# Expecting either one more string literal, or COLON, or RPAREN.
# clang requests ')'. gcc requests ':' or ')'.
@@ -3057,20 +3160,11 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 615.
-##
-## asm_statement(nop) -> ASM asm_attributes LPAREN . string_literals_list asm_arguments RPAREN SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## ASM asm_attributes LPAREN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN XOR_ASSIGN
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO ASM LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 457.
+## Ends in an error in state: 447.
##
-## asm_statement(close_context) -> ASM asm_attributes LPAREN . string_literals_list asm_arguments RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## asm_statement -> ASM asm_attributes LPAREN . string_literals_list asm_arguments RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ASM asm_attributes LPAREN
@@ -3081,192 +3175,80 @@ At this point, a string literal, representing an instruction, is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE BREAK XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE BREAK XOR_ASSIGN
##
-## Ends in an error in state: 611.
+## Ends in an error in state: 439.
##
-## jump_statement(nop) -> BREAK . SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## jump_statement -> BREAK . SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## BREAK
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO BREAK XOR_ASSIGN
-##
-## Ends in an error in state: 449.
-##
-## jump_statement(close_context) -> BREAK close_context . SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE CONTINUE XOR_ASSIGN
##
-## The known suffix of the stack is as follows:
-## BREAK close_context
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE CONTINUE XOR_ASSIGN
-##
-## Ends in an error in state: 606.
+## Ends in an error in state: 434.
##
-## jump_statement(nop) -> CONTINUE . SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## jump_statement -> CONTINUE . SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## CONTINUE
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO CONTINUE XOR_ASSIGN
-##
-## Ends in an error in state: 443.
-##
-## jump_statement(close_context) -> CONTINUE close_context . SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## CONTINUE close_context
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO SEMICOLON WHILE LPAREN VAR_NAME RPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 602.
-##
-## iteration_statement(open_context,statement_finish_close) -> DO open_context statement_finish_close WHILE open_context LPAREN expression RPAREN close_context . SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close WHILE open_context LPAREN expression RPAREN close_context
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO DO SEMICOLON WHILE LPAREN VAR_NAME RPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DO SEMICOLON WHILE LPAREN PRE_NAME VAR_NAME RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 507.
+## Ends in an error in state: 558.
##
-## iteration_statement(nop,statement_finish_close) -> DO open_context statement_finish_close WHILE LPAREN expression RPAREN close_context . SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## iteration_statement -> save_context do_statement1 WHILE LPAREN expression RPAREN . SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close WHILE LPAREN expression RPAREN close_context
+## save_context do_statement1 WHILE LPAREN expression RPAREN
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE GOTO VAR_NAME XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE GOTO PRE_NAME VAR_NAME XOR_ASSIGN
##
-## Ends in an error in state: 585.
+## Ends in an error in state: 430.
##
-## jump_statement(nop) -> GOTO general_identifier . SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## jump_statement -> GOTO general_identifier . SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## GOTO general_identifier
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO GOTO VAR_NAME XOR_ASSIGN
-##
-## Ends in an error in state: 414.
-##
-## jump_statement(close_context) -> GOTO general_identifier close_context . SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## GOTO general_identifier close_context
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN DO SEMICOLON WHILE LPAREN VAR_NAME RPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 539.
-##
-## iteration_statement(nop,statement_finish_close) -> DO open_context statement_finish_close WHILE LPAREN expression RPAREN close_context . SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> DO open_context statement_finish_close WHILE LPAREN expression RPAREN close_context . SEMICOLON [ ELSE ]
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL RPAREN XOR_ASSIGN
##
-## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close WHILE LPAREN expression RPAREN close_context
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL RPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 618.
+## Ends in an error in state: 471.
##
-## asm_statement(nop) -> ASM asm_attributes LPAREN string_literals_list asm_arguments RPAREN . SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## asm_statement -> ASM asm_attributes LPAREN string_literals_list asm_arguments RPAREN . SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ASM asm_attributes LPAREN string_literals_list asm_arguments RPAREN
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO ASM LPAREN STRING_LITERAL RPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 482.
-##
-## asm_statement(close_context) -> ASM asm_attributes LPAREN string_literals_list asm_arguments RPAREN close_context . SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## ASM asm_attributes LPAREN string_literals_list asm_arguments RPAREN close_context
-##
Ill-formed statement.
At this point, a semicolon ';' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE CASE CONSTANT COLON XOR_ASSIGN
-##
-## Ends in an error in state: 610.
-##
-## labeled_statement(statement_finish_noclose) -> CASE conditional_expression COLON . statement_finish_noclose [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## CASE conditional_expression COLON
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DEFAULT COLON XOR_ASSIGN
-##
-## Ends in an error in state: 605.
-##
-## labeled_statement(statement_finish_noclose) -> DEFAULT COLON . statement_finish_noclose [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## DEFAULT COLON
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN CASE CONSTANT COLON XOR_ASSIGN
-##
-## Ends in an error in state: 545.
-##
-## labeled_statement(statement_finish_close) -> CASE conditional_expression COLON . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## labeled_statement(statement_intern_close) -> CASE conditional_expression COLON . statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## CASE conditional_expression COLON
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE CASE CONSTANT COLON XOR_ASSIGN
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO CASE CONSTANT COLON XOR_ASSIGN
+## Ends in an error in state: 438.
##
-## Ends in an error in state: 447.
-##
-## labeled_statement(statement_finish_close) -> CASE conditional_expression COLON . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## labeled_statement -> CASE conditional_expression COLON . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## CASE conditional_expression COLON
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN DEFAULT COLON XOR_ASSIGN
-##
-## Ends in an error in state: 542.
-##
-## labeled_statement(statement_finish_close) -> DEFAULT COLON . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## labeled_statement(statement_intern_close) -> DEFAULT COLON . statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## DEFAULT COLON
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO DEFAULT COLON XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DEFAULT COLON XOR_ASSIGN
##
-## Ends in an error in state: 441.
+## Ends in an error in state: 433.
##
-## labeled_statement(statement_finish_close) -> DEFAULT COLON . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## labeled_statement -> DEFAULT COLON . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## DEFAULT COLON
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN VAR_NAME COLON XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE PRE_NAME VAR_NAME COLON XOR_ASSIGN
##
-## Ends in an error in state: 555.
+## Ends in an error in state: 487.
##
-## labeled_statement(statement_finish_close) -> general_identifier COLON . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## labeled_statement(statement_intern_close) -> general_identifier COLON . statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## general_identifier COLON
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE VAR_NAME COLON XOR_ASSIGN
-##
-## Ends in an error in state: 631.
-##
-## labeled_statement(statement_finish_noclose) -> general_identifier COLON . statement_finish_noclose [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## general_identifier COLON
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO VAR_NAME COLON XOR_ASSIGN
-##
-## Ends in an error in state: 496.
-##
-## labeled_statement(statement_finish_close) -> general_identifier COLON . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## labeled_statement -> general_identifier COLON . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## general_identifier COLON
@@ -3279,66 +3261,11 @@ At this point, a statement is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE CASE CONSTANT SEMICOLON
-##
-## Ends in an error in state: 609.
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE CASE CONSTANT SEMICOLON
##
-## labeled_statement(statement_finish_noclose) -> CASE conditional_expression . COLON statement_finish_noclose [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## CASE conditional_expression
-##
-## WARNING: This example involves spurious reductions.
-## This implies that, although the LR(1) items shown above provide an
-## accurate view of the past (what has been recognized so far), they
-## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 31, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN CASE CONSTANT SEMICOLON
-##
-## Ends in an error in state: 544.
-##
-## labeled_statement(statement_finish_close) -> CASE conditional_expression . COLON statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## labeled_statement(statement_intern_close) -> CASE conditional_expression . COLON statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## CASE conditional_expression
-##
-## WARNING: This example involves spurious reductions.
-## This implies that, although the LR(1) items shown above provide an
-## accurate view of the past (what has been recognized so far), they
-## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 31, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO CASE CONSTANT SEMICOLON
-##
-## Ends in an error in state: 446.
+## Ends in an error in state: 437.
##
-## labeled_statement(statement_finish_close) -> CASE conditional_expression . COLON statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## labeled_statement -> CASE conditional_expression . COLON statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## CASE conditional_expression
@@ -3347,19 +3274,19 @@ translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO CASE CO
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 31, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 57, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
##
Ill-formed labeled statement.
@@ -3370,30 +3297,11 @@ then at this point, a colon ':' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE CASE XOR_ASSIGN
-##
-## Ends in an error in state: 608.
-##
-## labeled_statement(statement_finish_noclose) -> CASE . conditional_expression COLON statement_finish_noclose [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## CASE
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN CASE XOR_ASSIGN
-##
-## Ends in an error in state: 543.
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE CASE XOR_ASSIGN
##
-## labeled_statement(statement_finish_close) -> CASE . conditional_expression COLON statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## labeled_statement(statement_intern_close) -> CASE . conditional_expression COLON statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## CASE
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO CASE XOR_ASSIGN
-##
-## Ends in an error in state: 445.
+## Ends in an error in state: 436.
##
-## labeled_statement(statement_finish_close) -> CASE . conditional_expression COLON statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## labeled_statement -> CASE . conditional_expression COLON statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## CASE
@@ -3404,62 +3312,24 @@ At this point, a constant expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DEFAULT XOR_ASSIGN
-##
-## Ends in an error in state: 604.
-##
-## labeled_statement(statement_finish_noclose) -> DEFAULT . COLON statement_finish_noclose [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## DEFAULT
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO TYPEDEF_NAME XOR_ASSIGN
-##
-## Ends in an error in state: 495.
-##
-## labeled_statement(statement_finish_close) -> general_identifier . COLON statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DEFAULT XOR_ASSIGN
##
-## The known suffix of the stack is as follows:
-## general_identifier
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN DEFAULT XOR_ASSIGN
-##
-## Ends in an error in state: 541.
+## Ends in an error in state: 432.
##
-## labeled_statement(statement_finish_close) -> DEFAULT . COLON statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## labeled_statement(statement_intern_close) -> DEFAULT . COLON statement_intern_close [ ELSE ]
+## labeled_statement -> DEFAULT . COLON statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## DEFAULT
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN TYPEDEF_NAME XOR_ASSIGN
-##
-## Ends in an error in state: 554.
-##
-## labeled_statement(statement_finish_close) -> general_identifier . COLON statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## labeled_statement(statement_intern_close) -> general_identifier . COLON statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## general_identifier
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE VAR_NAME COLON TYPEDEF_NAME XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DO PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 630.
+## Ends in an error in state: 486.
##
-## labeled_statement(statement_finish_noclose) -> general_identifier . COLON statement_finish_noclose [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## labeled_statement -> general_identifier . COLON statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## general_identifier
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO DEFAULT XOR_ASSIGN
-##
-## Ends in an error in state: 440.
-##
-## labeled_statement(statement_finish_close) -> DEFAULT . COLON statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## DEFAULT
-##
# gcc and clang apparently do not allow a TYPEDEF_NAME to be reclassified as a label.
@@ -3468,96 +3338,35 @@ At this point, a colon ':' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO SEMICOLON WHILE LPAREN VAR_NAME SEMICOLON
-##
-## Ends in an error in state: 600.
-##
-## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## iteration_statement(open_context,statement_finish_close) -> DO open_context statement_finish_close WHILE open_context LPAREN expression . RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close WHILE open_context LPAREN expression
-##
-## WARNING: This example involves spurious reductions.
-## This implies that, although the LR(1) items shown above provide an
-## accurate view of the past (what has been recognized so far), they
-## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN DO SEMICOLON WHILE LPAREN VAR_NAME SEMICOLON
-##
-## Ends in an error in state: 537.
-##
-## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## iteration_statement(nop,statement_finish_close) -> DO open_context statement_finish_close WHILE LPAREN expression . RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> DO open_context statement_finish_close WHILE LPAREN expression . RPAREN close_context SEMICOLON [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close WHILE LPAREN expression
-##
-## WARNING: This example involves spurious reductions.
-## This implies that, although the LR(1) items shown above provide an
-## accurate view of the past (what has been recognized so far), they
-## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO DO SEMICOLON WHILE LPAREN VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DO SEMICOLON WHILE LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 505.
+## Ends in an error in state: 557.
##
## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## iteration_statement(nop,statement_finish_close) -> DO open_context statement_finish_close WHILE LPAREN expression . RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## iteration_statement -> save_context do_statement1 WHILE LPAREN expression . RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close WHILE LPAREN expression
+## save_context do_statement1 WHILE LPAREN expression
##
## WARNING: This example involves spurious reductions.
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 127, spurious reduction of production expression -> assignment_expression
##
Ill-formed 'do' ... 'while' statement.
@@ -3568,33 +3377,14 @@ then at this point, a closing parenthesis ')' and a semicolon ';' are expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO SEMICOLON WHILE LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 599.
-##
-## iteration_statement(open_context,statement_finish_close) -> DO open_context statement_finish_close WHILE open_context LPAREN . expression RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close WHILE open_context LPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN DO SEMICOLON WHILE LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 536.
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DO SEMICOLON WHILE LPAREN XOR_ASSIGN
##
-## iteration_statement(nop,statement_finish_close) -> DO open_context statement_finish_close WHILE LPAREN . expression RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> DO open_context statement_finish_close WHILE LPAREN . expression RPAREN close_context SEMICOLON [ ELSE ]
+## Ends in an error in state: 556.
##
-## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close WHILE LPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO DO SEMICOLON WHILE LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 504.
-##
-## iteration_statement(nop,statement_finish_close) -> DO open_context statement_finish_close WHILE LPAREN . expression RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## iteration_statement -> save_context do_statement1 WHILE LPAREN . expression RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close WHILE LPAREN
+## save_context do_statement1 WHILE LPAREN
##
Ill-formed 'do' ... 'while' statement.
@@ -3602,33 +3392,14 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO SEMICOLON WHILE XOR_ASSIGN
-##
-## Ends in an error in state: 598.
-##
-## iteration_statement(open_context,statement_finish_close) -> DO open_context statement_finish_close WHILE open_context . LPAREN expression RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DO SEMICOLON WHILE XOR_ASSIGN
##
-## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close WHILE open_context
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN DO SEMICOLON WHILE XOR_ASSIGN
-##
-## Ends in an error in state: 535.
-##
-## iteration_statement(nop,statement_finish_close) -> DO open_context statement_finish_close WHILE . LPAREN expression RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> DO open_context statement_finish_close WHILE . LPAREN expression RPAREN close_context SEMICOLON [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close WHILE
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO DO SEMICOLON WHILE XOR_ASSIGN
-##
-## Ends in an error in state: 503.
+## Ends in an error in state: 555.
##
-## iteration_statement(nop,statement_finish_close) -> DO open_context statement_finish_close WHILE . LPAREN expression RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## iteration_statement -> save_context do_statement1 WHILE . LPAREN expression RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close WHILE
+## save_context do_statement1 WHILE
##
Ill-formed 'do' ... 'while' statement.
@@ -3636,33 +3407,14 @@ At this point, an opening parenthesis '(' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO SEMICOLON XOR_ASSIGN
-##
-## Ends in an error in state: 596.
-##
-## iteration_statement(open_context,statement_finish_close) -> DO open_context statement_finish_close . WHILE open_context LPAREN expression RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DO SEMICOLON XOR_ASSIGN
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN DO SEMICOLON XOR_ASSIGN
-##
-## Ends in an error in state: 534.
-##
-## iteration_statement(nop,statement_finish_close) -> DO open_context statement_finish_close . WHILE LPAREN expression RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> DO open_context statement_finish_close . WHILE LPAREN expression RPAREN close_context SEMICOLON [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO DO SEMICOLON XOR_ASSIGN
-##
-## Ends in an error in state: 502.
+## Ends in an error in state: 554.
##
-## iteration_statement(nop,statement_finish_close) -> DO open_context statement_finish_close . WHILE LPAREN expression RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## iteration_statement -> save_context do_statement1 . WHILE LPAREN expression RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## DO open_context statement_finish_close
+## save_context do_statement1
##
# Quite nicely, in this case, there is no doubt that the statement is
@@ -3676,33 +3428,14 @@ At this point, a 'while' keyword is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO XOR_ASSIGN
-##
-## Ends in an error in state: 595.
-##
-## iteration_statement(open_context,statement_finish_close) -> DO open_context . statement_finish_close WHILE open_context LPAREN expression RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## DO open_context
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN DO XOR_ASSIGN
-##
-## Ends in an error in state: 432.
-##
-## iteration_statement(nop,statement_finish_close) -> DO open_context . statement_finish_close WHILE LPAREN expression RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> DO open_context . statement_finish_close WHILE LPAREN expression RPAREN close_context SEMICOLON [ ELSE ]
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DO XOR_ASSIGN
##
-## The known suffix of the stack is as follows:
-## DO open_context
+## Ends in an error in state: 550.
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO DO XOR_ASSIGN
-##
-## Ends in an error in state: 439.
-##
-## iteration_statement(nop,statement_finish_close) -> DO open_context . statement_finish_close WHILE LPAREN expression RPAREN close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## do_statement1 -> save_context DO . statement [ WHILE ]
##
## The known suffix of the stack is as follows:
-## DO open_context
+## save_context DO
##
# gcc and clang expect an expression.
@@ -3712,33 +3445,14 @@ At this point, a statement (the loop body) is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE FOR LPAREN SEMICOLON SEMICOLON RPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE FOR LPAREN SEMICOLON SEMICOLON RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 592.
+## Ends in an error in state: 520.
##
-## iteration_statement(open_context,statement_finish_close) -> FOR open_context LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## iteration_statement -> save_context FOR LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## FOR open_context LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN)
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN FOR LPAREN SEMICOLON SEMICOLON RPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 430.
-##
-## iteration_statement(nop,statement_finish_close) -> FOR LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> FOR LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) . statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## FOR LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN)
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO FOR LPAREN SEMICOLON SEMICOLON RPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 437.
-##
-## iteration_statement(nop,statement_finish_close) -> FOR LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## FOR LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN)
+## save_context FOR LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN)
##
Ill-formed 'for' statement.
@@ -3746,12 +3460,12 @@ At this point, a statement (the loop body) is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE FOR LPAREN SEMICOLON SEMICOLON VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE FOR LPAREN SEMICOLON SEMICOLON PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 510.
+## Ends in an error in state: 522.
##
## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## optional(expression,RPAREN) -> expression . RPAREN [ WHILE VAR_NAME TYPEDEF_NAME TILDE SWITCH STRING_LITERAL STAR SIZEOF SEMICOLON RETURN PLUS MINUS LPAREN LBRACE INC IF GOTO FOR DO DEFAULT DEC CONTINUE CONSTANT CASE BUILTIN_VA_ARG BREAK BANG ASM AND ALIGNOF ]
+## optional(expression,RPAREN) -> expression . RPAREN [ WHILE TILDE SWITCH STRING_LITERAL STAR SIZEOF SEMICOLON RETURN PRE_NAME PLUS MINUS LPAREN LBRACE INC IF GOTO FOR DO DEFAULT DEC CONTINUE CONSTANT CASE BUILTIN_VA_ARG BREAK BANG ASM AND ALIGNOF ]
##
## The known suffix of the stack is as follows:
## expression
@@ -3760,21 +3474,21 @@ translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE FOR LPAREN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 127, spurious reduction of production expression -> assignment_expression
##
# The use of optional(expression,RPAREN) tells us that we are in a FOR statement.
@@ -3788,33 +3502,14 @@ then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE FOR LPAREN SEMICOLON SEMICOLON XOR_ASSIGN
-##
-## Ends in an error in state: 591.
-##
-## iteration_statement(open_context,statement_finish_close) -> FOR open_context LPAREN for_statement_header optional(expression,SEMICOLON) . optional(expression,RPAREN) statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## FOR open_context LPAREN for_statement_header optional(expression,SEMICOLON)
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN FOR LPAREN SEMICOLON SEMICOLON XOR_ASSIGN
-##
-## Ends in an error in state: 428.
-##
-## iteration_statement(nop,statement_finish_close) -> FOR LPAREN for_statement_header optional(expression,SEMICOLON) . optional(expression,RPAREN) statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> FOR LPAREN for_statement_header optional(expression,SEMICOLON) . optional(expression,RPAREN) statement_intern_close [ ELSE ]
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE FOR LPAREN SEMICOLON SEMICOLON XOR_ASSIGN
##
-## The known suffix of the stack is as follows:
-## FOR LPAREN for_statement_header optional(expression,SEMICOLON)
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO FOR LPAREN SEMICOLON SEMICOLON XOR_ASSIGN
-##
-## Ends in an error in state: 436.
+## Ends in an error in state: 518.
##
-## iteration_statement(nop,statement_finish_close) -> FOR LPAREN for_statement_header optional(expression,SEMICOLON) . optional(expression,RPAREN) statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## iteration_statement -> save_context FOR LPAREN for_statement_header optional(expression,SEMICOLON) . optional(expression,RPAREN) statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## FOR LPAREN for_statement_header optional(expression,SEMICOLON)
+## save_context FOR LPAREN for_statement_header optional(expression,SEMICOLON)
##
# Expecting the third part of the loop header -- the expression
@@ -3827,33 +3522,14 @@ followed with a closing parenthesis ')', is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE FOR LPAREN SEMICOLON XOR_ASSIGN
-##
-## Ends in an error in state: 590.
-##
-## iteration_statement(open_context,statement_finish_close) -> FOR open_context LPAREN for_statement_header . optional(expression,SEMICOLON) optional(expression,RPAREN) statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## FOR open_context LPAREN for_statement_header
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN FOR LPAREN SEMICOLON XOR_ASSIGN
-##
-## Ends in an error in state: 427.
-##
-## iteration_statement(nop,statement_finish_close) -> FOR LPAREN for_statement_header . optional(expression,SEMICOLON) optional(expression,RPAREN) statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> FOR LPAREN for_statement_header . optional(expression,SEMICOLON) optional(expression,RPAREN) statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## FOR LPAREN for_statement_header
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO FOR LPAREN SEMICOLON XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE FOR LPAREN SEMICOLON XOR_ASSIGN
##
-## Ends in an error in state: 435.
+## Ends in an error in state: 517.
##
-## iteration_statement(nop,statement_finish_close) -> FOR LPAREN for_statement_header . optional(expression,SEMICOLON) optional(expression,RPAREN) statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## iteration_statement -> save_context FOR LPAREN for_statement_header . optional(expression,SEMICOLON) optional(expression,RPAREN) statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## FOR LPAREN for_statement_header
+## save_context FOR LPAREN for_statement_header
##
# Expecting the second part of the loop header -- the controlling expression.
@@ -3865,12 +3541,12 @@ followed with a semicolon ';', is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE FOR LPAREN VAR_NAME RPAREN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE FOR LPAREN PRE_NAME VAR_NAME RPAREN
##
-## Ends in an error in state: 512.
+## Ends in an error in state: 524.
##
## expression -> expression . COMMA assignment_expression [ SEMICOLON COMMA ]
-## optional(expression,SEMICOLON) -> expression . SEMICOLON [ VAR_NAME TILDE STRING_LITERAL STAR SIZEOF SEMICOLON RPAREN PLUS MINUS LPAREN INC DEC CONSTANT BUILTIN_VA_ARG BANG AND ALIGNOF ]
+## optional(expression,SEMICOLON) -> expression . SEMICOLON [ TILDE STRING_LITERAL STAR SIZEOF SEMICOLON RPAREN PRE_NAME PLUS MINUS LPAREN INC DEC CONSTANT BUILTIN_VA_ARG BANG AND ALIGNOF ]
##
## The known suffix of the stack is as follows:
## expression
@@ -3879,21 +3555,21 @@ translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE FOR LPAREN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 127, spurious reduction of production expression -> assignment_expression
##
# At the time of writing, optional(expression,SEMICOLON) is used only in FOR
@@ -3907,33 +3583,14 @@ then at this point, a semicolon ';' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE FOR LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 589.
-##
-## iteration_statement(open_context,statement_finish_close) -> FOR open_context LPAREN . for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## FOR open_context LPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN FOR LPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE FOR LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 417.
-##
-## iteration_statement(nop,statement_finish_close) -> FOR LPAREN . for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> FOR LPAREN . for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## FOR LPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO FOR LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 434.
+## Ends in an error in state: 505.
##
-## iteration_statement(nop,statement_finish_close) -> FOR LPAREN . for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## iteration_statement -> save_context FOR LPAREN . for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## FOR LPAREN
+## save_context FOR LPAREN
##
# gcc and clang say they expect an expression, which is incomplete.
@@ -3947,33 +3604,14 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE FOR XOR_ASSIGN
-##
-## Ends in an error in state: 588.
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE FOR XOR_ASSIGN
##
-## iteration_statement(open_context,statement_finish_close) -> FOR open_context . LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## FOR open_context
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN FOR XOR_ASSIGN
-##
-## Ends in an error in state: 416.
-##
-## iteration_statement(nop,statement_finish_close) -> FOR . LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> FOR . LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## FOR
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO FOR XOR_ASSIGN
-##
-## Ends in an error in state: 433.
+## Ends in an error in state: 504.
##
-## iteration_statement(nop,statement_finish_close) -> FOR . LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## iteration_statement -> save_context FOR . LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN) statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## FOR
+## save_context FOR
##
Ill-formed 'for' statement.
@@ -3981,20 +3619,11 @@ At this point, an opening parenthesis '(' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE GOTO XOR_ASSIGN
-##
-## Ends in an error in state: 584.
-##
-## jump_statement(nop) -> GOTO . general_identifier SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## GOTO
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO GOTO XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE GOTO XOR_ASSIGN
##
-## Ends in an error in state: 412.
+## Ends in an error in state: 429.
##
-## jump_statement(close_context) -> GOTO . general_identifier close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## jump_statement -> GOTO . general_identifier SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## GOTO
@@ -4005,33 +3634,14 @@ At this point, an identifier (a 'goto' label) is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN SEMICOLON ELSE XOR_ASSIGN
-##
-## Ends in an error in state: 628.
-##
-## selection_statement_finish(open_context) -> if_else_statement_begin(open_context) ELSE . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## if_else_statement_begin(open_context) ELSE
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN IF LPAREN VAR_NAME RPAREN SEMICOLON ELSE XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DO IF LPAREN CONSTANT RPAREN SEMICOLON ELSE XOR_ASSIGN
##
## Ends in an error in state: 552.
##
-## selection_statement_finish(nop) -> if_else_statement_begin(nop) ELSE . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## selection_statement_intern_close -> if_else_statement_begin(nop) ELSE . statement_intern_close [ ELSE ]
+## selection_statement -> save_context ifelse_statement1 . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## if_else_statement_begin(nop) ELSE
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO IF LPAREN CONSTANT RPAREN SEMICOLON ELSE XOR_ASSIGN
-##
-## Ends in an error in state: 493.
-##
-## selection_statement_finish(nop) -> if_else_statement_begin(nop) ELSE . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## if_else_statement_begin(nop) ELSE
+## save_context ifelse_statement1
##
Ill-formed 'if' ... 'else' statement.
@@ -4039,25 +3649,15 @@ At this point, a statement is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN IF LPAREN VAR_NAME RPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE IF LPAREN PRE_NAME VAR_NAME RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 403.
+## Ends in an error in state: 501.
##
-## if_else_statement_begin(nop) -> IF LPAREN expression RPAREN save_contexts_stk . statement_intern_close [ ELSE ]
-## selection_statement_finish(nop) -> IF LPAREN expression RPAREN save_contexts_stk . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## ifelse_statement1 -> IF LPAREN expression RPAREN save_context . statement ELSE [ WHILE TILDE SWITCH STRING_LITERAL STAR SIZEOF SEMICOLON RETURN PRE_NAME PLUS MINUS LPAREN LBRACE INC IF GOTO FOR DO DEFAULT DEC CONTINUE CONSTANT CASE BUILTIN_VA_ARG BREAK BANG ASM AND ALIGNOF ]
+## selection_statement -> save_context IF LPAREN expression RPAREN save_context . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## IF LPAREN expression RPAREN save_contexts_stk
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 581.
-##
-## if_else_statement_begin(open_context) -> IF open_context LPAREN expression RPAREN save_contexts_stk . statement_intern_close [ ELSE ]
-## selection_statement_finish(open_context) -> IF open_context LPAREN expression RPAREN save_contexts_stk . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## IF open_context LPAREN expression RPAREN save_contexts_stk
+## save_context IF LPAREN expression RPAREN save_context
##
Ill-formed 'if' statement.
@@ -4065,67 +3665,36 @@ At this point, a statement is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN IF LPAREN VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE IF LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 401.
+## Ends in an error in state: 499.
##
## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## if_else_statement_begin(nop) -> IF LPAREN expression . RPAREN save_contexts_stk statement_intern_close [ ELSE ]
-## selection_statement_finish(nop) -> IF LPAREN expression . RPAREN save_contexts_stk statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## ifelse_statement1 -> IF LPAREN expression . RPAREN save_context statement ELSE [ WHILE TILDE SWITCH STRING_LITERAL STAR SIZEOF SEMICOLON RETURN PRE_NAME PLUS MINUS LPAREN LBRACE INC IF GOTO FOR DO DEFAULT DEC CONTINUE CONSTANT CASE BUILTIN_VA_ARG BREAK BANG ASM AND ALIGNOF ]
+## selection_statement -> save_context IF LPAREN expression . RPAREN save_context statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## IF LPAREN expression
+## save_context IF LPAREN expression
##
## WARNING: This example involves spurious reductions.
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME SEMICOLON
-##
-## Ends in an error in state: 579.
-##
-## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## if_else_statement_begin(open_context) -> IF open_context LPAREN expression . RPAREN save_contexts_stk statement_intern_close [ ELSE ]
-## selection_statement_finish(open_context) -> IF open_context LPAREN expression . RPAREN save_contexts_stk statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## IF open_context LPAREN expression
-##
-## WARNING: This example involves spurious reductions.
-## This implies that, although the LR(1) items shown above provide an
-## accurate view of the past (what has been recognized so far), they
-## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 127, spurious reduction of production expression -> assignment_expression
##
Ill-formed 'if' statement.
@@ -4136,25 +3705,15 @@ then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN IF LPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE IF LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 400.
-##
-## if_else_statement_begin(nop) -> IF LPAREN . expression RPAREN save_contexts_stk statement_intern_close [ ELSE ]
-## selection_statement_finish(nop) -> IF LPAREN . expression RPAREN save_contexts_stk statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## IF LPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 578.
+## Ends in an error in state: 498.
##
-## if_else_statement_begin(open_context) -> IF open_context LPAREN . expression RPAREN save_contexts_stk statement_intern_close [ ELSE ]
-## selection_statement_finish(open_context) -> IF open_context LPAREN . expression RPAREN save_contexts_stk statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## ifelse_statement1 -> IF LPAREN . expression RPAREN save_context statement ELSE [ WHILE TILDE SWITCH STRING_LITERAL STAR SIZEOF SEMICOLON RETURN PRE_NAME PLUS MINUS LPAREN LBRACE INC IF GOTO FOR DO DEFAULT DEC CONTINUE CONSTANT CASE BUILTIN_VA_ARG BREAK BANG ASM AND ALIGNOF ]
+## selection_statement -> save_context IF LPAREN . expression RPAREN save_context statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## IF open_context LPAREN
+## save_context IF LPAREN
##
Ill-formed 'if' statement.
@@ -4162,25 +3721,15 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN IF XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE IF XOR_ASSIGN
##
-## Ends in an error in state: 399.
+## Ends in an error in state: 497.
##
-## if_else_statement_begin(nop) -> IF . LPAREN expression RPAREN save_contexts_stk statement_intern_close [ ELSE ]
-## selection_statement_finish(nop) -> IF . LPAREN expression RPAREN save_contexts_stk statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## ifelse_statement1 -> IF . LPAREN expression RPAREN save_context statement ELSE [ WHILE TILDE SWITCH STRING_LITERAL STAR SIZEOF SEMICOLON RETURN PRE_NAME PLUS MINUS LPAREN LBRACE INC IF GOTO FOR DO DEFAULT DEC CONTINUE CONSTANT CASE BUILTIN_VA_ARG BREAK BANG ASM AND ALIGNOF ]
+## selection_statement -> save_context IF . LPAREN expression RPAREN save_context statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## IF
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF XOR_ASSIGN
-##
-## Ends in an error in state: 577.
-##
-## if_else_statement_begin(open_context) -> IF open_context . LPAREN expression RPAREN save_contexts_stk statement_intern_close [ ELSE ]
-## selection_statement_finish(open_context) -> IF open_context . LPAREN expression RPAREN save_contexts_stk statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## IF open_context
+## save_context IF
##
Ill-formed 'if' statement.
@@ -4188,35 +3737,17 @@ At this point, an opening parenthesis '(' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN SWITCH LPAREN VAR_NAME RPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 411.
-##
-## selection_statement_finish(nop) -> SWITCH LPAREN expression RPAREN . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## selection_statement_intern_close -> SWITCH LPAREN expression RPAREN . statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## SWITCH LPAREN expression RPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE SWITCH LPAREN VAR_NAME RPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 398.
-##
-## selection_statement_finish(open_context) -> SWITCH open_context LPAREN expression RPAREN . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE SWITCH LPAREN PRE_NAME VAR_NAME RPAREN XOR_ASSIGN
##
-## The known suffix of the stack is as follows:
-## SWITCH open_context LPAREN expression RPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO SWITCH LPAREN VAR_NAME RPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 384.
+## Ends in an error in state: 495.
##
-## selection_statement_finish(nop) -> SWITCH LPAREN expression RPAREN . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## selection_statement -> save_context SWITCH LPAREN expression RPAREN . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## SWITCH LPAREN expression RPAREN
+## save_context SWITCH LPAREN expression RPAREN
##
+
# Technically, the body of a 'switch' statement is just a statement,
# but in practice, it should be a list of labeled statements,
# enclosed in braces. (Unless someone is writing a Duff loop...)
@@ -4230,96 +3761,35 @@ enclosed within braces '{' and '}'.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN SWITCH LPAREN VAR_NAME SEMICOLON
-##
-## Ends in an error in state: 410.
-##
-## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## selection_statement_finish(nop) -> SWITCH LPAREN expression . RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## selection_statement_intern_close -> SWITCH LPAREN expression . RPAREN statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## SWITCH LPAREN expression
-##
-## WARNING: This example involves spurious reductions.
-## This implies that, although the LR(1) items shown above provide an
-## accurate view of the past (what has been recognized so far), they
-## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE SWITCH LPAREN VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE SWITCH LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 397.
-##
-## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## selection_statement_finish(open_context) -> SWITCH open_context LPAREN expression . RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## SWITCH open_context LPAREN expression
-##
-## WARNING: This example involves spurious reductions.
-## This implies that, although the LR(1) items shown above provide an
-## accurate view of the past (what has been recognized so far), they
-## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO SWITCH LPAREN VAR_NAME SEMICOLON
-##
-## Ends in an error in state: 383.
+## Ends in an error in state: 494.
##
## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## selection_statement_finish(nop) -> SWITCH LPAREN expression . RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## selection_statement -> save_context SWITCH LPAREN expression . RPAREN statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## SWITCH LPAREN expression
+## save_context SWITCH LPAREN expression
##
## WARNING: This example involves spurious reductions.
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 127, spurious reduction of production expression -> assignment_expression
##
Ill-formed 'switch' statement.
@@ -4330,33 +3800,14 @@ then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN SWITCH LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 409.
-##
-## selection_statement_finish(nop) -> SWITCH LPAREN . expression RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## selection_statement_intern_close -> SWITCH LPAREN . expression RPAREN statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## SWITCH LPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE SWITCH LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 396.
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE SWITCH LPAREN XOR_ASSIGN
##
-## selection_statement_finish(open_context) -> SWITCH open_context LPAREN . expression RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## SWITCH open_context LPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO SWITCH LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 382.
+## Ends in an error in state: 493.
##
-## selection_statement_finish(nop) -> SWITCH LPAREN . expression RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## selection_statement -> save_context SWITCH LPAREN . expression RPAREN statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## SWITCH LPAREN
+## save_context SWITCH LPAREN
##
Ill-formed 'switch' statement.
@@ -4364,33 +3815,14 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN SWITCH XOR_ASSIGN
-##
-## Ends in an error in state: 408.
-##
-## selection_statement_finish(nop) -> SWITCH . LPAREN expression RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## selection_statement_intern_close -> SWITCH . LPAREN expression RPAREN statement_intern_close [ ELSE ]
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE SWITCH XOR_ASSIGN
##
-## The known suffix of the stack is as follows:
-## SWITCH
+## Ends in an error in state: 492.
##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE SWITCH XOR_ASSIGN
-##
-## Ends in an error in state: 395.
-##
-## selection_statement_finish(open_context) -> SWITCH open_context . LPAREN expression RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## selection_statement -> save_context SWITCH . LPAREN expression RPAREN statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## SWITCH open_context
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO SWITCH XOR_ASSIGN
-##
-## Ends in an error in state: 381.
-##
-## selection_statement_finish(nop) -> SWITCH . LPAREN expression RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## SWITCH
+## save_context SWITCH
##
Ill-formed 'switch' statement.
@@ -4398,33 +3830,14 @@ At this point, an opening parenthesis '(' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN WHILE LPAREN VAR_NAME RPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 407.
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE WHILE LPAREN PRE_NAME VAR_NAME RPAREN XOR_ASSIGN
##
-## iteration_statement(nop,statement_finish_close) -> WHILE LPAREN expression RPAREN . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> WHILE LPAREN expression RPAREN . statement_intern_close [ ELSE ]
+## Ends in an error in state: 479.
##
-## The known suffix of the stack is as follows:
-## WHILE LPAREN expression RPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE WHILE LPAREN VAR_NAME RPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 375.
-##
-## iteration_statement(open_context,statement_finish_close) -> WHILE open_context LPAREN expression RPAREN . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## WHILE open_context LPAREN expression RPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO WHILE LPAREN VAR_NAME RPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 379.
-##
-## iteration_statement(nop,statement_finish_close) -> WHILE LPAREN expression RPAREN . statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## iteration_statement -> save_context WHILE LPAREN expression RPAREN . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## WHILE LPAREN expression RPAREN
+## save_context WHILE LPAREN expression RPAREN
##
Ill-formed 'while' statement.
@@ -4432,96 +3845,35 @@ At this point, a statement (the loop body) is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN WHILE LPAREN VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE WHILE LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 406.
+## Ends in an error in state: 478.
##
## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## iteration_statement(nop,statement_finish_close) -> WHILE LPAREN expression . RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> WHILE LPAREN expression . RPAREN statement_intern_close [ ELSE ]
+## iteration_statement -> save_context WHILE LPAREN expression . RPAREN statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## WHILE LPAREN expression
+## save_context WHILE LPAREN expression
##
## WARNING: This example involves spurious reductions.
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE WHILE LPAREN VAR_NAME SEMICOLON
-##
-## Ends in an error in state: 374.
-##
-## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## iteration_statement(open_context,statement_finish_close) -> WHILE open_context LPAREN expression . RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## WHILE open_context LPAREN expression
-##
-## WARNING: This example involves spurious reductions.
-## This implies that, although the LR(1) items shown above provide an
-## accurate view of the past (what has been recognized so far), they
-## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO WHILE LPAREN VAR_NAME SEMICOLON
-##
-## Ends in an error in state: 378.
-##
-## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## iteration_statement(nop,statement_finish_close) -> WHILE LPAREN expression . RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## WHILE LPAREN expression
-##
-## WARNING: This example involves spurious reductions.
-## This implies that, although the LR(1) items shown above provide an
-## accurate view of the past (what has been recognized so far), they
-## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 127, spurious reduction of production expression -> assignment_expression
##
Ill-formed 'while' statement.
@@ -4532,33 +3884,14 @@ then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN WHILE LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 405.
-##
-## iteration_statement(nop,statement_finish_close) -> WHILE LPAREN . expression RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> WHILE LPAREN . expression RPAREN statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## WHILE LPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE WHILE LPAREN XOR_ASSIGN
-##
-## Ends in an error in state: 373.
-##
-## iteration_statement(open_context,statement_finish_close) -> WHILE open_context LPAREN . expression RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## WHILE open_context LPAREN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO WHILE LPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE WHILE LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 377.
+## Ends in an error in state: 477.
##
-## iteration_statement(nop,statement_finish_close) -> WHILE LPAREN . expression RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## iteration_statement -> save_context WHILE LPAREN . expression RPAREN statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## WHILE LPAREN
+## save_context WHILE LPAREN
##
Ill-formed 'while' statement.
@@ -4566,33 +3899,14 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE IF LPAREN VAR_NAME RPAREN WHILE XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE WHILE XOR_ASSIGN
##
-## Ends in an error in state: 404.
-##
-## iteration_statement(nop,statement_finish_close) -> WHILE . LPAREN expression RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## iteration_statement(nop,statement_intern_close) -> WHILE . LPAREN expression RPAREN statement_intern_close [ ELSE ]
-##
-## The known suffix of the stack is as follows:
-## WHILE
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE WHILE XOR_ASSIGN
-##
-## Ends in an error in state: 372.
-##
-## iteration_statement(open_context,statement_finish_close) -> WHILE open_context . LPAREN expression RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## WHILE open_context
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO WHILE XOR_ASSIGN
-##
-## Ends in an error in state: 376.
+## Ends in an error in state: 476.
##
-## iteration_statement(nop,statement_finish_close) -> WHILE . LPAREN expression RPAREN statement_finish_close [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## iteration_statement -> save_context WHILE . LPAREN expression RPAREN statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## WHILE
+## save_context WHILE
##
Ill-formed 'while' statement.
@@ -4600,43 +3914,23 @@ At this point, an opening parenthesis '(' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE LBRACE XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE XOR_ASSIGN
##
-## Ends in an error in state: 575.
+## Ends in an error in state: 420.
##
-## block_item_list -> option(block_item_list) . block_item [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## compound_statement(open_context) -> LBRACE open_context option(block_item_list) . close_context RBRACE [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## block_item_list -> option(block_item_list) . block_item [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## compound_statement -> save_context LBRACE option(block_item_list) . RBRACE [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN EOF ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## LBRACE open_context option(block_item_list)
+## save_context LBRACE option(block_item_list)
##
# We are possibly at the end of a block.
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO LBRACE XOR_ASSIGN
-##
-## Ends in an error in state: 391.
-##
-## block_item_list -> option(block_item_list) . block_item [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## compound_statement(nop) -> LBRACE option(block_item_list) . close_context RBRACE [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## LBRACE option(block_item_list)
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE XOR_ASSIGN
-##
-## Ends in an error in state: 370.
-##
-## block_item_list -> option(block_item_list) . block_item [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-## function_definition -> function_definition_begin LBRACE option(block_item_list) . close_context RBRACE [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC SIGNED SHORT SEMICOLON RESTRICT REGISTER PRAGMA PACKED LONG INT INLINE FLOAT EXTERN EOF ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## function_definition_begin LBRACE option(block_item_list)
-##
-# We are possibly at the end of a function body.
#
-# Note that, because we have used %on_error_reduce to consider some statements
-# as complete even when they could be continued ELSE, we may end up here even
-# though ELSE is permitted. There is nothing we can do about it. We just omit
-# this permitted continuation in our message.
+# Note that, because we have used %on_error_reduce to consider some
+# statements as complete even when they could be continued with ELSE,
+# we may end up here even though ELSE is permitted. There is nothing
+# we can do about it. We just omit this permitted continuation in our
+# message.
# clang and gcc say an expression is expected.
@@ -4648,20 +3942,11 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE RETURN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE RETURN XOR_ASSIGN
##
-## Ends in an error in state: 569.
+## Ends in an error in state: 421.
##
-## jump_statement(nop) -> RETURN . option(expression) SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
-##
-## The known suffix of the stack is as follows:
-## RETURN
-##
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE DO RETURN XOR_ASSIGN
-##
-## Ends in an error in state: 385.
-##
-## jump_statement(close_context) -> RETURN . option(expression) close_context SEMICOLON [ WHILE VOLATILE VOID VAR_NAME UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## jump_statement -> RETURN . option(expression) SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## RETURN
@@ -4676,9 +3961,9 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE STRING_LITERAL RPAREN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE STRING_LITERAL RPAREN
##
-## Ends in an error in state: 389.
+## Ends in an error in state: 424.
##
## expression -> expression . COMMA assignment_expression [ SEMICOLON COMMA ]
## option(expression) -> expression . [ SEMICOLON ]
@@ -4690,23 +3975,23 @@ translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE STRING_LIT
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 32, spurious reduction of production primary_expression -> string_literals_list
-## In state 34, spurious reduction of production postfix_expression -> primary_expression
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 103, spurious reduction of production expression -> assignment_expression
+## In state 58, spurious reduction of production primary_expression -> string_literals_list
+## In state 60, spurious reduction of production postfix_expression -> primary_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 127, spurious reduction of production expression -> assignment_expression
##
Up to this point, an expression has been recognized:
@@ -4716,16 +4001,16 @@ then at this point, a semicolon ';' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE TYPEDEF_NAME XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 392.
+## Ends in an error in state: 561.
##
-## declaration_specifiers(declaration(block_item)) -> TYPEDEF_NAME . list(declaration_specifier_no_type) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## declaration_specifiers_typedef -> TYPEDEF_NAME . list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) [ VAR_NAME TYPEDEF_NAME STAR SEMICOLON LPAREN ]
-## general_identifier -> TYPEDEF_NAME . [ COLON ]
+## declaration_specifiers(declaration(block_item)) -> typedef_name . list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers_typedef -> typedef_name . list(declaration_specifier_no_type) TYPEDEF list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## general_identifier -> typedef_name . [ COLON ]
##
## The known suffix of the stack is as follows:
-## TYPEDEF_NAME
+## typedef_name
##
# We see a type name "foo" at the beginning of a block_item, it seems.
@@ -4751,9 +4036,9 @@ at this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME RPAREN LBRACE VAR_NAME COMMA XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 98.
+## Ends in an error in state: 122.
##
## expression -> expression COMMA . assignment_expression [ SEMICOLON RPAREN RBRACK COMMA COLON ]
##
@@ -4766,9 +4051,9 @@ At this point, an expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME COMMA VAR_NAME RPAREN
+translation_unit_file: INT PRE_NAME VAR_NAME COMMA PRE_NAME VAR_NAME RPAREN
##
-## Ends in an error in state: 526.
+## Ends in an error in state: 538.
##
## init_declarator_list -> init_declarator_list . COMMA init_declarator [ SEMICOLON COMMA ]
## option(init_declarator_list) -> init_declarator_list . [ SEMICOLON ]
@@ -4780,11 +4065,12 @@ translation_unit_file: INT VAR_NAME COMMA VAR_NAME RPAREN
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 275, spurious reduction of production attribute_specifier_list ->
-## In state 281, spurious reduction of production declarator -> direct_declarator attribute_specifier_list
-## In state 294, spurious reduction of production declare_varname(fst(declarator)) -> declarator
-## In state 529, spurious reduction of production init_declarator -> declare_varname(fst(declarator))
-## In state 528, spurious reduction of production init_declarator_list -> init_declarator_list COMMA init_declarator
+## In state 250, spurious reduction of production declarator_noattrend -> direct_declarator
+## In state 546, spurious reduction of production declare_varname(declarator_noattrend) -> declarator_noattrend
+## In state 541, spurious reduction of production save_context ->
+## In state 542, spurious reduction of production attribute_specifier_list ->
+## In state 543, spurious reduction of production init_declarator -> declare_varname(declarator_noattrend) save_context attribute_specifier_list
+## In state 540, spurious reduction of production init_declarator_list -> init_declarator_list COMMA init_declarator
##
Up to this point, a list of declarators has been recognized:
@@ -4794,9 +4080,9 @@ then at this point, a semicolon ';' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME COMMA XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 527.
+## Ends in an error in state: 539.
##
## init_declarator_list -> init_declarator_list COMMA . init_declarator [ SEMICOLON COMMA ]
##
@@ -4809,18 +4095,18 @@ At this point, an init declarator is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LBRACE DOT VAR_NAME EQ ALIGNAS
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE DOT PRE_NAME VAR_NAME EQ ALIGNAS
##
-## Ends in an error in state: 132.
+## Ends in an error in state: 321.
##
## initializer_list -> option(designation) . c_initializer [ RBRACE COMMA ]
##
## The known suffix of the stack is as follows:
## option(designation)
##
-translation_unit_file: INT VAR_NAME EQ LBRACE VAR_NAME COMMA DOT VAR_NAME EQ ALIGNAS
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE PRE_NAME VAR_NAME COMMA DOT PRE_NAME VAR_NAME EQ ALIGNAS
##
-## Ends in an error in state: 136.
+## Ends in an error in state: 325.
##
## initializer_list -> initializer_list COMMA option(designation) . c_initializer [ RBRACE COMMA ]
##
@@ -4833,11 +4119,11 @@ At this point, an initializer is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LBRACE DOT VAR_NAME XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE DOT PRE_NAME VAR_NAME XOR_ASSIGN
##
-## Ends in an error in state: 139.
+## Ends in an error in state: 328.
##
-## designation -> designator_list . EQ [ VAR_NAME TILDE STRING_LITERAL STAR SIZEOF PLUS MINUS LPAREN LBRACE INC DEC CONSTANT BUILTIN_VA_ARG BANG AND ALIGNOF ]
+## designation -> designator_list . EQ [ TILDE STRING_LITERAL STAR SIZEOF PRE_NAME PLUS MINUS LPAREN LBRACE INC DEC CONSTANT BUILTIN_VA_ARG BANG AND ALIGNOF ]
## option(designator_list) -> designator_list . [ LBRACK DOT ]
##
## The known suffix of the stack is as follows:
@@ -4855,9 +4141,9 @@ then at this point, an equals sign '=' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LBRACE DOT XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE DOT XOR_ASSIGN
##
-## Ends in an error in state: 129.
+## Ends in an error in state: 318.
##
## designator -> DOT . general_identifier [ LBRACK EQ DOT ]
##
@@ -4872,9 +4158,9 @@ At this point, the name of a struct or union member is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LBRACE LBRACK VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE LBRACK PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 127.
+## Ends in an error in state: 316.
##
## designator -> LBRACK conditional_expression . RBRACK [ LBRACK EQ DOT ]
##
@@ -4885,19 +4171,19 @@ translation_unit_file: INT VAR_NAME EQ LBRACE LBRACK VAR_NAME SEMICOLON
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 31, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 57, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
##
Ill-formed designator.
@@ -4908,9 +4194,9 @@ then at this point, a closing bracket ']' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LBRACE LBRACK XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE LBRACK XOR_ASSIGN
##
-## Ends in an error in state: 126.
+## Ends in an error in state: 315.
##
## designator -> LBRACK . conditional_expression RBRACK [ LBRACK EQ DOT ]
##
@@ -4923,9 +4209,9 @@ At this point, a constant expression is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LBRACE VAR_NAME COMMA XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 135.
+## Ends in an error in state: 324.
##
## initializer_list -> initializer_list COMMA . option(designation) c_initializer [ RBRACE COMMA ]
## option(COMMA) -> COMMA . [ RBRACE ]
@@ -4944,9 +4230,9 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LBRACE CONSTANT SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE CONSTANT SEMICOLON
##
-## Ends in an error in state: 134.
+## Ends in an error in state: 323.
##
## c_initializer -> LBRACE initializer_list . option(COMMA) RBRACE [ SEMICOLON RBRACE COMMA ]
## initializer_list -> initializer_list . COMMA option(designation) c_initializer [ RBRACE COMMA ]
@@ -4958,22 +4244,22 @@ translation_unit_file: INT VAR_NAME EQ LBRACE CONSTANT SEMICOLON
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
-## In state 138, spurious reduction of production c_initializer -> assignment_expression
-## In state 144, spurious reduction of production initializer_list -> option(designation) c_initializer
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
+## In state 327, spurious reduction of production c_initializer -> assignment_expression
+## In state 333, spurious reduction of production initializer_list -> option(designation) c_initializer
##
# Omitting the fact that the closing brace can be preceded with a comma.
@@ -4986,9 +4272,9 @@ then at this point, a closing brace '}' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ LBRACE XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE XOR_ASSIGN
##
-## Ends in an error in state: 133.
+## Ends in an error in state: 322.
##
## c_initializer -> LBRACE . initializer_list option(COMMA) RBRACE [ SEMICOLON RBRACE COMMA ]
##
@@ -5007,14 +4293,14 @@ followed with an initializer, is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME EQ XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME EQ XOR_ASSIGN
##
-## Ends in an error in state: 530.
+## Ends in an error in state: 544.
##
-## init_declarator -> declare_varname(fst(declarator)) EQ . c_initializer [ SEMICOLON COMMA ]
+## init_declarator -> declare_varname(declarator_noattrend) save_context attribute_specifier_list EQ . c_initializer [ SEMICOLON COMMA ]
##
## The known suffix of the stack is as follows:
-## declare_varname(fst(declarator)) EQ
+## declare_varname(declarator_noattrend) save_context attribute_specifier_list EQ
##
# clang and gcc expect an expression (incomplete).
@@ -5024,11 +4310,11 @@ At this point, an initializer is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LBRACK CONSTANT SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME LBRACK CONSTANT SEMICOLON
##
-## Ends in an error in state: 279.
+## Ends in an error in state: 238.
##
-## optional(assignment_expression,RBRACK) -> assignment_expression . RBRACK [ SEMICOLON RPAREN PACKED LPAREN LBRACK LBRACE EQ COMMA COLON ATTRIBUTE ALIGNAS ]
+## optional(assignment_expression,RBRACK) -> assignment_expression . RBRACK [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
## assignment_expression
@@ -5037,20 +4323,20 @@ translation_unit_file: INT VAR_NAME LBRACK CONSTANT SEMICOLON
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 35, spurious reduction of production unary_expression -> postfix_expression
-## In state 41, spurious reduction of production cast_expression -> unary_expression
-## In state 64, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 58, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 77, spurious reduction of production shift_expression -> additive_expression
-## In state 54, spurious reduction of production relational_expression -> shift_expression
-## In state 70, spurious reduction of production equality_expression -> relational_expression
-## In state 86, spurious reduction of production and_expression -> equality_expression
-## In state 94, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 95, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 96, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 80, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 78, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 99, spurious reduction of production assignment_expression -> conditional_expression
+## In state 61, spurious reduction of production unary_expression -> postfix_expression
+## In state 65, spurious reduction of production cast_expression -> unary_expression
+## In state 88, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 82, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 101, spurious reduction of production shift_expression -> additive_expression
+## In state 78, spurious reduction of production relational_expression -> shift_expression
+## In state 94, spurious reduction of production equality_expression -> relational_expression
+## In state 110, spurious reduction of production and_expression -> equality_expression
+## In state 118, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 119, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 120, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 104, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 102, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 123, spurious reduction of production assignment_expression -> conditional_expression
##
# At the time of writing, optional(expression,RBRACK) is used only in direct
@@ -5066,11 +4352,11 @@ then at this point, a closing bracket ']' is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN VAR_NAME COMMA XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 667.
+## Ends in an error in state: 277.
##
-## identifier_list -> identifier_list COMMA . VAR_NAME [ RPAREN COMMA ]
+## identifier_list -> identifier_list COMMA . PRE_NAME VAR_NAME [ RPAREN COMMA ]
##
## The known suffix of the stack is as follows:
## identifier_list COMMA
@@ -5083,39 +4369,141 @@ At this point, an identifier is expected.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN VAR_NAME RPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME COMMA PRE_NAME TYPEDEF_NAME
+##
+## Ends in an error in state: 278.
+##
+## identifier_list -> identifier_list COMMA PRE_NAME . VAR_NAME [ RPAREN COMMA ]
+##
+## The known suffix of the stack is as follows:
+## identifier_list COMMA PRE_NAME
+##
+
+Ill-formed K&R function definition.
+The following type name is used as a K&R parameter name:
+ $0
+
+# ------------------------------------------------------------------------------
+
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN INT XOR_ASSIGN
+##
+## Ends in an error in state: 580.
+##
+## declaration_specifiers(declaration(block_item)) -> type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
+##
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN VOLATILE INT XOR_ASSIGN
+##
+## Ends in an error in state: 585.
+##
+## declaration_specifiers(declaration(block_item)) -> rlist(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_typedef_name) -> list(declaration_specifier_no_typedef_name) . declaration_specifier_no_typedef_name [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC STAR SIGNED SHORT SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## rlist(declaration_specifier_no_type) type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name)
+##
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 674.
+## Ends in an error in state: 578.
##
-## declaration_list -> declaration_list . declaration(block_item) [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PACKED LONG LBRACE INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
-## function_definition_begin -> declaration_specifiers(declaration(external_declaration)) direct_declarator LPAREN identifier_list RPAREN open_context declaration_list . [ LBRACE ]
+## declaration_specifiers(declaration(block_item)) -> typedef_name list(declaration_specifier_no_type) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## declaration_specifiers(declaration(external_declaration)) direct_declarator LPAREN identifier_list RPAREN open_context declaration_list
+## typedef_name list(declaration_specifier_no_type)
##
-translation_unit_file: INT STAR VAR_NAME LPAREN VAR_NAME RPAREN XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN VOLATILE PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 665.
+## Ends in an error in state: 583.
##
-## declaration_list -> declaration_list . declaration(block_item) [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF_NAME TYPEDEF STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PACKED LONG LBRACE INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
-## function_definition_begin -> declaration_specifiers(declaration(external_declaration)) list(pointer1) STAR option(type_qualifier_list) direct_declarator LPAREN identifier_list RPAREN open_context declaration_list . [ LBRACE ]
+## declaration_specifiers(declaration(block_item)) -> rlist(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type) . [ STAR SEMICOLON PRE_NAME LPAREN ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . storage_class_specifier_no_typedef [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . type_qualifier_noattr [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . function_specifier [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
+## list(declaration_specifier_no_type) -> list(declaration_specifier_no_type) . attribute_specifier [ VOLATILE STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
##
## The known suffix of the stack is as follows:
-## declaration_specifiers(declaration(external_declaration)) list(pointer1) STAR option(type_qualifier_list) direct_declarator LPAREN identifier_list RPAREN open_context declaration_list
+## rlist(declaration_specifier_no_type) typedef_name list(declaration_specifier_no_type)
##
-# clang requests the function body; gcc requests a declaration :-)
+# We omit the case of the empty list of declarators
+# We omit the case of successive primitive type specifiers
-Ill-formed K&R function definition.
+Ill-formed K&R parameter declaration.
At this point, one of the following is expected:
- a declaration; or
- an opening brace '{' (for the function body).
+ a storage class specifier; or
+ a type qualifier; or
+ a list of declarators.
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT COMMA XOR_ASSIGN
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN VOLATILE XOR_ASSIGN
+##
+## Ends in an error in state: 581.
+##
+## declaration_specifiers(declaration(block_item)) -> rlist(declaration_specifier_no_type) . typedef_name list(declaration_specifier_no_type) [ STAR SEMICOLON PRE_NAME LPAREN ]
+## declaration_specifiers(declaration(block_item)) -> rlist(declaration_specifier_no_type) . type_specifier_no_typedef_name list(declaration_specifier_no_typedef_name) [ STAR SEMICOLON PRE_NAME LPAREN ]
+##
+## The known suffix of the stack is as follows:
+## rlist(declaration_specifier_no_type)
+##
+## WARNING: This example involves spurious reductions.
+## This implies that, although the LR(1) items shown above provide an
+## accurate view of the past (what has been recognized so far), they
+## may provide an INCOMPLETE view of the future (what was expected next).
+## In state 209, spurious reduction of production rlist(declaration_specifier_no_type) -> type_qualifier_noattr
+##
+
+Ill-formed K&R parameter declaration.
+At this point, one of the following is expected:
+ a storage class specifier; or
+ a type qualifier; or
+ a type specifier.
+
+# ------------------------------------------------------------------------------
+
+translation_unit_file: VOID PRE_NAME TYPEDEF_NAME PACKED LPAREN CONSTANT RPAREN XOR_ASSIGN
+##
+## Ends in an error in state: 594.
##
-## Ends in an error in state: 245.
+## attribute_specifier_list -> attribute_specifier . attribute_specifier_list [ SEMICOLON LBRACE EQ COMMA ]
+## rlist(declaration_specifier_no_type) -> attribute_specifier . [ VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT SIGNED SHORT PRE_NAME LONG INT FLOAT ENUM DOUBLE CHAR ]
+## rlist(declaration_specifier_no_type) -> attribute_specifier . rlist(declaration_specifier_no_type) [ VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT SIGNED SHORT PRE_NAME LONG INT FLOAT ENUM DOUBLE CHAR ]
+##
+## The known suffix of the stack is as follows:
+## attribute_specifier
+##
+
+# We have just parsed a list of attribute specifiers, but we cannot
+# print it because it is not available. We do not know wether it is
+# part of the declaration or whether it is part of the first K&R parameter
+# declaration.
+
+# We omit the possibility of one more attribute specifier
+
+Ill-formed declaration or function definition.
+Up to this point, a list of attribute specifiers has been recognized.
+If this is a declaration,
+ then at this point, a semicolon ';' is expected.
+If this is a function definition,
+ then at this point, an opening brace '{' is expected (for the function body).
+If this is the parameter declaration of a K&R function definition,
+ then at this point, one of the following is expected:
+ a storage class specifier; or
+ a type qualifier; or
+ a type specifier.
+
+# ------------------------------------------------------------------------------
+
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT COMMA XOR_ASSIGN
+##
+## Ends in an error in state: 225.
##
## parameter_list -> parameter_list COMMA . parameter_declaration [ RPAREN COMMA ]
## parameter_type_list -> parameter_list COMMA . ELLIPSIS [ RPAREN ]
@@ -5130,9 +4518,9 @@ At this point, one of the following is expected:
# ------------------------------------------------------------------------------
-translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME SEMICOLON
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 244.
+## Ends in an error in state: 224.
##
## parameter_list -> parameter_list . COMMA parameter_declaration [ RPAREN COMMA ]
## parameter_type_list -> parameter_list . [ RPAREN ]
@@ -5145,11 +4533,12 @@ translation_unit_file: INT VAR_NAME LPAREN INT VAR_NAME SEMICOLON
## This implies that, although the LR(1) items shown above provide an
## accurate view of the past (what has been recognized so far), they
## may provide an INCOMPLETE view of the future (what was expected next).
-## In state 275, spurious reduction of production attribute_specifier_list ->
-## In state 281, spurious reduction of production declarator -> direct_declarator attribute_specifier_list
-## In state 294, spurious reduction of production declare_varname(fst(declarator)) -> declarator
-## In state 293, spurious reduction of production parameter_declaration -> declaration_specifiers(parameter_declaration) declare_varname(fst(declarator))
-## In state 263, spurious reduction of production parameter_list -> parameter_declaration
+## In state 250, spurious reduction of production declarator_noattrend -> direct_declarator
+## In state 255, spurious reduction of production attribute_specifier_list ->
+## In state 256, spurious reduction of production declarator -> declarator_noattrend attribute_specifier_list
+## In state 272, spurious reduction of production declare_varname(declarator) -> declarator
+## In state 271, spurious reduction of production parameter_declaration -> declaration_specifiers(parameter_declaration) declare_varname(declarator)
+## In state 232, spurious reduction of production parameter_list -> parameter_declaration
##
# We omit the possibility of an ellipsis.
@@ -5165,6 +4554,115 @@ then at this point, a closing parenthesis ')' is expected.
# ------------------------------------------------------------------------------
+translation_unit_file: PRE_NAME VAR_NAME
+##
+## Ends in an error in state: 14.
+##
+## typedef_name -> PRE_NAME . TYPEDEF_NAME [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA COLON AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## PRE_NAME
+##
+
+# This can only happen in a declaration
+
+Ill-formed declaration.
+The following identifier is used as a type, but has not been defined as such:
+ $0
+
+# ------------------------------------------------------------------------------
+
+translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN INT SEMICOLON XOR_ASSIGN
+##
+## Ends in an error in state: 590.
+##
+## declaration_list -> declaration_list . kr_param_declaration [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PRE_NAME PACKED LONG LBRACE INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
+## function_definition1 -> declaration_specifiers(declaration(external_declaration)) declare_varname(declarator_noattrend) save_context declaration_list . [ LBRACE ]
+##
+## The known suffix of the stack is as follows:
+## declaration_specifiers(declaration(external_declaration)) declare_varname(declarator_noattrend) save_context declaration_list
+##
+# clang requests the function body; gcc requests a declaration :-)
+
+Ill-formed K&R function definition.
+At this point, one of the following is expected:
+ a declaration; or
+ an opening brace '{' (for the function body).
+
+# ------------------------------------------------------------------------------
+
+translation_unit_file: ALIGNAS LPAREN PRE_NAME XOR_ASSIGN
+##
+## Ends in an error in state: 27.
+##
+## primary_expression -> PRE_NAME . VAR_NAME [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RPAREN RIGHT_ASSIGN RIGHT QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
+## typedef_name -> PRE_NAME . TYPEDEF_NAME [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## PRE_NAME
+##
+translation_unit_file: ALIGNAS LPAREN VOID LPAREN VOID LPAREN PRE_NAME XOR_ASSIGN
+##
+## Ends in an error in state: 231.
+##
+## declarator_identifier -> PRE_NAME . low_prec TYPEDEF_NAME [ RPAREN PACKED LPAREN LBRACK ATTRIBUTE ALIGNAS ]
+## declarator_identifier -> PRE_NAME . VAR_NAME [ RPAREN PACKED LPAREN LBRACK ATTRIBUTE ALIGNAS ]
+## typedef_name -> PRE_NAME . TYPEDEF_NAME [ VOLATILE STATIC STAR RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## PRE_NAME
+##
+translation_unit_file: UNION PRE_NAME XOR_ASSIGN
+##
+## Ends in an error in state: 38.
+##
+## general_identifier -> PRE_NAME . VAR_NAME [ XOR_ASSIGN VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF SUB_ASSIGN STRUCT STATIC STAR SLASH SIGNED SHORT SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RESTRICT REGISTER RBRACK RBRACE QUESTION PTR PRE_NAME PLUS PERCENT PACKED OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LONG LEQ LEFT_ASSIGN LEFT LBRACK LBRACE INT INLINE INC HAT GT GEQ FLOAT EXTERN EQEQ EQ ENUM DOUBLE DOT DIV_ASSIGN DEC CONST COMMA COLON CHAR BARBAR BAR AUTO ATTRIBUTE AND_ASSIGN ANDAND AND ALIGNAS ADD_ASSIGN ]
+## typedef_name -> PRE_NAME . TYPEDEF_NAME [ XOR_ASSIGN VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF SUB_ASSIGN STRUCT STATIC STAR SLASH SIGNED SHORT SEMICOLON RPAREN RIGHT_ASSIGN RIGHT RESTRICT REGISTER RBRACK RBRACE QUESTION PTR PRE_NAME PLUS PERCENT PACKED OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LONG LEQ LEFT_ASSIGN LEFT LBRACK LBRACE INT INLINE INC HAT GT GEQ FLOAT EXTERN EQEQ EQ ENUM DOUBLE DOT DIV_ASSIGN DEC CONST COMMA COLON CHAR BARBAR BAR AUTO ATTRIBUTE AND_ASSIGN ANDAND AND ALIGNAS ADD_ASSIGN ]
+##
+## The known suffix of the stack is as follows:
+## PRE_NAME
+##
+translation_unit_file: VOID PRE_NAME TYPEDEF_NAME LBRACE PRE_NAME XOR_ASSIGN
+##
+## Ends in an error in state: 426.
+##
+## general_identifier -> PRE_NAME . VAR_NAME [ COLON ]
+## primary_expression -> PRE_NAME . VAR_NAME [ XOR_ASSIGN SUB_ASSIGN STAR SLASH SEMICOLON RIGHT_ASSIGN RIGHT QUESTION PTR PLUS PERCENT OR_ASSIGN NEQ MUL_ASSIGN MOD_ASSIGN MINUS LT LPAREN LEQ LEFT_ASSIGN LEFT LBRACK INC HAT GT GEQ EQEQ EQ DOT DIV_ASSIGN DEC COMMA BARBAR BAR AND_ASSIGN ANDAND AND ADD_ASSIGN ]
+## typedef_name -> PRE_NAME . TYPEDEF_NAME [ VOLATILE TYPEDEF STATIC STAR SEMICOLON RESTRICT REGISTER PRE_NAME PACKED LPAREN INLINE EXTERN CONST COLON AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## PRE_NAME
+##
+translation_unit_file: VOID PRE_NAME TYPEDEF_NAME LPAREN PRE_NAME XOR_ASSIGN
+##
+## Ends in an error in state: 190.
+##
+## identifier_list -> PRE_NAME . VAR_NAME [ RPAREN COMMA ]
+## typedef_name -> PRE_NAME . TYPEDEF_NAME [ VOLATILE STATIC STAR RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## PRE_NAME
+##
+translation_unit_file: VOID PRE_NAME XOR_ASSIGN
+##
+## Ends in an error in state: 178.
+##
+## declarator_identifier -> PRE_NAME . low_prec TYPEDEF_NAME [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
+## declarator_identifier -> PRE_NAME . VAR_NAME [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
+##
+## The known suffix of the stack is as follows:
+## PRE_NAME
+##
+
+# This is not supposed to be possible, since the Lexer can only emit a
+# VAR_NAME or a TYPEDEF_NAME after a PRE_NAME
+
+Internal error when printing a syntax error message. Please report.
+
+# ------------------------------------------------------------------------------
+
+
+
# Local Variables:
# mode: shell-script
# End:
diff --git a/cparser/pre_parser.mly b/cparser/pre_parser.mly
index 25e7a745..d217a7a4 100644
--- a/cparser/pre_parser.mly
+++ b/cparser/pre_parser.mly
@@ -14,6 +14,14 @@
/* */
/* *********************************************************************/
+(*
+ WARNING: The precedence declarations tend to silently solve
+ conflicts. So, if you change the grammar (especially for
+ statements), you should check that when you run "make correct"
+ in the cparser/ directory, Menhir should say:
+ 2 shift/reduce conflicts were silently solved.
+*)
+
%{
open Pre_parser_aux
@@ -26,8 +34,15 @@
let declare_typename (i,_,_) =
!declare_typename i
+ type 'id fun_declarator_ctx =
+ | Decl_ident
+ | Decl_other
+ | Decl_fun of (unit -> unit)
+ | Decl_krfun of 'id
+
%}
+%token<string> PRE_NAME
%token<string * Pre_parser_aux.identifier_type ref * Cabs.cabsloc>
VAR_NAME TYPEDEF_NAME
%token<Cabs.constant * Cabs.cabsloc> CONSTANT
@@ -46,19 +61,22 @@
%token EOF
-(* These precedences declarations solve the conflict in the following declaration :
+(* These precedence declarations solve the conflict in the following
+ declaration:
int f(int (a));
- when a is a TYPEDEF_NAME. It is specified by 6.7.5.3 11.
+ when a is a TYPEDEF_NAME. It is specified by 6.7.5.3 11: 'a' should
+ be taken as the type of parameter of the anonymous function.
- WARNING: These precedence declarations tend to silently solve other
- conflicts. So, if you change the grammar (especially or
- statements), you should check that without these declarations, it
- has ONLY ONE CONFLICT.
+ See below comment on [low_prec]
*)
+%nonassoc lowPrec1
%nonassoc TYPEDEF_NAME
-%nonassoc highPrec
+
+(* These precedence declarations solve the dangling else conflict. *)
+%nonassoc lowPrec2
+%nonassoc ELSE
%start<unit> translation_unit_file
@@ -90,8 +108,8 @@
expression
attribute_specifier_list
declarator
- statement_finish_close
- iteration_statement(nop,statement_finish_close)
+ declarator_noattrend
+ selection_statement
enum_specifier
struct_or_union_specifier
specifier_qualifier_list(struct_declaration)
@@ -102,6 +120,7 @@
asm_flags
asm_operands
init_declarator
+ rlist(declaration_specifier_no_type)
%%
@@ -152,16 +171,35 @@ optional(X, Y):
| list(X) X {}
list(X):
- xs = ilist(X)
- { xs }
+| ilist(X) {}
-%inline fst(X):
-| x = X
- { fst x }
+(* [rlist(X)] is right-recursive non-empty list. *)
+
+rlist(X):
+| X {}
+| X rlist(X) {}
+
+(* The kind of an identifier should not be determined when looking
+ ahead, because the context may not be up to date. For this reason,
+ when reading an identifier, the lexer emits two tokens: the first
+ one (PRE_NAME) is eaten as a lookahead token, the second one is the
+ actual identifier.
+*)
+(* For [var_name] we need more context on error reporting, so we use
+ %inline. Not using %inline for typedef_name helps foctorizing many
+ similar error messages. *)
+
+typedef_name:
+| PRE_NAME i = TYPEDEF_NAME
+ { i }
+
+%inline var_name:
+| PRE_NAME i = VAR_NAME
+ { i }
general_identifier:
-| i = VAR_NAME
-| i = TYPEDEF_NAME
+| i = typedef_name
+| i = var_name
{ i }
(* [other_identifier] is equivalent to [general_identifier], but adds
@@ -178,32 +216,13 @@ string_literals_list:
| string_literals_list STRING_LITERAL
{}
-(* WARNING : because of the lookahead token, the context might be
- opened or closed one token after the position of this non-terminal !
-
- Opening too late is not dangerous for us, because this does not
- change the token stream. However, we have to make sure the
- lookahead token present just after closing/declaring/restoring is
- not an identifier. An easy way to check that is to look at the
- follow set of the non-terminal in question. The follow sets are
- given by menhir with option -lg 3. *)
-
-%inline nop: (* empty *) {}
-
-open_context:
- (* empty *)%prec highPrec { !open_context () }
-close_context:
- (* empty *) { !close_context () }
-in_context(nt):
- open_context x = nt close_context { x }
-
-save_contexts_stk:
- (* empty *) { !save_contexts_stk () }
+save_context:
+ (* empty *) { !save_context () }
declare_varname(nt):
- i = nt { declare_varname i; i }
+ i = nt { declare_varname (fst i); i }
declare_typename(nt):
- i = nt { declare_typename i; i }
+ i = nt { declare_typename (fst i); i }
(* A note about phantom parameters. The definition of a non-terminal symbol
[nt] is sometimes parameterized with a parameter that is unused in the
@@ -216,7 +235,7 @@ declare_typename(nt):
states have been duplicated. In these states, more information about the
context is available, which allows better syntax error messages to be
given.
-
+
By convention, a formal phantom parameter is named [phantom], so as to be
easily recognizable. For clarity, we usually explicitly document which
actual values it can take. *)
@@ -224,7 +243,7 @@ declare_typename(nt):
(* Actual grammar *)
primary_expression:
-| VAR_NAME
+| var_name
| CONSTANT
| string_literals_list
| LPAREN expression RPAREN
@@ -392,8 +411,8 @@ init_declarator_list:
{}
init_declarator:
-| declare_varname(fst(declarator))
-| declare_varname(fst(declarator)) EQ c_initializer
+| declare_varname(declarator_noattrend) save_context attribute_specifier_list
+| declare_varname(declarator_noattrend) save_context attribute_specifier_list EQ c_initializer
{}
typedef_declarator_list:
@@ -402,7 +421,7 @@ typedef_declarator_list:
{}
typedef_declarator:
-| declare_typename(fst(declarator))
+| declare_typename(declarator)
{}
storage_class_specifier_no_typedef:
@@ -414,10 +433,11 @@ storage_class_specifier_no_typedef:
(* [declaration_specifier_no_type] matches declaration specifiers
that do not contain either "typedef" nor type specifiers. *)
-declaration_specifier_no_type:
+%inline declaration_specifier_no_type:
| storage_class_specifier_no_typedef
-| type_qualifier
+| type_qualifier_noattr
| function_specifier
+| attribute_specifier
{}
(* [declaration_specifier_no_typedef_name] matches declaration
@@ -431,10 +451,8 @@ declaration_specifier_no_typedef_name:
| type_specifier_no_typedef_name
{}
-(* [declaration_specifiers_no_type] matches declaration_specifiers
- that do not contains "typedef". Moreover, it makes sure that it
- contains either one typename and not other type specifier or no
- typename.
+(* [declaration_specifiers] makes sure one type specifier is given, and,
+ if a typedef_name is given, then no other type specifier is given.
This is a weaker condition than 6.7.2 2. It is necessary to enforce
this in the grammar to disambiguate the example in 6.7.7 6:
@@ -452,18 +470,18 @@ declaration_specifier_no_typedef_name:
[parameter_declaration], which means that this is the beginning of a
parameter declaration. *)
declaration_specifiers(phantom):
-| ilist(declaration_specifier_no_type) TYPEDEF_NAME declaration_specifier_no_type*
-| declaration_specifier_no_type* type_specifier_no_typedef_name declaration_specifier_no_typedef_name*
+| ioption(rlist(declaration_specifier_no_type)) typedef_name declaration_specifier_no_type*
+| ioption(rlist(declaration_specifier_no_type)) type_specifier_no_typedef_name declaration_specifier_no_typedef_name*
{}
(* This matches declaration_specifiers that do contains once the
"typedef" keyword. To avoid conflicts, we also encode the
constraint described in the comment for [declaration_specifiers]. *)
declaration_specifiers_typedef:
-| declaration_specifier_no_type* TYPEDEF declaration_specifier_no_type* TYPEDEF_NAME declaration_specifier_no_type*
-| ilist(declaration_specifier_no_type) TYPEDEF_NAME declaration_specifier_no_type* TYPEDEF declaration_specifier_no_type*
-| declaration_specifier_no_type* TYPEDEF declaration_specifier_no_type* type_specifier_no_typedef_name declaration_specifier_no_typedef_name*
-| declaration_specifier_no_type* type_specifier_no_typedef_name declaration_specifier_no_typedef_name* TYPEDEF declaration_specifier_no_typedef_name*
+| ioption(rlist(declaration_specifier_no_type)) TYPEDEF declaration_specifier_no_type* typedef_name declaration_specifier_no_type*
+| ioption(rlist(declaration_specifier_no_type)) typedef_name declaration_specifier_no_type* TYPEDEF declaration_specifier_no_type*
+| ioption(rlist(declaration_specifier_no_type)) TYPEDEF declaration_specifier_no_type* type_specifier_no_typedef_name declaration_specifier_no_typedef_name*
+| ioption(rlist(declaration_specifier_no_type)) type_specifier_no_typedef_name declaration_specifier_no_typedef_name* TYPEDEF declaration_specifier_no_typedef_name*
{}
(* A type specifier which is not a typedef name. *)
@@ -505,8 +523,8 @@ struct_declaration:
in the comment above [declaration_specifiers]. *)
(* The phantom parameter can be [struct_declaration] or [type_name]. *)
specifier_qualifier_list(phantom):
-| type_qualifier_list? TYPEDEF_NAME type_qualifier_list?
-| type_qualifier_list? type_specifier_no_typedef_name specifier_qualifier_no_typedef_name*
+| ioption(type_qualifier_list) typedef_name type_qualifier_list?
+| ioption(type_qualifier_list) type_specifier_no_typedef_name specifier_qualifier_no_typedef_name*
{}
specifier_qualifier_no_typedef_name:
@@ -537,29 +555,31 @@ enumerator_list:
enumerator:
| i = enumeration_constant
| i = enumeration_constant EQ constant_expression
- { i }
+ { (i, ()) }
enumeration_constant:
| i = general_identifier
{ set_id_type i VarId; i }
-type_qualifier:
+type_qualifier_noattr:
| CONST
| RESTRICT
| VOLATILE
+ {}
+
+%inline type_qualifier:
+| type_qualifier_noattr
| attribute_specifier
{}
attribute_specifier_list:
| /* empty */
-| attribute_specifier_list attribute_specifier
+| attribute_specifier attribute_specifier_list
{}
attribute_specifier:
| ATTRIBUTE LPAREN LPAREN gcc_attribute_list RPAREN RPAREN
| PACKED LPAREN argument_expression_list RPAREN
-(* TODO: slove conflict *)
-(* | PACKED *)
| ALIGNAS LPAREN argument_expression_list RPAREN
| ALIGNAS LPAREN type_name RPAREN
{}
@@ -574,7 +594,7 @@ gcc_attribute:
| gcc_attribute_word
| gcc_attribute_word LPAREN argument_expression_list? RPAREN
{}
-| gcc_attribute_word LPAREN i = TYPEDEF_NAME COMMA argument_expression_list RPAREN
+| gcc_attribute_word LPAREN i = typedef_name COMMA argument_expression_list RPAREN
(* This is to emulate GCC's attribute syntax : we make this identifier
a var name identifier, so that the parser will see it as a variable
reference *)
@@ -590,26 +610,74 @@ function_specifier:
| INLINE
{}
+(* We add this non-terminal here to force the resolution of the
+ conflict at the point of shifting the TYPEDEF_NAME. If we had
+ already shifted it, a reduce/reduce conflict appears, and menhir is
+ not able to solve them.
+
+ The conflict in question is when parsing :
+ int f(int (t
+ With lookahead ')', in a context where 't' is a type name.
+ In this case, we are able to reduce the two productions:
+ (1) "declarator_identifier -> PRE_NAME TYPEDEF_NAME"
+ followed by "direct_declarator -> declarator_identifier"
+ meaning that 't' is the parameter of function 'f'
+ (2) "list(declaration_specifier_no_type) -> "
+ followed by "list(declaration_specifier_no_type) -> PRE_NAME TYPEDEF_NAME list(declaration_specifier_no_type)"
+ followed by "declaration_specifiers(...) -> ..."
+ followed by "parameter_declaration -> ..."
+ meaning that 't' is the type of the parameter of a function
+ passed as parameter to 'f'
+
+ By adding this non-terminal at this point, we force this conflict to
+ be solved earlier: once we have seen "f(int (", followed by PRE_NAME
+ and with TYPEDEF_NAME in lookahead position, we know (1) can safely
+ be ignored (if (1) is still possible after reading the next token,
+ (2) will also be possible, and the conflict has to be solved in
+ favor of (2)). We add low_prec in declaration_identifier, but not in
+ typedef_name, so that it has to be reduced in (1) but not in (2).
+ This is a shift/reduce conflict that can be solved using
+ precedences.
+ *)
+low_prec : %prec lowPrec1 {}
+declarator_identifier:
+| PRE_NAME low_prec i = TYPEDEF_NAME
+| i = var_name
+ { i }
+
(* The semantic action returned by [declarator] is a pair of the
- identifier being defined and an option of the context stack that
- has to be restored if entering the body of the function being
+ identifier being defined and a value containing the context stack
+ that has to be restored if entering the body of the function being
defined, if so. *)
declarator:
-| ilist(pointer1) x = direct_declarator attribute_specifier_list
+| x = declarator_noattrend attribute_specifier_list
+ { x }
+
+declarator_noattrend:
+| x = direct_declarator
{ x }
+| pointer x = direct_declarator
+ { match snd x with
+ | Decl_ident -> (fst x, Decl_other)
+ | _ -> x }
direct_declarator:
-| i = general_identifier
- { set_id_type i VarId; (i, None) }
-| LPAREN x = declarator RPAREN
+| i = declarator_identifier
+ { set_id_type i VarId; (i, Decl_ident) }
+| LPAREN save_context x = declarator RPAREN
| x = direct_declarator LBRACK type_qualifier_list? optional(assignment_expression, RBRACK)
- { x }
-| x = direct_declarator LPAREN
- open_context parameter_type_list? restore_fun = save_contexts_stk
- close_context RPAREN
{ match snd x with
- | None -> (fst x, Some restore_fun)
- | Some _ -> x }
+ | Decl_ident -> (fst x, Decl_other)
+ | _ -> x }
+| x = direct_declarator LPAREN ctx = context_parameter_type_list RPAREN
+ { match snd x with
+ | Decl_ident -> (fst x, Decl_fun ctx)
+ | _ -> x }
+| x = direct_declarator LPAREN save_context il=identifier_list? RPAREN
+ { match snd x, il with
+ | Decl_ident, Some il -> (fst x, Decl_krfun il)
+ | Decl_ident, None -> (fst x, Decl_krfun [])
+ | _ -> x }
(* The C standard defines [pointer] as a right-recursive list. We prefer to
define it as a left-recursive list, because this provides better static
@@ -622,8 +690,7 @@ direct_declarator:
as [pointer1* pointer1].
When the C standard writes [pointer?], which represents a possibly empty
- list of [pointer1]'s, we write [pointer1*] or [ilist(pointer1)]. The two
- are equivalent, as long as there is no conflict. *)
+ list of [pointer1]'s, we write [pointer1*]. *)
%inline pointer1:
STAR type_qualifier_list?
@@ -637,22 +704,24 @@ type_qualifier_list:
| type_qualifier_list? type_qualifier
{}
+context_parameter_type_list:
+| ctx1 = save_context parameter_type_list ctx2 = save_context
+ { ctx1 (); ctx2 }
+
parameter_type_list:
-| l=parameter_list
-| l=parameter_list COMMA ELLIPSIS
- { l }
+| parameter_list
+| parameter_list COMMA ELLIPSIS
+ {}
parameter_list:
-| i=parameter_declaration
- { [i] }
-| l=parameter_list COMMA i=parameter_declaration
- { i::l }
+| parameter_declaration
+| parameter_list COMMA parameter_declaration
+ {}
parameter_declaration:
-| declaration_specifiers(parameter_declaration) id=declare_varname(fst(declarator))
- { Some id }
+| declaration_specifiers(parameter_declaration) declare_varname(declarator)
| declaration_specifiers(parameter_declaration) abstract_declarator(parameter_declaration)?
- { None }
+ {}
type_name:
| specifier_qualifier_list(type_name) abstract_declarator(type_name)?
@@ -664,13 +733,13 @@ type_name:
is permitted (and we do not wish to keep track of why it is permitted). *)
abstract_declarator(phantom):
| pointer
-| ilist(pointer1) direct_abstract_declarator
+| ioption(pointer) direct_abstract_declarator
{}
direct_abstract_declarator:
-| LPAREN abstract_declarator(type_name) RPAREN
+| LPAREN save_context abstract_declarator(type_name) RPAREN
| direct_abstract_declarator? LBRACK type_qualifier_list? optional(assignment_expression, RBRACK)
-| ioption(direct_abstract_declarator) LPAREN in_context(parameter_type_list?) RPAREN
+| ioption(direct_abstract_declarator) LPAREN context_parameter_type_list? RPAREN
{}
c_initializer:
@@ -696,64 +765,25 @@ designator:
| DOT other_identifier
{}
-(* The grammar of statements is replicated three times.
-
- [statement_finish_close] should close the current context just
- before its last token.
-
- [statement_finish_noclose] should not close the current context. It
- should modify it only if this modification actually changes the
- context of the current block.
-
- [statement_intern_close] is like [statement_finish_close], except
- it cannot reduce to a single-branch if statement.
-*)
-
-statement_finish_close:
-| labeled_statement(statement_finish_close)
-| compound_statement(nop)
-| expression_statement(close_context)
-| selection_statement_finish(nop)
-| iteration_statement(nop,statement_finish_close)
-| jump_statement(close_context)
-| asm_statement(close_context)
+statement:
+| labeled_statement
+| compound_statement
+| expression_statement
+| selection_statement
+| iteration_statement
+| jump_statement
+| asm_statement
{}
-statement_finish_noclose:
-| labeled_statement(statement_finish_noclose)
-| compound_statement(open_context)
-| expression_statement(nop)
-| selection_statement_finish(open_context)
-| iteration_statement(open_context,statement_finish_close)
-| jump_statement(nop)
-| asm_statement(nop)
+labeled_statement:
+| other_identifier COLON statement
+| CASE constant_expression COLON statement
+| DEFAULT COLON statement
{}
-statement_intern_close:
-| labeled_statement(statement_intern_close)
-| compound_statement(nop)
-| expression_statement(close_context)
-| selection_statement_intern_close
-| iteration_statement(nop,statement_intern_close)
-| jump_statement(close_context)
-| asm_statement(close_context)
- {}
-
-(* [labeled_statement(last_statement)] has the same effect on contexts
- as [last_statement]. *)
-labeled_statement(last_statement):
-| other_identifier COLON last_statement
-| CASE constant_expression COLON last_statement
-| DEFAULT COLON last_statement
- {}
-
-(* [compound_statement] uses a local context and closes it before its
- last token. It uses [openc] to open this local context if needed.
- That is, if a local context has already been opened, [openc] = [nop],
- otherwise, [openc] = [open_context]. *)
-compound_statement(openc):
-| LBRACE openc block_item_list? close_context RBRACE
- {}
+compound_statement:
+| ctx = save_context LBRACE block_item_list? RBRACE
+ { ctx() }
block_item_list:
| block_item_list? block_item
@@ -761,93 +791,44 @@ block_item_list:
block_item:
| declaration(block_item)
-| statement_finish_noclose
+| statement
| PRAGMA
{}
-(* [expression_statement], [jump_statement] and [asm_statement] close
- the local context if needed, depending of the close parameter. If
- there is no local context, [close] = [nop]. Otherwise,
- [close] = [close_context]. *)
-expression_statement(close):
-| expression? close SEMICOLON
+expression_statement:
+| expression? SEMICOLON
{}
-jump_statement(close):
-| GOTO other_identifier close SEMICOLON
-| CONTINUE close SEMICOLON
-| BREAK close SEMICOLON
-| RETURN expression? close SEMICOLON
+jump_statement:
+| GOTO other_identifier SEMICOLON
+| CONTINUE SEMICOLON
+| BREAK SEMICOLON
+| RETURN expression? SEMICOLON
{}
-asm_statement(close):
-| ASM asm_attributes LPAREN string_literals_list asm_arguments RPAREN close SEMICOLON
+asm_statement:
+| ASM asm_attributes LPAREN string_literals_list asm_arguments RPAREN SEMICOLON
{}
-(* [selection_statement_finish] and [selection_statement_intern] use a
- local context and close it before their last token.
+ifelse_statement1:
+| IF LPAREN expression RPAREN ctx = save_context statement ELSE
+ { ctx() }
- [selection_statement_finish(openc)] uses [openc] to open this local
- context if needed. That is, if a local context has already been
- opened, [openc] = [nop], otherwise, [openc] = [open_context].
+selection_statement:
+| ctx = save_context ifelse_statement1 statement
+| ctx = save_context IF LPAREN expression RPAREN save_context statement %prec lowPrec2
+| ctx = save_context SWITCH LPAREN expression RPAREN statement
+ { ctx() }
- [selection_statement_intern_close] is always called with a local
- context openned. It closes it before its last token. *)
+do_statement1:
+| ctx = save_context DO statement
+ { ctx () }
-(* It should be noted that the token [ELSE] should be lookaheaded
- /outside/ of the local context because if the lookaheaded token is
- not [ELSE], then this is the end of the statement.
-
- This is especially important to parse correctly the following
- example:
-
- typedef int a;
-
- int f() {
- for(int a; ;)
- if(1);
- a * x;
- }
-
- However, if the lookahead token is [ELSE], we should parse the
- second branch in the same context as the first branch, so we have
- to reopen the previously closed context. This is the reason for the
- save/restore system.
-*)
-
-if_else_statement_begin(openc):
-| IF openc LPAREN expression RPAREN restore_fun = save_contexts_stk
- statement_intern_close
- { restore_fun () }
-
-selection_statement_finish(openc):
-| IF openc LPAREN expression RPAREN save_contexts_stk statement_finish_close
-| if_else_statement_begin(openc) ELSE statement_finish_close
-| SWITCH openc LPAREN expression RPAREN statement_finish_close
- {}
-
-selection_statement_intern_close:
-| if_else_statement_begin(nop) ELSE statement_intern_close
-| SWITCH LPAREN expression RPAREN statement_intern_close
- {}
-
-(* [iteration_statement] uses a local context and closes it before
- their last token.
-
- [iteration_statement] uses [openc] to open this local context if
- needed. That is, if a local context has already been opened,
- [openc] = [nop], otherwise, [openc] = [open_context].
-
- [last_statement] is either [statement_intern_close] or
- [statement_finish_close]. That is, it should /always/ close the
- local context. *)
-
-iteration_statement(openc,last_statement):
-| WHILE openc LPAREN expression RPAREN last_statement
-| DO open_context statement_finish_close WHILE
- openc LPAREN expression RPAREN close_context SEMICOLON
-| FOR openc LPAREN for_statement_header optional(expression, SEMICOLON) optional(expression, RPAREN) last_statement
- {}
+iteration_statement:
+| ctx = save_context WHILE LPAREN expression RPAREN statement
+| ctx = save_context do_statement1 WHILE LPAREN expression RPAREN SEMICOLON
+| ctx = save_context FOR LPAREN for_statement_header optional(expression, SEMICOLON) optional(expression, RPAREN) statement
+ { ctx() }
for_statement_header:
| optional(expression, SEMICOLON)
@@ -909,40 +890,35 @@ external_declaration:
| PRAGMA
{}
-function_definition_begin:
-| declaration_specifiers(declaration(external_declaration))
- ilist(pointer1) x=direct_declarator
- { match x with
- | (_, None) -> !open_context()
- (* this case does not make sense, but we let it pass anyway;
- this error will be caught later on by a type check *)
- | (i, Some restore_fun) -> restore_fun ()
- }
-| declaration_specifiers(declaration(external_declaration))
- ilist(pointer1) x=direct_declarator
- LPAREN params=identifier_list RPAREN open_context declaration_list
- { match x with
- | (i, Some _) -> declare_varname i
- (* this case does not make sense; the syntax of K&R
- declarators is broken anyway, Jacques-Henri should
- propose a fix soon *)
- | (i, None) ->
- declare_varname i;
- List.iter declare_varname params
- }
-
identifier_list:
-| id = VAR_NAME
- { [id] }
-| idl = identifier_list COMMA id = VAR_NAME
- { id :: idl }
+| x = var_name
+ { [x] }
+| l = identifier_list COMMA x = var_name
+ { x::l }
-declaration_list:
-| /*empty*/
+kr_param_declaration:
+| declaration_specifiers(declaration(block_item)) init_declarator_list? SEMICOLON
{}
-| declaration_list declaration(block_item)
+
+declaration_list:
+| kr_param_declaration
+| declaration_list kr_param_declaration
{}
+function_definition1:
+| declaration_specifiers(declaration(external_declaration))
+ func = declare_varname(declarator_noattrend)
+ save_context attribute_specifier_list ctx = save_context
+| declaration_specifiers(declaration(external_declaration))
+ func = declare_varname(declarator_noattrend)
+ ctx = save_context declaration_list
+ { begin match snd func with
+ | Decl_fun ctx -> ctx (); declare_varname (fst func)
+ | Decl_krfun il -> List.iter declare_varname il
+ | _ -> ()
+ end;
+ ctx }
+
function_definition:
-| function_definition_begin LBRACE block_item_list? close_context RBRACE
- {}
+| ctx = function_definition1 compound_statement
+ { ctx () }
diff --git a/cparser/pre_parser_aux.ml b/cparser/pre_parser_aux.ml
index c6b48608..74bf494c 100644
--- a/cparser/pre_parser_aux.ml
+++ b/cparser/pre_parser_aux.ml
@@ -18,20 +18,11 @@ type identifier_type =
| TypedefId
| OtherId
-(* These functions push and pop a context on the contexts stack. *)
-let open_context:(unit -> unit) ref = ref (fun () -> assert false)
-let close_context:(unit -> unit) ref = ref (fun () -> assert false)
+(* Applying once this functions saves the current context, and
+ applying it the second time restores it. *)
+let save_context:(unit -> (unit -> unit)) ref = ref (fun _ -> assert false)
-(* Applying once this functions saves the whole contexts stack, and
- applying it the second time restores it.
-
- This is mainly used to rollback the context stack to a previous
- state. This is usefull for example when we pop too much contexts at
- the end of the first branch of an if statement. See
- pre_parser.mly. *)
-let save_contexts_stk:(unit -> (unit -> unit)) ref = ref (fun _ -> assert false)
-
-(* Change the context at the top of the top stack of context, by
- changing an identifier to be a varname or a typename*)
+(* Change the context by changing an identifier to be a varname or a
+ typename *)
let declare_varname:(string -> unit) ref = ref (fun _ -> assert false)
let declare_typename:(string -> unit) ref = ref (fun _ -> assert false)