aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
-rw-r--r--aarch64/Asm.v542
-rw-r--r--aarch64/Asmblock.v2046
-rw-r--r--aarch64/Asmblockgen.v1197
-rw-r--r--aarch64/Asmblockgenproof.v1158
-rw-r--r--aarch64/Asmblockgenproof1.v (renamed from aarch64/Asmgenproof1.v)3
-rw-r--r--aarch64/Asmgen.v1448
-rw-r--r--aarch64/Asmgenproof.v2880
-rwxr-xr-xconfigure15
-rw-r--r--kvx/lib/IterList.v96
-rw-r--r--kvx/lib/OptionMonad.v49
-rw-r--r--kvx/lib/PseudoAsmblock.v267
-rw-r--r--kvx/lib/PseudoAsmblockproof.v1173
12 files changed, 8297 insertions, 2577 deletions
diff --git a/aarch64/Asm.v b/aarch64/Asm.v
index 47cd3051..4de1a9d0 100644
--- a/aarch64/Asm.v
+++ b/aarch64/Asm.v
@@ -10,149 +10,28 @@
(* *)
(* *********************************************************************)
-(** Abstract syntax and semantics for AArch64 assembly language *)
+(** Abstract syntax and semantics for AArch64 assembly language
+
+
+W.r.t Asmblock, this is a "flat" syntax and semantics of "Aarch64" assembly:
+ - without the basic block structure
+ - without the hierarchy of instructions.
+
+
+TODO: with respect to original file, remove parts of the syntax/semantics
+that have been already defined in Asmblock.
+
+*)
Require Import Coqlib Zbits Maps.
Require Import AST Integers Floats.
Require Import Values Memory Events Globalenvs Smallstep.
Require Import Locations Conventions.
Require Stacklayout.
+Require Import OptionMonad Asmblock.
(** * Abstract syntax *)
-(** Integer registers, floating-point registers. *)
-
-(** In assembly files, [Xn] denotes the full 64-bit register
- and [Wn] the low 32 bits of [Xn]. *)
-
-Inductive ireg: Type :=
- | X0 | X1 | X2 | X3 | X4 | X5 | X6 | X7
- | X8 | X9 | X10 | X11 | X12 | X13 | X14 | X15
- | X16 | X17 | X18 | X19 | X20 | X21 | X22 | X23
- | X24 | X25 | X26 | X27 | X28 | X29 | X30.
-
-Inductive ireg0: Type :=
- | RR0 (r: ireg) | XZR.
-
-Inductive iregsp: Type :=
- | RR1 (r: ireg) | XSP.
-
-Coercion RR0: ireg >-> ireg0.
-Coercion RR1: ireg >-> iregsp.
-
-Lemma ireg_eq: forall (x y: ireg), {x=y} + {x<>y}.
-Proof. decide equality. Defined.
-
-(** In assembly files, [Dn] denotes the low 64-bit of a vector register,
- and [Sn] the low 32 bits. *)
-
-Inductive freg: Type :=
- | D0 | D1 | D2 | D3 | D4 | D5 | D6 | D7
- | D8 | D9 | D10 | D11 | D12 | D13 | D14 | D15
- | D16 | D17 | D18 | D19 | D20 | D21 | D22 | D23
- | D24 | D25 | D26 | D27 | D28 | D29 | D30 | D31.
-
-Lemma freg_eq: forall (x y: freg), {x=y} + {x<>y}.
-Proof. decide equality. Defined.
-
-(** Bits in the condition register. *)
-
-Inductive crbit: Type :=
- | CN: crbit (**r negative *)
- | CZ: crbit (**r zero *)
- | CC: crbit (**r carry *)
- | CV: crbit. (**r overflow *)
-
-Lemma crbit_eq: forall (x y: crbit), {x=y} + {x<>y}.
-Proof. decide equality. Defined.
-
-(** We model the following registers of the ARM architecture. *)
-
-Inductive preg: Type :=
- | IR: ireg -> preg (**r 64- or 32-bit integer registers *)
- | FR: freg -> preg (**r double- or single-precision float registers *)
- | CR: crbit -> preg (**r bits in the condition register *)
- | SP: preg (**r register X31 used as stack pointer *)
- | PC: preg. (**r program counter *)
-
-Coercion IR: ireg >-> preg.
-Coercion FR: freg >-> preg.
-Coercion CR: crbit >-> preg.
-
-Lemma preg_eq: forall (x y: preg), {x=y} + {x<>y}.
-Proof. decide equality. apply ireg_eq. apply freg_eq. apply crbit_eq. Defined.
-
-Module PregEq.
- Definition t := preg.
- Definition eq := preg_eq.
-End PregEq.
-
-Module Pregmap := EMap(PregEq).
-
-Definition preg_of_iregsp (r: iregsp) : preg :=
- match r with RR1 r => IR r | XSP => SP end.
-
-Coercion preg_of_iregsp: iregsp >-> preg.
-
-(** Conventional name for return address ([RA]) *)
-
-Notation "'RA'" := X30 (only parsing) : asm.
-
-(** The instruction set. Most instructions correspond exactly to
- actual AArch64 instructions. See the ARM reference manuals for more
- details. Some instructions, described below, are
- pseudo-instructions: they expand to canned instruction sequences
- during the printing of the assembly code. *)
-
-Definition label := positive.
-
-Inductive isize: Type :=
- | W (**r 32-bit integer operation *)
- | X. (**r 64-bit integer operation *)
-
-Inductive fsize: Type :=
- | S (**r 32-bit, single-precision FP operation *)
- | D. (**r 64-bit, double-precision FP operation *)
-
-Inductive testcond : Type :=
- | TCeq: testcond (**r equal *)
- | TCne: testcond (**r not equal *)
- | TChs: testcond (**r unsigned higher or same *)
- | TClo: testcond (**r unsigned lower *)
- | TCmi: testcond (**r negative *)
- | TCpl: testcond (**r positive *)
- | TChi: testcond (**r unsigned higher *)
- | TCls: testcond (**r unsigned lower or same *)
- | TCge: testcond (**r signed greater or equal *)
- | TClt: testcond (**r signed less than *)
- | TCgt: testcond (**r signed greater *)
- | TCle: testcond. (**r signed less than or equal *)
-
-Inductive addressing: Type :=
- | ADimm (base: iregsp) (n: int64) (**r base plus immediate offset *)
- | ADreg (base: iregsp) (r: ireg) (**r base plus reg *)
- | ADlsl (base: iregsp) (r: ireg) (n: int) (**r base plus reg LSL n *)
- | ADsxt (base: iregsp) (r: ireg) (n: int) (**r base plus SIGN-EXT(reg) LSL n *)
- | ADuxt (base: iregsp) (r: ireg) (n: int) (**r base plus ZERO-EXT(reg) LSL n *)
- | ADadr (base: iregsp) (id: ident) (ofs: ptrofs) (**r base plus low address of [id + ofs] *)
- | ADpostincr (base: iregsp) (n: int64). (**r base plus offset; base is updated after *)
-
-Inductive shift_op: Type :=
- | SOnone
- | SOlsl (n: int)
- | SOlsr (n: int)
- | SOasr (n: int)
- | SOror (n: int).
-
-Inductive extend_op: Type :=
- | EOsxtb (n: int)
- | EOsxth (n: int)
- | EOsxtw (n: int)
- | EOuxtb (n: int)
- | EOuxth (n: int)
- | EOuxtw (n: int)
- | EOuxtx (n: int).
-
Inductive instruction: Type :=
(** Branches *)
| Pb (lbl: label) (**r branch *)
@@ -307,65 +186,8 @@ Definition code := list instruction.
Record function : Type := mkfunction { fn_sig: signature; fn_code: code }.
Definition fundef := AST.fundef function.
Definition program := AST.program fundef unit.
-
-(** * Operational semantics *)
-
-(** The semantics operates over a single mapping from registers
- (type [preg]) to values. We maintain (but do not enforce)
- the convention that integer registers are mapped to values of
- type [Tint], float registers to values of type [Tfloat],
- and condition bits to either [Vzero] or [Vone]. *)
-
-Definition regset := Pregmap.t val.
Definition genv := Genv.t fundef unit.
-(** The value of an [ireg0] is either the value of the integer register,
- or 0. *)
-
-Definition ir0w (rs: regset) (r: ireg0) : val :=
- match r with RR0 r => rs (IR r) | XZR => Vint Int.zero end.
-Definition ir0x (rs: regset) (r: ireg0) : val :=
- match r with RR0 r => rs (IR r) | XZR => Vlong Int64.zero end.
-
-(** Concise notations for accessing and updating the values of registers. *)
-
-Notation "a # b" := (a b) (at level 1, only parsing) : asm.
-Notation "a # b <- c" := (Pregmap.set b c a) (at level 1, b at next level) : asm.
-Notation "a ## b" := (ir0w a b) (at level 1, only parsing) : asm.
-Notation "a ### b" := (ir0x a b) (at level 1, only parsing) : asm.
-
-Open Scope asm.
-
-(** Undefining some registers *)
-
-Fixpoint undef_regs (l: list preg) (rs: regset) : regset :=
- match l with
- | nil => rs
- | r :: l' => undef_regs l' (rs#r <- Vundef)
- end.
-
-(** Undefining the condition codes *)
-
-Definition undef_flags (rs: regset) : regset :=
- fun r => match r with CR _ => Vundef | _ => rs r end.
-
-(** Assigning a register pair *)
-
-Definition set_pair (p: rpair preg) (v: val) (rs: regset) : regset :=
- match p with
- | One r => rs#r <- v
- | Twolong rhi rlo => rs#rhi <- (Val.hiword v) #rlo <- (Val.loword v)
- end.
-
-(** Assigning the result of a builtin *)
-
-Fixpoint set_res (res: builtin_res preg) (v: val) (rs: regset) : regset :=
- match res with
- | BR r => rs#r <- v
- | BR_none => rs
- | BR_splitlong hi lo => set_res lo (Val.loword v) (set_res hi (Val.hiword v) rs)
- end.
-
(** The two functions below axiomatize how the linker processes
symbolic references [symbol + offset]. It computes the
difference between the address and the PC, and splits it into:
@@ -385,6 +207,8 @@ Axiom symbol_high_low:
forall (ge: genv) (id: ident) (ofs: ptrofs),
Val.addl (symbol_high ge id ofs) (symbol_low ge id ofs) = Genv.symbol_address ge id ofs.
+(** * Operational semantics *)
+
Section RELSEM.
Variable ge: genv.
@@ -416,22 +240,9 @@ Fixpoint label_pos (lbl: label) (pos: Z) (c: code) {struct c} : option Z :=
match c with
| nil => None
| instr :: c' =>
- if is_label lbl instr then Some (pos + 1) else label_pos lbl (pos + 1) c'
+ if is_label lbl instr then Some pos else label_pos lbl (pos + 1) c'
end.
-(** The semantics is purely small-step and defined as a function
- from the current state (a register set + a memory state)
- to either [Next rs' m'] where [rs'] and [m'] are the updated register
- set and memory state after execution of the instruction at [rs#PC],
- or [Stuck] if the processor is stuck. *)
-
-Inductive outcome: Type :=
- | Next: regset -> mem -> outcome
- | Stuck: outcome.
-
-(** Manipulations over the [PC] register: continuing with the next
- instruction ([nextinstr]) or branching to a label ([goto_label]). *)
-
Definition nextinstr (rs: regset) :=
rs#PC <- (Val.offset_ptr rs#PC Ptrofs.one).
@@ -445,88 +256,6 @@ Definition goto_label (f: function) (lbl: label) (rs: regset) (m: mem) :=
end
end.
-(** Testing a condition *)
-
-Definition eval_testcond (c: testcond) (rs: regset) : option bool :=
- match c with
- | TCeq => (**r equal *)
- match rs#CZ with
- | Vint n => Some (Int.eq n Int.one)
- | _ => None
- end
- | TCne => (**r not equal *)
- match rs#CZ with
- | Vint n => Some (Int.eq n Int.zero)
- | _ => None
- end
- | TClo => (**r unsigned less than *)
- match rs#CC with
- | Vint n => Some (Int.eq n Int.zero)
- | _ => None
- end
- | TCls => (**r unsigned less or equal *)
- match rs#CC, rs#CZ with
- | Vint c, Vint z => Some (Int.eq c Int.zero || Int.eq z Int.one)
- | _, _ => None
- end
- | TChs => (**r unsigned greater or equal *)
- match rs#CC with
- | Vint n => Some (Int.eq n Int.one)
- | _ => None
- end
- | TChi => (**r unsigned greater *)
- match rs#CC, rs#CZ with
- | Vint c, Vint z => Some (Int.eq c Int.one && Int.eq z Int.zero)
- | _, _ => None
- end
- | TClt => (**r signed less than *)
- match rs#CV, rs#CN with
- | Vint o, Vint s => Some (Int.eq (Int.xor o s) Int.one)
- | _, _ => None
- end
- | TCle => (**r signed less or equal *)
- match rs#CV, rs#CN, rs#CZ with
- | Vint o, Vint s, Vint z => Some (Int.eq (Int.xor o s) Int.one || Int.eq z Int.one)
- | _, _, _ => None
- end
- | TCge => (**r signed greater or equal *)
- match rs#CV, rs#CN with
- | Vint o, Vint s => Some (Int.eq (Int.xor o s) Int.zero)
- | _, _ => None
- end
- | TCgt => (**r signed greater *)
- match rs#CV, rs#CN, rs#CZ with
- | Vint o, Vint s, Vint z => Some (Int.eq (Int.xor o s) Int.zero && Int.eq z Int.zero)
- | _, _, _ => None
- end
- | TCpl => (**r positive *)
- match rs#CN with
- | Vint n => Some (Int.eq n Int.zero)
- | _ => None
- end
- | TCmi => (**r negative *)
- match rs#CN with
- | Vint n => Some (Int.eq n Int.one)
- | _ => None
- end
- end.
-
-(** Integer "is zero?" test *)
-
-Definition eval_testzero (sz: isize) (v: val) (m: mem): option bool :=
- match sz with
- | W => Val.cmpu_bool (Mem.valid_pointer m) Ceq v (Vint Int.zero)
- | X => Val.cmplu_bool (Mem.valid_pointer m) Ceq v (Vlong Int64.zero)
- end.
-
-(** Integer "bit is set?" test *)
-
-Definition eval_testbit (sz: isize) (v: val) (n: int): option bool :=
- match sz with
- | W => Val.cmp_bool Cne (Val.and v (Vint (Int.shl Int.one n))) (Vint Int.zero)
- | X => Val.cmpl_bool Cne (Val.andl v (Vlong (Int64.shl' Int64.one n))) (Vlong Int64.zero)
- end.
-
(** Evaluating an addressing mode *)
Definition eval_addressing (a: addressing) (rs: regset): val :=
@@ -557,130 +286,6 @@ Definition exec_store (chunk: memory_chunk)
| Some m' => Next (nextinstr rs) m'
end.
-(** Comparisons *)
-
-Definition compare_int (rs: regset) (v1 v2: val) (m: mem) :=
- rs#CN <- (Val.negative (Val.sub v1 v2))
- #CZ <- (Val.cmpu (Mem.valid_pointer m) Ceq v1 v2)
- #CC <- (Val.cmpu (Mem.valid_pointer m) Cge v1 v2)
- #CV <- (Val.sub_overflow v1 v2).
-
-Definition compare_long (rs: regset) (v1 v2: val) (m: mem) :=
- rs#CN <- (Val.negativel (Val.subl v1 v2))
- #CZ <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Ceq v1 v2))
- #CC <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Cge v1 v2))
- #CV <- (Val.subl_overflow v1 v2).
-
-(** Semantics of [fcmp] instructions:
-<<
-== N=0 Z=1 C=1 V=0
-< N=1 Z=0 C=0 V=0
-> N=0 Z=0 C=1 V=0
-unord N=0 Z=0 C=1 V=1
->>
-*)
-
-Definition compare_float (rs: regset) (v1 v2: val) :=
- match v1, v2 with
- | Vfloat f1, Vfloat f2 =>
- rs#CN <- (Val.of_bool (Float.cmp Clt f1 f2))
- #CZ <- (Val.of_bool (Float.cmp Ceq f1 f2))
- #CC <- (Val.of_bool (negb (Float.cmp Clt f1 f2)))
- #CV <- (Val.of_bool (negb (Float.ordered f1 f2)))
- | _, _ =>
- rs#CN <- Vundef
- #CZ <- Vundef
- #CC <- Vundef
- #CV <- Vundef
- end.
-
-Definition compare_single (rs: regset) (v1 v2: val) :=
- match v1, v2 with
- | Vsingle f1, Vsingle f2 =>
- rs#CN <- (Val.of_bool (Float32.cmp Clt f1 f2))
- #CZ <- (Val.of_bool (Float32.cmp Ceq f1 f2))
- #CC <- (Val.of_bool (negb (Float32.cmp Clt f1 f2)))
- #CV <- (Val.of_bool (negb (Float32.ordered f1 f2)))
- | _, _ =>
- rs#CN <- Vundef
- #CZ <- Vundef
- #CC <- Vundef
- #CV <- Vundef
- end.
-
-(** Insertion of bits into an integer *)
-
-Definition insert_in_int (x: val) (y: Z) (pos: Z) (len: Z) : val :=
- match x with
- | Vint n => Vint (Int.repr (Zinsert (Int.unsigned n) y pos len))
- | _ => Vundef
- end.
-
-Definition insert_in_long (x: val) (y: Z) (pos: Z) (len: Z) : val :=
- match x with
- | Vlong n => Vlong (Int64.repr (Zinsert (Int64.unsigned n) y pos len))
- | _ => Vundef
- end.
-
-(** Evaluation of shifted operands *)
-
-Definition eval_shift_op_int (v: val) (s: shift_op): val :=
- match s with
- | SOnone => v
- | SOlsl n => Val.shl v (Vint n)
- | SOlsr n => Val.shru v (Vint n)
- | SOasr n => Val.shr v (Vint n)
- | SOror n => Val.ror v (Vint n)
- end.
-
-Definition eval_shift_op_long (v: val) (s: shift_op): val :=
- match s with
- | SOnone => v
- | SOlsl n => Val.shll v (Vint n)
- | SOlsr n => Val.shrlu v (Vint n)
- | SOasr n => Val.shrl v (Vint n)
- | SOror n => Val.rorl v (Vint n)
- end.
-
-(** Evaluation of sign- or zero- extended operands *)
-
-Definition eval_extend (v: val) (x: extend_op): val :=
- match x with
- | EOsxtb n => Val.shll (Val.longofint (Val.sign_ext 8 v)) (Vint n)
- | EOsxth n => Val.shll (Val.longofint (Val.sign_ext 16 v)) (Vint n)
- | EOsxtw n => Val.shll (Val.longofint v) (Vint n)
- | EOuxtb n => Val.shll (Val.longofintu (Val.zero_ext 8 v)) (Vint n)
- | EOuxth n => Val.shll (Val.longofintu (Val.zero_ext 16 v)) (Vint n)
- | EOuxtw n => Val.shll (Val.longofintu v) (Vint n)
- | EOuxtx n => Val.shll v (Vint n)
- end.
-
-(** Bit-level conversion from integers to FP numbers *)
-
-Definition float32_of_bits (v: val): val :=
- match v with
- | Vint n => Vsingle (Float32.of_bits n)
- | _ => Vundef
- end.
-
-Definition float64_of_bits (v: val): val :=
- match v with
- | Vlong n => Vfloat (Float.of_bits n)
- | _ => Vundef
- end.
-
-(** Execution of a single instruction [i] in initial state
- [rs] and [m]. Return updated state. For instructions
- that correspond to actual AArch64 instructions, the cases are
- straightforward transliterations of the informal descriptions
- given in the ARMv8 reference manuals. For pseudo-instructions,
- refer to the informal descriptions given above.
-
- Note that we set to [Vundef] the registers used as temporaries by
- the expansions of the pseudo-instructions, so that the code we
- generate cannot use those registers to hold values that must
- survive the execution of the pseudo-instruction.
-*)
Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : outcome :=
match i with
@@ -704,13 +309,13 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
| Pret r =>
Next (rs#PC <- (rs#r)) m
| Pcbnz sz r lbl =>
- match eval_testzero sz rs#r m with
+ match eval_testzero sz rs#r with
| Some true => Next (nextinstr rs) m
| Some false => goto_label f lbl rs m
| None => Stuck
end
| Pcbz sz r lbl =>
- match eval_testzero sz rs#r m with
+ match eval_testzero sz rs#r with
| Some true => goto_label f lbl rs m
| Some false => Next (nextinstr rs) m
| None => Stuck
@@ -778,13 +383,13 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
| Psubimm X rd r1 n =>
Next (nextinstr (rs#rd <- (Val.subl rs#r1 (Vlong (Int64.repr n))))) m
| Pcmpimm W r1 n =>
- Next (nextinstr (compare_int rs rs#r1 (Vint (Int.repr n)) m)) m
+ Next (nextinstr (compare_int rs rs#r1 (Vint (Int.repr n)))) m
| Pcmpimm X r1 n =>
- Next (nextinstr (compare_long rs rs#r1 (Vlong (Int64.repr n)) m)) m
+ Next (nextinstr (compare_long rs rs#r1 (Vlong (Int64.repr n)))) m
| Pcmnimm W r1 n =>
- Next (nextinstr (compare_int rs rs#r1 (Vint (Int.neg (Int.repr n))) m)) m
+ Next (nextinstr (compare_int rs rs#r1 (Vint (Int.neg (Int.repr n))))) m
| Pcmnimm X r1 n =>
- Next (nextinstr (compare_long rs rs#r1 (Vlong (Int64.neg (Int64.repr n))) m)) m
+ Next (nextinstr (compare_long rs rs#r1 (Vlong (Int64.neg (Int64.repr n))))) m
(** Move integer register *)
| Pmov rd r1 =>
Next (nextinstr (rs#rd <- (rs#r1))) m
@@ -802,9 +407,9 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
| Porrimm X rd r1 n =>
Next (nextinstr (rs#rd <- (Val.orl rs###r1 (Vlong (Int64.repr n))))) m
| Ptstimm W r1 n =>
- Next (nextinstr (compare_int rs (Val.and rs#r1 (Vint (Int.repr n))) (Vint Int.zero) m)) m
+ Next (nextinstr (compare_int rs (Val.and rs#r1 (Vint (Int.repr n))) (Vint Int.zero))) m
| Ptstimm X r1 n =>
- Next (nextinstr (compare_long rs (Val.andl rs#r1 (Vlong (Int64.repr n))) (Vlong Int64.zero) m)) m
+ Next (nextinstr (compare_long rs (Val.andl rs#r1 (Vlong (Int64.repr n))) (Vlong Int64.zero))) m
(** Move wide immediate *)
| Pmovz W rd n pos =>
Next (nextinstr (rs#rd <- (Vint (Int.repr (Z.shiftl n pos))))) m
@@ -850,22 +455,22 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
| Psub X rd r1 r2 s =>
Next (nextinstr (rs#rd <- (Val.subl rs###r1 (eval_shift_op_long rs#r2 s)))) m
| Pcmp W r1 r2 s =>
- Next (nextinstr (compare_int rs rs##r1 (eval_shift_op_int rs#r2 s) m)) m
+ Next (nextinstr (compare_int rs rs##r1 (eval_shift_op_int rs#r2 s))) m
| Pcmp X r1 r2 s =>
- Next (nextinstr (compare_long rs rs###r1 (eval_shift_op_long rs#r2 s) m)) m
+ Next (nextinstr (compare_long rs rs###r1 (eval_shift_op_long rs#r2 s))) m
| Pcmn W r1 r2 s =>
- Next (nextinstr (compare_int rs rs##r1 (Val.neg (eval_shift_op_int rs#r2 s)) m)) m
+ Next (nextinstr (compare_int rs rs##r1 (Val.neg (eval_shift_op_int rs#r2 s)))) m
| Pcmn X r1 r2 s =>
- Next (nextinstr (compare_long rs rs###r1 (Val.negl (eval_shift_op_long rs#r2 s)) m)) m
+ Next (nextinstr (compare_long rs rs###r1 (Val.negl (eval_shift_op_long rs#r2 s)))) m
(** Integer arithmetic, extending register *)
| Paddext rd r1 r2 x =>
Next (nextinstr (rs#rd <- (Val.addl rs#r1 (eval_extend rs#r2 x)))) m
| Psubext rd r1 r2 x =>
Next (nextinstr (rs#rd <- (Val.subl rs#r1 (eval_extend rs#r2 x)))) m
| Pcmpext r1 r2 x =>
- Next (nextinstr (compare_long rs rs#r1 (eval_extend rs#r2 x) m)) m
+ Next (nextinstr (compare_long rs rs#r1 (eval_extend rs#r2 x))) m
| Pcmnext r1 r2 x =>
- Next (nextinstr (compare_long rs rs#r1 (Val.negl (eval_extend rs#r2 x)) m)) m
+ Next (nextinstr (compare_long rs rs#r1 (Val.negl (eval_extend rs#r2 x)))) m
(** Logical, shifted register *)
| Pand W rd r1 r2 s =>
Next (nextinstr (rs#rd <- (Val.and rs##r1 (eval_shift_op_int rs#r2 s)))) m
@@ -892,9 +497,9 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
| Porn X rd r1 r2 s =>
Next (nextinstr (rs#rd <- (Val.orl rs###r1 (Val.notl (eval_shift_op_long rs#r2 s))))) m
| Ptst W r1 r2 s =>
- Next (nextinstr (compare_int rs (Val.and rs##r1 (eval_shift_op_int rs#r2 s)) (Vint Int.zero) m)) m
+ Next (nextinstr (compare_int rs (Val.and rs##r1 (eval_shift_op_int rs#r2 s)) (Vint Int.zero))) m
| Ptst X r1 r2 s =>
- Next (nextinstr (compare_long rs (Val.andl rs###r1 (eval_shift_op_long rs#r2 s)) (Vlong Int64.zero) m)) m
+ Next (nextinstr (compare_long rs (Val.andl rs###r1 (eval_shift_op_long rs#r2 s)) (Vlong Int64.zero))) m
(** Variable shifts *)
| Pasrv W rd r1 r2 =>
Next (nextinstr (rs#rd <- (Val.shr rs#r1 rs#r2))) m
@@ -1118,80 +723,9 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
Stuck
end.
-(** Translation of the LTL/Linear/Mach view of machine registers
- to the AArch64 view. Note that no LTL register maps to [X16],
- [X18], nor [X30].
- [X18] is reserved as the platform register and never used by the
- code generated by CompCert.
- [X30] is used for the return address, and can also be used as temporary.
- [X16] can be used as temporary. *)
-
-Definition preg_of (r: mreg) : preg :=
- match r with
- | R0 => X0 | R1 => X1 | R2 => X2 | R3 => X3
- | R4 => X4 | R5 => X5 | R6 => X6 | R7 => X7
- | R8 => X8 | R9 => X9 | R10 => X10 | R11 => X11
- | R12 => X12 | R13 => X13 | R14 => X14 | R15 => X15
- | R17 => X17 | R19 => X19
- | R20 => X20 | R21 => X21 | R22 => X22 | R23 => X23
- | R24 => X24 | R25 => X25 | R26 => X26 | R27 => X27
- | R28 => X28 | R29 => X29
- | F0 => D0 | F1 => D1 | F2 => D2 | F3 => D3
- | F4 => D4 | F5 => D5 | F6 => D6 | F7 => D7
- | F8 => D8 | F9 => D9 | F10 => D10 | F11 => D11
- | F12 => D12 | F13 => D13 | F14 => D14 | F15 => D15
- | F16 => D16 | F17 => D17 | F18 => D18 | F19 => D19
- | F20 => D20 | F21 => D21 | F22 => D22 | F23 => D23
- | F24 => D24 | F25 => D25 | F26 => D26 | F27 => D27
- | F28 => D28 | F29 => D29 | F30 => D30 | F31 => D31
- end.
-
-(** Undefine all registers except SP and callee-save registers *)
-
-Definition undef_caller_save_regs (rs: regset) : regset :=
- fun r =>
- if preg_eq r SP
- || In_dec preg_eq r (List.map preg_of (List.filter is_callee_save all_mregs))
- then rs r
- else Vundef.
-
-(** Extract the values of the arguments of an external call.
- We exploit the calling conventions from module [Conventions], except that
- we use AArch64 registers instead of locations. *)
-
-Inductive extcall_arg (rs: regset) (m: mem): loc -> val -> Prop :=
- | extcall_arg_reg: forall r,
- extcall_arg rs m (R r) (rs (preg_of r))
- | extcall_arg_stack: forall ofs ty bofs v,
- bofs = Stacklayout.fe_ofs_arg + 4 * ofs ->
- Mem.loadv (chunk_of_type ty) m
- (Val.offset_ptr rs#SP (Ptrofs.repr bofs)) = Some v ->
- extcall_arg rs m (Locations.S Outgoing ofs ty) v.
-
-Inductive extcall_arg_pair (rs: regset) (m: mem): rpair loc -> val -> Prop :=
- | extcall_arg_one: forall l v,
- extcall_arg rs m l v ->
- extcall_arg_pair rs m (One l) v
- | extcall_arg_twolong: forall hi lo vhi vlo,
- extcall_arg rs m hi vhi ->
- extcall_arg rs m lo vlo ->
- extcall_arg_pair rs m (Twolong hi lo) (Val.longofwords vhi vlo).
-
-Definition extcall_arguments
- (rs: regset) (m: mem) (sg: signature) (args: list val) : Prop :=
- list_forall2 (extcall_arg_pair rs m) (loc_arguments sg) args.
-
-Definition loc_external_result (sg: signature) : rpair preg :=
- map_rpair preg_of (loc_result sg).
-
-(** Execution of the instruction at [rs#PC]. *)
-
-Inductive state: Type :=
- | State: regset -> mem -> state.
-
Inductive step: state -> trace -> state -> Prop :=
| exec_step_internal:
- forall b ofs f i rs m rs' m',
+ forall b ofs (f:function) i rs m rs' m',
rs PC = Vptr b ofs ->
Genv.find_funct_ptr ge b = Some (Internal f) ->
find_instr (Ptrofs.unsigned ofs) f.(fn_code) = Some i ->
@@ -1302,11 +836,11 @@ Qed.
Definition data_preg (r: preg) : bool :=
match r with
- | IR X16 => false
- | IR X30 => false
- | IR _ => true
- | FR _ => true
+ | DR (IR' X16) => false
+ | DR (IR' X30) => false
+ | DR (IR' _) => true
+ | DR (FR' _) => true
| CR _ => false
- | SP => true
+ (* | SP => true; subsumed by IR (iregsp) *)
| PC => false
end.
diff --git a/aarch64/Asmblock.v b/aarch64/Asmblock.v
new file mode 100644
index 00000000..950e28e9
--- /dev/null
+++ b/aarch64/Asmblock.v
@@ -0,0 +1,2046 @@
+(* *************************************************************)
+(* *)
+(* 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 *)
+(* *)
+(* Copyright Kalray. Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+
+(* Asmblock language for aarch64
+
+WORK IN PROGRESS: we want to define an Asmblock syntax, with an Asmblock semantics
+(e.g. we don't need the parallel semantics of Asmvliw)
+
+
+NOTE: this file is inspired from
+ - aarch64/Asm.v
+ - kvx/Asmvliw.v (only the Asmblock syntax)
+ - kvx/Asmblock.v
+*)
+
+
+(** Abstract syntax and semantics for AArch64 assembly language *)
+
+Require Import Coqlib Zbits Maps.
+Require Import AST Integers Floats.
+Require Import Values Memory Events Globalenvs Smallstep.
+Require Import Locations Conventions.
+Require Stacklayout.
+Require Import OptionMonad.
+
+(** * Abstract syntax *)
+
+(** Integer registers, floating-point registers. *)
+
+(** In assembly files, [Xn] denotes the full 64-bit register
+ and [Wn] the low 32 bits of [Xn]. *)
+
+Inductive ireg: Type :=
+ | X0 | X1 | X2 | X3 | X4 | X5 | X6 | X7
+ | X8 | X9 | X10 | X11 | X12 | X13 | X14 | X15
+ | X16 | X17 | X18 | X19 | X20 | X21 | X22 | X23
+ | X24 | X25 | X26 | X27 | X28 | X29 | X30.
+
+Inductive ireg0: Type :=
+ | RR0 (r: ireg) | XZR.
+
+Inductive iregsp: Type :=
+ | RR1 (r: ireg) | XSP.
+
+Coercion RR0: ireg >-> ireg0.
+Coercion RR1: ireg >-> iregsp.
+
+Lemma ireg_eq: forall (x y: ireg), {x=y} + {x<>y}.
+Proof. decide equality. Defined.
+
+Lemma iregsp_eq: forall (x y: iregsp), {x=y} + {x<>y}.
+Proof. decide equality; apply ireg_eq. Defined.
+
+(** In assembly files, [Dn] denotes the low 64-bit of a vector register,
+ and [Sn] the low 32 bits. *)
+
+Inductive freg: Type :=
+ | D0 | D1 | D2 | D3 | D4 | D5 | D6 | D7
+ | D8 | D9 | D10 | D11 | D12 | D13 | D14 | D15
+ | D16 | D17 | D18 | D19 | D20 | D21 | D22 | D23
+ | D24 | D25 | D26 | D27 | D28 | D29 | D30 | D31.
+
+Lemma freg_eq: forall (x y: freg), {x=y} + {x<>y}.
+Proof. decide equality. Defined.
+
+(** Bits in the condition register. *)
+
+Inductive crbit: Type :=
+ | CN: crbit (**r negative *)
+ | CZ: crbit (**r zero *)
+ | CC: crbit (**r carry *)
+ | CV: crbit. (**r overflow *)
+
+Lemma crbit_eq: forall (x y: crbit), {x=y} + {x<>y}.
+Proof. decide equality. Defined.
+
+Inductive dreg : Type :=
+ | IR': iregsp -> dreg (**r 64- or 32-bit integer registers *)
+ | FR': freg -> dreg. (**r double- or single-precision float registers *)
+
+(** We model the following registers of the ARM architecture. *)
+
+Inductive preg: Type :=
+ | DR: dreg -> preg (** see dreg **)
+ | CR: crbit -> preg (**r bits in the condition register *)
+ | PC: preg. (**r program counter *)
+
+(* XXX: ireg no longer coerces to preg *)
+Coercion IR': iregsp >-> dreg.
+Coercion FR': freg >-> dreg.
+Coercion DR: dreg >-> preg.
+Definition SP:preg := XSP.
+Coercion CR: crbit >-> preg.
+
+Lemma dreg_eq: forall (x y: dreg), {x=y} + {x<>y}.
+Proof. decide equality. apply iregsp_eq. apply freg_eq. Defined.
+
+Lemma preg_eq: forall (x y: preg), {x=y} + {x<>y}.
+Proof. decide equality. apply dreg_eq. apply crbit_eq. Defined.
+
+Module PregEq.
+ Definition t := preg.
+ Definition eq := preg_eq.
+End PregEq.
+
+Module Pregmap := EMap(PregEq).
+
+Definition preg_of_iregsp (r: iregsp) : preg :=
+ match r with RR1 r => IR' r | XSP => SP end.
+
+(** Conventional name for return address ([RA]) *)
+
+Notation "'RA'" := X30 (only parsing) : asm.
+
+(** The instruction set. Most instructions correspond exactly to
+ actual AArch64 instructions. See the ARM reference manuals for more
+ details. Some instructions, described below, are
+ pseudo-instructions: they expand to canned instruction sequences
+ during the printing of the assembly code. *)
+
+Definition label := positive.
+
+Inductive isize: Type :=
+ | W (**r 32-bit integer operation *)
+ | X. (**r 64-bit integer operation *)
+
+Inductive fsize: Type :=
+ | S (**r 32-bit, single-precision FP operation *)
+ | D. (**r 64-bit, double-precision FP operation *)
+
+Inductive testcond : Type :=
+ | TCeq: testcond (**r equal *)
+ | TCne: testcond (**r not equal *)
+ | TChs: testcond (**r unsigned higher or same *)
+ | TClo: testcond (**r unsigned lower *)
+ | TCmi: testcond (**r negative *)
+ | TCpl: testcond (**r positive *)
+ | TChi: testcond (**r unsigned higher *)
+ | TCls: testcond (**r unsigned lower or same *)
+ | TCge: testcond (**r signed greater or equal *)
+ | TClt: testcond (**r signed less than *)
+ | TCgt: testcond (**r signed greater *)
+ | TCle: testcond. (**r signed less than or equal *)
+
+Inductive addressing: Type :=
+ | ADimm (base: iregsp) (n: int64) (**r base plus immediate offset *)
+ | ADreg (base: iregsp) (r: ireg) (**r base plus reg *)
+ | ADlsl (base: iregsp) (r: ireg) (n: int) (**r base plus reg LSL n *)
+ | ADsxt (base: iregsp) (r: ireg) (n: int) (**r base plus SIGN-EXT(reg) LSL n *)
+ | ADuxt (base: iregsp) (r: ireg) (n: int) (**r base plus ZERO-EXT(reg) LSL n *)
+ | ADadr (base: iregsp) (id: ident) (ofs: ptrofs) (**r base plus low address of [id + ofs] *)
+ | ADpostincr (base: iregsp) (n: int64). (**r base plus offset; base is updated after *)
+
+Inductive shift_op: Type :=
+ | SOnone
+ | SOlsl (n: int)
+ | SOlsr (n: int)
+ | SOasr (n: int)
+ | SOror (n: int).
+
+Inductive extend_op: Type :=
+ | EOsxtb (n: int)
+ | EOsxth (n: int)
+ | EOsxtw (n: int)
+ | EOuxtb (n: int)
+ | EOuxth (n: int)
+ | EOuxtw (n: int)
+ | EOuxtx (n: int).
+
+(* First task: splitting the big [instruction] type below into CFI and basic instructions.
+ Actually a finer splitting in order to regroup "similar" instructions could be much better for automation of the scheduler proof!
+ e.g. "similar" means identical "interface" w.r.t. pseudo-registers when translated to AbstractBB,
+ or with a "similar" semantics.
+
+ see example of loads below.
+*)
+
+(** Control Flow instructions
+
+*)
+Inductive cf_instruction : Type :=
+ | Pb (lbl: label) (**r branch *)
+ | Pbc (c: testcond) (lbl: label) (**r conditional branch *)
+ | Pbl (id: ident) (sg: signature) (**r jump to function and link *)
+ | Pbs (id: ident) (sg: signature) (**r jump to function *)
+ | Pblr (r: ireg) (sg: signature) (**r indirect jump and link *)
+ | Pbr (r: ireg) (sg: signature) (**r indirect jump *)
+ | Pret (r: ireg) (**r return *)
+ | Pcbnz (sz: isize) (r: ireg) (lbl: label) (**r branch if not zero *)
+ | Pcbz (sz: isize) (r: ireg) (lbl: label) (**r branch if zero *)
+ | Ptbnz (sz: isize) (r: ireg) (n: int) (lbl: label) (**r branch if bit n is not zero *)
+ | Ptbz (sz: isize) (r: ireg) (n: int) (lbl: label) (**r branch if bit n is zero *)
+ (** Pseudo-instructions *)
+ | Pbtbl (r1: ireg) (tbl: list label) (**r N-way branch through a jump table *)
+ .
+
+(*
+A builtin is considered as a control-flow instruction, because it could emit a trace (cf. Machblock semantics).
+Here, we do not need to have builtins alone in basic-blocks (on the contrary to KVX bundles).
+*)
+
+Inductive control: Type :=
+ | PCtlFlow (i: cf_instruction)
+ (** Pseudo-instructions *)
+ | Pbuiltin (ef: external_function)
+ (args: list (builtin_arg preg)) (res: builtin_res preg) (**r built-in function (pseudo) *)
+ .
+
+Coercion PCtlFlow: cf_instruction >-> control.
+
+(** Basic instructions *)
+
+(* Loads waiting for (rd: dreg) (a: addressing)
+ * XXX Use dreg because exec_load is defined in terms of it, thus allowing us to
+ * treat integer and floating point loads the same. *)
+Inductive load_rd_a: Type :=
+ (* Integer load *)
+ | Pldrw (**r load int32 *)
+ | Pldrw_a (**r load int32 as any32 *)
+ | Pldrx (**r load int64 *)
+ | Pldrx_a (**r load int64 as any64 *)
+ | Pldrb (sz: isize) (**r load int8, zero-extend *)
+ | Pldrsb (sz: isize) (**r load int8, sign-extend *)
+ | Pldrh (sz: isize) (**r load int16, zero-extend *)
+ | Pldrsh (sz: isize) (**r load int16, sign-extend *)
+ | Pldrzw (**r load int32, zero-extend to int64 *)
+ | Pldrsw (**r load int32, sign-extend to int64 *)
+ (* Floating-point load *)
+ | Pldrs (**r load float32 (single precision) *)
+ | Pldrd (**r load float64 (double precision) *)
+ | Pldrd_a (**r load float64 as any64 *)
+ .
+
+(* Actually, Pldp is not used in the aarch64/Asm semantics!
+
+Thus we skip this particular case:
+
+Inductive ld_instruction: Type :=
+ | PLd_rd_a (ld: load_rd_a) (rd: ireg) (a: addressing)
+ | Pldp (rd1 rd2: ireg) (a: addressing) (**r load two int64 *)
+ .
+
+Coercion PLoad: ld_instruction >-> basic.
+
+*)
+
+Inductive store_rs_a : Type :=
+ (* Integer store *)
+ | Pstrw (**r store int32 *)
+ | Pstrw_a (**r store int32 as any32 *)
+ | Pstrx (**r store int64 *)
+ | Pstrx_a (**r store int64 as any64 *)
+ | Pstrb (**r store int8 *)
+ | Pstrh (**r store int16 *)
+ (* Floating-point store *)
+ | Pstrs (**r store float32 *)
+ | Pstrd (**r store float64 *)
+ | Pstrd_a (**r store float64 as any64 *)
+ .
+(* Actually, Pstp is not used in the aarch64/Asm semantics!
+ * Thus we skip this particular case:
+ * Inductive st_instruction : Type :=
+ * | PSt_rs_a (st : store_rs_a) (rs : ireg) (a : addressing)
+ * | Pstp (rs1 rs2: ireg) (a: addressing) (**r store two int64 *)
+ * .
+ * Coercion PStore : st_instruction >-> basic.
+ *)
+
+Inductive arith_p : Type :=
+ (** PC-relative addressing *)
+ | Padrp (id: ident) (ofs: ptrofs) (**r set [rd] to high address of [id + ofs] *)
+ (** Move wide immediate *)
+ | Pmovz (sz: isize) (n: Z) (pos: Z) (**r move [n << pos] to [rd] *)
+ | Pmovn (sz: isize) (n: Z) (pos: Z) (**r move [NOT(n << pos)] to [rd] *)
+ (** Floating-point move *)
+ | Pfmovimms (f: float32) (**r load float32 constant *)
+ | Pfmovimmd (f: float) (**r load float64 constant *)
+.
+
+Inductive arith_comparison_p : Type :=
+ (** Floating-point comparison *)
+ | Pfcmp0 (sz: fsize) (**r compare [r1] and [+0.0] *)
+ (** Integer arithmetic, immediate *)
+ | Pcmpimm (sz: isize) (n: Z) (**r compare *)
+ | Pcmnimm (sz: isize) (n: Z) (**r compare negative *)
+ (** Logical, immediate *)
+ | Ptstimm (sz: isize) (n: Z) (**r and, then set flags *)
+.
+
+Inductive arith_pp : Type :=
+ (** Move integer register *)
+ | Pmov
+ (** Move wide immediate *)
+ (* XXX: has to have the same register supplied both times *)
+ | Pmovk (sz: isize) (n: Z) (pos: Z) (**r insert 16 bits of [n] at [pos] in rd *)
+ (** PC-relative addressing *)
+ | Paddadr (id: ident) (ofs: ptrofs) (**r add the low address of [id + ofs] *)
+ (** Bit-field operations *)
+ | Psbfiz (sz: isize) (r: int) (s: Z) (**r sign extend and shift left *)
+ | Psbfx (sz: isize) (r: int) (s: Z) (**r shift right and sign extend *)
+ | Pubfiz (sz: isize) (r: int) (s: Z) (**r zero extend and shift left *)
+ | Pubfx (sz: isize) (r: int) (s: Z) (**r shift right and zero extend *)
+(* Bit operations are not used in the aarch64/asm semantics
+ (** Bit operations *)
+ | Pcls (sz: isize) (**r count leading sign bits *)
+ | Pclz (sz: isize) (**r count leading zero bits *)
+ | Prev (sz: isize) (**r reverse bytes *)
+ | Prev16 (sz: isize) (**r reverse bytes in each 16-bit word *)
+*)
+ (** Floating-point move *)
+ | Pfmov
+ (** Floating-point conversions *)
+ | Pfcvtds (**r convert float32 to float64 *)
+ | Pfcvtsd (**r convert float64 to float32 *)
+ (** Floating-point arithmetic *)
+ | Pfabs (sz: fsize) (**r absolute value *)
+ | Pfneg (sz: fsize) (**r negation *)
+ (* Pfsqrt is not used in the semantics of aarch64/asm
+ | Pfsqrt (sz: fsize) (**r square root *) *)
+ (** Floating-point conversions *)
+ | Pscvtf (fsz: fsize) (isz: isize) (**r convert signed int to float *)
+ | Pucvtf (fsz: fsize) (isz: isize) (**r convert unsigned int to float *)
+ | Pfcvtzs (isz: isize) (fsz: fsize) (**r convert float to signed int *)
+ | Pfcvtzu (isz: isize) (fsz: fsize) (**r convert float to unsigned int *)
+ (** Integer arithmetic, immediate *)
+ | Paddimm (sz: isize) (n: Z) (**r addition *)
+ | Psubimm (sz: isize) (n: Z) (**r subtraction *)
+.
+
+Inductive arith_comparison_r0r : Type :=
+ (** Integer arithmetic, shifted register *)
+ | Pcmp (is:isize) (s: shift_op) (**r compare *)
+ | Pcmn (is:isize) (s: shift_op) (**r compare negative *)
+ (** Logical, shifted register *)
+ | Ptst (is:isize) (s: shift_op) (**r and, then set flags *)
+.
+
+Inductive arith_comparison_pp : Type :=
+ (** Integer arithmetic, extending register *)
+ | Pcmpext (x: extend_op) (**r int64-int32 cmp *)
+ | Pcmnext (x: extend_op) (**r int64-int32 cmn *)
+ (** Floating-point comparison *)
+ | Pfcmp (sz: fsize) (**r compare [r1] and [r2] *)
+.
+
+Inductive arith_ppp : Type :=
+ (** Variable shifts *)
+ | Pasrv (sz: isize) (**r arithmetic right shift *)
+ | Plslv (sz: isize) (**r left shift *)
+ | Plsrv (sz: isize) (**r logical right shift *)
+ | Prorv (sz: isize) (**r rotate right *)
+ (** Integer multiply/divide *)
+ | Psmulh (**r signed multiply high *)
+ | Pumulh (**r unsigned multiply high *)
+ | Psdiv (sz: isize) (**r signed division *)
+ | Pudiv (sz: isize) (**r unsigned division *)
+ (** Integer arithmetic, extending register *)
+ | Paddext (x: extend_op) (**r int64-int32 add *)
+ | Psubext (x: extend_op) (**r int64-int32 sub *)
+ (** Floating-point arithmetic *)
+ | Pfadd (sz: fsize) (**r addition *)
+ | Pfdiv (sz: fsize) (**r division *)
+ | Pfmul (sz: fsize) (**r multiplication *)
+ | Pfsub (sz: fsize) (**r subtraction *)
+.
+
+Inductive arith_rr0r : Type :=
+ (** Integer arithmetic, shifted register *)
+ | Padd (sz:isize) (s: shift_op) (**r addition *)
+ | Psub (sz:isize) (s: shift_op) (**r subtraction *)
+ (** Logical, shifted register *)
+ | Pand (sz:isize) (s: shift_op) (**r and *)
+ | Pbic (sz:isize) (s: shift_op) (**r and-not *)
+ | Peon (sz:isize) (s: shift_op) (**r xor-not *)
+ | Peor (sz:isize) (s: shift_op) (**r xor *)
+ | Porr (sz:isize) (s: shift_op) (**r or *)
+ | Porn (sz:isize) (s: shift_op) (**r or-not *)
+.
+
+
+Inductive arith_rr0 : Type :=
+ (** Logical, immediate *)
+ | Pandimm (sz: isize) (n: Z) (**r and *)
+ | Peorimm (sz: isize) (n: Z) (**r xor *)
+ | Porrimm (sz: isize) (n: Z) (**r or *)
+.
+
+Inductive arith_arrrr0 : Type :=
+ (** Integer multiply/divide *)
+ | Pmadd (sz: isize) (**r multiply-add *)
+ | Pmsub (sz: isize) (**r multiply-sub *)
+.
+
+(* Currently not used by the semantics of aarch64/Asm
+ * Inductive arith_apppp : Type :=
+ * (** Floating-point arithmetic *)
+ * | Pfmadd (sz: fsize) (**r [rd = r3 + r1 * r2] *)
+ * | Pfmsub (sz: fsize) (**r [rd = r3 - r1 * r2] *)
+ * .
+
+ * Inductive arith_aapppp : Type :=
+ * (** Floating-point arithmetic *)
+ * | Pfnmadd (sz: fsize) (**r [rd = - r3 - r1 * r2] *)
+ * | Pfnmsub (sz: fsize) (**r [rd = - r3 + r1 * r2] *)
+ * . *)
+
+(* Notes on the naming scheme used here:
+ * R: ireg
+ * R0: ireg0
+ * Rsp: iregsp
+ * F: freg
+ * W/X: Occur in conjunction with R0, says whether an ireg0 should be evaluated
+ * as W regsiter (32 bit) or X register (64 bit)
+ * S/D: Used for completeness sake. Only used for copying an integer register
+ * to a floating point register. Could be removed.
+ * A: These instructions perform an additional arithmetic operation
+ XXX Does this interpretation match the use in kvx/Asmvliw?
+ * Comparison: For these instructions the first register is not the target.
+ * Instead, the condition register is mutated.
+ *)
+Inductive ar_instruction : Type :=
+ | PArithP (i : arith_p) (rd : dreg)
+ | PArithPP (i : arith_pp) (rd rs : dreg)
+ | PArithPPP (i : arith_ppp) (rd r1 r2 : dreg)
+ | PArithRR0R (i : arith_rr0r) (rd : ireg) (r1 : ireg0) (r2 : ireg)
+ | PArithRR0 (i : arith_rr0) (rd : ireg) (r1 : ireg0)
+ | PArithARRRR0 (i : arith_arrrr0) (rd r1 r2 : ireg) (r3 : ireg0)
+ (* Pfmadd and Pfmsub are currently not used by the semantics of aarch64/Asm
+ | PArithAPPPP (i : arith_apppp) (rd r1 r2 r3 : preg) *)
+ (* Pfnmadd and Pfnmsub are currently not used by the semantics of aarch64/Asm
+ | PArithAAPPPP (i : arith_aapppp) (rd r1 r2 r3 : preg) *)
+ | PArithComparisonPP (i : arith_comparison_pp) (r1 r2 : dreg)
+ | PArithComparisonR0R (i : arith_comparison_r0r) (r1 : ireg0) (r2 : ireg)
+ | PArithComparisonP (i : arith_comparison_p) (r1 : dreg)
+ (* Instructions without indirection sine they would be the only ones *)
+ (* PArithCP: Pcsetm is commented out by aarch64/Asm, so Pcset is alone *)
+ | Pcset (rd : ireg) (c : testcond) (**r set to 1/0 if cond is true/false *)
+ (* PArithFR0 *)
+ | Pfmovi (fsz : fsize) (rd : freg) (r1 : ireg0) (**r copy int reg to FP reg *)
+ (* PArithCPPP *)
+ | Pcsel (rd r1 r2 : dreg) (c : testcond) (**r int/float conditional move *)
+ (* PArithAFFF *)
+ | Pfnmul (fsz : fsize) (rd r1 r2 : freg) (**r multiply-negate *)
+.
+
+Inductive basic : Type :=
+ | PArith (i: ar_instruction) (**r TODO *)
+ | PLoad (ld: load_rd_a) (rd: dreg) (a: addressing) (**r TODO *)
+ | PStore (st: store_rs_a) (r: dreg) (a: addressing) (**r TODO *)
+ | Pallocframe (sz: Z) (linkofs: ptrofs) (**r allocate new stack frame *)
+ | Pfreeframe (sz: Z) (linkofs: ptrofs) (**r deallocate stack frame and restore previous frame *)
+ | Ploadsymbol (rd: ireg) (id: ident) (**r load the address of [id] *)
+ | Pcvtsw2x (rd: ireg) (r1: ireg) (**r sign-extend 32-bit int to 64-bit *)
+ | Pcvtuw2x (rd: ireg) (r1: ireg) (**r zero-extend 32-bit int to 64-bit *)
+ | Pcvtx2w (rd: ireg) (**r retype a 64-bit int as a 32-bit int *)
+(* NOT USED IN THE SEMANTICS !
+ | Pnop (**r no operation *)
+ | Pcfi_adjust (ofs: int) (**r .cfi_adjust debug directive *)
+ | Pcfi_rel_offset (ofs: int) (**r .cfi_rel_offset debug directive *)
+*)
+.
+
+(** * Definition of a bblock
+
+A bblock must contain at least one instruction.
+
+This choice simplifies the definition of [find_bblock] below:
+indeed, each address of a code block identifies at most one bblock.
+*)
+
+Definition non_empty_body (body: list basic): bool :=
+ match body with
+ | nil => false
+ | _ => true
+ end.
+
+Definition non_empty_exit (exit: option control): bool :=
+ match exit with
+ | None => false
+ | _ => true
+ end.
+
+Definition non_empty_bblockb (body: list basic) (exit: option control): bool := non_empty_body body || non_empty_exit exit.
+
+(** A bblock is well-formed if he contains at least one instruction. *)
+
+Record bblock := mk_bblock {
+ header: list label;
+ body: list basic;
+ exit: option control;
+ correct: Is_true (non_empty_bblockb body exit)
+}.
+
+(* FIXME? redundant with definition in Machblock *)
+Definition length_opt {A} (o: option A) : nat :=
+ match o with
+ | Some o => 1
+ | None => 0
+ end.
+
+(* This notion of size induces the notion of "valid" code address given by [find_bblock]
+
+ The result is in Z to be compatible with operations on PC.
+*)
+Definition size (b:bblock): Z := Z.of_nat (length (header b) + length (body b) + length_opt (exit b)).
+
+Definition bblocks := list bblock.
+
+Fixpoint size_blocks (l: bblocks): Z :=
+ match l with
+ | nil => 0
+ | b :: l =>
+ (size b) + (size_blocks l)
+ end
+ .
+
+Record function : Type := mkfunction { fn_sig: signature; fn_blocks: bblocks }.
+Definition fundef := AST.fundef function.
+Definition program := AST.program fundef unit.
+
+(* TODO: ORIGINAL ABSTRACT SYNTAX OF INSTRUCTIONS THAT NEEDS TO BE ADAPTED !
+
+Inductive instruction: Type :=
+ (** Branches *)
+ | Pb (lbl: label) (**r branch *)
+ | Pbc (c: testcond) (lbl: label) (**r conditional branch *)
+ | Pbl (id: ident) (sg: signature) (**r jump to function and link *)
+ | Pbs (id: ident) (sg: signature) (**r jump to function *)
+ | Pblr (r: ireg) (sg: signature) (**r indirect jump and link *)
+ | Pbr (r: ireg) (sg: signature) (**r indirect jump *)
+ | Pret (r: ireg) (**r return *)
+ | Pcbnz (sz: isize) (r: ireg) (lbl: label) (**r branch if not zero *)
+ | Pcbz (sz: isize) (r: ireg) (lbl: label) (**r branch if zero *)
+ | Ptbnz (sz: isize) (r: ireg) (n: int) (lbl: label) (**r branch if bit n is not zero *)
+ | Ptbz (sz: isize) (r: ireg) (n: int) (lbl: label) (**r branch if bit n is zero *)
+ (** Memory loads and stores *)
+ | Pldrw (rd: ireg) (a: addressing) (**r load int32 *)
+ | Pldrw_a (rd: ireg) (a: addressing) (**r load int32 as any32 *)
+ | Pldrx (rd: ireg) (a: addressing) (**r load int64 *)
+ | Pldrx_a (rd: ireg) (a: addressing) (**r load int64 as any64 *)
+ | Pldrb (sz: isize) (rd: ireg) (a: addressing) (**r load int8, zero-extend *)
+ | Pldrsb (sz: isize) (rd: ireg) (a: addressing) (**r load int8, sign-extend *)
+ | Pldrh (sz: isize) (rd: ireg) (a: addressing) (**r load int16, zero-extend *)
+ | Pldrsh (sz: isize) (rd: ireg) (a: addressing) (**r load int16, sign-extend *)
+ | Pldrzw (rd: ireg) (a: addressing) (**r load int32, zero-extend to int64 *)
+ | Pldrsw (rd: ireg) (a: addressing) (**r load int32, sign-extend to int64 *)
+ | Pldp (rd1 rd2: ireg) (a: addressing) (**r load two int64 *)
+ | Pstrw (rs: ireg) (a: addressing) (**r store int32 *)
+ | Pstrw_a (rs: ireg) (a: addressing) (**r store int32 as any32 *)
+ | Pstrx (rs: ireg) (a: addressing) (**r store int64 *)
+ | Pstrx_a (rs: ireg) (a: addressing) (**r store int64 as any64 *)
+ | Pstrb (rs: ireg) (a: addressing) (**r store int8 *)
+ | Pstrh (rs: ireg) (a: addressing) (**r store int16 *)
+ | Pstp (rs1 rs2: ireg) (a: addressing) (**r store two int64 *)
+ (** Integer arithmetic, immediate *)
+ | Paddimm (sz: isize) (rd: iregsp) (r1: iregsp) (n: Z) (**r addition *)
+ | Psubimm (sz: isize) (rd: iregsp) (r1: iregsp) (n: Z) (**r subtraction *)
+ | Pcmpimm (sz: isize) (r1: ireg) (n: Z) (**r compare *)
+ | Pcmnimm (sz: isize) (r1: ireg) (n: Z) (**r compare negative *)
+ (** Move integer register *)
+ | Pmov (rd: iregsp) (r1: iregsp)
+ (** Logical, immediate *)
+ | Pandimm (sz: isize) (rd: ireg) (r1: ireg0) (n: Z) (**r and *)
+ | Peorimm (sz: isize) (rd: ireg) (r1: ireg0) (n: Z) (**r xor *)
+ | Porrimm (sz: isize) (rd: ireg) (r1: ireg0) (n: Z) (**r or *)
+ | Ptstimm (sz: isize) (r1: ireg) (n: Z) (**r and, then set flags *)
+ (** Move wide immediate *)
+ | Pmovz (sz: isize) (rd: ireg) (n: Z) (pos: Z) (**r move [n << pos] to [rd] *)
+ | Pmovn (sz: isize) (rd: ireg) (n: Z) (pos: Z) (**r move [NOT(n << pos)] to [rd] *)
+ | Pmovk (sz: isize) (rd: ireg) (n: Z) (pos: Z) (**r insert 16 bits of [n] at [pos] in rd *)
+ (** PC-relative addressing *)
+ | Padrp (rd: ireg) (id: ident) (ofs: ptrofs) (**r set [rd] to high address of [id + ofs] *)
+ | Paddadr (rd: ireg) (r1: ireg) (id: ident) (ofs: ptrofs) (**r add the low address of [id + ofs] *)
+ (** Bit-field operations *)
+ | Psbfiz (sz: isize) (rd: ireg) (r1: ireg) (r: int) (s: Z) (**r sign extend and shift left *)
+ | Psbfx (sz: isize) (rd: ireg) (r1: ireg) (r: int) (s: Z) (**r shift right and sign extend *)
+ | Pubfiz (sz: isize) (rd: ireg) (r1: ireg) (r: int) (s: Z) (**r zero extend and shift left *)
+ | Pubfx (sz: isize) (rd: ireg) (r1: ireg) (r: int) (s: Z) (**r shift right and zero extend *)
+ (** Integer arithmetic, shifted register *)
+ | Padd (sz: isize) (rd: ireg) (r1: ireg0) (r2: ireg) (s: shift_op) (**r addition *)
+ | Psub (sz: isize) (rd: ireg) (r1: ireg0) (r2: ireg) (s: shift_op) (**r subtraction *)
+ | Pcmp (sz: isize) (r1: ireg0) (r2: ireg) (s: shift_op) (**r compare *)
+ | Pcmn (sz: isize) (r1: ireg0) (r2: ireg) (s: shift_op) (**r compare negative *)
+ (** Integer arithmetic, extending register *)
+ | Paddext (rd: iregsp) (r1: iregsp) (r2: ireg) (x: extend_op) (**r int64-int32 add *)
+ | Psubext (rd: iregsp) (r1: iregsp) (r2: ireg) (x: extend_op) (**r int64-int32 sub *)
+ | Pcmpext (r1: ireg) (r2: ireg) (x: extend_op) (**r int64-int32 cmp *)
+ | Pcmnext (r1: ireg) (r2: ireg) (x: extend_op) (**r int64-int32 cmn *)
+ (** Logical, shifted register *)
+ | Pand (sz: isize) (rd: ireg) (r1: ireg0) (r2: ireg) (s: shift_op) (**r and *)
+ | Pbic (sz: isize) (rd: ireg) (r1: ireg0) (r2: ireg) (s: shift_op) (**r and-not *)
+ | Peon (sz: isize) (rd: ireg) (r1: ireg0) (r2: ireg) (s: shift_op) (**r xor-not *)
+ | Peor (sz: isize) (rd: ireg) (r1: ireg0) (r2: ireg) (s: shift_op) (**r xor *)
+ | Porr (sz: isize) (rd: ireg) (r1: ireg0) (r2: ireg) (s: shift_op) (**r or *)
+ | Porn (sz: isize) (rd: ireg) (r1: ireg0) (r2: ireg) (s: shift_op) (**r or-not *)
+ | Ptst (sz: isize) (r1: ireg0) (r2: ireg) (s: shift_op) (**r and, then set flags *)
+ (** Variable shifts *)
+ | Pasrv (sz: isize) (rd: ireg) (r1 r2: ireg) (**r arithmetic right shift *)
+ | Plslv (sz: isize) (rd: ireg) (r1 r2: ireg) (**r left shift *)
+ | Plsrv (sz: isize) (rd: ireg) (r1 r2: ireg) (**r logical right shift *)
+ | Prorv (sz: isize) (rd: ireg) (r1 r2: ireg) (**r rotate right *)
+ (** Bit operations *)
+ | Pcls (sz: isize) (rd r1: ireg) (**r count leading sign bits *)
+ | Pclz (sz: isize) (rd r1: ireg) (**r count leading zero bits *)
+ | Prev (sz: isize) (rd r1: ireg) (**r reverse bytes *)
+ | Prev16 (sz: isize) (rd r1: ireg) (**r reverse bytes in each 16-bit word *)
+ (** Conditional data processing *)
+ | Pcsel (rd: ireg) (r1 r2: ireg) (c: testcond) (**r int conditional move *)
+ | Pcset (rd: ireg) (c: testcond) (**r set to 1/0 if cond is true/false *)
+(*
+ | Pcsetm (rd: ireg) (c: testcond) (**r set to -1/0 if cond is true/false *)
+*)
+ (** Integer multiply/divide *)
+ | Pmadd (sz: isize) (rd: ireg) (r1 r2: ireg) (r3: ireg0) (**r multiply-add *)
+ | Pmsub (sz: isize) (rd: ireg) (r1 r2: ireg) (r3: ireg0) (**r multiply-sub *)
+ | Psmulh (rd: ireg) (r1 r2: ireg) (**r signed multiply high *)
+ | Pumulh (rd: ireg) (r1 r2: ireg) (**r unsigned multiply high *)
+ | Psdiv (sz: isize) (rd: ireg) (r1 r2: ireg) (**r signed division *)
+ | Pudiv (sz: isize) (rd: ireg) (r1 r2: ireg) (**r unsigned division *)
+ (** Floating-point loads and stores *)
+ | Pldrs (rd: freg) (a: addressing) (**r load float32 (single precision) *)
+ | Pldrd (rd: freg) (a: addressing) (**r load float64 (double precision) *)
+ | Pldrd_a (rd: freg) (a: addressing) (**r load float64 as any64 *)
+ | Pstrs (rs: freg) (a: addressing) (**r store float32 *)
+ | Pstrd (rs: freg) (a: addressing) (**r store float64 *)
+ | Pstrd_a (rs: freg) (a: addressing) (**r store float64 as any64 *)
+ (** Floating-point move *)
+ | Pfmov (rd r1: freg)
+ | Pfmovimms (rd: freg) (f: float32) (**r load float32 constant *)
+ | Pfmovimmd (rd: freg) (f: float) (**r load float64 constant *)
+ | Pfmovi (fsz: fsize) (rd: freg) (r1: ireg0) (**r copy int reg to FP reg *)
+ (** Floating-point conversions *)
+ | Pfcvtds (rd r1: freg) (**r convert float32 to float64 *)
+ | Pfcvtsd (rd r1: freg) (**r convert float64 to float32 *)
+ | Pfcvtzs (isz: isize) (fsz: fsize) (rd: ireg) (r1: freg) (**r convert float to signed int *)
+ | Pfcvtzu (isz: isize) (fsz: fsize) (rd: ireg) (r1: freg) (**r convert float to unsigned int *)
+ | Pscvtf (fsz: fsize) (isz: isize) (rd: freg) (r1: ireg) (**r convert signed int to float *)
+ | Pucvtf (fsz: fsize) (isz: isize) (rd: freg) (r1: ireg) (**r convert unsigned int to float *)
+ (** Floating-point arithmetic *)
+ | Pfabs (sz: fsize) (rd r1: freg) (**r absolute value *)
+ | Pfneg (sz: fsize) (rd r1: freg) (**r negation *)
+ | Pfsqrt (sz: fsize) (rd r1: freg) (**r square root *)
+ | Pfadd (sz: fsize) (rd r1 r2: freg) (**r addition *)
+ | Pfdiv (sz: fsize) (rd r1 r2: freg) (**r division *)
+ | Pfmul (sz: fsize) (rd r1 r2: freg) (**r multiplication *)
+ | Pfnmul (sz: fsize) (rd r1 r2: freg) (**r multiply-negate *)
+ | Pfsub (sz: fsize) (rd r1 r2: freg) (**r subtraction *)
+ | Pfmadd (sz: fsize) (rd r1 r2 r3: freg) (**r [rd = r3 + r1 * r2] *)
+ | Pfmsub (sz: fsize) (rd r1 r2 r3: freg) (**r [rd = r3 - r1 * r2] *)
+ | Pfnmadd (sz: fsize) (rd r1 r2 r3: freg) (**r [rd = - r3 - r1 * r2] *)
+ | Pfnmsub (sz: fsize) (rd r1 r2 r3: freg) (**r [rd = - r3 + r1 * r2] *)
+ (** Floating-point comparison *)
+ | Pfcmp (sz: fsize) (r1 r2: freg) (**r compare [r1] and [r2] *)
+ | Pfcmp0 (sz: fsize) (r1: freg) (**r compare [r1] and [+0.0] *)
+ (** Floating-point conditional select *)
+ | Pfsel (rd r1 r2: freg) (cond: testcond)
+ (** Pseudo-instructions *)
+ | Pallocframe (sz: Z) (linkofs: ptrofs) (**r allocate new stack frame *)
+ | Pfreeframe (sz: Z) (linkofs: ptrofs) (**r deallocate stack frame and restore previous frame *)
+ | Plabel (lbl: label) (**r define a code label *)
+ | Ploadsymbol (rd: ireg) (id: ident) (**r load the address of [id] *)
+ | Pcvtsw2x (rd: ireg) (r1: ireg) (**r sign-extend 32-bit int to 64-bit *)
+ | Pcvtuw2x (rd: ireg) (r1: ireg) (**r zero-extend 32-bit int to 64-bit *)
+ | Pcvtx2w (rd: ireg) (**r retype a 64-bit int as a 32-bit int *)
+ | Pbtbl (r1: ireg) (tbl: list label) (**r N-way branch through a jump table *)
+ | Pbuiltin (ef: external_function)
+ (args: list (builtin_arg preg)) (res: builtin_res preg) (**r built-in function (pseudo) *)
+ | Pnop (**r no operation *)
+ | Pcfi_adjust (ofs: int) (**r .cfi_adjust debug directive *)
+ | Pcfi_rel_offset (ofs: int) (**r .cfi_rel_offset debug directive *)
+.
+*)
+
+
+(** * Operational semantics *)
+
+(** The semantics operates over a single mapping from registers
+ (type [preg]) to values. We maintain (but do not enforce)
+ the convention that integer registers are mapped to values of
+ type [Tint], float registers to values of type [Tfloat],
+ and condition bits to either [Vzero] or [Vone]. *)
+
+Definition regset := Pregmap.t val.
+Definition genv := Genv.t fundef unit.
+
+(** The value of an [ireg0] is either the value of the integer register,
+ or 0. *)
+
+Definition ir0 (is:isize) (rs: regset) (r: ireg0) : val :=
+ match r with RR0 r => rs (IR' r) | XZR => if is (* is W *) then Vint Int.zero else Vlong Int64.zero end.
+
+(** Concise notations for accessing and updating the values of registers. *)
+
+Notation "a # b" := (a b) (at level 1, only parsing) : asm.
+Notation "a # b <- c" := (Pregmap.set b c a) (at level 1, b at next level) : asm.
+Notation "a ## b" := (ir0 W a b) (at level 1, only parsing) : asm.
+Notation "a ### b" := (ir0 X a b) (at level 1, only parsing) : asm.
+
+Open Scope asm.
+
+(** Undefining some registers *)
+
+Fixpoint undef_regs (l: list preg) (rs: regset) : regset :=
+ match l with
+ | nil => rs
+ | r :: l' => undef_regs l' (rs#r <- Vundef)
+ end.
+
+(** Undefining the condition codes *)
+
+Definition undef_flags (rs: regset) : regset :=
+ fun r => match r with CR _ => Vundef | _ => rs r end.
+
+(** Assigning a register pair *)
+
+Definition set_pair (p: rpair preg) (v: val) (rs: regset) : regset :=
+ match p with
+ | One r => rs#r <- v
+ | Twolong rhi rlo => rs#rhi <- (Val.hiword v) #rlo <- (Val.loword v)
+ end.
+
+(** Assigning the result of a builtin *)
+
+Fixpoint set_res (res: builtin_res preg) (v: val) (rs: regset) : regset :=
+ match res with
+ | BR r => rs#r <- v
+ | BR_none => rs
+ | BR_splitlong hi lo => set_res lo (Val.loword v) (set_res hi (Val.hiword v) rs)
+ end.
+
+(** See "Parameters" of the same names in Asm.v *)
+Record aarch64_linker: Type := {
+ symbol_low: ident -> ptrofs -> val;
+ symbol_high: ident -> ptrofs -> val
+}.
+
+Section RELSEM.
+
+Variable lk: aarch64_linker.
+Variable ge: genv.
+
+
+(** The semantics is purely small-step and defined as a function
+ from the current state (a register set + a memory state)
+ to either [Next rs' m'] where [rs'] and [m'] are the updated register
+ set and memory state after execution of the instruction at [rs#PC],
+ or [Stuck] if the processor is stuck. *)
+
+Record state: Type := State { _rs: regset; _m: mem }.
+Definition outcome := option state.
+Definition Next rs m: outcome := Some (State rs m).
+Definition Stuck: outcome := None.
+Local Open Scope option_monad_scope.
+
+(* a few lemma on comparisons of unsigned (e.g. pointers)
+
+TODO: these lemma are a copy/paste of kvx/Asmvliw.v (sharing to improve!)
+*)
+
+Definition Val_cmpu_bool cmp v1 v2: option bool :=
+ Val.cmpu_bool (fun _ _ => true) cmp v1 v2.
+
+Lemma Val_cmpu_bool_correct (m:mem) (cmp: comparison) (v1 v2: val) b:
+ (Val.cmpu_bool (Mem.valid_pointer m) cmp v1 v2) = Some b
+ -> (Val_cmpu_bool cmp v1 v2) = Some b.
+Proof.
+ intros; eapply Val.cmpu_bool_lessdef; (econstructor 1 || eauto).
+Qed.
+
+Definition Val_cmpu cmp v1 v2 := Val.of_optbool (Val_cmpu_bool cmp v1 v2).
+
+Lemma Val_cmpu_correct (m:mem) (cmp: comparison) (v1 v2: val):
+ Val.lessdef (Val.cmpu (Mem.valid_pointer m) cmp v1 v2)
+ (Val_cmpu cmp v1 v2).
+Proof.
+ unfold Val.cmpu, Val_cmpu.
+ remember (Val.cmpu_bool (Mem.valid_pointer m) cmp v1 v2) as ob.
+ destruct ob; simpl.
+ - erewrite Val_cmpu_bool_correct; eauto.
+ econstructor.
+ - econstructor.
+Qed.
+
+Definition Val_cmplu_bool (cmp: comparison) (v1 v2: val)
+ := (Val.cmplu_bool (fun _ _ => true) cmp v1 v2).
+
+Lemma Val_cmplu_bool_correct (m:mem) (cmp: comparison) (v1 v2: val) b:
+ (Val.cmplu_bool (Mem.valid_pointer m) cmp v1 v2) = Some b
+ -> (Val_cmplu_bool cmp v1 v2) = Some b.
+Proof.
+ intros; eapply Val.cmplu_bool_lessdef; (econstructor 1 || eauto).
+Qed.
+
+Definition Val_cmplu cmp v1 v2 := Val.of_optbool (Val_cmplu_bool cmp v1 v2).
+
+Lemma Val_cmplu_correct (m:mem) (cmp: comparison) (v1 v2: val):
+ Val.lessdef (Val.maketotal (Val.cmplu (Mem.valid_pointer m) cmp v1 v2))
+ (Val_cmplu cmp v1 v2).
+Proof.
+ unfold Val.cmplu, Val_cmplu.
+ remember (Val.cmplu_bool (Mem.valid_pointer m) cmp v1 v2) as ob.
+ destruct ob as [b|]; simpl.
+ - erewrite Val_cmplu_bool_correct; eauto.
+ simpl. econstructor.
+ - econstructor.
+Qed.
+
+(** Testing a condition *)
+
+Definition eval_testcond (c: testcond) (rs: regset) : option bool :=
+ match c with
+ | TCeq => (**r equal *)
+ match rs#CZ with
+ | Vint n => Some (Int.eq n Int.one)
+ | _ => None
+ end
+ | TCne => (**r not equal *)
+ match rs#CZ with
+ | Vint n => Some (Int.eq n Int.zero)
+ | _ => None
+ end
+ | TClo => (**r unsigned less than *)
+ match rs#CC with
+ | Vint n => Some (Int.eq n Int.zero)
+ | _ => None
+ end
+ | TCls => (**r unsigned less or equal *)
+ match rs#CC, rs#CZ with
+ | Vint c, Vint z => Some (Int.eq c Int.zero || Int.eq z Int.one)
+ | _, _ => None
+ end
+ | TChs => (**r unsigned greater or equal *)
+ match rs#CC with
+ | Vint n => Some (Int.eq n Int.one)
+ | _ => None
+ end
+ | TChi => (**r unsigned greater *)
+ match rs#CC, rs#CZ with
+ | Vint c, Vint z => Some (Int.eq c Int.one && Int.eq z Int.zero)
+ | _, _ => None
+ end
+ | TClt => (**r signed less than *)
+ match rs#CV, rs#CN with
+ | Vint o, Vint s => Some (Int.eq (Int.xor o s) Int.one)
+ | _, _ => None
+ end
+ | TCle => (**r signed less or equal *)
+ match rs#CV, rs#CN, rs#CZ with
+ | Vint o, Vint s, Vint z => Some (Int.eq (Int.xor o s) Int.one || Int.eq z Int.one)
+ | _, _, _ => None
+ end
+ | TCge => (**r signed greater or equal *)
+ match rs#CV, rs#CN with
+ | Vint o, Vint s => Some (Int.eq (Int.xor o s) Int.zero)
+ | _, _ => None
+ end
+ | TCgt => (**r signed greater *)
+ match rs#CV, rs#CN, rs#CZ with
+ | Vint o, Vint s, Vint z => Some (Int.eq (Int.xor o s) Int.zero && Int.eq z Int.zero)
+ | _, _, _ => None
+ end
+ | TCpl => (**r positive *)
+ match rs#CN with
+ | Vint n => Some (Int.eq n Int.zero)
+ | _ => None
+ end
+ | TCmi => (**r negative *)
+ match rs#CN with
+ | Vint n => Some (Int.eq n Int.one)
+ | _ => None
+ end
+ end.
+
+(** Integer "is zero?" test *)
+
+Definition eval_testzero (sz: isize) (v: val): option bool :=
+ match sz with
+ | W => Val_cmpu_bool Ceq v (Vint Int.zero)
+ | X => Val_cmplu_bool Ceq v (Vlong Int64.zero)
+ end.
+
+(** Integer "bit is set?" test *)
+
+Definition eval_testbit (sz: isize) (v: val) (n: int): option bool :=
+ match sz with
+ | W => Val.cmp_bool Cne (Val.and v (Vint (Int.shl Int.one n))) (Vint Int.zero)
+ | X => Val.cmpl_bool Cne (Val.andl v (Vlong (Int64.shl' Int64.one n))) (Vlong Int64.zero)
+ end.
+
+
+(** Comparisons *)
+
+Definition compare_int (rs: regset) (v1 v2: val) :=
+ rs#CN <- (Val.negative (Val.sub v1 v2))
+ #CZ <- (Val_cmpu Ceq v1 v2)
+ #CC <- (Val_cmpu Cge v1 v2)
+ #CV <- (Val.sub_overflow v1 v2).
+
+Definition compare_long (rs: regset) (v1 v2: val) :=
+ rs#CN <- (Val.negativel (Val.subl v1 v2))
+ #CZ <- (Val_cmplu Ceq v1 v2)
+ #CC <- (Val_cmplu Cge v1 v2)
+ #CV <- (Val.subl_overflow v1 v2).
+
+(** Semantics of [fcmp] instructions:
+<<
+== N=0 Z=1 C=1 V=0
+< N=1 Z=0 C=0 V=0
+> N=0 Z=0 C=1 V=0
+unord N=0 Z=0 C=1 V=1
+>>
+*)
+
+Definition compare_float (rs: regset) (v1 v2: val) :=
+ match v1, v2 with
+ | Vfloat f1, Vfloat f2 =>
+ rs#CN <- (Val.of_bool (Float.cmp Clt f1 f2))
+ #CZ <- (Val.of_bool (Float.cmp Ceq f1 f2))
+ #CC <- (Val.of_bool (negb (Float.cmp Clt f1 f2)))
+ #CV <- (Val.of_bool (negb (Float.ordered f1 f2)))
+ | _, _ =>
+ rs#CN <- Vundef
+ #CZ <- Vundef
+ #CC <- Vundef
+ #CV <- Vundef
+ end.
+
+Definition compare_single (rs: regset) (v1 v2: val) :=
+ match v1, v2 with
+ | Vsingle f1, Vsingle f2 =>
+ rs#CN <- (Val.of_bool (Float32.cmp Clt f1 f2))
+ #CZ <- (Val.of_bool (Float32.cmp Ceq f1 f2))
+ #CC <- (Val.of_bool (negb (Float32.cmp Clt f1 f2)))
+ #CV <- (Val.of_bool (negb (Float32.ordered f1 f2)))
+ | _, _ =>
+ rs#CN <- Vundef
+ #CZ <- Vundef
+ #CC <- Vundef
+ #CV <- Vundef
+ end.
+
+
+(** Evaluating an addressing mode *)
+
+Definition eval_addressing (a: addressing) (rs: regset): val :=
+ match a with
+ | ADimm base n => Val.addl rs#base (Vlong n)
+ | ADreg base r => Val.addl rs#base rs#r
+ | ADlsl base r n => Val.addl rs#base (Val.shll rs#r (Vint n))
+ | ADsxt base r n => Val.addl rs#base (Val.shll (Val.longofint rs#r) (Vint n))
+ | ADuxt base r n => Val.addl rs#base (Val.shll (Val.longofintu rs#r) (Vint n))
+ | ADadr base id ofs => Val.addl rs#base (symbol_low lk id ofs)
+ | ADpostincr base n => Vundef (* not modeled yet *)
+ end.
+
+(** Auxiliaries for memory accesses *)
+
+Definition exec_load (chunk: memory_chunk) (transf: val -> val)
+ (a: addressing) (r: dreg) (rs: regset) (m: mem) :=
+ SOME v <- Mem.loadv chunk m (eval_addressing a rs) IN
+ Next (rs#r <- (transf v)) m.
+
+Definition exec_store (chunk: memory_chunk)
+ (a: addressing) (v: val)
+ (rs: regset) (m: mem) :=
+ SOME m' <- Mem.storev chunk m (eval_addressing a rs) v IN
+ Next rs m'.
+
+(** Insertion of bits into an integer *)
+
+Definition insert_in_int (x: val) (y: Z) (pos: Z) (len: Z) : val :=
+ match x with
+ | Vint n => Vint (Int.repr (Zinsert (Int.unsigned n) y pos len))
+ | _ => Vundef
+ end.
+
+Definition insert_in_long (x: val) (y: Z) (pos: Z) (len: Z) : val :=
+ match x with
+ | Vlong n => Vlong (Int64.repr (Zinsert (Int64.unsigned n) y pos len))
+ | _ => Vundef
+ end.
+
+(** Evaluation of shifted operands *)
+
+Definition eval_shift_op_int (v: val) (s: shift_op): val :=
+ match s with
+ | SOnone => v
+ | SOlsl n => Val.shl v (Vint n)
+ | SOlsr n => Val.shru v (Vint n)
+ | SOasr n => Val.shr v (Vint n)
+ | SOror n => Val.ror v (Vint n)
+ end.
+
+Definition eval_shift_op_long (v: val) (s: shift_op): val :=
+ match s with
+ | SOnone => v
+ | SOlsl n => Val.shll v (Vint n)
+ | SOlsr n => Val.shrlu v (Vint n)
+ | SOasr n => Val.shrl v (Vint n)
+ | SOror n => Val.rorl v (Vint n)
+ end.
+
+(** Evaluation of sign- or zero- extended operands *)
+
+Definition eval_extend (v: val) (x: extend_op): val :=
+ match x with
+ | EOsxtb n => Val.shll (Val.longofint (Val.sign_ext 8 v)) (Vint n)
+ | EOsxth n => Val.shll (Val.longofint (Val.sign_ext 16 v)) (Vint n)
+ | EOsxtw n => Val.shll (Val.longofint v) (Vint n)
+ | EOuxtb n => Val.shll (Val.longofintu (Val.zero_ext 8 v)) (Vint n)
+ | EOuxth n => Val.shll (Val.longofintu (Val.zero_ext 16 v)) (Vint n)
+ | EOuxtw n => Val.shll (Val.longofintu v) (Vint n)
+ | EOuxtx n => Val.shll v (Vint n)
+ end.
+
+(** Bit-level conversion from integers to FP numbers *)
+
+Definition float32_of_bits (v: val): val :=
+ match v with
+ | Vint n => Vsingle (Float32.of_bits n)
+ | _ => Vundef
+ end.
+
+Definition float64_of_bits (v: val): val :=
+ match v with
+ | Vlong n => Vfloat (Float.of_bits n)
+ | _ => Vundef
+ end.
+
+(** execution of loads
+*)
+
+Definition chunk_load_rd_a (ld: load_rd_a): memory_chunk :=
+ match ld with
+ | Pldrw => Mint32
+ | Pldrw_a => Many32
+ | Pldrx => Mint64
+ | Pldrx_a => Many64
+ | Pldrb _ => Mint8unsigned
+ | Pldrsb _ => Mint8signed
+ | Pldrh _ => Mint16unsigned
+ | Pldrsh _ => Mint16signed
+ | Pldrzw => Mint32
+ | Pldrsw => Mint32
+ | Pldrs => Mfloat32
+ | Pldrd => Mfloat64
+ | Pldrd_a => Many64
+ end.
+
+Definition chunk_store_rs_a (st: store_rs_a) : memory_chunk :=
+ match st with
+ | Pstrw => Mint32
+ | Pstrw_a => Many32
+ | Pstrx => Mint64
+ | Pstrx_a => Many64
+ | Pstrb => Mint8unsigned
+ | Pstrh => Mint16unsigned
+ | Pstrs => Mfloat32
+ | Pstrd => Mfloat64
+ | Pstrd_a => Many64
+ end.
+
+Definition interp_load_rd_a (ld: load_rd_a): val -> val :=
+ match ld with
+ | Pldrb X => Val.longofintu
+ | Pldrsb X => Val.longofint
+ | Pldrh X => Val.longofintu
+ | Pldrsh X => Val.longofint
+ | Pldrzw => Val.longofintu
+ | Pldrsw => Val.longofint
+ (* Changed to exhaustive list because I tend to forgot all the places I need
+ * to check when changing things. *)
+ | Pldrb W | Pldrsb W | Pldrh W | Pldrsh W
+ | Pldrw | Pldrw_a | Pldrx | Pldrx_a
+ | Pldrs | Pldrd | Pldrd_a => fun v => v
+ end.
+
+(** TODO: redundant w.r.t Machblock ?? *)
+Lemma in_dec (lbl: label) (l: list label): { List.In lbl l } + { ~(List.In lbl l) }.
+Proof.
+ apply List.in_dec.
+ apply Pos.eq_dec.
+Qed.
+
+(** Note: copy-paste from Machblock *)
+Definition is_label (lbl: label) (bb: bblock) : bool :=
+ if in_dec lbl (header bb) then true else false.
+
+Lemma is_label_correct_true lbl bb:
+ List.In lbl (header bb) <-> is_label lbl bb = true.
+Proof.
+ unfold is_label; destruct (in_dec lbl (header bb)); simpl; intuition.
+Qed.
+
+Lemma is_label_correct_false lbl bb:
+ ~(List.In lbl (header bb)) <-> is_label lbl bb = false.
+Proof.
+ unfold is_label; destruct (in_dec lbl (header bb)); simpl; intuition.
+Qed.
+
+(** convert a label into a position in the code *)
+Fixpoint label_pos (lbl: label) (pos: Z) (lb: bblocks) {struct lb} : option Z :=
+ match lb with
+ | nil => None
+ | b :: lb' => if is_label lbl b then Some pos else label_pos lbl (pos + (size b)) lb'
+ end.
+
+Definition goto_label (f: function) (lbl: label) (rs: regset) (m: mem) :=
+ SOME pos <- label_pos lbl 0 (fn_blocks f) IN
+ match rs#PC with
+ | Vptr b ofs => Next (rs#PC <- (Vptr b (Ptrofs.repr pos))) m
+ | _ => Stuck
+ end.
+
+(** Evaluating a branch
+
+Warning: PC is assumed to be already pointing on the next instruction !
+
+*)
+
+Definition eval_branch (f: function) (lbl: label) (rs: regset) (m: mem) (ores: option bool) :=
+ SOME res <- ores IN
+ if res then goto_label f lbl rs m else Next rs m.
+
+Definition eval_neg_branch (f: function) (lbl: label) (rs: regset) (m: mem) (ores: option bool) :=
+ SOME res <- ores IN
+ if res then Next rs m else goto_label f lbl rs m.
+
+Definition exec_cfi (f: function) (cfi: cf_instruction) (rs: regset) (m: mem) : outcome :=
+ match cfi with
+ (** Branches *)
+ | Pb lbl =>
+ goto_label f lbl rs m
+ | Pbc cond lbl =>
+ eval_branch f lbl rs m (eval_testcond cond rs)
+ | Pbl id sg =>
+ Next (rs#RA <- (rs#PC) #PC <- (Genv.symbol_address ge id Ptrofs.zero)) m
+ | Pbs id sg =>
+ Next (rs#PC <- (Genv.symbol_address ge id Ptrofs.zero)) m
+ | Pblr r sg =>
+ Next (rs#RA <- (rs#PC) #PC <- (rs#r)) m
+ | Pbr r sg =>
+ Next (rs#PC <- (rs#r)) m
+ | Pret r =>
+ Next (rs#PC <- (rs#r)) m
+ | Pcbnz sz r lbl =>
+ eval_neg_branch f lbl rs m (eval_testzero sz rs#r)
+ | Pcbz sz r lbl =>
+ eval_branch f lbl rs m (eval_testzero sz rs#r)
+ | Ptbnz sz r n lbl =>
+ eval_branch f lbl rs m (eval_testbit sz rs#r n)
+ | Ptbz sz r n lbl =>
+ eval_neg_branch f lbl rs m (eval_testbit sz rs#r n)
+ (** Pseudo-instructions *)
+ | Pbtbl r tbl =>
+ match (rs#X16 <- Vundef)#r with
+ | Vint n =>
+ SOME lbl <- list_nth_z tbl (Int.unsigned n) IN
+ goto_label f lbl (rs#X16 <- Vundef #X17 <- Vundef) m
+ | _ => Stuck
+ end
+ end.
+
+Definition arith_eval_p (i : arith_p) : val :=
+ match i with
+ | Padrp id ofs => symbol_high lk id ofs
+ (** Move wide immediate *)
+ | Pmovz W n pos => Vint (Int.repr (Z.shiftl n pos))
+ | Pmovz X n pos => Vlong (Int64.repr (Z.shiftl n pos))
+ | Pmovn W n pos => Vint (Int.repr (Z.lnot (Z.shiftl n pos)))
+ | Pmovn X n pos => Vlong (Int64.repr (Z.lnot (Z.shiftl n pos)))
+ (** Floating-point move *)
+ | Pfmovimms f => Vsingle f
+ | Pfmovimmd f => Vfloat f
+ end.
+
+Definition if_opt_bool_val (c: option bool) v1 v2: val :=
+ match c with
+ | Some true => v1
+ | Some false => v2
+ | None => Vundef
+ end.
+
+Definition arith_eval_pp i v :=
+ match i with
+ | Pmov => v
+ | Pmovk W n pos => insert_in_int v n pos 16
+ | Pmovk X n pos => insert_in_long v n pos 16
+ | Paddadr id ofs => Val.addl v (symbol_low lk id ofs)
+ | Psbfiz W r s => Val.shl (Val.sign_ext s v) (Vint r)
+ | Psbfiz X r s => Val.shll (Val.sign_ext_l s v) (Vint r)
+ | Psbfx W r s => Val.sign_ext s (Val.shr v (Vint r))
+ | Psbfx X r s => Val.sign_ext_l s (Val.shrl v (Vint r))
+ | Pubfiz W r s => Val.shl (Val.zero_ext s v) (Vint r)
+ | Pubfiz X r s => Val.shll (Val.zero_ext_l s v) (Vint r)
+ | Pubfx W r s => Val.zero_ext s (Val.shru v (Vint r))
+ | Pubfx X r s => Val.zero_ext_l s (Val.shrlu v (Vint r))
+ | Pfmov => v
+ | Pfcvtds => Val.floatofsingle v
+ | Pfcvtsd => Val.singleoffloat v
+ | Pfabs S => Val.absfs v
+ | Pfabs D => Val.absf v
+ | Pfneg S => Val.negfs v
+ | Pfneg D => Val.negf v
+ | Pfcvtzs W S => Val.maketotal (Val.intofsingle v)
+ | Pfcvtzs W D => Val.maketotal (Val.intoffloat v)
+ | Pfcvtzs X S => Val.maketotal (Val.longofsingle v)
+ | Pfcvtzs X D => Val.maketotal (Val.longoffloat v)
+ | Pfcvtzu W S => Val.maketotal (Val.intuofsingle v)
+ | Pfcvtzu W D => Val.maketotal (Val.intuoffloat v)
+ | Pfcvtzu X S => Val.maketotal (Val.longuofsingle v)
+ | Pfcvtzu X D => Val.maketotal (Val.longuoffloat v)
+ | Paddimm W n => Val.add v (Vint (Int.repr n))
+ | Paddimm X n => Val.addl v (Vlong (Int64.repr n))
+ | Psubimm W n => Val.sub v (Vint (Int.repr n))
+ | Psubimm X n => Val.subl v (Vlong (Int64.repr n))
+ | Pscvtf S W => Val.maketotal (Val.singleofint v)
+ | Pscvtf D W => Val.maketotal (Val.floatofint v)
+ | Pscvtf S X => Val.maketotal (Val.singleoflong v)
+ | Pscvtf D X => Val.maketotal (Val.floatoflong v)
+ | Pucvtf S W => Val.maketotal (Val.singleofintu v)
+ | Pucvtf D W => Val.maketotal (Val.floatofintu v)
+ | Pucvtf S X => Val.maketotal (Val.singleoflongu v)
+ | Pucvtf D X => Val.maketotal (Val.floatoflongu v)
+ end.
+
+Definition arith_eval_ppp i v1 v2 :=
+ match i with
+ | Pasrv W => Val.shr v1 v2
+ | Pasrv X => Val.shrl v1 v2
+ | Plslv W => Val.shl v1 v2
+ | Plslv X => Val.shll v1 v2
+ | Plsrv W => Val.shru v1 v2
+ | Plsrv X => Val.shrlu v1 v2
+ | Prorv W => Val.ror v1 v2
+ | Prorv X => Val.rorl v1 v2
+ | Psmulh => Val.mullhs v1 v2
+ | Pumulh => Val.mullhu v1 v2
+ | Psdiv W => Val.maketotal (Val.divs v1 v2)
+ | Psdiv X => Val.maketotal (Val.divls v1 v2)
+ | Pudiv W => Val.maketotal (Val.divu v1 v2)
+ | Pudiv X => Val.maketotal (Val.divlu v1 v2)
+ | Paddext x => Val.addl v1 (eval_extend v2 x)
+ | Psubext x => Val.subl v1 (eval_extend v2 x)
+ | Pfadd S => Val.addfs v1 v2
+ | Pfadd D => Val.addf v1 v2
+ | Pfdiv S => Val.divfs v1 v2
+ | Pfdiv D => Val.divf v1 v2
+ | Pfmul S => Val.mulfs v1 v2
+ | Pfmul D => Val.mulf v1 v2
+ | Pfsub S => Val.subfs v1 v2
+ | Pfsub D => Val.subf v1 v2
+ end.
+
+Definition arith_rr0r_isize (i: arith_rr0r) :=
+ match i with
+ | Padd is _ => is
+ | Psub is _ => is
+ | Pand is _ => is
+ | Pbic is _ => is
+ | Peon is _ => is
+ | Peor is _ => is
+ | Porr is _ => is
+ | Porn is _ => is
+ end.
+
+(* obtain v1 by [ir0 (arith_rr0r_isize i) rs s1] *)
+Definition arith_eval_rr0r i v1 v2 :=
+ match i with
+ | Padd W s => Val.add v1 (eval_shift_op_int v2 s)
+ | Padd X s => Val.addl v1 (eval_shift_op_long v2 s)
+ | Psub W s => Val.sub v1 (eval_shift_op_int v2 s)
+ | Psub X s => Val.subl v1 (eval_shift_op_long v2 s)
+ | Pand W s => Val.and v1 (eval_shift_op_int v2 s)
+ | Pand X s => Val.andl v1 (eval_shift_op_long v2 s)
+ | Pbic W s => Val.and v1 (Val.notint (eval_shift_op_int v2 s))
+ | Pbic X s => Val.andl v1 (Val.notl (eval_shift_op_long v2 s))
+ | Peon W s => Val.xor v1 (Val.notint (eval_shift_op_int v2 s))
+ | Peon X s => Val.xorl v1 (Val.notl (eval_shift_op_long v2 s))
+ | Peor W s => Val.xor v1 (eval_shift_op_int v2 s)
+ | Peor X s => Val.xorl v1 (eval_shift_op_long v2 s)
+ | Porr W s => Val.or v1 (eval_shift_op_int v2 s)
+ | Porr X s => Val.orl v1 (eval_shift_op_long v2 s)
+ | Porn W s => Val.or v1 (Val.notint (eval_shift_op_int v2 s))
+ | Porn X s => Val.orl v1 (Val.notl (eval_shift_op_long v2 s))
+ end.
+
+Definition arith_rr0_isize (i : arith_rr0) :=
+ match i with
+ | Pandimm is _ | Peorimm is _ | Porrimm is _ => is
+ end.
+
+(* obtain v by [ir0 (arith_rr0_isize i) rs s] *)
+Definition arith_eval_rr0 i v :=
+ match i with
+ | Pandimm W n => Val.and v (Vint (Int.repr n))
+ | Pandimm X n => Val.andl v (Vlong (Int64.repr n))
+ | Peorimm W n => Val.xor v (Vint (Int.repr n))
+ | Peorimm X n => Val.xorl v (Vlong (Int64.repr n))
+ | Porrimm W n => Val.or v (Vint (Int.repr n))
+ | Porrimm X n => Val.orl v (Vlong (Int64.repr n))
+ end.
+
+Definition arith_arrrr0_isize (i : arith_arrrr0) :=
+ match i with
+ | Pmadd is | Pmsub is => is
+ end.
+
+(* obtain v3 by [ir0 (arith_arrrr0_isize i) rs s3] *)
+Definition arith_eval_arrrr0 i v1 v2 v3 :=
+ match i with
+ | Pmadd W => Val.add v3 (Val.mul v1 v2)
+ | Pmadd X => Val.addl v3 (Val.mull v1 v2)
+ | Pmsub W => Val.sub v3 (Val.mul v1 v2)
+ | Pmsub X => Val.subl v3 (Val.mull v1 v2)
+ end.
+
+Definition arith_prepare_comparison_pp i (v1 v2 : val) :=
+ match i with
+ | Pcmpext x => (v1, (eval_extend v2 x))
+ | Pcmnext x => (v1, (Val.negl (eval_extend v2 x)))
+ | Pfcmp _ => (v1, v2)
+ end.
+
+Definition arith_comparison_r0r_isize i :=
+ match i with
+ | Pcmp is _ => is
+ | Pcmn is _ => is
+ | Ptst is _ => is
+ end.
+
+Definition arith_prepare_comparison_r0r i v1 v2 :=
+ match i with
+ | Pcmp W s => (v1, (eval_shift_op_int v2 s))
+ | Pcmp X s => (v1, (eval_shift_op_long v2 s))
+ | Pcmn W s => (v1, (Val.neg (eval_shift_op_int v2 s)))
+ | Pcmn X s => (v1, (Val.negl (eval_shift_op_long v2 s)))
+ | Ptst W s => ((Val.and v1 (eval_shift_op_int v2 s)), (Vint Int.zero))
+ | Ptst X s => ((Val.andl v1 (eval_shift_op_long v2 s)), (Vlong Int64.zero))
+ end.
+
+Definition arith_prepare_comparison_p i v :=
+ match i with
+ | Pcmpimm W n => (v, (Vint (Int.repr n)))
+ | Pcmpimm X n => (v, (Vlong (Int64.repr n)))
+ | Pcmnimm W n => (v, (Vint (Int.neg (Int.repr n))))
+ | Pcmnimm X n => (v, (Vlong (Int64.neg (Int64.repr n))))
+ | Ptstimm W n => ((Val.and v (Vint (Int.repr n))), (Vint Int.zero))
+ | Ptstimm X n => ((Val.andl v (Vlong (Int64.repr n))), (Vlong Int64.zero))
+ | Pfcmp0 S => (v, (Vsingle Float32.zero))
+ | Pfcmp0 D => (v, (Vfloat Float.zero))
+ end.
+
+Definition arith_comparison_pp_compare i :=
+ match i with
+ | Pcmpext _ | Pcmnext _ => compare_long
+ | Pfcmp S => compare_single
+ | Pfcmp D => compare_float
+ end.
+
+Definition arith_comparison_p_compare i :=
+ match i with
+ | Pcmpimm W _ | Pcmnimm W _ | Ptstimm W _ => compare_int
+ | Pcmpimm X _ | Pcmnimm X _ | Ptstimm X _ => compare_long
+ | Pfcmp0 S => compare_single
+ | Pfcmp0 D => compare_float
+ end.
+
+Definition exec_arith_instr (ai: ar_instruction) (rs: regset): regset :=
+ match ai with
+ | PArithP i d => rs#d <- (arith_eval_p i)
+ | PArithPP i d s => rs#d <- (arith_eval_pp i rs#s)
+ | PArithPPP i d s1 s2 => rs#d <- (arith_eval_ppp i rs#s1 rs#s2)
+
+ | PArithRR0R i d s1 s2 => rs#d <- (arith_eval_rr0r i (ir0 (arith_rr0r_isize i) rs s1) rs#s2)
+
+ | PArithRR0 i d s => rs#d <- (arith_eval_rr0 i (ir0 (arith_rr0_isize i) rs s))
+
+ | PArithARRRR0 i d s1 s2 s3 =>
+ rs#d <- (arith_eval_arrrr0 i rs#s1 rs#s2 (ir0 (arith_arrrr0_isize i) rs s3))
+
+ | PArithComparisonPP i s1 s2 =>
+ let (v1, v2) := arith_prepare_comparison_pp i rs#s1 rs#s2 in
+ arith_comparison_pp_compare i rs v1 v2
+ | PArithComparisonR0R i s1 s2 =>
+ let is := arith_comparison_r0r_isize i in
+ let (v1, v2) := arith_prepare_comparison_r0r i (ir0 is rs s1) rs#s2 in
+ (if is (* is W *) then compare_int else compare_long) rs v1 v2
+ | PArithComparisonP i s =>
+ let (v1, v2) := arith_prepare_comparison_p i rs#s in
+ arith_comparison_p_compare i rs v1 v2
+ | Pcset d c => rs#d <- (if_opt_bool_val (eval_testcond c rs) (Vint Int.one) (Vint Int.zero))
+ | Pfmovi S d s => rs#d <- (float32_of_bits rs##s)
+ | Pfmovi D d s => rs#d <- (float64_of_bits rs###s)
+ | Pcsel d s1 s2 c => rs#d <- (if_opt_bool_val (eval_testcond c rs) (rs#s1) (rs#s2))
+ | Pfnmul S d s1 s2 => rs#d <- (Val.negfs (Val.mulfs rs#s1 rs#s2))
+ | Pfnmul D d s1 s2 => rs#d <- (Val.negf (Val.mulf rs#s1 rs#s2))
+ end.
+
+(* basic exec *)
+Definition exec_basic (b: basic) (rs: regset) (m: mem): outcome :=
+ match b with
+ | PArith ai => Next (exec_arith_instr ai rs) m
+ | PLoad ld rd a => exec_load (chunk_load_rd_a ld) (interp_load_rd_a ld) a rd rs m
+ | PStore st r a => exec_store (chunk_store_rs_a st) a rs#r rs m
+ | Pallocframe sz pos =>
+ let (m1, stk) := Mem.alloc m 0 sz in
+ let sp := (Vptr stk Ptrofs.zero) in
+ SOME m2 <- Mem.storev Mint64 m1 (Val.offset_ptr sp pos) rs#SP IN
+ Next (rs #X29 <- (rs#SP) #SP <- sp #X16 <- Vundef) m2
+ | Pfreeframe sz pos =>
+ SOME v <- Mem.loadv Mint64 m (Val.offset_ptr rs#SP pos) IN
+ match rs#SP with
+ | Vptr stk ofs =>
+ SOME m' <- Mem.free m stk 0 sz IN
+ Next (rs#SP <- v #X16 <- Vundef) m'
+ | _ => Stuck
+ end
+ | Ploadsymbol rd id =>
+ Next (rs#rd <- (Genv.symbol_address ge id Ptrofs.zero)) m
+ | Pcvtsw2x rd r1 =>
+ Next (rs#rd <- (Val.longofint rs#r1)) m
+ | Pcvtuw2x rd r1 =>
+ Next (rs#rd <- (Val.longofintu rs#r1)) m
+ | Pcvtx2w rd =>
+ Next (rs#rd <- (Val.loword rs#rd)) m
+ end.
+
+(* TODO: ORIGINAL EXECUTION OF INSTRUCTIONS THAT NEEDS TO BE ADAPTED !
+
+(** Execution of a single instruction [i] in initial state
+ [rs] and [m]. Return updated state. For instructions
+ that correspond to actual AArch64 instructions, the cases are
+ straightforward transliterations of the informal descriptions
+ given in the ARMv8 reference manuals. For pseudo-instructions,
+ refer to the informal descriptions given above.
+
+ Note that we set to [Vundef] the registers used as temporaries by
+ the expansions of the pseudo-instructions, so that the code we
+ generate cannot use those registers to hold values that must
+ survive the execution of the pseudo-instruction.
+*)
+
+Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : outcome :=
+ match i with
+ (** Branches *)
+ | Pb lbl =>
+ goto_label f lbl rs m
+ | Pbc cond lbl =>
+ match eval_testcond cond rs with
+ | Some true => goto_label f lbl rs m
+ | Some false => Next (nextinstr rs) m
+ | None => Stuck
+ end
+ | Pbl id sg =>
+ Next (rs#RA <- (Val.offset_ptr rs#PC Ptrofs.one) #PC <- (Genv.symbol_address ge id Ptrofs.zero)) m
+ | Pbs id sg =>
+ Next (rs#PC <- (Genv.symbol_address ge id Ptrofs.zero)) m
+ | Pblr r sg =>
+ Next (rs#RA <- (Val.offset_ptr rs#PC Ptrofs.one) #PC <- (rs#r)) m
+ | Pbr r sg =>
+ Next (rs#PC <- (rs#r)) m
+ | Pret r =>
+ Next (rs#PC <- (rs#r)) m
+ | Pcbnz sz r lbl =>
+ match eval_testzero sz rs#r m with
+ | Some true => Next (nextinstr rs) m
+ | Some false => goto_label f lbl rs m
+ | None => Stuck
+ end
+ | Pcbz sz r lbl =>
+ match eval_testzero sz rs#r m with
+ | Some true => goto_label f lbl rs m
+ | Some false => Next (nextinstr rs) m
+ | None => Stuck
+ end
+ | Ptbnz sz r n lbl =>
+ match eval_testbit sz rs#r n with
+ | Some true => goto_label f lbl rs m
+ | Some false => Next (nextinstr rs) m
+ | None => Stuck
+ end
+ | Ptbz sz r n lbl =>
+ match eval_testbit sz rs#r n with
+ | Some true => Next (nextinstr rs) m
+ | Some false => goto_label f lbl rs m
+ | None => Stuck
+ end
+ (** Memory loads and stores *)
+ | Pldrw rd a =>
+ exec_load Mint32 (fun v => v) a rd rs m
+ | Pldrw_a rd a =>
+ exec_load Many32 (fun v => v) a rd rs m
+ | Pldrx rd a =>
+ exec_load Mint64 (fun v => v) a rd rs m
+ | Pldrx_a rd a =>
+ exec_load Many64 (fun v => v) a rd rs m
+ | Pldrb W rd a =>
+ exec_load Mint8unsigned (fun v => v) a rd rs m
+ | Pldrb X rd a =>
+ exec_load Mint8unsigned Val.longofintu a rd rs m
+ | Pldrsb W rd a =>
+ exec_load Mint8signed (fun v => v) a rd rs m
+ | Pldrsb X rd a =>
+ exec_load Mint8signed Val.longofint a rd rs m
+ | Pldrh W rd a =>
+ exec_load Mint16unsigned (fun v => v) a rd rs m
+ | Pldrh X rd a =>
+ exec_load Mint16unsigned Val.longofintu a rd rs m
+ | Pldrsh W rd a =>
+ exec_load Mint16signed (fun v => v) a rd rs m
+ | Pldrsh X rd a =>
+ exec_load Mint16signed Val.longofint a rd rs m
+ | Pldrzw rd a =>
+ exec_load Mint32 Val.longofintu a rd rs m
+ | Pldrsw rd a =>
+ exec_load Mint32 Val.longofint a rd rs m
+ | Pstrw r a =>
+ exec_store Mint32 a rs#r rs m
+ | Pstrw_a r a =>
+ exec_store Many32 a rs#r rs m
+ | Pstrx r a =>
+ exec_store Mint64 a rs#r rs m
+ | Pstrx_a r a =>
+ exec_store Many64 a rs#r rs m
+ | Pstrb r a =>
+ exec_store Mint8unsigned a rs#r rs m
+ | Pstrh r a =>
+ exec_store Mint16unsigned a rs#r rs m
+ (** Integer arithmetic, immediate *)
+ | Paddimm W rd r1 n =>
+ Next (nextinstr (rs#rd <- (Val.add rs#r1 (Vint (Int.repr n))))) m
+ | Paddimm X rd r1 n =>
+ Next (nextinstr (rs#rd <- (Val.addl rs#r1 (Vlong (Int64.repr n))))) m
+ | Psubimm W rd r1 n =>
+ Next (nextinstr (rs#rd <- (Val.sub rs#r1 (Vint (Int.repr n))))) m
+ | Psubimm X rd r1 n =>
+ Next (nextinstr (rs#rd <- (Val.subl rs#r1 (Vlong (Int64.repr n))))) m
+ | Pcmpimm W r1 n =>
+ Next (nextinstr (compare_int rs rs#r1 (Vint (Int.repr n)) m)) m
+ | Pcmpimm X r1 n =>
+ Next (nextinstr (compare_long rs rs#r1 (Vlong (Int64.repr n)) m)) m
+ | Pcmnimm W r1 n =>
+ Next (nextinstr (compare_int rs rs#r1 (Vint (Int.neg (Int.repr n))) m)) m
+ | Pcmnimm X r1 n =>
+ Next (nextinstr (compare_long rs rs#r1 (Vlong (Int64.neg (Int64.repr n))) m)) m
+ (** Move integer register *)
+ | Pmov rd r1 =>
+ Next (nextinstr (rs#rd <- (rs#r1))) m
+ (** Logical, immediate *)
+ | Pandimm W rd r1 n =>
+ Next (nextinstr (rs#rd <- (Val.and rs##r1 (Vint (Int.repr n))))) m
+ | Pandimm X rd r1 n =>
+ Next (nextinstr (rs#rd <- (Val.andl rs###r1 (Vlong (Int64.repr n))))) m
+ | Peorimm W rd r1 n =>
+ Next (nextinstr (rs#rd <- (Val.xor rs##r1 (Vint (Int.repr n))))) m
+ | Peorimm X rd r1 n =>
+ Next (nextinstr (rs#rd <- (Val.xorl rs###r1 (Vlong (Int64.repr n))))) m
+ | Porrimm W rd r1 n =>
+ Next (nextinstr (rs#rd <- (Val.or rs##r1 (Vint (Int.repr n))))) m
+ | Porrimm X rd r1 n =>
+ Next (nextinstr (rs#rd <- (Val.orl rs###r1 (Vlong (Int64.repr n))))) m
+ | Ptstimm W r1 n =>
+ Next (nextinstr (compare_int rs (Val.and rs#r1 (Vint (Int.repr n))) (Vint Int.zero) m)) m
+ | Ptstimm X r1 n =>
+ Next (nextinstr (compare_long rs (Val.andl rs#r1 (Vlong (Int64.repr n))) (Vlong Int64.zero) m)) m
+ (** Move wide immediate *)
+ | Pmovz W rd n pos =>
+ Next (nextinstr (rs#rd <- (Vint (Int.repr (Z.shiftl n pos))))) m
+ | Pmovz X rd n pos =>
+ Next (nextinstr (rs#rd <- (Vlong (Int64.repr (Z.shiftl n pos))))) m
+ | Pmovn W rd n pos =>
+ Next (nextinstr (rs#rd <- (Vint (Int.repr (Z.lnot (Z.shiftl n pos)))))) m
+ | Pmovn X rd n pos =>
+ Next (nextinstr (rs#rd <- (Vlong (Int64.repr (Z.lnot (Z.shiftl n pos)))))) m
+ | Pmovk W rd n pos =>
+ Next (nextinstr (rs#rd <- (insert_in_int rs#rd n pos 16))) m
+ | Pmovk X rd n pos =>
+ Next (nextinstr (rs#rd <- (insert_in_long rs#rd n pos 16))) m
+ (** PC-relative addressing *)
+ | Padrp rd id ofs =>
+ Next (nextinstr (rs#rd <- (symbol_high ge id ofs))) m
+ | Paddadr rd r1 id ofs =>
+ Next (nextinstr (rs#rd <- (Val.addl rs#r1 (symbol_low ge id ofs)))) m
+ (** Bit-field operations *)
+ | Psbfiz W rd r1 r s =>
+ Next (nextinstr (rs#rd <- (Val.shl (Val.sign_ext s rs#r1) (Vint r)))) m
+ | Psbfiz X rd r1 r s =>
+ Next (nextinstr (rs#rd <- (Val.shll (Val.sign_ext_l s rs#r1) (Vint r)))) m
+ | Psbfx W rd r1 r s =>
+ Next (nextinstr (rs#rd <- (Val.sign_ext s (Val.shr rs#r1 (Vint r))))) m
+ | Psbfx X rd r1 r s =>
+ Next (nextinstr (rs#rd <- (Val.sign_ext_l s (Val.shrl rs#r1 (Vint r))))) m
+ | Pubfiz W rd r1 r s =>
+ Next (nextinstr (rs#rd <- (Val.shl (Val.zero_ext s rs#r1) (Vint r)))) m
+ | Pubfiz X rd r1 r s =>
+ Next (nextinstr (rs#rd <- (Val.shll (Val.zero_ext_l s rs#r1) (Vint r)))) m
+ | Pubfx W rd r1 r s =>
+ Next (nextinstr (rs#rd <- (Val.zero_ext s (Val.shru rs#r1 (Vint r))))) m
+ | Pubfx X rd r1 r s =>
+ Next (nextinstr (rs#rd <- (Val.zero_ext_l s (Val.shrlu rs#r1 (Vint r))))) m
+ (** Integer arithmetic, shifted register *)
+ | Padd W rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.add rs##r1 (eval_shift_op_int rs#r2 s)))) m
+ | Padd X rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.addl rs###r1 (eval_shift_op_long rs#r2 s)))) m
+ | Psub W rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.sub rs##r1 (eval_shift_op_int rs#r2 s)))) m
+ | Psub X rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.subl rs###r1 (eval_shift_op_long rs#r2 s)))) m
+ | Pcmp W r1 r2 s =>
+ Next (nextinstr (compare_int rs rs##r1 (eval_shift_op_int rs#r2 s) m)) m
+ | Pcmp X r1 r2 s =>
+ Next (nextinstr (compare_long rs rs###r1 (eval_shift_op_long rs#r2 s) m)) m
+ | Pcmn W r1 r2 s =>
+ Next (nextinstr (compare_int rs rs##r1 (Val.neg (eval_shift_op_int rs#r2 s)) m)) m
+ | Pcmn X r1 r2 s =>
+ Next (nextinstr (compare_long rs rs###r1 (Val.negl (eval_shift_op_long rs#r2 s)) m)) m
+ (** Integer arithmetic, extending register *)
+ | Paddext rd r1 r2 x =>
+ Next (nextinstr (rs#rd <- (Val.addl rs#r1 (eval_extend rs#r2 x)))) m
+ | Psubext rd r1 r2 x =>
+ Next (nextinstr (rs#rd <- (Val.subl rs#r1 (eval_extend rs#r2 x)))) m
+ | Pcmpext r1 r2 x =>
+ Next (nextinstr (compare_long rs rs#r1 (eval_extend rs#r2 x) m)) m
+ | Pcmnext r1 r2 x =>
+ Next (nextinstr (compare_long rs rs#r1 (Val.negl (eval_extend rs#r2 x)) m)) m
+ (** Logical, shifted register *)
+ | Pand W rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.and rs##r1 (eval_shift_op_int rs#r2 s)))) m
+ | Pand X rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.andl rs###r1 (eval_shift_op_long rs#r2 s)))) m
+ | Pbic W rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.and rs##r1 (Val.notint (eval_shift_op_int rs#r2 s))))) m
+ | Pbic X rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.andl rs###r1 (Val.notl (eval_shift_op_long rs#r2 s))))) m
+ | Peon W rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.xor rs##r1 (Val.notint (eval_shift_op_int rs#r2 s))))) m
+ | Peon X rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.xorl rs###r1 (Val.notl (eval_shift_op_long rs#r2 s))))) m
+ | Peor W rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.xor rs##r1 (eval_shift_op_int rs#r2 s)))) m
+ | Peor X rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.xorl rs###r1 (eval_shift_op_long rs#r2 s)))) m
+ | Porr W rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.or rs##r1 (eval_shift_op_int rs#r2 s)))) m
+ | Porr X rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.orl rs###r1 (eval_shift_op_long rs#r2 s)))) m
+ | Porn W rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.or rs##r1 (Val.notint (eval_shift_op_int rs#r2 s))))) m
+ | Porn X rd r1 r2 s =>
+ Next (nextinstr (rs#rd <- (Val.orl rs###r1 (Val.notl (eval_shift_op_long rs#r2 s))))) m
+ | Ptst W r1 r2 s =>
+ Next (nextinstr (compare_int rs (Val.and rs##r1 (eval_shift_op_int rs#r2 s)) (Vint Int.zero) m)) m
+ | Ptst X r1 r2 s =>
+ Next (nextinstr (compare_long rs (Val.andl rs###r1 (eval_shift_op_long rs#r2 s)) (Vlong Int64.zero) m)) m
+ (** Variable shifts *)
+ | Pasrv W rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.shr rs#r1 rs#r2))) m
+ | Pasrv X rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.shrl rs#r1 rs#r2))) m
+ | Plslv W rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.shl rs#r1 rs#r2))) m
+ | Plslv X rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.shll rs#r1 rs#r2))) m
+ | Plsrv W rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.shru rs#r1 rs#r2))) m
+ | Plsrv X rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.shrlu rs#r1 rs#r2))) m
+ | Prorv W rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.ror rs#r1 rs#r2))) m
+ | Prorv X rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.rorl rs#r1 rs#r2))) m
+ (** Conditional data processing *)
+ | Pcsel rd r1 r2 cond =>
+ let v :=
+ match eval_testcond cond rs with
+ | Some true => rs#r1
+ | Some false => rs#r2
+ | None => Vundef
+ end in
+ Next (nextinstr (rs#rd <- v)) m
+ | Pcset rd cond =>
+ let v :=
+ match cond rs with
+ | Some true => Vint Int.one
+ | Some false => Vint Int.zero
+ | None => Vundef
+ end in
+ Next (nextinstr (rs#rd <- v)) m
+ (** Integer multiply/divide *)
+ | Pmadd W rd r1 r2 r3 =>
+ Next (nextinstr (rs#rd <- (Val.add rs##r3 (Val.mul rs#r1 rs#r2)))) m
+ | Pmadd X rd r1 r2 r3 =>
+ Next (nextinstr (rs#rd <- (Val.addl rs###r3 (Val.mull rs#r1 rs#r2)))) m
+ | Pmsub W rd r1 r2 r3 =>
+ Next (nextinstr (rs#rd <- (Val.sub rs##r3 (Val.mul rs#r1 rs#r2)))) m
+ | Pmsub X rd r1 r2 r3 =>
+ Next (nextinstr (rs#rd <- (Val.subl rs###r3 (Val.mull rs#r1 rs#r2)))) m
+ | Psmulh rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.mullhs rs#r1 rs#r2))) m
+ | Pumulh rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.mullhu rs#r1 rs#r2))) m
+ | Psdiv W rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.divs rs#r1 rs#r2)))) m
+ | Psdiv X rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.divls rs#r1 rs#r2)))) m
+ | Pudiv W rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.divu rs#r1 rs#r2)))) m
+ | Pudiv X rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.divlu rs#r1 rs#r2)))) m
+ (** Floating-point loads and stores *)
+ | Pldrs rd a =>
+ exec_load Mfloat32 (fun v => v) a rd rs m
+ | Pldrd rd a =>
+ exec_load Mfloat64 (fun v => v) a rd rs m
+ | Pldrd_a rd a =>
+ exec_load Many64 (fun v => v) a rd rs m
+ | Pstrs r a =>
+ exec_store Mfloat32 a rs#r rs m
+ | Pstrd r a =>
+ exec_store Mfloat64 a rs#r rs m
+ | Pstrd_a r a =>
+ exec_store Many64 a rs#r rs m
+ (** Floating-point move *)
+ | Pfmov rd r1 =>
+ Next (nextinstr (rs#rd <- (rs#r1))) m
+ | Pfmovimms rd f =>
+ Next (nextinstr (rs#rd <- (Vsingle f))) m
+ | Pfmovimmd rd f =>
+ Next (nextinstr (rs#rd <- (Vfloat f))) m
+ | Pfmovi S rd r1 =>
+ Next (nextinstr (rs#rd <- (float32_of_bits rs##r1))) m
+ | Pfmovi D rd r1 =>
+ Next (nextinstr (rs#rd <- (float64_of_bits rs###r1))) m
+ (** Floating-point conversions *)
+ | Pfcvtds rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.floatofsingle rs#r1))) m
+ | Pfcvtsd rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.singleoffloat rs#r1))) m
+ | Pfcvtzs W S rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.intofsingle rs#r1)))) m
+ | Pfcvtzs W D rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.intoffloat rs#r1)))) m
+ | Pfcvtzs X S rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.longofsingle rs#r1)))) m
+ | Pfcvtzs X D rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.longoffloat rs#r1)))) m
+ | Pfcvtzu W S rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.intuofsingle rs#r1)))) m
+ | Pfcvtzu W D rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.intuoffloat rs#r1)))) m
+ | Pfcvtzu X S rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.longuofsingle rs#r1)))) m
+ | Pfcvtzu X D rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.longuoffloat rs#r1)))) m
+ | Pscvtf S W rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.singleofint rs#r1)))) m
+ | Pscvtf D W rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.floatofint rs#r1)))) m
+ | Pscvtf S X rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.singleoflong rs#r1)))) m
+ | Pscvtf D X rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.floatoflong rs#r1)))) m
+ | Pucvtf S W rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.singleofintu rs#r1)))) m
+ | Pucvtf D W rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.floatofintu rs#r1)))) m
+ | Pucvtf S X rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.singleoflongu rs#r1)))) m
+ | Pucvtf D X rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.maketotal (Val.floatoflongu rs#r1)))) m
+ (** Floating-point arithmetic *)
+ | Pfabs S rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.absfs rs#r1))) m
+ | Pfabs D rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.absf rs#r1))) m
+ | Pfneg S rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.negfs rs#r1))) m
+ | Pfneg D rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.negf rs#r1))) m
+ | Pfadd S rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.addfs rs#r1 rs#r2))) m
+ | Pfadd D rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.addf rs#r1 rs#r2))) m
+ | Pfdiv S rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.divfs rs#r1 rs#r2))) m
+ | Pfdiv D rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.divf rs#r1 rs#r2))) m
+ | Pfmul S rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.mulfs rs#r1 rs#r2))) m
+ | Pfmul D rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.mulf rs#r1 rs#r2))) m
+ | Pfnmul S rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.negfs (Val.mulfs rs#r1 rs#r2)))) m
+ | Pfnmul D rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.negf (Val.mulf rs#r1 rs#r2)))) m
+ | Pfsub S rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.subfs rs#r1 rs#r2))) m
+ | Pfsub D rd r1 r2 =>
+ Next (nextinstr (rs#rd <- (Val.subf rs#r1 rs#r2))) m
+ (** Floating-point comparison *)
+ | Pfcmp S r1 r2 =>
+ Next (nextinstr (compare_single rs rs#r1 rs#r2)) m
+ | Pfcmp D r1 r2 =>
+ Next (nextinstr (compare_float rs rs#r1 rs#r2)) m
+ | Pfcmp0 S r1 =>
+ Next (nextinstr (compare_single rs rs#r1 (Vsingle Float32.zero))) m
+ | Pfcmp0 D r1 =>
+ Next (nextinstr (compare_float rs rs#r1 (Vfloat Float.zero))) m
+ (** Floating-point conditional select *)
+ | Pfsel rd r1 r2 cond =>
+ let v :=
+ match eval_testcond cond rs with
+ | Some true => rs#r1
+ | Some false => rs#r2
+ | None => Vundef
+ end in
+ Next (nextinstr (rs#rd <- v)) m
+ (** Pseudo-instructions *)
+ | Pallocframe sz pos =>
+ let (m1, stk) := Mem.alloc m 0 sz in
+ let sp := (Vptr stk Ptrofs.zero) in
+ match Mem.storev Mint64 m1 (Val.offset_ptr sp pos) rs#SP with
+ | None => Stuck
+ | Some m2 => Next (nextinstr (rs #X29 <- (rs#SP) #SP <- sp #X16 <- Vundef)) m2
+ end
+ | Pfreeframe sz pos =>
+ match Mem.loadv Mint64 m (Val.offset_ptr rs#SP pos) with
+ | None => Stuck
+ | Some v =>
+ match rs#SP with
+ | Vptr stk ofs =>
+ match Mem.free m stk 0 sz with
+ | None => Stuck
+ | Some m' => Next (nextinstr (rs#SP <- v #X16 <- Vundef)) m'
+ end
+ | _ => Stuck
+ end
+ end
+ | Plabel lbl =>
+ Next (nextinstr rs) m
+ | Ploadsymbol rd id =>
+ Next (nextinstr (rs#rd <- (Genv.symbol_address ge id Ptrofs.zero))) m
+ | Pcvtsw2x rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.longofint rs#r1))) m
+ | Pcvtuw2x rd r1 =>
+ Next (nextinstr (rs#rd <- (Val.longofintu rs#r1))) m
+ | Pcvtx2w rd =>
+ Next (nextinstr (rs#rd <- (Val.loword rs#rd))) m
+ | Pbtbl r tbl =>
+ match (rs#X16 <- Vundef)#r with
+ | Vint n =>
+ match list_nth_z tbl (Int.unsigned n) with
+ | None => Stuck
+ | Some lbl => goto_label f lbl (rs#X16 <- Vundef #X17 <- Vundef) m
+ end
+ | _ => Stuck
+ end
+ | Pbuiltin ef args res => Stuck (**r treated specially below *)
+ (** The following instructions and directives are not generated directly
+ by Asmgen, so we do not model them. *)
+ | Pldp _ _ _
+ | Pstp _ _ _
+ | Pcls _ _ _
+ | Pclz _ _ _
+ | Prev _ _ _
+ | Prev16 _ _ _
+ | Pfsqrt _ _ _
+ | Pfmadd _ _ _ _ _
+ | Pfmsub _ _ _ _ _
+ | Pfnmadd _ _ _ _ _
+ | Pfnmsub _ _ _ _ _
+ | Pnop
+ | Pcfi_adjust _
+ | Pcfi_rel_offset _ =>
+ Stuck
+ end.
+*)
+
+(** Translation of the LTL/Linear/Mach view of machine registers
+ to the AArch64 view. Note that no LTL register maps to [X16],
+ [X18], nor [X30].
+ [X18] is reserved as the platform register and never used by the
+ code generated by CompCert.
+ [X30] is used for the return address, and can also be used as temporary.
+ [X16] can be used as temporary. *)
+
+Definition preg_of (r: mreg) : preg :=
+ match r with
+ | R0 => X0 | R1 => X1 | R2 => X2 | R3 => X3
+ | R4 => X4 | R5 => X5 | R6 => X6 | R7 => X7
+ | R8 => X8 | R9 => X9 | R10 => X10 | R11 => X11
+ | R12 => X12 | R13 => X13 | R14 => X14 | R15 => X15
+ | R17 => X17 | R19 => X19
+ | R20 => X20 | R21 => X21 | R22 => X22 | R23 => X23
+ | R24 => X24 | R25 => X25 | R26 => X26 | R27 => X27
+ | R28 => X28 | R29 => X29
+ | F0 => D0 | F1 => D1 | F2 => D2 | F3 => D3
+ | F4 => D4 | F5 => D5 | F6 => D6 | F7 => D7
+ | F8 => D8 | F9 => D9 | F10 => D10 | F11 => D11
+ | F12 => D12 | F13 => D13 | F14 => D14 | F15 => D15
+ | F16 => D16 | F17 => D17 | F18 => D18 | F19 => D19
+ | F20 => D20 | F21 => D21 | F22 => D22 | F23 => D23
+ | F24 => D24 | F25 => D25 | F26 => D26 | F27 => D27
+ | F28 => D28 | F29 => D29 | F30 => D30 | F31 => D31
+ end.
+
+(** Undefine all registers except SP and callee-save registers *)
+
+Definition undef_caller_save_regs (rs: regset) : regset :=
+ fun r =>
+ if preg_eq r SP
+ || In_dec preg_eq r (List.map preg_of (List.filter is_callee_save all_mregs))
+ then rs r
+ else Vundef.
+
+(** Extract the values of the arguments of an external call.
+ We exploit the calling conventions from module [Conventions], except that
+ we use AArch64 registers instead of locations. *)
+
+Inductive extcall_arg (rs: regset) (m: mem): loc -> val -> Prop :=
+ | extcall_arg_reg: forall r,
+ extcall_arg rs m (R r) (rs (preg_of r))
+ | extcall_arg_stack: forall ofs ty bofs v,
+ bofs = Stacklayout.fe_ofs_arg + 4 * ofs ->
+ Mem.loadv (chunk_of_type ty) m
+ (Val.offset_ptr rs#SP (Ptrofs.repr bofs)) = Some v ->
+ extcall_arg rs m (Locations.S Outgoing ofs ty) v.
+
+Inductive extcall_arg_pair (rs: regset) (m: mem): rpair loc -> val -> Prop :=
+ | extcall_arg_one: forall l v,
+ extcall_arg rs m l v ->
+ extcall_arg_pair rs m (One l) v
+ | extcall_arg_twolong: forall hi lo vhi vlo,
+ extcall_arg rs m hi vhi ->
+ extcall_arg rs m lo vlo ->
+ extcall_arg_pair rs m (Twolong hi lo) (Val.longofwords vhi vlo).
+
+Definition extcall_arguments
+ (rs: regset) (m: mem) (sg: signature) (args: list val) : Prop :=
+ list_forall2 (extcall_arg_pair rs m) (loc_arguments sg) args.
+
+Definition loc_external_result (sg: signature) : rpair preg :=
+ map_rpair preg_of (loc_result sg).
+
+
+(** execution of the body of a bblock *)
+Fixpoint exec_body (body: list basic) (rs: regset) (m: mem): outcome :=
+ match body with
+ | nil => Next rs m
+ | bi::body' =>
+ SOME o <- exec_basic bi rs m IN
+ exec_body body' (_rs o) (_m o)
+ end.
+
+Definition incrPC size_b (rs: regset) :=
+ rs#PC <- (Val.offset_ptr rs#PC size_b).
+
+(** execution of the exit instruction of a bblock *)
+Inductive exec_exit (f: function) size_b (rs: regset) (m: mem): (option control) -> trace -> regset -> mem -> Prop :=
+ | none_step:
+ exec_exit f size_b rs m None E0 (incrPC size_b rs) m
+ | cfi_step (cfi: cf_instruction) rs' m':
+ exec_cfi f cfi (incrPC size_b rs) m = Next rs' m' ->
+ exec_exit f size_b rs m (Some (PCtlFlow cfi)) E0 rs' m'
+ | builtin_step ef args res vargs t vres rs' m':
+ eval_builtin_args ge rs rs#SP m args vargs ->
+ external_call ef ge vargs m t vres m' ->
+ rs' = incrPC size_b
+ (set_res res vres
+ (undef_regs (map preg_of (destroyed_by_builtin ef)) rs)) ->
+ exec_exit f size_b rs m (Some (Pbuiltin ef args res)) t rs' m'
+ .
+
+Definition exec_bblock (f: function) (b: bblock) (rs: regset) (m: mem) (t:trace) (rs':regset) (m':mem): Prop
+ := exists rs1 m1, exec_body (body b) rs m = Next rs1 m1 /\ exec_exit f (Ptrofs.repr (size b)) rs1 m1 (exit b) t rs' m'.
+
+Fixpoint find_bblock (pos: Z) (lb: bblocks) {struct lb} : option bblock :=
+ match lb with
+ | nil => None
+ | b :: il =>
+ if zlt pos 0 then None (* NOTE: It is impossible to branch inside a block *)
+ else if zeq pos 0 then Some b
+ else find_bblock (pos - (size b)) il
+ end.
+
+(** Execution of the instruction at [rs PC]. *)
+
+Inductive step: state -> trace -> state -> Prop :=
+ | exec_step_internal:
+ forall b ofs f bi rs m t rs' m',
+ rs PC = Vptr b ofs ->
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
+ find_bblock (Ptrofs.unsigned ofs) (fn_blocks f) = Some bi ->
+ exec_bblock f bi rs m t rs' m' ->
+ step (State rs m) t (State rs' m')
+ | exec_step_external:
+ forall b ef args res rs m t rs' m',
+ rs PC = Vptr b Ptrofs.zero ->
+ Genv.find_funct_ptr ge b = Some (External ef) ->
+ external_call ef ge args m t res m' ->
+ extcall_arguments rs m (ef_sig ef) args ->
+ rs' = (set_pair (loc_external_result (ef_sig ef) ) res (undef_caller_save_regs rs))#PC <- (rs RA) ->
+ step (State rs m) t (State rs' m')
+ .
+
+
+End RELSEM.
+
+
+(** Execution of whole programs. *)
+
+Inductive initial_state (p: program): state -> Prop :=
+ | initial_state_intro: forall m0,
+ Genv.init_mem p = Some m0 ->
+ let ge := Genv.globalenv p in
+ let rs0 :=
+ (Pregmap.init Vundef)
+ # PC <- (Genv.symbol_address ge p.(prog_main) Ptrofs.zero)
+ # RA <- Vnullptr
+ # SP <- Vnullptr in
+ initial_state p (State rs0 m0).
+
+Inductive final_state: state -> int -> Prop :=
+ | final_state_intro: forall rs m r,
+ rs#PC = Vnullptr ->
+ rs#X0 = Vint r ->
+ final_state (State rs m) r.
+
+Definition semantics (lk: aarch64_linker) (p: program) :=
+ Semantics (step lk) (initial_state p) final_state (Genv.globalenv p).
diff --git a/aarch64/Asmblockgen.v b/aarch64/Asmblockgen.v
new file mode 100644
index 00000000..83f2b96b
--- /dev/null
+++ b/aarch64/Asmblockgen.v
@@ -0,0 +1,1197 @@
+(* *************************************************************)
+(* *)
+(* 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 *)
+(* *)
+(* Copyright Kalray. Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+(** * Translation from Machblock to AArch64 assembly block language (Asmblock)
+ Inspired from the Mach->Asm pass of original Leroy's backend, but adapted to the block structure like the KVX backend. *)
+
+Require Import Recdef Coqlib Zwf Zbits.
+Require Import Errors AST Integers Floats Op.
+Require Import Locations Machblock Asmblock.
+
+(* STUB *)
+Definition transf_function (f: Machblock.function) : res Asmblock.function :=
+ Error (msg "Asmblockgen not yet implmented").
+
+Definition transf_fundef (f: Machblock.fundef) : res Asmblock.fundef :=
+ transf_partial_fundef transf_function f.
+
+Definition transf_program (p: Machblock.program) : res Asmblock.program :=
+ transform_partial_program transf_fundef p.
+
+(* ORIGINAL aarch64/Asmgen file that needs to be adapted
+
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Translation from Mach to AArch64. *)
+
+Require Import Recdef Coqlib Zwf Zbits.
+Require Import Errors AST Integers Floats Op.
+Require Import Locations Mach Asm.
+
+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).
+
+(** Arithmetic immediates are 12-bit unsigned numbers, possibly shifted left 12 bits *)
+
+Definition is_arith_imm32 (x: int) : bool :=
+ Int.eq x (Int.zero_ext 12 x)
+ || Int.eq x (Int.shl (Int.zero_ext 12 (Int.shru x (Int.repr 12))) (Int.repr 12)).
+
+Definition is_arith_imm64 (x: int64) : bool :=
+ Int64.eq x (Int64.zero_ext 12 x)
+ || Int64.eq x (Int64.shl (Int64.zero_ext 12 (Int64.shru x (Int64.repr 12))) (Int64.repr 12)).
+
+(** Decompose integer literals into 16-bit fragments *)
+
+Fixpoint decompose_int (N: nat) (n p: Z) {struct N} : list (Z * Z) :=
+ match N with
+ | Datatypes.O => nil
+ | Datatypes.S N =>
+ let frag := Zzero_ext 16 (Z.shiftr n p) in
+ if Z.eqb frag 0 then
+ decompose_int N n (p + 16)
+ else
+ (frag, p) :: decompose_int N (Z.ldiff n (Z.shiftl 65535 p)) (p + 16)
+ end.
+
+Definition negate_decomposition (l: list (Z * Z)) :=
+ List.map (fun np => (Z.lxor (fst np) 65535, snd np)) l.
+
+Definition loadimm_k (sz: isize) (rd: ireg) (l: list (Z * Z)) (k: code) : code :=
+ List.fold_right (fun np k => 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 => Pmovz sz rd 0 0 :: k
+ | (n1, p1) :: l => 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 => Pmovn sz rd 0 0 :: k
+ | (n1, p1) :: l => 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 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 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 addimm32 (rd r1: ireg) (n: int) (k: code) : code :=
+ let m := Int.neg n in
+ if Int.eq n (Int.zero_ext 24 n) then
+ addimm_aux (Paddimm W) rd r1 (Int.unsigned n) k
+ else if Int.eq m (Int.zero_ext 24 m) then
+ addimm_aux (Psubimm W) rd r1 (Int.unsigned m) k
+ else if Int.lt n Int.zero then
+ loadimm32 X16 m (Psub W rd r1 X16 SOnone :: k)
+ else
+ loadimm32 X16 n (Padd W rd r1 X16 SOnone :: 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 (Paddimm X) rd r1 (Int64.unsigned n) k
+ else if Int64.eq m (Int64.zero_ext 24 m) then
+ addimm_aux (Psubimm X) rd r1 (Int64.unsigned m) k
+ else if Int64.lt n Int64.zero then
+ 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 if Int.eq n Int.one then
+ Padd W X16 r1 r1 (SOlsr (Int.repr 31)) ::
+ Porr W rd XZR X16 (SOasr n) :: 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 if Int.eq n Int.one then
+ Padd X X16 r1 r1 (SOlsr (Int.repr 63)) ::
+ Porr X rd XZR X16 (SOasr n) :: 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 Archi.pic_code tt then
+ if Ptrofs.eq ofs Ptrofs.zero then
+ Ploadsymbol rd id :: k
+ else
+ Ploadsymbol rd id :: addimm64 rd rd (Ptrofs.to_int64 ofs) k
+ else
+ Padrp rd id ofs :: Paddadr rd rd id ofs :: k.
+
+(** Translate a shifted operand *)
+
+Definition transl_shift (s: Op.shift) (a: int): Asm.shift_op :=
+ match s with
+ | Slsl => SOlsl a
+ | Slsr => SOlsr a
+ | Sasr => SOasr a
+ | Sror => SOror a
+ end.
+
+(** Translation of a condition. Prepends to [k] the instructions
+ that evaluate the condition and leave its boolean result in one of
+ the bits of the condition register. The bit in question is
+ determined by the [crbit_for_cond] function. *)
+
+Definition transl_cond
+ (cond: condition) (args: list mreg) (k: code) :=
+ match cond, args with
+ | (Ccomp c | Ccompu c), a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pcmp W r1 r2 SOnone :: k)
+ | (Ccompshift c s a | Ccompushift c s a), a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pcmp W r1 r2 (transl_shift s a) :: k)
+ | (Ccompimm c n | Ccompuimm c n), a1 :: nil =>
+ do r1 <- ireg_of a1;
+ OK (if is_arith_imm32 n then
+ Pcmpimm W r1 (Int.unsigned n) :: k
+ else if is_arith_imm32 (Int.neg n) then
+ Pcmnimm W r1 (Int.unsigned (Int.neg n)) :: k
+ else
+ loadimm32 X16 n (Pcmp W r1 X16 SOnone :: k))
+ | (Cmaskzero n | Cmasknotzero n), a1 :: nil =>
+ do r1 <- ireg_of a1;
+ OK (if is_logical_imm32 n then
+ Ptstimm W r1 (Int.unsigned n) :: k
+ else
+ loadimm32 X16 n (Ptst W r1 X16 SOnone :: k))
+ | (Ccompl c | Ccomplu c), a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pcmp X r1 r2 SOnone :: k)
+ | (Ccomplshift c s a | Ccomplushift c s a), a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pcmp X r1 r2 (transl_shift s a) :: k)
+ | (Ccomplimm c n | Ccompluimm c n), a1 :: nil =>
+ do r1 <- ireg_of a1;
+ OK (if is_arith_imm64 n then
+ Pcmpimm X r1 (Int64.unsigned n) :: k
+ else if is_arith_imm64 (Int64.neg n) then
+ Pcmnimm X r1 (Int64.unsigned (Int64.neg n)) :: k
+ else
+ loadimm64 X16 n (Pcmp X r1 X16 SOnone :: k))
+ | (Cmasklzero n | Cmasklnotzero n), a1 :: nil =>
+ do r1 <- ireg_of a1;
+ OK (if is_logical_imm64 n then
+ Ptstimm X r1 (Int64.unsigned n) :: k
+ else
+ loadimm64 X16 n (Ptst X r1 X16 SOnone :: k))
+ | Ccompf cmp, a1 :: a2 :: nil =>
+ do r1 <- freg_of a1; do r2 <- freg_of a2;
+ OK (Pfcmp D r1 r2 :: k)
+ | Cnotcompf cmp, a1 :: a2 :: nil =>
+ do r1 <- freg_of a1; do r2 <- freg_of a2;
+ OK (Pfcmp D r1 r2 :: k)
+ | Ccompfzero cmp, a1 :: nil =>
+ do r1 <- freg_of a1;
+ OK (Pfcmp0 D r1 :: k)
+ | Cnotcompfzero cmp, a1 :: nil =>
+ do r1 <- freg_of a1;
+ OK (Pfcmp0 D r1 :: k)
+ | Ccompfs cmp, a1 :: a2 :: nil =>
+ do r1 <- freg_of a1; do r2 <- freg_of a2;
+ OK (Pfcmp S r1 r2 :: k)
+ | Cnotcompfs cmp, a1 :: a2 :: nil =>
+ do r1 <- freg_of a1; do r2 <- freg_of a2;
+ OK (Pfcmp S r1 r2 :: k)
+ | Ccompfszero cmp, a1 :: nil =>
+ do r1 <- freg_of a1;
+ OK (Pfcmp0 S r1 :: k)
+ | Cnotcompfszero cmp, a1 :: nil =>
+ do r1 <- freg_of a1;
+ OK (Pfcmp0 S r1 :: k)
+ | _, _ =>
+ Error(msg "Asmgen.transl_cond")
+ end.
+
+Definition cond_for_signed_cmp (cmp: comparison) :=
+ match cmp with
+ | Ceq => TCeq
+ | Cne => TCne
+ | Clt => TClt
+ | Cle => TCle
+ | Cgt => TCgt
+ | Cge => TCge
+ end.
+
+Definition cond_for_unsigned_cmp (cmp: comparison) :=
+ match cmp with
+ | Ceq => TCeq
+ | Cne => TCne
+ | Clt => TClo
+ | Cle => TCls
+ | Cgt => TChi
+ | Cge => TChs
+ end.
+
+Definition cond_for_float_cmp (cmp: comparison) :=
+ match cmp with
+ | Ceq => TCeq
+ | Cne => TCne
+ | Clt => TCmi
+ | Cle => TCls
+ | Cgt => TCgt
+ | Cge => TCge
+ end.
+
+Definition cond_for_float_not_cmp (cmp: comparison) :=
+ match cmp with
+ | Ceq => TCne
+ | Cne => TCeq
+ | Clt => TCpl
+ | Cle => TChi
+ | Cgt => TCle
+ | Cge => TClt
+ end.
+
+Definition cond_for_cond (cond: condition) :=
+ match cond with
+ | Ccomp cmp => cond_for_signed_cmp cmp
+ | Ccompu cmp => cond_for_unsigned_cmp cmp
+ | Ccompshift cmp s a => cond_for_signed_cmp cmp
+ | Ccompushift cmp s a => cond_for_unsigned_cmp cmp
+ | Ccompimm cmp n => cond_for_signed_cmp cmp
+ | Ccompuimm cmp n => cond_for_unsigned_cmp cmp
+ | Cmaskzero n => TCeq
+ | Cmasknotzero n => TCne
+ | Ccompl cmp => cond_for_signed_cmp cmp
+ | Ccomplu cmp => cond_for_unsigned_cmp cmp
+ | Ccomplshift cmp s a => cond_for_signed_cmp cmp
+ | Ccomplushift cmp s a => cond_for_unsigned_cmp cmp
+ | Ccomplimm cmp n => cond_for_signed_cmp cmp
+ | Ccompluimm cmp n => cond_for_unsigned_cmp cmp
+ | Cmasklzero n => TCeq
+ | Cmasklnotzero n => TCne
+ | Ccompf cmp => cond_for_float_cmp cmp
+ | Cnotcompf cmp => cond_for_float_not_cmp cmp
+ | Ccompfzero cmp => cond_for_float_cmp cmp
+ | Cnotcompfzero cmp => cond_for_float_not_cmp cmp
+ | Ccompfs cmp => cond_for_float_cmp cmp
+ | Cnotcompfs cmp => cond_for_float_not_cmp cmp
+ | Ccompfszero cmp => cond_for_float_cmp cmp
+ | Cnotcompfszero cmp => cond_for_float_not_cmp cmp
+ end.
+
+(** Translation of a conditional branch. Prepends to [k] the instructions
+ that evaluate the condition and ranch to [lbl] if it holds.
+ We recognize some conditional branches that can be implemented
+ without setting then testing condition flags. *)
+
+Definition transl_cond_branch_default
+ (c: condition) (args: list mreg) (lbl: label) (k: code) :=
+ transl_cond c args (Pbc (cond_for_cond c) lbl :: k).
+
+Definition transl_cond_branch
+ (c: condition) (args: list mreg) (lbl: label) (k: code) :=
+ match args, c with
+ | a1 :: nil, (Ccompimm Cne n | Ccompuimm Cne n) =>
+ if Int.eq n Int.zero
+ then (do r1 <- ireg_of a1; OK (Pcbnz W r1 lbl :: k))
+ else transl_cond_branch_default c args lbl k
+ | a1 :: nil, (Ccompimm Ceq n | Ccompuimm Ceq n) =>
+ if Int.eq n Int.zero
+ then (do r1 <- ireg_of a1; OK (Pcbz W r1 lbl :: k))
+ else transl_cond_branch_default c args lbl k
+ | a1 :: nil, (Ccomplimm Cne n | Ccompluimm Cne n) =>
+ if Int64.eq n Int64.zero
+ then (do r1 <- ireg_of a1; OK (Pcbnz X r1 lbl :: k))
+ else transl_cond_branch_default c args lbl k
+ | a1 :: nil, (Ccomplimm Ceq n | Ccompluimm Ceq n) =>
+ if Int64.eq n Int64.zero
+ then (do r1 <- ireg_of a1; OK (Pcbz X r1 lbl :: k))
+ else transl_cond_branch_default c args lbl k
+ | a1 :: nil, Cmaskzero n =>
+ match Int.is_power2 n with
+ | Some bit => do r1 <- ireg_of a1; OK (Ptbz W r1 bit lbl :: k)
+ | None => transl_cond_branch_default c args lbl k
+ end
+ | a1 :: nil, Cmasknotzero n =>
+ match Int.is_power2 n with
+ | Some bit => do r1 <- ireg_of a1; OK (Ptbnz W r1 bit lbl :: k)
+ | None => transl_cond_branch_default c args lbl k
+ end
+ | a1 :: nil, Cmasklzero n =>
+ match Int64.is_power2' n with
+ | Some bit => do r1 <- ireg_of a1; OK (Ptbz X r1 bit lbl :: k)
+ | None => transl_cond_branch_default c args lbl k
+ end
+ | a1 :: nil, Cmasklnotzero n =>
+ match Int64.is_power2' n with
+ | Some bit => do r1 <- ireg_of a1; OK (Ptbnz X r1 bit lbl :: k)
+ | None => transl_cond_branch_default c args lbl k
+ end
+ | _, _ =>
+ transl_cond_branch_default c args lbl k
+ end.
+
+(** Translation of the arithmetic operation [res <- op(args)].
+ The corresponding instructions are prepended to [k]. *)
+
+Definition transl_op
+ (op: operation) (args: list mreg) (res: mreg) (k: code) :=
+ match op, args with
+ | Omove, a1 :: nil =>
+ match preg_of res, preg_of a1 with
+ | IR r, IR a => OK (Pmov r a :: k)
+ | FR r, FR a => OK (Pfmov r a :: k)
+ | _ , _ => Error(msg "Asmgen.Omove")
+ end
+ | Ointconst n, nil =>
+ do rd <- ireg_of res;
+ OK (loadimm32 rd n k)
+ | Olongconst n, nil =>
+ do rd <- ireg_of res;
+ OK (loadimm64 rd n k)
+ | Ofloatconst f, nil =>
+ do rd <- freg_of res;
+ OK (if Float.eq_dec f Float.zero
+ then Pfmovi D rd XZR :: k
+ else Pfmovimmd rd f :: k)
+ | Osingleconst f, nil =>
+ do rd <- freg_of res;
+ OK (if Float32.eq_dec f Float32.zero
+ then Pfmovi S rd XZR :: k
+ else Pfmovimms rd f :: k)
+ | Oaddrsymbol id ofs, nil =>
+ do rd <- ireg_of res;
+ OK (loadsymbol rd id ofs k)
+ | Oaddrstack ofs, nil =>
+ do rd <- ireg_of res;
+ OK (addimm64 rd XSP (Ptrofs.to_int64 ofs) k)
+(** 32-bit integer arithmetic *)
+ | Oshift s a, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Porr W rd XZR r1 (transl_shift s a) :: k)
+ | Oadd, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Padd W rd r1 r2 SOnone :: k)
+ | Oaddshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Padd W rd r1 r2 (transl_shift s a) :: k)
+ | Oaddimm n, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (addimm32 rd r1 n k)
+ | Oneg, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Psub W rd XZR r1 SOnone :: k)
+ | Onegshift s a, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Psub W rd XZR r1 (transl_shift s a) :: k)
+ | Osub, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Psub W rd r1 r2 SOnone :: k)
+ | Osubshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Psub W rd r1 r2 (transl_shift s a) :: k)
+ | Omul, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pmadd W rd r1 r2 XZR :: k)
+ | Omuladd, a1 :: a2 :: a3 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r3 <- ireg_of a3;
+ OK (Pmadd W rd r2 r3 r1 :: k)
+ | Omulsub, a1 :: a2 :: a3 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r3 <- ireg_of a3;
+ OK (Pmsub W rd r2 r3 r1 :: k)
+ | Odiv, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Psdiv W rd r1 r2 :: k)
+ | Odivu, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pudiv W rd r1 r2 :: k)
+ | Oand, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pand W rd r1 r2 SOnone :: k)
+ | Oandshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pand W rd r1 r2 (transl_shift s a) :: k)
+ | Oandimm n, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (logicalimm32 (Pandimm W) (Pand W) rd r1 n k)
+ | Oor, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Porr W rd r1 r2 SOnone :: k)
+ | Oorshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Porr W rd r1 r2 (transl_shift s a) :: k)
+ | Oorimm n, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (logicalimm32 (Porrimm W) (Porr W) rd r1 n k)
+ | Oxor, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Peor W rd r1 r2 SOnone :: k)
+ | Oxorshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Peor W rd r1 r2 (transl_shift s a) :: k)
+ | Oxorimm n, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (logicalimm32 (Peorimm W) (Peor W) rd r1 n k)
+ | Onot, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Porn W rd XZR r1 SOnone :: k)
+ | Onotshift s a, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Porn W rd XZR r1 (transl_shift s a) :: k)
+ | Obic, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pbic W rd r1 r2 SOnone :: k)
+ | Obicshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pbic W rd r1 r2 (transl_shift s a) :: k)
+ | Oorn, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Porn W rd r1 r2 SOnone :: k)
+ | Oornshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Porn W rd r1 r2 (transl_shift s a) :: k)
+ | Oeqv, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Peon W rd r1 r2 SOnone :: k)
+ | Oeqvshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Peon W rd r1 r2 (transl_shift s a) :: k)
+ | Oshl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Plslv W rd r1 r2 :: k)
+ | Oshr, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pasrv W rd r1 r2 :: k)
+ | Oshru, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Plsrv W rd r1 r2 :: k)
+ | Oshrximm n, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (shrx32 rd r1 n k)
+ | Ozext s, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Pubfiz W rd r1 Int.zero s :: k)
+ | Osext s, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Psbfiz W rd r1 Int.zero s :: k)
+ | Oshlzext s a, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Pubfiz W rd r1 a (Z.min s (Int.zwordsize - Int.unsigned a)) :: k)
+ | Oshlsext s a, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Psbfiz W rd r1 a (Z.min s (Int.zwordsize - Int.unsigned a)) :: k)
+ | Ozextshr a s, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Pubfx W rd r1 a (Z.min s (Int.zwordsize - Int.unsigned a)) :: k)
+ | Osextshr a s, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Psbfx W rd r1 a (Z.min s (Int.zwordsize - Int.unsigned a)) :: k)
+(** 64-bit integer arithmetic *)
+ | Oshiftl s a, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Porr X rd XZR r1 (transl_shift s a) :: k)
+ | Oextend x a, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (move_extended rd r1 x a k)
+ (* [Omakelong] and [Ohighlong] should not occur *)
+ | Olowlong, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ assertion (ireg_eq rd r1);
+ OK (Pcvtx2w rd :: k)
+ | Oaddl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Padd X rd r1 r2 SOnone :: k)
+ | Oaddlshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Padd X rd r1 r2 (transl_shift s a) :: k)
+ | Oaddlext x a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (arith_extended Paddext (Padd X) rd r1 r2 x a k)
+ | Oaddlimm n, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (addimm64 rd r1 n k)
+ | Onegl, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Psub X rd XZR r1 SOnone :: k)
+ | Oneglshift s a, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Psub X rd XZR r1 (transl_shift s a) :: k)
+ | Osubl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Psub X rd r1 r2 SOnone :: k)
+ | Osublshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Psub X rd r1 r2 (transl_shift s a) :: k)
+ | Osublext x a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (arith_extended Psubext (Psub X) rd r1 r2 x a k)
+ | Omull, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pmadd X rd r1 r2 XZR :: k)
+ | Omulladd, a1 :: a2 :: a3 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r3 <- ireg_of a3;
+ OK (Pmadd X rd r2 r3 r1 :: k)
+ | Omullsub, a1 :: a2 :: a3 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r3 <- ireg_of a3;
+ OK (Pmsub X rd r2 r3 r1 :: k)
+ | Omullhs, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Psmulh rd r1 r2 :: k)
+ | Omullhu, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pumulh rd r1 r2 :: k)
+ | Odivl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Psdiv X rd r1 r2 :: k)
+ | Odivlu, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pudiv X rd r1 r2 :: k)
+ | Oandl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pand X rd r1 r2 SOnone :: k)
+ | Oandlshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pand X rd r1 r2 (transl_shift s a) :: k)
+ | Oandlimm n, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (logicalimm64 (Pandimm X) (Pand X) rd r1 n k)
+ | Oorl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Porr X rd r1 r2 SOnone :: k)
+ | Oorlshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Porr X rd r1 r2 (transl_shift s a) :: k)
+ | Oorlimm n, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (logicalimm64 (Porrimm X) (Porr X) rd r1 n k)
+ | Oxorl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Peor X rd r1 r2 SOnone :: k)
+ | Oxorlshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Peor X rd r1 r2 (transl_shift s a) :: k)
+ | Oxorlimm n, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (logicalimm64 (Peorimm X) (Peor X) rd r1 n k)
+ | Onotl, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Porn X rd XZR r1 SOnone :: k)
+ | Onotlshift s a, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Porn X rd XZR r1 (transl_shift s a) :: k)
+ | Obicl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pbic X rd r1 r2 SOnone :: k)
+ | Obiclshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pbic X rd r1 r2 (transl_shift s a) :: k)
+ | Oornl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Porn X rd r1 r2 SOnone :: k)
+ | Oornlshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Porn X rd r1 r2 (transl_shift s a) :: k)
+ | Oeqvl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Peon X rd r1 r2 SOnone :: k)
+ | Oeqvlshift s a, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Peon X rd r1 r2 (transl_shift s a) :: k)
+ | Oshll, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Plslv X rd r1 r2 :: k)
+ | Oshrl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Pasrv X rd r1 r2 :: k)
+ | Oshrlu, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (Plsrv X rd r1 r2 :: k)
+ | Oshrlximm n, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (shrx64 rd r1 n k)
+ | Ozextl s, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Pubfiz X rd r1 Int.zero s :: k)
+ | Osextl s, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Psbfiz X rd r1 Int.zero s :: k)
+ | Oshllzext s a, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Pubfiz X rd r1 a (Z.min s (Int64.zwordsize - Int.unsigned a)) :: k)
+ | Oshllsext s a, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Psbfiz X rd r1 a (Z.min s (Int64.zwordsize - Int.unsigned a)) :: k)
+ | Ozextshrl a s, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Pubfx X rd r1 a (Z.min s (Int64.zwordsize - Int.unsigned a)) :: k)
+ | Osextshrl a s, a1 :: nil =>
+ do rd <- ireg_of res; do r1 <- ireg_of a1;
+ OK (Psbfx X rd r1 a (Z.min s (Int64.zwordsize - Int.unsigned a)) :: k)
+(** 64-bit floating-point arithmetic *)
+ | Onegf, a1 :: nil =>
+ do rd <- freg_of res; do rs <- freg_of a1;
+ OK (Pfneg D rd rs :: k)
+ | Oabsf, a1 :: nil =>
+ do rd <- freg_of res; do rs <- freg_of a1;
+ OK (Pfabs D rd rs :: k)
+ | Oaddf, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfadd D rd rs1 rs2 :: k)
+ | Osubf, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfsub D rd rs1 rs2 :: k)
+ | Omulf, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfmul D rd rs1 rs2 :: k)
+ | Odivf, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfdiv D rd rs1 rs2 :: k)
+(** 32-bit floating-point arithmetic *)
+ | Onegfs, a1 :: nil =>
+ do rd <- freg_of res; do rs <- freg_of a1;
+ OK (Pfneg S rd rs :: k)
+ | Oabsfs, a1 :: nil =>
+ do rd <- freg_of res; do rs <- freg_of a1;
+ OK (Pfabs S rd rs :: k)
+ | Oaddfs, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfadd S rd rs1 rs2 :: k)
+ | Osubfs, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfsub S rd rs1 rs2 :: k)
+ | Omulfs, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfmul S rd rs1 rs2 :: k)
+ | Odivfs, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfdiv S rd rs1 rs2 :: k)
+ | Osingleoffloat, a1 :: nil =>
+ do rd <- freg_of res; do rs <- freg_of a1;
+ OK (Pfcvtsd rd rs :: k)
+ | Ofloatofsingle, a1 :: nil =>
+ do rd <- freg_of res; do rs <- freg_of a1;
+ OK (Pfcvtds rd rs :: k)
+(** Conversions between int and float *)
+ | Ointoffloat, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtzs W D rd rs :: k)
+ | Ointuoffloat, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtzu W D rd rs :: k)
+ | Ofloatofint, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pscvtf D W rd rs :: k)
+ | Ofloatofintu, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pucvtf D W rd rs :: k)
+ | Ointofsingle, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtzs W S rd rs :: k)
+ | Ointuofsingle, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtzu W S rd rs :: k)
+ | Osingleofint, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pscvtf S W rd rs :: k)
+ | Osingleofintu, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pucvtf S W rd rs :: k)
+ | Olongoffloat, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtzs X D rd rs :: k)
+ | Olonguoffloat, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtzu X D rd rs :: k)
+ | Ofloatoflong, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pscvtf D X rd rs :: k)
+ | Ofloatoflongu, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pucvtf D X rd rs :: k)
+ | Olongofsingle, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtzs X S rd rs :: k)
+ | Olonguofsingle, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtzu X S rd rs :: k)
+ | Osingleoflong, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pscvtf S X rd rs :: k)
+ | Osingleoflongu, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pucvtf S X rd rs :: k)
+(** Boolean tests *)
+ | Ocmp c, _ =>
+ do rd <- ireg_of res;
+ transl_cond c args (Pcset rd (cond_for_cond c) :: k)
+(** Conditional move *)
+ | Osel cmp ty, a1 :: a2 :: args =>
+ match preg_of res with
+ | IR r =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ transl_cond cmp args (Pcsel r r1 r2 (cond_for_cond cmp) :: k)
+ | FR r =>
+ do r1 <- freg_of a1; do r2 <- freg_of a2;
+ transl_cond cmp args (Pfsel r r1 r2 (cond_for_cond cmp) :: k)
+ | _ =>
+ Error(msg "Asmgen.Osel")
+ end
+ | _, _ =>
+ Error(msg "Asmgen.transl_op")
+ end.
+
+(** 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 (Archi.pic_code tt));
+ 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")
+ end.
+
+(** Translation of loads and stores *)
+
+Definition transl_load (trap: trapping_mode)
+ (chunk: memory_chunk) (addr: Op.addressing)
+ (args: list mreg) (dst: mreg) (k: code) : res code :=
+ match trap with
+ | NOTRAP => Error (msg "Asmgen.transl_load non-trapping loads unsupported on aarch64")
+ | TRAP =>
+ match chunk with
+ | Mint8unsigned =>
+ do rd <- ireg_of dst; transl_addressing 1 addr args (Pldrb W rd) k
+ | Mint8signed =>
+ do rd <- ireg_of dst; transl_addressing 1 addr args (Pldrsb W rd) k
+ | Mint16unsigned =>
+ do rd <- ireg_of dst; transl_addressing 2 addr args (Pldrh W rd) k
+ | Mint16signed =>
+ do rd <- ireg_of dst; transl_addressing 2 addr args (Pldrsh W rd) k
+ | Mint32 =>
+ do rd <- ireg_of dst; transl_addressing 4 addr args (Pldrw rd) k
+ | Mint64 =>
+ do rd <- ireg_of dst; transl_addressing 8 addr args (Pldrx rd) k
+ | Mfloat32 =>
+ do rd <- freg_of dst; transl_addressing 4 addr args (Pldrs rd) k
+ | Mfloat64 =>
+ do rd <- freg_of dst; transl_addressing 8 addr args (Pldrd rd) k
+ | Many32 =>
+ do rd <- ireg_of dst; transl_addressing 4 addr args (Pldrw_a rd) k
+ | Many64 =>
+ do rd <- ireg_of dst; transl_addressing 8 addr args (Pldrx_a rd) k
+ end
+ end.
+
+Definition transl_store (chunk: memory_chunk) (addr: Op.addressing)
+ (args: list mreg) (src: mreg) (k: code) : res code :=
+ match chunk with
+ | Mint8unsigned | Mint8signed =>
+ do r1 <- ireg_of src; transl_addressing 1 addr args (Pstrb r1) k
+ | Mint16unsigned | Mint16signed =>
+ do r1 <- ireg_of src; transl_addressing 2 addr args (Pstrh r1) k
+ | Mint32 =>
+ do r1 <- ireg_of src; transl_addressing 4 addr args (Pstrw r1) k
+ | Mint64 =>
+ do r1 <- ireg_of src; transl_addressing 8 addr args (Pstrx r1) k
+ | Mfloat32 =>
+ do r1 <- freg_of src; transl_addressing 4 addr args (Pstrs r1) k
+ | Mfloat64 =>
+ do r1 <- freg_of src; transl_addressing 8 addr args (Pstrd r1) k
+ | Many32 =>
+ do r1 <- ireg_of src; transl_addressing 4 addr args (Pstrw_a r1) k
+ | Many64 =>
+ do r1 <- ireg_of src; transl_addressing 8 addr args (Pstrx_a r1) k
+ end.
+
+(** Register-indexed loads and 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 loadind (base: iregsp) (ofs: ptrofs) (ty: typ) (dst: mreg) (k: code) :=
+ match ty, preg_of dst with
+ | Tint, IR rd => OK (indexed_memory_access (Pldrw rd) 4 base ofs k)
+ | Tlong, IR rd => OK (indexed_memory_access (Pldrx rd) 8 base ofs k)
+ | Tsingle, FR rd => OK (indexed_memory_access (Pldrs rd) 4 base ofs k)
+ | Tfloat, FR rd => OK (indexed_memory_access (Pldrd rd) 8 base ofs k)
+ | Tany32, IR rd => OK (indexed_memory_access (Pldrw_a rd) 4 base ofs k)
+ | Tany64, IR rd => OK (indexed_memory_access (Pldrx_a rd) 8 base ofs k)
+ | Tany64, FR rd => OK (indexed_memory_access (Pldrd_a rd) 8 base ofs k)
+ | _, _ => Error (msg "Asmgen.loadind")
+ end.
+
+Definition storeind (src: mreg) (base: iregsp) (ofs: ptrofs) (ty: typ) (k: code) :=
+ match ty, preg_of src with
+ | Tint, IR rd => OK (indexed_memory_access (Pstrw rd) 4 base ofs k)
+ | Tlong, IR rd => OK (indexed_memory_access (Pstrx rd) 8 base ofs k)
+ | Tsingle, FR rd => OK (indexed_memory_access (Pstrs rd) 4 base ofs k)
+ | Tfloat, FR rd => OK (indexed_memory_access (Pstrd rd) 8 base ofs k)
+ | Tany32, IR rd => OK (indexed_memory_access (Pstrw_a rd) 4 base ofs k)
+ | Tany64, IR rd => OK (indexed_memory_access (Pstrx_a rd) 8 base ofs k)
+ | Tany64, FR rd => OK (indexed_memory_access (Pstrd_a rd) 8 base ofs k)
+ | _, _ => Error (msg "Asmgen.storeind")
+ end.
+
+Definition loadptr (base: iregsp) (ofs: ptrofs) (dst: ireg) (k: code) :=
+ indexed_memory_access (Pldrx dst) 8 base ofs k.
+
+Definition storeptr (src: ireg) (base: iregsp) (ofs: ptrofs) (k: code) :=
+ indexed_memory_access (Pstrx src) 8 base ofs k.
+
+(** Function epilogue *)
+
+Definition make_epilogue (f: Mach.function) (k: code) :=
+ (* FIXME
+ Cannot be used because memcpy destroys X30;
+ issue being discussed with X. Leroy *)
+ (* if is_leaf_function f
+ then Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) :: k
+ else*) loadptr XSP f.(fn_retaddr_ofs) RA
+ (Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) :: k).
+
+(** Translation of a Mach instruction. *)
+
+Definition transl_instr (f: Mach.function) (i: Mach.instruction)
+ (r29_is_parent: bool) (k: code) : res code :=
+ match i with
+ | Mgetstack ofs ty dst =>
+ loadind XSP ofs ty dst k
+ | Msetstack src ofs ty =>
+ storeind src XSP ofs ty k
+ | Mgetparam ofs ty dst =>
+ (* load via the frame pointer if it is valid *)
+ do c <- loadind X29 ofs ty dst k;
+ OK (if r29_is_parent then c else loadptr XSP f.(fn_link_ofs) X29 c)
+ | Mop op args res =>
+ transl_op op args res k
+ | Mload trap chunk addr args dst =>
+ transl_load trap chunk addr args dst k
+ | Mstore chunk addr args src =>
+ transl_store chunk addr args src k
+ | Mcall sig (inl r) =>
+ do r1 <- ireg_of r; OK (Pblr r1 sig :: k)
+ | Mcall sig (inr symb) =>
+ OK (Pbl symb sig :: k)
+ | Mtailcall sig (inl r) =>
+ do r1 <- ireg_of r;
+ OK (make_epilogue f (Pbr r1 sig :: k))
+ | Mtailcall sig (inr symb) =>
+ OK (make_epilogue f (Pbs symb sig :: k))
+ | Mbuiltin ef args res =>
+ OK (Pbuiltin ef (List.map (map_builtin_arg preg_of) args) (map_builtin_res preg_of res) :: k)
+ | Mlabel lbl =>
+ OK (Plabel lbl :: k)
+ | Mgoto lbl =>
+ OK (Pb lbl :: k)
+ | Mcond cond args lbl =>
+ transl_cond_branch cond args lbl k
+ | Mjumptable arg tbl =>
+ do r <- ireg_of arg;
+ OK (Pbtbl r tbl :: k)
+ | Mreturn =>
+ OK (make_epilogue f (Pret RA :: k))
+ end.
+
+(** Translation of a code sequence *)
+
+Definition it1_is_parent (before: bool) (i: Mach.instruction) : bool :=
+ match i with
+ | Msetstack src ofs ty => before
+ | Mgetparam ofs ty dst => negb (mreg_eq dst R29)
+ | Mop op args res => before && negb (mreg_eq res R29)
+ | _ => false
+ end.
+
+(** This is the naive definition that we no longer use because it
+ is not tail-recursive. It is kept as specification. *)
+
+Fixpoint transl_code (f: Mach.function) (il: list Mach.instruction) (it1p: bool) :=
+ match il with
+ | nil => OK nil
+ | i1 :: il' =>
+ do k <- transl_code f il' (it1_is_parent it1p i1);
+ transl_instr f i1 it1p k
+ end.
+
+(** This is an equivalent definition in continuation-passing style
+ that runs in constant stack space. *)
+
+Fixpoint transl_code_rec (f: Mach.function) (il: list Mach.instruction)
+ (it1p: bool) (k: code -> res code) :=
+ match il with
+ | nil => k nil
+ | i1 :: il' =>
+ transl_code_rec f il' (it1_is_parent it1p i1)
+ (fun c1 => do c2 <- transl_instr f i1 it1p c1; k c2)
+ end.
+
+Definition transl_code' (f: Mach.function) (il: list Mach.instruction) (it1p: bool) :=
+ transl_code_rec f il it1p (fun c => OK c).
+
+(** Translation of a whole function. Note that we must check
+ that the generated code contains less than [2^32] instructions,
+ otherwise the offset part of the [PC] code pointer could wrap
+ around, leading to incorrect executions. *)
+
+Definition transl_function (f: Mach.function) :=
+ do c <- transl_code' f f.(Mach.fn_code) true;
+ OK (mkfunction f.(Mach.fn_sig)
+ (Pallocframe f.(fn_stacksize) f.(fn_link_ofs) ::
+ storeptr RA XSP f.(fn_retaddr_ofs) c)).
+
+*) \ No newline at end of file
diff --git a/aarch64/Asmblockgenproof.v b/aarch64/Asmblockgenproof.v
new file mode 100644
index 00000000..41082772
--- /dev/null
+++ b/aarch64/Asmblockgenproof.v
@@ -0,0 +1,1158 @@
+(** Correctness proof for aarch64/Asmblock generation: main proof.
+CURRENTLY A STUB !
+*)
+
+Require Import Coqlib Errors.
+Require Import Integers Floats AST Linking.
+Require Import Values Memory Events Globalenvs Smallstep.
+Require Import Op Locations Machblock Conventions PseudoAsmblock Asmblock IterList.
+Require Import PseudoAsmblockproof Asmblockgen.
+
+Module MB := Machblock.
+Module AB := Asmblock.
+
+Definition match_prog (p: MB.program) (tp: AB.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.
+
+Parameter next: MB.function -> Z -> Z.
+
+Section PRESERVATION.
+
+Lemma next_progress: forall (f:MB.function) (pos:Z), (pos < (next f pos))%Z.
+Admitted.
+
+Variable lk: aarch64_linker.
+Variable prog: Machblock.program.
+Variable tprog: Asmblock.program.
+Hypothesis TRANSF: match_prog prog tprog.
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+
+Hypothesis symbol_high_low: forall (id: ident) (ofs: ptrofs),
+ Val.addl (symbol_high lk id ofs) (symbol_low lk id ofs) = Genv.symbol_address tge id ofs.
+
+Lemma functions_bound_max_pos: forall fb f,
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ (max_pos next f) <= Ptrofs.max_unsigned.
+Admitted.
+
+
+
+Lemma transf_program_correct:
+ forward_simulation (PseudoAsmblock.semantics next prog) (AB.semantics lk tprog).
+Admitted.
+
+End PRESERVATION.
+
+
+(* ORIGINAL aarch64/Asmgenproof file that needs to be adapted
+
+(* *********************************************************************)
+(* *)
+(* 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. *)
+
+Require Import Coqlib Errors.
+Require Import Integers Floats AST Linking.
+Require Import Values Memory Events Globalenvs Smallstep.
+Require Import Op Locations Mach Conventions Asm.
+Require Import Asmgen Asmgenproof0 Asmgenproof1.
+
+Definition match_prog (p: Mach.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: Mach.program.
+Variable tprog: Asm.program.
+Hypothesis TRANSF: match_prog prog tprog.
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+Lemma symbols_preserved:
+ forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
+Proof (Genv.find_symbol_match TRANSF).
+
+Lemma senv_preserved:
+ Senv.equiv ge tge.
+Proof (Genv.senv_match TRANSF).
+
+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 functions_transl:
+ forall fb f tf,
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ transf_function f = OK tf ->
+ 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.
+ omega.
+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.
+*)
+
+Section TRANSL_LABEL.
+
+Remark loadimm_z_label: forall sz rd l k, tail_nolabel k (loadimm_z sz rd l k).
+Proof.
+ intros; destruct l as [ | [n1 p1] l]; simpl; TailNoLabel.
+ induction l as [ | [n p] l]; simpl; TailNoLabel.
+Qed.
+
+Remark loadimm_n_label: forall sz rd l k, tail_nolabel k (loadimm_n sz rd l k).
+Proof.
+ intros; destruct l as [ | [n1 p1] l]; simpl; TailNoLabel.
+ induction l as [ | [n p] l]; simpl; TailNoLabel.
+Qed.
+
+Remark loadimm_label: forall sz rd n k, tail_nolabel k (loadimm sz rd n k).
+Proof.
+ unfold loadimm; intros. destruct Nat.leb; [apply loadimm_z_label|apply loadimm_n_label].
+Qed.
+Hint Resolve loadimm_label: labels.
+
+Remark loadimm32_label: forall r n k, tail_nolabel k (loadimm32 r n k).
+Proof.
+ unfold loadimm32; intros. destruct (is_logical_imm32 n); TailNoLabel.
+Qed.
+Hint Resolve loadimm32_label: labels.
+
+Remark loadimm64_label: forall r n k, tail_nolabel k (loadimm64 r n k).
+Proof.
+ unfold loadimm64; intros. destruct (is_logical_imm64 n); TailNoLabel.
+Qed.
+Hint Resolve loadimm64_label: labels.
+
+Remark addimm_aux: forall insn rd r1 n k,
+ (forall rd r1 n, nolabel (insn rd r1 n)) ->
+ tail_nolabel k (addimm_aux insn rd r1 n k).
+Proof.
+ unfold addimm_aux; intros.
+ destruct Z.eqb. TailNoLabel. destruct Z.eqb; TailNoLabel.
+Qed.
+
+Remark addimm32_label: forall rd r1 n k, tail_nolabel k (addimm32 rd r1 n k).
+Proof.
+ unfold addimm32; intros.
+ destruct Int.eq. apply addimm_aux; intros; red; auto.
+ destruct Int.eq. apply addimm_aux; intros; red; auto.
+ destruct Int.lt; eapply tail_nolabel_trans; TailNoLabel.
+Qed.
+Hint Resolve addimm32_label: labels.
+
+Remark addimm64_label: forall rd r1 n k, tail_nolabel k (addimm64 rd r1 n k).
+Proof.
+ unfold addimm64; intros.
+ destruct Int64.eq. apply addimm_aux; intros; red; auto.
+ destruct Int64.eq. apply addimm_aux; intros; red; auto.
+ destruct Int64.lt; eapply tail_nolabel_trans; TailNoLabel.
+Qed.
+Hint Resolve addimm64_label: labels.
+
+Remark logicalimm32_label: forall insn1 insn2 rd r1 n k,
+ (forall rd r1 n, nolabel (insn1 rd r1 n)) ->
+ (forall rd r1 r2 s, nolabel (insn2 rd r1 r2 s)) ->
+ tail_nolabel k (logicalimm32 insn1 insn2 rd r1 n k).
+Proof.
+ unfold logicalimm32; intros.
+ destruct (is_logical_imm32 n). TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
+Qed.
+
+Remark logicalimm64_label: forall insn1 insn2 rd r1 n k,
+ (forall rd r1 n, nolabel (insn1 rd r1 n)) ->
+ (forall rd r1 r2 s, nolabel (insn2 rd r1 r2 s)) ->
+ tail_nolabel k (logicalimm64 insn1 insn2 rd r1 n k).
+Proof.
+ unfold logicalimm64; intros.
+ destruct (is_logical_imm64 n). TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
+Qed.
+
+Remark move_extended_label: forall rd r1 ex a k, tail_nolabel k (move_extended rd r1 ex a k).
+Proof.
+ unfold move_extended, move_extended_base; intros. destruct Int.eq, ex; TailNoLabel.
+Qed.
+Hint Resolve move_extended_label: labels.
+
+Remark arith_extended_label: forall insnX insnS rd r1 r2 ex a k,
+ (forall rd r1 r2 x, nolabel (insnX rd r1 r2 x)) ->
+ (forall rd r1 r2 s, nolabel (insnS rd r1 r2 s)) ->
+ tail_nolabel k (arith_extended insnX insnS rd r1 r2 ex a k).
+Proof.
+ unfold arith_extended; intros. destruct Int.ltu.
+ TailNoLabel.
+ destruct ex; simpl; TailNoLabel.
+Qed.
+
+Remark loadsymbol_label: forall r id ofs k, tail_nolabel k (loadsymbol r id ofs k).
+Proof.
+ intros; unfold loadsymbol.
+ destruct (Archi.pic_code tt); TailNoLabel. destruct Ptrofs.eq; TailNoLabel.
+Qed.
+Hint Resolve loadsymbol_label: labels.
+
+Remark transl_cond_label: forall cond args k c,
+ transl_cond cond args k = OK c -> tail_nolabel k c.
+Proof.
+ unfold transl_cond; intros; destruct cond; TailNoLabel.
+- destruct is_arith_imm32; TailNoLabel. destruct is_arith_imm32; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
+- destruct is_arith_imm32; TailNoLabel. destruct is_arith_imm32; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
+- destruct is_logical_imm32; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
+- destruct is_logical_imm32; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
+- destruct is_arith_imm64; TailNoLabel. destruct is_arith_imm64; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
+- destruct is_arith_imm64; TailNoLabel. destruct is_arith_imm64; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
+- destruct is_logical_imm64; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
+- destruct is_logical_imm64; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
+Qed.
+
+Remark transl_cond_branch_default_label: forall cond args lbl k c,
+ transl_cond_branch_default cond args lbl k = OK c -> tail_nolabel k c.
+Proof.
+ unfold transl_cond_branch_default; intros.
+ eapply tail_nolabel_trans; [eapply transl_cond_label;eauto|TailNoLabel].
+Qed.
+Hint Resolve transl_cond_branch_default_label: labels.
+
+Remark transl_cond_branch_label: forall cond args lbl k c,
+ transl_cond_branch cond args lbl k = OK c -> tail_nolabel k c.
+Proof.
+ unfold transl_cond_branch; intros; destruct args; TailNoLabel; destruct cond; TailNoLabel.
+- destruct c0; TailNoLabel.
+- destruct c0; TailNoLabel.
+- destruct (Int.is_power2 n); TailNoLabel.
+- destruct (Int.is_power2 n); TailNoLabel.
+- destruct c0; TailNoLabel.
+- destruct c0; TailNoLabel.
+- destruct (Int64.is_power2' n); TailNoLabel.
+- destruct (Int64.is_power2' n); TailNoLabel.
+Qed.
+
+Remark transl_op_label:
+ forall op args r k c,
+ transl_op op args r k = OK c -> tail_nolabel k c.
+Proof.
+ unfold transl_op; intros; destruct op; TailNoLabel.
+- destruct (preg_of r); try discriminate; destruct (preg_of m); inv H; TailNoLabel.
+- destruct (Float.eq_dec n Float.zero); TailNoLabel.
+- destruct (Float32.eq_dec n Float32.zero); TailNoLabel.
+- apply logicalimm32_label; unfold nolabel; auto.
+- apply logicalimm32_label; unfold nolabel; auto.
+- apply logicalimm32_label; unfold nolabel; auto.
+- unfold shrx32. destruct (Int.eq _ _); try destruct (Int.eq _ _); TailNoLabel.
+- apply arith_extended_label; unfold nolabel; auto.
+- apply arith_extended_label; unfold nolabel; auto.
+- apply logicalimm64_label; unfold nolabel; auto.
+- apply logicalimm64_label; unfold nolabel; auto.
+- apply logicalimm64_label; unfold nolabel; auto.
+- unfold shrx64. destruct (Int.eq _ _); try destruct (Int.eq _ _); TailNoLabel.
+- eapply tail_nolabel_trans. eapply transl_cond_label; eauto. TailNoLabel.
+- destruct (preg_of r); try discriminate; TailNoLabel;
+ (eapply tail_nolabel_trans; [eapply transl_cond_label; eauto | TailNoLabel]).
+Qed.
+
+Remark transl_addressing_label:
+ forall sz addr args insn k c,
+ transl_addressing sz addr args insn k = OK c ->
+ (forall ad, nolabel (insn ad)) ->
+ tail_nolabel k c.
+Proof.
+ unfold transl_addressing; intros; destruct addr; TailNoLabel;
+ eapply tail_nolabel_trans; TailNoLabel.
+ eapply tail_nolabel_trans. apply arith_extended_label; unfold nolabel; auto. TailNoLabel.
+Qed.
+
+Remark transl_load_label:
+ forall trap chunk addr args dst k c,
+ transl_load trap chunk addr args dst k = OK c -> tail_nolabel k c.
+Proof.
+ unfold transl_load; intros; destruct trap; try discriminate; destruct chunk; TailNoLabel; eapply transl_addressing_label; eauto; unfold nolabel; auto.
+Qed.
+
+Remark transl_store_label:
+ forall chunk addr args src k c,
+ transl_store chunk addr args src k = OK c -> tail_nolabel k c.
+Proof.
+ unfold transl_store; intros; destruct chunk; TailNoLabel; eapply transl_addressing_label; eauto; unfold nolabel; auto.
+Qed.
+
+Remark indexed_memory_access_label:
+ forall insn sz base ofs k,
+ (forall ad, nolabel (insn ad)) ->
+ tail_nolabel k (indexed_memory_access insn sz base ofs k).
+Proof.
+ unfold indexed_memory_access; intros. destruct offset_representable.
+ TailNoLabel.
+ eapply tail_nolabel_trans; TailNoLabel.
+Qed.
+
+Remark loadind_label:
+ forall base ofs ty dst k c,
+ loadind base ofs ty dst k = OK c -> tail_nolabel k c.
+Proof.
+ unfold loadind; intros.
+ destruct ty, (preg_of dst); inv H; apply indexed_memory_access_label; intros; exact I.
+Qed.
+
+Remark storeind_label:
+ forall src base ofs ty k c,
+ storeind src base ofs ty k = OK c -> tail_nolabel k c.
+Proof.
+ unfold storeind; intros.
+ destruct ty, (preg_of src); inv H; apply indexed_memory_access_label; intros; exact I.
+Qed.
+
+Remark loadptr_label:
+ forall base ofs dst k, tail_nolabel k (loadptr base ofs dst k).
+Proof.
+ intros. apply indexed_memory_access_label. unfold nolabel; auto.
+Qed.
+
+Remark storeptr_label:
+ forall src base ofs k, tail_nolabel k (storeptr src base ofs k).
+Proof.
+ intros. apply indexed_memory_access_label. unfold nolabel; auto.
+Qed.
+
+Remark make_epilogue_label:
+ forall f k, tail_nolabel k (make_epilogue f k).
+Proof.
+ unfold make_epilogue; intros.
+ (* FIXME destruct is_leaf_function.
+ { TailNoLabel. } *)
+ eapply tail_nolabel_trans.
+ apply loadptr_label.
+ TailNoLabel.
+Qed.
+
+Lemma transl_instr_label:
+ forall f i ep k c,
+ transl_instr f i ep k = OK c ->
+ match i with Mlabel lbl => c = Plabel lbl :: k | _ => tail_nolabel k c end.
+Proof.
+ unfold transl_instr; intros; destruct i; TailNoLabel.
+- eapply loadind_label; eauto.
+- eapply storeind_label; eauto.
+- destruct ep. eapply loadind_label; eauto.
+ eapply tail_nolabel_trans. apply loadptr_label. eapply loadind_label; eauto.
+- eapply transl_op_label; eauto.
+- eapply transl_load_label; eauto.
+- eapply transl_store_label; eauto.
+- destruct s0; monadInv H; TailNoLabel.
+- destruct s0; monadInv H; (eapply tail_nolabel_trans; [eapply make_epilogue_label|TailNoLabel]).
+- eapply transl_cond_branch_label; eauto.
+- eapply tail_nolabel_trans; [eapply make_epilogue_label|TailNoLabel].
+Qed.
+
+Lemma transl_instr_label':
+ forall lbl f i ep k c,
+ transl_instr f i ep k = OK c ->
+ find_label lbl c = if Mach.is_label lbl i then Some k else find_label lbl k.
+Proof.
+ intros. exploit transl_instr_label; eauto.
+ destruct i; try (intros [A B]; apply B).
+ intros. subst c. simpl. auto.
+Qed.
+
+Lemma transl_code_label:
+ forall lbl f c ep tc,
+ transl_code f c ep = OK tc ->
+ match Mach.find_label lbl c with
+ | None => find_label lbl tc = None
+ | Some c' => exists tc', find_label lbl tc = Some tc' /\ transl_code f c' false = OK tc'
+ end.
+Proof.
+ induction c; simpl; intros.
+ inv H. auto.
+ monadInv H. rewrite (transl_instr_label' lbl _ _ _ _ _ EQ0).
+ generalize (Mach.is_label_correct lbl a).
+ destruct (Mach.is_label lbl a); intros.
+ subst a. simpl in EQ. exists x; auto.
+ eapply IHc; eauto.
+Qed.
+
+Lemma transl_find_label:
+ forall lbl f tf,
+ transf_function f = OK tf ->
+ match Mach.find_label lbl f.(Mach.fn_code) with
+ | None => find_label lbl tf.(fn_code) = None
+ | Some c => exists tc, find_label lbl tf.(fn_code) = Some tc /\ transl_code f c false = OK tc
+ end.
+Proof.
+ intros. monadInv H. destruct (zlt Ptrofs.max_unsigned (list_length_z x.(fn_code))); inv EQ0.
+ monadInv EQ. rewrite transl_code'_transl_code in EQ0. unfold fn_code.
+ simpl. destruct (storeptr_label X30 XSP (fn_retaddr_ofs f) x) as [A B]; rewrite B.
+ eapply transl_code_label; eauto.
+Qed.
+
+End TRANSL_LABEL.
+
+(** A valid branch in a piece of Mach code translates to a valid ``go to''
+ transition in the generated Asm code. *)
+
+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. omega.
+ generalize (transf_function_no_overflow _ _ H0). omega.
+ intros. apply Pregmap.gso; auto.
+Qed.
+
+(** Existence of return addresses *)
+
+Lemma return_address_exists:
+ forall f sg ros c, is_tail (Mcall sg ros :: c) f.(Mach.fn_code) ->
+ exists ra, return_address_offset f c ra.
+Proof.
+ intros. eapply Asmgenproof0.return_address_exists; eauto.
+- intros. exploit transl_instr_label; eauto.
+ destruct i; try (intros [A B]; apply A). intros. subst c0. repeat constructor.
+- intros. monadInv H0.
+ destruct (zlt Ptrofs.max_unsigned (list_length_z x.(fn_code))); inv EQ0. monadInv EQ.
+ rewrite transl_code'_transl_code in EQ0.
+ exists x; exists true; split; auto. unfold fn_code.
+ constructor. apply (storeptr_label X30 XSP (fn_retaddr_ofs f0) x).
+- exact transf_function_no_overflow.
+Qed.
+
+(** * Proof of semantic preservation *)
+
+(** Semantic preservation is proved using simulation diagrams
+ of the following form.
+<<
+ st1 --------------- st2
+ | |
+ t| *|t
+ | |
+ v v
+ st1'--------------- st2'
+>>
+ The invariant is the [match_states] predicate below, which includes:
+- The Asm code pointed by the PC register is the translation of
+ the current Mach code sequence.
+- Mach register values and Asm register values agree.
+*)
+
+Inductive match_states: Mach.state -> Asm.state -> Prop :=
+ | match_states_intro:
+ forall s fb sp c ep ms m m' rs f tf tc
+ (STACKS: match_stack ge s)
+ (FIND: Genv.find_funct_ptr ge fb = Some (Internal f))
+ (MEXT: Mem.extends m m')
+ (AT: transl_code_at_pc ge (rs PC) fb f c ep tf tc)
+ (AG: agree ms sp rs)
+ (DXP: ep = true -> rs#X29 = parent_sp s)
+ (LEAF: is_leaf_function f = true -> rs#RA = parent_ra s),
+ match_states (Mach.State s fb sp c ms m)
+ (Asm.State rs m')
+ | match_states_call:
+ forall s fb ms m m' rs
+ (STACKS: match_stack ge s)
+ (MEXT: Mem.extends m m')
+ (AG: agree ms (parent_sp s) rs)
+ (ATPC: rs PC = Vptr fb Ptrofs.zero)
+ (ATLR: rs RA = parent_ra s),
+ match_states (Mach.Callstate s fb ms m)
+ (Asm.State rs m')
+ | match_states_return:
+ forall s ms m m' rs
+ (STACKS: match_stack ge s)
+ (MEXT: Mem.extends m m')
+ (AG: agree ms (parent_sp s) rs)
+ (ATPC: rs PC = parent_ra s),
+ match_states (Mach.Returnstate s ms m)
+ (Asm.State rs m').
+
+Lemma exec_straight_steps:
+ forall s fb f rs1 i c ep tf tc m1' m2 m2' sp ms2,
+ match_stack ge s ->
+ Mem.extends m2 m2' ->
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc ->
+ (forall k c (TR: transl_instr f i ep k = OK c),
+ exists rs2,
+ exec_straight tge tf c rs1 m1' k rs2 m2'
+ /\ agree ms2 sp rs2
+ /\ (it1_is_parent ep i = true -> rs2#X29 = parent_sp s)
+ /\ (is_leaf_function f = true -> rs2#RA = parent_ra s)) ->
+ exists st',
+ plus step tge (State rs1 m1') E0 st' /\
+ match_states (Mach.State s fb sp c ms2 m2) st'.
+Proof.
+ intros. inversion H2. subst. monadInv H7.
+ exploit H3; eauto. intros [rs2 [A [B [C D]]]].
+ exists (State rs2 m2'); split.
+ - eapply exec_straight_exec; eauto.
+ - econstructor; eauto. eapply exec_straight_at; eauto.
+Qed.
+
+Lemma exec_straight_steps_goto:
+ forall s fb f rs1 i c ep tf tc m1' m2 m2' sp ms2 lbl c',
+ match_stack ge s ->
+ Mem.extends m2 m2' ->
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ Mach.find_label lbl f.(Mach.fn_code) = Some c' ->
+ transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc ->
+ it1_is_parent ep i = false ->
+ (forall k c (TR: transl_instr f i ep k = OK c),
+ exists jmp, exists k', exists rs2,
+ exec_straight tge tf c rs1 m1' (jmp :: k') rs2 m2'
+ /\ agree ms2 sp rs2
+ /\ exec_instr tge tf jmp rs2 m2' = goto_label tf lbl rs2 m2'
+ /\ (is_leaf_function f = true -> rs2#RA = parent_ra s)) ->
+ exists st',
+ plus step tge (State rs1 m1') E0 st' /\
+ match_states (Mach.State s fb sp c' ms2 m2) st'.
+Proof.
+ intros. inversion H3. subst. monadInv H9.
+ exploit H5; eauto. intros [jmp [k' [rs2 [A [B [C D]]]]]].
+ generalize (functions_transl _ _ _ H7 H8); intro FN.
+ generalize (transf_function_no_overflow _ _ H8); intro NOOV.
+ exploit exec_straight_steps_2; eauto.
+ intros [ofs' [PC2 CT2]].
+ exploit find_label_goto_label; eauto.
+ intros [tc' [rs3 [GOTO [AT' OTH]]]].
+ exists (State rs3 m2'); split.
+ eapply plus_right'.
+ eapply exec_straight_steps_1; eauto.
+ econstructor; eauto.
+ eapply find_instr_tail. eauto.
+ rewrite C. eexact GOTO.
+ traceEq.
+ econstructor; eauto.
+ apply agree_exten with rs2; auto with asmgen.
+ congruence.
+ rewrite OTH by congruence; auto.
+Qed.
+
+Lemma exec_straight_opt_steps_goto:
+ forall s fb f rs1 i c ep tf tc m1' m2 m2' sp ms2 lbl c',
+ match_stack ge s ->
+ Mem.extends m2 m2' ->
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ Mach.find_label lbl f.(Mach.fn_code) = Some c' ->
+ transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc ->
+ it1_is_parent ep i = false ->
+ (forall k c (TR: transl_instr f i ep k = OK c),
+ exists jmp, exists k', exists rs2,
+ exec_straight_opt tge tf c rs1 m1' (jmp :: k') rs2 m2'
+ /\ agree ms2 sp rs2
+ /\ exec_instr tge tf jmp rs2 m2' = goto_label tf lbl rs2 m2'
+ /\ (is_leaf_function f = true -> rs2#RA = parent_ra s)) ->
+ exists st',
+ plus step tge (State rs1 m1') E0 st' /\
+ match_states (Mach.State s fb sp c' ms2 m2) st'.
+Proof.
+ intros. inversion H3. subst. monadInv H9.
+ exploit H5; eauto. intros [jmp [k' [rs2 [A [B [C D]]]]]].
+ generalize (functions_transl _ _ _ H7 H8); intro FN.
+ generalize (transf_function_no_overflow _ _ H8); intro NOOV.
+ inv A.
+- exploit find_label_goto_label; eauto.
+ intros [tc' [rs3 [GOTO [AT' OTH]]]].
+ exists (State rs3 m2'); split.
+ apply plus_one. econstructor; eauto.
+ eapply find_instr_tail. eauto.
+ rewrite C. eexact GOTO.
+ econstructor; eauto.
+ apply agree_exten with rs2; auto with asmgen.
+ congruence.
+ rewrite OTH by congruence; auto.
+- exploit exec_straight_steps_2; eauto.
+ intros [ofs' [PC2 CT2]].
+ exploit find_label_goto_label; eauto.
+ intros [tc' [rs3 [GOTO [AT' OTH]]]].
+ exists (State rs3 m2'); split.
+ eapply plus_right'.
+ eapply exec_straight_steps_1; eauto.
+ econstructor; eauto.
+ eapply find_instr_tail. eauto.
+ rewrite C. eexact GOTO.
+ traceEq.
+ econstructor; eauto.
+ apply agree_exten with rs2; auto with asmgen.
+ congruence.
+ rewrite OTH by congruence; auto.
+Qed.
+
+(** We need to show that, in the simulation diagram, we cannot
+ take infinitely many Mach transitions that correspond to zero
+ transitions on the Asm side. Actually, all Mach transitions
+ correspond to at least one Asm transition, except the
+ transition from [Machsem.Returnstate] to [Machsem.State].
+ So, the following integer measure will suffice to rule out
+ the unwanted behaviour. *)
+
+Definition measure (s: Mach.state) : nat :=
+ match s with
+ | Mach.State _ _ _ _ _ _ => 0%nat
+ | Mach.Callstate _ _ _ _ => 0%nat
+ | Mach.Returnstate _ _ _ => 1%nat
+ end.
+
+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') (WF: wf_state ge 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. }
+ split. { simpl; congruence. }
+ rewrite nextinstr_inv by congruence; assumption.
+
+- (* 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 S]]]].
+ exists rs'; split. eauto.
+ split. { eapply agree_set_mreg; eauto with asmgen. congruence. }
+ split. { simpl; congruence. }
+ rewrite S. assumption.
+
+- (* 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 R]]].
+ exists rs'; split. eauto.
+ split. eapply agree_undef_regs; eauto with asmgen.
+ simpl; intros.
+ split. rewrite Q; auto with asmgen.
+ rewrite R. assumption.
+
+- (* 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 S]]]].
+ exists rs1; split. eauto.
+ split. eapply agree_set_mreg. eapply agree_set_mreg; eauto. congruence. auto with asmgen.
+ simpl; split; intros.
+ { rewrite R; auto with asmgen.
+ apply preg_of_not_X29; auto.
+ }
+ { rewrite S; 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 V]]]].
+ 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.
+ split; simpl; intros. rewrite U; auto with asmgen.
+ apply preg_of_not_X29; auto.
+ rewrite V. rewrite R by congruence. 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 S]]]].
+ exists rs2; split. eauto. split.
+ apply agree_set_undef_mreg with rs0; auto.
+ apply Val.lessdef_trans with v'; auto.
+ split; simpl; intros. InvBooleans.
+ rewrite R; auto. apply preg_of_not_X29; auto.
+Local Transparent destroyed_by_op.
+ destruct op; try exact I; simpl; congruence.
+ rewrite S.
+ auto.
+- (* Mload *)
+ destruct trap.
+ {
+ 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 S]]]].
+ exists rs2; split. eauto.
+ split. eapply agree_set_undef_mreg; eauto. congruence.
+ split. simpl; congruence.
+ rewrite S. assumption.
+ }
+
+ (* Mload notrap1 *)
+ inv AT. simpl in *. unfold bind in *. destruct (transl_code _ _ _) in *; discriminate.
+
+- (* Mload notrap *)
+ inv AT. simpl in *. unfold bind in *. destruct (transl_code _ _ _) in *; discriminate.
+
+- (* Mload notrap *)
+ inv AT. simpl in *. unfold bind in *. destruct (transl_code _ _ _) in *; discriminate.
+
+- (* 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 R]]].
+ exists rs2; split. eauto.
+ split. eapply agree_undef_regs; eauto with asmgen.
+ split. simpl; congruence.
+ rewrite R. assumption.
+
+- (* 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_2.
+ rewrite <- H1. simpl. econstructor; eauto.
+ eapply code_tail_next_int; eauto.
+ rewrite preg_notin_charact. intros. auto with asmgen.
+ auto with asmgen.
+ apply agree_nextinstr. eapply agree_set_res; auto.
+ eapply agree_undef_regs; eauto. intros. rewrite undef_regs_other_2; auto.
+ congruence.
+
+ Simpl.
+ rewrite set_res_other by trivial.
+ rewrite undef_regs_other.
+ assumption.
+ intro.
+ rewrite in_map_iff.
+ intros (x0 & PREG & IN).
+ subst r'.
+ intro.
+ apply (preg_of_not_RA x0).
+ 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.
+
+ rewrite INV by congruence.
+ assumption.
+
+- (* 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 & D).
+ exists jmp; exists k; exists rs'.
+ split. eexact A.
+ split. apply agree_exten with rs0; auto with asmgen.
+ split.
+ exact B.
+ rewrite D. exact LEAF.
+
+- (* 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 & D).
+ 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.
+ split.
+ simpl; congruence.
+ Simpl. rewrite D.
+ exact LEAF.
+
+- (* 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 #X17 <- 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.
+
+ rewrite C by congruence.
+ repeat rewrite Pregmap.gso by congruence.
+ assumption.
+
+- (* 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 & W).
+ 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. omega. constructor.
+ intros (ofs' & X & Y).
+ left; exists (State rs3 m3'); split.
+ eapply exec_straight_steps_1; eauto. omega. 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.
+
+ rewrite W.
+ unfold rs2.
+ Simpl.
+
+- (* 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. omega. split. auto.
+ rewrite <- ATPC in H5.
+ econstructor; eauto. congruence.
+ inv WF.
+ inv STACK.
+ inv H1.
+ congruence.
+Qed.
+
+Lemma transf_initial_states:
+ forall st1, Mach.initial_state prog st1 ->
+ exists st2, Asm.initial_state tprog st2 /\ match_states st1 st2.
+Proof.
+ intros. inversion H. unfold ge0 in *.
+ econstructor; split.
+ econstructor.
+ eapply (Genv.init_mem_transf_partial TRANSF); eauto.
+ replace (Genv.symbol_address (Genv.globalenv tprog) (prog_main tprog) Ptrofs.zero)
+ with (Vptr fb Ptrofs.zero).
+ econstructor; eauto.
+ constructor.
+ apply Mem.extends_refl.
+ split. auto. simpl. unfold Vnullptr; destruct Archi.ptr64; congruence.
+ intros. rewrite Regmap.gi. auto.
+ unfold Genv.symbol_address.
+ rewrite (match_program_main TRANSF).
+ rewrite symbols_preserved.
+ unfold ge; rewrite H1. auto.
+Qed.
+
+Lemma transf_final_states:
+ forall st1 st2 r,
+ match_states st1 st2 -> Mach.final_state st1 r -> Asm.final_state st2 r.
+Proof.
+ intros. inv H0. inv H. constructor. assumption.
+ compute in H1. inv H1.
+ generalize (preg_val _ _ _ R0 AG). rewrite H2. intros LD; inv LD. auto.
+Qed.
+
+Theorem transf_program_correct:
+ forward_simulation (Mach.semantics return_address_offset prog) (Asm.semantics tprog).
+Proof.
+ eapply forward_simulation_star with (measure := measure)
+ (match_states := fun S1 S2 => match_states S1 S2 /\ wf_state ge S1).
+ - apply senv_preserved.
+ - simpl; intros. exploit transf_initial_states; eauto.
+ intros (s2 & A & B).
+ exists s2; intuition auto. apply wf_initial; auto.
+ - simpl; intros. destruct H as [MS WF]. eapply transf_final_states; eauto.
+ - simpl; intros. destruct H0 as [MS WF].
+ exploit step_simulation; eauto. intros [ (s2' & A & B) | (A & B & C) ].
+ + left; exists s2'; intuition auto. eapply wf_step; eauto.
+ + right; intuition auto. eapply wf_step; eauto.
+Qed.
+
+End PRESERVATION.
+*) \ No newline at end of file
diff --git a/aarch64/Asmgenproof1.v b/aarch64/Asmblockgenproof1.v
index 0e36bd05..b42309bc 100644
--- a/aarch64/Asmgenproof1.v
+++ b/aarch64/Asmblockgenproof1.v
@@ -1,3 +1,5 @@
+(* ORIGINAL aarch64/Asmgenproof1 file that needs to be adapted
+
(* *********************************************************************)
(* *)
(* The Compcert verified compiler *)
@@ -2136,3 +2138,4 @@ Proof.
Qed.
End CONSTRUCTORS.
+*) \ No newline at end of file
diff --git a/aarch64/Asmgen.v b/aarch64/Asmgen.v
index 024c9a17..4e81a936 100644
--- a/aarch64/Asmgen.v
+++ b/aarch64/Asmgen.v
@@ -1,1172 +1,362 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Translation from Mach to AArch64. *)
+(* *************************************************************)
+(* *)
+(* 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 *)
+(* *)
+(* 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.
-Require Import Locations Mach Asm.
-
-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).
-
-(** Arithmetic immediates are 12-bit unsigned numbers, possibly shifted left 12 bits *)
-
-Definition is_arith_imm32 (x: int) : bool :=
- Int.eq x (Int.zero_ext 12 x)
- || Int.eq x (Int.shl (Int.zero_ext 12 (Int.shru x (Int.repr 12))) (Int.repr 12)).
-
-Definition is_arith_imm64 (x: int64) : bool :=
- Int64.eq x (Int64.zero_ext 12 x)
- || Int64.eq x (Int64.shl (Int64.zero_ext 12 (Int64.shru x (Int64.repr 12))) (Int64.repr 12)).
-
-(** Decompose integer literals into 16-bit fragments *)
-
-Fixpoint decompose_int (N: nat) (n p: Z) {struct N} : list (Z * Z) :=
- match N with
- | Datatypes.O => nil
- | Datatypes.S N =>
- let frag := Zzero_ext 16 (Z.shiftr n p) in
- if Z.eqb frag 0 then
- decompose_int N n (p + 16)
- else
- (frag, p) :: decompose_int N (Z.ldiff n (Z.shiftl 65535 p)) (p + 16)
- end.
-
-Definition negate_decomposition (l: list (Z * Z)) :=
- List.map (fun np => (Z.lxor (fst np) 65535, snd np)) l.
-
-Definition loadimm_k (sz: isize) (rd: ireg) (l: list (Z * Z)) (k: code) : code :=
- List.fold_right (fun np k => 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 => Pmovz sz rd 0 0 :: k
- | (n1, p1) :: l => 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 => Pmovn sz rd 0 0 :: k
- | (n1, p1) :: l => 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 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 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 addimm32 (rd r1: ireg) (n: int) (k: code) : code :=
- let m := Int.neg n in
- if Int.eq n (Int.zero_ext 24 n) then
- addimm_aux (Paddimm W) rd r1 (Int.unsigned n) k
- else if Int.eq m (Int.zero_ext 24 m) then
- addimm_aux (Psubimm W) rd r1 (Int.unsigned m) k
- else if Int.lt n Int.zero then
- loadimm32 X16 m (Psub W rd r1 X16 SOnone :: k)
- else
- loadimm32 X16 n (Padd W rd r1 X16 SOnone :: 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 (Paddimm X) rd r1 (Int64.unsigned n) k
- else if Int64.eq m (Int64.zero_ext 24 m) then
- addimm_aux (Psubimm X) rd r1 (Int64.unsigned m) k
- else if Int64.lt n Int64.zero then
- loadimm64 X16 m (Psubext rd r1 X16 (EOuxtx Int.zero) :: k)
- else
- loadimm64 X16 n (Paddext rd r1 X16 (EOuxtx Int.zero) :: k).
+Require Import Locations Compopts.
+Require Import Mach Asm Asmblock Asmblockgen Machblockgen PseudoAsmblock PseudoAsmblockproof.
-(** 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 if Int.eq n Int.one then
- Padd W X16 r1 r1 (SOlsr (Int.repr 31)) ::
- Porr W rd XZR X16 (SOasr n) :: 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 if Int.eq n Int.one then
- Padd X X16 r1 r1 (SOlsr (Int.repr 63)) ::
- Porr X rd XZR X16 (SOasr n) :: 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 Archi.pic_code tt then
- if Ptrofs.eq ofs Ptrofs.zero then
- Ploadsymbol rd id :: k
- else
- Ploadsymbol rd id :: addimm64 rd rd (Ptrofs.to_int64 ofs) k
- else
- Padrp rd id ofs :: Paddadr rd rd id ofs :: k.
-
-(** Translate a shifted operand *)
-
-Definition transl_shift (s: Op.shift) (a: int): Asm.shift_op :=
- match s with
- | Slsl => SOlsl a
- | Slsr => SOlsr a
- | Sasr => SOasr a
- | Sror => SOror a
- end.
-
-(** Translation of a condition. Prepends to [k] the instructions
- that evaluate the condition and leave its boolean result in one of
- the bits of the condition register. The bit in question is
- determined by the [crbit_for_cond] function. *)
-
-Definition transl_cond
- (cond: condition) (args: list mreg) (k: code) :=
- match cond, args with
- | (Ccomp c | Ccompu c), a1 :: a2 :: nil =>
- do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pcmp W r1 r2 SOnone :: k)
- | (Ccompshift c s a | Ccompushift c s a), a1 :: a2 :: nil =>
- do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pcmp W r1 r2 (transl_shift s a) :: k)
- | (Ccompimm c n | Ccompuimm c n), a1 :: nil =>
- do r1 <- ireg_of a1;
- OK (if is_arith_imm32 n then
- Pcmpimm W r1 (Int.unsigned n) :: k
- else if is_arith_imm32 (Int.neg n) then
- Pcmnimm W r1 (Int.unsigned (Int.neg n)) :: k
- else
- loadimm32 X16 n (Pcmp W r1 X16 SOnone :: k))
- | (Cmaskzero n | Cmasknotzero n), a1 :: nil =>
- do r1 <- ireg_of a1;
- OK (if is_logical_imm32 n then
- Ptstimm W r1 (Int.unsigned n) :: k
- else
- loadimm32 X16 n (Ptst W r1 X16 SOnone :: k))
- | (Ccompl c | Ccomplu c), a1 :: a2 :: nil =>
- do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pcmp X r1 r2 SOnone :: k)
- | (Ccomplshift c s a | Ccomplushift c s a), a1 :: a2 :: nil =>
- do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pcmp X r1 r2 (transl_shift s a) :: k)
- | (Ccomplimm c n | Ccompluimm c n), a1 :: nil =>
- do r1 <- ireg_of a1;
- OK (if is_arith_imm64 n then
- Pcmpimm X r1 (Int64.unsigned n) :: k
- else if is_arith_imm64 (Int64.neg n) then
- Pcmnimm X r1 (Int64.unsigned (Int64.neg n)) :: k
- else
- loadimm64 X16 n (Pcmp X r1 X16 SOnone :: k))
- | (Cmasklzero n | Cmasklnotzero n), a1 :: nil =>
- do r1 <- ireg_of a1;
- OK (if is_logical_imm64 n then
- Ptstimm X r1 (Int64.unsigned n) :: k
- else
- loadimm64 X16 n (Ptst X r1 X16 SOnone :: k))
- | Ccompf cmp, a1 :: a2 :: nil =>
- do r1 <- freg_of a1; do r2 <- freg_of a2;
- OK (Pfcmp D r1 r2 :: k)
- | Cnotcompf cmp, a1 :: a2 :: nil =>
- do r1 <- freg_of a1; do r2 <- freg_of a2;
- OK (Pfcmp D r1 r2 :: k)
- | Ccompfzero cmp, a1 :: nil =>
- do r1 <- freg_of a1;
- OK (Pfcmp0 D r1 :: k)
- | Cnotcompfzero cmp, a1 :: nil =>
- do r1 <- freg_of a1;
- OK (Pfcmp0 D r1 :: k)
- | Ccompfs cmp, a1 :: a2 :: nil =>
- do r1 <- freg_of a1; do r2 <- freg_of a2;
- OK (Pfcmp S r1 r2 :: k)
- | Cnotcompfs cmp, a1 :: a2 :: nil =>
- do r1 <- freg_of a1; do r2 <- freg_of a2;
- OK (Pfcmp S r1 r2 :: k)
- | Ccompfszero cmp, a1 :: nil =>
- do r1 <- freg_of a1;
- OK (Pfcmp0 S r1 :: k)
- | Cnotcompfszero cmp, a1 :: nil =>
- do r1 <- freg_of a1;
- OK (Pfcmp0 S r1 :: k)
- | _, _ =>
- Error(msg "Asmgen.transl_cond")
- end.
+Local Open Scope error_monad_scope.
-Definition cond_for_signed_cmp (cmp: comparison) :=
- match cmp with
- | Ceq => TCeq
- | Cne => TCne
- | Clt => TClt
- | Cle => TCle
- | Cgt => TCgt
- | Cge => TCge
- end.
+(** * Translation from Asmblock to assembly language
+ Inspired from the KVX backend (see kvx/Asm.v and kvx/Asmgen.v) *)
-Definition cond_for_unsigned_cmp (cmp: comparison) :=
- match cmp with
- | Ceq => TCeq
- | Cne => TCne
- | Clt => TClo
- | Cle => TCls
- | Cgt => TChi
- | Cge => TChs
- end.
+Module Asmblock_TRANSF.
+(* STUB *)
-Definition cond_for_float_cmp (cmp: comparison) :=
- match cmp with
- | Ceq => TCeq
- | Cne => TCne
- | Clt => TCmi
- | Cle => TCls
- | Cgt => TCgt
- | Cge => TCge
+Definition ireg_of_preg (p : Asmblock.preg) : res ireg :=
+ match p with
+ | DR (IR' (RR1 r)) => OK r
+ | _ => Error (msg "Asmgen.ireg_of_preg")
end.
-Definition cond_for_float_not_cmp (cmp: comparison) :=
- match cmp with
- | Ceq => TCne
- | Cne => TCeq
- | Clt => TCpl
- | Cle => TChi
- | Cgt => TCle
- | Cge => TClt
+Definition freg_of_preg (p : Asmblock.preg) : res freg :=
+ match p with
+ | DR (FR' r) => OK r
+ | _ => Error (msg "Asmgen.freg_of_preg")
end.
-Definition cond_for_cond (cond: condition) :=
- match cond with
- | Ccomp cmp => cond_for_signed_cmp cmp
- | Ccompu cmp => cond_for_unsigned_cmp cmp
- | Ccompshift cmp s a => cond_for_signed_cmp cmp
- | Ccompushift cmp s a => cond_for_unsigned_cmp cmp
- | Ccompimm cmp n => cond_for_signed_cmp cmp
- | Ccompuimm cmp n => cond_for_unsigned_cmp cmp
- | Cmaskzero n => TCeq
- | Cmasknotzero n => TCne
- | Ccompl cmp => cond_for_signed_cmp cmp
- | Ccomplu cmp => cond_for_unsigned_cmp cmp
- | Ccomplshift cmp s a => cond_for_signed_cmp cmp
- | Ccomplushift cmp s a => cond_for_unsigned_cmp cmp
- | Ccomplimm cmp n => cond_for_signed_cmp cmp
- | Ccompluimm cmp n => cond_for_unsigned_cmp cmp
- | Cmasklzero n => TCeq
- | Cmasklnotzero n => TCne
- | Ccompf cmp => cond_for_float_cmp cmp
- | Cnotcompf cmp => cond_for_float_not_cmp cmp
- | Ccompfzero cmp => cond_for_float_cmp cmp
- | Cnotcompfzero cmp => cond_for_float_not_cmp cmp
- | Ccompfs cmp => cond_for_float_cmp cmp
- | Cnotcompfs cmp => cond_for_float_not_cmp cmp
- | Ccompfszero cmp => cond_for_float_cmp cmp
- | Cnotcompfszero cmp => cond_for_float_not_cmp cmp
+Definition iregsp_of_preg (p : Asmblock.preg) : res iregsp :=
+ match p with
+ | DR (IR' r) => OK r
+ | _ => Error (msg "Asmgen.iregsp_of_preg")
end.
-(** Translation of a conditional branch. Prepends to [k] the instructions
- that evaluate the condition and ranch to [lbl] if it holds.
- We recognize some conditional branches that can be implemented
- without setting then testing condition flags. *)
-
-Definition transl_cond_branch_default
- (c: condition) (args: list mreg) (lbl: label) (k: code) :=
- transl_cond c args (Pbc (cond_for_cond c) lbl :: k).
-
-Definition transl_cond_branch
- (c: condition) (args: list mreg) (lbl: label) (k: code) :=
- match args, c with
- | a1 :: nil, (Ccompimm Cne n | Ccompuimm Cne n) =>
- if Int.eq n Int.zero
- then (do r1 <- ireg_of a1; OK (Pcbnz W r1 lbl :: k))
- else transl_cond_branch_default c args lbl k
- | a1 :: nil, (Ccompimm Ceq n | Ccompuimm Ceq n) =>
- if Int.eq n Int.zero
- then (do r1 <- ireg_of a1; OK (Pcbz W r1 lbl :: k))
- else transl_cond_branch_default c args lbl k
- | a1 :: nil, (Ccomplimm Cne n | Ccompluimm Cne n) =>
- if Int64.eq n Int64.zero
- then (do r1 <- ireg_of a1; OK (Pcbnz X r1 lbl :: k))
- else transl_cond_branch_default c args lbl k
- | a1 :: nil, (Ccomplimm Ceq n | Ccompluimm Ceq n) =>
- if Int64.eq n Int64.zero
- then (do r1 <- ireg_of a1; OK (Pcbz X r1 lbl :: k))
- else transl_cond_branch_default c args lbl k
- | a1 :: nil, Cmaskzero n =>
- match Int.is_power2 n with
- | Some bit => do r1 <- ireg_of a1; OK (Ptbz W r1 bit lbl :: k)
- | None => transl_cond_branch_default c args lbl k
- end
- | a1 :: nil, Cmasknotzero n =>
- match Int.is_power2 n with
- | Some bit => do r1 <- ireg_of a1; OK (Ptbnz W r1 bit lbl :: k)
- | None => transl_cond_branch_default c args lbl k
- end
- | a1 :: nil, Cmasklzero n =>
- match Int64.is_power2' n with
- | Some bit => do r1 <- ireg_of a1; OK (Ptbz X r1 bit lbl :: k)
- | None => transl_cond_branch_default c args lbl k
+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 (Asmblock.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' <- ireg_of_preg rd;
+ do rs' <- ireg_of_preg rs;
+ OK (Asm.Paddimm sz rd' rs' n)
+ | PArith (PArithPP (Psubimm sz n) rd rs) => do rd' <- ireg_of_preg rd;
+ do rs' <- ireg_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' <- ireg_of_preg rd;
+ do r1' <- ireg_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' <- ireg_of_preg rd;
+ do r1' <- ireg_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', IR' 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
- | a1 :: nil, Cmasklnotzero n =>
- match Int64.is_power2' n with
- | Some bit => do r1 <- ireg_of a1; OK (Ptbnz X r1 bit lbl :: k)
- | None => transl_cond_branch_default c args lbl k
- end
- | _, _ =>
- transl_cond_branch_default c args lbl k
- end.
-
-(** Translation of the arithmetic operation [res <- op(args)].
- The corresponding instructions are prepended to [k]. *)
+ | PArith (Pfnmul fsz rd r1 r2) => OK (Asm.Pfnmul fsz rd r1 r2)
-Definition transl_op
- (op: operation) (args: list mreg) (res: mreg) (k: code) :=
- match op, args with
- | Omove, a1 :: nil =>
- match preg_of res, preg_of a1 with
- | IR r, IR a => OK (Pmov r a :: k)
- | FR r, FR a => OK (Pfmov r a :: k)
- | _ , _ => Error(msg "Asmgen.Omove")
- end
- | Ointconst n, nil =>
- do rd <- ireg_of res;
- OK (loadimm32 rd n k)
- | Olongconst n, nil =>
- do rd <- ireg_of res;
- OK (loadimm64 rd n k)
- | Ofloatconst f, nil =>
- do rd <- freg_of res;
- OK (if Float.eq_dec f Float.zero
- then Pfmovi D rd XZR :: k
- else Pfmovimmd rd f :: k)
- | Osingleconst f, nil =>
- do rd <- freg_of res;
- OK (if Float32.eq_dec f Float32.zero
- then Pfmovi S rd XZR :: k
- else Pfmovimms rd f :: k)
- | Oaddrsymbol id ofs, nil =>
- do rd <- ireg_of res;
- OK (loadsymbol rd id ofs k)
- | Oaddrstack ofs, nil =>
- do rd <- ireg_of res;
- OK (addimm64 rd XSP (Ptrofs.to_int64 ofs) k)
-(** 32-bit integer arithmetic *)
- | Oshift s a, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Porr W rd XZR r1 (transl_shift s a) :: k)
- | Oadd, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Padd W rd r1 r2 SOnone :: k)
- | Oaddshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Padd W rd r1 r2 (transl_shift s a) :: k)
- | Oaddimm n, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (addimm32 rd r1 n k)
- | Oneg, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Psub W rd XZR r1 SOnone :: k)
- | Onegshift s a, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Psub W rd XZR r1 (transl_shift s a) :: k)
- | Osub, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Psub W rd r1 r2 SOnone :: k)
- | Osubshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Psub W rd r1 r2 (transl_shift s a) :: k)
- | Omul, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pmadd W rd r1 r2 XZR :: k)
- | Omuladd, a1 :: a2 :: a3 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r3 <- ireg_of a3;
- OK (Pmadd W rd r2 r3 r1 :: k)
- | Omulsub, a1 :: a2 :: a3 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r3 <- ireg_of a3;
- OK (Pmsub W rd r2 r3 r1 :: k)
- | Odiv, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Psdiv W rd r1 r2 :: k)
- | Odivu, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pudiv W rd r1 r2 :: k)
- | Oand, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pand W rd r1 r2 SOnone :: k)
- | Oandshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pand W rd r1 r2 (transl_shift s a) :: k)
- | Oandimm n, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (logicalimm32 (Pandimm W) (Pand W) rd r1 n k)
- | Oor, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Porr W rd r1 r2 SOnone :: k)
- | Oorshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Porr W rd r1 r2 (transl_shift s a) :: k)
- | Oorimm n, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (logicalimm32 (Porrimm W) (Porr W) rd r1 n k)
- | Oxor, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Peor W rd r1 r2 SOnone :: k)
- | Oxorshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Peor W rd r1 r2 (transl_shift s a) :: k)
- | Oxorimm n, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (logicalimm32 (Peorimm W) (Peor W) rd r1 n k)
- | Onot, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Porn W rd XZR r1 SOnone :: k)
- | Onotshift s a, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Porn W rd XZR r1 (transl_shift s a) :: k)
- | Obic, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pbic W rd r1 r2 SOnone :: k)
- | Obicshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pbic W rd r1 r2 (transl_shift s a) :: k)
- | Oorn, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Porn W rd r1 r2 SOnone :: k)
- | Oornshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Porn W rd r1 r2 (transl_shift s a) :: k)
- | Oeqv, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Peon W rd r1 r2 SOnone :: k)
- | Oeqvshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Peon W rd r1 r2 (transl_shift s a) :: k)
- | Oshl, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Plslv W rd r1 r2 :: k)
- | Oshr, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pasrv W rd r1 r2 :: k)
- | Oshru, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Plsrv W rd r1 r2 :: k)
- | Oshrximm n, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (shrx32 rd r1 n k)
- | Ozext s, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Pubfiz W rd r1 Int.zero s :: k)
- | Osext s, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Psbfiz W rd r1 Int.zero s :: k)
- | Oshlzext s a, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Pubfiz W rd r1 a (Z.min s (Int.zwordsize - Int.unsigned a)) :: k)
- | Oshlsext s a, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Psbfiz W rd r1 a (Z.min s (Int.zwordsize - Int.unsigned a)) :: k)
- | Ozextshr a s, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Pubfx W rd r1 a (Z.min s (Int.zwordsize - Int.unsigned a)) :: k)
- | Osextshr a s, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Psbfx W rd r1 a (Z.min s (Int.zwordsize - Int.unsigned a)) :: k)
-(** 64-bit integer arithmetic *)
- | Oshiftl s a, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Porr X rd XZR r1 (transl_shift s a) :: k)
- | Oextend x a, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (move_extended rd r1 x a k)
- (* [Omakelong] and [Ohighlong] should not occur *)
- | Olowlong, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- assertion (ireg_eq rd r1);
- OK (Pcvtx2w rd :: k)
- | Oaddl, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Padd X rd r1 r2 SOnone :: k)
- | Oaddlshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Padd X rd r1 r2 (transl_shift s a) :: k)
- | Oaddlext x a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (arith_extended Paddext (Padd X) rd r1 r2 x a k)
- | Oaddlimm n, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (addimm64 rd r1 n k)
- | Onegl, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Psub X rd XZR r1 SOnone :: k)
- | Oneglshift s a, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Psub X rd XZR r1 (transl_shift s a) :: k)
- | Osubl, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Psub X rd r1 r2 SOnone :: k)
- | Osublshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Psub X rd r1 r2 (transl_shift s a) :: k)
- | Osublext x a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (arith_extended Psubext (Psub X) rd r1 r2 x a k)
- | Omull, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pmadd X rd r1 r2 XZR :: k)
- | Omulladd, a1 :: a2 :: a3 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r3 <- ireg_of a3;
- OK (Pmadd X rd r2 r3 r1 :: k)
- | Omullsub, a1 :: a2 :: a3 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r3 <- ireg_of a3;
- OK (Pmsub X rd r2 r3 r1 :: k)
- | Omullhs, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Psmulh rd r1 r2 :: k)
- | Omullhu, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pumulh rd r1 r2 :: k)
- | Odivl, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Psdiv X rd r1 r2 :: k)
- | Odivlu, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pudiv X rd r1 r2 :: k)
- | Oandl, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pand X rd r1 r2 SOnone :: k)
- | Oandlshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pand X rd r1 r2 (transl_shift s a) :: k)
- | Oandlimm n, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (logicalimm64 (Pandimm X) (Pand X) rd r1 n k)
- | Oorl, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Porr X rd r1 r2 SOnone :: k)
- | Oorlshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Porr X rd r1 r2 (transl_shift s a) :: k)
- | Oorlimm n, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (logicalimm64 (Porrimm X) (Porr X) rd r1 n k)
- | Oxorl, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Peor X rd r1 r2 SOnone :: k)
- | Oxorlshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Peor X rd r1 r2 (transl_shift s a) :: k)
- | Oxorlimm n, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (logicalimm64 (Peorimm X) (Peor X) rd r1 n k)
- | Onotl, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Porn X rd XZR r1 SOnone :: k)
- | Onotlshift s a, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Porn X rd XZR r1 (transl_shift s a) :: k)
- | Obicl, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pbic X rd r1 r2 SOnone :: k)
- | Obiclshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pbic X rd r1 r2 (transl_shift s a) :: k)
- | Oornl, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Porn X rd r1 r2 SOnone :: k)
- | Oornlshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Porn X rd r1 r2 (transl_shift s a) :: k)
- | Oeqvl, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Peon X rd r1 r2 SOnone :: k)
- | Oeqvlshift s a, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Peon X rd r1 r2 (transl_shift s a) :: k)
- | Oshll, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Plslv X rd r1 r2 :: k)
- | Oshrl, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Pasrv X rd r1 r2 :: k)
- | Oshrlu, a1 :: a2 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (Plsrv X rd r1 r2 :: k)
- | Oshrlximm n, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (shrx64 rd r1 n k)
- | Ozextl s, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Pubfiz X rd r1 Int.zero s :: k)
- | Osextl s, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Psbfiz X rd r1 Int.zero s :: k)
- | Oshllzext s a, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Pubfiz X rd r1 a (Z.min s (Int64.zwordsize - Int.unsigned a)) :: k)
- | Oshllsext s a, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Psbfiz X rd r1 a (Z.min s (Int64.zwordsize - Int.unsigned a)) :: k)
- | Ozextshrl a s, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Pubfx X rd r1 a (Z.min s (Int64.zwordsize - Int.unsigned a)) :: k)
- | Osextshrl a s, a1 :: nil =>
- do rd <- ireg_of res; do r1 <- ireg_of a1;
- OK (Psbfx X rd r1 a (Z.min s (Int64.zwordsize - Int.unsigned a)) :: k)
-(** 64-bit floating-point arithmetic *)
- | Onegf, a1 :: nil =>
- do rd <- freg_of res; do rs <- freg_of a1;
- OK (Pfneg D rd rs :: k)
- | Oabsf, a1 :: nil =>
- do rd <- freg_of res; do rs <- freg_of a1;
- OK (Pfabs D rd rs :: k)
- | Oaddf, a1 :: a2 :: nil =>
- do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
- OK (Pfadd D rd rs1 rs2 :: k)
- | Osubf, a1 :: a2 :: nil =>
- do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
- OK (Pfsub D rd rs1 rs2 :: k)
- | Omulf, a1 :: a2 :: nil =>
- do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
- OK (Pfmul D rd rs1 rs2 :: k)
- | Odivf, a1 :: a2 :: nil =>
- do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
- OK (Pfdiv D rd rs1 rs2 :: k)
-(** 32-bit floating-point arithmetic *)
- | Onegfs, a1 :: nil =>
- do rd <- freg_of res; do rs <- freg_of a1;
- OK (Pfneg S rd rs :: k)
- | Oabsfs, a1 :: nil =>
- do rd <- freg_of res; do rs <- freg_of a1;
- OK (Pfabs S rd rs :: k)
- | Oaddfs, a1 :: a2 :: nil =>
- do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
- OK (Pfadd S rd rs1 rs2 :: k)
- | Osubfs, a1 :: a2 :: nil =>
- do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
- OK (Pfsub S rd rs1 rs2 :: k)
- | Omulfs, a1 :: a2 :: nil =>
- do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
- OK (Pfmul S rd rs1 rs2 :: k)
- | Odivfs, a1 :: a2 :: nil =>
- do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
- OK (Pfdiv S rd rs1 rs2 :: k)
- | Osingleoffloat, a1 :: nil =>
- do rd <- freg_of res; do rs <- freg_of a1;
- OK (Pfcvtsd rd rs :: k)
- | Ofloatofsingle, a1 :: nil =>
- do rd <- freg_of res; do rs <- freg_of a1;
- OK (Pfcvtds rd rs :: k)
-(** Conversions between int and float *)
- | Ointoffloat, a1 :: nil =>
- do rd <- ireg_of res; do rs <- freg_of a1;
- OK (Pfcvtzs W D rd rs :: k)
- | Ointuoffloat, a1 :: nil =>
- do rd <- ireg_of res; do rs <- freg_of a1;
- OK (Pfcvtzu W D rd rs :: k)
- | Ofloatofint, a1 :: nil =>
- do rd <- freg_of res; do rs <- ireg_of a1;
- OK (Pscvtf D W rd rs :: k)
- | Ofloatofintu, a1 :: nil =>
- do rd <- freg_of res; do rs <- ireg_of a1;
- OK (Pucvtf D W rd rs :: k)
- | Ointofsingle, a1 :: nil =>
- do rd <- ireg_of res; do rs <- freg_of a1;
- OK (Pfcvtzs W S rd rs :: k)
- | Ointuofsingle, a1 :: nil =>
- do rd <- ireg_of res; do rs <- freg_of a1;
- OK (Pfcvtzu W S rd rs :: k)
- | Osingleofint, a1 :: nil =>
- do rd <- freg_of res; do rs <- ireg_of a1;
- OK (Pscvtf S W rd rs :: k)
- | Osingleofintu, a1 :: nil =>
- do rd <- freg_of res; do rs <- ireg_of a1;
- OK (Pucvtf S W rd rs :: k)
- | Olongoffloat, a1 :: nil =>
- do rd <- ireg_of res; do rs <- freg_of a1;
- OK (Pfcvtzs X D rd rs :: k)
- | Olonguoffloat, a1 :: nil =>
- do rd <- ireg_of res; do rs <- freg_of a1;
- OK (Pfcvtzu X D rd rs :: k)
- | Ofloatoflong, a1 :: nil =>
- do rd <- freg_of res; do rs <- ireg_of a1;
- OK (Pscvtf D X rd rs :: k)
- | Ofloatoflongu, a1 :: nil =>
- do rd <- freg_of res; do rs <- ireg_of a1;
- OK (Pucvtf D X rd rs :: k)
- | Olongofsingle, a1 :: nil =>
- do rd <- ireg_of res; do rs <- freg_of a1;
- OK (Pfcvtzs X S rd rs :: k)
- | Olonguofsingle, a1 :: nil =>
- do rd <- ireg_of res; do rs <- freg_of a1;
- OK (Pfcvtzu X S rd rs :: k)
- | Osingleoflong, a1 :: nil =>
- do rd <- freg_of res; do rs <- ireg_of a1;
- OK (Pscvtf S X rd rs :: k)
- | Osingleoflongu, a1 :: nil =>
- do rd <- freg_of res; do rs <- ireg_of a1;
- OK (Pucvtf S X rd rs :: k)
-(** Boolean tests *)
- | Ocmp c, _ =>
- do rd <- ireg_of res;
- transl_cond c args (Pcset rd (cond_for_cond c) :: k)
-(** Conditional move *)
- | Osel cmp ty, a1 :: a2 :: args =>
- match preg_of res with
- | IR r =>
- do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- transl_cond cmp args (Pcsel r r1 r2 (cond_for_cond cmp) :: k)
- | FR r =>
- do r1 <- freg_of a1; do r2 <- freg_of a2;
- transl_cond cmp args (Pfsel r r1 r2 (cond_for_cond cmp) :: k)
- | _ =>
- Error(msg "Asmgen.Osel")
- end
- | _, _ =>
- Error(msg "Asmgen.transl_op")
- end.
+ | PLoad Pldrw rd a => do rd' <- ireg_of_preg rd; OK (Asm.Pldrw rd' a)
+ | PLoad Pldrw_a rd a => do rd' <- ireg_of_preg rd; OK (Asm.Pldrw_a rd' a)
+ | PLoad Pldrx rd a => do rd' <- ireg_of_preg rd; OK (Asm.Pldrx rd' a)
+ | PLoad Pldrx_a rd a => do rd' <- ireg_of_preg rd; OK (Asm.Pldrx_a rd' a)
+ | PLoad (Pldrb sz) rd a => do rd' <- ireg_of_preg rd; OK (Asm.Pldrb sz rd' a)
+ | PLoad (Pldrsb sz) rd a => do rd' <- ireg_of_preg rd; OK (Asm.Pldrsb sz rd' a)
+ | PLoad (Pldrh sz) rd a => do rd' <- ireg_of_preg rd; OK (Asm.Pldrh sz rd' a)
+ | PLoad (Pldrsh sz) rd a => do rd' <- ireg_of_preg rd; OK (Asm.Pldrsh sz rd' a)
+ | PLoad Pldrzw rd a => do rd' <- ireg_of_preg rd; OK (Asm.Pldrzw rd' a)
+ | PLoad Pldrsw rd a => do rd' <- ireg_of_preg rd; OK (Asm.Pldrsw rd' a)
-(** Translation of addressing modes *)
+ | PLoad Pldrs rd a => do rd' <- freg_of_preg rd; OK (Asm.Pldrs rd' a)
+ | PLoad Pldrd rd a => do rd' <- freg_of_preg rd; OK (Asm.Pldrd rd' a)
+ | PLoad Pldrd_a rd a => do rd' <- freg_of_preg rd; OK (Asm.Pldrd_a rd' a)
-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 (Archi.pic_code tt));
- 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")
- end.
+ | PStore Pstrw r a => do r' <- ireg_of_preg r; OK (Asm.Pstrw r' a)
+ | PStore Pstrw_a r a => do r' <- ireg_of_preg r; OK (Asm.Pstrw_a r' a)
+ | PStore Pstrx r a => do r' <- ireg_of_preg r; OK (Asm.Pstrx r' a)
+ | PStore Pstrx_a r a => do r' <- ireg_of_preg r; OK (Asm.Pstrx_a r' a)
+ | PStore Pstrb r a => do r' <- ireg_of_preg r; OK (Asm.Pstrb r' a)
+ | PStore Pstrh r a => do r' <- ireg_of_preg r; OK (Asm.Pstrh r' a)
-(** Translation of loads and stores *)
+ | PStore Pstrs r a => do r' <- freg_of_preg r; OK (Asm.Pstrs r' a)
+ | PStore Pstrd r a => do r' <- freg_of_preg r; OK (Asm.Pstrd r' a)
+ | PStore Pstrd_a r a => do r' <- freg_of_preg r; OK (Asm.Pstrd_a r' a)
-Definition transl_load (trap: trapping_mode)
- (chunk: memory_chunk) (addr: Op.addressing)
- (args: list mreg) (dst: mreg) (k: code) : res code :=
- match trap with
- | NOTRAP => Error (msg "Asmgen.transl_load non-trapping loads unsupported on aarch64")
- | TRAP =>
- match chunk with
- | Mint8unsigned =>
- do rd <- ireg_of dst; transl_addressing 1 addr args (Pldrb W rd) k
- | Mint8signed =>
- do rd <- ireg_of dst; transl_addressing 1 addr args (Pldrsb W rd) k
- | Mint16unsigned =>
- do rd <- ireg_of dst; transl_addressing 2 addr args (Pldrh W rd) k
- | Mint16signed =>
- do rd <- ireg_of dst; transl_addressing 2 addr args (Pldrsh W rd) k
- | Mint32 =>
- do rd <- ireg_of dst; transl_addressing 4 addr args (Pldrw rd) k
- | Mint64 =>
- do rd <- ireg_of dst; transl_addressing 8 addr args (Pldrx rd) k
- | Mfloat32 =>
- do rd <- freg_of dst; transl_addressing 4 addr args (Pldrs rd) k
- | Mfloat64 =>
- do rd <- freg_of dst; transl_addressing 8 addr args (Pldrd rd) k
- | Many32 =>
- do rd <- ireg_of dst; transl_addressing 4 addr args (Pldrw_a rd) k
- | Many64 =>
- do rd <- ireg_of dst; transl_addressing 8 addr args (Pldrx_a rd) k
- end
- end.
+ | Pallocframe sz linkofs => OK (Asm.Pallocframe sz linkofs)
+ | Pfreeframe sz linkofs => OK (Asm.Pfreeframe sz linkofs)
-Definition transl_store (chunk: memory_chunk) (addr: Op.addressing)
- (args: list mreg) (src: mreg) (k: code) : res code :=
- match chunk with
- | Mint8unsigned | Mint8signed =>
- do r1 <- ireg_of src; transl_addressing 1 addr args (Pstrb r1) k
- | Mint16unsigned | Mint16signed =>
- do r1 <- ireg_of src; transl_addressing 2 addr args (Pstrh r1) k
- | Mint32 =>
- do r1 <- ireg_of src; transl_addressing 4 addr args (Pstrw r1) k
- | Mint64 =>
- do r1 <- ireg_of src; transl_addressing 8 addr args (Pstrx r1) k
- | Mfloat32 =>
- do r1 <- freg_of src; transl_addressing 4 addr args (Pstrs r1) k
- | Mfloat64 =>
- do r1 <- freg_of src; transl_addressing 8 addr args (Pstrd r1) k
- | Many32 =>
- do r1 <- ireg_of src; transl_addressing 4 addr args (Pstrw_a r1) k
- | Many64 =>
- do r1 <- ireg_of src; transl_addressing 8 addr args (Pstrx_a r1) k
- end.
+ | Ploadsymbol rd id => OK (Asm.Ploadsymbol rd id)
-(** Register-indexed loads and stores *)
+ | Pcvtsw2x rd r1 => OK (Asm.Pcvtsw2x rd r1)
-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).
+ | Pcvtuw2x rd r1 => OK (Asm.Pcvtuw2x rd r1)
-Definition loadind (base: iregsp) (ofs: ptrofs) (ty: typ) (dst: mreg) (k: code) :=
- match ty, preg_of dst with
- | Tint, IR rd => OK (indexed_memory_access (Pldrw rd) 4 base ofs k)
- | Tlong, IR rd => OK (indexed_memory_access (Pldrx rd) 8 base ofs k)
- | Tsingle, FR rd => OK (indexed_memory_access (Pldrs rd) 4 base ofs k)
- | Tfloat, FR rd => OK (indexed_memory_access (Pldrd rd) 8 base ofs k)
- | Tany32, IR rd => OK (indexed_memory_access (Pldrw_a rd) 4 base ofs k)
- | Tany64, IR rd => OK (indexed_memory_access (Pldrx_a rd) 8 base ofs k)
- | Tany64, FR rd => OK (indexed_memory_access (Pldrd_a rd) 8 base ofs k)
- | _, _ => Error (msg "Asmgen.loadind")
+ | Pcvtx2w rd => OK (Asm.Pcvtx2w rd)
end.
-Definition storeind (src: mreg) (base: iregsp) (ofs: ptrofs) (ty: typ) (k: code) :=
- match ty, preg_of src with
- | Tint, IR rd => OK (indexed_memory_access (Pstrw rd) 4 base ofs k)
- | Tlong, IR rd => OK (indexed_memory_access (Pstrx rd) 8 base ofs k)
- | Tsingle, FR rd => OK (indexed_memory_access (Pstrs rd) 4 base ofs k)
- | Tfloat, FR rd => OK (indexed_memory_access (Pstrd rd) 8 base ofs k)
- | Tany32, IR rd => OK (indexed_memory_access (Pstrw_a rd) 4 base ofs k)
- | Tany64, IR rd => OK (indexed_memory_access (Pstrx_a rd) 8 base ofs k)
- | Tany64, FR rd => OK (indexed_memory_access (Pstrd_a rd) 8 base ofs k)
- | _, _ => Error (msg "Asmgen.storeind")
+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
end.
-Definition loadptr (base: iregsp) (ofs: ptrofs) (dst: ireg) (k: code) :=
- indexed_memory_access (Pldrx dst) 8 base ofs k.
-
-Definition storeptr (src: ireg) (base: iregsp) (ofs: ptrofs) (k: code) :=
- indexed_memory_access (Pstrx src) 8 base ofs k.
-
-(** Function epilogue *)
-
-Definition make_epilogue (f: Mach.function) (k: code) :=
- (* FIXME
- Cannot be used because memcpy destroys X30;
- issue being discussed with X. Leroy *)
- (* if is_leaf_function f
- then Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) :: k
- else*) loadptr XSP f.(fn_retaddr_ofs) RA
- (Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) :: k).
-
-(** Translation of a Mach instruction. *)
-
-Definition transl_instr (f: Mach.function) (i: Mach.instruction)
- (r29_is_parent: bool) (k: code) : res code :=
- match i with
- | Mgetstack ofs ty dst =>
- loadind XSP ofs ty dst k
- | Msetstack src ofs ty =>
- storeind src XSP ofs ty k
- | Mgetparam ofs ty dst =>
- (* load via the frame pointer if it is valid *)
- do c <- loadind X29 ofs ty dst k;
- OK (if r29_is_parent then c else loadptr XSP f.(fn_link_ofs) X29 c)
- | Mop op args res =>
- transl_op op args res k
- | Mload trap chunk addr args dst =>
- transl_load trap chunk addr args dst k
- | Mstore chunk addr args src =>
- transl_store chunk addr args src k
- | Mcall sig (inl r) =>
- do r1 <- ireg_of r; OK (Pblr r1 sig :: k)
- | Mcall sig (inr symb) =>
- OK (Pbl symb sig :: k)
- | Mtailcall sig (inl r) =>
- do r1 <- ireg_of r;
- OK (make_epilogue f (Pbr r1 sig :: k))
- | Mtailcall sig (inr symb) =>
- OK (make_epilogue f (Pbs symb sig :: k))
- | Mbuiltin ef args res =>
- OK (Pbuiltin ef (List.map (map_builtin_arg preg_of) args) (map_builtin_res preg_of res) :: k)
- | Mlabel lbl =>
- OK (Plabel lbl :: k)
- | Mgoto lbl =>
- OK (Pb lbl :: k)
- | Mcond cond args lbl =>
- transl_cond_branch cond args lbl k
- | Mjumptable arg tbl =>
- do r <- ireg_of arg;
- OK (Pbtbl r tbl :: k)
- | Mreturn =>
- OK (make_epilogue f (Pret RA :: k))
+Definition control_to_instruction (c: control) :=
+ match c with
+ | PCtlFlow i => cf_instruction_to_instruction i
+ | Pbuiltin ef args res => Asm.Pbuiltin ef args res
end.
-(** Translation of a code sequence *)
-
-Definition it1_is_parent (before: bool) (i: Mach.instruction) : bool :=
- match i with
- | Msetstack src ofs ty => before
- | Mgetparam ofs ty dst => negb (mreg_eq dst R29)
- | Mop op args res => before && negb (mreg_eq res R29)
- | _ => false
+Fixpoint unfold_label (ll: list label) :=
+ match ll with
+ | nil => nil
+ | l :: ll => Plabel l :: unfold_label ll
end.
-(** This is the naive definition that we no longer use because it
- is not tail-recursive. It is kept as specification. *)
-
-Fixpoint transl_code (f: Mach.function) (il: list Mach.instruction) (it1p: bool) :=
- match il with
+Fixpoint unfold_body (lb: list basic) : res Asm.code :=
+ match lb with
| nil => OK nil
- | i1 :: il' =>
- do k <- transl_code f il' (it1_is_parent it1p i1);
- transl_instr f i1 it1p k
+ | 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.
-(** This is an equivalent definition in continuation-passing style
- that runs in constant stack space. *)
-
-Fixpoint transl_code_rec (f: Mach.function) (il: list Mach.instruction)
- (it1p: bool) (k: code -> res code) :=
- match il with
- | nil => k nil
- | i1 :: il' =>
- transl_code_rec f il' (it1_is_parent it1p i1)
- (fun c1 => do c2 <- transl_instr f i1 it1p c1; k c2)
+Definition unfold_exit (oc: option control) :=
+ match oc with
+ | None => nil
+ | Some c => control_to_instruction c :: nil
end.
-Definition transl_code' (f: Mach.function) (il: list Mach.instruction) (it1p: bool) :=
- transl_code_rec f il it1p (fun c => OK c).
+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.
-(** Translation of a whole function. Note that we must check
- that the generated code contains less than [2^32] instructions,
- otherwise the offset part of the [PC] code pointer could wrap
- around, leading to incorrect executions. *)
+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 transl_function (f: Mach.function) :=
- do c <- transl_code' f f.(Mach.fn_code) true;
- OK (mkfunction f.(Mach.fn_sig)
- (Pallocframe f.(fn_stacksize) f.(fn_link_ofs) ::
- storeptr RA XSP f.(fn_retaddr_ofs) c)).
+Definition transf_fundef (f: Asmblock.fundef) : res Asm.fundef :=
+ transf_partial_fundef transf_function f.
-Definition transf_function (f: Mach.function) : res Asm.function :=
- do tf <- transl_function f;
- if zlt Ptrofs.max_unsigned (list_length_z tf.(fn_code))
- then Error (msg "code size exceeded")
- else OK tf.
+Definition transf_program (p: Asmblock.program) : res Asm.program :=
+ transform_partial_program transf_fundef p.
-Definition transf_fundef (f: Mach.fundef) : res Asm.fundef :=
- transf_partial_fundef transf_function f.
+End Asmblock_TRANSF.
Definition transf_program (p: Mach.program) : res Asm.program :=
- transform_partial_program transf_fundef p.
+ let mbp := Machblockgen.transf_program p in
+ do mbp' <- PseudoAsmblockproof.transf_program mbp;
+ do abp <- Asmblockgen.transf_program mbp';
+ Asmblock_TRANSF.transf_program abp.
diff --git a/aarch64/Asmgenproof.v b/aarch64/Asmgenproof.v
index 6831509f..b2671972 100644
--- a/aarch64/Asmgenproof.v
+++ b/aarch64/Asmgenproof.v
@@ -1,24 +1,34 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
+(* *************************************************************)
+(* *)
+(* 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 *)
+(* *)
+(* 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 Mach Conventions Asm.
-Require Import Asmgen Asmgenproof0 Asmgenproof1.
+Require Import Op Locations Machblock Conventions PseudoAsmblock Asmblock.
+Require Machblockgenproof Asmblockgenproof.
+Require Import Asmgen.
+Require Import Axioms.
+Require Import IterList.
+Require Import Ring.
-Definition match_prog (p: Mach.program) (tp: Asm.program) :=
+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:
@@ -29,20 +39,36 @@ Qed.
Section PRESERVATION.
-Variable prog: Mach.program.
+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 ->
@@ -50,1052 +76,2020 @@ Lemma functions_translated:
Genv.find_funct_ptr tge b = Some tf /\ transf_fundef f = OK tf.
Proof (Genv.find_funct_ptr_transf_partial TRANSF).
-Lemma functions_transl:
- forall fb f tf,
+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.
+ omega.
+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 ->
+ 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; omega.
+Qed.
+
+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; omega.
+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 ->
- Genv.find_funct_ptr tge fb = Some (Internal 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; omega.
+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.
+ - omega.
+ - generalize (IHl (Z.succ acc)). omega.
+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 || omega);
+ unfold list_length_z; simpl;
+ generalize (list_length_z_aux_increase _ tl 1); omega.
+Qed.
+
+
+Lemma list_length_add_acc A (l : list A) acc:
+ list_length_z_aux l acc = (list_length_z l) + acc.
Proof.
- intros. exploit functions_translated; eauto. intros [tf' [A B]].
- monadInv B. rewrite H0 in EQ; inv EQ; auto.
+ unfold list_length_z, list_length_z_aux. simpl.
+ fold list_length_z_aux.
+ rewrite (list_length_z_aux_shift l acc 0).
+ omega.
Qed.
-(** * Properties of control flow *)
+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 length_agree A (l : list A):*)
+ (*list_length_z l = Z.of_nat (length l).*)
+(*Proof.*)
+ (*induction l as [| ? l IH]; intros.*)
+ (*- unfold list_length_z; reflexivity.*)
+ (*- simpl; rewrite list_length_z_cons, Zpos_P_of_succ_nat; omega.*)
+(*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 transf_function_no_overflow:
- forall f tf,
- transf_function f = OK tf -> list_length_z tf.(fn_code) <= Ptrofs.max_unsigned.
+Lemma header_size_lt_block_size bb:
+ list_length_z (header bb) < size bb.
Proof.
- intros. monadInv H. destruct (zlt Ptrofs.max_unsigned (list_length_z x.(fn_code))); inv EQ0.
+ 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; omega.
+ - destruct (exit bb); try contradiction; simpl; repeat rewrite list_length_z_nat; omega.
+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; omega.
+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)).
omega.
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.
-*)
+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.
-Section TRANSL_LABEL.
+Lemma unfold_cdr bb bbs tc:
+ unfold (bb :: bbs) = OK tc ->
+ exists tc', unfold bbs = OK tc'.
+Proof.
+ intros; exploit unfold_car_cdr; eauto. intros (_ & ? & _ & ? & _).
+ eexists; eauto.
+Qed.
-Remark loadimm_z_label: forall sz rd l k, tail_nolabel k (loadimm_z sz rd l k).
-Proof.
- intros; destruct l as [ | [n1 p1] l]; simpl; TailNoLabel.
- induction l as [ | [n p] l]; simpl; TailNoLabel.
+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.
-Remark loadimm_n_label: forall sz rd l k, tail_nolabel k (loadimm_n sz rd l k).
-Proof.
- intros; destruct l as [ | [n1 p1] l]; simpl; TailNoLabel.
- induction l as [ | [n p] l]; simpl; TailNoLabel.
+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.
-Remark loadimm_label: forall sz rd n k, tail_nolabel k (loadimm sz rd n k).
+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.
- unfold loadimm; intros. destruct Nat.leb; [apply loadimm_z_label|apply loadimm_n_label].
+ 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.
-Hint Resolve loadimm_label: labels.
-Remark loadimm32_label: forall r n k, tail_nolabel k (loadimm32 r n k).
+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.
- unfold loadimm32; intros. destruct (is_logical_imm32 n); TailNoLabel.
+ 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.
-Hint Resolve loadimm32_label: labels.
-Remark loadimm64_label: forall r n k, tail_nolabel k (loadimm64 r n k).
+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.
- unfold loadimm64; intros. destruct (is_logical_imm64 n); TailNoLabel.
+ intros; exploit bblock_in_bblocks; eauto; intros;
+ eapply all_blocks_translated; eauto.
Qed.
-Hint Resolve loadimm64_label: labels.
-Remark addimm_aux: forall insn rd r1 n k,
- (forall rd r1 n, nolabel (insn rd r1 n)) ->
- tail_nolabel k (addimm_aux insn rd r1 n k).
+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.
- unfold addimm_aux; intros.
- destruct Z.eqb. TailNoLabel. destruct Z.eqb; TailNoLabel.
+ 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.
-Remark addimm32_label: forall rd r1 n k, tail_nolabel k (addimm32 rd r1 n k).
+Lemma list_nth_z_neg A (l: list A): forall n,
+ n < 0 -> list_nth_z l n = None.
Proof.
- unfold addimm32; intros.
- destruct Int.eq. apply addimm_aux; intros; red; auto.
- destruct Int.eq. apply addimm_aux; intros; red; auto.
- destruct Int.lt; eapply tail_nolabel_trans; TailNoLabel.
+ induction l; simpl; auto.
+ intros n H; destruct (zeq _ _); (try eapply IHl); omega.
Qed.
-Hint Resolve addimm32_label: labels.
-Remark addimm64_label: forall rd r1 n k, tail_nolabel k (addimm64 rd r1 n k).
+Lemma find_bblock_neg bbs: forall pos,
+ pos < 0 -> find_bblock pos bbs = None.
Proof.
- unfold addimm64; intros.
- destruct Int64.eq. apply addimm_aux; intros; red; auto.
- destruct Int64.eq. apply addimm_aux; intros; red; auto.
- destruct Int64.lt; eapply tail_nolabel_trans; TailNoLabel.
+ induction bbs; simpl; auto.
+ intros. destruct (zlt pos 0). { reflexivity. }
+ destruct (zeq pos 0); contradiction.
Qed.
-Hint Resolve addimm64_label: labels.
-Remark logicalimm32_label: forall insn1 insn2 rd r1 n k,
- (forall rd r1 n, nolabel (insn1 rd r1 n)) ->
- (forall rd r1 r2 s, nolabel (insn2 rd r1 r2 s)) ->
- tail_nolabel k (logicalimm32 insn1 insn2 rd r1 n k).
+Lemma equal_header_size bb:
+ length (header bb) = length (unfold_label (header bb)).
Proof.
- unfold logicalimm32; intros.
- destruct (is_logical_imm32 n). TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
+ induction (header bb); auto.
+ simpl. rewrite IHl. auto.
Qed.
-Remark logicalimm64_label: forall insn1 insn2 rd r1 n k,
- (forall rd r1 n, nolabel (insn1 rd r1 n)) ->
- (forall rd r1 r2 s, nolabel (insn2 rd r1 r2 s)) ->
- tail_nolabel k (logicalimm64 insn1 insn2 rd r1 n k).
+Lemma equal_body_size:
+ forall bb tb,
+ unfold_body (body bb) = OK tb ->
+ length (body bb) = length tb.
Proof.
- unfold logicalimm64; intros.
- destruct (is_logical_imm64 n). TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
+ 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.
-Remark move_extended_label: forall rd r1 ex a k, tail_nolabel k (move_extended rd r1 ex a k).
+Lemma equal_exit_size bb:
+ length_opt (exit bb) = length (unfold_exit (exit bb)).
Proof.
- unfold move_extended, move_extended_base; intros. destruct Int.eq, ex; TailNoLabel.
+ destruct (exit bb); trivial.
Qed.
-Hint Resolve move_extended_label: labels.
-Remark arith_extended_label: forall insnX insnS rd r1 r2 ex a k,
- (forall rd r1 r2 x, nolabel (insnX rd r1 r2 x)) ->
- (forall rd r1 r2 s, nolabel (insnS rd r1 r2 s)) ->
- tail_nolabel k (arith_extended insnX insnS rd r1 r2 ex a k).
+Lemma bblock_size_preserved bb tb:
+ unfold_bblock bb = OK tb ->
+ size bb = list_length_z tb.
Proof.
- unfold arith_extended; intros. destruct Int.ltu.
- TailNoLabel.
- destruct ex; simpl; TailNoLabel.
+ 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.
-Remark loadsymbol_label: forall r id ofs k, tail_nolabel k (loadsymbol r id ofs k).
+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.
- intros; unfold loadsymbol.
- destruct (Archi.pic_code tt); TailNoLabel. destruct Ptrofs.eq; TailNoLabel.
-Qed.
-Hint Resolve loadsymbol_label: labels.
+ 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.
+ omega.
+ + 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).
+ omega.
+Qed.
-Remark transl_cond_label: forall cond args k c,
- transl_cond cond args k = OK c -> tail_nolabel k c.
+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 transl_cond; intros; destruct cond; TailNoLabel.
-- destruct is_arith_imm32; TailNoLabel. destruct is_arith_imm32; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
-- destruct is_arith_imm32; TailNoLabel. destruct is_arith_imm32; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
-- destruct is_logical_imm32; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
-- destruct is_logical_imm32; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
-- destruct is_arith_imm64; TailNoLabel. destruct is_arith_imm64; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
-- destruct is_arith_imm64; TailNoLabel. destruct is_arith_imm64; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
-- destruct is_logical_imm64; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
-- destruct is_logical_imm64; TailNoLabel. eapply tail_nolabel_trans; TailNoLabel.
+ 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.
-Remark transl_cond_branch_default_label: forall cond args lbl k c,
- transl_cond_branch_default cond args lbl k = OK c -> tail_nolabel k c.
+Lemma unfold_bblock_not_nil bb:
+ unfold_bblock bb = OK nil -> False.
Proof.
- unfold transl_cond_branch_default; intros.
- eapply tail_nolabel_trans; [eapply transl_cond_label;eauto|TailNoLabel].
+ intros.
+ exploit bblock_size_preserved; eauto. unfold list_length_z; simpl. intros SIZE.
+ generalize (bblock_size_pos bb). intros SIZE'. omega.
Qed.
-Hint Resolve transl_cond_branch_default_label: labels.
-Remark transl_cond_branch_label: forall cond args lbl k c,
- transl_cond_branch cond args lbl k = OK c -> tail_nolabel k c.
+(* 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.
- unfold transl_cond_branch; intros; destruct args; TailNoLabel; destruct cond; TailNoLabel.
-- destruct c0; TailNoLabel.
-- destruct c0; TailNoLabel.
-- destruct (Int.is_power2 n); TailNoLabel.
-- destruct (Int.is_power2 n); TailNoLabel.
-- destruct c0; TailNoLabel.
-- destruct c0; TailNoLabel.
-- destruct (Int64.is_power2' n); TailNoLabel.
-- destruct (Int64.is_power2' n); TailNoLabel.
+ induction c; simpl; intros.
+ discriminate.
+ rewrite list_length_z_cons. destruct (zeq n 0).
+ generalize (list_length_z_pos c); omega.
+ exploit IHc; eauto. omega.
Qed.
-Remark transl_op_label:
- forall op args r k c,
- transl_op op args r k = OK c -> tail_nolabel k c.
+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.
- unfold transl_op; intros; destruct op; TailNoLabel.
-- destruct (preg_of r); try discriminate; destruct (preg_of m); inv H; TailNoLabel.
-- destruct (Float.eq_dec n Float.zero); TailNoLabel.
-- destruct (Float32.eq_dec n Float32.zero); TailNoLabel.
-- apply logicalimm32_label; unfold nolabel; auto.
-- apply logicalimm32_label; unfold nolabel; auto.
-- apply logicalimm32_label; unfold nolabel; auto.
-- unfold shrx32. destruct (Int.eq _ _); try destruct (Int.eq _ _); TailNoLabel.
-- apply arith_extended_label; unfold nolabel; auto.
-- apply arith_extended_label; unfold nolabel; auto.
-- apply logicalimm64_label; unfold nolabel; auto.
-- apply logicalimm64_label; unfold nolabel; auto.
-- apply logicalimm64_label; unfold nolabel; auto.
-- unfold shrx64. destruct (Int.eq _ _); try destruct (Int.eq _ _); TailNoLabel.
-- eapply tail_nolabel_trans. eapply transl_cond_label; eauto. TailNoLabel.
-- destruct (preg_of r); try discriminate; TailNoLabel;
- (eapply tail_nolabel_trans; [eapply transl_cond_label; eauto | TailNoLabel]).
+ 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). omega.
+ + replace (pos + (list_length_z tbb + 1) - 1) with (pos + list_length_z tbb) by omega.
+ eapply IHtbb; eauto.
Qed.
-Remark transl_addressing_label:
- forall sz addr args insn k c,
- transl_addressing sz addr args insn k = OK c ->
- (forall ad, nolabel (insn ad)) ->
- tail_nolabel k c.
+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.
- unfold transl_addressing; intros; destruct addr; TailNoLabel;
- eapply tail_nolabel_trans; TailNoLabel.
- eapply tail_nolabel_trans. apply arith_extended_label; unfold nolabel; auto. TailNoLabel.
+ 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. }
+ omega.
Qed.
-Remark transl_load_label:
- forall trap chunk addr args dst k c,
- transl_load trap chunk addr args dst k = OK c -> tail_nolabel k c.
+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.
- unfold transl_load; intros; destruct trap; try discriminate; destruct chunk; TailNoLabel; eapply transl_addressing_label; eauto; unfold nolabel; auto.
+ 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. omega.
+ + 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 omega.
+ apply find_instr_tail; auto.
Qed.
-Remark transl_store_label:
- forall chunk addr args src k c,
- transl_store chunk addr args src k = OK c -> tail_nolabel k c.
+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.
- unfold transl_store; intros; destruct chunk; TailNoLabel; eapply transl_addressing_label; eauto; unfold nolabel; auto.
+ 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.
-Remark indexed_memory_access_label:
- forall insn sz base ofs k,
- (forall ad, nolabel (insn ad)) ->
- tail_nolabel k (indexed_memory_access insn sz base ofs k).
+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.
- unfold indexed_memory_access; intros. destruct offset_representable.
- TailNoLabel.
- eapply tail_nolabel_trans; TailNoLabel.
+ 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.
-Remark loadind_label:
- forall base ofs ty dst k c,
- loadind base ofs ty dst k = OK c -> tail_nolabel k c.
+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.
- unfold loadind; intros.
- destruct ty, (preg_of dst); inv H; apply indexed_memory_access_label; intros; exact I.
+ induction ll.
+ - unfold list_length_z. simpl. intros.
+ replace (n - 0) with n in H by omega. 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 omega.
+ generalize (list_length_z_pos ll). intros.
+ rewrite list_nth_z_neg in H; try omega. 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 omega.
+ eapply IHll; eauto.
Qed.
-Remark storeind_label:
- forall src base ofs ty k c,
- storeind src base ofs ty k = OK c -> tail_nolabel k c.
+(* 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.
- unfold storeind; intros.
- destruct ty, (preg_of src); inv H; apply indexed_memory_access_label; intros; exact I.
+ induction l.
+ - intros. unfold list_length_z in H. simpl in H. omega.
+ - 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 omega.
+ eapply IHl; omega.
Qed.
-Remark loadptr_label:
- forall base ofs dst k, tail_nolabel k (loadptr base ofs dst k).
+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.
- intros. apply indexed_memory_access_label. unfold nolabel; auto.
+ induction l.
+ - intros. unfold list_length_z. simpl. omega.
+ - 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). { omega. }
+ assert (0 <= n - 1). { omega. }
+ generalize (IHl (n - 1)). intros IH.
+ assert (n - 1 >= list_length_z l). { auto. }
+ assert (n > list_length_z l); omega.
Qed.
-Remark storeptr_label:
- forall src base ofs k, tail_nolabel k (storeptr src base ofs k).
+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.
- intros. apply indexed_memory_access_label. unfold nolabel; auto.
+ 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 omega.
+ simpl in H. destruct (zeq n 0). { contradiction. }
+ replace (Z.pred n) with (n - 1) in H by omega.
+ apply IH; auto.
Qed.
-Remark make_epilogue_label:
- forall f k, tail_nolabel k (make_epilogue f k).
+(* 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.
- unfold make_epilogue; intros.
- (* FIXME destruct is_leaf_function.
- { TailNoLabel. } *)
- eapply tail_nolabel_trans.
- apply loadptr_label.
- TailNoLabel.
+ 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 omega.
+ simpl in NTH. destruct (zeq n 0). { contradiction. }
+ replace (Z.pred n) with (n - 1) in NTH by omega.
+ apply IHlbi; auto.
Qed.
-Lemma transl_instr_label:
- forall f i ep k c,
- transl_instr f i ep k = OK c ->
- match i with Mlabel lbl => c = Plabel lbl :: k | _ => tail_nolabel k c end.
+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.
- unfold transl_instr; intros; destruct i; TailNoLabel.
-- eapply loadind_label; eauto.
-- eapply storeind_label; eauto.
-- destruct ep. eapply loadind_label; eauto.
- eapply tail_nolabel_trans. apply loadptr_label. eapply loadind_label; eauto.
-- eapply transl_op_label; eauto.
-- eapply transl_load_label; eauto.
-- eapply transl_store_label; eauto.
-- destruct s0; monadInv H; TailNoLabel.
-- destruct s0; monadInv H; (eapply tail_nolabel_trans; [eapply make_epilogue_label|TailNoLabel]).
-- eapply transl_cond_branch_label; eauto.
-- eapply tail_nolabel_trans; [eapply make_epilogue_label|TailNoLabel].
+ intros.
+ assert (0 <= n). { omega. }
+ 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)). { omega. }
+ 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)). { omega. }
+ rewrite Nat2Z.inj_add.
+ repeat (rewrite <- list_length_z_nat). assumption.
Qed.
-Lemma transl_instr_label':
- forall lbl f i ep k c,
- transl_instr f i ep k = OK c ->
- find_label lbl c = if Mach.is_label lbl i then Some k else find_label lbl k.
+Lemma exec_arith_instr_dont_move_PC ai rs rs': forall
+ (BASIC: exec_arith_instr lk ai rs = rs'),
+ rs PC = rs' PC.
Proof.
- intros. exploit transl_instr_label; eauto.
- destruct i; try (intros [A B]; apply B).
- intros. subst c. simpl. auto.
+ destruct ai; simpl; intros;
+ try (rewrite <- BASIC; rewrite Pregmap.gso; auto; discriminate).
+ - destruct i; simpl in BASIC.
+ + unfold compare_long in BASIC; rewrite <- BASIC.
+ repeat rewrite Pregmap.gso; try discriminate. reflexivity.
+ + unfold compare_long in BASIC; 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 transl_code_label:
- forall lbl f c ep tc,
- transl_code f c ep = OK tc ->
- match Mach.find_label lbl c with
- | None => find_label lbl tc = None
- | Some c' => exists tc', find_label lbl tc = Some tc' /\ transl_code f c' false = OK tc'
- end.
+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.
- induction c; simpl; intros.
- inv H. auto.
- monadInv H. rewrite (transl_instr_label' lbl _ _ _ _ _ EQ0).
- generalize (Mach.is_label_correct lbl a).
- destruct (Mach.is_label lbl a); intros.
- subst a. simpl in EQ. exists x; auto.
- eapply IHc; eauto.
+ destruct bi; simpl; intros.
+ - inv BASIC. exploit exec_arith_instr_dont_move_PC; eauto.
+ - unfold exec_load in BASIC.
+ destruct Mem.loadv. 2: { discriminate BASIC. }
+ inv BASIC. rewrite Pregmap.gso; try discriminate; auto.
+ - unfold exec_store in BASIC.
+ destruct Mem.storev. 2: { 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.
Qed.
-Lemma transl_find_label:
- forall lbl f tf,
- transf_function f = OK tf ->
- match Mach.find_label lbl f.(Mach.fn_code) with
- | None => find_label lbl tf.(fn_code) = None
- | Some c => exists tc, find_label lbl tf.(fn_code) = Some tc /\ transl_code f c false = OK tc
+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)). { omega. }
+ assert (n = size bb - 1). {
+ unfold size. rewrite <- EXIT. simpl.
+ repeat (rewrite Nat2Z.inj_add). repeat (rewrite <- list_length_z_nat). simpl. omega.
+ }
+ 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 *) omega.
+ ++ (* absurd *)
+ unfold size in SIZE. rewrite <- EXIT in SIZE. simpl in SIZE.
+ destruct SIZE as (? & SIZE'). rewrite Nat.add_0_r in SIZE'. omega.
+ + 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 omega.
+ eapply find_instr_bblock_tail; try assumption.
+ replace (Z.pos p + n - size b) with (Z.pos p - size b + n) by omega.
+ 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). omega. }
+ intros (i & NTH & FIND_INSTR).
+ inv NTH.
+ * rewrite EQhead in H; simpl in H. inv H.
+ cutrewrite (Ptrofs.unsigned ofs + 0 = Ptrofs.unsigned ofs) in FIND_INSTR; try omega.
+ 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; omega].
+ * (* absurd case *)
+ rewrite bblock_size_aux, Lhead in *. generalize (bblock_size_aux_pos bb). omega.
+ + (* absurd case *)
+ unfold list_length_z in BNDhead. simpl in *.
+ generalize (list_length_z_aux_increase _ l1 2); omega.
+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. monadInv H. destruct (zlt Ptrofs.max_unsigned (list_length_z x.(fn_code))); inv EQ0.
- monadInv EQ. rewrite transl_code'_transl_code in EQ0. unfold fn_code.
- simpl. destruct (storeptr_label X30 XSP (fn_retaddr_ofs f) x) as [A B]; rewrite B.
- eapply transl_code_label; eauto.
+ intros; unfold Asm.nextinstr in *; rewrite Pregmap.gso in *; eauto.
Qed.
-End TRANSL_LABEL.
+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.
-(** A valid branch in a piece of Mach code translates to a valid ``go to''
- transition in the generated Asm code. *)
+Lemma load_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 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, 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 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. omega.
- generalize (transf_function_no_overflow _ _ H0). omega.
- intros. apply Pregmap.gso; auto.
-Qed.
-
-(** Existence of return addresses *)
+Lemma store_preserved n rs1 m1 rs1' m1' rs2 m2 rd chk a: forall
+ (BOUNDED: 0 <= n <= Ptrofs.max_unsigned)
+ (MATCHI: match_internal n (State rs1 m1) (State rs2 m2))
+ (HSTORE: exec_store lk chk a rd rs1 m1 = Next rs1' m1'),
+ exists (rs2' : regset) (m2' : mem), Asm.exec_store tge chk a rd rs2 m2 = Next rs2' m2'
+ /\ match_internal (n + 1) (State rs1' m1') (State rs2' m2').
+Proof.
+ intros.
+ unfold exec_store, 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 return_address_exists:
- forall f sg ros c, is_tail (Mcall sg ros :: c) f.(Mach.fn_code) ->
- exists ra, return_address_offset f c ra.
+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. eapply Asmgenproof0.return_address_exists; eauto.
-- intros. exploit transl_instr_label; eauto.
- destruct i; try (intros [A B]; apply A). intros. subst c0. repeat constructor.
-- intros. monadInv H0.
- destruct (zlt Ptrofs.max_unsigned (list_length_z x.(fn_code))); inv EQ0. monadInv EQ.
- rewrite transl_code'_transl_code in EQ0.
- exists x; exists true; split; auto. unfold fn_code.
- constructor. apply (storeptr_label X30 XSP (fn_retaddr_ofs f0) x).
-- exact transf_function_no_overflow.
-Qed.
-
-(** * Proof of semantic preservation *)
-
-(** Semantic preservation is proved using simulation diagrams
- of the following form.
-<<
- st1 --------------- st2
- | |
- t| *|t
- | |
- v v
- st1'--------------- st2'
->>
- The invariant is the [match_states] predicate below, which includes:
-- The Asm code pointed by the PC register is the translation of
- the current Mach code sequence.
-- Mach register values and Asm register values agree.
-*)
+ 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.
-Inductive match_states: Mach.state -> Asm.state -> Prop :=
- | match_states_intro:
- forall s fb sp c ep ms m m' rs f tf tc
- (STACKS: match_stack ge s)
- (FIND: Genv.find_funct_ptr ge fb = Some (Internal f))
- (MEXT: Mem.extends m m')
- (AT: transl_code_at_pc ge (rs PC) fb f c ep tf tc)
- (AG: agree ms sp rs)
- (DXP: ep = true -> rs#X29 = parent_sp s)
- (LEAF: is_leaf_function f = true -> rs#RA = parent_ra s),
- match_states (Mach.State s fb sp c ms m)
- (Asm.State rs m')
- | match_states_call:
- forall s fb ms m m' rs
- (STACKS: match_stack ge s)
- (MEXT: Mem.extends m m')
- (AG: agree ms (parent_sp s) rs)
- (ATPC: rs PC = Vptr fb Ptrofs.zero)
- (ATLR: rs RA = parent_ra s),
- match_states (Mach.Callstate s fb ms m)
- (Asm.State rs m')
- | match_states_return:
- forall s ms m m' rs
- (STACKS: match_stack ge s)
- (MEXT: Mem.extends m m')
- (AG: agree ms (parent_sp s) rs)
- (ATPC: rs PC = parent_ra s),
- match_states (Mach.Returnstate s ms m)
- (Asm.State rs m').
-
-Lemma exec_straight_steps:
- forall s fb f rs1 i c ep tf tc m1' m2 m2' sp ms2,
- match_stack ge s ->
- Mem.extends m2 m2' ->
- Genv.find_funct_ptr ge fb = Some (Internal f) ->
- transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc ->
- (forall k c (TR: transl_instr f i ep k = OK c),
- exists rs2,
- exec_straight tge tf c rs1 m1' k rs2 m2'
- /\ agree ms2 sp rs2
- /\ (it1_is_parent ep i = true -> rs2#X29 = parent_sp s)
- /\ (is_leaf_function f = true -> rs2#RA = parent_ra s)) ->
- exists st',
- plus step tge (State rs1 m1') E0 st' /\
- match_states (Mach.State s fb sp c ms2 m2) st'.
-Proof.
- intros. inversion H2. subst. monadInv H7.
- exploit H3; eauto. intros [rs2 [A [B [C D]]]].
- exists (State rs2 m2'); split.
- - eapply exec_straight_exec; eauto.
- - econstructor; eauto. eapply exec_straight_at; eauto.
-Qed.
-
-Lemma exec_straight_steps_goto:
- forall s fb f rs1 i c ep tf tc m1' m2 m2' sp ms2 lbl c',
- match_stack ge s ->
- Mem.extends m2 m2' ->
- Genv.find_funct_ptr ge fb = Some (Internal f) ->
- Mach.find_label lbl f.(Mach.fn_code) = Some c' ->
- transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc ->
- it1_is_parent ep i = false ->
- (forall k c (TR: transl_instr f i ep k = OK c),
- exists jmp, exists k', exists rs2,
- exec_straight tge tf c rs1 m1' (jmp :: k') rs2 m2'
- /\ agree ms2 sp rs2
- /\ exec_instr tge tf jmp rs2 m2' = goto_label tf lbl rs2 m2'
- /\ (is_leaf_function f = true -> rs2#RA = parent_ra s)) ->
- exists st',
- plus step tge (State rs1 m1') E0 st' /\
- match_states (Mach.State s fb sp c' ms2 m2) st'.
-Proof.
- intros. inversion H3. subst. monadInv H9.
- exploit H5; eauto. intros [jmp [k' [rs2 [A [B [C D]]]]]].
- generalize (functions_transl _ _ _ H7 H8); intro FN.
- generalize (transf_function_no_overflow _ _ H8); intro NOOV.
- exploit exec_straight_steps_2; eauto.
- intros [ofs' [PC2 CT2]].
- exploit find_label_goto_label; eauto.
- intros [tc' [rs3 [GOTO [AT' OTH]]]].
- exists (State rs3 m2'); split.
- eapply plus_right'.
- eapply exec_straight_steps_1; eauto.
- econstructor; eauto.
- eapply find_instr_tail. eauto.
- rewrite C. eexact GOTO.
- traceEq.
- econstructor; eauto.
- apply agree_exten with rs2; auto with asmgen.
- congruence.
- rewrite OTH by congruence; auto.
-Qed.
-
-Lemma exec_straight_opt_steps_goto:
- forall s fb f rs1 i c ep tf tc m1' m2 m2' sp ms2 lbl c',
- match_stack ge s ->
- Mem.extends m2 m2' ->
- Genv.find_funct_ptr ge fb = Some (Internal f) ->
- Mach.find_label lbl f.(Mach.fn_code) = Some c' ->
- transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc ->
- it1_is_parent ep i = false ->
- (forall k c (TR: transl_instr f i ep k = OK c),
- exists jmp, exists k', exists rs2,
- exec_straight_opt tge tf c rs1 m1' (jmp :: k') rs2 m2'
- /\ agree ms2 sp rs2
- /\ exec_instr tge tf jmp rs2 m2' = goto_label tf lbl rs2 m2'
- /\ (is_leaf_function f = true -> rs2#RA = parent_ra s)) ->
- exists st',
- plus step tge (State rs1 m1') E0 st' /\
- match_states (Mach.State s fb sp c' ms2 m2) st'.
-Proof.
- intros. inversion H3. subst. monadInv H9.
- exploit H5; eauto. intros [jmp [k' [rs2 [A [B [C D]]]]]].
- generalize (functions_transl _ _ _ H7 H8); intro FN.
- generalize (transf_function_no_overflow _ _ H8); intro NOOV.
- inv A.
-- exploit find_label_goto_label; eauto.
- intros [tc' [rs3 [GOTO [AT' OTH]]]].
- exists (State rs3 m2'); split.
- apply plus_one. econstructor; eauto.
- eapply find_instr_tail. eauto.
- rewrite C. eexact GOTO.
- econstructor; eauto.
- apply agree_exten with rs2; auto with asmgen.
- congruence.
- rewrite OTH by congruence; auto.
-- exploit exec_straight_steps_2; eauto.
- intros [ofs' [PC2 CT2]].
- exploit find_label_goto_label; eauto.
- intros [tc' [rs3 [GOTO [AT' OTH]]]].
- exists (State rs3 m2'); split.
- eapply plus_right'.
- eapply exec_straight_steps_1; eauto.
- econstructor; eauto.
- eapply find_instr_tail. eauto.
- rewrite C. eexact GOTO.
- traceEq.
- econstructor; eauto.
- apply agree_exten with rs2; auto with asmgen.
- congruence.
- rewrite OTH by congruence; auto.
-Qed.
-
-(** We need to show that, in the simulation diagram, we cannot
- take infinitely many Mach transitions that correspond to zero
- transitions on the Asm side. Actually, all Mach transitions
- correspond to at least one Asm transition, except the
- transition from [Machsem.Returnstate] to [Machsem.State].
- So, the following integer measure will suffice to rule out
- the unwanted behaviour. *)
-
-Definition measure (s: Mach.state) : nat :=
- match s with
- | Mach.State _ _ _ _ _ _ => 0%nat
- | Mach.Callstate _ _ _ _ => 0%nat
- | Mach.Returnstate _ _ _ => 1%nat
- end.
+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,2,3,4,5,6: (* 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; monadInv TRANSBI; try destruct_ireg_inv; exploit load_preserved; eauto;
+ intros; simpl in *; destruct sz; eauto. }
+ { (* PStore *)
+ destruct st; monadInv TRANSBI; try destruct_ireg_inv; exploit store_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).
+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; omega.
+ - repeat (rewrite list_length_z_nat). repeat (rewrite Nat2Z.inj_add). omega. }
+ intros (i & NTH & FIND_INSTR).
+ exists i; intros.
+ inv NTH.
+ - (* absurd *) apply list_nth_z_range in H; omega.
+ - 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; omega.
+Qed.
+
+(** "is_tail" auxiliary lemma about is_tail to move in IterList ou Coqlib (déplacer aussi Machblockgenproof.is_tail_app_inv) ? *)
+
+Lemma is_tail_app_right A (l2 l1: list A): is_tail l1 (l2++l1).
+Proof.
+ intros; eapply Machblockgenproof.is_tail_app_inv; econstructor.
+Qed.
+
+Lemma is_tail_app_def A (l1 l2: list A):
+ is_tail l1 l2 -> exists l3, l2 = l3 ++ l1.
+Proof.
+ induction 1 as [|x l1 l2]; simpl.
+ - exists nil; simpl; auto.
+ - destruct IHis_tail as (l3 & EQ); rewrite EQ.
+ exists (x::l3); simpl; auto.
+Qed.
+
+Lemma is_tail_bound A (l1 l2: list A):
+ is_tail l1 l2 -> (length l1 <= length l2)%nat.
+Proof.
+ intros H; destruct (is_tail_app_def _ _ _ H) as (l3 & EQ).
+ subst; rewrite app_length.
+ omega.
+Qed.
+
+Lemma is_tail_list_nth_z A (l1 l2: list A):
+ is_tail l1 l2 -> list_nth_z l2 ((list_length_z l2) - (list_length_z l1)) = list_nth_z l1 0.
+Proof.
+ induction 1; simpl.
+ - replace (list_length_z c - list_length_z c) with 0; omega || auto.
+ - assert (X: list_length_z (i :: c2) > list_length_z c1).
+ { rewrite !list_length_z_nat, <- Nat2Z.inj_gt.
+ exploit is_tail_bound; simpl; eauto.
+ omega. }
+ destruct (zeq (list_length_z (i :: c2) - list_length_z c1) 0) as [Y|Y]; try omega.
+ replace (Z.pred (list_length_z (i :: c2) - list_length_z c1)) with (list_length_z c2 - list_length_z c1); auto.
+ rewrite list_length_z_cons.
+ omega.
+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 omega. }
+ exploit find_basic_instructions; eauto.
+ rewrite Z2Nat.id; try omega. 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 omega.
+ unfold size in BBSIZE.
+ rewrite !Nat2Z.inj_add in BBSIZE.
+ rewrite <- !list_length_z_nat in BBSIZE.
+ omega.
+Qed.
-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') (WF: wf_state ge 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. }
- split. { simpl; congruence. }
- rewrite nextinstr_inv by congruence; assumption.
-
-- (* 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 S]]]].
- exists rs'; split. eauto.
- split. { eapply agree_set_mreg; eauto with asmgen. congruence. }
- split. { simpl; congruence. }
- rewrite S. assumption.
-
-- (* 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 R]]].
- exists rs'; split. eauto.
- split. eapply agree_undef_regs; eauto with asmgen.
- simpl; intros.
- split. rewrite Q; auto with asmgen.
- rewrite R. assumption.
-
-- (* 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 S]]]].
- exists rs1; split. eauto.
- split. eapply agree_set_mreg. eapply agree_set_mreg; eauto. congruence. auto with asmgen.
- simpl; split; intros.
- { rewrite R; auto with asmgen.
- apply preg_of_not_X29; auto.
+(* 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 omega.
+ rewrite !list_length_z_nat; split.
+ - rewrite <- Nat2Z.inj_lt. simpl. omega.
+ - rewrite <- Nat2Z.inj_le; eapply is_tail_bound; eauto.
}
- { rewrite S; 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 V]]]].
- 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.
- split; simpl; intros. rewrite U; auto with asmgen.
- apply preg_of_not_X29; auto.
- rewrite V. rewrite R by congruence. 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 S]]]].
- exists rs2; split. eauto. split.
- apply agree_set_undef_mreg with rs0; auto.
- apply Val.lessdef_trans with v'; auto.
- split; simpl; intros. InvBooleans.
- rewrite R; auto. apply preg_of_not_X29; auto.
-Local Transparent destroyed_by_op.
- destruct op; try exact I; simpl; congruence.
- rewrite S.
- auto.
-- (* Mload *)
- destruct trap.
+ 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.
{
- 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 S]]]].
- exists rs2; split. eauto.
- split. eapply agree_set_undef_mreg; eauto. congruence.
- split. simpl; congruence.
- rewrite S. assumption.
+ rewrite Ptrofs.add_unsigned.
+ repeat (rewrite Ptrofs.unsigned_repr); try omega.
+ 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. }
+ omega. }
+ 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 omega;
+ try assumption; try omega. }
+
+ (* 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.
+ omega.
+ - 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.
+ omega.
+ + 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.
+ omega.
+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. omega. }
+ 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. omega.
+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 omega.
+ rewrite list_length_z_nat in H2.
+ assert (Datatypes.length l0 = 0%nat) by omega.
+ 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; 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. omega.
+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 omega.
+ 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 as [l2|l1]; try congruence. rewrite !list_length_z_cons in HSIZE.
+ assert (list_length_z l2 >= 0) by eapply list_length_z_pos.
+ assert (list_length_z l2 + 1 + 1 >= 2) by omega.
+ assert (2 <= 1) by omega. contradiction H1. omega.
+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. omega.
+ * 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 omega;
+ 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 omega.
+ rewrite Ptrofs.unsigned_one.
+ replace (size bb - 1 + 1) with (size bb) by omega.
+ 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 omega.
+ replace (size bb - 1 + 1) with (size bb) by omega. 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 omega. rewrite Ptrofs.unsigned_repr; try omega.
+ 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) # X17 <- 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) # X17 <- Vundef) # PC <- (Vptr b0 (Ptrofs.repr z))) with
+ ((((rs1 # PC <- (Vptr b0 (Ptrofs.add i1 (Ptrofs.repr (size bb))))) # X16 <-
+ Vundef) # X17 <- 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.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 X17); subst.
+ * rewrite !Pregmap.gss; auto.
+ * rewrite !(Pregmap.gso (i := x) (j:= X17)); 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). omega.
}
-
- (* Mload notrap1 *)
- inv AT. simpl in *. unfold bind in *. destruct (transl_code _ _ _) in *; discriminate.
-
-- (* Mload notrap *)
- inv AT. simpl in *. unfold bind in *. destruct (transl_code _ _ _) in *; discriminate.
-
-- (* Mload notrap *)
- inv AT. simpl in *. unfold bind in *. destruct (transl_code _ _ _) in *; discriminate.
-
-- (* 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 R]]].
- exists rs2; split. eauto.
- split. eapply agree_undef_regs; eauto with asmgen.
- split. simpl; congruence.
- rewrite R. assumption.
-
-- (* 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_2.
- rewrite <- H1. simpl. econstructor; eauto.
- eapply code_tail_next_int; eauto.
- rewrite preg_notin_charact. intros. auto with asmgen.
- auto with asmgen.
- apply agree_nextinstr. eapply agree_set_res; auto.
- eapply agree_undef_regs; eauto. intros. rewrite undef_regs_other_2; auto.
- congruence.
-
- Simpl.
- rewrite set_res_other by trivial.
- rewrite undef_regs_other.
- assumption.
- intro.
- rewrite in_map_iff.
- intros (x0 & PREG & IN).
- subst r'.
- intro.
- apply (preg_of_not_RA x0).
- 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.
-
- rewrite INV by congruence.
- assumption.
-
-- (* 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 & D).
- exists jmp; exists k; exists rs'.
- split. eexact A.
- split. apply agree_exten with rs0; auto with asmgen.
- split.
- exact B.
- rewrite D. exact LEAF.
-
-- (* 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 & D).
- 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.
- split.
- simpl; congruence.
- Simpl. rewrite D.
- exact LEAF.
-
-- (* 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 #X17 <- 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.
-
- rewrite C by congruence.
- repeat rewrite Pregmap.gso by congruence.
- assumption.
-
-- (* 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 & W).
- 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. omega. constructor.
- intros (ofs' & X & Y).
- left; exists (State rs3 m3'); split.
- eapply exec_straight_steps_1; eauto. omega. 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.
-
- rewrite W.
- unfold rs2.
- Simpl.
-
-- (* 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. omega. split. auto.
- rewrite <- ATPC in H5.
- econstructor; eauto. congruence.
- inv WF.
- inv STACK.
- inv H1.
- congruence.
+ remember (list_nth_z (header bb) (size bb - 1)) as label_opt; destruct label_opt; auto;
+ exploit list_nth_z_range; eauto; omega.
Qed.
-Lemma transf_initial_states:
- forall st1, Mach.initial_state prog st1 ->
- exists st2, Asm.initial_state tprog st2 /\ match_states st1 st2.
+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. inversion H. unfold ge0 in *.
- econstructor; split.
- econstructor.
- eapply (Genv.init_mem_transf_partial TRANSF); eauto.
- replace (Genv.symbol_address (Genv.globalenv tprog) (prog_main tprog) Ptrofs.zero)
- with (Vptr fb Ptrofs.zero).
- econstructor; eauto.
- constructor.
- apply Mem.extends_refl.
- split. auto. simpl. unfold Vnullptr; destruct Archi.ptr64; congruence.
- intros. rewrite Regmap.gi. auto.
- unfold Genv.symbol_address.
- rewrite (match_program_main TRANSF).
- rewrite symbols_preserved.
- unfold ge; rewrite H1. auto.
+ intros.
+ exploit internal_functions_unfold; eauto.
+ intros (tc & FINDtf & TRANStf & _).
+
+ exploit (find_instr_bblock (size bb - 1)); eauto.
+ { generalize (bblock_size_pos bb). omega. }
+ 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 omega.
+ rewrite list_nth_z_range_exceeded in H; try omega. 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); omega. }
+ 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 *)
+ (* TODO: ligne ci-dessous à remplacer avec un intro (rs2' & m2' & ...) *)
+ intro H3; destruct H3; destruct H3; destruct H3. rewrite H5.
+ inversion MATCHI.
+ apply plus_one.
+ eapply Asm.exec_step_internal; eauto.
+ - rewrite <- AGPC. rewrite ATPC. unfold Val.offset_ptr. eauto.
+ - 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 omega.
+ rewrite Ptrofs.unsigned_repr; omega.
+ * (* 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.
+ - inv MATCHI. rewrite <- AGPC. rewrite ATPC. unfold Val.offset_ptr. eauto.
+ - (* XXX copy-paste from case above *)
+ simpl. replace (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr (size bb - 1))))
+ with (Ptrofs.unsigned ofs + (size bb - 1)); eauto.
+ generalize (bblock_size_pos bb); generalize (Ptrofs.unsigned_range_2 ofs); intros.
+ unfold Ptrofs.add.
+ rewrite Ptrofs.unsigned_repr. rewrite Ptrofs.unsigned_repr; try omega.
+ rewrite Ptrofs.unsigned_repr; omega.
+ - simpl. (* TODO something like eval_builtin_args_match *) admit.
+ - inv MATCHI; eauto.
+ - rewrite H11. (* TODO lemmas for set_res, undef_regs ... *) admit.
+Admitted.
+
+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 omega. reflexivity.
+ - rewrite AG; try assumption. reflexivity.
+ }
+ destruct X.
+ subst; eapply star_refl; eauto.
Qed.
-Lemma transf_final_states:
- forall st1 st2 r,
- match_states st1 st2 -> Mach.final_state st1 r -> Asm.final_state st2 r.
+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. inv H0. inv H. constructor. assumption.
- compute in H1. inv H1.
- generalize (preg_val _ _ _ R0 AG). rewrite H2. intros LD; inv LD. auto.
+ 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.
+ - (* TODO step_simulation *)
+ 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 PseudoAsmblockproof.match_prog
+ ::: mkpass Asmblockgenproof.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. inversion H2. remember (Machblockgen.transf_program p) as mbp.
+ unfold match_prog; simpl.
+ exists mbp; split. apply Machblockgenproof.transf_program_match; auto.
+ exists x; split. apply PseudoAsmblockproof.transf_program_match; auto.
+ exists x0; split. apply Asmblockgenproof.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 (PseudoAsmblockproof.rao Asmblockgenproof.next).
+
+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; eapply Machblockgenproof.Mach_return_address_exists; eauto.
+Admitted.
+
+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.
- eapply forward_simulation_star with (measure := measure)
- (match_states := fun S1 S2 => match_states S1 S2 /\ wf_state ge S1).
- - apply senv_preserved.
- - simpl; intros. exploit transf_initial_states; eauto.
- intros (s2 & A & B).
- exists s2; intuition auto. apply wf_initial; auto.
- - simpl; intros. destruct H as [MS WF]. eapply transf_final_states; eauto.
- - simpl; intros. destruct H0 as [MS WF].
- exploit step_simulation; eauto. intros [ (s2' & A & B) | (A & B & C) ].
- + left; exists s2'; intuition auto. eapply wf_step; eauto.
- + right; intuition auto. eapply wf_step; eauto.
+ 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 PseudoAsmblockproof.transf_program_correct; eauto.
+ - intros; apply Asmblockgenproof.next_progress.
+ - intros; eapply Asmblockgenproof.functions_bound_max_pos; eauto.
+ { intros; eapply Asmblock_PRESERVATION.symbol_high_low; eauto. }
+ + eapply compose_forward_simulations. apply Asmblockgenproof.transf_program_correct; eauto.
+ { intros; eapply Asmblock_PRESERVATION.symbol_high_low; 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/configure b/configure
index d0bbd0c1..a73e7879 100755
--- a/configure
+++ b/configure
@@ -696,6 +696,9 @@ echo "-R lib compcert.lib \
-R cparser compcert.cparser \
-R MenhirLib compcert.MenhirLib" > _CoqProject
case $arch in
+ aarch64)
+ echo "-R kvx/lib compcert.kvx.lib -R kvx/abstractbb compcert.kvx.abstractbb" >> _CoqProject # import KVX libraries for aarch64
+ ;;
x86)
echo "-R x86_${bitsize} compcert.x86_${bitsize}" >> _CoqProject
;;
@@ -838,13 +841,23 @@ RESPONSEFILE="none"
EOF
fi
+if [ "$arch" = "aarch64" ]; then # for aarch64 scheduling
+cat >> Makefile.config <<EOF
+ARCHDIRS=$arch kvx/lib kvx/abstractbb kvx/abstractbb/Impure
+BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v OptionMonad.v IterList.v PseudoAsmblock.v PseudoAsmblockproof.v ForwardSimulationBlock.v\\
+ AbstractBasicBlocksDef.v SeqSimuTheory.v ImpSimuTest.v\\
+ ImpConfig.v ImpCore.v ImpExtern.v ImpHCons.v ImpIO.v ImpLoops.v ImpMonads.v ImpPrelude.v\\
+ Asmblock.v Asmblockgen.v Asmblockgenproof.v # TODO: UPDATE THIS
+EOF
+fi
+
if [ "$arch" = "kvx" ]; then
cat >> Makefile.config <<EOF
ARCHDIRS=$arch $arch/lib $arch/abstractbb $arch/abstractbb/Impure
EXECUTE=kvx-cluster --syscall=libstd_scalls.so --
CFLAGS= -D __KVX_COS__
SIMU=kvx-cluster --
-BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v\\
+BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v OptionMonad.v IterList.v PseudoAsmblock.v PseudoAsmblockproof.v\\
Asmblock.v Asmblockgen.v Asmblockgenproof0.v Asmblockgenproof1.v Asmblockgenproof.v Asmvliw.v Asmblockprops.v\\
ForwardSimulationBlock.v PostpassScheduling.v PostpassSchedulingproof.v\\
Asmblockdeps.v DecBoolOps.v Chunks.v Peephole.v ExtValues.v ExtFloats.v\\
diff --git a/kvx/lib/IterList.v b/kvx/lib/IterList.v
new file mode 100644
index 00000000..49beb1c5
--- /dev/null
+++ b/kvx/lib/IterList.v
@@ -0,0 +1,96 @@
+Require Import Coqlib.
+
+(** TODO: are these def and lemma already defined in the standard library ?
+
+In this case, it should be better to reuse those of the standard library !
+
+*)
+
+Fixpoint iter {A} (n:nat) (f: A -> A) (x: A) {struct n}: A :=
+ match n with
+ | O => x
+ | S n0 => iter n0 f (f x)
+ end.
+
+Lemma iter_S A (n:nat) (f: A -> A): forall x, iter (S n) f x = f (iter n f x).
+Proof.
+ induction n; simpl; auto.
+ intros; erewrite <- IHn; simpl; auto.
+Qed.
+
+Lemma iter_plus A (n m:nat) (f: A -> A): forall x, iter (n+m) f x = iter m f (iter n f x).
+Proof.
+ induction n; simpl; auto.
+Qed.
+
+Definition iter_tail {A} (n:nat) (l: list A) := iter n (@tl A) l.
+
+Lemma iter_tail_S {A} (n:nat) (l: list A): iter_tail (S n) l = tl (iter_tail n l).
+Proof.
+ apply iter_S.
+Qed.
+
+Lemma iter_tail_plus A (n m:nat) (l: list A): iter_tail (n+m) l = iter_tail m (iter_tail n l).
+Proof.
+ apply iter_plus.
+Qed.
+
+Lemma iter_tail_length A l1: forall (l2: list A), iter_tail (length l1) (l1 ++ l2) = l2.
+Proof.
+ induction l1; auto.
+Qed.
+
+Lemma iter_tail_nil A n: @iter_tail A n nil = nil.
+Proof.
+ unfold iter_tail; induction n; simpl; auto.
+Qed.
+
+Lemma iter_tail_reach_nil A (l: list A): iter_tail (length l) l = nil.
+Proof.
+ rewrite (app_nil_end l) at 2.
+ rewrite iter_tail_length.
+ auto.
+Qed.
+
+Lemma length_iter_tail {A} (n:nat): forall (l: list A), (n <= List.length l)%nat -> (List.length l = n + List.length (iter_tail n l))%nat.
+Proof.
+ unfold iter_tail; induction n; auto.
+ intros l; destruct l. { simpl; omega. }
+ intros; simpl. erewrite IHn; eauto.
+ simpl in *; omega.
+Qed.
+
+Lemma iter_tail_S_ex {A} (n:nat): forall (l: list A), (n < length l)%nat -> exists x, iter_tail n l = x::(iter_tail (S n) l).
+Proof.
+ unfold iter_tail; induction n; simpl.
+ - intros l; destruct l; simpl; omega || eauto.
+ - intros l H; destruct (IHn (tl l)) as (x & H1).
+ + destruct l; simpl in *; try omega.
+ + rewrite H1; eauto.
+Qed.
+
+Lemma iter_tail_inject1 {A} (n1 n2:nat) (l: list A): (n1 <= List.length l)%nat -> (n2 <= List.length l)%nat -> iter_tail n1 l = iter_tail n2 l -> n1=n2.
+Proof.
+ intros H1 H2 EQ; exploit (length_iter_tail n1 l); eauto.
+ rewrite EQ.
+ rewrite (length_iter_tail n2 l); eauto.
+ omega.
+Qed.
+
+Lemma iter_tail_nil_inject {A} (n:nat) (l: list A): iter_tail n l = nil -> (List.length l <= n)%nat.
+Proof.
+ destruct (le_lt_dec n (List.length l)); try omega.
+ intros; exploit (iter_tail_inject1 n (length l) l); try omega.
+ rewrite iter_tail_reach_nil. auto.
+Qed.
+
+Lemma list_length_z_nat (A: Type) (l: list A): list_length_z l = Z.of_nat (length l).
+Proof.
+ induction l; auto.
+ rewrite list_length_z_cons. simpl. rewrite Zpos_P_of_succ_nat. omega.
+Qed.
+
+Lemma list_length_nat_z (A: Type) (l: list A): length l = Z.to_nat (list_length_z l).
+Proof.
+ intros; rewrite list_length_z_nat, Nat2Z.id. auto.
+Qed.
diff --git a/kvx/lib/OptionMonad.v b/kvx/lib/OptionMonad.v
new file mode 100644
index 00000000..824a9c2f
--- /dev/null
+++ b/kvx/lib/OptionMonad.v
@@ -0,0 +1,49 @@
+(* Declare Scope option_monad_scope. *)
+
+Notation "'SOME' X <- A 'IN' B" := (match A with Some X => B | None => None end)
+ (at level 200, X ident, A at level 100, B at level 200)
+ : option_monad_scope.
+
+Notation "'ASSERT' A 'IN' B" := (if A then B else None)
+ (at level 200, A at level 100, B at level 200)
+ : option_monad_scope.
+
+Local Open Scope option_monad_scope.
+
+
+(** Simple tactics for option-monad *)
+
+Lemma destruct_SOME A B (P: option B -> Prop) (e: option A) (f: A -> option B):
+ (forall x, e = Some x -> P (f x)) -> (e = None -> P None) -> (P (SOME x <- e IN f x)).
+Proof.
+ intros; destruct e; simpl; auto.
+Qed.
+
+Lemma destruct_ASSERT B (P: option B -> Prop) (e: bool) (x: option B):
+ (e = true -> P x) -> (e = false -> P None) -> (P (ASSERT e IN x)).
+Proof.
+ intros; destruct e; simpl; auto.
+Qed.
+
+Ltac inversion_SOME x :=
+ try (eapply destruct_SOME; [ let x := fresh x in intro x | simpl; try congruence ]).
+
+Ltac inversion_ASSERT :=
+ try (eapply destruct_ASSERT; [ idtac | simpl; try congruence ]).
+
+Ltac simplify_someHyp :=
+ match goal with
+ | H: None = Some _ |- _ => inversion H; clear H; subst
+ | H: Some _ = None |- _ => inversion H; clear H; subst
+ | H: ?t = ?t |- _ => clear H
+ | H: Some _ = Some _ |- _ => inversion H; clear H; subst
+ | H: Some _ <> None |- _ => clear H
+ | H: None <> Some _ |- _ => clear H
+ | H: _ = Some _ |- _ => (try rewrite !H in * |- *); generalize H; clear H
+ end.
+
+Ltac simplify_someHyps :=
+ repeat (simplify_someHyp; simpl in * |- *).
+
+Ltac try_simplify_someHyps :=
+ try (intros; simplify_someHyps; eauto).
diff --git a/kvx/lib/PseudoAsmblock.v b/kvx/lib/PseudoAsmblock.v
new file mode 100644
index 00000000..b33ea1bd
--- /dev/null
+++ b/kvx/lib/PseudoAsmblock.v
@@ -0,0 +1,267 @@
+Require Import Coqlib Maps AST Integers Values Memory Events Globalenvs Smallstep.
+Require Import Op Machregs Locations Stacklayout Conventions.
+Require Import Mach Machblock OptionMonad.
+
+
+(** Registers and States *)
+
+Inductive preg: Type :=
+ | PC: preg (* program counter *)
+ | RA: preg (* return-address *)
+ | SP: preg (* stack-pointer *)
+ | preg_of: mreg -> preg.
+Coercion preg_of: mreg >-> preg.
+
+Lemma preg_eq: forall (x y: preg), {x=y} + {x<>y}.
+Proof.
+ decide equality.
+ apply mreg_eq.
+Defined.
+
+Module PregEq.
+ Definition t := preg.
+ Definition eq := preg_eq.
+End PregEq.
+
+Module Pregmap := EMap(PregEq).
+
+Definition regset := Pregmap.t val.
+
+Module AsmNotations.
+
+(* Declare Scope asm. *)
+Notation "a # b" := (a b) (at level 1, only parsing) : asm.
+Notation "a # b <- c" := (Pregmap.set b c a) (at level 1, b at next level) : asm.
+Open Scope asm.
+
+End AsmNotations.
+
+Import AsmNotations.
+
+Definition to_Machrs (rs: regset): Mach.regset :=
+ fun (r:mreg) => rs r.
+Coercion to_Machrs: regset >-> Mach.regset.
+
+Definition set_from_Machrs (mrs: Mach.regset) (rs: regset): regset :=
+ fun (r:preg) =>
+ match r with
+ | preg_of mr => mrs mr
+ | _ => rs r
+ end.
+
+Local Open Scope option_monad_scope.
+
+Record state: Type := State { _rs: regset; _m: mem }.
+Definition outcome := option state.
+Definition Next rs m: outcome := Some (State rs m).
+Definition Stuck: outcome := None.
+
+(* Asm semantic on Mach *)
+
+Section RELSEM.
+
+(** "oracle" stating the successive position (ie offset) of each block in the final assembly function *)
+Variable next: function -> Z -> Z.
+
+Inductive is_pos (f: function): Z -> code -> Prop :=
+ | First_pos: is_pos f 0 (fn_code f)
+ | Next_pos pos b c: is_pos f pos (b::c) -> is_pos f (next f pos) c.
+
+Fixpoint label_pos (f: function) (lbl: label) (pos: Z) (c: code) {struct c} : option Z :=
+ match c with
+ | nil => None
+ | b :: c' =>
+ if is_label lbl b then Some pos else label_pos f lbl (next f pos) c'
+ end.
+
+Definition goto_label (f: function) (lbl: label) (rs: regset) : option val :=
+ SOME pos <- label_pos f lbl 0 (fn_code f) IN
+ match rs#PC with
+ | Vptr b _ => Some (Vptr b (Ptrofs.repr pos))
+ | _ => None
+ end.
+
+Definition next_addr (f: function) (rs: regset): option val :=
+ match rs#PC with
+ | Vptr b ofs => Some (Vptr b (Ptrofs.repr (next f (Ptrofs.unsigned ofs))))
+ | _ => None
+ end.
+
+Variable ge:genv.
+
+Inductive basic_step (f: function) (rs: regset) (m:mem): basic_inst -> regset -> mem -> Prop :=
+ | exec_MBgetstack:
+ forall ofs ty dst v,
+ load_stack m (rs#SP) ty ofs = Some v ->
+ basic_step f rs m (MBgetstack ofs ty dst) (rs#dst <- v) m
+ | exec_MBsetstack:
+ forall (src: mreg) ofs ty m' rs',
+ store_stack m (rs#SP) ty ofs (rs src) = Some m' ->
+ rs' = set_from_Machrs (undef_regs (destroyed_by_setstack ty) rs) rs ->
+ basic_step f rs m (MBsetstack src ofs ty) rs' m'
+ | exec_MBgetparam:
+ forall ofs ty (dst: mreg) v rs' psp,
+ load_stack m (rs#SP) Tptr f.(fn_link_ofs) = Some psp ->
+ load_stack m psp ty ofs = Some v ->
+ rs' = (rs # temp_for_parent_frame <- Vundef # dst <- v) ->
+ basic_step f rs m (MBgetparam ofs ty dst) rs' m
+ | exec_MBop:
+ forall op args v rs' (res: mreg),
+ eval_operation ge (rs#SP) op (to_Machrs rs)##args m = Some v ->
+ rs' = (set_from_Machrs (undef_regs (destroyed_by_op op) rs) rs)#res <- v ->
+ basic_step f rs m (MBop op args res) rs' m
+ | exec_MBload:
+ forall addr args a v rs' trap chunk (dst: mreg),
+ eval_addressing ge (rs#SP) addr (to_Machrs rs)##args = Some a ->
+ Mem.loadv chunk m a = Some v ->
+ rs' = (set_from_Machrs (undef_regs (destroyed_by_load chunk addr) rs) rs)#dst <- v ->
+ basic_step f rs m (MBload trap chunk addr args dst) rs' m
+ | exec_MBload_notrap1:
+ forall addr args rs' chunk (dst: mreg),
+ eval_addressing ge (rs#SP) addr (to_Machrs rs)##args = None ->
+ rs' = (set_from_Machrs (undef_regs (destroyed_by_load chunk addr) rs) rs)#dst <- (default_notrap_load_value chunk) ->
+ basic_step f rs m (MBload NOTRAP chunk addr args dst) rs' m
+ | exec_MBload_notrap2:
+ forall addr args a rs' chunk (dst: mreg),
+ eval_addressing ge (rs#SP) addr (to_Machrs rs)##args = Some a ->
+ Mem.loadv chunk m a = None ->
+ rs' = (set_from_Machrs (undef_regs (destroyed_by_load chunk addr) rs) rs)#dst <- (default_notrap_load_value chunk) ->
+ basic_step f rs m (MBload NOTRAP chunk addr args dst) rs' m
+ | exec_MBstore:
+ forall chunk addr args (src: mreg) m' a rs',
+ eval_addressing ge (rs#SP) addr (to_Machrs rs)##args = Some a ->
+ Mem.storev chunk m a (rs src) = Some m' ->
+ rs' = set_from_Machrs (undef_regs (destroyed_by_store chunk addr) rs) rs ->
+ basic_step f rs m (MBstore chunk addr args src) rs' m'
+ .
+
+
+Inductive body_step (f: function) : bblock_body -> regset -> mem -> regset -> mem -> Prop :=
+ | exec_nil_body rs m:
+ body_step f nil rs m rs m
+ | exec_cons_body rs m bi p rs' m' rs'' m'':
+ basic_step f rs m bi rs' m' ->
+ body_step f p rs' m' rs'' m'' ->
+ body_step f (bi::p) rs m rs'' m''
+ .
+
+
+Definition find_function_ptr (ros: mreg + ident) (rs: regset) : option val :=
+ match ros with
+ | inl r => Some (rs#r)
+ | inr symb =>
+ SOME b <- Genv.find_symbol ge symb IN
+ Some (Vptr b Ptrofs.zero)
+ end.
+
+Definition exec_epilogue (f: function) (rs: regset) (m: mem) : outcome :=
+ SOME sp <- load_stack m rs#SP Tptr f.(fn_link_ofs) IN
+ SOME ra <- load_stack m rs#SP Tptr f.(fn_retaddr_ofs) IN
+ match rs#SP with
+ | Vptr stk ofs =>
+ SOME m' <- Mem.free m stk 0 f.(fn_stacksize) IN
+ Next (rs#SP <- sp #RA <- ra) m'
+ | _ => Stuck
+ end.
+
+Inductive cfi_step (f: function): control_flow_inst -> regset -> mem -> trace -> regset -> mem -> Prop :=
+ | exec_MBcall sig ros rs m pc:
+ find_function_ptr ros rs = Some pc ->
+ cfi_step f (MBcall sig ros) rs m E0 ((rs#RA <- (rs#PC))#PC <- pc) m
+ | exec_MBtailcall sig ros rs m rs' m' pc:
+ find_function_ptr ros rs = Some pc ->
+ exec_epilogue f rs m = Next rs' m' ->
+ cfi_step f (MBtailcall sig ros) rs m E0 (rs'#PC <- pc) m'
+ | exec_MBreturn rs m rs' m':
+ exec_epilogue f rs m = Next rs' m' ->
+ cfi_step f MBreturn rs m E0 (rs'#PC <- (rs'#RA)) m'
+ | exec_MBgoto lbl rs m pc:
+ goto_label f lbl rs = Some pc ->
+ cfi_step f (MBgoto lbl) rs m E0 (rs#PC <- pc) m
+ | exec_MBcond_true cond args lbl rs m pc rs':
+ eval_condition cond (to_Machrs rs)##args m = Some true ->
+ goto_label f lbl rs = Some pc ->
+ rs' = set_from_Machrs (undef_regs (destroyed_by_cond cond) rs) rs ->
+ cfi_step f (MBcond cond args lbl) rs m E0 (rs'#PC <- pc) m
+ | exec_MBcond_false cond args lbl rs m rs':
+ eval_condition cond (to_Machrs rs)##args m = Some false ->
+ rs' = set_from_Machrs (undef_regs (destroyed_by_cond cond) rs) rs ->
+ cfi_step f (MBcond cond args lbl) rs m E0 rs' m
+ | exec_MBjumptable arg tbl rs m index lbl pc rs':
+ to_Machrs rs arg = Vint index ->
+ list_nth_z tbl (Int.unsigned index) = Some lbl ->
+ goto_label f lbl rs = Some pc ->
+ rs' = set_from_Machrs (undef_regs destroyed_by_jumptable rs) rs ->
+ cfi_step f (MBjumptable arg tbl) rs m E0 (rs'#PC <- pc) m
+ | exec_MBbuiltin rs m ef args res vargs t vres rs' m':
+ eval_builtin_args ge (to_Machrs rs) (rs#SP) m args vargs ->
+ external_call ef ge vargs m t vres m' ->
+ rs' = set_from_Machrs (set_res res vres (undef_regs (destroyed_by_builtin ef) rs)) rs ->
+ cfi_step f (MBbuiltin ef args res) rs m t rs' m'
+ .
+
+Inductive exit_step f : option control_flow_inst -> regset -> mem -> trace -> regset -> mem -> Prop :=
+ | exec_Some_exit ctl rs m t rs' m':
+ cfi_step f ctl rs m t rs' m' ->
+ exit_step f (Some ctl) rs m t rs' m'
+ | exec_None_exit rs m:
+ exit_step f None rs m E0 rs m
+ .
+
+Inductive exec_bblock f: bblock -> regset -> mem -> trace -> regset -> mem -> Prop :=
+ | exec_bblock_all (bb: bblock) rs0 m0 rs1 m1 pc t rs2 m2:
+ body_step f (body bb) rs0 m0 rs1 m1 ->
+ next_addr f rs1 = Some pc ->
+ exit_step f (exit bb) (rs1#PC <- pc) m1 t rs2 m2 ->
+ exec_bblock f bb rs0 m0 t rs2 m2.
+
+Definition exec_prologue f (pos:Z) (rs: regset) (m:mem) : option state :=
+ if Z.eq_dec pos 0 then
+ let (m1, stk) := Mem.alloc m 0 f.(fn_stacksize) in
+ let sp := Vptr stk Ptrofs.zero in
+ SOME m2 <- store_stack m1 sp Tptr f.(fn_link_ofs) rs#SP IN
+ SOME m3 <- store_stack m2 sp Tptr f.(fn_retaddr_ofs) rs#RA IN
+ Next ((set_from_Machrs (undef_regs destroyed_at_function_entry rs) rs)#SP <- sp) m3
+ else
+ Next rs m.
+
+
+Inductive step: state -> trace -> state -> Prop :=
+ | exec_step_internal b ofs f bb c rs0 m0 rs1 m1 t rs2 m2:
+ rs0 PC = Vptr b ofs ->
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
+ is_pos f (Ptrofs.unsigned ofs) (bb::c) ->
+ exec_prologue f (Ptrofs.unsigned ofs) rs0 m0 = Next rs1 m1 ->
+ exec_bblock f bb rs1 m1 t rs2 m2 ->
+ step (State rs0 m0) t (State rs2 m2)
+ | exec_step_external:
+ forall b ef args res rs m t rs' m',
+ rs PC = Vptr b Ptrofs.zero ->
+ Genv.find_funct_ptr ge b = Some (External ef) ->
+ extcall_arguments (to_Machrs rs) m (rs#SP) (ef_sig ef) args ->
+ external_call ef ge args m t res m' ->
+ rs' = (set_from_Machrs (set_pair (loc_result (ef_sig ef)) res (undef_caller_save_regs rs)) rs) #PC <- (rs RA) ->
+ step (State rs m) t (State rs' m').
+
+End RELSEM.
+
+Inductive initial_state (p: program): state -> Prop :=
+ | initial_state_intro: forall m0,
+ let ge := Genv.globalenv p in
+ Genv.init_mem p = Some m0 ->
+ let rs0 :=
+ (Pregmap.init Vundef)
+ # PC <- (Genv.symbol_address ge p.(prog_main) Ptrofs.zero)
+ # SP <- Vnullptr
+ # RA <- Vnullptr in
+ initial_state p (State rs0 m0).
+
+Inductive final_state: state -> int -> Prop :=
+ | final_state_intro: forall rs m r retcode,
+ loc_result signature_main = One r ->
+ rs PC = Vnullptr ->
+ rs r = Vint retcode ->
+ final_state (State rs m) retcode.
+
+Definition semantics (next: function -> Z -> Z) (p: program) :=
+ Semantics (step next) (initial_state p) final_state (Genv.globalenv p).
diff --git a/kvx/lib/PseudoAsmblockproof.v b/kvx/lib/PseudoAsmblockproof.v
new file mode 100644
index 00000000..67308278
--- /dev/null
+++ b/kvx/lib/PseudoAsmblockproof.v
@@ -0,0 +1,1173 @@
+Require Import Coqlib Maps AST Integers Values Memory Events Globalenvs Smallstep.
+Require Import Op Machregs Locations Stacklayout Conventions.
+Require Import Mach Machblock OptionMonad.
+Require Import Errors Datatypes PseudoAsmblock IterList.
+
+(** Tiny translation from Machblock semantics to PseudoAsmblock semantics (needs additional checks)
+*)
+
+Section TRANSLATION.
+
+(* In the actual Asmblock code, the prologue will be inserted in the first block of the function.
+ But, this block should have an empty header.
+*)
+
+Definition has_header (c: code) : bool :=
+ match c with
+ | nil => false
+ | bb::_ => match header bb with
+ | nil => false
+ | _ => true
+ end
+ end.
+
+Definition insert_implicit_prologue c :=
+ if has_header c then {| header := nil; body := nil; exit := None |}::c else c.
+
+Definition transl_function (f: function) : function :=
+ {| fn_sig:=fn_sig f;
+ fn_code:=insert_implicit_prologue (fn_code f);
+ fn_stacksize := fn_stacksize f;
+ fn_link_ofs := fn_link_ofs f;
+ fn_retaddr_ofs := fn_retaddr_ofs f
+ |}.
+
+Definition transf_function (f: function) : res function :=
+ let tf := transl_function f in
+ (* removed because it is simpler or/and more efficient to perform this test in Asmblockgen !
+ if zlt Ptrofs.max_unsigned (max_pos tf)
+ then Error (msg "code size exceeded")
+ else *)
+ OK tf.
+
+Definition transf_fundef (f: fundef) : res fundef :=
+ transf_partial_fundef transf_function f.
+
+Definition transf_program (p: program) : res program :=
+ transform_partial_program transf_fundef p.
+
+End TRANSLATION.
+
+(** Proof of the translation
+*)
+
+Require Import Linking.
+Import PseudoAsmblock.AsmNotations.
+
+Section PRESERVATION.
+
+Definition match_prog (p: program) (tp: 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.
+
+Local Open Scope Z_scope.
+Local Open Scope option_monad_scope.
+
+Variable prog: program.
+Variable tprog: program.
+
+Hypothesis TRANSF: match_prog prog tprog.
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+Variable next: function -> Z -> Z.
+
+Hypothesis next_progress: forall (f:function) (pos:Z), (pos < (next f pos))%Z.
+
+Definition max_pos (f:function) := iter (S (length f.(fn_code))) (next f) 0.
+
+(* This hypothesis expresses that Asmgen checks for each tf
+ that (max_pos tf) represents a valid address
+*)
+Hypothesis functions_bound_max_pos: forall fb tf,
+ Genv.find_funct_ptr tge fb = Some (Internal tf) ->
+ (max_pos tf) <= Ptrofs.max_unsigned.
+
+(** * Agreement between Mach registers and PseudoAsm registers *)
+Record agree (ms: Mach.regset) (sp: val) (rs: regset) : Prop := mkagree {
+ agree_sp: rs#SP = sp;
+ agree_sp_def: sp <> Vundef;
+ agree_mregs: forall r: mreg, (ms r)=(rs#r)
+}.
+
+(** [transl_code_at_pc pc fb f tf c] holds if the code pointer [pc] points
+ within the code generated by Machblock function (originally [f] -- but translated as [tf]),
+ and [c] is the tail of the code at the position corresponding to the code pointer [pc].
+*)
+Inductive transl_code_at_pc (b:block) (f:function) (tf:function) (c:code): val -> Prop :=
+ transl_code_at_pc_intro ofs:
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
+ transf_function f = OK tf ->
+ (* we have passed the first block containing the prologue *)
+ (0 < (Ptrofs.unsigned ofs))%Z ->
+ (* the code is identical in the two functions *)
+ is_pos next tf (Ptrofs.unsigned ofs) c ->
+ transl_code_at_pc b f tf c (Vptr b ofs).
+
+Inductive match_stack: list stackframe -> Prop :=
+ | match_stack_nil:
+ match_stack nil
+ | match_stack_cons: forall fb sp ra c s f tf,
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ transl_code_at_pc fb f tf c ra ->
+ sp <> Vundef ->
+ match_stack s ->
+ match_stack (Stackframe fb sp ra c :: s).
+
+(** Semantic preservation is proved using simulation diagrams
+ of the following form.
+<<
+ s1 ---------------- s2
+ | |
+ t| *|t
+ | |
+ v v
+ s1'---------------- s2'
+>>
+ The invariant is the [match_states] predicate below...
+
+*)
+
+Inductive match_states: Machblock.state -> state -> Prop :=
+ | match_states_internal s fb sp c ms m rs f tf
+ (STACKS: match_stack s)
+ (FIND: Genv.find_funct_ptr ge fb = Some (Internal f))
+ (AT: transl_code_at_pc fb f tf c (rs PC))
+ (AG: agree ms sp rs):
+ match_states (Machblock.State s fb sp c ms m)
+ (State rs m)
+ | match_states_prologue s sp fb ms rs0 m0 f rs1 m1
+ (STACKS: match_stack s)
+ (AG: agree ms sp rs1)
+ (ATPC: rs0 PC = Vptr fb Ptrofs.zero)
+ (ATLR: rs0 RA = parent_ra s)
+ (FIND: Genv.find_funct_ptr ge fb = Some (Internal f))
+ (PROL: exec_prologue f 0 rs0 m0 = Next rs1 m1):
+ match_states (Machblock.State s fb sp (fn_code f) ms m1)
+ (State rs0 m0)
+ | match_states_call s fb ms m rs
+ (STACKS: match_stack s)
+ (AG: agree ms (parent_sp s) rs)
+ (ATPC: rs PC = Vptr fb Ptrofs.zero)
+ (ATLR: rs RA = parent_ra s):
+ match_states (Machblock.Callstate s fb ms m)
+ (State rs m)
+ | match_states_return s ms m rs
+ (STACKS: match_stack s)
+ (AG: agree ms (parent_sp s) rs)
+ (ATPC: rs PC = parent_ra s):
+ match_states (Machblock.Returnstate s ms m)
+ (State rs m).
+
+Definition measure (s: Machblock.state) : nat :=
+ match s with
+ | Machblock.State _ _ _ _ _ _ => 0%nat
+ | Machblock.Callstate _ _ _ _ => 1%nat
+ | Machblock.Returnstate _ _ _ => 1%nat
+ end.
+
+Definition rao (f: function) (c: code) (ofs: ptrofs) : Prop :=
+ forall tf,
+ transf_function f = OK tf ->
+ is_pos next tf (Ptrofs.unsigned ofs) c.
+
+Lemma symbols_preserved:
+ forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
+Proof (Genv.find_symbol_match TRANSF).
+
+Lemma senv_preserved:
+ Senv.equiv ge tge.
+Proof (Genv.senv_match TRANSF).
+
+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 functions_transl fb f tf:
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ transf_function f = OK tf ->
+ Genv.find_funct_ptr tge fb = Some (Internal tf).
+Proof.
+ intros. exploit functions_translated; eauto. intros [tf' [A B]].
+ monadInv B. inv H0; auto.
+Qed.
+
+Lemma function_bound fb f tf:
+ Genv.find_funct_ptr ge fb = Some (Internal f) -> transf_function f = OK tf -> (max_pos tf) <= Ptrofs.max_unsigned.
+Proof.
+ intros; eapply functions_bound_max_pos; eauto.
+ eapply functions_transl; eauto.
+Qed.
+
+Lemma transf_function_def f tf:
+ transf_function f = OK tf -> tf.(fn_code) = insert_implicit_prologue f.(fn_code).
+Proof.
+ unfold transf_function.
+ intros EQ; inv EQ.
+ auto.
+Qed.
+
+Lemma stackinfo_preserved f tf:
+ transf_function f = OK tf ->
+ tf.(fn_stacksize) = f.(fn_stacksize)
+ /\ tf.(fn_retaddr_ofs) = f.(fn_retaddr_ofs)
+ /\ tf.(fn_link_ofs) = f.(fn_link_ofs).
+Proof.
+ unfold transf_function.
+ intros EQ0; inv EQ0. simpl; intuition.
+Qed.
+
+Lemma transf_initial_states st1: Machblock.initial_state prog st1 ->
+ exists st2, initial_state tprog st2 /\ match_states st1 st2.
+Proof.
+ intro H. inversion H. unfold ge0 in *.
+ econstructor; split.
+ - econstructor.
+ eapply (Genv.init_mem_transf_partial TRANSF); eauto.
+ - replace (Genv.symbol_address (Genv.globalenv tprog) (prog_main tprog) Ptrofs.zero)
+ with (Vptr fb Ptrofs.zero).
+ + econstructor; eauto.
+ * constructor.
+ * split; auto; simpl; unfold Vnullptr; destruct Archi.ptr64; congruence.
+ + unfold Genv.symbol_address.
+ rewrite (match_program_main TRANSF).
+ rewrite symbols_preserved.
+ unfold ge. simplify_someHyp. auto.
+Qed.
+
+Lemma transf_final_states st1 st2 r:
+ match_states st1 st2 -> Machblock.final_state st1 r -> final_state st2 r.
+Proof.
+ intros H H0. inv H0. inv H.
+ econstructor; eauto.
+ exploit agree_mregs; eauto.
+ erewrite H2. intro H3; inversion H3.
+ auto.
+Qed.
+
+(** Lemma on [is_pos]. *)
+
+Lemma is_pos_alt_def f pos code: is_pos next f pos code ->
+ exists n, (n <= List.length (fn_code f))%nat /\ pos = (iter n (next f) 0) /\ code = iter_tail n (fn_code f).
+Proof.
+ induction 1.
+ - unfold iter_tail; exists O; simpl; intuition.
+ - destruct IHis_pos as (n & H0 & H1 & H2).
+ exists (S n). repeat split.
+ + rewrite (length_iter_tail n); eauto.
+ rewrite <- H2; simpl; omega.
+ + rewrite iter_S; congruence.
+ + unfold iter_tail in *; rewrite iter_S, <- H2; auto.
+Qed.
+
+Local Hint Resolve First_pos Next_pos: core.
+
+Lemma is_pos_alt_def_recip f n: (n <= List.length (fn_code f))%nat ->
+ is_pos next f (iter n (next f) 0) (iter_tail n (fn_code f)).
+Proof.
+ induction n.
+ - unfold iter_tail; simpl; eauto.
+ - intros H; destruct (iter_tail_S_ex n (fn_code f)) as (x & H1); try omega.
+ rewrite iter_S; lapply IHn; try omega.
+ rewrite H1; eauto.
+Qed.
+
+Lemma is_pos_inject1 f pos1 pos2 code:
+ is_pos next f pos1 code -> is_pos next f pos2 code -> pos1=pos2.
+Proof.
+ intros H1 H2.
+ destruct (is_pos_alt_def f pos1 code) as (n1 & B1 & POS1 & CODE1); eauto.
+ destruct (is_pos_alt_def f pos2 code) as (n2 & B2 & POS2 & CODE2); eauto.
+ clear H1 H2; subst.
+ erewrite (iter_tail_inject1 n1 n2); eauto.
+Qed.
+
+Lemma iter_next_strict_monotonic f n m x: (n < m)%nat -> iter n (next f) x < iter m (next f) x.
+Proof.
+ induction 1; rewrite iter_S; auto.
+ generalize (next_progress f (iter m (next f) x)).
+ omega.
+Qed.
+
+Lemma iter_next_monotonic f n m x: (n <= m)%nat -> iter n (next f) x <= iter m (next f) x.
+Proof.
+ destruct 1.
+ - omega.
+ - generalize (iter_next_strict_monotonic f n (S m) x). omega.
+Qed.
+
+Lemma is_pos_bound_pos f pos code:
+ is_pos next f pos code -> 0 <= pos <= max_pos f.
+Proof.
+ intros H; exploit is_pos_alt_def; eauto.
+ intros (n & H1 & H2 & H3).
+ rewrite H2. unfold max_pos. split.
+ - cutrewrite (0 = iter O (next f) 0); auto.
+ apply iter_next_monotonic; omega.
+ - apply iter_next_monotonic; omega.
+Qed.
+
+Lemma is_pos_unsigned_repr f pos code:
+ is_pos next f pos code ->
+ max_pos f <= Ptrofs.max_unsigned ->
+ Ptrofs.unsigned (Ptrofs.repr pos) = pos.
+Proof.
+ intros; eapply Ptrofs.unsigned_repr.
+ exploit is_pos_bound_pos; eauto.
+ omega.
+Qed.
+
+Lemma is_pos_simplify f pos code:
+ is_pos next f pos code ->
+ max_pos f <= Ptrofs.max_unsigned ->
+ is_pos next f (Ptrofs.unsigned (Ptrofs.repr pos)) code.
+Proof.
+ intros; erewrite is_pos_unsigned_repr; eauto.
+Qed.
+
+Lemma find_label_label_pos f lbl c: forall pos c',
+ find_label lbl c = Some c' ->
+ exists n,
+ label_pos next f lbl pos c = Some (iter n (next f) pos)
+ /\ c' = iter_tail n c
+ /\ (n <= List.length c)%nat.
+Proof.
+ induction c.
+ - simpl; intros. discriminate.
+ - simpl; intros pos c'.
+ destruct (is_label lbl a).
+ + intro EQ; injection EQ; intro; subst c'.
+ exists O; simpl; intuition.
+ + intros. generalize (IHc (next f pos) c' H). intros (n' & A & B & C).
+ exists (S n'). intuition.
+Qed.
+
+Lemma find_label_insert_implicit_prologue lbl c:
+ find_label lbl c = find_label lbl (insert_implicit_prologue c).
+Proof.
+ unfold insert_implicit_prologue.
+ destruct (has_header c); simpl; auto.
+ unfold is_label; simpl.
+ destruct (in_dec lbl nil); auto.
+ simpl in *. tauto.
+Qed.
+
+Lemma no_header_insert_implicit_prologue c:
+ has_header (insert_implicit_prologue c) = false.
+Proof.
+ unfold insert_implicit_prologue.
+ destruct (has_header c) eqn: H; simpl; auto.
+Qed.
+
+Lemma find_label_has_header lbl c c':
+ find_label lbl c = Some c' ->
+ has_header c' = true.
+Proof.
+ induction c; simpl; try congruence.
+ destruct (is_label lbl a) eqn:LAB; auto.
+ intros X; inv X; simpl.
+ unfold is_label in LAB.
+ destruct (in_dec lbl (header a)); try congruence.
+ destruct (header a); try congruence.
+ simpl in *; tauto.
+Qed.
+
+Lemma find_label_label_pos_no_header f lbl c pos c':
+ (has_header c) = false ->
+ find_label lbl c = Some c' ->
+ exists n,
+ label_pos next f lbl pos c = Some (iter (S n) (next f) pos)
+ /\ c' = iter_tail (S n) c
+ /\ ((S n) <= List.length c)%nat.
+Proof.
+ intros H H0; exploit find_label_label_pos; eauto.
+ intros ([|n] & H1 & H2 & H3); try (exists n; intuition eauto).
+ unfold iter_tail in *; simpl in *; subst.
+ erewrite find_label_has_header in H; eauto.
+ congruence.
+Qed.
+
+Hint Resolve is_pos_simplify is_pos_alt_def_recip function_bound: core.
+
+Lemma find_label_goto_label f tf lbl rs c' b ofs:
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
+ transf_function f = OK tf ->
+ Vptr b ofs = rs PC ->
+ find_label lbl f.(fn_code) = Some c' ->
+ exists pc,
+ goto_label next tf lbl rs = Some pc
+ /\ transl_code_at_pc b f tf c' pc.
+Proof.
+ intros FINDF T HPC FINDL.
+ erewrite find_label_insert_implicit_prologue, <- transf_function_def in FINDL; eauto.
+ exploit find_label_label_pos_no_header; eauto.
+ { erewrite transf_function_def; eauto.
+ apply no_header_insert_implicit_prologue.
+ }
+ intros (n & LAB & CODE & BOUND); subst.
+ exists (Vptr b (Ptrofs.repr (iter (S n) (next tf) 0))).
+ unfold goto_label; intuition.
+ - simplify_someHyps; rewrite <- HPC. auto.
+ - econstructor; eauto.
+ erewrite is_pos_unsigned_repr; eauto.
+ generalize (iter_next_strict_monotonic tf O (S n) 0); simpl.
+ omega.
+Qed.
+
+(** Preservation of register agreement under various assignments. *)
+
+Lemma agree_mregs_list ms sp rs:
+ agree ms sp rs ->
+ forall l, (ms##l)=(to_Machrs rs)##l.
+Proof.
+ unfold to_Machrs. intros AG; induction l; simpl; eauto.
+ erewrite agree_mregs; eauto.
+ congruence.
+Qed.
+
+Lemma agree_set_mreg ms sp rs r v rs':
+ agree ms sp rs ->
+ v=(rs'#(preg_of r)) ->
+ (forall r', r' <> preg_of r -> rs'#r' = rs#r') ->
+ agree (Regmap.set r v ms) sp rs'.
+Proof.
+ intros H H0 H1. destruct H. split; auto.
+ - rewrite H1; auto. destruct r; simpl; congruence.
+ - intros. unfold Regmap.set. destruct (RegEq.eq r0 r). congruence.
+ rewrite H1; auto. destruct r; simpl; congruence.
+Qed.
+
+Corollary agree_set_mreg_parallel:
+ forall ms sp rs r v,
+ agree ms sp rs ->
+ agree (Regmap.set r v ms) sp (Pregmap.set (preg_of r) v rs).
+Proof.
+ intros. eapply agree_set_mreg; eauto.
+ - rewrite Pregmap.gss; auto.
+ - intros; apply Pregmap.gso; auto.
+Qed.
+
+Corollary agree_set_mreg_parallel2:
+ forall ms sp rs r v ms',
+ agree ms sp (set_from_Machrs ms' rs)->
+ agree (Regmap.set r v ms) sp (set_from_Machrs (Regmap.set r v ms') rs).
+Proof.
+ intros. unfold set_from_Machrs in *. eapply agree_set_mreg; eauto.
+ - rewrite Regmap.gss; auto.
+ - intros r' X. destruct r'; try congruence. rewrite Regmap.gso; try congruence.
+Qed.
+
+Definition data_preg (r: preg) : bool :=
+ match r with
+ | preg_of _ | SP => true
+ | _ => false
+ end.
+
+Lemma agree_exten ms sp rs rs':
+ agree ms sp rs ->
+ (forall r, data_preg r = true -> rs'#r = rs#r) ->
+ agree ms sp rs'.
+Proof.
+ intros H H0. destruct H. split; intros; try rewrite H0; auto.
+Qed.
+
+Lemma agree_set_from_Machrs ms sp ms' rs:
+ agree ms sp rs ->
+ (forall (r:mreg), (ms' r) = rs#r) ->
+ agree ms sp (set_from_Machrs ms' rs).
+Proof.
+ unfold set_from_Machrs; intros.
+ eapply agree_exten; eauto.
+ intros r; destruct r; simpl; try congruence.
+Qed.
+
+Lemma agree_set_other ms sp rs r v:
+ agree ms sp rs ->
+ data_preg r = false ->
+ agree ms sp (rs#r <- v).
+Proof.
+ intros; apply agree_exten with rs; auto.
+ intros. apply Pregmap.gso. congruence.
+Qed.
+
+
+Lemma agree_next_addr f ms sp b pos rs:
+ agree ms sp rs ->
+ agree ms sp (rs#PC <- (Vptr b (Ptrofs.repr (next f pos)))).
+Proof.
+ intros. apply agree_set_other; auto.
+Qed.
+
+Local Hint Resolve agree_set_mreg_parallel2: core.
+
+Lemma agree_set_pair sp p v ms ms' rs:
+ agree ms sp (set_from_Machrs ms' rs) ->
+ agree (set_pair p v ms) sp (set_from_Machrs (set_pair p v ms') rs).
+Proof.
+ intros H; destruct p; simpl; auto.
+Qed.
+
+Lemma agree_undef_caller_save_regs:
+ forall ms sp ms' rs,
+ agree ms sp (set_from_Machrs ms' rs) ->
+ agree (undef_caller_save_regs ms) sp (set_from_Machrs (undef_caller_save_regs ms') rs).
+Proof.
+ intros. destruct H as [H0 H1 H2]. unfold undef_caller_save_regs. split; auto.
+ intros.
+ unfold set_from_Machrs in * |- *.
+ rewrite H2. auto.
+Qed.
+
+Lemma agree_change_sp ms sp rs sp':
+ agree ms sp rs -> sp' <> Vundef ->
+ agree ms sp' (rs#SP <- sp').
+Proof.
+ intros H H0. inv H. split; auto.
+Qed.
+
+Lemma agree_undef_regs ms sp rl ms' rs:
+ agree ms sp (set_from_Machrs ms' rs) ->
+ agree (Mach.undef_regs rl ms) sp (set_from_Machrs (Mach.undef_regs rl ms') rs).
+Proof.
+ unfold set_from_Machrs; intros H. destruct H; subst. split; auto.
+ intros. destruct (In_dec mreg_eq r rl).
+ + rewrite! undef_regs_same; auto.
+ + rewrite! undef_regs_other; auto.
+Qed.
+
+(** Translation of arguments and results to builtins. *)
+
+Remark builtin_arg_match:
+ forall ms rs sp m a v,
+ agree ms sp rs ->
+ eval_builtin_arg ge ms sp m a v ->
+ eval_builtin_arg ge (to_Machrs rs) sp m a v.
+Proof.
+ induction 2; simpl; eauto with barg.
+ unfold to_Machrs; erewrite agree_mregs; eauto.
+ econstructor.
+Qed.
+
+Lemma builtin_args_match:
+ forall ms sp rs m,
+ agree ms sp rs ->
+ forall al vl, eval_builtin_args ge ms sp m al vl ->
+ eval_builtin_args ge (to_Machrs rs) sp m al vl.
+Proof.
+ induction 2; intros; simpl; try (constructor; auto).
+ eapply eval_builtin_arg_preserved; eauto.
+ eapply builtin_arg_match; eauto.
+Qed.
+
+Lemma agree_set_res res: forall ms sp rs v ms',
+ agree ms sp (set_from_Machrs ms' rs) ->
+ agree (set_res res v ms) sp (set_from_Machrs (set_res res v ms') rs).
+Proof.
+ induction res; simpl; auto.
+Qed.
+
+Lemma find_function_ptr_agree ros ms rs sp b:
+ agree ms sp rs ->
+ Machblock.find_function_ptr ge ros ms = Some b ->
+ find_function_ptr tge ros rs = Some (Vptr b Ptrofs.zero).
+Proof.
+ intros AG; unfold Mach.find_function_ptr; destruct ros as [r|s]; simpl; auto.
+ - generalize (agree_mregs _ _ _ AG r). destruct (ms r); simpl; try congruence.
+ intros H; inv H; try congruence.
+ inversion_ASSERT. intros H; rewrite (Ptrofs.same_if_eq _ _ H); eauto.
+ try_simplify_someHyps.
+ - intros H; rewrite symbols_preserved, H. auto.
+Qed.
+
+Lemma parent_sp_def: forall s, match_stack s -> parent_sp s <> Vundef.
+Proof.
+ induction 1; simpl.
+ unfold Vnullptr; destruct Archi.ptr64; congruence.
+ auto.
+Qed.
+
+Lemma extcall_arg_match ms sp rs m l v:
+ agree ms sp rs ->
+ extcall_arg ms m sp l v ->
+ extcall_arg rs m (rs#SP) l v.
+Proof.
+ destruct 2.
+ - erewrite agree_mregs; eauto. constructor.
+ - unfold load_stack in *. econstructor; eauto.
+ erewrite agree_sp; eauto.
+Qed.
+
+Local Hint Resolve extcall_arg_match: core.
+
+Lemma extcall_arg_pair_match:
+ forall ms sp rs m p v,
+ agree ms sp rs ->
+ extcall_arg_pair ms m sp p v ->
+ extcall_arg_pair rs m (rs#SP) p v.
+Proof.
+ destruct 2; constructor; eauto.
+Qed.
+
+Local Hint Resolve extcall_arg_pair_match: core.
+
+Lemma extcall_args_match:
+ forall ms sp rs m, agree ms sp rs ->
+ forall ll vl,
+ list_forall2 (extcall_arg_pair ms m sp) ll vl ->
+ list_forall2 (extcall_arg_pair rs m rs#SP) ll vl.
+Proof.
+ induction 2; constructor; eauto.
+Qed.
+
+Lemma extcall_arguments_match:
+ forall ms m sp rs sg args,
+ agree ms sp rs ->
+ extcall_arguments ms m sp sg args ->
+ extcall_arguments rs m (rs#SP) sg args.
+Proof.
+ unfold extcall_arguments, extcall_arguments; intros.
+ eapply extcall_args_match; eauto.
+Qed.
+
+(** A few tactics *)
+
+Local Hint Resolve functions_transl symbols_preserved
+ agree_next_addr agree_mregs agree_set_mreg_parallel agree_undef_regs agree_set_other agree_change_sp
+ agree_sp_def agree_set_from_Machrs agree_set_pair agree_undef_caller_save_regs agree_set_res f_equal Ptrofs.repr_unsigned parent_sp_def
+ builtin_args_match external_call_symbols_preserved: core.
+
+Ltac simplify_regmap :=
+ repeat (rewrite ?Pregmap.gss; try (rewrite Pregmap.gso; try congruence)).
+
+Ltac simplify_next_addr :=
+ match goal with
+ | [ HPC: Vptr _ _ = _ PC |- _ ] => try (unfold next_addr, Val.offset_ptr); simplify_regmap; try (rewrite <- HPC)
+ end.
+
+Ltac discharge_match_states :=
+ econstructor; eauto; try ( simplify_next_addr; econstructor; eauto ).
+
+
+(** Instruction step simulation lemma: the simulation lemma for stepping one instruction
+
+<<
+ s1 ---------------- s2
+ | |
+ t| +|t
+ | |
+ v v
+ s1'---------------- s2'
+>>
+
+*)
+
+Lemma trivial_exec_prologue:
+ forall tf ofs rs m,
+ 0 < Ptrofs.unsigned ofs ->
+ exec_prologue tf (Ptrofs.unsigned ofs) rs m = Next rs m.
+Proof.
+ unfold exec_prologue. intros.
+ destruct (Z.eq_dec); eauto. omega.
+Qed.
+
+Lemma basic_step_simulation: forall ms sp rs s f tf fb ms' bi m m',
+ agree ms sp rs ->
+ match_stack s ->
+ transf_function f = OK tf ->
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ Machblock.basic_step ge s fb sp ms m bi ms' m' ->
+ exists rs', basic_step tge tf rs m bi rs' m' /\ agree ms' sp rs' /\ rs' PC = rs PC.
+Proof.
+ destruct 5.
+ (* MBgetstack *)
+ - eexists; split.
+ + econstructor; eauto. erewrite agree_sp; eauto.
+ + eauto.
+ (* MBsetstack *)
+ - eexists; split.
+ + econstructor; eauto.
+ erewrite agree_sp; eauto.
+ erewrite <- agree_mregs; eauto.
+ + rewrite H4; split; eauto.
+ (* MBgetparam *)
+ - eexists; split.
+ + econstructor; eauto.
+ erewrite agree_sp; eauto.
+ assert (f = f0). { rewrite H2 in H3. inversion H3. reflexivity. }
+ assert (fn_link_ofs tf = fn_link_ofs f0). {
+ rewrite <- H7.
+ eapply stackinfo_preserved. eauto.
+ }
+ rewrite H8. eauto.
+ + rewrite H6; split; eauto.
+ (* MBop *)
+ - eexists; split.
+ + econstructor; eauto.
+ erewrite agree_sp; eauto.
+ erewrite agree_mregs_list in H3.
+ * erewrite <- H3.
+ eapply eval_operation_preserved; trivial.
+ * eauto.
+ + rewrite H4; split; eauto.
+ (* MBload *)
+ - eexists; split.
+ + econstructor; eauto.
+ erewrite agree_sp; eauto. erewrite <- agree_mregs_list; eauto.
+ erewrite <- H3.
+ eapply eval_addressing_preserved; trivial.
+ + rewrite H5; split; eauto.
+ (* MBload notrap1 *)
+ - eexists; split.
+ + eapply exec_MBload_notrap1; eauto.
+ erewrite agree_sp; eauto.
+ erewrite agree_mregs_list in H3; eauto.
+ * erewrite eval_addressing_preserved; eauto.
+ + rewrite H4; eauto.
+ (* MBload notrap2 *)
+ - eexists; split.
+ + eapply exec_MBload_notrap2; eauto.
+ erewrite agree_sp; eauto.
+ erewrite agree_mregs_list in H3; eauto.
+ * erewrite eval_addressing_preserved; eauto.
+ + rewrite H5; eauto.
+ (* MBstore *)
+ - eexists; split.
+ + econstructor; eauto.
+ * erewrite agree_sp; eauto.
+ erewrite agree_mregs_list in H3.
+ erewrite eval_addressing_preserved; eauto.
+ eauto.
+ * erewrite <- agree_mregs; eauto.
+ + rewrite H5; eauto.
+Qed.
+
+Lemma body_step_simulation: forall ms sp s f tf fb ms' bb m m',
+ match_stack s ->
+ transf_function f = OK tf ->
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ Machblock.body_step ge s fb sp bb ms m ms' m' ->
+ forall rs, agree ms sp rs ->
+ exists rs', body_step tge tf bb rs m rs' m' /\ agree ms' sp rs' /\ rs' PC = rs PC.
+Proof.
+ induction 4.
+ - repeat (econstructor; eauto).
+ - intros. exploit basic_step_simulation; eauto.
+ intros (rs'0 & BASIC & AG1' & AGPC1).
+ exploit IHbody_step; eauto.
+ intros (rs'1 & BODY & AG2' & AGPC2).
+ repeat (econstructor; eauto).
+ congruence.
+Qed.
+
+Local Hint Resolve trivial_exec_prologue: core.
+
+Lemma exit_step_simulation s fb f sp c t bb ms m s1' tf rs pc
+ (STEP: Machblock.exit_step rao ge (exit bb) (Machblock.State s fb sp (bb :: c) ms m) t s1')
+ (STACKS: match_stack s)
+ (AG: agree ms sp rs)
+ (NXT: next_addr next tf rs = Some pc)
+ (AT: transl_code_at_pc fb f tf c pc):
+ exists rs' m', exit_step next tge tf (exit bb) (rs#PC <- pc) m t rs' m' /\ match_states s1' (State rs' m').
+Proof.
+ inv AT.
+ inv STEP.
+ (* cfi_step currently only defined for exec_MBcall, exec_MBreturn, and exec_MBgoto *)
+ - inversion H4; subst. clear H4. (* inversion_clear H4 does not work so well: it clears an important hypothesis about "sp" in the Mreturn case *)
+ (* MBcall *)
+ + eexists. eexists. split.
+ * apply exec_Some_exit.
+ apply exec_MBcall.
+ eapply find_function_ptr_agree; eauto.
+ * assert (f0 = f). { congruence. } subst.
+ assert (Ptrofs.unsigned ra = Ptrofs.unsigned ofs). {
+ eapply is_pos_inject1; eauto.
+ }
+ assert (ofs = ra). {
+ apply Ptrofs.same_if_eq. unfold Ptrofs.eq.
+ rewrite H4. rewrite zeq_true. reflexivity.
+ }
+ repeat econstructor; eauto.
+ -- unfold rao in *. congruence.
+ -- simpl. simplify_regmap.
+ erewrite agree_sp; eauto.
+ -- simpl. simplify_regmap. auto.
+ (* MBtailcall *)
+ + assert (f0 = f). { congruence. } subst.
+ eexists. eexists. split.
+ * repeat econstructor.
+ -- eapply find_function_ptr_agree; eauto.
+ -- unfold exec_epilogue. erewrite agree_sp; eauto.
+ apply stackinfo_preserved in H0 as ( SS & RA & LO ).
+ rewrite SS, RA, LO.
+ try_simplify_someHyps.
+ * repeat econstructor; eauto.
+ intros r.
+ eapply agree_mregs; eapply agree_set_other; eauto.
+ (* MBbuiltin *)
+ +eexists. eexists. split.
+ * repeat econstructor.
+ -- simplify_regmap. erewrite agree_sp; eauto.
+ eapply eval_builtin_args_preserved; eauto.
+ -- eapply external_call_symbols_preserved; eauto.
+ exact senv_preserved.
+ * repeat econstructor; eauto.
+ -- assert (transl_function f = tf). {
+ unfold transf_function in *. congruence.
+ }
+ erewrite H5. assumption.
+ -- eapply agree_sp. eapply agree_set_res. eapply agree_undef_regs.
+ eapply agree_set_from_Machrs; eauto.
+ -- intros; simpl.
+ eapply agree_set_res. eapply agree_undef_regs.
+ eapply agree_set_from_Machrs; eauto.
+ (* MBgoto *)
+ + simplify_someHyps. intros.
+ assert ((rs # PC <- (Vptr fb ofs)) PC = Vptr fb ofs). {
+ rewrite Pregmap.gss. reflexivity.
+ }
+ eassert (exists pc, goto_label next tf lbl rs # PC <- (Vptr fb ofs) = Some pc
+ /\ transl_code_at_pc fb f tf c' pc) as (pc & GOTO_LABEL & _). {
+ eapply find_label_goto_label; eauto.
+ }
+ eexists. eexists. split.
+ * apply exec_Some_exit.
+ apply exec_MBgoto.
+ rewrite GOTO_LABEL. trivial.
+ * repeat econstructor; eauto.
+ -- simplify_regmap.
+ exploit find_label_goto_label; eauto. intros (pc' & GOTO_LABEL' & TRANSL).
+ assert(pc' = pc). { congruence. } subst. eauto.
+ -- simplify_regmap. erewrite agree_sp; eauto.
+ (* MBcond true *)
+ (* Mostly copy and paste from MBgoto *)
+ + simplify_someHyps. intros.
+ assert ((rs # PC <- (Vptr fb ofs)) PC = Vptr fb ofs). {
+ rewrite Pregmap.gss. reflexivity.
+ }
+ eassert (exists pc, goto_label next tf lbl rs # PC <- (Vptr fb ofs) = Some pc
+ /\ transl_code_at_pc fb f tf c' pc) as (pc & GOTO_LABEL & _). {
+ eapply find_label_goto_label; eauto.
+ }
+ eexists. eexists. split.
+ * apply exec_Some_exit. eapply exec_MBcond_true; eauto.
+ erewrite agree_mregs_list in H14; eauto.
+ * repeat econstructor; eauto.
+ -- simplify_regmap.
+ exploit find_label_goto_label; eauto. intros (pc' & GOTO_LABEL' & TRANSL).
+ assert(pc' = pc). { congruence. } subst. eauto.
+ -- simplify_regmap. erewrite agree_sp; eauto.
+ (* MBcond false *)
+ + inv H0. eexists. eexists. split.
+ * apply exec_Some_exit. apply exec_MBcond_false.
+ -- erewrite agree_mregs_list in H15; eauto.
+ -- trivial.
+ * repeat econstructor; eauto. erewrite agree_sp; eauto.
+ (* MBjumptable *)
+ + simplify_someHyps. intros.
+ assert ((rs # PC <- (Vptr fb ofs)) PC = Vptr fb ofs). {
+ rewrite Pregmap.gss. reflexivity.
+ }
+ eassert (exists pc, goto_label next tf lbl rs # PC <- (Vptr fb ofs) = Some pc
+ /\ transl_code_at_pc fb f tf c' pc) as (pc & GOTO_LABEL & _). {
+ eapply find_label_goto_label; eauto.
+ }
+ eexists. eexists. split.
+ * repeat econstructor; eauto.
+ rewrite <- H14.
+ symmetry. eapply agree_mregs. eapply agree_set_other; eauto.
+ * repeat econstructor; eauto.
+ (* copy paste from MBgoto *)
+ -- simplify_regmap.
+ exploit find_label_goto_label; eauto. intros (pc' & GOTO_LABEL' & TRANSL).
+ assert(pc' = pc). { congruence. } subst. eauto.
+ -- simplify_regmap. erewrite agree_sp; eauto.
+ -- intros; simplify_regmap. eauto.
+ + (* MBreturn *)
+ try_simplify_someHyps.
+ eexists. eexists. split.
+ * apply exec_Some_exit.
+ apply exec_MBreturn.
+ unfold exec_epilogue.
+ erewrite agree_sp; eauto.
+ apply stackinfo_preserved in H0 as ( SS & RA & LO ).
+ rewrite SS, RA, LO.
+ try_simplify_someHyps.
+ * repeat econstructor; eauto. intros r.
+ simplify_regmap. eapply agree_mregs; eauto.
+ - inv H0; repeat econstructor; eauto.
+ erewrite agree_sp; eauto.
+Qed.
+
+Lemma inst_step_simulation s fb f sp c t ms m s1' tf rs
+ (STEP: Machblock.step rao ge (Machblock.State s fb sp c ms m) t s1')
+ (STACKS: match_stack s)
+ (AT: transl_code_at_pc fb f tf c (rs PC))
+ (AG: agree ms sp rs):
+ exists s2' : state, plus (step next) tge (State rs m) t s2' /\ match_states s1' s2'.
+Proof.
+ inv STEP.
+ inv AT.
+ exploit body_step_simulation; eauto. intros (rs0' & BODY & AG0 & AGPC).
+ assert (NXT: next_addr next tf rs0' = Some (Vptr fb (Ptrofs.repr (next tf (Ptrofs.unsigned ofs))))).
+ { unfold next_addr; rewrite AGPC, <- H; simpl; eauto. }
+ exploit exit_step_simulation; eauto.
+ { econstructor; eauto.
+ erewrite is_pos_unsigned_repr; eauto.
+ generalize (next_progress tf (Ptrofs.unsigned ofs)); omega. }
+ intros (rs2 & m2 & STEP & MS).
+ eexists.
+ split; eauto.
+ eapply plus_one.
+ eapply exec_step_internal; eauto.
+ econstructor; eauto.
+Qed.
+
+Lemma prologue_preserves_pc: forall f rs0 m0 rs1 m1,
+ exec_prologue f 0 rs0 m0 = Next rs1 m1 ->
+ rs1 PC = rs0 PC.
+Proof.
+ unfold exec_prologue; simpl;
+ intros f rs0 m0 rs1 m1 H.
+ destruct (Mem.alloc m0 0 (fn_stacksize f)) in H; unfold Next in H.
+ simplify_someHyps; inversion_SOME ignored; inversion_SOME ignored';
+ intros ? ? H'; inversion H'; trivial.
+Qed.
+
+Lemma is_pos_next_zero bb c fb f
+ (FIND : Genv.find_funct_ptr ge fb = Some (Internal f))
+ (FN_HEAD : bb :: c = fn_code f):
+ is_pos next (transl_function f) (next (transl_function f) (Ptrofs.unsigned Ptrofs.zero)) (if has_header (fn_code f) then bb::c else c).
+Proof.
+ exploit (transf_function_def f). unfold transf_function; auto.
+ unfold insert_implicit_prologue.
+ intros fn_code_tf; destruct (has_header (fn_code f));
+ eapply Next_pos; rewrite Ptrofs.unsigned_zero;
+ rewrite FN_HEAD; rewrite <- fn_code_tf; apply First_pos.
+Qed.
+
+Lemma prologue_simulation_no_header_step t s1' s sp fb ms f m1 rs0 m0 rs1
+ (STACKS : match_stack s)
+ (AG : agree ms sp rs1)
+ (ATPC : rs0 PC = Vptr fb Ptrofs.zero)
+ (ATLR : rs0 RA = parent_ra s)
+ (FIND : Genv.find_funct_ptr ge fb = Some (Internal f))
+ (PROL : exec_prologue f 0 rs0 m0 = Next rs1 m1)
+ (STEP : Machblock.step rao ge (Machblock.State s fb sp (fn_code f) ms m1) t s1')
+ (NO_HEADER: has_header (fn_code f) = false):
+ exists s2' : state, step next tge {| _rs := rs0; _m := m0 |} t s2' /\ match_states s1' s2'.
+Proof.
+ inv STEP.
+
+ exploit functions_translated; eauto;
+ intros (tf & FINDtf & TRANSf); monadInv TRANSf.
+ assert (fn_code f = fn_code (transl_function f)) as TF_CODE. {
+ unfold transl_function; simpl; unfold insert_implicit_prologue;
+ rewrite NO_HEADER; trivial.
+ }
+
+ exploit body_step_simulation; eauto; unfold transf_function; auto.
+ intros (rs0' & BODY & AG0 & AGPC).
+
+ exploit prologue_preserves_pc; eauto.
+ intros AGPC'.
+
+ exploit is_pos_next_zero; eauto; rewrite NO_HEADER.
+ intros IS_POS.
+
+ exploit transl_code_at_pc_intro; eauto; unfold transf_function; auto. {
+ rewrite Ptrofs.unsigned_zero; erewrite is_pos_unsigned_repr; eauto.
+ } intros TRANSL_CODE.
+
+ assert (next_addr next (transl_function f) rs0' =
+ Some (Vptr fb (Ptrofs.repr (next (transl_function f)
+ (Ptrofs.unsigned Ptrofs.zero))))) as NEXT_ADDR. {
+ unfold next_addr; rewrite AGPC; rewrite AGPC'; rewrite ATPC; reflexivity.
+ }
+
+ exploit exit_step_simulation; eauto.
+ intros (? & ? & EXIT_STEP & MATCH_EXIT).
+
+ exploit exec_bblock_all; eauto.
+ intros EXEC_BBLOCK.
+
+ exploit exec_step_internal; eauto.
+ apply is_pos_simplify; eauto. rewrite H3; rewrite TF_CODE; apply First_pos.
+Qed.
+
+Lemma prologue_simulation_header_step t s1' s sp fb ms f m1 rs0 m0 rs1
+ (STACKS : match_stack s)
+ (AG : agree ms sp rs1)
+ (ATPC : rs0 PC = Vptr fb Ptrofs.zero)
+ (ATLR : rs0 RA = parent_ra s)
+ (FIND : Genv.find_funct_ptr ge fb = Some (Internal f))
+ (PROL : exec_prologue f 0 rs0 m0 = Next rs1 m1)
+ (STEP : Machblock.step rao ge (Machblock.State s fb sp (fn_code f) ms m1) t s1')
+ (HEADER: has_header (fn_code f) = true):
+ exists s2' : state, plus (step next) tge {| _rs := rs0; _m := m0 |} t s2' /\ match_states s1' s2'.
+Proof.
+ inv STEP.
+
+ (* FIRST STEP *)
+ exploit functions_translated; eauto;
+ intros (tf & FINDtf & TRANSf); monadInv TRANSf.
+ exploit transf_function_def; eauto; unfold transf_function; auto;
+ unfold insert_implicit_prologue; rewrite HEADER; intros TF_CODE.
+
+ exploit is_pos_next_zero; eauto; rewrite HEADER; rewrite H3;
+ intros IS_POS.
+
+ exploit prologue_preserves_pc; eauto.
+ intros AGPC'.
+
+ assert ( next_addr next (transl_function f) rs1
+ = Some (Vptr fb (Ptrofs.repr (next (transl_function f) (Ptrofs.unsigned Ptrofs.zero))))
+ ) as NEXT_ADDR0. { unfold next_addr; rewrite AGPC'; rewrite ATPC; trivial. }
+
+ exploit exec_nil_body; intros BODY0.
+ assert ((body {| header := nil; body := nil; exit := None |}) = nil) as NIL; auto.
+ rewrite <- NIL in BODY0.
+
+ exploit exec_None_exit; intros EXIT0.
+ assert ((exit {| header := nil; body := nil; exit := None |}) = None) as NONE; auto.
+ rewrite <- NONE in EXIT0.
+
+ exploit exec_bblock_all; eauto;
+ intros BBLOCK0.
+
+ exploit exec_step_internal; eauto. rewrite <- TF_CODE; apply First_pos.
+ intros STEP0.
+
+ clear BODY0 BBLOCK0 EXIT0.
+
+ (* SECOND step *)
+
+ exploit (mkagree ms sp
+ (rs1 # PC <- (Vptr fb (Ptrofs.repr (next (transl_function f)
+ (Ptrofs.unsigned Ptrofs.zero)))))); eauto.
+ apply (agree_sp ms); apply agree_set_other; eauto.
+ intros AG'.
+
+ exploit body_step_simulation; eauto; unfold transf_function; auto.
+ intros (rs1' & BODY1 & AGRS1' & AGPC).
+
+ assert ( next_addr next (transl_function f) rs1'
+ = Some (Vptr fb (Ptrofs.repr (next (transl_function f) (Ptrofs.unsigned
+ (Ptrofs.repr (next (transl_function f) (Ptrofs.unsigned Ptrofs.zero)))))))
+ ) as NEXT_ADDR1. { unfold next_addr; rewrite AGPC; reflexivity. }
+
+ assert (IS_POS' := IS_POS).
+ rewrite <- H3 in IS_POS'; apply Next_pos in IS_POS'.
+ exploit transl_code_at_pc_intro; eauto; unfold transf_function; auto. {
+ rewrite Ptrofs.unsigned_zero; erewrite is_pos_unsigned_repr; eauto.
+ assert (0 < next (transl_function f) 0) as Z0. { apply next_progress. }
+ assert ( next (transl_function f) 0
+ < next (transl_function f) (next (transl_function f) 0)
+ ) as Z1. { apply next_progress. }
+ rewrite <- Z1. exact Z0.
+ } intros TRANSL_CODE1.
+
+ exploit exit_step_simulation; eauto.
+ rewrite Ptrofs.unsigned_repr.
+ 2: {
+ assert(max_pos (transl_function f) <= Ptrofs.max_unsigned) as MAX_POS. {
+ eapply functions_bound_max_pos; eauto.
+ }
+ rewrite <- MAX_POS.
+ eapply is_pos_bound_pos; eauto.
+ }
+ exact TRANSL_CODE1.
+ intros (? & ? & EXIT_STEP & MATCH_EXIT).
+
+ exploit (trivial_exec_prologue (transl_function f) (Ptrofs.repr (next (transl_function f) (Ptrofs.unsigned Ptrofs.zero)))). {
+ rewrite Ptrofs.unsigned_zero; erewrite is_pos_unsigned_repr; eauto.
+ } intros NO_PROL.
+
+ exploit exec_bblock_all; eauto; intros BBLOCK1.
+
+ clear AGPC.
+ rewrite <- H3 in IS_POS.
+ exploit (exec_step_internal next tge fb
+ (Ptrofs.repr (next (transl_function f) (Ptrofs.unsigned Ptrofs.zero)))
+ (transl_function f)
+ bb c); simplify_regmap; eauto.
+
+ intros STEP1.
+
+ eexists; split.
+ + eapply plus_two.
+ * exact STEP0.
+ * exact STEP1.
+ * trivial.
+ + assumption.
+Qed.
+
+Lemma step_simulation s1 t s1':
+ Machblock.step rao ge s1 t s1' ->
+ forall s2, match_states s1 s2 ->
+ (exists s2', plus (step next) tge s2 t s2' /\ match_states s1' s2') \/ ((measure s1' < measure s1)%nat /\ t = E0 /\ match_states s1' s2).
+Proof.
+ intros STEP s2 MATCH; destruct MATCH.
+ - exploit inst_step_simulation; eauto.
+ - left.
+ destruct (has_header (fn_code f)) eqn:NO_HEADER.
+ + eapply prologue_simulation_header_step; eauto.
+ + exploit prologue_simulation_no_header_step; eauto;
+ intros (s2' & NO_HEADER_STEP & MATCH_STATES).
+ eexists; split; eauto.
+ apply plus_one; eauto.
+ - inv STEP; simpl; exploit functions_translated; eauto;
+ intros (tf0 & FINDtf & TRANSf);
+ monadInv TRANSf.
+ * (* internal calls *)
+ right.
+ intuition.
+ econstructor; eauto.
+ -- exploit
+ (mkagree (undef_regs destroyed_at_function_entry ms)
+ sp
+ (set_from_Machrs (undef_regs destroyed_at_function_entry rs) rs) # SP <- sp
+ ); eauto.
+ ++ unfold sp; discriminate.
+ ++ intros mr; unfold undef_regs;
+ induction destroyed_at_function_entry as [ | ? ? IH].
+ ** inversion AG as [_ _ AG_MREGS]; apply AG_MREGS.
+ ** fold undef_regs in *;
+ unfold Regmap.set; simpl; rewrite IH; reflexivity.
+ -- unfold exec_prologue;
+ inversion AG as (RS_SP & ?); rewrite RS_SP; fold sp;
+ rewrite H4; fold sp; rewrite H7; rewrite ATLR; rewrite H8; eauto.
+ * (* external calls *)
+ left.
+ exploit extcall_arguments_match; eauto.
+ eexists; split.
+ + eapply plus_one.
+ eapply exec_step_external; eauto.
+ eapply external_call_symbols_preserved; eauto. apply senv_preserved.
+ + econstructor; eauto.
+ - (* Returnstate *)
+ inv STEP; simpl; right.
+ inv STACKS; simpl in *; subst.
+ repeat (econstructor; eauto).
+Qed.
+
+(** * The main simulation theorem *)
+
+Theorem transf_program_correct:
+ forward_simulation (Machblock.semantics rao prog) (semantics next tprog).
+Proof.
+ eapply forward_simulation_star with (measure := measure).
+ apply senv_preserved.
+ - eexact transf_initial_states.
+ - eexact transf_final_states.
+ - eexact step_simulation.
+Qed.
+
+End PRESERVATION.