aboutsummaryrefslogtreecommitdiffstats
path: root/cfrontend/Initializers.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2016-10-06 09:49:50 +0200
committerXavier Leroy <xavier.leroy@inria.fr>2016-10-06 09:49:50 +0200
commitfecf2472ee5c2f6da25eb1ea340e0f07dcdaec69 (patch)
tree5bbe7afa05895ba7857b86e7c5d7d40c29076e1b /cfrontend/Initializers.v
parent67c58f9ee025fb2fc0de6d932366b7e5db7a6678 (diff)
downloadcompcert-kvx-fecf2472ee5c2f6da25eb1ea340e0f07dcdaec69.tar.gz
compcert-kvx-fecf2472ee5c2f6da25eb1ea340e0f07dcdaec69.zip
Regression: compile-time evaluation of ((struct s *)0)->field
This somewhat useful idiom was broken when Val.add was replaced by Val.offset_ptr in the Efield case. This commit restores the previous behavior. This used to be supported (and is useful
Diffstat (limited to 'cfrontend/Initializers.v')
-rw-r--r--cfrontend/Initializers.v4
1 files changed, 3 insertions, 1 deletions
diff --git a/cfrontend/Initializers.v b/cfrontend/Initializers.v
index 5b7e52c8..19518aea 100644
--- a/cfrontend/Initializers.v
+++ b/cfrontend/Initializers.v
@@ -119,7 +119,9 @@ Fixpoint constval (ce: composite_env) (a: expr) : res val :=
do co <- lookup_composite ce id;
do delta <- field_offset ce f (co_members co);
do v <- constval ce l;
- OK (Val.offset_ptr v (Ptrofs.repr delta))
+ OK (if Archi.ptr64
+ then Val.addl v (Vlong (Int64.repr delta))
+ else Val.add v (Vint (Int.repr delta)))
| Tunion id _ =>
constval ce l
| _ =>