aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
-rw-r--r--cparser/Cabs.v4
-rw-r--r--cparser/Cabshelper.ml10
-rw-r--r--cparser/Elab.ml196
-rw-r--r--cparser/Parser.vy146
-rw-r--r--cparser/pre_parser.mly178
-rw-r--r--test/regression/Results/parsing10
-rw-r--r--test/regression/parsing.c62
7 files changed, 345 insertions, 261 deletions
diff --git a/cparser/Cabs.v b/cparser/Cabs.v
index 6d9e95d5..ab53a3a8 100644
--- a/cparser/Cabs.v
+++ b/cparser/Cabs.v
@@ -81,6 +81,7 @@ with decl_type :=
| PTR : list cvspec -> decl_type -> decl_type
(* The bool is true for variable length parameters. *)
| PROTO : decl_type -> list parameter * bool -> decl_type
+ | PROTO_OLD : decl_type -> list string -> decl_type
with parameter :=
| PARAM : list spec_elem -> option string -> decl_type -> list attribute -> cabsloc -> parameter
@@ -190,8 +191,7 @@ Definition asm_flag := (bool * list char_code)%type.
** Declaration definition (at toplevel)
*)
Inductive definition :=
- | FUNDEF : list spec_elem -> name -> statement -> cabsloc -> definition
- | KRFUNDEF : list spec_elem -> name -> list string -> list definition -> statement -> cabsloc -> definition
+ | FUNDEF : list spec_elem -> name -> list definition -> statement -> cabsloc -> definition
| DECDEF : init_name_group -> cabsloc -> definition (* global variable(s), or function prototype *)
| PRAGMA : string -> cabsloc -> definition
diff --git a/cparser/Cabshelper.ml b/cparser/Cabshelper.ml
index 5e6a19d0..b3782ba8 100644
--- a/cparser/Cabshelper.ml
+++ b/cparser/Cabshelper.ml
@@ -46,8 +46,7 @@ let rec isTypedef = function
let get_definitionloc (d : definition) : cabsloc =
match d with
- | FUNDEF(_, _, _, l) -> l
- | KRFUNDEF(_, _, _, _, _, l) -> l
+ | FUNDEF(_, _, _, _, l) -> l
| DECDEF(_, l) -> l
| PRAGMA(_, l) -> l
@@ -78,10 +77,3 @@ let string_of_cabsloc l =
let format_cabsloc pp l =
Format.fprintf pp "%s:%d" l.filename l.lineno
-
-let rec append_decltype dt1 dt2 =
- match dt1 with
- | JUSTBASE -> dt2
- | ARRAY(dt, attr, sz) -> ARRAY(append_decltype dt dt2, attr, sz)
- | PTR(attr, dt) -> PTR(attr, append_decltype dt dt2)
- | PROTO(dt, params) -> PROTO(append_decltype dt dt2, params)
diff --git a/cparser/Elab.ml b/cparser/Elab.ml
index 4d3d1d02..d078cdac 100644
--- a/cparser/Elab.ml
+++ b/cparser/Elab.ml
@@ -518,9 +518,9 @@ and elab_cvspecs env cv_specs =
(* Elaboration of a type declarator. C99 section 6.7.5. *)
-and elab_type_declarator loc env ty = function
+and elab_type_declarator loc env ty kr_ok = function
| Cabs.JUSTBASE ->
- (ty, env)
+ ((ty, None), env)
| Cabs.ARRAY(d, cv_specs, sz) ->
let a = elab_cvspecs env cv_specs in
let sz' =
@@ -536,32 +536,41 @@ and elab_type_declarator loc env ty = function
| None ->
error loc "array size is not a compile-time constant";
Some 1L in (* produces better error messages later *)
- elab_type_declarator loc env (TArray(ty, sz', a)) d
+ elab_type_declarator loc env (TArray(ty, sz', a)) kr_ok d
| Cabs.PTR(cv_specs, d) ->
let a = elab_cvspecs env cv_specs in
- elab_type_declarator loc env (TPtr(ty, a)) d
+ elab_type_declarator loc env (TPtr(ty, a)) kr_ok d
| Cabs.PROTO(d, (params, vararg)) ->
- begin match unroll env ty with
- | TArray _ | TFun _ ->
- error loc "illegal function return type@ %a" Cprint.typ ty
- | _ -> ()
- end;
- let params' = elab_parameters env params in
- elab_type_declarator loc env (TFun(ty, params', vararg, [])) d
+ begin match unroll env ty with
+ | TArray _ | TFun _ ->
+ error loc "Illegal function return type@ %a" Cprint.typ ty
+ | _ -> ()
+ end;
+ let params' = elab_parameters env params in
+ elab_type_declarator loc env (TFun(ty, Some params', vararg, [])) kr_ok d
+ | Cabs.PROTO_OLD(d, params) ->
+ begin match unroll env ty with
+ | TArray _ | TFun _ ->
+ error loc "Illegal function return type@ %a" Cprint.typ ty
+ | _ -> ()
+ end;
+ match params with
+ | [] ->
+ elab_type_declarator loc env (TFun(ty, None, false, [])) kr_ok d
+ | _ ->
+ if not kr_ok || d <> Cabs.JUSTBASE then
+ fatal_error loc "Illegal old-style K&R function definition";
+ ((TFun(ty, None, false, []), Some params), env)
(* Elaboration of parameters in a prototype *)
and elab_parameters env params =
- match params with
- | [] -> (* old-style K&R prototype *)
- None
- | _ ->
- (* Prototype introduces a new scope *)
- let (vars, _) = mmap elab_parameter (Env.new_scope env) params in
- (* Catch special case f(t) where t is void or a typedef to void *)
- match vars with
- | [ ( {name=""}, t) ] when is_void_type env t -> Some []
- | _ -> Some vars
+ (* Prototype introduces a new scope *)
+ let (vars, _) = mmap elab_parameter (Env.new_scope env) params in
+ (* Catch special case f(t) where t is void or a typedef to void *)
+ match vars with
+ | [ ( {name=""}, t) ] when is_void_type env t -> []
+ | _ -> vars
(* Elaboration of a function parameter *)
@@ -569,7 +578,7 @@ and elab_parameter env (PARAM (spec, id, decl, attr, loc)) =
let (sto, inl, tydef, bty, env1) = elab_specifier loc env spec in
if tydef then
error loc "'typedef' used in function parameter";
- let (ty, env2) = elab_type_declarator loc env1 bty decl in
+ let ((ty, _), env2) = elab_type_declarator loc env1 bty false decl in
let ty = add_attributes_type (elab_attributes env attr) ty in
if sto <> Storage_default && sto <> Storage_register then
error loc
@@ -586,13 +595,13 @@ and elab_parameter env (PARAM (spec, id, decl, attr, loc)) =
(* Elaboration of a (specifier, Cabs "name") pair *)
-and elab_name env spec (Name (id, decl, attr, loc)) =
+and elab_fundef_name env spec (Name (id, decl, attr, loc)) =
let (sto, inl, tydef, bty, env') = elab_specifier loc env spec in
if tydef then
error loc "'typedef' is forbidden here";
- let (ty, env'') = elab_type_declarator loc env' bty decl in
+ let ((ty, kr_params), env'') = elab_type_declarator loc env' bty true decl in
let a = elab_attributes env attr in
- (id, sto, inl, add_attributes_type a ty, env'')
+ (id, sto, inl, add_attributes_type a ty, kr_params, env'')
(* Elaboration of a name group. C99 section 6.7.6 *)
@@ -604,8 +613,8 @@ and elab_name_group loc env (spec, namelist) =
if inl then
error loc "'inline' is forbidden here";
let elab_one_name env (Name (id, decl, attr, loc)) =
- let (ty, env1) =
- elab_type_declarator loc env bty decl in
+ let ((ty, _), env1) =
+ elab_type_declarator loc env bty false decl in
let a = elab_attributes env attr in
((id, add_attributes_type a ty), env1) in
(mmap elab_one_name env' namelist, sto)
@@ -616,8 +625,8 @@ and elab_init_name_group loc env (spec, namelist) =
let (sto, inl, tydef, bty, env') =
elab_specifier ~only:(namelist=[]) loc env spec in
let elab_one_name env (Init_name (Name (id, decl, attr, loc), init)) =
- let (ty, env1) =
- elab_type_declarator loc env bty decl in
+ let ((ty, _), env1) =
+ elab_type_declarator loc env bty false decl in
let a = elab_attributes env attr in
if inl && not (is_function_type env ty) then
error loc "'inline' can only appear on functions";
@@ -818,7 +827,7 @@ and elab_enum only loc tag optmembers attrs env =
let elab_type loc env spec decl =
let (sto, inl, tydef, bty, env') = elab_specifier loc env spec in
- let (ty, env'') = elab_type_declarator loc env' bty decl in
+ let ((ty, _), env'') = elab_type_declarator loc env' bty false decl in
if sto <> Storage_default || inl || tydef then
error loc "'typedef', 'extern', 'static', 'register' and 'inline' are meaningless in cast";
(ty, env'')
@@ -1737,8 +1746,8 @@ let elab_expr loc env a =
match args, params with
| [], [] -> []
| [], _::_ -> err "not enough arguments in function call"; []
- | _::_, [] ->
- if vararg
+ | _::_, [] ->
+ if vararg
then args
else (err "too many arguments in function call"; args)
| arg1 :: argl, (_, ty_p) :: paraml ->
@@ -1881,20 +1890,70 @@ let enter_decdefs local loc env sto dl =
let (decls, env') = List.fold_left enter_decdef ([], env) dl in
(List.rev decls, env')
-let elab_fundef env spec name body loc =
- let (s, sto, inline, ty, env1) = elab_name env spec name in
+let elab_fundef env spec name defs body loc =
+ let (s, sto, inline, ty, kr_params, env1) = elab_fundef_name env spec name in
if sto = Storage_register then
- fatal_error loc "a function definition cannot have 'register' storage class";
- (* Fix up the type. We can have params = None but only for an
- old-style parameterless function "int f() {...}" *)
- let ty =
- match ty with
- | TFun(ty_ret, None, vararg, attr) -> TFun(ty_ret, Some [], vararg, attr)
- | _ -> ty in
+ fatal_error loc "A function definition cannot have 'register' storage class";
+ begin match kr_params, defs with
+ | None, d::_ ->
+ error (get_definitionloc d)
+ "Old-style parameter declaration in a new-style function definition"
+ | _ -> ()
+ end;
+ let (ty, env1) =
+ match ty, kr_params with
+ | TFun(ty_ret, None, vararg, attr), None ->
+ (TFun(ty_ret, Some [], vararg, attr), env1)
+ | ty, None ->
+ (ty, env1)
+ | TFun(ty_ret, None, false, attr), Some params ->
+ warning loc "Non-prototype, pre-standard function definition.@ Converting to prototype form";
+ (* Check that the parameters have unique names *)
+ List.iter (fun id ->
+ if List.length (List.filter ((=) id) params) > 1 then
+ fatal_error loc "Parameter '%s' appears more than once in function declaration" id)
+ params;
+ (* Check that the declarations only declare parameters *)
+ let extract_name (Init_name(Name(s, dty, attrs, loc') as name, ie)) =
+ if not (List.mem s params) then
+ error loc' "Declaration of '%s' which is not a function parameter" s;
+ if ie <> NO_INIT then
+ error loc' "Illegal initialization of function parameter '%s'" s;
+ name
+ in
+ (* Convert old-style K&R function definition to modern prototyped form *)
+ let elab_kr_param_def env = function
+ | DECDEF((spec', name_init_list), loc') ->
+ let name_list = List.map extract_name name_init_list in
+ let (paramsenv, sto) = elab_name_group loc' env (spec', name_list) in
+ if sto <> Storage_default && sto <> Storage_register then
+ error loc' "'extern' or 'static' storage not supported for function parameter";
+ paramsenv
+ | d ->
+ (* Should never be produced by the parser *)
+ fatal_error (get_definitionloc d)
+ "Illegal declaration of function parameter" in
+ let (kr_params_defs, env1) = mmap elab_kr_param_def env1 defs in
+ let kr_params_defs = List.concat kr_params_defs in
+ let rec search_param_type param =
+ match List.filter (fun (p, _) -> p = param) kr_params_defs with
+ | [] ->
+ (* Parameter is not declared, defaults to "int" in ISO C90,
+ is an error in ISO C99. Just emit a warning. *)
+ warning loc "Type of '%s' defaults to 'int'" param;
+ (Env.fresh_ident param, TInt (IInt, []))
+ | (_,ty)::q ->
+ if q <> [] then error loc "Parameter '%s' defined more than once" param;
+ (Env.fresh_ident param, argument_conversion env1 ty)
+ in
+ let params' = List.map search_param_type params in
+ (TFun(ty_ret, Some params', false, attr), env1)
+ | _, Some params -> assert false
+ in
(* Extract info from type *)
let (ty_ret, params, vararg, attr) =
match ty with
- | TFun(ty_ret, Some params, vararg, attr) ->
+ | TFun(ty_ret, Some params, vararg, attr) ->
if wrap incomplete_type loc env1 ty_ret && not (is_void_type env ty_ret) then
fatal_error loc "return type is an incomplete type";
(ty_ret, params, vararg, attr)
@@ -1938,66 +1997,19 @@ let elab_fundef env spec name body loc =
emit_elab env1 loc (Gfundef fn);
env1
-let elab_kr_fundef env spec name params defs body loc =
- warning loc "Non-prototype, pre-standard function definition.@ Converting to prototype form";
- (* Check that the declarations only declare parameters *)
- let check_one_decl (Init_name(Name(s, dty, attrs, loc'), ie)) =
- if not (List.mem s params) then
- error loc' "Declaration of '%s' which is not a function parameter" s;
- if ie <> NO_INIT then
- error loc' "Illegal initialization of function parameter '%s'" s in
- let check_decl = function
- | DECDEF((spec', name_init_list), loc') ->
- List.iter check_one_decl name_init_list
- | d ->
- (* Should never be produced by the parser *)
- fatal_error (get_definitionloc d)
- "Illegal declaration of function parameter" in
- List.iter check_decl defs;
- (* Convert old-style K&R function definition to modern prototyped form *)
- let rec convert_param param = function
- | [] ->
- (* Parameter is not declared, defaults to "int" in ISO C90,
- is an error in ISO C99. Just emit a warning. *)
- warning loc "Type of '%s' defaults to 'int'" param;
- PARAM([SpecType Tint], Some param, JUSTBASE, [], loc)
- | DECDEF((spec', name_init_list), loc') :: defs ->
- let rec convert = function
- | [] -> convert_param param defs
- | Init_name(Name(s, dty, attrs, loc''), ie) :: l ->
- if s = param
- then PARAM(spec', Some param, dty, attrs, loc'')
- else convert l
- in convert name_init_list
- | _ ->
- assert false (* checked earlier *) in
- let params' =
- List.map (fun p -> convert_param p defs) params in
- let name' =
- let (Name(s, dty, attr, loc')) = name in
- Name(s, append_decltype dty (PROTO(JUSTBASE, (params', false))),
- attr, loc') in
- (* Elaborate the prototyped form *)
- elab_fundef env spec name' body loc
-
let rec elab_definition (local: bool) (env: Env.t) (def: Cabs.definition)
: decl list * Env.t =
match def with
(* "int f(int x) { ... }" *)
- | FUNDEF(spec, name, body, loc) ->
- if local then error loc "local definition of a function";
- let env1 = elab_fundef env spec name body loc in
- ([], env1)
-
(* "int f(x, y) double y; { ... }" *)
- | KRFUNDEF(spec, name, params, defs, body, loc) ->
+ | FUNDEF(spec, name, defs, body, loc) ->
if local then error loc "local definition of a function";
- let env1 = elab_kr_fundef env spec name params defs body loc in
+ let env1 = elab_fundef env spec name defs body loc in
([], env1)
(* "int x = 12, y[10], *z" *)
| DECDEF(init_name_group, loc) ->
- let ((dl, env1), sto, tydef) =
+ let ((dl, env1), sto, tydef) =
elab_init_name_group loc env init_name_group in
if tydef then
let env2 = enter_typedefs loc env1 sto dl
diff --git a/cparser/Parser.vy b/cparser/Parser.vy
index 7c0bfb55..16f6a0ef 100644
--- a/cparser/Parser.vy
+++ b/cparser/Parser.vy
@@ -52,6 +52,7 @@ Require Import List.
%type<list expression (* Reverse order *)> argument_expression_list
%type<definition> declaration
%type<list spec_elem * cabsloc> declaration_specifiers
+%type<list spec_elem> declaration_specifiers_typespec_opt
%type<list init_name (* Reverse order *)> init_declarator_list
%type<init_name> init_declarator
%type<storage * cabsloc> storage_class_specifier
@@ -65,9 +66,9 @@ Require Import List.
%type<list (string * option expression * cabsloc) (* Reverse order *)> enumerator_list
%type<string * option expression * cabsloc> enumerator
%type<string * cabsloc> enumeration_constant
-%type<cvspec * cabsloc> type_qualifier
+%type<cvspec * cabsloc> type_qualifier type_qualifier_noattr
%type<cabsloc> function_specifier
-%type<name> declarator direct_declarator
+%type<name> declarator declarator_noattrend direct_declarator
%type<(decl_type -> decl_type) * cabsloc> pointer
%type<list cvspec (* Reverse order *)> type_qualifier_list
%type<list parameter * bool> parameter_type_list
@@ -95,7 +96,6 @@ Require Import List.
%type<gcc_attribute> gcc_attribute
%type<list gcc_attribute> gcc_attribute_list
%type<gcc_attribute_word> gcc_attribute_word
-%type<name * list string> old_function_declarator direct_old_function_declarator
%type<list string (* Reverse order *)> identifier_list
%type<list asm_flag> asm_flags
%type<option string> asm_op_name
@@ -337,23 +337,34 @@ declaration:
| decspec = declaration_specifiers SEMICOLON
{ DECDEF (fst decspec, []) (snd decspec) }
+declaration_specifiers_typespec_opt:
+| storage = storage_class_specifier rest = declaration_specifiers_typespec_opt
+ { SpecStorage (fst storage)::rest }
+| typ = type_specifier rest = declaration_specifiers_typespec_opt
+ { SpecType (fst typ)::rest }
+| qual = type_qualifier rest = declaration_specifiers_typespec_opt
+ { SpecCV (fst qual)::rest }
+| loc = function_specifier rest = declaration_specifiers_typespec_opt
+ { SpecInline::rest }
+| /* empty */
+ { [] }
+
+
+(* We impose a lighter constraint on declaration specifiers than in the
+ pre_parser: declaration specifiers should have at least one type
+ specifier. *)
declaration_specifiers:
| storage = storage_class_specifier rest = declaration_specifiers
{ (SpecStorage (fst storage)::fst rest, snd storage) }
-| storage = storage_class_specifier
- { ([SpecStorage (fst storage)], snd storage) }
-| typ = type_specifier rest = declaration_specifiers
- { (SpecType (fst typ)::fst rest, snd typ) }
-| typ = type_specifier
- { ([SpecType (fst typ)], snd typ) }
-| qual = type_qualifier rest = declaration_specifiers
+| typ = type_specifier rest = declaration_specifiers_typespec_opt
+ { (SpecType (fst typ)::rest, snd typ) }
+(* We have to inline type_qualifier in order to avoid a conflict. *)
+| qual = type_qualifier_noattr rest = declaration_specifiers
{ (SpecCV (fst qual)::fst rest, snd qual) }
-| qual = type_qualifier
- { ([SpecCV (fst qual)], snd qual) }
+| attr = attribute_specifier rest = declaration_specifiers
+ { (SpecCV (CV_ATTR (fst attr))::fst rest, snd attr) }
| loc = function_specifier rest = declaration_specifiers
{ (SpecInline::fst rest, loc) }
-| loc = function_specifier
- { ([SpecInline], loc) }
init_declarator_list:
| init = init_declarator
@@ -411,12 +422,17 @@ type_specifier:
(* 6.7.2.1 *)
struct_or_union_specifier:
-| str_uni = struct_or_union attrs = attribute_specifier_list id = OTHER_NAME LBRACE decls = struct_declaration_list RBRACE
- { (Tstruct_union (fst str_uni) (Some (fst id)) (Some (rev' decls)) (rev' attrs), snd str_uni) }
-| str_uni = struct_or_union attrs = attribute_specifier_list LBRACE decls = struct_declaration_list RBRACE
- { (Tstruct_union (fst str_uni) None (Some (rev' decls)) (rev' attrs), snd str_uni) }
| str_uni = struct_or_union attrs = attribute_specifier_list id = OTHER_NAME
- { (Tstruct_union (fst str_uni) (Some (fst id)) None (rev' attrs), snd str_uni) }
+ LBRACE decls = struct_declaration_list RBRACE
+ { (Tstruct_union (fst str_uni) (Some (fst id)) (Some (rev' decls)) attrs,
+ snd str_uni) }
+| str_uni = struct_or_union attrs = attribute_specifier_list
+ LBRACE decls = struct_declaration_list RBRACE
+ { (Tstruct_union (fst str_uni) None (Some (rev' decls)) attrs,
+ snd str_uni) }
+| str_uni = struct_or_union attrs = attribute_specifier_list id = OTHER_NAME
+ { (Tstruct_union (fst str_uni) (Some (fst id)) None attrs,
+ snd str_uni) }
struct_or_union:
| loc = STRUCT
@@ -463,16 +479,20 @@ struct_declarator:
(* 6.7.2.2 *)
enum_specifier:
-| loc = ENUM attrs = attribute_specifier_list name = OTHER_NAME LBRACE enum_list = enumerator_list RBRACE
- { (Tenum (Some (fst name)) (Some (rev' enum_list)) (rev' attrs), loc) }
-| loc = ENUM attrs = attribute_specifier_list LBRACE enum_list = enumerator_list RBRACE
- { (Tenum None (Some (rev' enum_list)) (rev' attrs), loc) }
-| loc = ENUM attrs = attribute_specifier_list name = OTHER_NAME LBRACE enum_list = enumerator_list COMMA RBRACE
- { (Tenum (Some (fst name)) (Some (rev' enum_list)) (rev' attrs), loc) }
-| loc = ENUM attrs = attribute_specifier_list LBRACE enum_list = enumerator_list COMMA RBRACE
- { (Tenum None (Some (rev' enum_list)) (rev' attrs), loc) }
| loc = ENUM attrs = attribute_specifier_list name = OTHER_NAME
- { (Tenum (Some (fst name)) None (rev' attrs), loc) }
+ LBRACE enum_list = enumerator_list RBRACE
+ { (Tenum (Some (fst name)) (Some (rev' enum_list)) attrs, loc) }
+| loc = ENUM attrs = attribute_specifier_list
+ LBRACE enum_list = enumerator_list RBRACE
+ { (Tenum None (Some (rev' enum_list)) attrs, loc) }
+| loc = ENUM attrs = attribute_specifier_list name = OTHER_NAME
+ LBRACE enum_list = enumerator_list COMMA RBRACE
+ { (Tenum (Some (fst name)) (Some (rev' enum_list)) attrs, loc) }
+| loc = ENUM attrs = attribute_specifier_list
+ LBRACE enum_list = enumerator_list COMMA RBRACE
+ { (Tenum None (Some (rev' enum_list)) attrs, loc) }
+| loc = ENUM attrs = attribute_specifier_list name = OTHER_NAME
+ { (Tenum (Some (fst name)) None attrs, loc) }
enumerator_list:
| enum = enumerator
@@ -491,13 +511,17 @@ enumeration_constant:
{ cst }
(* 6.7.3 *)
-type_qualifier:
+type_qualifier_noattr:
| loc = CONST
{ (CV_CONST, loc) }
| loc = RESTRICT
{ (CV_RESTRICT, loc) }
| loc = VOLATILE
{ (CV_VOLATILE, loc) }
+
+type_qualifier:
+| qual = type_qualifier_noattr
+ { qual }
(* Non-standard *)
| attr = attribute_specifier
{ (CV_ATTR (fst attr), snd attr) }
@@ -507,17 +531,14 @@ type_qualifier:
attribute_specifier_list:
| /* empty */
{ [] }
-| attrs = attribute_specifier_list attr = attribute_specifier
- { fst attr :: attrs }
+| attr = attribute_specifier attrs = attribute_specifier_list
+ { fst attr :: attrs }
attribute_specifier:
| loc = ATTRIBUTE LPAREN LPAREN attr = gcc_attribute_list RPAREN RPAREN
{ (GCC_ATTR (rev' attr) loc, loc) }
| loc = PACKED LPAREN args = argument_expression_list RPAREN
{ (PACKED_ATTR (rev' args) loc, loc) }
-/* TODO: slove conflict */
-/*| loc = PACKED
- { (PACKED_ATTR [] loc, loc) }*/
| loc = ALIGNAS LPAREN args = argument_expression_list RPAREN
{ (ALIGNAS_ATTR (rev' args) loc, loc) }
| loc = ALIGNAS LPAREN typ = type_name RPAREN
@@ -554,12 +575,16 @@ function_specifier:
(* 6.7.5 *)
declarator:
-| decl = direct_declarator attrs = attribute_specifier_list
+| decl = declarator_noattrend attrs = attribute_specifier_list
{ match decl with Name name typ attr loc =>
Name name typ (List.app attr attrs) loc end }
-| pt = pointer decl = direct_declarator attrs = attribute_specifier_list
+
+declarator_noattrend:
+| decl = direct_declarator
+ { decl }
+| pt = pointer decl = direct_declarator
{ match decl with Name name typ attr _ =>
- Name name ((fst pt) typ) (List.app attr attrs) (snd pt) end }
+ Name name ((fst pt) typ) attr (snd pt) end }
direct_declarator:
| id = VAR_NAME
@@ -585,7 +610,10 @@ direct_declarator:
Name name (PROTO typ params) attr loc end }
| decl = direct_declarator LPAREN RPAREN
{ match decl with Name name typ attr loc =>
- Name name (PROTO typ ([], false)) attr loc end }
+ Name name (PROTO_OLD typ []) attr loc end }
+| decl = direct_declarator LPAREN params = identifier_list RPAREN
+ { match decl with Name name typ attr loc =>
+ Name name (PROTO_OLD typ (rev' params)) attr loc end }
pointer:
| loc = STAR
@@ -624,6 +652,12 @@ parameter_declaration:
| specs = declaration_specifiers
{ PARAM (fst specs) None JUSTBASE [] (snd specs) }
+identifier_list:
+| id = VAR_NAME
+ { [fst id] }
+| idl = identifier_list COMMA id = VAR_NAME
+ { fst id :: idl }
+
(* 6.7.6 *)
type_name:
| specqual = specifier_qualifier_list
@@ -902,40 +936,18 @@ external_declaration:
(* 6.9.1 *)
function_definition:
| specs = declaration_specifiers
- decl = declarator
+ decl = declarator_noattrend
+ dlist = declaration_list
stmt = compound_statement
- { FUNDEF (fst specs) decl stmt (snd specs) }
+ { FUNDEF (fst specs) decl (List.rev' dlist) stmt (snd specs) }
| specs = declaration_specifiers
- decl = old_function_declarator
- dlist = declaration_list
+ decl = declarator
stmt = compound_statement
- { KRFUNDEF (fst specs) (fst decl) (snd decl) (List.rev' dlist) stmt (snd specs) }
-
-(* Hack to parse old-style, unprototyped function definitions
- without causing ambiguities between a regular declarator
- (which can end with attributes) and a declaration list
- (which can start with attributes). *)
-
-old_function_declarator:
-| decl = direct_old_function_declarator
- { decl }
-| pt = pointer decl = direct_old_function_declarator
- { match decl with (Name name typ attr _, params) =>
- (Name name ((fst pt) typ) attr (snd pt), params) end }
-
-direct_old_function_declarator:
-| decl = direct_declarator LPAREN params = identifier_list RPAREN
- { (decl, List.rev' params) }
-
-identifier_list:
-| id = VAR_NAME
- { [fst id] }
-| idl = identifier_list COMMA id = VAR_NAME
- { fst id :: idl }
+ { FUNDEF (fst specs) decl [] stmt (snd specs) }
declaration_list:
-| /*empty*/
- { [] }
+| d = declaration
+ { [d] }
| dl = declaration_list d = declaration
{ d :: dl }
diff --git a/cparser/pre_parser.mly b/cparser/pre_parser.mly
index eacd59c8..41b068de 100644
--- a/cparser/pre_parser.mly
+++ b/cparser/pre_parser.mly
@@ -13,6 +13,13 @@
/* */
/* *********************************************************************/
+(*
+ WARNING: The precedence declarations tend to silently solve
+ conflicts. So, if you change the grammar (especially for
+ statements), you should check that without these declarations, it
+ has ONLY 3 CONFLICTS in 3 STATES.
+*)
+
%{
open Pre_parser_aux
@@ -35,6 +42,12 @@
Cerrors.fatal_error "%s:%d: this is the location of the unclosed '%s'"
pos1.Lexing.pos_fname pos1.Lexing.pos_lnum opening
+ type 'id fun_declarator_ctx =
+ | Decl_ident
+ | Decl_other
+ | Decl_fun of (unit -> unit)
+ | Decl_krfun of 'id
+
%}
%token<string> PRE_NAME
@@ -57,7 +70,7 @@
%token EOF
(* These precedence declarations solve the conflict in the following
- declaration :
+ declaration:
int f(int (a));
@@ -71,13 +84,6 @@
%nonassoc lowPrec2
%nonassoc ELSE
-(*
- WARNING: These precedence declarations tend to silently solve other
- conflicts. So, if you change the grammar (especially or
- statements), you should check that without these declarations, it
- has ONLY 3 CONFLICTS.
-*)
-
%start<unit> translation_unit_file
%%
@@ -89,10 +95,6 @@
| x = X
{ Some x }
-%inline fst(X):
-| x = X
- { fst x }
-
(* The kind of an identifier should not be determined when looking
ahead, because the context may not be up to date. For this reason,
when reading an identifier, the lexer emits two tokens: the first
@@ -113,16 +115,6 @@ general_identifier:
| i = var_name
{ i }
-(* We add this non-terminal here to force the resolution of the
- conflict at the point of shifting the TYPEDEF_NAME. If we had
- already shifted it, reduce/reduce conflict appear, and menhir is
- not able to solve them. *)
-low_prec : %prec lowPrec1 {}
-general_identifier_red:
-| PRE_NAME low_prec i = TYPEDEF_NAME
-| PRE_NAME i = VAR_NAME
- { i }
-
string_literals_list:
| string_literals_list? STRING_LITERAL
{}
@@ -131,9 +123,9 @@ save_context:
(* empty *) { !save_context () }
declare_varname(nt):
- i = nt { declare_varname i; i }
+ i = nt { declare_varname (fst i); i }
declare_typename(nt):
- i = nt { declare_typename i; i }
+ i = nt { declare_typename (fst i); i }
(* Actual grammar *)
@@ -301,7 +293,6 @@ constant_expression:
declaration:
| declaration_specifiers init_declarator_list? SEMICOLON
- {}
| declaration_specifiers_typedef typedef_declarator_list? SEMICOLON
{}
@@ -311,9 +302,9 @@ init_declarator_list:
{}
init_declarator:
-| declare_varname(fst(declarator))
-| declare_varname(fst(declarator)) EQ c_initializer
- { }
+| declare_varname(declarator_noattrend) save_context attribute_specifier_list
+| declare_varname(declarator_noattrend) save_context attribute_specifier_list EQ c_initializer
+ {}
typedef_declarator_list:
| typedef_declarator
@@ -321,8 +312,8 @@ typedef_declarator_list:
{}
typedef_declarator:
-| declare_typename(fst(declarator))
- { }
+| declare_typename(declarator)
+ {}
storage_class_specifier_no_typedef:
| EXTERN
@@ -334,9 +325,9 @@ storage_class_specifier_no_typedef:
(* [declaration_specifiers_no_type] matches declaration specifiers
that do not contain either "typedef" nor type specifiers. *)
declaration_specifiers_no_type:
-| declaration_specifiers_no_type? storage_class_specifier_no_typedef
-| declaration_specifiers_no_type? type_qualifier
-| declaration_specifiers_no_type? function_specifier
+| storage_class_specifier_no_typedef declaration_specifiers_no_type?
+| type_qualifier declaration_specifiers_no_type?
+| function_specifier declaration_specifiers_no_type?
{}
(* [declaration_specifiers_no_typedef_name] matches declaration
@@ -350,10 +341,8 @@ declaration_specifiers_no_typedef_name:
| declaration_specifiers_no_typedef_name? type_specifier_no_typedef_name
{}
-(* [declaration_specifiers_no_type] matches declaration_specifiers
- that do not contains "typedef". Moreover, it makes sure that it
- contains either one typename and not other type specifier or no
- typename.
+(* [declaration_specifiers] makes sure one type specifier is given, and,
+ if a typedef_name is given, then no other type specifier is given.
This is a weaker condition than 6.7.2 2. It is necessary to enforce
this in the grammar to disambiguate the example in 6.7.7 6:
@@ -474,13 +463,13 @@ enumerator_list:
enumerator:
| i = enumeration_constant
| i = enumeration_constant EQ constant_expression
- { i }
+ { (i, ()) }
enumeration_constant:
| i = general_identifier
{ set_id_type i VarId; i }
-type_qualifier:
+%inline type_qualifier:
| CONST
| RESTRICT
| VOLATILE
@@ -489,7 +478,7 @@ type_qualifier:
attribute_specifier_list:
| /* empty */
-| attribute_specifier_list attribute_specifier
+| attribute_specifier attribute_specifier_list
{}
attribute_specifier:
@@ -528,24 +517,50 @@ function_specifier:
| INLINE
{}
+(* We add this non-terminal here to force the resolution of the
+ conflict at the point of shifting the TYPEDEF_NAME. If we had
+ already shifted it, reduce/reduce conflict appear, and menhir is
+ not able to solve them. *)
+low_prec : %prec lowPrec1 {}
+declarator_identifier:
+| PRE_NAME low_prec i = TYPEDEF_NAME
+| PRE_NAME i = VAR_NAME
+ { i }
+
(* The semantic action returned by [declarator] is a pair of the
- identifier being defined and an option of the context stack that
- has to be restored if entering the body of the function being
+ identifier being defined and a value containing the context stack
+ that has to be restored if entering the body of the function being
defined, if so. *)
declarator:
-| pointer? x = direct_declarator attribute_specifier_list
+| x = declarator_noattrend attribute_specifier_list
{ x }
+declarator_noattrend:
+| x = direct_declarator
+ { x }
+| pointer x = direct_declarator
+ { match snd x with
+ | Decl_ident -> (fst x, Decl_other)
+ | _ -> x }
+
direct_declarator:
-| i = general_identifier_red
- { set_id_type i VarId; (i, None) }
+| i = declarator_identifier
+ { set_id_type i VarId; (i, Decl_ident) }
| LPAREN save_context x = declarator RPAREN
-| x = direct_declarator LBRACK type_qualifier_list? assignment_expression? RBRACK
{ x }
+| x = direct_declarator LBRACK type_qualifier_list? assignment_expression? RBRACK
+ { match snd x with
+ | Decl_ident -> (fst x, Decl_other)
+ | _ -> x }
| x = direct_declarator LPAREN ctx = context_parameter_type_list RPAREN
{ match snd x with
- | None -> (fst x, Some ctx)
- | Some _ -> x }
+ | Decl_ident -> (fst x, Decl_fun ctx)
+ | _ -> x }
+| x = direct_declarator LPAREN save_context il=identifier_list? RPAREN
+ { match snd x, il with
+ | Decl_ident, Some il -> (fst x, Decl_krfun il)
+ | Decl_ident, None -> (fst x, Decl_krfun [])
+ | _ -> x }
pointer:
| STAR type_qualifier_list?
@@ -557,25 +572,23 @@ type_qualifier_list:
{}
context_parameter_type_list:
-| ctx1 = save_context parameter_type_list? ctx2 = save_context
+| ctx1 = save_context parameter_type_list ctx2 = save_context
{ ctx1 (); ctx2 }
parameter_type_list:
-| l=parameter_list
-| l=parameter_list COMMA ELLIPSIS
- { l }
+| parameter_list
+| parameter_list COMMA ELLIPSIS
+ {}
parameter_list:
-| i=parameter_declaration
- { [i] }
-| l=parameter_list COMMA i=parameter_declaration
- { i::l }
+| parameter_declaration
+| parameter_list COMMA parameter_declaration
+ {}
parameter_declaration:
-| declaration_specifiers id=declare_varname(fst(declarator))
- { Some id }
+| declaration_specifiers declare_varname(declarator)
| declaration_specifiers abstract_declarator?
- { None }
+ {}
type_name:
| specifier_qualifier_list abstract_declarator?
@@ -745,39 +758,28 @@ external_declaration:
{}
identifier_list:
-| id = var_name
- { [id] }
-| idl = identifier_list COMMA id = var_name
- { id :: idl }
+| x = var_name
+ { [x] }
+| l = identifier_list COMMA x = var_name
+ { x::l }
declaration_list:
-| /*empty*/
- { }
+| declaration
| declaration_list declaration
- { }
+ {}
function_definition1:
-| declaration_specifiers pointer? x=direct_declarator ctx = save_context
- { match x with
- | (_, None) -> $syntaxerror
- | (_, Some ctx') -> ctx'(); ctx
- }
-| declaration_specifiers pointer? x=direct_declarator
- LPAREN save_context params=identifier_list RPAREN ctx = save_context declaration_list
- { match x with
- | (_, Some _) -> $syntaxerror
- | (i, None) ->
- declare_varname i;
- List.iter declare_varname params;
- ctx
- }
-
-function_definition2:
-| ctx = function_definition1 LBRACE block_item_list?
- { ctx() }
-| ctx = function_definition1 LBRACE block_item_list? error
- { unclosed "{" "}" $startpos($2) $endpos }
+| declaration_specifiers func = declare_varname(declarator_noattrend)
+ save_context attribute_specifier_list ctx = save_context
+| declaration_specifiers func = declare_varname(declarator_noattrend)
+ ctx = save_context declaration_list
+ { begin match snd func with
+ | Decl_fun ctx -> ctx (); declare_varname (fst func)
+ | Decl_krfun il -> List.iter declare_varname il
+ | _ -> ()
+ end;
+ ctx }
function_definition:
-| function_definition2 RBRACE
- { }
+| ctx = function_definition1 compound_statement
+ { ctx () }
diff --git a/test/regression/Results/parsing b/test/regression/Results/parsing
index e69de29b..94923756 100644
--- a/test/regression/Results/parsing
+++ b/test/regression/Results/parsing
@@ -0,0 +1,10 @@
+2
+3
+4
+5 4.500000 9.000000
+23
+3
+12 13 14
+4
+12 13 14
+4
diff --git a/test/regression/parsing.c b/test/regression/parsing.c
index 8687d6aa..ba093225 100644
--- a/test/regression/parsing.c
+++ b/test/regression/parsing.c
@@ -42,9 +42,9 @@ struct S {
unsigned T:3;
const T:3;
};
-
+struct S stru;
void i() {
- struct S s;
+ struct S s = stru;
s.T = -1;
if(s.T < 0) printf("ERROR i\n");
}
@@ -106,16 +106,72 @@ int j() {
T T;
}
-int k() {
+T k() {
{ T T; }
T t;
for(T T; ; );
T u;
}
+void krf(a)
+ int a;
+{
+ printf("%d\n", a);
+}
+
+void krg();
+void krg(int a)
+{
+ printf("%d\n", a);
+}
+
+void krh(int);
+void krh(b)
+ T b;
+{
+ printf("%d\n", b);
+}
+
+void kri();
+void kri(b, c)
+ int b;
+ double c;
+{
+ printf("%d %f %f\n", b, c, 2*c);
+}
+
+void krj();
+void krj(a, aa)
+ int a[];
+ void aa(int);
+{
+ printf("%d\n", *a);
+ aa(3);
+}
+
+void aa(int x) {
+ printf("%d\n", x);
+}
+
+void (*krk(a, b, c))(int)
+ int b, a, c;
+{
+ printf("%d %d %d\n", a, b, c);
+ return aa;
+}
+
int main () {
f(g);
i();
m();
+
+ krf(2);
+ krg(3);
+ krh(4);
+ kri(5.5, 4.5);
+ int x = 23;
+ krj(&x, aa);
+ krk(12, 13, 14)(4);
+ (*krk(12, 13, 14))(4);
return 0;
}