Library compcert.ia32.standard.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 int_caller_save_regs := AX :: CX :: DX :: nil.
Definition float_caller_save_regs := X0 :: X1 :: X2 :: X3 :: X4 :: X5 :: X6 :: X7 :: nil.
Definition int_callee_save_regs := BX :: SI :: DI :: BP :: nil.
Definition float_callee_save_regs : list mreg := nil.
Definition destroyed_at_call :=
FP0 :: int_caller_save_regs ++ float_caller_save_regs.
Definition dummy_int_reg := AX. Definition dummy_float_reg := X0.
The index_int_callee_save and index_float_callee_save associate
a unique positive integer to callee-save registers. This integer is
used in Stacking to determine where to save these registers in
the activation record if they are used by the current function.
Definition index_int_callee_save (r: mreg) :=
match r with
| BX => 0 | SI => 1 | DI => 2 | BP => 3 | _ => -1
end.
Definition index_float_callee_save (r: mreg) := -1.
Ltac ElimOrEq :=
match goal with
| |- (?x = ?y) \/ _ -> _ =>
let H := fresh in
(intro H; elim H; clear H;
[intro H; rewrite <- H; clear H | ElimOrEq])
| |- False -> _ =>
let H := fresh in (intro H; contradiction)
end.
Ltac OrEq :=
match goal with
| |- (?x = ?x) \/ _ => left; reflexivity
| |- (?x = ?y) \/ _ => right; OrEq
| |- False => fail
end.
Ltac NotOrEq :=
match goal with
| |- (?x = ?y) \/ _ -> False =>
let H := fresh in (
intro H; elim H; clear H; [intro; discriminate | NotOrEq])
| |- False -> False =>
contradiction
end.
Lemma index_int_callee_save_pos:
forall r, In r int_callee_save_regs -> index_int_callee_save r >= 0.
Proof.
intro r. simpl; ElimOrEq; unfold index_int_callee_save; omega.
Qed.
Lemma index_float_callee_save_pos:
forall r, In r float_callee_save_regs -> index_float_callee_save r >= 0.
Proof.
intro r. simpl; ElimOrEq; unfold index_float_callee_save; omega.
Qed.
Lemma index_int_callee_save_pos2:
forall r, index_int_callee_save r >= 0 -> In r int_callee_save_regs.
Proof.
destruct r; simpl; intro; omegaContradiction || OrEq.
Qed.
Lemma index_float_callee_save_pos2:
forall r, index_float_callee_save r >= 0 -> In r float_callee_save_regs.
Proof.
unfold index_float_callee_save; intros. omegaContradiction.
Qed.
Lemma index_int_callee_save_inj:
forall r1 r2,
In r1 int_callee_save_regs ->
In r2 int_callee_save_regs ->
r1 <> r2 ->
index_int_callee_save r1 <> index_int_callee_save r2.
Proof.
intros r1 r2.
simpl; ElimOrEq; ElimOrEq; unfold index_int_callee_save;
intros; congruence.
Qed.
Lemma index_float_callee_save_inj:
forall r1 r2,
In r1 float_callee_save_regs ->
In r2 float_callee_save_regs ->
r1 <> r2 ->
index_float_callee_save r1 <> index_float_callee_save r2.
Proof.
simpl; intros. contradiction.
Qed.
The following lemmas show that
(destroyed at call, integer callee-save, float callee-save)
is a partition of the set of machine registers.
Lemma int_float_callee_save_disjoint:
list_disjoint int_callee_save_regs float_callee_save_regs.
Proof.
red; intros r1 r2. simpl; ElimOrEq; ElimOrEq; discriminate.
Qed.
Lemma register_classification:
forall r,
In r destroyed_at_call \/ In r int_callee_save_regs \/ In r float_callee_save_regs.
Proof.
destruct r;
try (left; simpl; OrEq);
try (right; left; simpl; OrEq);
try (right; right; simpl; OrEq).
Qed.
Lemma int_callee_save_not_destroyed:
forall r,
In r destroyed_at_call -> In r int_callee_save_regs -> False.
Proof.
intros. revert H0 H. simpl. ElimOrEq; NotOrEq.
Qed.
Lemma float_callee_save_not_destroyed:
forall r,
In r destroyed_at_call -> In r float_callee_save_regs -> False.
Proof.
intros. revert H0 H. simpl. ElimOrEq; NotOrEq.
Qed.
Lemma int_callee_save_type:
forall r, In r int_callee_save_regs -> mreg_type r = Tint.
Proof.
intro. simpl; ElimOrEq; reflexivity.
Qed.
Lemma float_callee_save_type:
forall r, In r float_callee_save_regs -> mreg_type r = Tfloat.
Proof.
intro. simpl; ElimOrEq; reflexivity.
Qed.
Ltac NoRepet :=
match goal with
| |- list_norepet nil =>
apply list_norepet_nil
| |- list_norepet (?a :: ?b) =>
apply list_norepet_cons; [simpl; intuition discriminate | NoRepet]
end.
Lemma int_callee_save_norepet:
list_norepet int_callee_save_regs.
Proof.
unfold int_callee_save_regs; NoRepet.
Qed.
Lemma float_callee_save_norepet:
list_norepet float_callee_save_regs.
Proof.
unfold float_callee_save_regs; NoRepet.
Qed.
Function calling conventions
Location of function result
Definition loc_result (s: signature) : list mreg :=
match s.(sig_res) with
| None => AX :: nil
| Some Tint => AX :: nil
| Some (Tfloat | Tsingle) => FP0 :: nil
| Some Tlong => DX :: AX :: nil
end.
The result locations are caller-save registers
Lemma loc_result_caller_save:
forall (s: signature) (r: mreg),
In r (loc_result s) -> In r destroyed_at_call.
Proof.
intros.
assert (r = AX \/ r = DX \/ r = FP0).
unfold loc_result in H. destruct (sig_res s); [destruct t|idtac]; simpl in H; intuition.
destruct H0 as [A | [A | A]]; subst r; simpl; OrEq.
Qed.
Fixpoint loc_arguments_rec
(tyl: list typ) (ofs: Z) {struct tyl} : list loc :=
match tyl with
| nil => nil
| Tint :: tys => S Outgoing ofs Tint :: loc_arguments_rec tys (ofs + 1)
| Tfloat :: tys => S Outgoing ofs Tfloat :: loc_arguments_rec tys (ofs + 2)
| Tsingle :: tys => S Outgoing ofs Tsingle :: loc_arguments_rec tys (ofs + 1)
| Tlong :: tys => S Outgoing (ofs + 1) Tint :: S Outgoing ofs Tint :: loc_arguments_rec tys (ofs + 2)
end.
loc_arguments s returns the list of locations where to store arguments
when calling a function with signature s.
size_arguments s returns the number of Outgoing slots used
to call a function with signature s.
Fixpoint size_arguments_rec
(tyl: list typ) (ofs: Z) {struct tyl} : Z :=
match tyl with
| nil => ofs
| ty :: tys => size_arguments_rec tys (ofs + typesize ty)
end.
Definition size_arguments (s: signature) : Z :=
size_arguments_rec 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 => In r destroyed_at_call
| S Outgoing ofs ty => ofs >= 0 /\ ty <> Tlong
| _ => False
end.
Remark loc_arguments_rec_charact:
forall tyl ofs l,
In l (loc_arguments_rec tyl ofs) ->
match l with
| S Outgoing ofs´ ty => ofs´ >= ofs /\ ty <> Tlong
| _ => False
end.
Proof.
induction tyl; simpl loc_arguments_rec; intros.
destruct H.
destruct a.
- destruct H. subst l. split. omega. congruence.
exploit IHtyl; eauto.
destruct l; auto. destruct sl; auto. intuition omega.
- destruct H. subst l. split. omega. congruence.
exploit IHtyl; eauto.
destruct l; auto. destruct sl; auto. intuition omega.
- destruct H. subst l; split; [omega|congruence].
destruct H. subst l; split; [omega|congruence].
exploit IHtyl; eauto.
destruct l; auto. destruct sl; auto. intuition omega.
- destruct H. subst l. split. omega. congruence.
exploit IHtyl; eauto.
destruct l; auto. destruct sl; auto. intuition omega.
Qed.
Lemma loc_arguments_acceptable:
forall (s: signature) (l: loc),
In l (loc_arguments s) -> loc_argument_acceptable l.
Proof.
unfold loc_arguments; intros.
exploit loc_arguments_rec_charact; eauto.
unfold loc_argument_acceptable.
destruct l; tauto.
Qed.
Hint Resolve loc_arguments_acceptable: locs.
The offsets of Outgoing arguments are below size_arguments s.
Remark size_arguments_rec_above:
forall tyl ofs0, ofs0 <= size_arguments_rec tyl ofs0.
Proof.
induction tyl; simpl; intros.
omega.
apply Zle_trans with (ofs0 + typesize a); auto.
generalize (typesize_pos a); omega.
Qed.
Lemma size_arguments_above:
forall s, size_arguments s >= 0.
Proof.
intros; unfold size_arguments. apply Zle_ge.
apply size_arguments_rec_above.
Qed.
Lemma loc_arguments_bounded:
forall (s: signature) (ofs: Z) (ty: typ),
In (S Outgoing ofs ty) (loc_arguments s) ->
ofs + typesize ty <= size_arguments s.
Proof.
intros until ty. unfold loc_arguments, size_arguments. generalize (sig_args s) 0.
induction l; simpl; intros.
destruct H.
destruct a.
destruct H. inv H. apply size_arguments_rec_above. auto.
destruct H. inv H. apply size_arguments_rec_above. auto.
destruct H. inv H.
simpl typesize. replace (z + 1 + 1) with (z + 2) by omega. apply size_arguments_rec_above.
destruct H. inv H.
simpl typesize. apply Zle_trans with (ofs + 2). omega. apply size_arguments_rec_above.
auto.
destruct H. inv H. apply size_arguments_rec_above. auto.
Qed.