summaryrefslogtreecommitdiff
path: root/cparser/Parser.vy
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2014-05-12 15:52:42 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2014-05-12 15:52:42 +0000
commitedc00e0c90a5598f653add89f42a095d8ee1b629 (patch)
tree2d2539335cc7e916a8964847b2ed7489f9340d00 /cparser/Parser.vy
parent951bf7bdb208f500c86e8d45c45247cd25adb4ab (diff)
Assorted fixes to fix parsing issues and be more GCC-like:
- Moved scanning of char constants and string literals entirely to Lexer - Parser: separate STRING_LITERAL from CONSTANT to be closer to ISO C99 grammar - pre_parser: adapted + "asm" takes string_literal, not CONSTANT - Revised errors "inline doesnt belong here" git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2492 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'cparser/Parser.vy')
-rw-r--r--cparser/Parser.vy8
1 files changed, 6 insertions, 2 deletions
diff --git a/cparser/Parser.vy b/cparser/Parser.vy
index f32fe8f..b59997d 100644
--- a/cparser/Parser.vy
+++ b/cparser/Parser.vy
@@ -22,6 +22,7 @@ Require Import List.
%token<string * cabsloc> VAR_NAME TYPEDEF_NAME OTHER_NAME
%token<string * cabsloc> PRAGMA
+%token<bool * list char_code * cabsloc> STRING_LITERAL
%token<constant * cabsloc> CONSTANT
%token<cabsloc> 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
@@ -108,6 +109,9 @@ primary_expression:
{ (VARIABLE (fst var), snd var) }
| cst = CONSTANT
{ (CONSTANT (fst cst), snd cst) }
+| str = STRING_LITERAL
+ { let '((wide, chars), loc) := str in
+ (CONSTANT (CONST_STRING wide chars), loc) }
| loc = LPAREN expr = expression RPAREN
{ (fst expr, loc)}
@@ -816,8 +820,8 @@ jump_statement:
(* Non-standard *)
asm_statement:
-| loc = ASM LPAREN template = CONSTANT RPAREN SEMICOLON
- { ASM (fst template) loc }
+| loc = ASM LPAREN template = STRING_LITERAL RPAREN SEMICOLON
+ { let '(wide, chars, _) := template in ASM wide chars loc }
(* 6.9 *)
translation_unit_file: