aboutsummaryrefslogtreecommitdiffstats
path: root/cparser/Cabs.v
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 /cparser/Cabs.v
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.
Diffstat (limited to 'cparser/Cabs.v')
-rw-r--r--cparser/Cabs.v8
1 files changed, 7 insertions, 1 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.