Skip to content

Commit f66711d

Browse files
authored
Merge pull request #22 from AbsIntPrivate/arm_large_offsets
Issue #P18: handle large offsets when accessing return address and back link in the stack frame
2 parents ab6d5e9 + fc1b2bf commit f66711d

8 files changed

+201
-91
lines changed

arm/Asm.v

+6-3
Original file line numberDiff line numberDiff line change
@@ -199,15 +199,16 @@ Inductive instruction : Type :=
199199
| Pfsts: freg -> ireg -> int -> instruction (**r float32 store *)
200200

201201
(* Pseudo-instructions *)
202-
| Pallocframe: Z -> ptrofs -> instruction (**r allocate new stack frame *)
203-
| Pfreeframe: Z -> ptrofs -> instruction (**r deallocate stack frame and restore previous frame *)
204-
| Plabel: label -> instruction (**r define a code label *)
202+
| Pallocframe: Z -> ptrofs -> instruction (**r allocate new stack frame *)
203+
| Pfreeframe: Z -> ptrofs -> instruction (**r deallocate stack frame and restore previous frame *)
204+
| Plabel: label -> instruction (**r define a code label *)
205205
| Ploadsymbol: ireg -> ident -> ptrofs -> instruction (**r load the address of a symbol *)
206206
| Pmovite: testcond -> ireg -> shift_op -> shift_op -> instruction (**r integer conditional move *)
207207
| Pbtbl: ireg -> list label -> instruction (**r N-way branch through a jump table *)
208208
| Pbuiltin: external_function -> list (builtin_arg preg) -> builtin_res preg -> instruction (**r built-in function (pseudo) *)
209209
| Padc: ireg -> ireg -> shift_op -> instruction (**r add with carry *)
210210
| Pcfi_adjust: int -> instruction (**r .cfi_adjust debug directive *)
211+
| Pcfi_rel_offset: int -> instruction (**r .cfi_rel_offset debug directive *)
211212
| Pclz: ireg -> ireg -> instruction (**r count leading zeros. *)
212213
| Pfsqrt: freg -> freg -> instruction (**r floating-point square root. *)
213214
| Prev: ireg -> ireg -> instruction (**r reverse bytes and reverse bits. *)
@@ -757,6 +758,8 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
757758
end
758759
| _ => Stuck
759760
end
761+
| Pcfi_rel_offset ofs =>
762+
Next (nextinstr rs) m
760763
| Pbuiltin ef args res => Stuck (**r treated specially below *)
761764
(** The following instructions and directives are not generated directly by Asmgen,
762765
so we do not model them. *)

arm/Asmexpand.ml

+37-17
Original file line numberDiff line numberDiff line change
@@ -46,18 +46,22 @@ let expand_movimm dst n =
4646
(fun n -> emit (Porr (dst,dst, SOimm n))) tl
4747

4848
let expand_subimm dst src n =
49-
match Asmgen.decompose_int n with
50-
| [] -> assert false
51-
| hd::tl ->
52-
emit (Psub(dst,src,SOimm hd));
53-
List.iter (fun n -> emit (Psub (dst,dst,SOimm n))) tl
49+
if dst <> src || n <> _0 then begin
50+
match Asmgen.decompose_int n with
51+
| [] -> assert false
52+
| hd::tl ->
53+
emit (Psub(dst,src,SOimm hd));
54+
List.iter (fun n -> emit (Psub (dst,dst,SOimm n))) tl
55+
end
5456

5557
let expand_addimm dst src n =
56-
match Asmgen.decompose_int n with
57-
| [] -> assert false
58-
| hd::tl ->
59-
emit (Padd (dst,src,SOimm hd));
60-
List.iter (fun n -> emit (Padd (dst,dst,SOimm n))) tl
58+
if dst <> src || n <> _0 then begin
59+
match Asmgen.decompose_int n with
60+
| [] -> assert false
61+
| hd::tl ->
62+
emit (Padd (dst,src,SOimm hd));
63+
List.iter (fun n -> emit (Padd (dst,dst,SOimm n))) tl
64+
end
6165

