From abe2bb5c40260a31ce5ee27b841bcbd647ff8b88 Mon Sep 17 00:00:00 2001 From: xleroy Date: Sat, 9 Apr 2011 16:59:13 +0000 Subject: Merge of branch "unsigned-offsets": - In pointer values "Vptr b ofs", interpret "ofs" as an unsigned int. (Fixes issue with wrong comparison of pointers across 0x8000_0000) - Revised Stacking pass to not use negative SP offsets. - Add pointer validity checks to Cminor ... Mach to support the use of memory injections in Stacking. - Cleaned up Stacklayout modules. - IA32: improved code generation for Mgetparam. - ARM: improved code generation for op-immediate instructions. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1632 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e --- backend/RTL.v | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) (limited to 'backend/RTL.v') diff --git a/backend/RTL.v b/backend/RTL.v index 208c7b13..2cb27196 100644 --- a/backend/RTL.v +++ b/backend/RTL.v @@ -217,7 +217,7 @@ Inductive step: state -> trace -> state -> Prop := | exec_Iop: forall s f sp pc rs m op args res pc' v, (fn_code f)!pc = Some(Iop op args res pc') -> - eval_operation ge sp op rs##args = Some v -> + eval_operation ge sp op rs##args m = Some v -> step (State s f sp pc rs m) E0 (State s f sp pc' (rs#res <- v) m) | exec_Iload: @@ -258,20 +258,20 @@ Inductive step: state -> trace -> state -> Prop := | exec_Icond_true: forall s f sp pc rs m cond args ifso ifnot, (fn_code f)!pc = Some(Icond cond args ifso ifnot) -> - eval_condition cond rs##args = Some true -> + eval_condition cond rs##args m = Some true -> step (State s f sp pc rs m) E0 (State s f sp ifso rs m) | exec_Icond_false: forall s f sp pc rs m cond args ifso ifnot, (fn_code f)!pc = Some(Icond cond args ifso ifnot) -> - eval_condition cond rs##args = Some false -> + eval_condition cond rs##args m = Some false -> step (State s f sp pc rs m) E0 (State s f sp ifnot rs m) | exec_Ijumptable: forall s f sp pc rs m arg tbl n pc', (fn_code f)!pc = Some(Ijumptable arg tbl) -> rs#arg = Vint n -> - list_nth_z tbl (Int.signed n) = Some pc' -> + list_nth_z tbl (Int.unsigned n) = Some pc' -> step (State s f sp pc rs m) E0 (State s f sp pc' rs m) | exec_Ireturn: @@ -303,7 +303,7 @@ Inductive step: state -> trace -> state -> Prop := Lemma exec_Iop': forall s f sp pc rs m op args res pc' rs' v, (fn_code f)!pc = Some(Iop op args res pc') -> - eval_operation ge sp op rs##args = Some v -> + eval_operation ge sp op rs##args m = Some v -> rs' = (rs#res <- v) -> step (State s f sp pc rs m) E0 (State s f sp pc' rs' m). -- cgit