Library compcert.x86.Conventions1
Function calling conventions and other conventions regarding the use of
machine registers and stack slots.
Classification of machine registers
- Callee-save registers, whose value is preserved across a function call.
- Caller-save registers that can be modified during a function call.
Definition is_callee_save (r: mreg) : bool :=
match r with
| AX | CX | DX ⇒ false
| BX | BP ⇒ true
| SI | DI ⇒ negb Archi.ptr64
| R8 | R9 | R10 | R11 ⇒ false
| R12 | R13 | R14 | R15 ⇒ true
| X0 | X1 | X2 | X3 | X4 | X5 | X6 | X7 ⇒ false
| X8 | X9 | X10 | X11 | X12 | X13 | X14 | X15 ⇒ false
| FP0 ⇒ false
end.
Definition int_caller_save_regs :=
if Archi.ptr64
then AX :: CX :: DX :: SI :: DI :: R8 :: R9 :: R10 :: R11 :: nil
else AX :: CX :: DX :: nil.
Definition float_caller_save_regs :=
if Archi.ptr64
then X0 :: X1 :: X2 :: X3 :: X4 :: X5 :: X6 :: X7 ::
X8 :: X9 :: X10 :: X11 :: X12 :: X13 :: X14 :: X15 :: nil
else X0 :: X1 :: X2 :: X3 :: X4 :: X5 :: X6 :: X7 :: nil.
Definition int_callee_save_regs :=
if Archi.ptr64
then BX :: BP :: R12 :: R13 :: R14 :: R15 :: nil
else BX :: SI :: DI :: BP :: nil.
Definition float_callee_save_regs : list mreg := nil.
Definition destroyed_at_call :=
List.filter (fun r ⇒ negb (is_callee_save r)) all_mregs.
Definition dummy_int_reg := AX. Definition dummy_float_reg := X0.
Definition is_float_reg (r: mreg) :=
match r with
| AX | BX | CX | DX | SI | DI | BP
| R8 | R9 | R10 | R11 | R12 | R13 | R14 | R15 ⇒ false
| X0 | X1 | X2 | X3 | X4 | X5 | X6 | X7
| X8 | X9 | X10 | X11 | X12 | X13 | X14 | X15 | FP0 ⇒ true
end.
Function calling conventions
Location of function result
Definition loc_result_32 (s: signature) : rpair mreg :=
match s.(sig_res) with
| None ⇒ One AX
| Some (Tint | Tany32) ⇒ One AX
| Some (Tfloat | Tsingle) ⇒ One FP0
| Some Tany64 ⇒ One X0
| Some Tlong ⇒ Twolong DX AX
end.
Definition loc_result_64 (s: signature) : rpair mreg :=
match s.(sig_res) with
| None ⇒ One AX
| Some (Tint | Tlong | Tany32 | Tany64) ⇒ One AX
| Some (Tfloat | Tsingle) ⇒ One X0
end.
Definition loc_result :=
if Archi.ptr64 then loc_result_64 else loc_result_32.
The result registers have types compatible with that given in the signature.
Lemma loc_result_type:
∀ sig,
subtype (proj_sig_res sig) (typ_rpair mreg_type (loc_result sig)) = true.
Proof.
intros. unfold proj_sig_res, loc_result, loc_result_32, loc_result_64, mreg_type;
destruct Archi.ptr64; destruct (sig_res sig) as [[]|]; auto.
Qed.
The result locations are caller-save registers
Lemma loc_result_caller_save:
∀ (s: signature),
forall_rpair (fun r ⇒ is_callee_save r = false) (loc_result s).
Proof.
intros. unfold loc_result, loc_result_32, loc_result_64, is_callee_save;
destruct Archi.ptr64; destruct (sig_res s) as [[]|]; simpl; auto.
Qed.
If the result is in a pair of registers, those registers are distinct and have type Tint at least.
Lemma loc_result_pair:
∀ sg,
match loc_result sg with
| One _ ⇒ True
| Twolong r1 r2 ⇒
r1 ≠ r2 ∧ sg.(sig_res) = Some Tlong
∧ subtype Tint (mreg_type r1) = true ∧ subtype Tint (mreg_type r2) = true
∧ Archi.splitlong = true
end.
Proof.
intros. change Archi.splitlong with (negb Archi.ptr64).
unfold loc_result, loc_result_32, loc_result_64, mreg_type;
destruct Archi.ptr64; destruct (sig_res sg) as [[]|]; auto.
split; auto. congruence.
Qed.
The location of the result depends only on the result part of the signature
Lemma loc_result_exten:
∀ s1 s2, s1.(sig_res) = s2.(sig_res) → loc_result s1 = loc_result s2.
Proof.
intros. unfold loc_result, loc_result_32, loc_result_64.
destruct Archi.ptr64; rewrite H; auto.
Qed.
Fixpoint loc_arguments_32
(tyl: list typ) (ofs: Z) {struct tyl} : list (rpair loc) :=
match tyl with
| nil ⇒ nil
| ty :: tys ⇒
match ty with
| Tlong ⇒ Twolong (S Outgoing (ofs + 1) Tint) (S Outgoing ofs Tint)
| _ ⇒ One (S Outgoing ofs ty)
end
:: loc_arguments_32 tys (ofs + typesize ty)
end.
In the x86-64 ABI:
- The first 6 integer arguments are passed in registers DI, SI, DX, CX, R8, R9.
- The first 8 floating-point arguments are passed in registers X0 to X7.
- Extra arguments are passed on the stack, in Outgoing slots. Consecutive stack slots are separated by 8 bytes, even if only 4 bytes of data is used in a slot.
Definition int_param_regs := DI :: SI :: DX :: CX :: R8 :: R9 :: nil.
Definition float_param_regs := X0 :: X1 :: X2 :: X3 :: X4 :: X5 :: X6 :: X7 :: nil.
Fixpoint loc_arguments_64
(tyl: list typ) (ir fr ofs: Z) {struct tyl} : list (rpair loc) :=
match tyl with
| nil ⇒ nil
| (Tint | Tlong | Tany32 | Tany64) as ty :: tys ⇒
match list_nth_z int_param_regs ir with
| None ⇒
One (S Outgoing ofs ty) :: loc_arguments_64 tys ir fr (ofs + 2)
| Some ireg ⇒
One (R ireg) :: loc_arguments_64 tys (ir + 1) fr ofs
end
| (Tfloat | Tsingle) as ty :: tys ⇒
match list_nth_z float_param_regs fr with
| None ⇒
One (S Outgoing ofs ty) :: loc_arguments_64 tys ir fr (ofs + 2)
| Some freg ⇒
One (R freg) :: loc_arguments_64 tys ir (fr + 1) ofs
end
end.
loc_arguments s returns the list of locations where to store arguments
when calling a function with signature s.
Definition loc_arguments (s: signature) : list (rpair loc) :=
if Archi.ptr64
then loc_arguments_64 s.(sig_args) 0 0 0
else loc_arguments_32 s.(sig_args) 0.
Fixpoint size_arguments_32
(tyl: list typ) (ofs: Z) {struct tyl} : Z :=
match tyl with
| nil ⇒ ofs
| ty :: tys ⇒ size_arguments_32 tys (ofs + typesize ty)
end.
Fixpoint size_arguments_64 (tyl: list typ) (ir fr ofs: Z) {struct tyl} : Z :=
match tyl with
| nil ⇒ ofs
| (Tint | Tlong | Tany32 | Tany64) :: tys ⇒
match list_nth_z int_param_regs ir with
| None ⇒ size_arguments_64 tys ir fr (ofs + 2)
| Some ireg ⇒ size_arguments_64 tys (ir + 1) fr ofs
end
| (Tfloat | Tsingle) :: tys ⇒
match list_nth_z float_param_regs fr with
| None ⇒ size_arguments_64 tys ir fr (ofs + 2)
| Some freg ⇒ size_arguments_64 tys ir (fr + 1) ofs
end
end.
Definition size_arguments (s: signature) : Z :=
if Archi.ptr64
then size_arguments_64 s.(sig_args) 0 0 0
else size_arguments_32 s.(sig_args) 0.
Argument locations are either caller-save registers or Outgoing
stack slots at nonnegative offsets.
Definition loc_argument_acceptable (l: loc) : Prop :=
match l with
| R r ⇒ is_callee_save r = false
| S Outgoing ofs ty ⇒ ofs ≥ 0 ∧ (typealign ty | ofs)
| _ ⇒ False
end.
Definition loc_argument_32_charact (ofs: Z) (l: loc) : Prop :=
match l with
| S Outgoing ofs' ty ⇒ ofs' ≥ ofs ∧ typealign ty = 1
| _ ⇒ False
end.
Definition loc_argument_64_charact (ofs: Z) (l: loc) : Prop :=
match l with
| R r ⇒ In r int_param_regs ∨ In r float_param_regs
| S Outgoing ofs' ty ⇒ ofs' ≥ ofs ∧ (2 | ofs')
| _ ⇒ False
end.
Remark loc_arguments_32_charact:
∀ tyl ofs p,
In p (loc_arguments_32 tyl ofs) → forall_rpair (loc_argument_32_charact ofs) p.
Proof.
assert (X: ∀ ofs1 ofs2 l, loc_argument_32_charact ofs2 l → ofs1 ≤ ofs2 → loc_argument_32_charact ofs1 l).
{ destruct l; simpl; intros; auto. destruct sl; auto. intuition omega. }
induction tyl as [ | ty tyl]; simpl loc_arguments_32; intros.
- contradiction.
- destruct H.
+ destruct ty; subst p; simpl; omega.
+ apply IHtyl in H. generalize (typesize_pos ty); intros. destruct p; simpl in ×.
× eapply X; eauto; omega.
× destruct H; split; eapply X; eauto; omega.
Qed.
Remark loc_arguments_64_charact:
∀ tyl ir fr ofs p,
In p (loc_arguments_64 tyl ir fr ofs) → (2 | ofs) → forall_rpair (loc_argument_64_charact ofs) p.
Proof.
assert (X: ∀ ofs1 ofs2 l, loc_argument_64_charact ofs2 l → ofs1 ≤ ofs2 → loc_argument_64_charact ofs1 l).
{ destruct l; simpl; intros; auto. destruct sl; auto. intuition omega. }
assert (Y: ∀ ofs1 ofs2 p, forall_rpair (loc_argument_64_charact ofs2) p → ofs1 ≤ ofs2 → forall_rpair (loc_argument_64_charact ofs1) p).
{ destruct p; simpl; intuition eauto. }
assert (Z: ∀ ofs, (2 | ofs) → (2 | ofs + 2)).
{ intros. apply Z.divide_add_r; auto. apply Zdivide_refl. }
Opaque list_nth_z.
induction tyl; simpl loc_arguments_64; intros.
elim H.
assert (A: ∀ ty, In p
match list_nth_z int_param_regs ir with
| Some ireg ⇒ One (R ireg) :: loc_arguments_64 tyl (ir + 1) fr ofs
| None ⇒ One (S Outgoing ofs ty) :: loc_arguments_64 tyl ir fr (ofs + 2)
end →
forall_rpair (loc_argument_64_charact ofs) p).
{ intros. destruct (list_nth_z int_param_regs ir) as [r|] eqn:E; destruct H1.
subst. left. eapply list_nth_z_in; eauto.
eapply IHtyl; eauto.
subst. split. omega. assumption.
eapply Y; eauto. omega. }
assert (B: ∀ ty, In p
match list_nth_z float_param_regs fr with
| Some ireg ⇒ One (R ireg) :: loc_arguments_64 tyl ir (fr + 1) ofs
| None ⇒ One (S Outgoing ofs ty) :: loc_arguments_64 tyl ir fr (ofs + 2)
end →
forall_rpair (loc_argument_64_charact ofs) p).
{ intros. destruct (list_nth_z float_param_regs fr) as [r|] eqn:E; destruct H1.
subst. right. eapply list_nth_z_in; eauto.
eapply IHtyl; eauto.
subst. split. omega. assumption.
eapply Y; eauto. omega. }
destruct a; eauto.
Qed.
Lemma loc_arguments_acceptable:
∀ (s: signature) (p: rpair loc),
In p (loc_arguments s) → forall_rpair loc_argument_acceptable p.
Proof.
unfold loc_arguments; intros. destruct Archi.ptr64 eqn:SF.
-
assert (A: ∀ r, In r int_param_regs → is_callee_save r = false) by (unfold is_callee_save; rewrite SF; decide_goal).
assert (B: ∀ r, In r float_param_regs → is_callee_save r = false) by decide_goal.
assert (X: ∀ l, loc_argument_64_charact 0 l → loc_argument_acceptable l).
{ unfold loc_argument_64_charact, loc_argument_acceptable.
destruct l as [r | [] ofs ty]; auto. intros [C|C]; auto.
intros [C D]. split; auto. apply Zdivide_trans with 2; auto.
∃ (2 / typealign ty); destruct ty; reflexivity.
}
exploit loc_arguments_64_charact; eauto using Zdivide_0.
unfold forall_rpair; destruct p; intuition auto.
-
assert (X: ∀ l, loc_argument_32_charact 0 l → loc_argument_acceptable l).
{ destruct l as [r | [] ofs ty]; simpl; intuition auto. rewrite H2; apply Z.divide_1_l. }
exploit loc_arguments_32_charact; eauto.
unfold forall_rpair; destruct p; intuition auto.
Qed.
Hint Resolve loc_arguments_acceptable: locs.
Remark size_arguments_32_above:
∀ tyl ofs0, ofs0 ≤ size_arguments_32 tyl ofs0.
Proof.
induction tyl; simpl; intros.
omega.
apply Zle_trans with (ofs0 + typesize a); auto.
generalize (typesize_pos a); omega.
Qed.
Remark size_arguments_64_above:
∀ tyl ir fr ofs0,
ofs0 ≤ size_arguments_64 tyl ir fr ofs0.
Proof.
induction tyl; simpl; intros.
omega.
assert (A: ofs0 ≤
match list_nth_z int_param_regs ir with
| Some _ ⇒ size_arguments_64 tyl (ir + 1) fr ofs0
| None ⇒ size_arguments_64 tyl ir fr (ofs0 + 2)
end).
{ destruct (list_nth_z int_param_regs ir); eauto.
apply Zle_trans with (ofs0 + 2); auto. omega. }
assert (B: ofs0 ≤
match list_nth_z float_param_regs fr with
| Some _ ⇒ size_arguments_64 tyl ir (fr + 1) ofs0
| None ⇒ size_arguments_64 tyl ir fr (ofs0 + 2)
end).
{ destruct (list_nth_z float_param_regs fr); eauto.
apply Zle_trans with (ofs0 + 2); auto. omega. }
destruct a; auto.
Qed.
Lemma size_arguments_above:
∀ s, size_arguments s ≥ 0.
Proof.
intros; unfold size_arguments. apply Zle_ge.
destruct Archi.ptr64; [apply size_arguments_64_above|apply size_arguments_32_above].
Qed.
Lemma loc_arguments_32_bounded:
∀ ofs ty tyl ofs0,
In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments_32 tyl ofs0)) →
ofs + typesize ty ≤ size_arguments_32 tyl ofs0.
Proof.
induction tyl as [ | t l]; simpl; intros x IN.
- contradiction.
- rewrite in_app_iff in IN; destruct IN as [IN|IN].
+ apply Zle_trans with (x + typesize t); [|apply size_arguments_32_above].
Ltac decomp :=
match goal with
| [ H: _ ∨ _ |- _ ] ⇒ destruct H; decomp
| [ H: S _ _ _ = S _ _ _ |- _ ] ⇒ inv H
| [ H: False |- _ ] ⇒ contradiction
end.
destruct t; simpl in IN; decomp; simpl; omega.
+ apply IHl; auto.
Qed.
Lemma loc_arguments_64_bounded:
∀ ofs ty tyl ir fr ofs0,
In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments_64 tyl ir fr ofs0)) →
ofs + typesize ty ≤ size_arguments_64 tyl ir fr ofs0.
Proof.
induction tyl; simpl; intros.
contradiction.
assert (T: ∀ ty0, typesize ty0 ≤ 2).
{ destruct ty0; simpl; omega. }
assert (A: ∀ ty0,
In (S Outgoing ofs ty) (regs_of_rpairs
match list_nth_z int_param_regs ir with
| Some ireg ⇒
One (R ireg) :: loc_arguments_64 tyl (ir + 1) fr ofs0
| None ⇒ One (S Outgoing ofs0 ty0) :: loc_arguments_64 tyl ir fr (ofs0 + 2)
end) →
ofs + typesize ty ≤
match list_nth_z int_param_regs ir with
| Some _ ⇒ size_arguments_64 tyl (ir + 1) fr ofs0
| None ⇒ size_arguments_64 tyl ir fr (ofs0 + 2)
end).
{ intros. destruct (list_nth_z int_param_regs ir); simpl in H0; destruct H0.
- discriminate.
- eapply IHtyl; eauto.
- inv H0. apply Zle_trans with (ofs + 2). specialize (T ty). omega. apply size_arguments_64_above.
- eapply IHtyl; eauto. }
assert (B: ∀ ty0,
In (S Outgoing ofs ty) (regs_of_rpairs
match list_nth_z float_param_regs fr with
| Some ireg ⇒
One (R ireg) :: loc_arguments_64 tyl ir (fr + 1) ofs0
| None ⇒ One (S Outgoing ofs0 ty0) :: loc_arguments_64 tyl ir fr (ofs0 + 2)
end) →
ofs + typesize ty ≤
match list_nth_z float_param_regs fr with
| Some _ ⇒ size_arguments_64 tyl ir (fr + 1) ofs0
| None ⇒ size_arguments_64 tyl ir fr (ofs0 + 2)
end).
{ intros. destruct (list_nth_z float_param_regs fr); simpl in H0; destruct H0.
- discriminate.
- eapply IHtyl; eauto.
- inv H0. apply Zle_trans with (ofs + 2). specialize (T ty). omega. apply size_arguments_64_above.
- eapply IHtyl; eauto. }
destruct a; eauto.
Qed.
Lemma loc_arguments_bounded:
∀ (s: signature) (ofs: Z) (ty: typ),
In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments s)) →
ofs + typesize ty ≤ size_arguments s.
Proof.
unfold loc_arguments, size_arguments; intros.
destruct Archi.ptr64; eauto using loc_arguments_32_bounded, loc_arguments_64_bounded.
Qed.
Lemma loc_arguments_main:
loc_arguments signature_main = nil.
Proof.
unfold loc_arguments; destruct Archi.ptr64; reflexivity.
Qed.
Lemma loc_arguments_32_charact':
∀ tyl ofs p,
In p (regs_of_rpairs (loc_arguments_32 tyl ofs)) →
loc_argument_32_charact ofs p.
Proof.
intros tyl ofs p IN.
apply in_regs_of_rpairs_inv in IN. destruct IN as (P & IN & IN').
apply loc_arguments_32_charact in IN.
destruct P; simpl in *; auto.
destruct IN'; subst; auto. easy.
destruct IN' as [A|[A|A]]; inv A; intuition.
Qed.
Lemma loc_arguments_64_charact':
∀ tyl ir fr ofs p,
(2|ofs) →
In p (regs_of_rpairs (loc_arguments_64 tyl ir fr ofs)) →
loc_argument_64_charact ofs p.
Proof.
intros tyl ir fr ofs p DIV IN.
apply in_regs_of_rpairs_inv in IN. destruct IN as (P & IN & IN').
apply loc_arguments_64_charact in IN; auto.
destruct P; simpl in *; auto.
destruct IN'; subst; auto. easy.
destruct IN' as [A|[A|A]]; inv A; intuition.
Qed.
Lemma loc_arguments_32_norepet sg z:
Loc.norepet (regs_of_rpairs (loc_arguments_32 sg z)).
Proof.
revert z.
induction sg; simpl; auto using Loc.norepet_nil.
intros z.
assert (H32: ∀ ty,
Loc.norepet
(regs_of_rpair (One (S Outgoing z ty)) ++
regs_of_rpairs (loc_arguments_32 sg (z + typesize ty)))).
{
simpl. intros ty. apply Loc.norepet_cons; auto.
rewrite Loc.notin_iff.
intros l' H.
apply loc_arguments_32_charact' in H.
destruct l' ; try contradiction. simpl in H.
destruct sl; try contradiction.
right. simpl. omega.
}
destruct a; auto.
simpl in ×. apply Loc.norepet_cons.
- rewrite Loc.notin_iff.
inversion 1; subst.
simpl; right; omega.
apply loc_arguments_32_charact' in H0.
destruct l' ; try contradiction.
destruct sl; try contradiction.
destruct H0.
red.
right; simpl in *; omega.
- apply Loc.norepet_cons; auto.
rewrite Loc.notin_iff.
intros l' H0.
apply loc_arguments_32_charact' in H0.
destruct l' ; try contradiction.
destruct sl; try contradiction.
destruct H0.
red; right; simpl in *; omega.
Qed.
Definition loc_argument_64_charact' ofs ir fr l :=
match l with
| R r ⇒ (In r int_param_regs ∧ ∃ i, i ≥ ir ∧ list_nth_z int_param_regs i = Some r)
∨ (In r float_param_regs ∧ ∃ i, i ≥ fr ∧ list_nth_z float_param_regs i = Some r)
| S Local _ _ ⇒ False
| S Incoming _ _ ⇒ False
| S Outgoing ofs' _ ⇒ ofs' ≥ ofs ∧ (2 | ofs')
end.
Remark loc_arguments_64_charact'':
∀ tyl ir fr ofs p,
In p (loc_arguments_64 tyl ir fr ofs) → (2 | ofs) → forall_rpair (loc_argument_64_charact' ofs ir fr) p.
Proof.
assert (X: ∀ ofs1 ofs2 ir fr l, loc_argument_64_charact' ofs2 ir fr l → ofs1 ≤ ofs2 → loc_argument_64_charact' ofs1 ir fr l).
{ destruct l; simpl; intros; auto. destruct sl; auto. intuition omega. }
assert (Y: ∀ ofs1 ofs2 ir fr p, forall_rpair (loc_argument_64_charact' ofs2 ir fr) p → ofs1 ≤ ofs2 → forall_rpair (loc_argument_64_charact' ofs1 ir fr) p).
{ destruct p; simpl; intuition eauto. }
assert (Z: ∀ ofs, (2 | ofs) → (2 | ofs + 2)).
{ intros. apply Z.divide_add_r; auto. apply Zdivide_refl. }
assert (C: ∀ ofs ir1 ir2 fr l, loc_argument_64_charact' ofs ir2 fr l → ir1 ≤ ir2 → loc_argument_64_charact' ofs ir1 fr l).
{
Opaque int_param_regs float_param_regs.
destruct l; simpl; intros; auto. destruct H.
- destruct H.
left; split; auto. destruct H1 as (i & LE & NTH); ∃ i; split; eauto. omega.
- destruct H.
right; split; auto.
}
assert (D: ∀ ofs ir1 ir2 fr p, forall_rpair (loc_argument_64_charact' ofs ir2 fr) p → ir1 ≤ ir2 → forall_rpair (loc_argument_64_charact' ofs ir1 fr) p).
{ destruct p; simpl; intuition eauto.
}
assert (E: ∀ ofs ir fr1 fr2 l, loc_argument_64_charact' ofs ir fr2 l → fr1 ≤ fr2 → loc_argument_64_charact' ofs ir fr1 l).
{
Opaque int_param_regs float_param_regs.
destruct l; simpl; intros; auto. destruct H.
- destruct H.
left; split; auto.
- destruct H.
right; split; auto.
destruct H1 as (i & LE & NTH). ∃ i; split; eauto. omega.
}
assert (F: ∀ ofs ir fr1 fr2 p, forall_rpair (loc_argument_64_charact' ofs ir fr2) p → fr1 ≤ fr2 → forall_rpair (loc_argument_64_charact' ofs ir fr1) p).
{ destruct p; simpl; intuition eauto.
}
Opaque list_nth_z.
induction tyl; simpl loc_arguments_64; intros.
elim H.
assert (A: ∀ ty, In p
match list_nth_z int_param_regs ir with
| Some ireg ⇒ One (R ireg) :: loc_arguments_64 tyl (ir + 1) fr ofs
| None ⇒ One (S Outgoing ofs ty) :: loc_arguments_64 tyl ir fr (ofs + 2)
end →
forall_rpair (loc_argument_64_charact' ofs ir fr) p).
{ intros. destruct (list_nth_z int_param_regs ir) as [r|] eqn:G; destruct H1.
subst. left. split. eapply list_nth_z_in; eauto. ∃ ir; split. omega. auto.
eapply D.
eapply IHtyl. eauto. auto. omega.
subst. split. omega. assumption.
eapply Y; eauto. omega. }
assert (B: ∀ ty, In p
match list_nth_z float_param_regs fr with
| Some ireg ⇒ One (R ireg) :: loc_arguments_64 tyl ir (fr + 1) ofs
| None ⇒ One (S Outgoing ofs ty) :: loc_arguments_64 tyl ir fr (ofs + 2)
end →
forall_rpair (loc_argument_64_charact' ofs ir fr) p).
{ intros. destruct (list_nth_z float_param_regs fr) as [r|] eqn:I; destruct H1.
subst. right. split. eapply list_nth_z_in; eauto. ∃ fr; split. omega. auto.
eapply F. eapply IHtyl; eauto. omega.
subst. split. omega. assumption.
eapply Y; eauto. omega. }
destruct a; eauto.
Qed.
Remark loc_arguments_64_charact''':
∀ tyl ir fr ofs p,
In p (regs_of_rpairs (loc_arguments_64 tyl ir fr ofs)) →
(2 | ofs) →
loc_argument_64_charact' ofs ir fr p.
Proof.
intros.
apply in_regs_of_rpairs_inv in H.
destruct H as (p0 & IN & ROR).
exploit loc_arguments_64_charact''; eauto.
intros.
destruct p0; simpl in ×.
destruct ROR; subst; auto. easy.
intuition subst; auto.
Qed.
Lemma list_nth_z_rew:
∀ {A} (l: list A) a n,
list_nth_z (a::l) n =
if zeq n 0
then Some a
else list_nth_z l (Z.pred n).
Proof. simpl; intros. reflexivity. Qed.
Lemma list_nth_z_norepet_same:
∀ {A} (l: list A) (lnr: list_norepet l) r i1 i2,
list_nth_z l i1 = Some r →
list_nth_z l i2 = Some r →
i1 = i2.
Proof.
induction 1; simpl; intros; eauto.
discriminate.
rewrite list_nth_z_rew in H0, H1.
destruct (zeq i1 0). inv H0.
destruct (zeq i2 0). auto. apply list_nth_z_in in H1. congruence.
destruct (zeq i2 0). inv H1. apply list_nth_z_in in H0. congruence.
eapply IHlnr in H0. 2: exact H1. apply f_equal with (f:=Z.succ) in H0.
rewrite <- ! Zsucc_pred in H0.
eauto.
Qed.
Lemma loc_arguments_64_norepet sg:
∀ ir fr ofs,
(2 | ofs) →
Loc.norepet (regs_of_rpairs (loc_arguments_64 sg ir fr ofs)).
Proof.
Opaque int_param_regs float_param_regs.
induction sg; simpl; auto using Loc.norepet_nil.
assert (H64: ∀ ty ofs ir fr,
(2 | ofs + 2) →
Loc.norepet
(regs_of_rpair (One (S Outgoing ofs ty)) ++
regs_of_rpairs (loc_arguments_64 sg ir fr (ofs + 2)))).
{
simpl. intros ty ofs ir fr DIV.
apply Loc.norepet_cons; auto.
- rewrite Loc.notin_iff.
intros l' H.
apply loc_arguments_64_charact' in H; auto.
destruct l' ; try contradiction. simpl. auto.
destruct sl; try contradiction. simpl.
simpl in H.
right. simpl.
cut (typesize ty ≤ 2). omega.
destruct ty ; simpl; omega.
}
generalize (loc_arguments_64_charact''' sg).
unfold loc_argument_64_charact'.
assert ( ∀ r, In r int_param_regs → In r float_param_regs → False).
{
Transparent int_param_regs float_param_regs.
simpl. intuition congruence.
}
assert( list_norepet int_param_regs).
repeat (constructor; [ simpl; intuition congruence | ]); constructor.
assert( list_norepet float_param_regs).
repeat (constructor; [ simpl; intuition congruence | ]); constructor.
revert H H0 H1.
generalize int_param_regs as IPR, float_param_regs as FPR.
intros IPR FPR DISJ LNR1 LNR2 CHARACT.
assert (H64_reg: ∀ ir fr ofs ireg,
(2 | ofs ) →
list_nth_z IPR ir = Some ireg →
Loc.norepet
(regs_of_rpair (One (R ireg)) ++
regs_of_rpairs (loc_arguments_64 sg (ir + 1) fr ofs))).
{
simpl. intros ir fr ofs ireg DIV NTH.
apply Loc.norepet_cons; auto.
rewrite Loc.notin_iff.
intros l' H.
apply CHARACT in H; auto. simpl.
destruct l'; auto.
destruct H as [(IN & (i & SUP & EQ)) | (IN & (i & SUP & EQ))].
intro; subst.
exploit (list_nth_z_norepet_same (A:=mreg)). 2: apply NTH. 2: apply EQ. auto. intro; subst. omega.
apply list_nth_z_in in NTH. intro; subst. eauto.
}
assert (H64_freg: ∀ ir fr ofs ireg,
(2 | ofs ) →
list_nth_z FPR fr = Some ireg →
Loc.norepet
(regs_of_rpair (One (R ireg)) ++
regs_of_rpairs (loc_arguments_64 sg ir (fr + 1) ofs))).
{
simpl. intros ir fr ofs ireg DIV NTH.
apply Loc.norepet_cons; auto.
rewrite Loc.notin_iff.
intros l' H.
apply CHARACT in H; auto. simpl.
destruct l'; auto.
destruct H as [(IN & (i & SUP & EQ)) | (IN & (i & SUP & EQ))].
intro; subst. apply list_nth_z_in in NTH. eauto.
intro; subst.
exploit (list_nth_z_norepet_same (A:=mreg)). 2: apply NTH. 2: apply EQ. auto. intro; subst. omega.
}
intros.
destruct a; auto.
- destruct (list_nth_z IPR ir) eqn:EQ.
+ simpl in ×. apply H64_reg. auto. auto.
+ simpl in ×. apply H64. apply Z.divide_add_r; auto. apply Zdivide_refl.
- destruct (list_nth_z FPR fr) eqn:EQ.
+ simpl in ×. apply H64_freg. auto. auto.
+ simpl in ×. apply H64. apply Z.divide_add_r; auto. apply Zdivide_refl.
- destruct (list_nth_z IPR ir) eqn:EQ.
+ simpl in ×. apply H64_reg. auto. auto.
+ simpl in ×. apply H64. apply Z.divide_add_r; auto. apply Zdivide_refl.
- destruct (list_nth_z FPR fr) eqn:EQ.
+ simpl in ×. apply H64_freg. auto. auto.
+ simpl in ×. apply H64. apply Z.divide_add_r; auto. apply Zdivide_refl.
- destruct (list_nth_z IPR ir) eqn:EQ.
+ simpl in ×. apply H64_reg. auto. auto.
+ simpl in ×. apply H64. apply Z.divide_add_r; auto. apply Zdivide_refl.
- destruct (list_nth_z IPR ir) eqn:EQ.
+ simpl in ×. apply H64_reg. auto. auto.
+ simpl in ×. apply H64. apply Z.divide_add_r; auto. apply Zdivide_refl.
Qed.
Lemma loc_arguments_norepet:
∀ sg,
Loc.norepet (regs_of_rpairs (loc_arguments sg)).
Proof.
unfold loc_arguments; intros.
destruct Archi.ptr64.
apply loc_arguments_64_norepet. apply Z.divide_0_r.
apply loc_arguments_32_norepet.
Qed.