6266
let expand_int64_arith conflict rl fn =
6367
if conflict then
@@ -410,12 +414,22 @@ let expand_instruction instr =
410414
| Pallocframe (sz, ofs) ->
411415
emit (Pmov (IR12,SOreg IR13));
412416
if (is_current_function_variadic ()) then begin
413-
emit (Ppush [IR0;IR1;IR2;IR3]);
414-
emit (Pcfi_adjust _16);
415-
end;
416-
expand_subimm IR13 IR13 sz;
417-
emit (Pcfi_adjust sz);
418-
emit (Pstr (IR12,IR13,SOimm ofs));
417+
emit (Ppush [IR0;IR1;IR2;IR3]);
418+
emit (Pcfi_adjust _16);
419+
end;
420+
let sz' = camlint_of_coqint sz in
421+
let ofs' = camlint_of_coqint ofs in
422+
if ofs' >= 4096l && sz' >= ofs' then begin
423+
expand_subimm IR13 IR13 (coqint_of_camlint (Int32.sub sz' (Int32.add ofs' 4l)));
424+
emit (Ppush [IR12]);
425+
expand_subimm IR13 IR13 ofs;
426+
emit (Pcfi_adjust sz);
427+
end else begin
428+
assert (ofs' < 4096l);
429+
expand_subimm IR13 IR13 sz;
430+
emit (Pcfi_adjust sz);
431+
emit (Pstr (IR12,IR13,SOimm ofs));
432+
end;
419433
PrintAsmaux.current_function_stacksize := camlint_of_coqint sz
420434
| Pfreeframe (sz, ofs) ->
421435
let sz =
@@ -424,7 +438,13 @@ let expand_instruction instr =
424438
else sz in
425439
if Asmgen.is_immed_arith sz
426440
then emit (Padd (IR13,IR13,SOimm sz))
427-
else emit (Pldr (IR13,IR13,SOimm ofs))
441+
else begin
442+
if camlint_of_coqint ofs >= 4096l then begin
443+
expand_addimm IR13 IR13 ofs;
444+
emit (Pldr (IR13,IR13,SOimm _0))
445+
end else
446+
emit (Pldr (IR13,IR13,SOimm ofs));
447+
end
428448
| Pbuiltin (ef,args,res) ->
429449
begin match ef with
430450
| EF_builtin (name,sg) ->

arm/Asmgen.v

+20-2
Original file line numberDiff line numberDiff line change
@@ -602,6 +602,22 @@ Definition storeind (src: mreg) (base: ireg) (ofs: ptrofs) (ty: typ) (k: code) :
602602
Error (msg "Asmgen.storeind")
603603
end.
604604

605+
(** This is a variant of [storeind] that is used to save the return address
606+
at the beginning of a function. It uses [R12] instead of [R14] as
607+
temporary register. *)
608+
609+
Definition save_lr (ofs: ptrofs) (k: code) :=
610+
let n := Ptrofs.to_int ofs in
611+
let n1 := mk_immed_mem_word n in
612+
if Int.eq n n1
613+
then Pstr IR14 IR13 (SOimm n) :: k
614+
else addimm IR12 IR13 (Int.sub n n1) (Pstr IR14 IR12 (SOimm n1) :: k).
615+
616+
Definition save_lr_preserves_R12 (ofs: ptrofs) : bool :=
617+
let n := Ptrofs.to_int ofs in
618+
let n1 := mk_immed_mem_word n in
619+
Int.eq n n1.
620+
605621
(** Translation of memory accesses *)
606622

607623
Definition transl_memory_access
@@ -787,10 +803,12 @@ Definition transl_code' (f: Mach.function) (il: list Mach.instruction) (it1p: bo
787803
around, leading to incorrect executions. *)
788804

789805
Definition transl_function (f: Mach.function) :=
790-
do c <- transl_code f f.(Mach.fn_code) true;
806+
do c <- transl_code f f.(Mach.fn_code)
807+
(save_lr_preserves_R12 f.(fn_retaddr_ofs));
791808
OK (mkfunction f.(Mach.fn_sig)
792809
(Pallocframe f.(fn_stacksize) f.(fn_link_ofs) ::
793-
Pstr IR14 IR13 (SOimm (Ptrofs.to_int f.(fn_retaddr_ofs))) :: c)).
810+
save_lr f.(fn_retaddr_ofs)
811+
(Pcfi_rel_offset (Ptrofs.to_int f.(fn_retaddr_ofs)):: c))).
794812

795813
Definition transf_function (f: Mach.function) : res Asm.function :=
796814
do tf <- transl_function f;

arm/Asmgenproof.v

+41-20
Original file line numberDiff line numberDiff line change
@@ -241,6 +241,15 @@ Proof.
241241
destruct ty, (preg_of src); inv H; TailNoLabel.
242242
Qed.
243243

