aboutsummaryrefslogtreecommitdiffstats
path: root/backend
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2016-10-01 17:25:18 +0200
committerXavier Leroy <xavier.leroy@inria.fr>2016-10-01 17:25:18 +0200
commite637d041c5c2ee3a3ed395a7dab6c9101e8eb16c (patch)
tree518d558674d3e1c6ff41c46d84c784e727ed5d04 /backend
parentad2a2c862abef3aee701d1bca0524fcbf2d07b30 (diff)
downloadcompcert-kvx-e637d041c5c2ee3a3ed395a7dab6c9101e8eb16c.tar.gz
compcert-kvx-e637d041c5c2ee3a3ed395a7dab6c9101e8eb16c.zip
Support for 64-bit architectures: generic support
- Introduce Archi.ptr64 parameter. - Define module Ptrofs of integers as wide as a pointer (64 if Archi.ptr64, 32 otherwise). - Use Ptrofs.int as the offset type for Vptr values and anywhere pointer offsets are manipulated. - Modify Val operations that handle pointers (e.g. Val.add, Val.sub, Val.cmpu) so that in 64-bit pointer mode it is the "long" operation (e.g. Val.addl, Val.subl, Val.cmplu) that handles pointers. - Update the memory model accordingly. - Modify C operations that handle pointers (e.g. addition, subtraction, comparisons) accordingly. - Make it possible to turn off the splitting of 64-bit integers into pairs of 32-bit integers. - Update the compiler front-end and back-end accordingly.
Diffstat (limited to 'backend')
-rw-r--r--backend/Allocation.v120
-rw-r--r--backend/Allocproof.v400
-rw-r--r--backend/Asmgenproof0.v75
-rw-r--r--backend/Bounds.v43
-rw-r--r--backend/CSE.v10
-rw-r--r--backend/CSEproof.v64
-rw-r--r--backend/Cminor.v45
-rw-r--r--backend/CminorSel.v22
-rw-r--r--backend/Constprop.v12
-rw-r--r--backend/Constpropproof.v42
-rw-r--r--backend/Deadcodeproof.v32
-rw-r--r--backend/Debugvar.v2
-rw-r--r--backend/IRC.ml41
-rw-r--r--backend/IRC.mli4
-rw-r--r--backend/Inlining.v12
-rw-r--r--backend/Inliningproof.v50
-rw-r--r--backend/Inliningspec.v4
-rw-r--r--backend/LTL.v21
-rw-r--r--backend/Linear.v19
-rw-r--r--backend/Lineartyping.v3
-rw-r--r--backend/Mach.v46
-rw-r--r--backend/NeedDomain.v64
-rw-r--r--backend/PrintAsmaux.ml3
-rw-r--r--backend/RTL.v20
-rw-r--r--backend/RTLtyping.v18
-rw-r--r--backend/Regalloc.ml88
-rw-r--r--backend/SelectDiv.vp36
-rw-r--r--backend/SelectDivproof.v45
-rw-r--r--backend/Selection.v64
-rw-r--r--backend/Selectionproof.v157
-rw-r--r--backend/SplitLong.vp (renamed from backend/SelectLong.vp)60
-rw-r--r--backend/SplitLongproof.v (renamed from backend/SelectLongproof.v)150
-rw-r--r--backend/Stacking.v30
-rw-r--r--backend/Stackingproof.v180
-rw-r--r--backend/Tailcallproof.v26
-rw-r--r--backend/Unusedglobproof.v42
-rw-r--r--backend/ValueAnalysis.v40
-rw-r--r--backend/ValueDomain.v731
38 files changed, 1624 insertions, 1197 deletions
diff --git a/backend/Allocation.v b/backend/Allocation.v
index 0d25d84a..f561ef4e 100644
--- a/backend/Allocation.v
+++ b/backend/Allocation.v
@@ -39,7 +39,8 @@ Require Import Op Registers RTL Locations Conventions RTLtyping LTL.
maching between an RTL instruction and an LTL basic block.
*)
-Definition moves := list (loc * loc)%type.
+Definition move := (loc * loc)%type.
+Definition moves := list move.
Inductive block_shape: Type :=
| BSnop (mv: moves) (s: node)
@@ -89,6 +90,25 @@ Inductive block_shape: Type :=
| BSreturn (arg: option reg)
(mv: moves).
+(** Classify operations into moves, 64-bit split integer operations, and other
+ arithmetic/logical operations. *)
+
+Inductive operation_kind {A: Type}: operation -> list A -> Type :=
+ | operation_Omove: forall arg, operation_kind Omove (arg :: nil)
+ | operation_Omakelong: forall arg1 arg2, operation_kind Omakelong (arg1 :: arg2 :: nil)
+ | operation_Olowlong: forall arg, operation_kind Olowlong (arg :: nil)
+ | operation_Ohighlong: forall arg, operation_kind Ohighlong (arg :: nil)
+ | operation_other: forall op args, operation_kind op args.
+
+Definition classify_operation {A: Type} (op: operation) (args: list A) : operation_kind op args :=
+ match op, args with
+ | Omove, arg::nil => operation_Omove arg
+ | Omakelong, arg1::arg2::nil => operation_Omakelong arg1 arg2
+ | Olowlong, arg::nil => operation_Olowlong arg
+ | Ohighlong, arg::nil => operation_Ohighlong arg
+ | op, args => operation_other op args
+ end.
+
(** Extract the move instructions at the beginning of block [b].
Return the list of moves and the suffix of [b] after the moves. *)
@@ -100,8 +120,10 @@ Fixpoint extract_moves (accu: moves) (b: bblock) {struct b} : moves * bblock :=
extract_moves ((R src, S sl ofs ty) :: accu) b'
| Lop op args res :: b' =>
match is_move_operation op args with
- | Some arg => extract_moves ((R arg, R res) :: accu) b'
- | None => (List.rev accu, b)
+ | Some arg =>
+ extract_moves ((R arg, R res) :: accu) b'
+ | None =>
+ (List.rev accu, b)
end
| _ =>
(List.rev accu, b)
@@ -123,29 +145,23 @@ Notation "'assertion' A ; B" := (if A then B else None)
Local Open Scope option_monad_scope.
-(** Classify operations into moves, 64-bit integer operations, and other
- arithmetic/logical operations. *)
-
-Inductive operation_kind: operation -> list reg -> Type :=
- | operation_Omove: forall arg, operation_kind Omove (arg :: nil)
- | operation_Omakelong: forall arg1 arg2, operation_kind Omakelong (arg1 :: arg2 :: nil)
- | operation_Olowlong: forall arg, operation_kind Olowlong (arg :: nil)
- | operation_Ohighlong: forall arg, operation_kind Ohighlong (arg :: nil)
- | operation_other: forall op args, operation_kind op args.
-
-Definition classify_operation (op: operation) (args: list reg) : operation_kind op args :=
- match op, args with
- | Omove, arg::nil => operation_Omove arg
- | Omakelong, arg1::arg2::nil => operation_Omakelong arg1 arg2
- | Olowlong, arg::nil => operation_Olowlong arg
- | Ohighlong, arg::nil => operation_Ohighlong arg
- | op, args => operation_other op args
- end.
-
(** Check RTL instruction [i] against LTL basic block [b].
On success, return [Some] with a [block_shape] describing the correspondence.
On error, return [None]. *)
+Definition pair_Iop_block (op: operation) (args: list reg) (res: reg) (s: node) (b: LTL.bblock) :=
+ let (mv1, b1) := extract_moves nil b in
+ match b1 with
+ | Lop op' args' res' :: b2 =>
+ let (mv2, b3) := extract_moves nil b2 in
+ assertion (eq_operation op op');
+ assertion (check_succ s b3);
+ Some(BSop op args res mv1 args' res' mv2 s)
+ | _ =>
+ assertion (check_succ s b1);
+ Some(BSopdead op args res mv1 s)
+ end.
+
Definition pair_instr_block
(i: RTL.instruction) (b: LTL.bblock) : option block_shape :=
match i with
@@ -158,32 +174,31 @@ Definition pair_instr_block
let (mv, b1) := extract_moves nil b in
assertion (check_succ s b1); Some(BSmove arg res mv s)
| operation_Omakelong arg1 arg2 =>
- let (mv, b1) := extract_moves nil b in
- assertion (check_succ s b1); Some(BSmakelong arg1 arg2 res mv s)
+ if Archi.splitlong then
+ (let (mv, b1) := extract_moves nil b in
+ assertion (check_succ s b1); Some(BSmakelong arg1 arg2 res mv s))
+ else
+ pair_Iop_block op args res s b
| operation_Olowlong arg =>
- let (mv, b1) := extract_moves nil b in
- assertion (check_succ s b1); Some(BSlowlong arg res mv s)
+ if Archi.splitlong then
+ (let (mv, b1) := extract_moves nil b in
+ assertion (check_succ s b1); Some(BSlowlong arg res mv s))
+ else
+ pair_Iop_block op args res s b
| operation_Ohighlong arg =>
- let (mv, b1) := extract_moves nil b in
- assertion (check_succ s b1); Some(BShighlong arg res mv s)
+ if Archi.splitlong then
+ (let (mv, b1) := extract_moves nil b in
+ assertion (check_succ s b1); Some(BShighlong arg res mv s))
+ else
+ pair_Iop_block op args res s b
| operation_other _ _ =>
- let (mv1, b1) := extract_moves nil b in
- match b1 with
- | Lop op' args' res' :: b2 =>
- let (mv2, b3) := extract_moves nil b2 in
- assertion (eq_operation op op');
- assertion (check_succ s b3);
- Some(BSop op args res mv1 args' res' mv2 s)
- | _ =>
- assertion (check_succ s b1);
- Some(BSopdead op args res mv1 s)
- end
+ pair_Iop_block op args res s b
end
| Iload chunk addr args dst s =>
let (mv1, b1) := extract_moves nil b in
match b1 with
| Lload chunk' addr' args' dst' :: b2 =>
- if chunk_eq chunk Mint64 then
+ if chunk_eq chunk Mint64 && Archi.splitlong then
assertion (chunk_eq chunk' Mint32);
let (mv2, b3) := extract_moves nil b2 in
match b3 with
@@ -191,7 +206,7 @@ Definition pair_instr_block
let (mv3, b5) := extract_moves nil b4 in
assertion (chunk_eq chunk'' Mint32);
assertion (eq_addressing addr addr');
- assertion (option_eq eq_addressing (offset_addressing addr (Int.repr 4)) (Some addr''));
+ assertion (option_eq eq_addressing (offset_addressing addr 4) (Some addr''));
assertion (check_succ s b5);
Some(BSload2 addr addr'' args dst mv1 args' dst' mv2 args'' dst'' mv3 s)
| _ =>
@@ -199,7 +214,7 @@ Definition pair_instr_block
if (eq_addressing addr addr') then
Some(BSload2_1 addr args dst mv1 args' dst' mv2 s)
else
- (assertion (option_eq eq_addressing (offset_addressing addr (Int.repr 4)) (Some addr'));
+ (assertion (option_eq eq_addressing (offset_addressing addr 4) (Some addr'));
Some(BSload2_2 addr addr' args dst mv1 args' dst' mv2 s))
end
else (
@@ -216,14 +231,14 @@ Definition pair_instr_block
let (mv1, b1) := extract_moves nil b in
match b1 with
| Lstore chunk' addr' args' src' :: b2 =>
- if chunk_eq chunk Mint64 then
+ if chunk_eq chunk Mint64 && Archi.splitlong then
let (mv2, b3) := extract_moves nil b2 in
match b3 with
| Lstore chunk'' addr'' args'' src'' :: b4 =>
assertion (chunk_eq chunk' Mint32);
assertion (chunk_eq chunk'' Mint32);
assertion (eq_addressing addr addr');
- assertion (option_eq eq_addressing (offset_addressing addr (Int.repr 4)) (Some addr''));
+ assertion (option_eq eq_addressing (offset_addressing addr 4) (Some addr''));
assertion (check_succ s b4);
Some(BSstore2 addr addr'' args src mv1 args' src' mv2 args'' src'' s)
| _ => None
@@ -622,7 +637,9 @@ Function add_equations_args (rl: list reg) (tyl: list typ) (ll: list (rpair loc)
| r1 :: rl, ty :: tyl, One l1 :: ll =>
add_equations_args rl tyl ll (add_equation (Eq Full r1 l1) e)
| r1 :: rl, Tlong :: tyl, Twolong l1 l2 :: ll =>
- add_equations_args rl tyl ll (add_equation (Eq Low r1 l2) (add_equation (Eq High r1 l1) e))
+ if Archi.splitlong then
+ add_equations_args rl tyl ll (add_equation (Eq Low r1 l2) (add_equation (Eq High r1 l1) e))
+ else None
| _, _, _ => None
end.
@@ -634,7 +651,9 @@ Function add_equations_res (r: reg) (oty: option typ) (p: rpair mreg) (e: eqs) :
| One mr, _ =>
Some (add_equation (Eq Full r (R mr)) e)
| Twolong mr1 mr2, Some Tlong =>
- Some (add_equation (Eq Low r (R mr2)) (add_equation (Eq High r (R mr1)) e))
+ if Archi.splitlong then
+ Some (add_equation (Eq Low r (R mr2)) (add_equation (Eq High r (R mr1)) e))
+ else None
| _, _ =>
None
end.
@@ -673,6 +692,7 @@ Fixpoint add_equations_builtin_arg
Some (add_equation (Eq Full r l) e)
| BA r, BA_splitlong (BA lhi) (BA llo) =>
assertion (typ_eq (env r) Tlong);
+ assertion (Archi.splitlong);
Some (add_equation (Eq Low r llo) (add_equation (Eq High r lhi) e))
| BA_int n, BA_int n' =>
assertion (Int.eq_dec n n'); Some e
@@ -684,19 +704,19 @@ Fixpoint add_equations_builtin_arg
assertion (Float32.eq_dec f f'); Some e
| BA_loadstack chunk ofs, BA_loadstack chunk' ofs' =>
assertion (chunk_eq chunk chunk');
- assertion (Int.eq_dec ofs ofs');
+ assertion (Ptrofs.eq_dec ofs ofs');
Some e
| BA_addrstack ofs, BA_addrstack ofs' =>
- assertion (Int.eq_dec ofs ofs');
+ assertion (Ptrofs.eq_dec ofs ofs');
Some e
| BA_loadglobal chunk id ofs, BA_loadglobal chunk' id' ofs' =>
assertion (chunk_eq chunk chunk');
assertion (ident_eq id id');
- assertion (Int.eq_dec ofs ofs');
+ assertion (Ptrofs.eq_dec ofs ofs');
Some e
| BA_addrglobal id ofs, BA_addrglobal id' ofs' =>
assertion (ident_eq id id');
- assertion (Int.eq_dec ofs ofs');
+ assertion (Ptrofs.eq_dec ofs ofs');
Some e
| BA_splitlong hi lo, BA_splitlong hi' lo' =>
do e1 <- add_equations_builtin_arg env hi hi' e;
diff --git a/backend/Allocproof.v b/backend/Allocproof.v
index 47dac12f..888945ec 100644
--- a/backend/Allocproof.v
+++ b/backend/Allocproof.v
@@ -32,8 +32,8 @@ Qed.
(** * Soundness of structural checks *)
-Definition expand_move (sd: loc * loc) : instruction :=
- match sd with
+Definition expand_move (m: move) : instruction :=
+ match m with
| (R src, R dst) => Lop Omove (src::nil) dst
| (S sl ofs ty, R dst) => Lgetstack sl ofs ty dst
| (R src, S sl ofs ty) => Lsetstack src sl ofs ty
@@ -43,14 +43,14 @@ Definition expand_move (sd: loc * loc) : instruction :=
Definition expand_moves (mv: moves) (k: bblock) : bblock :=
List.map expand_move mv ++ k.
-Definition wf_move (sd: loc * loc) : Prop :=
- match sd with
+Definition wf_move (m: move) : Prop :=
+ match m with
| (S _ _ _, S _ _ _) => False
| _ => True
end.
Definition wf_moves (mv: moves) : Prop :=
- forall sd, In sd mv -> wf_move sd.
+ List.Forall wf_move mv.
Inductive expand_block_shape: block_shape -> RTL.instruction -> LTL.bblock -> Prop :=
| ebs_nop: forall mv s k,
@@ -64,17 +64,17 @@ Inductive expand_block_shape: block_shape -> RTL.instruction -> LTL.bblock -> Pr
(Iop Omove (src :: nil) dst s)
(expand_moves mv (Lbranch s :: k))
| ebs_makelong: forall src1 src2 dst mv s k,
- wf_moves mv ->
+ wf_moves mv -> Archi.splitlong = true ->
expand_block_shape (BSmakelong src1 src2 dst mv s)
(Iop Omakelong (src1 :: src2 :: nil) dst s)
(expand_moves mv (Lbranch s :: k))
| ebs_lowlong: forall src dst mv s k,
- wf_moves mv ->
+ wf_moves mv -> Archi.splitlong = true ->
expand_block_shape (BSlowlong src dst mv s)
(Iop Olowlong (src :: nil) dst s)
(expand_moves mv (Lbranch s :: k))
| ebs_highlong: forall src dst mv s k,
- wf_moves mv ->
+ wf_moves mv -> Archi.splitlong = true ->
expand_block_shape (BShighlong src dst mv s)
(Iop Ohighlong (src :: nil) dst s)
(expand_moves mv (Lbranch s :: k))
@@ -97,7 +97,7 @@ Inductive expand_block_shape: block_shape -> RTL.instruction -> LTL.bblock -> Pr
(Lload chunk addr args' dst' :: expand_moves mv2 (Lbranch s :: k)))
| ebs_load2: forall addr addr2 args dst mv1 args1' dst1' mv2 args2' dst2' mv3 s k,
wf_moves mv1 -> wf_moves mv2 -> wf_moves mv3 ->
- offset_addressing addr (Int.repr 4) = Some addr2 ->
+ Archi.splitlong = true -> offset_addressing addr 4 = Some addr2 ->
expand_block_shape (BSload2 addr addr2 args dst mv1 args1' dst1' mv2 args2' dst2' mv3 s)
(Iload Mint64 addr args dst s)
(expand_moves mv1
@@ -107,6 +107,7 @@ Inductive expand_block_shape: block_shape -> RTL.instruction -> LTL.bblock -> Pr
expand_moves mv3 (Lbranch s :: k))))
| ebs_load2_1: forall addr args dst mv1 args' dst' mv2 s k,
wf_moves mv1 -> wf_moves mv2 ->
+ Archi.splitlong = true ->
expand_block_shape (BSload2_1 addr args dst mv1 args' dst' mv2 s)
(Iload Mint64 addr args dst s)
(expand_moves mv1
@@ -114,7 +115,7 @@ Inductive expand_block_shape: block_shape -> RTL.instruction -> LTL.bblock -> Pr
expand_moves mv2 (Lbranch s :: k)))
| ebs_load2_2: forall addr addr2 args dst mv1 args' dst' mv2 s k,
wf_moves mv1 -> wf_moves mv2 ->
- offset_addressing addr (Int.repr 4) = Some addr2 ->
+ Archi.splitlong = true -> offset_addressing addr 4 = Some addr2 ->
expand_block_shape (BSload2_2 addr addr2 args dst mv1 args' dst' mv2 s)
(Iload Mint64 addr args dst s)
(expand_moves mv1
@@ -133,7 +134,7 @@ Inductive expand_block_shape: block_shape -> RTL.instruction -> LTL.bblock -> Pr
(Lstore chunk addr args' src' :: Lbranch s :: k))
| ebs_store2: forall addr addr2 args src mv1 args1' src1' mv2 args2' src2' s k,
wf_moves mv1 -> wf_moves mv2 ->
- offset_addressing addr (Int.repr 4) = Some addr2 ->
+ Archi.splitlong = true -> offset_addressing addr 4 = Some addr2 ->
expand_block_shape (BSstore2 addr addr2 args src mv1 args1' src1' mv2 args2' src2' s)
(Istore Mint64 addr args src s)
(expand_moves mv1
@@ -196,6 +197,13 @@ Ltac MonadInv :=
idtac
end.
+Remark expand_moves_cons:
+ forall m accu b,
+ expand_moves (rev (m :: accu)) b = expand_moves (rev accu) (expand_move m :: b).
+Proof.
+ unfold expand_moves; intros. simpl. rewrite map_app. rewrite app_ass. auto.
+Qed.
+
Lemma extract_moves_sound:
forall b mv b',
extract_moves nil b = (mv, b') ->
@@ -205,39 +213,27 @@ Proof.
forall accu b,
wf_moves accu ->
wf_moves (List.rev accu) /\ expand_moves (List.rev accu) b = expand_moves (List.rev accu) b).
- intros; split; auto.
- red; intros. apply H. rewrite <- in_rev in H0; auto.
+ { intros; split; auto. unfold wf_moves in *; rewrite Forall_forall in *.
+ intros. apply H. rewrite <- in_rev in H0; auto. }
assert (IND: forall b accu mv b',
extract_moves accu b = (mv, b') ->
wf_moves accu ->
wf_moves mv /\ expand_moves (List.rev accu) b = expand_moves mv b').
- induction b; simpl; intros.
- inv H. auto.
- destruct a; try (inv H; apply BASE; auto; fail).
- destruct (is_move_operation op args) as [arg|] eqn:E.
+ { induction b; simpl; intros.
+ - inv H. auto.
+ - destruct a; try (inv H; apply BASE; auto; fail).
+ + destruct (is_move_operation op args) as [arg|] eqn:E.
exploit is_move_operation_correct; eauto. intros [A B]; subst.
(* reg-reg move *)
- exploit IHb; eauto.
- red; intros. destruct H1; auto. subst sd; exact I.
- intros [P Q].
- split; auto. rewrite <- Q. simpl. unfold expand_moves. rewrite map_app.
- rewrite app_ass. simpl. auto.
+ exploit IHb; eauto. constructor; auto. exact I. rewrite expand_moves_cons; auto.
inv H; apply BASE; auto.
- (* stack-reg move *)
- exploit IHb; eauto.
- red; intros. destruct H1; auto. subst sd; exact I.
- intros [P Q].
- split; auto. rewrite <- Q. simpl. unfold expand_moves. rewrite map_app.
- rewrite app_ass. simpl. auto.
- (* reg-stack move *)
- exploit IHb; eauto.
- red; intros. destruct H1; auto. subst sd; exact I.
- intros [P Q].
- split; auto. rewrite <- Q. simpl. unfold expand_moves. rewrite map_app.
- rewrite app_ass. simpl. auto.
-
- intros. exploit IND; eauto. red; intros. elim H0.
+ + (* stack-reg move *)
+ exploit IHb; eauto. constructor; auto. exact I. rewrite expand_moves_cons; auto.
+ + (* reg-stack move *)
+ exploit IHb; eauto. constructor; auto. exact I. rewrite expand_moves_cons; auto.
+ }
+ intros. exploit IND; eauto. constructor.
Qed.
Lemma check_succ_sound:
@@ -253,7 +249,7 @@ Ltac UseParsingLemmas :=
| [ H: extract_moves nil _ = (_, _) |- _ ] =>
destruct (extract_moves_sound _ _ _ H); clear H; subst; UseParsingLemmas
| [ H: check_succ _ _ = true |- _ ] =>
- try discriminate;
+ try (discriminate H);
destruct (check_succ_sound _ _ H); clear H; subst; UseParsingLemmas
| _ => idtac
end.
@@ -262,59 +258,64 @@ Lemma pair_instr_block_sound:
forall i b bsh,
pair_instr_block i b = Some bsh -> expand_block_shape bsh i b.
Proof.
+ assert (OP: forall op args res s b bsh,
+ pair_Iop_block op args res s b = Some bsh -> expand_block_shape bsh (Iop op args res s) b).
+ {
+ unfold pair_Iop_block; intros. MonadInv. destruct b0.
+ MonadInv; UseParsingLemmas.
+ destruct i; MonadInv; UseParsingLemmas.
+ eapply ebs_op; eauto.
+ inv H0. eapply ebs_op_dead; eauto. }
+
intros; destruct i; simpl in H; MonadInv; UseParsingLemmas.
-(* nop *)
+- (* nop *)
econstructor; eauto.
-(* op *)
+- (* op *)
destruct (classify_operation o l).
- (* move *)
++ (* move *)
MonadInv; UseParsingLemmas. econstructor; eauto.
- (* makelong *)
++ (* makelong *)
+ destruct Archi.splitlong eqn:SL; eauto.
MonadInv; UseParsingLemmas. econstructor; eauto.
- (* lowlong *)
++ (* lowlong *)
+ destruct Archi.splitlong eqn:SL; eauto.
MonadInv; UseParsingLemmas. econstructor; eauto.
- (* highlong *)
++ (* highlong *)
+ destruct Archi.splitlong eqn:SL; eauto.
MonadInv; UseParsingLemmas. econstructor; eauto.
- (* other ops *)
- MonadInv. destruct b0.
- MonadInv; UseParsingLemmas.
- destruct i; MonadInv; UseParsingLemmas.
- eapply ebs_op; eauto.
- inv H0. eapply ebs_op_dead; eauto.
-(* load *)
- destruct b0.
- MonadInv; UseParsingLemmas.
- destruct i; MonadInv; UseParsingLemmas.
- destruct (chunk_eq m Mint64).
- MonadInv; UseParsingLemmas.
- destruct b; MonadInv; UseParsingLemmas. destruct i; MonadInv; UseParsingLemmas.
- eapply ebs_load2; eauto.
- destruct (eq_addressing a addr).
- MonadInv. inv H2. eapply ebs_load2_1; eauto.
- MonadInv. inv H2. eapply ebs_load2_2; eauto.
- MonadInv; UseParsingLemmas. eapply ebs_load; eauto.
++ (* other ops *)
+ eauto.
+- (* load *)
+ destruct b0 as [ | [] b0]; MonadInv; UseParsingLemmas.
+ destruct (chunk_eq m Mint64 && Archi.splitlong) eqn:A; MonadInv; UseParsingLemmas.
+ destruct b as [ | [] b]; MonadInv; UseParsingLemmas.
+ InvBooleans. subst m. eapply ebs_load2; eauto.
+ InvBooleans. subst m.
+ destruct (eq_addressing a addr).
+ inv H; inv H2. eapply ebs_load2_1; eauto.
+ destruct (option_eq eq_addressing (offset_addressing a 4) (Some addr)).
+ inv H; inv H2. eapply ebs_load2_2; eauto.
+ discriminate.
+ eapply ebs_load; eauto.
inv H. eapply ebs_load_dead; eauto.
-(* store *)
+- (* store *)
destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas.
- destruct (chunk_eq m Mint64).
- MonadInv; UseParsingLemmas.
- destruct b; MonadInv. destruct i; MonadInv; UseParsingLemmas.
- eapply ebs_store2; eauto.
- MonadInv; UseParsingLemmas.
+ destruct (chunk_eq m Mint64 && Archi.splitlong) eqn:A; MonadInv; UseParsingLemmas.
+ destruct b as [ | [] b]; MonadInv; UseParsingLemmas.
+ InvBooleans. subst m. eapply ebs_store2; eauto.
eapply ebs_store; eauto.
-(* call *)
- destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas. econstructor; eauto.
-(* tailcall *)
- destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas. econstructor; eauto.
-(* builtin *)
- destruct b1; MonadInv. destruct i; MonadInv; UseParsingLemmas.
- econstructor; eauto.
-(* cond *)
- destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas. econstructor; eauto.
-(* jumptable *)
- destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas. econstructor; eauto.
-(* return *)
- destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas. econstructor; eauto.
+- (* call *)
+ destruct b0 as [|[] ]; MonadInv; UseParsingLemmas. econstructor; eauto.
+- (* tailcall *)
+ destruct b0 as [|[] ]; MonadInv; UseParsingLemmas. econstructor; eauto.
+- (* builtin *)
+ destruct b1 as [|[] ]; MonadInv; UseParsingLemmas. econstructor; eauto.
+- (* cond *)
+ destruct b0 as [|[]]; MonadInv; UseParsingLemmas. econstructor; eauto.
+- (* jumptable *)
+ destruct b0 as [|[]]; MonadInv; UseParsingLemmas. econstructor; eauto.
+- (* return *)
+ destruct b0 as [|[]]; MonadInv; UseParsingLemmas. econstructor; eauto.
Qed.
Lemma matching_instr_block:
@@ -419,16 +420,18 @@ Proof.
- eapply add_equation_satisf; eauto.
- eapply add_equation_satisf. eapply add_equation_satisf. eauto.
- congruence.
+- congruence.
Qed.
Lemma val_longofwords_eq:
forall v,
- Val.has_type v Tlong ->
+ Val.has_type v Tlong -> Archi.splitlong = true ->
Val.longofwords (Val.hiword v) (Val.loword v) = v.
Proof.
intros. red in H. destruct v; try contradiction.
- reflexivity.
- simpl. rewrite Int64.ofwords_recompose. auto.
+- reflexivity.
+- simpl. rewrite Int64.ofwords_recompose. auto.
+- rewrite Archi.splitlong_ptr32 in H by auto. congruence.
Qed.
Lemma add_equations_args_lessdef:
@@ -443,12 +446,13 @@ Proof.
- destruct H1. constructor; auto.
eapply add_equation_lessdef with (q := Eq Full r1 l1). eapply add_equations_args_satisf; eauto.
- destruct H1. constructor; auto.
- rewrite <- (val_longofwords_eq (rs#r1)); auto. apply Val.longofwords_lessdef.
+ rewrite <- (val_longofwords_eq (rs#r1)) by auto. apply Val.longofwords_lessdef.
eapply add_equation_lessdef with (q := Eq High r1 l1).
eapply add_equation_satisf. eapply add_equations_args_satisf; eauto.
eapply add_equation_lessdef with (q := Eq Low r1 l2).
eapply add_equations_args_satisf; eauto.
- discriminate.
+- discriminate.
Qed.
Lemma add_equation_ros_satisf:
@@ -694,6 +698,14 @@ Proof.
eapply reg_unconstrained_sound; eauto.
Qed.
+Remark in_elements_between_1:
+ forall r1 s q,
+ EqSet.In q (EqSet.elements_between (select_reg_l r1) (select_reg_h r1) s) <-> EqSet.In q s /\ ereg q = r1.
+Proof.
+ intros. rewrite EqSet.elements_between_iff, select_reg_charact. tauto.
+ exact (select_reg_l_monotone r1). exact (select_reg_h_monotone r1).
+Qed.
+
Lemma in_subst_reg:
forall r1 r2 q (e: eqs),
EqSet.In q e ->
@@ -702,14 +714,9 @@ Lemma in_subst_reg:
Proof.
intros r1 r2 q e0 IN0. unfold subst_reg.
set (f := fun (q: EqSet.elt) e => add_equation (Eq (ekind q) r2 (eloc q)) (remove_equation q e)).
+ generalize (in_elements_between_1 r1 e0).
set (elt := EqSet.elements_between (select_reg_l r1) (select_reg_h r1) e0).
- assert (IN_ELT: forall q, EqSet.In q elt <-> EqSet.In q e0 /\ ereg q = r1).
- {
- intros. unfold elt. rewrite EqSet.elements_between_iff.
- rewrite select_reg_charact. tauto.
- exact (select_reg_l_monotone r1).
- exact (select_reg_h_monotone r1).
- }
+ intros IN_ELT.
set (P := fun e1 e2 =>
EqSet.In q e1 ->
EqSet.In (Eq (ekind q) r2 (eloc q)) e2).
@@ -730,9 +737,7 @@ Proof.
{
apply ESP.fold_rec; unfold Q; intros.
- auto.
- - simpl. red in H2. rewrite H2 in H4.
- rewrite ESF.add_iff. rewrite ESF.remove_iff.
- right. split. apply H3. tauto. tauto.
+ - simpl. red in H2. rewrite H2 in H4. ESD.fsetdec.
}
destruct (ESP.In_dec q elt).
left. split. apply IN_ELT. auto. apply H. auto.
@@ -761,14 +766,9 @@ Proof.
if IndexedEqKind.eq (ekind q) k1
then add_equation (Eq k2 r2 (eloc q)) (remove_equation q e)
else e).
+ generalize (in_elements_between_1 r1 e0).
set (elt := EqSet.elements_between (select_reg_l r1) (select_reg_h r1) e0).
- assert (IN_ELT: forall q, EqSet.In q elt <-> EqSet.In q e0 /\ ereg q = r1).
- {
- intros. unfold elt. rewrite EqSet.elements_between_iff.
- rewrite select_reg_charact. tauto.
- exact (select_reg_l_monotone r1).
- exact (select_reg_h_monotone r1).
- }
+ intros IN_ELT.
set (P := fun e1 e2 =>
EqSet.In q e1 -> ekind q = k1 ->
EqSet.In (Eq k2 r2 (eloc q)) e2).
@@ -796,7 +796,7 @@ Proof.
destruct (IndexedEqKind.eq (ekind x) k1).
simpl. rewrite ESF.add_iff. rewrite ESF.remove_iff.
right. split. apply H3. tauto. intuition congruence.
- apply H3. intuition.
+ apply H3. intuition auto.
}
destruct (ESP.In_dec q elt).
destruct (IndexedEqKind.eq (ekind q) k1).
@@ -863,68 +863,65 @@ Module ESF2 := FSetFacts.Facts(EqSet2).
Module ESP2 := FSetProperties.Properties(EqSet2).
Module ESD2 := FSetDecide.Decide(EqSet2).
+Lemma partial_fold_ind:
+ forall (A: Type) (P: EqSet2.t -> A -> Prop) f init final s,
+ EqSet2.fold
+ (fun q opte =>
+ match opte with
+ | None => None
+ | Some e => f q e
+ end)
+ s (Some init) = Some final ->
+ (forall s', EqSet2.Empty s' -> P s' init) ->
+ (forall x a' a'' s' s'',
+ EqSet2.In x s -> ~EqSet2.In x s' -> ESP2.Add x s' s'' ->
+ f x a' = Some a'' -> P s' a' -> P s'' a'') ->
+ P s final.
+Proof.
+ intros.
+ set (g := fun q opte => match opte with Some e => f q e | None => None end) in *.
+ set (Q := fun s1 opte => match opte with None => True | Some e => P s1 e end).
+ change (Q s (Some final)).
+ rewrite <- H. apply ESP2.fold_rec; unfold Q, g; intros.
+ - auto.
+ - destruct a as [e|]; auto. destruct (f x e) as [e'|] eqn:F; auto. eapply H1; eauto.
+Qed.
+
Lemma in_subst_loc:
forall l1 l2 q (e e': eqs),
EqSet.In q e ->
subst_loc l1 l2 e = Some e' ->
(eloc q = l1 /\ EqSet.In (Eq (ekind q) (ereg q) l2) e') \/ (Loc.diff l1 (eloc q) /\ EqSet.In q e').
Proof.
- intros l1 l2 q e0 e0'.
- unfold subst_loc.
- set (f := fun (q0 : EqSet2.elt) (opte : option eqs) =>
- match opte with
- | Some e =>
- if Loc.eq l1 (eloc q0)
- then
- Some
- (add_equation {| ekind := ekind q0; ereg := ereg q0; eloc := l2 |}
- (remove_equation q0 e))
- else None
- | None => None
- end).
- set (elt := EqSet2.elements_between (select_loc_l l1) (select_loc_h l1) (eqs2 e0)).
- intros IN SUBST.
- set (P := fun e1 (opte: option eqs) =>
- match opte with
- | None => True
- | Some e2 =>
- EqSet2.In q e1 ->
- eloc q = l1 /\ EqSet.In (Eq (ekind q) (ereg q) l2) e2
- end).
- assert (P elt (EqSet2.fold f elt (Some e0))).
- {
- apply ESP2.fold_rec; unfold P; intros.
- - ESD2.fsetdec.
- - destruct a as [e2|]; simpl; auto.
- destruct (Loc.eq l1 (eloc x)); auto.
- unfold add_equation, remove_equation; simpl.
- red in H1. rewrite H1. intros [A|A].
- + subst x. split. auto. ESD.fsetdec.
- + exploit H2; eauto. intros [B C]. split. auto.
- rewrite ESF.add_iff. rewrite ESF.remove_iff.
- destruct (OrderedEquation.eq_dec x {| ekind := ekind q; ereg := ereg q; eloc := l2 |}).
- left. rewrite e1; auto.
- right; auto.
- }
- set (Q := fun e1 (opte: option eqs) =>
- match opte with
- | None => True
- | Some e2 => ~EqSet2.In q e1 -> EqSet.In q e2
- end).
- assert (Q elt (EqSet2.fold f elt (Some e0))).
- {
- apply ESP2.fold_rec; unfold Q; intros.
- - auto.
- - destruct a as [e2|]; simpl; auto.
- destruct (Loc.eq l1 (eloc x)); auto.
- red in H2. rewrite H2; intros.
- unfold add_equation, remove_equation; simpl.
- rewrite ESF.add_iff. rewrite ESF.remove_iff.
- right; split. apply H3. tauto. tauto.
+ unfold subst_loc; intros l1 l2 q e0 e0' IN SUBST.
+ set (elt := EqSet2.elements_between (select_loc_l l1) (select_loc_h l1) (eqs2 e0)) in *.
+ set (f := fun q0 e =>
+ if Loc.eq l1 (eloc q0) then
+ Some (add_equation
+ {| ekind := ekind q0; ereg := ereg q0; eloc := l2 |}
+ (remove_equation q0 e))
+ else None).
+ set (P := fun e1 e2 => EqSet2.In q e1 -> eloc q = l1 /\ EqSet.In (Eq (ekind q) (ereg q) l2) e2).
+ assert (A: P elt e0').
+ { eapply partial_fold_ind with (f := f) (s := elt) (final := e0'). eexact SUBST.
+ - unfold P; intros. ESD2.fsetdec.
+ - unfold P, f; intros. destruct (Loc.eq l1 (eloc x)); inversion H2; subst a''; clear H2.
+ simpl. rewrite ESF.add_iff, ESF.remove_iff.
+ apply H1 in H4; destruct H4.
+ subst x; rewrite e; auto.
+ apply H3 in H2; destruct H2. split. congruence.
+ destruct (OrderedEquation.eq_dec x {| ekind := ekind q; ereg := ereg q; eloc := l2 |}); auto.
+ subst x; auto.
}
- rewrite SUBST in H; rewrite SUBST in H0; simpl in *.
+ set (Q := fun e1 e2 => ~EqSet2.In q e1 -> EqSet.In q e2).
+ assert (B: Q elt e0').
+ { eapply partial_fold_ind with (f := f) (s := elt) (final := e0'). eexact SUBST.
+ - unfold Q; intros. auto.
+ - unfold Q, f; intros. destruct (Loc.eq l1 (eloc x)); inversion H2; subst a''; clear H2.
+ simpl. rewrite ESF.add_iff, ESF.remove_iff.
+ red in H1. rewrite H1 in H4. intuition auto. }
destruct (ESP2.In_dec q elt).
- left. apply H; auto.
+ left. apply A; auto.
right. split; auto.
rewrite <- select_loc_charact.
destruct (select_loc_l l1 q) eqn: LL; auto.
@@ -1287,14 +1284,15 @@ Qed.
Lemma loadv_int64_split:
forall m a v,
- Mem.loadv Mint64 m a = Some v ->
+ Mem.loadv Mint64 m a = Some v -> Archi.splitlong = true ->
exists v1 v2,
Mem.loadv Mint32 m a = Some (if Archi.big_endian then v1 else v2)
- /\ Mem.loadv Mint32 m (Val.add a (Vint (Int.repr 4))) = Some (if Archi.big_endian then v2 else v1)
+ /\ Mem.loadv Mint32 m (Val.add a (Vint (Int.repr 4))) = Some (if Archi.big_endian then v2 else v1)
/\ Val.lessdef (Val.hiword v) v1
/\ Val.lessdef (Val.loword v) v2.
Proof.
- intros. exploit Mem.loadv_int64_split; eauto. intros (v1 & v2 & A & B & C).
+ intros. apply Archi.splitlong_ptr32 in H0.
+ exploit Mem.loadv_int64_split; eauto. intros (v1 & v2 & A & B & C).
exists v1, v2. split; auto. split; auto.
inv C; auto. destruct v1, v2; simpl; auto.
rewrite Int64.hi_ofwords, Int64.lo_ofwords; auto.
@@ -1328,6 +1326,7 @@ Proof.
exists (Val.longofwords (ls x0) (ls x1)); split; auto with barg.
rewrite <- (val_longofwords_eq rs#x). apply Val.longofwords_lessdef; auto.
rewrite <- e0; apply WT.
+ assumption.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
@@ -1639,24 +1638,23 @@ Opaque destroyed_by_op.
(* base *)
- unfold expand_moves; simpl. inv H. exists ls; split. apply star_refl. auto.
(* step *)
-- destruct a as [src dst]. unfold expand_moves. simpl.
- destruct (track_moves env mv e) as [e1|] eqn:?; MonadInv.
- assert (wf_moves mv). red; intros. apply H0; auto with coqlib.
+- assert (wf_moves mv) by (inv H0; auto).
+ destruct a as (src, dst); unfold expand_moves; simpl; MonadInv.
destruct src as [rsrc | ssrc]; destruct dst as [rdst | sdst].
- (* reg-reg *)
-+ exploit IHmv; eauto. eapply subst_loc_undef_satisf; eauto.
+* (* reg-reg *)
+ exploit IHmv; eauto. eapply subst_loc_undef_satisf; eauto.
intros [ls' [A B]]. exists ls'; split; auto. eapply star_left; eauto.
econstructor. simpl. eauto. auto. auto.
- (* reg->stack *)
-+ exploit IHmv; eauto. eapply subst_loc_undef_satisf; eauto.
+* (* reg->stack *)
+ exploit IHmv; eauto. eapply subst_loc_undef_satisf; eauto.
intros [ls' [A B]]. exists ls'; split; auto. eapply star_left; eauto.
econstructor. simpl. eauto. auto.
- (* stack->reg *)
-+ simpl in Heqb. exploit IHmv; eauto. eapply subst_loc_undef_satisf; eauto.
+* (* stack->reg *)
+ simpl in Heqb. exploit IHmv; eauto. eapply subst_loc_undef_satisf; eauto.
intros [ls' [A B]]. exists ls'; split; auto. eapply star_left; eauto.
econstructor. auto. auto.
- (* stack->stack *)
-+ exploit H0; auto with coqlib. unfold wf_move. tauto.
+* (* stack->stack *)
+ inv H0. simpl in H6. contradiction.
Qed.
(** The simulation relation *)
@@ -1730,7 +1728,7 @@ Proof.
constructor. congruence.
econstructor; eauto.
unfold proj_sig_res in *. rewrite H0; auto.
- intros. unfold loc_result in H; rewrite H0 in H; eauto.
+ intros. rewrite (loc_result_exten sg' sg) in H by auto. eauto.
Qed.
Ltac UseShape :=
@@ -1742,22 +1740,17 @@ Ltac UseShape :=
Remark addressing_not_long:
forall env f addr args dst s r,
- wt_instr f env (Iload Mint64 addr args dst s) ->
+ wt_instr f env (Iload Mint64 addr args dst s) -> Archi.splitlong = true ->
In r args -> r <> dst.
Proof.
- intros.
- assert (forall ty, In ty (type_of_addressing addr) -> ty = Tint).
- { intros. destruct addr; simpl in H1; intuition. }
- inv H.
- assert (env r = Tint).
- { generalize args (type_of_addressing addr) H0 H1 H5.
- induction args0; simpl; intros.
- contradiction.
- destruct l. discriminate. inv H4.
- destruct H2. subst a. apply H3; auto with coqlib.
- eauto with coqlib.
- }
- red; intros; subst r. rewrite H in H8; discriminate.
+ intros. inv H.
+ assert (A: forall ty, In ty (type_of_addressing addr) -> ty = Tptr).
+ { intros. destruct addr; simpl in H; intuition. }
+ assert (B: In (env r) (type_of_addressing addr)).
+ { rewrite <- H5. apply in_map; auto. }
+ assert (C: env r = Tint).
+ { apply A in B. rewrite B. unfold Tptr. rewrite Archi.splitlong_ptr32 by auto. auto. }
+ red; intros; subst r. rewrite C in H8; discriminate.
Qed.
(** The proof of semantic preservation is a simulation argument of the
@@ -1771,7 +1764,7 @@ Proof.
induction 1; intros WT S1' MS; inv MS; try UseShape.
(* nop *)
- exploit exec_moves; eauto. intros [ls1 [X Y]].
+- exploit exec_moves; eauto. intros [ls1 [X Y]].
econstructor; split.
eapply plus_left. econstructor; eauto.
eapply star_right. eexact X. econstructor; eauto.
@@ -1901,8 +1894,11 @@ Proof.
eapply addressing_not_long; eauto.
}
exploit eval_addressing_lessdef. eexact LD3.
- eapply eval_offset_addressing; eauto. intros [a2' [F2 G2]].
- exploit Mem.loadv_extends. eauto. eexact LOAD2. eexact G2. intros (v2'' & LOAD2' & LD4).
+ eapply eval_offset_addressing; eauto; apply Archi.splitlong_ptr32; auto.
+ intros [a2' [F2 G2]].
+ assert (LOADX: exists v2'', Mem.loadv Mint32 m' a2' = Some v2'' /\ Val.lessdef v2' v2'').
+ { discriminate || (eapply Mem.loadv_extends; [eauto|eexact LOAD2|eexact G2]). }
+ destruct LOADX as (v2'' & LOAD2' & LD4).
set (ls4 := Locmap.set (R dst2') v2'' (undef_regs (destroyed_by_load Mint32 addr2) ls3)).
assert (SAT4: satisf (rs#dst <- v) ls4 e0).
{ eapply loc_unconstrained_satisf. eapply can_undef_satisf; eauto.
@@ -1966,8 +1962,11 @@ Proof.
assert (LD1: Val.lessdef_list rs##args (reglist ls1 args')).
{ eapply add_equations_lessdef; eauto. }
exploit eval_addressing_lessdef. eexact LD1.
- eapply eval_offset_addressing; eauto. intros [a1' [F1 G1]].
- exploit Mem.loadv_extends. eauto. eexact LOAD2. eexact G1. intros (v2'' & LOAD2' & LD2).
+ eapply eval_offset_addressing; eauto; apply Archi.splitlong_ptr32; auto.
+ intros [a1' [F1 G1]].
+ assert (LOADX: exists v2'', Mem.loadv Mint32 m' a1' = Some v2'' /\ Val.lessdef v2' v2'').
+ { discriminate || (eapply Mem.loadv_extends; [eauto|eexact LOAD2|eexact G1]). }
+ destruct LOADX as (v2'' & LOAD2' & LD2).
set (ls2 := Locmap.set (R dst') v2'' (undef_regs (destroyed_by_load Mint32 addr2) ls1)).
assert (SAT2: satisf (rs#dst <- v) ls2 e0).
{ eapply parallel_assignment_satisf; eauto.
@@ -2015,7 +2014,8 @@ Proof.
econstructor; eauto.
(* store 2 *)
-- exploit Mem.storev_int64_split; eauto.
+- assert (SF: Archi.ptr64 = false) by (apply Archi.splitlong_ptr32; auto).
+ exploit Mem.storev_int64_split; eauto.
replace (if Archi.big_endian then Val.hiword rs#src else Val.loword rs#src)
with (sel_val kind_first_word rs#src)
by (unfold kind_first_word; destruct Archi.big_endian; reflexivity).
@@ -2043,10 +2043,12 @@ Proof.
exploit eval_addressing_lessdef. eexact LD3. eauto. intros [a2' [F2 G2]].
assert (F2': eval_addressing tge sp addr (reglist ls3 args2') = Some a2').
rewrite <- F2. apply eval_addressing_preserved. exact symbols_preserved.
- exploit eval_offset_addressing. eauto. eexact F2'. intros F2''.
- exploit Mem.storev_extends. eexact EXT1. eexact STORE2.
- apply Val.add_lessdef. eexact G2. eauto. eauto.
- intros [m2' [STORE2' EXT2]].
+ exploit (eval_offset_addressing tge); eauto. intros F2''.
+ assert (STOREX: exists m2', Mem.storev Mint32 m1' (Val.add a2' (Vint (Int.repr 4))) (ls3 (R src2')) = Some m2' /\ Mem.extends m' m2').
+ { try discriminate;
+ (eapply Mem.storev_extends;
+ [eexact EXT1 | eexact STORE2 | apply Val.add_lessdef; [eexact G2|eauto] | eauto]). }
+ destruct STOREX as [m2' [STORE2' EXT2]].
econstructor; split.
eapply plus_left. econstructor; eauto.
eapply star_trans. eexact X.
@@ -2054,7 +2056,7 @@ Proof.
econstructor. eexact F1'. eexact STORE1'. instantiate (1 := ls2). auto.
eapply star_trans. eexact U.
eapply star_two.
- econstructor. eexact F2''. eexact STORE2'. eauto.
+ eapply exec_Lstore with (m' := m2'). eexact F2''. discriminate||exact STORE2'. eauto.
constructor. eauto. eauto. eauto. eauto. traceEq.
exploit satisf_successors; eauto. simpl; eauto.
eapply can_undef_satisf. eauto.
@@ -2229,7 +2231,7 @@ Proof.
econstructor; eauto.
simpl. destruct (loc_result (ef_sig ef)) eqn:RES; simpl.
rewrite Locmap.gss; auto.
- generalize (loc_result_pair (ef_sig ef)); rewrite RES; intros (A & B & C & D).
+ generalize (loc_result_pair (ef_sig ef)); rewrite RES; intros (A & B & C & D & E).
exploit external_call_well_typed; eauto. unfold proj_sig_res; rewrite B. intros WTRES'.
rewrite Locmap.gss. rewrite Locmap.gso by (red; auto). rewrite Locmap.gss.
rewrite val_longofwords_eq by auto. auto.
@@ -2275,8 +2277,8 @@ Lemma final_states_simulation:
match_states st1 st2 -> RTL.final_state st1 r -> LTL.final_state st2 r.
Proof.
intros. inv H0. inv H. inv STACKS.
- econstructor.
- unfold loc_result in RES; rewrite H in RES. simpl in RES. inv RES. auto.
+ econstructor. rewrite <- (loc_result_exten sg). inv RES; auto.
+ rewrite H; auto.
Qed.
Lemma wt_prog: wt_program prog.
diff --git a/backend/Asmgenproof0.v b/backend/Asmgenproof0.v
index 30d6990e..2c7994e9 100644
--- a/backend/Asmgenproof0.v
+++ b/backend/Asmgenproof0.v
@@ -81,7 +81,7 @@ Qed.
Hint Resolve preg_of_not_SP preg_of_not_PC: asmgen.
Lemma nextinstr_pc:
- forall rs, (nextinstr rs)#PC = Val.add rs#PC Vone.
+ forall rs, (nextinstr rs)#PC = Val.offset_ptr rs#PC Ptrofs.one.
Proof.
intros. apply Pregmap.gss.
Qed.
@@ -100,7 +100,7 @@ Qed.
Lemma nextinstr_set_preg:
forall rs m v,
- (nextinstr (rs#(preg_of m) <- v))#PC = Val.add rs#PC Vone.
+ (nextinstr (rs#(preg_of m) <- v))#PC = Val.offset_ptr rs#PC Ptrofs.one.
Proof.
intros. unfold nextinstr. rewrite Pregmap.gss.
rewrite Pregmap.gso. auto. apply sym_not_eq. apply preg_of_not_PC.
@@ -491,13 +491,12 @@ Qed.
Lemma code_tail_next_int:
forall fn ofs i c,
- list_length_z fn <= Int.max_unsigned ->
- code_tail (Int.unsigned ofs) fn (i :: c) ->
- code_tail (Int.unsigned (Int.add ofs Int.one)) fn c.
+ list_length_z fn <= Ptrofs.max_unsigned ->
+ code_tail (Ptrofs.unsigned ofs) fn (i :: c) ->
+ code_tail (Ptrofs.unsigned (Ptrofs.add ofs Ptrofs.one)) fn c.
Proof.
- intros. rewrite Int.add_unsigned.
- change (Int.unsigned Int.one) with 1.
- rewrite Int.unsigned_repr. apply code_tail_next with i; auto.
+ intros. rewrite Ptrofs.add_unsigned, Ptrofs.unsigned_one.
+ rewrite Ptrofs.unsigned_repr. apply code_tail_next with i; auto.
generalize (code_tail_bounds_2 _ _ _ _ H0). omega.
Qed.
@@ -513,7 +512,7 @@ Inductive transl_code_at_pc (ge: Mach.genv):
Genv.find_funct_ptr ge b = Some(Internal f) ->
transf_function f = Errors.OK tf ->
transl_code f c ep = OK tc ->
- code_tail (Int.unsigned ofs) (fn_code tf) tc ->
+ code_tail (Ptrofs.unsigned ofs) (fn_code tf) tc ->
transl_code_at_pc ge (Vptr b ofs) b f c ep tf tc.
(** Equivalence between [transl_code] and [transl_code']. *)
@@ -563,11 +562,11 @@ Qed.
>>
*)
-Definition return_address_offset (f: Mach.function) (c: Mach.code) (ofs: int) : Prop :=
+Definition return_address_offset (f: Mach.function) (c: Mach.code) (ofs: ptrofs) : Prop :=
forall tf tc,
transf_function f = OK tf ->
transl_code f c false = OK tc ->
- code_tail (Int.unsigned ofs) (fn_code tf) tc.
+ code_tail (Ptrofs.unsigned ofs) (fn_code tf) tc.
(** We now show that such an offset always exists if the Mach code [c]
is a suffix of [f.(fn_code)]. This holds because the translation
@@ -590,7 +589,7 @@ Hypothesis transf_function_inv:
forall f tf, transf_function f = OK tf ->
exists tc, exists ep, transl_code f (Mach.fn_code f) ep = OK tc /\ is_tail tc (fn_code tf).
Hypothesis transf_function_len:
- forall f tf, transf_function f = OK tf -> list_length_z (fn_code tf) <= Int.max_unsigned.
+ forall f tf, transf_function f = OK tf -> list_length_z (fn_code tf) <= Ptrofs.max_unsigned.
Lemma transl_code_tail:
forall f c1 c2, is_tail c1 c2 ->
@@ -618,11 +617,11 @@ Opaque transl_instr.
apply is_tail_trans with tc2; auto.
eapply transl_instr_tail; eauto. }
exploit is_tail_code_tail. eexact TL3. intros [ofs CT].
- exists (Int.repr ofs). red; intros.
- rewrite Int.unsigned_repr. congruence.
+ exists (Ptrofs.repr ofs). red; intros.
+ rewrite Ptrofs.unsigned_repr. congruence.
exploit code_tail_bounds_1; eauto.
apply transf_function_len in TF. omega.
-+ exists Int.zero; red; intros. congruence.
++ exists Ptrofs.zero; red; intros. congruence.
Qed.
End RETADDR_EXISTS.
@@ -651,8 +650,8 @@ Lemma return_address_offset_correct:
Proof.
intros. inv H. red in H0.
exploit code_tail_unique. eexact H12. eapply H0; eauto. intro.
- rewrite <- (Int.repr_unsigned ofs).
- rewrite <- (Int.repr_unsigned ofs').
+ rewrite <- (Ptrofs.repr_unsigned ofs).
+ rewrite <- (Ptrofs.repr_unsigned ofs').
congruence.
Qed.
@@ -758,12 +757,12 @@ Inductive exec_straight: code -> regset -> mem ->
| exec_straight_one:
forall i1 c rs1 m1 rs2 m2,
exec_instr ge fn i1 rs1 m1 = Next rs2 m2 ->
- rs2#PC = Val.add rs1#PC Vone ->
+ rs2#PC = Val.offset_ptr rs1#PC Ptrofs.one ->
exec_straight (i1 :: c) rs1 m1 c rs2 m2
| exec_straight_step:
forall i c rs1 m1 rs2 m2 c' rs3 m3,
exec_instr ge fn i rs1 m1 = Next rs2 m2 ->
- rs2#PC = Val.add rs1#PC Vone ->
+ rs2#PC = Val.offset_ptr rs1#PC Ptrofs.one ->
exec_straight c rs2 m2 c' rs3 m3 ->
exec_straight (i :: c) rs1 m1 c' rs3 m3.
@@ -782,8 +781,8 @@ Lemma exec_straight_two:
forall i1 i2 c rs1 m1 rs2 m2 rs3 m3,
exec_instr ge fn i1 rs1 m1 = Next rs2 m2 ->
exec_instr ge fn i2 rs2 m2 = Next rs3 m3 ->
- rs2#PC = Val.add rs1#PC Vone ->
- rs3#PC = Val.add rs2#PC Vone ->
+ rs2#PC = Val.offset_ptr rs1#PC Ptrofs.one ->
+ rs3#PC = Val.offset_ptr rs2#PC Ptrofs.one ->
exec_straight (i1 :: i2 :: c) rs1 m1 c rs3 m3.
Proof.
intros. apply exec_straight_step with rs2 m2; auto.
@@ -795,9 +794,9 @@ Lemma exec_straight_three:
exec_instr ge fn i1 rs1 m1 = Next rs2 m2 ->
exec_instr ge fn i2 rs2 m2 = Next rs3 m3 ->
exec_instr ge fn i3 rs3 m3 = Next rs4 m4 ->
- rs2#PC = Val.add rs1#PC Vone ->
- rs3#PC = Val.add rs2#PC Vone ->
- rs4#PC = Val.add rs3#PC Vone ->
+ rs2#PC = Val.offset_ptr rs1#PC Ptrofs.one ->
+ rs3#PC = Val.offset_ptr rs2#PC Ptrofs.one ->
+ rs4#PC = Val.offset_ptr rs3#PC Ptrofs.one ->
exec_straight (i1 :: i2 :: i3 :: c) rs1 m1 c rs4 m4.
Proof.
intros. apply exec_straight_step with rs2 m2; auto.
@@ -810,11 +809,11 @@ Qed.
Lemma exec_straight_steps_1:
forall c rs m c' rs' m',
exec_straight c rs m c' rs' m' ->
- list_length_z (fn_code fn) <= Int.max_unsigned ->
+ list_length_z (fn_code fn) <= Ptrofs.max_unsigned ->
forall b ofs,
rs#PC = Vptr b ofs ->
Genv.find_funct_ptr ge b = Some (Internal fn) ->
- code_tail (Int.unsigned ofs) (fn_code fn) c ->
+ code_tail (Ptrofs.unsigned ofs) (fn_code fn) c ->
plus step ge (State rs m) E0 (State rs' m').
Proof.
induction 1; intros.
@@ -824,7 +823,7 @@ Proof.
eapply plus_left'.
econstructor; eauto.
eapply find_instr_tail. eauto.
- apply IHexec_straight with b (Int.add ofs Int.one).
+ apply IHexec_straight with b (Ptrofs.add ofs Ptrofs.one).
auto. rewrite H0. rewrite H3. reflexivity.
auto.
apply code_tail_next_int with i; auto.
@@ -834,20 +833,20 @@ Qed.
Lemma exec_straight_steps_2:
forall c rs m c' rs' m',
exec_straight c rs m c' rs' m' ->
- list_length_z (fn_code fn) <= Int.max_unsigned ->
+ list_length_z (fn_code fn) <= Ptrofs.max_unsigned ->
forall b ofs,
rs#PC = Vptr b ofs ->
Genv.find_funct_ptr ge b = Some (Internal fn) ->
- code_tail (Int.unsigned ofs) (fn_code fn) c ->
+ code_tail (Ptrofs.unsigned ofs) (fn_code fn) c ->
exists ofs',
rs'#PC = Vptr b ofs'
- /\ code_tail (Int.unsigned ofs') (fn_code fn) c'.
+ /\ code_tail (Ptrofs.unsigned ofs') (fn_code fn) c'.
Proof.
induction 1; intros.
- exists (Int.add ofs Int.one). split.
+ exists (Ptrofs.add ofs Ptrofs.one). split.
rewrite H0. rewrite H2. auto.
apply code_tail_next_int with i1; auto.
- apply IHexec_straight with (Int.add ofs Int.one).
+ apply IHexec_straight with (Ptrofs.add ofs Ptrofs.one).
auto. rewrite H0. rewrite H3. reflexivity. auto.
apply code_tail_next_int with i; auto.
Qed.
@@ -871,10 +870,18 @@ Inductive match_stack: list Mach.stackframe -> Prop :=
match_stack (Stackframe fb sp ra c :: s).
Lemma parent_sp_def: forall s, match_stack s -> parent_sp s <> Vundef.
-Proof. induction 1; simpl. unfold Vzero; congruence. auto. Qed.
+Proof.
+ induction 1; simpl.
+ unfold Vnullptr; destruct Archi.ptr64; congruence.
+ auto.
+Qed.
Lemma parent_ra_def: forall s, match_stack s -> parent_ra s <> Vundef.
-Proof. induction 1; simpl. unfold Vzero; congruence. inv H0. congruence. Qed.
+Proof.
+ induction 1; simpl.
+ unfold Vnullptr; destruct Archi.ptr64; congruence.
+ inv H0. congruence.
+Qed.
Lemma lessdef_parent_sp:
forall s v,
diff --git a/backend/Bounds.v b/backend/Bounds.v
index 178ff6ed..8a383380 100644
--- a/backend/Bounds.v
+++ b/backend/Bounds.v
@@ -444,6 +444,7 @@ Definition size_callee_save_area (b: bounds) (ofs: Z) : Z :=
Lemma size_callee_save_area_rec_incr:
forall l ofs, ofs <= size_callee_save_area_rec l ofs.
Proof.
+Local Opaque mreg_type.
induction l as [ | r l]; intros; simpl.
- omega.
- eapply Zle_trans. 2: apply IHl.
@@ -472,45 +473,3 @@ Record frame_env : Type := mk_frame_env {
fe_stack_data: Z;
fe_used_callee_save: list mreg
}.
-
-(*
-Record frame_env_properties (b: bounds) (fe: frame_env) (fe_ofs_arg: Z) := mk_frame_env_properties {
- (** Separation property *)
- fe_separated:
- Intv.pairwise_disjoint (
- (fe.(fe_ofs_link), fe.(fe_ofs_link) + 4)
- :: (fe.(fe_ofs_retaddr), fe.(fe_ofs_retaddr) + 4)
- :: (fe.(fe_ofs_local), fe.(fe_ofs_local) + 4 * b.(bound_local))
- :: (fe_ofs_arg, fe_ofs_arg + 4 * b.(bound_outgoing))
- :: (fe.(fe_ofs_callee_save), size_callee_save_area b fe.(fe_ofs_callee_save))
- :: (fe.(fe_stack_data), fe.(fe_stack_data) + b.(bound_stack_data))
- :: nil);
- (** Inclusion properties *)
- fe_incl_link:
- Intv.incl (fe.(fe_ofs_link), fe.(fe_ofs_link) + 4) (0, fe.(fe_size));
- fe_incl_retaddr:
- Intv.incl (fe.(fe_ofs_retaddr), fe.(fe_ofs_retaddr) + 4) (0, fe.(fe_size));
- fe_incl_local:
- Intv.incl (fe.(fe_ofs_local), fe.(fe_ofs_local) + 4 * b.(bound_local)) (0, fe.(fe_size));
- fe_incl_outgoing:
- Intv.incl (fe_ofs_arg, fe_ofs_arg + 4 * b.(bound_outgoing)) (0, fe.(fe_size));
- fe_incl_callee_save:
- Intv.incl (fe.(fe_ofs_callee_save), size_callee_save_area b fe.(fe_ofs_callee_save)) (0, fe.(fe_size));
- fe_incl_stack_data:
- Intv.incl (fe.(fe_stack_data), fe.(fe_stack_data) + b.(bound_stack_data)) (0, fe.(fe_size));
- (** Alignment properties *)
- fe_align_link:
- (4 | fe.(fe_ofs_link));
- fe_align_retaddr:
- (4 | fe.(fe_ofs_retaddr));
- fe_align_local:
- (8 | fe.(fe_ofs_local));
- fe_align_stack_data:
- (8 | fe.(fe_stack_data));
- fe_align_size:
- (4 | fe.(fe_size));
- (** Callee-save registers *)
- fe_used_callee_save_eq:
- fe.(fe_used_callee_save) = b.(used_callee_save)
-}.
-*)
diff --git a/backend/CSE.v b/backend/CSE.v
index d6b89557..4fa1bd6c 100644
--- a/backend/CSE.v
+++ b/backend/CSE.v
@@ -327,14 +327,14 @@ Definition kill_loads_after_storebytes
Definition shift_memcpy_eq (src sz delta: Z) (e: equation) :=
match e with
| Eq l strict (Load chunk (Ainstack i) _) =>
- let i := Int.unsigned i in
+ let i := Ptrofs.unsigned i in
let j := i + delta in
if zle src i
&& zle (i + size_chunk chunk) (src + sz)
&& zeq (Zmod delta (align_chunk chunk)) 0
&& zle 0 j
- && zle j Int.max_unsigned
- then Some(Eq l strict (Load chunk (Ainstack (Int.repr j)) nil))
+ && zle j Ptrofs.max_unsigned
+ then Some(Eq l strict (Load chunk (Ainstack (Ptrofs.repr j)) nil))
else None
| _ => None
end.
@@ -353,8 +353,8 @@ Definition add_memcpy (n1 n2: numbering) (asrc adst: aptr) (sz: Z) :=
match asrc, adst with
| Stk src, Stk dst =>
{| num_next := n2.(num_next);
- num_eqs := add_memcpy_eqs (Int.unsigned src) sz
- (Int.unsigned dst - Int.unsigned src)
+ num_eqs := add_memcpy_eqs (Ptrofs.unsigned src) sz
+ (Ptrofs.unsigned dst - Ptrofs.unsigned src)
n1.(num_eqs) n2.(num_eqs);
num_reg := n2.(num_reg);
num_val := n2.(num_val) |}
diff --git a/backend/CSEproof.v b/backend/CSEproof.v
index 2c144249..bf152e82 100644
--- a/backend/CSEproof.v
+++ b/backend/CSEproof.v
@@ -462,14 +462,14 @@ Qed.
Lemma kill_loads_after_store_holds:
forall valu ge sp rs m n addr args a chunk v m' bc approx ae am,
- numbering_holds valu ge (Vptr sp Int.zero) rs m n ->
- eval_addressing ge (Vptr sp Int.zero) addr rs##args = Some a ->
+ numbering_holds valu ge (Vptr sp Ptrofs.zero) rs m n ->
+ eval_addressing ge (Vptr sp Ptrofs.zero) addr rs##args = Some a ->
Mem.storev chunk m a v = Some m' ->
genv_match bc ge ->
bc sp = BCstack ->
ematch bc rs ae ->
approx = VA.State ae am ->
- numbering_holds valu ge (Vptr sp Int.zero) rs m'
+ numbering_holds valu ge (Vptr sp Ptrofs.zero) rs m'
(kill_loads_after_store approx n chunk addr args).
Proof.
intros. apply kill_equations_hold with m; auto.
@@ -493,11 +493,15 @@ Lemma store_normalized_range_sound:
vmatch bc v (store_normalized_range chunk) ->
Val.lessdef (Val.load_result chunk v) v.
Proof.
- intros. destruct chunk; simpl in *; destruct v; auto.
+ intros. unfold Val.load_result; remember Archi.ptr64 as ptr64.
+ destruct chunk; simpl in *; destruct v; auto.
- inv H. rewrite is_sgn_sign_ext in H4 by omega. rewrite H4; auto.
- inv H. rewrite is_uns_zero_ext in H4 by omega. rewrite H4; auto.
- inv H. rewrite is_sgn_sign_ext in H4 by omega. rewrite H4; auto.
- inv H. rewrite is_uns_zero_ext in H4 by omega. rewrite H4; auto.
+- destruct ptr64; auto.
+- destruct ptr64; auto.
+- destruct ptr64; auto.
Qed.
Lemma add_store_result_hold:
@@ -533,15 +537,15 @@ Qed.
Lemma kill_loads_after_storebytes_holds:
forall valu ge sp rs m n dst b ofs bytes m' bc approx ae am sz,
- numbering_holds valu ge (Vptr sp Int.zero) rs m n ->
+ numbering_holds valu ge (Vptr sp Ptrofs.zero) rs m n ->
pmatch bc b ofs dst ->
- Mem.storebytes m b (Int.unsigned ofs) bytes = Some m' ->
+ Mem.storebytes m b (Ptrofs.unsigned ofs) bytes = Some m' ->
genv_match bc ge ->
bc sp = BCstack ->
ematch bc rs ae ->
approx = VA.State ae am ->
length bytes = nat_of_Z sz -> sz >= 0 ->
- numbering_holds valu ge (Vptr sp Int.zero) rs m'
+ numbering_holds valu ge (Vptr sp Ptrofs.zero) rs m'
(kill_loads_after_storebytes approx n dst sz).
Proof.
intros. apply kill_equations_hold with m; auto.
@@ -619,10 +623,11 @@ Lemma shift_memcpy_eq_wf:
Proof with (try discriminate).
unfold shift_memcpy_eq; intros.
destruct e. destruct r... destruct a...
- destruct (zle src (Int.unsigned i) &&
- zle (Int.unsigned i + size_chunk m) (src + sz) &&
- zeq (delta mod align_chunk m) 0 && zle 0 (Int.unsigned i + delta) &&
- zle (Int.unsigned i + delta) Int.max_unsigned)...
+ try (rename i into ofs).
+ destruct (zle src (Ptrofs.unsigned ofs) &&
+ zle (Ptrofs.unsigned ofs + size_chunk m) (src + sz) &&
+ zeq (delta mod align_chunk m) 0 && zle 0 (Ptrofs.unsigned ofs + delta) &&
+ zle (Ptrofs.unsigned ofs + delta) Ptrofs.max_unsigned)...
inv H. destruct H0. split. auto. red; simpl; tauto.
Qed.
@@ -631,35 +636,40 @@ Lemma shift_memcpy_eq_holds:
shift_memcpy_eq src sz (dst - src) e = Some e' ->
Mem.loadbytes m sp src sz = Some bytes ->
Mem.storebytes m sp dst bytes = Some m' ->
- equation_holds valu ge (Vptr sp Int.zero) m e ->
- equation_holds valu ge (Vptr sp Int.zero) m' e'.
+ equation_holds valu ge (Vptr sp Ptrofs.zero) m e ->
+ equation_holds valu ge (Vptr sp Ptrofs.zero) m' e'.
Proof with (try discriminate).
intros. set (delta := dst - src) in *. unfold shift_memcpy_eq in H.
destruct e as [l strict rhs] eqn:E.
destruct rhs as [op vl | chunk addr vl]...
destruct addr...
- set (i1 := Int.unsigned i) in *. set (j := i1 + delta) in *.
+ try (rename i into ofs).
+ set (i1 := Ptrofs.unsigned ofs) in *. set (j := i1 + delta) in *.
destruct (zle src i1)...
destruct (zle (i1 + size_chunk chunk) (src + sz))...
destruct (zeq (delta mod align_chunk chunk) 0)...
destruct (zle 0 j)...
- destruct (zle j Int.max_unsigned)...
+ destruct (zle j Ptrofs.max_unsigned)...
simpl in H; inv H.
assert (LD: forall v,
- Mem.loadv chunk m (Vptr sp i) = Some v ->
- Mem.loadv chunk m' (Vptr sp (Int.repr j)) = Some v).
+ Mem.loadv chunk m (Vptr sp ofs) = Some v ->
+ Mem.loadv chunk m' (Vptr sp (Ptrofs.repr j)) = Some v).
{
- simpl; intros. rewrite Int.unsigned_repr by omega.
+ simpl; intros. rewrite Ptrofs.unsigned_repr by omega.
unfold j, delta. eapply load_memcpy; eauto.
apply Zmod_divide; auto. generalize (align_chunk_pos chunk); omega.
}
inv H2.
-+ inv H3. destruct vl... simpl in H6. rewrite Int.add_zero_l in H6. inv H6.
- apply eq_holds_strict. econstructor. simpl. rewrite Int.add_zero_l. eauto.
++ inv H3. exploit eval_addressing_Ainstack_inv; eauto. intros [E1 E2].
+ simpl in E2; rewrite Ptrofs.add_zero_l in E2. subst a.
+ apply eq_holds_strict. econstructor. rewrite eval_addressing_Ainstack.
+ simpl. rewrite Ptrofs.add_zero_l. eauto.
apply LD; auto.
-+ inv H4. destruct vl... simpl in H7. rewrite Int.add_zero_l in H7. inv H7.
++ inv H4. exploit eval_addressing_Ainstack_inv; eauto. intros [E1 E2].
+ simpl in E2; rewrite Ptrofs.add_zero_l in E2. subst a.
apply eq_holds_lessdef with v; auto.
- econstructor. simpl. rewrite Int.add_zero_l. eauto. apply LD; auto.
+ econstructor. rewrite eval_addressing_Ainstack. simpl. rewrite Ptrofs.add_zero_l. eauto.
+ apply LD; auto.
Qed.
Lemma add_memcpy_eqs_charact:
@@ -677,15 +687,15 @@ Qed.
Lemma add_memcpy_holds:
forall m bsrc osrc sz bytes bdst odst m' valu ge sp rs n1 n2 bc asrc adst,
- Mem.loadbytes m bsrc (Int.unsigned osrc) sz = Some bytes ->
- Mem.storebytes m bdst (Int.unsigned odst) bytes = Some m' ->
- numbering_holds valu ge (Vptr sp Int.zero) rs m n1 ->
- numbering_holds valu ge (Vptr sp Int.zero) rs m' n2 ->
+ Mem.loadbytes m bsrc (Ptrofs.unsigned osrc) sz = Some bytes ->
+ Mem.storebytes m bdst (Ptrofs.unsigned odst) bytes = Some m' ->
+ numbering_holds valu ge (Vptr sp Ptrofs.zero) rs m n1 ->
+ numbering_holds valu ge (Vptr sp Ptrofs.zero) rs m' n2 ->
pmatch bc bsrc osrc asrc ->
pmatch bc bdst odst adst ->
bc sp = BCstack ->
Ple (num_next n1) (num_next n2) ->
- numbering_holds valu ge (Vptr sp Int.zero) rs m' (add_memcpy n1 n2 asrc adst sz).
+ numbering_holds valu ge (Vptr sp Ptrofs.zero) rs m' (add_memcpy n1 n2 asrc adst sz).
Proof.
intros. unfold add_memcpy.
destruct asrc; auto; destruct adst; auto.
diff --git a/backend/Cminor.v b/backend/Cminor.v
index 0d959531..e238140b 100644
--- a/backend/Cminor.v
+++ b/backend/Cminor.v
@@ -38,8 +38,8 @@ Inductive constant : Type :=
| Ofloatconst: float -> constant (**r double-precision floating-point constant *)
| Osingleconst: float32 -> constant (**r single-precision floating-point constant *)
| Olongconst: int64 -> constant (**r long integer constant *)
- | Oaddrsymbol: ident -> int -> constant (**r address of the symbol plus the offset *)
- | Oaddrstack: int -> constant. (**r stack pointer plus the given offset *)
+ | Oaddrsymbol: ident -> ptrofs -> constant (**r address of the symbol plus the offset *)
+ | Oaddrstack: ptrofs -> constant. (**r stack pointer plus the given offset *)
Inductive unary_operation : Type :=
| Ocast8unsigned: unary_operation (**r 8-bit zero extension *)
@@ -257,11 +257,8 @@ Definition eval_constant (sp: val) (cst: constant) : option val :=
| Ofloatconst n => Some (Vfloat n)
| Osingleconst n => Some (Vsingle n)
| Olongconst n => Some (Vlong n)
- | Oaddrsymbol s ofs =>
- Some(match Genv.find_symbol ge s with
- | None => Vundef
- | Some b => Vptr b ofs end)
- | Oaddrstack ofs => Some (Val.add sp (Vint ofs))
+ | Oaddrsymbol s ofs => Some (Genv.symbol_address ge s ofs)
+ | Oaddrstack ofs => Some (Val.offset_ptr sp ofs)
end.
Definition eval_unop (op: unary_operation) (arg: val) : option val :=
@@ -343,7 +340,7 @@ Definition eval_binop
| Ocmpf c => Some (Val.cmpf c arg1 arg2)
| Ocmpfs c => Some (Val.cmpfs c arg1 arg2)
| Ocmpl c => Val.cmpl c arg1 arg2
- | Ocmplu c => Val.cmplu c arg1 arg2
+ | Ocmplu c => Val.cmplu (Mem.valid_pointer m) c arg1 arg2
end.
(** Evaluation of an expression: [eval_expr ge sp e m a v]
@@ -444,7 +441,7 @@ Inductive step: state -> trace -> state -> Prop :=
| step_skip_call: forall f k sp e m m',
is_call_cont k ->
Mem.free m sp 0 f.(fn_stackspace) = Some m' ->
- step (State f Sskip k (Vptr sp Int.zero) e m)
+ step (State f Sskip k (Vptr sp Ptrofs.zero) e m)
E0 (Returnstate Vundef k m')
| step_assign: forall f id a k sp e m v,
@@ -468,12 +465,12 @@ Inductive step: state -> trace -> state -> Prop :=
E0 (Callstate fd vargs (Kcall optid f sp e k) m)
| step_tailcall: forall f sig a bl k sp e m vf vargs fd m',
- eval_expr (Vptr sp Int.zero) e m a vf ->
- eval_exprlist (Vptr sp Int.zero) e m bl vargs ->
+ eval_expr (Vptr sp Ptrofs.zero) e m a vf ->
+ eval_exprlist (Vptr sp Ptrofs.zero) e m bl vargs ->
Genv.find_funct ge vf = Some fd ->
funsig fd = sig ->
Mem.free m sp 0 f.(fn_stackspace) = Some m' ->
- step (State f (Stailcall sig a bl) k (Vptr sp Int.zero) e m)
+ step (State f (Stailcall sig a bl) k (Vptr sp Ptrofs.zero) e m)
E0 (Callstate fd vargs (call_cont k) m')
| step_builtin: forall f optid ef bl k sp e m vargs t vres m',
@@ -518,12 +515,12 @@ Inductive step: state -> trace -> state -> Prop :=
| step_return_0: forall f k sp e m m',
Mem.free m sp 0 f.(fn_stackspace) = Some m' ->
- step (State f (Sreturn None) k (Vptr sp Int.zero) e m)
+ step (State f (Sreturn None) k (Vptr sp Ptrofs.zero) e m)
E0 (Returnstate Vundef (call_cont k) m')
| step_return_1: forall f a k sp e m v m',
- eval_expr (Vptr sp Int.zero) e m a v ->
+ eval_expr (Vptr sp Ptrofs.zero) e m a v ->
Mem.free m sp 0 f.(fn_stackspace) = Some m' ->
- step (State f (Sreturn (Some a)) k (Vptr sp Int.zero) e m)
+ step (State f (Sreturn (Some a)) k (Vptr sp Ptrofs.zero) e m)
E0 (Returnstate v (call_cont k) m')
| step_label: forall f lbl s k sp e m,
@@ -539,7 +536,7 @@ Inductive step: state -> trace -> state -> Prop :=
Mem.alloc m 0 f.(fn_stackspace) = (m', sp) ->
set_locals f.(fn_vars) (set_params vargs f.(fn_params)) = e ->
step (Callstate (Internal f) vargs k m)
- E0 (State f f.(fn_body) k (Vptr sp Int.zero) e m')
+ E0 (State f f.(fn_body) k (Vptr sp Ptrofs.zero) e m')
| step_external_function: forall ef vargs k m t vres m',
external_call ef ge vargs m t vres m' ->
step (Callstate (External ef) vargs k m)
@@ -649,7 +646,7 @@ Inductive eval_funcall:
forall m f vargs m1 sp e t e2 m2 out vres m3,
Mem.alloc m 0 f.(fn_stackspace) = (m1, sp) ->
set_locals f.(fn_vars) (set_params vargs f.(fn_params)) = e ->
- exec_stmt f (Vptr sp Int.zero) e m1 f.(fn_body) t e2 m2 out ->
+ exec_stmt f (Vptr sp Ptrofs.zero) e m1 f.(fn_body) t e2 m2 out ->
outcome_result_value out f.(fn_sig).(sig_res) vres ->
outcome_free_mem out m2 sp f.(fn_stackspace) m3 ->
eval_funcall m (Internal f) vargs t m3 vres
@@ -748,13 +745,13 @@ with exec_stmt:
exec_stmt f sp e m (Sreturn (Some a)) E0 e m (Out_return (Some v))
| exec_Stailcall:
forall f sp e m sig a bl vf vargs fd t m' m'' vres,
- eval_expr ge (Vptr sp Int.zero) e m a vf ->
- eval_exprlist ge (Vptr sp Int.zero) e m bl vargs ->
+ eval_expr ge (Vptr sp Ptrofs.zero) e m a vf ->
+ eval_exprlist ge (Vptr sp Ptrofs.zero) e m bl vargs ->
Genv.find_funct ge vf = Some fd ->
funsig fd = sig ->
Mem.free m sp 0 f.(fn_stackspace) = Some m' ->
eval_funcall m' fd vargs t m'' vres ->
- exec_stmt f (Vptr sp Int.zero) e m (Stailcall sig a bl) t e m'' (Out_tailcall_return vres).
+ exec_stmt f (Vptr sp Ptrofs.zero) e m (Stailcall sig a bl) t e m'' (Out_tailcall_return vres).
Scheme eval_funcall_ind2 := Minimality for eval_funcall Sort Prop
with exec_stmt_ind2 := Minimality for exec_stmt Sort Prop.
@@ -774,7 +771,7 @@ CoInductive evalinf_funcall:
forall m f vargs m1 sp e t,
Mem.alloc m 0 f.(fn_stackspace) = (m1, sp) ->
set_locals f.(fn_vars) (set_params vargs f.(fn_params)) = e ->
- execinf_stmt f (Vptr sp Int.zero) e m1 f.(fn_body) t ->
+ execinf_stmt f (Vptr sp Ptrofs.zero) e m1 f.(fn_body) t ->
evalinf_funcall m (Internal f) vargs t
(** [execinf_stmt ge sp e m s t] means that statement [s] diverges.
@@ -823,13 +820,13 @@ with execinf_stmt:
execinf_stmt f sp e m (Sblock s) t
| execinf_Stailcall:
forall f sp e m sig a bl vf vargs fd m' t,
- eval_expr ge (Vptr sp Int.zero) e m a vf ->
- eval_exprlist ge (Vptr sp Int.zero) e m bl vargs ->
+ eval_expr ge (Vptr sp Ptrofs.zero) e m a vf ->
+ eval_exprlist ge (Vptr sp Ptrofs.zero) e m bl vargs ->
Genv.find_funct ge vf = Some fd ->
funsig fd = sig ->
Mem.free m sp 0 f.(fn_stackspace) = Some m' ->
evalinf_funcall m' fd vargs t ->
- execinf_stmt f (Vptr sp Int.zero) e m (Stailcall sig a bl) t.
+ execinf_stmt f (Vptr sp Ptrofs.zero) e m (Stailcall sig a bl) t.
End NATURALSEM.
diff --git a/backend/CminorSel.v b/backend/CminorSel.v
index d654502b..9439c269 100644
--- a/backend/CminorSel.v
+++ b/backend/CminorSel.v
@@ -246,7 +246,7 @@ Inductive eval_expr_or_symbol: letenv -> expr + ident -> val -> Prop :=
eval_expr_or_symbol le (inl _ e) v
| eval_eos_s: forall le id b,
Genv.find_symbol ge id = Some b ->
- eval_expr_or_symbol le (inr _ id) (Vptr b Int.zero).
+ eval_expr_or_symbol le (inr _ id) (Vptr b Ptrofs.zero).
Inductive eval_builtin_arg: builtin_arg expr -> val -> Prop :=
| eval_BA: forall a v,
@@ -261,10 +261,10 @@ Inductive eval_builtin_arg: builtin_arg expr -> val -> Prop :=
| eval_BA_single: forall n,
eval_builtin_arg (BA_single n) (Vsingle n)
| eval_BA_loadstack: forall chunk ofs v,
- Mem.loadv chunk m (Val.add sp (Vint ofs)) = Some v ->
+ Mem.loadv chunk m (Val.offset_ptr sp ofs) = Some v ->
eval_builtin_arg (BA_loadstack chunk ofs) v
| eval_BA_addrstack: forall ofs,
- eval_builtin_arg (BA_addrstack ofs) (Val.add sp (Vint ofs))
+ eval_builtin_arg (BA_addrstack ofs) (Val.offset_ptr sp ofs)
| eval_BA_loadglobal: forall chunk id ofs v,
Mem.loadv chunk m (Genv.symbol_address ge id ofs) = Some v ->
eval_builtin_arg (BA_loadglobal chunk id ofs) v
@@ -338,7 +338,7 @@ Inductive step: state -> trace -> state -> Prop :=
| step_skip_call: forall f k sp e m m',
is_call_cont k ->
Mem.free m sp 0 f.(fn_stackspace) = Some m' ->
- step (State f Sskip k (Vptr sp Int.zero) e m)
+ step (State f Sskip k (Vptr sp Ptrofs.zero) e m)
E0 (Returnstate Vundef k m')
| step_assign: forall f id a k sp e m v,
@@ -363,12 +363,12 @@ Inductive step: state -> trace -> state -> Prop :=
E0 (Callstate fd vargs (Kcall optid f sp e k) m)
| step_tailcall: forall f sig a bl k sp e m vf vargs fd m',
- eval_expr_or_symbol (Vptr sp Int.zero) e m nil a vf ->
- eval_exprlist (Vptr sp Int.zero) e m nil bl vargs ->
+ eval_expr_or_symbol (Vptr sp Ptrofs.zero) e m nil a vf ->
+ eval_exprlist (Vptr sp Ptrofs.zero) e m nil bl vargs ->
Genv.find_funct ge vf = Some fd ->
funsig fd = sig ->
Mem.free m sp 0 f.(fn_stackspace) = Some m' ->
- step (State f (Stailcall sig a bl) k (Vptr sp Int.zero) e m)
+ step (State f (Stailcall sig a bl) k (Vptr sp Ptrofs.zero) e m)
E0 (Callstate fd vargs (call_cont k) m')
| step_builtin: forall f res ef al k sp e m vl t v m',
@@ -411,12 +411,12 @@ Inductive step: state -> trace -> state -> Prop :=
| step_return_0: forall f k sp e m m',
Mem.free m sp 0 f.(fn_stackspace) = Some m' ->
- step (State f (Sreturn None) k (Vptr sp Int.zero) e m)
+ step (State f (Sreturn None) k (Vptr sp Ptrofs.zero) e m)
E0 (Returnstate Vundef (call_cont k) m')
| step_return_1: forall f a k sp e m v m',
- eval_expr (Vptr sp Int.zero) e m nil a v ->
+ eval_expr (Vptr sp Ptrofs.zero) e m nil a v ->
Mem.free m sp 0 f.(fn_stackspace) = Some m' ->
- step (State f (Sreturn (Some a)) k (Vptr sp Int.zero) e m)
+ step (State f (Sreturn (Some a)) k (Vptr sp Ptrofs.zero) e m)
E0 (Returnstate v (call_cont k) m')
| step_label: forall f lbl s k sp e m,
@@ -432,7 +432,7 @@ Inductive step: state -> trace -> state -> Prop :=
Mem.alloc m 0 f.(fn_stackspace) = (m', sp) ->
set_locals f.(fn_vars) (set_params vargs f.(fn_params)) = e ->
step (Callstate (Internal f) vargs k m)
- E0 (State f f.(fn_body) k (Vptr sp Int.zero) e m')
+ E0 (State f f.(fn_body) k (Vptr sp Ptrofs.zero) e m')
| step_external_function: forall ef vargs k m t vres m',
external_call ef ge vargs m t vres m' ->
step (Callstate (External ef) vargs k m)
diff --git a/backend/Constprop.v b/backend/Constprop.v
index 4de80b7a..151f8418 100644
--- a/backend/Constprop.v
+++ b/backend/Constprop.v
@@ -56,22 +56,12 @@ Definition transf_ros (ae: AE.t) (ros: reg + ident) : reg + ident :=
match ros with
| inl r =>
match areg ae r with
- | Ptr(Gl symb ofs) => if Int.eq ofs Int.zero then inr _ symb else ros
+ | Ptr(Gl symb ofs) => if Ptrofs.eq ofs Ptrofs.zero then inr _ symb else ros
| _ => ros
end
| inr s => ros
end.
-Definition const_for_result (a: aval) : option operation :=
- match a with
- | I n => Some(Ointconst n)
- | F n => if Compopts.generate_float_constants tt then Some(Ofloatconst n) else None
- | FS n => if Compopts.generate_float_constants tt then Some(Osingleconst n) else None
- | Ptr(Gl symb ofs) => Some(Oaddrsymbol symb ofs)
- | Ptr(Stk ofs) => Some(Oaddrstack ofs)
- | _ => None
- end.
-
Fixpoint successor_rec (n: nat) (f: function) (ae: AE.t) (pc: node) : node :=
match n with
| O => pc
diff --git a/backend/Constpropproof.v b/backend/Constpropproof.v
index 4e76c641..fd9cfaa5 100644
--- a/backend/Constpropproof.v
+++ b/backend/Constpropproof.v
@@ -107,7 +107,7 @@ Proof.
simpl. inv LD. apply functions_translated; auto. rewrite <- H0 in FF; discriminate.
}
destruct (areg ae r); auto. destruct p; auto.
- predSpec Int.eq Int.eq_spec ofs Int.zero; intros; auto.
+ predSpec Ptrofs.eq Ptrofs.eq_spec ofs Ptrofs.zero; intros; auto.
subst ofs. exploit vmatch_ptr_gl; eauto. intros LD'. inv LD'; try discriminate.
rewrite H1 in FF. unfold Genv.symbol_address in FF.
simpl. rewrite symbols_preserved.
@@ -127,26 +127,12 @@ Lemma const_for_result_correct:
vmatch bc v a ->
bc sp = BCstack ->
genv_match bc ge ->
- exists v', eval_operation tge (Vptr sp Int.zero) op nil m = Some v' /\ Val.lessdef v v'.
+ exists v', eval_operation tge (Vptr sp Ptrofs.zero) op nil m = Some v' /\ Val.lessdef v v'.
Proof.
- unfold const_for_result; intros.
- destruct a; try discriminate.
-- (* integer *)
- inv H. inv H0. exists (Vint n); auto.
-- (* float *)
- destruct (Compopts.generate_float_constants tt); inv H. inv H0. exists (Vfloat f); auto.
-- (* single *)
- destruct (Compopts.generate_float_constants tt); inv H. inv H0. exists (Vsingle f); auto.
-- (* pointer *)
- destruct p; try discriminate.
- + (* global *)
- inv H. exists (Genv.symbol_address ge id ofs); split.
- unfold Genv.symbol_address. rewrite <- symbols_preserved. reflexivity.
- eapply vmatch_ptr_gl; eauto.
- + (* stack *)
- inv H. exists (Vptr sp ofs); split.
- simpl; rewrite Int.add_zero_l; auto.
- eapply vmatch_ptr_stk; eauto.
+ intros. exploit ConstpropOpproof.const_for_result_correct; eauto. intros (v' & A & B).
+ exists v'; split.
+ rewrite <- A; apply eval_operation_preserved. exact symbols_preserved.
+ auto.
Qed.
Inductive match_pc (f: function) (rs: regset) (m: mem): nat -> node -> node -> Prop :=
@@ -399,12 +385,12 @@ Proof.
assert(OP:
let (op', args') := op_strength_reduction op args (aregs ae args) in
exists v',
- eval_operation ge (Vptr sp0 Int.zero) op' rs ## args' m = Some v' /\
+ eval_operation ge (Vptr sp0 Ptrofs.zero) op' rs ## args' m = Some v' /\
Val.lessdef v v').
{ eapply op_strength_reduction_correct with (ae := ae); eauto with va. }
destruct (op_strength_reduction op args (aregs ae args)) as [op' args'].
destruct OP as [v' [EV' LD']].
- assert (EV'': exists v'', eval_operation ge (Vptr sp0 Int.zero) op' rs'##args' m' = Some v'' /\ Val.lessdef v' v'').
+ assert (EV'': exists v'', eval_operation ge (Vptr sp0 Ptrofs.zero) op' rs'##args' m' = Some v'' /\ Val.lessdef v' v'').
{ eapply eval_operation_lessdef; eauto. eapply regs_lessdef_regs; eauto. }
destruct EV'' as [v'' [EV'' LD'']].
left; econstructor; econstructor; split.
@@ -431,14 +417,14 @@ Proof.
assert (ADDR:
let (addr', args') := addr_strength_reduction addr args (aregs ae args) in
exists a',
- eval_addressing ge (Vptr sp0 Int.zero) addr' rs ## args' = Some a' /\
+ eval_addressing ge (Vptr sp0 Ptrofs.zero) addr' rs ## args' = Some a' /\
Val.lessdef a a').
{ eapply addr_strength_reduction_correct with (ae := ae); eauto with va. }
destruct (addr_strength_reduction addr args (aregs ae args)) as [addr' args'].
destruct ADDR as (a' & P & Q).
exploit eval_addressing_lessdef. eapply regs_lessdef_regs; eauto. eexact P.
intros (a'' & U & V).
- assert (W: eval_addressing tge (Vptr sp0 Int.zero) addr' rs'##args' = Some a'').
+ assert (W: eval_addressing tge (Vptr sp0 Ptrofs.zero) addr' rs'##args' = Some a'').
{ rewrite <- U. apply eval_addressing_preserved. exact symbols_preserved. }
exploit Mem.loadv_extends. eauto. eauto. apply Val.lessdef_trans with a'; eauto.
intros (v' & X & Y).
@@ -451,14 +437,14 @@ Proof.
assert (ADDR:
let (addr', args') := addr_strength_reduction addr args (aregs ae args) in
exists a',
- eval_addressing ge (Vptr sp0 Int.zero) addr' rs ## args' = Some a' /\
+ eval_addressing ge (Vptr sp0 Ptrofs.zero) addr' rs ## args' = Some a' /\
Val.lessdef a a').
{ eapply addr_strength_reduction_correct with (ae := ae); eauto with va. }
destruct (addr_strength_reduction addr args (aregs ae args)) as [addr' args'].
destruct ADDR as (a' & P & Q).
exploit eval_addressing_lessdef. eapply regs_lessdef_regs; eauto. eexact P.
intros (a'' & U & V).
- assert (W: eval_addressing tge (Vptr sp0 Int.zero) addr' rs'##args' = Some a'').
+ assert (W: eval_addressing tge (Vptr sp0 Ptrofs.zero) addr' rs'##args' = Some a'').
{ rewrite <- U. apply eval_addressing_preserved. exact symbols_preserved. }
exploit Mem.storev_extends. eauto. eauto. apply Val.lessdef_trans with a'; eauto. apply REGS.
intros (m2' & X & Y).
@@ -510,7 +496,7 @@ Opaque builtin_strength_reduction.
generalize (cond_strength_reduction_correct bc ae rs m EM cond args (aregs ae args) (refl_equal _)).
destruct (cond_strength_reduction cond args (aregs ae args)) as [cond' args'].
intros EV1 TCODE.
- left; exists O; exists (State s' (transf_function (romem_for cu) f) (Vptr sp0 Int.zero) (if b then ifso else ifnot) rs' m'); split.
+ left; exists O; exists (State s' (transf_function (romem_for cu) f) (Vptr sp0 Ptrofs.zero) (if b then ifso else ifnot) rs' m'); split.
destruct (resolve_branch ac) eqn: RB.
assert (b0 = b) by (eapply resolve_branch_sound; eauto). subst b0.
destruct b; eapply exec_Inop; eauto.
@@ -534,7 +520,7 @@ Opaque builtin_strength_reduction.
rewrite H1. auto. }
assert (rs'#arg = Vint n).
{ generalize (REGS arg). rewrite H0. intros LD; inv LD; auto. }
- left; exists O; exists (State s' (transf_function (romem_for cu) f) (Vptr sp0 Int.zero) pc' rs' m'); split.
+ left; exists O; exists (State s' (transf_function (romem_for cu) f) (Vptr sp0 Ptrofs.zero) pc' rs' m'); split.
destruct A. eapply exec_Ijumptable; eauto. eapply exec_Inop; eauto.
eapply match_states_succ; eauto.
diff --git a/backend/Deadcodeproof.v b/backend/Deadcodeproof.v
index 5c293ee1..52f1f112 100644
--- a/backend/Deadcodeproof.v
+++ b/backend/Deadcodeproof.v
@@ -489,8 +489,8 @@ Inductive match_stackframes: stackframe -> stackframe -> Prop :=
Val.lessdef v tv ->
eagree (e#res <- v) (te#res<- tv)
(fst (transfer f (vanalyze cu f) pc an!!pc))),
- match_stackframes (Stackframe res f (Vptr sp Int.zero) pc e)
- (Stackframe res tf (Vptr sp Int.zero) pc te).
+ match_stackframes (Stackframe res f (Vptr sp Ptrofs.zero) pc e)
+ (Stackframe res tf (Vptr sp Ptrofs.zero) pc te).
Inductive match_states: state -> state -> Prop :=
| match_regular_states:
@@ -501,8 +501,8 @@ Inductive match_states: state -> state -> Prop :=
(ANL: analyze (vanalyze cu f) f = Some an)
(ENV: eagree e te (fst (transfer f (vanalyze cu f) pc an!!pc)))
(MEM: magree m tm (nlive ge sp (snd (transfer f (vanalyze cu f) pc an!!pc)))),
- match_states (State s f (Vptr sp Int.zero) pc e m)
- (State ts tf (Vptr sp Int.zero) pc te tm)
+ match_states (State s f (Vptr sp Ptrofs.zero) pc e m)
+ (State ts tf (Vptr sp Ptrofs.zero) pc te tm)
| match_call_states:
forall s f args m ts tf targs tm cu
(STACKS: list_forall2 match_stackframes s ts)
@@ -544,8 +544,8 @@ Lemma match_succ_states:
(ANPC: an!!pc = (ne, nm))
(ENV: eagree e te ne)
(MEM: magree m tm (nlive ge sp nm)),
- match_states (State s f (Vptr sp Int.zero) pc' e m)
- (State ts tf (Vptr sp Int.zero) pc' te tm).
+ match_states (State s f (Vptr sp Ptrofs.zero) pc' e m)
+ (State ts tf (Vptr sp Ptrofs.zero) pc' te tm).
Proof.
intros. exploit analyze_successors; eauto. rewrite ANPC; simpl. intros [A B].
econstructor; eauto.
@@ -567,7 +567,7 @@ Qed.
Lemma transfer_builtin_arg_sound:
forall bc e e' sp m m' a v,
- eval_builtin_arg ge (fun r => e#r) (Vptr sp Int.zero) m a v ->
+ eval_builtin_arg ge (fun r => e#r) (Vptr sp Ptrofs.zero) m a v ->
forall nv ne1 nm1 ne2 nm2,
transfer_builtin_arg nv (ne1, nm1) a = (ne2, nm2) ->
eagree e e' ne2 ->
@@ -575,7 +575,7 @@ Lemma transfer_builtin_arg_sound:
genv_match bc ge ->
bc sp = BCstack ->
exists v',
- eval_builtin_arg ge (fun r => e'#r) (Vptr sp Int.zero) m' a v'
+ eval_builtin_arg ge (fun r => e'#r) (Vptr sp Ptrofs.zero) m' a v'
/\ vagree v v' nv
/\ eagree e e' ne1
/\ magree m m' (nlive ge sp nm1).
@@ -587,11 +587,11 @@ Proof.
- exists (Vfloat n); intuition auto. constructor. apply vagree_same.
- exists (Vsingle n); intuition auto. constructor. apply vagree_same.
- simpl in H. exploit magree_load; eauto.
- intros. eapply nlive_add; eauto with va. rewrite Int.add_zero_l in H0; auto.
+ intros. eapply nlive_add; eauto with va. rewrite Ptrofs.add_zero_l in H0; auto.
intros (v' & A & B).
exists v'; intuition auto. constructor; auto. apply vagree_lessdef; auto.
eapply magree_monotone; eauto. intros; eapply incl_nmem_add; eauto.
-- exists (Vptr sp (Int.add Int.zero ofs)); intuition auto with na. constructor.
+- exists (Vptr sp (Ptrofs.add Ptrofs.zero ofs)); intuition auto with na. constructor.
- unfold Senv.symbol_address in H; simpl in H.
destruct (Genv.find_symbol ge id) as [b|] eqn:FS; simpl in H; try discriminate.
exploit magree_load; eauto.
@@ -613,7 +613,7 @@ Qed.
Lemma transfer_builtin_args_sound:
forall e sp m e' m' bc al vl,
- eval_builtin_args ge (fun r => e#r) (Vptr sp Int.zero) m al vl ->
+ eval_builtin_args ge (fun r => e#r) (Vptr sp Ptrofs.zero) m al vl ->
forall ne1 nm1 ne2 nm2,
transfer_builtin_args (ne1, nm1) al = (ne2, nm2) ->
eagree e e' ne2 ->
@@ -621,7 +621,7 @@ Lemma transfer_builtin_args_sound:
genv_match bc ge ->
bc sp = BCstack ->
exists vl',
- eval_builtin_args ge (fun r => e'#r) (Vptr sp Int.zero) m' al vl'
+ eval_builtin_args ge (fun r => e'#r) (Vptr sp Ptrofs.zero) m' al vl'
/\ Val.lessdef_list vl vl'
/\ eagree e e' ne1
/\ magree m m' (nlive ge sp nm1).
@@ -639,8 +639,8 @@ Lemma can_eval_builtin_arg:
forall sp e m e' m' P,
magree m m' P ->
forall a v,
- eval_builtin_arg ge (fun r => e#r) (Vptr sp Int.zero) m a v ->
- exists v', eval_builtin_arg tge (fun r => e'#r) (Vptr sp Int.zero) m' a v'.
+ eval_builtin_arg ge (fun r => e#r) (Vptr sp Ptrofs.zero) m a v ->
+ exists v', eval_builtin_arg tge (fun r => e'#r) (Vptr sp Ptrofs.zero) m' a v'.
Proof.
intros until P; intros MA.
assert (LD: forall chunk addr v,
@@ -663,8 +663,8 @@ Lemma can_eval_builtin_args:
forall sp e m e' m' P,
magree m m' P ->
forall al vl,
- eval_builtin_args ge (fun r => e#r) (Vptr sp Int.zero) m al vl ->
- exists vl', eval_builtin_args tge (fun r => e'#r) (Vptr sp Int.zero) m' al vl'.
+ eval_builtin_args ge (fun r => e#r) (Vptr sp Ptrofs.zero) m al vl ->
+ exists vl', eval_builtin_args tge (fun r => e'#r) (Vptr sp Ptrofs.zero) m' al vl'.
Proof.
induction 2.
- exists (@nil val); constructor.
diff --git a/backend/Debugvar.v b/backend/Debugvar.v
index 5d31831a..1f361030 100644
--- a/backend/Debugvar.v
+++ b/backend/Debugvar.v
@@ -136,7 +136,7 @@ Definition kill_at_call (s: avail) : avail :=
Definition eq_arg (a1 a2: builtin_arg loc) : {a1=a2} + {a1<>a2}.
Proof.
- generalize Loc.eq ident_eq Int.eq_dec Int64.eq_dec Float.eq_dec Float32.eq_dec chunk_eq;
+ generalize Loc.eq ident_eq Int.eq_dec Int64.eq_dec Ptrofs.eq_dec Float.eq_dec Float32.eq_dec chunk_eq;
decide equality.
Defined.
Global Opaque eq_arg.
diff --git a/backend/IRC.ml b/backend/IRC.ml
index 036b4ac5..43955897 100644
--- a/backend/IRC.ml
+++ b/backend/IRC.ml
@@ -238,9 +238,16 @@ type graph = {
by giving it a negative spill cost. *)
let class_of_type = function
- | Tint | Tany32 -> 0
- | Tfloat | Tsingle | Tany64 -> 1
- | Tlong -> assert false
+ | Tint | Tlong -> 0
+ | Tfloat | Tsingle -> 1
+ | Tany32 | Tany64 -> assert false
+
+let class_of_reg r =
+ if Conventions1.is_float_reg r then 1 else 0
+
+let class_of_loc = function
+ | R r -> class_of_reg r
+ | S(_, _, ty) -> class_of_type ty
let no_spill_class = 2
@@ -319,7 +326,7 @@ let newNodeOfLoc g l =
let ty = Loc.coq_type l in
g.nextIdent <- g.nextIdent + 1;
{ ident = g.nextIdent; typ = ty;
- var = L l; regclass = class_of_type ty;
+ var = L l; regclass = class_of_loc l;
accesses = 0; spillcost = 0.0;
adjlist = []; degree = 0; movelist = []; extra_adj = []; extra_pref = [];
alias = None;
@@ -828,20 +835,26 @@ let compare_slots s1 s2 =
| S(_, ofs1, _), S(_, ofs2, _) -> Z.compare ofs1 ofs2
| _, _ -> assert false
+let align a b = (a + b - 1) land (-b) (* assuming b is a power of 2 *)
+
let find_slot conflicts typ =
+ let sz = Z.to_int (Locations.typesize typ) in
+ let al = Z.to_int (Locations.typealign typ) in
let rec find curr = function
| [] ->
- S(Local, curr, typ)
+ S(Local, Z.of_uint curr, typ)
| S(Local, ofs, typ') :: l ->
- if Z.le (Z.add curr (Locations.typesize typ)) ofs then
- S(Local, curr, typ)
+ let ofs = Z.to_int ofs in
+ if curr + sz <= ofs then
+ S(Local, Z.of_uint curr, typ)
else begin
- let ofs' = Z.add ofs (Locations.typesize typ') in
- find (if Z.le ofs' curr then curr else ofs') l
+ let sz' = Z.to_int (Locations.typesize typ') in
+ let ofs' = align (ofs + sz') al in
+ find (if ofs' <= curr then curr else ofs') l
end
| _ :: l ->
find curr l
- in find Z.zero (List.stable_sort compare_slots conflicts)
+ in find 0 (List.stable_sort compare_slots conflicts)
(* Record locations assigned to interfering nodes *)
@@ -891,10 +904,10 @@ let location_of_var g v =
| None -> assert false
| Some l -> l
with Not_found ->
- match ty with
- | Tint -> R dummy_int_reg
- | Tfloat | Tsingle -> R dummy_float_reg
- | Tlong | Tany32 | Tany64 -> assert false
+ match class_of_type ty with
+ | 0 -> R dummy_int_reg
+ | 1 -> R dummy_float_reg
+ | _ -> assert false
(* The exported interface *)
diff --git a/backend/IRC.mli b/backend/IRC.mli
index d27dedaa..30b6d5c1 100644
--- a/backend/IRC.mli
+++ b/backend/IRC.mli
@@ -41,3 +41,7 @@ val coloring: graph -> (var -> loc)
(* Machine registers that are reserved and not available for allocation. *)
val reserved_registers: mreg list ref
+
+(* Auxiliaries to deal with register classes *)
+val class_of_type: AST.typ -> int
+val class_of_loc: loc -> int
diff --git a/backend/Inlining.v b/backend/Inlining.v
index 5c8f4419..61776743 100644
--- a/backend/Inlining.v
+++ b/backend/Inlining.v
@@ -192,16 +192,16 @@ Definition sregs (ctx: context) (rl: list reg) := List.map (sreg ctx) rl.
Definition sros (ctx: context) (ros: reg + ident) := sum_left_map (sreg ctx) ros.
Definition sop (ctx: context) (op: operation) :=
- shift_stack_operation (Int.repr ctx.(dstk)) op.
+ shift_stack_operation ctx.(dstk) op.
Definition saddr (ctx: context) (addr: addressing) :=
- shift_stack_addressing (Int.repr ctx.(dstk)) addr.
+ shift_stack_addressing ctx.(dstk) addr.
Fixpoint sbuiltinarg (ctx: context) (a: builtin_arg reg) : builtin_arg reg :=
match a with
| BA x => BA (sreg ctx x)
- | BA_loadstack chunk ofs => BA_loadstack chunk (Int.add ofs (Int.repr ctx.(dstk)))
- | BA_addrstack ofs => BA_addrstack (Int.add ofs (Int.repr ctx.(dstk)))
+ | BA_loadstack chunk ofs => BA_loadstack chunk (Ptrofs.add ofs (Ptrofs.repr ctx.(dstk)))
+ | BA_addrstack ofs => BA_addrstack (Ptrofs.add ofs (Ptrofs.repr ctx.(dstk)))
| BA_splitlong hi lo => BA_splitlong (sbuiltinarg ctx hi) (sbuiltinarg ctx lo)
| _ => a
end.
@@ -437,13 +437,13 @@ Definition expand_function (fenv: funenv) (f: function): mon context :=
Local Open Scope string_scope.
(** Inlining can increase the size of the function's stack block. We must
- make sure that the new size does not exceed [Int.max_unsigned], otherwise
+ make sure that the new size does not exceed [Ptrofs.max_unsigned], otherwise
address computations within the stack would overflow and produce incorrect
results. *)
Definition transf_function (fenv: funenv) (f: function) : Errors.res function :=
let '(R ctx s _) := expand_function fenv f initstate in
- if zlt s.(st_stksize) Int.max_unsigned then
+ if zlt s.(st_stksize) Ptrofs.max_unsigned then
OK (mkfunction f.(fn_sig)
(sregs ctx f.(fn_params))
s.(st_stksize)
diff --git a/backend/Inliningproof.v b/backend/Inliningproof.v
index 91f4a3f5..d06fa997 100644
--- a/backend/Inliningproof.v
+++ b/backend/Inliningproof.v
@@ -411,8 +411,8 @@ Lemma tr_builtin_arg:
F sp = Some(sp', ctx.(dstk)) ->
Mem.inject F m m' ->
forall a v,
- eval_builtin_arg ge (fun r => rs#r) (Vptr sp Int.zero) m a v ->
- exists v', eval_builtin_arg tge (fun r => rs'#r) (Vptr sp' Int.zero) m' (sbuiltinarg ctx a) v'
+ eval_builtin_arg ge (fun r => rs#r) (Vptr sp Ptrofs.zero) m a v ->
+ exists v', eval_builtin_arg tge (fun r => rs'#r) (Vptr sp' Ptrofs.zero) m' (sbuiltinarg ctx a) v'
/\ Val.inject F v v'.
Proof.
intros until m'; intros MG AG SP MI. induction 1; simpl.
@@ -422,20 +422,20 @@ Proof.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
- exploit Mem.loadv_inject; eauto.
- instantiate (1 := Vptr sp' (Int.add ofs (Int.repr (dstk ctx)))).
- simpl. econstructor; eauto. rewrite Int.add_zero_l; auto.
- intros (v' & A & B). exists v'; split; auto. constructor. simpl. rewrite Int.add_zero_l; auto.
-- econstructor; split. constructor. simpl. econstructor; eauto. rewrite ! Int.add_zero_l; auto.
+ instantiate (1 := Vptr sp' (Ptrofs.add ofs (Ptrofs.repr (dstk ctx)))).
+ simpl. econstructor; eauto. rewrite Ptrofs.add_zero_l; auto.
+ intros (v' & A & B). exists v'; split; auto. constructor. simpl. rewrite Ptrofs.add_zero_l; auto.
+- econstructor; split. constructor. simpl. econstructor; eauto. rewrite ! Ptrofs.add_zero_l; auto.
- assert (Val.inject F (Senv.symbol_address ge id ofs) (Senv.symbol_address tge id ofs)).
{ unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
rewrite symbols_preserved. destruct (Genv.find_symbol ge id) as [b|] eqn:FS; auto.
- inv MG. econstructor. eauto. rewrite Int.add_zero; auto. }
+ inv MG. econstructor. eauto. rewrite Ptrofs.add_zero; auto. }
exploit Mem.loadv_inject; eauto. intros (v' & A & B).
exists v'; eauto with barg.
- econstructor; split. constructor.
unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
rewrite symbols_preserved. destruct (Genv.find_symbol ge id) as [b|] eqn:FS; auto.
- inv MG. econstructor. eauto. rewrite Int.add_zero; auto.
+ inv MG. econstructor. eauto. rewrite Ptrofs.add_zero; auto.
- destruct IHeval_builtin_arg1 as (v1 & A1 & B1).
destruct IHeval_builtin_arg2 as (v2 & A2 & B2).
econstructor; split. eauto with barg. apply Val.longofwords_inject; auto.
@@ -448,8 +448,8 @@ Lemma tr_builtin_args:
F sp = Some(sp', ctx.(dstk)) ->
Mem.inject F m m' ->
forall al vl,
- eval_builtin_args ge (fun r => rs#r) (Vptr sp Int.zero) m al vl ->
- exists vl', eval_builtin_args tge (fun r => rs'#r) (Vptr sp' Int.zero) m' (map (sbuiltinarg ctx) al) vl'
+ eval_builtin_args ge (fun r => rs#r) (Vptr sp Ptrofs.zero) m al vl ->
+ exists vl', eval_builtin_args tge (fun r => rs'#r) (Vptr sp' Ptrofs.zero) m' (map (sbuiltinarg ctx) al) vl'
/\ Val.inject_list F vl vl'.
Proof.
induction 5; simpl.
@@ -474,24 +474,24 @@ Inductive match_stacks (F: meminj) (m m': mem):
(AG: agree_regs F ctx rs rs')
(SP: F sp = Some(sp', ctx.(dstk)))
(PRIV: range_private F m m' sp' (ctx.(dstk) + ctx.(mstk)) f'.(fn_stacksize))
- (SSZ1: 0 <= f'.(fn_stacksize) < Int.max_unsigned)
+ (SSZ1: 0 <= f'.(fn_stacksize) < Ptrofs.max_unsigned)
(SSZ2: forall ofs, Mem.perm m' sp' ofs Max Nonempty -> 0 <= ofs <= f'.(fn_stacksize))
(RES: Ple res ctx.(mreg))
(BELOW: Plt sp' bound),
match_stacks F m m'
- (Stackframe res f (Vptr sp Int.zero) pc rs :: stk)
- (Stackframe (sreg ctx res) f' (Vptr sp' Int.zero) (spc ctx pc) rs' :: stk')
+ (Stackframe res f (Vptr sp Ptrofs.zero) pc rs :: stk)
+ (Stackframe (sreg ctx res) f' (Vptr sp' Ptrofs.zero) (spc ctx pc) rs' :: stk')
bound
| match_stacks_untailcall: forall stk res f' sp' rpc rs' stk' bound ctx
(MS: match_stacks_inside F m m' stk stk' f' ctx sp' rs')
(PRIV: range_private F m m' sp' ctx.(dstk) f'.(fn_stacksize))
- (SSZ1: 0 <= f'.(fn_stacksize) < Int.max_unsigned)
+ (SSZ1: 0 <= f'.(fn_stacksize) < Ptrofs.max_unsigned)
(SSZ2: forall ofs, Mem.perm m' sp' ofs Max Nonempty -> 0 <= ofs <= f'.(fn_stacksize))
(RET: ctx.(retinfo) = Some (rpc, res))
(BELOW: Plt sp' bound),
match_stacks F m m'
stk
- (Stackframe res f' (Vptr sp' Int.zero) rpc rs' :: stk')
+ (Stackframe res f' (Vptr sp' Ptrofs.zero) rpc rs' :: stk')
bound
with match_stacks_inside (F: meminj) (m m': mem):
@@ -512,7 +512,7 @@ with match_stacks_inside (F: meminj) (m m': mem):
(RET: ctx.(retinfo) = Some (spc ctx' pc, sreg ctx' res))
(BELOW: context_below ctx' ctx)
(SBELOW: context_stack_call ctx' ctx),
- match_stacks_inside F m m' (Stackframe res f (Vptr sp Int.zero) pc rs :: stk)
+ match_stacks_inside F m m' (Stackframe res f (Vptr sp Ptrofs.zero) pc rs :: stk)
stk' f' ctx sp' rs'.
(** Properties of match_stacks *)
@@ -863,10 +863,10 @@ Inductive match_states: RTL.state -> RTL.state -> Prop :=
(MINJ: Mem.inject F m m')
(VB: Mem.valid_block m' sp')
(PRIV: range_private F m m' sp' (ctx.(dstk) + ctx.(mstk)) f'.(fn_stacksize))
- (SSZ1: 0 <= f'.(fn_stacksize) < Int.max_unsigned)
+ (SSZ1: 0 <= f'.(fn_stacksize) < Ptrofs.max_unsigned)
(SSZ2: forall ofs, Mem.perm m' sp' ofs Max Nonempty -> 0 <= ofs <= f'.(fn_stacksize)),
- match_states (State stk f (Vptr sp Int.zero) pc rs m)
- (State stk' f' (Vptr sp' Int.zero) (spc ctx pc) rs' m')
+ match_states (State stk f (Vptr sp Ptrofs.zero) pc rs m)
+ (State stk' f' (Vptr sp' Ptrofs.zero) (spc ctx pc) rs' m')
| match_call_states: forall stk fd args m stk' fd' args' m' cunit F
(MS: match_stacks F m m' stk stk' (Mem.nextblock m'))
(LINK: linkorder cunit prog)
@@ -886,10 +886,10 @@ Inductive match_states: RTL.state -> RTL.state -> Prop :=
(MINJ: Mem.inject F m m')
(VB: Mem.valid_block m' sp')
(PRIV: range_private F m m' sp' ctx.(dstk) f'.(fn_stacksize))
- (SSZ1: 0 <= f'.(fn_stacksize) < Int.max_unsigned)
+ (SSZ1: 0 <= f'.(fn_stacksize) < Ptrofs.max_unsigned)
(SSZ2: forall ofs, Mem.perm m' sp' ofs Max Nonempty -> 0 <= ofs <= f'.(fn_stacksize)),
match_states (Callstate stk (Internal f) vargs m)
- (State stk' f' (Vptr sp' Int.zero) pc' rs' m')
+ (State stk' f' (Vptr sp' Ptrofs.zero) pc' rs' m')
| match_return_states: forall stk v m stk' v' m' F
(MS: match_stacks F m m' stk stk' (Mem.nextblock m'))
(VINJ: Val.inject F v v')
@@ -904,10 +904,10 @@ Inductive match_states: RTL.state -> RTL.state -> Prop :=
(MINJ: Mem.inject F m m')
(VB: Mem.valid_block m' sp')
(PRIV: range_private F m m' sp' ctx.(dstk) f'.(fn_stacksize))
- (SSZ1: 0 <= f'.(fn_stacksize) < Int.max_unsigned)
+ (SSZ1: 0 <= f'.(fn_stacksize) < Ptrofs.max_unsigned)
(SSZ2: forall ofs, Mem.perm m' sp' ofs Max Nonempty -> 0 <= ofs <= f'.(fn_stacksize)),
match_states (Returnstate stk v m)
- (State stk' f' (Vptr sp' Int.zero) pc' rs' m').
+ (State stk' f' (Vptr sp' Ptrofs.zero) pc' rs' m').
(** ** Forward simulation *)
@@ -964,7 +964,7 @@ Proof.
eauto.
fold (saddr ctx addr). intros [a' [P Q]].
exploit Mem.loadv_inject; eauto. intros [v' [U V]].
- assert (eval_addressing tge (Vptr sp' Int.zero) (saddr ctx addr) rs' ## (sregs ctx args) = Some a').
+ assert (eval_addressing tge (Vptr sp' Ptrofs.zero) (saddr ctx addr) rs' ## (sregs ctx args) = Some a').
rewrite <- P. apply eval_addressing_preserved. exact symbols_preserved.
left; econstructor; split.
eapply plus_one. eapply exec_Iload; eauto.
@@ -982,7 +982,7 @@ Proof.
fold saddr. intros [a' [P Q]].
exploit Mem.storev_mapped_inject; eauto. eapply agree_val_reg; eauto.
intros [m1' [U V]].
- assert (eval_addressing tge (Vptr sp' Int.zero) (saddr ctx addr) rs' ## (sregs ctx args) = Some a').
+ assert (eval_addressing tge (Vptr sp' Ptrofs.zero) (saddr ctx addr) rs' ## (sregs ctx args) = Some a').
rewrite <- P. apply eval_addressing_preserved. exact symbols_preserved.
left; econstructor; split.
eapply plus_one. eapply exec_Istore; eauto.
diff --git a/backend/Inliningspec.v b/backend/Inliningspec.v
index f56d6d18..331f8b06 100644
--- a/backend/Inliningspec.v
+++ b/backend/Inliningspec.v
@@ -693,7 +693,7 @@ Inductive tr_function: program -> function -> function -> Prop :=
f'.(fn_sig) = f.(fn_sig) ->
f'.(fn_params) = sregs ctx f.(fn_params) ->
f'.(fn_entrypoint) = spc ctx f.(fn_entrypoint) ->
- 0 <= fn_stacksize f' < Int.max_unsigned ->
+ 0 <= fn_stacksize f' < Ptrofs.max_unsigned ->
tr_function p f f'.
Lemma tr_function_linkorder:
@@ -713,7 +713,7 @@ Proof.
intros. unfold transf_function in H.
set (fenv := funenv_program cunit) in *.
destruct (expand_function fenv f initstate) as [ctx s i] eqn:?.
- destruct (zlt (st_stksize s) Int.max_unsigned); inv H.
+ destruct (zlt (st_stksize s) Ptrofs.max_unsigned); inv H.
monadInv Heqr. set (ctx := initcontext x x0 (max_reg_function f) (fn_stacksize f)) in *.
Opaque initstate.
destruct INCR3. inversion EQ1. inversion EQ.
diff --git a/backend/LTL.v b/backend/LTL.v
index 5f7116ae..8567a891 100644
--- a/backend/LTL.v
+++ b/backend/LTL.v
@@ -15,18 +15,9 @@
LTL (``Location Transfer Language'') is the target language
for register allocation and the source language for linearization. *)
-Require Import Coqlib.
-Require Import Maps.
-Require Import AST.
-Require Import Integers.
-Require Import Values.
-Require Import Events.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Smallstep.
-Require Import Op.
-Require Import Locations.
-Require Import Conventions.
+Require Import Coqlib Maps.
+Require Import AST Integers Values Events Memory Globalenvs Smallstep.
+Require Import Op Locations Conventions.
(** * Abstract syntax *)
@@ -233,7 +224,7 @@ Inductive step: state -> trace -> state -> Prop :=
find_function ros rs' = Some fd ->
funsig fd = sig ->
Mem.free m sp 0 f.(fn_stacksize) = Some m' ->
- step (Block s f (Vptr sp Int.zero) (Ltailcall sig ros :: bb) rs m)
+ step (Block s f (Vptr sp Ptrofs.zero) (Ltailcall sig ros :: bb) rs m)
E0 (Callstate s fd rs' m')
| exec_Lbuiltin: forall s f sp ef args res bb rs m vargs t vres rs' m',
eval_builtin_args ge rs sp m args vargs ->
@@ -258,13 +249,13 @@ Inductive step: state -> trace -> state -> Prop :=
E0 (State s f sp pc rs' m)
| exec_Lreturn: forall s f sp bb rs m m',
Mem.free m sp 0 f.(fn_stacksize) = Some m' ->
- step (Block s f (Vptr sp Int.zero) (Lreturn :: bb) rs m)
+ step (Block s f (Vptr sp Ptrofs.zero) (Lreturn :: bb) rs m)
E0 (Returnstate s (return_regs (parent_locset s) rs) m')
| exec_function_internal: forall s f rs m m' sp rs',
Mem.alloc m 0 f.(fn_stacksize) = (m', sp) ->
rs' = undef_regs destroyed_at_function_entry (call_regs rs) ->
step (Callstate s (Internal f) rs m)
- E0 (State s f (Vptr sp Int.zero) f.(fn_entrypoint) rs' m')
+ E0 (State s f (Vptr sp Ptrofs.zero) f.(fn_entrypoint) rs' m')
| exec_function_external: forall s ef t args res rs m rs' m',
args = map (fun p => Locmap.getpair p rs) (loc_arguments (ef_sig ef)) ->
external_call ef ge args m t res m' ->
diff --git a/backend/Linear.v b/backend/Linear.v
index da1b4c04..55f92d16 100644
--- a/backend/Linear.v
+++ b/backend/Linear.v
@@ -17,17 +17,8 @@
instructions with explicit labels and ``goto'' instructions. *)
Require Import Coqlib.
-Require Import AST.
-Require Import Integers.
-Require Import Values.
-Require Import Memory.
-Require Import Events.
-Require Import Globalenvs.
-Require Import Smallstep.
-Require Import Op.
-Require Import Locations.
-Require Import LTL.
-Require Import Conventions.
+Require Import AST Integers Values Memory Events Globalenvs Smallstep.
+Require Import Op Locations LTL Conventions.
(** * Abstract syntax *)
@@ -194,7 +185,7 @@ Inductive step: state -> trace -> state -> Prop :=
find_function ros rs' = Some f' ->
sig = funsig f' ->
Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step (State s f (Vptr stk Int.zero) (Ltailcall sig ros :: b) rs m)
+ step (State s f (Vptr stk Ptrofs.zero) (Ltailcall sig ros :: b) rs m)
E0 (Callstate s f' rs' m')
| exec_Lbuiltin:
forall s f sp rs m ef args res b vargs t vres rs' m',
@@ -236,14 +227,14 @@ Inductive step: state -> trace -> state -> Prop :=
| exec_Lreturn:
forall s f stk b rs m m',
Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step (State s f (Vptr stk Int.zero) (Lreturn :: b) rs m)
+ step (State s f (Vptr stk Ptrofs.zero) (Lreturn :: b) rs m)
E0 (Returnstate s (return_regs (parent_locset s) rs) m')
| exec_function_internal:
forall s f rs m rs' m' stk,
Mem.alloc m 0 f.(fn_stacksize) = (m', stk) ->
rs' = undef_regs destroyed_at_function_entry (call_regs rs) ->
step (Callstate s (Internal f) rs m)
- E0 (State s f (Vptr stk Int.zero) f.(fn_code) rs' m')
+ E0 (State s f (Vptr stk Ptrofs.zero) f.(fn_code) rs' m')
| exec_function_external:
forall s ef args res rs1 rs2 m t m',
args = map (fun p => Locmap.getpair p rs1) (loc_arguments (ef_sig ef)) ->
diff --git a/backend/Lineartyping.v b/backend/Lineartyping.v
index 123c6b5a..e13ffb40 100644
--- a/backend/Lineartyping.v
+++ b/backend/Lineartyping.v
@@ -164,7 +164,7 @@ Proof.
intros. generalize (loc_result_pair sg) (loc_result_type sg).
destruct (loc_result sg); simpl Locmap.setpair.
- intros. apply wt_setreg; auto. eapply Val.has_subtype; eauto.
-- intros (A & B & C & D) E.
+- intros (A & B & C & D & E) F.
apply wt_setreg. eapply Val.has_subtype; eauto. destruct v; exact I.
apply wt_setreg. eapply Val.has_subtype; eauto. destruct v; exact I.
auto.
@@ -267,6 +267,7 @@ Qed.
Theorem step_type_preservation:
forall S1 t S2, step ge S1 t S2 -> wt_state S1 -> wt_state S2.
Proof.
+Local Opaque mreg_type.
induction 1; intros WTS; inv WTS.
- (* getstack *)
simpl in *; InvBooleans.
diff --git a/backend/Mach.v b/backend/Mach.v
index 3e15c97c..212088f3 100644
--- a/backend/Mach.v
+++ b/backend/Mach.v
@@ -52,9 +52,9 @@ Require Stacklayout.
Definition label := positive.
Inductive instruction: Type :=
- | Mgetstack: int -> typ -> mreg -> instruction
- | Msetstack: mreg -> int -> typ -> instruction
- | Mgetparam: int -> typ -> mreg -> instruction
+ | Mgetstack: ptrofs -> typ -> mreg -> instruction
+ | Msetstack: mreg -> ptrofs -> typ -> instruction
+ | Mgetparam: ptrofs -> typ -> mreg -> instruction
| Mop: operation -> list mreg -> mreg -> instruction
| Mload: memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Mstore: memory_chunk -> addressing -> list mreg -> mreg -> instruction
@@ -73,8 +73,8 @@ Record function: Type := mkfunction
{ fn_sig: signature;
fn_code: code;
fn_stacksize: Z;
- fn_link_ofs: int;
- fn_retaddr_ofs: int }.
+ fn_link_ofs: ptrofs;
+ fn_retaddr_ofs: ptrofs }.
Definition fundef := AST.fundef function.
@@ -118,11 +118,11 @@ value of the return address that the Asm code generated later will
store in the reserved location.
*)
-Definition load_stack (m: mem) (sp: val) (ty: typ) (ofs: int) :=
- Mem.loadv (chunk_of_type ty) m (Val.add sp (Vint ofs)).
+Definition load_stack (m: mem) (sp: val) (ty: typ) (ofs: ptrofs) :=
+ Mem.loadv (chunk_of_type ty) m (Val.offset_ptr sp ofs).
-Definition store_stack (m: mem) (sp: val) (ty: typ) (ofs: int) (v: val) :=
- Mem.storev (chunk_of_type ty) m (Val.add sp (Vint ofs)) v.
+Definition store_stack (m: mem) (sp: val) (ty: typ) (ofs: ptrofs) (v: val) :=
+ Mem.storev (chunk_of_type ty) m (Val.offset_ptr sp ofs) v.
Module RegEq.
Definition t := mreg.
@@ -198,7 +198,7 @@ Qed.
Section RELSEM.
-Variable return_address_offset: function -> code -> int -> Prop.
+Variable return_address_offset: function -> code -> ptrofs -> Prop.
Variable ge: genv.
@@ -207,7 +207,7 @@ Definition find_function_ptr
match ros with
| inl r =>
match rs r with
- | Vptr b ofs => if Int.eq ofs Int.zero then Some b else None
+ | Vptr b ofs => if Ptrofs.eq ofs Ptrofs.zero then Some b else None
| _ => None
end
| inr symb =>
@@ -220,7 +220,7 @@ Inductive extcall_arg (rs: regset) (m: mem) (sp: val): loc -> val -> Prop :=
| extcall_arg_reg: forall r,
extcall_arg rs m sp (R r) (rs r)
| extcall_arg_stack: forall ofs ty v,
- load_stack m sp ty (Int.repr (Stacklayout.fe_ofs_arg + 4 * ofs)) = Some v ->
+ load_stack m sp ty (Ptrofs.repr (Stacklayout.fe_ofs_arg + 4 * ofs)) = Some v ->
extcall_arg rs m sp (S Outgoing ofs ty) v.
Inductive extcall_arg_pair (rs: regset) (m: mem) (sp: val): rpair loc -> val -> Prop :=
@@ -271,13 +271,13 @@ Inductive state: Type :=
Definition parent_sp (s: list stackframe) : val :=
match s with
- | nil => Vzero
+ | nil => Vnullptr
| Stackframe f sp ra c :: s' => sp
end.
Definition parent_ra (s: list stackframe) : val :=
match s with
- | nil => Vzero
+ | nil => Vnullptr
| Stackframe f sp ra c :: s' => ra
end.
@@ -300,7 +300,7 @@ Inductive step: state -> trace -> state -> Prop :=
| exec_Mgetparam:
forall s fb f sp ofs ty dst c rs m v rs',
Genv.find_funct_ptr ge fb = Some (Internal f) ->
- load_stack m sp Tint f.(fn_link_ofs) = Some (parent_sp s) ->
+ load_stack m sp Tptr f.(fn_link_ofs) = Some (parent_sp s) ->
load_stack m (parent_sp s) ty ofs = Some v ->
rs' = (rs # temp_for_parent_frame <- Vundef # dst <- v) ->
step (State s fb sp (Mgetparam ofs ty dst :: c) rs m)
@@ -337,8 +337,8 @@ Inductive step: state -> trace -> state -> Prop :=
forall s fb stk soff sig ros c rs m f f' m',
find_function_ptr ge ros rs = Some f' ->
Genv.find_funct_ptr ge fb = Some (Internal f) ->
- load_stack m (Vptr stk soff) Tint f.(fn_link_ofs) = Some (parent_sp s) ->
- load_stack m (Vptr stk soff) Tint f.(fn_retaddr_ofs) = Some (parent_ra s) ->
+ load_stack m (Vptr stk soff) Tptr f.(fn_link_ofs) = Some (parent_sp s) ->
+ load_stack m (Vptr stk soff) Tptr f.(fn_retaddr_ofs) = Some (parent_ra s) ->
Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
step (State s fb (Vptr stk soff) (Mtailcall sig ros :: c) rs m)
E0 (Callstate s f' rs m')
@@ -381,8 +381,8 @@ Inductive step: state -> trace -> state -> Prop :=
| exec_Mreturn:
forall s fb stk soff c rs m f m',
Genv.find_funct_ptr ge fb = Some (Internal f) ->
- load_stack m (Vptr stk soff) Tint f.(fn_link_ofs) = Some (parent_sp s) ->
- load_stack m (Vptr stk soff) Tint f.(fn_retaddr_ofs) = Some (parent_ra s) ->
+ load_stack m (Vptr stk soff) Tptr f.(fn_link_ofs) = Some (parent_sp s) ->
+ load_stack m (Vptr stk soff) Tptr f.(fn_retaddr_ofs) = Some (parent_ra s) ->
Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
step (State s fb (Vptr stk soff) (Mreturn :: c) rs m)
E0 (Returnstate s rs m')
@@ -390,9 +390,9 @@ Inductive step: state -> trace -> state -> Prop :=
forall s fb rs m f m1 m2 m3 stk rs',
Genv.find_funct_ptr ge fb = Some (Internal f) ->
Mem.alloc m 0 f.(fn_stacksize) = (m1, stk) ->
- let sp := Vptr stk Int.zero in
- store_stack m1 sp Tint f.(fn_link_ofs) (parent_sp s) = Some m2 ->
- store_stack m2 sp Tint f.(fn_retaddr_ofs) (parent_ra s) = Some m3 ->
+ let sp := Vptr stk Ptrofs.zero in
+ store_stack m1 sp Tptr f.(fn_link_ofs) (parent_sp s) = Some m2 ->
+ store_stack m2 sp Tptr f.(fn_retaddr_ofs) (parent_ra s) = Some m3 ->
rs' = undef_regs destroyed_at_function_entry rs ->
step (Callstate s fb rs m)
E0 (State s fb sp f.(fn_code) rs' m3)
@@ -424,5 +424,5 @@ Inductive final_state: state -> int -> Prop :=
rs r = Vint retcode ->
final_state (Returnstate nil rs m) retcode.
-Definition semantics (rao: function -> code -> int -> Prop) (p: program) :=
+Definition semantics (rao: function -> code -> ptrofs -> Prop) (p: program) :=
Semantics (step rao) (initial_state p) final_state (Genv.globalenv p).
diff --git a/backend/NeedDomain.v b/backend/NeedDomain.v
index 442352e7..a53040f9 100644
--- a/backend/NeedDomain.v
+++ b/backend/NeedDomain.v
@@ -379,6 +379,7 @@ Ltac InvAgree :=
match goal with
| [ H: False |- _ ] => contradiction
| [ H: match ?v with Vundef => _ | Vint _ => _ | Vlong _ => _ | Vfloat _ => _ | Vsingle _ => _ | Vptr _ _ => _ end |- _ ] => destruct v
+ | [ |- context [if Archi.ptr64 then _ else _] ] => destruct Archi.ptr64 eqn:?
end).
(** And immediate, or immediate *)
@@ -608,7 +609,8 @@ Lemma add_sound:
Proof.
unfold modarith; intros. destruct x; simpl in *.
- auto.
-- unfold Val.add; InvAgree. apply eqmod_iagree. apply Int.eqmod_add; apply iagree_eqmod; auto.
+- unfold Val.add; InvAgree.
+ apply eqmod_iagree. apply Int.eqmod_add; apply iagree_eqmod; auto.
- inv H; auto. inv H0; auto. destruct w1; auto.
Qed.
@@ -802,20 +804,20 @@ Hypothesis PERM: forall b ofs k p, Mem.perm m1 b ofs k p -> Mem.perm m2 b ofs k
Let valid_pointer_inj:
forall b1 ofs b2 delta,
inject_id b1 = Some(b2, delta) ->
- Mem.valid_pointer m1 b1 (Int.unsigned ofs) = true ->
- Mem.valid_pointer m2 b2 (Int.unsigned (Int.add ofs (Int.repr delta))) = true.
+ Mem.valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
+ Mem.valid_pointer m2 b2 (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr delta))) = true.
Proof.
- unfold inject_id; intros. inv H. rewrite Int.add_zero.
+ unfold inject_id; intros. inv H. rewrite Ptrofs.add_zero.
rewrite Mem.valid_pointer_nonempty_perm in *. eauto.
Qed.
Let weak_valid_pointer_inj:
forall b1 ofs b2 delta,
inject_id b1 = Some(b2, delta) ->
- Mem.weak_valid_pointer m1 b1 (Int.unsigned ofs) = true ->
- Mem.weak_valid_pointer m2 b2 (Int.unsigned (Int.add ofs (Int.repr delta))) = true.
+ Mem.weak_valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
+ Mem.weak_valid_pointer m2 b2 (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr delta))) = true.
Proof.
- unfold inject_id; intros. inv H. rewrite Int.add_zero.
+ unfold inject_id; intros. inv H. rewrite Ptrofs.add_zero.
rewrite Mem.weak_valid_pointer_spec in *.
rewrite ! Mem.valid_pointer_nonempty_perm in *.
destruct H0; [left|right]; eauto.
@@ -824,21 +826,21 @@ Qed.
Let weak_valid_pointer_no_overflow:
forall b1 ofs b2 delta,
inject_id b1 = Some(b2, delta) ->
- Mem.weak_valid_pointer m1 b1 (Int.unsigned ofs) = true ->
- 0 <= Int.unsigned ofs + Int.unsigned (Int.repr delta) <= Int.max_unsigned.
+ Mem.weak_valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
+ 0 <= Ptrofs.unsigned ofs + Ptrofs.unsigned (Ptrofs.repr delta) <= Ptrofs.max_unsigned.
Proof.
- unfold inject_id; intros. inv H. rewrite Zplus_0_r. apply Int.unsigned_range_2.
+ unfold inject_id; intros. inv H. rewrite Zplus_0_r. apply Ptrofs.unsigned_range_2.
Qed.
Let valid_different_pointers_inj:
forall b1 ofs1 b2 ofs2 b1' delta1 b2' delta2,
b1 <> b2 ->
- Mem.valid_pointer m1 b1 (Int.unsigned ofs1) = true ->
- Mem.valid_pointer m1 b2 (Int.unsigned ofs2) = true ->
+ Mem.valid_pointer m1 b1 (Ptrofs.unsigned ofs1) = true ->
+ Mem.valid_pointer m1 b2 (Ptrofs.unsigned ofs2) = true ->
inject_id b1 = Some (b1', delta1) ->
inject_id b2 = Some (b2', delta2) ->
b1' <> b2' \/
- Int.unsigned (Int.add ofs1 (Int.repr delta1)) <> Int.unsigned (Int.add ofs2 (Int.repr delta2)).
+ Ptrofs.unsigned (Ptrofs.add ofs1 (Ptrofs.repr delta1)) <> Ptrofs.unsigned (Ptrofs.add ofs2 (Ptrofs.repr delta2)).
Proof.
unfold inject_id; intros. left; congruence.
Qed.
@@ -855,13 +857,13 @@ Qed.
Lemma default_needs_of_operation_sound:
forall op args1 v1 args2 nv,
- eval_operation ge (Vptr sp Int.zero) op args1 m1 = Some v1 ->
+ eval_operation ge (Vptr sp Ptrofs.zero) op args1 m1 = Some v1 ->
vagree_list args1 args2 nil
\/ vagree_list args1 args2 (default nv :: nil)
\/ vagree_list args1 args2 (default nv :: default nv :: nil) ->
nv <> Nothing ->
exists v2,
- eval_operation ge (Vptr sp Int.zero) op args2 m2 = Some v2
+ eval_operation ge (Vptr sp Ptrofs.zero) op args2 m2 = Some v2
/\ vagree v1 v2 nv.
Proof.
intros. assert (default nv = All) by (destruct nv; simpl; congruence).
@@ -875,7 +877,7 @@ Proof.
exploit (@eval_operation_inj _ _ _ _ ge ge inject_id).
eassumption. auto. auto. auto.
instantiate (1 := op). intros. apply val_inject_lessdef; auto.
- apply val_inject_lessdef. instantiate (1 := Vptr sp Int.zero). instantiate (1 := Vptr sp Int.zero). auto.
+ apply val_inject_lessdef. instantiate (1 := Vptr sp Ptrofs.zero). instantiate (1 := Vptr sp Ptrofs.zero). auto.
apply val_inject_list_lessdef; eauto.
eauto.
intros (v2 & A & B). exists v2; split; auto.
@@ -1135,13 +1137,13 @@ Definition nmem_add (nm: nmem) (p: aptr) (sz: Z) : nmem :=
match p with
| Gl id ofs =>
match gl!id with
- | Some iv => NMem stk (PTree.set id (ISet.remove (Int.unsigned ofs) (Int.unsigned ofs + sz) iv) gl)
+ | Some iv => NMem stk (PTree.set id (ISet.remove (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) iv) gl)
| None => nm
end
| Glo id =>
NMem stk (PTree.remove id gl)
| Stk ofs =>
- NMem (ISet.remove (Int.unsigned ofs) (Int.unsigned ofs + sz) stk) gl
+ NMem (ISet.remove (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) stk) gl
| Stack =>
NMem ISet.empty gl
| _ => nmem_all
@@ -1153,7 +1155,7 @@ Lemma nlive_add:
genv_match bc ge ->
bc sp = BCstack ->
pmatch bc b ofs p ->
- Int.unsigned ofs <= i < Int.unsigned ofs + sz ->
+ Ptrofs.unsigned ofs <= i < Ptrofs.unsigned ofs + sz ->
nlive (nmem_add nm p sz) b i.
Proof.
intros. unfold nmem_add. destruct nm. apply nlive_all.
@@ -1221,12 +1223,12 @@ Definition nmem_remove (nm: nmem) (p: aptr) (sz: Z) : nmem :=
| Gl id ofs =>
let iv' :=
match gl!id with
- | Some iv => ISet.add (Int.unsigned ofs) (Int.unsigned ofs + sz) iv
- | None => ISet.interval (Int.unsigned ofs) (Int.unsigned ofs + sz)
+ | Some iv => ISet.add (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) iv
+ | None => ISet.interval (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz)
end in
NMem stk (PTree.set id iv' gl)
| Stk ofs =>
- NMem (ISet.add (Int.unsigned ofs) (Int.unsigned ofs + sz) stk) gl
+ NMem (ISet.add (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) stk) gl
| _ => nm
end
end.
@@ -1237,17 +1239,17 @@ Lemma nlive_remove:
bc sp = BCstack ->
pmatch bc b ofs p ->
nlive nm b' i ->
- b' <> b \/ i < Int.unsigned ofs \/ Int.unsigned ofs + sz <= i ->
+ b' <> b \/ i < Ptrofs.unsigned ofs \/ Ptrofs.unsigned ofs + sz <= i ->
nlive (nmem_remove nm p sz) b' i.
Proof.
intros. inversion H2; subst. unfold nmem_remove; inv H1; auto.
- (* Gl id ofs *)
set (iv' := match gl!id with
| Some iv =>
- ISet.add (Int.unsigned ofs) (Int.unsigned ofs + sz) iv
+ ISet.add (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) iv
| None =>
- ISet.interval (Int.unsigned ofs)
- (Int.unsigned ofs + sz)
+ ISet.interval (Ptrofs.unsigned ofs)
+ (Ptrofs.unsigned ofs + sz)
end).
assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto).
split; simpl; auto; intros.
@@ -1272,11 +1274,11 @@ Definition nmem_contains (nm: nmem) (p: aptr) (sz: Z) :=
match p with
| Gl id ofs =>
match gl!id with
- | Some iv => negb (ISet.contains (Int.unsigned ofs) (Int.unsigned ofs + sz) iv)
+ | Some iv => negb (ISet.contains (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) iv)
| None => true
end
| Stk ofs =>
- negb (ISet.contains (Int.unsigned ofs) (Int.unsigned ofs + sz) stk)
+ negb (ISet.contains (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) stk)
| _ => true (**r conservative answer *)
end
end.
@@ -1287,7 +1289,7 @@ Lemma nlive_contains:
bc sp = BCstack ->
pmatch bc b ofs p ->
nmem_contains nm p sz = false ->
- Int.unsigned ofs <= i < Int.unsigned ofs + sz ->
+ Ptrofs.unsigned ofs <= i < Ptrofs.unsigned ofs + sz ->
~(nlive nm b i).
Proof.
unfold nmem_contains; intros. red; intros L; inv L.
@@ -1295,10 +1297,10 @@ Proof.
- (* Gl id ofs *)
assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto).
destruct gl!id as [iv|] eqn:HG; inv H2.
- destruct (ISet.contains (Int.unsigned ofs) (Int.unsigned ofs + sz) iv) eqn:IC; try discriminate.
+ destruct (ISet.contains (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) iv) eqn:IC; try discriminate.
rewrite ISet.contains_spec in IC. eelim GL; eauto.
- (* Stk ofs *)
- destruct (ISet.contains (Int.unsigned ofs) (Int.unsigned ofs + sz) stk) eqn:IC; try discriminate.
+ destruct (ISet.contains (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) stk) eqn:IC; try discriminate.
rewrite ISet.contains_spec in IC. eelim STK; eauto. eapply bc_stack; eauto.
Qed.
diff --git a/backend/PrintAsmaux.ml b/backend/PrintAsmaux.ml
index 148c5300..b220659c 100644
--- a/backend/PrintAsmaux.ml
+++ b/backend/PrintAsmaux.ml
@@ -138,6 +138,9 @@ let cfi_section =
let coqint oc n =
fprintf oc "%ld" (camlint_of_coqint n)
+let coqint64 oc n =
+ fprintf oc "%Ld" (camlint64_of_coqint n)
+
(** Programmer-supplied annotations (__builtin_annot). *)
let re_annot_param = Str.regexp "%%\\|%[1-9][0-9]*"
diff --git a/backend/RTL.v b/backend/RTL.v
index a39d37cb..d191918c 100644
--- a/backend/RTL.v
+++ b/backend/RTL.v
@@ -16,17 +16,9 @@
intermediate language after Cminor and CminorSel.
*)
-Require Import Coqlib.
-Require Import Maps.
-Require Import AST.
-Require Import Integers.
-Require Import Values.
-Require Import Events.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Smallstep.
-Require Import Op.
-Require Import Registers.
+Require Import Coqlib Maps.
+Require Import AST Integers Values Events Memory Globalenvs Smallstep.
+Require Import Op Registers.
(** * Abstract syntax *)
@@ -246,7 +238,7 @@ Inductive step: state -> trace -> state -> Prop :=
find_function ros rs = Some fd ->
funsig fd = sig ->
Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step (State s f (Vptr stk Int.zero) pc rs m)
+ step (State s f (Vptr stk Ptrofs.zero) pc rs m)
E0 (Callstate s fd rs##args m')
| exec_Ibuiltin:
forall s f sp pc rs m ef args res pc' vargs t vres m',
@@ -273,7 +265,7 @@ Inductive step: state -> trace -> state -> Prop :=
forall s f stk pc rs m or m',
(fn_code f)!pc = Some(Ireturn or) ->
Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step (State s f (Vptr stk Int.zero) pc rs m)
+ step (State s f (Vptr stk Ptrofs.zero) pc rs m)
E0 (Returnstate s (regmap_optget or Vundef rs) m')
| exec_function_internal:
forall s f args m m' stk,
@@ -281,7 +273,7 @@ Inductive step: state -> trace -> state -> Prop :=
step (Callstate s (Internal f) args m)
E0 (State s
f
- (Vptr stk Int.zero)
+ (Vptr stk Ptrofs.zero)
f.(fn_entrypoint)
(init_regs args f.(fn_params))
m')
diff --git a/backend/RTLtyping.v b/backend/RTLtyping.v
index dec1b988..f9f01d49 100644
--- a/backend/RTLtyping.v
+++ b/backend/RTLtyping.v
@@ -73,9 +73,9 @@ Definition type_of_builtin_arg (a: builtin_arg reg) : typ :=
| BA_float _ => Tfloat
| BA_single _ => Tsingle
| BA_loadstack chunk ofs => type_of_chunk chunk
- | BA_addrstack ofs => Tint
+ | BA_addrstack ofs => Tptr
| BA_loadglobal chunk id ofs => type_of_chunk chunk
- | BA_addrglobal id ofs => Tint
+ | BA_addrglobal id ofs => Tptr
| BA_splitlong hi lo => Tlong
end.
@@ -116,14 +116,14 @@ Inductive wt_instr : instruction -> Prop :=
wt_instr (Istore chunk addr args src s)
| wt_Icall:
forall sig ros args res s,
- match ros with inl r => env r = Tint | inr s => True end ->
+ match ros with inl r => env r = Tptr | inr s => True end ->
map env args = sig.(sig_args) ->
env res = proj_sig_res sig ->
valid_successor s ->
wt_instr (Icall sig ros args res s)
| wt_Itailcall:
forall sig ros args,
- match ros with inl r => env r = Tint | inr s => True end ->
+ match ros with inl r => env r = Tptr | inr s => True end ->
map env args = sig.(sig_args) ->
sig.(sig_res) = funct.(fn_sig).(sig_res) ->
tailcall_possible sig ->
@@ -227,7 +227,7 @@ Fixpoint check_successors (sl: list node): res unit :=
Definition type_ros (e: S.typenv) (ros: reg + ident) : res S.typenv :=
match ros with
- | inl r => S.set e r Tint
+ | inl r => S.set e r Tptr
| inr s => OK e
end.
@@ -245,9 +245,9 @@ Definition type_builtin_arg (e: S.typenv) (a: builtin_arg reg) (ty: typ) : res S
| BA_float _ => type_expect e ty Tfloat
| BA_single _ => type_expect e ty Tsingle
| BA_loadstack chunk ofs => type_expect e ty (type_of_chunk chunk)
- | BA_addrstack ofs => type_expect e ty Tint
+ | BA_addrstack ofs => type_expect e ty Tptr
| BA_loadglobal chunk id ofs => type_expect e ty (type_of_chunk chunk)
- | BA_addrglobal id ofs => type_expect e ty Tint
+ | BA_addrglobal id ofs => type_expect e ty Tptr
| BA_splitlong hi lo => type_expect e ty Tlong
end.
@@ -367,7 +367,7 @@ Hint Resolve type_ros_incr: ty.
Lemma type_ros_sound:
forall e ros e' te, type_ros e ros = OK e' -> S.satisf te e' ->
- match ros with inl r => te r = Tint | inr s => True end.
+ match ros with inl r => te r = Tptr | inr s => True end.
Proof.
unfold type_ros; intros. destruct ros.
eapply S.set_sound; eauto.
@@ -594,7 +594,7 @@ Qed.
Lemma type_ros_complete:
forall te ros e,
S.satisf te e ->
- match ros with inl r => te r = Tint | inr s => True end ->
+ match ros with inl r => te r = Tptr | inr s => True end ->
exists e', type_ros e ros = OK e' /\ S.satisf te e'.
Proof.
intros; destruct ros; simpl.
diff --git a/backend/Regalloc.ml b/backend/Regalloc.ml
index b91bad27..200d0237 100644
--- a/backend/Regalloc.ml
+++ b/backend/Regalloc.ml
@@ -102,50 +102,61 @@ let parmove_regs2locs tyenv srcs dsts k =
assert (List.length srcs = List.length dsts);
let rec expand srcs' dsts' rl ll =
match rl, ll with
- | [], [] -> (srcs', dsts')
+ | [], [] ->
+ begin match srcs', dsts' with
+ | [], [] -> k
+ | [src], [dst] -> move src dst k
+ | _, _ -> Xparmove(srcs', dsts', new_temp Tint, new_temp Tfloat) :: k
+ end
| r :: rl, One l :: ll ->
let ty = tyenv r in
expand (V(r, ty) :: srcs') (L l :: dsts') rl ll
| r :: rl, Twolong(l1, l2) :: ll ->
assert (tyenv r = Tlong);
- expand (V(r, Tint) :: V(twin_reg r, Tint) :: srcs')
- (L l1 :: L l2 :: dsts')
- rl ll
+ if Archi.splitlong then
+ expand (V(r, Tint) :: V(twin_reg r, Tint) :: srcs')
+ (L l1 :: L l2 :: dsts')
+ rl ll
+ else
+ Xop(Ohighlong, [V(r, Tlong)], L l1) ::
+ Xop(Olowlong, [V(r, Tlong)], L l2) ::
+ expand srcs' dsts' rl ll
| _, _ ->
assert false in
- let (srcs', dsts') = expand [] [] srcs dsts in
- match srcs', dsts' with
- | [], [] -> k
- | [src], [dst] -> move src dst k
- | _, _ -> Xparmove(srcs', dsts', new_temp Tint, new_temp Tfloat) :: k
+ expand [] [] srcs dsts
let parmove_locs2regs tyenv srcs dsts k =
assert (List.length srcs = List.length dsts);
let rec expand srcs' dsts' ll rl =
match ll, rl with
- | [], [] -> (srcs', dsts')
+ | [], [] ->
+ begin match srcs', dsts' with
+ | [], [] -> k
+ | [src], [dst] -> move src dst k
+ | _, _ -> Xparmove(srcs', dsts', new_temp Tint, new_temp Tfloat) :: k
+ end
| One l :: ll, r :: rl ->
let ty = tyenv r in
expand (L l :: srcs') (V(r, ty) :: dsts') ll rl
| Twolong(l1, l2) :: ll, r :: rl ->
assert (tyenv r = Tlong);
- expand (L l1 :: L l2 :: srcs')
- (V(r, Tint) :: V(twin_reg r, Tint) :: dsts')
- ll rl
+ if Archi.splitlong then
+ expand (L l1 :: L l2 :: srcs')
+ (V(r, Tint) :: V(twin_reg r, Tint) :: dsts')
+ ll rl
+ else
+ Xop(Omakelong, [L l1; L l2], V(r, Tlong)) ::
+ expand srcs' dsts' ll rl
| _, _ ->
assert false in
- let (srcs', dsts') = expand [] [] srcs dsts in
- match srcs', dsts' with
- | [], [] -> k
- | [src], [dst] -> move src dst k
- | _, _ -> Xparmove(srcs', dsts', new_temp Tint, new_temp Tfloat) :: k
+ expand [] [] srcs dsts
let rec convert_builtin_arg tyenv = function
| BA r ->
- begin match tyenv r with
- | Tlong -> BA_splitlong(BA(V(r, Tint)), BA(V(twin_reg r, Tint)))
- | ty -> BA(V(r, ty))
- end
+ let ty = tyenv r in
+ if Archi.splitlong && ty = Tlong
+ then BA_splitlong(BA(V(r, Tint)), BA(V(twin_reg r, Tint)))
+ else BA(V(r, ty))
| BA_int n -> BA_int n
| BA_long n -> BA_long n
| BA_float n -> BA_float n
@@ -159,10 +170,10 @@ let rec convert_builtin_arg tyenv = function
let convert_builtin_res tyenv = function
| BR r ->
- begin match tyenv r with
- | Tlong -> BR_splitlong(BR(V(r, Tint)), BR(V(twin_reg r, Tint)))
- | ty -> BR(V(r, ty))
- end
+ let ty = tyenv r in
+ if Archi.splitlong && ty = Tlong
+ then BR_splitlong(BR(V(r, Tint)), BR(V(twin_reg r, Tint)))
+ else BR(V(r, ty))
| BR_none -> BR_none
| BR_splitlong _ -> assert false
@@ -197,25 +208,26 @@ let rec constrain_builtin_res a cl =
(* Return the XTL basic block corresponding to the given RTL instruction.
Move and parallel move instructions are introduced to honor calling
conventions and register constraints on some operations.
- 64-bit integer variables are split in two 32-bit halves. *)
+ 64-bit integer variables are split in two 32-bit halves
+ if [Archi.splitlong] is true. *)
let block_of_RTL_instr funsig tyenv = function
| RTL.Inop s ->
[Xbranch s]
| RTL.Iop(Omove, [arg], res, s) ->
- if tyenv arg = Tlong then
+ if Archi.splitlong && tyenv arg = Tlong then
[Xmove(V(arg, Tint), V(res, Tint));
Xmove(V(twin_reg arg, Tint), V(twin_reg res, Tint));
Xbranch s]
else
[Xmove(vreg tyenv arg, vreg tyenv res); Xbranch s]
- | RTL.Iop(Omakelong, [arg1; arg2], res, s) ->
+ | RTL.Iop(Omakelong, [arg1; arg2], res, s) when Archi.splitlong ->
[Xmove(V(arg1, Tint), V(res, Tint));
Xmove(V(arg2, Tint), V(twin_reg res, Tint));
Xbranch s]
- | RTL.Iop(Olowlong, [arg], res, s) ->
+ | RTL.Iop(Olowlong, [arg], res, s) when Archi.splitlong ->
[Xmove(V(twin_reg arg, Tint), V(res, Tint)); Xbranch s]
- | RTL.Iop(Ohighlong, [arg], res, s) ->
+ | RTL.Iop(Ohighlong, [arg], res, s) when Archi.splitlong ->
[Xmove(V(arg, Tint), V(res, Tint)); Xbranch s]
| RTL.Iop(op, args, res, s) ->
let (cargs, cres) = mregs_for_operation op in
@@ -232,7 +244,7 @@ let block_of_RTL_instr funsig tyenv = function
let t = new_temp (tyenv res) in (t :: args2', t) in
movelist args1 args3 (Xop(op, args3, res3) :: move res3 res1 [Xbranch s])
| RTL.Iload(chunk, addr, args, dst, s) ->
- if chunk = Mint64 then begin
+ if Archi.splitlong && chunk = Mint64 then begin
match offset_addressing addr (coqint_of_camlint 4l) with
| None -> assert false
| Some addr' ->
@@ -244,7 +256,7 @@ let block_of_RTL_instr funsig tyenv = function
end else
[Xload(chunk, addr, vregs tyenv args, vreg tyenv dst); Xbranch s]
| RTL.Istore(chunk, addr, args, src, s) ->
- if chunk = Mint64 then begin
+ if Archi.splitlong && chunk = Mint64 then begin
match offset_addressing addr (coqint_of_camlint 4l) with
| None -> assert false
| Some addr' ->
@@ -1024,10 +1036,8 @@ let make_parmove srcs dsts itmp ftmp k =
let n = Array.length src in
assert (Array.length dst = n);
let status = Array.make n To_move in
- let temp_for =
- function (Tint|Tany32) -> itmp
- | (Tfloat|Tsingle|Tany64) -> ftmp
- | Tlong -> assert false in
+ let temp_for cls =
+ match cls with 0 -> itmp | 1 -> ftmp | _ -> assert false in
let code = ref [] in
let add_move s d =
match s, d with
@@ -1038,7 +1048,7 @@ let make_parmove srcs dsts itmp ftmp k =
| Locations.S(sl, ofs, ty), R rd ->
code := LTL.Lgetstack(sl, ofs, ty, rd) :: !code
| Locations.S(sls, ofss, tys), Locations.S(sld, ofsd, tyd) ->
- let tmp = temp_for tys in
+ let tmp = temp_for (class_of_type tys) in
(* code will be reversed at the end *)
code := LTL.Lsetstack(tmp, sld, ofsd, tyd) ::
LTL.Lgetstack(sls, ofss, tys, tmp) :: !code
@@ -1052,7 +1062,7 @@ let make_parmove srcs dsts itmp ftmp k =
| To_move ->
move_one j
| Being_moved ->
- let tmp = R (temp_for (Loc.coq_type src.(j))) in
+ let tmp = R (temp_for (class_of_loc src.(j))) in
add_move src.(j) tmp;
src.(j) <- tmp
| Moved ->
diff --git a/backend/SelectDiv.vp b/backend/SelectDiv.vp
index a275a850..d708afb7 100644
--- a/backend/SelectDiv.vp
+++ b/backend/SelectDiv.vp
@@ -36,7 +36,7 @@ Fixpoint find_div_mul_params (fuel: nat) (nc: Z) (d: Z) (p: Z) : option (Z * Z)
| S fuel' =>
let twp := two_p p in
if zlt (nc * (d - twp mod d)) twp
- then Some(p - 32, (twp + d - twp mod d) / d)
+ then Some(p, (twp + d - twp mod d) / d)
else find_div_mul_params fuel' nc d (p + 1)
end.
@@ -47,6 +47,7 @@ Definition divs_mul_params (d: Z) : option (Z * Z) :=
d 32 with
| None => None
| Some(p, m) =>
+ let p := p - 32 in
if zlt 0 d
&& zlt (two_p (32 + p)) (m * d)
&& zle (m * d) (two_p (32 + p) + two_p (p + 1))
@@ -62,6 +63,7 @@ Definition divu_mul_params (d: Z) : option (Z * Z) :=
d 32 with
| None => None
| Some(p, m) =>
+ let p := p - 32 in
if zlt 0 d
&& zle (two_p (32 + p)) (m * d)
&& zle (m * d) (two_p (32 + p) + two_p p)
@@ -70,6 +72,38 @@ Definition divu_mul_params (d: Z) : option (Z * Z) :=
then Some(p, m) else None
end.
+Definition divls_mul_params (d: Z) : option (Z * Z) :=
+ match find_div_mul_params
+ Int64.wordsize
+ (Int64.half_modulus - Int64.half_modulus mod d - 1)
+ d 64 with
+ | None => None
+ | Some(p, m) =>
+ let p := p - 64 in
+ if zlt 0 d
+ && zlt (two_p (64 + p)) (m * d)
+ && zle (m * d) (two_p (64 + p) + two_p (p + 1))
+ && zle 0 m && zlt m Int64.modulus
+ && zle 0 p && zlt p 64
+ then Some(p, m) else None
+ end.
+
+Definition divlu_mul_params (d: Z) : option (Z * Z) :=
+ match find_div_mul_params
+ Int64.wordsize
+ (Int64.modulus - Int64.modulus mod d - 1)
+ d 64 with
+ | None => None
+ | Some(p, m) =>
+ let p := p - 64 in
+ if zlt 0 d
+ && zle (two_p (64 + p)) (m * d)
+ && zle (m * d) (two_p (64 + p) + two_p p)
+ && zle 0 m && zlt m Int64.modulus
+ && zle 0 p && zlt p 64
+ then Some(p, m) else None
+ end.
+
Definition divu_mul (p: Z) (m: Z) :=
shruimm (Eop Omulhu (Eletvar O ::: Eop (Ointconst (Int.repr m)) Enil ::: Enil))
(Int.repr p).
diff --git a/backend/SelectDivproof.v b/backend/SelectDivproof.v
index ffe607e4..5621acd5 100644
--- a/backend/SelectDivproof.v
+++ b/backend/SelectDivproof.v
@@ -12,21 +12,10 @@
(** Correctness of instruction selection for integer division *)
-Require Import Coqlib.
-Require Import Zquot.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Events.
-Require Import Cminor.
-Require Import Op.
-Require Import CminorSel.
-Require Import SelectOp.
-Require Import SelectOpproof.
-Require Import SelectDiv.
+Require Import Zquot Coqlib.
+Require Import AST Integers Floats Values Memory Globalenvs Events.
+Require Import Cminor Op CminorSel.
+Require Import SelectOp SelectOpproof SelectDiv.
Open Local Scope cminorsel_scope.
@@ -191,18 +180,19 @@ Lemma divs_mul_params_sound:
Int.min_signed <= n <= Int.max_signed ->
Z.quot n d = Zdiv (m * n) (two_p (32 + p)) + (if zlt n 0 then 1 else 0).
Proof with (try discriminate).
- unfold divs_mul_params; intros d m' p' EQ.
+ unfold divs_mul_params; intros d m' p'.
destruct (find_div_mul_params Int.wordsize
(Int.half_modulus - Int.half_modulus mod d - 1) d 32)
as [[p m] | ]...
+ generalize (p - 32). intro p1.
destruct (zlt 0 d)...
- destruct (zlt (two_p (32 + p)) (m * d))...
- destruct (zle (m * d) (two_p (32 + p) + two_p (p + 1)))...
+ destruct (zlt (two_p (32 + p1)) (m * d))...
+ destruct (zle (m * d) (two_p (32 + p1) + two_p (p1 + 1)))...
destruct (zle 0 m)...
destruct (zlt m Int.modulus)...
- destruct (zle 0 p)...
- destruct (zlt p 32)...
- simpl in EQ. inv EQ.
+ destruct (zle 0 p1)...
+ destruct (zlt p1 32)...
+ intros EQ; inv EQ.
split. auto. split. auto. intros.
replace (32 + p') with (31 + (p' + 1)) by omega.
apply Zquot_mul; try omega.
@@ -219,18 +209,19 @@ Lemma divu_mul_params_sound:
0 <= n < Int.modulus ->
Zdiv n d = Zdiv (m * n) (two_p (32 + p)).
Proof with (try discriminate).
- unfold divu_mul_params; intros d m' p' EQ.
+ unfold divu_mul_params; intros d m' p'.
destruct (find_div_mul_params Int.wordsize
(Int.modulus - Int.modulus mod d - 1) d 32)
as [[p m] | ]...
+ generalize (p - 32); intro p1.
destruct (zlt 0 d)...
- destruct (zle (two_p (32 + p)) (m * d))...
- destruct (zle (m * d) (two_p (32 + p) + two_p p))...
+ destruct (zle (two_p (32 + p1)) (m * d))...
+ destruct (zle (m * d) (two_p (32 + p1) + two_p p1))...
destruct (zle 0 m)...
destruct (zlt m Int.modulus)...
- destruct (zle 0 p)...
- destruct (zlt p 32)...
- simpl in EQ. inv EQ.
+ destruct (zle 0 p1)...
+ destruct (zlt p1 32)...
+ intros EQ; inv EQ.
split. auto. split. auto. intros.
apply Zdiv_mul_pos; try omega. assumption.
Qed.
diff --git a/backend/Selection.v b/backend/Selection.v
index 02b37c48..3aff446e 100644
--- a/backend/Selection.v
+++ b/backend/Selection.v
@@ -23,19 +23,11 @@
The source language is Cminor and the target language is CminorSel. *)
Require String.
-Require Import Coqlib.
-Require Import Maps.
-Require Import AST.
-Require Import Errors.
-Require Import Integers.
-Require Import Globalenvs.
-Require Import Switch.
+Require Import Coqlib Maps.
+Require Import AST Errors Integers Globalenvs Switch.
Require Cminor.
-Require Import Op.
-Require Import CminorSel.
-Require Import SelectOp.
-Require Import SelectDiv.
-Require Import SelectLong.
+Require Import Op CminorSel.
+Require Import SelectOp SelectDiv SplitLong SelectLong.
Require Machregs.
Local Open Scope cminorsel_scope.
@@ -71,7 +63,7 @@ Section SELECTION.
Definition globdef := AST.globdef Cminor.fundef unit.
Variable defmap: PTree.t globdef.
-Variable hf: helper_functions.
+Context {hf: helper_functions}.
Definition sel_constant (cst: Cminor.constant) : expr :=
match cst with
@@ -110,14 +102,14 @@ Definition sel_unop (op: Cminor.unary_operation) (arg: expr) : expr :=
| Cminor.Ointoflong => intoflong arg
| Cminor.Olongofint => longofint arg
| Cminor.Olongofintu => longofintu arg
- | Cminor.Olongoffloat => longoffloat hf arg
- | Cminor.Olonguoffloat => longuoffloat hf arg
- | Cminor.Ofloatoflong => floatoflong hf arg
- | Cminor.Ofloatoflongu => floatoflongu hf arg
- | Cminor.Olongofsingle => longofsingle hf arg
- | Cminor.Olonguofsingle => longuofsingle hf arg
- | Cminor.Osingleoflong => singleoflong hf arg
- | Cminor.Osingleoflongu => singleoflongu hf arg
+ | Cminor.Olongoffloat => longoffloat arg
+ | Cminor.Olonguoffloat => longuoffloat arg
+ | Cminor.Ofloatoflong => floatoflong arg
+ | Cminor.Ofloatoflongu => floatoflongu arg
+ | Cminor.Olongofsingle => longofsingle arg
+ | Cminor.Olonguofsingle => longuofsingle arg
+ | Cminor.Osingleoflong => singleoflong arg
+ | Cminor.Osingleoflongu => singleoflongu arg
end.
Definition sel_binop (op: Cminor.binary_operation) (arg1 arg2: expr) : expr :=
@@ -145,17 +137,17 @@ Definition sel_binop (op: Cminor.binary_operation) (arg1 arg2: expr) : expr :=
| Cminor.Odivfs => divfs arg1 arg2
| Cminor.Oaddl => addl arg1 arg2
| Cminor.Osubl => subl arg1 arg2
- | Cminor.Omull => mull hf arg1 arg2
- | Cminor.Odivl => divl hf arg1 arg2
- | Cminor.Odivlu => divlu hf arg1 arg2
- | Cminor.Omodl => modl hf arg1 arg2
- | Cminor.Omodlu => modlu hf arg1 arg2
+ | Cminor.Omull => mull arg1 arg2
+ | Cminor.Odivl => divl arg1 arg2
+ | Cminor.Odivlu => divlu arg1 arg2
+ | Cminor.Omodl => modl arg1 arg2
+ | Cminor.Omodlu => modlu arg1 arg2
| Cminor.Oandl => andl arg1 arg2
| Cminor.Oorl => orl arg1 arg2
| Cminor.Oxorl => xorl arg1 arg2
- | Cminor.Oshll => shll hf arg1 arg2
- | Cminor.Oshrl => shrl hf arg1 arg2
- | Cminor.Oshrlu => shrlu hf arg1 arg2
+ | Cminor.Oshll => shll arg1 arg2
+ | Cminor.Oshrl => shrl arg1 arg2
+ | Cminor.Oshrlu => shrlu arg1 arg2
| Cminor.Ocmp c => comp c arg1 arg2
| Cminor.Ocmpu c => compu c arg1 arg2
| Cminor.Ocmpf c => compf c arg1 arg2
@@ -192,7 +184,7 @@ Inductive call_kind : Type :=
Definition expr_is_addrof_ident (e: Cminor.expr) : option ident :=
match e with
| Cminor.Econst (Cminor.Oaddrsymbol id ofs) =>
- if Int.eq ofs Int.zero then Some id else None
+ if Ptrofs.eq ofs Ptrofs.zero then Some id else None
| _ => None
end.
@@ -326,10 +318,12 @@ Fixpoint sel_stmt (s: Cminor.stmt) : res stmt :=
| Cminor.Sgoto lbl => OK (Sgoto lbl)
end.
+End SELECTION.
+
(** Conversion of functions. *)
-Definition sel_function (f: Cminor.function) : res function :=
- do body' <- sel_stmt f.(Cminor.fn_body);
+Definition sel_function (dm: PTree.t globdef) (hf: helper_functions) (f: Cminor.function) : res function :=
+ do body' <- sel_stmt dm f.(Cminor.fn_body);
OK (mkfunction
f.(Cminor.fn_sig)
f.(Cminor.fn_params)
@@ -337,10 +331,8 @@ Definition sel_function (f: Cminor.function) : res function :=
f.(Cminor.fn_stackspace)
body').
-Definition sel_fundef (f: Cminor.fundef) : res fundef :=
- transf_partial_fundef sel_function f.
-
-End SELECTION.
+Definition sel_fundef (dm: PTree.t globdef) (hf: helper_functions) (f: Cminor.fundef) : res fundef :=
+ transf_partial_fundef (sel_function dm hf) f.
(** Setting up the helper functions. *)
diff --git a/backend/Selectionproof.v b/backend/Selectionproof.v
index a57e5ea6..34157553 100644
--- a/backend/Selectionproof.v
+++ b/backend/Selectionproof.v
@@ -12,28 +12,11 @@
(** Correctness of instruction selection *)
-Require Import Coqlib.
-Require Import Maps.
-Require Import AST.
-Require Import Linking.
-Require Import Errors.
-Require Import Integers.
-Require Import Values.
-Require Import Memory.
-Require Import Events.
-Require Import Globalenvs.
-Require Import Smallstep.
-Require Import Switch.
-Require Import Cminor.
-Require Import Op.
-Require Import CminorSel.
-Require Import SelectOp.
-Require Import SelectDiv.
-Require Import SelectLong.
-Require Import Selection.
-Require Import SelectOpproof.
-Require Import SelectDivproof.
-Require Import SelectLongproof.
+Require Import Coqlib Maps.
+Require Import AST Linking Errors Integers Values Memory Events Globalenvs Smallstep.
+Require Import Switch Cminor Op CminorSel.
+Require Import SelectOp SelectDiv SplitLong SelectLong Selection.
+Require Import SelectOpproof SelectDivproof SplitLongproof SelectLongproof.
Local Open Scope cminorsel_scope.
Local Open Scope error_monad_scope.
@@ -252,8 +235,7 @@ Lemma eval_sel_unop:
forall le op a1 v1 v,
eval_expr tge sp e m le a1 v1 ->
eval_unop op v1 = Some v ->
-
- exists v', eval_expr tge sp e m le (sel_unop hf op a1) v' /\ Val.lessdef v v'.
+ exists v', eval_expr tge sp e m le (sel_unop op a1) v' /\ Val.lessdef v v'.
Proof.
destruct op; simpl; intros; FuncInv; try subst v.
apply eval_cast8unsigned; auto.
@@ -296,7 +278,7 @@ Lemma eval_sel_binop:
eval_expr tge sp e m le a1 v1 ->
eval_expr tge sp e m le a2 v2 ->
eval_binop op v1 v2 m = Some v ->
- exists v', eval_expr tge sp e m le (sel_binop hf op a1 a2) v' /\ Val.lessdef v v'.
+ exists v', eval_expr tge sp e m le (sel_binop op a1 a2) v' /\ Val.lessdef v v'.
Proof.
destruct op; simpl; intros; FuncInv; try subst v.
apply eval_add; auto.
@@ -348,12 +330,12 @@ End CMCONSTR.
Lemma expr_is_addrof_ident_correct:
forall e id,
expr_is_addrof_ident e = Some id ->
- e = Cminor.Econst (Cminor.Oaddrsymbol id Int.zero).
+ e = Cminor.Econst (Cminor.Oaddrsymbol id Ptrofs.zero).
Proof.
intros e id. unfold expr_is_addrof_ident.
destruct e; try congruence.
destruct c; try congruence.
- predSpec Int.eq Int.eq_spec i0 Int.zero; congruence.
+ predSpec Ptrofs.eq Ptrofs.eq_spec i0 Ptrofs.zero; congruence.
Qed.
Lemma classify_call_correct:
@@ -363,17 +345,17 @@ Lemma classify_call_correct:
Genv.find_funct ge v = Some fd ->
match classify_call (prog_defmap unit) a with
| Call_default => True
- | Call_imm id => exists b, Genv.find_symbol ge id = Some b /\ v = Vptr b Int.zero
+ | Call_imm id => exists b, Genv.find_symbol ge id = Some b /\ v = Vptr b Ptrofs.zero
| Call_builtin ef => fd = External ef
end.
Proof.
unfold classify_call; intros.
destruct (expr_is_addrof_ident a) as [id|] eqn:EA; auto.
exploit expr_is_addrof_ident_correct; eauto. intros EQ; subst a.
- inv H0. inv H3.
+ inv H0. inv H3. unfold Genv.symbol_address in *.
destruct (Genv.find_symbol ge id) as [b|] eqn:FS; try discriminate.
rewrite Genv.find_funct_find_funct_ptr in H1.
- assert (DFL: exists b1, Genv.find_symbol ge id = Some b1 /\ Vptr b Int.zero = Vptr b1 Int.zero) by (exists b; auto).
+ assert (DFL: exists b1, Genv.find_symbol ge id = Some b1 /\ Vptr b Ptrofs.zero = Vptr b1 Ptrofs.zero) by (exists b; auto).
unfold globdef; destruct (prog_defmap unit)!id as [[[f|ef] |gv] |] eqn:G; auto.
destruct (ef_inline ef) eqn:INLINE; auto.
destruct (prog_defmap_linkorder _ _ _ _ H G) as (gd & P & Q).
@@ -530,12 +512,12 @@ Proof.
rewrite Int64.unsigned_repr. destruct (zeq (Int64.unsigned n0) n); auto.
unfold Int64.max_unsigned; omega.
- intros until n; intros EVAL R RANGE.
- eapply eval_cmplu. eexact EVAL. apply eval_longconst with (n := Int64.repr n).
+ eapply eval_cmplu; auto. eexact EVAL. apply eval_longconst with (n := Int64.repr n).
inv R. unfold Val.cmplu. simpl. f_equal; f_equal. unfold Int64.ltu.
rewrite Int64.unsigned_repr. destruct (zlt (Int64.unsigned n0) n); auto.
unfold Int64.max_unsigned; omega.
- intros until n; intros EVAL R RANGE.
- exploit eval_subl. eexact EVAL. apply eval_longconst with (n := Int64.repr n).
+ exploit eval_subl; auto. eexact EVAL. apply eval_longconst with (n := Int64.repr n).
intros (vb & A & B).
inv R. simpl in B. inv B. econstructor; split; eauto.
replace ((Int64.unsigned n0 - n) mod Int64.modulus)
@@ -579,13 +561,25 @@ Lemma eval_binop_lessdef:
Proof.
intros until m'; intros EV LD1 LD2 ME.
assert (exists v', eval_binop op v1' v2' m = Some v' /\ Val.lessdef v v').
- inv LD1. inv LD2. exists v; auto.
- destruct op; destruct v1'; simpl in *; inv EV; TrivialExists.
- destruct op; simpl in *; inv EV; TrivialExists.
- destruct op; try (exact H).
- simpl in *. TrivialExists. inv EV. apply Val.of_optbool_lessdef.
- intros. apply Val.cmpu_bool_lessdef with (Mem.valid_pointer m) v1 v2; auto.
- intros; eapply Mem.valid_pointer_extends; eauto.
+ { inv LD1. inv LD2. exists v; auto.
+ destruct op; destruct v1'; simpl in *; inv EV; TrivialExists.
+ destruct op; simpl in *; inv EV; TrivialExists. }
+ assert (CMPU: forall c,
+ eval_binop (Ocmpu c) v1 v2 m = Some v ->
+ exists v' : val, eval_binop (Ocmpu c) v1' v2' m' = Some v' /\ Val.lessdef v v').
+ { intros c A. simpl in *. inv A. econstructor; split. eauto.
+ apply Val.of_optbool_lessdef.
+ intros. apply Val.cmpu_bool_lessdef with (Mem.valid_pointer m) v1 v2; auto.
+ intros; eapply Mem.valid_pointer_extends; eauto. }
+ assert (CMPLU: forall c,
+ eval_binop (Ocmplu c) v1 v2 m = Some v ->
+ exists v' : val, eval_binop (Ocmplu c) v1' v2' m' = Some v' /\ Val.lessdef v v').
+ { intros c A. simpl in *. unfold Val.cmplu in *.
+ destruct (Val.cmplu_bool (Mem.valid_pointer m) c v1 v2) as [b|] eqn:C; simpl in A; inv A.
+ eapply Val.cmplu_bool_lessdef with (valid_ptr' := (Mem.valid_pointer m')) in C;
+ eauto using Mem.valid_pointer_extends.
+ rewrite C. exists (Val.of_bool b); auto. }
+ destruct op; auto.
Qed.
(** * Semantic preservation for instruction selection. *)
@@ -644,7 +638,7 @@ Lemma sel_expr_correct:
Cminor.eval_expr ge sp e m a v ->
forall e' le m',
env_lessdef e e' -> Mem.extends m m' ->
- exists v', eval_expr tge sp e' m' le (sel_expr hf a) v' /\ Val.lessdef v v'.
+ exists v', eval_expr tge sp e' m' le (sel_expr a) v' /\ Val.lessdef v v'.
Proof.
induction 1; intros; simpl.
(* Evar *)
@@ -654,10 +648,8 @@ Proof.
exists (Vint i); split; auto. econstructor. constructor. auto.
exists (Vfloat f); split; auto. econstructor. constructor. auto.
exists (Vsingle f); split; auto. econstructor. constructor. auto.
- exists (Val.longofwords (Vint (Int64.hiword i)) (Vint (Int64.loword i))); split.
- eapply eval_Eop. constructor. EvalOp. simpl; eauto. constructor. EvalOp. simpl; eauto. constructor. auto.
- simpl. rewrite Int64.ofwords_recompose. auto.
- rewrite <- symbols_preserved. fold (Genv.symbol_address tge i i0). apply eval_addrsymbol.
+ exists (Vlong i); split; auto. apply eval_longconst.
+ unfold Genv.symbol_address; rewrite <- symbols_preserved; fold (Genv.symbol_address tge i i0). apply eval_addrsymbol.
apply eval_addrstack.
(* Eunop *)
exploit IHeval_expr; eauto. intros [v1' [A B]].
@@ -668,7 +660,9 @@ Proof.
exploit IHeval_expr1; eauto. intros [v1' [A B]].
exploit IHeval_expr2; eauto. intros [v2' [C D]].
exploit eval_binop_lessdef; eauto. intros [v' [E F]].
- exploit eval_sel_binop. eexact LINK. eexact HF. eexact A. eexact C. eauto. intros [v'' [P Q]].
+ assert (G: exists v'', eval_expr tge sp e' m' le (sel_binop op (sel_expr a1) (sel_expr a2)) v'' /\ Val.lessdef v' v'')
+ by (eapply eval_sel_binop; eauto).
+ destruct G as [v'' [P Q]].
exists v''; split; eauto. eapply Val.lessdef_trans; eauto.
(* Eload *)
exploit IHeval_expr; eauto. intros [vaddr' [A B]].
@@ -681,7 +675,7 @@ Lemma sel_exprlist_correct:
Cminor.eval_exprlist ge sp e m a v ->
forall e' le m',
env_lessdef e e' -> Mem.extends m m' ->
- exists v', eval_exprlist tge sp e' m' le (sel_exprlist hf a) v' /\ Val.lessdef_list v v'.
+ exists v', eval_exprlist tge sp e' m' le (sel_exprlist a) v' /\ Val.lessdef_list v v'.
Proof.
induction 1; intros; simpl.
exists (@nil val); split; auto. constructor.
@@ -695,13 +689,13 @@ Lemma sel_builtin_arg_correct:
env_lessdef e e' -> Mem.extends m m' ->
Cminor.eval_expr ge sp e m a v ->
exists v',
- CminorSel.eval_builtin_arg tge sp e' m' (sel_builtin_arg hf a c) v'
+ CminorSel.eval_builtin_arg tge sp e' m' (sel_builtin_arg a c) v'
/\ Val.lessdef v v'.
Proof.
intros. unfold sel_builtin_arg.
exploit sel_expr_correct; eauto. intros (v1 & A & B).
exists v1; split; auto.
- destruct (builtin_arg_ok (builtin_arg (sel_expr hf a)) c).
+ destruct (builtin_arg_ok (builtin_arg (sel_expr a)) c).
apply eval_builtin_arg; eauto.
constructor; auto.
Qed.
@@ -714,7 +708,7 @@ Lemma sel_builtin_args_correct:
forall cl,
exists vl',
list_forall2 (CminorSel.eval_builtin_arg tge sp e' m')
- (sel_builtin_args hf al cl)
+ (sel_builtin_args al cl)
vl'
/\ Val.lessdef_list vl vl'.
Proof.
@@ -737,37 +731,11 @@ End EXPRESSIONS.
(** Semantic preservation for functions and statements. *)
-(*
-Inductive match_call_cont: Cminor.cont -> CminorSel.cont -> Prop :=
- | match_call_cont_stop:
- match_call_cont Cminor.Kstop Kstop
- | match_call_cont_call: forall cunit hf id f sp e k f' e' k',
- linkorder cunit prog ->
- helper_functions_declared cunit hf ->
- sel_function (prog_defmap cunit) hf f = OK f' ->
- match_cont cunit hf k k' -> env_lessdef e e' ->
- match_call_cont (Cminor.Kcall id f sp e k) (Kcall id f' sp e' k')
-
-with match_cont: Cminor.program -> helper_functions -> Cminor.cont -> CminorSel.cont -> Prop :=
- | match_cont_stop: forall cunit hf,
- match_cont cunit hf Cminor.Kstop Kstop
- | match_cont_seq: forall cunit hf s s' k k',
- sel_stmt (prog_defmap cunit) hf s = OK s' ->
- match_cont cunit hf k k' ->
- match_cont cunit hf (Cminor.Kseq s k) (Kseq s' k')
- | match_cont_block: forall cunit hf k k',
- match_cont cunit hf k k' ->
- match_cont cunit hf (Cminor.Kblock k) (Kblock k')
- | match_cont_call: forall cunit hf id f sp e k f' e' k',
- match_call_cont (Cminor.Kcall id f sp e k) (Kcall id f' sp e' k') ->
- match_cont cunit hf (Cminor.Kcall id f sp e k) (Kcall id f' sp e' k').
-*)
-
Inductive match_cont: Cminor.program -> helper_functions -> Cminor.cont -> CminorSel.cont -> Prop :=
| match_cont_stop: forall cunit hf,
match_cont cunit hf Cminor.Kstop Kstop
| match_cont_seq: forall cunit hf s s' k k',
- sel_stmt (prog_defmap cunit) hf s = OK s' ->
+ sel_stmt (prog_defmap cunit) s = OK s' ->
match_cont cunit hf k k' ->
match_cont cunit hf (Cminor.Kseq s k) (Kseq s' k')
| match_cont_block: forall cunit hf k k',
@@ -788,7 +756,7 @@ Inductive match_states: Cminor.state -> CminorSel.state -> Prop :=
(LINK: linkorder cunit prog)
(HF: helper_functions_declared cunit hf)
(TF: sel_function (prog_defmap cunit) hf f = OK f')
- (TS: sel_stmt (prog_defmap cunit) hf s = OK s')
+ (TS: sel_stmt (prog_defmap cunit) s = OK s')
(MC: match_cont cunit hf k k')
(LD: env_lessdef e e')
(ME: Mem.extends m m'),
@@ -835,31 +803,20 @@ Inductive match_states: Cminor.state -> CminorSel.state -> Prop :=
(Cminor.Returnstate v (Cminor.Kcall optid f sp e k) m)
(State f' Sskip k' sp (set_builtin_res (sel_builtin_res optid) v' e') m').
-(*
-Remark call_cont_commut_1:
- forall cunit hf k k', match_cont cunit hf k k' ->
- forall cunit' hf', match_cont cunit' hf' (Cminor.call_cont k) (call_cont k').
-Proof.
- induction 1; simpl; auto; intros; econstructor; eauto.
-Qed.
-
-Remark call_cont_commut_2:
- forall cunit hf k k', match_cont cunit hf k k' -> match_cont cunit hf (Cminor.call_cont k) (call_cont k').
-Proof.
- intros. eapply call_cont_commut_1; eauto.
-Qed.
-*)
-
Remark call_cont_commut:
forall cunit hf k k', match_cont cunit hf k k' -> match_call_cont (Cminor.call_cont k) (call_cont k').
Proof.
- induction 1; simpl; auto; red; intros; econstructor; eauto.
+ induction 1; simpl; auto; red; intros.
+- constructor.
+- eapply match_cont_call with (hf := hf); eauto.
Qed.
Remark match_is_call_cont:
forall cunit hf k k', match_cont cunit hf k k' -> Cminor.is_call_cont k -> match_call_cont k k'.
Proof.
- destruct 1; intros; try contradiction; red; intros; econstructor; eauto.
+ destruct 1; intros; try contradiction; red; intros.
+- constructor.
+- eapply match_cont_call with (hf := hf); eauto.
Qed.
Remark match_call_cont_cont:
@@ -875,16 +832,16 @@ Qed.
Remark find_label_commut:
forall cunit hf lbl s k s' k',
match_cont cunit hf k k' ->
- sel_stmt (prog_defmap cunit) hf s = OK s' ->
+ sel_stmt (prog_defmap cunit) s = OK s' ->
match Cminor.find_label lbl s k, find_label lbl s' k' with
| None, None => True
- | Some(s1, k1), Some(s1', k1') => sel_stmt (prog_defmap cunit) hf s1 = OK s1' /\ match_cont cunit hf k1 k1'
+ | Some(s1, k1), Some(s1', k1') => sel_stmt (prog_defmap cunit) s1 = OK s1' /\ match_cont cunit hf k1 k1'
| _, _ => False
end.
Proof.
induction s; intros until k'; simpl; intros MC SE; try (monadInv SE); simpl; auto.
(* store *)
- unfold store. destruct (addressing m (sel_expr hf e)); simpl; auto.
+ unfold store. destruct (addressing m (sel_expr e)); simpl; auto.
(* call *)
destruct (classify_call (prog_defmap cunit) e); simpl; auto.
(* tailcall *)
@@ -963,7 +920,7 @@ Proof.
econstructor; eauto. econstructor; eauto.
eapply sig_function_translated; eauto.
eapply match_callstate with (cunit := cunit'); eauto.
- red; intros. eapply match_cont_call with (cunit := cunit); eauto.
+ red; intros. eapply match_cont_call with (cunit := cunit) (hf := hf); eauto.
+ (* direct *)
intros [b [U V]].
exploit sel_exprlist_correct; eauto. intros [vargs' [C D]].
@@ -973,7 +930,7 @@ Proof.
subst vf. econstructor; eauto. rewrite symbols_preserved; eauto.
eapply sig_function_translated; eauto.
eapply match_callstate with (cunit := cunit'); eauto.
- red; intros; eapply match_cont_call with (cunit := cunit); eauto.
+ red; intros; eapply match_cont_call with (cunit := cunit) (hf := hf); eauto.
+ (* turned into Sbuiltin *)
intros EQ. subst fd.
exploit sel_builtin_args_correct; eauto. intros [vargs' [C D]].
@@ -1052,7 +1009,7 @@ Proof.
- (* Slabel *)
left; econstructor; split. constructor. econstructor; eauto.
- (* Sgoto *)
- assert (sel_stmt (prog_defmap cunit) hf (Cminor.fn_body f) = OK (fn_body f')).
+ assert (sel_stmt (prog_defmap cunit) (Cminor.fn_body f) = OK (fn_body f')).
{ monadInv TF; simpl; auto. }
exploit (find_label_commut cunit hf lbl (Cminor.fn_body f) (Cminor.call_cont k)).
eapply call_cont_commut; eauto. eauto.
diff --git a/backend/SelectLong.vp b/backend/SplitLong.vp
index 105b284c..305e20f3 100644
--- a/backend/SelectLong.vp
+++ b/backend/SplitLong.vp
@@ -14,21 +14,18 @@
Require String.
Require Import Coqlib.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Op.
-Require Import CminorSel.
+Require Import AST Integers Floats.
+Require Import Op CminorSel.
Require Import SelectOp.
Local Open Scope cminorsel_scope.
Local Open Scope string_scope.
(** Some operations on 64-bit integers are transformed into calls to
- runtime library functions. The following record type collects
+ runtime library functions. The following type class collects
the names of these functions. *)
-Record helper_functions : Type := mk_helper_functions {
+Class helper_functions := mk_helper_functions {
i64_dtos: ident; (**r float64 -> signed long *)
i64_dtou: ident; (**r float64 -> unsigned long *)
i64_stod: ident; (**r signed long -> float64 *)
@@ -54,7 +51,7 @@ Definition sig_ii_l := mksignature (Tint :: Tint :: nil) (Some Tlong) cc_default
Section SELECT.
-Variable hf: helper_functions.
+Context {hf: helper_functions}.
Definition makelong (h l: expr): expr :=
Eop Omakelong (h ::: l ::: Enil).
@@ -129,21 +126,21 @@ Definition notl (e: expr) :=
splitlong e (fun h l => makelong (notint h) (notint l)).
Definition longoffloat (arg: expr) :=
- Eexternal hf.(i64_dtos) sig_f_l (arg ::: Enil).
+ Eexternal i64_dtos sig_f_l (arg ::: Enil).
Definition longuoffloat (arg: expr) :=
- Eexternal hf.(i64_dtou) sig_f_l (arg ::: Enil).
+ Eexternal i64_dtou sig_f_l (arg ::: Enil).
Definition floatoflong (arg: expr) :=
- Eexternal hf.(i64_stod) sig_l_f (arg ::: Enil).
+ Eexternal i64_stod sig_l_f (arg ::: Enil).
Definition floatoflongu (arg: expr) :=
- Eexternal hf.(i64_utod) sig_l_f (arg ::: Enil).
+ Eexternal i64_utod sig_l_f (arg ::: Enil).
Definition longofsingle (arg: expr) :=
longoffloat (floatofsingle arg).
Definition longuofsingle (arg: expr) :=
longuoffloat (floatofsingle arg).
Definition singleoflong (arg: expr) :=
- Eexternal hf.(i64_stof) sig_l_s (arg ::: Enil).
+ Eexternal i64_stof sig_l_s (arg ::: Enil).
Definition singleoflongu (arg: expr) :=
- Eexternal hf.(i64_utof) sig_l_s (arg ::: Enil).
+ Eexternal i64_utof sig_l_s (arg ::: Enil).
Definition andl (e1 e2: expr) :=
splitlong2 e1 e2 (fun h1 l1 h2 l2 => makelong (and h1 h2) (and l1 l2)).
@@ -164,7 +161,7 @@ Definition shllimm (e1: expr) (n: int) :=
makelong (shlimm (lowlong e1) (Int.sub n Int.iwordsize))
(Eop (Ointconst Int.zero) Enil)
else
- Eexternal hf.(i64_shl) sig_li_l (e1 ::: Eop (Ointconst n) Enil ::: Enil).
+ Eexternal i64_shl sig_li_l (e1 ::: Eop (Ointconst n) Enil ::: Enil).
Definition shrluimm (e1: expr) (n: int) :=
if Int.eq n Int.zero then e1 else
@@ -176,7 +173,7 @@ Definition shrluimm (e1: expr) (n: int) :=
makelong (Eop (Ointconst Int.zero) Enil)
(shruimm (highlong e1) (Int.sub n Int.iwordsize))
else
- Eexternal hf.(i64_shr) sig_li_l (e1 ::: Eop (Ointconst n) Enil ::: Enil).
+ Eexternal i64_shr sig_li_l (e1 ::: Eop (Ointconst n) Enil ::: Enil).
Definition shrlimm (e1: expr) (n: int) :=
if Int.eq n Int.zero then e1 else
@@ -189,7 +186,7 @@ Definition shrlimm (e1: expr) (n: int) :=
(makelong (shrimm (Eletvar 0) (Int.repr 31))
(shrimm (Eletvar 0) (Int.sub n Int.iwordsize)))
else
- Eexternal hf.(i64_sar) sig_li_l (e1 ::: Eop (Ointconst n) Enil ::: Enil).
+ Eexternal i64_sar sig_li_l (e1 ::: Eop (Ointconst n) Enil ::: Enil).
Definition is_intconst (e: expr) :=
match e with
@@ -200,19 +197,19 @@ Definition is_intconst (e: expr) :=
Definition shll (e1 e2: expr) :=
match is_intconst e2 with
| Some n => shllimm e1 n
- | None => Eexternal hf.(i64_shl) sig_li_l (e1 ::: e2 ::: Enil)
+ | None => Eexternal i64_shl sig_li_l (e1 ::: e2 ::: Enil)
end.
Definition shrlu (e1 e2: expr) :=
match is_intconst e2 with
| Some n => shrluimm e1 n
- | None => Eexternal hf.(i64_shr) sig_li_l (e1 ::: e2 ::: Enil)
+ | None => Eexternal i64_shr sig_li_l (e1 ::: e2 ::: Enil)
end.
Definition shrl (e1 e2: expr) :=
match is_intconst e2 with
| Some n => shrlimm e1 n
- | None => Eexternal hf.(i64_sar) sig_li_l (e1 ::: e2 ::: Enil)
+ | None => Eexternal i64_sar sig_li_l (e1 ::: e2 ::: Enil)
end.
Definition addl (e1 e2: expr) :=
@@ -245,8 +242,8 @@ Definition mull_base (e1 e2: expr) :=
Definition mullimm (e: expr) (n: int64) :=
if Int64.eq n Int64.zero then longconst Int64.zero else
if Int64.eq n Int64.one then e else
- match Int64.is_power2 n with
- | Some l => shllimm e (Int.repr (Int64.unsigned l))
+ match Int64.is_power2' n with
+ | Some l => shllimm e l
| None => mull_base e (longconst n)
end.
@@ -264,23 +261,23 @@ Definition binop_long (id: ident) (sem: int64 -> int64 -> int64) (e1 e2: expr) :
| _, _ => Eexternal id sig_ll_l (e1 ::: e2 ::: Enil)
end.
-Definition divl e1 e2 := binop_long hf.(i64_sdiv) Int64.divs e1 e2.
-Definition modl e1 e2 := binop_long hf.(i64_smod) Int64.mods e1 e2.
+Definition divl e1 e2 := binop_long i64_sdiv Int64.divs e1 e2.
+Definition modl e1 e2 := binop_long i64_smod Int64.mods e1 e2.
Definition divlu (e1 e2: expr) :=
- let default := Eexternal hf.(i64_udiv) sig_ll_l (e1 ::: e2 ::: Enil) in
+ let default := Eexternal i64_udiv sig_ll_l (e1 ::: e2 ::: Enil) in
match is_longconst e1, is_longconst e2 with
| Some n1, Some n2 => longconst (Int64.divu n1 n2)
| _, Some n2 =>
- match Int64.is_power2 n2 with
- | Some l => shrluimm e1 (Int.repr (Int64.unsigned l))
+ match Int64.is_power2' n2 with
+ | Some l => shrluimm e1 l
| None => default
end
| _, _ => default
end.
Definition modlu (e1 e2: expr) :=
- let default := Eexternal hf.(i64_umod) sig_ll_l (e1 ::: e2 ::: Enil) in
+ let default := Eexternal i64_umod sig_ll_l (e1 ::: e2 ::: Enil) in
match is_longconst e1, is_longconst e2 with
| Some n1, Some n2 => longconst (Int64.modu n1 n2)
| _, Some n2 =>
@@ -307,15 +304,8 @@ Definition cmplu (c: comparison) (e1 e2: expr) :=
match c with
| Ceq =>
cmpl_eq_zero (xorl e1 e2)
-(*
- (if is_longconst_zero e2 then e1
- else if is_longconst_zero e1 then e2
- else xorl e1 e2) *)
| Cne =>
cmpl_ne_zero (xorl e1 e2)
-(* (if is_longconst_zero e2 then e1
- else if is_longconst_zero e1 then e2
- else xorl e1 e2) *)
| Clt =>
cmplu_gen Clt Clt e1 e2
| Cle =>
diff --git a/backend/SelectLongproof.v b/backend/SplitLongproof.v
index f15015e8..1dbe25bd 100644
--- a/backend/SelectLongproof.v
+++ b/backend/SplitLongproof.v
@@ -13,22 +13,10 @@
(** Correctness of instruction selection for integer division *)
Require Import String.
-Require Import Coqlib.
-Require Import Maps.
-Require Import AST.
-Require Import Errors.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Events.
-Require Import Cminor.
-Require Import Op.
-Require Import CminorSel.
-Require Import SelectOp.
-Require Import SelectOpproof.
-Require Import SelectLong.
+Require Import Coqlib Maps.
+Require Import AST Errors Integers Floats.
+Require Import Values Memory Globalenvs Events Cminor Op CminorSel.
+Require Import SelectOp SelectOpproof SplitLong.
Open Local Scope cminorsel_scope.
Open Local Scope string_scope.
@@ -66,19 +54,19 @@ Definition helper_declared {F V: Type} (p: AST.program (AST.fundef F) V) (id: id
(prog_defmap p)!id = Some (Gfun (External (EF_runtime name sg))).
Definition helper_functions_declared {F V: Type} (p: AST.program (AST.fundef F) V) (hf: helper_functions) : Prop :=
- helper_declared p hf.(i64_dtos) "__i64_dtos" sig_f_l
- /\ helper_declared p hf.(i64_dtou) "__i64_dtou" sig_f_l
- /\ helper_declared p hf.(i64_stod) "__i64_stod" sig_l_f
- /\ helper_declared p hf.(i64_utod) "__i64_utod" sig_l_f
- /\ helper_declared p hf.(i64_stof) "__i64_stof" sig_l_s
- /\ helper_declared p hf.(i64_utof) "__i64_utof" sig_l_s
- /\ helper_declared p hf.(i64_sdiv) "__i64_sdiv" sig_ll_l
- /\ helper_declared p hf.(i64_udiv) "__i64_udiv" sig_ll_l
- /\ helper_declared p hf.(i64_smod) "__i64_smod" sig_ll_l
- /\ helper_declared p hf.(i64_umod) "__i64_umod" sig_ll_l
- /\ helper_declared p hf.(i64_shl) "__i64_shl" sig_li_l
- /\ helper_declared p hf.(i64_shr) "__i64_shr" sig_li_l
- /\ helper_declared p hf.(i64_sar) "__i64_sar" sig_li_l.
+ helper_declared p i64_dtos "__i64_dtos" sig_f_l
+ /\ helper_declared p i64_dtou "__i64_dtou" sig_f_l
+ /\ helper_declared p i64_stod "__i64_stod" sig_l_f
+ /\ helper_declared p i64_utod "__i64_utod" sig_l_f
+ /\ helper_declared p i64_stof "__i64_stof" sig_l_s
+ /\ helper_declared p i64_utof "__i64_utof" sig_l_s
+ /\ helper_declared p i64_sdiv "__i64_sdiv" sig_ll_l
+ /\ helper_declared p i64_udiv "__i64_udiv" sig_ll_l
+ /\ helper_declared p i64_smod "__i64_smod" sig_ll_l
+ /\ helper_declared p i64_umod "__i64_umod" sig_ll_l
+ /\ helper_declared p i64_shl "__i64_shl" sig_li_l
+ /\ helper_declared p i64_shr "__i64_shr" sig_li_l
+ /\ helper_declared p i64_sar "__i64_sar" sig_li_l.
(** * Correctness of the instruction selection functions for 64-bit operators *)
@@ -414,7 +402,7 @@ Theorem eval_longoffloat:
forall le a x y,
eval_expr ge sp e m le a x ->
Val.longoffloat x = Some y ->
- exists v, eval_expr ge sp e m le (longoffloat hf a) v /\ Val.lessdef y v.
+ exists v, eval_expr ge sp e m le (longoffloat a) v /\ Val.lessdef y v.
Proof.
intros; unfold longoffloat. econstructor; split.
eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
@@ -424,7 +412,7 @@ Theorem eval_longuoffloat:
forall le a x y,
eval_expr ge sp e m le a x ->
Val.longuoffloat x = Some y ->
- exists v, eval_expr ge sp e m le (longuoffloat hf a) v /\ Val.lessdef y v.
+ exists v, eval_expr ge sp e m le (longuoffloat a) v /\ Val.lessdef y v.
Proof.
intros; unfold longuoffloat. econstructor; split.
eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
@@ -434,7 +422,7 @@ Theorem eval_floatoflong:
forall le a x y,
eval_expr ge sp e m le a x ->
Val.floatoflong x = Some y ->
- exists v, eval_expr ge sp e m le (floatoflong hf a) v /\ Val.lessdef y v.
+ exists v, eval_expr ge sp e m le (floatoflong a) v /\ Val.lessdef y v.
Proof.
intros; unfold floatoflong. econstructor; split.
eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
@@ -444,7 +432,7 @@ Theorem eval_floatoflongu:
forall le a x y,
eval_expr ge sp e m le a x ->
Val.floatoflongu x = Some y ->
- exists v, eval_expr ge sp e m le (floatoflongu hf a) v /\ Val.lessdef y v.
+ exists v, eval_expr ge sp e m le (floatoflongu a) v /\ Val.lessdef y v.
Proof.
intros; unfold floatoflongu. econstructor; split.
eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
@@ -454,7 +442,7 @@ Theorem eval_longofsingle:
forall le a x y,
eval_expr ge sp e m le a x ->
Val.longofsingle x = Some y ->
- exists v, eval_expr ge sp e m le (longofsingle hf a) v /\ Val.lessdef y v.
+ exists v, eval_expr ge sp e m le (longofsingle a) v /\ Val.lessdef y v.
Proof.
intros; unfold longofsingle.
destruct x; simpl in H0; inv H0. destruct (Float32.to_long f) as [n|] eqn:EQ; simpl in H2; inv H2.
@@ -468,7 +456,7 @@ Theorem eval_longuofsingle:
forall le a x y,
eval_expr ge sp e m le a x ->
Val.longuofsingle x = Some y ->
- exists v, eval_expr ge sp e m le (longuofsingle hf a) v /\ Val.lessdef y v.
+ exists v, eval_expr ge sp e m le (longuofsingle a) v /\ Val.lessdef y v.
Proof.
intros; unfold longuofsingle.
destruct x; simpl in H0; inv H0. destruct (Float32.to_longu f) as [n|] eqn:EQ; simpl in H2; inv H2.
@@ -482,7 +470,7 @@ Theorem eval_singleoflong:
forall le a x y,
eval_expr ge sp e m le a x ->
Val.singleoflong x = Some y ->
- exists v, eval_expr ge sp e m le (singleoflong hf a) v /\ Val.lessdef y v.
+ exists v, eval_expr ge sp e m le (singleoflong a) v /\ Val.lessdef y v.
Proof.
intros; unfold singleoflong. econstructor; split.
eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
@@ -492,7 +480,7 @@ Theorem eval_singleoflongu:
forall le a x y,
eval_expr ge sp e m le a x ->
Val.singleoflongu x = Some y ->
- exists v, eval_expr ge sp e m le (singleoflongu hf a) v /\ Val.lessdef y v.
+ exists v, eval_expr ge sp e m le (singleoflongu a) v /\ Val.lessdef y v.
Proof.
intros; unfold singleoflongu. econstructor; split.
eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
@@ -592,7 +580,7 @@ Qed.
Lemma eval_shllimm:
forall n,
- unary_constructor_sound (fun e => shllimm hf e n) (fun v => Val.shll v (Vint n)).
+ unary_constructor_sound (fun e => shllimm e n) (fun v => Val.shll v (Vint n)).
Proof.
unfold shllimm; red; intros.
apply eval_shift_imm; intros.
@@ -625,7 +613,7 @@ Proof.
econstructor; split. eapply eval_helper_2; eauto. EvalOp. DeclHelper. UseHelper. auto.
Qed.
-Theorem eval_shll: binary_constructor_sound (shll hf) Val.shll.
+Theorem eval_shll: binary_constructor_sound shll Val.shll.
Proof.
unfold shll; red; intros.
destruct (is_intconst b) as [n|] eqn:IC.
@@ -638,7 +626,7 @@ Qed.
Lemma eval_shrluimm:
forall n,
- unary_constructor_sound (fun e => shrluimm hf e n) (fun v => Val.shrlu v (Vint n)).
+ unary_constructor_sound (fun e => shrluimm e n) (fun v => Val.shrlu v (Vint n)).
Proof.
unfold shrluimm; red; intros. apply eval_shift_imm; intros.
+ (* n = 0 *)
@@ -670,7 +658,7 @@ Proof.
econstructor; split. eapply eval_helper_2; eauto. EvalOp. DeclHelper. UseHelper. auto.
Qed.
-Theorem eval_shrlu: binary_constructor_sound (shrlu hf) Val.shrlu.
+Theorem eval_shrlu: binary_constructor_sound shrlu Val.shrlu.
Proof.
unfold shrlu; red; intros.
destruct (is_intconst b) as [n|] eqn:IC.
@@ -683,7 +671,7 @@ Qed.
Lemma eval_shrlimm:
forall n,
- unary_constructor_sound (fun e => shrlimm hf e n) (fun v => Val.shrl v (Vint n)).
+ unary_constructor_sound (fun e => shrlimm e n) (fun v => Val.shrl v (Vint n)).
Proof.
unfold shrlimm; red; intros. apply eval_shift_imm; intros.
+ (* n = 0 *)
@@ -719,7 +707,7 @@ Proof.
econstructor; split. eapply eval_helper_2; eauto. EvalOp. DeclHelper. UseHelper. auto.
Qed.
-Theorem eval_shrl: binary_constructor_sound (shrl hf) Val.shrl.
+Theorem eval_shrl: binary_constructor_sound shrl Val.shrl.
Proof.
unfold shrl; red; intros.
destruct (is_intconst b) as [n|] eqn:IC.
@@ -730,9 +718,9 @@ Proof.
econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto.
Qed.
-Theorem eval_addl: binary_constructor_sound addl Val.addl.
+Theorem eval_addl: Archi.ptr64 = false -> binary_constructor_sound addl Val.addl.
Proof.
- unfold addl; red; intros.
+ unfold addl; red; intros.
set (default := Ebuiltin (EF_builtin "__builtin_addl" sig_ll_l) (a ::: b ::: Enil)).
assert (DEFAULT:
exists v, eval_expr ge sp e m le default v /\ Val.lessdef (Val.addl x y) v).
@@ -746,14 +734,14 @@ Proof.
econstructor; split. apply eval_longconst. simpl; auto.
- predSpec Int64.eq Int64.eq_spec p Int64.zero; auto.
subst p. exploit (is_longconst_sound le a); eauto. intros EQ; subst x.
- exists y; split; auto. simpl. destruct y; auto. rewrite Int64.add_zero_l; auto.
+ exists y; split; auto. unfold Val.addl; rewrite H; destruct y; auto. rewrite Int64.add_zero_l; auto.
- predSpec Int64.eq Int64.eq_spec q Int64.zero; auto.
subst q. exploit (is_longconst_sound le b); eauto. intros EQ; subst y.
- exists x; split; auto. destruct x; simpl; auto. rewrite Int64.add_zero; auto.
+ exists x; split; auto. unfold Val.addl; rewrite H; destruct x; simpl; auto. rewrite Int64.add_zero; auto.
- auto.
Qed.
-Theorem eval_subl: binary_constructor_sound subl Val.subl.
+Theorem eval_subl: Archi.ptr64 = false -> binary_constructor_sound subl Val.subl.
Proof.
unfold subl; red; intros.
set (default := Ebuiltin (EF_builtin "__builtin_subl" sig_ll_l) (a ::: b ::: Enil)).
@@ -773,7 +761,7 @@ Proof.
destruct y; simpl; auto.
- predSpec Int64.eq Int64.eq_spec q Int64.zero; auto.
subst q. exploit (is_longconst_sound le b); eauto. intros EQ; subst y.
- exists x; split; auto. destruct x; simpl; auto. rewrite Int64.sub_zero_l; auto.
+ exists x; split; auto. unfold Val.subl; rewrite H; destruct x; simpl; auto. rewrite Int64.sub_zero_l; auto.
- auto.
Qed.
@@ -799,7 +787,7 @@ Proof.
Qed.
Lemma eval_mullimm:
- forall n, unary_constructor_sound (fun a => mullimm hf a n) (fun v => Val.mull v (Vlong n)).
+ forall n, unary_constructor_sound (fun a => mullimm a n) (fun v => Val.mull v (Vlong n)).
Proof.
unfold mullimm; red; intros.
predSpec Int64.eq Int64.eq_spec n Int64.zero.
@@ -808,28 +796,17 @@ Proof.
predSpec Int64.eq Int64.eq_spec n Int64.one.
subst n. exists x; split; auto.
destruct x; simpl; auto. rewrite Int64.mul_one. auto.
- destruct (Int64.is_power2 n) as [l|] eqn:P2.
- exploit eval_shllimm. eauto. instantiate (1 := Int.repr (Int64.unsigned l)).
- intros [v [A B]].
+ destruct (Int64.is_power2' n) as [l|] eqn:P2.
+ exploit eval_shllimm. eauto. instantiate (1 := l). intros [v [A B]].
exists v; split; auto.
destruct x; simpl; auto.
- erewrite Int64.mul_pow2 by eauto.
- assert (EQ: Int.unsigned (Int.repr (Int64.unsigned l)) = Int64.unsigned l).
- { apply Int.unsigned_repr.
- exploit Int64.is_power2_rng; eauto.
- assert (Int64.zwordsize < Int.max_unsigned) by (compute; auto).
- omega.
- }
- simpl in B.
- replace (Int.ltu (Int.repr (Int64.unsigned l)) Int64.iwordsize')
- with (Int64.ltu l Int64.iwordsize) in B.
- erewrite Int64.is_power2_range in B by eauto.
- unfold Int64.shl' in B. rewrite EQ in B. auto.
- unfold Int64.ltu, Int.ltu. rewrite EQ. auto.
+ erewrite Int64.mul_pow2' by eauto.
+ simpl in B. erewrite Int64.is_power2'_range in B by eauto.
+ exact B.
apply eval_mull_base; auto. apply eval_longconst.
Qed.
-Theorem eval_mull: binary_constructor_sound (mull hf) Val.mull.
+Theorem eval_mull: binary_constructor_sound mull Val.mull.
Proof.
unfold mull; red; intros.
destruct (is_longconst a) as [p|] eqn:LC1;
@@ -870,7 +847,7 @@ Theorem eval_divl:
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.divls x y = Some z ->
- exists v, eval_expr ge sp e m le (divl hf a b) v /\ Val.lessdef z v.
+ exists v, eval_expr ge sp e m le (divl a b) v /\ Val.lessdef z v.
Proof.
intros. eapply eval_binop_long; eauto.
intros; subst; simpl in H1.
@@ -885,7 +862,7 @@ Theorem eval_modl:
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.modls x y = Some z ->
- exists v, eval_expr ge sp e m le (modl hf a b) v /\ Val.lessdef z v.
+ exists v, eval_expr ge sp e m le (modl a b) v /\ Val.lessdef z v.
Proof.
intros. eapply eval_binop_long; eauto.
intros; subst; simpl in H1.
@@ -900,10 +877,10 @@ Theorem eval_divlu:
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.divlu x y = Some z ->
- exists v, eval_expr ge sp e m le (divlu hf a b) v /\ Val.lessdef z v.
+ exists v, eval_expr ge sp e m le (divlu a b) v /\ Val.lessdef z v.
Proof.
intros. unfold divlu.
- set (default := Eexternal hf.(i64_udiv) sig_ll_l (a ::: b ::: Enil)).
+ set (default := Eexternal i64_udiv sig_ll_l (a ::: b ::: Enil)).
assert (DEFAULT:
exists v, eval_expr ge sp e m le default v /\ Val.lessdef z v).
{
@@ -916,25 +893,15 @@ Proof.
econstructor; split. apply eval_longconst.
simpl in H1. destruct (Int64.eq q Int64.zero); inv H1. auto.
- auto.
-- destruct (Int64.is_power2 q) as [l|] eqn:P2; auto.
+- destruct (Int64.is_power2' q) as [l|] eqn:P2; auto.
exploit (is_longconst_sound le b); eauto. intros EQ; subst y.
- replace z with (Val.shrlu x (Vint (Int.repr (Int64.unsigned l)))).
+ replace z with (Val.shrlu x (Vint l)).
apply eval_shrluimm. auto.
destruct x; simpl in H1; try discriminate.
destruct (Int64.eq q Int64.zero); inv H1.
- simpl.
- assert (EQ: Int.unsigned (Int.repr (Int64.unsigned l)) = Int64.unsigned l).
- { apply Int.unsigned_repr.
- exploit Int64.is_power2_rng; eauto.
- assert (Int64.zwordsize < Int.max_unsigned) by (compute; auto).
- omega.
- }
- replace (Int.ltu (Int.repr (Int64.unsigned l)) Int64.iwordsize')
- with (Int64.ltu l Int64.iwordsize).
- erewrite Int64.is_power2_range by eauto.
- erewrite Int64.divu_pow2 by eauto.
- unfold Int64.shru', Int64.shru. rewrite EQ. auto.
- unfold Int64.ltu, Int.ltu. rewrite EQ. auto.
+ simpl. erewrite Int64.is_power2'_range by eauto.
+ erewrite Int64.divu_pow2' by eauto.
+ auto.
- auto.
Qed.
@@ -943,10 +910,10 @@ Theorem eval_modlu:
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.modlu x y = Some z ->
- exists v, eval_expr ge sp e m le (modlu hf a b) v /\ Val.lessdef z v.
+ exists v, eval_expr ge sp e m le (modlu a b) v /\ Val.lessdef z v.
Proof.
intros. unfold modlu.
- set (default := Eexternal hf.(i64_umod) sig_ll_l (a ::: b ::: Enil)).
+ set (default := Eexternal i64_umod sig_ll_l (a ::: b ::: Enil)).
assert (DEFAULT:
exists v, eval_expr ge sp e m le default v /\ Val.lessdef z v).
{
@@ -1058,11 +1025,12 @@ Theorem eval_cmplu:
forall c le a x b y v,
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
- Val.cmplu c x y = Some v ->
+ Val.cmplu (Mem.valid_pointer m) c x y = Some v ->
+ Archi.ptr64 = false ->
eval_expr ge sp e m le (cmplu c a b) v.
Proof.
- intros. unfold Val.cmplu in H1.
- destruct x; simpl in H1; try discriminate. destruct y; inv H1.
+ intros. unfold Val.cmplu, Val.cmplu_bool in H1. rewrite H2 in H1. simpl in H1.
+ destruct x; simpl in H1; try discriminate H1; destruct y; inv H1.
rename i into x. rename i0 into y.
destruct c; simpl.
- (* Ceq *)
diff --git a/backend/Stacking.v b/backend/Stacking.v
index d1c17029..700025c2 100644
--- a/backend/Stacking.v
+++ b/backend/Stacking.v
@@ -39,7 +39,7 @@ Fixpoint save_callee_save_rec (rl: list mreg) (ofs: Z) (k: Mach.code) :=
let ty := mreg_type r in
let sz := AST.typesize ty in
let ofs1 := align ofs sz in
- Msetstack r (Int.repr ofs1) ty :: save_callee_save_rec rl (ofs1 + sz) k
+ Msetstack r (Ptrofs.repr ofs1) ty :: save_callee_save_rec rl (ofs1 + sz) k
end.
Definition save_callee_save (fe: frame_env) (k: Mach.code) :=
@@ -56,7 +56,7 @@ Fixpoint restore_callee_save_rec (rl: list mreg) (ofs: Z) (k: Mach.code) :=
let ty := mreg_type r in
let sz := AST.typesize ty in
let ofs1 := align ofs sz in
- Mgetstack (Int.repr ofs1) ty r :: restore_callee_save_rec rl (ofs1 + sz) k
+ Mgetstack (Ptrofs.repr ofs1) ty r :: restore_callee_save_rec rl (ofs1 + sz) k
end.
Definition restore_callee_save (fe: frame_env) (k: Mach.code) :=
@@ -72,10 +72,10 @@ Definition restore_callee_save (fe: frame_env) (k: Mach.code) :=
behaviour. *)
Definition transl_op (fe: frame_env) (op: operation) :=
- shift_stack_operation (Int.repr fe.(fe_stack_data)) op.
+ shift_stack_operation fe.(fe_stack_data) op.
Definition transl_addr (fe: frame_env) (addr: addressing) :=
- shift_stack_addressing (Int.repr fe.(fe_stack_data)) addr.
+ shift_stack_addressing fe.(fe_stack_data) addr.
(** Translation of a builtin argument. *)
@@ -83,16 +83,16 @@ Fixpoint transl_builtin_arg (fe: frame_env) (a: builtin_arg loc) : builtin_arg m
match a with
| BA (R r) => BA r
| BA (S Local ofs ty) =>
- BA_loadstack (chunk_of_type ty) (Int.repr (offset_local fe ofs))
+ BA_loadstack (chunk_of_type ty) (Ptrofs.repr (offset_local fe ofs))
| BA (S _ _ _) => BA_int Int.zero (**r never happens *)
| BA_int n => BA_int n
| BA_long n => BA_long n
| BA_float n => BA_float n
| BA_single n => BA_single n
| BA_loadstack chunk ofs =>
- BA_loadstack chunk (Int.add ofs (Int.repr fe.(fe_stack_data)))
+ BA_loadstack chunk (Ptrofs.add ofs (Ptrofs.repr fe.(fe_stack_data)))
| BA_addrstack ofs =>
- BA_addrstack (Int.add ofs (Int.repr fe.(fe_stack_data)))
+ BA_addrstack (Ptrofs.add ofs (Ptrofs.repr fe.(fe_stack_data)))
| BA_loadglobal chunk id ofs => BA_loadglobal chunk id ofs
| BA_addrglobal id ofs => BA_addrglobal id ofs
| BA_splitlong hi lo =>
@@ -114,20 +114,20 @@ Definition transl_instr
| Lgetstack sl ofs ty r =>
match sl with
| Local =>
- Mgetstack (Int.repr (offset_local fe ofs)) ty r :: k
+ Mgetstack (Ptrofs.repr (offset_local fe ofs)) ty r :: k
| Incoming =>
- Mgetparam (Int.repr (offset_arg ofs)) ty r :: k
+ Mgetparam (Ptrofs.repr (offset_arg ofs)) ty r :: k
| Outgoing =>
- Mgetstack (Int.repr (offset_arg ofs)) ty r :: k
+ Mgetstack (Ptrofs.repr (offset_arg ofs)) ty r :: k
end
| Lsetstack r sl ofs ty =>
match sl with
| Local =>
- Msetstack r (Int.repr (offset_local fe ofs)) ty :: k
+ Msetstack r (Ptrofs.repr (offset_local fe ofs)) ty :: k
| Incoming =>
k (* should not happen *)
| Outgoing =>
- Msetstack r (Int.repr (offset_arg ofs)) ty :: k
+ Msetstack r (Ptrofs.repr (offset_arg ofs)) ty :: k
end
| Lop op args res =>
Mop (transl_op fe op) args res :: k
@@ -175,15 +175,15 @@ Definition transf_function (f: Linear.function) : res Mach.function :=
let fe := make_env (function_bounds f) in
if negb (wt_function f) then
Error (msg "Ill-formed Linear code")
- else if zlt Int.max_unsigned fe.(fe_size) then
+ else if zlt Ptrofs.max_unsigned fe.(fe_size) then
Error (msg "Too many spilled variables, stack size exceeded")
else
OK (Mach.mkfunction
f.(Linear.fn_sig)
(transl_body f fe)
fe.(fe_size)
- (Int.repr fe.(fe_ofs_link))
- (Int.repr fe.(fe_ofs_retaddr))).
+ (Ptrofs.repr fe.(fe_ofs_link))
+ (Ptrofs.repr fe.(fe_ofs_retaddr))).
Definition transf_fundef (f: Linear.fundef) : res Mach.fundef :=
AST.transf_partial_fundef transf_function f.
diff --git a/backend/Stackingproof.v b/backend/Stackingproof.v
index 0e9c58b3..d8d916de 100644
--- a/backend/Stackingproof.v
+++ b/backend/Stackingproof.v
@@ -67,12 +67,14 @@ Lemma load_result_inject:
forall j ty v v',
Val.inject j v v' -> Val.has_type v ty -> Val.inject j v (Val.load_result (chunk_of_type ty) v').
Proof.
- destruct 1; intros; auto; destruct ty; simpl; try contradiction; econstructor; eauto.
+ intros until v'; unfold Val.has_type, Val.load_result; destruct Archi.ptr64;
+ destruct 1; intros; auto; destruct ty; simpl;
+ try contradiction; try discriminate; econstructor; eauto.
Qed.
Section PRESERVATION.
-Variable return_address_offset: Mach.function -> Mach.code -> int -> Prop.
+Variable return_address_offset: Mach.function -> Mach.code -> ptrofs -> Prop.
Hypothesis return_address_offset_exists:
forall f sg ros c,
@@ -100,12 +102,12 @@ Lemma unfold_transf_function:
f.(Linear.fn_sig)
(transl_body f fe)
fe.(fe_size)
- (Int.repr fe.(fe_ofs_link))
- (Int.repr fe.(fe_ofs_retaddr)).
+ (Ptrofs.repr fe.(fe_ofs_link))
+ (Ptrofs.repr fe.(fe_ofs_retaddr)).
Proof.
generalize TRANSF_F. unfold transf_function.
destruct (wt_function f); simpl negb.
- destruct (zlt Int.max_unsigned (fe_size (make_env (function_bounds f)))).
+ destruct (zlt Ptrofs.max_unsigned (fe_size (make_env (function_bounds f)))).
intros; discriminate.
intros. unfold fe. unfold b. congruence.
intros; discriminate.
@@ -118,11 +120,11 @@ Proof.
destruct (wt_function f); simpl negb. auto. intros; discriminate.
Qed.
-Lemma size_no_overflow: fe.(fe_size) <= Int.max_unsigned.
+Lemma size_no_overflow: fe.(fe_size) <= Ptrofs.max_unsigned.
Proof.
generalize TRANSF_F. unfold transf_function.
destruct (wt_function f); simpl negb.
- destruct (zlt Int.max_unsigned (fe_size (make_env (function_bounds f)))).
+ destruct (zlt Ptrofs.max_unsigned (fe_size (make_env (function_bounds f)))).
intros; discriminate.
intros. unfold fe. unfold b. omega.
intros; discriminate.
@@ -143,18 +145,18 @@ Local Opaque Z.add Z.mul Z.divide.
Lemma contains_get_stack:
forall spec m ty sp ofs,
m |= contains (chunk_of_type ty) sp ofs spec ->
- exists v, load_stack m (Vptr sp Int.zero) ty (Int.repr ofs) = Some v /\ spec v.
+ exists v, load_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr ofs) = Some v /\ spec v.
Proof.
intros. unfold load_stack.
- replace (Val.add (Vptr sp Int.zero) (Vint (Int.repr ofs))) with (Vptr sp (Int.repr ofs)).
+ replace (Val.offset_ptr (Vptr sp Ptrofs.zero) (Ptrofs.repr ofs)) with (Vptr sp (Ptrofs.repr ofs)).
eapply loadv_rule; eauto.
- simpl. rewrite Int.add_zero_l; auto.
+ simpl. rewrite Ptrofs.add_zero_l; auto.
Qed.
Lemma hasvalue_get_stack:
forall ty m sp ofs v,
m |= hasvalue (chunk_of_type ty) sp ofs v ->
- load_stack m (Vptr sp Int.zero) ty (Int.repr ofs) = Some v.
+ load_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr ofs) = Some v.
Proof.
intros. exploit contains_get_stack; eauto. intros (v' & A & B). congruence.
Qed.
@@ -164,13 +166,13 @@ Lemma contains_set_stack:
m |= contains (chunk_of_type ty) sp ofs spec1 ** P ->
spec (Val.load_result (chunk_of_type ty) v) ->
exists m',
- store_stack m (Vptr sp Int.zero) ty (Int.repr ofs) v = Some m'
+ store_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr ofs) v = Some m'
/\ m' |= contains (chunk_of_type ty) sp ofs spec ** P.
Proof.
intros. unfold store_stack.
- replace (Val.add (Vptr sp Int.zero) (Vint (Int.repr ofs))) with (Vptr sp (Int.repr ofs)).
+ replace (Val.offset_ptr (Vptr sp Ptrofs.zero) (Ptrofs.repr ofs)) with (Vptr sp (Ptrofs.repr ofs)).
eapply storev_rule; eauto.
- simpl. rewrite Int.add_zero_l; auto.
+ simpl. rewrite Ptrofs.add_zero_l; auto.
Qed.
(** [contains_locations j sp pos bound sl ls] is a separation logic assertion
@@ -184,7 +186,7 @@ Qed.
Program Definition contains_locations (j: meminj) (sp: block) (pos bound: Z) (sl: slot) (ls: locset) : massert := {|
m_pred := fun m =>
- (8 | pos) /\ 0 <= pos /\ pos + 4 * bound <= Int.modulus /\
+ (8 | pos) /\ 0 <= pos /\ pos + 4 * bound <= Ptrofs.modulus /\
Mem.range_perm m sp pos (pos + 4 * bound) Cur Freeable /\
forall ofs ty, 0 <= ofs -> ofs + typesize ty <= bound -> (typealign ty | ofs) ->
exists v, Mem.load (chunk_of_type ty) m sp (pos + 4 * ofs) = Some v
@@ -225,13 +227,13 @@ Lemma get_location:
m |= contains_locations j sp pos bound sl ls ->
0 <= ofs -> ofs + typesize ty <= bound -> (typealign ty | ofs) ->
exists v,
- load_stack m (Vptr sp Int.zero) ty (Int.repr (pos + 4 * ofs)) = Some v
+ load_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr (pos + 4 * ofs)) = Some v
/\ Val.inject j (ls (S sl ofs ty)) v.
Proof.
intros. destruct H as (D & E & F & G & H).
exploit H; eauto. intros (v & U & V). exists v; split; auto.
- unfold load_stack; simpl. rewrite Int.add_zero_l, Int.unsigned_repr; auto.
- unfold Int.max_unsigned. generalize (typesize_pos ty). omega.
+ unfold load_stack; simpl. rewrite Ptrofs.add_zero_l, Ptrofs.unsigned_repr; auto.
+ unfold Ptrofs.max_unsigned. generalize (typesize_pos ty). omega.
Qed.
Lemma set_location:
@@ -240,7 +242,7 @@ Lemma set_location:
0 <= ofs -> ofs + typesize ty <= bound -> (typealign ty | ofs) ->
Val.inject j v v' ->
exists m',
- store_stack m (Vptr sp Int.zero) ty (Int.repr (pos + 4 * ofs)) v' = Some m'
+ store_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr (pos + 4 * ofs)) v' = Some m'
/\ m' |= contains_locations j sp pos bound sl (Locmap.set (S sl ofs ty) v ls) ** P.
Proof.
intros. destruct H as (A & B & C). destruct A as (D & E & F & G & H).
@@ -249,8 +251,8 @@ Proof.
assert (PERM: Mem.range_perm m' sp pos (pos + 4 * bound) Cur Freeable).
{ red; intros; eauto with mem. }
exists m'; split.
-- unfold store_stack; simpl. rewrite Int.add_zero_l, Int.unsigned_repr; eauto.
- unfold Int.max_unsigned. generalize (typesize_pos ty). omega.
+- unfold store_stack; simpl. rewrite Ptrofs.add_zero_l, Ptrofs.unsigned_repr; eauto.
+ unfold Ptrofs.max_unsigned. generalize (typesize_pos ty). omega.
- simpl. intuition auto.
+ unfold Locmap.set.
destruct (Loc.eq (S sl ofs ty) (S sl ofs0 ty0)); [|destruct (Loc.diff_dec (S sl ofs ty) (S sl ofs0 ty0))].
@@ -258,7 +260,7 @@ Proof.
inv e. rename ofs0 into ofs. rename ty0 into ty.
exists (Val.load_result (chunk_of_type ty) v'); split.
eapply Mem.load_store_similar_2; eauto. omega.
- inv H3; destruct (chunk_of_type ty); simpl; econstructor; eauto.
+ apply Val.load_result_inject; auto.
* (* different locations *)
exploit H; eauto. intros (v0 & X & Y). exists v0; split; auto.
rewrite <- X; eapply Mem.load_store_other; eauto.
@@ -366,8 +368,8 @@ represents the Linear stack data. *)
Definition frame_contents_1 (j: meminj) (sp: block) (ls ls0: locset) (parent retaddr: val) :=
contains_locations j sp fe.(fe_ofs_local) b.(bound_local) Local ls
** contains_locations j sp fe_ofs_arg b.(bound_outgoing) Outgoing ls
- ** hasvalue Mint32 sp fe.(fe_ofs_link) parent
- ** hasvalue Mint32 sp fe.(fe_ofs_retaddr) retaddr
+ ** hasvalue Mptr sp fe.(fe_ofs_link) parent
+ ** hasvalue Mptr sp fe.(fe_ofs_retaddr) retaddr
** contains_callee_saves j sp fe.(fe_ofs_callee_save) b.(used_callee_save) ls0.
Definition frame_contents (j: meminj) (sp: block) (ls ls0: locset) (parent retaddr: val) :=
@@ -382,7 +384,7 @@ Lemma frame_get_local:
m |= frame_contents j sp ls ls0 parent retaddr ** P ->
slot_within_bounds b Local ofs ty -> slot_valid f Local ofs ty = true ->
exists v,
- load_stack m (Vptr sp Int.zero) ty (Int.repr (offset_local fe ofs)) = Some v
+ load_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr (offset_local fe ofs)) = Some v
/\ Val.inject j (ls (S Local ofs ty)) v.
Proof.
unfold frame_contents, frame_contents_1; intros. unfold slot_valid in H1; InvBooleans.
@@ -395,7 +397,7 @@ Lemma frame_get_outgoing:
m |= frame_contents j sp ls ls0 parent retaddr ** P ->
slot_within_bounds b Outgoing ofs ty -> slot_valid f Outgoing ofs ty = true ->
exists v,
- load_stack m (Vptr sp Int.zero) ty (Int.repr (offset_arg ofs)) = Some v
+ load_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr (offset_arg ofs)) = Some v
/\ Val.inject j (ls (S Outgoing ofs ty)) v.
Proof.
unfold frame_contents, frame_contents_1; intros. unfold slot_valid in H1; InvBooleans.
@@ -406,20 +408,20 @@ Qed.
Lemma frame_get_parent:
forall j sp ls ls0 parent retaddr m P,
m |= frame_contents j sp ls ls0 parent retaddr ** P ->
- load_stack m (Vptr sp Int.zero) Tint (Int.repr fe.(fe_ofs_link)) = Some parent.
+ load_stack m (Vptr sp Ptrofs.zero) Tptr (Ptrofs.repr fe.(fe_ofs_link)) = Some parent.
Proof.
unfold frame_contents, frame_contents_1; intros.
- apply mconj_proj1 in H. apply sep_proj1 in H. apply sep_pick3 in H.
+ apply mconj_proj1 in H. apply sep_proj1 in H. apply sep_pick3 in H. rewrite <- chunk_of_Tptr in H.
eapply hasvalue_get_stack; eauto.
Qed.
Lemma frame_get_retaddr:
forall j sp ls ls0 parent retaddr m P,
m |= frame_contents j sp ls ls0 parent retaddr ** P ->
- load_stack m (Vptr sp Int.zero) Tint (Int.repr fe.(fe_ofs_retaddr)) = Some retaddr.
+ load_stack m (Vptr sp Ptrofs.zero) Tptr (Ptrofs.repr fe.(fe_ofs_retaddr)) = Some retaddr.
Proof.
unfold frame_contents, frame_contents_1; intros.
- apply mconj_proj1 in H. apply sep_proj1 in H. apply sep_pick4 in H.
+ apply mconj_proj1 in H. apply sep_proj1 in H. apply sep_pick4 in H. rewrite <- chunk_of_Tptr in H.
eapply hasvalue_get_stack; eauto.
Qed.
@@ -431,7 +433,7 @@ Lemma frame_set_local:
slot_within_bounds b Local ofs ty -> slot_valid f Local ofs ty = true ->
Val.inject j v v' ->
exists m',
- store_stack m (Vptr sp Int.zero) ty (Int.repr (offset_local fe ofs)) v' = Some m'
+ store_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr (offset_local fe ofs)) v' = Some m'
/\ m' |= frame_contents j sp (Locmap.set (S Local ofs ty) v ls) ls0 parent retaddr ** P.
Proof.
intros. unfold frame_contents in H.
@@ -456,7 +458,7 @@ Lemma frame_set_outgoing:
slot_within_bounds b Outgoing ofs ty -> slot_valid f Outgoing ofs ty = true ->
Val.inject j v v' ->
exists m',
- store_stack m (Vptr sp Int.zero) ty (Int.repr (offset_arg ofs)) v' = Some m'
+ store_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr (offset_arg ofs)) v' = Some m'
/\ m' |= frame_contents j sp (Locmap.set (S Outgoing ofs ty) v ls) ls0 parent retaddr ** P.
Proof.
intros. unfold frame_contents in H.
@@ -855,7 +857,8 @@ Qed.
Remark destroyed_by_store_caller_save:
forall chunk addr, no_callee_saves (destroyed_by_store chunk addr).
Proof.
- unfold no_callee_saves; destruct chunk; reflexivity.
+Local Transparent destroyed_by_store.
+ unfold no_callee_saves, destroyed_by_store; intros; destruct chunk; try reflexivity; destruct Archi.ptr64; reflexivity.
Qed.
Remark destroyed_by_cond_caller_save:
@@ -939,12 +942,13 @@ Lemma save_callee_save_rec_correct:
agree_regs j ls rs ->
exists rs', exists m',
star step tge
- (State cs fb (Vptr sp Int.zero) (save_callee_save_rec l pos k) rs m)
- E0 (State cs fb (Vptr sp Int.zero) k rs' m')
+ (State cs fb (Vptr sp Ptrofs.zero) (save_callee_save_rec l pos k) rs m)
+ E0 (State cs fb (Vptr sp Ptrofs.zero) k rs' m')
/\ m' |= contains_callee_saves j sp pos l ls ** P
/\ (forall ofs k p, Mem.perm m sp ofs k p -> Mem.perm m' sp ofs k p)
/\ agree_regs j ls rs'.
Proof.
+Local Opaque mreg_type.
induction l as [ | r l]; simpl; intros until P; intros CS SEP AG.
- exists rs, m.
split. apply star_refl.
@@ -1029,8 +1033,8 @@ Lemma save_callee_save_correct:
let rs1 := undef_regs destroyed_at_function_entry rs in
exists rs', exists m',
star step tge
- (State cs fb (Vptr sp Int.zero) (save_callee_save fe k) rs1 m)
- E0 (State cs fb (Vptr sp Int.zero) k rs' m')
+ (State cs fb (Vptr sp Ptrofs.zero) (save_callee_save fe k) rs1 m)
+ E0 (State cs fb (Vptr sp Ptrofs.zero) k rs' m')
/\ m' |= contains_callee_saves j sp fe.(fe_ofs_callee_save) b.(used_callee_save) ls0 ** P
/\ (forall ofs k p, Mem.perm m sp ofs k p -> Mem.perm m' sp ofs k p)
/\ agree_regs j ls1 rs'.
@@ -1071,15 +1075,15 @@ Lemma function_prologue_correct:
ls1 = LTL.undef_regs destroyed_at_function_entry (LTL.call_regs ls) ->
rs1 = undef_regs destroyed_at_function_entry rs ->
Mem.alloc m1 0 f.(Linear.fn_stacksize) = (m2, sp) ->
- Val.has_type parent Tint -> Val.has_type ra Tint ->
+ Val.has_type parent Tptr -> Val.has_type ra Tptr ->
m1' |= minjection j m1 ** globalenv_inject ge j ** P ->
exists j', exists rs', exists m2', exists sp', exists m3', exists m4', exists m5',
Mem.alloc m1' 0 tf.(fn_stacksize) = (m2', sp')
- /\ store_stack m2' (Vptr sp' Int.zero) Tint tf.(fn_link_ofs) parent = Some m3'
- /\ store_stack m3' (Vptr sp' Int.zero) Tint tf.(fn_retaddr_ofs) ra = Some m4'
+ /\ store_stack m2' (Vptr sp' Ptrofs.zero) Tptr tf.(fn_link_ofs) parent = Some m3'
+ /\ store_stack m3' (Vptr sp' Ptrofs.zero) Tptr tf.(fn_retaddr_ofs) ra = Some m4'
/\ star step tge
- (State cs fb (Vptr sp' Int.zero) (save_callee_save fe k) rs1 m4')
- E0 (State cs fb (Vptr sp' Int.zero) k rs' m5')
+ (State cs fb (Vptr sp' Ptrofs.zero) (save_callee_save fe k) rs1 m4')
+ E0 (State cs fb (Vptr sp' Ptrofs.zero) k rs' m5')
/\ agree_regs j' ls1 rs'
/\ agree_locs ls1 ls0
/\ m5' |= frame_contents j' sp' ls1 ls0 parent ra ** minjection j' m2 ** globalenv_inject ge j' ** P
@@ -1113,17 +1117,17 @@ Local Opaque b fe.
(* Dividing up the frame *)
apply (frame_env_separated b) in SEP. replace (make_env b) with fe in SEP by auto.
(* Store of parent *)
- rewrite sep_swap3 in SEP.
- apply (range_contains Mint32) in SEP; [|tauto].
- exploit (contains_set_stack (fun v' => v' = parent) parent (fun _ => True) m2' Tint).
- eexact SEP. apply Val.load_result_same; auto.
+ rewrite sep_swap3 in SEP.
+ apply (range_contains Mptr) in SEP; [|tauto].
+ exploit (contains_set_stack (fun v' => v' = parent) parent (fun _ => True) m2' Tptr).
+ rewrite chunk_of_Tptr; eexact SEP. apply Val.load_result_same; auto.
clear SEP; intros (m3' & STORE_PARENT & SEP).
rewrite sep_swap3 in SEP.
(* Store of return address *)
rewrite sep_swap4 in SEP.
- apply (range_contains Mint32) in SEP; [|tauto].
- exploit (contains_set_stack (fun v' => v' = ra) ra (fun _ => True) m3' Tint).
- eexact SEP. apply Val.load_result_same; auto.
+ apply (range_contains Mptr) in SEP; [|tauto].
+ exploit (contains_set_stack (fun v' => v' = ra) ra (fun _ => True) m3' Tptr).
+ rewrite chunk_of_Tptr; eexact SEP. apply Val.load_result_same; auto.
clear SEP; intros (m4' & STORE_RETADDR & SEP).
rewrite sep_swap4 in SEP.
(* Saving callee-save registers *)
@@ -1147,7 +1151,8 @@ Local Opaque b fe.
rewrite sep_swap in SEP.
(* Now we frame this *)
assert (SEPFINAL: m5' |= frame_contents j' sp' ls1 ls0 parent ra ** minjection j' m2 ** globalenv_inject ge j' ** P).
- { eapply frame_mconj. eexact SEPCONJ.
+ { eapply frame_mconj. eexact SEPCONJ.
+ rewrite chunk_of_Tptr in SEP.
unfold frame_contents_1; rewrite ! sep_assoc. exact SEP.
assert (forall ofs k p, Mem.perm m2' sp' ofs k p -> Mem.perm m5' sp' ofs k p).
{ intros. apply PERMS.
@@ -1198,12 +1203,13 @@ Lemma restore_callee_save_rec_correct:
(forall r, In r l -> mreg_within_bounds b r) ->
exists rs',
star step tge
- (State cs fb (Vptr sp Int.zero) (restore_callee_save_rec l ofs k) rs m)
- E0 (State cs fb (Vptr sp Int.zero) k rs' m)
+ (State cs fb (Vptr sp Ptrofs.zero) (restore_callee_save_rec l ofs k) rs m)
+ E0 (State cs fb (Vptr sp Ptrofs.zero) k rs' m)
/\ (forall r, In r l -> Val.inject j (ls0 (R r)) (rs' r))
/\ (forall r, ~(In r l) -> rs' r = rs r)
/\ agree_unused ls0 rs'.
Proof.
+Local Opaque mreg_type.
induction l as [ | r l]; simpl; intros.
- (* base case *)
exists rs. intuition auto. apply star_refl.
@@ -1242,8 +1248,8 @@ Lemma restore_callee_save_correct:
agree_unused j ls0 rs ->
exists rs',
star step tge
- (State cs fb (Vptr sp Int.zero) (restore_callee_save fe k) rs m)
- E0 (State cs fb (Vptr sp Int.zero) k rs' m)
+ (State cs fb (Vptr sp Ptrofs.zero) (restore_callee_save fe k) rs m)
+ E0 (State cs fb (Vptr sp Ptrofs.zero) k rs' m)
/\ (forall r,
is_callee_save r = true -> Val.inject j (ls0 (R r)) (rs' r))
/\ (forall r,
@@ -1277,12 +1283,12 @@ Lemma function_epilogue_correct:
j sp = Some(sp', fe.(fe_stack_data)) ->
Mem.free m sp 0 f.(Linear.fn_stacksize) = Some m1 ->
exists rs1, exists m1',
- load_stack m' (Vptr sp' Int.zero) Tint tf.(fn_link_ofs) = Some pa
- /\ load_stack m' (Vptr sp' Int.zero) Tint tf.(fn_retaddr_ofs) = Some ra
+ load_stack m' (Vptr sp' Ptrofs.zero) Tptr tf.(fn_link_ofs) = Some pa
+ /\ load_stack m' (Vptr sp' Ptrofs.zero) Tptr tf.(fn_retaddr_ofs) = Some ra
/\ Mem.free m' sp' 0 tf.(fn_stacksize) = Some m1'
/\ star step tge
- (State cs fb (Vptr sp' Int.zero) (restore_callee_save fe k) rs m')
- E0 (State cs fb (Vptr sp' Int.zero) k rs1 m')
+ (State cs fb (Vptr sp' Ptrofs.zero) (restore_callee_save fe k) rs m')
+ E0 (State cs fb (Vptr sp' Ptrofs.zero) k rs1 m')
/\ agree_regs j (return_regs ls0 ls) rs1
/\ agree_callee_save (return_regs ls0 ls) ls0
/\ m1' |= minjection j m1 ** P.
@@ -1304,8 +1310,8 @@ Proof.
(* Reloading the back link and return address *)
unfold frame_contents in SEP; apply mconj_proj1 in SEP.
unfold frame_contents_1 in SEP; rewrite ! sep_assoc in SEP.
- exploit (hasvalue_get_stack Tint). eapply sep_pick3; eexact SEP. intros LOAD_LINK.
- exploit (hasvalue_get_stack Tint). eapply sep_pick4; eexact SEP. intros LOAD_RETADDR.
+ exploit (hasvalue_get_stack Tptr). rewrite chunk_of_Tptr. eapply sep_pick3; eexact SEP. intros LOAD_LINK.
+ exploit (hasvalue_get_stack Tptr). rewrite chunk_of_Tptr. eapply sep_pick4; eexact SEP. intros LOAD_RETADDR.
clear SEP.
(* Conclusions *)
rewrite unfold_transf_function; simpl.
@@ -1353,15 +1359,15 @@ Inductive match_stacks (j: meminj):
(TRF: transf_function f = OK trf)
(TRC: transl_code (make_env (function_bounds f)) c = c')
(INJ: j sp = Some(sp', (fe_stack_data (make_env (function_bounds f)))))
- (TY_RA: Val.has_type ra Tint)
+ (TY_RA: Val.has_type ra Tptr)
(AGL: agree_locs f ls (parent_locset cs))
(ARGS: forall ofs ty,
In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments sg)) ->
slot_within_bounds (function_bounds f) Outgoing ofs ty)
(STK: match_stacks j cs cs' (Linear.fn_sig f)),
match_stacks j
- (Linear.Stackframe f (Vptr sp Int.zero) ls c :: cs)
- (Stackframe fb (Vptr sp' Int.zero) ra c' :: cs')
+ (Linear.Stackframe f (Vptr sp Ptrofs.zero) ls c :: cs)
+ (Stackframe fb (Vptr sp' Ptrofs.zero) ra c' :: cs')
sg.
(** Invariance with respect to change of memory injection. *)
@@ -1409,17 +1415,17 @@ Qed.
Lemma match_stacks_type_sp:
forall j cs cs' sg,
match_stacks j cs cs' sg ->
- Val.has_type (parent_sp cs') Tint.
+ Val.has_type (parent_sp cs') Tptr.
Proof.
- induction 1; simpl; auto.
-Qed.
+ induction 1; unfold parent_sp. apply Val.Vnullptr_has_type. apply Val.Vptr_has_type.
+Qed.
Lemma match_stacks_type_retaddr:
forall j cs cs' sg,
match_stacks j cs cs' sg ->
- Val.has_type (parent_ra cs') Tint.
+ Val.has_type (parent_ra cs') Tptr.
Proof.
- induction 1; simpl; auto.
+ induction 1; unfold parent_ra. apply Val.Vnullptr_has_type. auto.
Qed.
(** * Syntactic properties of the translation *)
@@ -1700,11 +1706,11 @@ Hypothesis SEP: m' |= frame_contents f j sp' ls ls0 parent retaddr ** minjection
Lemma transl_builtin_arg_correct:
forall a v,
- eval_builtin_arg ge ls (Vptr sp Int.zero) m a v ->
+ eval_builtin_arg ge ls (Vptr sp Ptrofs.zero) m a v ->
(forall l, In l (params_of_builtin_arg a) -> loc_valid f l = true) ->
(forall sl ofs ty, In (S sl ofs ty) (params_of_builtin_arg a) -> slot_within_bounds b sl ofs ty) ->
exists v',
- eval_builtin_arg ge rs (Vptr sp' Int.zero) m' (transl_builtin_arg fe a) v'
+ eval_builtin_arg ge rs (Vptr sp' Ptrofs.zero) m' (transl_builtin_arg fe a) v'
/\ Val.inject j v v'.
Proof.
assert (SYMB: forall id ofs, Val.inject j (Senv.symbol_address ge id ofs) (Senv.symbol_address ge id ofs)).
@@ -1712,7 +1718,7 @@ Proof.
{ eapply globalenv_inject_preserves_globals. eapply sep_proj2. eapply sep_proj2. eexact SEP. }
intros; unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
destruct (Genv.find_symbol ge id) eqn:FS; auto.
- destruct G. econstructor. eauto. rewrite Int.add_zero; auto. }
+ destruct G. econstructor. eauto. rewrite Ptrofs.add_zero; auto. }
Local Opaque fe.
induction 1; simpl; intros VALID BOUNDS.
- assert (loc_valid f x = true) by auto.
@@ -1724,13 +1730,13 @@ Local Opaque fe.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
-- set (ofs' := Int.add ofs (Int.repr (fe_stack_data fe))).
+- set (ofs' := Ptrofs.add ofs (Ptrofs.repr (fe_stack_data fe))).
apply sep_proj2 in SEP. apply sep_proj1 in SEP. exploit loadv_parallel_rule; eauto.
- instantiate (1 := Val.add (Vptr sp' Int.zero) (Vint ofs')).
- simpl. rewrite ! Int.add_zero_l. econstructor; eauto.
+ instantiate (1 := Val.offset_ptr (Vptr sp' Ptrofs.zero) ofs').
+ simpl. rewrite ! Ptrofs.add_zero_l. econstructor; eauto.
intros (v' & A & B). exists v'; split; auto. constructor; auto.
- econstructor; split; eauto with barg.
- unfold Val.add. rewrite ! Int.add_zero_l. econstructor; eauto.
+ unfold Val.offset_ptr. rewrite ! Ptrofs.add_zero_l. econstructor; eauto.
- apply sep_proj2 in SEP. apply sep_proj1 in SEP. exploit loadv_parallel_rule; eauto.
intros (v' & A & B). exists v'; auto with barg.
- econstructor; split; eauto with barg.
@@ -1742,11 +1748,11 @@ Qed.
Lemma transl_builtin_args_correct:
forall al vl,
- eval_builtin_args ge ls (Vptr sp Int.zero) m al vl ->
+ eval_builtin_args ge ls (Vptr sp Ptrofs.zero) m al vl ->
(forall l, In l (params_of_builtin_args al) -> loc_valid f l = true) ->
(forall sl ofs ty, In (S sl ofs ty) (params_of_builtin_args al) -> slot_within_bounds b sl ofs ty) ->
exists vl',
- eval_builtin_args ge rs (Vptr sp' Int.zero) m' (List.map (transl_builtin_arg fe) al) vl'
+ eval_builtin_args ge rs (Vptr sp' Ptrofs.zero) m' (List.map (transl_builtin_arg fe) al) vl'
/\ Val.inject_list j vl vl'.
Proof.
induction 1; simpl; intros VALID BOUNDS.
@@ -1798,8 +1804,8 @@ Inductive match_states: Linear.state -> Mach.state -> Prop :=
** stack_contents j cs cs'
** minjection j m
** globalenv_inject ge j),
- match_states (Linear.State cs f (Vptr sp Int.zero) c ls m)
- (Mach.State cs' fb (Vptr sp' Int.zero) (transl_code (make_env (function_bounds f)) c) rs m')
+ match_states (Linear.State cs f (Vptr sp Ptrofs.zero) c ls m)
+ (Mach.State cs' fb (Vptr sp' Ptrofs.zero) (transl_code (make_env (function_bounds f)) c) rs m')
| match_states_call:
forall cs f ls m cs' fb rs m' j tf
(STACKS: match_stacks j cs cs' (Linear.funsig f))
@@ -1882,7 +1888,7 @@ Proof.
end).
eapply frame_undef_regs with (rl := destroyed_by_setstack ty) in SEP.
assert (A: exists m'',
- store_stack m' (Vptr sp' Int.zero) ty (Int.repr ofs') (rs0 src) = Some m''
+ store_stack m' (Vptr sp' Ptrofs.zero) ty (Ptrofs.repr ofs') (rs0 src) = Some m''
/\ m'' |= frame_contents f j sp' (Locmap.set (S sl ofs ty) (rs (R src))
(LTL.undef_regs (destroyed_by_setstack ty) rs))
(parent_locset s) (parent_sp cs') (parent_ra cs')
@@ -1902,7 +1908,7 @@ Proof.
- (* Lop *)
assert (exists v',
- eval_operation ge (Vptr sp' Int.zero) (transl_op (make_env (function_bounds f)) op) rs0##args m' = Some v'
+ eval_operation ge (Vptr sp' Ptrofs.zero) (transl_op (make_env (function_bounds f)) op) rs0##args m' = Some v'
/\ Val.inject j v v').
eapply eval_operation_inject; eauto.
eapply globalenv_inject_preserves_globals. eapply sep_proj2. eapply sep_proj2. eapply sep_proj2. eexact SEP.
@@ -1921,7 +1927,7 @@ Proof.
- (* Lload *)
assert (exists a',
- eval_addressing ge (Vptr sp' Int.zero) (transl_addr (make_env (function_bounds f)) addr) rs0##args = Some a'
+ eval_addressing ge (Vptr sp' Ptrofs.zero) (transl_addr (make_env (function_bounds f)) addr) rs0##args = Some a'
/\ Val.inject j a a').
eapply eval_addressing_inject; eauto.
eapply globalenv_inject_preserves_globals. eapply sep_proj2. eapply sep_proj2. eapply sep_proj2. eexact SEP.
@@ -1941,7 +1947,7 @@ Proof.
- (* Lstore *)
assert (exists a',
- eval_addressing ge (Vptr sp' Int.zero) (transl_addr (make_env (function_bounds f)) addr) rs0##args = Some a'
+ eval_addressing ge (Vptr sp' Ptrofs.zero) (transl_addr (make_env (function_bounds f)) addr) rs0##args = Some a'
/\ Val.inject j a a').
eapply eval_addressing_inject; eauto.
eapply globalenv_inject_preserves_globals. eapply sep_proj2. eapply sep_proj2. eapply sep_proj2. eexact SEP.
@@ -1972,7 +1978,7 @@ Proof.
apply plus_one. econstructor; eauto.
econstructor; eauto.
econstructor; eauto with coqlib.
- simpl; auto.
+ apply Val.Vptr_has_type.
intros; red.
apply Zle_trans with (size_arguments (Linear.funsig f')); auto.
apply loc_arguments_bounded; auto.
@@ -2150,7 +2156,11 @@ Lemma transf_final_states:
match_states st1 st2 -> Linear.final_state st1 r -> Mach.final_state st2 r.
Proof.
intros. inv H0. inv H. inv STACKS.
- assert (R: exists r, loc_result signature_main = One r) by (econstructor; reflexivity).
+ assert (R: exists r, loc_result signature_main = One r).
+ { destruct (loc_result signature_main) as [r1 | r1 r2] eqn:LR.
+ - exists r1; auto.
+ - generalize (loc_result_type signature_main). rewrite LR. discriminate.
+ }
destruct R as [rres EQ]. rewrite EQ in H1. simpl in H1.
generalize (AGREGS rres). rewrite H1. intros A; inv A.
econstructor; eauto.
diff --git a/backend/Tailcallproof.v b/backend/Tailcallproof.v
index 793dc861..1dcdfb64 100644
--- a/backend/Tailcallproof.v
+++ b/backend/Tailcallproof.v
@@ -310,14 +310,14 @@ Inductive match_stackframes: list stackframe -> list stackframe -> Prop :=
match_stackframes stk stk' ->
regs_lessdef rs rs' ->
match_stackframes
- (Stackframe res f (Vptr sp Int.zero) pc rs :: stk)
- (Stackframe res (transf_function f) (Vptr sp Int.zero) pc rs' :: stk')
+ (Stackframe res f (Vptr sp Ptrofs.zero) pc rs :: stk)
+ (Stackframe res (transf_function f) (Vptr sp Ptrofs.zero) pc rs' :: stk')
| match_stackframes_tail: forall stk stk' res sp pc rs f,
match_stackframes stk stk' ->
is_return_spec f pc res ->
f.(fn_stacksize) = 0 ->
match_stackframes
- (Stackframe res f (Vptr sp Int.zero) pc rs :: stk)
+ (Stackframe res f (Vptr sp Ptrofs.zero) pc rs :: stk)
stk'.
(** Here is the invariant relating two states. The first three
@@ -331,8 +331,8 @@ Inductive match_states: state -> state -> Prop :=
(STACKS: match_stackframes s s')
(RLD: regs_lessdef rs rs')
(MLD: Mem.extends m m'),
- match_states (State s f (Vptr sp Int.zero) pc rs m)
- (State s' (transf_function f) (Vptr sp Int.zero) pc rs' m')
+ match_states (State s f (Vptr sp Ptrofs.zero) pc rs m)
+ (State s' (transf_function f) (Vptr sp Ptrofs.zero) pc rs' m')
| match_states_call:
forall s f args m s' args' m',
match_stackframes s s' ->
@@ -354,7 +354,7 @@ Inductive match_states: state -> state -> Prop :=
is_return_spec f pc r ->
f.(fn_stacksize) = 0 ->
Val.lessdef (rs#r) v' ->
- match_states (State s f (Vptr sp Int.zero) pc rs m)
+ match_states (State s f (Vptr sp Ptrofs.zero) pc rs m)
(Returnstate s' v' m').
(** The last case of [match_states] corresponds to the execution
@@ -417,7 +417,7 @@ Proof.
assert (Val.lessdef_list (rs##args) (rs'##args)). apply regs_lessdef_regs; auto.
exploit eval_operation_lessdef; eauto.
intros [v' [EVAL' VLD]].
- left. exists (State s' (transf_function f) (Vptr sp0 Int.zero) pc' (rs'#res <- v') m'); split.
+ left. exists (State s' (transf_function f) (Vptr sp0 Ptrofs.zero) pc' (rs'#res <- v') m'); split.
eapply exec_Iop; eauto. rewrite <- EVAL'.
apply eval_operation_preserved. exact symbols_preserved.
econstructor; eauto. apply set_reg_lessdef; auto.
@@ -433,7 +433,7 @@ Proof.
intros [a' [ADDR' ALD]].
exploit Mem.loadv_extends; eauto.
intros [v' [LOAD' VLD]].
- left. exists (State s' (transf_function f) (Vptr sp0 Int.zero) pc' (rs'#dst <- v') m'); split.
+ left. exists (State s' (transf_function f) (Vptr sp0 Ptrofs.zero) pc' (rs'#dst <- v') m'); split.
eapply exec_Iload with (a := a'). eauto. rewrite <- ADDR'.
apply eval_addressing_preserved. exact symbols_preserved. eauto.
econstructor; eauto. apply set_reg_lessdef; auto.
@@ -445,7 +445,7 @@ Proof.
intros [a' [ADDR' ALD]].
exploit Mem.storev_extends. 2: eexact H1. eauto. eauto. apply RLD.
intros [m'1 [STORE' MLD']].
- left. exists (State s' (transf_function f) (Vptr sp0 Int.zero) pc' rs' m'1); split.
+ left. exists (State s' (transf_function f) (Vptr sp0 Ptrofs.zero) pc' rs' m'1); split.
eapply exec_Istore with (a := a'). eauto. rewrite <- ADDR'.
apply eval_addressing_preserved. exact symbols_preserved. eauto.
destruct a; simpl in H1; try discriminate.
@@ -465,7 +465,7 @@ Proof.
eapply Mem.free_right_extends; eauto.
rewrite stacksize_preserved. rewrite H7. intros. omegaContradiction.
+ (* call that remains a call *)
- left. exists (Callstate (Stackframe res (transf_function f) (Vptr sp0 Int.zero) pc' rs' :: s')
+ left. exists (Callstate (Stackframe res (transf_function f) (Vptr sp0 Ptrofs.zero) pc' rs' :: s')
(transf_fundef fd) (rs'##args) m'); split.
eapply exec_Icall; eauto. apply sig_preserved.
constructor. constructor; auto. apply regs_lessdef_regs; auto. auto.
@@ -485,7 +485,7 @@ Proof.
intros (vargs' & P & Q).
exploit external_call_mem_extends; eauto.
intros [v' [m'1 [A [B [C D]]]]].
- left. exists (State s' (transf_function f) (Vptr sp0 Int.zero) pc' (regmap_setres res v' rs') m'1); split.
+ left. exists (State s' (transf_function f) (Vptr sp0 Ptrofs.zero) pc' (regmap_setres res v' rs') m'1); split.
eapply exec_Ibuiltin; eauto.
eapply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved.
eapply external_call_symbols_preserved; eauto. apply senv_preserved.
@@ -493,14 +493,14 @@ Proof.
- (* cond *)
TransfInstr.
- left. exists (State s' (transf_function f) (Vptr sp0 Int.zero) (if b then ifso else ifnot) rs' m'); split.
+ left. exists (State s' (transf_function f) (Vptr sp0 Ptrofs.zero) (if b then ifso else ifnot) rs' m'); split.
eapply exec_Icond; eauto.
apply eval_condition_lessdef with (rs##args) m; auto. apply regs_lessdef_regs; auto.
constructor; auto.
- (* jumptable *)
TransfInstr.
- left. exists (State s' (transf_function f) (Vptr sp0 Int.zero) pc' rs' m'); split.
+ left. exists (State s' (transf_function f) (Vptr sp0 Ptrofs.zero) pc' rs' m'); split.
eapply exec_Ijumptable; eauto.
generalize (RLD arg). rewrite H0. intro. inv H2. auto.
constructor; auto.
diff --git a/backend/Unusedglobproof.v b/backend/Unusedglobproof.v
index 44cf1e8a..7e9c3ca0 100644
--- a/backend/Unusedglobproof.v
+++ b/backend/Unusedglobproof.v
@@ -627,7 +627,7 @@ Lemma symbol_address_inject:
Proof.
intros. unfold Genv.symbol_address. destruct (Genv.find_symbol ge id) as [b|] eqn:FS; auto.
exploit symbols_inject_2; eauto. intros (b' & TFS & INJ). rewrite TFS.
- econstructor; eauto. rewrite Int.add_zero; auto.
+ econstructor; eauto. rewrite Ptrofs.add_zero; auto.
Qed.
(** Semantic preservation *)
@@ -691,8 +691,8 @@ Inductive match_stacks (j: meminj):
(REGINJ: regset_inject j rs trs)
(BELOW: Plt sp bound)
(TBELOW: Plt tsp tbound),
- match_stacks j (Stackframe res f (Vptr sp Int.zero) pc rs :: s)
- (Stackframe res f (Vptr tsp Int.zero) pc trs :: ts)
+ match_stacks j (Stackframe res f (Vptr sp Ptrofs.zero) pc rs :: s)
+ (Stackframe res f (Vptr tsp Ptrofs.zero) pc trs :: ts)
bound tbound.
Lemma match_stacks_preserves_globals:
@@ -759,8 +759,8 @@ Inductive match_states: state -> state -> Prop :=
(SPINJ: j sp = Some(tsp, 0))
(REGINJ: regset_inject j rs trs)
(MEMINJ: Mem.inject j m tm),
- match_states (State s f (Vptr sp Int.zero) pc rs m)
- (State ts f (Vptr tsp Int.zero) pc trs tm)
+ match_states (State s f (Vptr sp Ptrofs.zero) pc rs m)
+ (State ts f (Vptr tsp Ptrofs.zero) pc trs tm)
| match_states_call: forall s fd args m ts targs tm j
(STACKS: match_stacks j s ts (Mem.nextblock m) (Mem.nextblock tm))
(KEPT: forall id, ref_fundef fd id -> kept id)
@@ -819,14 +819,14 @@ Qed.
Lemma eval_builtin_arg_inject:
forall rs sp m j rs' sp' m' a v,
- eval_builtin_arg ge (fun r => rs#r) (Vptr sp Int.zero) m a v ->
+ eval_builtin_arg ge (fun r => rs#r) (Vptr sp Ptrofs.zero) m a v ->
j sp = Some(sp', 0) ->
meminj_preserves_globals j ->
regset_inject j rs rs' ->
Mem.inject j m m' ->
(forall id, In id (globals_of_builtin_arg a) -> kept id) ->
exists v',
- eval_builtin_arg tge (fun r => rs'#r) (Vptr sp' Int.zero) m' a v'
+ eval_builtin_arg tge (fun r => rs'#r) (Vptr sp' Ptrofs.zero) m' a v'
/\ Val.inject j v v'.
Proof.
induction 1; intros SP GL RS MI K; simpl in K.
@@ -837,18 +837,18 @@ Proof.
- econstructor; eauto with barg.
- simpl in H. exploit Mem.load_inject; eauto. rewrite Zplus_0_r.
intros (v' & A & B). exists v'; auto with barg.
-- econstructor; split; eauto with barg. simpl. econstructor; eauto. rewrite Int.add_zero; auto.
+- econstructor; split; eauto with barg. simpl. econstructor; eauto. rewrite Ptrofs.add_zero; auto.
- assert (Val.inject j (Senv.symbol_address ge id ofs) (Senv.symbol_address tge id ofs)).
{ unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
destruct (Genv.find_symbol ge id) as [b|] eqn:FS; auto.
exploit symbols_inject_2; eauto. intros (b' & A & B). rewrite A.
- econstructor; eauto. rewrite Int.add_zero; auto. }
+ econstructor; eauto. rewrite Ptrofs.add_zero; auto. }
exploit Mem.loadv_inject; eauto. intros (v' & A & B). exists v'; auto with barg.
- econstructor; split; eauto with barg.
unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
destruct (Genv.find_symbol ge id) as [b|] eqn:FS; auto.
exploit symbols_inject_2; eauto. intros (b' & A & B). rewrite A.
- econstructor; eauto. rewrite Int.add_zero; auto.
+ econstructor; eauto. rewrite Ptrofs.add_zero; auto.
- destruct IHeval_builtin_arg1 as (v1' & A1 & B1); eauto using in_or_app.
destruct IHeval_builtin_arg2 as (v2' & A2 & B2); eauto using in_or_app.
exists (Val.longofwords v1' v2'); split; auto with barg.
@@ -857,14 +857,14 @@ Qed.
Lemma eval_builtin_args_inject:
forall rs sp m j rs' sp' m' al vl,
- eval_builtin_args ge (fun r => rs#r) (Vptr sp Int.zero) m al vl ->
+ eval_builtin_args ge (fun r => rs#r) (Vptr sp Ptrofs.zero) m al vl ->
j sp = Some(sp', 0) ->
meminj_preserves_globals j ->
regset_inject j rs rs' ->
Mem.inject j m m' ->
(forall id, In id (globals_of_builtin_args al) -> kept id) ->
exists vl',
- eval_builtin_args tge (fun r => rs'#r) (Vptr sp' Int.zero) m' al vl'
+ eval_builtin_args tge (fun r => rs'#r) (Vptr sp' Ptrofs.zero) m' al vl'
/\ Val.inject_list j vl vl'.
Proof.
induction 1; intros.
@@ -889,9 +889,9 @@ Proof.
- (* op *)
assert (A: exists tv,
- eval_operation tge (Vptr tsp Int.zero) op trs##args tm = Some tv
+ eval_operation tge (Vptr tsp Ptrofs.zero) op trs##args tm = Some tv
/\ Val.inject j v tv).
- { apply eval_operation_inj with (ge1 := ge) (m1 := m) (sp1 := Vptr sp0 Int.zero) (vl1 := rs##args).
+ { apply eval_operation_inj with (ge1 := ge) (m1 := m) (sp1 := Vptr sp0 Ptrofs.zero) (vl1 := rs##args).
intros; eapply Mem.valid_pointer_inject_val; eauto.
intros; eapply Mem.weak_valid_pointer_inject_val; eauto.
intros; eapply Mem.weak_valid_pointer_inject_no_overflow; eauto.
@@ -907,9 +907,9 @@ Proof.
- (* load *)
assert (A: exists ta,
- eval_addressing tge (Vptr tsp Int.zero) addr trs##args = Some ta
+ eval_addressing tge (Vptr tsp Ptrofs.zero) addr trs##args = Some ta
/\ Val.inject j a ta).
- { apply eval_addressing_inj with (ge1 := ge) (sp1 := Vptr sp0 Int.zero) (vl1 := rs##args).
+ { apply eval_addressing_inj with (ge1 := ge) (sp1 := Vptr sp0 Ptrofs.zero) (vl1 := rs##args).
intros. apply symbol_address_inject. eapply match_stacks_preserves_globals; eauto.
apply KEPT. red. exists pc, (Iload chunk addr args dst pc'); auto.
econstructor; eauto.
@@ -922,9 +922,9 @@ Proof.
- (* store *)
assert (A: exists ta,
- eval_addressing tge (Vptr tsp Int.zero) addr trs##args = Some ta
+ eval_addressing tge (Vptr tsp Ptrofs.zero) addr trs##args = Some ta
/\ Val.inject j a ta).
- { apply eval_addressing_inj with (ge1 := ge) (sp1 := Vptr sp0 Int.zero) (vl1 := rs##args).
+ { apply eval_addressing_inj with (ge1 := ge) (sp1 := Vptr sp0 Ptrofs.zero) (vl1 := rs##args).
intros. apply symbol_address_inject. eapply match_stacks_preserves_globals; eauto.
apply KEPT. red. exists pc, (Istore chunk addr args src pc'); auto.
econstructor; eauto.
@@ -1104,11 +1104,11 @@ Proof.
assert (kept i). { apply H. red. exists i0; auto with coqlib. }
exploit symbols_inject_2. apply init_meminj_preserves_globals. eauto. eauto.
intros (b' & A & B). rewrite A. apply inj_value_inject.
- econstructor; eauto. symmetry; apply Int.add_zero.
+ econstructor; eauto. symmetry; apply Ptrofs.add_zero.
destruct (Genv.find_symbol tge i) as [b'|] eqn:FS'.
exploit symbols_inject_3. apply init_meminj_preserves_globals. eauto.
intros (b & A & B). congruence.
- apply repeat_Undef_inject_self with (n := 4%nat).
+ apply repeat_Undef_inject_self.
+ apply IHil. intros id [ofs IN]. apply H. exists ofs; auto with coqlib.
Qed.
@@ -1177,7 +1177,7 @@ Proof.
exploit init_meminj_invert. eexact H1. intros (A2 & id2 & B2 & C2).
destruct (ident_eq id1 id2). congruence. left; eapply Genv.global_addresses_distinct; eauto.
- exploit init_meminj_invert; eauto. intros (A & id & B & C). subst delta.
- split. omega. generalize (Int.unsigned_range_2 ofs). omega.
+ split. omega. generalize (Ptrofs.unsigned_range_2 ofs). omega.
- exploit init_meminj_invert_strong; eauto. intros (A & id & gd & B & C & D & E & F).
exploit (Genv.init_mem_characterization_gen p); eauto.
exploit (Genv.init_mem_characterization_gen tp); eauto.
diff --git a/backend/ValueAnalysis.v b/backend/ValueAnalysis.v
index a4d34279..c89f8435 100644
--- a/backend/ValueAnalysis.v
+++ b/backend/ValueAnalysis.v
@@ -187,7 +187,7 @@ Definition store_init_data (ab: ablock) (p: Z) (id: init_data) : ablock :=
(if propagate_float_constants tt then FS n else ntop)
| Init_float64 n => ablock_store Mfloat64 ab p
(if propagate_float_constants tt then F n else ntop)
- | Init_addrof symb ofs => ablock_store Mint32 ab p (Ptr (Gl symb ofs))
+ | Init_addrof symb ofs => ablock_store Mptr ab p (Ptr (Gl symb ofs))
| Init_space n => ab
end.
@@ -329,13 +329,13 @@ Lemma abuiltin_arg_sound:
genv_match bc ge ->
bc sp = BCstack ->
forall a v,
- eval_builtin_arg ge (fun r => rs#r) (Vptr sp Int.zero) m a v ->
+ eval_builtin_arg ge (fun r => rs#r) (Vptr sp Ptrofs.zero) m a v ->
vmatch bc v (abuiltin_arg ae am rm a).
Proof.
intros until am; intros EM RM MM GM SP.
induction 1; simpl; eauto with va.
-- eapply loadv_sound; eauto. simpl. rewrite Int.add_zero_l. auto with va.
-- simpl. rewrite Int.add_zero_l. auto with va.
+- eapply loadv_sound; eauto. simpl. rewrite Ptrofs.add_zero_l. auto with va.
+- simpl. rewrite Ptrofs.add_zero_l. auto with va.
- eapply loadv_sound; eauto. apply symbol_address_sound; auto.
- apply symbol_address_sound; auto.
Qed.
@@ -348,7 +348,7 @@ Lemma abuiltin_args_sound:
genv_match bc ge ->
bc sp = BCstack ->
forall al vl,
- eval_builtin_args ge (fun r => rs#r) (Vptr sp Int.zero) m al vl ->
+ eval_builtin_args ge (fun r => rs#r) (Vptr sp Ptrofs.zero) m al vl ->
list_forall2 (vmatch bc) vl (map (abuiltin_arg ae am rm) al).
Proof.
intros until am; intros EM RM MM GM SP.
@@ -1050,7 +1050,7 @@ Inductive sound_stack: block_classification -> list stackframe -> mem -> block -
(GE: genv_match bc' ge)
(AN: VA.ge (analyze rm f)!!pc (VA.State (AE.set res Vtop ae) mafter_public_call))
(EM: ematch bc' e ae),
- sound_stack bc (Stackframe res f (Vptr sp Int.zero) pc e :: stk) m bound
+ sound_stack bc (Stackframe res f (Vptr sp Ptrofs.zero) pc e :: stk) m bound
| sound_stack_private_call:
forall (bc: block_classification) res f sp pc e stk m bound bc' bound' ae am
(STK: sound_stack bc' stk m sp)
@@ -1063,7 +1063,7 @@ Inductive sound_stack: block_classification -> list stackframe -> mem -> block -
(AN: VA.ge (analyze rm f)!!pc (VA.State (AE.set res (Ifptr Nonstack) ae) (mafter_private_call am)))
(EM: ematch bc' e ae)
(CONTENTS: bmatch bc' m sp am.(am_stack)),
- sound_stack bc (Stackframe res f (Vptr sp Int.zero) pc e :: stk) m bound.
+ sound_stack bc (Stackframe res f (Vptr sp Ptrofs.zero) pc e :: stk) m bound.
Inductive sound_state_base: state -> Prop :=
| sound_regular_state:
@@ -1075,7 +1075,7 @@ Inductive sound_state_base: state -> Prop :=
(MM: mmatch bc m am)
(GE: genv_match bc ge)
(SP: bc sp = BCstack),
- sound_state_base (State s f (Vptr sp Int.zero) pc e m)
+ sound_state_base (State s f (Vptr sp Ptrofs.zero) pc e m)
| sound_call_state:
forall s fd args m bc
(STK: sound_stack bc s m (Mem.nextblock m))
@@ -1143,7 +1143,7 @@ Qed.
Lemma sound_stack_storebytes:
forall m b ofs bytes m' bc aaddr stk bound,
- Mem.storebytes m b (Int.unsigned ofs) bytes = Some m' ->
+ Mem.storebytes m b (Ptrofs.unsigned ofs) bytes = Some m' ->
vmatch bc (Vptr b ofs) aaddr ->
sound_stack bc stk m bound ->
sound_stack bc stk m' bound.
@@ -1209,7 +1209,7 @@ Lemma sound_succ_state:
genv_match bc ge ->
bc sp = BCstack ->
sound_stack bc s m' sp ->
- sound_state_base (State s f (Vptr sp Int.zero) pc' e' m').
+ sound_state_base (State s f (Vptr sp Ptrofs.zero) pc' e' m').
Proof.
intros. exploit analyze_succ; eauto. intros (ae'' & am'' & AN & EM & MM).
econstructor; eauto.
@@ -1296,7 +1296,7 @@ Proof.
assert (DEFAULT:
transfer f rm pc ae am = transfer_builtin_default ae am rm args res ->
sound_state_base
- (State s f (Vptr sp0 Int.zero) pc' (regmap_setres res vres rs) m')).
+ (State s f (Vptr sp0 Ptrofs.zero) pc' (regmap_setres res vres rs) m')).
{ unfold transfer_builtin_default, analyze_call; intros TR'.
set (aargs := map (abuiltin_arg ae am rm) args) in *.
assert (ARGS: list_forall2 (vmatch bc) vargs aargs) by (eapply abuiltin_args_sound; eauto).
@@ -1603,9 +1603,13 @@ Lemma store_init_data_sound:
bmatch bc m' b (store_init_data ab p id).
Proof.
intros. destruct id; try (eapply ablock_store_sound; eauto; constructor).
+- (* float32 *)
simpl. destruct (propagate_float_constants tt); eapply ablock_store_sound; eauto; constructor.
+- (* float64 *)
simpl. destruct (propagate_float_constants tt); eapply ablock_store_sound; eauto; constructor.
+- (* space *)
simpl in H. inv H. auto.
+- (* addrof *)
simpl in H. destruct (Genv.find_symbol ge i) as [b'|] eqn:FS; try discriminate.
eapply ablock_store_sound; eauto. constructor. constructor. apply GMATCH; auto.
Qed.
@@ -1882,7 +1886,7 @@ Definition avalue (a: VA.t) (r: reg) : aval :=
Lemma avalue_sound:
forall cunit prog s f sp pc e m r,
- sound_state prog (State s f (Vptr sp Int.zero) pc e m) ->
+ sound_state prog (State s f (Vptr sp Ptrofs.zero) pc e m) ->
linkorder cunit prog ->
exists bc,
vmatch bc e#r (avalue (analyze (romem_for cunit) f)!!pc r)
@@ -1900,7 +1904,7 @@ Definition aaddr (a: VA.t) (r: reg) : aptr :=
Lemma aaddr_sound:
forall cunit prog s f sp pc e m r b ofs,
- sound_state prog (State s f (Vptr sp Int.zero) pc e m) ->
+ sound_state prog (State s f (Vptr sp Ptrofs.zero) pc e m) ->
linkorder cunit prog ->
e#r = Vptr b ofs ->
exists bc,
@@ -1920,9 +1924,9 @@ Definition aaddressing (a: VA.t) (addr: addressing) (args: list reg) : aptr :=
Lemma aaddressing_sound:
forall cunit prog s f sp pc e m addr args b ofs,
- sound_state prog (State s f (Vptr sp Int.zero) pc e m) ->
+ sound_state prog (State s f (Vptr sp Ptrofs.zero) pc e m) ->
linkorder cunit prog ->
- eval_addressing (Genv.globalenv prog) (Vptr sp Int.zero) addr e##args = Some (Vptr b ofs) ->
+ eval_addressing (Genv.globalenv prog) (Vptr sp Ptrofs.zero) addr e##args = Some (Vptr b ofs) ->
exists bc,
pmatch bc b ofs (aaddressing (analyze (romem_for cunit) f)!!pc addr args)
/\ genv_match bc (Genv.globalenv prog)
@@ -1955,7 +1959,7 @@ Lemma aaddr_arg_sound_1:
mmatch bc m am ->
genv_match bc ge ->
bc sp = BCstack ->
- eval_builtin_arg ge (fun r : positive => rs # r) (Vptr sp Int.zero) m a (Vptr b ofs) ->
+ eval_builtin_arg ge (fun r : positive => rs # r) (Vptr sp Ptrofs.zero) m a (Vptr b ofs) ->
pmatch bc b ofs (aaddr_arg (VA.State ae am) a).
Proof.
intros.
@@ -1966,9 +1970,9 @@ Qed.
Lemma aaddr_arg_sound:
forall cunit prog s f sp pc e m a b ofs,
- sound_state prog (State s f (Vptr sp Int.zero) pc e m) ->
+ sound_state prog (State s f (Vptr sp Ptrofs.zero) pc e m) ->
linkorder cunit prog ->
- eval_builtin_arg (Genv.globalenv prog) (fun r => e#r) (Vptr sp Int.zero) m a (Vptr b ofs) ->
+ eval_builtin_arg (Genv.globalenv prog) (fun r => e#r) (Vptr sp Ptrofs.zero) m a (Vptr b ofs) ->
exists bc,
pmatch bc b ofs (aaddr_arg (analyze (romem_for cunit) f)!!pc a)
/\ genv_match bc (Genv.globalenv prog)
diff --git a/backend/ValueDomain.v b/backend/ValueDomain.v
index bc09c3dc..6b314904 100644
--- a/backend/ValueDomain.v
+++ b/backend/ValueDomain.v
@@ -116,21 +116,21 @@ Qed.
(** * Abstracting pointers *)
Inductive aptr : Type :=
- | Pbot (**r bottom (empty set of pointers) *)
- | Gl (id: ident) (ofs: int) (**r pointer into the block for global variable [id] at offset [ofs] *)
- | Glo (id: ident) (**r pointer anywhere into the block for global [id] *)
- | Glob (**r pointer into any global variable *)
- | Stk (ofs: int) (**r pointer into the current stack frame at offset [ofs] *)
- | Stack (**r pointer anywhere into the current stack frame *)
- | Nonstack (**r pointer anywhere but into the current stack frame *)
- | Ptop. (**r any valid pointer *)
+ | Pbot (**r bottom (empty set of pointers) *)
+ | Gl (id: ident) (ofs: ptrofs) (**r pointer into the block for global variable [id] at offset [ofs] *)
+ | Glo (id: ident) (**r pointer anywhere into the block for global [id] *)
+ | Glob (**r pointer into any global variable *)
+ | Stk (ofs: ptrofs) (**r pointer into the current stack frame at offset [ofs] *)
+ | Stack (**r pointer anywhere into the current stack frame *)
+ | Nonstack (**r pointer anywhere but into the current stack frame *)
+ | Ptop. (**r any valid pointer *)
Definition eq_aptr: forall (p1 p2: aptr), {p1=p2} + {p1<>p2}.
Proof.
- intros. generalize ident_eq, Int.eq_dec; intros. decide equality.
+ intros. generalize ident_eq, Ptrofs.eq_dec; intros. decide equality.
Defined.
-Inductive pmatch (b: block) (ofs: int): aptr -> Prop :=
+Inductive pmatch (b: block) (ofs: ptrofs): aptr -> Prop :=
| pmatch_gl: forall id,
bc b = BCglob id ->
pmatch b ofs (Gl id ofs)
@@ -191,7 +191,7 @@ Definition plub (p q: aptr) : aptr :=
| Pbot, _ => q
| _, Pbot => p
| Gl id1 ofs1, Gl id2 ofs2 =>
- if ident_eq id1 id2 then if Int.eq_dec ofs1 ofs2 then p else Glo id1 else Glob
+ if ident_eq id1 id2 then if Ptrofs.eq_dec ofs1 ofs2 then p else Glo id1 else Glob
| Gl id1 ofs1, Glo id2 =>
if ident_eq id1 id2 then q else Glob
| Glo id1, Gl id2 ofs2 =>
@@ -205,7 +205,7 @@ Definition plub (p q: aptr) : aptr :=
| Nonstack, (Gl _ _ | Glo _ | Glob) =>
Nonstack
| Stk ofs1, Stk ofs2 =>
- if Int.eq_dec ofs1 ofs2 then p else Stack
+ if Ptrofs.eq_dec ofs1 ofs2 then p else Stack
| (Stk _ | Stack), Stack =>
Stack
| Stack, Stk _ =>
@@ -219,7 +219,7 @@ Proof.
intros; unfold plub; destruct p; destruct q; auto.
destruct (ident_eq id id0). subst id0.
rewrite dec_eq_true.
- destruct (Int.eq_dec ofs ofs0). subst ofs0. rewrite dec_eq_true. auto.
+ destruct (Ptrofs.eq_dec ofs ofs0). subst ofs0. rewrite dec_eq_true. auto.
rewrite dec_eq_false by auto. auto.
rewrite dec_eq_false by auto. auto.
destruct (ident_eq id id0). subst id0.
@@ -231,7 +231,7 @@ Proof.
destruct (ident_eq id id0). subst id0.
rewrite dec_eq_true; auto.
rewrite dec_eq_false; auto.
- destruct (Int.eq_dec ofs ofs0). subst ofs0. rewrite dec_eq_true; auto.
+ destruct (Ptrofs.eq_dec ofs ofs0). subst ofs0. rewrite dec_eq_true; auto.
rewrite dec_eq_false; auto.
Qed.
@@ -240,12 +240,12 @@ Lemma pge_lub_l:
Proof.
unfold plub; destruct p, q; auto with va.
- destruct (ident_eq id id0).
- destruct (Int.eq_dec ofs ofs0); subst; constructor.
+ destruct (Ptrofs.eq_dec ofs ofs0); subst; constructor.
constructor.
- destruct (ident_eq id id0); subst; constructor.
- destruct (ident_eq id id0); subst; constructor.
- destruct (ident_eq id id0); subst; constructor.
-- destruct (Int.eq_dec ofs ofs0); subst; constructor.
+- destruct (Ptrofs.eq_dec ofs ofs0); subst; constructor.
Qed.
Lemma pge_lub_r:
@@ -274,27 +274,27 @@ Proof.
- unfold plub; destruct q; repeat rewrite dec_eq_true; constructor.
- rewrite dec_eq_true; constructor.
- rewrite dec_eq_true; constructor.
-- rewrite dec_eq_true. destruct (Int.eq_dec ofs ofs0); constructor.
-- destruct (ident_eq id id0). destruct (Int.eq_dec ofs ofs0); constructor. constructor.
+- rewrite dec_eq_true. destruct (Ptrofs.eq_dec ofs ofs0); constructor.
+- destruct (ident_eq id id0). destruct (Ptrofs.eq_dec ofs ofs0); constructor. constructor.
- destruct (ident_eq id id0); constructor.
- destruct (ident_eq id id0); constructor.
- destruct (ident_eq id id0); constructor.
-- destruct (ident_eq id id0). destruct (Int.eq_dec ofs ofs0); constructor. constructor.
+- destruct (ident_eq id id0). destruct (Ptrofs.eq_dec ofs ofs0); constructor. constructor.
- destruct (ident_eq id id0); constructor.
- destruct (ident_eq id id0); constructor.
- destruct (ident_eq id id0); constructor.
-- destruct (Int.eq_dec ofs ofs0); constructor.
+- destruct (Ptrofs.eq_dec ofs ofs0); constructor.
Qed.
Definition pincl (p q: aptr) : bool :=
match p, q with
| Pbot, _ => true
- | Gl id1 ofs1, Gl id2 ofs2 => peq id1 id2 && Int.eq_dec ofs1 ofs2
+ | Gl id1 ofs1, Gl id2 ofs2 => peq id1 id2 && Ptrofs.eq_dec ofs1 ofs2
| Gl id1 ofs1, Glo id2 => peq id1 id2
| Glo id1, Glo id2 => peq id1 id2
| (Gl _ _ | Glo _ | Glob), Glob => true
| (Gl _ _ | Glo _ | Glob | Nonstack), Nonstack => true
- | Stk ofs1, Stk ofs2 => Int.eq_dec ofs1 ofs2
+ | Stk ofs1, Stk ofs2 => Ptrofs.eq_dec ofs1 ofs2
| Stk ofs1, Stack => true
| Stack, Stack => true
| _, Ptop => true
@@ -322,32 +322,32 @@ Proof.
intros. eapply pmatch_ge; eauto. apply pincl_ge; auto.
Qed.
-Definition padd (p: aptr) (n: int) : aptr :=
+Definition padd (p: aptr) (n: ptrofs) : aptr :=
match p with
- | Gl id ofs => Gl id (Int.add ofs n)
- | Stk ofs => Stk (Int.add ofs n)
+ | Gl id ofs => Gl id (Ptrofs.add ofs n)
+ | Stk ofs => Stk (Ptrofs.add ofs n)
| _ => p
end.
Lemma padd_sound:
forall b ofs p delta,
pmatch b ofs p ->
- pmatch b (Int.add ofs delta) (padd p delta).
+ pmatch b (Ptrofs.add ofs delta) (padd p delta).
Proof.
intros. inv H; simpl padd; eauto with va.
Qed.
-Definition psub (p: aptr) (n: int) : aptr :=
+Definition psub (p: aptr) (n: ptrofs) : aptr :=
match p with
- | Gl id ofs => Gl id (Int.sub ofs n)
- | Stk ofs => Stk (Int.sub ofs n)
+ | Gl id ofs => Gl id (Ptrofs.sub ofs n)
+ | Stk ofs => Stk (Ptrofs.sub ofs n)
| _ => p
end.
Lemma psub_sound:
forall b ofs p delta,
pmatch b ofs p ->
- pmatch b (Int.sub ofs delta) (psub p delta).
+ pmatch b (Ptrofs.sub ofs delta) (psub p delta).
Proof.
intros. inv H; simpl psub; eauto with va.
Qed.
@@ -367,29 +367,6 @@ Proof.
intros. inv H; simpl poffset; eauto with va.
Qed.
-Definition psub2 (p q: aptr) : option int :=
- match p, q with
- | Gl id1 ofs1, Gl id2 ofs2 =>
- if peq id1 id2 then Some (Int.sub ofs1 ofs2) else None
- | Stk ofs1, Stk ofs2 =>
- Some (Int.sub ofs1 ofs2)
- | _, _ =>
- None
- end.
-
-Lemma psub2_sound:
- forall b1 ofs1 p1 b2 ofs2 p2 delta,
- psub2 p1 p2 = Some delta ->
- pmatch b1 ofs1 p1 ->
- pmatch b2 ofs2 p2 ->
- b1 = b2 /\ delta = Int.sub ofs1 ofs2.
-Proof.
- intros. destruct p1; try discriminate; destruct p2; try discriminate; simpl in H.
-- destruct (peq id id0); inv H. inv H0; inv H1.
- split. eapply bc_glob; eauto. reflexivity.
-- inv H; inv H0; inv H1. split. eapply bc_stack; eauto. reflexivity.
-Qed.
-
Definition cmp_different_blocks (c: comparison) : abool :=
match c with
| Ceq => Maybe false
@@ -413,7 +390,7 @@ Definition pcmp (c: comparison) (p1 p2: aptr) : abool :=
match p1, p2 with
| Pbot, _ | _, Pbot => Bnone
| Gl id1 ofs1, Gl id2 ofs2 =>
- if peq id1 id2 then Maybe (Int.cmpu c ofs1 ofs2)
+ if peq id1 id2 then Maybe (Ptrofs.cmpu c ofs1 ofs2)
else cmp_different_blocks c
| Gl id1 ofs1, Glo id2 =>
if peq id1 id2 then Btop else cmp_different_blocks c
@@ -421,7 +398,7 @@ Definition pcmp (c: comparison) (p1 p2: aptr) : abool :=
if peq id1 id2 then Btop else cmp_different_blocks c
| Glo id1, Glo id2 =>
if peq id1 id2 then Btop else cmp_different_blocks c
- | Stk ofs1, Stk ofs2 => Maybe (Int.cmpu c ofs1 ofs2)
+ | Stk ofs1, Stk ofs2 => Maybe (Ptrofs.cmpu c ofs1 ofs2)
| (Gl _ _ | Glo _ | Glob | Nonstack), (Stk _ | Stack) => cmp_different_blocks c
| (Stk _ | Stack), (Gl _ _ | Glo _ | Glob | Nonstack) => cmp_different_blocks c
| _, _ => Btop
@@ -438,17 +415,59 @@ Proof.
(cmp_different_blocks c)).
{
intros. simpl. rewrite dec_eq_false by assumption.
- destruct (valid b1 (Int.unsigned ofs1) && valid b2 (Int.unsigned ofs2)); simpl.
+ destruct Archi.ptr64.
+ apply cmp_different_blocks_none.
+ destruct (valid b1 (Ptrofs.unsigned ofs1) && valid b2 (Ptrofs.unsigned ofs2)); simpl.
apply cmp_different_blocks_sound.
apply cmp_different_blocks_none.
}
assert (SAME: b1 = b2 ->
cmatch (Val.cmpu_bool valid c (Vptr b1 ofs1) (Vptr b2 ofs2))
- (Maybe (Int.cmpu c ofs1 ofs2))).
+ (Maybe (Ptrofs.cmpu c ofs1 ofs2))).
{
- intros. subst b2. simpl. rewrite dec_eq_true.
- destruct ((valid b1 (Int.unsigned ofs1) || valid b1 (Int.unsigned ofs1 - 1)) &&
- (valid b1 (Int.unsigned ofs2) || valid b1 (Int.unsigned ofs2 - 1))); simpl.
+ intros. subst b2. simpl. destruct Archi.ptr64.
+ constructor.
+ rewrite dec_eq_true.
+ destruct ((valid b1 (Ptrofs.unsigned ofs1) || valid b1 (Ptrofs.unsigned ofs1 - 1)) &&
+ (valid b1 (Ptrofs.unsigned ofs2) || valid b1 (Ptrofs.unsigned ofs2 - 1))); simpl.
+ constructor.
+ constructor.
+ }
+ unfold pcmp; inv H; inv H0; (apply cmatch_top || (apply DIFF; congruence) || idtac).
+ - destruct (peq id id0). subst id0. apply SAME. eapply bc_glob; eauto.
+ auto with va.
+ - destruct (peq id id0); auto with va.
+ - destruct (peq id id0); auto with va.
+ - destruct (peq id id0); auto with va.
+ - apply SAME. eapply bc_stack; eauto.
+Qed.
+
+Lemma pcmp_sound_64:
+ forall valid c b1 ofs1 p1 b2 ofs2 p2,
+ pmatch b1 ofs1 p1 -> pmatch b2 ofs2 p2 ->
+ cmatch (Val.cmplu_bool valid c (Vptr b1 ofs1) (Vptr b2 ofs2)) (pcmp c p1 p2).
+Proof.
+ intros.
+ assert (DIFF: b1 <> b2 ->
+ cmatch (Val.cmplu_bool valid c (Vptr b1 ofs1) (Vptr b2 ofs2))
+ (cmp_different_blocks c)).
+ {
+ intros. simpl. rewrite dec_eq_false by assumption.
+ destruct Archi.ptr64; simpl.
+ destruct (valid b1 (Ptrofs.unsigned ofs1) && valid b2 (Ptrofs.unsigned ofs2)); simpl.
+ apply cmp_different_blocks_sound.
+ apply cmp_different_blocks_none.
+ apply cmp_different_blocks_none.
+ }
+ assert (SAME: b1 = b2 ->
+ cmatch (Val.cmplu_bool valid c (Vptr b1 ofs1) (Vptr b2 ofs2))
+ (Maybe (Ptrofs.cmpu c ofs1 ofs2))).
+ {
+ intros. subst b2. simpl. destruct Archi.ptr64.
+ rewrite dec_eq_true.
+ destruct ((valid b1 (Ptrofs.unsigned ofs1) || valid b1 (Ptrofs.unsigned ofs1 - 1)) &&
+ (valid b1 (Ptrofs.unsigned ofs2) || valid b1 (Ptrofs.unsigned ofs2 - 1))); simpl.
+ constructor.
constructor.
constructor.
}
@@ -475,15 +494,15 @@ Definition pdisjoint (p1: aptr) (sz1: Z) (p2: aptr) (sz2: Z) : bool :=
| _, Pbot => true
| Gl id1 ofs1, Gl id2 ofs2 =>
if peq id1 id2
- then zle (Int.unsigned ofs1 + sz1) (Int.unsigned ofs2)
- || zle (Int.unsigned ofs2 + sz2) (Int.unsigned ofs1)
+ then zle (Ptrofs.unsigned ofs1 + sz1) (Ptrofs.unsigned ofs2)
+ || zle (Ptrofs.unsigned ofs2 + sz2) (Ptrofs.unsigned ofs1)
else true
| Gl id1 ofs1, Glo id2 => negb(peq id1 id2)
| Glo id1, Gl id2 ofs2 => negb(peq id1 id2)
| Glo id1, Glo id2 => negb(peq id1 id2)
| Stk ofs1, Stk ofs2 =>
- zle (Int.unsigned ofs1 + sz1) (Int.unsigned ofs2)
- || zle (Int.unsigned ofs2 + sz2) (Int.unsigned ofs1)
+ zle (Ptrofs.unsigned ofs1 + sz1) (Ptrofs.unsigned ofs2)
+ || zle (Ptrofs.unsigned ofs2 + sz2) (Ptrofs.unsigned ofs1)
| (Gl _ _ | Glo _ | Glob | Nonstack), (Stk _ | Stack) => true
| (Stk _ | Stack), (Gl _ _ | Glo _ | Glob | Nonstack) => true
| _, _ => false
@@ -493,7 +512,7 @@ Lemma pdisjoint_sound:
forall sz1 b1 ofs1 p1 sz2 b2 ofs2 p2,
pdisjoint p1 sz1 p2 sz2 = true ->
pmatch b1 ofs1 p1 -> pmatch b2 ofs2 p2 ->
- b1 <> b2 \/ Int.unsigned ofs1 + sz1 <= Int.unsigned ofs2 \/ Int.unsigned ofs2 + sz2 <= Int.unsigned ofs1.
+ b1 <> b2 \/ Ptrofs.unsigned ofs1 + sz1 <= Ptrofs.unsigned ofs2 \/ Ptrofs.unsigned ofs2 + sz2 <= Ptrofs.unsigned ofs1.
Proof.
intros. inv H0; inv H1; simpl in H; try discriminate; try (left; congruence).
- destruct (peq id id0). subst id0. destruct (orb_true_elim _ _ H); InvBooleans; auto.
@@ -1154,6 +1173,28 @@ Proof.
intros. unfold binop_int; inv H; auto with va; inv H0; auto with va.
Qed.
+Definition unop_long (sem: int64 -> int64) (x: aval) :=
+ match x with L n => L (sem n) | _ => ntop1 x end.
+
+Lemma unop_long_sound:
+ forall sem v x,
+ vmatch v x ->
+ vmatch (match v with Vlong i => Vlong(sem i) | _ => Vundef end) (unop_long sem x).
+Proof.
+ intros. unfold unop_long; inv H; auto with va.
+Qed.
+
+Definition binop_long (sem: int64 -> int64 -> int64) (x y: aval) :=
+ match x, y with L n, L m => L (sem n m) | _, _ => ntop2 x y end.
+
+Lemma binop_long_sound:
+ forall sem v x w y,
+ vmatch v x -> vmatch w y ->
+ vmatch (match v, w with Vlong i, Vlong j => Vlong(sem i j) | _, _ => Vundef end) (binop_long sem x y).
+Proof.
+ intros. unfold binop_long; inv H; auto with va; inv H0; auto with va.
+Qed.
+
Definition unop_float (sem: float -> float) (x: aval) :=
match x with F n => F (sem n) | _ => ntop1 x end.
@@ -1502,9 +1543,9 @@ Proof (unop_int_sound Int.neg).
Definition add (x y: aval) :=
match x, y with
| I i, I j => I (Int.add i j)
- | Ptr p, I i | I i, Ptr p => Ptr (padd p i)
+ | Ptr p, I i | I i, Ptr p => Ptr (if Archi.ptr64 then poffset p else padd p (Ptrofs.of_int i))
| Ptr p, _ | _, Ptr p => Ptr (poffset p)
- | Ifptr p, I i | I i, Ifptr p => Ifptr (padd p i)
+ | Ifptr p, I i | I i, Ifptr p => Ifptr (if Archi.ptr64 then poffset p else padd p (Ptrofs.of_int i))
| Ifptr p, Ifptr q => Ifptr (plub (poffset p) (poffset q))
| Ifptr p, _ | _, Ifptr p => Ifptr (poffset p)
| _, _ => ntop2 x y
@@ -1513,7 +1554,9 @@ Definition add (x y: aval) :=
Lemma add_sound:
forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.add v w) (add x y).
Proof.
- intros. unfold Val.add, add; inv H; inv H0; constructor;
+ intros. unfold Val.add, add. destruct Archi.ptr64.
+- inv H; inv H0; constructor.
+- inv H; inv H0; constructor;
((apply padd_sound; assumption) || (eapply poffset_sound; eassumption) || idtac).
apply pmatch_lub_r. eapply poffset_sound; eauto.
apply pmatch_lub_l. eapply poffset_sound; eauto.
@@ -1522,13 +1565,9 @@ Qed.
Definition sub (v w: aval) :=
match v, w with
| I i1, I i2 => I (Int.sub i1 i2)
- | Ptr p, I i => Ptr (psub p i)
-(* problem with undefs *)
-(*
- | Ptr p1, Ptr p2 => match psub2 p1 p2 with Some n => I n | _ => itop end
-*)
+ | Ptr p, I i => if Archi.ptr64 then Ifptr (poffset p) else Ptr (psub p (Ptrofs.of_int i))
| Ptr p, _ => Ifptr (poffset p)
- | Ifptr p, I i => Ifptr (psub p i)
+ | Ifptr p, I i => if Archi.ptr64 then Ifptr (plub (poffset p) (provenance w)) else Ifptr (psub p (Ptrofs.of_int i))
| Ifptr p, _ => Ifptr (plub (poffset p) (provenance w))
| _, _ => ntop2 v w
end.
@@ -1536,9 +1575,9 @@ Definition sub (v w: aval) :=
Lemma sub_sound:
forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.sub v w) (sub x y).
Proof.
- intros. inv H; subst; inv H0; subst; simpl;
- try (destruct (eq_block b b0));
- eauto using psub_sound, poffset_sound, pmatch_lub_l with va.
+ intros. unfold Val.sub, sub. destruct Archi.ptr64.
+- inv H; inv H0; eauto with va.
+- inv H; inv H0; try (destruct (eq_block b b0)); eauto using psub_sound, poffset_sound, pmatch_lub_l with va.
Qed.
Definition mul := binop_int Int.mul.
@@ -1659,6 +1698,246 @@ Proof.
rewrite LTU; auto with va.
Qed.
+(** 64-bit integer operations *)
+
+Definition shift_long (sem: int64 -> int -> int64) (v w: aval) :=
+ match w with
+ | I amount =>
+ if Int.ltu amount Int64.iwordsize' then
+ match v with
+ | L i => L (sem i amount)
+ | _ => ntop1 v
+ end
+ else ntop1 v
+ | _ => ntop1 v
+ end.
+
+Lemma shift_long_sound:
+ forall sem v w x y,
+ vmatch v x -> vmatch w y ->
+ vmatch (match v, w with
+ | Vlong i, Vint j => if Int.ltu j Int64.iwordsize'
+ then Vlong (sem i j) else Vundef
+ | _, _ => Vundef end)
+ (shift_long sem x y).
+Proof.
+ intros.
+ assert (DEFAULT:
+ vmatch (match v, w with
+ | Vlong i, Vint j => if Int.ltu j Int64.iwordsize'
+ then Vlong (sem i j) else Vundef
+ | _, _ => Vundef end)
+ (ntop1 x)).
+ { destruct v; try constructor; destruct w; try constructor.
+ destruct (Int.ltu i0 Int64.iwordsize'); constructor. }
+ unfold shift_long. destruct y; auto.
+ destruct (Int.ltu n Int64.iwordsize') eqn:LT; auto.
+ destruct x; auto.
+ inv H; inv H0. rewrite LT. constructor.
+Qed.
+
+Definition shll := shift_long Int64.shl'.
+
+Lemma shll_sound:
+ forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.shll v w) (shll x y).
+Proof (shift_long_sound Int64.shl').
+
+Definition shrl := shift_long Int64.shr'.
+
+Lemma shrl_sound:
+ forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.shrl v w) (shrl x y).
+Proof (shift_long_sound Int64.shr').
+
+Definition shrlu := shift_long Int64.shru'.
+
+Lemma shrlu_sound:
+ forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.shrlu v w) (shrlu x y).
+Proof (shift_long_sound Int64.shru').
+
+Definition andl := binop_long Int64.and.
+
+Lemma andl_sound:
+ forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.andl v w) (andl x y).
+Proof (binop_long_sound Int64.and).
+
+Definition orl := binop_long Int64.or.
+
+Lemma orl_sound:
+ forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.orl v w) (orl x y).
+Proof (binop_long_sound Int64.or).
+
+Definition xorl := binop_long Int64.xor.
+
+Lemma xorl_sound:
+ forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.xorl v w) (xorl x y).
+Proof (binop_long_sound Int64.xor).
+
+Definition notl := unop_long Int64.not.
+
+Lemma notl_sound:
+ forall v x, vmatch v x -> vmatch (Val.notl v) (notl x).
+Proof (unop_long_sound Int64.not).
+
+Definition rotate_long (sem: int64 -> int64 -> int64) (v w: aval) :=
+ match v, w with
+ | L i, I amount => L (sem i (Int64.repr (Int.unsigned amount)))
+ | _, _ => ntop1 v
+ end.
+
+Lemma rotate_long_sound:
+ forall sem v w x y,
+ vmatch v x -> vmatch w y ->
+ vmatch (match v, w with
+ | Vlong i, Vint j => Vlong (sem i (Int64.repr (Int.unsigned j)))
+ | _, _ => Vundef end)
+ (rotate_long sem x y).
+Proof.
+ intros.
+ assert (DEFAULT:
+ vmatch (match v, w with
+ | Vlong i, Vint j => Vlong (sem i (Int64.repr (Int.unsigned j)))
+ | _, _ => Vundef end)
+ (ntop1 x)).
+ { destruct v; try constructor. destruct w; constructor. }
+ unfold rotate_long. destruct x; auto. destruct y; auto. inv H; inv H0. constructor.
+Qed.
+
+Definition roll := rotate_long Int64.rol.
+
+Lemma roll_sound:
+ forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.roll v w) (roll x y).
+Proof (rotate_long_sound Int64.rol).
+
+Definition rorl := rotate_long Int64.ror.
+
+Lemma rorl_sound:
+ forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.rorl v w) (rorl x y).
+Proof (rotate_long_sound Int64.ror).
+
+Definition negl := unop_long Int64.neg.
+
+Lemma negl_sound:
+ forall v x, vmatch v x -> vmatch (Val.negl v) (negl x).
+Proof (unop_long_sound Int64.neg).
+
+Definition addl (x y: aval) :=
+ match x, y with
+ | L i, L j => L (Int64.add i j)
+ | Ptr p, L i | L i, Ptr p => Ptr (if Archi.ptr64 then padd p (Ptrofs.of_int64 i) else poffset p)
+ | Ptr p, _ | _, Ptr p => Ptr (poffset p)
+ | Ifptr p, L i | L i, Ifptr p => Ifptr (if Archi.ptr64 then padd p (Ptrofs.of_int64 i) else poffset p)
+ | Ifptr p, Ifptr q => Ifptr (plub (poffset p) (poffset q))
+ | Ifptr p, _ | _, Ifptr p => Ifptr (poffset p)
+ | _, _ => ntop2 x y
+ end.
+
+Lemma addl_sound:
+ forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.addl v w) (addl x y).
+Proof.
+ intros. unfold Val.addl, addl. destruct Archi.ptr64.
+- inv H; inv H0; constructor;
+ ((apply padd_sound; assumption) || (eapply poffset_sound; eassumption) || idtac).
+ apply pmatch_lub_r. eapply poffset_sound; eauto.
+ apply pmatch_lub_l. eapply poffset_sound; eauto.
+- inv H; inv H0; constructor.
+Qed.
+
+Definition subl (v w: aval) :=
+ match v, w with
+ | L i1, L i2 => L (Int64.sub i1 i2)
+ | Ptr p, L i => if Archi.ptr64 then Ptr (psub p (Ptrofs.of_int64 i)) else Ifptr (poffset p)
+ | Ptr p, _ => Ifptr (poffset p)
+ | Ifptr p, L i => if Archi.ptr64 then Ifptr (psub p (Ptrofs.of_int64 i)) else Ifptr (plub (poffset p) (provenance w))
+ | Ifptr p, _ => Ifptr (plub (poffset p) (provenance w))
+ | _, _ => ntop2 v w
+ end.
+
+Lemma subl_sound:
+ forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.subl v w) (subl x y).
+Proof.
+ intros. unfold Val.subl, subl. destruct Archi.ptr64.
+- inv H; inv H0; try (destruct (eq_block b b0)); eauto using psub_sound, poffset_sound, pmatch_lub_l with va.
+- inv H; inv H0; eauto with va.
+Qed.
+
+Definition mull := binop_long Int64.mul.
+
+Lemma mull_sound:
+ forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.mull v w) (mull x y).
+Proof (binop_long_sound Int64.mul).
+
+Definition divls (v w: aval) :=
+ match w, v with
+ | L i2, L i1 =>
+ if Int64.eq i2 Int64.zero
+ || Int64.eq i1 (Int64.repr Int64.min_signed) && Int64.eq i2 Int64.mone
+ then if va_strict tt then Vbot else ntop
+ else L (Int64.divs i1 i2)
+ | _, _ => ntop2 v w
+ end.
+
+Lemma divls_sound:
+ forall v w u x y, vmatch v x -> vmatch w y -> Val.divls v w = Some u -> vmatch u (divls x y).
+Proof.
+ intros. destruct v; destruct w; try discriminate; simpl in H1.
+ destruct (Int64.eq i0 Int64.zero
+ || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone) eqn:E; inv H1.
+ inv H; inv H0; auto with va. simpl. rewrite E. constructor.
+Qed.
+
+Definition divlu (v w: aval) :=
+ match w, v with
+ | L i2, L i1 =>
+ if Int64.eq i2 Int64.zero
+ then if va_strict tt then Vbot else ntop
+ else L (Int64.divu i1 i2)
+ | _, _ => ntop2 v w
+ end.
+
+Lemma divlu_sound:
+ forall v w u x y, vmatch v x -> vmatch w y -> Val.divlu v w = Some u -> vmatch u (divlu x y).
+Proof.
+ intros. destruct v; destruct w; try discriminate; simpl in H1.
+ destruct (Int64.eq i0 Int64.zero) eqn:E; inv H1.
+ inv H; inv H0; auto with va. simpl. rewrite E. constructor.
+Qed.
+
+Definition modls (v w: aval) :=
+ match w, v with
+ | L i2, L i1 =>
+ if Int64.eq i2 Int64.zero
+ || Int64.eq i1 (Int64.repr Int64.min_signed) && Int64.eq i2 Int64.mone
+ then if va_strict tt then Vbot else ntop
+ else L (Int64.mods i1 i2)
+ | _, _ => ntop2 v w
+ end.
+
+Lemma modls_sound:
+ forall v w u x y, vmatch v x -> vmatch w y -> Val.modls v w = Some u -> vmatch u (modls x y).
+Proof.
+ intros. destruct v; destruct w; try discriminate; simpl in H1.
+ destruct (Int64.eq i0 Int64.zero
+ || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone) eqn:E; inv H1.
+ inv H; inv H0; auto with va. simpl. rewrite E. constructor.
+Qed.
+
+Definition modlu (v w: aval) :=
+ match w, v with
+ | L i2, L i1 =>
+ if Int64.eq i2 Int64.zero
+ then if va_strict tt then Vbot else ntop
+ else L (Int64.modu i1 i2)
+ | _, _ => ntop2 v w
+ end.
+
+Lemma modlu_sound:
+ forall v w u x y, vmatch v x -> vmatch w y -> Val.modlu v w = Some u -> vmatch u (modlu x y).
+Proof.
+ intros. destruct v; destruct w; try discriminate; simpl in H1.
+ destruct (Int64.eq i0 Int64.zero) eqn:E; inv H1.
+ inv H; inv H0; auto with va. simpl. rewrite E. constructor.
+Qed.
+
(** Floating-point arithmetic operations *)
Definition negf := unop_float Float.neg.
@@ -1778,6 +2057,30 @@ Proof.
apply Z.min_case; auto with va.
Qed.
+Definition longofint (v: aval) :=
+ match v with
+ | I i => L (Int64.repr (Int.signed i))
+ | _ => ntop1 v
+ end.
+
+Lemma longofint_sound:
+ forall v x, vmatch v x -> vmatch (Val.longofint v) (longofint x).
+Proof.
+ unfold Val.longofint, longofint; intros; inv H; auto with va.
+Qed.
+
+Definition longofintu (v: aval) :=
+ match v with
+ | I i => L (Int64.repr (Int.unsigned i))
+ | _ => ntop1 v
+ end.
+
+Lemma longofintu_sound:
+ forall v x, vmatch v x -> vmatch (Val.longofintu v) (longofintu x).
+Proof.
+ unfold Val.longofintu, longofintu; intros; inv H; auto with va.
+Qed.
+
Definition singleoffloat (v: aval) :=
match v with
| F f => FS (Float.to_single f)
@@ -1932,6 +2235,130 @@ Proof.
inv H; simpl; auto with va.
Qed.
+Definition longoffloat (x: aval) :=
+ match x with
+ | F f =>
+ match Float.to_long f with
+ | Some i => L i
+ | None => if va_strict tt then Vbot else ntop
+ end
+ | _ => ntop1 x
+ end.
+
+Lemma longoffloat_sound:
+ forall v x w, vmatch v x -> Val.longoffloat v = Some w -> vmatch w (longoffloat x).
+Proof.
+ unfold Val.longoffloat; intros. destruct v; try discriminate.
+ destruct (Float.to_long f) as [i|] eqn:E; simpl in H0; inv H0.
+ inv H; simpl; auto with va. rewrite E; constructor.
+Qed.
+
+Definition longuoffloat (x: aval) :=
+ match x with
+ | F f =>
+ match Float.to_longu f with
+ | Some i => L i
+ | None => if va_strict tt then Vbot else ntop
+ end
+ | _ => ntop1 x
+ end.
+
+Lemma longuoffloat_sound:
+ forall v x w, vmatch v x -> Val.longuoffloat v = Some w -> vmatch w (longuoffloat x).
+Proof.
+ unfold Val.longuoffloat; intros. destruct v; try discriminate.
+ destruct (Float.to_longu f) as [i|] eqn:E; simpl in H0; inv H0.
+ inv H; simpl; auto with va. rewrite E; constructor.
+Qed.
+
+Definition floatoflong (x: aval) :=
+ match x with
+ | L i => F(Float.of_long i)
+ | _ => ntop1 x
+ end.
+
+Lemma floatoflong_sound:
+ forall v x w, vmatch v x -> Val.floatoflong v = Some w -> vmatch w (floatoflong x).
+Proof.
+ unfold Val.floatoflong; intros. destruct v; inv H0.
+ inv H; simpl; auto with va.
+Qed.
+
+Definition floatoflongu (x: aval) :=
+ match x with
+ | L i => F(Float.of_longu i)
+ | _ => ntop1 x
+ end.
+
+Lemma floatoflongu_sound:
+ forall v x w, vmatch v x -> Val.floatoflongu v = Some w -> vmatch w (floatoflongu x).
+Proof.
+ unfold Val.floatoflongu; intros. destruct v; inv H0.
+ inv H; simpl; auto with va.
+Qed.
+
+Definition longofsingle (x: aval) :=
+ match x with
+ | FS f =>
+ match Float32.to_long f with
+ | Some i => L i
+ | None => if va_strict tt then Vbot else ntop
+ end
+ | _ => ntop1 x
+ end.
+
+Lemma longofsingle_sound:
+ forall v x w, vmatch v x -> Val.longofsingle v = Some w -> vmatch w (longofsingle x).
+Proof.
+ unfold Val.longofsingle; intros. destruct v; try discriminate.
+ destruct (Float32.to_long f) as [i|] eqn:E; simpl in H0; inv H0.
+ inv H; simpl; auto with va. rewrite E; constructor.
+Qed.
+
+Definition longuofsingle (x: aval) :=
+ match x with
+ | FS f =>
+ match Float32.to_longu f with
+ | Some i => L i
+ | None => if va_strict tt then Vbot else ntop
+ end
+ | _ => ntop1 x
+ end.
+
+Lemma longuofsingle_sound:
+ forall v x w, vmatch v x -> Val.longuofsingle v = Some w -> vmatch w (longuofsingle x).
+Proof.
+ unfold Val.longuofsingle; intros. destruct v; try discriminate.
+ destruct (Float32.to_longu f) as [i|] eqn:E; simpl in H0; inv H0.
+ inv H; simpl; auto with va. rewrite E; constructor.
+Qed.
+
+Definition singleoflong (x: aval) :=
+ match x with
+ | L i => FS(Float32.of_long i)
+ | _ => ntop1 x
+ end.
+
+Lemma singleoflong_sound:
+ forall v x w, vmatch v x -> Val.singleoflong v = Some w -> vmatch w (singleoflong x).
+Proof.
+ unfold Val.singleoflong; intros. destruct v; inv H0.
+ inv H; simpl; auto with va.
+Qed.
+
+Definition singleoflongu (x: aval) :=
+ match x with
+ | L i => FS(Float32.of_longu i)
+ | _ => ntop1 x
+ end.
+
+Lemma singleoflongu_sound:
+ forall v x w, vmatch v x -> Val.singleoflongu v = Some w -> vmatch w (singleoflongu x).
+Proof.
+ unfold Val.singleoflongu; intros. destruct v; inv H0.
+ inv H; simpl; auto with va.
+Qed.
+
Definition floatofwords (x y: aval) :=
match x, y with
| I i, I j => F(Float.from_words i j)
@@ -2166,13 +2593,17 @@ Proof.
assert (IP: forall i b ofs,
cmatch (Val.cmpu_bool valid c (Vint i) (Vptr b ofs)) (cmp_different_blocks c)).
{
- intros. simpl. destruct (Int.eq i Int.zero && (valid b (Int.unsigned ofs) || valid b (Int.unsigned ofs - 1))).
+ intros. simpl. destruct Archi.ptr64.
+ apply cmp_different_blocks_none.
+ destruct (Int.eq i Int.zero && (valid b (Ptrofs.unsigned ofs) || valid b (Ptrofs.unsigned ofs - 1))).
apply cmp_different_blocks_sound. apply cmp_different_blocks_none.
}
assert (PI: forall i b ofs,
cmatch (Val.cmpu_bool valid c (Vptr b ofs) (Vint i)) (cmp_different_blocks c)).
{
- intros. simpl. destruct (Int.eq i Int.zero && (valid b (Int.unsigned ofs) || valid b (Int.unsigned ofs - 1))).
+ intros. simpl. destruct Archi.ptr64.
+ apply cmp_different_blocks_none.
+ destruct (Int.eq i Int.zero && (valid b (Ptrofs.unsigned ofs) || valid b (Ptrofs.unsigned ofs - 1))).
apply cmp_different_blocks_sound. apply cmp_different_blocks_none.
}
unfold cmpu_bool; inversion H; subst; inversion H0; subst;
@@ -2199,6 +2630,58 @@ Proof.
- constructor.
Qed.
+Definition cmplu_bool (c: comparison) (v w: aval) : abool :=
+ match v, w with
+ | L i1, L i2 => Just (Int64.cmpu c i1 i2)
+ | Ptr _, L _ => cmp_different_blocks c
+ | L _, Ptr _ => cmp_different_blocks c
+ | Ptr p1, Ptr p2 => pcmp c p1 p2
+ | Ptr p1, Ifptr p2 => club (pcmp c p1 p2) (cmp_different_blocks c)
+ | Ifptr p1, Ptr p2 => club (pcmp c p1 p2) (cmp_different_blocks c)
+ | _, _ => Btop
+ end.
+
+Lemma cmplu_bool_sound:
+ forall valid c v w x y, vmatch v x -> vmatch w y -> cmatch (Val.cmplu_bool valid c v w) (cmplu_bool c x y).
+Proof.
+ intros.
+ assert (IP: forall i b ofs,
+ cmatch (Val.cmplu_bool valid c (Vlong i) (Vptr b ofs)) (cmp_different_blocks c)).
+ {
+ intros. simpl. destruct Archi.ptr64; simpl.
+ destruct (Int64.eq i Int64.zero && (valid b (Ptrofs.unsigned ofs) || valid b (Ptrofs.unsigned ofs - 1))).
+ apply cmp_different_blocks_sound. apply cmp_different_blocks_none.
+ apply cmp_different_blocks_none.
+ }
+ assert (PI: forall i b ofs,
+ cmatch (Val.cmplu_bool valid c (Vptr b ofs) (Vlong i)) (cmp_different_blocks c)).
+ {
+ intros. simpl. destruct Archi.ptr64; simpl.
+ destruct (Int64.eq i Int64.zero && (valid b (Ptrofs.unsigned ofs) || valid b (Ptrofs.unsigned ofs - 1))).
+ apply cmp_different_blocks_sound. apply cmp_different_blocks_none.
+ apply cmp_different_blocks_none.
+ }
+ unfold cmplu_bool; inversion H; subst; inversion H0; subst;
+ auto using cmatch_top, cmp_different_blocks_none, pcmp_none,
+ cmatch_lub_l, cmatch_lub_r, pcmp_sound_64.
+- constructor.
+Qed.
+
+Definition cmpl_bool (c: comparison) (v w: aval) : abool :=
+ match v, w with
+ | L i1, L i2 => Just (Int64.cmp c i1 i2)
+ | _, _ => Btop
+ end.
+
+Lemma cmpl_bool_sound:
+ forall c v w x y, vmatch v x -> vmatch w y -> cmatch (Val.cmpl_bool c v w) (cmpl_bool c x y).
+Proof.
+ intros.
+ unfold cmpl_bool; inversion H; subst; inversion H0; subst;
+ auto using cmatch_top.
+- constructor.
+Qed.
+
Definition cmpf_bool (c: comparison) (v w: aval) : abool :=
match v, w with
| F f1, F f2 => Just (Float.cmp c f1 f2)
@@ -2298,12 +2781,15 @@ Definition vnormalize (chunk: memory_chunk) (v: aval) :=
| Mint16unsigned, I i => I (Int.zero_ext 16 i)
| Mint16unsigned, Uns p n => Uns (provenance v) (Z.min n 16)
| Mint16unsigned, _ => Uns (provenance v) 16
- | Mint32, (I _ | Uns _ _ | Sgn _ _ | Ptr _ | Ifptr _) => v
- | Mint64, L _ => v
- | Mint64, (Ptr p | Ifptr p | Uns p _ | Sgn p _) => Ifptr (if va_strict tt then Pbot else p)
+ | Mint32, (I _ | Uns _ _ | Sgn _ _ | Ifptr _) => v
+ | Mint32, Ptr p => if Archi.ptr64 then Ifptr p else v
+ | Mint64, (L _ | Ifptr _) => v
+ | Mint64, (Uns p _ | Sgn p _) => Ifptr p
+ | Mint64, Ptr p => if Archi.ptr64 then v else Ifptr p
| Mfloat32, FS f => v
| Mfloat64, F f => v
- | Many32, (I _ | Uns _ _ | Sgn _ _ | Ptr _ | Ifptr _ | FS _) => v
+ | Many32, (I _ | Uns _ _ | Sgn _ _ | FS _ | Ifptr _) => v
+ | Many32, Ptr p => if Archi.ptr64 then Ifptr p else v
| Many64, _ => v
| _, _ => Ifptr (provenance v)
end.
@@ -2311,7 +2797,8 @@ Definition vnormalize (chunk: memory_chunk) (v: aval) :=
Lemma vnormalize_sound:
forall chunk v x, vmatch v x -> vmatch (Val.load_result chunk v) (vnormalize chunk x).
Proof.
- unfold Val.load_result, vnormalize; induction 1; destruct chunk; auto with va.
+ unfold Val.load_result, vnormalize; generalize Archi.ptr64; intros ptr64;
+ induction 1; destruct chunk; auto with va.
- destruct (zlt n 8); constructor; auto with va.
apply is_sign_ext_uns; auto.
apply is_sign_ext_sgn; auto with va.
@@ -2326,10 +2813,19 @@ Proof.
- constructor. omega. apply is_zero_ext_uns; auto with va.
- constructor. xomega. apply is_sign_ext_sgn; auto with va. apply Z.min_case; auto with va.
- constructor. omega. apply is_zero_ext_uns; auto with va.
+- destruct ptr64; auto with va.
+- destruct ptr64; auto with va.
+- destruct ptr64; auto with va.
+- destruct ptr64; auto with va.
+- destruct ptr64; auto with va.
+- destruct ptr64; auto with va.
- constructor. omega. apply is_sign_ext_sgn; auto with va.
- constructor. omega. apply is_zero_ext_uns; auto with va.
- constructor. omega. apply is_sign_ext_sgn; auto with va.
- constructor. omega. apply is_zero_ext_uns; auto with va.
+- destruct ptr64; auto with va.
+- destruct ptr64; auto with va.
+- destruct ptr64; auto with va.
Qed.
Lemma vnormalize_cast:
@@ -2351,13 +2847,13 @@ Proof.
- (* int32 *)
auto.
- (* int64 *)
- destruct v; try contradiction; constructor.
+ auto.
- (* float32 *)
destruct v; try contradiction; constructor.
- (* float64 *)
destruct v; try contradiction; constructor.
- (* any32 *)
- auto.
+ destruct Archi.ptr64; auto.
- (* any64 *)
auto.
Qed.
@@ -2379,7 +2875,7 @@ Lemma vnormalize_monotone:
forall chunk x y,
vge x y -> vge (vnormalize chunk x) (vnormalize chunk y).
Proof with (auto using provenance_monotone with va).
- intros chunk x y V; inversion V; subst; destruct chunk; simpl...
+ intros chunk x y V; unfold vnormalize; generalize Archi.ptr64; intro ptr64; inversion V; subst; destruct chunk eqn:C; simpl...
- destruct (zlt n 8); constructor...
apply is_sign_ext_uns...
apply is_sign_ext_sgn...
@@ -2393,19 +2889,19 @@ Proof with (auto using provenance_monotone with va).
destruct (zlt n2 8)...
- destruct (zlt n1 16). rewrite zlt_true by omega...
destruct (zlt n2 16)...
-- destruct (va_strict tt)...
- constructor... apply is_sign_ext_sgn... apply Z.min_case...
- constructor... apply is_zero_ext_uns...
- constructor... apply is_sign_ext_sgn... apply Z.min_case...
- constructor... apply is_zero_ext_uns...
- unfold provenance; destruct (va_strict tt)...
-- destruct (va_strict tt)...
- destruct (zlt n2 8); constructor...
- destruct (zlt n2 16); constructor...
-- destruct (va_strict tt)...
-- destruct (va_strict tt)...
-- destruct (va_strict tt)...
-- destruct (va_strict tt)...
+- destruct ptr64...
+- destruct ptr64...
+- destruct ptr64...
+- destruct ptr64...
+- destruct ptr64...
+- destruct ptr64...
- constructor... apply is_sign_ext_sgn...
- constructor... apply is_zero_ext_uns...
- constructor... apply is_sign_ext_sgn...
@@ -2420,8 +2916,6 @@ Proof with (auto using provenance_monotone with va).
- unfold provenance; destruct (va_strict tt)...
- destruct (zlt n 8)...
- destruct (zlt n 16)...
-- destruct (va_strict tt)...
-- destruct (va_strict tt)...
Qed.
(** Abstracting memory blocks *)
@@ -2648,7 +3142,7 @@ Lemma store_provenance:
forall chunk m b ofs v m' b' ofs' b'' ofs'' q i,
Mem.store chunk m b ofs v = Some m' ->
Mem.loadbytes m' b' ofs' 1 = Some (Fragment (Vptr b'' ofs'') q i :: nil) ->
- v = Vptr b'' ofs'' /\ (chunk = Mint32 \/ chunk = Many32 \/ chunk = Many64)
+ v = Vptr b'' ofs'' /\ (chunk = Mint32 \/ chunk = Many32 \/ chunk = Mint64 \/ chunk = Many64)
\/ Mem.loadbytes m b' ofs' 1 = Some (Fragment (Vptr b'' ofs'') q i :: nil).
Proof.
intros. exploit storebytes_provenance; eauto. eapply Mem.store_storebytes; eauto.
@@ -2704,7 +3198,8 @@ Proof.
generalize (decode_val_shape chunk byte1' bytes'). rewrite <- Q.
intros DEC; inv DEC; try contradiction.
assert (v = Vptr bx ofsx).
- { destruct H5 as [E|[E|E]]; rewrite E in H4; destruct v; simpl in H4; congruence. }
+ { destruct H5 as [E|[E|[E|E]]]; rewrite E in H4; destruct v; simpl in H4;
+ try congruence; destruct Archi.ptr64; congruence. }
exploit In_loadbytes; eauto. eauto with coqlib.
intros (ofs' & X & Y). subst v.
exploit storebytes_provenance; eauto. intros [Z | Z].
@@ -3177,10 +3672,10 @@ Definition load (chunk: memory_chunk) (rm: romem) (m: amem) (p: aptr) : aval :=
| Pbot => if va_strict tt then Vbot else Vtop
| Gl id ofs =>
match rm!id with
- | Some ab => ablock_load chunk ab (Int.unsigned ofs)
+ | Some ab => ablock_load chunk ab (Ptrofs.unsigned ofs)
| None =>
match m.(am_glob)!id with
- | Some ab => ablock_load chunk ab (Int.unsigned ofs)
+ | Some ab => ablock_load chunk ab (Ptrofs.unsigned ofs)
| None => vnormalize chunk (Ifptr m.(am_nonstack))
end
end
@@ -3193,7 +3688,7 @@ Definition load (chunk: memory_chunk) (rm: romem) (m: amem) (p: aptr) : aval :=
| None => vnormalize chunk (Ifptr m.(am_nonstack))
end
end
- | Stk ofs => ablock_load chunk m.(am_stack) (Int.unsigned ofs)
+ | Stk ofs => ablock_load chunk m.(am_stack) (Ptrofs.unsigned ofs)
| Stack => ablock_load_anywhere chunk m.(am_stack)
| Glob | Nonstack => vnormalize chunk (Ifptr m.(am_nonstack))
| Ptop => vnormalize chunk (Ifptr m.(am_top))
@@ -3205,7 +3700,7 @@ Definition loadv (chunk: memory_chunk) (rm: romem) (m: amem) (addr: aval) : aval
Definition store (chunk: memory_chunk) (m: amem) (p: aptr) (av: aval) : amem :=
{| am_stack :=
match p with
- | Stk ofs => ablock_store chunk m.(am_stack) (Int.unsigned ofs) av
+ | Stk ofs => ablock_store chunk m.(am_stack) (Ptrofs.unsigned ofs) av
| Stack | Ptop => ablock_store_anywhere chunk m.(am_stack) av
| _ => m.(am_stack)
end;
@@ -3213,7 +3708,7 @@ Definition store (chunk: memory_chunk) (m: amem) (p: aptr) (av: aval) : amem :=
match p with
| Gl id ofs =>
let ab := match m.(am_glob)!id with Some ab => ab | None => ablock_init m.(am_nonstack) end in
- PTree.set id (ablock_store chunk ab (Int.unsigned ofs) av) m.(am_glob)
+ PTree.set id (ablock_store chunk ab (Ptrofs.unsigned ofs) av) m.(am_glob)
| Glo id =>
let ab := match m.(am_glob)!id with Some ab => ab | None => ablock_init m.(am_nonstack) end in
PTree.set id (ablock_store_anywhere chunk ab av) m.(am_glob)
@@ -3251,7 +3746,7 @@ Definition loadbytes (m: amem) (rm: romem) (p: aptr) : aptr :=
Definition storebytes (m: amem) (dst: aptr) (sz: Z) (p: aptr) : amem :=
{| am_stack :=
match dst with
- | Stk ofs => ablock_storebytes m.(am_stack) p (Int.unsigned ofs) sz
+ | Stk ofs => ablock_storebytes m.(am_stack) p (Ptrofs.unsigned ofs) sz
| Stack | Ptop => ablock_storebytes_anywhere m.(am_stack) p
| _ => m.(am_stack)
end;
@@ -3259,7 +3754,7 @@ Definition storebytes (m: amem) (dst: aptr) (sz: Z) (p: aptr) : amem :=
match dst with
| Gl id ofs =>
let ab := match m.(am_glob)!id with Some ab => ab | None => ablock_init m.(am_nonstack) end in
- PTree.set id (ablock_storebytes ab p (Int.unsigned ofs) sz) m.(am_glob)
+ PTree.set id (ablock_storebytes ab p (Ptrofs.unsigned ofs) sz) m.(am_glob)
| Glo id =>
let ab := match m.(am_glob)!id with Some ab => ab | None => ablock_init m.(am_nonstack) end in
PTree.set id (ablock_storebytes_anywhere ab p) m.(am_glob)
@@ -3276,7 +3771,7 @@ Definition storebytes (m: amem) (dst: aptr) (sz: Z) (p: aptr) : amem :=
Theorem load_sound:
forall chunk m b ofs v rm am p,
- Mem.load chunk m b (Int.unsigned ofs) = Some v ->
+ Mem.load chunk m b (Ptrofs.unsigned ofs) = Some v ->
romatch m rm ->
mmatch m am ->
pmatch b ofs p ->
@@ -3321,7 +3816,7 @@ Qed.
Theorem store_sound:
forall chunk m b ofs v m' am p av,
- Mem.store chunk m b (Int.unsigned ofs) v = Some m' ->
+ Mem.store chunk m b (Ptrofs.unsigned ofs) v = Some m' ->
mmatch m am ->
pmatch b ofs p ->
vmatch v av ->
@@ -3399,7 +3894,7 @@ Qed.
Theorem loadbytes_sound:
forall m b ofs sz bytes am rm p,
- Mem.loadbytes m b (Int.unsigned ofs) sz = Some bytes ->
+ Mem.loadbytes m b (Ptrofs.unsigned ofs) sz = Some bytes ->
romatch m rm ->
mmatch m am ->
pmatch b ofs p ->
@@ -3432,7 +3927,7 @@ Qed.
Theorem storebytes_sound:
forall m b ofs bytes m' am p sz q,
- Mem.storebytes m b (Int.unsigned ofs) bytes = Some m' ->
+ Mem.storebytes m b (Ptrofs.unsigned ofs) bytes = Some m' ->
mmatch m am ->
pmatch b ofs p ->
length bytes = nat_of_Z sz ->
@@ -3716,8 +4211,8 @@ Lemma vmatch_inj:
forall bc v x, vmatch bc v x -> Val.inject (inj_of_bc bc) v v.
Proof.
induction 1; econstructor.
- eapply pmatch_inj; eauto. rewrite Int.add_zero; auto.
- eapply pmatch_inj; eauto. rewrite Int.add_zero; auto.
+ eapply pmatch_inj; eauto. rewrite Ptrofs.add_zero; auto.
+ eapply pmatch_inj; eauto. rewrite Ptrofs.add_zero; auto.
Qed.
Lemma vmatch_list_inj:
@@ -3752,7 +4247,7 @@ Proof.
{ exploit mmatch_top; eauto. intros [P Q].
eapply pmatch_top'. eapply Q; eauto. }
inv PM; auto.
- rewrite Int.add_zero; auto.
+ rewrite Ptrofs.add_zero; auto.
- (* free blocks *)
intros. unfold inj_of_bc. erewrite bc_below_invalid; eauto.
- (* mapped blocks *)
@@ -3765,7 +4260,7 @@ Proof.
auto.
- (* overflow *)
intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst.
- rewrite Zplus_0_r. split. omega. apply Int.unsigned_range_2.
+ rewrite Zplus_0_r. split. omega. apply Ptrofs.unsigned_range_2.
- (* perm inv *)
intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst.
rewrite Zplus_0_r in H2. auto.
@@ -4046,13 +4541,21 @@ Hint Resolve cnot_sound symbol_address_sound
neg_sound add_sound sub_sound
mul_sound mulhs_sound mulhu_sound
divs_sound divu_sound mods_sound modu_sound shrx_sound
+ shll_sound shrl_sound shrlu_sound
+ andl_sound orl_sound xorl_sound notl_sound roll_sound rorl_sound
+ negl_sound addl_sound subl_sound mull_sound
+ divls_sound divlu_sound modls_sound modlu_sound
negf_sound absf_sound
addf_sound subf_sound mulf_sound divf_sound
negfs_sound absfs_sound
addfs_sound subfs_sound mulfs_sound divfs_sound
- zero_ext_sound sign_ext_sound singleoffloat_sound floatofsingle_sound
+ zero_ext_sound sign_ext_sound longofint_sound longofintu_sound
+ singleoffloat_sound floatofsingle_sound
intoffloat_sound intuoffloat_sound floatofint_sound floatofintu_sound
intofsingle_sound intuofsingle_sound singleofint_sound singleofintu_sound
+ longoffloat_sound longuoffloat_sound floatoflong_sound floatoflongu_sound
+ longofsingle_sound longuofsingle_sound singleoflong_sound singleoflongu_sound
longofwords_sound loword_sound hiword_sound
- cmpu_bool_sound cmp_bool_sound cmpf_bool_sound cmpfs_bool_sound
+ cmpu_bool_sound cmp_bool_sound cmplu_bool_sound cmpl_bool_sound
+ cmpf_bool_sound cmpfs_bool_sound
maskzero_sound : va.