aboutsummaryrefslogtreecommitdiffstats
path: root/aarch64/TO_MERGE
diff options
context:
space:
mode:
Diffstat (limited to 'aarch64/TO_MERGE')
-rw-r--r--aarch64/TO_MERGE/Archi.v100
-rw-r--r--aarch64/TO_MERGE/Asmgen.v712
-rw-r--r--aarch64/TO_MERGE/Asmgenproof.v2787
-rw-r--r--aarch64/TO_MERGE/Asmgenproof1.v1836
-rw-r--r--aarch64/TO_MERGE/TargetPrinter.ml862
-rw-r--r--aarch64/TO_MERGE/extractionMachdep.v45
6 files changed, 6342 insertions, 0 deletions
diff --git a/aarch64/TO_MERGE/Archi.v b/aarch64/TO_MERGE/Archi.v
new file mode 100644
index 00000000..eb022db9
--- /dev/null
+++ b/aarch64/TO_MERGE/Archi.v
@@ -0,0 +1,100 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, Collège de France and INRIA Paris *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the GNU General Public License as published by *)
+(* the Free Software Foundation, either version 2 of the License, or *)
+(* (at your option) any later version. This file is also distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(** Architecture-dependent parameters for AArch64 *)
+
+From Flocq Require Import Binary Bits.
+Require Import ZArith List.
+
+Definition ptr64 := true.
+
+Definition big_endian := false.
+
+Definition align_int64 := 8%Z.
+Definition align_float64 := 8%Z.
+
+Definition splitlong := false.
+
+Lemma splitlong_ptr32: splitlong = true -> ptr64 = false.
+Proof.
+ unfold splitlong, ptr64; congruence.
+Qed.
+
+Definition default_nan_64 := (false, iter_nat 51 _ xO xH).
+Definition default_nan_32 := (false, iter_nat 22 _ xO xH).
+
+(** Choose the first signaling NaN, if any;
+ otherwise choose the first NaN;
+ otherwise use default. *)
+
+Definition choose_nan (is_signaling: positive -> bool)
+ (default: bool * positive)
+ (l0: list (bool * positive)) : bool * positive :=
+ let fix choose_snan (l1: list (bool * positive)) :=
+ match l1 with
+ | nil =>
+ match l0 with nil => default | n :: _ => n end
+ | ((s, p) as n) :: l1 =>
+ if is_signaling p then n else choose_snan l1
+ end
+ in choose_snan l0.
+
+Lemma choose_nan_idem: forall is_signaling default n,
+ choose_nan is_signaling default (n :: n :: nil) =
+ choose_nan is_signaling default (n :: nil).
+Proof.
+ intros. destruct n as [s p]; unfold choose_nan; simpl.
+ destruct (is_signaling p); auto.
+Qed.
+
+Definition choose_nan_64 :=
+ choose_nan (fun p => negb (Pos.testbit p 51)) default_nan_64.
+
+Definition choose_nan_32 :=
+ choose_nan (fun p => negb (Pos.testbit p 22)) default_nan_32.
+
+Lemma choose_nan_64_idem: forall n,
+ choose_nan_64 (n :: n :: nil) = choose_nan_64 (n :: nil).
+Proof. intros; apply choose_nan_idem. Qed.
+
+Lemma choose_nan_32_idem: forall n,
+ choose_nan_32 (n :: n :: nil) = choose_nan_32 (n :: nil).
+Proof. intros; apply choose_nan_idem. Qed.
+
+Definition fma_order {A: Type} (x y z: A) := (z, x, y).
+
+Definition fma_invalid_mul_is_nan := true.
+
+Definition float_of_single_preserves_sNaN := false.
+
+Global Opaque ptr64 big_endian splitlong
+ default_nan_64 choose_nan_64
+ default_nan_32 choose_nan_32
+ fma_order fma_invalid_mul_is_nan
+ float_of_single_preserves_sNaN.
+
+(** Which ABI to implement *)
+
+<<<<<<< HEAD
+Parameter pic_code: unit -> bool.
+
+Definition has_notrap_loads := false.
+=======
+Inductive abi_kind: Type :=
+ | AAPCS64 (**r ARM's standard as used in Linux and other ELF platforms *)
+ | Apple. (**r the variant used in macOS and iOS *)
+
+Parameter abi: abi_kind.
+>>>>>>> master
diff --git a/aarch64/TO_MERGE/Asmgen.v b/aarch64/TO_MERGE/Asmgen.v
new file mode 100644
index 00000000..c8e48b40
--- /dev/null
+++ b/aarch64/TO_MERGE/Asmgen.v
@@ -0,0 +1,712 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Sylvain Boulmé Grenoble-INP, VERIMAG *)
+(* Justus Fasse UGA, VERIMAG *)
+(* Xavier Leroy INRIA Paris-Rocquencourt *)
+(* David Monniaux CNRS, VERIMAG *)
+(* Cyril Six Kalray *)
+(* Léo Gourdin UGA, VERIMAG *)
+(* *)
+(* Copyright Kalray. Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+Require Import Recdef Coqlib Zwf Zbits.
+Require Import Errors AST Integers Floats Op.
+<<<<<<< HEAD
+Require Import Locations Compopts.
+Require Import Mach Asm Asmblock Asmblockgen Machblockgen PostpassScheduling.
+=======
+Require Import Locations Mach Asm.
+Require SelectOp.
+
+Local Open Scope string_scope.
+Local Open Scope list_scope.
+Local Open Scope error_monad_scope.
+
+(** Alignment check for symbols *)
+
+Parameter symbol_is_aligned : ident -> Z -> bool.
+(** [symbol_is_aligned id sz] checks whether the symbol [id] is [sz] aligned *)
+
+(** Extracting integer or float registers. *)
+
+Definition ireg_of (r: mreg) : res ireg :=
+ match preg_of r with IR mr => OK mr | _ => Error(msg "Asmgen.ireg_of") end.
+
+Definition freg_of (r: mreg) : res freg :=
+ match preg_of r with FR mr => OK mr | _ => Error(msg "Asmgen.freg_of") end.
+
+(** Recognition of immediate arguments for logical integer operations.*)
+
+(** Valid immediate arguments are repetitions of a bit pattern [B]
+ of length [e] = 2, 4, 8, 16, 32 or 64.
+ The bit pattern [B] must be of the form [0*1*0*] or [1*0*1*]
+ but must not be all zeros or all ones. *)
+
+(** The following automaton recognizes [0*1*0*|1*0*1*].
+<<
+ 0 1 0
+ / \ / \ / \
+ \ / \ / \ /
+ -0--> [B] --1--> [D] --0--> [F]
+ /
+ [A]
+ \
+ -1--> [C] --0--> [E] --1--> [G]
+ / \ / \ / \
+ \ / \ / \ /
+ 1 0 1
+>>
+*)
+
+Module Automaton.
+
+Inductive state : Type := SA | SB | SC | SD | SE | SF | SG | Sbad.
+
+Definition start := SA.
+
+Definition next (s: state) (b: bool) :=
+ match s, b with
+ | SA,false => SB | SA,true => SC
+ | SB,false => SB | SB,true => SD
+ | SC,false => SE | SC,true => SC
+ | SD,false => SF | SD,true => SD
+ | SE,false => SE | SE,true => SG
+ | SF,false => SF | SF,true => Sbad
+ | SG,false => Sbad | SG,true => SG
+ | Sbad,_ => Sbad
+ end.
+
+Definition accepting (s: state) :=
+ match s with
+ | SA | SB | SC | SD | SE | SF | SG => true
+ | Sbad => false
+ end.
+
+Fixpoint run (len: nat) (s: state) (x: Z) : bool :=
+ match len with
+ | Datatypes.O => accepting s
+ | Datatypes.S len => run len (next s (Z.odd x)) (Z.div2 x)
+ end.
+
+End Automaton.
+
+(** The following function determines the candidate length [e],
+ ensuring that [x] is a repetition [BB...B]
+ of a bit pattern [B] of length [e]. *)
+
+Definition logical_imm_length (x: Z) (sixtyfour: bool) : nat :=
+ (** [test n] checks that the low [2n] bits of [x] are of the
+ form [BB], that is, two occurrences of the same [n] bits *)
+ let test (n: Z) : bool :=
+ Z.eqb (Zzero_ext n x) (Zzero_ext n (Z.shiftr x n)) in
+ (** If [test n] fails, we know that the candidate length [e] is
+ at least [2n]. Hence we test with decreasing values of [n]:
+ 32, 16, 8, 4, 2. *)
+ if sixtyfour && negb (test 32) then 64%nat
+ else if negb (test 16) then 32%nat
+ else if negb (test 8) then 16%nat
+ else if negb (test 4) then 8%nat
+ else if negb (test 2) then 4%nat
+ else 2%nat.
+
+(** A valid logical immediate is
+- neither [0] nor [-1];
+- composed of a repetition [BBBBB] of a bit-pattern [B] of length [e]
+- the low [e] bits of the number, that is, [B], match [0*1*0*] or [1*0*1*].
+*)
+
+Definition is_logical_imm32 (x: int) : bool :=
+ negb (Int.eq x Int.zero) && negb (Int.eq x Int.mone) &&
+ Automaton.run (logical_imm_length (Int.unsigned x) false)
+ Automaton.start (Int.unsigned x).
+
+Definition is_logical_imm64 (x: int64) : bool :=
+ negb (Int64.eq x Int64.zero) && negb (Int64.eq x Int64.mone) &&
+ Automaton.run (logical_imm_length (Int64.unsigned x) true)
+ Automaton.start (Int64.unsigned x).
+>>>>>>> master
+
+
+Local Open Scope error_monad_scope.
+
+(** Functions called by the Asmexpand ocaml file, inspired and adapted from Asmblockgen.v *)
+
+Module Asmgen_expand.
+
+(* Load immediate *)
+
+Definition loadimm_k (sz: isize) (rd: ireg) (l: list (Z * Z)) (k: code) : code :=
+ List.fold_right (fun np k => Asm.Pmovk sz rd (fst np) (snd np) :: k) k l.
+
+Definition loadimm_z (sz: isize) (rd: ireg) (l: list (Z * Z)) (k: code) : code :=
+ match l with
+ | nil => Asm.Pmovz sz rd 0 0 :: k
+ | (n1, p1) :: l => Asm.Pmovz sz rd n1 p1 :: loadimm_k sz rd l k
+ end.
+
+Definition loadimm_n (sz: isize) (rd: ireg) (l: list (Z * Z)) (k: code) : code :=
+ match l with
+ | nil => Asm.Pmovn sz rd 0 0 :: k
+ | (n1, p1) :: l => Asm.Pmovn sz rd n1 p1 :: loadimm_k sz rd (negate_decomposition l) k
+ end.
+
+Definition loadimm (sz: isize) (rd: ireg) (n: Z) (k: code) : code :=
+ let N := match sz with W => 2%nat | X => 4%nat end in
+ let dz := decompose_int N n 0 in
+ let dn := decompose_int N (Z.lnot n) 0 in
+ if Nat.leb (List.length dz) (List.length dn)
+ then loadimm_z sz rd dz k
+ else loadimm_n sz rd dn k.
+
+Definition loadimm32 (rd: ireg) (n: int) (k: code) : code :=
+ if is_logical_imm32 n
+ then Asm.Porrimm W rd XZR (Int.unsigned n) :: k
+ else loadimm W rd (Int.unsigned n) k.
+
+Definition loadimm64 (rd: ireg) (n: int64) (k: code) : code :=
+ if is_logical_imm64 n
+ then Asm.Porrimm X rd XZR (Int64.unsigned n) :: k
+ else loadimm X rd (Int64.unsigned n) k.
+
+(* Add immediate *)
+
+Definition addimm_aux (insn: iregsp -> iregsp -> Z -> instruction)
+ (rd r1: iregsp) (n: Z) (k: code) :=
+ let nlo := Zzero_ext 12 n in
+ let nhi := n - nlo in
+ if Z.eqb nhi 0 then
+ insn rd r1 nlo :: k
+ else if Z.eqb nlo 0 then
+ insn rd r1 nhi :: k
+ else
+ insn rd r1 nhi :: insn rd rd nlo :: k.
+
+Definition addimm64 (rd r1: iregsp) (n: int64) (k: code) : code :=
+ let m := Int64.neg n in
+ if Int64.eq n (Int64.zero_ext 24 n) then
+ addimm_aux (Asm.Paddimm X) rd r1 (Int64.unsigned n) k
+ else if Int64.eq m (Int64.zero_ext 24 m) then
+ addimm_aux (Asm.Psubimm X) rd r1 (Int64.unsigned m) k
+ else if Int64.lt n Int64.zero then
+<<<<<<< HEAD
+ loadimm64 X16 m (Asm.Psubext rd r1 X16 (EOuxtx Int.zero) :: k)
+=======
+ loadimm64 X16 m (Psubext rd r1 X16 (EOuxtx Int.zero) :: k)
+ else
+ loadimm64 X16 n (Paddext rd r1 X16 (EOuxtx Int.zero) :: k).
+
+(** Logical immediate *)
+
+Definition logicalimm32
+ (insn1: ireg -> ireg0 -> Z -> instruction)
+ (insn2: ireg -> ireg0 -> ireg -> shift_op -> instruction)
+ (rd r1: ireg) (n: int) (k: code) : code :=
+ if is_logical_imm32 n
+ then insn1 rd r1 (Int.unsigned n) :: k
+ else loadimm32 X16 n (insn2 rd r1 X16 SOnone :: k).
+
+Definition logicalimm64
+ (insn1: ireg -> ireg0 -> Z -> instruction)
+ (insn2: ireg -> ireg0 -> ireg -> shift_op -> instruction)
+ (rd r1: ireg) (n: int64) (k: code) : code :=
+ if is_logical_imm64 n
+ then insn1 rd r1 (Int64.unsigned n) :: k
+ else loadimm64 X16 n (insn2 rd r1 X16 SOnone :: k).
+
+(** Sign- or zero-extended arithmetic *)
+
+Definition transl_extension (ex: extension) (a: int) : extend_op :=
+ match ex with Xsgn32 => EOsxtw a | Xuns32 => EOuxtw a end.
+
+Definition move_extended_base
+ (rd: ireg) (r1: ireg) (ex: extension) (k: code) : code :=
+ match ex with
+ | Xsgn32 => Pcvtsw2x rd r1 :: k
+ | Xuns32 => Pcvtuw2x rd r1 :: k
+ end.
+
+Definition move_extended
+ (rd: ireg) (r1: ireg) (ex: extension) (a: int) (k: code) : code :=
+ if Int.eq a Int.zero then
+ move_extended_base rd r1 ex k
+ else
+ move_extended_base rd r1 ex (Padd X rd XZR rd (SOlsl a) :: k).
+
+Definition arith_extended
+ (insnX: iregsp -> iregsp -> ireg -> extend_op -> instruction)
+ (insnS: ireg -> ireg0 -> ireg -> shift_op -> instruction)
+ (rd r1 r2: ireg) (ex: extension) (a: int) (k: code) : code :=
+ if Int.ltu a (Int.repr 5) then
+ insnX rd r1 r2 (transl_extension ex a) :: k
+ else
+ move_extended_base X16 r2 ex (insnS rd r1 X16 (SOlsl a) :: k).
+
+(** Extended right shift *)
+
+Definition shrx32 (rd r1: ireg) (n: int) (k: code) : code :=
+ if Int.eq n Int.zero then
+ Pmov rd r1 :: k
+ else
+ Porr W X16 XZR r1 (SOasr (Int.repr 31)) ::
+ Padd W X16 r1 X16 (SOlsr (Int.sub Int.iwordsize n)) ::
+ Porr W rd XZR X16 (SOasr n) :: k.
+
+Definition shrx64 (rd r1: ireg) (n: int) (k: code) : code :=
+ if Int.eq n Int.zero then
+ Pmov rd r1 :: k
+ else
+ Porr X X16 XZR r1 (SOasr (Int.repr 63)) ::
+ Padd X X16 r1 X16 (SOlsr (Int.sub Int64.iwordsize' n)) ::
+ Porr X rd XZR X16 (SOasr n) :: k.
+
+(** Load the address [id + ofs] in [rd] *)
+
+Definition loadsymbol (rd: ireg) (id: ident) (ofs: ptrofs) (k: code) : code :=
+ if SelectOp.symbol_is_relocatable id then
+ if Ptrofs.eq ofs Ptrofs.zero then
+ Ploadsymbol rd id :: k
+ else
+ Ploadsymbol rd id :: addimm64 rd rd (Ptrofs.to_int64 ofs) k
+>>>>>>> master
+ else
+ loadimm64 X16 n (Asm.Paddext rd r1 X16 (EOuxtx Int.zero) :: k).
+
+(** Register-indexed stores *)
+
+Definition indexed_memory_access (insn: Asm.addressing -> instruction)
+ (sz: Z) (base: iregsp) (ofs: ptrofs) (k: code) :=
+ let ofs := Ptrofs.to_int64 ofs in
+ if offset_representable sz ofs
+ then insn (ADimm base ofs) :: k
+ else loadimm64 X16 ofs (insn (ADreg base X16) :: k).
+
+Definition storeptr (src: ireg) (base: iregsp) (ofs: ptrofs) (k: code) :=
+ indexed_memory_access (Asm.Pstrx src) 8 base ofs k.
+
+End Asmgen_expand.
+
+(** * Translation from Asmblock to assembly language
+ Inspired from the KVX backend (see kvx/Asm.v and kvx/Asmgen.v) *)
+
+Module Asmblock_TRANSF.
+(* STUB *)
+
+Definition ireg_of_preg (p : Asm.preg) : res ireg :=
+ match p with
+ | DR (IR (RR1 r)) => OK r
+ | _ => Error (msg "Asmgen.ireg_of_preg")
+ end.
+
+Definition freg_of_preg (p : Asm.preg) : res freg :=
+ match p with
+ | DR (FR r) => OK r
+ | _ => Error (msg "Asmgen.freg_of_preg")
+ end.
+
+Definition iregsp_of_preg (p : Asm.preg) : res iregsp :=
+ match p with
+ | DR (IR r) => OK r
+ | _ => Error (msg "Asmgen.iregsp_of_preg")
+ end.
+
+Definition basic_to_instruction (b: basic) : res Asm.instruction :=
+ match b with
+ (* Aithmetic instructions *)
+ | PArith (PArithP (Padrp id ofs) rd) => do rd' <- ireg_of_preg rd;
+ OK (Asm.Padrp rd' id ofs)
+ | PArith (PArithP (Pmovz sz n pos) rd) => do rd' <- ireg_of_preg rd;
+ OK (Asm.Pmovz sz rd' n pos)
+ | PArith (PArithP (Pmovn sz n pos) rd) => do rd' <- ireg_of_preg rd;
+ OK (Asm.Pmovn sz rd' n pos)
+ | PArith (PArithP (Pfmovimms f) rd) => do rd' <- freg_of_preg rd;
+ OK (Asm.Pfmovimms rd' f)
+ | PArith (PArithP (Pfmovimmd f) rd) => do rd' <- freg_of_preg rd;
+ OK (Asm.Pfmovimmd rd' f)
+
+ | PArith (PArithPP (Pmovk sz n pos) rd rs) =>
+ if (Asm.preg_eq rd rs) then (
+ do rd' <- ireg_of_preg rd;
+ OK (Asm.Pmovk sz rd' n pos)
+ ) else
+ Error (msg "Asmgen.basic_to_instruction: Pmovk uses a single register as both source and target")
+ | PArith (PArithPP Pmov rd rs) => do rd' <- iregsp_of_preg rd;
+ do rs' <- iregsp_of_preg rs;
+ OK (Asm.Pmov rd' rs')
+ | PArith (PArithPP (Paddadr id ofs) rd rs) => do rd' <- ireg_of_preg rd;
+ do rs' <- ireg_of_preg rs;
+ OK (Asm.Paddadr rd' rs' id ofs)
+ | PArith (PArithPP (Psbfiz sz r s) rd rs) => do rd' <- ireg_of_preg rd;
+ do rs' <- ireg_of_preg rs;
+ OK (Asm.Psbfiz sz rd' rs' r s)
+ | PArith (PArithPP (Psbfx sz r s) rd rs) => do rd' <- ireg_of_preg rd;
+ do rs' <- ireg_of_preg rs;
+ OK (Asm.Psbfx sz rd' rs' r s)
+ | PArith (PArithPP (Pubfiz sz r s) rd rs) => do rd' <- ireg_of_preg rd;
+ do rs' <- ireg_of_preg rs;
+ OK (Asm.Pubfiz sz rd' rs' r s)
+ | PArith (PArithPP (Pubfx sz r s) rd rs) => do rd' <- ireg_of_preg rd;
+ do rs' <- ireg_of_preg rs;
+ OK (Asm.Pubfx sz rd' rs' r s)
+ | PArith (PArithPP Pfmov rd rs) => do rd' <- freg_of_preg rd;
+ do rs' <- freg_of_preg rs;
+ OK (Asm.Pfmov rd' rs')
+ | PArith (PArithPP Pfcvtds rd rs) => do rd' <- freg_of_preg rd;
+ do rs' <- freg_of_preg rs;
+ OK (Asm.Pfcvtds rd' rs')
+ | PArith (PArithPP Pfcvtsd rd rs) => do rd' <- freg_of_preg rd;
+ do rs' <- freg_of_preg rs;
+ OK (Asm.Pfcvtsd rd' rs')
+ | PArith (PArithPP (Pfabs sz) rd rs) => do rd' <- freg_of_preg rd;
+ do rs' <- freg_of_preg rs;
+ OK (Asm.Pfabs sz rd' rs')
+ | PArith (PArithPP (Pfneg sz) rd rs) => do rd' <- freg_of_preg rd;
+ do rs' <- freg_of_preg rs;
+ OK (Asm.Pfneg sz rd' rs')
+ | PArith (PArithPP (Pscvtf fsz isz) rd rs) => do rd' <- freg_of_preg rd;
+ do rs' <- ireg_of_preg rs;
+ OK (Asm.Pscvtf fsz isz rd' rs')
+ | PArith (PArithPP (Pucvtf fsz isz) rd rs) => do rd' <- freg_of_preg rd;
+ do rs' <- ireg_of_preg rs;
+ OK (Asm.Pucvtf fsz isz rd' rs')
+ | PArith (PArithPP (Pfcvtzs isz fsz) rd rs) => do rd' <- ireg_of_preg rd;
+ do rs' <- freg_of_preg rs;
+ OK (Asm.Pfcvtzs isz fsz rd' rs')
+ | PArith (PArithPP (Pfcvtzu isz fsz) rd rs) => do rd' <- ireg_of_preg rd;
+ do rs' <- freg_of_preg rs;
+ OK (Asm.Pfcvtzu isz fsz rd' rs')
+ | PArith (PArithPP (Paddimm sz n) rd rs) => do rd' <- iregsp_of_preg rd;
+ do rs' <- iregsp_of_preg rs;
+ OK (Asm.Paddimm sz rd' rs' n)
+ | PArith (PArithPP (Psubimm sz n) rd rs) => do rd' <- iregsp_of_preg rd;
+ do rs' <- iregsp_of_preg rs;
+ OK (Asm.Psubimm sz rd' rs' n)
+
+ | PArith (PArithPPP (Pasrv sz) rd r1 r2) => do rd' <- ireg_of_preg rd;
+ do r1' <- ireg_of_preg r1;
+ do r2' <- ireg_of_preg r2;
+ OK (Asm.Pasrv sz rd' r1' r2')
+ | PArith (PArithPPP (Plslv sz) rd r1 r2) => do rd' <- ireg_of_preg rd;
+ do r1' <- ireg_of_preg r1;
+ do r2' <- ireg_of_preg r2;
+ OK (Asm.Plslv sz rd' r1' r2')
+ | PArith (PArithPPP (Plsrv sz) rd r1 r2) => do rd' <- ireg_of_preg rd;
+ do r1' <- ireg_of_preg r1;
+ do r2' <- ireg_of_preg r2;
+ OK (Asm.Plsrv sz rd' r1' r2')
+ | PArith (PArithPPP (Prorv sz) rd r1 r2) => do rd' <- ireg_of_preg rd;
+ do r1' <- ireg_of_preg r1;
+ do r2' <- ireg_of_preg r2;
+ OK (Asm.Prorv sz rd' r1' r2')
+ | PArith (PArithPPP Psmulh rd r1 r2) => do rd' <- ireg_of_preg rd;
+ do r1' <- ireg_of_preg r1;
+ do r2' <- ireg_of_preg r2;
+ OK (Asm.Psmulh rd' r1' r2')
+ | PArith (PArithPPP Pumulh rd r1 r2) => do rd' <- ireg_of_preg rd;
+ do r1' <- ireg_of_preg r1;
+ do r2' <- ireg_of_preg r2;
+ OK (Asm.Pumulh rd' r1' r2')
+ | PArith (PArithPPP (Psdiv sz) rd r1 r2) => do rd' <- ireg_of_preg rd;
+ do r1' <- ireg_of_preg r1;
+ do r2' <- ireg_of_preg r2;
+ OK (Asm.Psdiv sz rd' r1' r2')
+ | PArith (PArithPPP (Pudiv sz) rd r1 r2) => do rd' <- ireg_of_preg rd;
+ do r1' <- ireg_of_preg r1;
+ do r2' <- ireg_of_preg r2;
+ OK (Asm.Pudiv sz rd' r1' r2')
+ | PArith (PArithPPP (Paddext x) rd r1 r2) => do rd' <- iregsp_of_preg rd;
+ do r1' <- iregsp_of_preg r1;
+ do r2' <- ireg_of_preg r2;
+ OK (Asm.Paddext rd' r1' r2' x)
+ | PArith (PArithPPP (Psubext x) rd r1 r2) => do rd' <- iregsp_of_preg rd;
+ do r1' <- iregsp_of_preg r1;
+ do r2' <- ireg_of_preg r2;
+ OK (Asm.Psubext rd' r1' r2' x)
+ | PArith (PArithPPP (Pfadd sz) rd r1 r2) => do rd' <- freg_of_preg rd;
+ do r1' <- freg_of_preg r1;
+ do r2' <- freg_of_preg r2;
+ OK (Asm.Pfadd sz rd' r1' r2')
+ | PArith (PArithPPP (Pfdiv sz) rd r1 r2) => do rd' <- freg_of_preg rd;
+ do r1' <- freg_of_preg r1;
+ do r2' <- freg_of_preg r2;
+ OK (Asm.Pfdiv sz rd' r1' r2')
+ | PArith (PArithPPP (Pfmul sz) rd r1 r2) => do rd' <- freg_of_preg rd;
+ do r1' <- freg_of_preg r1;
+ do r2' <- freg_of_preg r2;
+ OK (Asm.Pfmul sz rd' r1' r2')
+ | PArith (PArithPPP (Pfsub sz) rd r1 r2) => do rd' <- freg_of_preg rd;
+ do r1' <- freg_of_preg r1;
+ do r2' <- freg_of_preg r2;
+ OK (Asm.Pfsub sz rd' r1' r2')
+
+ | PArith (PArithRR0 (Pandimm sz n) rd r1) => OK (Asm.Pandimm sz rd r1 n)
+ | PArith (PArithRR0 (Peorimm sz n) rd r1) => OK (Asm.Peorimm sz rd r1 n)
+ | PArith (PArithRR0 (Porrimm sz n) rd r1) => OK (Asm.Porrimm sz rd r1 n)
+
+
+ | PArith (PArithRR0R (Padd sz s) rd r1 r2) => OK (Asm.Padd sz rd r1 r2 s)
+ | PArith (PArithRR0R (Psub sz s) rd r1 r2) => OK (Asm.Psub sz rd r1 r2 s)
+ | PArith (PArithRR0R (Pand sz s) rd r1 r2) => OK (Asm.Pand sz rd r1 r2 s)
+ | PArith (PArithRR0R (Pbic sz s) rd r1 r2) => OK (Asm.Pbic sz rd r1 r2 s)
+ | PArith (PArithRR0R (Peon sz s) rd r1 r2) => OK (Asm.Peon sz rd r1 r2 s)
+ | PArith (PArithRR0R (Peor sz s) rd r1 r2) => OK (Asm.Peor sz rd r1 r2 s)
+ | PArith (PArithRR0R (Porr sz s) rd r1 r2) => OK (Asm.Porr sz rd r1 r2 s)
+ | PArith (PArithRR0R (Porn sz s) rd r1 r2) => OK (Asm.Porn sz rd r1 r2 s)
+
+ | PArith (PArithARRRR0 (Pmadd sz) rd r1 r2 r3) => OK (Asm.Pmadd sz rd r1 r2 r3)
+ | PArith (PArithARRRR0 (Pmsub sz) rd r1 r2 r3) => OK (Asm.Pmsub sz rd r1 r2 r3)
+
+ | PArith (PArithComparisonPP (Pcmpext x) r1 r2) => do r1' <- ireg_of_preg r1;
+ do r2' <- ireg_of_preg r2;
+ OK (Asm.Pcmpext r1' r2' x)
+ | PArith (PArithComparisonPP (Pcmnext x) r1 r2) => do r1' <- ireg_of_preg r1;
+ do r2' <- ireg_of_preg r2;
+ OK (Asm.Pcmnext r1' r2' x)
+ | PArith (PArithComparisonPP (Pfcmp sz) r1 r2) => do r1' <- freg_of_preg r1;
+ do r2' <- freg_of_preg r2;
+ OK (Asm.Pfcmp sz r1' r2')
+
+ | PArith (PArithComparisonR0R (Pcmp is s) r1 r2) => OK (Asm.Pcmp is r1 r2 s)
+ | PArith (PArithComparisonR0R (Pcmn is s) r1 r2) => OK (Asm.Pcmn is r1 r2 s)
+ | PArith (PArithComparisonR0R (Ptst is s) r1 r2) => OK (Asm.Ptst is r1 r2 s)
+
+ | PArith (PArithComparisonP (Pcmpimm sz n) r1) => do r1' <- ireg_of_preg r1;
+ OK (Asm.Pcmpimm sz r1' n)
+ | PArith (PArithComparisonP (Pcmnimm sz n) r1) => do r1' <- ireg_of_preg r1;
+ OK (Asm.Pcmnimm sz r1' n)
+ | PArith (PArithComparisonP (Ptstimm sz n) r1) => do r1' <- ireg_of_preg r1;
+ OK (Asm.Ptstimm sz r1' n)
+ | PArith (PArithComparisonP (Pfcmp0 sz) r1) => do r1' <- freg_of_preg r1;
+ OK (Asm.Pfcmp0 sz r1')
+
+ | PArith (Pcset rd c) => OK (Asm.Pcset rd c)
+ | PArith (Pfmovi fsz rd r1) => OK (Asm.Pfmovi fsz rd r1)
+ | PArith (Pcsel rd r1 r2 c) =>
+ match r1, r2 with
+ | IR r1', IR r2' => do rd' <- ireg_of_preg rd;
+ do r1'' <- ireg_of_preg r1';
+ do r2'' <- ireg_of_preg r2';
+ OK (Asm.Pcsel rd' r1'' r2'' c)
+ | FR r1', FR r2' => do rd' <- freg_of_preg rd;
+ do r1'' <- freg_of_preg r1';
+ do r2'' <- freg_of_preg r2';
+ OK (Asm.Pfsel rd' r1'' r2'' c)
+ | _, _ => Error (msg "Asmgen.basic_to_instruction: Pcsel is only defind on iregs and fregs.")
+ end
+ | PArith (Pfnmul fsz rd r1 r2) => OK (Asm.Pfnmul fsz rd r1 r2)
+
+ | PLoad (PLd_rd_a Pldrw rd a) => do rd' <- ireg_of_preg rd; OK (Asm.Pldrw rd' a)
+ | PLoad (PLd_rd_a Pldrw_a rd a) => do rd' <- ireg_of_preg rd; OK (Asm.Pldrw_a rd' a)
+ | PLoad (PLd_rd_a Pldrx rd a) => do rd' <- ireg_of_preg rd; OK (Asm.Pldrx rd' a)
+ | PLoad (PLd_rd_a Pldrx_a rd a) => do rd' <- ireg_of_preg rd; OK (Asm.Pldrx_a rd' a)
+ | PLoad (PLd_rd_a (Pldrb sz) rd a) => do rd' <- ireg_of_preg rd; OK (Asm.Pldrb sz rd' a)
+ | PLoad (PLd_rd_a (Pldrsb sz) rd a) => do rd' <- ireg_of_preg rd; OK (Asm.Pldrsb sz rd' a)
+ | PLoad (PLd_rd_a (Pldrh sz) rd a) => do rd' <- ireg_of_preg rd; OK (Asm.Pldrh sz rd' a)
+ | PLoad (PLd_rd_a (Pldrsh sz) rd a) => do rd' <- ireg_of_preg rd; OK (Asm.Pldrsh sz rd' a)
+ | PLoad (PLd_rd_a Pldrzw rd a) => do rd' <- ireg_of_preg rd; OK (Asm.Pldrzw rd' a)
+ | PLoad (PLd_rd_a Pldrsw rd a) => do rd' <- ireg_of_preg rd; OK (Asm.Pldrsw rd' a)
+
+ | PLoad (PLd_rd_a Pldrs rd a) => do rd' <- freg_of_preg rd; OK (Asm.Pldrs rd' a)
+ | PLoad (PLd_rd_a Pldrd rd a) => do rd' <- freg_of_preg rd; OK (Asm.Pldrd rd' a)
+ | PLoad (PLd_rd_a Pldrd_a rd a) => do rd' <- freg_of_preg rd; OK (Asm.Pldrd_a rd' a)
+
+ | PLoad (Pldp Pldpw rd1 rd2 chk1 chk2 a) => do rd1' <- ireg_of_preg rd1;
+ do rd2' <- ireg_of_preg rd2;
+ OK (Asm.Pldpw rd1' rd2' chk1 chk2 a)
+ | PLoad (Pldp Pldpx rd1 rd2 chk1 chk2 a) => do rd1' <- ireg_of_preg rd1;
+ do rd2' <- ireg_of_preg rd2;
+ OK (Asm.Pldpx rd1' rd2' chk1 chk2 a)
+ | PLoad (Pldp Pldps rd1 rd2 chk1 chk2 a) => do rd1' <- freg_of_preg rd1;
+ do rd2' <- freg_of_preg rd2;
+ OK (Asm.Pldps rd1' rd2' chk1 chk2 a)
+ | PLoad (Pldp Pldpd rd1 rd2 chk1 chk2 a) => do rd1' <- freg_of_preg rd1;
+ do rd2' <- freg_of_preg rd2;
+ OK (Asm.Pldpd rd1' rd2' chk1 chk2 a)
+
+ | PStore (PSt_rs_a Pstrw r a) => do r' <- ireg_of_preg r; OK (Asm.Pstrw r' a)
+ | PStore (PSt_rs_a Pstrw_a r a) => do r' <- ireg_of_preg r; OK (Asm.Pstrw_a r' a)
+ | PStore (PSt_rs_a Pstrx r a) => do r' <- ireg_of_preg r; OK (Asm.Pstrx r' a)
+ | PStore (PSt_rs_a Pstrx_a r a) => do r' <- ireg_of_preg r; OK (Asm.Pstrx_a r' a)
+ | PStore (PSt_rs_a Pstrb r a) => do r' <- ireg_of_preg r; OK (Asm.Pstrb r' a)
+ | PStore (PSt_rs_a Pstrh r a) => do r' <- ireg_of_preg r; OK (Asm.Pstrh r' a)
+
+ | PStore (PSt_rs_a Pstrs r a) => do r' <- freg_of_preg r; OK (Asm.Pstrs r' a)
+ | PStore (PSt_rs_a Pstrd r a) => do r' <- freg_of_preg r; OK (Asm.Pstrd r' a)
+ | PStore (PSt_rs_a Pstrd_a r a) => do r' <- freg_of_preg r; OK (Asm.Pstrd_a r' a)
+
+ | PStore (Pstp Pstpw rs1 rs2 chk1 chk2 a) => do rs1' <- ireg_of_preg rs1;
+ do rs2' <- ireg_of_preg rs2;
+ OK (Asm.Pstpw rs1' rs2' chk1 chk2 a)
+ | PStore (Pstp Pstpx rs1 rs2 chk1 chk2 a) => do rs1' <- ireg_of_preg rs1;
+ do rs2' <- ireg_of_preg rs2;
+ OK (Asm.Pstpx rs1' rs2' chk1 chk2 a)
+ | PStore (Pstp Pstps rs1 rs2 chk1 chk2 a) => do rs1' <- freg_of_preg rs1;
+ do rs2' <- freg_of_preg rs2;
+ OK (Asm.Pstps rs1' rs2' chk1 chk2 a)
+ | PStore (Pstp Pstpd rs1 rs2 chk1 chk2 a) => do rs1' <- freg_of_preg rs1;
+ do rs2' <- freg_of_preg rs2;
+ OK (Asm.Pstpd rs1' rs2' chk1 chk2 a)
+
+ | Pallocframe sz linkofs => OK (Asm.Pallocframe sz linkofs)
+ | Pfreeframe sz linkofs => OK (Asm.Pfreeframe sz linkofs)
+
+ | Ploadsymbol rd id => OK (Asm.Ploadsymbol rd id)
+
+ | Pcvtsw2x rd r1 => OK (Asm.Pcvtsw2x rd r1)
+
+ | Pcvtuw2x rd r1 => OK (Asm.Pcvtuw2x rd r1)
+
+ | Pcvtx2w rd => OK (Asm.Pcvtx2w rd)
+ | Pnop => OK (Asm.Pnop)
+ end.
+
+<<<<<<< HEAD
+Definition cf_instruction_to_instruction (cfi: cf_instruction) : Asm.instruction :=
+ match cfi with
+ | Pb l => Asm.Pb l
+ | Pbc c lbl => Asm.Pbc c lbl
+ | Pbl id sg => Asm.Pbl id sg
+ | Pbs id sg => Asm.Pbs id sg
+ | Pblr r sg => Asm.Pblr r sg
+ | Pbr r sg => Asm.Pbr r sg
+ | Pret r => Asm.Pret r
+ | Pcbnz sz r lbl => Asm.Pcbnz sz r lbl
+ | Pcbz sz r lbl => Asm.Pcbz sz r lbl
+ | Ptbnz sz r n lbl => Asm.Ptbnz sz r n lbl
+ | Ptbz sz r n lbl => Asm.Ptbz sz r n lbl
+ | Pbtbl r1 tbl => Asm.Pbtbl r1 tbl
+=======
+(** Translation of addressing modes *)
+
+Definition offset_representable (sz: Z) (ofs: int64) : bool :=
+ let isz := Int64.repr sz in
+ (** either unscaled 9-bit signed *)
+ Int64.eq ofs (Int64.sign_ext 9 ofs) ||
+ (** or scaled 12-bit unsigned *)
+ (Int64.eq (Int64.modu ofs isz) Int64.zero
+ && Int64.ltu ofs (Int64.shl isz (Int64.repr 12))).
+
+Definition transl_addressing (sz: Z) (addr: Op.addressing) (args: list mreg)
+ (insn: Asm.addressing -> instruction) (k: code) : res code :=
+ match addr, args with
+ | Aindexed ofs, a1 :: nil =>
+ do r1 <- ireg_of a1;
+ if offset_representable sz ofs then
+ OK (insn (ADimm r1 ofs) :: k)
+ else
+ OK (loadimm64 X16 ofs (insn (ADreg r1 X16) :: k))
+ | Aindexed2, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (insn (ADreg r1 r2) :: k)
+ | Aindexed2shift a, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ if Int.eq a Int.zero then
+ OK (insn (ADreg r1 r2) :: k)
+ else if Int.eq (Int.shl Int.one a) (Int.repr sz) then
+ OK (insn (ADlsl r1 r2 a) :: k)
+ else
+ OK (Padd X X16 r1 r2 (SOlsl a) :: insn (ADimm X16 Int64.zero) :: k)
+ | Aindexed2ext x a, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ if Int.eq a Int.zero || Int.eq (Int.shl Int.one a) (Int.repr sz) then
+ OK (insn (match x with Xsgn32 => ADsxt r1 r2 a
+ | Xuns32 => ADuxt r1 r2 a end) :: k)
+ else
+ OK (arith_extended Paddext (Padd X) X16 r1 r2 x a
+ (insn (ADimm X16 Int64.zero) :: k))
+ | Aglobal id ofs, nil =>
+ assertion (negb (SelectOp.symbol_is_relocatable id));
+ if Ptrofs.eq (Ptrofs.modu ofs (Ptrofs.repr sz)) Ptrofs.zero && symbol_is_aligned id sz
+ then OK (Padrp X16 id ofs :: insn (ADadr X16 id ofs) :: k)
+ else OK (loadsymbol X16 id ofs (insn (ADimm X16 Int64.zero) :: k))
+ | Ainstack ofs, nil =>
+ let ofs := Ptrofs.to_int64 ofs in
+ if offset_representable sz ofs then
+ OK (insn (ADimm XSP ofs) :: k)
+ else
+ OK (loadimm64 X16 ofs (insn (ADreg XSP X16) :: k))
+ | _, _ =>
+ Error(msg "Asmgen.transl_addressing")
+>>>>>>> master
+ end.
+
+Definition control_to_instruction (c: control) :=
+ match c with
+ | PCtlFlow i => cf_instruction_to_instruction i
+ | Pbuiltin ef args res => Asm.Pbuiltin ef (List.map (map_builtin_arg DR) args) (map_builtin_res DR res)
+ end.
+
+Fixpoint unfold_label (ll: list label) :=
+ match ll with
+ | nil => nil
+ | l :: ll => Plabel l :: unfold_label ll
+ end.
+
+Fixpoint unfold_body (lb: list basic) : res Asm.code :=
+ match lb with
+ | nil => OK nil
+ | b :: lb =>
+ (* x_is: x's instructions *)
+ do b_is <- basic_to_instruction b;
+ do lb_is <- unfold_body lb;
+ OK (b_is :: lb_is)
+ end.
+
+Definition unfold_exit (oc: option control) :=
+ match oc with
+ | None => nil
+ | Some c => control_to_instruction c :: nil
+ end.
+
+Definition unfold_bblock (bb: bblock) :=
+ let lbl := unfold_label (header bb) in
+ (*
+ * With this dynamically checked assumption on a previous optimization we
+ * can show that [Asmblock.label_pos] and [Asm.label_pos] retrieve the same
+ * exact address. Maintaining this property allows us to use the simple
+ * formulation of match_states defined as equality.
+ * Otherwise we would have to deal with the case of a basic block header
+ * that has multiple labels. Asmblock.label_pos will, for all labels, point
+ * to the same location at the beginning of the basic block. Asm.label_pos
+ * on the other hand could return a position pointing into the original
+ * basic block.
+ *)
+ if zle (list_length_z (header bb)) 1 then
+ do bo_is <- unfold_body (body bb);
+ OK (lbl ++ bo_is ++ unfold_exit (exit bb))
+ else
+ Error (msg "Asmgen.unfold_bblock: Multiple labels were generated.").
+
+Fixpoint unfold (bbs: Asmblock.bblocks) : res Asm.code :=
+ match bbs with
+ | nil => OK (nil)
+ | bb :: bbs' =>
+ do bb_is <- unfold_bblock bb;
+ do bbs'_is <- unfold bbs';
+ OK (bb_is ++ bbs'_is)
+ end.
+
+Definition transf_function (f: Asmblock.function) : res Asm.function :=
+ do c <- unfold (Asmblock.fn_blocks f);
+ if zlt Ptrofs.max_unsigned (list_length_z c)
+ then Error (msg "Asmgen.trans_function: code size exceeded")
+ else OK {| Asm.fn_sig := Asmblock.fn_sig f; Asm.fn_code := c |}.
+
+Definition transf_fundef (f: Asmblock.fundef) : res Asm.fundef :=
+ transf_partial_fundef transf_function f.
+
+Definition transf_program (p: Asmblock.program) : res Asm.program :=
+ transform_partial_program transf_fundef p.
+
+End Asmblock_TRANSF.
+
+Definition transf_program (p: Mach.program) : res Asm.program :=
+ let mbp := Machblockgen.transf_program p in
+ do abp <- Asmblockgen.transf_program mbp;
+ do abp' <- (time "PostpassScheduling total oracle+verification" PostpassScheduling.transf_program) abp;
+ Asmblock_TRANSF.transf_program abp'.
diff --git a/aarch64/TO_MERGE/Asmgenproof.v b/aarch64/TO_MERGE/Asmgenproof.v
new file mode 100644
index 00000000..8af013fd
--- /dev/null
+++ b/aarch64/TO_MERGE/Asmgenproof.v
@@ -0,0 +1,2787 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Sylvain Boulmé Grenoble-INP, VERIMAG *)
+(* Léo Gourdin UGA, VERIMAG *)
+(* Justus Fasse UGA, VERIMAG *)
+(* Xavier Leroy INRIA Paris-Rocquencourt *)
+(* David Monniaux CNRS, VERIMAG *)
+(* Cyril Six Kalray *)
+(* *)
+(* Copyright Kalray. Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+Require Import Coqlib Errors.
+Require Import Integers Floats AST Linking.
+Require Import Values Memory Events Globalenvs Smallstep.
+Require Import Op Locations Machblock Conventions Asm Asmblock.
+Require Machblockgenproof Asmblockgenproof PostpassSchedulingproof.
+Require Import Asmgen.
+Require Import Axioms.
+Require Import IterList.
+Require Import Ring Lia.
+
+Module Asmblock_PRESERVATION.
+
+Import Asmblock_TRANSF.
+
+Definition match_prog (p: Asmblock.program) (tp: Asm.program) :=
+ match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp.
+
+Lemma transf_program_match:
+ forall p tp, transf_program p = OK tp -> match_prog p tp.
+Proof.
+ intros. eapply match_transform_partial_program; eauto.
+Qed.
+
+Section PRESERVATION.
+
+Variable prog: Asmblock.program.
+Variable tprog: Asm.program.
+Hypothesis TRANSF: match_prog prog tprog.
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+Definition lk :aarch64_linker := {| Asmblock.symbol_low:=Asm.symbol_low tge; Asmblock.symbol_high:=Asm.symbol_high tge|}.
+
+Lemma symbols_preserved:
+ forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
+Proof (Genv.find_symbol_match TRANSF).
+
+Lemma symbol_addresses_preserved:
+ forall (s: ident) (ofs: ptrofs),
+ Genv.symbol_address tge s ofs = Genv.symbol_address ge s ofs.
+Proof.
+ intros; unfold Genv.symbol_address; rewrite symbols_preserved; reflexivity.
+Qed.
+
+Lemma senv_preserved:
+ Senv.equiv ge tge.
+Proof (Genv.senv_match TRANSF).
+
+Lemma symbol_high_low: forall (id: ident) (ofs: ptrofs),
+ Val.addl (Asmblock.symbol_high lk id ofs) (Asmblock.symbol_low lk id ofs) = Genv.symbol_address ge id ofs.
+Proof.
+ unfold lk; simpl. intros; rewrite Asm.symbol_high_low; unfold Genv.symbol_address;
+ rewrite symbols_preserved; reflexivity.
+Qed.
+
+Lemma functions_translated:
+ forall b f,
+ Genv.find_funct_ptr ge b = Some f ->
+ exists tf,
+ Genv.find_funct_ptr tge b = Some tf /\ transf_fundef f = OK tf.
+Proof (Genv.find_funct_ptr_transf_partial TRANSF).
+
+Lemma internal_functions_translated:
+ forall b f,
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
+ exists tf,
+ Genv.find_funct_ptr tge b = Some (Internal tf) /\ transf_function f = OK tf.
+Proof.
+ intros; exploit functions_translated; eauto.
+ intros (x & FIND & TRANSf).
+ apply bind_inversion in TRANSf.
+ destruct TRANSf as (tf & TRANSf & X).
+ inv X.
+ eauto.
+Qed.
+
+Lemma internal_functions_unfold:
+ forall b f,
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
+ exists tc,
+ Genv.find_funct_ptr tge b = Some (Internal (Asm.mkfunction (fn_sig f) tc))
+ /\ unfold (fn_blocks f) = OK tc
+ /\ list_length_z tc <= Ptrofs.max_unsigned.
+Proof.
+ intros.
+ exploit internal_functions_translated; eauto.
+ intros (tf & FINDtf & TRANStf).
+ unfold transf_function in TRANStf.
+ monadInv TRANStf.
+ destruct (zlt _ _); try congruence.
+ inv EQ. inv EQ0.
+ eexists; intuition eauto.
+ lia.
+Qed.
+
+
+Inductive is_nth_inst (bb: bblock) (n:Z) (i:Asm.instruction): Prop :=
+ | is_nth_label l:
+ list_nth_z (header bb) n = Some l ->
+ i = Asm.Plabel l ->
+ is_nth_inst bb n i
+ | is_nth_basic bi:
+ list_nth_z (body bb) (n - list_length_z (header bb)) = Some bi ->
+ basic_to_instruction bi = OK i ->
+ is_nth_inst bb n i
+ | is_nth_ctlflow cfi:
+ (exit bb) = Some cfi ->
+ n = size bb - 1 ->
+ i = control_to_instruction cfi ->
+ is_nth_inst bb n i.
+
+(* Asmblock and Asm share the same definition of state *)
+Definition match_states (s1 s2 : state) := s1 = s2.
+
+Inductive match_internal: forall n, state -> state -> Prop :=
+ | match_internal_intro n rs1 m1 rs2 m2
+ (MEM: m1 = m2)
+ (AG: forall r, r <> PC -> rs1 r = rs2 r)
+ (AGPC: Val.offset_ptr (rs1 PC) (Ptrofs.repr n) = rs2 PC)
+ : match_internal n (State rs1 m1) (State rs2 m2).
+
+Lemma match_internal_set_parallel:
+ forall n rs1 m1 rs2 m2 r val,
+ match_internal n (State rs1 m1) (State rs2 m2) ->
+ r <> PC ->
+ match_internal n (State (rs1#r <- val) m1) (State (rs2#r <- val ) m2).
+Proof.
+ intros n rs1 m1 rs2 m2 r v MI.
+ inversion MI; constructor; auto.
+ - intros r' NOTPC.
+ unfold Pregmap.set; rewrite AG. reflexivity. assumption.
+ - unfold Pregmap.set; destruct (PregEq.eq PC r); congruence.
+Qed.
+
+Lemma agree_match_states:
+ forall rs1 m1 rs2 m2,
+ match_states (State rs1 m1) (State rs2 m2) ->
+ forall r : preg, rs1#r = rs2#r.
+Proof.
+ intros.
+ unfold match_states in *.
+ assert (rs1 = rs2) as EQ. { congruence. }
+ rewrite EQ. reflexivity.
+Qed.
+
+Lemma match_states_set_parallel:
+ forall rs1 m1 rs2 m2 r v,
+ match_states (State rs1 m1) (State rs2 m2) ->
+ match_states (State (rs1#r <- v) m1) (State (rs2#r <- v) m2).
+Proof.
+ intros; unfold match_states in *.
+ assert (rs1 = rs2) as RSEQ. { congruence. }
+ assert (m1 = m2) as MEQ. { congruence. }
+ rewrite RSEQ in *; rewrite MEQ in *; unfold Pregmap.set; reflexivity.
+Qed.
+
+(* match_internal from match_states *)
+Lemma mi_from_ms:
+ forall rs1 m1 rs2 m2 b ofs,
+ match_states (State rs1 m1) (State rs2 m2) ->
+ rs1#PC = Vptr b ofs ->
+ match_internal 0 (State rs1 m1) (State rs2 m2).
+Proof.
+ intros rs1 m1 rs2 m2 b ofs MS PCVAL.
+ inv MS; constructor; auto; unfold Val.offset_ptr;
+ rewrite PCVAL; rewrite Ptrofs.add_zero; reflexivity.
+Qed.
+
+Lemma transf_initial_states:
+ forall s1, Asmblock.initial_state prog s1 ->
+ exists s2, Asm.initial_state tprog s2 /\ match_states s1 s2.
+Proof.
+ intros ? INIT_s1.
+ inversion INIT_s1 as (m, ?, ge0, rs). unfold ge0 in *.
+ econstructor; split.
+ - econstructor.
+ eapply (Genv.init_mem_transf_partial TRANSF); eauto.
+ - rewrite (match_program_main TRANSF); rewrite symbol_addresses_preserved.
+ unfold rs; reflexivity.
+Qed.
+
+Lemma transf_final_states:
+ forall s1 s2 r,
+ match_states s1 s2 -> Asmblock.final_state s1 r -> Asm.final_state s2 r.
+Proof.
+ intros s1 s2 r MATCH FINAL_s1.
+ inv FINAL_s1; inv MATCH; constructor; assumption.
+Qed.
+
+Definition max_pos (f : Asm.function) := list_length_z f.(Asm.fn_code).
+
+Lemma functions_bound_max_pos: forall fb f tf,
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ transf_function f = OK tf ->
+<<<<<<< HEAD
+ max_pos tf <= Ptrofs.max_unsigned.
+Proof.
+ intros fb f tf FINDf TRANSf.
+ unfold transf_function in TRANSf.
+ apply bind_inversion in TRANSf.
+ destruct TRANSf as (c & TRANSf).
+ destruct TRANSf as (_ & TRANSf).
+ destruct (zlt _ _).
+ - inversion TRANSf.
+ - unfold max_pos.
+ assert (Asm.fn_code tf = c) as H. { inversion TRANSf as (H'); auto. }
+ rewrite H; lia.
+Qed.
+=======
+ Genv.find_funct_ptr tge fb = Some (Internal tf).
+Proof.
+ intros. exploit functions_translated; eauto. intros [tf' [A B]].
+ monadInv B. rewrite H0 in EQ; inv EQ; auto.
+Qed.
+
+(** * Properties of control flow *)
+
+Lemma transf_function_no_overflow:
+ forall f tf,
+ transf_function f = OK tf -> list_length_z tf.(fn_code) <= Ptrofs.max_unsigned.
+Proof.
+ intros. monadInv H. destruct (zlt Ptrofs.max_unsigned (list_length_z x.(fn_code))); inv EQ0.
+ lia.
+Qed.
+
+Lemma exec_straight_exec:
+ forall fb f c ep tf tc c' rs m rs' m',
+ transl_code_at_pc ge (rs PC) fb f c ep tf tc ->
+ exec_straight tge tf tc rs m c' rs' m' ->
+ plus step tge (State rs m) E0 (State rs' m').
+Proof.
+ intros. inv H.
+ eapply exec_straight_steps_1; eauto.
+ eapply transf_function_no_overflow; eauto.
+ eapply functions_transl; eauto.
+Qed.
+
+Lemma exec_straight_at:
+ forall fb f c ep tf tc c' ep' tc' rs m rs' m',
+ transl_code_at_pc ge (rs PC) fb f c ep tf tc ->
+ transl_code f c' ep' = OK tc' ->
+ exec_straight tge tf tc rs m tc' rs' m' ->
+ transl_code_at_pc ge (rs' PC) fb f c' ep' tf tc'.
+Proof.
+ intros. inv H.
+ exploit exec_straight_steps_2; eauto.
+ eapply transf_function_no_overflow; eauto.
+ eapply functions_transl; eauto.
+ intros [ofs' [PC' CT']].
+ rewrite PC'. constructor; auto.
+Qed.
+
+(** The following lemmas show that the translation from Mach to Asm
+ preserves labels, in the sense that the following diagram commutes:
+<<
+ translation
+ Mach code ------------------------ Asm instr sequence
+ | |
+ | Mach.find_label lbl find_label lbl |
+ | |
+ v v
+ Mach code tail ------------------- Asm instr seq tail
+ translation
+>>
+ The proof demands many boring lemmas showing that Asm constructor
+ functions do not introduce new labels.
+*)
+>>>>>>> master
+
+Lemma one_le_max_unsigned:
+ 1 <= Ptrofs.max_unsigned.
+Proof.
+ unfold Ptrofs.max_unsigned; simpl; unfold Ptrofs.wordsize;
+ unfold Wordsize_Ptrofs.wordsize; destruct Archi.ptr64; simpl; lia.
+Qed.
+
+(* NB: does not seem useful anymore, with the [exec_header_simulation] proof below
+Lemma match_internal_exec_label:
+ forall n rs1 m1 rs2 m2 l fb f tf,
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ transf_function f = OK tf ->
+ match_internal n (State rs1 m1) (State rs2 m2) ->
+ n >= 0 ->
+ (* There is no step if n is already max_pos *)
+ n < (max_pos tf) ->
+ exists rs2' m2', Asm.exec_instr tge tf (Asm.Plabel l) rs2 m2 = Next rs2' m2'
+ /\ match_internal (n+1) (State rs1 m1) (State rs2' m2').
+Proof.
+ intros. (* XXX auto generated names *)
+ unfold Asm.exec_instr.
+ eexists; eexists; split; eauto.
+ inversion H1; constructor; auto.
+ - intros; unfold Asm.nextinstr; unfold Pregmap.set;
+ destruct (PregEq.eq r PC); auto; contradiction.
+ - unfold Asm.nextinstr; rewrite Pregmap.gss; unfold Ptrofs.one.
+ rewrite <- AGPC; rewrite Val.offset_ptr_assoc; unfold Ptrofs.add;
+ rewrite Ptrofs.unsigned_repr. rewrite Ptrofs.unsigned_repr; trivial.
+ + split.
+ * apply Z.le_0_1.
+ * apply one_le_max_unsigned.
+ + split.
+ * apply Z.ge_le; assumption.
+ * rewrite <- functions_bound_max_pos; eauto; lia.
+Qed.
+*)
+
+Lemma incrPC_agree_but_pc:
+ forall rs r ofs,
+ r <> PC ->
+ (incrPC ofs rs)#r = rs#r.
+Proof.
+ intros rs r ofs NOTPC.
+ unfold incrPC; unfold Pregmap.set; destruct (PregEq.eq r PC).
+ - contradiction.
+ - reflexivity.
+Qed.
+
+Lemma bblock_non_empty bb: body bb <> nil \/ exit bb <> None.
+Proof.
+ destruct bb. simpl.
+ unfold non_empty_bblockb in correct.
+ unfold non_empty_body, non_empty_exit, Is_true in correct.
+ destruct body, exit.
+ - right. discriminate.
+ - contradiction.
+ - right. discriminate.
+ - left. discriminate.
+Qed.
+
+Lemma list_length_z_aux_increase A (l: list A): forall acc,
+ list_length_z_aux l acc >= acc.
+Proof.
+ induction l; simpl; intros.
+ - lia.
+ - generalize (IHl (Z.succ acc)). lia.
+Qed.
+
+Lemma bblock_size_aux_pos bb: list_length_z (body bb) + Z.of_nat (length_opt (exit bb)) >= 1.
+Proof.
+ destruct (bblock_non_empty bb), (body bb) as [|hd tl], (exit bb); simpl;
+ try (congruence || lia);
+ unfold list_length_z; simpl;
+ generalize (list_length_z_aux_increase _ tl 1); lia.
+Qed.
+
+
+Lemma list_length_add_acc A (l : list A) acc:
+ list_length_z_aux l acc = (list_length_z l) + acc.
+Proof.
+ unfold list_length_z, list_length_z_aux. simpl.
+ fold list_length_z_aux.
+ rewrite (list_length_z_aux_shift l acc 0).
+ lia.
+Qed.
+
+Lemma list_length_z_cons A hd (tl : list A):
+ list_length_z (hd :: tl) = list_length_z tl + 1.
+Proof.
+ unfold list_length_z; simpl; rewrite list_length_add_acc; reflexivity.
+Qed.
+
+Lemma bblock_size_aux bb: size bb = list_length_z (header bb) + list_length_z (body bb) + Z.of_nat (length_opt (exit bb)).
+Proof.
+ unfold size.
+ repeat (rewrite list_length_z_nat). repeat (rewrite Nat2Z.inj_add). reflexivity.
+Qed.
+
+Lemma header_size_lt_block_size bb:
+ list_length_z (header bb) < size bb.
+Proof.
+ rewrite bblock_size_aux.
+ generalize (bblock_non_empty bb); intros NEMPTY; destruct NEMPTY as [HDR|EXIT].
+ - destruct (body bb); try contradiction; rewrite list_length_z_cons;
+ repeat rewrite list_length_z_nat; lia.
+ - destruct (exit bb); try contradiction; simpl; repeat rewrite list_length_z_nat; lia.
+Qed.
+
+Lemma body_size_le_block_size bb:
+ list_length_z (body bb) <= size bb.
+Proof.
+ rewrite bblock_size_aux; repeat rewrite list_length_z_nat; lia.
+Qed.
+
+
+Lemma bblock_size_pos bb: size bb >= 1.
+Proof.
+ rewrite (bblock_size_aux bb).
+ generalize (bblock_size_aux_pos bb).
+ generalize (list_length_z_pos (header bb)).
+ lia.
+Qed.
+
+Lemma unfold_car_cdr bb bbs tc:
+ unfold (bb :: bbs) = OK tc ->
+ exists tbb tc', unfold_bblock bb = OK tbb
+ /\ unfold bbs = OK tc'
+ /\ unfold (bb :: bbs) = OK (tbb ++ tc').
+Proof.
+ intros UNFOLD.
+ assert (UF := UNFOLD).
+ unfold unfold in UNFOLD.
+ apply bind_inversion in UNFOLD. destruct UNFOLD as (? & UBB). destruct UBB as (UBB & REST).
+ apply bind_inversion in REST. destruct REST as (? & UNFOLD').
+ fold unfold in UNFOLD'. destruct UNFOLD' as (UNFOLD' & UNFOLD).
+ rewrite <- UNFOLD in UF.
+ eauto.
+Qed.
+
+Lemma unfold_cdr bb bbs tc:
+ unfold (bb :: bbs) = OK tc ->
+ exists tc', unfold bbs = OK tc'.
+Proof.
+<<<<<<< HEAD
+ intros; exploit unfold_car_cdr; eauto. intros (_ & ? & _ & ? & _).
+ eexists; eauto.
+Qed.
+=======
+ intros; unfold loadsymbol.
+ destruct (SelectOp.symbol_is_relocatable id); TailNoLabel. destruct Ptrofs.eq; TailNoLabel.
+Qed.
+Hint Resolve loadsymbol_label: labels.
+>>>>>>> master
+
+Lemma unfold_car bb bbs tc:
+ unfold (bb :: bbs) = OK tc ->
+ exists tbb, unfold_bblock bb = OK tbb.
+Proof.
+ intros; exploit unfold_car_cdr; eauto. intros (? & _ & ? & _ & _).
+ eexists; eauto.
+Qed.
+
+Lemma all_blocks_translated:
+ forall bbs tc,
+ unfold bbs = OK tc ->
+ forall bb, In bb bbs ->
+ exists c, unfold_bblock bb = OK c.
+Proof.
+ induction bbs as [| bb bbs IHbbs].
+ - contradiction.
+ - intros ? UNFOLD ? IN.
+ (* unfold proceeds by unfolding the basic block at the head of the list and
+ * then recurring *)
+ exploit unfold_car_cdr; eauto. intros (? & ? & ? & ? & _).
+ (* basic block is either in head or tail *)
+ inversion IN as [EQ | NEQ].
+ + rewrite <- EQ; eexists; eauto.
+ + eapply IHbbs; eauto.
+Qed.
+
+Lemma entire_body_translated:
+ forall lbi tc,
+ unfold_body lbi = OK tc ->
+ forall bi, In bi lbi ->
+ exists bi', basic_to_instruction bi = OK bi'.
+Proof.
+ induction lbi as [| a lbi IHlbi].
+ - intros. contradiction.
+ - intros tc UNFOLD_BODY bi IN.
+ unfold unfold_body in UNFOLD_BODY. apply bind_inversion in UNFOLD_BODY.
+ destruct UNFOLD_BODY as (? & TRANSbi & REST).
+ apply bind_inversion in REST. destruct REST as (? & UNFOLD_BODY' & ?).
+ fold unfold_body in UNFOLD_BODY'.
+
+ inversion IN as [EQ | NEQ].
+ + rewrite <- EQ; eauto.
+ + eapply IHlbi; eauto.
+Qed.
+
+Lemma bblock_in_bblocks bbs bb: forall
+ tc pos
+ (UNFOLD: unfold bbs = OK tc)
+ (FINDBB: find_bblock pos bbs = Some bb),
+ In bb bbs.
+Proof.
+ induction bbs as [| b bbs IH].
+ - intros. inversion FINDBB.
+ - destruct pos.
+ + intros. inversion FINDBB as (EQ). rewrite <- EQ. apply in_eq.
+ + intros.
+ exploit unfold_cdr; eauto. intros (tc' & UNFOLD').
+ unfold find_bblock in FINDBB. simpl in FINDBB.
+ fold find_bblock in FINDBB.
+ apply in_cons. eapply IH; eauto.
+ + intros. inversion FINDBB.
+Qed.
+
+Lemma blocks_translated tc pos bbs bb: forall
+ (UNFOLD: unfold bbs = OK tc)
+ (FINDBB: find_bblock pos bbs = Some bb),
+ exists tbb, unfold_bblock bb = OK tbb.
+Proof.
+ intros; exploit bblock_in_bblocks; eauto; intros;
+ eapply all_blocks_translated; eauto.
+Qed.
+
+Lemma size_header b pos f bb: forall
+ (FINDF: Genv.find_funct_ptr ge b = Some (Internal f))
+ (FINDBB: find_bblock pos (fn_blocks f) = Some bb),
+ list_length_z (header bb) <= 1.
+Proof.
+ intros.
+ exploit internal_functions_unfold; eauto.
+ intros (tc & FINDtf & TRANStf & ?).
+ exploit blocks_translated; eauto. intros TBB.
+
+ unfold unfold_bblock in TBB.
+ destruct (zle (list_length_z (header bb)) 1).
+ - assumption.
+ - destruct TBB as (? & TBB). discriminate TBB.
+Qed.
+
+Lemma list_nth_z_neg A (l: list A): forall n,
+ n < 0 -> list_nth_z l n = None.
+Proof.
+ induction l; simpl; auto.
+ intros n H; destruct (zeq _ _); (try eapply IHl); lia.
+Qed.
+
+Lemma find_bblock_neg bbs: forall pos,
+ pos < 0 -> find_bblock pos bbs = None.
+Proof.
+ induction bbs; simpl; auto.
+ intros. destruct (zlt pos 0). { reflexivity. }
+ destruct (zeq pos 0); contradiction.
+Qed.
+
+Lemma equal_header_size bb:
+ length (header bb) = length (unfold_label (header bb)).
+Proof.
+ induction (header bb); auto.
+ simpl. rewrite IHl. auto.
+Qed.
+
+Lemma equal_body_size:
+ forall bb tb,
+ unfold_body (body bb) = OK tb ->
+ length (body bb) = length tb.
+Proof.
+ intros bb. induction (body bb).
+ - simpl. intros ? H. inversion H. auto.
+ - intros tb H. simpl in H. apply bind_inversion in H. destruct H as (? & BI & TAIL).
+ apply bind_inversion in TAIL. destruct TAIL as (tb' & BODY' & CONS). inv CONS.
+ simpl. specialize (IHl tb' BODY'). rewrite IHl. reflexivity.
+Qed.
+
+Lemma equal_exit_size bb:
+ length_opt (exit bb) = length (unfold_exit (exit bb)).
+Proof.
+ destruct (exit bb); trivial.
+Qed.
+
+Lemma bblock_size_preserved bb tb:
+ unfold_bblock bb = OK tb ->
+ size bb = list_length_z tb.
+Proof.
+ unfold unfold_bblock. intros UNFOLD_BBLOCK.
+ destruct (zle (list_length_z (header bb)) 1). 2: { inversion UNFOLD_BBLOCK. }
+ apply bind_inversion in UNFOLD_BBLOCK. destruct UNFOLD_BBLOCK as (? & UNFOLD_BODY & CONS).
+ inversion CONS.
+ unfold size.
+ rewrite equal_header_size, equal_exit_size.
+ erewrite equal_body_size; eauto.
+ rewrite list_length_z_nat.
+ repeat (rewrite app_length).
+ rewrite plus_assoc. auto.
+Qed.
+
+Lemma size_of_blocks_max_pos_aux:
+ forall bbs tbbs pos bb,
+ find_bblock pos bbs = Some bb ->
+ unfold bbs = OK tbbs ->
+ pos + size bb <= list_length_z tbbs.
+Proof.
+ induction bbs as [| bb ? IHbbs].
+ - intros tbbs ? ? FINDBB; inversion FINDBB.
+ - simpl; intros tbbs pos bb' FINDBB UNFOLD.
+ apply bind_inversion in UNFOLD; destruct UNFOLD as (tbb & UNFOLD_BBLOCK & H).
+ apply bind_inversion in H; destruct H as (tbbs' & UNFOLD & CONS).
+ inv CONS.
+ destruct (zlt pos 0). { discriminate FINDBB. }
+ destruct (zeq pos 0).
+ + inv FINDBB.
+ exploit bblock_size_preserved; eauto; intros SIZE; rewrite SIZE.
+ repeat (rewrite list_length_z_nat). rewrite app_length, Nat2Z.inj_add.
+ lia.
+ + generalize (IHbbs tbbs' (pos - size bb) bb' FINDBB UNFOLD). intros IH.
+ exploit bblock_size_preserved; eauto; intros SIZE.
+ repeat (rewrite list_length_z_nat); rewrite app_length.
+ rewrite Nat2Z.inj_add; repeat (rewrite <- list_length_z_nat).
+ lia.
+Qed.
+
+Lemma size_of_blocks_max_pos pos f tf bi:
+ find_bblock pos (fn_blocks f) = Some bi ->
+ transf_function f = OK tf ->
+ pos + size bi <= max_pos tf.
+Proof.
+ unfold transf_function, max_pos.
+ intros FINDBB UNFOLD.
+ apply bind_inversion in UNFOLD. destruct UNFOLD as (? & UNFOLD & H).
+ destruct (zlt Ptrofs.max_unsigned (list_length_z x)). { discriminate H. }
+ inv H. simpl.
+ eapply size_of_blocks_max_pos_aux; eauto.
+Qed.
+
+Lemma unfold_bblock_not_nil bb:
+ unfold_bblock bb = OK nil -> False.
+Proof.
+ intros.
+ exploit bblock_size_preserved; eauto. unfold list_length_z; simpl. intros SIZE.
+ generalize (bblock_size_pos bb). intros SIZE'. lia.
+Qed.
+
+(* same proof as list_nth_z_range (Coqlib) *)
+Lemma find_instr_range:
+ forall c n i,
+ Asm.find_instr n c = Some i -> 0 <= n < list_length_z c.
+Proof.
+ induction c; simpl; intros.
+ discriminate.
+ rewrite list_length_z_cons. destruct (zeq n 0).
+ generalize (list_length_z_pos c); lia.
+ exploit IHc; eauto. lia.
+Qed.
+
+Lemma find_instr_tail:
+ forall tbb pos c i,
+ Asm.find_instr pos c = Some i ->
+ Asm.find_instr (pos + list_length_z tbb) (tbb ++ c) = Some i.
+Proof.
+ induction tbb as [| ? ? IHtbb].
+ - intros. unfold list_length_z; simpl. rewrite Z.add_0_r. assumption.
+ - intros. rewrite list_length_z_cons. simpl.
+ destruct (zeq (pos + (list_length_z tbb + 1)) 0).
+ + exploit find_instr_range; eauto. intros POS_RANGE.
+ generalize (list_length_z_pos tbb). lia.
+ + replace (pos + (list_length_z tbb + 1) - 1) with (pos + list_length_z tbb) by lia.
+ eapply IHtbb; eauto.
+Qed.
+
+Lemma size_of_blocks_bounds fb pos f bi:
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ find_bblock pos (fn_blocks f) = Some bi ->
+ pos + size bi <= Ptrofs.max_unsigned.
+Proof.
+ intros; exploit internal_functions_translated; eauto.
+ intros (tf & _ & TRANSf).
+ assert (pos + size bi <= max_pos tf). { eapply size_of_blocks_max_pos; eauto. }
+ assert (max_pos tf <= Ptrofs.max_unsigned). { eapply functions_bound_max_pos; eauto. }
+ lia.
+Qed.
+
+Lemma find_instr_bblock_tail:
+ forall tbb bb pos c i,
+ Asm.find_instr pos c = Some i ->
+ unfold_bblock bb = OK tbb ->
+ Asm.find_instr (pos + size bb ) (tbb ++ c) = Some i.
+Proof.
+ induction tbb.
+ - intros. exploit unfold_bblock_not_nil; eauto. intros. contradiction.
+ - intros. simpl.
+ destruct (zeq (pos + size bb) 0).
+ + (* absurd *)
+ exploit find_instr_range; eauto. intros POS_RANGE.
+ generalize (bblock_size_pos bb). intros SIZE. lia.
+ + erewrite bblock_size_preserved; eauto.
+ rewrite list_length_z_cons.
+ replace (pos + (list_length_z tbb + 1) - 1) with (pos + list_length_z tbb) by lia.
+ apply find_instr_tail; auto.
+Qed.
+
+Lemma list_nth_z_find_label:
+ forall (ll : list label) il n l,
+ list_nth_z ll n = Some l ->
+ Asm.find_instr n ((unfold_label ll) ++ il) = Some (Asm.Plabel l).
+Proof.
+ induction ll.
+ - intros. inversion H.
+ - intros. simpl.
+ destruct (zeq n 0) as [Z | NZ].
+ + inversion H as (H'). rewrite Z in H'. simpl in H'. inv H'. reflexivity.
+ + simpl in H. destruct (zeq n 0). { contradiction. }
+ apply IHll; auto.
+Qed.
+
+Lemma list_nth_z_find_bi:
+ forall lbi bi tlbi n bi' exit,
+ list_nth_z lbi n = Some bi ->
+ unfold_body lbi = OK tlbi ->
+ basic_to_instruction bi = OK bi' ->
+ Asm.find_instr n (tlbi ++ exit) = Some bi'.
+Proof.
+ induction lbi.
+ - intros. inversion H.
+ - simpl. intros.
+ apply bind_inversion in H0. destruct H0 as (? & ? & ?).
+ apply bind_inversion in H2. destruct H2 as (? & ? & ?).
+ destruct (zeq n 0) as [Z | NZ].
+ + destruct n.
+ * inversion H as (BI). rewrite BI in *.
+ inversion H3. simpl. congruence.
+ * (* absurd *) congruence.
+ * (* absurd *) congruence.
+ + inv H3. simpl. destruct (zeq n 0). { contradiction. }
+ eapply IHlbi; eauto.
+Qed.
+
+Lemma list_nth_z_find_bi_with_header:
+ forall ll lbi bi tlbi n bi' (rest : list Asm.instruction),
+ list_nth_z lbi (n - list_length_z ll) = Some bi ->
+ unfold_body lbi = OK tlbi ->
+ basic_to_instruction bi = OK bi' ->
+ Asm.find_instr n ((unfold_label ll) ++ (tlbi) ++ (rest)) = Some bi'.
+Proof.
+ induction ll.
+ - unfold list_length_z. simpl. intros.
+ replace (n - 0) with n in H by lia. eapply list_nth_z_find_bi; eauto.
+ - intros. simpl. destruct (zeq n 0).
+ + rewrite list_length_z_cons in H. rewrite e in H.
+ replace (0 - (list_length_z ll + 1)) with (-1 - (list_length_z ll)) in H by lia.
+ generalize (list_length_z_pos ll). intros.
+ rewrite list_nth_z_neg in H; try lia. inversion H.
+ + rewrite list_length_z_cons in H.
+ replace (n - (list_length_z ll + 1)) with (n -1 - (list_length_z ll)) in H by lia.
+ eapply IHll; eauto.
+Qed.
+
+(* XXX unused *)
+Lemma range_list_nth_z:
+ forall (A: Type) (l: list A) n,
+ 0 <= n < list_length_z l ->
+ exists x, list_nth_z l n = Some x.
+Proof.
+ induction l.
+ - intros. unfold list_length_z in H. simpl in H. lia.
+ - intros n. destruct (zeq n 0).
+ + intros. simpl. destruct (zeq n 0). { eauto. } contradiction.
+ + intros H. rewrite list_length_z_cons in H.
+ simpl. destruct (zeq n 0). { contradiction. }
+ replace (Z.pred n) with (n - 1) by lia.
+ eapply IHl; lia.
+Qed.
+
+Lemma list_nth_z_n_too_big:
+ forall (A: Type) (l: list A) n,
+ 0 <= n ->
+ list_nth_z l n = None ->
+ n >= list_length_z l.
+Proof.
+ induction l.
+ - intros. unfold list_length_z. simpl. lia.
+ - intros. rewrite list_length_z_cons.
+ simpl in H0.
+ destruct (zeq n 0) as [N | N].
+ + inversion H0.
+ + (* XXX there must be a more elegant way to prove this simple fact *)
+ assert (n > 0). { lia. }
+ assert (0 <= n - 1). { lia. }
+ generalize (IHl (n - 1)). intros IH.
+ assert (n - 1 >= list_length_z l). { auto. }
+ assert (n > list_length_z l); lia.
+Qed.
+
+Lemma find_instr_past_header:
+ forall labels n rest,
+ list_nth_z labels n = None ->
+ Asm.find_instr n (unfold_label labels ++ rest) =
+ Asm.find_instr (n - list_length_z labels) rest.
+Proof.
+ induction labels as [| label labels' IH].
+ - unfold list_length_z; simpl; intros; rewrite Z.sub_0_r; reflexivity.
+ - intros. simpl. destruct (zeq n 0) as [N | N].
+ + rewrite N in H. inversion H.
+ + rewrite list_length_z_cons.
+ replace (n - (list_length_z labels' + 1)) with (n - 1 - list_length_z labels') by lia.
+ simpl in H. destruct (zeq n 0). { contradiction. }
+ replace (Z.pred n) with (n - 1) in H by lia.
+ apply IH; auto.
+Qed.
+
+(* very similar to find_instr_past_header *)
+Lemma find_instr_past_body:
+ forall lbi n tlbi rest,
+ list_nth_z lbi n = None ->
+ unfold_body lbi = OK tlbi ->
+ Asm.find_instr n (tlbi ++ rest) =
+ Asm.find_instr (n - list_length_z lbi) rest.
+Proof.
+ induction lbi.
+ - unfold list_length_z; simpl; intros ? ? ? ? H. inv H; rewrite Z.sub_0_r; reflexivity.
+ - intros n tlib ? NTH UNFOLD_BODY.
+ unfold unfold_body in UNFOLD_BODY. apply bind_inversion in UNFOLD_BODY.
+ destruct UNFOLD_BODY as (? & BI & H).
+ apply bind_inversion in H. destruct H as (? & UNFOLD_BODY' & CONS).
+ fold unfold_body in UNFOLD_BODY'. inv CONS.
+ simpl; destruct (zeq n 0) as [N|N].
+ + rewrite N in NTH; inversion NTH.
+ + rewrite list_length_z_cons.
+ replace (n - (list_length_z lbi + 1)) with (n - 1 - list_length_z lbi) by lia.
+ simpl in NTH. destruct (zeq n 0). { contradiction. }
+ replace (Z.pred n) with (n - 1) in NTH by lia.
+ apply IHlbi; auto.
+Qed.
+
+Lemma n_beyond_body:
+ forall bb n,
+ 0 <= n < size bb ->
+ list_nth_z (header bb) n = None ->
+ list_nth_z (body bb) (n - list_length_z (header bb)) = None ->
+ n >= Z.of_nat (length (header bb) + length (body bb)).
+Proof.
+ intros.
+ assert (0 <= n). { lia. }
+ generalize (list_nth_z_n_too_big label (header bb) n H2 H0). intros.
+ generalize (list_nth_z_n_too_big _ (body bb) (n - list_length_z (header bb))). intros.
+ unfold size in H.
+
+ assert (0 <= n - list_length_z (header bb)). { lia. }
+ assert (n - list_length_z (header bb) >= list_length_z (body bb)). { apply H4; auto. }
+
+ assert (n >= list_length_z (header bb) + list_length_z (body bb)). { lia. }
+ rewrite Nat2Z.inj_add.
+ repeat (rewrite <- list_length_z_nat). assumption.
+Qed.
+
+Lemma exec_arith_instr_dont_move_PC ai rs rs': forall
+ (BASIC: exec_arith_instr lk ai rs = rs'),
+ rs PC = rs' PC.
+Proof.
+ destruct ai; simpl; intros;
+ try (rewrite <- BASIC; rewrite Pregmap.gso; auto; discriminate).
+ - destruct i; simpl in BASIC;
+ try destruct (negb _); rewrite <- BASIC;
+ repeat rewrite Pregmap.gso; try discriminate; reflexivity.
+ - destruct i; simpl in BASIC.
+ 1,2: rewrite <- BASIC; repeat rewrite Pregmap.gso; try discriminate; reflexivity.
+ destruct sz;
+ try (unfold compare_single in BASIC || unfold compare_float in BASIC);
+ destruct (rs r1), (rs r2);
+ try (rewrite <- BASIC; repeat rewrite Pregmap.gso; try (discriminate || reflexivity)).
+ - destruct i; simpl in BASIC;
+ destruct is;
+ try (unfold compare_int in BASIC || unfold compare_long in BASIC);
+ try (rewrite <- BASIC; repeat rewrite Pregmap.gso; try (discriminate || reflexivity)).
+ - destruct i; simpl in BASIC; destruct sz;
+ try (unfold compare_single in BASIC || unfold compare_float in BASIC);
+ destruct (rs r1);
+ try (rewrite <- BASIC; repeat rewrite Pregmap.gso; try (discriminate || reflexivity)).
+ - destruct fsz; rewrite <- BASIC; rewrite Pregmap.gso; try (discriminate || reflexivity).
+ - destruct fsz; rewrite <- BASIC; rewrite Pregmap.gso; try (discriminate || reflexivity).
+Qed.
+
+Lemma exec_basic_dont_move_PC bi rs m rs' m': forall
+ (BASIC: exec_basic lk ge bi rs m = Next rs' m'),
+ rs PC = rs' PC.
+Proof.
+ destruct bi; simpl; intros.
+ - inv BASIC. exploit exec_arith_instr_dont_move_PC; eauto.
+ - unfold exec_load in BASIC.
+ destruct ld.
+ + unfold exec_load_rd_a in BASIC.
+ destruct Mem.loadv. 2: { discriminate BASIC. }
+ inv BASIC. rewrite Pregmap.gso; try discriminate; auto.
+ + unfold exec_load_double, is_pair_addressing_mode_correct in BASIC.
+ destruct a; try discriminate BASIC.
+ do 2 (destruct Mem.loadv; try discriminate BASIC).
+ inv BASIC. rewrite Pregmap.gso; try discriminate; auto.
+ - unfold exec_store in BASIC.
+ destruct st.
+ + unfold exec_store_rs_a in BASIC.
+ destruct Mem.storev. 2: { discriminate BASIC. }
+ inv BASIC; reflexivity.
+ + unfold exec_store_double in BASIC.
+ destruct a; try discriminate BASIC.
+ do 2 (destruct Mem.storev; try discriminate BASIC).
+ inv BASIC; reflexivity.
+ - destruct Mem.alloc, Mem.store. 2: { discriminate BASIC. }
+ inv BASIC. repeat (rewrite Pregmap.gso; try discriminate). reflexivity.
+ - destruct Mem.loadv. 2: { discriminate BASIC. }
+ destruct rs, Mem.free; try discriminate BASIC.
+ inv BASIC; rewrite Pregmap.gso; try discriminate; auto.
+ - inv BASIC; rewrite Pregmap.gso; try discriminate; auto.
+ - inv BASIC; rewrite Pregmap.gso; try discriminate; auto.
+ - inv BASIC; rewrite Pregmap.gso; try discriminate; auto.
+ - inv BASIC; rewrite Pregmap.gso; try discriminate; auto.
+ - inv BASIC; auto.
+Qed.
+
+Lemma exec_body_dont_move_PC_aux:
+ forall bis rs m rs' m'
+ (BODY: exec_body lk ge bis rs m = Next rs' m'),
+ rs PC = rs' PC.
+Proof.
+ induction bis.
+ - intros; inv BODY; reflexivity.
+ - simpl; intros.
+ remember (exec_basic lk ge a rs m) as bi eqn:BI; destruct bi. 2: { discriminate BODY. }
+ symmetry in BI; destruct s in BODY, BI; simpl in BODY, BI.
+ exploit exec_basic_dont_move_PC; eauto; intros AGPC; rewrite AGPC.
+ eapply IHbis; eauto.
+Qed.
+
+Lemma exec_body_dont_move_PC bb rs m rs' m': forall
+ (BODY: exec_body lk ge (body bb) rs m = Next rs' m'),
+ rs PC = rs' PC.
+Proof. apply exec_body_dont_move_PC_aux. Qed.
+
+Lemma find_instr_bblock:
+ forall n lb pos bb tlb
+ (FINDBB: find_bblock pos lb = Some bb)
+ (UNFOLD: unfold lb = OK tlb)
+ (SIZE: 0 <= n < size bb),
+ exists i, is_nth_inst bb n i /\ Asm.find_instr (pos+n) tlb = Some i.
+Proof.
+ induction lb as [| b lb IHlb].
+ - intros. inversion FINDBB.
+ - intros pos bb tlb FINDBB UNFOLD SIZE.
+ destruct pos.
+ + inv FINDBB. simpl.
+ exploit unfold_car_cdr; eauto. intros (tbb & tlb' & UNFOLD_BBLOCK & UNFOLD' & UNFOLD_cons).
+ rewrite UNFOLD in UNFOLD_cons. inversion UNFOLD_cons.
+ unfold unfold_bblock in UNFOLD_BBLOCK.
+ destruct (zle (list_length_z (header bb)) 1). 2: { inversion UNFOLD_BBLOCK. }
+ apply bind_inversion in UNFOLD_BBLOCK.
+ destruct UNFOLD_BBLOCK as (? & UNFOLD_BODY & H).
+ inversion H as (UNFOLD_BBLOCK).
+ remember (list_nth_z (header bb) n) as label_opt eqn:LBL. destruct label_opt.
+ * (* nth instruction is a label *)
+ eexists; split. { eapply is_nth_label; eauto. }
+ inversion UNFOLD_cons.
+ symmetry in LBL.
+ rewrite <- app_assoc.
+ apply list_nth_z_find_label; auto.
+ * remember (list_nth_z (body bb) (n - list_length_z (header bb))) as bi_opt eqn:BI.
+ destruct bi_opt.
+ -- (* nth instruction is a basic instruction *)
+ exploit list_nth_z_in; eauto. intros INBB.
+ exploit entire_body_translated; eauto. intros BI'.
+ destruct BI'.
+ eexists; split.
+ ++ eapply is_nth_basic; eauto.
+ ++ repeat (rewrite <- app_assoc). eapply list_nth_z_find_bi_with_header; eauto.
+ -- (* nth instruction is the exit instruction *)
+ generalize n_beyond_body. intros TEMP.
+ assert (n >= Z.of_nat (Datatypes.length (header bb)
+ + Datatypes.length (body bb))) as NGE. { auto. } clear TEMP.
+ remember (exit bb) as exit_opt eqn:EXIT. destruct exit_opt.
+ ++ rewrite <- app_assoc. rewrite find_instr_past_header; auto.
+ rewrite <- app_assoc. erewrite find_instr_past_body; eauto.
+ assert (SIZE' := SIZE).
+ unfold size in SIZE. rewrite <- EXIT in SIZE. simpl in SIZE.
+ destruct SIZE as (LOWER & UPPER).
+ repeat (rewrite Nat2Z.inj_add in UPPER).
+ repeat (rewrite <- list_length_z_nat in UPPER). repeat (rewrite Nat2Z.inj_add in NGE).
+ repeat (rewrite <- list_length_z_nat in NGE). simpl in UPPER.
+ assert (n = list_length_z (header bb) + list_length_z (body bb)). { lia. }
+ assert (n = size bb - 1). {
+ unfold size. rewrite <- EXIT. simpl.
+ repeat (rewrite Nat2Z.inj_add). repeat (rewrite <- list_length_z_nat). simpl. lia.
+ }
+ symmetry in EXIT.
+ eexists; split.
+ ** eapply is_nth_ctlflow; eauto.
+ ** simpl.
+ destruct (zeq (n - list_length_z (header bb) - list_length_z (body bb)) 0). { reflexivity. }
+ (* absurd *) lia.
+ ++ (* absurd *)
+ unfold size in SIZE. rewrite <- EXIT in SIZE. simpl in SIZE.
+ destruct SIZE as (? & SIZE'). rewrite Nat.add_0_r in SIZE'. lia.
+ + unfold find_bblock in FINDBB; simpl in FINDBB; fold find_bblock in FINDBB.
+ inversion UNFOLD as (UNFOLD').
+ apply bind_inversion in UNFOLD'. destruct UNFOLD' as (? & (UNFOLD_BBLOCK' & UNFOLD')).
+ apply bind_inversion in UNFOLD'. destruct UNFOLD' as (? & (UNFOLD' & TLB)).
+ inversion TLB.
+ generalize (IHlb _ _ _ FINDBB UNFOLD'). intros IH.
+ destruct IH as (? & (IH_is_nth & IH_find_instr)); eauto.
+ eexists; split.
+ * apply IH_is_nth.
+ * replace (Z.pos p + n) with (Z.pos p + n - size b + size b) by lia.
+ eapply find_instr_bblock_tail; try assumption.
+ replace (Z.pos p + n - size b) with (Z.pos p - size b + n) by lia.
+ apply IH_find_instr.
+ + (* absurd *)
+ generalize (Pos2Z.neg_is_neg p). intros. exploit (find_bblock_neg (b :: lb)); eauto.
+ rewrite FINDBB. intros CONTRA. inversion CONTRA.
+Qed.
+
+Lemma exec_header_simulation b ofs f bb rs m: forall
+ (ATPC: rs PC = Vptr b ofs)
+ (FINDF: Genv.find_funct_ptr ge b = Some (Internal f))
+ (FINDBB: find_bblock (Ptrofs.unsigned ofs) (fn_blocks f) = Some bb),
+ exists s', star Asm.step tge (State rs m) E0 s'
+ /\ match_internal (list_length_z (header bb)) (State rs m) s'.
+Proof.
+ intros.
+ exploit internal_functions_unfold; eauto.
+ intros (tc & FINDtf & TRANStf & _).
+ assert (BNDhead: list_length_z (header bb) <= 1). { eapply size_header; eauto. }
+ destruct (header bb) as [|l[|]] eqn: EQhead.
+ + (* header nil *)
+ eexists; split.
+ - eapply star_refl.
+ - split; eauto.
+ unfold list_length_z; rewrite !ATPC; simpl.
+ rewrite Ptrofs.add_zero; auto.
+ + (* header one *)
+ assert (Lhead: list_length_z (header bb) = 1). { rewrite EQhead; unfold list_length_z; simpl. auto. }
+ exploit (find_instr_bblock 0); eauto.
+ { generalize (bblock_size_pos bb). lia. }
+ intros (i & NTH & FIND_INSTR).
+ inv NTH.
+ * rewrite EQhead in H; simpl in H. inv H.
+ replace (Ptrofs.unsigned ofs + 0) with (Ptrofs.unsigned ofs) in FIND_INSTR by lia.
+ eexists. split.
+ - eapply star_one.
+ eapply Asm.exec_step_internal; eauto.
+ simpl; eauto.
+ - unfold list_length_z; simpl. split; eauto.
+ intros r; destruct r; simpl; congruence || auto.
+ * (* absurd case *)
+ erewrite list_nth_z_neg in * |-; [ congruence | rewrite Lhead; lia].
+ * (* absurd case *)
+ rewrite bblock_size_aux, Lhead in *. generalize (bblock_size_aux_pos bb). lia.
+ + (* absurd case *)
+ unfold list_length_z in BNDhead. simpl in *.
+ generalize (list_length_z_aux_increase _ l1 2); lia.
+Qed.
+
+Lemma eval_addressing_preserved a rs1 rs2:
+ (forall r : preg, r <> PC -> rs1 r = rs2 r) ->
+ eval_addressing lk a rs1 = Asm.eval_addressing tge a rs2.
+Proof.
+ intros EQ.
+ destruct a; simpl; try (rewrite !EQ; congruence).
+ auto.
+Qed.
+
+Ltac next_stuck_cong := try (unfold Next, Stuck in *; congruence).
+
+Ltac inv_ok_eq :=
+ repeat match goal with
+ | [EQ: OK ?x = OK ?y |- _ ]
+ => inversion EQ; clear EQ; subst
+ end.
+
+Ltac reg_rwrt :=
+ match goal with
+ | [e: DR _ = DR _ |- _ ]
+ => rewrite e in *
+ end.
+
+Ltac destruct_reg_inv :=
+ repeat match goal with
+ | [ H : match ?reg with _ => _ end = _ |- _ ]
+ => simpl in *; destruct reg; try congruence; try inv_ok_eq; try reg_rwrt
+ end.
+
+Ltac destruct_ireg_inv :=
+ repeat match goal with
+ | [ H : match ?reg with _ => _ end = _ |- _ ]
+ => destruct reg as [[r|]|]; try congruence; try inv_ok_eq; subst
+ end.
+
+Ltac destruct_reg_size :=
+ simpl in *;
+ match goal with
+ | [ |- context [ match ?reg with _ => _ end ] ]
+ => destruct reg; try congruence
+ end.
+
+Ltac find_rwrt_ag :=
+ simpl in *;
+ match goal with
+ | [ AG: forall r, r <> ?PC -> _ r = _ r |- _ ]
+ => repeat rewrite <- AG; try congruence
+ end.
+
+Ltac inv_matchi :=
+ match goal with
+ | [ MATCHI : match_internal _ _ _ |- _ ]
+ => inversion MATCHI; subst; find_rwrt_ag
+ end.
+
+Ltac destruct_ir0_reg :=
+ match goal with
+ | [ |- context [ ir0 _ _ ?r ] ]
+ => unfold ir0 in *; destruct r; find_rwrt_ag; eauto
+ end.
+
+Ltac pc_not_sp :=
+ match goal with
+ | [ |- ?PC <> ?SP ]
+ => destruct (PregEq.eq SP PC); repeat congruence; discriminate
+ end.
+
+Ltac update_x_access_x :=
+ subst; rewrite !Pregmap.gss; auto.
+
+Ltac update_x_access_r :=
+ rewrite !Pregmap.gso; auto.
+
+Lemma nextinstr_agree_but_pc rs1 rs2: forall
+ (AG: forall r, r <> PC -> rs1 r = rs2 r),
+ forall r, r <> PC -> rs1 r = Asm.nextinstr rs2 r.
+Proof.
+ intros; unfold Asm.nextinstr in *; rewrite Pregmap.gso in *; eauto.
+Qed.
+
+Lemma ptrofs_nextinstr_agree rs1 rs2 n: forall
+ (BOUNDED : 0 <= n <= Ptrofs.max_unsigned)
+ (AGPC : Val.offset_ptr (rs1 PC) (Ptrofs.repr n) = rs2 PC),
+ Val.offset_ptr (rs1 PC) (Ptrofs.repr (n + 1)) = Asm.nextinstr rs2 PC.
+Proof.
+ intros; unfold Asm.nextinstr; rewrite Pregmap.gss.
+ rewrite <- Ptrofs.unsigned_one; rewrite <- (Ptrofs.unsigned_repr n); eauto;
+ rewrite <- Ptrofs.add_unsigned; rewrite <- Val.offset_ptr_assoc; rewrite AGPC; eauto.
+Qed.
+
+Lemma load_rd_a_preserved n rs1 m1 rs1' m1' rs2 m2 rd chk f a: forall
+ (BOUNDED: 0 <= n <= Ptrofs.max_unsigned)
+ (MATCHI: match_internal n (State rs1 m1) (State rs2 m2))
+ (HLOAD: exec_load_rd_a lk chk f a rd rs1 m1 = Next rs1' m1'),
+ exists (rs2' : regset) (m2' : mem), Asm.exec_load tge chk f a rd rs2 m2 = Next rs2' m2'
+ /\ match_internal (n + 1) (State rs1' m1') (State rs2' m2').
+Proof.
+ intros.
+ unfold exec_load_rd_a, Asm.exec_load in *.
+ inversion MATCHI as [n0 r1 mx1 r2 mx2 EQM EQR EQPC]; subst.
+ rewrite <- (eval_addressing_preserved a rs1 rs2); auto.
+ destruct (Mem.loadv _ _ _).
+ + inversion HLOAD; auto. repeat (econstructor; eauto).
+ * eapply nextinstr_agree_but_pc; intros.
+ destruct (PregEq.eq r rd); try update_x_access_x; try update_x_access_r.
+ * eapply ptrofs_nextinstr_agree; eauto.
+ + next_stuck_cong.
+Qed.
+
+Lemma load_double_preserved n rs1 m1 rs1' m1' rs2 m2 rd1 rd2 chk1 chk2 f a: forall
+ (BOUNDED: 0 <= n <= Ptrofs.max_unsigned)
+ (MATCHI: match_internal n (State rs1 m1) (State rs2 m2))
+ (HLOAD: exec_load_double lk chk1 chk2 f a rd1 rd2 rs1 m1 = Next rs1' m1'),
+ exists (rs2' : regset) (m2' : mem), Asm.exec_load_double tge chk1 chk2 f a rd1 rd2 rs2 m2 = Next rs2' m2'
+ /\ match_internal (n + 1) (State rs1' m1') (State rs2' m2').
+Proof.
+ intros.
+ unfold exec_load_double, Asm.exec_load_double in *.
+ inversion MATCHI as [n0 r1 mx1 r2 mx2 EQM EQR EQPC]; subst.
+ erewrite <- !eval_addressing_preserved; eauto.
+ destruct (is_pair_addressing_mode_correct a); try discriminate.
+ destruct (Mem.loadv _ _ _);
+ destruct (Mem.loadv chk2 m2
+ (eval_addressing lk
+ (get_offset_addr a match chk1 with
+ | Mint32 | Mfloat32| Many32 => 4
+ | _ => 8
+ end) rs1));
+ inversion HLOAD; auto.
+ repeat (econstructor; eauto).
+ * eapply nextinstr_agree_but_pc; intros.
+ destruct (PregEq.eq r rd2); destruct (PregEq.eq r rd1).
+ - try update_x_access_x.
+ - try update_x_access_x.
+ - subst; repeat rewrite Pregmap.gso, Pregmap.gss; auto.
+ - try update_x_access_r.
+ * eapply ptrofs_nextinstr_agree; eauto.
+Qed.
+
+<<<<<<< HEAD
+Lemma store_rs_a_preserved n rs1 m1 rs1' m1' rs2 m2 v chk a: forall
+ (BOUNDED: 0 <= n <= Ptrofs.max_unsigned)
+ (MATCHI: match_internal n (State rs1 m1) (State rs2 m2))
+ (HSTORE: exec_store_rs_a lk chk a v rs1 m1 = Next rs1' m1'),
+ exists (rs2' : regset) (m2' : mem), Asm.exec_store tge chk a v rs2 m2 = Next rs2' m2'
+ /\ match_internal (n + 1) (State rs1' m1') (State rs2' m2').
+Proof.
+ intros.
+ unfold exec_store_rs_a, Asm.exec_store in *.
+ inversion MATCHI as [n0 r1 mx1 r2 mx2 EQM EQR EQPC]; subst.
+ rewrite <- (eval_addressing_preserved a rs1 rs2); auto.
+ destruct (Mem.storev _ _ _ _).
+ + inversion HSTORE; auto. repeat (econstructor; eauto).
+ * eapply nextinstr_agree_but_pc; intros.
+ subst. apply EQR. auto.
+ * eapply ptrofs_nextinstr_agree; subst; eauto.
+ + next_stuck_cong.
+Qed.
+=======
+Lemma find_label_goto_label:
+ forall f tf lbl rs m c' b ofs,
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
+ transf_function f = OK tf ->
+ rs PC = Vptr b ofs ->
+ Mach.find_label lbl f.(Mach.fn_code) = Some c' ->
+ exists tc', exists rs',
+ goto_label tf lbl rs m = Next rs' m
+ /\ transl_code_at_pc ge (rs' PC) b f c' false tf tc'
+ /\ forall r, r <> PC -> rs'#r = rs#r.
+Proof.
+ intros. exploit (transl_find_label lbl f tf); eauto. rewrite H2.
+ intros [tc [A B]].
+ exploit label_pos_code_tail; eauto. instantiate (1 := 0).
+ intros [pos' [P [Q R]]].
+ exists tc; exists (rs#PC <- (Vptr b (Ptrofs.repr pos'))).
+ split. unfold goto_label. rewrite P. rewrite H1. auto.
+ split. rewrite Pregmap.gss. constructor; auto.
+ rewrite Ptrofs.unsigned_repr. replace (pos' - 0) with pos' in Q.
+ auto. lia.
+ generalize (transf_function_no_overflow _ _ H0). lia.
+ intros. apply Pregmap.gso; auto.
+Qed.
+
+(** Existence of return addresses *)
+>>>>>>> master
+
+Lemma store_double_preserved n rs1 m1 rs1' m1' rs2 m2 v1 v2 chk1 chk2 a: forall
+ (BOUNDED: 0 <= n <= Ptrofs.max_unsigned)
+ (MATCHI: match_internal n (State rs1 m1) (State rs2 m2))
+ (HSTORE: exec_store_double lk chk1 chk2 a v1 v2 rs1 m1 = Next rs1' m1'),
+ exists (rs2' : regset) (m2' : mem), Asm.exec_store_double tge chk1 chk2 a v1 v2 rs2 m2 = Next rs2' m2'
+ /\ match_internal (n + 1) (State rs1' m1') (State rs2' m2').
+Proof.
+ intros.
+ unfold exec_store_double, Asm.exec_store_double in *.
+ inversion MATCHI as [n0 r1 mx1 r2 mx2 EQM EQR EQPC]; subst.
+ erewrite <- !eval_addressing_preserved; eauto.
+ destruct (is_pair_addressing_mode_correct a); try discriminate.
+ destruct (Mem.storev _ _ _ _);
+ try destruct (Mem.storev chk2 m
+ (eval_addressing lk
+ (get_offset_addr a
+ match chk1 with
+ | Mint32 | Mfloat32 | Many32 => 4
+ | _ => 8
+ end) rs1) v2);
+ inversion HSTORE; auto.
+ repeat (econstructor; eauto).
+ * eapply nextinstr_agree_but_pc; intros.
+ subst. apply EQR. auto.
+ * eapply ptrofs_nextinstr_agree; subst; eauto.
+Qed.
+
+Lemma next_inst_preserved n rs1 m1 rs1' m1' rs2 m2 (x: dreg) v: forall
+ (BOUNDED: 0 <= n <= Ptrofs.max_unsigned)
+ (MATCHI: match_internal n (State rs1 m1) (State rs2 m2))
+ (NEXTI: Next rs1 # x <- v m1 = Next rs1' m1'),
+ exists (rs2' : regset) (m2' : mem),
+ Next (Asm.nextinstr rs2 # x <- v) m2 = Next rs2' m2'
+ /\ match_internal (n + 1) (State rs1' m1') (State rs2' m2').
+Proof.
+ intros.
+ inversion MATCHI as [n0 r1 mx1 r2 mx2 EQM EQR EQPC]; subst.
+ inversion NEXTI. repeat (econstructor; eauto).
+ * eapply nextinstr_agree_but_pc; intros.
+ destruct (PregEq.eq r x); try update_x_access_x; try update_x_access_r.
+ * eapply ptrofs_nextinstr_agree; eauto.
+Qed.
+
+Lemma match_internal_nextinstr_switch:
+ forall n s rs2 m2 r v,
+ r <> PC ->
+ match_internal n s (State ((Asm.nextinstr rs2)#r <- v) m2) ->
+ match_internal n s (State (Asm.nextinstr (rs2#r <- v)) m2).
+Proof.
+ unfold Asm.nextinstr; intros n s rs2 m2 r v NOTPC1 MI.
+ inversion MI; subst; constructor; auto.
+ - eapply nextinstr_agree_but_pc; intros.
+ rewrite AG; try congruence.
+ destruct (PregEq.eq r r0); try update_x_access_x; try update_x_access_r.
+ - rewrite !Pregmap.gss, !Pregmap.gso; try congruence.
+ rewrite AGPC.
+ rewrite Pregmap.gso, Pregmap.gss; try congruence.
+Qed.
+
+Lemma match_internal_nextinstr_set_parallel:
+ forall n rs1 m1 rs2 m2 r v1 v2,
+ r <> PC ->
+ match_internal n (State rs1 m1) (State (Asm.nextinstr rs2) m2) ->
+ v1 = v2 ->
+ match_internal n (State (rs1#r <- v1) m1) (State (Asm.nextinstr (rs2#r <- v2)) m2).
+Proof.
+ intros; subst; eapply match_internal_nextinstr_switch; eauto.
+ intros; eapply match_internal_set_parallel; eauto.
+Qed.
+
+Lemma exec_basic_simulation:
+ forall tf n rs1 m1 rs1' m1' rs2 m2 bi tbi
+ (BOUNDED: 0 <= n <= Ptrofs.max_unsigned)
+ (BASIC: exec_basic lk ge bi rs1 m1 = Next rs1' m1')
+ (MATCHI: match_internal n (State rs1 m1) (State rs2 m2))
+ (TRANSBI: basic_to_instruction bi = OK tbi),
+ exists rs2' m2', Asm.exec_instr tge tf tbi
+ rs2 m2 = Next rs2' m2'
+ /\ match_internal (n + 1) (State rs1' m1') (State rs2' m2').
+Proof.
+ intros.
+ destruct bi.
+ { (* PArith *)
+ simpl in *; destruct i.
+ 1: {
+ destruct i.
+ 1,2,3:
+ try (destruct sumbool_rec; try congruence);
+ try (monadInv TRANSBI);
+ try (destruct_reg_inv);
+ try (inv_matchi);
+ try (exploit next_inst_preserved; eauto);
+ try (repeat destruct_reg_size);
+ try (destruct_ir0_reg).
+ 1,2: (* Special case for Pfmovimmd / Pfmovimms *)
+ try (monadInv TRANSBI);
+ try (destruct_reg_inv);
+ try (inv_matchi);
+ inversion BASIC; clear BASIC; subst;
+ try (destruct (is_immediate_float64 _));
+ try (destruct (is_immediate_float32 _));
+ eexists; eexists; split; eauto;
+ repeat (eapply match_internal_nextinstr_set_parallel; try congruence);
+ try (econstructor; eauto);
+ try (eapply nextinstr_agree_but_pc; eauto);
+ try (eapply ptrofs_nextinstr_agree; eauto).
+ }
+ 1,2,3,4,5: (* PArithP, PArithPP, PArithPPP, PArithRR0R, PArithRR0, PArithARRRR0 *)
+ destruct i;
+ try (destruct sumbool_rec; try congruence);
+ try (monadInv TRANSBI);
+ try (destruct_reg_inv);
+ try (inv_matchi);
+ try (exploit next_inst_preserved; eauto);
+ try (repeat destruct_reg_size);
+ try (destruct_ir0_reg).
+ { (* PArithComparisonPP *)
+ destruct i;
+ try (monadInv TRANSBI);
+ try (inv_matchi);
+ try (destruct_reg_inv);
+ simpl in *.
+ 1,2: (* compare_long *)
+ inversion BASIC; clear BASIC; subst;
+ eexists; eexists; split; eauto;
+ unfold compare_long;
+ repeat (eapply match_internal_nextinstr_set_parallel; [ congruence | idtac | try (rewrite !AG; congruence)]);
+ try (econstructor; eauto);
+ try (eapply nextinstr_agree_but_pc; eauto);
+ try (eapply ptrofs_nextinstr_agree; eauto).
+
+ destruct sz.
+ - (* compare_single *)
+ unfold compare_single in BASIC.
+ destruct (rs1 x), (rs1 x0);
+ inversion BASIC;
+ eexists; eexists; split; eauto;
+ repeat (eapply match_internal_nextinstr_set_parallel; [ congruence | idtac | try (rewrite !AG; congruence)]);
+ try (econstructor; eauto);
+ try (eapply nextinstr_agree_but_pc; eauto);
+ try (eapply ptrofs_nextinstr_agree; eauto).
+ - (* compare_float *)
+ unfold compare_float in BASIC.
+ destruct (rs1 x), (rs1 x0);
+ inversion BASIC;
+ eexists; eexists; split; eauto;
+ repeat (eapply match_internal_nextinstr_set_parallel; [ congruence | idtac | try (rewrite !AG; congruence)]);
+ try (econstructor; eauto);
+ try (eapply nextinstr_agree_but_pc; eauto);
+ try (eapply ptrofs_nextinstr_agree; eauto). }
+ 1,2: (* PArithComparisonR0R, PArithComparisonP *)
+ destruct i;
+ try (monadInv TRANSBI);
+ try (inv_matchi);
+ try (destruct_reg_inv);
+ try (destruct_reg_size);
+ simpl in *;
+ inversion BASIC; clear BASIC; subst;
+ eexists; eexists; split; eauto;
+ unfold compare_long, compare_int, compare_float, compare_single;
+ try (destruct_reg_size);
+ repeat (eapply match_internal_nextinstr_set_parallel; [ congruence | idtac | try (rewrite !AG; congruence)]);
+ try (econstructor; eauto);
+ try (destruct_ir0_reg);
+ try (eapply nextinstr_agree_but_pc; eauto);
+ try (eapply ptrofs_nextinstr_agree; eauto).
+ { (* Pcset *)
+ try (monadInv TRANSBI);
+ try (inv_matchi).
+ try (exploit next_inst_preserved; eauto);
+ try (simpl in *; intros;
+ unfold if_opt_bool_val in *; unfold eval_testcond in *;
+ rewrite <- !AG; try congruence; eauto). }
+ { (* Pfmovi *)
+ try (monadInv TRANSBI);
+ try (inv_matchi);
+ try (destruct_reg_size);
+ try (destruct_ir0_reg);
+ try (exploit next_inst_preserved; eauto). }
+ { (* Pcsel *)
+ try (destruct_reg_inv);
+ try (monadInv TRANSBI);
+ try (destruct_reg_inv);
+ try (inv_matchi);
+ try (exploit next_inst_preserved; eauto);
+ simpl in *; intros;
+ unfold if_opt_bool_val in *; unfold eval_testcond in *;
+ rewrite <- !AG; try congruence; eauto. }
+ { (* Pfnmul *)
+ try (monadInv TRANSBI);
+ try (inv_matchi);
+ try (destruct_reg_size);
+ try (exploit next_inst_preserved; eauto);
+ try (find_rwrt_ag). } }
+ { (* PLoad *)
+ destruct ld.
+ - destruct ld; monadInv TRANSBI; try destruct_ireg_inv; exploit load_rd_a_preserved; eauto;
+ intros; simpl in *; destruct sz; eauto.
+ - destruct ld; monadInv TRANSBI; destruct rd1 as [[rd1'|]|]; destruct rd2 as [[rd2'|]|];
+ inv EQ; inv EQ1; exploit load_double_preserved; eauto. }
+ { (* PStore *)
+ destruct st.
+ - destruct st; monadInv TRANSBI; try destruct_ireg_inv; exploit store_rs_a_preserved; eauto;
+ simpl in *; inv_matchi; find_rwrt_ag.
+ - destruct st; monadInv TRANSBI; destruct rs0 as [[rs0'|]|]; destruct rs3 as [[rs3'|]|];
+ inv EQ; inv EQ1; exploit store_double_preserved; eauto;
+ simpl in *; inv_matchi; find_rwrt_ag. }
+ { (* Pallocframe *)
+ monadInv TRANSBI;
+ inv_matchi; try pc_not_sp;
+ destruct sz eqn:EQSZ;
+ destruct Mem.alloc eqn:EQALLOC;
+ destruct Mem.store eqn:EQSTORE; inversion BASIC; try pc_not_sp;
+ eexists; eexists; split; eauto;
+ repeat (eapply match_internal_nextinstr_set_parallel; [ try (pc_not_sp; congruence) | idtac | try (reflexivity)]);
+ try (econstructor; eauto);
+ try (eapply nextinstr_agree_but_pc; eauto);
+ try (eapply ptrofs_nextinstr_agree; eauto). }
+ { (* Pfreeframe *)
+ monadInv TRANSBI;
+ inv_matchi; try pc_not_sp;
+ destruct sz eqn:EQSZ;
+ destruct Mem.loadv eqn:EQLOAD;
+ destruct (rs1 SP) eqn:EQRS1SP;
+ try (destruct Mem.free eqn:EQFREE);
+ inversion BASIC; try pc_not_sp;
+ eexists; eexists; split; eauto;
+ repeat (eapply match_internal_nextinstr_set_parallel; [ try (pc_not_sp; congruence) | idtac | try (reflexivity)]);
+ try (econstructor; eauto);
+ try (eapply nextinstr_agree_but_pc; eauto);
+ try (eapply ptrofs_nextinstr_agree; eauto). }
+ 1,2,3,4: (* Ploadsymbol, Pcvtsw2x, Pcvtuw2x, Pcvtx2w *)
+ try (monadInv TRANSBI);
+ try (inv_matchi);
+ try (exploit next_inst_preserved; eauto);
+ rewrite symbol_addresses_preserved; eauto;
+ try (find_rwrt_ag).
+ { (* Pnop *)
+ monadInv TRANSBI; inv_matchi.
+ inversion BASIC.
+ repeat (econstructor; eauto).
+ eapply nextinstr_agree_but_pc; intros;
+ try rewrite <- H0, AG; auto.
+ try eapply ptrofs_nextinstr_agree; auto; rewrite <- H0;
+ assumption. }
+Qed.
+
+Lemma find_basic_instructions b ofs f bb tc: forall
+ (FINDF: Genv.find_funct_ptr ge b = Some (Internal f))
+ (FINDBB: find_bblock (Ptrofs.unsigned ofs) (fn_blocks f) = Some bb)
+ (UNFOLD: unfold (fn_blocks f) = OK tc),
+ forall n,
+ (n < length (body bb))%nat ->
+ exists (i : Asm.instruction) (bi : basic),
+ list_nth_z (body bb) (Z.of_nat n) = Some bi
+ /\ basic_to_instruction bi = OK i
+ /\ Asm.find_instr (Ptrofs.unsigned ofs
+ + (list_length_z (header bb))
+ + Z.of_nat n) tc
+ = Some i.
+Proof.
+ intros until n; intros NLT.
+ exploit internal_functions_unfold; eauto.
+ intros (tc' & FINDtf & TRANStf & _).
+ assert (tc' = tc) by congruence; subst.
+ exploit (find_instr_bblock (list_length_z (header bb) + Z.of_nat n)); eauto.
+ { unfold size; split.
+ - rewrite list_length_z_nat; lia.
+ - repeat (rewrite list_length_z_nat). repeat (rewrite Nat2Z.inj_add). lia. }
+ intros (i & NTH & FIND_INSTR).
+ exists i; intros.
+ inv NTH.
+ - (* absurd *) apply list_nth_z_range in H; lia.
+ - exists bi;
+ rewrite Z.add_simpl_l in H;
+ rewrite Z.add_assoc in FIND_INSTR;
+ intuition.
+ - (* absurd *) rewrite bblock_size_aux in H0;
+ rewrite H in H0; simpl in H0; repeat rewrite list_length_z_nat in H0; lia.
+Qed.
+
+(* TODO: remplacer find_basic_instructions directement par ce lemme ? *)
+Lemma find_basic_instructions_alt b ofs f bb tc n: forall
+ (FINDF: Genv.find_funct_ptr ge b = Some (Internal f))
+ (FINDBB: find_bblock (Ptrofs.unsigned ofs) (fn_blocks f) = Some bb)
+ (UNFOLD: unfold (fn_blocks f) = OK tc)
+ (BOUND: 0 <= n < list_length_z (body bb)),
+ exists (i : Asm.instruction) (bi : basic),
+ list_nth_z (body bb) n = Some bi
+ /\ basic_to_instruction bi = OK i
+ /\ Asm.find_instr (Ptrofs.unsigned ofs
+ + (list_length_z (header bb))
+ + n) tc
+ = Some i.
+Proof.
+ intros; assert ((Z.to_nat n) < length (body bb))%nat.
+ { rewrite Nat2Z.inj_lt, <- list_length_z_nat, Z2Nat.id; try lia. }
+ exploit find_basic_instructions; eauto.
+ rewrite Z2Nat.id; try lia. intros (i & bi & X).
+ eexists; eexists; intuition eauto.
+Qed.
+
+Lemma header_body_tail_bound: forall (a: basic) (li: list basic) bb ofs
+ (BOUNDBB : Ptrofs.unsigned ofs + size bb <= Ptrofs.max_unsigned)
+ (BDYLENPOS : 0 <= list_length_z (body bb) - list_length_z (a :: li) <
+ list_length_z (body bb)),
+0 <= list_length_z (header bb) + list_length_z (body bb) - list_length_z (a :: li) <=
+Ptrofs.max_unsigned.
+Proof.
+ intros.
+ assert (HBBPOS: list_length_z (header bb) >= 0) by eapply list_length_z_pos.
+ assert (HBBSIZE: list_length_z (header bb) < size bb) by eapply header_size_lt_block_size.
+ assert (OFSBOUND: 0 <= Ptrofs.unsigned ofs <= Ptrofs.max_unsigned) by eapply Ptrofs.unsigned_range_2.
+ assert (BBSIZE: size bb <= Ptrofs.max_unsigned) by lia.
+ unfold size in BBSIZE.
+ rewrite !Nat2Z.inj_add in BBSIZE.
+ rewrite <- !list_length_z_nat in BBSIZE.
+ lia.
+Qed.
+
+(* A more general version of the exec_body_simulation_plus lemma below.
+ This generalization is necessary for the induction proof inside the body.
+*)
+Lemma exec_body_simulation_plus_gen li: forall b ofs f bb rs m s2 rs' m'
+ (BLI: is_tail li (body bb))
+ (ATPC: rs PC = Vptr b ofs)
+ (FINDF: Genv.find_funct_ptr ge b = Some (Internal f))
+ (FINDBB: find_bblock (Ptrofs.unsigned ofs) (fn_blocks f) = Some bb)
+ (NEMPTY_BODY: li <> nil)
+ (MATCHI: match_internal ((list_length_z (header bb)) + (list_length_z (body bb)) - (list_length_z li)) (State rs m) s2)
+ (BODY: exec_body lk ge li rs m = Next rs' m'),
+ exists s2', plus Asm.step tge s2 E0 s2'
+ /\ match_internal (size bb - (Z.of_nat (length_opt (exit bb)))) (State rs' m') s2'.
+Proof.
+ induction li as [|a li]; simpl; try congruence.
+ intros.
+ assert (BDYLENPOS: 0 <= (list_length_z (body bb) - list_length_z (a::li)) < list_length_z (body bb)). {
+ assert (Z.of_nat O < list_length_z (a::li) <= list_length_z (body bb)); try lia.
+ rewrite !list_length_z_nat; split.
+ - rewrite <- Nat2Z.inj_lt. simpl. lia.
+ - rewrite <- Nat2Z.inj_le; eapply is_tail_bound; eauto.
+ }
+ exploit internal_functions_unfold; eauto.
+ intros (tc & FINDtf & TRANStf & _).
+ exploit find_basic_instructions_alt; eauto.
+ intros (tbi & (bi & (NTHBI & TRANSBI & FIND_INSTR))).
+ exploit is_tail_list_nth_z; eauto.
+ rewrite NTHBI; simpl.
+ intros X; inversion X; subst; clear X NTHBI.
+ destruct (exec_basic _ _ _ _ _) eqn:EXEC_BASIC; next_stuck_cong.
+ destruct s as (rs1 & m1); simpl in *.
+ destruct s2 as (rs2 & m2); simpl in *.
+ assert (BOUNDBBMAX: Ptrofs.unsigned ofs + size bb <= Ptrofs.max_unsigned)
+ by (eapply size_of_blocks_bounds; eauto).
+ exploit header_body_tail_bound; eauto. intros BDYTAIL.
+ exploit exec_basic_simulation; eauto.
+ intros (rs_next' & m_next' & EXEC_INSTR & MI_NEXT).
+ exploit exec_basic_dont_move_PC; eauto. intros AGPC.
+ inversion MI_NEXT as [A B C D E M_NEXT_AGREE RS_NEXT_AGREE ATPC_NEXT PC_OFS_NEXT RS RS'].
+ subst A. subst B. subst C. subst D. subst E.
+ rewrite ATPC in AGPC. symmetry in AGPC, ATPC_NEXT.
+
+ inv MATCHI. symmetry in AGPC0.
+ rewrite ATPC in AGPC0.
+ unfold Val.offset_ptr in AGPC0.
+
+ simpl in FIND_INSTR.
+ (* Execute internal step. *)
+ exploit (Asm.exec_step_internal tge b); eauto.
+ {
+ rewrite Ptrofs.add_unsigned.
+ repeat (rewrite Ptrofs.unsigned_repr); try lia.
+ 2: {
+ assert (BOUNDOFS: 0 <= Ptrofs.unsigned ofs <= Ptrofs.max_unsigned) by eapply Ptrofs.unsigned_range_2.
+ assert (list_length_z (body bb) <= size bb) by eapply body_size_le_block_size.
+ assert (list_length_z (header bb) <= 1). { eapply size_header; eauto. }
+ lia. }
+ try rewrite list_length_z_nat; try split;
+ simpl; rewrite <- !list_length_z_nat;
+ replace (Ptrofs.unsigned ofs + (list_length_z (header bb) + list_length_z (body bb) -
+ list_length_z (a :: li))) with (Ptrofs.unsigned ofs + list_length_z (header bb) +
+ (list_length_z (body bb) - list_length_z (a :: li))) by lia;
+ try assumption; try lia. }
+
+ (* This is our STEP hypothesis. *)
+ intros STEP_NEXT.
+ destruct li as [|a' li]; simpl in *.
+ - (* case of a single instruction in li: this our base case in the induction *)
+ inversion BODY; subst.
+ eexists; split.
+ + apply plus_one. eauto.
+ + constructor; auto.
+ rewrite ATPC_NEXT.
+ apply f_equal.
+ apply f_equal.
+ rewrite bblock_size_aux, list_length_z_cons; simpl.
+ lia.
+ - exploit (IHli b ofs f bb rs1 m_next' (State rs_next' m_next')); congruence || eauto.
+ + exploit is_tail_app_def; eauto.
+ intros (l3 & EQ); rewrite EQ.
+ exploit (is_tail_app_right (l3 ++ a::nil)).
+ rewrite <- app_assoc; simpl; eauto.
+ + constructor; auto.
+ rewrite ATPC_NEXT.
+ apply f_equal.
+ apply f_equal.
+ rewrite! list_length_z_cons; simpl.
+ lia.
+ + intros (s2' & LAST_STEPS & LAST_MATCHS).
+ eexists. split; eauto.
+ eapply plus_left'; eauto.
+Qed.
+
+Lemma exec_body_simulation_plus b ofs f bb rs m s2 rs' m': forall
+ (ATPC: rs PC = Vptr b ofs)
+ (FINDF: Genv.find_funct_ptr ge b = Some (Internal f))
+ (FINDBB: find_bblock (Ptrofs.unsigned ofs) (fn_blocks f) = Some bb)
+ (NEMPTY_BODY: body bb <> nil)
+ (MATCHI: match_internal (list_length_z (header bb)) (State rs m) s2)
+ (BODY: exec_body lk ge (body bb) rs m = Next rs' m'),
+ exists s2', plus Asm.step tge s2 E0 s2'
+ /\ match_internal (size bb - (Z.of_nat (length_opt (exit bb)))) (State rs' m') s2'.
+Proof.
+ intros.
+ exploit exec_body_simulation_plus_gen; eauto.
+ - constructor.
+ - replace (list_length_z (header bb) + list_length_z (body bb) - list_length_z (body bb)) with (list_length_z (header bb)); auto.
+ lia.
+Qed.
+
+Lemma exec_body_simulation_star b ofs f bb rs m s2 rs' m': forall
+ (ATPC: rs PC = Vptr b ofs)
+ (FINDF: Genv.find_funct_ptr ge b = Some (Internal f))
+ (FINDBB: find_bblock (Ptrofs.unsigned ofs) (fn_blocks f) = Some bb)
+ (MATCHI: match_internal (list_length_z (header bb)) (State rs m) s2)
+ (BODY: exec_body lk ge (body bb) rs m = Next rs' m'),
+ exists s2', star Asm.step tge s2 E0 s2'
+ /\ match_internal (size bb - (Z.of_nat (length_opt (exit bb)))) (State rs' m') s2'.
+Proof.
+ intros.
+ destruct (body bb) eqn: Hbb.
+ - simpl in BODY. inv BODY.
+ eexists. split.
+ eapply star_refl; eauto.
+ assert (EQ: (size bb - Z.of_nat (length_opt (exit bb))) = list_length_z (header bb)).
+ { rewrite bblock_size_aux. rewrite Hbb; unfold list_length_z; simpl. lia. }
+ rewrite EQ; eauto.
+ - exploit exec_body_simulation_plus; congruence || eauto.
+ { rewrite Hbb; eauto. }
+ intros (s2' & PLUS & MATCHI').
+ eexists; split; eauto.
+ eapply plus_star; eauto.
+Qed.
+
+Lemma list_nth_z_range_exceeded A (l : list A) n:
+ n >= list_length_z l ->
+ list_nth_z l n = None.
+Proof.
+ intros N.
+ remember (list_nth_z l n) as opt eqn:H. symmetry in H.
+ destruct opt; auto.
+ exploit list_nth_z_range; eauto. lia.
+Qed.
+
+Lemma label_in_header_list lbl a:
+ is_label lbl a = true -> list_length_z (header a) <= 1 -> header a = lbl :: nil.
+Proof.
+ intros.
+ eapply is_label_correct_true in H.
+ destruct (header a).
+ - eapply in_nil in H. contradiction.
+ - rewrite list_length_z_cons in H0.
+ assert (list_length_z l0 >= 0) by eapply list_length_z_pos.
+ assert (list_length_z l0 = 0) by lia.
+ rewrite list_length_z_nat in H2.
+ assert (Datatypes.length l0 = 0%nat) by lia.
+ eapply length_zero_iff_nil in H3. subst.
+ unfold In in H. destruct H.
+ + subst; eauto.
+ + destruct H.
+Qed.
+
+Lemma no_label_in_basic_inst: forall a lbl x,
+ basic_to_instruction a = OK x -> Asm.is_label lbl x = false.
+Proof.
+ intros.
+ destruct a; simpl in *;
+ repeat destruct i;
+ repeat destruct ld; repeat destruct st;
+ simpl in *;
+ try (try destruct_reg_inv; monadInv H; simpl in *; reflexivity).
+Qed.
+
+Lemma label_pos_body bdy: forall c1 c2 z ex lbl
+ (HUNF : unfold_body bdy = OK c2),
+ Asm.label_pos lbl (z + Z.of_nat ((Datatypes.length bdy) + length_opt ex)) c1 = Asm.label_pos lbl (z) ((c2 ++ unfold_exit ex) ++ c1).
+Proof.
+ induction bdy.
+ - intros. inversion HUNF. simpl in *.
+ destruct ex eqn:EQEX.
+ + simpl in *. unfold Asm.is_label. destruct c; simpl; try congruence.
+ destruct i; simpl; try congruence.
+ + simpl in *. ring_simplify (z + 0). auto.
+ - intros. inversion HUNF; clear HUNF. monadInv H0. simpl in *.
+ erewrite no_label_in_basic_inst; eauto. rewrite <- IHbdy; eauto.
+ erewrite Zpos_P_of_succ_nat.
+ apply f_equal2; auto. lia.
+Qed.
+
+Lemma asm_label_pos_header: forall z a x0 x1 lbl
+ (HUNF: unfold_body (body a) = OK x1),
+ Asm.label_pos lbl (z + size a) x0 =
+ Asm.label_pos lbl (z + list_length_z (header a)) ((x1 ++ unfold_exit (exit a)) ++ x0).
+Proof.
+ intros.
+ unfold size.
+ rewrite <- plus_assoc. rewrite Nat2Z.inj_add.
+ rewrite list_length_z_nat.
+ replace (z + (Z.of_nat (Datatypes.length (header a)) + Z.of_nat (Datatypes.length (body a) + length_opt (exit a)))) with (z + Z.of_nat (Datatypes.length (header a)) + Z.of_nat (Datatypes.length (body a) + length_opt (exit a))) by lia.
+ eapply (label_pos_body (body a) x0 x1 (z + Z.of_nat (Datatypes.length (header a))) (exit a) lbl). auto.
+Qed.
+
+Lemma header_size_cons_nil: forall (l0: label) (l1: list label)
+ (HSIZE: list_length_z (l0 :: l1) <= 1),
+ l1 = nil.
+Proof.
+ intros.
+ destruct l1; try congruence. rewrite !list_length_z_cons in HSIZE.
+ assert (list_length_z l1 >= 0) by eapply list_length_z_pos.
+ assert (list_length_z l1 + 1 + 1 >= 2) by lia.
+ assert (2 <= 1) by lia. contradiction H1. lia.
+Qed.
+
+Lemma label_pos_preserved_gen bbs: forall lbl c z
+ (HUNF: unfold bbs = OK c),
+ label_pos lbl z bbs = Asm.label_pos lbl z c.
+Proof.
+ induction bbs.
+ - intros. simpl in *. inversion HUNF. simpl. reflexivity.
+ - intros. simpl in *. monadInv HUNF. unfold unfold_bblock in EQ.
+ destruct (zle _ _); try congruence. monadInv EQ.
+ destruct (is_label _ _) eqn:EQLBL.
+ + erewrite label_in_header_list; eauto.
+ simpl in *. destruct (peq lbl lbl); try congruence.
+ + erewrite IHbbs; eauto.
+ rewrite (asm_label_pos_header z a x0 x1 lbl); auto.
+ unfold is_label in *.
+ destruct (header a).
+ * replace (z + list_length_z (@nil label)) with (z); eauto.
+ unfold list_length_z. simpl. lia.
+ * eapply header_size_cons_nil in l as HL1.
+ subst. simpl in *. destruct (in_dec _ _); try congruence.
+ simpl in *.
+ destruct (peq _ _); try intuition congruence.
+Qed.
+
+Lemma label_pos_preserved f lbl z tf: forall
+ (FINDF: transf_function f = OK tf),
+ label_pos lbl z (fn_blocks f) = Asm.label_pos lbl z (Asm.fn_code tf).
+Proof.
+ intros.
+ eapply label_pos_preserved_gen.
+ unfold transf_function in FINDF. monadInv FINDF.
+ destruct zlt; try congruence. inversion EQ0. eauto.
+Qed.
+
+Lemma goto_label_preserved bb rs1 m1 rs1' m1' rs2 m2 lbl f tf v: forall
+ (FINDF: transf_function f = OK tf)
+ (BOUNDED: size bb <= Ptrofs.max_unsigned)
+ (MATCHI: match_internal (size bb - 1) (State rs1 m1) (State rs2 m2))
+ (HGOTO: goto_label f lbl (incrPC v rs1) m1 = Next rs1' m1'),
+ exists (rs2' : regset) (m2' : mem), Asm.goto_label tf lbl rs2 m2 = Next rs2' m2'
+ /\ match_states (State rs1' m1') (State rs2' m2').
+Proof.
+ intros.
+ unfold goto_label, Asm.goto_label in *.
+ rewrite <- (label_pos_preserved f); auto.
+ inversion MATCHI as [n0 r1 mx1 r2 mx2 EQM EQR EQPC]; subst.
+ destruct label_pos; next_stuck_cong.
+ destruct (incrPC v rs1 PC) eqn:INCRPC; next_stuck_cong.
+ inversion HGOTO; auto. repeat (econstructor; eauto).
+ rewrite <- EQPC.
+ unfold incrPC in *.
+ rewrite !Pregmap.gss in *.
+ destruct (rs1 PC) eqn:EQRS1; simpl in *; try congruence.
+ replace (rs2 # PC <- (Vptr b0 (Ptrofs.repr z))) with ((rs1 # PC <- (Vptr b0 (Ptrofs.add i0 v))) # PC <- (Vptr b (Ptrofs.repr z))); auto.
+ eapply functional_extensionality. intros.
+ destruct (PregEq.eq x PC); subst.
+ rewrite !Pregmap.gss. congruence.
+ rewrite !Pregmap.gso; auto.
+Qed.
+
+Lemma next_inst_incr_pc_preserved bb rs1 m1 rs1' m1' rs2 m2 f tf: forall
+ (FINDF: transf_function f = OK tf)
+ (BOUNDED: size bb <= Ptrofs.max_unsigned)
+ (MATCHI: match_internal (size bb - 1) (State rs1 m1) (State rs2 m2))
+ (NEXT: Next (incrPC (Ptrofs.repr (size bb)) rs1) m2 = Next rs1' m1'),
+ exists (rs2' : regset) (m2' : mem),
+ Next (Asm.nextinstr rs2) m2 = Next rs2' m2'
+ /\ match_states (State rs1' m1') (State rs2' m2').
+Proof.
+ intros; simpl in *; unfold incrPC in NEXT;
+ inv_matchi;
+ assert (size bb >= 1) by eapply bblock_size_pos;
+ assert (0 <= size bb - 1 <= Ptrofs.max_unsigned) by lia;
+ inversion NEXT; subst;
+ eexists; eexists; split; eauto.
+ assert (rs1 # PC <- (Val.offset_ptr (rs1 PC) (Ptrofs.repr (size bb))) = Asm.nextinstr rs2). {
+ unfold Pregmap.set. apply functional_extensionality.
+ intros x. destruct (PregEq.eq x PC).
+ -- unfold Asm.nextinstr. rewrite <- AGPC.
+ rewrite Val.offset_ptr_assoc. rewrite Ptrofs.add_unsigned.
+ rewrite (Ptrofs.unsigned_repr (size bb - 1)); try lia.
+ rewrite Ptrofs.unsigned_one.
+ replace (size bb - 1 + 1) with (size bb) by lia.
+ rewrite e. rewrite Pregmap.gss.
+ reflexivity.
+ -- eapply nextinstr_agree_but_pc; eauto. }
+ rewrite H1. econstructor.
+Qed.
+
+Lemma pc_reg_overwrite: forall (r: ireg) rs1 m1 rs2 m2 bb
+ (MATCHI: match_internal (size bb - 1) (State rs1 m1) (State rs2 m2)),
+ rs2 # PC <- (rs2 r) =
+ (rs1 # PC <- (Val.offset_ptr (rs1 PC) (Ptrofs.repr (size bb)))) # PC <-
+ (rs1 r).
+Proof.
+ intros.
+ unfold Pregmap.set; apply functional_extensionality.
+ intros x; destruct (PregEq.eq x PC) as [X | X]; try discriminate; inv_matchi.
+Qed.
+
+Lemma exec_cfi_simulation:
+ forall bb f tf rs1 m1 rs1' m1' rs2 m2 cfi
+ (SIZE: size bb <= Ptrofs.max_unsigned)
+ (FINDF: transf_function f = OK tf)
+ (* Warning: Asmblock's PC is assumed to be already pointing on the next instruction ! *)
+ (CFI: exec_cfi ge f cfi (incrPC (Ptrofs.repr (size bb)) rs1) m1 = Next rs1' m1')
+ (MATCHI: match_internal (size bb - 1) (State rs1 m1) (State rs2 m2)),
+ exists rs2' m2', Asm.exec_instr tge tf (cf_instruction_to_instruction cfi)
+ rs2 m2 = Next rs2' m2'
+ /\ match_states (State rs1' m1') (State rs2' m2').
+Proof.
+ intros.
+ assert (BBPOS: size bb >= 1) by eapply bblock_size_pos.
+ destruct cfi; inv CFI; simpl.
+ - (* Pb *)
+ exploit goto_label_preserved; eauto.
+ - (* Pbc *)
+ inv_matchi.
+ unfold eval_testcond in *. destruct c;
+ erewrite !incrPC_agree_but_pc in H0; try rewrite <- !AG; try congruence.
+ all:
+ destruct_reg_size;
+ try destruct b eqn:EQB.
+ 1,4,7,10,13,16,19,22,25,28,31,34:
+ exploit goto_label_preserved; eauto.
+ 1,3,5,7,9,11,13,15,17,19,21,23:
+ exploit next_inst_incr_pc_preserved; eauto.
+ all: repeat (econstructor; eauto).
+ - (* Pbl *)
+ eexists; eexists; split; eauto.
+ assert ( ((incrPC (Ptrofs.repr (size bb)) rs1) # X30 <- (incrPC (Ptrofs.repr (size bb)) rs1 PC))
+ # PC <- (Genv.symbol_address ge id Ptrofs.zero)
+ = (rs2 # X30 <- (Val.offset_ptr (rs2 PC) Ptrofs.one))
+ # PC <- (Genv.symbol_address tge id Ptrofs.zero)
+ ) as EQRS. {
+ unfold incrPC. unfold Pregmap.set. simpl. apply functional_extensionality.
+ intros x. destruct (PregEq.eq x PC).
+ * rewrite symbol_addresses_preserved. reflexivity.
+ * destruct (PregEq.eq x X30).
+ -- inv MATCHI. rewrite <- AGPC. rewrite Val.offset_ptr_assoc.
+ unfold Ptrofs.add, Ptrofs.one. repeat (rewrite Ptrofs.unsigned_repr); try lia.
+ replace (size bb - 1 + 1) with (size bb) by lia. reflexivity.
+ -- inv MATCHI; rewrite AG; try assumption; reflexivity.
+ } rewrite EQRS; inv MATCHI; reflexivity.
+ - (* Pbs *)
+ eexists; eexists; split; eauto.
+ assert ( (incrPC (Ptrofs.repr (size bb)) rs1) # PC <-
+ (Genv.symbol_address ge id Ptrofs.zero)
+ = rs2 # PC <- (Genv.symbol_address tge id Ptrofs.zero)
+ ) as EQRS. {
+ unfold incrPC, Pregmap.set. rewrite symbol_addresses_preserved. inv MATCHI.
+ apply functional_extensionality. intros x. destruct (PregEq.eq x PC); auto.
+ } rewrite EQRS; inv MATCHI; reflexivity.
+ - (* Pblr *)
+ eexists; eexists; split; eauto.
+ unfold incrPC. rewrite Pregmap.gss. rewrite Pregmap.gso; try discriminate.
+ assert ( (rs2 # X30 <- (Val.offset_ptr (rs2 PC) Ptrofs.one)) # PC <- (rs2 r)
+ = ((rs1 # PC <- (Val.offset_ptr (rs1 PC) (Ptrofs.repr (size bb))))
+ # X30 <- (Val.offset_ptr (rs1 PC) (Ptrofs.repr (size bb))))
+ # PC <- (rs1 r)
+ ) as EQRS. {
+ unfold Pregmap.set. apply functional_extensionality.
+ intros x; destruct (PregEq.eq x PC) as [X | X].
+ - inv_matchi; rewrite AG; auto.
+ - destruct (PregEq.eq x X30) as [X' | X'].
+ + inversion MATCHI; subst. rewrite <- AGPC.
+ rewrite Val.offset_ptr_assoc. unfold Ptrofs.one.
+ rewrite Ptrofs.add_unsigned. rewrite Ptrofs.unsigned_repr; try lia. rewrite Ptrofs.unsigned_repr; try lia.
+ rewrite Z.sub_add; reflexivity.
+ + inv_matchi.
+ } rewrite EQRS. inv_matchi.
+ - (* Pbr *)
+ eexists; eexists; split; eauto.
+ unfold incrPC. rewrite Pregmap.gso; try discriminate.
+ rewrite (pc_reg_overwrite r rs1 m1' rs2 m2 bb); auto.
+ inv_matchi.
+ - (* Pret *)
+ eexists; eexists; split; eauto.
+ unfold incrPC. rewrite Pregmap.gso; try discriminate.
+ rewrite (pc_reg_overwrite r rs1 m1' rs2 m2 bb); auto.
+ inv_matchi.
+ - (* Pcbnz *)
+ inv_matchi.
+ unfold eval_neg_branch in *.
+ erewrite incrPC_agree_but_pc in H0; try congruence.
+ destruct eval_testzero; next_stuck_cong.
+ destruct b.
+ * exploit next_inst_incr_pc_preserved; eauto.
+ * exploit goto_label_preserved; eauto.
+ - (* Pcbz *)
+ inv_matchi.
+ unfold eval_branch in *.
+ erewrite incrPC_agree_but_pc in H0; try congruence.
+ destruct eval_testzero; next_stuck_cong.
+ destruct b.
+ * exploit goto_label_preserved; eauto.
+ * exploit next_inst_incr_pc_preserved; eauto.
+ - (* Ptbnbz *)
+ inv_matchi.
+ unfold eval_branch in *.
+ erewrite incrPC_agree_but_pc in H0; try congruence.
+ destruct eval_testbit; next_stuck_cong.
+ destruct b.
+ * exploit goto_label_preserved; eauto.
+ * exploit next_inst_incr_pc_preserved; eauto.
+ - (* Ptbz *)
+ inv_matchi.
+ unfold eval_neg_branch in *.
+ erewrite incrPC_agree_but_pc in H0; try congruence.
+ destruct eval_testbit; next_stuck_cong.
+ destruct b.
+ * exploit next_inst_incr_pc_preserved; eauto.
+ * exploit goto_label_preserved; eauto.
+ - (* Pbtbl *)
+ assert (rs2 # X16 <- Vundef r1 = (incrPC (Ptrofs.repr (size bb)) rs1) # X16 <- Vundef r1)
+ as EQUNDEFX16. {
+ unfold incrPC, Pregmap.set.
+ destruct (PregEq.eq r1 X16) as [X16 | X16]; auto.
+ destruct (PregEq.eq r1 PC) as [PC' | PC']; try discriminate.
+ inv MATCHI; rewrite AG; auto.
+ } rewrite <- EQUNDEFX16 in H0.
+ destruct_reg_inv; next_stuck_cong.
+ unfold goto_label, Asm.goto_label in *.
+ rewrite <- (label_pos_preserved f); auto.
+ inversion MATCHI; subst.
+ destruct label_pos; next_stuck_cong.
+ destruct ((incrPC (Ptrofs.repr (size bb)) rs1) # X16 <- Vundef PC) eqn:INCRPC; next_stuck_cong.
+ inversion H0; auto. repeat (econstructor; eauto).
+ rewrite !Pregmap.gso; try congruence.
+ rewrite <- AGPC.
+ unfold incrPC in *.
+ destruct (rs1 PC) eqn:EQRS1; simpl in *; try discriminate.
+ replace ((rs2 # X16 <- Vundef) # PC <- (Vptr b0 (Ptrofs.repr z))) with
+ (((rs1 # PC <- (Vptr b0 (Ptrofs.add i1 (Ptrofs.repr (size bb))))) # X16 <-
+ Vundef) # PC <- (Vptr b (Ptrofs.repr z))); auto.
+ eapply functional_extensionality; intros x.
+ destruct (PregEq.eq x PC); subst.
+ + rewrite Pregmap.gso in INCRPC; try congruence.
+ rewrite Pregmap.gss in INCRPC.
+ rewrite !Pregmap.gss in *; congruence.
+ + rewrite Pregmap.gso; auto.
+ rewrite (Pregmap.gso (i := x) (j := PC)); auto.
+ destruct (PregEq.eq x X16); subst.
+ * rewrite !Pregmap.gss; auto.
+ * rewrite !Pregmap.gso; auto.
+Qed.
+
+Lemma last_instruction_cannot_be_label bb:
+ list_nth_z (header bb) (size bb - 1) = None.
+Proof.
+ assert (list_length_z (header bb) <= size bb - 1). {
+ rewrite bblock_size_aux. generalize (bblock_size_aux_pos bb). lia.
+ }
+ remember (list_nth_z (header bb) (size bb - 1)) as label_opt; destruct label_opt; auto;
+ exploit list_nth_z_range; eauto; lia.
+Qed.
+
+Lemma pc_ptr_exec_step: forall ofs bb b rs m _rs _m
+ (ATPC : rs PC = Vptr b ofs)
+ (MATCHI : match_internal (size bb - 1)
+ {| _rs := rs; _m := m |}
+ {| _rs := _rs; _m := _m |}),
+ _rs PC = Vptr b (Ptrofs.add ofs (Ptrofs.repr (size bb - 1))).
+Proof.
+ intros; inv MATCHI. rewrite <- AGPC; rewrite ATPC; unfold Val.offset_ptr; eauto.
+Qed.
+
+Lemma find_instr_ofs_somei: forall ofs bb f tc asmi rs m _rs _m
+ (BOUNDOFS : Ptrofs.unsigned ofs + size bb <= Ptrofs.max_unsigned)
+ (FIND_INSTR : Asm.find_instr (Ptrofs.unsigned ofs + (size bb - 1)) tc =
+ Some (asmi))
+ (MATCHI : match_internal (size bb - 1)
+ {| _rs := rs; _m := m |}
+ {| _rs := _rs; _m := _m |}),
+ Asm.find_instr (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr (size bb - 1))))
+ (Asm.fn_code {| Asm.fn_sig := fn_sig f; Asm.fn_code := tc |}) =
+ Some (asmi).
+Proof.
+ intros; simpl.
+ replace (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr (size bb - 1))))
+ with (Ptrofs.unsigned ofs + (size bb - 1)); try assumption.
+ generalize (bblock_size_pos bb); generalize (Ptrofs.unsigned_range_2 ofs); intros.
+ unfold Ptrofs.add.
+ rewrite Ptrofs.unsigned_repr. rewrite Ptrofs.unsigned_repr; try lia.
+ rewrite Ptrofs.unsigned_repr; lia.
+Qed.
+
+Lemma eval_builtin_arg_match: forall rs _m _rs a1 b1
+ (AG : forall r : preg, r <> PC -> rs r = _rs r)
+ (EVAL : eval_builtin_arg tge (fun r : dreg => rs r) (rs SP) _m a1 b1),
+ eval_builtin_arg tge _rs (_rs SP) _m (map_builtin_arg DR a1) b1.
+Proof.
+ intros; induction EVAL; simpl in *; try rewrite AG; try rewrite AG in EVAL; try discriminate; try congruence; eauto with barg.
+ econstructor. rewrite <- AG; try discriminate; auto.
+Qed.
+
+Lemma eval_builtin_args_match: forall bb rs m _rs _m args vargs
+ (MATCHI : match_internal (size bb - 1)
+ {| _rs := rs; _m := m |}
+ {| _rs := _rs; _m := _m |})
+ (EVAL : eval_builtin_args tge (fun r : dreg => rs r) (rs SP) m args vargs),
+ eval_builtin_args tge _rs (_rs SP) _m (map (map_builtin_arg DR) args) vargs.
+Proof.
+ intros; inv MATCHI.
+ induction EVAL; subst.
+ - econstructor.
+ - econstructor.
+ + eapply eval_builtin_arg_match; eauto.
+ + eauto.
+Qed.
+
+Lemma pc_both_sides: forall (rs _rs: regset) v
+ (AG : forall r : preg, r <> PC -> rs r = _rs r),
+ rs # PC <- v = _rs # PC <- v.
+Proof.
+ intros; unfold Pregmap.set; apply functional_extensionality; intros y.
+ destruct (PregEq.eq y PC); try rewrite AG; eauto.
+Qed.
+
+Lemma set_buitin_res_sym res: forall vres rs _rs r
+ (NPC: r <> PC)
+ (AG : forall r : preg, r <> PC -> rs r = _rs r),
+ set_res res vres rs r = set_res res vres _rs r.
+Proof.
+ induction res; simpl; intros; unfold Pregmap.set; try rewrite AG; eauto.
+Qed.
+
+Lemma set_builtin_res_dont_move_pc_gen res: forall vres rs _rs v1 v2
+ (HV: v1 = v2)
+ (AG : forall r : preg, r <> PC -> rs r = _rs r),
+ (set_res res vres rs) # PC <- v1 =
+ (set_res res vres _rs) # PC <- v2.
+Proof.
+ intros. rewrite HV. generalize res vres rs _rs AG v2.
+ clear res vres rs _rs AG v1 v2 HV.
+ induction res.
+ - simpl; intros. apply pc_both_sides; intros.
+ unfold Pregmap.set; try rewrite AG; eauto.
+ - simpl; intros; apply pc_both_sides; eauto.
+ - simpl; intros.
+ erewrite IHres2; eauto; intros.
+ eapply set_buitin_res_sym; eauto.
+Qed.
+
+Lemma set_builtin_map_not_pc (res: builtin_res dreg): forall vres rs,
+ set_res (map_builtin_res DR res) vres rs PC = rs PC.
+Proof.
+ induction res.
+ - intros; simpl. unfold Pregmap.set. destruct (PregEq.eq PC x); try congruence.
+ - intros; simpl; congruence.
+ - intros; simpl in *. rewrite IHres2. rewrite IHres1. reflexivity.
+Qed.
+
+Lemma undef_reg_preserved (rl: list mreg): forall rs _rs r
+ (NPC: r <> PC)
+ (AG : forall r : preg, r <> PC -> rs r = _rs r),
+ undef_regs (map preg_of rl) rs r = undef_regs (map preg_of rl) _rs r.
+Proof.
+ induction rl.
+ - simpl; auto.
+ - simpl; intros. erewrite IHrl; eauto.
+ intros. unfold Pregmap.set. destruct (PregEq.eq r0 (preg_of a)); try rewrite AG; eauto.
+Qed.
+
+Lemma undef_regs_other:
+ forall r rl rs,
+ (forall r', In r' rl -> r <> r') ->
+ undef_regs rl rs r = rs r.
+Proof.
+ induction rl; simpl; intros. auto.
+ rewrite IHrl by auto. rewrite Pregmap.gso; auto.
+Qed.
+
+Fixpoint preg_notin (r: preg) (rl: list mreg) : Prop :=
+ match rl with
+ | nil => True
+ | r1 :: nil => r <> preg_of r1
+ | r1 :: rl => r <> preg_of r1 /\ preg_notin r rl
+ end.
+
+<<<<<<< HEAD
+Remark preg_notin_charact:
+ forall r rl,
+ preg_notin r rl <-> (forall mr, In mr rl -> r <> preg_of mr).
+Proof.
+ induction rl; simpl; intros.
+ tauto.
+ destruct rl.
+ simpl. split. intros. intuition congruence. auto.
+ rewrite IHrl. split.
+ intros [A B]. intros. destruct H. congruence. auto.
+ auto.
+=======
+Remark preg_of_not_X29: forall r, negb (mreg_eq r R29) = true -> IR X29 <> preg_of r.
+Proof.
+ intros. change (IR X29) with (preg_of R29). red; intros.
+ exploit preg_of_injective; eauto. intros; subst r; discriminate.
+Qed.
+
+Lemma sp_val': forall ms sp rs, agree ms sp rs -> sp = rs XSP.
+Proof.
+ intros. eapply sp_val; eauto.
+Qed.
+
+(** This is the simulation diagram. We prove it by case analysis on the Mach transition. *)
+
+Theorem step_simulation:
+ forall S1 t S2, Mach.step return_address_offset ge S1 t S2 ->
+ forall S1' (MS: match_states S1 S1'),
+ (exists S2', plus step tge S1' t S2' /\ match_states S2 S2')
+ \/ (measure S2 < measure S1 /\ t = E0 /\ match_states S2 S1')%nat.
+Proof.
+ induction 1; intros; inv MS.
+
+- (* Mlabel *)
+ left; eapply exec_straight_steps; eauto; intros.
+ monadInv TR. econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split. apply agree_nextinstr; auto. simpl; congruence.
+
+- (* Mgetstack *)
+ unfold load_stack in H.
+ exploit Mem.loadv_extends; eauto. intros [v' [A B]].
+ rewrite (sp_val _ _ _ AG) in A.
+ left; eapply exec_straight_steps; eauto. intros. simpl in TR.
+ exploit loadind_correct; eauto with asmgen. intros [rs' [P [Q R]]].
+ exists rs'; split. eauto.
+ split. eapply agree_set_mreg; eauto with asmgen. congruence.
+ simpl; congruence.
+
+- (* Msetstack *)
+ unfold store_stack in H.
+ assert (Val.lessdef (rs src) (rs0 (preg_of src))) by (eapply preg_val; eauto).
+ exploit Mem.storev_extends; eauto. intros [m2' [A B]].
+ left; eapply exec_straight_steps; eauto.
+ rewrite (sp_val _ _ _ AG) in A. intros. simpl in TR.
+ exploit storeind_correct; eauto with asmgen. intros [rs' [P Q]].
+ exists rs'; split. eauto.
+ split. eapply agree_undef_regs; eauto with asmgen.
+ simpl; intros. rewrite Q; auto with asmgen.
+
+- (* Mgetparam *)
+ assert (f0 = f) by congruence; subst f0.
+ unfold load_stack in *.
+ exploit Mem.loadv_extends. eauto. eexact H0. auto.
+ intros [parent' [A B]]. rewrite (sp_val' _ _ _ AG) in A.
+ exploit lessdef_parent_sp; eauto. clear B; intros B; subst parent'.
+ exploit Mem.loadv_extends. eauto. eexact H1. auto.
+ intros [v' [C D]].
+Opaque loadind.
+ left; eapply exec_straight_steps; eauto; intros. monadInv TR.
+ destruct ep.
+(* X30 contains parent *)
+ exploit loadind_correct. eexact EQ.
+ instantiate (2 := rs0). simpl; rewrite DXP; eauto. simpl; congruence.
+ intros [rs1 [P [Q R]]].
+ exists rs1; split. eauto.
+ split. eapply agree_set_mreg. eapply agree_set_mreg; eauto. congruence. auto with asmgen.
+ simpl; intros. rewrite R; auto with asmgen.
+ apply preg_of_not_X29; auto.
+(* X30 does not contain parent *)
+ exploit loadptr_correct. eexact A. simpl; congruence. intros [rs1 [P [Q R]]].
+ exploit loadind_correct. eexact EQ. instantiate (2 := rs1). simpl; rewrite Q. eauto. simpl; congruence.
+ intros [rs2 [S [T U]]].
+ exists rs2; split. eapply exec_straight_trans; eauto.
+ split. eapply agree_set_mreg. eapply agree_set_mreg. eauto. eauto.
+ instantiate (1 := rs1#X29 <- (rs2#X29)). intros.
+ rewrite Pregmap.gso; auto with asmgen.
+ congruence.
+ intros. unfold Pregmap.set. destruct (PregEq.eq r' X29). congruence. auto with asmgen.
+ simpl; intros. rewrite U; auto with asmgen.
+ apply preg_of_not_X29; auto.
+
+- (* Mop *)
+ assert (eval_operation tge sp op (map rs args) m = Some v).
+ { rewrite <- H. apply eval_operation_preserved. exact symbols_preserved. }
+ exploit eval_operation_lessdef. eapply preg_vals; eauto. eauto. eexact H0.
+ intros [v' [A B]]. rewrite (sp_val _ _ _ AG) in A.
+ left; eapply exec_straight_steps; eauto; intros. simpl in TR.
+ exploit transl_op_correct; eauto. intros [rs2 [P [Q R]]].
+ exists rs2; split. eauto. split.
+ apply agree_set_undef_mreg with rs0; auto.
+ apply Val.lessdef_trans with v'; auto.
+ simpl; intros. InvBooleans.
+ rewrite R; auto. apply preg_of_not_X29; auto.
+Local Transparent destroyed_by_op.
+ destruct op; try exact I; simpl; congruence.
+
+- (* Mload *)
+ assert (Op.eval_addressing tge sp addr (map rs args) = Some a).
+ { rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved. }
+ exploit eval_addressing_lessdef. eapply preg_vals; eauto. eexact H1.
+ intros [a' [A B]]. rewrite (sp_val _ _ _ AG) in A.
+ exploit Mem.loadv_extends; eauto. intros [v' [C D]].
+ left; eapply exec_straight_steps; eauto; intros. simpl in TR.
+ exploit transl_load_correct; eauto. intros [rs2 [P [Q R]]].
+ exists rs2; split. eauto.
+ split. eapply agree_set_undef_mreg; eauto. congruence.
+ simpl; congruence.
+
+- (* Mstore *)
+ assert (Op.eval_addressing tge sp addr (map rs args) = Some a).
+ { rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved. }
+ exploit eval_addressing_lessdef. eapply preg_vals; eauto. eexact H1.
+ intros [a' [A B]]. rewrite (sp_val _ _ _ AG) in A.
+ assert (Val.lessdef (rs src) (rs0 (preg_of src))) by (eapply preg_val; eauto).
+ exploit Mem.storev_extends; eauto. intros [m2' [C D]].
+ left; eapply exec_straight_steps; eauto.
+ intros. simpl in TR. exploit transl_store_correct; eauto. intros [rs2 [P Q]].
+ exists rs2; split. eauto.
+ split. eapply agree_undef_regs; eauto with asmgen.
+ simpl; congruence.
+
+- (* Mcall *)
+ assert (f0 = f) by congruence. subst f0.
+ inv AT.
+ assert (NOOV: list_length_z tf.(fn_code) <= Ptrofs.max_unsigned).
+ { eapply transf_function_no_overflow; eauto. }
+ destruct ros as [rf|fid]; simpl in H; monadInv H5.
++ (* Indirect call *)
+ assert (rs rf = Vptr f' Ptrofs.zero).
+ { destruct (rs rf); try discriminate.
+ revert H; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. }
+ assert (rs0 x0 = Vptr f' Ptrofs.zero).
+ { exploit ireg_val; eauto. rewrite H5; intros LD; inv LD; auto. }
+ generalize (code_tail_next_int _ _ _ _ NOOV H6). intro CT1.
+ assert (TCA: transl_code_at_pc ge (Vptr fb (Ptrofs.add ofs Ptrofs.one)) fb f c false tf x).
+ { econstructor; eauto. }
+ exploit return_address_offset_correct; eauto. intros; subst ra.
+ left; econstructor; split.
+ apply plus_one. eapply exec_step_internal. Simpl. rewrite <- H2; simpl; eauto.
+ eapply functions_transl; eauto. eapply find_instr_tail; eauto.
+ simpl. eauto.
+ econstructor; eauto.
+ econstructor; eauto.
+ eapply agree_sp_def; eauto.
+ simpl. eapply agree_exten; eauto. intros. Simpl.
+ Simpl. rewrite <- H2. auto.
++ (* Direct call *)
+ generalize (code_tail_next_int _ _ _ _ NOOV H6). intro CT1.
+ assert (TCA: transl_code_at_pc ge (Vptr fb (Ptrofs.add ofs Ptrofs.one)) fb f c false tf x).
+ econstructor; eauto.
+ exploit return_address_offset_correct; eauto. intros; subst ra.
+ left; econstructor; split.
+ apply plus_one. eapply exec_step_internal. eauto.
+ eapply functions_transl; eauto. eapply find_instr_tail; eauto.
+ simpl. unfold Genv.symbol_address. rewrite symbols_preserved. rewrite H. eauto.
+ econstructor; eauto.
+ econstructor; eauto.
+ eapply agree_sp_def; eauto.
+ simpl. eapply agree_exten; eauto. intros. Simpl.
+ Simpl. rewrite <- H2. auto.
+
+- (* Mtailcall *)
+ assert (f0 = f) by congruence. subst f0.
+ inversion AT; subst.
+ assert (NOOV: list_length_z tf.(fn_code) <= Ptrofs.max_unsigned).
+ { eapply transf_function_no_overflow; eauto. }
+ exploit Mem.loadv_extends. eauto. eexact H1. auto. simpl. intros [parent' [A B]].
+ destruct ros as [rf|fid]; simpl in H; monadInv H7.
++ (* Indirect call *)
+ assert (rs rf = Vptr f' Ptrofs.zero).
+ { destruct (rs rf); try discriminate.
+ revert H; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. }
+ assert (rs0 x0 = Vptr f' Ptrofs.zero).
+ { exploit ireg_val; eauto. rewrite H7; intros LD; inv LD; auto. }
+ exploit make_epilogue_correct; eauto. intros (rs1 & m1 & U & V & W & X & Y & Z).
+ exploit exec_straight_steps_2; eauto using functions_transl.
+ intros (ofs' & P & Q).
+ left; econstructor; split.
+ (* execution *)
+ eapply plus_right'. eapply exec_straight_exec; eauto.
+ econstructor. eexact P. eapply functions_transl; eauto. eapply find_instr_tail. eexact Q.
+ simpl. reflexivity.
+ traceEq.
+ (* match states *)
+ econstructor; eauto.
+ apply agree_set_other; auto with asmgen.
+ Simpl. rewrite Z by (rewrite <- (ireg_of_eq _ _ EQ1); eauto with asmgen). assumption.
++ (* Direct call *)
+ exploit make_epilogue_correct; eauto. intros (rs1 & m1 & U & V & W & X & Y & Z).
+ exploit exec_straight_steps_2; eauto using functions_transl.
+ intros (ofs' & P & Q).
+ left; econstructor; split.
+ (* execution *)
+ eapply plus_right'. eapply exec_straight_exec; eauto.
+ econstructor. eexact P. eapply functions_transl; eauto. eapply find_instr_tail. eexact Q.
+ simpl. reflexivity.
+ traceEq.
+ (* match states *)
+ econstructor; eauto.
+ apply agree_set_other; auto with asmgen.
+ Simpl. unfold Genv.symbol_address. rewrite symbols_preserved. rewrite H. auto.
+
+- (* Mbuiltin *)
+ inv AT. monadInv H4.
+ exploit functions_transl; eauto. intro FN.
+ generalize (transf_function_no_overflow _ _ H3); intro NOOV.
+ exploit builtin_args_match; eauto. intros [vargs' [P Q]].
+ exploit external_call_mem_extends; eauto.
+ intros [vres' [m2' [A [B [C D]]]]].
+ left. econstructor; split. apply plus_one.
+ eapply exec_step_builtin. eauto. eauto.
+ eapply find_instr_tail; eauto.
+ erewrite <- sp_val by eauto.
+ eapply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved.
+ eapply external_call_symbols_preserved; eauto. apply senv_preserved.
+ eauto.
+ econstructor; eauto.
+ instantiate (2 := tf); instantiate (1 := x).
+ unfold nextinstr. rewrite Pregmap.gss.
+ rewrite set_res_other. rewrite undef_regs_other.
+ rewrite <- H1. simpl. econstructor; eauto.
+ eapply code_tail_next_int; eauto.
+ simpl; intros. destruct H4. congruence. destruct H4. congruence.
+ exploit list_in_map_inv; eauto. intros (mr & U & V). subst.
+ auto with asmgen.
+ auto with asmgen.
+ apply agree_nextinstr. eapply agree_set_res; auto.
+ eapply agree_undef_regs; eauto. intros.
+ simpl. rewrite undef_regs_other_2; auto. Simpl.
+ congruence.
+
+- (* Mgoto *)
+ assert (f0 = f) by congruence. subst f0.
+ inv AT. monadInv H4.
+ exploit find_label_goto_label; eauto. intros [tc' [rs' [GOTO [AT2 INV]]]].
+ left; exists (State rs' m'); split.
+ apply plus_one. econstructor; eauto.
+ eapply functions_transl; eauto.
+ eapply find_instr_tail; eauto.
+ simpl; eauto.
+ econstructor; eauto.
+ eapply agree_exten; eauto with asmgen.
+ congruence.
+
+- (* Mcond true *)
+ assert (f0 = f) by congruence. subst f0.
+ exploit eval_condition_lessdef. eapply preg_vals; eauto. eauto. eauto. intros EC.
+ left; eapply exec_straight_opt_steps_goto; eauto.
+ intros. simpl in TR.
+ exploit transl_cond_branch_correct; eauto. intros (rs' & jmp & A & B & C).
+ exists jmp; exists k; exists rs'.
+ split. eexact A.
+ split. apply agree_exten with rs0; auto with asmgen.
+ exact B.
+
+- (* Mcond false *)
+ exploit eval_condition_lessdef. eapply preg_vals; eauto. eauto. eauto. intros EC.
+ left; eapply exec_straight_steps; eauto. intros. simpl in TR.
+ exploit transl_cond_branch_correct; eauto. intros (rs' & jmp & A & B & C).
+ econstructor; split.
+ eapply exec_straight_opt_right. eexact A. apply exec_straight_one. eexact B. auto.
+ split. apply agree_exten with rs0; auto. intros. Simpl.
+ simpl; congruence.
+
+- (* Mjumptable *)
+ assert (f0 = f) by congruence. subst f0.
+ inv AT. monadInv H6.
+ exploit functions_transl; eauto. intro FN.
+ generalize (transf_function_no_overflow _ _ H5); intro NOOV.
+ exploit find_label_goto_label. eauto. eauto.
+ instantiate (2 := rs0#X16 <- Vundef).
+ Simpl. eauto.
+ eauto.
+ intros [tc' [rs' [A [B C]]]].
+ exploit ireg_val; eauto. rewrite H. intros LD; inv LD.
+ left; econstructor; split.
+ apply plus_one. econstructor; eauto.
+ eapply find_instr_tail; eauto.
+ simpl. Simpl. rewrite <- H9. unfold Mach.label in H0; unfold label; rewrite H0. eexact A.
+ econstructor; eauto.
+ eapply agree_undef_regs; eauto.
+ simpl. intros. rewrite C; auto with asmgen. Simpl.
+ congruence.
+
+- (* Mreturn *)
+ assert (f0 = f) by congruence. subst f0.
+ inversion AT; subst. simpl in H6; monadInv H6.
+ assert (NOOV: list_length_z tf.(fn_code) <= Ptrofs.max_unsigned).
+ eapply transf_function_no_overflow; eauto.
+ exploit make_epilogue_correct; eauto. intros (rs1 & m1 & U & V & W & X & Y & Z).
+ exploit exec_straight_steps_2; eauto using functions_transl.
+ intros (ofs' & P & Q).
+ left; econstructor; split.
+ (* execution *)
+ eapply plus_right'. eapply exec_straight_exec; eauto.
+ econstructor. eexact P. eapply functions_transl; eauto. eapply find_instr_tail. eexact Q.
+ simpl. reflexivity.
+ traceEq.
+ (* match states *)
+ econstructor; eauto.
+ apply agree_set_other; auto with asmgen.
+
+- (* internal function *)
+
+ exploit functions_translated; eauto. intros [tf [A B]]. monadInv B.
+ generalize EQ; intros EQ'. monadInv EQ'.
+ destruct (zlt Ptrofs.max_unsigned (list_length_z x0.(fn_code))); inversion EQ1. clear EQ1. subst x0.
+ unfold store_stack in *.
+ exploit Mem.alloc_extends. eauto. eauto. apply Z.le_refl. apply Z.le_refl.
+ intros [m1' [C D]].
+ exploit Mem.storev_extends. eexact D. eexact H1. eauto. eauto.
+ intros [m2' [F G]].
+ simpl chunk_of_type in F.
+ exploit Mem.storev_extends. eexact G. eexact H2. eauto. eauto.
+ intros [m3' [P Q]].
+ change (chunk_of_type Tptr) with Mint64 in *.
+ (* Execution of function prologue *)
+ monadInv EQ0. rewrite transl_code'_transl_code in EQ1.
+ set (tfbody := Pallocframe (fn_stacksize f) (fn_link_ofs f) ::
+ storeptr RA XSP (fn_retaddr_ofs f) x0) in *.
+ set (tf := {| fn_sig := Mach.fn_sig f; fn_code := tfbody |}) in *.
+ set (rs2 := nextinstr (rs0#X29 <- (parent_sp s) #SP <- sp #X16 <- Vundef)).
+ exploit (storeptr_correct tge tf XSP (fn_retaddr_ofs f) RA x0 m2' m3' rs2).
+ simpl preg_of_iregsp. change (rs2 X30) with (rs0 X30). rewrite ATLR.
+ change (rs2 X2) with sp. eexact P.
+ simpl; congruence. congruence.
+ intros (rs3 & U & V).
+ assert (EXEC_PROLOGUE:
+ exec_straight tge tf
+ tf.(fn_code) rs0 m'
+ x0 rs3 m3').
+ { change (fn_code tf) with tfbody; unfold tfbody.
+ apply exec_straight_step with rs2 m2'.
+ unfold exec_instr. rewrite C. fold sp.
+ rewrite <- (sp_val _ _ _ AG). rewrite F. reflexivity.
+ reflexivity.
+ eexact U. }
+ exploit exec_straight_steps_2; eauto using functions_transl. lia. constructor.
+ intros (ofs' & X & Y).
+ left; exists (State rs3 m3'); split.
+ eapply exec_straight_steps_1; eauto. lia. constructor.
+ econstructor; eauto.
+ rewrite X; econstructor; eauto.
+ apply agree_exten with rs2; eauto with asmgen.
+ unfold rs2.
+ apply agree_nextinstr. apply agree_set_other; auto with asmgen.
+ apply agree_change_sp with (parent_sp s).
+ apply agree_undef_regs with rs0. auto.
+Local Transparent destroyed_at_function_entry. simpl.
+ simpl; intros; Simpl.
+ unfold sp; congruence.
+ intros. rewrite V by auto with asmgen. reflexivity.
+
+- (* external function *)
+ exploit functions_translated; eauto.
+ intros [tf [A B]]. simpl in B. inv B.
+ exploit extcall_arguments_match; eauto.
+ intros [args' [C D]].
+ exploit external_call_mem_extends; eauto.
+ intros [res' [m2' [P [Q [R S]]]]].
+ left; econstructor; split.
+ apply plus_one. eapply exec_step_external; eauto.
+ eapply external_call_symbols_preserved; eauto. apply senv_preserved.
+ econstructor; eauto.
+ unfold loc_external_result. apply agree_set_other; auto. apply agree_set_pair; auto.
+ apply agree_undef_caller_save_regs; auto.
+
+- (* return *)
+ inv STACKS. simpl in *.
+ right. split. lia. split. auto.
+ rewrite <- ATPC in H5.
+ econstructor; eauto. congruence.
+>>>>>>> master
+Qed.
+
+Lemma undef_regs_other_2:
+ forall r rl rs,
+ preg_notin r rl ->
+ undef_regs (map preg_of rl) rs r = rs r.
+Proof.
+ intros. apply undef_regs_other. intros.
+ exploit list_in_map_inv; eauto. intros [mr [A B]]. subst.
+ rewrite preg_notin_charact in H. auto.
+Qed.
+
+Lemma exec_exit_simulation_plus b ofs f bb s2 t rs m rs' m': forall
+ (FINDF: Genv.find_funct_ptr ge b = Some (Internal f))
+ (FINDBB: find_bblock (Ptrofs.unsigned ofs) (fn_blocks f) = Some bb)
+ (NEMPTY_EXIT: exit bb <> None)
+ (MATCHI: match_internal (size bb - Z.of_nat (length_opt (exit bb))) (State rs m) s2)
+ (EXIT: exec_exit ge f (Ptrofs.repr (size bb)) rs m (exit bb) t rs' m')
+ (ATPC: rs PC = Vptr b ofs),
+ plus Asm.step tge s2 t (State rs' m').
+Proof.
+ intros.
+ exploit internal_functions_unfold; eauto.
+ intros (tc & FINDtf & TRANStf & _).
+
+ exploit (find_instr_bblock (size bb - 1)); eauto.
+ { generalize (bblock_size_pos bb). lia. }
+ intros (i' & NTH & FIND_INSTR).
+
+ inv NTH.
+ + rewrite last_instruction_cannot_be_label in *. discriminate.
+ + destruct (exit bb) as [ctrl |] eqn:NEMPTY_EXIT'. 2: { contradiction. }
+ rewrite bblock_size_aux in *. rewrite NEMPTY_EXIT' in *. simpl in *.
+ (* XXX: Is there a better way to simplify this expression i.e. automatically? *)
+ replace (list_length_z (header bb) + list_length_z (body bb) + 1 - 1 -
+ list_length_z (header bb)) with (list_length_z (body bb)) in H by lia.
+ rewrite list_nth_z_range_exceeded in H; try lia. discriminate.
+ + assert (Ptrofs.unsigned ofs + size bb <= Ptrofs.max_unsigned). {
+ eapply size_of_blocks_bounds; eauto.
+ }
+ assert (size bb <= Ptrofs.max_unsigned). { generalize (Ptrofs.unsigned_range_2 ofs); lia. }
+ destruct cfi.
+ * (* control flow instruction *)
+ destruct s2.
+ rewrite H in EXIT. (* exit bb is a cfi *)
+ inv EXIT.
+ rewrite H in MATCHI. simpl in MATCHI.
+ exploit internal_functions_translated; eauto.
+ rewrite FINDtf.
+ intros (tf & FINDtf' & TRANSf). inversion FINDtf'; subst; clear FINDtf'.
+ exploit exec_cfi_simulation; eauto.
+ (* extract exec_cfi_simulation's conclusion as separate hypotheses *)
+ intros (rs2' & m2' & EXECI & MATCHS); rewrite MATCHS.
+ apply plus_one.
+ eapply Asm.exec_step_internal; eauto.
+ - eapply pc_ptr_exec_step; eauto.
+ - eapply find_instr_ofs_somei; eauto.
+ * (* builtin *)
+ destruct s2.
+ rewrite H in EXIT.
+ rewrite H in MATCHI. simpl in MATCHI.
+ simpl in FIND_INSTR.
+ inversion EXIT.
+ apply plus_one.
+ eapply external_call_symbols_preserved in H10; try (apply senv_preserved).
+ eapply eval_builtin_args_preserved in H6; try (apply symbols_preserved).
+ eapply Asm.exec_step_builtin; eauto.
+ - eapply pc_ptr_exec_step; eauto.
+ - eapply find_instr_ofs_somei; eauto.
+ - eapply eval_builtin_args_match; eauto.
+ - inv MATCHI; eauto.
+ - inv MATCHI.
+ unfold Asm.nextinstr, incrPC.
+ assert (HPC: Val.offset_ptr (rs PC) (Ptrofs.repr (size bb))
+ = Val.offset_ptr (_rs PC) Ptrofs.one).
+ { rewrite <- AGPC. rewrite ATPC. unfold Val.offset_ptr.
+ rewrite Ptrofs.add_assoc. unfold Ptrofs.add.
+ assert (BBPOS: size bb >= 1) by eapply bblock_size_pos.
+ rewrite (Ptrofs.unsigned_repr (size bb - 1)); try lia.
+ rewrite Ptrofs.unsigned_one.
+ replace (size bb - 1 + 1) with (size bb) by lia.
+ reflexivity. }
+ apply set_builtin_res_dont_move_pc_gen.
+ -- erewrite !set_builtin_map_not_pc.
+ erewrite !undef_regs_other.
+ rewrite HPC; auto.
+ all: intros; simpl in *; destruct H3 as [HX16 | [HX30 | HDES]]; subst; try discriminate;
+ exploit list_in_map_inv; eauto; intros [mr [A B]]; subst; discriminate.
+ -- intros. eapply undef_reg_preserved; eauto.
+ intros. destruct (PregEq.eq X16 r0); destruct (PregEq.eq X30 r0); subst.
+ rewrite Pregmap.gso, Pregmap.gss; try congruence.
+ do 2 (rewrite Pregmap.gso, Pregmap.gss; try discriminate; auto).
+ rewrite 2Pregmap.gss; auto.
+ rewrite !Pregmap.gso; auto.
+Qed.
+
+Lemma exec_exit_simulation_star b ofs f bb s2 t rs m rs' m': forall
+ (FINDF: Genv.find_funct_ptr ge b = Some (Internal f))
+ (FINDBB: find_bblock (Ptrofs.unsigned ofs) (fn_blocks f) = Some bb)
+ (MATCHI: match_internal (size bb - Z.of_nat (length_opt (exit bb))) (State rs m) s2)
+ (EXIT: exec_exit ge f (Ptrofs.repr (size bb)) rs m (exit bb) t rs' m')
+ (ATPC: rs PC = Vptr b ofs),
+ star Asm.step tge s2 t (State rs' m').
+Proof.
+ intros.
+ destruct (exit bb) eqn: Hex.
+ - eapply plus_star.
+ eapply exec_exit_simulation_plus; try rewrite Hex; congruence || eauto.
+ - inv MATCHI.
+ inv EXIT.
+ assert (X: rs2 = incrPC (Ptrofs.repr (size bb)) rs). {
+ unfold incrPC. unfold Pregmap.set.
+ apply functional_extensionality. intros x.
+ destruct (PregEq.eq x PC) as [X|].
+ - rewrite X. rewrite <- AGPC. simpl.
+ replace (size bb - 0) with (size bb) by lia. reflexivity.
+ - rewrite AG; try assumption. reflexivity.
+ }
+ destruct X.
+ subst; eapply star_refl; eauto.
+Qed.
+
+Lemma exec_bblock_simulation b ofs f bb t rs m rs' m': forall
+ (ATPC: rs PC = Vptr b ofs)
+ (FINDF: Genv.find_funct_ptr ge b = Some (Internal f))
+ (FINDBB: find_bblock (Ptrofs.unsigned ofs) (fn_blocks f) = Some bb)
+ (EXECBB: exec_bblock lk ge f bb rs m t rs' m'),
+ plus Asm.step tge (State rs m) t (State rs' m').
+Proof.
+ intros; destruct EXECBB as (rs1 & m1 & BODY & CTL).
+ exploit exec_header_simulation; eauto.
+ intros (s0 & STAR & MATCH0).
+ eapply star_plus_trans; traceEq || eauto.
+ destruct (bblock_non_empty bb).
+ - (* body bb <> nil *)
+ exploit exec_body_simulation_plus; eauto.
+ intros (s1 & PLUS & MATCH1).
+ eapply plus_star_trans; traceEq || eauto.
+ eapply exec_exit_simulation_star; eauto.
+ erewrite <- exec_body_dont_move_PC; eauto.
+ - (* exit bb <> None *)
+ exploit exec_body_simulation_star; eauto.
+ intros (s1 & STAR1 & MATCH1).
+ eapply star_plus_trans; traceEq || eauto.
+ eapply exec_exit_simulation_plus; eauto.
+ erewrite <- exec_body_dont_move_PC; eauto.
+Qed.
+
+Lemma step_simulation s t s':
+ Asmblock.step lk ge s t s' -> plus Asm.step tge s t s'.
+Proof.
+ intros STEP.
+ inv STEP; simpl; exploit functions_translated; eauto;
+ intros (tf0 & FINDtf & TRANSf);
+ monadInv TRANSf.
+ - (* internal step *) eapply exec_bblock_simulation; eauto.
+ - (* external step *)
+ apply plus_one.
+ exploit external_call_symbols_preserved; eauto. apply senv_preserved.
+ intros ?.
+ eapply Asm.exec_step_external; eauto.
+Qed.
+
+Lemma transf_program_correct:
+ forward_simulation (Asmblock.semantics lk prog) (Asm.semantics tprog).
+Proof.
+ eapply forward_simulation_plus.
+ - apply senv_preserved.
+ - eexact transf_initial_states.
+ - eexact transf_final_states.
+ - unfold match_states.
+ simpl; intros; subst; eexists; split; eauto.
+ eapply step_simulation; eauto.
+Qed.
+
+End PRESERVATION.
+
+End Asmblock_PRESERVATION.
+
+
+Local Open Scope linking_scope.
+
+Definition block_passes :=
+ mkpass Machblockgenproof.match_prog
+ ::: mkpass Asmblockgenproof.match_prog
+ ::: mkpass PostpassSchedulingproof.match_prog
+ ::: mkpass Asmblock_PRESERVATION.match_prog
+ ::: pass_nil _.
+
+Definition match_prog := pass_match (compose_passes block_passes).
+
+Lemma transf_program_match:
+ forall p tp, Asmgen.transf_program p = OK tp -> match_prog p tp.
+Proof.
+ intros p tp H.
+ unfold Asmgen.transf_program in H. apply bind_inversion in H. destruct H.
+ inversion_clear H. apply bind_inversion in H1. destruct H1.
+ inversion_clear H.
+ unfold Compopts.time in *. remember (Machblockgen.transf_program p) as mbp.
+ unfold match_prog; simpl.
+ exists mbp; split. apply Machblockgenproof.transf_program_match; auto.
+ exists x; split. apply Asmblockgenproof.transf_program_match; auto.
+ exists x0; split. apply PostpassSchedulingproof.transf_program_match; auto.
+ exists tp; split. apply Asmblock_PRESERVATION.transf_program_match; auto. auto.
+Qed.
+
+(** Return Address Offset *)
+
+Definition return_address_offset: Mach.function -> Mach.code -> ptrofs -> Prop :=
+ Machblockgenproof.Mach_return_address_offset (Asmblockgenproof.return_address_offset).
+
+Lemma return_address_exists:
+ forall f sg ros c, is_tail (Mach.Mcall sg ros :: c) f.(Mach.fn_code) ->
+ exists ra, return_address_offset f c ra.
+Proof.
+ intros; unfold return_address_offset; eapply Machblockgenproof.Mach_return_address_exists; eauto.
+ intros; eapply Asmblockgenproof.return_address_exists; eauto.
+Qed.
+
+Section PRESERVATION.
+
+Variable prog: Mach.program.
+Variable tprog: Asm.program.
+Hypothesis TRANSF: match_prog prog tprog.
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+Theorem transf_program_correct:
+ forward_simulation (Mach.semantics return_address_offset prog) (Asm.semantics tprog).
+Proof.
+ unfold match_prog in TRANSF. simpl in TRANSF.
+ inv TRANSF. inv H. inv H1. inv H. inv H2. inv H. inv H3. inv H.
+ eapply compose_forward_simulations.
+ { exploit Machblockgenproof.transf_program_correct; eauto. }
+
+ eapply compose_forward_simulations.
+ + apply Asmblockgenproof.transf_program_correct; eauto.
+ { intros.
+ unfold Genv.symbol_address.
+ erewrite <- PostpassSchedulingproof.symbols_preserved; eauto.
+ erewrite Asmblock_PRESERVATION.symbol_high_low; eauto.
+ reflexivity.
+ }
+ + eapply compose_forward_simulations.
+ - apply PostpassSchedulingproof.transf_program_correct; eauto.
+ - apply Asmblock_PRESERVATION.transf_program_correct; eauto.
+Qed.
+
+End PRESERVATION.
+
+Instance TransfAsm: TransfLink match_prog := pass_match_link (compose_passes block_passes).
+
+(*******************************************)
+(* Stub actually needed by driver/Compiler *)
+
+Module Asmgenproof0.
+
+Definition return_address_offset := return_address_offset.
+
+End Asmgenproof0.
diff --git a/aarch64/TO_MERGE/Asmgenproof1.v b/aarch64/TO_MERGE/Asmgenproof1.v
new file mode 100644
index 00000000..93c1f1ed
--- /dev/null
+++ b/aarch64/TO_MERGE/Asmgenproof1.v
@@ -0,0 +1,1836 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, Collège de France and INRIA Paris *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(** Correctness proof for AArch64 code generation: auxiliary results. *)
+
+Require Import Recdef Coqlib Zwf Zbits.
+Require Import Maps Errors AST Integers Floats Values Memory Globalenvs.
+Require Import Op Locations Mach Asm Conventions.
+Require Import Asmgen.
+Require Import Asmgenproof0.
+
+Local Transparent Archi.ptr64.
+
+(** Properties of registers *)
+
+Lemma preg_of_iregsp_not_PC: forall r, preg_of_iregsp r <> PC.
+Proof.
+ destruct r; simpl; congruence.
+Qed.
+Global Hint Resolve preg_of_iregsp_not_PC: asmgen.
+
+Lemma preg_of_not_X16: forall r, preg_of r <> X16.
+Proof.
+ destruct r; simpl; congruence.
+Qed.
+
+Lemma ireg_of_not_X16: forall r x, ireg_of r = OK x -> x <> X16.
+Proof.
+ unfold ireg_of; intros. destruct (preg_of r) eqn:E; inv H.
+ red; intros; subst x. elim (preg_of_not_X16 r); auto.
+Qed.
+
+Lemma ireg_of_not_X16': forall r x, ireg_of r = OK x -> IR x <> IR X16.
+Proof.
+ intros. apply ireg_of_not_X16 in H. congruence.
+Qed.
+
+Global Hint Resolve preg_of_not_X16 ireg_of_not_X16 ireg_of_not_X16': asmgen.
+
+(** Useful simplification tactic *)
+
+
+Ltac Simplif :=
+ ((rewrite nextinstr_inv by eauto with asmgen)
+ || (rewrite nextinstr_inv1 by eauto with asmgen)
+ || (rewrite Pregmap.gss)
+ || (rewrite nextinstr_pc)
+ || (rewrite Pregmap.gso by eauto with asmgen)); auto with asmgen.
+
+Ltac Simpl := repeat Simplif.
+
+(** * Correctness of ARM constructor functions *)
+
+Section CONSTRUCTORS.
+
+Variable ge: genv.
+Variable fn: function.
+
+(** Decomposition of integer literals *)
+
+Inductive wf_decomposition: list (Z * Z) -> Prop :=
+ | wf_decomp_nil:
+ wf_decomposition nil
+ | wf_decomp_cons: forall m n p l,
+ n = Zzero_ext 16 m -> 0 <= p -> wf_decomposition l ->
+ wf_decomposition ((n, p) :: l).
+
+Lemma decompose_int_wf:
+ forall N n p, 0 <= p -> wf_decomposition (decompose_int N n p).
+Proof.
+Local Opaque Zzero_ext.
+ induction N as [ | N]; simpl; intros.
+- constructor.
+- set (frag := Zzero_ext 16 (Z.shiftr n p)) in *. destruct (Z.eqb frag 0).
++ apply IHN. lia.
++ econstructor. reflexivity. lia. apply IHN; lia.
+Qed.
+
+Fixpoint recompose_int (accu: Z) (l: list (Z * Z)) : Z :=
+ match l with
+ | nil => accu
+ | (n, p) :: l => recompose_int (Zinsert accu n p 16) l
+ end.
+
+Lemma decompose_int_correct:
+ forall N n p accu,
+ 0 <= p ->
+ (forall i, p <= i -> Z.testbit accu i = false) ->
+ (forall i, 0 <= i < p + Z.of_nat N * 16 ->
+ Z.testbit (recompose_int accu (decompose_int N n p)) i =
+ if zlt i p then Z.testbit accu i else Z.testbit n i).
+Proof.
+ induction N as [ | N]; intros until accu; intros PPOS ABOVE i RANGE.
+- simpl. rewrite zlt_true; auto. extlia.
+- rewrite inj_S in RANGE. simpl.
+ set (frag := Zzero_ext 16 (Z.shiftr n p)).
+ assert (FRAG: forall i, p <= i < p + 16 -> Z.testbit n i = Z.testbit frag (i - p)).
+ { unfold frag; intros. rewrite Zzero_ext_spec by lia. rewrite zlt_true by lia.
+ rewrite Z.shiftr_spec by lia. f_equal; lia. }
+ destruct (Z.eqb_spec frag 0).
++ rewrite IHN.
+* destruct (zlt i p). rewrite zlt_true by lia. auto.
+ destruct (zlt i (p + 16)); auto.
+ rewrite ABOVE by lia. rewrite FRAG by lia. rewrite e, Z.testbit_0_l. auto.
+* lia.
+* intros; apply ABOVE; lia.
+* extlia.
++ simpl. rewrite IHN.
+* destruct (zlt i (p + 16)).
+** rewrite Zinsert_spec by lia. unfold proj_sumbool.
+ rewrite zlt_true by lia.
+ destruct (zlt i p).
+ rewrite zle_false by lia. auto.
+ rewrite zle_true by lia. simpl. symmetry; apply FRAG; lia.
+** rewrite Z.ldiff_spec, Z.shiftl_spec by lia.
+ change 65535 with (two_p 16 - 1). rewrite Ztestbit_two_p_m1 by lia.
+ rewrite zlt_false by lia. rewrite zlt_false by lia. apply andb_true_r.
+* lia.
+* intros. rewrite Zinsert_spec by lia. unfold proj_sumbool.
+ rewrite zle_true by lia. rewrite zlt_false by lia. simpl.
+ apply ABOVE. lia.
+* extlia.
+Qed.
+
+Corollary decompose_int_eqmod: forall N n,
+ eqmod (two_power_nat (N * 16)%nat) (recompose_int 0 (decompose_int N n 0)) n.
+Proof.
+ intros; apply eqmod_same_bits; intros.
+ rewrite decompose_int_correct. apply zlt_false; lia.
+ lia. intros; apply Z.testbit_0_l. extlia.
+Qed.
+
+Corollary decompose_notint_eqmod: forall N n,
+ eqmod (two_power_nat (N * 16)%nat)
+ (Z.lnot (recompose_int 0 (decompose_int N (Z.lnot n) 0))) n.
+Proof.
+ intros; apply eqmod_same_bits; intros.
+ rewrite Z.lnot_spec, decompose_int_correct.
+ rewrite zlt_false by lia. rewrite Z.lnot_spec by lia. apply negb_involutive.
+ lia. intros; apply Z.testbit_0_l. extlia. lia.
+Qed.
+
+Lemma negate_decomposition_wf:
+ forall l, wf_decomposition l -> wf_decomposition (negate_decomposition l).
+Proof.
+ induction 1; simpl; econstructor; auto.
+ instantiate (1 := (Z.lnot m)).
+ apply equal_same_bits; intros.
+ rewrite H. change 65535 with (two_p 16 - 1).
+ rewrite Z.lxor_spec, !Zzero_ext_spec, Z.lnot_spec, Ztestbit_two_p_m1 by lia.
+ destruct (zlt i 16).
+ apply xorb_true_r.
+ auto.
+Qed.
+
+Lemma Zinsert_eqmod:
+ forall n x1 x2 y p l, 0 <= p -> 0 <= l ->
+ eqmod (two_power_nat n) x1 x2 ->
+ eqmod (two_power_nat n) (Zinsert x1 y p l) (Zinsert x2 y p l).
+Proof.
+ intros. apply eqmod_same_bits; intros. rewrite ! Zinsert_spec by lia.
+ destruct (zle p i && zlt i (p + l)); auto.
+ apply same_bits_eqmod with n; auto.
+Qed.
+
+Lemma Zinsert_0_l:
+ forall y p l,
+ 0 <= p -> 0 <= l ->
+ Z.shiftl (Zzero_ext l y) p = Zinsert 0 (Zzero_ext l y) p l.
+Proof.
+ intros. apply equal_same_bits; intros.
+ rewrite Zinsert_spec by lia. unfold proj_sumbool.
+ destruct (zlt i p); [rewrite zle_false by lia|rewrite zle_true by lia]; simpl.
+- rewrite Z.testbit_0_l, Z.shiftl_spec_low by auto. auto.
+- rewrite Z.shiftl_spec by lia.
+ destruct (zlt i (p + l)); auto.
+ rewrite Zzero_ext_spec, zlt_false, Z.testbit_0_l by lia. auto.
+Qed.
+
+Lemma recompose_int_negated:
+ forall l, wf_decomposition l ->
+ forall accu, recompose_int (Z.lnot accu) (negate_decomposition l) = Z.lnot (recompose_int accu l).
+Proof.
+ induction 1; intros accu; simpl.
+- auto.
+- rewrite <- IHwf_decomposition. f_equal. apply equal_same_bits; intros.
+ rewrite Z.lnot_spec, ! Zinsert_spec, Z.lxor_spec, Z.lnot_spec by lia.
+ unfold proj_sumbool.
+ destruct (zle p i); simpl; auto.
+ destruct (zlt i (p + 16)); simpl; auto.
+ change 65535 with (two_p 16 - 1).
+ rewrite Ztestbit_two_p_m1 by lia. rewrite zlt_true by lia.
+ apply xorb_true_r.
+Qed.
+
+Lemma exec_loadimm_k_w:
+ forall (rd: ireg) k m l,
+ wf_decomposition l ->
+ forall (rs: regset) accu,
+ rs#rd = Vint (Int.repr accu) ->
+ exists rs',
+ exec_straight_opt ge fn (loadimm_k W rd l k) rs m k rs' m
+ /\ rs'#rd = Vint (Int.repr (recompose_int accu l))
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ induction 1; intros rs accu ACCU; simpl.
+- exists rs; split. apply exec_straight_opt_refl. auto.
+- destruct (IHwf_decomposition
+ (nextinstr (rs#rd <- (insert_in_int rs#rd n p 16)))
+ (Zinsert accu n p 16))
+ as (rs' & P & Q & R).
+ Simpl. rewrite ACCU. simpl. f_equal. apply Int.eqm_samerepr.
+ apply Zinsert_eqmod. auto. lia. apply Int.eqm_sym; apply Int.eqm_unsigned_repr.
+ exists rs'; split.
+ eapply exec_straight_opt_step_opt. simpl; eauto. auto. exact P.
+ split. exact Q. intros; Simpl. rewrite R by auto. Simpl.
+Qed.
+
+Lemma exec_loadimm_z_w:
+ forall rd l k rs m,
+ wf_decomposition l ->
+ exists rs',
+ exec_straight ge fn (loadimm_z W rd l k) rs m k rs' m
+ /\ rs'#rd = Vint (Int.repr (recompose_int 0 l))
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold loadimm_z; destruct 1.
+- econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split. Simpl.
+ intros; Simpl.
+- set (accu0 := Zinsert 0 n p 16).
+ set (rs1 := nextinstr (rs#rd <- (Vint (Int.repr accu0)))).
+ destruct (exec_loadimm_k_w rd k m l H1 rs1 accu0) as (rs2 & P & Q & R); auto.
+ unfold rs1; Simpl.
+ exists rs2; split.
+ eapply exec_straight_opt_step; eauto.
+ simpl. unfold rs1. do 5 f_equal. unfold accu0. rewrite H. apply Zinsert_0_l; lia.
+ reflexivity.
+ split. exact Q.
+ intros. rewrite R by auto. unfold rs1; Simpl.
+Qed.
+
+Lemma exec_loadimm_n_w:
+ forall rd l k rs m,
+ wf_decomposition l ->
+ exists rs',
+ exec_straight ge fn (loadimm_n W rd l k) rs m k rs' m
+ /\ rs'#rd = Vint (Int.repr (Z.lnot (recompose_int 0 l)))
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold loadimm_n; destruct 1.
+- econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split. Simpl.
+ intros; Simpl.
+- set (accu0 := Z.lnot (Zinsert 0 n p 16)).
+ set (rs1 := nextinstr (rs#rd <- (Vint (Int.repr accu0)))).
+ destruct (exec_loadimm_k_w rd k m (negate_decomposition l)
+ (negate_decomposition_wf l H1)
+ rs1 accu0) as (rs2 & P & Q & R).
+ unfold rs1; Simpl.
+ exists rs2; split.
+ eapply exec_straight_opt_step; eauto.
+ simpl. unfold rs1. do 5 f_equal.
+ unfold accu0. f_equal. rewrite H. apply Zinsert_0_l; lia.
+ reflexivity.
+ split. unfold accu0 in Q; rewrite recompose_int_negated in Q by auto. exact Q.
+ intros. rewrite R by auto. unfold rs1; Simpl.
+Qed.
+
+Lemma exec_loadimm32:
+ forall rd n k rs m,
+ exists rs',
+ exec_straight ge fn (loadimm32 rd n k) rs m k rs' m
+ /\ rs'#rd = Vint n
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold loadimm32, loadimm; intros.
+ destruct (is_logical_imm32 n).
+- econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split. Simpl. rewrite Int.repr_unsigned, Int.or_zero_l; auto.
+ intros; Simpl.
+- set (dz := decompose_int 2%nat (Int.unsigned n) 0).
+ set (dn := decompose_int 2%nat (Z.lnot (Int.unsigned n)) 0).
+ assert (A: Int.repr (recompose_int 0 dz) = n).
+ { transitivity (Int.repr (Int.unsigned n)).
+ apply Int.eqm_samerepr. apply decompose_int_eqmod.
+ apply Int.repr_unsigned. }
+ assert (B: Int.repr (Z.lnot (recompose_int 0 dn)) = n).
+ { transitivity (Int.repr (Int.unsigned n)).
+ apply Int.eqm_samerepr. apply decompose_notint_eqmod.
+ apply Int.repr_unsigned. }
+ destruct Nat.leb.
++ rewrite <- A. apply exec_loadimm_z_w. apply decompose_int_wf; lia.
++ rewrite <- B. apply exec_loadimm_n_w. apply decompose_int_wf; lia.
+Qed.
+
+Lemma exec_loadimm_k_x:
+ forall (rd: ireg) k m l,
+ wf_decomposition l ->
+ forall (rs: regset) accu,
+ rs#rd = Vlong (Int64.repr accu) ->
+ exists rs',
+ exec_straight_opt ge fn (loadimm_k X rd l k) rs m k rs' m
+ /\ rs'#rd = Vlong (Int64.repr (recompose_int accu l))
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ induction 1; intros rs accu ACCU; simpl.
+- exists rs; split. apply exec_straight_opt_refl. auto.
+- destruct (IHwf_decomposition
+ (nextinstr (rs#rd <- (insert_in_long rs#rd n p 16)))
+ (Zinsert accu n p 16))
+ as (rs' & P & Q & R).
+ Simpl. rewrite ACCU. simpl. f_equal. apply Int64.eqm_samerepr.
+ apply Zinsert_eqmod. auto. lia. apply Int64.eqm_sym; apply Int64.eqm_unsigned_repr.
+ exists rs'; split.
+ eapply exec_straight_opt_step_opt. simpl; eauto. auto. exact P.
+ split. exact Q. intros; Simpl. rewrite R by auto. Simpl.
+Qed.
+
+Lemma exec_loadimm_z_x:
+ forall rd l k rs m,
+ wf_decomposition l ->
+ exists rs',
+ exec_straight ge fn (loadimm_z X rd l k) rs m k rs' m
+ /\ rs'#rd = Vlong (Int64.repr (recompose_int 0 l))
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold loadimm_z; destruct 1.
+- econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split. Simpl.
+ intros; Simpl.
+- set (accu0 := Zinsert 0 n p 16).
+ set (rs1 := nextinstr (rs#rd <- (Vlong (Int64.repr accu0)))).
+ destruct (exec_loadimm_k_x rd k m l H1 rs1 accu0) as (rs2 & P & Q & R); auto.
+ unfold rs1; Simpl.
+ exists rs2; split.
+ eapply exec_straight_opt_step; eauto.
+ simpl. unfold rs1. do 5 f_equal. unfold accu0. rewrite H. apply Zinsert_0_l; lia.
+ reflexivity.
+ split. exact Q.
+ intros. rewrite R by auto. unfold rs1; Simpl.
+Qed.
+
+Lemma exec_loadimm_n_x:
+ forall rd l k rs m,
+ wf_decomposition l ->
+ exists rs',
+ exec_straight ge fn (loadimm_n X rd l k) rs m k rs' m
+ /\ rs'#rd = Vlong (Int64.repr (Z.lnot (recompose_int 0 l)))
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold loadimm_n; destruct 1.
+- econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split. Simpl.
+ intros; Simpl.
+- set (accu0 := Z.lnot (Zinsert 0 n p 16)).
+ set (rs1 := nextinstr (rs#rd <- (Vlong (Int64.repr accu0)))).
+ destruct (exec_loadimm_k_x rd k m (negate_decomposition l)
+ (negate_decomposition_wf l H1)
+ rs1 accu0) as (rs2 & P & Q & R).
+ unfold rs1; Simpl.
+ exists rs2; split.
+ eapply exec_straight_opt_step; eauto.
+ simpl. unfold rs1. do 5 f_equal.
+ unfold accu0. f_equal. rewrite H. apply Zinsert_0_l; lia.
+ reflexivity.
+ split. unfold accu0 in Q; rewrite recompose_int_negated in Q by auto. exact Q.
+ intros. rewrite R by auto. unfold rs1; Simpl.
+Qed.
+
+Lemma exec_loadimm64:
+ forall rd n k rs m,
+ exists rs',
+ exec_straight ge fn (loadimm64 rd n k) rs m k rs' m
+ /\ rs'#rd = Vlong n
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold loadimm64, loadimm; intros.
+ destruct (is_logical_imm64 n).
+- econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split. Simpl. rewrite Int64.repr_unsigned, Int64.or_zero_l; auto.
+ intros; Simpl.
+- set (dz := decompose_int 4%nat (Int64.unsigned n) 0).
+ set (dn := decompose_int 4%nat (Z.lnot (Int64.unsigned n)) 0).
+ assert (A: Int64.repr (recompose_int 0 dz) = n).
+ { transitivity (Int64.repr (Int64.unsigned n)).
+ apply Int64.eqm_samerepr. apply decompose_int_eqmod.
+ apply Int64.repr_unsigned. }
+ assert (B: Int64.repr (Z.lnot (recompose_int 0 dn)) = n).
+ { transitivity (Int64.repr (Int64.unsigned n)).
+ apply Int64.eqm_samerepr. apply decompose_notint_eqmod.
+ apply Int64.repr_unsigned. }
+ destruct Nat.leb.
++ rewrite <- A. apply exec_loadimm_z_x. apply decompose_int_wf; lia.
++ rewrite <- B. apply exec_loadimm_n_x. apply decompose_int_wf; lia.
+Qed.
+
+(** Add immediate *)
+
+Lemma exec_addimm_aux_32:
+ forall (insn: iregsp -> iregsp -> Z -> instruction) (sem: val -> val -> val),
+ (forall rd r1 n rs m,
+ exec_instr ge fn (insn rd r1 n) rs m =
+ Next (nextinstr (rs#rd <- (sem rs#r1 (Vint (Int.repr n))))) m) ->
+ (forall v n1 n2, sem (sem v (Vint n1)) (Vint n2) = sem v (Vint (Int.add n1 n2))) ->
+ forall rd r1 n k rs m,
+ exists rs',
+ exec_straight ge fn (addimm_aux insn rd r1 (Int.unsigned n) k) rs m k rs' m
+ /\ rs'#rd = sem rs#r1 (Vint n)
+ /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r.
+Proof.
+ intros insn sem SEM ASSOC; intros. unfold addimm_aux.
+ set (nlo := Zzero_ext 12 (Int.unsigned n)). set (nhi := Int.unsigned n - nlo).
+ assert (E: Int.unsigned n = nhi + nlo) by (unfold nhi; lia).
+ rewrite <- (Int.repr_unsigned n).
+ destruct (Z.eqb_spec nhi 0); [|destruct (Z.eqb_spec nlo 0)].
+- econstructor; split. apply exec_straight_one. apply SEM. Simpl.
+ split. Simpl. do 3 f_equal; lia.
+ intros; Simpl.
+- econstructor; split. apply exec_straight_one. apply SEM. Simpl.
+ split. Simpl. do 3 f_equal; lia.
+ intros; Simpl.
+- econstructor; split. eapply exec_straight_two.
+ apply SEM. apply SEM. Simpl. Simpl.
+ split. Simpl. rewrite ASSOC. do 2 f_equal. apply Int.eqm_samerepr.
+ rewrite E. auto with ints.
+ intros; Simpl.
+Qed.
+
+Lemma exec_addimm32:
+ forall rd r1 n k rs m,
+ r1 <> X16 ->
+ exists rs',
+ exec_straight ge fn (addimm32 rd r1 n k) rs m k rs' m
+ /\ rs'#rd = Val.add rs#r1 (Vint n)
+ /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r.
+Proof.
+ intros. unfold addimm32. set (nn := Int.neg n).
+ destruct (Int.eq n (Int.zero_ext 24 n)); [| destruct (Int.eq nn (Int.zero_ext 24 nn))].
+- apply exec_addimm_aux_32 with (sem := Val.add). auto. intros; apply Val.add_assoc.
+- rewrite <- Val.sub_opp_add.
+ apply exec_addimm_aux_32 with (sem := Val.sub). auto.
+ intros. rewrite ! Val.sub_add_opp, Val.add_assoc. rewrite Int.neg_add_distr. auto.
+- destruct (Int.lt n Int.zero).
++ rewrite <- Val.sub_opp_add; fold nn.
+ edestruct (exec_loadimm32 X16 nn) as (rs1 & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. eapply exec_straight_one. simpl; eauto. auto.
+ split. Simpl. rewrite B, C; eauto with asmgen.
+ intros; Simpl.
++ edestruct (exec_loadimm32 X16 n) as (rs1 & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. eapply exec_straight_one. simpl; eauto. auto.
+ split. Simpl. rewrite B, C; eauto with asmgen.
+ intros; Simpl.
+Qed.
+
+Lemma exec_addimm_aux_64:
+ forall (insn: iregsp -> iregsp -> Z -> instruction) (sem: val -> val -> val),
+ (forall rd r1 n rs m,
+ exec_instr ge fn (insn rd r1 n) rs m =
+ Next (nextinstr (rs#rd <- (sem rs#r1 (Vlong (Int64.repr n))))) m) ->
+ (forall v n1 n2, sem (sem v (Vlong n1)) (Vlong n2) = sem v (Vlong (Int64.add n1 n2))) ->
+ forall rd r1 n k rs m,
+ exists rs',
+ exec_straight ge fn (addimm_aux insn rd r1 (Int64.unsigned n) k) rs m k rs' m
+ /\ rs'#rd = sem rs#r1 (Vlong n)
+ /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r.
+Proof.
+ intros insn sem SEM ASSOC; intros. unfold addimm_aux.
+ set (nlo := Zzero_ext 12 (Int64.unsigned n)). set (nhi := Int64.unsigned n - nlo).
+ assert (E: Int64.unsigned n = nhi + nlo) by (unfold nhi; lia).
+ rewrite <- (Int64.repr_unsigned n).
+ destruct (Z.eqb_spec nhi 0); [|destruct (Z.eqb_spec nlo 0)].
+- econstructor; split. apply exec_straight_one. apply SEM. Simpl.
+ split. Simpl. do 3 f_equal; lia.
+ intros; Simpl.
+- econstructor; split. apply exec_straight_one. apply SEM. Simpl.
+ split. Simpl. do 3 f_equal; lia.
+ intros; Simpl.
+- econstructor; split. eapply exec_straight_two.
+ apply SEM. apply SEM. Simpl. Simpl.
+ split. Simpl. rewrite ASSOC. do 2 f_equal. apply Int64.eqm_samerepr.
+ rewrite E. auto with ints.
+ intros; Simpl.
+Qed.
+
+Lemma exec_addimm64:
+ forall rd r1 n k rs m,
+ preg_of_iregsp r1 <> X16 ->
+ exists rs',
+ exec_straight ge fn (addimm64 rd r1 n k) rs m k rs' m
+ /\ rs'#rd = Val.addl rs#r1 (Vlong n)
+ /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r.
+Proof.
+ intros.
+ unfold addimm64. set (nn := Int64.neg n).
+ destruct (Int64.eq n (Int64.zero_ext 24 n)); [| destruct (Int64.eq nn (Int64.zero_ext 24 nn))].
+- apply exec_addimm_aux_64 with (sem := Val.addl). auto. intros; apply Val.addl_assoc.
+- rewrite <- Val.subl_opp_addl.
+ apply exec_addimm_aux_64 with (sem := Val.subl). auto.
+ intros. rewrite ! Val.subl_addl_opp, Val.addl_assoc. rewrite Int64.neg_add_distr. auto.
+- destruct (Int64.lt n Int64.zero).
++ rewrite <- Val.subl_opp_addl; fold nn.
+ edestruct (exec_loadimm64 X16 nn) as (rs1 & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. eapply exec_straight_one. simpl; eauto. Simpl.
+ split. Simpl. rewrite B, C; eauto with asmgen. simpl. rewrite Int64.shl'_zero. auto.
+ intros; Simpl.
++ edestruct (exec_loadimm64 X16 n) as (rs1 & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. eapply exec_straight_one. simpl; eauto. Simpl.
+ split. Simpl. rewrite B, C; eauto with asmgen. simpl. rewrite Int64.shl'_zero. auto.
+ intros; Simpl.
+Qed.
+
+(** Logical immediate *)
+
+Lemma exec_logicalimm32:
+ forall (insn1: ireg -> ireg0 -> Z -> instruction)
+ (insn2: ireg -> ireg0 -> ireg -> shift_op -> instruction)
+ (sem: val -> val -> val),
+ (forall rd r1 n rs m,
+ exec_instr ge fn (insn1 rd r1 n) rs m =
+ Next (nextinstr (rs#rd <- (sem rs##r1 (Vint (Int.repr n))))) m) ->
+ (forall rd r1 r2 s rs m,
+ exec_instr ge fn (insn2 rd r1 r2 s) rs m =
+ Next (nextinstr (rs#rd <- (sem rs##r1 (eval_shift_op_int rs#r2 s)))) m) ->
+ forall rd r1 n k rs m,
+ r1 <> X16 ->
+ exists rs',
+ exec_straight ge fn (logicalimm32 insn1 insn2 rd r1 n k) rs m k rs' m
+ /\ rs'#rd = sem rs#r1 (Vint n)
+ /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r.
+Proof.
+ intros until sem; intros SEM1 SEM2; intros. unfold logicalimm32.
+ destruct (is_logical_imm32 n).
+- econstructor; split.
+ apply exec_straight_one. apply SEM1. reflexivity.
+ split. Simpl. rewrite Int.repr_unsigned; auto. intros; Simpl.
+- edestruct (exec_loadimm32 X16 n) as (rs1 & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A.
+ apply exec_straight_one. apply SEM2. reflexivity.
+ split. Simpl. f_equal; auto. apply C; auto with asmgen.
+ intros; Simpl.
+Qed.
+
+Lemma exec_logicalimm64:
+ forall (insn1: ireg -> ireg0 -> Z -> instruction)
+ (insn2: ireg -> ireg0 -> ireg -> shift_op -> instruction)
+ (sem: val -> val -> val),
+ (forall rd r1 n rs m,
+ exec_instr ge fn (insn1 rd r1 n) rs m =
+ Next (nextinstr (rs#rd <- (sem rs###r1 (Vlong (Int64.repr n))))) m) ->
+ (forall rd r1 r2 s rs m,
+ exec_instr ge fn (insn2 rd r1 r2 s) rs m =
+ Next (nextinstr (rs#rd <- (sem rs###r1 (eval_shift_op_long rs#r2 s)))) m) ->
+ forall rd r1 n k rs m,
+ r1 <> X16 ->
+ exists rs',
+ exec_straight ge fn (logicalimm64 insn1 insn2 rd r1 n k) rs m k rs' m
+ /\ rs'#rd = sem rs#r1 (Vlong n)
+ /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r.
+Proof.
+ intros until sem; intros SEM1 SEM2; intros. unfold logicalimm64.
+ destruct (is_logical_imm64 n).
+- econstructor; split.
+ apply exec_straight_one. apply SEM1. reflexivity.
+ split. Simpl. rewrite Int64.repr_unsigned. auto. intros; Simpl.
+- edestruct (exec_loadimm64 X16 n) as (rs1 & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A.
+ apply exec_straight_one. apply SEM2. reflexivity.
+ split. Simpl. f_equal; auto. apply C; auto with asmgen.
+ intros; Simpl.
+Qed.
+
+(** Load address of symbol *)
+
+Lemma exec_loadsymbol: forall rd s ofs k rs m,
+ rd <> X16 \/ SelectOp.symbol_is_relocatable s = false ->
+ exists rs',
+ exec_straight ge fn (loadsymbol rd s ofs k) rs m k rs' m
+ /\ rs'#rd = Genv.symbol_address ge s ofs
+ /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold loadsymbol; intros. destruct (SelectOp.symbol_is_relocatable s).
+- predSpec Ptrofs.eq Ptrofs.eq_spec ofs Ptrofs.zero.
++ subst ofs. econstructor; split.
+ apply exec_straight_one; [simpl; eauto | reflexivity].
+ split. Simpl. intros; Simpl.
++ exploit exec_addimm64. instantiate (1 := rd). simpl. destruct H; congruence.
+ intros (rs1 & A & B & C).
+ econstructor; split.
+ econstructor. simpl; eauto. auto. eexact A.
+ split. simpl in B; rewrite B. Simpl.
+ rewrite <- Genv.shift_symbol_address_64 by auto.
+ rewrite Ptrofs.add_zero_l, Ptrofs.of_int64_to_int64 by auto. auto.
+ intros. rewrite C by auto. Simpl.
+- econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split. Simpl. rewrite symbol_high_low; auto.
+ intros; Simpl.
+Qed.
+
+(** Shifted operands *)
+
+Remark transl_shift_not_none:
+ forall s a, transl_shift s a <> SOnone.
+Proof.
+ destruct s; intros; simpl; congruence.
+Qed.
+
+Remark or_zero_eval_shift_op_int:
+ forall v s, s <> SOnone -> Val.or (Vint Int.zero) (eval_shift_op_int v s) = eval_shift_op_int v s.
+Proof.
+ intros; destruct s; try congruence; destruct v; auto; simpl;
+ destruct (Int.ltu n Int.iwordsize); auto; rewrite Int.or_zero_l; auto.
+Qed.
+
+Remark or_zero_eval_shift_op_long:
+ forall v s, s <> SOnone -> Val.orl (Vlong Int64.zero) (eval_shift_op_long v s) = eval_shift_op_long v s.
+Proof.
+ intros; destruct s; try congruence; destruct v; auto; simpl;
+ destruct (Int.ltu n Int64.iwordsize'); auto; rewrite Int64.or_zero_l; auto.
+Qed.
+
+Remark add_zero_eval_shift_op_long:
+ forall v s, s <> SOnone -> Val.addl (Vlong Int64.zero) (eval_shift_op_long v s) = eval_shift_op_long v s.
+Proof.
+ intros; destruct s; try congruence; destruct v; auto; simpl;
+ destruct (Int.ltu n Int64.iwordsize'); auto; rewrite Int64.add_zero_l; auto.
+Qed.
+
+Lemma transl_eval_shift: forall s v (a: amount32),
+ eval_shift_op_int v (transl_shift s a) = eval_shift s v a.
+Proof.
+ intros. destruct s; simpl; auto.
+Qed.
+
+Lemma transl_eval_shift': forall s v (a: amount32),
+ Val.or (Vint Int.zero) (eval_shift_op_int v (transl_shift s a)) = eval_shift s v a.
+Proof.
+ intros. rewrite or_zero_eval_shift_op_int by (apply transl_shift_not_none).
+ apply transl_eval_shift.
+Qed.
+
+Lemma transl_eval_shiftl: forall s v (a: amount64),
+ eval_shift_op_long v (transl_shift s a) = eval_shiftl s v a.
+Proof.
+ intros. destruct s; simpl; auto.
+Qed.
+
+Lemma transl_eval_shiftl': forall s v (a: amount64),
+ Val.orl (Vlong Int64.zero) (eval_shift_op_long v (transl_shift s a)) = eval_shiftl s v a.
+Proof.
+ intros. rewrite or_zero_eval_shift_op_long by (apply transl_shift_not_none).
+ apply transl_eval_shiftl.
+Qed.
+
+Lemma transl_eval_shiftl'': forall s v (a: amount64),
+ Val.addl (Vlong Int64.zero) (eval_shift_op_long v (transl_shift s a)) = eval_shiftl s v a.
+Proof.
+ intros. rewrite add_zero_eval_shift_op_long by (apply transl_shift_not_none).
+ apply transl_eval_shiftl.
+Qed.
+
+(** Zero- and Sign- extensions *)
+
+Lemma exec_move_extended_base: forall rd r1 ex k rs m,
+ exists rs',
+ exec_straight ge fn (move_extended_base rd r1 ex k) rs m k rs' m
+ /\ rs' rd = match ex with Xsgn32 => Val.longofint rs#r1 | Xuns32 => Val.longofintu rs#r1 end
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold move_extended_base; destruct ex; econstructor;
+ (split; [apply exec_straight_one; [simpl;eauto|auto] | split; [Simpl|intros;Simpl]]).
+Qed.
+
+Lemma exec_move_extended: forall rd r1 ex (a: amount64) k rs m,
+ exists rs',
+ exec_straight ge fn (move_extended rd r1 ex a k) rs m k rs' m
+ /\ rs' rd = Op.eval_extend ex rs#r1 a
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold move_extended; intros. predSpec Int.eq Int.eq_spec a Int.zero.
+- exploit (exec_move_extended_base rd r1 ex). intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split. unfold Op.eval_extend. rewrite H. rewrite B.
+ destruct ex, (rs r1); simpl; auto; rewrite Int64.shl'_zero; auto.
+ auto.
+- Local Opaque Val.addl.
+ exploit (exec_move_extended_base rd r1 ex). intros (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. apply exec_straight_one.
+ unfold exec_instr. change (SOlsl a) with (transl_shift Slsl a). rewrite transl_eval_shiftl''. eauto. auto.
+ split. Simpl. rewrite B. auto.
+ intros; Simpl.
+Qed.
+
+Lemma exec_arith_extended:
+ forall (sem: val -> val -> val)
+ (insnX: iregsp -> iregsp -> ireg -> extend_op -> instruction)
+ (insnS: ireg -> ireg0 -> ireg -> shift_op -> instruction),
+ (forall rd r1 r2 x rs m,
+ exec_instr ge fn (insnX rd r1 r2 x) rs m =
+ Next (nextinstr (rs#rd <- (sem rs#r1 (eval_extend rs#r2 x)))) m) ->
+ (forall rd r1 r2 s rs m,
+ exec_instr ge fn (insnS rd r1 r2 s) rs m =
+ Next (nextinstr (rs#rd <- (sem rs###r1 (eval_shift_op_long rs#r2 s)))) m) ->
+ forall (rd r1 r2: ireg) (ex: extension) (a: amount64) (k: code) rs m,
+ r1 <> X16 ->
+ exists rs',
+ exec_straight ge fn (arith_extended insnX insnS rd r1 r2 ex a k) rs m k rs' m
+ /\ rs'#rd = sem rs#r1 (Op.eval_extend ex rs#r2 a)
+ /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r.
+Proof.
+ intros sem insnX insnS EX ES; intros. unfold arith_extended. destruct (Int.ltu a (Int.repr 5)).
+- econstructor; split.
+ apply exec_straight_one. rewrite EX; eauto. auto.
+ split. Simpl. f_equal. destruct ex; auto.
+ intros; Simpl.
+- exploit (exec_move_extended_base X16 r2 ex). intros (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. apply exec_straight_one.
+ rewrite ES. eauto. auto.
+ split. Simpl. unfold ir0x. rewrite C by eauto with asmgen. f_equal.
+ rewrite B. destruct ex; auto.
+ intros; Simpl.
+Qed.
+
+(** Extended right shift *)
+
+Lemma exec_shrx32: forall (rd r1: ireg) (n: int) k v (rs: regset) m,
+ Val.shrx rs#r1 (Vint n) = Some v ->
+ r1 <> X16 ->
+ exists rs',
+ exec_straight ge fn (shrx32 rd r1 n k) rs m k rs' m
+ /\ rs'#rd = v
+ /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold shrx32; intros. apply Val.shrx_shr_2 in H.
+ destruct (Int.eq n Int.zero) eqn:E.
+- econstructor; split. apply exec_straight_one; [simpl;eauto|auto].
+ split. Simpl. subst v; auto. intros; Simpl.
+- econstructor; split. eapply exec_straight_three.
+ unfold exec_instr. rewrite or_zero_eval_shift_op_int by congruence. eauto.
+ simpl; eauto.
+ unfold exec_instr. rewrite or_zero_eval_shift_op_int by congruence. eauto.
+ auto. auto. auto.
+ split. subst v; Simpl. intros; Simpl.
+Qed.
+
+Lemma exec_shrx64: forall (rd r1: ireg) (n: int) k v (rs: regset) m,
+ Val.shrxl rs#r1 (Vint n) = Some v ->
+ r1 <> X16 ->
+ exists rs',
+ exec_straight ge fn (shrx64 rd r1 n k) rs m k rs' m
+ /\ rs'#rd = v
+ /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold shrx64; intros. apply Val.shrxl_shrl_2 in H.
+ destruct (Int.eq n Int.zero) eqn:E.
+- econstructor; split. apply exec_straight_one; [simpl;eauto|auto].
+ split. Simpl. subst v; auto. intros; Simpl.
+- econstructor; split. eapply exec_straight_three.
+ unfold exec_instr. rewrite or_zero_eval_shift_op_long by congruence. eauto.
+ simpl; eauto.
+ unfold exec_instr. rewrite or_zero_eval_shift_op_long by congruence. eauto.
+ auto. auto. auto.
+ split. subst v; Simpl. intros; Simpl.
+Qed.
+
+(** Condition bits *)
+
+Lemma compare_int_spec: forall rs v1 v2 m,
+ let rs' := compare_int rs v1 v2 m in
+ rs'#CN = (Val.negative (Val.sub v1 v2))
+ /\ rs'#CZ = (Val.cmpu (Mem.valid_pointer m) Ceq v1 v2)
+ /\ rs'#CC = (Val.cmpu (Mem.valid_pointer m) Cge v1 v2)
+ /\ rs'#CV = (Val.sub_overflow v1 v2).
+Proof.
+ intros; unfold rs'; auto.
+Qed.
+
+Lemma eval_testcond_compare_sint: forall c v1 v2 b rs m,
+ Val.cmp_bool c v1 v2 = Some b ->
+ eval_testcond (cond_for_signed_cmp c) (compare_int rs v1 v2 m) = Some b.
+Proof.
+ intros. generalize (compare_int_spec rs v1 v2 m).
+ set (rs' := compare_int rs v1 v2 m). intros (B & C & D & E).
+ unfold eval_testcond; rewrite B, C, D, E.
+ destruct v1; try discriminate; destruct v2; try discriminate.
+ simpl in H; inv H.
+ unfold Val.cmpu; simpl. destruct c; simpl.
+- destruct (Int.eq i i0); auto.
+- destruct (Int.eq i i0); auto.
+- rewrite Int.lt_sub_overflow. destruct (Int.lt i i0); auto.
+- rewrite Int.lt_sub_overflow, Int.not_lt.
+ destruct (Int.eq i i0), (Int.lt i i0); auto.
+- rewrite Int.lt_sub_overflow, (Int.lt_not i).
+ destruct (Int.eq i i0), (Int.lt i i0); auto.
+- rewrite Int.lt_sub_overflow. destruct (Int.lt i i0); auto.
+Qed.
+
+Lemma eval_testcond_compare_uint: forall c v1 v2 b rs m,
+ Val.cmpu_bool (Mem.valid_pointer m) c v1 v2 = Some b ->
+ eval_testcond (cond_for_unsigned_cmp c) (compare_int rs v1 v2 m) = Some b.
+Proof.
+ intros. generalize (compare_int_spec rs v1 v2 m).
+ set (rs' := compare_int rs v1 v2 m). intros (B & C & D & E).
+ unfold eval_testcond; rewrite B, C, D, E.
+ destruct v1; try discriminate; destruct v2; try discriminate.
+ simpl in H; inv H.
+ unfold Val.cmpu; simpl. destruct c; simpl.
+- destruct (Int.eq i i0); auto.
+- destruct (Int.eq i i0); auto.
+- destruct (Int.ltu i i0); auto.
+- rewrite (Int.not_ltu i). destruct (Int.eq i i0), (Int.ltu i i0); auto.
+- rewrite (Int.ltu_not i). destruct (Int.eq i i0), (Int.ltu i i0); auto.
+- destruct (Int.ltu i i0); auto.
+Qed.
+
+Lemma compare_long_spec: forall rs v1 v2 m,
+ let rs' := compare_long rs v1 v2 m in
+ rs'#CN = (Val.negativel (Val.subl v1 v2))
+ /\ rs'#CZ = (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Ceq v1 v2))
+ /\ rs'#CC = (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Cge v1 v2))
+ /\ rs'#CV = (Val.subl_overflow v1 v2).
+Proof.
+ intros; unfold rs'; auto.
+Qed.
+
+Remark int64_sub_overflow:
+ forall x y,
+ Int.xor (Int.repr (Int64.unsigned (Int64.sub_overflow x y Int64.zero)))
+ (Int.repr (Int64.unsigned (Int64.negative (Int64.sub x y)))) =
+ (if Int64.lt x y then Int.one else Int.zero).
+Proof.
+ intros.
+ transitivity (Int.repr (Int64.unsigned (if Int64.lt x y then Int64.one else Int64.zero))).
+ rewrite <- (Int64.lt_sub_overflow x y).
+ unfold Int64.sub_overflow, Int64.negative.
+ set (s := Int64.signed x - Int64.signed y - Int64.signed Int64.zero).
+ destruct (zle Int64.min_signed s && zle s Int64.max_signed);
+ destruct (Int64.lt (Int64.sub x y) Int64.zero);
+ auto.
+ destruct (Int64.lt x y); auto.
+Qed.
+
+Lemma eval_testcond_compare_slong: forall c v1 v2 b rs m,
+ Val.cmpl_bool c v1 v2 = Some b ->
+ eval_testcond (cond_for_signed_cmp c) (compare_long rs v1 v2 m) = Some b.
+Proof.
+ intros. generalize (compare_long_spec rs v1 v2 m).
+ set (rs' := compare_long rs v1 v2 m). intros (B & C & D & E).
+ unfold eval_testcond; rewrite B, C, D, E.
+ destruct v1; try discriminate; destruct v2; try discriminate.
+ simpl in H; inv H.
+ unfold Val.cmplu; simpl. destruct c; simpl.
+- destruct (Int64.eq i i0); auto.
+- destruct (Int64.eq i i0); auto.
+- rewrite int64_sub_overflow. destruct (Int64.lt i i0); auto.
+- rewrite int64_sub_overflow, Int64.not_lt.
+ destruct (Int64.eq i i0), (Int64.lt i i0); auto.
+- rewrite int64_sub_overflow, (Int64.lt_not i).
+ destruct (Int64.eq i i0), (Int64.lt i i0); auto.
+- rewrite int64_sub_overflow. destruct (Int64.lt i i0); auto.
+Qed.
+
+Lemma eval_testcond_compare_ulong: forall c v1 v2 b rs m,
+ Val.cmplu_bool (Mem.valid_pointer m) c v1 v2 = Some b ->
+ eval_testcond (cond_for_unsigned_cmp c) (compare_long rs v1 v2 m) = Some b.
+Proof.
+ intros. generalize (compare_long_spec rs v1 v2 m).
+ set (rs' := compare_long rs v1 v2 m). intros (B & C & D & E).
+ unfold eval_testcond; rewrite B, C, D, E; unfold Val.cmplu.
+ destruct v1; try discriminate; destruct v2; try discriminate; simpl in H.
+- (* int-int *)
+ inv H. destruct c; simpl.
++ destruct (Int64.eq i i0); auto.
++ destruct (Int64.eq i i0); auto.
++ destruct (Int64.ltu i i0); auto.
++ rewrite (Int64.not_ltu i). destruct (Int64.eq i i0), (Int64.ltu i i0); auto.
++ rewrite (Int64.ltu_not i). destruct (Int64.eq i i0), (Int64.ltu i i0); auto.
++ destruct (Int64.ltu i i0); auto.
+- (* int-ptr *)
+ simpl.
+ destruct (Int64.eq i Int64.zero &&
+ (Mem.valid_pointer m b0 (Ptrofs.unsigned i0)
+ || Mem.valid_pointer m b0 (Ptrofs.unsigned i0 - 1))); try discriminate.
+ destruct c; simpl in H; inv H; reflexivity.
+- (* ptr-int *)
+ simpl.
+ destruct (Int64.eq i0 Int64.zero &&
+ (Mem.valid_pointer m b0 (Ptrofs.unsigned i)
+ || Mem.valid_pointer m b0 (Ptrofs.unsigned i - 1))); try discriminate.
+ destruct c; simpl in H; inv H; reflexivity.
+- (* ptr-ptr *)
+ simpl.
+ destruct (eq_block b0 b1).
++ destruct ((Mem.valid_pointer m b0 (Ptrofs.unsigned i)
+ || Mem.valid_pointer m b0 (Ptrofs.unsigned i - 1)) &&
+ (Mem.valid_pointer m b1 (Ptrofs.unsigned i0)
+ || Mem.valid_pointer m b1 (Ptrofs.unsigned i0 - 1)));
+ inv H.
+ destruct c; simpl.
+* destruct (Ptrofs.eq i i0); auto.
+* destruct (Ptrofs.eq i i0); auto.
+* destruct (Ptrofs.ltu i i0); auto.
+* rewrite (Ptrofs.not_ltu i). destruct (Ptrofs.eq i i0), (Ptrofs.ltu i i0); auto.
+* rewrite (Ptrofs.ltu_not i). destruct (Ptrofs.eq i i0), (Ptrofs.ltu i i0); auto.
+* destruct (Ptrofs.ltu i i0); auto.
++ destruct (Mem.valid_pointer m b0 (Ptrofs.unsigned i) &&
+ Mem.valid_pointer m b1 (Ptrofs.unsigned i0)); try discriminate.
+ destruct c; simpl in H; inv H; reflexivity.
+Qed.
+
+Lemma compare_float_spec: forall rs f1 f2,
+ let rs' := compare_float rs (Vfloat f1) (Vfloat f2) in
+ rs'#CN = (Val.of_bool (Float.cmp Clt f1 f2))
+ /\ rs'#CZ = (Val.of_bool (Float.cmp Ceq f1 f2))
+ /\ rs'#CC = (Val.of_bool (negb (Float.cmp Clt f1 f2)))
+ /\ rs'#CV = (Val.of_bool (negb (Float.ordered f1 f2))).
+Proof.
+ intros; auto.
+Qed.
+
+Lemma eval_testcond_compare_float: forall c v1 v2 b rs,
+ Val.cmpf_bool c v1 v2 = Some b ->
+ eval_testcond (cond_for_float_cmp c) (compare_float rs v1 v2) = Some b.
+Proof.
+ intros. destruct v1; try discriminate; destruct v2; simpl in H; inv H.
+ generalize (compare_float_spec rs f f0).
+ set (rs' := compare_float rs (Vfloat f) (Vfloat f0)).
+ intros (B & C & D & E).
+ unfold eval_testcond; rewrite B, C, D, E.
+Local Transparent Float.cmp Float.ordered.
+ unfold Float.cmp, Float.ordered;
+ destruct c; destruct (Float.compare f f0) as [[]|]; reflexivity.
+Qed.
+
+Lemma eval_testcond_compare_not_float: forall c v1 v2 b rs,
+ option_map negb (Val.cmpf_bool c v1 v2) = Some b ->
+ eval_testcond (cond_for_float_not_cmp c) (compare_float rs v1 v2) = Some b.
+Proof.
+ intros. destruct v1; try discriminate; destruct v2; simpl in H; inv H.
+ generalize (compare_float_spec rs f f0).
+ set (rs' := compare_float rs (Vfloat f) (Vfloat f0)).
+ intros (B & C & D & E).
+ unfold eval_testcond; rewrite B, C, D, E.
+Local Transparent Float.cmp Float.ordered.
+ unfold Float.cmp, Float.ordered;
+ destruct c; destruct (Float.compare f f0) as [[]|]; reflexivity.
+Qed.
+
+Lemma compare_single_spec: forall rs f1 f2,
+ let rs' := compare_single rs (Vsingle f1) (Vsingle f2) in
+ rs'#CN = (Val.of_bool (Float32.cmp Clt f1 f2))
+ /\ rs'#CZ = (Val.of_bool (Float32.cmp Ceq f1 f2))
+ /\ rs'#CC = (Val.of_bool (negb (Float32.cmp Clt f1 f2)))
+ /\ rs'#CV = (Val.of_bool (negb (Float32.ordered f1 f2))).
+Proof.
+ intros; auto.
+Qed.
+
+Lemma eval_testcond_compare_single: forall c v1 v2 b rs,
+ Val.cmpfs_bool c v1 v2 = Some b ->
+ eval_testcond (cond_for_float_cmp c) (compare_single rs v1 v2) = Some b.
+Proof.
+ intros. destruct v1; try discriminate; destruct v2; simpl in H; inv H.
+ generalize (compare_single_spec rs f f0).
+ set (rs' := compare_single rs (Vsingle f) (Vsingle f0)).
+ intros (B & C & D & E).
+ unfold eval_testcond; rewrite B, C, D, E.
+Local Transparent Float32.cmp Float32.ordered.
+ unfold Float32.cmp, Float32.ordered;
+ destruct c; destruct (Float32.compare f f0) as [[]|]; reflexivity.
+Qed.
+
+Lemma eval_testcond_compare_not_single: forall c v1 v2 b rs,
+ option_map negb (Val.cmpfs_bool c v1 v2) = Some b ->
+ eval_testcond (cond_for_float_not_cmp c) (compare_single rs v1 v2) = Some b.
+Proof.
+ intros. destruct v1; try discriminate; destruct v2; simpl in H; inv H.
+ generalize (compare_single_spec rs f f0).
+ set (rs' := compare_single rs (Vsingle f) (Vsingle f0)).
+ intros (B & C & D & E).
+ unfold eval_testcond; rewrite B, C, D, E.
+Local Transparent Float32.cmp Float32.ordered.
+ unfold Float32.cmp, Float32.ordered;
+ destruct c; destruct (Float32.compare f f0) as [[]|]; reflexivity.
+Qed.
+
+Remark compare_float_inv: forall rs v1 v2 r,
+ match r with CR _ => False | _ => True end ->
+ (nextinstr (compare_float rs v1 v2))#r = (nextinstr rs)#r.
+Proof.
+ intros; unfold compare_float.
+ destruct r; try contradiction; destruct v1; auto; destruct v2; auto.
+Qed.
+
+Remark compare_single_inv: forall rs v1 v2 r,
+ match r with CR _ => False | _ => True end ->
+ (nextinstr (compare_single rs v1 v2))#r = (nextinstr rs)#r.
+Proof.
+ intros; unfold compare_single.
+ destruct r; try contradiction; destruct v1; auto; destruct v2; auto.
+Qed.
+
+(** Translation of conditionals *)
+
+Ltac ArgsInv :=
+ repeat (match goal with
+ | [ H: Error _ = OK _ |- _ ] => discriminate
+ | [ H: match ?args with nil => _ | _ :: _ => _ end = OK _ |- _ ] => destruct args
+ | [ H: bind _ _ = OK _ |- _ ] => monadInv H
+ | [ H: match _ with left _ => _ | right _ => assertion_failed end = OK _ |- _ ] => monadInv H; ArgsInv
+ | [ H: match _ with true => _ | false => assertion_failed end = OK _ |- _ ] => monadInv H; ArgsInv
+ end);
+ subst;
+ repeat (match goal with
+ | [ H: ireg_of _ = OK _ |- _ ] => simpl in *; rewrite (ireg_of_eq _ _ H) in *
+ | [ H: freg_of _ = OK _ |- _ ] => simpl in *; rewrite (freg_of_eq _ _ H) in *
+ end).
+
+Lemma transl_cond_correct:
+ forall cond args k c rs m,
+ transl_cond cond args k = OK c ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m
+ /\ (forall b,
+ eval_condition cond (map rs (map preg_of args)) m = Some b ->
+ eval_testcond (cond_for_cond cond) rs' = Some b)
+ /\ forall r, data_preg r = true -> rs'#r = rs#r.
+Proof.
+ intros until m; intros TR. destruct cond; simpl in TR; ArgsInv.
+- (* Ccomp *)
+ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. apply eval_testcond_compare_sint; auto.
+ destruct r; reflexivity || discriminate.
+- (* Ccompu *)
+ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. apply eval_testcond_compare_uint; auto.
+ destruct r; reflexivity || discriminate.
+- (* Ccompimm *)
+ destruct (is_arith_imm32 n); [|destruct (is_arith_imm32 (Int.neg n))].
++ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. rewrite Int.repr_unsigned. apply eval_testcond_compare_sint; auto.
+ destruct r; reflexivity || discriminate.
++ econstructor; split.
+ apply exec_straight_one. simpl. rewrite Int.repr_unsigned, Int.neg_involutive. eauto. auto.
+ split; intros. apply eval_testcond_compare_sint; auto.
+ destruct r; reflexivity || discriminate.
++ exploit (exec_loadimm32 X16 n). intros (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. apply exec_straight_one.
+ simpl. rewrite B, C by eauto with asmgen. eauto. auto.
+ split; intros. apply eval_testcond_compare_sint; auto.
+ transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen.
+- (* Ccompuimm *)
+ destruct (is_arith_imm32 n); [|destruct (is_arith_imm32 (Int.neg n))].
++ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. rewrite Int.repr_unsigned. apply eval_testcond_compare_uint; auto.
+ destruct r; reflexivity || discriminate.
++ econstructor; split.
+ apply exec_straight_one. simpl. rewrite Int.repr_unsigned, Int.neg_involutive. eauto. auto.
+ split; intros. apply eval_testcond_compare_uint; auto.
+ destruct r; reflexivity || discriminate.
++ exploit (exec_loadimm32 X16 n). intros (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. apply exec_straight_one.
+ simpl. rewrite B, C by eauto with asmgen. eauto. auto.
+ split; intros. apply eval_testcond_compare_uint; auto.
+ transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen.
+- (* Ccompshift *)
+ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. rewrite transl_eval_shift. apply eval_testcond_compare_sint; auto.
+ destruct r; reflexivity || discriminate.
+- (* Ccompushift *)
+ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. rewrite transl_eval_shift. apply eval_testcond_compare_uint; auto.
+ destruct r; reflexivity || discriminate.
+- (* Cmaskzero *)
+ destruct (is_logical_imm32 n).
++ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. rewrite Int.repr_unsigned. apply (eval_testcond_compare_sint Ceq); auto.
+ destruct r; reflexivity || discriminate.
++ exploit (exec_loadimm32 X16 n). intros (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A.
+ apply exec_straight_one. simpl. rewrite B, C by eauto with asmgen. eauto. auto.
+ split; intros. apply (eval_testcond_compare_sint Ceq); auto.
+ transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen.
+- (* Cmasknotzero *)
+ destruct (is_logical_imm32 n).
++ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. rewrite Int.repr_unsigned. apply (eval_testcond_compare_sint Cne); auto.
+ destruct r; reflexivity || discriminate.
++ exploit (exec_loadimm32 X16 n). intros (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A.
+ apply exec_straight_one. simpl. rewrite B, C by eauto with asmgen. eauto. auto.
+ split; intros. apply (eval_testcond_compare_sint Cne); auto.
+ transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen.
+- (* Ccompl *)
+ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. apply eval_testcond_compare_slong; auto.
+ destruct r; reflexivity || discriminate.
+- (* Ccomplu *)
+ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. apply eval_testcond_compare_ulong; auto.
+ destruct r; reflexivity || discriminate.
+- (* Ccomplimm *)
+ destruct (is_arith_imm64 n); [|destruct (is_arith_imm64 (Int64.neg n))].
++ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. rewrite Int64.repr_unsigned. apply eval_testcond_compare_slong; auto.
+ destruct r; reflexivity || discriminate.
++ econstructor; split.
+ apply exec_straight_one. simpl. rewrite Int64.repr_unsigned, Int64.neg_involutive. eauto. auto.
+ split; intros. apply eval_testcond_compare_slong; auto.
+ destruct r; reflexivity || discriminate.
++ exploit (exec_loadimm64 X16 n). intros (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. apply exec_straight_one.
+ simpl. rewrite B, C by eauto with asmgen. eauto. auto.
+ split; intros. apply eval_testcond_compare_slong; auto.
+ transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen.
+- (* Ccompluimm *)
+ destruct (is_arith_imm64 n); [|destruct (is_arith_imm64 (Int64.neg n))].
++ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. rewrite Int64.repr_unsigned. apply eval_testcond_compare_ulong; auto.
+ destruct r; reflexivity || discriminate.
++ econstructor; split.
+ apply exec_straight_one. simpl. rewrite Int64.repr_unsigned, Int64.neg_involutive. eauto. auto.
+ split; intros. apply eval_testcond_compare_ulong; auto.
+ destruct r; reflexivity || discriminate.
++ exploit (exec_loadimm64 X16 n). intros (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. apply exec_straight_one.
+ simpl. rewrite B, C by eauto with asmgen. eauto. auto.
+ split; intros. apply eval_testcond_compare_ulong; auto.
+ transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen.
+- (* Ccomplshift *)
+ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. rewrite transl_eval_shiftl. apply eval_testcond_compare_slong; auto.
+ destruct r; reflexivity || discriminate.
+- (* Ccomplushift *)
+ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. rewrite transl_eval_shiftl. apply eval_testcond_compare_ulong; auto.
+ destruct r; reflexivity || discriminate.
+- (* Cmasklzero *)
+ destruct (is_logical_imm64 n).
++ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. rewrite Int64.repr_unsigned. apply (eval_testcond_compare_slong Ceq); auto.
+ destruct r; reflexivity || discriminate.
++ exploit (exec_loadimm64 X16 n). intros (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A.
+ apply exec_straight_one. simpl. rewrite B, C by eauto with asmgen. eauto. auto.
+ split; intros. apply (eval_testcond_compare_slong Ceq); auto.
+ transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen.
+- (* Cmasknotzero *)
+ destruct (is_logical_imm64 n).
++ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros. rewrite Int64.repr_unsigned. apply (eval_testcond_compare_slong Cne); auto.
+ destruct r; reflexivity || discriminate.
++ exploit (exec_loadimm64 X16 n). intros (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A.
+ apply exec_straight_one. simpl. rewrite B, C by eauto with asmgen. eauto. auto.
+ split; intros. apply (eval_testcond_compare_slong Cne); auto.
+ transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen.
+- (* Ccompf *)
+ econstructor; split. apply exec_straight_one. simpl; eauto.
+ rewrite compare_float_inv; auto.
+ split; intros. apply eval_testcond_compare_float; auto.
+ destruct r; discriminate || rewrite compare_float_inv; auto.
+- (* Cnotcompf *)
+ econstructor; split. apply exec_straight_one. simpl; eauto.
+ rewrite compare_float_inv; auto.
+ split; intros. apply eval_testcond_compare_not_float; auto.
+ destruct r; discriminate || rewrite compare_float_inv; auto.
+- (* Ccompfzero *)
+ econstructor; split. apply exec_straight_one. simpl; eauto.
+ rewrite compare_float_inv; auto.
+ split; intros. apply eval_testcond_compare_float; auto.
+ destruct r; discriminate || rewrite compare_float_inv; auto.
+- (* Cnotcompfzero *)
+ econstructor; split. apply exec_straight_one. simpl; eauto.
+ rewrite compare_float_inv; auto.
+ split; intros. apply eval_testcond_compare_not_float; auto.
+ destruct r; discriminate || rewrite compare_float_inv; auto.
+- (* Ccompfs *)
+ econstructor; split. apply exec_straight_one. simpl; eauto.
+ rewrite compare_single_inv; auto.
+ split; intros. apply eval_testcond_compare_single; auto.
+ destruct r; discriminate || rewrite compare_single_inv; auto.
+- (* Cnotcompfs *)
+ econstructor; split. apply exec_straight_one. simpl; eauto.
+ rewrite compare_single_inv; auto.
+ split; intros. apply eval_testcond_compare_not_single; auto.
+ destruct r; discriminate || rewrite compare_single_inv; auto.
+- (* Ccompfszero *)
+ econstructor; split. apply exec_straight_one. simpl; eauto.
+ rewrite compare_single_inv; auto.
+ split; intros. apply eval_testcond_compare_single; auto.
+ destruct r; discriminate || rewrite compare_single_inv; auto.
+- (* Cnotcompfszero *)
+ econstructor; split. apply exec_straight_one. simpl; eauto.
+ rewrite compare_single_inv; auto.
+ split; intros. apply eval_testcond_compare_not_single; auto.
+ destruct r; discriminate || rewrite compare_single_inv; auto.
+Qed.
+
+(** Translation of conditional branches *)
+
+Lemma transl_cond_branch_correct:
+ forall cond args lbl k c rs m b,
+ transl_cond_branch cond args lbl k = OK c ->
+ eval_condition cond (map rs (map preg_of args)) m = Some b ->
+ exists rs' insn,
+ exec_straight_opt ge fn c rs m (insn :: k) rs' m
+ /\ exec_instr ge fn insn rs' m =
+ (if b then goto_label fn lbl rs' m else Next (nextinstr rs') m)
+ /\ forall r, data_preg r = true -> rs'#r = rs#r.
+Proof.
+ intros until b; intros TR EV.
+ assert (DFL:
+ transl_cond_branch_default cond args lbl k = OK c ->
+ exists rs' insn,
+ exec_straight_opt ge fn c rs m (insn :: k) rs' m
+ /\ exec_instr ge fn insn rs' m =
+ (if b then goto_label fn lbl rs' m else Next (nextinstr rs') m)
+ /\ forall r, data_preg r = true -> rs'#r = rs#r).
+ {
+ unfold transl_cond_branch_default; intros.
+ exploit transl_cond_correct; eauto. intros (rs' & A & B & C).
+ exists rs', (Pbc (cond_for_cond cond) lbl); split.
+ apply exec_straight_opt_intro. eexact A.
+ split; auto. simpl. rewrite (B b) by auto. auto.
+ }
+Local Opaque transl_cond transl_cond_branch_default.
+ destruct args as [ | a1 args]; simpl in TR; auto.
+ destruct args as [ | a2 args]; simpl in TR; auto.
+ destruct cond; simpl in TR; auto.
+- (* Ccompimm *)
+ destruct c0; auto; destruct (Int.eq n Int.zero) eqn:N0; auto;
+ apply Int.same_if_eq in N0; subst n; ArgsInv.
++ (* Ccompimm Cne 0 *)
+ do 2 econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl. destruct (rs x); simpl in EV; inv EV. simpl. auto.
++ (* Ccompimm Ceq 0 *)
+ do 2 econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl. destruct (rs x); simpl in EV; inv EV. simpl. destruct (Int.eq i Int.zero); auto.
+- (* Ccompuimm *)
+ destruct c0; auto; destruct (Int.eq n Int.zero) eqn:N0; auto;
+ apply Int.same_if_eq in N0; subst n; ArgsInv.
++ (* Ccompuimm Cne 0 *)
+ do 2 econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl. rewrite EV. auto.
++ (* Ccompuimm Ceq 0 *)
+ do 2 econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl. rewrite (Val.negate_cmpu_bool (Mem.valid_pointer m) Cne), EV. destruct b; auto.
+- (* Cmaskzero *)
+ destruct (Int.is_power2 n) as [bit|] eqn:P2; auto. ArgsInv.
+ do 2 econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl.
+ erewrite <- Int.mul_pow2, Int.mul_commut, Int.mul_one by eauto.
+ rewrite (Val.negate_cmp_bool Ceq), EV. destruct b; auto.
+- (* Cmasknotzero *)
+ destruct (Int.is_power2 n) as [bit|] eqn:P2; auto. ArgsInv.
+ do 2 econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl.
+ erewrite <- Int.mul_pow2, Int.mul_commut, Int.mul_one by eauto.
+ rewrite EV. auto.
+- (* Ccomplimm *)
+ destruct c0; auto; destruct (Int64.eq n Int64.zero) eqn:N0; auto;
+ apply Int64.same_if_eq in N0; subst n; ArgsInv.
++ (* Ccomplimm Cne 0 *)
+ do 2 econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl. destruct (rs x); simpl in EV; inv EV. simpl. auto.
++ (* Ccomplimm Ceq 0 *)
+ do 2 econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl. destruct (rs x); simpl in EV; inv EV. simpl. destruct (Int64.eq i Int64.zero); auto.
+- (* Ccompluimm *)
+ destruct c0; auto; destruct (Int64.eq n Int64.zero) eqn:N0; auto;
+ apply Int64.same_if_eq in N0; subst n; ArgsInv.
++ (* Ccompluimm Cne 0 *)
+ do 2 econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl. rewrite EV. auto.
++ (* Ccompluimm Ceq 0 *)
+ do 2 econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl. rewrite (Val.negate_cmplu_bool (Mem.valid_pointer m) Cne), EV. destruct b; auto.
+- (* Cmasklzero *)
+ destruct (Int64.is_power2' n) as [bit|] eqn:P2; auto. ArgsInv.
+ do 2 econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl.
+ erewrite <- Int64.mul_pow2', Int64.mul_commut, Int64.mul_one by eauto.
+ rewrite (Val.negate_cmpl_bool Ceq), EV. destruct b; auto.
+- (* Cmasklnotzero *)
+ destruct (Int64.is_power2' n) as [bit|] eqn:P2; auto. ArgsInv.
+ do 2 econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl.
+ erewrite <- Int64.mul_pow2', Int64.mul_commut, Int64.mul_one by eauto.
+ rewrite EV. auto.
+Qed.
+
+(** Translation of arithmetic operations *)
+
+Ltac SimplEval H :=
+ match type of H with
+ | Some _ = None _ => discriminate
+ | Some _ = Some _ => inv H
+ | ?a = Some ?b => let A := fresh in assert (A: Val.maketotal a = b) by (rewrite H; reflexivity)
+end.
+
+Ltac TranslOpSimpl :=
+ econstructor; split;
+ [ apply exec_straight_one; [simpl; eauto | reflexivity]
+ | split; [ rewrite ? transl_eval_shift, ? transl_eval_shiftl;
+ apply Val.lessdef_same; Simpl; fail
+ | intros; Simpl; fail ] ].
+
+Ltac TranslOpBase :=
+ econstructor; split;
+ [ apply exec_straight_one; [simpl; eauto | reflexivity]
+ | split; [ rewrite ? transl_eval_shift, ? transl_eval_shiftl; Simpl
+ | intros; Simpl; fail ] ].
+
+Lemma transl_op_correct:
+ forall op args res k (rs: regset) m v c,
+ transl_op op args res k = OK c ->
+ eval_operation ge (rs#SP) op (map rs (map preg_of args)) m = Some v ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m
+ /\ Val.lessdef v rs'#(preg_of res)
+ /\ forall r, data_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs' r = rs r.
+Proof.
+Local Opaque Int.eq Int64.eq Val.add Val.addl Int.zwordsize Int64.zwordsize.
+ intros until c; intros TR EV.
+ unfold transl_op in TR; destruct op; ArgsInv; simpl in EV; SimplEval EV; try TranslOpSimpl.
+- (* move *)
+ destruct (preg_of res) eqn:RR; try discriminate; destruct (preg_of m0) eqn:R1; inv TR.
++ TranslOpSimpl.
++ TranslOpSimpl.
+- (* intconst *)
+ exploit exec_loadimm32. intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split. rewrite B; auto. intros; auto with asmgen.
+- (* longconst *)
+ exploit exec_loadimm64. intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split. rewrite B; auto. intros; auto with asmgen.
+- (* floatconst *)
+ destruct (Float.eq_dec n Float.zero).
++ subst n. TranslOpSimpl.
++ TranslOpSimpl.
+- (* singleconst *)
+ destruct (Float32.eq_dec n Float32.zero).
++ subst n. TranslOpSimpl.
++ TranslOpSimpl.
+- (* loadsymbol *)
+ exploit (exec_loadsymbol x id ofs). eauto with asmgen. intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split. rewrite B; auto. auto.
+- (* addrstack *)
+ exploit (exec_addimm64 x XSP (Ptrofs.to_int64 ofs)). simpl; eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split. simpl in B; rewrite B.
+Local Transparent Val.addl.
+ destruct (rs SP); simpl; auto. rewrite Ptrofs.of_int64_to_int64 by auto. auto.
+ auto.
+- (* shift *)
+ rewrite <- transl_eval_shift'. TranslOpSimpl.
+- (* addimm *)
+ exploit (exec_addimm32 x x0 n). eauto with asmgen. intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split. rewrite B; auto. auto.
+- (* mul *)
+ TranslOpBase.
+Local Transparent Val.add.
+ destruct (rs x0); auto; destruct (rs x1); auto. simpl. rewrite Int.add_zero_l; auto.
+- (* andimm *)
+ exploit (exec_logicalimm32 (Pandimm W) (Pand W)).
+ intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split. rewrite B; auto. auto.
+- (* orimm *)
+ exploit (exec_logicalimm32 (Porrimm W) (Porr W)).
+ intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split. rewrite B; auto. auto.
+- (* xorimm *)
+ exploit (exec_logicalimm32 (Peorimm W) (Peor W)).
+ intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split. rewrite B; auto. auto.
+- (* not *)
+ TranslOpBase.
+ destruct (rs x0); auto. simpl. rewrite Int.or_zero_l; auto.
+- (* notshift *)
+ TranslOpBase.
+ destruct (eval_shift s (rs x0) a); auto. simpl. rewrite Int.or_zero_l; auto.
+- (* shrx *)
+ exploit (exec_shrx32 x x0 n); eauto with asmgen. intros (rs' & A & B & C).
+ econstructor; split. eexact A. split. rewrite B; auto. auto.
+- (* zero-ext *)
+ TranslOpBase.
+ destruct (rs x0); auto; simpl. rewrite Int.shl_zero. auto.
+- (* sign-ext *)
+ TranslOpBase.
+ destruct (rs x0); auto; simpl. rewrite Int.shl_zero. auto.
+- (* shlzext *)
+ TranslOpBase.
+ destruct (rs x0); simpl; auto. rewrite <- Int.shl_zero_ext_min; auto using a32_range.
+- (* shlsext *)
+ TranslOpBase.
+ destruct (rs x0); simpl; auto. rewrite <- Int.shl_sign_ext_min; auto using a32_range.
+- (* zextshr *)
+ TranslOpBase.
+ destruct (rs x0); simpl; auto. rewrite ! a32_range; simpl. rewrite <- Int.zero_ext_shru_min; auto using a32_range.
+- (* sextshr *)
+ TranslOpBase.
+ destruct (rs x0); simpl; auto. rewrite ! a32_range; simpl. rewrite <- Int.sign_ext_shr_min; auto using a32_range.
+- (* shiftl *)
+ rewrite <- transl_eval_shiftl'. TranslOpSimpl.
+- (* extend *)
+ exploit (exec_move_extended x0 x1 x a k). intros (rs' & A & B & C).
+ econstructor; split. eexact A.
+ split. rewrite B; auto. eauto with asmgen.
+- (* addext *)
+ exploit (exec_arith_extended Val.addl Paddext (Padd X)).
+ auto. auto. instantiate (1 := x1). eauto with asmgen. intros (rs' & A & B & C).
+ econstructor; split. eexact A. split. rewrite B; auto. auto.
+- (* addlimm *)
+ exploit (exec_addimm64 x x0 n). simpl. generalize (ireg_of_not_X16 _ _ EQ1). congruence.
+ intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split. simpl in B; rewrite B; auto. auto.
+- (* subext *)
+ exploit (exec_arith_extended Val.subl Psubext (Psub X)).
+ auto. auto. instantiate (1 := x1). eauto with asmgen. intros (rs' & A & B & C).
+ econstructor; split. eexact A. split. rewrite B; auto. auto.
+- (* mull *)
+ TranslOpBase.
+ destruct (rs x0); auto; destruct (rs x1); auto. simpl. rewrite Int64.add_zero_l; auto.
+- (* andlimm *)
+ exploit (exec_logicalimm64 (Pandimm X) (Pand X)).
+ intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split. rewrite B; auto. auto.
+- (* orlimm *)
+ exploit (exec_logicalimm64 (Porrimm X) (Porr X)).
+ intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split. rewrite B; auto. auto.
+- (* xorlimm *)
+ exploit (exec_logicalimm64 (Peorimm X) (Peor X)).
+ intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split. rewrite B; auto. auto.
+- (* notl *)
+ TranslOpBase.
+ destruct (rs x0); auto. simpl. rewrite Int64.or_zero_l; auto.
+- (* notlshift *)
+ TranslOpBase.
+ destruct (eval_shiftl s (rs x0) a); auto. simpl. rewrite Int64.or_zero_l; auto.
+- (* shrx *)
+ exploit (exec_shrx64 x x0 n); eauto with asmgen. intros (rs' & A & B & C).
+ econstructor; split. eexact A. split. rewrite B; auto. auto.
+- (* zero-ext-l *)
+ TranslOpBase.
+ destruct (rs x0); auto; simpl. rewrite Int64.shl'_zero. auto.
+- (* sign-ext-l *)
+ TranslOpBase.
+ destruct (rs x0); auto; simpl. rewrite Int64.shl'_zero. auto.
+- (* shllzext *)
+ TranslOpBase.
+ destruct (rs x0); simpl; auto. rewrite <- Int64.shl'_zero_ext_min; auto using a64_range.
+- (* shllsext *)
+ TranslOpBase.
+ destruct (rs x0); simpl; auto. rewrite <- Int64.shl'_sign_ext_min; auto using a64_range.
+- (* zextshrl *)
+ TranslOpBase.
+ destruct (rs x0); simpl; auto. rewrite ! a64_range; simpl. rewrite <- Int64.zero_ext_shru'_min; auto using a64_range.
+- (* sextshrl *)
+ TranslOpBase.
+ destruct (rs x0); simpl; auto. rewrite ! a64_range; simpl. rewrite <- Int64.sign_ext_shr'_min; auto using a64_range.
+- (* condition *)
+ exploit (transl_cond_correct cond args); eauto. intros (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; eauto. auto.
+ split. Simpl. destruct (eval_condition cond (map rs (map preg_of args)) m) as [b|]; simpl in *.
+ rewrite (B b) by auto. auto.
+ auto.
+ intros; Simpl.
+- (* select *)
+ destruct (preg_of res) eqn:RES; monadInv TR.
+ + (* integer *)
+ generalize (ireg_of_eq _ _ EQ) (ireg_of_eq _ _ EQ1); intros E1 E2; rewrite E1, E2.
+ exploit (transl_cond_correct cond args); eauto. intros (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; eauto. auto.
+ split. Simpl. destruct (eval_condition cond (map rs (map preg_of args)) m) as [b|]; simpl in *.
+ rewrite (B b) by auto. rewrite !C. apply Val.lessdef_normalize.
+ rewrite <- E2; auto with asmgen. rewrite <- E1; auto with asmgen.
+ auto.
+ intros; Simpl.
+ + (* FP *)
+ generalize (freg_of_eq _ _ EQ) (freg_of_eq _ _ EQ1); intros E1 E2; rewrite E1, E2.
+ exploit (transl_cond_correct cond args); eauto. intros (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; eauto. auto.
+ split. Simpl. destruct (eval_condition cond (map rs (map preg_of args)) m) as [b|]; simpl in *.
+ rewrite (B b) by auto. rewrite !C. apply Val.lessdef_normalize.
+ rewrite <- E2; auto with asmgen. rewrite <- E1; auto with asmgen.
+ auto.
+ intros; Simpl.
+Qed.
+
+(** Translation of addressing modes, loads, stores *)
+
+Lemma transl_addressing_correct:
+ forall sz addr args (insn: Asm.addressing -> instruction) k (rs: regset) m c b o,
+ transl_addressing sz addr args insn k = OK c ->
+ Op.eval_addressing ge (rs#SP) addr (map rs (map preg_of args)) = Some (Vptr b o) ->
+ exists ad rs',
+ exec_straight_opt ge fn c rs m (insn ad :: k) rs' m
+ /\ Asm.eval_addressing ge ad rs' = Vptr b o
+ /\ forall r, data_preg r = true -> rs' r = rs r.
+Proof.
+ intros until o; intros TR EV.
+ unfold transl_addressing in TR; destruct addr; ArgsInv; SimplEval EV.
+- (* Aindexed *)
+ destruct (offset_representable sz ofs); inv EQ0.
++ econstructor; econstructor; split. apply exec_straight_opt_refl.
+ auto.
++ exploit (exec_loadimm64 X16 ofs). intros (rs' & A & B & C).
+ econstructor; exists rs'; split. apply exec_straight_opt_intro; eexact A.
+ split. simpl. rewrite B, C by eauto with asmgen. auto.
+ eauto with asmgen.
+- (* Aindexed2 *)
+ econstructor; econstructor; split. apply exec_straight_opt_refl.
+ auto.
+- (* Aindexed2shift *)
+ destruct (Int.eq a Int.zero) eqn:E; [|destruct (Int.eq (Int.shl Int.one a) (Int.repr sz))]; inv EQ2.
++ apply Int.same_if_eq in E. rewrite E.
+ econstructor; econstructor; split. apply exec_straight_opt_refl.
+ split; auto. simpl.
+ rewrite Val.addl_commut in H0. destruct (rs x0); try discriminate.
+ unfold Val.shll. rewrite Int64.shl'_zero. auto.
++ econstructor; econstructor; split. apply exec_straight_opt_refl.
+ auto.
++ econstructor; econstructor; split.
+ apply exec_straight_opt_intro. apply exec_straight_one. simpl; eauto. auto.
+ split. simpl. Simpl. rewrite H0. simpl. rewrite Ptrofs.add_zero. auto.
+ intros; Simpl.
+- (* Aindexed2ext *)
+ destruct (Int.eq a Int.zero || Int.eq (Int.shl Int.one a) (Int.repr sz)); inv EQ2.
++ econstructor; econstructor; split. apply exec_straight_opt_refl.
+ split; auto. destruct x; auto.
++ exploit (exec_arith_extended Val.addl Paddext (Padd X)); auto.
+ instantiate (1 := x0). eauto with asmgen.
+ intros (rs' & A & B & C).
+ econstructor; exists rs'; split.
+ apply exec_straight_opt_intro. eexact A.
+ split. simpl. rewrite B. rewrite Val.addl_assoc. f_equal.
+ unfold Op.eval_extend; destruct x, (rs x1); simpl; auto; rewrite ! a64_range;
+ simpl; rewrite Int64.add_zero; auto.
+ intros. apply C; eauto with asmgen.
+- (* Aglobal *)
+ destruct (Ptrofs.eq (Ptrofs.modu ofs (Ptrofs.repr sz)) Ptrofs.zero && symbol_is_aligned id sz); inv TR.
++ econstructor; econstructor; split.
+ apply exec_straight_opt_intro. apply exec_straight_one. simpl; eauto. auto.
+ split. simpl. Simpl. rewrite symbol_high_low. simpl in EV. congruence.
+ intros; Simpl.
++ exploit (exec_loadsymbol X16 id ofs). auto. intros (rs' & A & B & C).
+ econstructor; exists rs'; split.
+ apply exec_straight_opt_intro. eexact A.
+ split. simpl.
+ rewrite B. rewrite <- Genv.shift_symbol_address_64, Ptrofs.add_zero by auto.
+ simpl in EV. congruence.
+ auto with asmgen.
+- (* Ainstrack *)
+ assert (E: Val.addl (rs SP) (Vlong (Ptrofs.to_int64 ofs)) = Vptr b o).
+ { simpl in EV. inv EV. destruct (rs SP); simpl in H1; inv H1. simpl.
+ rewrite Ptrofs.of_int64_to_int64 by auto. auto. }
+ destruct (offset_representable sz (Ptrofs.to_int64 ofs)); inv TR.
++ econstructor; econstructor; split. apply exec_straight_opt_refl.
+ auto.
++ exploit (exec_loadimm64 X16 (Ptrofs.to_int64 ofs)). intros (rs' & A & B & C).
+ econstructor; exists rs'; split.
+ apply exec_straight_opt_intro. eexact A.
+ split. simpl. rewrite B, C by eauto with asmgen. auto.
+ auto with asmgen.
+Qed.
+
+Lemma transl_load_correct:
+ forall chunk addr args dst k c (rs: regset) m vaddr v,
+ transl_load chunk addr args dst k = OK c ->
+ Op.eval_addressing ge (rs#SP) addr (map rs (map preg_of args)) = Some vaddr ->
+ Mem.loadv chunk m vaddr = Some v ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m
+ /\ rs'#(preg_of dst) = v
+ /\ forall r, data_preg r = true -> r <> preg_of dst -> rs' r = rs r.
+Proof.
+ intros. destruct vaddr; try discriminate.
+ assert (A: exists sz insn,
+ transl_addressing sz addr args insn k = OK c
+ /\ (forall ad rs', exec_instr ge fn (insn ad) rs' m =
+ exec_load ge chunk (fun v => v) ad (preg_of dst) rs' m)).
+ {
+ destruct chunk; monadInv H;
+ try rewrite (ireg_of_eq _ _ EQ); try rewrite (freg_of_eq _ _ EQ);
+ do 2 econstructor; (split; [eassumption|auto]).
+ }
+ destruct A as (sz & insn & B & C).
+ exploit transl_addressing_correct. eexact B. eexact H0. intros (ad & rs' & P & Q & R).
+ assert (X: exec_load ge chunk (fun v => v) ad (preg_of dst) rs' m =
+ Next (nextinstr (rs'#(preg_of dst) <- v)) m).
+ { unfold exec_load. rewrite Q, H1. auto. }
+ econstructor; split.
+ eapply exec_straight_opt_right. eexact P.
+ apply exec_straight_one. rewrite C, X; eauto. Simpl.
+ split. Simpl. intros; Simpl.
+Qed.
+
+Lemma transl_store_correct:
+ forall chunk addr args src k c (rs: regset) m vaddr m',
+ transl_store chunk addr args src k = OK c ->
+ Op.eval_addressing ge (rs#SP) addr (map rs (map preg_of args)) = Some vaddr ->
+ Mem.storev chunk m vaddr rs#(preg_of src) = Some m' ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m'
+ /\ forall r, data_preg r = true -> rs' r = rs r.
+Proof.
+ intros. destruct vaddr; try discriminate.
+ set (chunk' := match chunk with Mint8signed => Mint8unsigned
+ | Mint16signed => Mint16unsigned
+ | _ => chunk end).
+ assert (A: exists sz insn,
+ transl_addressing sz addr args insn k = OK c
+ /\ (forall ad rs', exec_instr ge fn (insn ad) rs' m =
+ exec_store ge chunk' ad rs'#(preg_of src) rs' m)).
+ {
+ unfold chunk'; destruct chunk; monadInv H;
+ try rewrite (ireg_of_eq _ _ EQ); try rewrite (freg_of_eq _ _ EQ);
+ do 2 econstructor; (split; [eassumption|auto]).
+ }
+ destruct A as (sz & insn & B & C).
+ exploit transl_addressing_correct. eexact B. eexact H0. intros (ad & rs' & P & Q & R).
+ assert (X: Mem.storev chunk' m (Vptr b i) rs#(preg_of src) = Some m').
+ { rewrite <- H1. unfold chunk'. destruct chunk; auto; simpl; symmetry.
+ apply Mem.store_signed_unsigned_8.
+ apply Mem.store_signed_unsigned_16. }
+ assert (Y: exec_store ge chunk' ad rs'#(preg_of src) rs' m =
+ Next (nextinstr rs') m').
+ { unfold exec_store. rewrite Q, R, X by auto with asmgen. auto. }
+ econstructor; split.
+ eapply exec_straight_opt_right. eexact P.
+ apply exec_straight_one. rewrite C, Y; eauto. Simpl.
+ intros; Simpl.
+Qed.
+
+(** Translation of indexed memory accesses *)
+
+Lemma indexed_memory_access_correct: forall insn sz (base: iregsp) ofs k (rs: regset) m b i,
+ preg_of_iregsp base <> IR X16 ->
+ Val.offset_ptr rs#base ofs = Vptr b i ->
+ exists ad rs',
+ exec_straight_opt ge fn (indexed_memory_access insn sz base ofs k) rs m (insn ad :: k) rs' m
+ /\ Asm.eval_addressing ge ad rs' = Vptr b i
+ /\ forall r, r <> PC -> r <> X16 -> rs' r = rs r.
+Proof.
+ unfold indexed_memory_access; intros.
+ assert (Val.addl rs#base (Vlong (Ptrofs.to_int64 ofs)) = Vptr b i).
+ { destruct (rs base); try discriminate. simpl in *. rewrite Ptrofs.of_int64_to_int64 by auto. auto. }
+ destruct offset_representable.
+- econstructor; econstructor; split. apply exec_straight_opt_refl. auto.
+- exploit (exec_loadimm64 X16); eauto. intros (rs' & A & B & C).
+ econstructor; econstructor; split. apply exec_straight_opt_intro; eexact A.
+ split. simpl. rewrite B, C by eauto with asmgen. auto. auto.
+Qed.
+
+Lemma loadptr_correct: forall (base: iregsp) ofs dst k m v (rs: regset),
+ Mem.loadv Mint64 m (Val.offset_ptr rs#base ofs) = Some v ->
+ preg_of_iregsp base <> IR X16 ->
+ exists rs',
+ exec_straight ge fn (loadptr base ofs dst k) rs m k rs' m
+ /\ rs'#dst = v
+ /\ forall r, r <> PC -> r <> X16 -> r <> dst -> rs' r = rs r.
+Proof.
+ intros.
+ destruct (Val.offset_ptr rs#base ofs) eqn:V; try discriminate.
+ exploit indexed_memory_access_correct; eauto. intros (ad & rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_opt_right. eexact A.
+ apply exec_straight_one. simpl. unfold exec_load. rewrite B, H. eauto. auto.
+ split. Simpl. intros; Simpl.
+Qed.
+
+Lemma storeptr_correct: forall (base: iregsp) ofs (src: ireg) k m m' (rs: regset),
+ Mem.storev Mint64 m (Val.offset_ptr rs#base ofs) rs#src = Some m' ->
+ preg_of_iregsp base <> IR X16 ->
+ src <> X16 ->
+ exists rs',
+ exec_straight ge fn (storeptr src base ofs k) rs m k rs' m'
+ /\ forall r, r <> PC -> r <> X16 -> rs' r = rs r.
+Proof.
+ intros.
+ destruct (Val.offset_ptr rs#base ofs) eqn:V; try discriminate.
+ exploit indexed_memory_access_correct; eauto. intros (ad & rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_opt_right. eexact A.
+ apply exec_straight_one. simpl. unfold exec_store. rewrite B, C, H by eauto with asmgen. eauto. auto.
+ intros; Simpl.
+Qed.
+
+Lemma loadind_correct: forall (base: iregsp) ofs ty dst k c (rs: regset) m v,
+ loadind base ofs ty dst k = OK c ->
+ Mem.loadv (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) = Some v ->
+ preg_of_iregsp base <> IR X16 ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m
+ /\ rs'#(preg_of dst) = v
+ /\ forall r, data_preg r = true -> r <> preg_of dst -> rs' r = rs r.
+Proof.
+ intros.
+ destruct (Val.offset_ptr rs#base ofs) eqn:V; try discriminate.
+ assert (X: exists sz insn,
+ c = indexed_memory_access insn sz base ofs k
+ /\ (forall ad rs', exec_instr ge fn (insn ad) rs' m =
+ exec_load ge (chunk_of_type ty) (fun v => v) ad (preg_of dst) rs' m)).
+ {
+ unfold loadind in H; destruct ty; destruct (preg_of dst); inv H; do 2 econstructor; eauto.
+ }
+ destruct X as (sz & insn & EQ & SEM). subst c.
+ exploit indexed_memory_access_correct; eauto. intros (ad & rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_opt_right. eexact A.
+ apply exec_straight_one. rewrite SEM. unfold exec_load. rewrite B, H0. eauto. Simpl.
+ split. Simpl. intros; Simpl.
+Qed.
+
+Lemma storeind_correct: forall (base: iregsp) ofs ty src k c (rs: regset) m m',
+ storeind src base ofs ty k = OK c ->
+ Mem.storev (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) rs#(preg_of src) = Some m' ->
+ preg_of_iregsp base <> IR X16 ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m'
+ /\ forall r, data_preg r = true -> rs' r = rs r.
+Proof.
+ intros.
+ destruct (Val.offset_ptr rs#base ofs) eqn:V; try discriminate.
+ assert (X: exists sz insn,
+ c = indexed_memory_access insn sz base ofs k
+ /\ (forall ad rs', exec_instr ge fn (insn ad) rs' m =
+ exec_store ge (chunk_of_type ty) ad rs'#(preg_of src) rs' m)).
+ {
+ unfold storeind in H; destruct ty; destruct (preg_of src); inv H; do 2 econstructor; eauto.
+ }
+ destruct X as (sz & insn & EQ & SEM). subst c.
+ exploit indexed_memory_access_correct; eauto. intros (ad & rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_opt_right. eexact A.
+ apply exec_straight_one. rewrite SEM.
+ unfold exec_store. rewrite B, C, H0 by eauto with asmgen. eauto.
+ Simpl.
+ intros; Simpl.
+Qed.
+
+Lemma make_epilogue_correct:
+ forall ge0 f m stk soff cs m' ms rs k tm,
+ load_stack m (Vptr stk soff) Tptr f.(fn_link_ofs) = Some (parent_sp cs) ->
+ load_stack m (Vptr stk soff) Tptr f.(fn_retaddr_ofs) = Some (parent_ra cs) ->
+ Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
+ agree ms (Vptr stk soff) rs ->
+ Mem.extends m tm ->
+ match_stack ge0 cs ->
+ exists rs', exists tm',
+ exec_straight ge fn (make_epilogue f k) rs tm k rs' tm'
+ /\ agree ms (parent_sp cs) rs'
+ /\ Mem.extends m' tm'
+ /\ rs'#RA = parent_ra cs
+ /\ rs'#SP = parent_sp cs
+ /\ (forall r, r <> PC -> r <> SP -> r <> X30 -> r <> X16 -> rs'#r = rs#r).
+Proof.
+ intros until tm; intros LP LRA FREE AG MEXT MCS.
+ exploit Mem.loadv_extends. eauto. eexact LP. auto. simpl. intros (parent' & LP' & LDP').
+ exploit Mem.loadv_extends. eauto. eexact LRA. auto. simpl. intros (ra' & LRA' & LDRA').
+ exploit lessdef_parent_sp; eauto. intros EQ; subst parent'; clear LDP'.
+ exploit lessdef_parent_ra; eauto. intros EQ; subst ra'; clear LDRA'.
+ exploit Mem.free_parallel_extends; eauto. intros (tm' & FREE' & MEXT').
+ unfold make_epilogue.
+ exploit (loadptr_correct XSP (fn_retaddr_ofs f)).
+ instantiate (2 := rs). simpl. rewrite <- (sp_val _ _ _ AG). simpl. eexact LRA'. simpl; congruence.
+ intros (rs1 & A1 & B1 & C1).
+ econstructor; econstructor; split.
+ eapply exec_straight_trans. eexact A1. apply exec_straight_one. simpl.
+ simpl; rewrite (C1 SP) by auto with asmgen. rewrite <- (sp_val _ _ _ AG). simpl; rewrite LP'.
+ rewrite FREE'. eauto. auto.
+ split. apply agree_nextinstr. apply agree_set_other; auto.
+ apply agree_change_sp with (Vptr stk soff).
+ apply agree_exten with rs; auto. intros; apply C1; auto with asmgen.
+ eapply parent_sp_def; eauto.
+ split. auto.
+ split. Simpl.
+ split. Simpl.
+ intros. Simpl.
+Qed.
+
+End CONSTRUCTORS.
diff --git a/aarch64/TO_MERGE/TargetPrinter.ml b/aarch64/TO_MERGE/TargetPrinter.ml
new file mode 100644
index 00000000..bc4279a0
--- /dev/null
+++ b/aarch64/TO_MERGE/TargetPrinter.ml
@@ -0,0 +1,862 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, Collège de France and INRIA Paris *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(* Printing AArch64 assembly code in asm syntax *)
+
+open Printf
+open Camlcoq
+open Sections
+open AST
+open Asm
+open AisAnnot
+open PrintAsmaux
+open Fileinfo
+
+<<<<<<< HEAD
+(* Module containing the printing functions *)
+
+module Target (*: TARGET*) =
+=======
+(* Recognition of FP numbers that are supported by the fmov #imm instructions:
+ "a normalized binary floating point encoding with 1 sign bit,
+ 4 bits of fraction and a 3-bit exponent"
+*)
+
+let is_immediate_float64 bits =
+ let exp = (Int64.(to_int (shift_right_logical bits 52)) land 0x7FF) - 1023 in
+ let mant = Int64.logand bits 0xF_FFFF_FFFF_FFFFL in
+ exp >= -3 && exp <= 4 && Int64.logand mant 0xF_0000_0000_0000L = mant
+
+let is_immediate_float32 bits =
+ let exp = (Int32.(to_int (shift_right_logical bits 23)) land 0xFF) - 127 in
+ let mant = Int32.logand bits 0x7F_FFFFl in
+ exp >= -3 && exp <= 4 && Int32.logand mant 0x78_0000l = mant
+
+(* Naming and printing registers *)
+
+let intsz oc (sz, n) =
+ match sz with X -> coqint64 oc n | W -> coqint oc n
+
+let xreg_name = function
+ | X0 -> "x0" | X1 -> "x1" | X2 -> "x2" | X3 -> "x3"
+ | X4 -> "x4" | X5 -> "x5" | X6 -> "x6" | X7 -> "x7"
+ | X8 -> "x8" | X9 -> "x9" | X10 -> "x10" | X11 -> "x11"
+ | X12 -> "x12" | X13 -> "x13" | X14 -> "x14" | X15 -> "x15"
+ | X16 -> "x16" | X17 -> "x17" | X18 -> "x18" | X19 -> "x19"
+ | X20 -> "x20" | X21 -> "x21" | X22 -> "x22" | X23 -> "x23"
+ | X24 -> "x24" | X25 -> "x25" | X26 -> "x26" | X27 -> "x27"
+ | X28 -> "x28" | X29 -> "x29" | X30 -> "x30"
+
+let wreg_name = function
+ | X0 -> "w0" | X1 -> "w1" | X2 -> "w2" | X3 -> "w3"
+ | X4 -> "w4" | X5 -> "w5" | X6 -> "w6" | X7 -> "w7"
+ | X8 -> "w8" | X9 -> "w9" | X10 -> "w10" | X11 -> "w11"
+ | X12 -> "w12" | X13 -> "w13" | X14 -> "w14" | X15 -> "w15"
+ | X16 -> "w16" | X17 -> "w17" | X18 -> "w18" | X19 -> "w19"
+ | X20 -> "w20" | X21 -> "w21" | X22 -> "w22" | X23 -> "w23"
+ | X24 -> "w24" | X25 -> "w25" | X26 -> "w26" | X27 -> "w27"
+ | X28 -> "w28" | X29 -> "w29" | X30 -> "w30"
+
+let xreg0_name = function RR0 r -> xreg_name r | XZR -> "xzr"
+let wreg0_name = function RR0 r -> wreg_name r | XZR -> "wzr"
+
+let xregsp_name = function RR1 r -> xreg_name r | XSP -> "sp"
+let wregsp_name = function RR1 r -> wreg_name r | XSP -> "wsp"
+
+let dreg_name = function
+| D0 -> "d0" | D1 -> "d1" | D2 -> "d2" | D3 -> "d3"
+| D4 -> "d4" | D5 -> "d5" | D6 -> "d6" | D7 -> "d7"
+| D8 -> "d8" | D9 -> "d9" | D10 -> "d10" | D11 -> "d11"
+| D12 -> "d12" | D13 -> "d13" | D14 -> "d14" | D15 -> "d15"
+| D16 -> "d16" | D17 -> "d17" | D18 -> "d18" | D19 -> "d19"
+| D20 -> "d20" | D21 -> "d21" | D22 -> "d22" | D23 -> "d23"
+| D24 -> "d24" | D25 -> "d25" | D26 -> "d26" | D27 -> "d27"
+| D28 -> "d28" | D29 -> "d29" | D30 -> "d30" | D31 -> "d31"
+
+let sreg_name = function
+| D0 -> "s0" | D1 -> "s1" | D2 -> "s2" | D3 -> "s3"
+| D4 -> "s4" | D5 -> "s5" | D6 -> "s6" | D7 -> "s7"
+| D8 -> "s8" | D9 -> "s9" | D10 -> "s10" | D11 -> "s11"
+| D12 -> "s12" | D13 -> "s13" | D14 -> "s14" | D15 -> "s15"
+| D16 -> "s16" | D17 -> "s17" | D18 -> "s18" | D19 -> "s19"
+| D20 -> "s20" | D21 -> "s21" | D22 -> "s22" | D23 -> "s23"
+| D24 -> "s24" | D25 -> "s25" | D26 -> "s26" | D27 -> "s27"
+| D28 -> "s28" | D29 -> "s29" | D30 -> "s30" | D31 -> "s31"
+
+let xreg oc r = output_string oc (xreg_name r)
+let wreg oc r = output_string oc (wreg_name r)
+let ireg oc (sz, r) =
+ output_string oc (match sz with X -> xreg_name r | W -> wreg_name r)
+
+let xreg0 oc r = output_string oc (xreg0_name r)
+let wreg0 oc r = output_string oc (wreg0_name r)
+let ireg0 oc (sz, r) =
+ output_string oc (match sz with X -> xreg0_name r | W -> wreg0_name r)
+
+let xregsp oc r = output_string oc (xregsp_name r)
+let iregsp oc (sz, r) =
+ output_string oc (match sz with X -> xregsp_name r | W -> wregsp_name r)
+
+let dreg oc r = output_string oc (dreg_name r)
+let sreg oc r = output_string oc (sreg_name r)
+let freg oc (sz, r) =
+ output_string oc (match sz with D -> dreg_name r | S -> sreg_name r)
+
+let preg_asm oc ty = function
+ | IR r -> if ty = Tint then wreg oc r else xreg oc r
+ | FR r -> if ty = Tsingle then sreg oc r else dreg oc r
+ | _ -> assert false
+
+let preg_annot = function
+ | IR r -> xreg_name r
+ | FR r -> dreg_name r
+ | _ -> assert false
+
+(* Base-2 log of a Caml integer *)
+let rec log2 n =
+ assert (n > 0);
+ if n = 1 then 0 else 1 + log2 (n lsr 1)
+
+(* System dependent printer functions *)
+
+module type SYSTEM =
+ sig
+ val comment: string
+ val raw_symbol: out_channel -> string -> unit
+ val symbol: out_channel -> P.t -> unit
+ val symbol_offset_high: out_channel -> P.t * Z.t -> unit
+ val symbol_offset_low: out_channel -> P.t * Z.t -> unit
+ val label: out_channel -> int -> unit
+ val label_high: out_channel -> int -> unit
+ val label_low: out_channel -> int -> unit
+ val load_symbol_address: out_channel -> ireg -> P.t -> unit
+ val name_of_section: section_name -> string
+ val print_fun_info: out_channel -> P.t -> unit
+ val print_var_info: out_channel -> P.t -> unit
+ val print_comm_decl: out_channel -> P.t -> Z.t -> int -> unit
+ val print_lcomm_decl: out_channel -> P.t -> Z.t -> int -> unit
+ end
+
+module ELF_System : SYSTEM =
+>>>>>>> master
+ struct
+ let comment = "//"
+ let raw_symbol = output_string
+ let symbol = elf_symbol
+ let symbol_offset_high = elf_symbol_offset
+ let symbol_offset_low oc id_ofs =
+ fprintf oc "#:lo12:%a" elf_symbol_offset id_ofs
+
+ let label = elf_label
+ let label_high = elf_label
+ let label_low oc lbl =
+ fprintf oc "#:lo12:%a" elf_label lbl
+
+<<<<<<< HEAD
+ let print_label oc lbl = label oc (transl_label lbl)
+
+ let intsz oc (sz, n) =
+ match sz with X -> coqint64 oc n | W -> coqint oc n
+
+ let xreg_name = function
+ | X0 -> "x0" | X1 -> "x1" | X2 -> "x2" | X3 -> "x3"
+ | X4 -> "x4" | X5 -> "x5" | X6 -> "x6" | X7 -> "x7"
+ | X8 -> "x8" | X9 -> "x9" | X10 -> "x10" | X11 -> "x11"
+ | X12 -> "x12" | X13 -> "x13" | X14 -> "x14" | X15 -> "x15"
+ | X16 -> "x16" | X17 -> "x17" | X18 -> "x18" | X19 -> "x19"
+ | X20 -> "x20" | X21 -> "x21" | X22 -> "x22" | X23 -> "x23"
+ | X24 -> "x24" | X25 -> "x25" | X26 -> "x26" | X27 -> "x27"
+ | X28 -> "x28" | X29 -> "x29" | X30 -> "x30"
+
+ let wreg_name = function
+ | X0 -> "w0" | X1 -> "w1" | X2 -> "w2" | X3 -> "w3"
+ | X4 -> "w4" | X5 -> "w5" | X6 -> "w6" | X7 -> "w7"
+ | X8 -> "w8" | X9 -> "w9" | X10 -> "w10" | X11 -> "w11"
+ | X12 -> "w12" | X13 -> "w13" | X14 -> "w14" | X15 -> "w15"
+ | X16 -> "w16" | X17 -> "w17" | X18 -> "w18" | X19 -> "w19"
+ | X20 -> "w20" | X21 -> "w21" | X22 -> "w22" | X23 -> "w23"
+ | X24 -> "w24" | X25 -> "w25" | X26 -> "w26" | X27 -> "w27"
+ | X28 -> "w28" | X29 -> "w29" | X30 -> "w30"
+
+ let xreg0_name = function RR0 r -> xreg_name r | XZR -> "xzr"
+ let wreg0_name = function RR0 r -> wreg_name r | XZR -> "wzr"
+
+ let xregsp_name = function RR1 r -> xreg_name r | XSP -> "sp"
+ let wregsp_name = function RR1 r -> wreg_name r | XSP -> "wsp"
+
+ let dreg_name = function
+ | D0 -> "d0" | D1 -> "d1" | D2 -> "d2" | D3 -> "d3"
+ | D4 -> "d4" | D5 -> "d5" | D6 -> "d6" | D7 -> "d7"
+ | D8 -> "d8" | D9 -> "d9" | D10 -> "d10" | D11 -> "d11"
+ | D12 -> "d12" | D13 -> "d13" | D14 -> "d14" | D15 -> "d15"
+ | D16 -> "d16" | D17 -> "d17" | D18 -> "d18" | D19 -> "d19"
+ | D20 -> "d20" | D21 -> "d21" | D22 -> "d22" | D23 -> "d23"
+ | D24 -> "d24" | D25 -> "d25" | D26 -> "d26" | D27 -> "d27"
+ | D28 -> "d28" | D29 -> "d29" | D30 -> "d30" | D31 -> "d31"
+
+ let sreg_name = function
+ | D0 -> "s0" | D1 -> "s1" | D2 -> "s2" | D3 -> "s3"
+ | D4 -> "s4" | D5 -> "s5" | D6 -> "s6" | D7 -> "s7"
+ | D8 -> "s8" | D9 -> "s9" | D10 -> "s10" | D11 -> "s11"
+ | D12 -> "s12" | D13 -> "s13" | D14 -> "s14" | D15 -> "s15"
+ | D16 -> "s16" | D17 -> "s17" | D18 -> "s18" | D19 -> "s19"
+ | D20 -> "s20" | D21 -> "s21" | D22 -> "s22" | D23 -> "s23"
+ | D24 -> "s24" | D25 -> "s25" | D26 -> "s26" | D27 -> "s27"
+ | D28 -> "s28" | D29 -> "s29" | D30 -> "s30" | D31 -> "s31"
+
+ let xreg oc r = output_string oc (xreg_name r)
+ let wreg oc r = output_string oc (wreg_name r)
+ let ireg oc (sz, r) =
+ output_string oc (match sz with X -> xreg_name r | W -> wreg_name r)
+
+ let xreg0 oc r = output_string oc (xreg0_name r)
+ let wreg0 oc r = output_string oc (wreg0_name r)
+ let ireg0 oc (sz, r) =
+ output_string oc (match sz with X -> xreg0_name r | W -> wreg0_name r)
+
+ let xregsp oc r = output_string oc (xregsp_name r)
+ let iregsp oc (sz, r) =
+ output_string oc (match sz with X -> xregsp_name r | W -> wregsp_name r)
+
+ let dreg oc r = output_string oc (dreg_name r)
+ let sreg oc r = output_string oc (sreg_name r)
+ let freg oc (sz, r) =
+ output_string oc (match sz with D -> dreg_name r | S -> sreg_name r)
+
+ let preg_asm oc ty = function
+ | DR (IR (RR1 r)) -> if ty = Tint then wreg oc r else xreg oc r
+ | DR (FR r) -> if ty = Tsingle then sreg oc r else dreg oc r
+ | _ -> assert false
+
+ let preg_annot = function
+ | DR (IR (RR1 r)) -> xreg_name r
+ | DR (FR r) -> dreg_name r
+ | _ -> assert false
+
+(* Names of sections *)
+
+ let name_of_section = function
+ | Section_text -> ".text"
+ | Section_data(i, true) ->
+ failwith "_Thread_local unsupported on this platform"
+ | Section_data(i, false) | Section_small_data i ->
+ if i then ".data" else common_section ()
+=======
+ let load_symbol_address oc rd id =
+ fprintf oc " adrp %a, :got:%a\n" xreg rd symbol id;
+ fprintf oc " ldr %a, [%a, #:got_lo12:%a]\n" xreg rd xreg rd symbol id
+
+ let name_of_section = function
+ | Section_text -> ".text"
+ | Section_data i | Section_small_data i ->
+ variable_section ~sec:".data" ~bss:".bss" i
+>>>>>>> master
+ | Section_const i | Section_small_const i ->
+ variable_section ~sec:".section .rodata" i
+ | Section_string -> ".section .rodata"
+ | Section_literal -> ".section .rodata"
+ | Section_jumptable -> ".section .rodata"
+ | Section_debug_info _ -> ".section .debug_info,\"\",%progbits"
+ | Section_debug_loc -> ".section .debug_loc,\"\",%progbits"
+ | Section_debug_abbrev -> ".section .debug_abbrev,\"\",%progbits"
+ | Section_debug_line _ -> ".section .debug_line,\"\",%progbits"
+ | Section_debug_ranges -> ".section .debug_ranges,\"\",%progbits"
+ | Section_debug_str -> ".section .debug_str,\"MS\",%progbits,1"
+ | Section_user(s, wr, ex) ->
+ sprintf ".section \"%s\",\"a%s%s\",%%progbits"
+ s (if wr then "w" else "") (if ex then "x" else "")
+ | Section_ais_annotation -> sprintf ".section \"__compcert_ais_annotations\",\"\",@note"
+
+ let print_fun_info = elf_print_fun_info
+ let print_var_info = elf_print_var_info
+
+ let print_comm_decl oc name sz al =
+ fprintf oc " .comm %a, %s, %d\n" symbol name (Z.to_string sz) al
+
+ let print_lcomm_decl oc name sz al =
+ fprintf oc " .local %a\n" symbol name;
+ print_comm_decl oc name sz al
+
+ end
+
+module MacOS_System : SYSTEM =
+ struct
+ let comment = ";"
+
+ let raw_symbol oc s =
+ fprintf oc "_%s" s
+
+ let symbol oc symb =
+ raw_symbol oc (extern_atom symb)
+
+ let symbol_offset_gen kind oc (id, ofs) =
+ fprintf oc "%a@%s" symbol id kind;
+ let ofs = camlint64_of_ptrofs ofs in
+ if ofs <> 0L then fprintf oc " + %Ld" ofs
+
+ let symbol_offset_high = symbol_offset_gen "PAGE"
+ let symbol_offset_low = symbol_offset_gen "PAGEOFF"
+
+ let label oc lbl =
+ fprintf oc "L%d" lbl
+
+ let label_high oc lbl =
+ fprintf oc "%a@PAGE" label lbl
+ let label_low oc lbl =
+ fprintf oc "%a@PAGEOFF" label lbl
+
+ let load_symbol_address oc rd id =
+ fprintf oc " adrp %a, %a@GOTPAGE\n" xreg rd symbol id;
+ fprintf oc " ldr %a, [%a, %a@GOTPAGEOFF]\n" xreg rd xreg rd symbol id
+
+ let name_of_section = function
+ | Section_text -> ".text"
+ | Section_data i | Section_small_data i ->
+ variable_section ~sec:".data" i
+ | Section_const i | Section_small_const i ->
+ variable_section ~sec:".const" ~reloc:".const_data" i
+ | Section_string -> ".const"
+ | Section_literal -> ".const"
+ | Section_jumptable -> ".text"
+ | Section_user(s, wr, ex) ->
+ sprintf ".section \"%s\", %s, %s"
+ (if wr then "__DATA" else "__TEXT") s
+ (if ex then "regular, pure_instructions" else "regular")
+ | Section_debug_info _ -> ".section __DWARF,__debug_info,regular,debug"
+ | Section_debug_loc -> ".section __DWARF,__debug_loc,regular,debug"
+ | Section_debug_line _ -> ".section __DWARF,__debug_line,regular,debug"
+ | Section_debug_str -> ".section __DWARF,__debug_str,regular,debug"
+ | Section_debug_ranges -> ".section __DWARF,__debug_ranges,regular,debug"
+ | Section_debug_abbrev -> ".section __DWARF,__debug_abbrev,regular,debug"
+ | Section_ais_annotation -> assert false (* Not supported under MacOS *)
+
+ let print_fun_info _ _ = ()
+ let print_var_info _ _ = ()
+
+ let print_comm_decl oc name sz al =
+ fprintf oc " .comm %a, %s, %d\n"
+ symbol name (Z.to_string sz) (log2 al)
+
+ let print_lcomm_decl oc name sz al =
+ fprintf oc " .lcomm %a, %s, %d\n"
+ symbol name (Z.to_string sz) (log2 al)
+
+ end
+
+(* Module containing the printing functions *)
+
+module Target(System: SYSTEM): TARGET =
+ struct
+ include System
+
+(* Basic printing functions *)
+
+ let print_label oc lbl = label oc (transl_label lbl)
+
+(* Names of sections *)
+
+ let section oc sec =
+ fprintf oc " %s\n" (name_of_section sec)
+
+(* Associate labels to floating-point constants and to symbols. *)
+
+ let emit_constants oc lit =
+ if exists_constants () then begin
+ section oc lit;
+ if Hashtbl.length literal64_labels > 0 then
+ begin
+ fprintf oc " .balign 8\n";
+ Hashtbl.iter
+ (fun bf lbl -> fprintf oc "%a: .quad 0x%Lx\n" label lbl bf)
+ literal64_labels
+ end;
+ if Hashtbl.length literal32_labels > 0 then
+ begin
+ fprintf oc " .balign 4\n";
+ Hashtbl.iter
+ (fun bf lbl ->
+ fprintf oc "%a: .long 0x%lx\n" label lbl bf)
+ literal32_labels
+ end;
+ reset_literals ()
+ end
+
+(* Emit .file / .loc debugging directives *)
+
+ let print_file_line oc file line =
+ print_file_line oc comment file line
+
+(* Name of testable condition *)
+
+ let condition_name = function
+ | TCeq -> "eq"
+ | TCne -> "ne"
+ | TChs -> "hs"
+ | TClo -> "lo"
+ | TCmi -> "mi"
+ | TCpl -> "pl"
+ | TChi -> "hi"
+ | TCls -> "ls"
+ | TCge -> "ge"
+ | TClt -> "lt"
+ | TCgt -> "gt"
+ | TCle -> "le"
+
+(* Print an addressing mode *)
+
+ let addressing oc = function
+ | ADimm(base, n) -> fprintf oc "[%a, #%a]" xregsp base coqint64 n
+ | ADreg(base, r) -> fprintf oc "[%a, %a]" xregsp base xreg r
+ | ADlsl(base, r, n) -> fprintf oc "[%a, %a, lsl #%a]" xregsp base xreg r coqint n
+ | ADsxt(base, r, n) -> fprintf oc "[%a, %a, sxtw #%a]" xregsp base wreg r coqint n
+ | ADuxt(base, r, n) -> fprintf oc "[%a, %a, uxtw #%a]" xregsp base wreg r coqint n
+ | ADadr(base, id, ofs) -> fprintf oc "[%a, %a]" xregsp base symbol_offset_low (id, ofs)
+ | ADpostincr(base, n) -> fprintf oc "[%a], #%a" xregsp base coqint64 n
+
+(* Print a shifted operand *)
+ let shiftop oc = function
+ | SOnone -> ()
+ | SOlsl n -> fprintf oc ", lsl #%a" coqint n
+ | SOlsr n -> fprintf oc ", lsr #%a" coqint n
+ | SOasr n -> fprintf oc ", asr #%a" coqint n
+ | SOror n -> fprintf oc ", ror #%a" coqint n
+
+(* Print a sign- or zero-extended register operand *)
+ let regextend oc = function
+ | (r, EOsxtb n) -> fprintf oc "%a, sxtb #%a" wreg r coqint n
+ | (r, EOsxth n) -> fprintf oc "%a, sxth #%a" wreg r coqint n
+ | (r, EOsxtw n) -> fprintf oc "%a, sxtw #%a" wreg r coqint n
+ | (r, EOuxtb n) -> fprintf oc "%a, uxtb #%a" wreg r coqint n
+ | (r, EOuxth n) -> fprintf oc "%a, uxth #%a" wreg r coqint n
+ | (r, EOuxtw n) -> fprintf oc "%a, uxtw #%a" wreg r coqint n
+ | (r, EOuxtx n) -> fprintf oc "%a, uxtx #%a" xreg r coqint n
+
+ let next_profiling_label =
+ let atomic_incr_counter = ref 0 in
+ fun () ->
+ let r = sprintf ".Lcompcert_atomic_incr%d" !atomic_incr_counter in
+ incr atomic_incr_counter; r;;
+
+ let print_profiling_logger oc id kind =
+ assert (kind >= 0);
+ assert (kind <= 1);
+ fprintf oc "%s begin profiling %a %d: atomic increment\n" comment
+ Profilingaux.pp_id id kind;
+ let ofs = profiling_offset id kind and lbl = next_profiling_label () in
+ fprintf oc " adrp x15, %s+%d\n" profiling_counter_table_name ofs;
+ fprintf oc " add x15, x15, :lo12:(%s+%d)\n" profiling_counter_table_name ofs;
+ fprintf oc "%s:\n" lbl;
+ fprintf oc " ldaxr x17, [x15]\n";
+ fprintf oc " add x17, x17, 1\n";
+ fprintf oc " stlxr w17, x17, [x15]\n";
+ fprintf oc " cbnz w17, %s\n" lbl;
+ fprintf oc "%s end profiling %a %d\n" comment
+ Profilingaux.pp_id id kind;;
+
+(* Printing of instructions *)
+ let print_instruction oc = function
+ (* Branches *)
+ | Pb lbl ->
+ fprintf oc " b %a\n" print_label lbl
+ | Pbc(c, lbl) ->
+ fprintf oc " b.%s %a\n" (condition_name c) print_label lbl
+ | Pbl(id, sg) ->
+ fprintf oc " bl %a\n" symbol id
+ | Pbs(id, sg) ->
+ fprintf oc " b %a\n" symbol id
+ | Pblr(r, sg) ->
+ fprintf oc " blr %a\n" xreg r
+ | Pbr(r, sg) ->
+ fprintf oc " br %a\n" xreg r
+ | Pret r ->
+ fprintf oc " ret %a\n" xreg r
+ | Pcbnz(sz, r, lbl) ->
+ fprintf oc " cbnz %a, %a\n" ireg (sz, r) print_label lbl
+ | Pcbz(sz, r, lbl) ->
+ fprintf oc " cbz %a, %a\n" ireg (sz, r) print_label lbl
+ | Ptbnz(sz, r, n, lbl) ->
+ fprintf oc " tbnz %a, #%a, %a\n" ireg (sz, r) coqint n print_label lbl
+ | Ptbz(sz, r, n, lbl) ->
+ fprintf oc " tbz %a, #%a, %a\n" ireg (sz, r) coqint n print_label lbl
+ (* Memory loads and stores *)
+ | Pldrw(rd, a) | Pldrw_a(rd, a) ->
+ fprintf oc " ldr %a, %a\n" wreg rd addressing a
+ | Pldrx(rd, a) | Pldrx_a(rd, a) ->
+ fprintf oc " ldr %a, %a\n" xreg rd addressing a
+ | Pldrb(sz, rd, a) ->
+ fprintf oc " ldrb %a, %a\n" wreg rd addressing a
+ | Pldrsb(sz, rd, a) ->
+ fprintf oc " ldrsb %a, %a\n" ireg (sz, rd) addressing a
+ | Pldrh(sz, rd, a) ->
+ fprintf oc " ldrh %a, %a\n" wreg rd addressing a
+ | Pldrsh(sz, rd, a) ->
+ fprintf oc " ldrsh %a, %a\n" ireg (sz, rd) addressing a
+ | Pldrzw(rd, a) ->
+ fprintf oc " ldr %a, %a\n" wreg rd addressing a
+ (* the upper 32 bits of Xrd are set to 0, performing zero-extension *)
+ | Pldrsw(rd, a) ->
+ fprintf oc " ldrsw %a, %a\n" xreg rd addressing a
+ | Pldpw(rd1, rd2, _, _, a) ->
+ fprintf oc " ldp %a, %a, %a\n" wreg rd1 wreg rd2 addressing a
+ | Pldpx(rd1, rd2, _, _, a) ->
+ fprintf oc " ldp %a, %a, %a\n" xreg rd1 xreg rd2 addressing a
+ | Pstrw(rs, a) | Pstrw_a(rs, a) ->
+ fprintf oc " str %a, %a\n" wreg rs addressing a
+ | Pstrx(rs, a) | Pstrx_a(rs, a) ->
+ fprintf oc " str %a, %a\n" xreg rs addressing a
+ | Pstrb(rs, a) ->
+ fprintf oc " strb %a, %a\n" wreg rs addressing a
+ | Pstrh(rs, a) ->
+ fprintf oc " strh %a, %a\n" wreg rs addressing a
+ | Pstpw(rs1, rs2, _, _, a) ->
+ fprintf oc " stp %a, %a, %a\n" wreg rs1 wreg rs2 addressing a
+ | Pstpx(rs1, rs2, _, _, a) ->
+ fprintf oc " stp %a, %a, %a\n" xreg rs1 xreg rs2 addressing a
+ (* Integer arithmetic, immediate *)
+ | Paddimm(sz, rd, r1, n) ->
+ fprintf oc " add %a, %a, #%a\n" iregsp (sz, rd) iregsp (sz, r1) intsz (sz, n)
+ | Psubimm(sz, rd, r1, n) ->
+ fprintf oc " sub %a, %a, #%a\n" iregsp (sz, rd) iregsp (sz, r1) intsz (sz, n)
+ | Pcmpimm(sz, r1, n) ->
+ fprintf oc " cmp %a, #%a\n" ireg (sz, r1) intsz (sz, n)
+ | Pcmnimm(sz, r1, n) ->
+ fprintf oc " cmn %a, #%a\n" ireg (sz, r1) intsz (sz, n)
+ (* Move integer register *)
+ | Pmov(rd, r1) ->
+ fprintf oc " mov %a, %a\n" xregsp rd xregsp r1
+ (* Logical, immediate *)
+ | Pandimm(sz, rd, r1, n) ->
+ fprintf oc " and %a, %a, #%a\n" ireg (sz, rd) ireg0 (sz, r1) intsz (sz, n)
+ | Peorimm(sz, rd, r1, n) ->
+ fprintf oc " eor %a, %a, #%a\n" ireg (sz, rd) ireg0 (sz, r1) intsz (sz, n)
+ | Porrimm(sz, rd, r1, n) ->
+ fprintf oc " orr %a, %a, #%a\n" ireg (sz, rd) ireg0 (sz, r1) intsz (sz, n)
+ | Ptstimm(sz, r1, n) ->
+ fprintf oc " tst %a, #%a\n" ireg (sz, r1) intsz (sz, n)
+ (* Move wide immediate *)
+ | Pmovz(sz, rd, n, pos) ->
+ fprintf oc " movz %a, #%d, lsl #%d\n" ireg (sz, rd) (Z.to_int n) (Z.to_int pos)
+ | Pmovn(sz, rd, n, pos) ->
+ fprintf oc " movn %a, #%d, lsl #%d\n" ireg (sz, rd) (Z.to_int n) (Z.to_int pos)
+ | Pmovk(sz, rd, n, pos) ->
+ fprintf oc " movk %a, #%d, lsl #%d\n" ireg (sz, rd) (Z.to_int n) (Z.to_int pos)
+ (* PC-relative addressing *)
+ | Padrp(rd, id, ofs) ->
+ fprintf oc " adrp %a, %a\n" xreg rd symbol_offset_high (id, ofs)
+ | Paddadr(rd, r1, id, ofs) ->
+ fprintf oc " add %a, %a, %a\n" xreg rd xreg r1 symbol_offset_low (id, ofs)
+ (* Bit-field operations *)
+ | Psbfiz(sz, rd, r1, r, s) ->
+ fprintf oc " sbfiz %a, %a, %a, %d\n" ireg (sz, rd) ireg (sz, r1) coqint r (Z.to_int s)
+ | Psbfx(sz, rd, r1, r, s) ->
+ fprintf oc " sbfx %a, %a, %a, %d\n" ireg (sz, rd) ireg (sz, r1) coqint r (Z.to_int s)
+ | Pubfiz(sz, rd, r1, r, s) ->
+ fprintf oc " ubfiz %a, %a, %a, %d\n" ireg (sz, rd) ireg (sz, r1) coqint r (Z.to_int s)
+ | Pubfx(sz, rd, r1, r, s) ->
+ fprintf oc " ubfx %a, %a, %a, %d\n" ireg (sz, rd) ireg (sz, r1) coqint r (Z.to_int s)
+ (* Integer arithmetic, shifted register *)
+ | Padd(sz, rd, r1, r2, s) ->
+ fprintf oc " add %a, %a, %a%a\n" ireg (sz, rd) ireg0 (sz, r1) ireg (sz, r2) shiftop s
+ | Psub(sz, rd, r1, r2, s) ->
+ fprintf oc " sub %a, %a, %a%a\n" ireg (sz, rd) ireg0 (sz, r1) ireg (sz, r2) shiftop s
+ | Pcmp(sz, r1, r2, s) ->
+ fprintf oc " cmp %a, %a%a\n" ireg0 (sz, r1) ireg (sz, r2) shiftop s
+ | Pcmn(sz, r1, r2, s) ->
+ fprintf oc " cmn %a, %a%a\n" ireg0 (sz, r1) ireg (sz, r2) shiftop s
+ (* Integer arithmetic, extending register *)
+ | Paddext(rd, r1, r2, x) ->
+ fprintf oc " add %a, %a, %a\n" xregsp rd xregsp r1 regextend (r2, x)
+ | Psubext(rd, r1, r2, x) ->
+ fprintf oc " sub %a, %a, %a\n" xregsp rd xregsp r1 regextend (r2, x)
+ | Pcmpext(r1, r2, x) ->
+ fprintf oc " cmp %a, %a\n" xreg r1 regextend (r2, x)
+ | Pcmnext(r1, r2, x) ->
+ fprintf oc " cmn %a, %a\n" xreg r1 regextend (r2, x)
+ (* Logical, shifted register *)
+ | Pand(sz, rd, r1, r2, s) ->
+ fprintf oc " and %a, %a, %a%a\n" ireg (sz, rd) ireg0 (sz, r1) ireg (sz, r2) shiftop s
+ | Pbic(sz, rd, r1, r2, s) ->
+ fprintf oc " bic %a, %a, %a%a\n" ireg (sz, rd) ireg0 (sz, r1) ireg (sz, r2) shiftop s
+ | Peon(sz, rd, r1, r2, s) ->
+ fprintf oc " eon %a, %a, %a%a\n" ireg (sz, rd) ireg0 (sz, r1) ireg (sz, r2) shiftop s
+ | Peor(sz, rd, r1, r2, s) ->
+ fprintf oc " eor %a, %a, %a%a\n" ireg (sz, rd) ireg0 (sz, r1) ireg (sz, r2) shiftop s
+ | Porr(sz, rd, r1, r2, s) ->
+ fprintf oc " orr %a, %a, %a%a\n" ireg (sz, rd) ireg0 (sz, r1) ireg (sz, r2) shiftop s
+ | Porn(sz, rd, r1, r2, s) ->
+ fprintf oc " orn %a, %a, %a%a\n" ireg (sz, rd) ireg0 (sz, r1) ireg (sz, r2) shiftop s
+ | Ptst(sz, r1, r2, s) ->
+ fprintf oc " tst %a, %a%a\n" ireg0 (sz, r1) ireg (sz, r2) shiftop s
+ (* Variable shifts *)
+ | Pasrv(sz, rd, r1, r2) ->
+ fprintf oc " asr %a, %a, %a\n" ireg (sz, rd) ireg (sz, r1) ireg (sz, r2)
+ | Plslv(sz, rd, r1, r2) ->
+ fprintf oc " lsl %a, %a, %a\n" ireg (sz, rd) ireg (sz, r1) ireg (sz, r2)
+ | Plsrv(sz, rd, r1, r2) ->
+ fprintf oc " lsr %a, %a, %a\n" ireg (sz, rd) ireg (sz, r1) ireg (sz, r2)
+ | Prorv(sz, rd, r1, r2) ->
+ fprintf oc " ror %a, %a, %a\n" ireg (sz, rd) ireg (sz, r1) ireg (sz, r2)
+ (* Bit operations *)
+ | Pcls(sz, rd, r1) ->
+ fprintf oc " cls %a, %a\n" ireg (sz, rd) ireg (sz, r1)
+ | Pclz(sz, rd, r1) ->
+ fprintf oc " clz %a, %a\n" ireg (sz, rd) ireg (sz, r1)
+ | Prev(sz, rd, r1) ->
+ fprintf oc " rev %a, %a\n" ireg (sz, rd) ireg (sz, r1)
+ | Prev16(sz, rd, r1) ->
+ fprintf oc " rev16 %a, %a\n" ireg (sz, rd) ireg (sz, r1)
+ | Prbit(sz, rd, r1) ->
+ fprintf oc " rbit %a, %a\n" ireg (sz, rd) ireg (sz, r1)
+ (* Conditional data processing *)
+ | Pcsel(rd, r1, r2, c) ->
+ fprintf oc " csel %a, %a, %a, %s\n" xreg rd xreg r1 xreg r2 (condition_name c)
+ | Pcset(rd, c) ->
+ fprintf oc " cset %a, %s\n" xreg rd (condition_name c)
+ (* Integer multiply/divide *)
+ | Pmadd(sz, rd, r1, r2, r3) ->
+ fprintf oc " madd %a, %a, %a, %a\n" ireg (sz, rd) ireg (sz, r1) ireg (sz, r2) ireg0 (sz, r3)
+ | Pmsub(sz, rd, r1, r2, r3) ->
+ fprintf oc " msub %a, %a, %a, %a\n" ireg (sz, rd) ireg (sz, r1) ireg (sz, r2) ireg0 (sz, r3)
+ | Psmulh(rd, r1, r2) ->
+ fprintf oc " smulh %a, %a, %a\n" xreg rd xreg r1 xreg r2
+ | Pumulh(rd, r1, r2) ->
+ fprintf oc " umulh %a, %a, %a\n" xreg rd xreg r1 xreg r2
+ | Psdiv(sz, rd, r1, r2) ->
+ fprintf oc " sdiv %a, %a, %a\n" ireg (sz, rd) ireg (sz, r1) ireg (sz, r2)
+ | Pudiv(sz, rd, r1, r2) ->
+ fprintf oc " udiv %a, %a, %a\n" ireg (sz, rd) ireg (sz, r1) ireg (sz, r2)
+ (* Floating-point loads and stores *)
+ | Pldrs(rd, a) ->
+ fprintf oc " ldr %a, %a\n" sreg rd addressing a
+ | Pldrd(rd, a) | Pldrd_a(rd, a) ->
+ fprintf oc " ldr %a, %a\n" dreg rd addressing a
+ | Pstrs(rd, a) ->
+ fprintf oc " str %a, %a\n" sreg rd addressing a
+ | Pstrd(rd, a) | Pstrd_a(rd, a) ->
+ fprintf oc " str %a, %a\n" dreg rd addressing a
+ | Pldps(rd1, rd2, _, _, a) ->
+ fprintf oc " ldp %a, %a, %a\n" sreg rd1 sreg rd2 addressing a
+ | Pldpd(rd1, rd2, _, _, a) ->
+ fprintf oc " ldp %a, %a, %a\n" dreg rd1 dreg rd2 addressing a
+ | Pstps(rd1, rd2, _, _, a) ->
+ fprintf oc " stp %a, %a, %a\n" sreg rd1 sreg rd2 addressing a
+ | Pstpd(rd1, rd2, _, _, a) ->
+ fprintf oc " stp %a, %a, %a\n" dreg rd1 dreg rd2 addressing a
+ (* Floating-point move *)
+ | Pfmov(rd, r1) ->
+ fprintf oc " fmov %a, %a\n" dreg rd dreg r1
+ | Pfmovimmd(rd, f) ->
+ let d = camlint64_of_coqint (Floats.Float.to_bits f) in
+ if is_immediate_float64 f then
+ fprintf oc " fmov %a, #%.7f\n" dreg rd (Int64.float_of_bits d)
+ else begin
+ let lbl = label_literal64 d in
+ fprintf oc " adrp x16, %a\n" label_high lbl;
+ fprintf oc " ldr %a, [x16, %a] %s %.18g\n" dreg rd label_low lbl comment (Int64.float_of_bits d)
+ end
+ | Pfmovimms(rd, f) ->
+ let d = camlint_of_coqint (Floats.Float32.to_bits f) in
+ if is_immediate_float32 f then
+ fprintf oc " fmov %a, #%.7f\n" sreg rd (Int32.float_of_bits d)
+ else begin
+ let lbl = label_literal32 d in
+ fprintf oc " adrp x16, %a\n" label_high lbl;
+ fprintf oc " ldr %a, [x16, %a] %s %.18g\n" sreg rd label_low lbl comment (Int32.float_of_bits d)
+ end
+ | Pfmovi(D, rd, r1) ->
+ fprintf oc " fmov %a, %a\n" dreg rd xreg0 r1
+ | Pfmovi(S, rd, r1) ->
+ fprintf oc " fmov %a, %a\n" sreg rd wreg0 r1
+ (* Floating-point conversions *)
+ | Pfcvtds(rd, r1) ->
+ fprintf oc " fcvt %a, %a\n" dreg rd sreg r1
+ | Pfcvtsd(rd, r1) ->
+ fprintf oc " fcvt %a, %a\n" sreg rd dreg r1
+ | Pfcvtzs(isz, fsz, rd, r1) ->
+ fprintf oc " fcvtzs %a, %a\n" ireg (isz, rd) freg (fsz, r1)
+ | Pfcvtzu(isz, fsz, rd, r1) ->
+ fprintf oc " fcvtzu %a, %a\n" ireg (isz, rd) freg (fsz, r1)
+ | Pscvtf(fsz, isz, rd, r1) ->
+ fprintf oc " scvtf %a, %a\n" freg (fsz, rd) ireg (isz, r1)
+ | Pucvtf(fsz, isz, rd, r1) ->
+ fprintf oc " ucvtf %a, %a\n" freg (fsz, rd) ireg (isz, r1)
+ (* Floating-point arithmetic *)
+ | Pfabs(sz, rd, r1) ->
+ fprintf oc " fabs %a, %a\n" freg (sz, rd) freg (sz, r1)
+ | Pfneg(sz, rd, r1) ->
+ fprintf oc " fneg %a, %a\n" freg (sz, rd) freg (sz, r1)
+ | Pfsqrt(sz, rd, r1) ->
+ fprintf oc " fsqrt %a, %a\n" freg (sz, rd) freg (sz, r1)
+ | Pfadd(sz, rd, r1, r2) ->
+ fprintf oc " fadd %a, %a, %a\n" freg (sz, rd) freg (sz, r1) freg (sz, r2)
+ | Pfdiv(sz, rd, r1, r2) ->
+ fprintf oc " fdiv %a, %a, %a\n" freg (sz, rd) freg (sz, r1) freg (sz, r2)
+ | Pfmul(sz, rd, r1, r2) ->
+ fprintf oc " fmul %a, %a, %a\n" freg (sz, rd) freg (sz, r1) freg (sz, r2)
+ | Pfnmul(sz, rd, r1, r2) ->
+ fprintf oc " fnmul %a, %a, %a\n" freg (sz, rd) freg (sz, r1) freg (sz, r2)
+ | Pfsub(sz, rd, r1, r2) ->
+ fprintf oc " fsub %a, %a, %a\n" freg (sz, rd) freg (sz, r1) freg (sz, r2)
+ | Pfmadd(sz, rd, r1, r2, r3) ->
+ fprintf oc " fmadd %a, %a, %a, %a\n" freg (sz, rd) freg (sz, r1) freg (sz, r2) freg (sz, r3)
+ | Pfmsub(sz, rd, r1, r2, r3) ->
+ fprintf oc " fmsub %a, %a, %a, %a\n" freg (sz, rd) freg (sz, r1) freg (sz, r2) freg (sz, r3)
+ | Pfnmadd(sz, rd, r1, r2, r3) ->
+ fprintf oc " fnmadd %a, %a, %a, %a\n" freg (sz, rd) freg (sz, r1) freg (sz, r2) freg (sz, r3)
+ | Pfnmsub(sz, rd, r1, r2, r3) ->
+ fprintf oc " fnmsub %a, %a, %a, %a\n" freg (sz, rd) freg (sz, r1) freg (sz, r2) freg (sz, r3)
+ | Pfmax (sz, rd, r1, r2) ->
+ fprintf oc " fmax %a, %a, %a\n" freg (sz, rd) freg (sz, r1) freg (sz, r2)
+ | Pfmin (sz, rd, r1, r2) ->
+ fprintf oc " fmin %a, %a, %a\n" freg (sz, rd) freg (sz, r1) freg (sz, r2)
+ (* Floating-point comparison *)
+ | Pfcmp(sz, r1, r2) ->
+ fprintf oc " fcmp %a, %a\n" freg (sz, r1) freg (sz, r2)
+ | Pfcmp0(sz, r1) ->
+ fprintf oc " fcmp %a, #0.0\n" freg (sz, r1)
+ (* Floating-point conditional select *)
+ | Pfsel(rd, r1, r2, c) ->
+ fprintf oc " fcsel %a, %a, %a, %s\n" dreg rd dreg r1 dreg r2 (condition_name c)
+ (* No-op *)
+ | Pnop -> ()
+ (*fprintf oc " nop\n"*)
+ (* Pseudo-instructions expanded in Asmexpand *)
+ | Pallocframe(sz, linkofs) -> assert false
+ | Pfreeframe(sz, linkofs) -> assert false
+ | Pcvtx2w rd -> assert false
+ (* Pseudo-instructions not yet expanded *)
+ | Plabel lbl ->
+ fprintf oc "%a:\n" print_label lbl
+ | Ploadsymbol(rd, id) ->
+ load_symbol_address oc rd id
+ | Pcvtsw2x(rd, r1) ->
+ fprintf oc " sxtw %a, %a\n" xreg rd wreg r1
+ | Pcvtuw2x(rd, r1) ->
+ fprintf oc " uxtw %a, %a\n" xreg rd wreg r1
+ | Pbtbl(r1, tbl) ->
+ let lbl = new_label() in
+ fprintf oc " adr x16, %a\n" label lbl;
+ fprintf oc " add x16, x16, %a, uxtw #2\n" wreg r1;
+ fprintf oc " br x16\n";
+ fprintf oc "%a:" label lbl;
+ List.iter (fun l -> fprintf oc " b %a\n" print_label l) tbl
+ | Pcfi_adjust sz ->
+ cfi_adjust oc (camlint_of_coqint sz)
+ | Pcfi_rel_offset ofs ->
+ cfi_rel_offset oc "lr" (camlint_of_coqint ofs)
+ | Pbuiltin(ef, args, res) ->
+ begin match ef with
+ | EF_annot(kind,txt, targs) ->
+ begin match (P.to_int kind) with
+ | 1 -> let annot = annot_text preg_annot "sp" (camlstring_of_coqstring txt) args in
+ fprintf oc "%s annotation: %S\n" comment annot
+ | 2 -> let lbl = new_label () in
+ fprintf oc "%a:\n" label lbl;
+ add_ais_annot lbl preg_annot "sp" (camlstring_of_coqstring txt) args
+ | _ -> assert false
+ end
+ | EF_debug(kind, txt, targs) ->
+ print_debug_info comment print_file_line preg_annot "sp" oc
+ (P.to_int kind) (extern_atom txt) args
+ | EF_inline_asm(txt, sg, clob) ->
+ fprintf oc "%s begin inline assembly\n\t" comment;
+ print_inline_asm preg_asm oc (camlstring_of_coqstring txt) sg args res;
+ fprintf oc "%s end inline assembly\n" comment
+ | EF_profiling (id, coq_kind) ->
+ print_profiling_logger oc id (Z.to_int coq_kind)
+ | _ ->
+ assert false
+ end
+
+ let get_section_names name =
+ let (text, lit) =
+ match C2C.atom_sections name with
+ | t :: l :: _ -> (t, l)
+ | _ -> (Section_text, Section_literal) in
+ text,lit,Section_jumptable
+
+ let print_align oc alignment =
+ fprintf oc " .balign %d\n" alignment
+
+ let print_jumptable oc jmptbl =
+ let print_tbl oc (lbl, tbl) =
+ fprintf oc "%a:\n" label lbl;
+ List.iter
+ (fun l -> fprintf oc " .long %a - %a\n"
+ print_label l label lbl)
+ tbl in
+ if !jumptables <> [] then
+ begin
+ section oc jmptbl;
+ fprintf oc " .balign 4\n";
+ List.iter (print_tbl oc) !jumptables;
+ jumptables := []
+ end
+
+ let print_optional_fun_info _ = ()
+
+ let print_comm_symb oc sz name align =
+ if C2C.atom_is_static name
+ then print_lcomm_decl oc name sz align
+ else print_comm_decl oc name sz align
+
+ let print_instructions oc fn =
+ current_function_sig := fn.fn_sig;
+ List.iter (print_instruction oc) fn.fn_code
+
+(* Data *)
+
+ let address = ".quad"
+
+ let print_prologue oc =
+ if !Clflags.option_g then begin
+ section oc Section_text;
+ end
+
+ let aarch64_profiling_stub oc nr_items
+ profiling_id_table_name
+ profiling_counter_table_name =
+ fprintf oc " adrp x2, %s\n" profiling_counter_table_name;
+ fprintf oc " adrp x1, %s\n" profiling_id_table_name;
+ fprintf oc " add x2, x2, :lo12:%s\n" profiling_counter_table_name;
+ fprintf oc " add x1, x1, :lo12:%s\n" profiling_id_table_name;
+ fprintf oc " mov w0, %d\n" nr_items;
+ fprintf oc " b %s\n" profiling_write_table_helper ;;
+
+ let print_atexit oc to_be_called =
+ fprintf oc " adrp x0, %s\n" to_be_called;
+ fprintf oc " add x0, x0, :lo12:%s\n" to_be_called;
+ fprintf oc " b atexit\n";;
+
+
+ let print_epilogue oc =
+ print_profiling_epilogue elf_text_print_fun_info (Init_atexit print_atexit) aarch64_profiling_stub oc;
+ if !Clflags.option_g then begin
+ Debug.compute_gnu_file_enum (fun f -> ignore (print_file oc f));
+ section oc Section_text;
+ end
+
+ let default_falignment = 4
+
+ let cfi_startproc oc = ()
+ let cfi_endproc oc = ()
+
+ end
+
+let sel_target () =
+ let module S =
+ (val (match Configuration.system with
+ | "linux" -> (module ELF_System : SYSTEM)
+ | "macos" -> (module MacOS_System : SYSTEM)
+ | _ -> invalid_arg ("System " ^ Configuration.system ^ " not supported"))
+ : SYSTEM) in
+ (module Target(S) : TARGET)
diff --git a/aarch64/TO_MERGE/extractionMachdep.v b/aarch64/TO_MERGE/extractionMachdep.v
new file mode 100644
index 00000000..947fa38b
--- /dev/null
+++ b/aarch64/TO_MERGE/extractionMachdep.v
@@ -0,0 +1,45 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, Collège de France and INRIA Paris *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the GNU General Public License as published by *)
+(* the Free Software Foundation, either version 2 of the License, or *)
+(* (at your option) any later version. This file is also distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(* Additional extraction directives specific to the AArch64 port *)
+
+Require Archi Asm Asmgen SelectOp.
+
+(* Archi *)
+
+Extract Constant Archi.abi =>
+ "match Configuration.abi with
+ | ""apple"" -> Apple
+ | _ -> AAPCS64".
+
+(* SelectOp *)
+
+Extract Constant SelectOp.symbol_is_relocatable =>
+ "match Configuration.system with
+ | ""macos"" -> C2C.atom_is_extern
+ | _ -> (fun _ -> false)".
+
+(* Asm *)
+
+Extract Constant Asm.symbol_low => "fun _ _ _ -> assert false".
+Extract Constant Asm.symbol_high => "fun _ _ _ -> assert false".
+<<<<<<< HEAD
+Extract Constant Asmblockgen.symbol_is_aligned => "C2C.atom_is_aligned".
+=======
+
+(* Asmgen *)
+
+Extract Constant Asmgen.symbol_is_aligned => "C2C.atom_is_aligned".
+>>>>>>> master