244+
Remark save_lr_label:
245+
forall ofs k, tail_nolabel k (save_lr ofs k).
246+
Proof.
247+
unfold save_lr; intros.
248+
destruct (Int.eq (Ptrofs.to_int ofs) (mk_immed_mem_word (Ptrofs.to_int ofs))).
249+
TailNoLabel.
250+
eapply tail_nolabel_trans; TailNoLabel.
251+
Qed.
252+
244253
Remark transl_cond_label:
245254
forall cond args k c, transl_cond cond args k = OK c -> tail_nolabel k c.
246255
Proof.
@@ -338,7 +347,7 @@ Lemma transl_find_label:
338347
end.
339348
Proof.
340349
intros. monadInv H. destruct (zlt Ptrofs.max_unsigned (list_length_z (fn_code x))); inv EQ0.
341-
monadInv EQ. simpl.
350+
monadInv EQ. simpl. erewrite tail_nolabel_find_label by (apply save_lr_label). simpl.
342351
eapply transl_code_label; eauto.
343352
Qed.
344353

@@ -382,7 +391,8 @@ Proof.
382391
destruct i; try (intros [A B]; apply A). intros. subst c0. repeat constructor.
383392
- intros. monadInv H0.
384393
destruct (zlt Ptrofs.max_unsigned (list_length_z (fn_code x))); inv EQ0. monadInv EQ.
385-
exists x; exists true; split; auto. repeat constructor.
394+
exists x; exists (save_lr_preserves_R12 (fn_retaddr_ofs f0)); split; auto.
395+
constructor. eapply is_tail_trans. 2: apply tail_nolabel_is_tail; apply save_lr_label. repeat constructor.
386396
- exact transf_function_no_overflow.
387397
Qed.
388398

