aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@college-de-france.fr>2022-05-02 17:31:35 +0200
committerXavier Leroy <xavier.leroy@college-de-france.fr>2022-05-13 11:21:18 +0200
commite44143ad023400c7a8193b7e9fc3b62b9f9614e1 (patch)
treee887f1fac6525441ab4ebdb9de56aa3525eeea0d
parent1d572b330362711c808094333134ba94fcd7b768 (diff)
downloadcompcert-e44143ad023400c7a8193b7e9fc3b62b9f9614e1.tar.gz
compcert-e44143ad023400c7a8193b7e9fc3b62b9f9614e1.zip
Support _Generic from ISO C 2011
Entirely handled during elaboration. No impact on the verified part of CompCert.
-rw-r--r--cparser/Cabs.v8
-rw-r--r--cparser/Elab.ml56
-rw-r--r--cparser/Lexer.mll2
-rw-r--r--cparser/Parser.vy26
-rw-r--r--cparser/deLexer.ml1
-rw-r--r--cparser/handcrafted.messages1632
-rw-r--r--cparser/pre_parser.mly17
7 files changed, 1031 insertions, 711 deletions
diff --git a/cparser/Cabs.v b/cparser/Cabs.v
index accb95a0..bf8c8c74 100644
--- a/cparser/Cabs.v
+++ b/cparser/Cabs.v
@@ -141,8 +141,10 @@ with expression :=
| MEMBEROF : expression -> string -> expression
| MEMBEROFPTR : expression -> string -> expression
- (* Non-standard *)
+ (* C11 *)
| ALIGNOF : (list spec_elem * decl_type) -> expression
+ | GENERIC : expression -> list (option (list spec_elem * decl_type) * expression) -> expression
+ (* Non-standard *)
| BUILTIN_OFFSETOF : (list spec_elem * decl_type) -> list initwhat -> expression
with constant :=
@@ -186,6 +188,10 @@ Definition init_name_group := (list spec_elem * list init_name)%type.
(* e.g.: int x, y; *)
Definition name_group := (list spec_elem * list name)%type.
+(* Useful type abbreviations *)
+Definition type_name := (list spec_elem * decl_type)%type.
+Definition generic_assoc := (option type_name * expression)%type.
+
(* GCC extended asm *)
Inductive asm_operand :=
| ASMOPERAND: option string -> bool -> list char_code -> expression -> asm_operand.
diff --git a/cparser/Elab.ml b/cparser/Elab.ml
index eff6f3ba..40f7eb73 100644
--- a/cparser/Elab.ml
+++ b/cparser/Elab.ml
@@ -1765,6 +1765,32 @@ let elab_expr ctx loc env a =
let cst' = elab_constant loc cst in
{ edesc = EConst cst'; etyp = type_of_constant cst' },env
+(* 6.5.1.1 Generic selection *)
+
+ | GENERIC(a1, assoc) ->
+ let b1,env = elab env a1 in
+ let bssoc,env = elab_generic_association env assoc in
+ let ty = erase_attributes_type env (pointer_decay env b1.etyp) in
+ let exact_match = function
+ | (None, _) -> false
+ | (Some ty', _) -> compatible_types AttrCompat env ty ty'
+ and default_match = function
+ | (None, _) -> true
+ | (Some _, _) -> false in
+ begin match List.filter exact_match bssoc with
+ | (_, b) :: others ->
+ if others <> [] then
+ error "'_Generic' selector of type %a is compatible with several associations"
+ (print_typ env) ty;
+ (b,env)
+ | [] ->
+ match List.find_opt default_match bssoc with
+ | Some (_, b) -> (b,env)
+ | None ->
+ fatal_error "'_Generic' selector of type %a is not compatible with any association"
+ (print_typ env) ty
+ end
+
(* 6.5.2 Postfix expressions *)
| INDEX(a1, a2) -> (* e1[e2] *)
@@ -2449,6 +2475,36 @@ let elab_expr ctx loc env a =
end;
let rest,env = elab_arguments (argno + 1) (argl,env) paraml vararg in
arg1 :: rest,env
+
+ (* Elaboration of _Generic association lists *)
+ and elab_generic_association env assoc =
+ let rec elab_gen env accu = function
+ | [] -> (List.rev accu, env)
+ | (None, a) :: l ->
+ if List.exists (fun (oty, _) -> oty = None) accu then
+ error "duplicate default generic association";
+ let b,env = elab env a in
+ elab_gen env ((None, b) :: accu) l
+ | (Some(spec, dcl), a) :: l ->
+ let ty,env = elab_type loc env spec dcl in
+ if wrap is_function_type loc env ty then
+ error "function type %a in generic association"
+ (print_typ env) ty
+ else if wrap incomplete_type loc env ty then
+ error "incomplete type %a in generic association"
+ (print_typ env) ty;
+ List.iter
+ (function
+ | (None, _) -> ()
+ | (Some ty', _) ->
+ if compatible_types AttrCompat env ty ty' then
+ error "type %a in generic association compatible with previously specified type %a"
+ (print_typ env) ty (print_typ env) ty')
+ accu;
+ let b,env = elab env a in
+ elab_gen env ((Some ty, b) :: accu) l
+ in elab_gen env [] assoc
+
in elab env a
(* Filling in forward declaration *)
diff --git a/cparser/Lexer.mll b/cparser/Lexer.mll
index 42980d30..9f7fba1e 100644
--- a/cparser/Lexer.mll
+++ b/cparser/Lexer.mll
@@ -34,6 +34,7 @@ let () =
("_Alignas", fun loc -> ALIGNAS loc);
("_Alignof", fun loc -> ALIGNOF loc);
("_Bool", fun loc -> UNDERSCORE_BOOL loc);
+ ("_Generic", fun loc -> GENERIC loc);
("_Complex", fun loc -> reserved_keyword loc "_Complex");
("_Imaginary", fun loc -> reserved_keyword loc "_Imaginary");
("_Static_assert", fun loc -> STATIC_ASSERT loc);
@@ -538,6 +539,7 @@ and singleline_comment = parse
| Pre_parser.EXTERN loc -> loop (Parser.EXTERN loc)
| Pre_parser.FLOAT loc -> loop (Parser.FLOAT loc)
| Pre_parser.FOR loc -> loop (Parser.FOR loc)
+ | Pre_parser.GENERIC loc -> loop (Parser.GENERIC loc)
| Pre_parser.GEQ loc -> loop (Parser.GEQ loc)
| Pre_parser.GOTO loc -> loop (Parser.GOTO loc)
| Pre_parser.GT loc -> loop (Parser.GT loc)
diff --git a/cparser/Parser.vy b/cparser/Parser.vy
index d489d339..8d7cd055 100644
--- a/cparser/Parser.vy
+++ b/cparser/Parser.vy
@@ -27,7 +27,7 @@ Require Cabs.
%token<Cabs.constant * Cabs.loc> CONSTANT
%token<Cabs.loc> SIZEOF PTR INC DEC LEFT RIGHT LEQ GEQ EQEQ EQ NEQ LT GT
ANDAND BARBAR PLUS MINUS STAR TILDE BANG SLASH PERCENT HAT BAR QUESTION
- COLON AND ALIGNOF
+ COLON AND ALIGNOF GENERIC
%token<Cabs.loc> MUL_ASSIGN DIV_ASSIGN MOD_ASSIGN ADD_ASSIGN SUB_ASSIGN
LEFT_ASSIGN RIGHT_ASSIGN AND_ASSIGN XOR_ASSIGN OR_ASSIGN
@@ -47,7 +47,9 @@ Require Cabs.
shift_expression relational_expression equality_expression AND_expression
exclusive_OR_expression inclusive_OR_expression logical_AND_expression
logical_OR_expression conditional_expression assignment_expression
- constant_expression expression
+ constant_expression expression generic_selection
+%type<list Cabs.generic_assoc> (* Reverse order *) generic_assoc_list
+%type<Cabs.generic_assoc> generic_association
%type<Cabs.unary_operator * Cabs.loc> unary_operator
%type<Cabs.binary_operator> assignment_operator
%type<list Cabs.expression (* Reverse order *)> argument_expression_list
@@ -126,6 +128,26 @@ primary_expression:
(Cabs.CONSTANT (Cabs.CONST_STRING wide chars), loc) }
| loc = LPAREN expr = expression RPAREN
{ (fst expr, loc)}
+| sel = generic_selection
+ { sel }
+
+(* 6.5.1.1 *)
+generic_selection:
+| loc = GENERIC LPAREN expr = assignment_expression COMMA
+ alist = generic_assoc_list RPAREN
+ { (Cabs.GENERIC (fst expr) (rev' alist), loc) }
+
+generic_assoc_list:
+| a = generic_association
+ { [a] }
+| l = generic_assoc_list COMMA a = generic_association
+ { a :: l }
+
+generic_association:
+| tname = type_name COLON expr = assignment_expression
+ { (Some tname, fst expr) }
+| DEFAULT COLON expr = assignment_expression
+ { (None, fst expr) }
(* 6.5.2 *)
postfix_expression:
diff --git a/cparser/deLexer.ml b/cparser/deLexer.ml
index e3ab3291..c4a2c03e 100644
--- a/cparser/deLexer.ml
+++ b/cparser/deLexer.ml
@@ -48,6 +48,7 @@ let delex (symbol : string) : string =
| "EXTERN" -> "extern"
| "FLOAT" -> "float"
| "FOR" -> "for"
+ | "GENERIC" -> "_Generic"
| "GOTO" -> "goto"
| "IF" -> "if"
| "INT" -> "int"
diff --git a/cparser/handcrafted.messages b/cparser/handcrafted.messages
index db7318c4..079937f3 100644
--- a/cparser/handcrafted.messages
+++ b/cparser/handcrafted.messages
@@ -182,9 +182,9 @@
translation_unit_file: ALIGNAS LPAREN INT XOR_ASSIGN
##
-## Ends in an error in state: 322.
+## Ends in an error in state: 315.
##
-## 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ALIGNAS LPAREN type_name
@@ -193,9 +193,9 @@ translation_unit_file: ALIGNAS LPAREN INT 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 67, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
-## In state 314, spurious reduction of production option(abstract_declarator(type_name)) ->
-## In state 320, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+## In state 159, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 317, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 323, 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
@@ -215,7 +215,7 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ ALIGNOF LPAREN VOID XOR_ASSIGN
##
-## Ends in an error in state: 312.
+## Ends in an error in state: 327.
##
## 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 ]
##
@@ -226,13 +226,13 @@ translation_unit_file: INT PRE_NAME 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 67, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
-## In state 314, spurious reduction of production option(abstract_declarator(type_name)) ->
-## In state 320, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+## In state 159, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 317, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 323, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
##
translation_unit_file: INT PRE_NAME VAR_NAME EQ SIZEOF LPAREN VOID XOR_ASSIGN
##
-## Ends in an error in state: 396.
+## Ends in an error in state: 412.
##
## 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 ]
@@ -244,9 +244,9 @@ translation_unit_file: INT PRE_NAME 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 67, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
-## In state 314, spurious reduction of production option(abstract_declarator(type_name)) ->
-## In state 320, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+## In state 159, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 317, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 323, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
##
Ill-formed use of $2.
@@ -259,7 +259,7 @@ then at this point, a closing parenthesis ')' is expected.
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: 341.
+## Ends in an error in state: 371.
##
## 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 ]
##
@@ -270,9 +270,9 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ BUILTIN_VA_ARG LPAREN PRE_NAME V
## 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 67, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
-## In state 314, spurious reduction of production option(abstract_declarator(type_name)) ->
-## In state 320, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+## In state 159, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 317, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 323, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
##
Ill-formed use of __builtin_va_arg.
@@ -285,7 +285,7 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ INC LPAREN VOID XOR_ASSIGN
##
-## Ends in an error in state: 371.
+## Ends in an error in state: 387.
##
## 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 ]
##
@@ -296,9 +296,9 @@ translation_unit_file: INT PRE_NAME 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 67, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
-## In state 314, spurious reduction of production option(abstract_declarator(type_name)) ->
-## In state 320, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+## In state 159, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 317, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 323, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
##
# gcc simply says it expects a closing parenthesis,
@@ -314,7 +314,7 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ LPAREN VOID XOR_ASSIGN
##
-## Ends in an error in state: 393.
+## Ends in an error in state: 409.
##
## 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 ]
@@ -326,9 +326,9 @@ translation_unit_file: INT PRE_NAME 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 67, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
-## In state 314, spurious reduction of production option(abstract_declarator(type_name)) ->
-## In state 320, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+## In state 159, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 317, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 323, 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.
@@ -342,10 +342,10 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: ALIGNAS LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 324.
+## Ends in an error in state: 325.
##
## argument_expression_list -> argument_expression_list . COMMA assignment_expression [ RPAREN COMMA ]
-## 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ALIGNAS LPAREN argument_expression_list
@@ -354,21 +354,21 @@ translation_unit_file: ALIGNAS LPAREN PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 155, spurious reduction of production argument_expression_list -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 146, spurious reduction of production argument_expression_list -> assignment_expression
##
# We are trying to recognize an alignas specifier.
@@ -392,22 +392,22 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: ALIGNAS LPAREN INT LBRACK RPAREN
##
-## Ends in an error in state: 248.
+## Ends in an error in state: 251.
##
-## 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_noattr [ VOLATILE TILDE STRING_LITERAL STAR SIZEOF RESTRICT RBRACK PRE_NAME PLUS PACKED MINUS LPAREN INC DEC CONSTANT CONST BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 BUILTIN_OFFSETOF BANG ATTRIBUTE AND ALIGNOF ALIGNAS ]
+## direct_abstract_declarator -> option(direct_abstract_declarator) LBRACK option(type_qualifier_list) . optional(assignment_expression,RBRACK) [ RPAREN LPAREN LBRACK COMMA COLON ]
+## 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 GENERIC DEC CONSTANT CONST BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 GENERIC DEC CONSTANT CONST BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 PRE_NAME VAR_NAME LBRACK RPAREN
##
-## Ends in an error in state: 265.
+## Ends in an error in state: 268.
##
## 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 NORETURN 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 BUILTIN_OFFSETOF 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 BUILTIN_OFFSETOF 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 GENERIC DEC CONSTANT CONST BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 GENERIC DEC CONSTANT CONST BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## direct_declarator LBRACK option(type_qualifier_list)
@@ -437,25 +437,25 @@ At this point, one of the following is expected:
translation_unit_file: ALIGNAS LPAREN INT LPAREN INT COMMA ELLIPSIS XOR_ASSIGN
##
-## Ends in an error in state: 276.
+## Ends in an error in state: 279.
##
-## direct_abstract_declarator -> LPAREN option(context_parameter_type_list) . RPAREN [ RPAREN LPAREN LBRACK COMMA ]
+## direct_abstract_declarator -> LPAREN option(context_parameter_type_list) . RPAREN [ RPAREN LPAREN LBRACK COMMA COLON ]
##
## The known suffix of the stack is as follows:
## LPAREN option(context_parameter_type_list)
##
translation_unit_file: ALIGNAS LPAREN INT LBRACK RBRACK LPAREN INT COMMA ELLIPSIS XOR_ASSIGN
##
-## Ends in an error in state: 259.
+## Ends in an error in state: 262.
##
-## direct_abstract_declarator -> direct_abstract_declarator LPAREN option(context_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 COLON ]
##
## The known suffix of the stack is as follows:
## direct_abstract_declarator LPAREN option(context_parameter_type_list)
##
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT COMMA ELLIPSIS XOR_ASSIGN
##
-## Ends in an error in state: 293.
+## Ends in an error in state: 296.
##
## 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 NORETURN LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
@@ -472,9 +472,9 @@ At this point, a closing parenthesis ')' is expected.
translation_unit_file: ALIGNAS LPAREN INT LPAREN LPAREN RPAREN COMMA
##
-## Ends in an error in state: 274.
+## Ends in an error in state: 277.
##
-## direct_abstract_declarator -> LPAREN save_context abstract_declarator(type_name) . RPAREN [ RPAREN LPAREN LBRACK COMMA ]
+## direct_abstract_declarator -> LPAREN save_context abstract_declarator(type_name) . RPAREN [ RPAREN LPAREN LBRACK COMMA COLON ]
##
## The known suffix of the stack is as follows:
## LPAREN save_context abstract_declarator(type_name)
@@ -483,7 +483,7 @@ translation_unit_file: ALIGNAS LPAREN INT LPAREN LPAREN RPAREN COMMA
## 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 267, spurious reduction of production abstract_declarator(type_name) -> direct_abstract_declarator
+## In state 270, 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.
@@ -514,10 +514,10 @@ At this point, a closing parenthesis ')' is expected.
translation_unit_file: ALIGNAS LPAREN INT LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 315.
+## Ends in an error in state: 318.
##
-## 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_abstract_declarator -> LPAREN . save_context abstract_declarator(type_name) RPAREN [ RPAREN LPAREN LBRACK COMMA COLON ]
+## direct_abstract_declarator -> LPAREN . option(context_parameter_type_list) RPAREN [ RPAREN LPAREN LBRACK COMMA COLON ]
##
## The known suffix of the stack is as follows:
## LPAREN
@@ -537,7 +537,7 @@ At this point, one of the following is expected:
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 242.
+## Ends in an error in state: 245.
##
## 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 ]
@@ -560,11 +560,11 @@ At this point, one of the following is expected:
translation_unit_file: ALIGNAS LPAREN VOLATILE ADD_ASSIGN
##
-## Ends in an error in state: 307.
+## Ends in an error in state: 310.
##
## 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 COMMA ]
-## specifier_qualifier_list(type_name) -> type_qualifier_list . type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name) [ STAR RPAREN LPAREN LBRACK COMMA ]
+## specifier_qualifier_list(type_name) -> type_qualifier_list . typedef_name option(type_qualifier_list) [ STAR RPAREN LPAREN LBRACK COMMA COLON ]
+## specifier_qualifier_list(type_name) -> type_qualifier_list . type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name) [ STAR RPAREN LPAREN LBRACK COMMA COLON ]
##
## The known suffix of the stack is as follows:
## type_qualifier_list
@@ -585,10 +585,10 @@ At this point, one of the following is expected:
translation_unit_file: ALIGNAS LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 61.
+## Ends in an error in state: 67.
##
-## 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ALIGNAS LPAREN
@@ -607,10 +607,10 @@ At this point, one of the following is expected:
translation_unit_file: ALIGNAS XOR_ASSIGN
##
-## Ends in an error in state: 60.
+## Ends in an error in state: 66.
##
-## 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ALIGNAS
@@ -625,7 +625,7 @@ At this point, an opening parenthesis '(' is expected.
translation_unit_file: ATTRIBUTE LPAREN LPAREN COMMA XOR_ASSIGN
##
-## Ends in an error in state: 353.
+## Ends in an error in state: 338.
##
## gcc_attribute_list -> gcc_attribute_list COMMA . gcc_attribute [ RPAREN COMMA ]
##
@@ -647,9 +647,9 @@ At this point, a gcc attribute is expected.
translation_unit_file: ATTRIBUTE LPAREN LPAREN RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 351.
+## Ends in an error in state: 336.
##
-## 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ATTRIBUTE LPAREN LPAREN gcc_attribute_list RPAREN
@@ -662,9 +662,9 @@ At this point, a second closing parenthesis ')' is expected.
translation_unit_file: ATTRIBUTE LPAREN LPAREN PRE_NAME VAR_NAME LPAREN RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 350.
+## Ends in an error in state: 335.
##
-## 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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:
@@ -685,7 +685,7 @@ At this point, one of the following is expected:
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: 346.
+## Ends in an error in state: 331.
##
## 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 ]
@@ -697,21 +697,21 @@ translation_unit_file: ATTRIBUTE LPAREN LPAREN PRE_NAME VAR_NAME LPAREN PRE_NAME
## 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 155, spurious reduction of production argument_expression_list -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 146, spurious reduction of production argument_expression_list -> assignment_expression
##
# We know for sure that we are parsing a gcc attribute.
@@ -729,7 +729,7 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: ATTRIBUTE LPAREN LPAREN PRE_NAME VAR_NAME LPAREN PRE_NAME TYPEDEF_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 345.
+## Ends in an error in state: 330.
##
## gcc_attribute -> gcc_attribute_word LPAREN typedef_name COMMA . argument_expression_list RPAREN [ RPAREN COMMA ]
##
@@ -746,7 +746,7 @@ At this point, an expression is expected.
translation_unit_file: ATTRIBUTE LPAREN LPAREN PRE_NAME VAR_NAME LPAREN PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 344.
+## Ends in an error in state: 329.
##
## gcc_attribute -> gcc_attribute_word LPAREN typedef_name . COMMA argument_expression_list RPAREN [ RPAREN COMMA ]
##
@@ -763,7 +763,7 @@ At this point, a comma ',' is expected.
translation_unit_file: ATTRIBUTE LPAREN LPAREN PRE_NAME VAR_NAME LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 47.
+## Ends in an error in state: 55.
##
## 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 ]
@@ -783,7 +783,7 @@ At this point, a list of expressions is expected.
translation_unit_file: ATTRIBUTE LPAREN LPAREN PRE_NAME VAR_NAME XOR_ASSIGN
##
-## Ends in an error in state: 46.
+## Ends in an error in state: 54.
##
## gcc_attribute -> gcc_attribute_word . [ RPAREN COMMA ]
## gcc_attribute -> gcc_attribute_word . LPAREN option(argument_expression_list) RPAREN [ RPAREN COMMA ]
@@ -809,9 +809,9 @@ At this point, one of the following is expected:
translation_unit_file: ATTRIBUTE LPAREN LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 39.
+## Ends in an error in state: 47.
##
-## 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ATTRIBUTE LPAREN LPAREN
@@ -831,9 +831,9 @@ At this point, a gcc attribute is expected.
translation_unit_file: ATTRIBUTE LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 38.
+## Ends in an error in state: 46.
##
-## 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ATTRIBUTE LPAREN
@@ -846,9 +846,9 @@ At this point, a second opening parenthesis '(' is expected.
translation_unit_file: ATTRIBUTE XOR_ASSIGN
##
-## Ends in an error in state: 37.
+## Ends in an error in state: 45.
##
-## 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ATTRIBUTE
@@ -861,7 +861,7 @@ At this point, two opening parentheses '((' are expected.
translation_unit_file: ENUM LBRACE PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 361.
+## Ends in an error in state: 346.
##
## enumerator_list -> enumerator_list COMMA . declare_varname(enumerator) [ RBRACE COMMA ]
## option(COMMA) -> COMMA . [ RBRACE ]
@@ -882,7 +882,7 @@ At this point, an enumerator is expected.
translation_unit_file: ENUM LBRACE PRE_NAME VAR_NAME EQ CONSTANT SEMICOLON
##
-## Ends in an error in state: 360.
+## Ends in an error in state: 345.
##
## 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 NORETURN 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 ]
@@ -894,22 +894,22 @@ translation_unit_file: ENUM LBRACE PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 79, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 365, spurious reduction of production enumerator -> enumeration_constant EQ conditional_expression
-## In state 362, spurious reduction of production declare_varname(enumerator) -> enumerator
-## In state 369, spurious reduction of production enumerator_list -> declare_varname(enumerator)
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 69, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 350, spurious reduction of production enumerator -> enumeration_constant EQ conditional_expression
+## In state 347, spurious reduction of production declare_varname(enumerator) -> enumerator
+## In state 354, spurious reduction of production enumerator_list -> declare_varname(enumerator)
##
#
# At first sight, it seems that the last enumerator that we have recognized
@@ -940,7 +940,7 @@ then at this point, a closing brace '}' is expected.
translation_unit_file: ENUM LBRACE PRE_NAME VAR_NAME EQ XOR_ASSIGN
##
-## Ends in an error in state: 364.
+## Ends in an error in state: 349.
##
## enumerator -> enumeration_constant EQ . conditional_expression [ RBRACE COMMA ]
##
@@ -955,7 +955,7 @@ At this point, a constant expression is expected.
translation_unit_file: ENUM LBRACE PRE_NAME VAR_NAME XOR_ASSIGN
##
-## Ends in an error in state: 363.
+## Ends in an error in state: 348.
##
## enumerator -> enumeration_constant . [ RBRACE COMMA ]
## enumerator -> enumeration_constant . EQ conditional_expression [ RBRACE COMMA ]
@@ -976,7 +976,7 @@ At this point, one of the following is expected:
translation_unit_file: ENUM LBRACE XOR_ASSIGN
##
-## Ends in an error in state: 358.
+## Ends in an error in state: 343.
##
## 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 NORETURN LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
@@ -994,7 +994,7 @@ At this point, an enumerator is expected.
translation_unit_file: ENUM XOR_ASSIGN
##
-## Ends in an error in state: 356.
+## Ends in an error in state: 341.
##
## 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 NORETURN 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 NORETURN LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
@@ -1006,7 +1006,7 @@ translation_unit_file: ENUM 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 36, spurious reduction of production attribute_specifier_list ->
+## In state 44, spurious reduction of production attribute_specifier_list ->
##
# Here, both clang and gcc give an incomplete diagnostic message.
@@ -1021,7 +1021,7 @@ At this point, one of the following is expected:
translation_unit_file: INT PRE_NAME VAR_NAME EQ ALIGNOF LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 65.
+## Ends in an error in state: 59.
##
## 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 ]
##
@@ -1054,7 +1054,7 @@ At this point, an expression is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ ALIGNOF XOR_ASSIGN
##
-## Ends in an error in state: 64.
+## Ends in an error in state: 58.
##
## 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 ]
##
@@ -1092,7 +1092,7 @@ followed with an expression or a type name.
translation_unit_file: INT PRE_NAME VAR_NAME EQ BUILTIN_VA_ARG LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 339.
+## Ends in an error in state: 369.
##
## 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 ]
##
@@ -1103,20 +1103,20 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ BUILTIN_VA_ARG LPAREN PRE_NAME V
## 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
##
Ill-formed use of $2.
@@ -1129,7 +1129,7 @@ then at this point, a comma ',' is expected.
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: 340.
+## Ends in an error in state: 370.
##
## 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 ]
##
@@ -1144,7 +1144,7 @@ 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: 51.
+## Ends in an error in state: 40.
##
## 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 ]
##
@@ -1159,7 +1159,7 @@ At this point, an expression is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ BUILTIN_VA_ARG XOR_ASSIGN
##
-## Ends in an error in state: 50.
+## Ends in an error in state: 39.
##
## 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 ]
##
@@ -1174,7 +1174,7 @@ At this point, an opening parenthesis '(' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ DEC XOR_ASSIGN
##
-## Ends in an error in state: 48.
+## Ends in an error in state: 37.
##
## 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 ]
##
@@ -1198,7 +1198,7 @@ At this point, an expression is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ INC LPAREN INT RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 372.
+## Ends in an error in state: 388.
##
## 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 ]
##
@@ -1250,7 +1250,7 @@ At this point, one of the following is expected:
translation_unit_file: INT PRE_NAME VAR_NAME EQ LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 390.
+## Ends in an error in state: 406.
##
## 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 ]
@@ -1262,21 +1262,21 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ LPAREN PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 149, spurious reduction of production expression -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 140, spurious reduction of production expression -> assignment_expression
##
# Since we are saying "if this expression is complete",
@@ -1294,7 +1294,7 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ LPAREN INT RPAREN LBRACE PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 387.
+## Ends in an error in state: 403.
##
## 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 ]
@@ -1306,22 +1306,22 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ LPAREN INT RPAREN LBRACE PRE_NAM
## 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 380, spurious reduction of production c_initializer -> assignment_expression
-## In state 386, spurious reduction of production initializer_list -> option(designation) c_initializer
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 396, spurious reduction of production c_initializer -> assignment_expression
+## In state 402, spurious reduction of production initializer_list -> option(designation) c_initializer
##
# Let's ignore the fact that a comma can precede a closing brace.
@@ -1336,7 +1336,7 @@ then at this point, a closing brace '}' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ LPAREN INT RPAREN LBRACE XOR_ASSIGN
##
-## Ends in an error in state: 373.
+## Ends in an error in state: 389.
##
## 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 ]
##
@@ -1353,7 +1353,7 @@ At this point, an initializer is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ LPAREN INT RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 394.
+## Ends in an error in state: 410.
##
## 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 ]
@@ -1397,7 +1397,7 @@ At this point, one of the following is expected:
translation_unit_file: INT PRE_NAME VAR_NAME EQ TILDE XOR_ASSIGN
##
-## Ends in an error in state: 78.
+## Ends in an error in state: 68.
##
## 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 ]
##
@@ -1414,7 +1414,7 @@ At this point, an expression is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME AND XOR_ASSIGN
##
-## Ends in an error in state: 138.
+## Ends in an error in state: 129.
##
## and_expression -> and_expression AND . equality_expression [ SEMICOLON RPAREN RBRACK RBRACE QUESTION HAT COMMA COLON BARBAR BAR ANDAND AND ]
##
@@ -1423,7 +1423,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME AND XOR_ASSIGN
##
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME ANDAND XOR_ASSIGN
##
-## Ends in an error in state: 127.
+## Ends in an error in state: 118.
##
## logical_and_expression -> logical_and_expression ANDAND . inclusive_or_expression [ SEMICOLON RPAREN RBRACK RBRACE QUESTION COMMA COLON BARBAR ANDAND ]
##
@@ -1432,7 +1432,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME ANDAND XOR_ASS
##
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME BAR XOR_ASSIGN
##
-## Ends in an error in state: 129.
+## Ends in an error in state: 120.
##
## inclusive_or_expression -> inclusive_or_expression BAR . exclusive_or_expression [ SEMICOLON RPAREN RBRACK RBRACE QUESTION COMMA COLON BARBAR BAR ANDAND ]
##
@@ -1441,7 +1441,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME BAR XOR_ASSIGN
##
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME BARBAR XOR_ASSIGN
##
-## Ends in an error in state: 150.
+## Ends in an error in state: 141.
##
## logical_or_expression -> logical_or_expression BARBAR . logical_and_expression [ SEMICOLON RPAREN RBRACK RBRACE QUESTION COMMA COLON BARBAR ]
##
@@ -1450,7 +1450,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME BARBAR XOR_ASS
##
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME HAT XOR_ASSIGN
##
-## Ends in an error in state: 131.
+## Ends in an error in state: 122.
##
## exclusive_or_expression -> exclusive_or_expression HAT . and_expression [ SEMICOLON RPAREN RBRACK RBRACE QUESTION HAT COMMA COLON BARBAR BAR ANDAND ]
##
@@ -1459,7 +1459,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME HAT XOR_ASSIGN
##
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME LT XOR_ASSIGN
##
-## Ends in an error in state: 121.
+## Ends in an error in state: 112.
##
## 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 ]
##
@@ -1468,7 +1468,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME LT XOR_ASSIGN
##
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME NEQ XOR_ASSIGN
##
-## Ends in an error in state: 135.
+## Ends in an error in state: 126.
##
## equality_expression -> equality_expression equality_operator . relational_expression [ SEMICOLON RPAREN RBRACK RBRACE QUESTION NEQ HAT EQEQ COMMA COLON BARBAR BAR ANDAND AND ]
##
@@ -1477,7 +1477,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME NEQ XOR_ASSIGN
##
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME PLUS XOR_ASSIGN
##
-## Ends in an error in state: 114.
+## Ends in an error in state: 105.
##
## 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 ]
##
@@ -1486,7 +1486,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME PLUS XOR_ASSIG
##
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME RIGHT XOR_ASSIGN
##
-## Ends in an error in state: 103.
+## Ends in an error in state: 93.
##
## 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 ]
##
@@ -1495,7 +1495,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME RIGHT XOR_ASSI
##
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME STAR XOR_ASSIGN
##
-## Ends in an error in state: 108.
+## Ends in an error in state: 98.
##
## 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 ]
##
@@ -1512,7 +1512,7 @@ At this point, an expression is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME XOR_ASSIGN XOR_ASSIGN
##
-## Ends in an error in state: 99.
+## Ends in an error in state: 89.
##
## assignment_expression -> unary_expression assignment_operator . assignment_expression [ SEMICOLON RPAREN RBRACK RBRACE COMMA COLON ]
##
@@ -1529,7 +1529,7 @@ At this point, an expression is expected.
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: 157.
+## Ends in an error in state: 148.
##
## argument_expression_list -> argument_expression_list COMMA . assignment_expression [ RPAREN COMMA ]
##
@@ -1550,7 +1550,7 @@ At this point, an expression is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME DOT XOR_ASSIGN
##
-## Ends in an error in state: 163.
+## Ends in an error in state: 154.
##
## 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 ]
##
@@ -1559,7 +1559,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME DOT XOR_ASSIGN
##
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME PTR XOR_ASSIGN
##
-## Ends in an error in state: 84.
+## Ends in an error in state: 74.
##
## 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 ]
##
@@ -1576,7 +1576,7 @@ At this point, the name of a struct or union member is expected.
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: 160.
+## Ends in an error in state: 151.
##
## 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 ]
@@ -1588,21 +1588,21 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME LBRACK PRE_NAM
## 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 149, spurious reduction of production expression -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 140, spurious reduction of production expression -> assignment_expression
##
# We know for sure that an array subscript expression has begun, and
@@ -1621,7 +1621,7 @@ then at this point, a closing bracket ']' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME LBRACK XOR_ASSIGN
##
-## Ends in an error in state: 159.
+## Ends in an error in state: 150.
##
## 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 ]
##
@@ -1636,7 +1636,7 @@ At this point, an expression is expected.
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: 156.
+## Ends in an error in state: 147.
##
## argument_expression_list -> argument_expression_list . COMMA assignment_expression [ RPAREN COMMA ]
## option(argument_expression_list) -> argument_expression_list . [ RPAREN ]
@@ -1648,21 +1648,21 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME LPAREN PRE_NAM
## 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 155, spurious reduction of production argument_expression_list -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 146, spurious reduction of production argument_expression_list -> assignment_expression
##
Up to this point, a list of expressions has been recognized:
@@ -1674,7 +1674,7 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 86.
+## Ends in an error in state: 76.
##
## 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 ]
##
@@ -1692,7 +1692,7 @@ followed with a closing parenthesis ')', is expected.
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: 147.
+## Ends in an error in state: 138.
##
## conditional_expression -> logical_or_expression QUESTION expression COLON . conditional_expression [ SEMICOLON RPAREN RBRACK RBRACE COMMA COLON ]
##
@@ -1701,7 +1701,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME QUESTION PRE_N
##
translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME QUESTION XOR_ASSIGN
##
-## Ends in an error in state: 125.
+## Ends in an error in state: 116.
##
## conditional_expression -> logical_or_expression QUESTION . expression COLON conditional_expression [ SEMICOLON RPAREN RBRACK RBRACE COMMA COLON ]
##
@@ -1716,7 +1716,7 @@ At this point, an expression is expected.
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: 143.
+## Ends in an error in state: 134.
##
## conditional_expression -> logical_or_expression QUESTION expression . COLON conditional_expression [ SEMICOLON RPAREN RBRACK RBRACE COMMA COLON ]
## expression -> expression . COMMA assignment_expression [ COMMA COLON ]
@@ -1728,21 +1728,21 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ PRE_NAME VAR_NAME QUESTION PRE_N
## 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 149, spurious reduction of production expression -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 140, spurious reduction of production expression -> assignment_expression
##
# gcc and clang simply expect a colon.
@@ -1759,10 +1759,10 @@ then at this point, a colon ':' is expected.
translation_unit_file: PACKED LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 399.
+## Ends in an error in state: 415.
##
## argument_expression_list -> argument_expression_list . COMMA assignment_expression [ RPAREN COMMA ]
-## 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## PACKED LPAREN argument_expression_list
@@ -1771,21 +1771,21 @@ translation_unit_file: PACKED LPAREN PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 155, spurious reduction of production argument_expression_list -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 146, spurious reduction of production argument_expression_list -> assignment_expression
##
Ill-formed $2 attribute.
@@ -1800,7 +1800,7 @@ translation_unit_file: PACKED LPAREN XOR_ASSIGN
##
## Ends in an error in state: 19.
##
-## 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## PACKED LPAREN
@@ -1835,7 +1835,7 @@ translation_unit_file: PACKED XOR_ASSIGN
##
## Ends in an error in state: 18.
##
-## 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 NORETURN MINUS LPAREN LONG LBRACK LBRACE INT INLINE INC GENERIC FLOAT EXTERN EQ ENUM DOUBLE DEC CONSTANT CONST COMMA COLON CHAR BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AUTO ATTRIBUTE AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## PACKED
@@ -1876,7 +1876,7 @@ At this point, one of the following is expected:
translation_unit_file: TYPEDEF PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 402.
+## Ends in an error in state: 418.
##
## 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 NORETURN LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
@@ -1889,7 +1889,7 @@ translation_unit_file: TYPEDEF PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
translation_unit_file: PRE_NAME TYPEDEF_NAME TYPEDEF XOR_ASSIGN
##
-## Ends in an error in state: 411.
+## Ends in an error in state: 427.
##
## 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 NORETURN LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
@@ -1902,7 +1902,7 @@ translation_unit_file: PRE_NAME TYPEDEF_NAME TYPEDEF XOR_ASSIGN
##
translation_unit_file: VOLATILE TYPEDEF PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 422.
+## Ends in an error in state: 438.
##
## 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 NORETURN LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
@@ -1915,7 +1915,7 @@ translation_unit_file: VOLATILE TYPEDEF PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
translation_unit_file: VOLATILE PRE_NAME TYPEDEF_NAME TYPEDEF XOR_ASSIGN
##
-## Ends in an error in state: 428.
+## Ends in an error in state: 444.
##
## 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 NORETURN LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
@@ -1928,7 +1928,7 @@ translation_unit_file: VOLATILE PRE_NAME TYPEDEF_NAME TYPEDEF XOR_ASSIGN
##
translation_unit_file: TYPEDEF INT XOR_ASSIGN
##
-## Ends in an error in state: 404.
+## Ends in an error in state: 420.
##
## 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 NORETURN LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
@@ -1938,7 +1938,7 @@ translation_unit_file: TYPEDEF INT XOR_ASSIGN
##
translation_unit_file: INT TYPEDEF XOR_ASSIGN
##
-## Ends in an error in state: 415.
+## Ends in an error in state: 431.
##
## 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 NORETURN LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
@@ -1948,7 +1948,7 @@ translation_unit_file: INT TYPEDEF XOR_ASSIGN
##
translation_unit_file: VOLATILE TYPEDEF INT XOR_ASSIGN
##
-## Ends in an error in state: 424.
+## Ends in an error in state: 440.
##
## 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 NORETURN LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
@@ -1958,7 +1958,7 @@ translation_unit_file: VOLATILE TYPEDEF INT XOR_ASSIGN
##
translation_unit_file: VOLATILE INT TYPEDEF XOR_ASSIGN
##
-## Ends in an error in state: 432.
+## Ends in an error in state: 448.
##
## 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 NORETURN LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
@@ -2011,7 +2011,7 @@ translation_unit_file: TYPEDEF XOR_ASSIGN
##
translation_unit_file: VOLATILE TYPEDEF XOR_ASSIGN
##
-## Ends in an error in state: 420.
+## Ends in an error in state: 436.
##
## 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 ]
@@ -2045,7 +2045,7 @@ At this point, one of the following is expected:
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN VOLATILE XOR_ASSIGN
##
-## Ends in an error in state: 230.
+## Ends in an error in state: 233.
##
## 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 ]
@@ -2057,7 +2057,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN VOLATILE 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 222, spurious reduction of production rlist(declaration_specifier_no_type) -> type_qualifier_noattr
+## In state 225, 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,
@@ -2081,7 +2081,7 @@ At this point, one of the following is expected:
translation_unit_file: PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 409.
+## Ends in an error in state: 425.
##
## 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 ]
@@ -2095,7 +2095,7 @@ translation_unit_file: PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
translation_unit_file: VOLATILE PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 426.
+## Ends in an error in state: 442.
##
## 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 ]
@@ -2109,7 +2109,7 @@ translation_unit_file: VOLATILE PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
translation_unit_file: INT XOR_ASSIGN
##
-## Ends in an error in state: 413.
+## Ends in an error in state: 429.
##
## 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 ]
@@ -2120,7 +2120,7 @@ translation_unit_file: INT XOR_ASSIGN
##
translation_unit_file: VOLATILE INT XOR_ASSIGN
##
-## Ends in an error in state: 430.
+## Ends in an error in state: 446.
##
## 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 ]
@@ -2181,7 +2181,7 @@ At this point, one of the following is expected:
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 209.
+## Ends in an error in state: 212.
##
## 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 NORETURN LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
@@ -2194,7 +2194,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME TYPEDEF_NAME XOR_AS
##
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN VOLATILE PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 232.
+## Ends in an error in state: 235.
##
## 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 NORETURN LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
@@ -2207,7 +2207,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN VOLATILE PRE_NAME TYPEDEF_NA
##
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT XOR_ASSIGN
##
-## Ends in an error in state: 215.
+## Ends in an error in state: 218.
##
## 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 NORETURN LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA CHAR AUTO ATTRIBUTE ALIGNAS ]
@@ -2217,7 +2217,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT XOR_ASSIGN
##
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN VOLATILE INT XOR_ASSIGN
##
-## Ends in an error in state: 234.
+## Ends in an error in state: 237.
##
## 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 NORETURN LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA CHAR AUTO ATTRIBUTE ALIGNAS ]
@@ -2266,7 +2266,7 @@ At this point, one of the following is expected:
translation_unit_file: VOLATILE XOR_ASSIGN
##
-## Ends in an error in state: 418.
+## Ends in an error in state: 434.
##
## 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 ]
@@ -2282,7 +2282,7 @@ translation_unit_file: VOLATILE 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 222, spurious reduction of production rlist(declaration_specifier_no_type) -> type_qualifier_noattr
+## In state 225, 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
@@ -2311,7 +2311,7 @@ At this point, one of the following is expected:
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: 528.
+## Ends in an error in state: 544.
##
## 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 ]
@@ -2327,7 +2327,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME 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 222, spurious reduction of production rlist(declaration_specifier_no_type) -> type_qualifier_noattr
+## In state 225, 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.
@@ -2342,7 +2342,7 @@ At this point, one of the following is expected:
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: 524.
+## Ends in an error in state: 540.
##
## 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 ]
@@ -2356,7 +2356,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN
##
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: 530.
+## Ends in an error in state: 546.
##
## 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 ]
@@ -2370,7 +2370,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN
##
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: 526.
+## Ends in an error in state: 542.
##
## 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 ]
@@ -2381,7 +2381,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN
##
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: 532.
+## Ends in an error in state: 548.
##
## 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 ]
@@ -2421,7 +2421,7 @@ At this point, one of the following is expected:
translation_unit_file: UNION LBRACE PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 189.
+## Ends in an error in state: 192.
##
## struct_declaration -> specifier_qualifier_list(struct_declaration) . option(struct_declarator_list) SEMICOLON [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC_ASSERT SIGNED SHORT RESTRICT RBRACE PRE_NAME PACKED LONG INT FLOAT ENUM DOUBLE CONST CHAR ATTRIBUTE ALIGNAS ]
##
@@ -2432,7 +2432,7 @@ translation_unit_file: UNION LBRACE PRE_NAME 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 174, spurious reduction of production specifier_qualifier_list(struct_declaration) -> typedef_name option(type_qualifier_list)
+## In state 177, spurious reduction of production specifier_qualifier_list(struct_declaration) -> typedef_name option(type_qualifier_list)
##
# We have (spuriously) recognized a specifier_qualifier_list,
@@ -2466,7 +2466,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: 295.
+## Ends in an error in state: 298.
##
## option(struct_declarator_list) -> struct_declarator_list . [ SEMICOLON ]
## struct_declarator_list -> struct_declarator_list . COMMA struct_declarator [ SEMICOLON COMMA ]
@@ -2478,21 +2478,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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 79, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 300, spurious reduction of production struct_declarator -> option(declarator) COLON conditional_expression
-## In state 302, spurious reduction of production struct_declarator_list -> struct_declarator
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 69, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 303, spurious reduction of production struct_declarator -> option(declarator) COLON conditional_expression
+## In state 305, spurious reduction of production struct_declarator_list -> struct_declarator
##
# We have seen a non-empty struct_declarator_list.
@@ -2510,7 +2510,7 @@ then at this point, a semicolon ';' is expected.
translation_unit_file: UNION LBRACE INT COLON XOR_ASSIGN
##
-## Ends in an error in state: 299.
+## Ends in an error in state: 302.
##
## struct_declarator -> option(declarator) COLON . conditional_expression [ SEMICOLON COMMA ]
##
@@ -2525,7 +2525,7 @@ At this point, a constant expression is expected.
translation_unit_file: UNION LBRACE INT PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 296.
+## Ends in an error in state: 299.
##
## struct_declarator_list -> struct_declarator_list COMMA . struct_declarator [ SEMICOLON COMMA ]
##
@@ -2540,7 +2540,7 @@ At this point, a struct declarator is expected.
translation_unit_file: UNION LBRACE INT PRE_NAME VAR_NAME RPAREN
##
-## Ends in an error in state: 301.
+## Ends in an error in state: 304.
##
## option(declarator) -> declarator . [ COLON ]
## struct_declarator -> declarator . [ SEMICOLON COMMA ]
@@ -2552,9 +2552,9 @@ translation_unit_file: UNION LBRACE INT PRE_NAME 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 263, spurious reduction of production declarator_noattrend -> direct_declarator
-## In state 268, spurious reduction of production attribute_specifier_list ->
-## In state 269, spurious reduction of production declarator -> declarator_noattrend attribute_specifier_list
+## In state 266, spurious reduction of production declarator_noattrend -> direct_declarator
+## In state 271, spurious reduction of production attribute_specifier_list ->
+## In state 272, spurious reduction of production declarator -> declarator_noattrend attribute_specifier_list
##
# Assuming the declarator so far is complete, we expect
@@ -2577,7 +2577,7 @@ then at this point, one of the following is expected:
translation_unit_file: UNION LBRACE VOLATILE ADD_ASSIGN
##
-## Ends in an error in state: 182.
+## Ends in an error in state: 185.
##
## 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 ]
@@ -2600,7 +2600,7 @@ At this point, one of the following is expected:
translation_unit_file: UNION LBRACE XOR_ASSIGN
##
-## Ends in an error in state: 75.
+## Ends in an error in state: 167.
##
## struct_declaration_list -> struct_declaration_list . struct_declaration [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC_ASSERT 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 NORETURN LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
@@ -2620,7 +2620,7 @@ At this point, one of the following is expected:
translation_unit_file: UNION XOR_ASSIGN
##
-## Ends in an error in state: 72.
+## Ends in an error in state: 164.
##
## 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 NORETURN 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 NORETURN LPAREN LONG LBRACK INT INLINE FLOAT EXTERN ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
@@ -2632,7 +2632,7 @@ translation_unit_file: UNION 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 71, spurious reduction of production attribute_specifier_list ->
+## In state 163, spurious reduction of production attribute_specifier_list ->
##
# gcc expects '{'.
@@ -2649,7 +2649,7 @@ At this point, one of the following is expected:
translation_unit_file: INT LPAREN PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 272.
+## Ends in an error in state: 275.
##
## direct_declarator -> LPAREN save_context declarator . RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED NORETURN LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
@@ -2660,9 +2660,9 @@ translation_unit_file: INT LPAREN PRE_NAME 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 263, spurious reduction of production declarator_noattrend -> direct_declarator
-## In state 268, spurious reduction of production attribute_specifier_list ->
-## In state 269, spurious reduction of production declarator -> declarator_noattrend attribute_specifier_list
+## In state 266, spurious reduction of production declarator_noattrend -> direct_declarator
+## In state 271, spurious reduction of production attribute_specifier_list ->
+## In state 272, spurious reduction of production declarator -> declarator_noattrend attribute_specifier_list
##
Up to this point, a declarator has been recognized:
@@ -2674,7 +2674,7 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: INT LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 195.
+## Ends in an error in state: 198.
##
## direct_declarator -> LPAREN save_context . declarator RPAREN [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED NORETURN LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
@@ -2691,7 +2691,7 @@ At this point, a declarator is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME XOR_ASSIGN
##
-## Ends in an error in state: 289.
+## Ends in an error in state: 292.
##
## identifier_list -> identifier_list . COMMA PRE_NAME VAR_NAME [ RPAREN COMMA ]
## option(identifier_list) -> identifier_list . [ RPAREN ]
@@ -2710,7 +2710,7 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 201.
+## Ends in an error in state: 204.
##
## 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 NORETURN LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
@@ -2734,7 +2734,7 @@ followed with a closing parenthesis ')', is expected.
translation_unit_file: INT STAR RPAREN
##
-## Ends in an error in state: 198.
+## Ends in an error in state: 201.
##
## 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 NORETURN 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 ]
@@ -2767,7 +2767,7 @@ At this point, one of the following is expected:
translation_unit_file: TYPEDEF INT PRE_NAME VAR_NAME XOR_ASSIGN
##
-## Ends in an error in state: 544.
+## Ends in an error in state: 560.
##
## option(typedef_declarator_list) -> typedef_declarator_list . [ SEMICOLON ]
## typedef_declarator_list -> typedef_declarator_list . COMMA typedef_declarator [ SEMICOLON COMMA ]
@@ -2779,12 +2779,12 @@ translation_unit_file: TYPEDEF INT PRE_NAME 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 263, spurious reduction of production declarator_noattrend -> direct_declarator
-## In state 268, spurious reduction of production attribute_specifier_list ->
-## In state 269, spurious reduction of production declarator -> declarator_noattrend attribute_specifier_list
-## In state 548, spurious reduction of production declare_typename(declarator) -> declarator
-## In state 547, spurious reduction of production typedef_declarator -> declare_typename(declarator)
-## In state 549, spurious reduction of production typedef_declarator_list -> typedef_declarator
+## In state 266, spurious reduction of production declarator_noattrend -> direct_declarator
+## In state 271, spurious reduction of production attribute_specifier_list ->
+## In state 272, spurious reduction of production declarator -> declarator_noattrend attribute_specifier_list
+## In state 564, spurious reduction of production declare_typename(declarator) -> declarator
+## In state 563, spurious reduction of production typedef_declarator -> declare_typename(declarator)
+## In state 565, spurious reduction of production typedef_declarator_list -> typedef_declarator
##
# Because attribute_specifier_list, declarator and declarator_noattrend have been marked
@@ -2810,7 +2810,7 @@ then at this point, a semicolon ';' is expected.
translation_unit_file: TYPEDEF INT PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 545.
+## Ends in an error in state: 561.
##
## typedef_declarator_list -> typedef_declarator_list COMMA . typedef_declarator [ SEMICOLON COMMA ]
##
@@ -2824,9 +2824,9 @@ At this point, a declarator is expected.
translation_unit_file: ALIGNAS LPAREN INT LPAREN RPAREN LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 257.
+## Ends in an error in state: 260.
##
-## direct_abstract_declarator -> direct_abstract_declarator LPAREN . option(context_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 COLON ]
##
## The known suffix of the stack is as follows:
## direct_abstract_declarator LPAREN
@@ -2839,7 +2839,7 @@ followed with a closing parenthesis ')', is expected.
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: 459.
+## Ends in an error in state: 475.
##
## asm_attributes -> CONST . asm_attributes [ LPAREN ]
##
@@ -2848,7 +2848,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN
##
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: 458.
+## Ends in an error in state: 474.
##
## asm_attributes -> VOLATILE . asm_attributes [ LPAREN ]
##
@@ -2857,9 +2857,9 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN
##
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: 457.
+## Ends in an error in state: 473.
##
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ASM
@@ -2874,7 +2874,7 @@ At this point, one of the following is expected:
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: 483.
+## Ends in an error in state: 499.
##
## asm_flags -> asm_flags COMMA . string_literals_list [ RPAREN COMMA ]
##
@@ -2886,7 +2886,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN
# first(asm_flags) = STRING_LITERAL
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: 480.
+## Ends in an error in state: 496.
##
## asm_arguments -> COLON asm_operands COLON asm_operands COLON . asm_flags [ RPAREN ]
##
@@ -2906,7 +2906,7 @@ Examples of clobbered resources:
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: 482.
+## Ends in an error in state: 498.
##
## asm_arguments -> COLON asm_operands COLON asm_operands COLON asm_flags . [ RPAREN ]
## asm_flags -> asm_flags . COMMA string_literals_list [ RPAREN COMMA ]
@@ -2918,7 +2918,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME 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 481, spurious reduction of production asm_flags -> string_literals_list
+## In state 497, spurious reduction of production asm_flags -> string_literals_list
##
# Let's ignore the possibility of concatenating string literals.
@@ -2935,7 +2935,7 @@ then at this point, a closing parenthesis ')' is expected.
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: 477.
+## Ends in an error in state: 493.
##
## asm_arguments -> COLON asm_operands . [ RPAREN ]
## asm_arguments -> COLON asm_operands . COLON asm_operands [ RPAREN ]
@@ -2948,7 +2948,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME 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 469, spurious reduction of production asm_operands -> asm_operands_ne
+## In state 485, spurious reduction of production asm_operands -> asm_operands_ne
##
# We have seen one COLON, hence the outputs. (The list of outputs may be empty.)
@@ -2965,7 +2965,7 @@ then at this point, one of the following is expected:
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: 479.
+## Ends in an error in state: 495.
##
## asm_arguments -> COLON asm_operands COLON asm_operands . [ RPAREN ]
## asm_arguments -> COLON asm_operands COLON asm_operands . COLON asm_flags [ RPAREN ]
@@ -2977,7 +2977,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME 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 478, spurious reduction of production asm_operands ->
+## In state 494, spurious reduction of production asm_operands ->
##
# We have seen two COLONs, hence the outputs and inputs. (The list of inputs may be empty.)
@@ -2997,7 +2997,7 @@ then at this point, one of the following is expected:
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: 472.
+## Ends in an error in state: 488.
##
## asm_operand -> asm_op_name . string_literals_list LPAREN expression RPAREN [ RPAREN COMMA COLON ]
##
@@ -3016,7 +3016,7 @@ At this point, a string literal, representing a constraint, is expected.
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: 467.
+## Ends in an error in state: 483.
##
## asm_op_name -> LBRACK general_identifier . RBRACK [ STRING_LITERAL ]
##
@@ -3031,7 +3031,7 @@ At this point, a closing bracket ']' is expected.
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: 466.
+## Ends in an error in state: 482.
##
## asm_op_name -> LBRACK . general_identifier RBRACK [ STRING_LITERAL ]
##
@@ -3046,7 +3046,7 @@ At this point, an identifier is expected.
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: 470.
+## Ends in an error in state: 486.
##
## asm_operands_ne -> asm_operands_ne COMMA . asm_operand [ RPAREN COMMA COLON ]
##
@@ -3063,7 +3063,7 @@ At this point, an assembly operand is expected.
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: 475.
+## Ends in an error in state: 491.
##
## asm_operand -> asm_op_name string_literals_list LPAREN expression . RPAREN [ RPAREN COMMA COLON ]
## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
@@ -3075,21 +3075,21 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 149, spurious reduction of production expression -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 140, spurious reduction of production expression -> assignment_expression
##
Ill-formed assembly operand.
@@ -3102,7 +3102,7 @@ then at this point, a closing parenthesis ')' is expected.
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: 474.
+## Ends in an error in state: 490.
##
## asm_operand -> asm_op_name string_literals_list LPAREN . expression RPAREN [ RPAREN COMMA COLON ]
##
@@ -3117,7 +3117,7 @@ At this point, an expression is expected.
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: 473.
+## Ends in an error in state: 489.
##
## 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 ]
@@ -3137,9 +3137,9 @@ followed with an expression and a closing parenthesis ')', is expected.
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: 464.
+## Ends in an error in state: 480.
##
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF 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:
@@ -3159,9 +3159,9 @@ At this point, one of the following is expected:
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 463.
+## Ends in an error in state: 479.
##
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## ASM asm_attributes LPAREN
@@ -3174,45 +3174,45 @@ At this point, a string literal, representing an instruction, is expected.
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: 455.
+## Ends in an error in state: 471.
##
-## jump_statement -> BREAK . SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## BREAK
##
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE CONTINUE XOR_ASSIGN
##
-## Ends in an error in state: 450.
+## Ends in an error in state: 466.
##
-## jump_statement -> CONTINUE . SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## CONTINUE
##
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: 575.
+## Ends in an error in state: 591.
##
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context do_statement1 WHILE LPAREN expression RPAREN
##
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: 446.
+## Ends in an error in state: 462.
##
-## jump_statement -> GOTO general_identifier . SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## GOTO general_identifier
##
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE ASM LPAREN STRING_LITERAL RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 487.
+## Ends in an error in state: 503.
##
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF 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
@@ -3225,27 +3225,27 @@ At this point, a semicolon ';' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE CASE CONSTANT COLON XOR_ASSIGN
##
-## Ends in an error in state: 454.
+## Ends in an error in state: 470.
##
-## labeled_statement -> CASE conditional_expression COLON . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF 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 PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DEFAULT COLON XOR_ASSIGN
##
-## Ends in an error in state: 449.
+## Ends in an error in state: 465.
##
-## labeled_statement -> DEFAULT COLON . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## DEFAULT COLON
##
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: 503.
+## Ends in an error in state: 519.
##
-## labeled_statement -> general_identifier COLON . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## general_identifier COLON
@@ -3260,9 +3260,9 @@ At this point, a statement is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE CASE CONSTANT SEMICOLON
##
-## Ends in an error in state: 453.
+## Ends in an error in state: 469.
##
-## labeled_statement -> CASE conditional_expression . COLON statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## CASE conditional_expression
@@ -3271,19 +3271,19 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 79, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 69, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
##
Ill-formed labeled statement.
@@ -3296,9 +3296,9 @@ then at this point, a colon ':' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE CASE XOR_ASSIGN
##
-## Ends in an error in state: 452.
+## Ends in an error in state: 468.
##
-## labeled_statement -> CASE . conditional_expression COLON statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## CASE
@@ -3311,18 +3311,18 @@ At this point, a constant expression is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DEFAULT XOR_ASSIGN
##
-## Ends in an error in state: 448.
+## Ends in an error in state: 464.
##
-## labeled_statement -> DEFAULT . COLON statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## DEFAULT
##
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: 502.
+## Ends in an error in state: 518.
##
-## labeled_statement -> general_identifier . COLON statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## general_identifier
@@ -3337,10 +3337,10 @@ At this point, a colon ':' is expected.
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: 574.
+## Ends in an error in state: 590.
##
## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context do_statement1 WHILE LPAREN expression
@@ -3349,21 +3349,21 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 149, spurious reduction of production expression -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 140, spurious reduction of production expression -> assignment_expression
##
Ill-formed 'do' ... 'while' statement.
@@ -3376,9 +3376,9 @@ then at this point, a closing parenthesis ')' and a semicolon ';' are expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DO SEMICOLON WHILE LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 573.
+## Ends in an error in state: 589.
##
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context do_statement1 WHILE LPAREN
@@ -3391,9 +3391,9 @@ At this point, an expression is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DO SEMICOLON WHILE XOR_ASSIGN
##
-## Ends in an error in state: 572.
+## Ends in an error in state: 588.
##
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context do_statement1 WHILE
@@ -3406,9 +3406,9 @@ At this point, an opening parenthesis '(' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DO SEMICOLON XOR_ASSIGN
##
-## Ends in an error in state: 571.
+## Ends in an error in state: 587.
##
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context do_statement1
@@ -3427,7 +3427,7 @@ At this point, a 'while' keyword is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE DO XOR_ASSIGN
##
-## Ends in an error in state: 567.
+## Ends in an error in state: 583.
##
## do_statement1 -> save_context DO . statement [ WHILE ]
##
@@ -3444,9 +3444,9 @@ At this point, a statement (the loop body) is expected.
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: 537.
+## Ends in an error in state: 553.
##
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context FOR LPAREN for_statement_header optional(expression,SEMICOLON) optional(expression,RPAREN)
@@ -3459,10 +3459,10 @@ At this point, a statement (the loop body) is expected.
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: 539.
+## Ends in an error in state: 555.
##
## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## 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 BUILTIN_OFFSETOF 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 GENERIC FOR DO DEFAULT DEC CONTINUE CONSTANT CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG ASM AND ALIGNOF ]
##
## The known suffix of the stack is as follows:
## expression
@@ -3471,21 +3471,21 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 149, spurious reduction of production expression -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 140, spurious reduction of production expression -> assignment_expression
##
# The use of optional(expression,RPAREN) tells us that we are in a FOR statement.
@@ -3501,9 +3501,9 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE FOR LPAREN SEMICOLON SEMICOLON XOR_ASSIGN
##
-## Ends in an error in state: 535.
+## Ends in an error in state: 551.
##
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context FOR LPAREN for_statement_header optional(expression,SEMICOLON)
@@ -3521,9 +3521,9 @@ followed with a closing parenthesis ')', is expected.
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: 534.
+## Ends in an error in state: 550.
##
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context FOR LPAREN for_statement_header
@@ -3540,10 +3540,10 @@ followed with a semicolon ';', is expected.
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: 541.
+## Ends in an error in state: 557.
##
## expression -> expression . COMMA assignment_expression [ SEMICOLON COMMA ]
-## optional(expression,SEMICOLON) -> expression . SEMICOLON [ TILDE STRING_LITERAL STAR SIZEOF SEMICOLON RPAREN PRE_NAME PLUS MINUS LPAREN INC DEC CONSTANT BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AND ALIGNOF ]
+## optional(expression,SEMICOLON) -> expression . SEMICOLON [ TILDE STRING_LITERAL STAR SIZEOF SEMICOLON RPAREN PRE_NAME PLUS MINUS LPAREN INC GENERIC DEC CONSTANT BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AND ALIGNOF ]
##
## The known suffix of the stack is as follows:
## expression
@@ -3552,21 +3552,21 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 149, spurious reduction of production expression -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 140, spurious reduction of production expression -> assignment_expression
##
# At the time of writing, optional(expression,SEMICOLON) is used only in FOR
@@ -3582,9 +3582,9 @@ then at this point, a semicolon ';' is expected.
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: 521.
+## Ends in an error in state: 537.
##
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context FOR LPAREN
@@ -3603,9 +3603,9 @@ At this point, one of the following is expected:
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE FOR XOR_ASSIGN
##
-## Ends in an error in state: 520.
+## Ends in an error in state: 536.
##
-## 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context FOR
@@ -3618,9 +3618,9 @@ At this point, an opening parenthesis '(' is expected.
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: 445.
+## Ends in an error in state: 461.
##
-## jump_statement -> GOTO . general_identifier SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## GOTO
@@ -3633,9 +3633,9 @@ At this point, an identifier (a 'goto' label) is expected.
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: 569.
+## Ends in an error in state: 585.
##
-## selection_statement -> save_context ifelse_statement1 . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## selection_statement -> save_context ifelse_statement1 . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context ifelse_statement1
@@ -3648,10 +3648,10 @@ At this point, a statement is expected.
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: 517.
+## Ends in an error in state: 533.
##
-## 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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 GENERIC FOR DO DEFAULT DEC CONTINUE CONSTANT CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context IF LPAREN expression RPAREN save_context
@@ -3664,11 +3664,11 @@ At this point, a statement is expected.
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: 515.
+## Ends in an error in state: 531.
##
## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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 GENERIC FOR DO DEFAULT DEC CONTINUE CONSTANT CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context IF LPAREN expression
@@ -3677,21 +3677,21 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 149, spurious reduction of production expression -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 140, spurious reduction of production expression -> assignment_expression
##
Ill-formed 'if' statement.
@@ -3704,10 +3704,10 @@ then at this point, a closing parenthesis ')' is expected.
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: 514.
+## Ends in an error in state: 530.
##
-## 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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 GENERIC FOR DO DEFAULT DEC CONTINUE CONSTANT CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context IF LPAREN
@@ -3720,10 +3720,10 @@ At this point, an expression is expected.
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: 513.
+## Ends in an error in state: 529.
##
-## 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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 GENERIC FOR DO DEFAULT DEC CONTINUE CONSTANT CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context IF
@@ -3736,9 +3736,9 @@ At this point, an opening parenthesis '(' is expected.
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
##
-## Ends in an error in state: 511.
+## Ends in an error in state: 527.
##
-## selection_statement -> save_context SWITCH LPAREN expression RPAREN . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context SWITCH LPAREN expression RPAREN
@@ -3760,10 +3760,10 @@ enclosed within braces '{' and '}'.
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: 510.
+## Ends in an error in state: 526.
##
## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## selection_statement -> save_context SWITCH LPAREN expression . RPAREN statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context SWITCH LPAREN expression
@@ -3772,21 +3772,21 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 149, spurious reduction of production expression -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 140, spurious reduction of production expression -> assignment_expression
##
Ill-formed 'switch' statement.
@@ -3799,9 +3799,9 @@ then at this point, a closing parenthesis ')' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE SWITCH LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 509.
+## Ends in an error in state: 525.
##
-## selection_statement -> save_context SWITCH LPAREN . expression RPAREN statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context SWITCH LPAREN
@@ -3814,9 +3814,9 @@ At this point, an expression is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE SWITCH XOR_ASSIGN
##
-## Ends in an error in state: 508.
+## Ends in an error in state: 524.
##
-## selection_statement -> save_context SWITCH . LPAREN expression RPAREN statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context SWITCH
@@ -3829,9 +3829,9 @@ At this point, an opening parenthesis '(' is expected.
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
##
-## Ends in an error in state: 495.
+## Ends in an error in state: 511.
##
-## iteration_statement -> save_context WHILE LPAREN expression RPAREN . statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context WHILE LPAREN expression RPAREN
@@ -3844,10 +3844,10 @@ At this point, a statement (the loop body) is expected.
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: 494.
+## Ends in an error in state: 510.
##
## expression -> expression . COMMA assignment_expression [ RPAREN COMMA ]
-## iteration_statement -> save_context WHILE LPAREN expression . RPAREN statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context WHILE LPAREN expression
@@ -3856,21 +3856,21 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 149, spurious reduction of production expression -> assignment_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 140, spurious reduction of production expression -> assignment_expression
##
Ill-formed 'while' statement.
@@ -3883,9 +3883,9 @@ then at this point, a closing parenthesis ')' is expected.
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: 493.
+## Ends in an error in state: 509.
##
-## iteration_statement -> save_context WHILE LPAREN . expression RPAREN statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context WHILE LPAREN
@@ -3898,9 +3898,9 @@ At this point, an expression is expected.
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: 492.
+## Ends in an error in state: 508.
##
-## iteration_statement -> save_context WHILE . LPAREN expression RPAREN statement [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context WHILE
@@ -3913,10 +3913,10 @@ At this point, an opening parenthesis '(' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME RPAREN LBRACE XOR_ASSIGN
##
-## Ends in an error in state: 436.
+## Ends in an error in state: 452.
##
-## block_item_list -> option(block_item_list) . block_item [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN EOF ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## save_context LBRACE option(block_item_list)
@@ -3941,9 +3941,9 @@ At this point, one of the following is expected:
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: 437.
+## Ends in an error in state: 453.
##
-## jump_statement -> RETURN . option(expression) SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN 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 BUILTIN_OFFSETOF 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_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN ENUM ELSE DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## RETURN
@@ -3960,7 +3960,7 @@ At this point, one of the following is expected:
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: 440.
+## Ends in an error in state: 456.
##
## expression -> expression . COMMA assignment_expression [ SEMICOLON COMMA ]
## option(expression) -> expression . [ SEMICOLON ]
@@ -3972,23 +3972,23 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME 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 80, spurious reduction of production primary_expression -> string_literals_list
-## In state 82, spurious reduction of production postfix_expression -> primary_expression
-## In state 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 149, spurious reduction of production expression -> assignment_expression
+## In state 70, spurious reduction of production primary_expression -> string_literals_list
+## In state 72, spurious reduction of production postfix_expression -> primary_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 140, spurious reduction of production expression -> assignment_expression
##
Up to this point, an expression has been recognized:
@@ -4000,7 +4000,7 @@ then at this point, a semicolon ';' is expected.
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: 578.
+## Ends in an error in state: 594.
##
## 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 ]
@@ -4035,7 +4035,7 @@ at this point, one of the following is expected:
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: 144.
+## Ends in an error in state: 135.
##
## expression -> expression COMMA . assignment_expression [ SEMICOLON RPAREN RBRACK COMMA COLON ]
##
@@ -4050,7 +4050,7 @@ At this point, an expression is expected.
translation_unit_file: INT PRE_NAME VAR_NAME COMMA PRE_NAME VAR_NAME RPAREN
##
-## Ends in an error in state: 555.
+## Ends in an error in state: 571.
##
## init_declarator_list -> init_declarator_list . COMMA init_declarator [ SEMICOLON COMMA ]
## option(init_declarator_list) -> init_declarator_list . [ SEMICOLON ]
@@ -4062,12 +4062,12 @@ translation_unit_file: INT PRE_NAME VAR_NAME COMMA PRE_NAME 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 263, spurious reduction of production declarator_noattrend -> direct_declarator
-## In state 563, spurious reduction of production declare_varname(declarator_noattrend) -> declarator_noattrend
-## In state 558, spurious reduction of production save_context ->
-## In state 559, spurious reduction of production attribute_specifier_list ->
-## In state 560, spurious reduction of production init_declarator -> declare_varname(declarator_noattrend) save_context attribute_specifier_list
-## In state 557, spurious reduction of production init_declarator_list -> init_declarator_list COMMA init_declarator
+## In state 266, spurious reduction of production declarator_noattrend -> direct_declarator
+## In state 579, spurious reduction of production declare_varname(declarator_noattrend) -> declarator_noattrend
+## In state 574, spurious reduction of production save_context ->
+## In state 575, spurious reduction of production attribute_specifier_list ->
+## In state 576, spurious reduction of production init_declarator -> declare_varname(declarator_noattrend) save_context attribute_specifier_list
+## In state 573, spurious reduction of production init_declarator_list -> init_declarator_list COMMA init_declarator
##
Up to this point, a list of declarators has been recognized:
@@ -4079,7 +4079,7 @@ then at this point, a semicolon ';' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 556.
+## Ends in an error in state: 572.
##
## init_declarator_list -> init_declarator_list COMMA . init_declarator [ SEMICOLON COMMA ]
##
@@ -4094,7 +4094,7 @@ At this point, an init declarator is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE DOT PRE_NAME VAR_NAME EQ ALIGNAS
##
-## Ends in an error in state: 374.
+## Ends in an error in state: 390.
##
## initializer_list -> option(designation) . c_initializer [ RBRACE COMMA ]
##
@@ -4103,7 +4103,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE DOT PRE_NAME VAR_NAME EQ
##
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: 378.
+## Ends in an error in state: 394.
##
## initializer_list -> initializer_list COMMA option(designation) . c_initializer [ RBRACE COMMA ]
##
@@ -4118,9 +4118,9 @@ At this point, an initializer is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE DOT PRE_NAME VAR_NAME XOR_ASSIGN
##
-## Ends in an error in state: 381.
+## Ends in an error in state: 397.
##
-## designation -> designator_list . EQ [ TILDE STRING_LITERAL STAR SIZEOF PRE_NAME PLUS MINUS LPAREN LBRACE INC DEC CONSTANT BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AND ALIGNOF ]
+## designation -> designator_list . EQ [ TILDE STRING_LITERAL STAR SIZEOF PRE_NAME PLUS MINUS LPAREN LBRACE INC GENERIC DEC CONSTANT BUILTIN_VA_ARG BUILTIN_OFFSETOF BANG AND ALIGNOF ]
## option(designator_list) -> designator_list . [ LBRACK DOT ]
##
## The known suffix of the stack is as follows:
@@ -4140,7 +4140,7 @@ then at this point, an equals sign '=' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE DOT XOR_ASSIGN
##
-## Ends in an error in state: 334.
+## Ends in an error in state: 364.
##
## designator -> DOT . general_identifier [ RPAREN LBRACK EQ DOT ]
##
@@ -4157,7 +4157,7 @@ At this point, the name of a struct or union member is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE LBRACK PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 332.
+## Ends in an error in state: 362.
##
## designator -> LBRACK conditional_expression . RBRACK [ RPAREN LBRACK EQ DOT ]
##
@@ -4168,19 +4168,19 @@ translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE LBRACK PRE_NAME VAR_NAME
## 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 79, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 69, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
##
Ill-formed designator.
@@ -4193,7 +4193,7 @@ then at this point, a closing bracket ']' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE LBRACK XOR_ASSIGN
##
-## Ends in an error in state: 331.
+## Ends in an error in state: 361.
##
## designator -> LBRACK . conditional_expression RBRACK [ RPAREN LBRACK EQ DOT ]
##
@@ -4208,7 +4208,7 @@ At this point, a constant expression is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 377.
+## Ends in an error in state: 393.
##
## initializer_list -> initializer_list COMMA . option(designation) c_initializer [ RBRACE COMMA ]
## option(COMMA) -> COMMA . [ RBRACE ]
@@ -4229,7 +4229,7 @@ At this point, one of the following is expected:
translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE CONSTANT SEMICOLON
##
-## Ends in an error in state: 376.
+## Ends in an error in state: 392.
##
## c_initializer -> LBRACE initializer_list . option(COMMA) RBRACE [ SEMICOLON RBRACE COMMA ]
## initializer_list -> initializer_list . COMMA option(designation) c_initializer [ RBRACE COMMA ]
@@ -4241,22 +4241,22 @@ translation_unit_file: INT PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
-## In state 380, spurious reduction of production c_initializer -> assignment_expression
-## In state 386, spurious reduction of production initializer_list -> option(designation) c_initializer
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 396, spurious reduction of production c_initializer -> assignment_expression
+## In state 402, spurious reduction of production initializer_list -> option(designation) c_initializer
##
# Omitting the fact that the closing brace can be preceded with a comma.
@@ -4271,7 +4271,7 @@ then at this point, a closing brace '}' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ LBRACE XOR_ASSIGN
##
-## Ends in an error in state: 375.
+## Ends in an error in state: 391.
##
## c_initializer -> LBRACE . initializer_list option(COMMA) RBRACE [ SEMICOLON RBRACE COMMA ]
##
@@ -4292,7 +4292,7 @@ followed with an initializer, is expected.
translation_unit_file: INT PRE_NAME VAR_NAME EQ XOR_ASSIGN
##
-## Ends in an error in state: 561.
+## Ends in an error in state: 577.
##
## init_declarator -> declare_varname(declarator_noattrend) save_context attribute_specifier_list EQ . c_initializer [ SEMICOLON COMMA ]
##
@@ -4309,7 +4309,7 @@ At this point, an initializer is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LBRACK CONSTANT SEMICOLON
##
-## Ends in an error in state: 251.
+## Ends in an error in state: 254.
##
## optional(assignment_expression,RBRACK) -> assignment_expression . RBRACK [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT SEMICOLON RPAREN RESTRICT REGISTER PRE_NAME PACKED NORETURN LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
##
@@ -4320,20 +4320,20 @@ translation_unit_file: INT PRE_NAME 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 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 87, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
-## In state 145, spurious reduction of production assignment_expression -> conditional_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
##
# At the time of writing, optional(expression,RBRACK) is used only in direct
@@ -4351,7 +4351,7 @@ then at this point, a closing bracket ']' is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME COMMA XOR_ASSIGN
##
-## Ends in an error in state: 290.
+## Ends in an error in state: 293.
##
## identifier_list -> identifier_list COMMA . PRE_NAME VAR_NAME [ RPAREN COMMA ]
##
@@ -4368,7 +4368,7 @@ At this point, an identifier is expected.
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME COMMA PRE_NAME TYPEDEF_NAME
##
-## Ends in an error in state: 291.
+## Ends in an error in state: 294.
##
## identifier_list -> identifier_list COMMA PRE_NAME . VAR_NAME [ RPAREN COMMA ]
##
@@ -4384,7 +4384,7 @@ The following type name is used as a K&R parameter name:
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN INT XOR_ASSIGN
##
-## Ends in an error in state: 596.
+## Ends in an error in state: 612.
##
## 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 NORETURN LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
@@ -4394,7 +4394,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN INT
##
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN VOLATILE INT XOR_ASSIGN
##
-## Ends in an error in state: 601.
+## Ends in an error in state: 617.
##
## 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 NORETURN LPAREN LONG INT INLINE FLOAT EXTERN ENUM DOUBLE CONST CHAR AUTO ATTRIBUTE ALIGNAS ]
@@ -4404,7 +4404,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN VOL
##
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: 594.
+## Ends in an error in state: 610.
##
## 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 NORETURN LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
@@ -4417,7 +4417,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN PRE
##
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: 599.
+## Ends in an error in state: 615.
##
## 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 NORETURN LPAREN INLINE EXTERN CONST AUTO ATTRIBUTE ALIGNAS ]
@@ -4442,7 +4442,7 @@ At this point, one of the following is expected:
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN VOLATILE XOR_ASSIGN
##
-## Ends in an error in state: 597.
+## Ends in an error in state: 613.
##
## 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 ]
@@ -4454,7 +4454,7 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN VOL
## 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 222, spurious reduction of production rlist(declaration_specifier_no_type) -> type_qualifier_noattr
+## In state 225, spurious reduction of production rlist(declaration_specifier_no_type) -> type_qualifier_noattr
##
Ill-formed K&R parameter declaration.
@@ -4467,7 +4467,7 @@ At this point, one of the following is expected:
translation_unit_file: VOID PRE_NAME TYPEDEF_NAME PACKED LPAREN CONSTANT RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 610.
+## Ends in an error in state: 626.
##
## 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 ]
@@ -4500,7 +4500,7 @@ If this is the parameter declaration of a K&R function definition,
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT COMMA XOR_ASSIGN
##
-## Ends in an error in state: 238.
+## Ends in an error in state: 241.
##
## parameter_list -> parameter_list COMMA . parameter_declaration [ RPAREN COMMA ]
## parameter_type_list -> parameter_list COMMA . ELLIPSIS [ RPAREN ]
@@ -4517,7 +4517,7 @@ At this point, one of the following is expected:
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME SEMICOLON
##
-## Ends in an error in state: 237.
+## Ends in an error in state: 240.
##
## parameter_list -> parameter_list . COMMA parameter_declaration [ RPAREN COMMA ]
## parameter_type_list -> parameter_list . [ RPAREN ]
@@ -4530,12 +4530,12 @@ translation_unit_file: INT PRE_NAME VAR_NAME LPAREN INT PRE_NAME VAR_NAME SEMICO
## 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 263, spurious reduction of production declarator_noattrend -> direct_declarator
-## In state 268, spurious reduction of production attribute_specifier_list ->
-## In state 269, spurious reduction of production declarator -> declarator_noattrend attribute_specifier_list
-## In state 285, spurious reduction of production declare_varname(declarator) -> declarator
-## In state 284, spurious reduction of production parameter_declaration -> declaration_specifiers(parameter_declaration) declare_varname(declarator)
-## In state 245, spurious reduction of production parameter_list -> parameter_declaration
+## In state 266, spurious reduction of production declarator_noattrend -> direct_declarator
+## In state 271, spurious reduction of production attribute_specifier_list ->
+## In state 272, spurious reduction of production declarator -> declarator_noattrend attribute_specifier_list
+## In state 288, spurious reduction of production declare_varname(declarator) -> declarator
+## In state 287, spurious reduction of production parameter_declaration -> declaration_specifiers(parameter_declaration) declare_varname(declarator)
+## In state 248, spurious reduction of production parameter_list -> parameter_declaration
##
# We omit the possibility of an ellipsis.
@@ -4571,7 +4571,7 @@ The following identifier is used as a type, but has not been defined as such:
translation_unit_file: INT PRE_NAME VAR_NAME LPAREN PRE_NAME VAR_NAME RPAREN INT SEMICOLON XOR_ASSIGN
##
-## Ends in an error in state: 606.
+## Ends in an error in state: 622.
##
## declaration_list -> declaration_list . kr_param_declaration [ VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL STRUCT STATIC SIGNED SHORT RESTRICT REGISTER PRE_NAME PACKED NORETURN 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 ]
@@ -4590,7 +4590,7 @@ At this point, one of the following is expected:
translation_unit_file: PACKED LPAREN BUILTIN_OFFSETOF XOR_ASSIGN
##
-## Ends in an error in state: 52.
+## Ends in an error in state: 41.
##
## postfix_expression -> BUILTIN_OFFSETOF . LPAREN type_name COMMA general_identifier 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 ]
## postfix_expression -> BUILTIN_OFFSETOF . LPAREN type_name COMMA general_identifier designator_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 ]
@@ -4606,7 +4606,7 @@ At this point, an opening parenthesis '(' is expected.
translation_unit_file: PACKED LPAREN BUILTIN_OFFSETOF LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 53.
+## Ends in an error in state: 42.
##
## postfix_expression -> BUILTIN_OFFSETOF LPAREN . type_name COMMA general_identifier 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 ]
## postfix_expression -> BUILTIN_OFFSETOF LPAREN . type_name COMMA general_identifier designator_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 ]
@@ -4622,7 +4622,7 @@ At this point, a struct or union name is expected.
translation_unit_file: PACKED LPAREN BUILTIN_OFFSETOF LPAREN VOID XOR_ASSIGN
##
-## Ends in an error in state: 326.
+## Ends in an error in state: 356.
##
## postfix_expression -> BUILTIN_OFFSETOF LPAREN type_name . COMMA general_identifier 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 ]
## postfix_expression -> BUILTIN_OFFSETOF LPAREN type_name . COMMA general_identifier designator_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 ]
@@ -4634,9 +4634,9 @@ translation_unit_file: PACKED LPAREN BUILTIN_OFFSETOF 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 67, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
-## In state 314, spurious reduction of production option(abstract_declarator(type_name)) ->
-## In state 320, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+## In state 159, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 317, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 323, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
##
Ill-formed __builtin_offsetof.
@@ -4646,7 +4646,7 @@ At this point, a colon ',' is expected
translation_unit_file: PACKED LPAREN BUILTIN_OFFSETOF LPAREN VOID COMMA XOR_ASSIGN
##
-## Ends in an error in state: 327.
+## Ends in an error in state: 357.
##
## postfix_expression -> BUILTIN_OFFSETOF LPAREN type_name COMMA . general_identifier 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 ]
## postfix_expression -> BUILTIN_OFFSETOF LPAREN type_name COMMA . general_identifier designator_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 ]
@@ -4662,7 +4662,7 @@ At this point, a member-designator is expected.
translation_unit_file: PACKED LPAREN BUILTIN_OFFSETOF LPAREN VOID COMMA PRE_NAME TYPEDEF_NAME XOR_ASSIGN
##
-## Ends in an error in state: 328.
+## Ends in an error in state: 358.
##
## postfix_expression -> BUILTIN_OFFSETOF LPAREN type_name COMMA general_identifier . 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 ]
## postfix_expression -> BUILTIN_OFFSETOF LPAREN type_name COMMA general_identifier . designator_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 ]
@@ -4678,7 +4678,7 @@ At this point, a member-designator is expected.
translation_unit_file: PACKED LPAREN BUILTIN_OFFSETOF LPAREN VOID COMMA PRE_NAME TYPEDEF_NAME LBRACK STRING_LITERAL RBRACK XOR_ASSIGN
##
-## Ends in an error in state: 337.
+## Ends in an error in state: 367.
##
## option(designator_list) -> designator_list . [ LBRACK DOT ]
## postfix_expression -> BUILTIN_OFFSETOF LPAREN type_name COMMA general_identifier designator_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 ]
@@ -4694,9 +4694,9 @@ At this point, a member-designator is expected.
translation_unit_file: STATIC_ASSERT XOR_ASSIGN
##
-## Ends in an error in state: 76.
+## Ends in an error in state: 168.
##
-## static_assert_declaration -> STATIC_ASSERT . LPAREN conditional_expression COMMA string_literals_list RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN EOF ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## static_assert_declaration -> STATIC_ASSERT . LPAREN conditional_expression COMMA string_literals_list RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN EOF ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## STATIC_ASSERT
@@ -4709,9 +4709,9 @@ At this point, an opening parenthesis '(' is expected.
translation_unit_file: STATIC_ASSERT LPAREN XOR_ASSIGN
##
-## Ends in an error in state: 77.
+## Ends in an error in state: 169.
##
-## static_assert_declaration -> STATIC_ASSERT LPAREN . conditional_expression COMMA string_literals_list RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN EOF ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## static_assert_declaration -> STATIC_ASSERT LPAREN . conditional_expression COMMA string_literals_list RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN EOF ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## STATIC_ASSERT LPAREN
@@ -4724,9 +4724,9 @@ At this point, a constant expression is expected.
translation_unit_file: STATIC_ASSERT LPAREN STRING_LITERAL XOR_ASSIGN
##
-## Ends in an error in state: 167.
+## Ends in an error in state: 170.
##
-## static_assert_declaration -> STATIC_ASSERT LPAREN conditional_expression . COMMA string_literals_list RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN EOF ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## static_assert_declaration -> STATIC_ASSERT LPAREN conditional_expression . COMMA string_literals_list RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN EOF ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## STATIC_ASSERT LPAREN conditional_expression
@@ -4735,21 +4735,21 @@ translation_unit_file: STATIC_ASSERT LPAREN STRING_LITERAL 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 80, spurious reduction of production primary_expression -> string_literals_list
-## In state 82, spurious reduction of production postfix_expression -> primary_expression
-## In state 83, spurious reduction of production unary_expression -> postfix_expression
-## In state 79, spurious reduction of production cast_expression -> unary_expression
-## In state 110, spurious reduction of production multiplicative_expression -> cast_expression
-## In state 104, spurious reduction of production additive_expression -> multiplicative_expression
-## In state 123, spurious reduction of production shift_expression -> additive_expression
-## In state 100, spurious reduction of production relational_expression -> shift_expression
-## In state 116, spurious reduction of production equality_expression -> relational_expression
-## In state 132, spurious reduction of production and_expression -> equality_expression
-## In state 140, spurious reduction of production exclusive_or_expression -> and_expression
-## In state 141, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
-## In state 142, spurious reduction of production logical_and_expression -> inclusive_or_expression
-## In state 126, spurious reduction of production logical_or_expression -> logical_and_expression
-## In state 124, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 70, spurious reduction of production primary_expression -> string_literals_list
+## In state 72, spurious reduction of production postfix_expression -> primary_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 69, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
##
Ill-formed _Static_assert.
@@ -4759,9 +4759,9 @@ At this point, a comma ',' is expected.
translation_unit_file: STATIC_ASSERT LPAREN STRING_LITERAL COMMA XOR_ASSIGN
##
-## Ends in an error in state: 168.
+## Ends in an error in state: 171.
##
-## static_assert_declaration -> STATIC_ASSERT LPAREN conditional_expression COMMA . string_literals_list RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN EOF ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## static_assert_declaration -> STATIC_ASSERT LPAREN conditional_expression COMMA . string_literals_list RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN EOF ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## STATIC_ASSERT LPAREN conditional_expression COMMA
@@ -4774,9 +4774,9 @@ At this point, a string literal is expected.
translation_unit_file: STATIC_ASSERT LPAREN STRING_LITERAL COMMA STRING_LITERAL XOR_ASSIGN
##
-## Ends in an error in state: 169.
+## Ends in an error in state: 172.
##
-## static_assert_declaration -> STATIC_ASSERT LPAREN conditional_expression COMMA string_literals_list . RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN EOF ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## static_assert_declaration -> STATIC_ASSERT LPAREN conditional_expression COMMA string_literals_list . RPAREN SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN EOF ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
## string_literals_list -> string_literals_list . STRING_LITERAL [ STRING_LITERAL RPAREN ]
##
## The known suffix of the stack is as follows:
@@ -4790,9 +4790,9 @@ At this point, a closing parenthesis ')' is expected.
translation_unit_file: STATIC_ASSERT LPAREN STRING_LITERAL COMMA STRING_LITERAL RPAREN XOR_ASSIGN
##
-## Ends in an error in state: 170.
+## Ends in an error in state: 173.
##
-## static_assert_declaration -> STATIC_ASSERT LPAREN conditional_expression COMMA string_literals_list RPAREN . SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO FOR FLOAT EXTERN EOF ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
+## static_assert_declaration -> STATIC_ASSERT LPAREN conditional_expression COMMA string_literals_list RPAREN . SEMICOLON [ WHILE VOLATILE VOID UNSIGNED UNION UNDERSCORE_BOOL TYPEDEF TILDE SWITCH STRUCT STRING_LITERAL STATIC_ASSERT STATIC STAR SIZEOF SIGNED SHORT SEMICOLON RETURN RESTRICT REGISTER RBRACE PRE_NAME PRAGMA PLUS PACKED NORETURN MINUS LPAREN LONG LBRACE INT INLINE INC IF GOTO GENERIC FOR FLOAT EXTERN EOF ENUM DOUBLE DO DEFAULT DEC CONTINUE CONSTANT CONST CHAR CASE BUILTIN_VA_ARG BUILTIN_OFFSETOF BREAK BANG AUTO ATTRIBUTE ASM AND ALIGNOF ALIGNAS ]
##
## The known suffix of the stack is as follows:
## STATIC_ASSERT LPAREN conditional_expression COMMA string_literals_list RPAREN
@@ -4803,6 +4803,224 @@ At this point, a semicolon ';' is expected.
#------------------------------------------------------------------------------
+translation_unit_file: ALIGNAS LPAREN GENERIC XOR_ASSIGN
+##
+## Ends in an error in state: 35.
+##
+## generic_selection -> GENERIC . LPAREN assignment_expression COMMA generic_assoc_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 ]
+##
+## The known suffix of the stack is as follows:
+## GENERIC
+##
+
+Ill-formed _Generic expression.
+At this point, an opening parenthesis '(' is expected.
+
+#------------------------------------------------------------------------------
+
+translation_unit_file: ALIGNAS LPAREN GENERIC LPAREN XOR_ASSIGN
+##
+## Ends in an error in state: 36.
+##
+## generic_selection -> GENERIC LPAREN . assignment_expression COMMA generic_assoc_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 ]
+##
+## The known suffix of the stack is as follows:
+## GENERIC LPAREN
+##
+
+Ill-formed _Generic expression.
+At this point, an expression is expected.
+
+#------------------------------------------------------------------------------
+
+translation_unit_file: ALIGNAS LPAREN GENERIC LPAREN STRING_LITERAL WHILE
+##
+## Ends in an error in state: 374.
+##
+## generic_selection -> GENERIC LPAREN assignment_expression . COMMA generic_assoc_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 ]
+##
+## The known suffix of the stack is as follows:
+## GENERIC LPAREN assignment_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 70, spurious reduction of production primary_expression -> string_literals_list
+## In state 72, spurious reduction of production postfix_expression -> primary_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+##
+
+Ill-formed _Generic expression.
+Up to this point, an expression has been recognized:
+ $0
+If this expression is complete,
+then at this point, a comma ',' is expected.
+
+#------------------------------------------------------------------------------
+
+translation_unit_file: ALIGNAS LPAREN GENERIC LPAREN STRING_LITERAL COMMA XOR_ASSIGN
+##
+## Ends in an error in state: 375.
+##
+## generic_selection -> GENERIC LPAREN assignment_expression COMMA . generic_assoc_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 ]
+##
+## The known suffix of the stack is as follows:
+## GENERIC LPAREN assignment_expression COMMA
+##
+
+Ill-formed _Generic expression.
+At this point, either a type name or 'default' are expected.
+
+#------------------------------------------------------------------------------
+
+translation_unit_file: ALIGNAS LPAREN GENERIC LPAREN STRING_LITERAL COMMA DEFAULT XOR_ASSIGN
+##
+## Ends in an error in state: 376.
+##
+## generic_association -> DEFAULT . COLON assignment_expression [ RPAREN COMMA ]
+##
+## The known suffix of the stack is as follows:
+## DEFAULT
+##
+
+Ill-formed _Generic expression.
+At this point, a colon ':' is expected.
+
+#------------------------------------------------------------------------------
+
+translation_unit_file: ALIGNAS LPAREN GENERIC LPAREN STRING_LITERAL COMMA DEFAULT COLON XOR_ASSIGN
+##
+## Ends in an error in state: 377.
+##
+## generic_association -> DEFAULT COLON . assignment_expression [ RPAREN COMMA ]
+##
+## The known suffix of the stack is as follows:
+## DEFAULT COLON
+##
+
+Ill-formed _Generic expression.
+At this point, an expression is expected.
+
+#------------------------------------------------------------------------------
+
+translation_unit_file: ALIGNAS LPAREN GENERIC LPAREN STRING_LITERAL COMMA CHAR XOR_ASSIGN
+##
+## Ends in an error in state: 379.
+##
+## generic_association -> type_name . COLON assignment_expression [ RPAREN COMMA ]
+##
+## The known suffix of the stack is as follows:
+## 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 159, spurious reduction of production specifier_qualifier_list(type_name) -> type_specifier_no_typedef_name list(specifier_qualifier_no_typedef_name)
+## In state 317, spurious reduction of production option(abstract_declarator(type_name)) ->
+## In state 323, spurious reduction of production type_name -> specifier_qualifier_list(type_name) option(abstract_declarator(type_name))
+##
+
+Ill-formed _Generic expression.
+Up to this point, a type name has been recognized:
+ $0
+At this point, a colon ':' is expected.
+
+#------------------------------------------------------------------------------
+
+translation_unit_file: ALIGNAS LPAREN GENERIC LPAREN STRING_LITERAL COMMA CHAR COLON XOR_ASSIGN
+##
+## Ends in an error in state: 380.
+##
+## generic_association -> type_name COLON . assignment_expression [ RPAREN COMMA ]
+##
+## The known suffix of the stack is as follows:
+## type_name COLON
+##
+
+Ill-formed _Generic expression.
+Up to this point, a type name followed by a colon have been recognized:
+ $0
+At this point, an expression is expected.
+
+#------------------------------------------------------------------------------
+
+translation_unit_file: ALIGNAS LPAREN GENERIC LPAREN STRING_LITERAL COMMA CHAR COLON STRING_LITERAL WHILE
+##
+## Ends in an error in state: 383.
+##
+## generic_assoc_list -> generic_assoc_list . COMMA generic_association [ RPAREN COMMA ]
+## generic_selection -> GENERIC LPAREN assignment_expression COMMA generic_assoc_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 ]
+##
+## The known suffix of the stack is as follows:
+## GENERIC LPAREN assignment_expression COMMA generic_assoc_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 70, spurious reduction of production primary_expression -> string_literals_list
+## In state 72, spurious reduction of production postfix_expression -> primary_expression
+## In state 73, spurious reduction of production unary_expression -> postfix_expression
+## In state 77, spurious reduction of production cast_expression -> unary_expression
+## In state 101, spurious reduction of production multiplicative_expression -> cast_expression
+## In state 94, spurious reduction of production additive_expression -> multiplicative_expression
+## In state 114, spurious reduction of production shift_expression -> additive_expression
+## In state 90, spurious reduction of production relational_expression -> shift_expression
+## In state 107, spurious reduction of production equality_expression -> relational_expression
+## In state 123, spurious reduction of production and_expression -> equality_expression
+## In state 131, spurious reduction of production exclusive_or_expression -> and_expression
+## In state 132, spurious reduction of production inclusive_or_expression -> exclusive_or_expression
+## In state 133, spurious reduction of production logical_and_expression -> inclusive_or_expression
+## In state 117, spurious reduction of production logical_or_expression -> logical_and_expression
+## In state 115, spurious reduction of production conditional_expression -> logical_or_expression
+## In state 136, spurious reduction of production assignment_expression -> conditional_expression
+## In state 381, spurious reduction of production generic_association -> type_name COLON assignment_expression
+## In state 382, spurious reduction of production generic_assoc_list -> generic_association
+##
+
+Ill-formed _Generic expression.
+Up to this point, one or several generic associations have been recognized:
+ $0
+At this point, either a closing parenthesis ')' or a comma ',' are expected.
+Use a closing parenthesis ')' to terminate the _Generic expression.
+Use a comma to add another generic association.
+
+#------------------------------------------------------------------------------
+
+translation_unit_file: ALIGNAS LPAREN GENERIC LPAREN STRING_LITERAL COMMA CHAR COLON STRING_LITERAL COMMA XOR_ASSIGN
+##
+## Ends in an error in state: 385.
+##
+## generic_assoc_list -> generic_assoc_list COMMA . generic_association [ RPAREN COMMA ]
+##
+## The known suffix of the stack is as follows:
+## generic_assoc_list COMMA
+##
+
+Ill-formed _Generic expression.
+Up to this point, one or several generic associations have been recognized,
+followed by a comma ',':
+ $0
+At this point, another generic association is expected.
+It should start either with a type name or with the 'default' keyword.
+
+#------------------------------------------------------------------------------
+
translation_unit_file: ALIGNAS LPAREN PRE_NAME XOR_ASSIGN
##
## Ends in an error in state: 29.
@@ -4815,7 +5033,7 @@ translation_unit_file: ALIGNAS LPAREN PRE_NAME XOR_ASSIGN
##
translation_unit_file: ALIGNAS LPAREN VOID LPAREN VOID LPAREN PRE_NAME XOR_ASSIGN
##
-## Ends in an error in state: 244.
+## Ends in an error in state: 247.
##
## 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 ]
@@ -4826,7 +5044,7 @@ translation_unit_file: ALIGNAS LPAREN VOID LPAREN VOID LPAREN PRE_NAME XOR_ASSIG
##
translation_unit_file: UNION PRE_NAME XOR_ASSIGN
##
-## Ends in an error in state: 40.
+## Ends in an error in state: 48.
##
## 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 NORETURN 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 NORETURN 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 ]
@@ -4836,7 +5054,7 @@ translation_unit_file: UNION PRE_NAME XOR_ASSIGN
##
translation_unit_file: VOID PRE_NAME TYPEDEF_NAME LBRACE PRE_NAME XOR_ASSIGN
##
-## Ends in an error in state: 442.
+## Ends in an error in state: 458.
##
## 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 ]
@@ -4847,7 +5065,7 @@ translation_unit_file: VOID PRE_NAME TYPEDEF_NAME LBRACE PRE_NAME XOR_ASSIGN
##
translation_unit_file: VOID PRE_NAME TYPEDEF_NAME LPAREN PRE_NAME XOR_ASSIGN
##
-## Ends in an error in state: 202.
+## Ends in an error in state: 205.
##
## identifier_list -> PRE_NAME . VAR_NAME [ RPAREN COMMA ]
## typedef_name -> PRE_NAME . TYPEDEF_NAME [ VOLATILE STATIC STAR RPAREN RESTRICT REGISTER PRE_NAME PACKED NORETURN LPAREN LBRACK INLINE EXTERN CONST COMMA AUTO ATTRIBUTE ALIGNAS ]
@@ -4857,7 +5075,7 @@ translation_unit_file: VOID PRE_NAME TYPEDEF_NAME LPAREN PRE_NAME XOR_ASSIGN
##
translation_unit_file: VOID PRE_NAME XOR_ASSIGN
##
-## Ends in an error in state: 190.
+## Ends in an error in state: 193.
##
## 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 NORETURN 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 NORETURN LPAREN LONG LBRACK LBRACE INT INLINE FLOAT EXTERN EQ ENUM DOUBLE CONST COMMA COLON CHAR AUTO ATTRIBUTE ALIGNAS ]
diff --git a/cparser/pre_parser.mly b/cparser/pre_parser.mly
index cffbd192..ad294398 100644
--- a/cparser/pre_parser.mly
+++ b/cparser/pre_parser.mly
@@ -58,7 +58,7 @@
AUTO REGISTER INLINE NORETURN CHAR SHORT INT LONG SIGNED UNSIGNED FLOAT DOUBLE
UNDERSCORE_BOOL CONST VOLATILE VOID STRUCT UNION ENUM CASE DEFAULT IF ELSE
SWITCH WHILE DO FOR GOTO CONTINUE BREAK RETURN BUILTIN_VA_ARG ALIGNOF
- ATTRIBUTE ALIGNAS PACKED ASM BUILTIN_OFFSETOF STATIC_ASSERT
+ ATTRIBUTE ALIGNAS PACKED ASM BUILTIN_OFFSETOF STATIC_ASSERT GENERIC
%token EOF
@@ -248,6 +248,21 @@ primary_expression:
| CONSTANT
| string_literals_list
| LPAREN expression RPAREN
+| generic_selection
+ {}
+
+generic_selection:
+| GENERIC LPAREN assignment_expression COMMA generic_assoc_list RPAREN
+ {}
+
+generic_assoc_list:
+| generic_association
+| generic_assoc_list COMMA generic_association
+ {}
+
+generic_association:
+| type_name COLON assignment_expression
+| DEFAULT COLON assignment_expression
{}
postfix_expression: