summaryrefslogtreecommitdiff
path: root/cfrontend/SimplExprspec.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2012-02-26 10:41:07 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2012-02-26 10:41:07 +0000
commit2570ddd61b1c98b62c8d97fce862654535696844 (patch)
treee9a652b115045a3b2c4ade69ec3cc3fdad429b54 /cfrontend/SimplExprspec.v
parent65cc3738e7436e46f70c0508638a71fbb49c50a8 (diff)
- Support for _Alignof(ty) operator from ISO C 2011
and __alignof__(ty), __alignof__(expr) from GCC. - Resurrected __builtin_memcpy_aligned, useful for files generated by Scade KCG 6. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1827 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'cfrontend/SimplExprspec.v')
-rw-r--r--cfrontend/SimplExprspec.v7
1 files changed, 7 insertions, 0 deletions
diff --git a/cfrontend/SimplExprspec.v b/cfrontend/SimplExprspec.v
index 647e048..bd7f336 100644
--- a/cfrontend/SimplExprspec.v
+++ b/cfrontend/SimplExprspec.v
@@ -89,6 +89,10 @@ Inductive tr_expr: temp_env -> destination -> C.expr -> list statement -> expr -
tr_expr le dst (C.Esizeof ty' ty)
(final dst (Esizeof ty' ty))
(Esizeof ty' ty) tmp
+ | tr_alignof: forall le dst ty' ty tmp,
+ tr_expr le dst (C.Ealignof ty' ty)
+ (final dst (Ealignof ty' ty))
+ (Ealignof ty' ty) tmp
| tr_valof: forall le dst e1 ty tmp sl1 a1 tmp1 sl2 a2 tmp2,
tr_expr le For_val e1 sl1 a1 tmp1 ->
tr_rvalof (C.typeof e1) a1 sl2 a2 tmp2 ->
@@ -719,6 +723,9 @@ Opaque makeif.
(* sizeof *)
monadInv H. UseFinish.
exists (@nil ident); split; auto with gensym. constructor.
+(* alignof *)
+ monadInv H. UseFinish.
+ exists (@nil ident); split; auto with gensym. constructor.
(* assign *)
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
exploit H0; eauto. intros [tmp2 [C D]].