@@ -854,7 +864,10 @@ Opaque loadind.
854864
generalize EQ; intros EQ'. monadInv EQ'.
855865
destruct (zlt Ptrofs.max_unsigned (list_length_z (fn_code x0))); inversion EQ1. clear EQ1. subst x0.
856866
monadInv EQ0.
857-
set (tfbody := Pallocframe (fn_stacksize f) (fn_link_ofs f) :: Pstr IR14 IR13 (SOimm (Ptrofs.to_int (fn_retaddr_ofs f))) :: x0) in *.
867+
set (ra_ofs := fn_retaddr_ofs f) in *.
868+
set (ra_ofs' := Ptrofs.to_int ra_ofs) in *.
869+
set (tfbody := Pallocframe (fn_stacksize f) (fn_link_ofs f) ::
870+
save_lr ra_ofs (Pcfi_rel_offset ra_ofs' :: x0)) in *.
858871
set (tf := {| fn_sig := Mach.fn_sig f; fn_code := tfbody |}) in *.
859872
unfold store_stack in *.
860873
exploit Mem.alloc_extends. eauto. eauto. apply Zle_refl. apply Zle_refl.
@@ -865,32 +878,40 @@ Opaque loadind.
865878
intros [m3' [P Q]].
866879
(* Execution of function prologue *)
867880
set (rs2 := nextinstr (rs0#IR12 <- (parent_sp s) #IR13 <- (Vptr stk Ptrofs.zero))).
868-
set (rs3 := nextinstr rs2).
881+
edestruct (save_lr_correct tge tf ra_ofs (Pcfi_rel_offset ra_ofs' :: x0) rs2) as (rs3 & X & Y & Z).
882+
change (rs2 IR13) with sp. change (rs2 IR14) with (rs0 IR14). rewrite ATLR. eexact P.
883+
set (rs4 := nextinstr rs3).
869884
assert (EXEC_PROLOGUE:
870885
exec_straight tge tf
871886
(fn_code tf) rs0 m'
872-
x0 rs3 m3').
887+
x0 rs4 m3').
888+
{
873889
change (fn_code tf) with tfbody; unfold tfbody.
874-
apply exec_straight_two with rs2 m2'.
890+
eapply exec_straight_trans with (rs2 := rs2) (m2 := m2').
891+
apply exec_straight_one.
875892
unfold exec_instr. rewrite C. fold sp.
876893
rewrite <- (sp_val _ _ _ AG). unfold Tptr, chunk_of_type, Archi.ptr64 in F. rewrite F. auto.
877-
simpl. auto.
878-
simpl. unfold exec_store. change (rs2 IR14) with (rs0 IR14).
879-
rewrite Ptrofs.add_zero_l. simpl. unfold Tptr, chunk_of_type, Archi.ptr64 in P. simpl in P.
880-
rewrite Ptrofs.add_zero_l in P. rewrite ATLR. rewrite Ptrofs.of_int_to_int by auto.
881-
rewrite P. auto. auto. auto.
882-
left; exists (State rs3 m3'); split.
894+
auto.
895+
eapply exec_straight_trans with (rs2 := rs3) (m2 := m3').
896+
eexact X.
897+
apply exec_straight_one.
898+
simpl; reflexivity. reflexivity.
899+
}
900+
(* After the function prologue is the code for the function body *)
901+
exploit exec_straight_steps_2; eauto using functions_transl. omega. constructor.
902+
intros (ofsbody & U & V).
903+
(* Conclusions *)
904+
left; exists (State rs4 m3'); split.
883905
eapply exec_straight_steps_1; eauto. omega. constructor.
884-
econstructor; eauto.
885-
change (rs3 PC) with (Val.offset_ptr (Val.offset_ptr (rs0 PC) Ptrofs.one) Ptrofs.one).
886-
rewrite ATPC. simpl. constructor; eauto.
887-
eapply code_tail_next_int. omega.
888-
eapply code_tail_next_int. omega. constructor.
889-
unfold rs3, rs2.
890-
apply agree_nextinstr. apply agree_nextinstr.
906+
econstructor; eauto. rewrite U. econstructor; eauto.
907+
apply agree_nextinstr.
908+
apply agree_undef_regs2 with rs2.
909+
apply agree_nextinstr.
891910
eapply agree_change_sp.
892911
apply agree_undef_regs with rs0; eauto.
893-
intros. Simpl. congruence.
912+
intros; Simpl.
913+
congruence.
914+
intros; apply Y; eauto with asmgen.
894915

895916
- (* external function *)
896917
exploit functions_translated; eauto.

arm/Asmgenproof1.v

+31
Original file line numberDiff line numberDiff line change
@@ -616,6 +616,37 @@ Proof.
616616
intros; Simpl.
617617
Qed.
618618

619+
(** Saving the link register *)
620+
621+
Lemma save_lr_correct:
622+
forall ofs k (rs: regset) m m',
623+
Mem.storev Mint32 m (Val.offset_ptr rs#IR13 ofs) (rs#IR14) = Some m' ->
624+
exists rs',
625+
exec_straight ge fn (save_lr ofs k) rs m k rs' m'
626+
/\ (forall r, if_preg r = true -> r <> IR12 -> rs'#r = rs#r)
627+
/\ (save_lr_preserves_R12 ofs = true -> rs'#IR12 = rs#IR12).
628+
Proof.
629+
intros; unfold save_lr, save_lr_preserves_R12.
630+
set (n := Ptrofs.to_int ofs). set (n1 := mk_immed_mem_word n).
631+
assert (EQ: Val.offset_ptr rs#IR13 ofs = Val.add rs#IR13 (Vint n)).
632+
{ destruct rs#IR13; try discriminate. simpl. f_equal; f_equal. unfold n; symmetry; auto with ptrofs. }
633+
destruct (Int.eq n n1).
634+
- econstructor; split. apply exec_straight_one. simpl; unfold exec_store. rewrite <- EQ, H; reflexivity. auto.
635+
split. intros; Simpl. intros; Simpl.
636+
- destruct (addimm_correct IR12 IR13 (Int.sub n n1) (Pstr IR14 IR12 (SOimm n1) :: k) rs m)
637+
as (rs1 & A & B & C).
638+
econstructor; split.
639+
eapply exec_straight_trans. eexact A.
640+
apply exec_straight_one. simpl; unfold exec_store.
641+
rewrite B. rewrite Val.add_assoc. simpl.
642+
rewrite Int.sub_add_opp. rewrite Int.add_assoc.
643+
rewrite (Int.add_commut (Int.neg n1)).
644+
rewrite Int.add_neg_zero. rewrite Int.add_zero.
645+
rewrite <- EQ. rewrite C by eauto with asmgen. rewrite H. reflexivity.
646+
auto.
647+
split. intros; Simpl. congruence.
648+
Qed.
649+
619650
(** Translation of shift immediates *)
620651

621652
Lemma transl_shift_correct:

0 commit comments

Comments
 (0)