summaryrefslogtreecommitdiff
path: root/common/Globalenvs.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2012-05-26 07:32:01 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2012-05-26 07:32:01 +0000
commit132e36fa0be63eb5672eda9168403d3fb74af2fa (patch)
tree33955e0ccb4210271c82326b941523e6e4b2c289 /common/Globalenvs.v
parent9ea00d39bb32c1f188f1af2745c3368da6a349c1 (diff)
CSE: add recognition of some combined operators, conditions, and addressing modes (cf. CombineOp.v)
Memory model: cleaning up Memdata Inlining and new Constprop: updated for ARM. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1902 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'common/Globalenvs.v')
-rw-r--r--common/Globalenvs.v13
1 files changed, 6 insertions, 7 deletions
diff --git a/common/Globalenvs.v b/common/Globalenvs.v
index d7449f9..ba038f8 100644
--- a/common/Globalenvs.v
+++ b/common/Globalenvs.v
@@ -897,7 +897,6 @@ Proof.
assert (A: forall chunk v m b p m1 il m',
Mem.store chunk m b p v = Some m1 ->
store_init_data_list m1 b (p + size_chunk chunk) il = Some m' ->
- Val.has_type v (type_of_chunk chunk) ->
Mem.load chunk m' b p = Some(Val.load_result chunk v)).
intros. transitivity (Mem.load chunk m1 b p).
eapply store_init_data_list_outside; eauto. right. omega.
@@ -909,13 +908,13 @@ Proof.
intros m1 B C.
exploit IHil; eauto. intro D.
destruct a; simpl in B; intuition.
- eapply (A Mint8unsigned (Vint i)); eauto. simpl; auto.
- eapply (A Mint16unsigned (Vint i)); eauto. simpl; auto.
- eapply (A Mint32 (Vint i)); eauto. simpl; auto.
- eapply (A Mfloat32 (Vfloat f)); eauto. simpl; auto.
- eapply (A Mfloat64 (Vfloat f)); eauto. simpl; auto.
+ eapply (A Mint8unsigned (Vint i)); eauto.
+ eapply (A Mint16unsigned (Vint i)); eauto.
+ eapply (A Mint32 (Vint i)); eauto.
+ eapply (A Mfloat32 (Vfloat f)); eauto.
+ eapply (A Mfloat64 (Vfloat f)); eauto.
destruct (find_symbol ge i); try congruence. exists b0; split; auto.
- eapply (A Mint32 (Vptr b0 i0)); eauto. simpl; auto.
+ eapply (A Mint32 (Vptr b0 i0)); eauto.
Qed.
Remark load_alloc_variables: