aboutsummaryrefslogtreecommitdiffstats
path: root/backend/Asmgenproof0.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2019-08-08 11:18:38 +0200
committerXavier Leroy <xavier.leroy@college-de-france.fr>2019-08-08 11:18:38 +0200
commit7cdd676d002e33015b496f609538a9e86d77c543 (patch)
treef4d105bce152445334613e857d4a672976a56f3e /backend/Asmgenproof0.v
parenteb85803875c5a4e90be60d870f01fac380ca18b0 (diff)
downloadcompcert-kvx-7cdd676d002e33015b496f609538a9e86d77c543.tar.gz
compcert-kvx-7cdd676d002e33015b496f609538a9e86d77c543.zip
AArch64 port
This commit adds a back-end for the AArch64 architecture, namely ARMv8 in 64-bit mode.
Diffstat (limited to 'backend/Asmgenproof0.v')
-rw-r--r--backend/Asmgenproof0.v51
1 files changed, 37 insertions, 14 deletions
diff --git a/backend/Asmgenproof0.v b/backend/Asmgenproof0.v
index 111e435f..3638c465 100644
--- a/backend/Asmgenproof0.v
+++ b/backend/Asmgenproof0.v
@@ -899,30 +899,53 @@ Qed.
(** A variant that supports zero steps of execution *)
-Inductive exec_straight0: code -> regset -> mem ->
- code -> regset -> mem -> Prop :=
- | exec_straight0_none:
- forall c rs m,
- exec_straight0 c rs m c rs m
- | exec_straight0_step:
- forall i c rs1 m1 rs2 m2 c' rs3 m3,
- exec_instr ge fn i rs1 m1 = Next rs2 m2 ->
- rs2#PC = Val.offset_ptr rs1#PC Ptrofs.one ->
- exec_straight0 c rs2 m2 c' rs3 m3 ->
- exec_straight0 (i :: c) rs1 m1 c' rs3 m3.
+Inductive exec_straight_opt: code -> regset -> mem -> code -> regset -> mem -> Prop :=
+ | exec_straight_opt_refl: forall c rs m,
+ exec_straight_opt c rs m c rs m
+ | exec_straight_opt_intro: forall c1 rs1 m1 c2 rs2 m2,
+ exec_straight c1 rs1 m1 c2 rs2 m2 ->
+ exec_straight_opt c1 rs1 m1 c2 rs2 m2.
+
+Lemma exec_straight_opt_left:
+ forall c3 rs3 m3 c1 rs1 m1 c2 rs2 m2,
+ exec_straight c1 rs1 m1 c2 rs2 m2 ->
+ exec_straight_opt c2 rs2 m2 c3 rs3 m3 ->
+ exec_straight c1 rs1 m1 c3 rs3 m3.
+Proof.
+ destruct 2; intros. auto. eapply exec_straight_trans; eauto.
+Qed.
+
+Lemma exec_straight_opt_right:
+ forall c3 rs3 m3 c1 rs1 m1 c2 rs2 m2,
+ exec_straight_opt c1 rs1 m1 c2 rs2 m2 ->
+ exec_straight c2 rs2 m2 c3 rs3 m3 ->
+ exec_straight c1 rs1 m1 c3 rs3 m3.
+Proof.
+ destruct 1; intros. auto. eapply exec_straight_trans; eauto.
+Qed.
-Lemma exec_straight_step':
+Lemma exec_straight_opt_step:
forall i c rs1 m1 rs2 m2 c' rs3 m3,
exec_instr ge fn i rs1 m1 = Next rs2 m2 ->
rs2#PC = Val.offset_ptr rs1#PC Ptrofs.one ->
- exec_straight0 c rs2 m2 c' rs3 m3 ->
+ exec_straight_opt c rs2 m2 c' rs3 m3 ->
exec_straight (i :: c) rs1 m1 c' rs3 m3.
Proof.
- intros. revert i rs1 m1 H H0. revert H1. induction 1; intros.
+ intros. inv H1.
- apply exec_straight_one; auto.
- eapply exec_straight_step; eauto.
Qed.
+Lemma exec_straight_opt_step_opt:
+ forall i c rs1 m1 rs2 m2 c' rs3 m3,
+ exec_instr ge fn i rs1 m1 = Next rs2 m2 ->
+ rs2#PC = Val.offset_ptr rs1#PC Ptrofs.one ->
+ exec_straight_opt c rs2 m2 c' rs3 m3 ->
+ exec_straight_opt (i :: c) rs1 m1 c' rs3 m3.
+Proof.
+ intros. apply exec_straight_opt_intro. eapply exec_straight_opt_step; eauto.
+Qed.
+
End STRAIGHTLINE.
(** * Properties of the Mach call stack *)