summaryrefslogtreecommitdiff
path: root/backend/ValueAnalysis.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2014-02-19 09:55:45 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2014-02-19 09:55:45 +0000
commit3ec022950ec233a2af418aacd1755fce4d701724 (patch)
tree154256c5c73fda06e874fb05695e14e610ba8ad4 /backend/ValueAnalysis.v
parent9aeba45962e8ba5cde5d81fb701a4c9a3963f4a5 (diff)
Add option -Os to optimize for code size rather than for execution speed.
Refactored compilation flags that affect the Coq part (module Compopts). Added support for C99 for loops with declarations. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2410 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'backend/ValueAnalysis.v')
-rw-r--r--backend/ValueAnalysis.v7
1 files changed, 4 insertions, 3 deletions
diff --git a/backend/ValueAnalysis.v b/backend/ValueAnalysis.v
index f580438..e293028 100644
--- a/backend/ValueAnalysis.v
+++ b/backend/ValueAnalysis.v
@@ -12,6 +12,7 @@
Require Import Coqlib.
Require Import Maps.
+Require Import Compopts.
Require Import AST.
Require Import Integers.
Require Import Floats.
@@ -74,7 +75,7 @@ Definition transfer_builtin (ae: aenv) (am: amem) (rm: romem) (ef: external_func
match classify_builtin ef args ae with
| Builtin_vload chunk aaddr =>
let a :=
- if strict
+ if va_strict tt
then vlub (loadv chunk rm am aaddr) (vnormalize chunk (Ifptr Glob))
else vnormalize chunk Vtop in
VA.State (AE.set res a ae) am
@@ -1255,11 +1256,11 @@ Proof.
* (* true volatile access *)
assert (V: vmatch bc v0 (Ifptr Glob)).
{ inv H4; constructor. econstructor. eapply GE; eauto. }
- destruct strict. apply vmatch_lub_r. apply vnormalize_sound. auto.
+ destruct (va_strict tt). apply vmatch_lub_r. apply vnormalize_sound. auto.
apply vnormalize_sound. eapply vmatch_ge; eauto. constructor. constructor.
* (* normal memory access *)
exploit loadv_sound; eauto. simpl; eauto. intros V.
- destruct strict.
+ destruct (va_strict tt).
apply vmatch_lub_l. auto.
eapply vnormalize_cast; eauto. eapply vmatch_top; eauto.
+ (* volatile store *)