aboutsummaryrefslogtreecommitdiffstats
path: root/arm/SelectOpproof.v
diff options
context:
space:
mode:
authorxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2014-05-24 09:46:07 +0000
committerxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2014-05-24 09:46:07 +0000
commit3fa79790e617d87584598746296e626e0ce3b256 (patch)
treedcdc926130d9ed8d302eedf8215d065c0e787eed /arm/SelectOpproof.v
parent285d908c5dbd90bff7f03618c7a9e0fa5e287c94 (diff)
downloadcompcert-kvx-3fa79790e617d87584598746296e626e0ce3b256.tar.gz
compcert-kvx-3fa79790e617d87584598746296e626e0ce3b256.zip
Refactoring: move symbol_offset into Genv.
git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2503 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'arm/SelectOpproof.v')
-rw-r--r--arm/SelectOpproof.v4
1 files changed, 2 insertions, 2 deletions
diff --git a/arm/SelectOpproof.v b/arm/SelectOpproof.v
index d10e7fc3..4b891191 100644
--- a/arm/SelectOpproof.v
+++ b/arm/SelectOpproof.v
@@ -114,7 +114,7 @@ Definition binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> va
Theorem eval_addrsymbol:
forall le id ofs,
- exists v, eval_expr ge sp e m le (addrsymbol id ofs) v /\ Val.lessdef (symbol_address ge id ofs) v.
+ exists v, eval_expr ge sp e m le (addrsymbol id ofs) v /\ Val.lessdef (Genv.symbol_address ge id ofs) v.
Proof.
intros. unfold addrsymbol. econstructor; split.
EvalOp. simpl; eauto.
@@ -150,7 +150,7 @@ Proof.
destruct x; simpl; auto. rewrite Int.add_zero. auto. rewrite Int.add_zero. auto.
case (addimm_match a); intros; InvEval; simpl; TrivialExists; simpl.
rewrite Int.add_commut. auto.
- unfold symbol_address. destruct (Genv.find_symbol ge s); simpl; auto. rewrite Int.add_commut; auto.
+ unfold Genv.symbol_address. destruct (Genv.find_symbol ge s); simpl; auto. rewrite Int.add_commut; auto.
rewrite Val.add_assoc. rewrite Int.add_commut. auto.
subst x. rewrite Val.add_assoc. rewrite Int.add_commut. auto.
Qed.