aboutsummaryrefslogtreecommitdiffstats
path: root/backend/Inlining.v
diff options
context:
space:
mode:
authorxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2013-10-18 06:55:06 +0000
committerxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2013-10-18 06:55:06 +0000
commitca281a5ff122f136db761581f95110465b5eea31 (patch)
tree4f67e8604a6b18b2e3a700490793a6eb7bd2b66a /backend/Inlining.v
parentc857f0b02463f4b0bc8100434eecdd46ce2ecbd1 (diff)
downloadcompcert-kvx-ca281a5ff122f136db761581f95110465b5eea31.tar.gz
compcert-kvx-ca281a5ff122f136db761581f95110465b5eea31.zip
Revised renumbering of nodes and registers so that main function is not shifted by 1
git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2346 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'backend/Inlining.v')
-rw-r--r--backend/Inlining.v10
1 files changed, 6 insertions, 4 deletions
diff --git a/backend/Inlining.v b/backend/Inlining.v
index 683d1900..c6df3de0 100644
--- a/backend/Inlining.v
+++ b/backend/Inlining.v
@@ -143,9 +143,9 @@ Qed.
Program Definition add_instr (i: instruction): mon node :=
fun s =>
- let pc := Psucc s.(st_nextnode) in
+ let pc := s.(st_nextnode) in
R pc
- (mkstate s.(st_nextreg) pc (PTree.set pc i s.(st_code)) s.(st_stksize))
+ (mkstate s.(st_nextreg) (Psucc pc) (PTree.set pc i s.(st_code)) s.(st_stksize))
_.
Next Obligation.
intros; constructor; simpl; xomega.
@@ -213,9 +213,11 @@ Record context: Type := mkcontext {
(** The following functions "shift" (relocate) PCs, registers, operations, etc. *)
-Definition spc (ctx: context) (pc: node) := Pplus pc ctx.(dpc).
+Definition shiftpos (p amount: positive) := Ppred (Pplus p amount).
-Definition sreg (ctx: context) (r: reg) := Pplus r ctx.(dreg).
+Definition spc (ctx: context) (pc: node) := shiftpos pc ctx.(dpc).
+
+Definition sreg (ctx: context) (r: reg) := shiftpos r ctx.(dreg).
Definition sregs (ctx: context) (rl: list reg) := List.map (sreg ctx) rl.