aboutsummaryrefslogtreecommitdiffstats
path: root/ia32
diff options
context:
space:
mode:
Diffstat (limited to 'ia32')
-rw-r--r--ia32/Archi.v54
-rw-r--r--ia32/Asm.v1204
-rw-r--r--ia32/AsmToJSON.ml17
-rw-r--r--ia32/AsmToJSON.mli13
-rw-r--r--ia32/Asmexpand.ml638
-rw-r--r--ia32/Asmgen.v754
-rw-r--r--ia32/Asmgenproof.v916
-rw-r--r--ia32/Asmgenproof1.v1481
-rw-r--r--ia32/CBuiltins.ml94
-rw-r--r--ia32/CombineOp.v150
-rw-r--r--ia32/CombineOpproof.v179
-rw-r--r--ia32/ConstpropOp.vp404
-rw-r--r--ia32/ConstpropOpproof.v881
-rw-r--r--ia32/Conventions1.v473
-rw-r--r--ia32/Machregs.v367
-rw-r--r--ia32/Machregsaux.ml33
-rw-r--r--ia32/Machregsaux.mli18
-rw-r--r--ia32/NeedOp.v254
-rw-r--r--ia32/Op.v1457
-rw-r--r--ia32/PrintOp.ml169
-rw-r--r--ia32/SelectLong.vp347
-rw-r--r--ia32/SelectLongproof.v557
-rw-r--r--ia32/SelectOp.vp530
-rw-r--r--ia32/SelectOpproof.v958
-rw-r--r--ia32/Stacklayout.v148
-rw-r--r--ia32/TargetPrinter.ml881
-rw-r--r--ia32/ValueAOp.v266
-rw-r--r--ia32/extractionMachdep.v31
28 files changed, 0 insertions, 13274 deletions
diff --git a/ia32/Archi.v b/ia32/Archi.v
deleted file mode 100644
index 936bacb3..00000000
--- a/ia32/Archi.v
+++ /dev/null
@@ -1,54 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris *)
-(* Jacques-Henri Jourdan, INRIA Paris *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-(** Architecture-dependent parameters for IA32 *)
-
-Require Import ZArith.
-Require Import Fappli_IEEE.
-Require Import Fappli_IEEE_bits.
-
-Parameter ptr64: bool.
-
-Definition big_endian := false.
-
-Definition align_int64 := if ptr64 then 8%Z else 4%Z.
-Definition align_float64 := if ptr64 then 8%Z else 4%Z.
-
-Definition splitlong := negb ptr64.
-
-Lemma splitlong_ptr32: splitlong = true -> ptr64 = false.
-Proof.
- unfold splitlong. destruct ptr64; simpl; congruence.
-Qed.
-
-Program Definition default_pl_64 : bool * nan_pl 53 :=
- (true, iter_nat 51 _ xO xH).
-
-Definition choose_binop_pl_64 (s1: bool) (pl1: nan_pl 53) (s2: bool) (pl2: nan_pl 53) :=
- false. (**r always choose first NaN *)
-
-Program Definition default_pl_32 : bool * nan_pl 24 :=
- (true, iter_nat 22 _ xO xH).
-
-Definition choose_binop_pl_32 (s1: bool) (pl1: nan_pl 24) (s2: bool) (pl2: nan_pl 24) :=
- false. (**r always choose first NaN *)
-
-Definition float_of_single_preserves_sNaN := false.
-
-Global Opaque ptr64 big_endian splitlong
- default_pl_64 choose_binop_pl_64
- default_pl_32 choose_binop_pl_32
- float_of_single_preserves_sNaN.
diff --git a/ia32/Asm.v b/ia32/Asm.v
deleted file mode 100644
index 304cb8e4..00000000
--- a/ia32/Asm.v
+++ /dev/null
@@ -1,1204 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Abstract syntax and semantics for IA32 assembly language *)
-
-Require Import Coqlib Maps.
-Require Import AST Integers Floats Values Memory Events Globalenvs Smallstep.
-Require Import Locations Stacklayout Conventions.
-
-(** * Abstract syntax *)
-
-(** ** Registers. *)
-
-(** Integer registers. *)
-
-Inductive ireg: Type :=
- | RAX | RBX | RCX | RDX | RSI | RDI | RBP | RSP
- | R8 | R9 | R10 | R11 | R12 | R13 | R14 | R15.
-
-(** Floating-point registers, i.e. SSE2 registers *)
-
-Inductive freg: Type :=
- | XMM0 | XMM1 | XMM2 | XMM3 | XMM4 | XMM5 | XMM6 | XMM7
- | XMM8 | XMM9 | XMM10 | XMM11 | XMM12 | XMM13 | XMM14 | XMM15.
-
-Lemma ireg_eq: forall (x y: ireg), {x=y} + {x<>y}.
-Proof. decide equality. Defined.
-
-Lemma freg_eq: forall (x y: freg), {x=y} + {x<>y}.
-Proof. decide equality. Defined.
-
-(** Bits of the flags register. *)
-
-Inductive crbit: Type :=
- | ZF | CF | PF | SF | OF.
-
-(** All registers modeled here. *)
-
-Inductive preg: Type :=
- | PC: preg (**r program counter *)
- | IR: ireg -> preg (**r integer register *)
- | FR: freg -> preg (**r XMM register *)
- | ST0: preg (**r top of FP stack *)
- | CR: crbit -> preg (**r bit of the flags register *)
- | RA: preg. (**r pseudo-reg representing return address *)
-
-Coercion IR: ireg >-> preg.
-Coercion FR: freg >-> preg.
-Coercion CR: crbit >-> preg.
-
-(** Conventional names for stack pointer ([SP]) and return address ([RA]) *)
-
-Notation SP := RSP (only parsing).
-
-(** ** Instruction set. *)
-
-Definition label := positive.
-
-(** General form of an addressing mode. *)
-
-Inductive addrmode: Type :=
- | Addrmode (base: option ireg)
- (ofs: option (ireg * Z))
- (const: Z + ident * ptrofs).
-
-(** Testable conditions (for conditional jumps and more). *)
-
-Inductive testcond: Type :=
- | Cond_e | Cond_ne
- | Cond_b | Cond_be | Cond_ae | Cond_a
- | Cond_l | Cond_le | Cond_ge | Cond_g
- | Cond_p | Cond_np.
-
-(** Instructions. IA32 instructions accept many combinations of
- registers, memory references and immediate constants as arguments.
- Here, we list only the combinations that we actually use.
-
- Naming conventions for types:
-- [b]: 8 bits
-- [w]: 16 bits ("word")
-- [l]: 32 bits ("longword")
-- [q]: 64 bits ("quadword")
-- [d] or [sd]: FP double precision (64 bits)
-- [s] or [ss]: FP single precision (32 bits)
-
- Naming conventions for operands:
-- [r]: integer register operand
-- [f]: XMM register operand
-- [m]: memory operand
-- [i]: immediate integer operand
-- [s]: immediate symbol operand
-- [l]: immediate label operand
-- [cl]: the [CL] register
-
- For two-operand instructions, the first suffix describes the result
- (and first argument), the second suffix describes the second argument.
-*)
-
-Inductive instruction: Type :=
- (** Moves *)
- | Pmov_rr (rd: ireg) (r1: ireg) (**r [mov] (integer) *)
- | Pmovl_ri (rd: ireg) (n: int)
- | Pmovq_ri (rd: ireg) (n: int64)
- | Pmov_rs (rd: ireg) (id: ident)
- | Pmovl_rm (rd: ireg) (a: addrmode)
- | Pmovq_rm (rd: ireg) (a: addrmode)
- | Pmovl_mr (a: addrmode) (rs: ireg)
- | Pmovq_mr (a: addrmode) (rs: ireg)
- | Pmovsd_ff (rd: freg) (r1: freg) (**r [movsd] (single 64-bit float) *)
- | Pmovsd_fi (rd: freg) (n: float) (**r (pseudo-instruction) *)
- | Pmovsd_fm (rd: freg) (a: addrmode)
- | Pmovsd_mf (a: addrmode) (r1: freg)
- | Pmovss_fi (rd: freg) (n: float32) (**r [movss] (single 32-bit float) *)
- | Pmovss_fm (rd: freg) (a: addrmode)
- | Pmovss_mf (a: addrmode) (r1: freg)
- | Pfldl_m (a: addrmode) (**r [fld] double precision *)
- | Pfstpl_m (a: addrmode) (**r [fstp] double precision *)
- | Pflds_m (a: addrmode) (**r [fld] simple precision *)
- | Pfstps_m (a: addrmode) (**r [fstp] simple precision *)
- (** Moves with conversion *)
- | Pmovb_mr (a: addrmode) (rs: ireg) (**r [mov] (8-bit int) *)
- | Pmovw_mr (a: addrmode) (rs: ireg) (**r [mov] (16-bit int) *)
- | Pmovzb_rr (rd: ireg) (rs: ireg) (**r [movzb] (8-bit zero-extension) *)
- | Pmovzb_rm (rd: ireg) (a: addrmode)
- | Pmovsb_rr (rd: ireg) (rs: ireg) (**r [movsb] (8-bit sign-extension) *)
- | Pmovsb_rm (rd: ireg) (a: addrmode)
- | Pmovzw_rr (rd: ireg) (rs: ireg) (**r [movzw] (16-bit zero-extension) *)
- | Pmovzw_rm (rd: ireg) (a: addrmode)
- | Pmovsw_rr (rd: ireg) (rs: ireg) (**r [movsw] (16-bit sign-extension) *)
- | Pmovsw_rm (rd: ireg) (a: addrmode)
- | Pmovzl_rr (rd: ireg) (rs: ireg) (**r [movzl] (32-bit zero-extension) *)
- | Pmovsl_rr (rd: ireg) (rs: ireg) (**r [movsl] (32-bit sign-extension) *)
- | Pmovls_rr (rd: ireg) (** 64 to 32 bit conversion (pseudo) *)
- | Pcvtsd2ss_ff (rd: freg) (r1: freg) (**r conversion to single float *)
- | Pcvtss2sd_ff (rd: freg) (r1: freg) (**r conversion to double float *)
- | Pcvttsd2si_rf (rd: ireg) (r1: freg) (**r double to signed int *)
- | Pcvtsi2sd_fr (rd: freg) (r1: ireg) (**r signed int to double *)
- | Pcvttss2si_rf (rd: ireg) (r1: freg) (**r single to signed int *)
- | Pcvtsi2ss_fr (rd: freg) (r1: ireg) (**r signed int to single *)
- | Pcvttsd2sl_rf (rd: ireg) (r1: freg) (**r double to signed long *)
- | Pcvtsl2sd_fr (rd: freg) (r1: ireg) (**r signed long to double *)
- | Pcvttss2sl_rf (rd: ireg) (r1: freg) (**r single to signed long *)
- | Pcvtsl2ss_fr (rd: freg) (r1: ireg) (**r signed long to single *)
- (** Integer arithmetic *)
- | Pleal (rd: ireg) (a: addrmode)
- | Pleaq (rd: ireg) (a: addrmode)
- | Pnegl (rd: ireg)
- | Pnegq (rd: ireg)
- | Paddl_ri (rd: ireg) (n: int)
- | Paddq_ri (rd: ireg) (n: int64)
- | Psubl_rr (rd: ireg) (r1: ireg)
- | Psubq_rr (rd: ireg) (r1: ireg)
- | Pimull_rr (rd: ireg) (r1: ireg)
- | Pimulq_rr (rd: ireg) (r1: ireg)
- | Pimull_ri (rd: ireg) (n: int)
- | Pimulq_ri (rd: ireg) (n: int64)
- | Pimull_r (r1: ireg)
- | Pimulq_r (r1: ireg)
- | Pmull_r (r1: ireg)
- | Pmulq_r (r1: ireg)
- | Pcltd
- | Pcqto
- | Pdivl (r1: ireg)
- | Pdivq (r1: ireg)
- | Pidivl (r1: ireg)
- | Pidivq (r1: ireg)
- | Pandl_rr (rd: ireg) (r1: ireg)
- | Pandq_rr (rd: ireg) (r1: ireg)
- | Pandl_ri (rd: ireg) (n: int)
- | Pandq_ri (rd: ireg) (n: int64)
- | Porl_rr (rd: ireg) (r1: ireg)
- | Porq_rr (rd: ireg) (r1: ireg)
- | Porl_ri (rd: ireg) (n: int)
- | Porq_ri (rd: ireg) (n: int64)
- | Pxorl_r (rd: ireg) (**r [xor] with self = set to zero *)
- | Pxorq_r (rd: ireg)
- | Pxorl_rr (rd: ireg) (r1: ireg)
- | Pxorq_rr (rd: ireg) (r1: ireg)
- | Pxorl_ri (rd: ireg) (n: int)
- | Pxorq_ri (rd: ireg) (n: int64)
- | Pnotl (rd: ireg)
- | Pnotq (rd: ireg)
- | Psall_rcl (rd: ireg)
- | Psalq_rcl (rd: ireg)
- | Psall_ri (rd: ireg) (n: int)
- | Psalq_ri (rd: ireg) (n: int)
- | Pshrl_rcl (rd: ireg)
- | Pshrq_rcl (rd: ireg)
- | Pshrl_ri (rd: ireg) (n: int)
- | Pshrq_ri (rd: ireg) (n: int)
- | Psarl_rcl (rd: ireg)
- | Psarq_rcl (rd: ireg)
- | Psarl_ri (rd: ireg) (n: int)
- | Psarq_ri (rd: ireg) (n: int)
- | Pshld_ri (rd: ireg) (r1: ireg) (n: int)
- | Prorl_ri (rd: ireg) (n: int)
- | Prorq_ri (rd: ireg) (n: int)
- | Pcmpl_rr (r1 r2: ireg)
- | Pcmpq_rr (r1 r2: ireg)
- | Pcmpl_ri (r1: ireg) (n: int)
- | Pcmpq_ri (r1: ireg) (n: int64)
- | Ptestl_rr (r1 r2: ireg)
- | Ptestq_rr (r1 r2: ireg)
- | Ptestl_ri (r1: ireg) (n: int)
- | Ptestq_ri (r1: ireg) (n: int64)
- | Pcmov (c: testcond) (rd: ireg) (r1: ireg)
- | Psetcc (c: testcond) (rd: ireg)
- (** Floating-point arithmetic *)
- | Paddd_ff (rd: freg) (r1: freg)
- | Psubd_ff (rd: freg) (r1: freg)
- | Pmuld_ff (rd: freg) (r1: freg)
- | Pdivd_ff (rd: freg) (r1: freg)
- | Pnegd (rd: freg)
- | Pabsd (rd: freg)
- | Pcomisd_ff (r1 r2: freg)
- | Pxorpd_f (rd: freg) (**r [xor] with self = set to zero *)
- | Padds_ff (rd: freg) (r1: freg)
- | Psubs_ff (rd: freg) (r1: freg)
- | Pmuls_ff (rd: freg) (r1: freg)
- | Pdivs_ff (rd: freg) (r1: freg)
- | Pnegs (rd: freg)
- | Pabss (rd: freg)
- | Pcomiss_ff (r1 r2: freg)
- | Pxorps_f (rd: freg) (**r [xor] with self = set to zero *)
- (** Branches and calls *)
- | Pjmp_l (l: label)
- | Pjmp_s (symb: ident) (sg: signature)
- | Pjmp_r (r: ireg) (sg: signature)
- | Pjcc (c: testcond)(l: label)
- | Pjcc2 (c1 c2: testcond)(l: label) (**r pseudo *)
- | Pjmptbl (r: ireg) (tbl: list label) (**r pseudo *)
- | Pcall_s (symb: ident) (sg: signature)
- | Pcall_r (r: ireg) (sg: signature)
- | Pret
- (** Saving and restoring registers *)
- | Pmov_rm_a (rd: ireg) (a: addrmode) (**r like [Pmov_rm], using [Many64] chunk *)
- | Pmov_mr_a (a: addrmode) (rs: ireg) (**r like [Pmov_mr], using [Many64] chunk *)
- | Pmovsd_fm_a (rd: freg) (a: addrmode) (**r like [Pmovsd_fm], using [Many64] chunk *)
- | Pmovsd_mf_a (a: addrmode) (r1: freg) (**r like [Pmovsd_mf], using [Many64] chunk *)
- (** Pseudo-instructions *)
- | Plabel(l: label)
- | Pallocframe(sz: Z)(ofs_ra ofs_link: ptrofs)
- | Pfreeframe(sz: Z)(ofs_ra ofs_link: ptrofs)
- | Pbuiltin(ef: external_function)(args: list (builtin_arg preg))(res: builtin_res preg)
- (** Instructions not generated by [Asmgen] -- TO CHECK *)
- | Padcl_ri (rd: ireg) (n: int)
- | Padcl_rr (rd: ireg) (r2: ireg)
- | Paddl_mi (a: addrmode) (n: int)
- | Paddl_rr (rd: ireg) (r2: ireg)
- | Pbsfl (rd: ireg) (r1: ireg)
- | Pbsfq (rd: ireg) (r1: ireg)
- | Pbsrl (rd: ireg) (r1: ireg)
- | Pbsrq (rd: ireg) (r1: ireg)
- | Pbswap64 (rd: ireg)
- | Pbswap32 (rd: ireg)
- | Pbswap16 (rd: ireg)
- | Pcfi_adjust (n: int)
- | Pfmadd132 (rd: freg) (r2: freg) (r3: freg)
- | Pfmadd213 (rd: freg) (r2: freg) (r3: freg)
- | Pfmadd231 (rd: freg) (r2: freg) (r3: freg)
- | Pfmsub132 (rd: freg) (r2: freg) (r3: freg)
- | Pfmsub213 (rd: freg) (r2: freg) (r3: freg)
- | Pfmsub231 (rd: freg) (r2: freg) (r3: freg)
- | Pfnmadd132 (rd: freg) (r2: freg) (r3: freg)
- | Pfnmadd213 (rd: freg) (r2: freg) (r3: freg)
- | Pfnmadd231 (rd: freg) (r2: freg) (r3: freg)
- | Pfnmsub132 (rd: freg) (r2: freg) (r3: freg)
- | Pfnmsub213 (rd: freg) (r2: freg) (r3: freg)
- | Pfnmsub231 (rd: freg) (r2: freg) (r3: freg)
- | Pmaxsd (rd: freg) (r2: freg)
- | Pminsd (rd: freg) (r2: freg)
- | Pmovb_rm (rd: ireg) (a: addrmode)
- | Pmovsq_mr (a: addrmode) (rs: freg)
- | Pmovsq_rm (rd: freg) (a: addrmode)
- | Pmovsb
- | Pmovsw
- | Pmovw_rm (rd: ireg) (ad: addrmode)
- | Prep_movsl
- | Psbbl_rr (rd: ireg) (r2: ireg)
- | Psqrtsd (rd: freg) (r1: freg)
- | Psubl_ri (rd: ireg) (n: int)
- | Psubq_ri (rd: ireg) (n: int64).
-
-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 *)
-
-Lemma preg_eq: forall (x y: preg), {x=y} + {x<>y}.
-Proof. decide equality. apply ireg_eq. apply freg_eq. decide equality. Defined.
-
-Module PregEq.
- Definition t := preg.
- Definition eq := preg_eq.
-End PregEq.
-
-Module Pregmap := EMap(PregEq).
-
-Definition regset := Pregmap.t val.
-Definition genv := Genv.t fundef unit.
-
-Notation "a # b" := (a b) (at level 1, only parsing).
-Notation "a # b <- c" := (Pregmap.set b c a) (at level 1, b at next level).
-
-(** 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.
-
-(** 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.
-
-Section RELSEM.
-
-(** Looking up instructions in a code sequence by position. *)
-
-Fixpoint find_instr (pos: Z) (c: code) {struct c} : option instruction :=
- match c with
- | nil => None
- | i :: il => if zeq pos 0 then Some i else find_instr (pos - 1) il
- end.
-
-(** Position corresponding to a label *)
-
-Definition is_label (lbl: label) (instr: instruction) : bool :=
- match instr with
- | Plabel lbl' => if peq lbl lbl' then true else false
- | _ => false
- end.
-
-Lemma is_label_correct:
- forall lbl instr,
- if is_label lbl instr then instr = Plabel lbl else instr <> Plabel lbl.
-Proof.
- intros. destruct instr; simpl; try discriminate.
- case (peq lbl l); intro; congruence.
-Qed.
-
-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'
- end.
-
-Variable ge: genv.
-
-(** Evaluating an addressing mode *)
-
-Definition eval_addrmode32 (a: addrmode) (rs: regset) : val :=
- let '(Addrmode base ofs const) := a in
- Val.add (match base with
- | None => Vint Int.zero
- | Some r => rs r
- end)
- (Val.add (match ofs with
- | None => Vint Int.zero
- | Some(r, sc) =>
- if zeq sc 1
- then rs r
- else Val.mul (rs r) (Vint (Int.repr sc))
- end)
- (match const with
- | inl ofs => Vint (Int.repr ofs)
- | inr(id, ofs) => Genv.symbol_address ge id ofs
- end)).
-
-Definition eval_addrmode64 (a: addrmode) (rs: regset) : val :=
- let '(Addrmode base ofs const) := a in
- Val.addl (match base with
- | None => Vlong Int64.zero
- | Some r => rs r
- end)
- (Val.addl (match ofs with
- | None => Vlong Int64.zero
- | Some(r, sc) =>
- if zeq sc 1
- then rs r
- else Val.mull (rs r) (Vlong (Int64.repr sc))
- end)
- (match const with
- | inl ofs => Vlong (Int64.repr ofs)
- | inr(id, ofs) => Genv.symbol_address ge id ofs
- end)).
-
-Definition eval_addrmode (a: addrmode) (rs: regset) : val :=
- if Archi.ptr64 then eval_addrmode64 a rs else eval_addrmode32 a rs.
-
-(** Performing a comparison *)
-
-(** Integer comparison between x and y:
-- ZF = 1 if x = y, 0 if x != y
-- CF = 1 if x <u y, 0 if x >=u y
-- SF = 1 if x - y is negative, 0 if x - y is positive
-- OF = 1 if x - y overflows (signed), 0 if not
-- PF is undefined
-*)
-
-Definition compare_ints (x y: val) (rs: regset) (m: mem): regset :=
- rs #ZF <- (Val.cmpu (Mem.valid_pointer m) Ceq x y)
- #CF <- (Val.cmpu (Mem.valid_pointer m) Clt x y)
- #SF <- (Val.negative (Val.sub x y))
- #OF <- (Val.sub_overflow x y)
- #PF <- Vundef.
-
-Definition compare_longs (x y: val) (rs: regset) (m: mem): regset :=
- rs #ZF <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Ceq x y))
- #CF <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Clt x y))
- #SF <- (Val.negativel (Val.subl x y))
- #OF <- (Val.subl_overflow x y)
- #PF <- Vundef.
-
-(** Floating-point comparison between x and y:
-- ZF = 1 if x=y or unordered, 0 if x<>y
-- CF = 1 if x<y or unordered, 0 if x>=y
-- PF = 1 if unordered, 0 if ordered.
-- SF and 0F are undefined
-*)
-
-Definition compare_floats (vx vy: val) (rs: regset) : regset :=
- match vx, vy with
- | Vfloat x, Vfloat y =>
- rs #ZF <- (Val.of_bool (negb (Float.cmp Cne x y)))
- #CF <- (Val.of_bool (negb (Float.cmp Cge x y)))
- #PF <- (Val.of_bool (negb (Float.cmp Ceq x y || Float.cmp Clt x y || Float.cmp Cgt x y)))
- #SF <- Vundef
- #OF <- Vundef
- | _, _ =>
- undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF :: nil) rs
- end.
-
-Definition compare_floats32 (vx vy: val) (rs: regset) : regset :=
- match vx, vy with
- | Vsingle x, Vsingle y =>
- rs #ZF <- (Val.of_bool (negb (Float32.cmp Cne x y)))
- #CF <- (Val.of_bool (negb (Float32.cmp Cge x y)))
- #PF <- (Val.of_bool (negb (Float32.cmp Ceq x y || Float32.cmp Clt x y || Float32.cmp Cgt x y)))
- #SF <- Vundef
- #OF <- Vundef
- | _, _ =>
- undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF :: nil) rs
- end.
-
-(** Testing a condition *)
-
-Definition eval_testcond (c: testcond) (rs: regset) : option bool :=
- match c with
- | Cond_e =>
- match rs ZF with
- | Vint n => Some (Int.eq n Int.one)
- | _ => None
- end
- | Cond_ne =>
- match rs ZF with
- | Vint n => Some (Int.eq n Int.zero)
- | _ => None
- end
- | Cond_b =>
- match rs CF with
- | Vint n => Some (Int.eq n Int.one)
- | _ => None
- end
- | Cond_be =>
- match rs CF, rs ZF with
- | Vint c, Vint z => Some (Int.eq c Int.one || Int.eq z Int.one)
- | _, _ => None
- end
- | Cond_ae =>
- match rs CF with
- | Vint n => Some (Int.eq n Int.zero)
- | _ => None
- end
- | Cond_a =>
- match rs CF, rs ZF with
- | Vint c, Vint z => Some (Int.eq c Int.zero && Int.eq z Int.zero)
- | _, _ => None
- end
- | Cond_l =>
- match rs OF, rs SF with
- | Vint o, Vint s => Some (Int.eq (Int.xor o s) Int.one)
- | _, _ => None
- end
- | Cond_le =>
- match rs OF, rs SF, rs ZF with
- | Vint o, Vint s, Vint z => Some (Int.eq (Int.xor o s) Int.one || Int.eq z Int.one)
- | _, _, _ => None
- end
- | Cond_ge =>
- match rs OF, rs SF with
- | Vint o, Vint s => Some (Int.eq (Int.xor o s) Int.zero)
- | _, _ => None
- end
- | Cond_g =>
- match rs OF, rs SF, rs ZF with
- | Vint o, Vint s, Vint z => Some (Int.eq (Int.xor o s) Int.zero && Int.eq z Int.zero)
- | _, _, _ => None
- end
- | Cond_p =>
- match rs PF with
- | Vint n => Some (Int.eq n Int.one)
- | _ => None
- end
- | Cond_np =>
- match rs PF with
- | Vint n => Some (Int.eq n Int.zero)
- | _ => None
- end
- 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]).
- [nextinstr_nf] is a variant of [nextinstr] that sets condition flags
- to [Vundef] in addition to incrementing the [PC]. *)
-
-Definition nextinstr (rs: regset) :=
- rs#PC <- (Val.offset_ptr rs#PC Ptrofs.one).
-
-Definition nextinstr_nf (rs: regset) : regset :=
- nextinstr (undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF :: nil) rs).
-
-Definition goto_label (f: function) (lbl: label) (rs: regset) (m: mem) :=
- match label_pos lbl 0 (fn_code f) with
- | None => Stuck
- | Some pos =>
- match rs#PC with
- | Vptr b ofs => Next (rs#PC <- (Vptr b (Ptrofs.repr pos))) m
- | _ => Stuck
- end
- end.
-
-(** Auxiliaries for memory accesses. *)
-
-Definition exec_load (chunk: memory_chunk) (m: mem)
- (a: addrmode) (rs: regset) (rd: preg) :=
- match Mem.loadv chunk m (eval_addrmode a rs) with
- | Some v => Next (nextinstr_nf (rs#rd <- v)) m
- | None => Stuck
- end.
-
-Definition exec_store (chunk: memory_chunk) (m: mem)
- (a: addrmode) (rs: regset) (r1: preg)
- (destroyed: list preg) :=
- match Mem.storev chunk m (eval_addrmode a rs) (rs r1) with
- | Some m' => Next (nextinstr_nf (undef_regs destroyed rs)) m'
- | None => Stuck
- end.
-
-(** Execution of a single instruction [i] in initial state
- [rs] and [m]. Return updated state. For instructions
- that correspond to actual IA32 instructions, the cases are
- straightforward transliterations of the informal descriptions
- given in the IA32 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 IA32 code
- we generate cannot use those registers to hold values that must
- survive the execution of the pseudo-instruction.
-
- Concerning condition flags, the comparison instructions set them
- accurately; other instructions (moves, [lea]) preserve them;
- and all other instruction set those flags to [Vundef], to reflect
- the fact that the processor updates some or all of those flags,
- but we do not need to model this precisely.
-*)
-
-Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : outcome :=
- match i with
- (** Moves *)
- | Pmov_rr rd r1 =>
- Next (nextinstr (rs#rd <- (rs r1))) m
- | Pmovl_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Vint n))) m
- | Pmovq_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Vlong n))) m
- | Pmov_rs rd id =>
- Next (nextinstr_nf (rs#rd <- (Genv.symbol_address ge id Ptrofs.zero))) m
- | Pmovl_rm rd a =>
- exec_load Mint32 m a rs rd
- | Pmovq_rm rd a =>
- exec_load Mint64 m a rs rd
- | Pmovl_mr a r1 =>
- exec_store Mint32 m a rs r1 nil
- | Pmovq_mr a r1 =>
- exec_store Mint64 m a rs r1 nil
- | Pmovsd_ff rd r1 =>
- Next (nextinstr (rs#rd <- (rs r1))) m
- | Pmovsd_fi rd n =>
- Next (nextinstr (rs#rd <- (Vfloat n))) m
- | Pmovsd_fm rd a =>
- exec_load Mfloat64 m a rs rd
- | Pmovsd_mf a r1 =>
- exec_store Mfloat64 m a rs r1 nil
- | Pmovss_fi rd n =>
- Next (nextinstr (rs#rd <- (Vsingle n))) m
- | Pmovss_fm rd a =>
- exec_load Mfloat32 m a rs rd
- | Pmovss_mf a r1 =>
- exec_store Mfloat32 m a rs r1 nil
- | Pfldl_m a =>
- exec_load Mfloat64 m a rs ST0
- | Pfstpl_m a =>
- exec_store Mfloat64 m a rs ST0 (ST0 :: nil)
- | Pflds_m a =>
- exec_load Mfloat32 m a rs ST0
- | Pfstps_m a =>
- exec_store Mfloat32 m a rs ST0 (ST0 :: nil)
- (** Moves with conversion *)
- | Pmovb_mr a r1 =>
- exec_store Mint8unsigned m a rs r1 nil
- | Pmovw_mr a r1 =>
- exec_store Mint16unsigned m a rs r1 nil
- | Pmovzb_rr rd r1 =>
- Next (nextinstr (rs#rd <- (Val.zero_ext 8 rs#r1))) m
- | Pmovzb_rm rd a =>
- exec_load Mint8unsigned m a rs rd
- | Pmovsb_rr rd r1 =>
- Next (nextinstr (rs#rd <- (Val.sign_ext 8 rs#r1))) m
- | Pmovsb_rm rd a =>
- exec_load Mint8signed m a rs rd
- | Pmovzw_rr rd r1 =>
- Next (nextinstr (rs#rd <- (Val.zero_ext 16 rs#r1))) m
- | Pmovzw_rm rd a =>
- exec_load Mint16unsigned m a rs rd
- | Pmovsw_rr rd r1 =>
- Next (nextinstr (rs#rd <- (Val.sign_ext 16 rs#r1))) m
- | Pmovsw_rm rd a =>
- exec_load Mint16signed m a rs rd
- | Pmovzl_rr rd r1 =>
- Next (nextinstr (rs#rd <- (Val.longofintu rs#r1))) m
- | Pmovsl_rr rd r1 =>
- Next (nextinstr (rs#rd <- (Val.longofint rs#r1))) m
- | Pmovls_rr rd =>
- Next (nextinstr (rs#rd <- (Val.loword rs#rd))) m
- | Pcvtsd2ss_ff rd r1 =>
- Next (nextinstr (rs#rd <- (Val.singleoffloat rs#r1))) m
- | Pcvtss2sd_ff rd r1 =>
- Next (nextinstr (rs#rd <- (Val.floatofsingle rs#r1))) m
- | Pcvttsd2si_rf rd r1 =>
- Next (nextinstr (rs#rd <- (Val.maketotal (Val.intoffloat rs#r1)))) m
- | Pcvtsi2sd_fr rd r1 =>
- Next (nextinstr (rs#rd <- (Val.maketotal (Val.floatofint rs#r1)))) m
- | Pcvttss2si_rf rd r1 =>
- Next (nextinstr (rs#rd <- (Val.maketotal (Val.intofsingle rs#r1)))) m
- | Pcvtsi2ss_fr rd r1 =>
- Next (nextinstr (rs#rd <- (Val.maketotal (Val.singleofint rs#r1)))) m
- | Pcvttsd2sl_rf rd r1 =>
- Next (nextinstr (rs#rd <- (Val.maketotal (Val.longoffloat rs#r1)))) m
- | Pcvtsl2sd_fr rd r1 =>
- Next (nextinstr (rs#rd <- (Val.maketotal (Val.floatoflong rs#r1)))) m
- | Pcvttss2sl_rf rd r1 =>
- Next (nextinstr (rs#rd <- (Val.maketotal (Val.longofsingle rs#r1)))) m
- | Pcvtsl2ss_fr rd r1 =>
- Next (nextinstr (rs#rd <- (Val.maketotal (Val.singleoflong rs#r1)))) m
- (** Integer arithmetic *)
- | Pleal rd a =>
- Next (nextinstr (rs#rd <- (eval_addrmode32 a rs))) m
- | Pleaq rd a =>
- Next (nextinstr (rs#rd <- (eval_addrmode64 a rs))) m
- | Pnegl rd =>
- Next (nextinstr_nf (rs#rd <- (Val.neg rs#rd))) m
- | Pnegq rd =>
- Next (nextinstr_nf (rs#rd <- (Val.negl rs#rd))) m
- | Paddl_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.add rs#rd (Vint n)))) m
- | Paddq_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.addl rs#rd (Vlong n)))) m
- | Psubl_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.sub rs#rd rs#r1))) m
- | Psubq_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.subl rs#rd rs#r1))) m
- | Pimull_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.mul rs#rd rs#r1))) m
- | Pimulq_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.mull rs#rd rs#r1))) m
- | Pimull_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.mul rs#rd (Vint n)))) m
- | Pimulq_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.mull rs#rd (Vlong n)))) m
- | Pimull_r r1 =>
- Next (nextinstr_nf (rs#RAX <- (Val.mul rs#RAX rs#r1)
- #RDX <- (Val.mulhs rs#RAX rs#r1))) m
- | Pimulq_r r1 =>
- Next (nextinstr_nf (rs#RAX <- (Val.mull rs#RAX rs#r1)
- #RDX <- (Val.mullhs rs#RAX rs#r1))) m
- | Pmull_r r1 =>
- Next (nextinstr_nf (rs#RAX <- (Val.mul rs#RAX rs#r1)
- #RDX <- (Val.mulhu rs#RAX rs#r1))) m
- | Pmulq_r r1 =>
- Next (nextinstr_nf (rs#RAX <- (Val.mull rs#RAX rs#r1)
- #RDX <- (Val.mullhu rs#RAX rs#r1))) m
- | Pcltd =>
- Next (nextinstr_nf (rs#RDX <- (Val.shr rs#RAX (Vint (Int.repr 31))))) m
- | Pcqto =>
- Next (nextinstr_nf (rs#RDX <- (Val.shrl rs#RAX (Vint (Int.repr 63))))) m
- | Pdivl r1 =>
- match rs#RDX, rs#RAX, rs#r1 with
- | Vint nh, Vint nl, Vint d =>
- match Int.divmodu2 nh nl d with
- | Some(q, r) => Next (nextinstr_nf (rs#RAX <- (Vint q) #RDX <- (Vint r))) m
- | None => Stuck
- end
- | _, _, _ => Stuck
- end
- | Pdivq r1 =>
- match rs#RDX, rs#RAX, rs#r1 with
- | Vlong nh, Vlong nl, Vlong d =>
- match Int64.divmodu2 nh nl d with
- | Some(q, r) => Next (nextinstr_nf (rs#RAX <- (Vlong q) #RDX <- (Vlong r))) m
- | None => Stuck
- end
- | _, _, _ => Stuck
- end
- | Pidivl r1 =>
- match rs#RDX, rs#RAX, rs#r1 with
- | Vint nh, Vint nl, Vint d =>
- match Int.divmods2 nh nl d with
- | Some(q, r) => Next (nextinstr_nf (rs#RAX <- (Vint q) #RDX <- (Vint r))) m
- | None => Stuck
- end
- | _, _, _ => Stuck
- end
- | Pidivq r1 =>
- match rs#RDX, rs#RAX, rs#r1 with
- | Vlong nh, Vlong nl, Vlong d =>
- match Int64.divmods2 nh nl d with
- | Some(q, r) => Next (nextinstr_nf (rs#RAX <- (Vlong q) #RDX <- (Vlong r))) m
- | None => Stuck
- end
- | _, _, _ => Stuck
- end
- | Pandl_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.and rs#rd rs#r1))) m
- | Pandq_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.andl rs#rd rs#r1))) m
- | Pandl_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.and rs#rd (Vint n)))) m
- | Pandq_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.andl rs#rd (Vlong n)))) m
- | Porl_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.or rs#rd rs#r1))) m
- | Porq_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.orl rs#rd rs#r1))) m
- | Porl_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.or rs#rd (Vint n)))) m
- | Porq_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.orl rs#rd (Vlong n)))) m
- | Pxorl_r rd =>
- Next (nextinstr_nf (rs#rd <- Vzero)) m
- | Pxorq_r rd =>
- Next (nextinstr_nf (rs#rd <- (Vlong Int64.zero))) m
- | Pxorl_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.xor rs#rd rs#r1))) m
- | Pxorq_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.xorl rs#rd rs#r1))) m
- | Pxorl_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.xor rs#rd (Vint n)))) m
- | Pxorq_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.xorl rs#rd (Vlong n)))) m
- | Pnotl rd =>
- Next (nextinstr_nf (rs#rd <- (Val.notint rs#rd))) m
- | Pnotq rd =>
- Next (nextinstr_nf (rs#rd <- (Val.notl rs#rd))) m
- | Psall_rcl rd =>
- Next (nextinstr_nf (rs#rd <- (Val.shl rs#rd rs#RCX))) m
- | Psalq_rcl rd =>
- Next (nextinstr_nf (rs#rd <- (Val.shll rs#rd rs#RCX))) m
- | Psall_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.shl rs#rd (Vint n)))) m
- | Psalq_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.shll rs#rd (Vint n)))) m
- | Pshrl_rcl rd =>
- Next (nextinstr_nf (rs#rd <- (Val.shru rs#rd rs#RCX))) m
- | Pshrq_rcl rd =>
- Next (nextinstr_nf (rs#rd <- (Val.shrlu rs#rd rs#RCX))) m
- | Pshrl_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.shru rs#rd (Vint n)))) m
- | Pshrq_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.shrlu rs#rd (Vint n)))) m
- | Psarl_rcl rd =>
- Next (nextinstr_nf (rs#rd <- (Val.shr rs#rd rs#RCX))) m
- | Psarq_rcl rd =>
- Next (nextinstr_nf (rs#rd <- (Val.shrl rs#rd rs#RCX))) m
- | Psarl_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.shr rs#rd (Vint n)))) m
- | Psarq_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.shrl rs#rd (Vint n)))) m
- | Pshld_ri rd r1 n =>
- Next (nextinstr_nf
- (rs#rd <- (Val.or (Val.shl rs#rd (Vint n))
- (Val.shru rs#r1 (Vint (Int.sub Int.iwordsize n)))))) m
- | Prorl_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.ror rs#rd (Vint n)))) m
- | Prorq_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.rorl rs#rd (Vint n)))) m
- | Pcmpl_rr r1 r2 =>
- Next (nextinstr (compare_ints (rs r1) (rs r2) rs m)) m
- | Pcmpq_rr r1 r2 =>
- Next (nextinstr (compare_longs (rs r1) (rs r2) rs m)) m
- | Pcmpl_ri r1 n =>
- Next (nextinstr (compare_ints (rs r1) (Vint n) rs m)) m
- | Pcmpq_ri r1 n =>
- Next (nextinstr (compare_longs (rs r1) (Vlong n) rs m)) m
- | Ptestl_rr r1 r2 =>
- Next (nextinstr (compare_ints (Val.and (rs r1) (rs r2)) Vzero rs m)) m
- | Ptestq_rr r1 r2 =>
- Next (nextinstr (compare_longs (Val.andl (rs r1) (rs r2)) (Vlong Int64.zero) rs m)) m
- | Ptestl_ri r1 n =>
- Next (nextinstr (compare_ints (Val.and (rs r1) (Vint n)) Vzero rs m)) m
- | Ptestq_ri r1 n =>
- Next (nextinstr (compare_longs (Val.andl (rs r1) (Vlong n)) (Vlong Int64.zero) rs m)) m
- | Pcmov c rd r1 =>
- match eval_testcond c rs with
- | Some true => Next (nextinstr (rs#rd <- (rs#r1))) m
- | Some false => Next (nextinstr rs) m
- | None => Next (nextinstr (rs#rd <- Vundef)) m
- end
- | Psetcc c rd =>
- Next (nextinstr (rs#rd <- (Val.of_optbool (eval_testcond c rs)))) m
- (** Arithmetic operations over double-precision floats *)
- | Paddd_ff rd r1 =>
- Next (nextinstr (rs#rd <- (Val.addf rs#rd rs#r1))) m
- | Psubd_ff rd r1 =>
- Next (nextinstr (rs#rd <- (Val.subf rs#rd rs#r1))) m
- | Pmuld_ff rd r1 =>
- Next (nextinstr (rs#rd <- (Val.mulf rs#rd rs#r1))) m
- | Pdivd_ff rd r1 =>
- Next (nextinstr (rs#rd <- (Val.divf rs#rd rs#r1))) m
- | Pnegd rd =>
- Next (nextinstr (rs#rd <- (Val.negf rs#rd))) m
- | Pabsd rd =>
- Next (nextinstr (rs#rd <- (Val.absf rs#rd))) m
- | Pcomisd_ff r1 r2 =>
- Next (nextinstr (compare_floats (rs r1) (rs r2) rs)) m
- | Pxorpd_f rd =>
- Next (nextinstr_nf (rs#rd <- (Vfloat Float.zero))) m
- (** Arithmetic operations over single-precision floats *)
- | Padds_ff rd r1 =>
- Next (nextinstr (rs#rd <- (Val.addfs rs#rd rs#r1))) m
- | Psubs_ff rd r1 =>
- Next (nextinstr (rs#rd <- (Val.subfs rs#rd rs#r1))) m
- | Pmuls_ff rd r1 =>
- Next (nextinstr (rs#rd <- (Val.mulfs rs#rd rs#r1))) m
- | Pdivs_ff rd r1 =>
- Next (nextinstr (rs#rd <- (Val.divfs rs#rd rs#r1))) m
- | Pnegs rd =>
- Next (nextinstr (rs#rd <- (Val.negfs rs#rd))) m
- | Pabss rd =>
- Next (nextinstr (rs#rd <- (Val.absfs rs#rd))) m
- | Pcomiss_ff r1 r2 =>
- Next (nextinstr (compare_floats32 (rs r1) (rs r2) rs)) m
- | Pxorps_f rd =>
- Next (nextinstr_nf (rs#rd <- (Vsingle Float32.zero))) m
- (** Branches and calls *)
- | Pjmp_l lbl =>
- goto_label f lbl rs m
- | Pjmp_s id sg =>
- Next (rs#PC <- (Genv.symbol_address ge id Ptrofs.zero)) m
- | Pjmp_r r sg =>
- Next (rs#PC <- (rs r)) m
- | Pjcc 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
- | Pjcc2 cond1 cond2 lbl =>
- match eval_testcond cond1 rs, eval_testcond cond2 rs with
- | Some true, Some true => goto_label f lbl rs m
- | Some _, Some _ => Next (nextinstr rs) m
- | _, _ => Stuck
- end
- | Pjmptbl r tbl =>
- match rs#r with
- | Vint n =>
- match list_nth_z tbl (Int.unsigned n) with
- | None => Stuck
- | Some lbl => goto_label f lbl (rs #RAX <- Vundef #RDX <- Vundef) m
- end
- | _ => Stuck
- end
- | Pcall_s id sg =>
- Next (rs#RA <- (Val.offset_ptr rs#PC Ptrofs.one) #PC <- (Genv.symbol_address ge id Ptrofs.zero)) m
- | Pcall_r r sg =>
- Next (rs#RA <- (Val.offset_ptr rs#PC Ptrofs.one) #PC <- (rs r)) m
- | Pret =>
- Next (rs#PC <- (rs#RA)) m
- (** Saving and restoring registers *)
- | Pmov_rm_a rd a =>
- exec_load (if Archi.ptr64 then Many64 else Many32) m a rs rd
- | Pmov_mr_a a r1 =>
- exec_store (if Archi.ptr64 then Many64 else Many32) m a rs r1 nil
- | Pmovsd_fm_a rd a =>
- exec_load Many64 m a rs rd
- | Pmovsd_mf_a a r1 =>
- exec_store Many64 m a rs r1 nil
- (** Pseudo-instructions *)
- | Plabel lbl =>
- Next (nextinstr rs) m
- | Pallocframe sz ofs_ra ofs_link =>
- let (m1, stk) := Mem.alloc m 0 sz in
- let sp := Vptr stk Ptrofs.zero in
- match Mem.storev Mptr m1 (Val.offset_ptr sp ofs_link) rs#RSP with
- | None => Stuck
- | Some m2 =>
- match Mem.storev Mptr m2 (Val.offset_ptr sp ofs_ra) rs#RA with
- | None => Stuck
- | Some m3 => Next (nextinstr (rs #RAX <- (rs#RSP) #RSP <- sp)) m3
- end
- end
- | Pfreeframe sz ofs_ra ofs_link =>
- match Mem.loadv Mptr m (Val.offset_ptr rs#RSP ofs_ra) with
- | None => Stuck
- | Some ra =>
- match Mem.loadv Mptr m (Val.offset_ptr rs#RSP ofs_link) with
- | None => Stuck
- | Some sp =>
- match rs#RSP with
- | Vptr stk ofs =>
- match Mem.free m stk 0 sz with
- | None => Stuck
- | Some m' => Next (nextinstr (rs#RSP <- sp #RA <- ra)) m'
- end
- | _ => Stuck
- end
- end
- 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. *)
- | Padcl_ri _ _
- | Padcl_rr _ _
- | Paddl_mi _ _
- | Paddl_rr _ _
- | Pbsfl _ _
- | Pbsfq _ _
- | Pbsrl _ _
- | Pbsrq _ _
- | Pbswap64 _
- | Pbswap32 _
- | Pbswap16 _
- | Pcfi_adjust _
- | Pfmadd132 _ _ _
- | Pfmadd213 _ _ _
- | Pfmadd231 _ _ _
- | Pfmsub132 _ _ _
- | Pfmsub213 _ _ _
- | Pfmsub231 _ _ _
- | Pfnmadd132 _ _ _
- | Pfnmadd213 _ _ _
- | Pfnmadd231 _ _ _
- | Pfnmsub132 _ _ _
- | Pfnmsub213 _ _ _
- | Pfnmsub231 _ _ _
- | Pmaxsd _ _
- | Pminsd _ _
- | Pmovb_rm _ _
- | Pmovsq_rm _ _
- | Pmovsq_mr _ _
- | Pmovsb
- | Pmovsw
- | Pmovw_rm _ _
- | Prep_movsl
- | Psbbl_rr _ _
- | Psqrtsd _ _
- | Psubl_ri _ _
- | Psubq_ri _ _ => Stuck
- end.
-
-(** Translation of the LTL/Linear/Mach view of machine registers
- to the Asm view. *)
-
-Definition preg_of (r: mreg) : preg :=
- match r with
- | AX => IR RAX
- | BX => IR RBX
- | CX => IR RCX
- | DX => IR RDX
- | SI => IR RSI
- | DI => IR RDI
- | BP => IR RBP
- | Machregs.R8 => IR R8
- | Machregs.R9 => IR R9
- | Machregs.R10 => IR R10
- | Machregs.R11 => IR R11
- | Machregs.R12 => IR R12
- | Machregs.R13 => IR R13
- | Machregs.R14 => IR R14
- | Machregs.R15 => IR R15
- | X0 => FR XMM0
- | X1 => FR XMM1
- | X2 => FR XMM2
- | X3 => FR XMM3
- | X4 => FR XMM4
- | X5 => FR XMM5
- | X6 => FR XMM6
- | X7 => FR XMM7
- | X8 => FR XMM8
- | X9 => FR XMM9
- | X10 => FR XMM10
- | X11 => FR XMM11
- | X12 => FR XMM12
- | X13 => FR XMM13
- | X14 => FR XMM14
- | X15 => FR XMM15
- | FP0 => ST0
- end.
-
-(** Extract the values of the arguments of an external call.
- We exploit the calling conventions from module [Conventions], except that
- we use machine 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 (IR RSP)) (Ptrofs.repr bofs)) = Some v ->
- extcall_arg rs m (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',
- rs PC = Vptr b ofs ->
- Genv.find_funct_ptr ge b = Some (Internal f) ->
- find_instr (Ptrofs.unsigned ofs) f.(fn_code) = Some i ->
- exec_instr f i rs m = Next rs' m' ->
- step (State rs m) E0 (State rs' m')
- | exec_step_builtin:
- forall b ofs f ef args res rs m vargs t vres 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 (Pbuiltin ef args res) ->
- eval_builtin_args ge rs (rs RSP) m args vargs ->
- external_call ef ge vargs m t vres m' ->
- rs' = nextinstr_nf
- (set_res res vres
- (undef_regs (map preg_of (destroyed_by_builtin ef)) rs)) ->
- 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) ->
- extcall_arguments rs m (ef_sig ef) args ->
- external_call ef ge args m t res m' ->
- rs' = (set_pair (loc_external_result (ef_sig ef)) res 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
- # RSP <- 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#RAX = Vint r ->
- final_state (State rs m) r.
-
-Definition semantics (p: program) :=
- Semantics step (initial_state p) final_state (Genv.globalenv p).
-
-(** Determinacy of the [Asm] semantics. *)
-
-Remark extcall_arguments_determ:
- forall rs m sg args1 args2,
- extcall_arguments rs m sg args1 -> extcall_arguments rs m sg args2 -> args1 = args2.
-Proof.
- intros until m.
- assert (A: forall l v1 v2,
- extcall_arg rs m l v1 -> extcall_arg rs m l v2 -> v1 = v2).
- { intros. inv H; inv H0; congruence. }
- assert (B: forall p v1 v2,
- extcall_arg_pair rs m p v1 -> extcall_arg_pair rs m p v2 -> v1 = v2).
- { intros. inv H; inv H0.
- eapply A; eauto.
- f_equal; eapply A; eauto. }
- assert (C: forall ll vl1, list_forall2 (extcall_arg_pair rs m) ll vl1 ->
- forall vl2, list_forall2 (extcall_arg_pair rs m) ll vl2 -> vl1 = vl2).
- {
- induction 1; intros vl2 EA; inv EA.
- auto.
- f_equal; eauto. }
- intros. eapply C; eauto.
-Qed.
-
-Lemma semantics_determinate: forall p, determinate (semantics p).
-Proof.
-Ltac Equalities :=
- match goal with
- | [ H1: ?a = ?b, H2: ?a = ?c |- _ ] =>
- rewrite H1 in H2; inv H2; Equalities
- | _ => idtac
- end.
- intros; constructor; simpl; intros.
-- (* determ *)
- inv H; inv H0; Equalities.
-+ split. constructor. auto.
-+ discriminate.
-+ discriminate.
-+ assert (vargs0 = vargs) by (eapply eval_builtin_args_determ; eauto). subst vargs0.
- exploit external_call_determ. eexact H5. eexact H11. intros [A B].
- split. auto. intros. destruct B; auto. subst. auto.
-+ assert (args0 = args) by (eapply extcall_arguments_determ; eauto). subst args0.
- exploit external_call_determ. eexact H4. eexact H9. intros [A B].
- split. auto. intros. destruct B; auto. subst. auto.
-- (* trace length *)
- red; intros; inv H; simpl.
- omega.
- eapply external_call_trace_length; eauto.
- eapply external_call_trace_length; eauto.
-- (* initial states *)
- inv H; inv H0. f_equal. congruence.
-- (* final no step *)
- assert (NOTNULL: forall b ofs, Vnullptr <> Vptr b ofs).
- { intros; unfold Vnullptr; destruct Archi.ptr64; congruence. }
- inv H. red; intros; red; intros. inv H; rewrite H0 in *; eelim NOTNULL; eauto.
-- (* final states *)
- inv H; inv H0. congruence.
-Qed.
-
-(** Classification functions for processor registers (used in Asmgenproof). *)
-
-Definition data_preg (r: preg) : bool :=
- match r with
- | PC => false
- | IR _ => true
- | FR _ => true
- | ST0 => true
- | CR _ => false
- | RA => false
- end.
-
diff --git a/ia32/AsmToJSON.ml b/ia32/AsmToJSON.ml
deleted file mode 100644
index 3214491f..00000000
--- a/ia32/AsmToJSON.ml
+++ /dev/null
@@ -1,17 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Bernhard Schommer, AbsInt Angewandte Informatik GmbH *)
-(* *)
-(* AbsInt Angewandte Informatik GmbH. All rights reserved. This file *)
-(* is distributed under the terms of the INRIA Non-Commercial *)
-(* License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-(* Simple functions to serialize ia32 Asm to JSON *)
-
-(* Dummy function *)
-let p_program oc prog =
- ()
diff --git a/ia32/AsmToJSON.mli b/ia32/AsmToJSON.mli
deleted file mode 100644
index 20bcba5e..00000000
--- a/ia32/AsmToJSON.mli
+++ /dev/null
@@ -1,13 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Bernhard Schommer, AbsInt Angewandte Informatik GmbH *)
-(* *)
-(* AbsInt Angewandte Informatik GmbH. All rights reserved. This file *)
-(* is distributed under the terms of the INRIA Non-Commercial *)
-(* License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-val p_program: out_channel -> (Asm.coq_function AST.fundef, 'a) AST.program -> unit
diff --git a/ia32/Asmexpand.ml b/ia32/Asmexpand.ml
deleted file mode 100644
index 0436bc86..00000000
--- a/ia32/Asmexpand.ml
+++ /dev/null
@@ -1,638 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* Bernhard Schommer, AbsInt Angewandte Informatik GmbH *)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(* Expanding built-ins and some pseudo-instructions by rewriting
- of the IA32 assembly code. *)
-
-open Asm
-open Asmexpandaux
-open AST
-open Camlcoq
-open Datatypes
-
-exception Error of string
-
-(* Useful constants and helper functions *)
-
-let _0 = Integers.Int.zero
-let _1 = Integers.Int.one
-let _2 = coqint_of_camlint 2l
-let _4 = coqint_of_camlint 4l
-let _8 = coqint_of_camlint 8l
-
-let _0z = Z.zero
-let _1z = Z.one
-let _2z = Z.of_sint 2
-let _4z = Z.of_sint 4
-let _8z = Z.of_sint 8
-let _16z = Z.of_sint 16
-
-let stack_alignment () = 16
-
-(* Pseudo instructions for 32/64 bit compatibility *)
-
-let _Plea (r, addr) =
- if Archi.ptr64 then Pleaq (r, addr) else Pleal (r, addr)
-
-(* SP adjustment to allocate or free a stack frame *)
-
-let align n a =
- if n >= 0 then (n + a - 1) land (-a) else n land (-a)
-
-let sp_adjustment_32 sz =
- let sz = Z.to_int sz in
- (* Preserve proper alignment of the stack *)
- let sz = align sz (stack_alignment ()) in
- (* The top 4 bytes have already been allocated by the "call" instruction. *)
- sz - 4
-
-let sp_adjustment_64 sz =
- let sz = Z.to_int sz in
- if is_current_function_variadic() then begin
- (* If variadic, add room for register save area, which must be 16-aligned *)
- let ofs = align (sz - 8) 16 in
- let sz = ofs + 176 (* save area *) + 8 (* return address *) in
- (* Preserve proper alignment of the stack *)
- let sz = align sz 16 in
- (* The top 8 bytes have already been allocated by the "call" instruction. *)
- (sz - 8, ofs)
- end else begin
- (* Preserve proper alignment of the stack *)
- let sz = align sz 16 in
- (* The top 8 bytes have already been allocated by the "call" instruction. *)
- (sz - 8, -1)
- end
-
-(* Built-ins. They come in two flavors:
- - annotation statements: take their arguments in registers or stack
- locations; generate no code;
- - inlined by the compiler: take their arguments in arbitrary
- registers; preserve all registers except ECX, EDX, XMM6 and XMM7. *)
-
-(* Handling of annotations *)
-
-let expand_annot_val txt targ args res =
- emit (Pbuiltin (EF_annot(txt,[targ]), args, BR_none));
- match args, res with
- | [BA(IR src)], BR(IR dst) ->
- if dst <> src then emit (Pmov_rr (dst,src))
- | [BA(FR src)], BR(FR dst) ->
- if dst <> src then emit (Pmovsd_ff (dst,src))
- | _, _ ->
- raise (Error "ill-formed __builtin_annot_intval")
-
-(* Operations on addressing modes *)
-
-let offset_addressing (Addrmode(base, ofs, cst)) delta =
- Addrmode(base, ofs,
- match cst with
- | Coq_inl n -> Coq_inl(Z.add n delta)
- | Coq_inr(id, n) -> Coq_inr(id, Integers.Ptrofs.add n delta))
-
-let linear_addr reg ofs = Addrmode(Some reg, None, Coq_inl ofs)
-let global_addr id ofs = Addrmode(None, None, Coq_inr(id, ofs))
-
-(* Translate a builtin argument into an addressing mode *)
-
-let addressing_of_builtin_arg = function
- | BA (IR r) -> linear_addr r Z.zero
- | BA_addrstack ofs -> linear_addr RSP (Integers.Ptrofs.unsigned ofs)
- | BA_addrglobal(id, ofs) -> global_addr id ofs
- | _ -> assert false
-
-(* Handling of memcpy *)
-
-(* Unaligned memory accesses are quite fast on IA32, so use large
- memory accesses regardless of alignment. *)
-
-let expand_builtin_memcpy_small sz al src dst =
- let rec copy src dst sz =
- if sz >= 8 && Archi.ptr64 then begin
- emit (Pmovq_rm (RCX, src));
- emit (Pmovq_mr (dst, RCX));
- copy (offset_addressing src _8z) (offset_addressing dst _8z) (sz - 8)
- end else if sz >= 8 && !Clflags.option_ffpu then begin
- emit (Pmovsq_rm (XMM7, src));
- emit (Pmovsq_mr (dst, XMM7));
- copy (offset_addressing src _8z) (offset_addressing dst _8z) (sz - 8)
- end else if sz >= 4 then begin
- emit (Pmovl_rm (RCX, src));
- emit (Pmovl_mr (dst, RCX));
- copy (offset_addressing src _4z) (offset_addressing dst _4z) (sz - 4)
- end else if sz >= 2 then begin
- emit (Pmovw_rm (RCX, src));
- emit (Pmovw_mr (dst, RCX));
- copy (offset_addressing src _2z) (offset_addressing dst _2z) (sz - 2)
- end else if sz >= 1 then begin
- emit (Pmovb_rm (RCX, src));
- emit (Pmovb_mr (dst, RCX));
- copy (offset_addressing src _1z) (offset_addressing dst _1z) (sz - 1)
- end in
- copy (addressing_of_builtin_arg src) (addressing_of_builtin_arg dst) sz
-
-let expand_builtin_memcpy_big sz al src dst =
- if src <> BA (IR RSI) then emit (_Plea (RSI, addressing_of_builtin_arg src));
- if dst <> BA (IR RDI) then emit (_Plea (RDI, addressing_of_builtin_arg dst));
- (* TODO: movsq? *)
- emit (Pmovl_ri (RCX,coqint_of_camlint (Int32.of_int (sz / 4))));
- emit Prep_movsl;
- if sz mod 4 >= 2 then emit Pmovsw;
- if sz mod 2 >= 1 then emit Pmovsb
-
-let expand_builtin_memcpy sz al args =
- let (dst, src) = match args with [d; s] -> (d, s) | _ -> assert false in
- if sz <= 32
- then expand_builtin_memcpy_small sz al src dst
- else expand_builtin_memcpy_big sz al src dst
-
-(* Handling of volatile reads and writes *)
-
-let expand_builtin_vload_common chunk addr res =
- match chunk, res with
- | Mint8unsigned, BR(IR res) ->
- emit (Pmovzb_rm (res,addr))
- | Mint8signed, BR(IR res) ->
- emit (Pmovsb_rm (res,addr))
- | Mint16unsigned, BR(IR res) ->
- emit (Pmovzw_rm (res,addr))
- | Mint16signed, BR(IR res) ->
- emit (Pmovsw_rm (res,addr))
- | Mint32, BR(IR res) ->
- emit (Pmovl_rm (res,addr))
- | Mint64, BR(IR res) ->
- emit (Pmovq_rm (res,addr))
- | Mint64, BR_splitlong(BR(IR res1), BR(IR res2)) ->
- let addr' = offset_addressing addr _4z in
- if not (Asmgen.addressing_mentions addr res2) then begin
- emit (Pmovl_rm (res2,addr));
- emit (Pmovl_rm (res1,addr'))
- end else begin
- emit (Pmovl_rm (res1,addr'));
- emit (Pmovl_rm (res2,addr))
- end
- | Mfloat32, BR(FR res) ->
- emit (Pmovss_fm (res,addr))
- | Mfloat64, BR(FR res) ->
- emit (Pmovsd_fm (res,addr))
- | _ ->
- assert false
-
-let expand_builtin_vload chunk args res =
- match args with
- | [addr] ->
- expand_builtin_vload_common chunk (addressing_of_builtin_arg addr) res
- | _ ->
- assert false
-
-let expand_builtin_vstore_common chunk addr src tmp =
- match chunk, src with
- | (Mint8signed | Mint8unsigned), BA(IR src) ->
- if Archi.ptr64 || Asmgen.low_ireg src then
- emit (Pmovb_mr (addr,src))
- else begin
- emit (Pmov_rr (tmp,src));
- emit (Pmovb_mr (addr,tmp))
- end
- | (Mint16signed | Mint16unsigned), BA(IR src) ->
- emit (Pmovw_mr (addr,src))
- | Mint32, BA(IR src) ->
- emit (Pmovl_mr (addr,src))
- | Mint64, BA(IR src) ->
- emit (Pmovq_mr (addr,src))
- | Mint64, BA_splitlong(BA(IR src1), BA(IR src2)) ->
- let addr' = offset_addressing addr _4z in
- emit (Pmovl_mr (addr,src2));
- emit (Pmovl_mr (addr',src1))
- | Mfloat32, BA(FR src) ->
- emit (Pmovss_mf (addr,src))
- | Mfloat64, BA(FR src) ->
- emit (Pmovsd_mf (addr,src))
- | _ ->
- assert false
-
-let expand_builtin_vstore chunk args =
- match args with
- | [addr; src] ->
- let addr = addressing_of_builtin_arg addr in
- expand_builtin_vstore_common chunk addr src
- (if Asmgen.addressing_mentions addr RAX then RCX else RAX)
- | _ -> assert false
-
-(* Handling of varargs *)
-
-let rec next_arg_locations ir fr ofs = function
- | [] ->
- (ir, fr, ofs)
- | (Tint | Tlong | Tany32 | Tany64) :: l ->
- if ir < 6
- then next_arg_locations (ir + 1) fr ofs l
- else next_arg_locations ir fr (ofs + 8) l
- | (Tfloat | Tsingle) :: l ->
- if fr < 8
- then next_arg_locations ir (fr + 1) ofs l
- else next_arg_locations ir fr (ofs + 8) l
-
-let current_function_stacksize = ref 0L
-
-let expand_builtin_va_start_32 r =
- if not (is_current_function_variadic ()) then
- invalid_arg "Fatal error: va_start used in non-vararg function";
- let ofs =
- Int32.(add (add !PrintAsmaux.current_function_stacksize 4l)
- (mul 4l (Z.to_int32 (Conventions1.size_arguments
- (get_current_function_sig ()))))) in
- emit (Pleal (RAX, linear_addr RSP (Z.of_uint32 ofs)));
- emit (Pmovl_mr (linear_addr r _0z, RAX))
-
-let expand_builtin_va_start_64 r =
- if not (is_current_function_variadic ()) then
- invalid_arg "Fatal error: va_start used in non-vararg function";
- let (ir, fr, ofs) =
- next_arg_locations 0 0 0 (get_current_function_args ()) in
- (* [r] points to the following struct:
- struct {
- unsigned int gp_offset;
- unsigned int fp_offset;
- void *overflow_arg_area;
- void *reg_save_area;
- }
- gp_offset is initialized to ir * 8
- fp_offset is initialized to 6 * 8 + fr * 16
- overflow_arg_area is initialized to sp + current stacksize + ofs
- reg_save_area is initialized to
- sp + current stacksize - 16 - save area size (6 * 8 + 8 * 16) *)
- let gp_offset = Int32.of_int (ir * 8)
- and fp_offset = Int32.of_int (6 * 8 + fr * 16)
- and overflow_arg_area = Int64.(add !current_function_stacksize (of_int ofs))
- and reg_save_area = Int64.(sub !current_function_stacksize 192L) in
- assert (r <> RAX);
- emit (Pmovl_ri (RAX, coqint_of_camlint gp_offset));
- emit (Pmovl_mr (linear_addr r _0z, RAX));
- emit (Pmovl_ri (RAX, coqint_of_camlint fp_offset));
- emit (Pmovl_mr (linear_addr r _4z, RAX));
- emit (Pleaq (RAX, linear_addr RSP (Z.of_uint64 overflow_arg_area)));
- emit (Pmovq_mr (linear_addr r _8z, RAX));
- emit (Pleaq (RAX, linear_addr RSP (Z.of_uint64 reg_save_area)));
- emit (Pmovq_mr (linear_addr r _16z, RAX))
-
-(* FMA operations *)
-
-(* vfmadd<i><j><k> r1, r2, r3 performs r1 := ri * rj + rk
- hence
- vfmadd132 r1, r2, r3 performs r1 := r1 * r3 + r2
- vfmadd213 r1, r2, r3 performs r1 := r2 * r1 + r3
- vfmadd231 r1, r2, r3 performs r1 := r2 * r3 + r1
-*)
-
-let expand_fma args res i132 i213 i231 =
- match args, res with
- | [BA(FR a1); BA(FR a2); BA(FR a3)], BR(FR res) ->
- if res = a1 then emit (i132 a1 a3 a2) (* a1 * a2 + a3 *)
- else if res = a2 then emit (i213 a2 a1 a3) (* a1 * a2 + a3 *)
- else if res = a3 then emit (i231 a3 a1 a2) (* a1 * a2 + a3 *)
- else begin
- emit (Pmovsd_ff(res, a3));
- emit (i231 res a1 a2) (* a1 * a2 + res *)
- end
- | _ ->
- invalid_arg ("ill-formed fma builtin")
-
-(* Handling of compiler-inlined builtins *)
-
-let expand_builtin_inline name args res =
- match name, args, res with
- (* Integer arithmetic *)
- | ("__builtin_bswap"| "__builtin_bswap32"), [BA(IR a1)], BR(IR res) ->
- if a1 <> res then
- emit (Pmov_rr (res,a1));
- emit (Pbswap32 res)
- | "__builtin_bswap64", [BA(IR a1)], BR(IR res) ->
- if a1 <> res then
- emit (Pmov_rr (res,a1));
- emit (Pbswap64 res)
- | "__builtin_bswap64", [BA_splitlong(BA(IR ah), BA(IR al))],
- BR_splitlong(BR(IR rh), BR(IR rl)) ->
- assert (ah = RAX && al = RDX && rh = RDX && rl = RAX);
- emit (Pbswap32 RAX);
- emit (Pbswap32 RDX)
- | "__builtin_bswap16", [BA(IR a1)], BR(IR res) ->
- if a1 <> res then
- emit (Pmov_rr (res,a1));
- emit (Pbswap16 res)
- | ("__builtin_clz"|"__builtin_clzl"), [BA(IR a1)], BR(IR res) ->
- emit (Pbsrl (res,a1));
- emit (Pxorl_ri(res,coqint_of_camlint 31l))
- | "__builtin_clzll", [BA(IR a1)], BR(IR res) ->
- emit (Pbsrq (res,a1));
- emit (Pxorl_ri(res,coqint_of_camlint 63l))
- | "__builtin_clzll", [BA_splitlong(BA (IR ah), BA (IR al))], BR(IR res) ->
- let lbl1 = new_label() in
- let lbl2 = new_label() in
- emit (Ptestl_rr(ah, ah));
- emit (Pjcc(Cond_e, lbl1));
- emit (Pbsrl(res, ah));
- emit (Pxorl_ri(res, coqint_of_camlint 31l));
- emit (Pjmp_l lbl2);
- emit (Plabel lbl1);
- emit (Pbsrl(res, al));
- emit (Pxorl_ri(res, coqint_of_camlint 63l));
- emit (Plabel lbl2)
- | ("__builtin_ctz" | "__builtin_ctzl"), [BA(IR a1)], BR(IR res) ->
- emit (Pbsfl (res,a1))
- | "__builtin_ctzll", [BA(IR a1)], BR(IR res) ->
- emit (Pbsfq (res,a1))
- | "__builtin_ctzll", [BA_splitlong(BA (IR ah), BA (IR al))], BR(IR res) ->
- let lbl1 = new_label() in
- let lbl2 = new_label() in
- emit (Ptestl_rr(al, al));
- emit (Pjcc(Cond_e, lbl1));
- emit (Pbsfl(res, al));
- emit (Pjmp_l lbl2);
- emit (Plabel lbl1);
- emit (Pbsfl(res, ah));
- emit (Paddl_ri(res, coqint_of_camlint 32l));
- emit (Plabel lbl2)
- (* Float arithmetic *)
- | "__builtin_fabs", [BA(FR a1)], BR(FR res) ->
- if a1 <> res then
- emit (Pmovsd_ff (res,a1));
- emit (Pabsd res) (* This ensures that need_masks is set to true *)
- | "__builtin_fsqrt", [BA(FR a1)], BR(FR res) ->
- emit (Psqrtsd (res,a1))
- | "__builtin_fmax", [BA(FR a1); BA(FR a2)], BR(FR res) ->
- if res = a1 then
- emit (Pmaxsd (res,a2))
- else if res = a2 then
- emit (Pmaxsd (res,a1))
- else begin
- emit (Pmovsd_ff (res,a1));
- emit (Pmaxsd (res,a2))
- end
- | "__builtin_fmin", [BA(FR a1); BA(FR a2)], BR(FR res) ->
- if res = a1 then
- emit (Pminsd (res,a2))
- else if res = a2 then
- emit (Pminsd (res,a1))
- else begin
- emit (Pmovsd_ff (res,a1));
- emit (Pminsd (res,a2))
- end
- | "__builtin_fmadd", _, _ ->
- expand_fma args res
- (fun r1 r2 r3 -> Pfmadd132(r1, r2, r3))
- (fun r1 r2 r3 -> Pfmadd213(r1, r2, r3))
- (fun r1 r2 r3 -> Pfmadd231(r1, r2, r3))
- | "__builtin_fmsub", _, _ ->
- expand_fma args res
- (fun r1 r2 r3 -> Pfmsub132(r1, r2, r3))
- (fun r1 r2 r3 -> Pfmsub213(r1, r2, r3))
- (fun r1 r2 r3 -> Pfmsub231(r1, r2, r3))
- | "__builtin_fnmadd", _, _ ->
- expand_fma args res
- (fun r1 r2 r3 -> Pfnmadd132(r1, r2, r3))
- (fun r1 r2 r3 -> Pfnmadd213(r1, r2, r3))
- (fun r1 r2 r3 -> Pfnmadd231(r1, r2, r3))
- | "__builtin_fnmsub", _, _ ->
- expand_fma args res
- (fun r1 r2 r3 -> Pfnmsub132(r1, r2, r3))
- (fun r1 r2 r3 -> Pfnmsub213(r1, r2, r3))
- (fun r1 r2 r3 -> Pfnmsub231(r1, r2, r3))
- (* 64-bit integer arithmetic *)
- | "__builtin_negl", [BA_splitlong(BA(IR ah), BA(IR al))],
- BR_splitlong(BR(IR rh), BR(IR rl)) ->
- assert (ah = RDX && al = RAX && rh = RDX && rl = RAX);
- emit (Pnegl RAX);
- emit (Padcl_ri (RDX,_0));
- emit (Pnegl RDX)
- | "__builtin_addl", [BA_splitlong(BA(IR ah), BA(IR al));
- BA_splitlong(BA(IR bh), BA(IR bl))],
- BR_splitlong(BR(IR rh), BR(IR rl)) ->
- assert (ah = RDX && al = RAX && bh = RCX && bl = RBX && rh = RDX && rl = RAX);
- emit (Paddl_rr (RAX,RBX));
- emit (Padcl_rr (RDX,RCX))
- | "__builtin_subl", [BA_splitlong(BA(IR ah), BA(IR al));
- BA_splitlong(BA(IR bh), BA(IR bl))],
- BR_splitlong(BR(IR rh), BR(IR rl)) ->
- assert (ah = RDX && al = RAX && bh = RCX && bl = RBX && rh = RDX && rl = RAX);
- emit (Psubl_rr (RAX,RBX));
- emit (Psbbl_rr (RDX,RCX))
- | "__builtin_mull", [BA(IR a); BA(IR b)],
- BR_splitlong(BR(IR rh), BR(IR rl)) ->
- assert (a = RAX && b = RDX && rh = RDX && rl = RAX);
- emit (Pmull_r RDX)
- (* Memory accesses *)
- | "__builtin_read16_reversed", [BA(IR a1)], BR(IR res) ->
- emit (Pmovzw_rm (res, linear_addr a1 _0));
- emit (Pbswap16 res)
- | "__builtin_read32_reversed", [BA(IR a1)], BR(IR res) ->
- emit (Pmovl_rm (res, linear_addr a1 _0));
- emit (Pbswap32 res)
- | "__builtin_write16_reversed", [BA(IR a1); BA(IR a2)], _ ->
- let tmp = if a1 = RCX then RDX else RCX in
- if a2 <> tmp then
- emit (Pmov_rr (tmp,a2));
- emit (Pbswap16 tmp);
- emit (Pmovw_mr (linear_addr a1 _0z, tmp))
- | "__builtin_write32_reversed", [BA(IR a1); BA(IR a2)], _ ->
- let tmp = if a1 = RCX then RDX else RCX in
- if a2 <> tmp then
- emit (Pmov_rr (tmp,a2));
- emit (Pbswap32 tmp);
- emit (Pmovl_mr (linear_addr a1 _0z, tmp))
- (* Vararg stuff *)
- | "__builtin_va_start", [BA(IR a)], _ ->
- assert (a = RDX);
- if Archi.ptr64
- then expand_builtin_va_start_64 a
- else expand_builtin_va_start_32 a
- (* Synchronization *)
- | "__builtin_membar", [], _ ->
- ()
- (* no operation *)
- | "__builtin_nop", [], _ ->
- emit (Pmov_rr (RAX,RAX))
- (* Catch-all *)
- | _ ->
- raise (Error ("unrecognized builtin " ^ name))
-
-(* Calls to variadic functions for x86-64: register AL must contain
- the number of XMM registers used for parameter passing. To be on
- the safe side. do the same if the called function is
- unprototyped. *)
-
-let set_al sg =
- if Archi.ptr64 && (sg.sig_cc.cc_vararg || sg.sig_cc.cc_unproto) then begin
- let (ir, fr, ofs) = next_arg_locations 0 0 0 sg.sig_args in
- emit (Pmovl_ri (RAX, coqint_of_camlint (Int32.of_int fr)))
- end
-
-(* Expansion of instructions *)
-
-let expand_instruction instr =
- match instr with
- | Pallocframe (sz, ofs_ra, ofs_link) ->
- if Archi.ptr64 then begin
- let (sz, save_regs) = sp_adjustment_64 sz in
- (* Allocate frame *)
- let sz' = Z.of_uint sz in
- emit (Psubq_ri (RSP, sz'));
- emit (Pcfi_adjust sz');
- if save_regs >= 0 then begin
- (* Save the registers *)
- emit (Pleaq (R10, linear_addr RSP (Z.of_uint save_regs)));
- emit (Pcall_s (intern_string "__compcert_va_saveregs",
- {sig_args = []; sig_res = None; sig_cc = cc_default}))
- end;
- (* Stack chaining *)
- let fullsz = sz + 8 in
- let addr1 = linear_addr RSP (Z.of_uint fullsz) in
- let addr2 = linear_addr RSP ofs_link in
- emit (Pleaq (RAX, addr1));
- emit (Pmovq_mr (addr2, RAX));
- current_function_stacksize := Int64.of_int fullsz
- end else begin
- let sz = sp_adjustment_32 sz in
- (* Allocate frame *)
- let sz' = Z.of_uint sz in
- emit (Psubl_ri (RSP, sz'));
- emit (Pcfi_adjust sz');
- (* Stack chaining *)
- let addr1 = linear_addr RSP (Z.of_uint (sz + 4)) in
- let addr2 = linear_addr RSP ofs_link in
- emit (Pleal (RAX,addr1));
- emit (Pmovl_mr (addr2,RAX));
- PrintAsmaux.current_function_stacksize := Int32.of_int sz
- end
- | Pfreeframe(sz, ofs_ra, ofs_link) ->
- if Archi.ptr64 then begin
- let (sz, _) = sp_adjustment_64 sz in
- emit (Paddq_ri (RSP, Z.of_uint sz))
- end else begin
- let sz = sp_adjustment_32 sz in
- emit (Paddl_ri (RSP, Z.of_uint sz))
- end
- | Pjmp_s(_, sg) | Pjmp_r(_, sg) | Pcall_s(_, sg) | Pcall_r(_, sg) ->
- set_al sg;
- emit instr
- | Pbuiltin (ef,args, res) ->
- begin
- match ef with
- | EF_builtin(name, sg) ->
- expand_builtin_inline (camlstring_of_coqstring name) args res
- | EF_vload chunk ->
- expand_builtin_vload chunk args res
- | EF_vstore chunk ->
- expand_builtin_vstore chunk args
- | EF_memcpy(sz, al) ->
- expand_builtin_memcpy (Z.to_int sz) (Z.to_int al) args
- | EF_annot_val(txt, targ) ->
- expand_annot_val txt targ args res
- | EF_annot _ | EF_debug _ | EF_inline_asm _ ->
- emit instr
- | _ ->
- assert false
- end
- | _ -> emit instr
-
-let int_reg_to_dwarf_32 = function
- | RAX -> 0
- | RBX -> 3
- | RCX -> 1
- | RDX -> 2
- | RSI -> 6
- | RDI -> 7
- | RBP -> 5
- | RSP -> 4
- | _ -> assert false
-
-let int_reg_to_dwarf_64 = function
- | RAX -> 0
- | RDX -> 1
- | RCX -> 2
- | RBX -> 3
- | RSI -> 4
- | RDI -> 5
- | RBP -> 6
- | RSP -> 7
- | R8 -> 8
- | R9 -> 9
- | R10 -> 10
- | R11 -> 11
- | R12 -> 12
- | R13 -> 13
- | R14 -> 14
- | R15 -> 15
-
-let int_reg_to_dwarf =
- if Archi.ptr64 then int_reg_to_dwarf_64 else int_reg_to_dwarf_32
-
-let float_reg_to_dwarf_32 = function
- | XMM0 -> 21
- | XMM1 -> 22
- | XMM2 -> 23
- | XMM3 -> 24
- | XMM4 -> 25
- | XMM5 -> 26
- | XMM6 -> 27
- | XMM7 -> 28
- | _ -> assert false
-
-let float_reg_to_dwarf_64 = function
- | XMM0 -> 17
- | XMM1 -> 18
- | XMM2 -> 19
- | XMM3 -> 20
- | XMM4 -> 21
- | XMM5 -> 22
- | XMM6 -> 23
- | XMM7 -> 24
- | XMM8 -> 25
- | XMM9 -> 26
- | XMM10 -> 27
- | XMM11 -> 28
- | XMM12 -> 29
- | XMM13 -> 30
- | XMM14 -> 31
- | XMM15 -> 32
-
-let float_reg_to_dwarf =
- if Archi.ptr64 then float_reg_to_dwarf_64 else float_reg_to_dwarf_32
-
-let preg_to_dwarf = function
- | IR r -> int_reg_to_dwarf r
- | FR r -> float_reg_to_dwarf r
- | _ -> assert false
-
-
-let expand_function id fn =
- try
- set_current_function fn;
- if !Clflags.option_g then
- expand_debug id 4 preg_to_dwarf expand_instruction fn.fn_code
- else
- List.iter expand_instruction fn.fn_code;
- Errors.OK (get_current_function ())
- with Error s ->
- Errors.Error (Errors.msg (coqstring_of_camlstring s))
-
-let expand_fundef id = function
- | Internal f ->
- begin match expand_function id f with
- | Errors.OK tf -> Errors.OK (Internal tf)
- | Errors.Error msg -> Errors.Error msg
- end
- | External ef ->
- Errors.OK (External ef)
-
-let expand_program (p: Asm.program) : Asm.program Errors.res =
- AST.transform_partial_program2 expand_fundef (fun id v -> Errors.OK v) p
diff --git a/ia32/Asmgen.v b/ia32/Asmgen.v
deleted file mode 100644
index bb26d507..00000000
--- a/ia32/Asmgen.v
+++ /dev/null
@@ -1,754 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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 IA32 assembly language *)
-
-Require Import Coqlib Errors.
-Require Import AST Integers Floats Memdata.
-Require Import Op Locations Mach Asm.
-
-Open Local Scope string_scope.
-Open Local Scope error_monad_scope.
-
-(** The code generation functions take advantage of several characteristics of the [Mach] code generated by earlier passes of the compiler:
-- Argument and result registers are of the correct type.
-- For two-address instructions, the result and the first argument
- are in the same register. (True by construction in [RTLgen], and preserved by [Reload].)
-- The top of the floating-point stack ([ST0], a.k.a. [FP0]) can only
- appear in [mov] instructions, but never in arithmetic instructions.
-
-All these properties are true by construction, but it is painful to track them statically. Instead, we recheck them during code generation and fail if they do not hold.
-*)
-
-(** 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.
-
-(** Smart constructors for some operations. *)
-
-Definition mk_mov (rd rs: preg) (k: code) : res code :=
- match rd, rs with
- | IR rd, IR rs => OK (Pmov_rr rd rs :: k)
- | FR rd, FR rs => OK (Pmovsd_ff rd rs :: k)
- | _, _ => Error(msg "Asmgen.mk_mov")
- end.
-
-Definition mk_shrximm (n: int) (k: code) : res code :=
- let p := Int.sub (Int.shl Int.one n) Int.one in
- OK (Ptestl_rr RAX RAX ::
- Pleal RCX (Addrmode (Some RAX) None (inl _ (Int.unsigned p))) ::
- Pcmov Cond_l RAX RCX ::
- Psarl_ri RAX n :: k).
-
-Definition mk_shrxlimm (n: int) (k: code) : res code :=
- OK (if Int.eq n Int.zero then Pmov_rr RAX RAX :: k else
- Pcqto ::
- Pshrq_ri RDX (Int.sub (Int.repr 64) n) ::
- Pleaq RAX (Addrmode (Some RAX) (Some(RDX, 1)) (inl _ 0)) ::
- Psarq_ri RAX n :: k).
-
-Definition low_ireg (r: ireg) : bool :=
- match r with RAX | RBX | RCX | RDX => true | _ => false end.
-
-Definition mk_intconv (mk: ireg -> ireg -> instruction) (rd rs: ireg) (k: code) :=
- if Archi.ptr64 || low_ireg rs then
- OK (mk rd rs :: k)
- else
- OK (Pmov_rr RAX rs :: mk rd RAX :: k).
-
-Definition addressing_mentions (addr: addrmode) (r: ireg) : bool :=
- match addr with Addrmode base displ const =>
- match base with Some r' => ireg_eq r r' | None => false end
- || match displ with Some(r', sc) => ireg_eq r r' | None => false end
- end.
-
-Definition mk_storebyte (addr: addrmode) (rs: ireg) (k: code) :=
- if Archi.ptr64 || low_ireg rs then
- OK (Pmovb_mr addr rs :: k)
- else if addressing_mentions addr RAX then
- OK (Pleal RCX addr :: Pmov_rr RAX rs ::
- Pmovb_mr (Addrmode (Some RCX) None (inl _ 0)) RAX :: k)
- else
- OK (Pmov_rr RAX rs :: Pmovb_mr addr RAX :: k).
-
-(** Accessing slots in the stack frame. *)
-
-Definition loadind (base: ireg) (ofs: ptrofs) (ty: typ) (dst: mreg) (k: code) :=
- let a := Addrmode (Some base) None (inl _ (Ptrofs.unsigned ofs)) in
- match ty, preg_of dst with
- | Tint, IR r => OK (Pmovl_rm r a :: k)
- | Tlong, IR r => OK (Pmovq_rm r a :: k)
- | Tsingle, FR r => OK (Pmovss_fm r a :: k)
- | Tsingle, ST0 => OK (Pflds_m a :: k)
- | Tfloat, FR r => OK (Pmovsd_fm r a :: k)
- | Tfloat, ST0 => OK (Pfldl_m a :: k)
- | Tany32, IR r => if Archi.ptr64 then Error (msg "Asmgen.loadind1") else OK (Pmov_rm_a r a :: k)
- | Tany64, IR r => if Archi.ptr64 then OK (Pmov_rm_a r a :: k) else Error (msg "Asmgen.loadind2")
- | Tany64, FR r => OK (Pmovsd_fm_a r a :: k)
- | _, _ => Error (msg "Asmgen.loadind")
- end.
-
-Definition storeind (src: mreg) (base: ireg) (ofs: ptrofs) (ty: typ) (k: code) :=
- let a := Addrmode (Some base) None (inl _ (Ptrofs.unsigned ofs)) in
- match ty, preg_of src with
- | Tint, IR r => OK (Pmovl_mr a r :: k)
- | Tlong, IR r => OK (Pmovq_mr a r :: k)
- | Tsingle, FR r => OK (Pmovss_mf a r :: k)
- | Tsingle, ST0 => OK (Pfstps_m a :: k)
- | Tfloat, FR r => OK (Pmovsd_mf a r :: k)
- | Tfloat, ST0 => OK (Pfstpl_m a :: k)
- | Tany32, IR r => if Archi.ptr64 then Error (msg "Asmgen.storeind1") else OK (Pmov_mr_a a r :: k)
- | Tany64, IR r => if Archi.ptr64 then OK (Pmov_mr_a a r :: k) else Error (msg "Asmgen.storeind2")
- | Tany64, FR r => OK (Pmovsd_mf_a a r :: k)
- | _, _ => Error (msg "Asmgen.storeind")
- end.
-
-(** Translation of addressing modes *)
-
-Definition transl_addressing (a: addressing) (args: list mreg): res addrmode :=
- match a, args with
- | Aindexed n, a1 :: nil =>
- do r1 <- ireg_of a1; OK(Addrmode (Some r1) None (inl _ n))
- | Aindexed2 n, a1 :: a2 :: nil =>
- do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK(Addrmode (Some r1) (Some(r2, 1)) (inl _ n))
- | Ascaled sc n, a1 :: nil =>
- do r1 <- ireg_of a1; OK(Addrmode None (Some(r1, sc)) (inl _ n))
- | Aindexed2scaled sc n, a1 :: a2 :: nil =>
- do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK(Addrmode (Some r1) (Some(r2, sc)) (inl _ n))
- | Aglobal id ofs, nil =>
- OK(Addrmode None None (inr _ (id, ofs)))
- | Abased id ofs, a1 :: nil =>
- do r1 <- ireg_of a1; OK(Addrmode (Some r1) None (inr _ (id, ofs)))
- | Abasedscaled sc id ofs, a1 :: nil =>
- do r1 <- ireg_of a1; OK(Addrmode None (Some(r1, sc)) (inr _ (id, ofs)))
- | Ainstack n, nil =>
- OK(Addrmode (Some RSP) None (inl _ (Ptrofs.signed n)))
- | _, _ =>
- Error(msg "Asmgen.transl_addressing")
- end.
-
-Definition normalize_addrmode_32 (a: addrmode) :=
- match a with
- | Addrmode base ofs (inl n) =>
- Addrmode base ofs (inl _ (Int.signed (Int.repr n)))
- | Addrmode base ofs (inr _) =>
- a
- end.
-
-Definition normalize_addrmode_64 (a: addrmode) :=
- match a with
- | Addrmode base ofs (inl n) =>
- let n' := Int.signed (Int.repr n) in
- if zeq n' n
- then (a, None)
- else (Addrmode base ofs (inl _ 0), Some (Int64.repr n))
- | Addrmode base ofs (inr _) =>
- (a, None)
- end.
-
-(** Floating-point comparison. We swap the operands in some cases
- to simplify the handling of the unordered case. *)
-
-Definition floatcomp (cmp: comparison) (r1 r2: freg) : instruction :=
- match cmp with
- | Clt | Cle => Pcomisd_ff r2 r1
- | Ceq | Cne | Cgt | Cge => Pcomisd_ff r1 r2
- end.
-
-Definition floatcomp32 (cmp: comparison) (r1 r2: freg) : instruction :=
- match cmp with
- | Clt | Cle => Pcomiss_ff r2 r1
- | Ceq | Cne | Cgt | Cge => Pcomiss_ff r1 r2
- end.
-
-(** Translation of a condition. Prepends to [k] the instructions
- that evaluate the condition and leave its boolean result in bits
- of the condition register. *)
-
-Definition transl_cond
- (cond: condition) (args: list mreg) (k: code) : res code :=
- match cond, args with
- | Ccomp c, a1 :: a2 :: nil =>
- do r1 <- ireg_of a1; do r2 <- ireg_of a2; OK (Pcmpl_rr r1 r2 :: k)
- | Ccompu c, a1 :: a2 :: nil =>
- do r1 <- ireg_of a1; do r2 <- ireg_of a2; OK (Pcmpl_rr r1 r2 :: k)
- | Ccompimm c n, a1 :: nil =>
- do r1 <- ireg_of a1;
- OK (if Int.eq_dec n Int.zero then Ptestl_rr r1 r1 :: k else Pcmpl_ri r1 n :: k)
- | Ccompuimm c n, a1 :: nil =>
- do r1 <- ireg_of a1; OK (Pcmpl_ri r1 n :: k)
- | Ccompl c, a1 :: a2 :: nil =>
- do r1 <- ireg_of a1; do r2 <- ireg_of a2; OK (Pcmpq_rr r1 r2 :: k)
- | Ccomplu c, a1 :: a2 :: nil =>
- do r1 <- ireg_of a1; do r2 <- ireg_of a2; OK (Pcmpq_rr r1 r2 :: k)
- | Ccomplimm c n, a1 :: nil =>
- do r1 <- ireg_of a1;
- OK (if Int64.eq_dec n Int64.zero then Ptestq_rr r1 r1 :: k else Pcmpq_ri r1 n :: k)
- | Ccompluimm c n, a1 :: nil =>
- do r1 <- ireg_of a1; OK (Pcmpq_ri r1 n :: k)
- | Ccompf cmp, a1 :: a2 :: nil =>
- do r1 <- freg_of a1; do r2 <- freg_of a2; OK (floatcomp cmp r1 r2 :: k)
- | Cnotcompf cmp, a1 :: a2 :: nil =>
- do r1 <- freg_of a1; do r2 <- freg_of a2; OK (floatcomp cmp r1 r2 :: k)
- | Ccompfs cmp, a1 :: a2 :: nil =>
- do r1 <- freg_of a1; do r2 <- freg_of a2; OK (floatcomp32 cmp r1 r2 :: k)
- | Cnotcompfs cmp, a1 :: a2 :: nil =>
- do r1 <- freg_of a1; do r2 <- freg_of a2; OK (floatcomp32 cmp r1 r2 :: k)
- | Cmaskzero n, a1 :: nil =>
- do r1 <- ireg_of a1; OK (Ptestl_ri r1 n :: k)
- | Cmasknotzero n, a1 :: nil =>
- do r1 <- ireg_of a1; OK (Ptestl_ri r1 n :: k)
- | _, _ =>
- Error(msg "Asmgen.transl_cond")
- end.
-
-(** What processor condition to test for a given Mach condition. *)
-
-Definition testcond_for_signed_comparison (cmp: comparison) :=
- match cmp with
- | Ceq => Cond_e
- | Cne => Cond_ne
- | Clt => Cond_l
- | Cle => Cond_le
- | Cgt => Cond_g
- | Cge => Cond_ge
- end.
-
-Definition testcond_for_unsigned_comparison (cmp: comparison) :=
- match cmp with
- | Ceq => Cond_e
- | Cne => Cond_ne
- | Clt => Cond_b
- | Cle => Cond_be
- | Cgt => Cond_a
- | Cge => Cond_ae
- end.
-
-Inductive extcond: Type :=
- | Cond_base (c: testcond)
- | Cond_or (c1 c2: testcond)
- | Cond_and (c1 c2: testcond).
-
-Definition testcond_for_condition (cond: condition) : extcond :=
- match cond with
- | Ccomp c => Cond_base(testcond_for_signed_comparison c)
- | Ccompu c => Cond_base(testcond_for_unsigned_comparison c)
- | Ccompimm c n => Cond_base(testcond_for_signed_comparison c)
- | Ccompuimm c n => Cond_base(testcond_for_unsigned_comparison c)
- | Ccompl c => Cond_base(testcond_for_signed_comparison c)
- | Ccomplu c => Cond_base(testcond_for_unsigned_comparison c)
- | Ccomplimm c n => Cond_base(testcond_for_signed_comparison c)
- | Ccompluimm c n => Cond_base(testcond_for_unsigned_comparison c)
- | Ccompf c | Ccompfs c =>
- match c with
- | Ceq => Cond_and Cond_np Cond_e
- | Cne => Cond_or Cond_p Cond_ne
- | Clt => Cond_base Cond_a
- | Cle => Cond_base Cond_ae
- | Cgt => Cond_base Cond_a
- | Cge => Cond_base Cond_ae
- end
- | Cnotcompf c | Cnotcompfs c =>
- match c with
- | Ceq => Cond_or Cond_p Cond_ne
- | Cne => Cond_and Cond_np Cond_e
- | Clt => Cond_base Cond_be
- | Cle => Cond_base Cond_b
- | Cgt => Cond_base Cond_be
- | Cge => Cond_base Cond_b
- end
- | Cmaskzero n => Cond_base Cond_e
- | Cmasknotzero n => Cond_base Cond_ne
- end.
-
-(** Acting upon extended conditions. *)
-
-Definition mk_setcc_base (cond: extcond) (rd: ireg) (k: code) :=
- match cond with
- | Cond_base c =>
- Psetcc c rd :: k
- | Cond_and c1 c2 =>
- if ireg_eq rd RAX
- then Psetcc c1 RAX :: Psetcc c2 RCX :: Pandl_rr RAX RCX :: k
- else Psetcc c1 RAX :: Psetcc c2 rd :: Pandl_rr rd RAX :: k
- | Cond_or c1 c2 =>
- if ireg_eq rd RAX
- then Psetcc c1 RAX :: Psetcc c2 RCX :: Porl_rr RAX RCX :: k
- else Psetcc c1 RAX :: Psetcc c2 rd :: Porl_rr rd RAX :: k
- end.
-
-Definition mk_setcc (cond: extcond) (rd: ireg) (k: code) :=
- if Archi.ptr64 || low_ireg rd
- then mk_setcc_base cond rd k
- else mk_setcc_base cond RAX (Pmov_rr rd RAX :: k).
-
-Definition mk_jcc (cond: extcond) (lbl: label) (k: code) :=
- match cond with
- | Cond_base c => Pjcc c lbl :: k
- | Cond_and c1 c2 => Pjcc2 c1 c2 lbl :: k
- | Cond_or c1 c2 => Pjcc c1 lbl :: Pjcc c2 lbl :: k
- end.
-
-(** Translation of the arithmetic operation [r <- op(args)].
- The corresponding instructions are prepended to [k]. *)
-
-Definition transl_op
- (op: operation) (args: list mreg) (res: mreg) (k: code) : Errors.res code :=
- match op, args with
- | Omove, a1 :: nil =>
- mk_mov (preg_of res) (preg_of a1) k
- | Ointconst n, nil =>
- do r <- ireg_of res;
- OK ((if Int.eq_dec n Int.zero then Pxorl_r r else Pmovl_ri r n) :: k)
- | Olongconst n, nil =>
- do r <- ireg_of res;
- OK ((if Int64.eq_dec n Int64.zero then Pxorq_r r else Pmovq_ri r n) :: k)
- | Ofloatconst f, nil =>
- do r <- freg_of res;
- OK ((if Float.eq_dec f Float.zero then Pxorpd_f r else Pmovsd_fi r f) :: k)
- | Osingleconst f, nil =>
- do r <- freg_of res;
- OK ((if Float32.eq_dec f Float32.zero then Pxorps_f r else Pmovss_fi r f) :: k)
- | Oindirectsymbol id, nil =>
- do r <- ireg_of res;
- OK (Pmov_rs r id :: k)
- | Ocast8signed, a1 :: nil =>
- do r1 <- ireg_of a1; do r <- ireg_of res; mk_intconv Pmovsb_rr r r1 k
- | Ocast8unsigned, a1 :: nil =>
- do r1 <- ireg_of a1; do r <- ireg_of res; mk_intconv Pmovzb_rr r r1 k
- | Ocast16signed, a1 :: nil =>
- do r1 <- ireg_of a1; do r <- ireg_of res; OK (Pmovsw_rr r r1 :: k)
- | Ocast16unsigned, a1 :: nil =>
- do r1 <- ireg_of a1; do r <- ireg_of res; OK (Pmovzw_rr r r1 :: k)
- | Oneg, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pnegl r :: k)
- | Osub, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Psubl_rr r r2 :: k)
- | Omul, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Pimull_rr r r2 :: k)
- | Omulimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pimull_ri r n :: k)
- | Omulhs, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq res DX);
- do r2 <- ireg_of a2; OK (Pimull_r r2 :: k)
- | Omulhu, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq res DX);
- do r2 <- ireg_of a2; OK (Pmull_r r2 :: k)
- | Odiv, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq a2 CX);
- assertion (mreg_eq res AX);
- OK(Pcltd :: Pidivl RCX :: k)
- | Odivu, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq a2 CX);
- assertion (mreg_eq res AX);
- OK(Pxorl_r RDX :: Pdivl RCX :: k)
- | Omod, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq a2 CX);
- assertion (mreg_eq res DX);
- OK(Pcltd :: Pidivl RCX :: k)
- | Omodu, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq a2 CX);
- assertion (mreg_eq res DX);
- OK(Pxorl_r RDX :: Pdivl RCX :: k)
- | Oand, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Pandl_rr r r2 :: k)
- | Oandimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pandl_ri r n :: k)
- | Oor, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Porl_rr r r2 :: k)
- | Oorimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Porl_ri r n :: k)
- | Oxor, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Pxorl_rr r r2 :: k)
- | Oxorimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pxorl_ri r n :: k)
- | Onot, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pnotl r :: k)
- | Oshl, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- assertion (mreg_eq a2 CX);
- do r <- ireg_of res; OK (Psall_rcl r :: k)
- | Oshlimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Psall_ri r n :: k)
- | Oshr, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- assertion (mreg_eq a2 CX);
- do r <- ireg_of res; OK (Psarl_rcl r :: k)
- | Oshrimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Psarl_ri r n :: k)
- | Oshru, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- assertion (mreg_eq a2 CX);
- do r <- ireg_of res; OK (Pshrl_rcl r :: k)
- | Oshruimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pshrl_ri r n :: k)
- | Oshrximm n, a1 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq res AX);
- mk_shrximm n k
- | Ororimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Prorl_ri r n :: k)
- | Oshldimm n, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Pshld_ri r r2 n :: k)
- | Olea addr, _ =>
- do am <- transl_addressing addr args; do r <- ireg_of res;
- OK (Pleal r (normalize_addrmode_32 am) :: k)
-(* 64-bit integer operations *)
- | Olowlong, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pmovls_rr r :: k)
- | Ocast32signed, a1 :: nil =>
- do r1 <- ireg_of a1; do r <- ireg_of res; OK (Pmovsl_rr r r1 :: k)
- | Ocast32unsigned, a1 :: nil =>
- do r1 <- ireg_of a1; do r <- ireg_of res; OK (Pmovzl_rr r r1 :: k)
- | Onegl, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pnegq r :: k)
- | Oaddlimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Paddq_ri r n :: k)
- | Osubl, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Psubq_rr r r2 :: k)
- | Omull, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Pimulq_rr r r2 :: k)
- | Omullimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pimulq_ri r n :: k)
- | Omullhs, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq res DX);
- do r2 <- ireg_of a2; OK (Pimulq_r r2 :: k)
- | Omullhu, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq res DX);
- do r2 <- ireg_of a2; OK (Pmulq_r r2 :: k)
- | Odivl, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq a2 CX);
- assertion (mreg_eq res AX);
- OK(Pcqto :: Pidivq RCX :: k)
- | Odivlu, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq a2 CX);
- assertion (mreg_eq res AX);
- OK(Pxorq_r RDX :: Pdivq RCX :: k)
- | Omodl, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq a2 CX);
- assertion (mreg_eq res DX);
- OK(Pcqto :: Pidivq RCX :: k)
- | Omodlu, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq a2 CX);
- assertion (mreg_eq res DX);
- OK(Pxorq_r RDX :: Pdivq RCX :: k)
- | Oandl, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Pandq_rr r r2 :: k)
- | Oandlimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pandq_ri r n :: k)
- | Oorl, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Porq_rr r r2 :: k)
- | Oorlimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Porq_ri r n :: k)
- | Oxorl, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Pxorq_rr r r2 :: k)
- | Oxorlimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pxorq_ri r n :: k)
- | Onotl, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pnotq r :: k)
- | Oshll, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- assertion (mreg_eq a2 CX);
- do r <- ireg_of res; OK (Psalq_rcl r :: k)
- | Oshllimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Psalq_ri r n :: k)
- | Oshrl, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- assertion (mreg_eq a2 CX);
- do r <- ireg_of res; OK (Psarq_rcl r :: k)
- | Oshrlimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Psarq_ri r n :: k)
- | Oshrxlimm n, a1 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq res AX);
- mk_shrxlimm n k
- | Oshrlu, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- assertion (mreg_eq a2 CX);
- do r <- ireg_of res; OK (Pshrq_rcl r :: k)
- | Oshrluimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pshrq_ri r n :: k)
- | Ororlimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Prorq_ri r n :: k)
- | Oleal addr, _ =>
- do am <- transl_addressing addr args; do r <- ireg_of res;
- OK (match normalize_addrmode_64 am with
- | (am', None) => Pleaq r am' :: k
- | (am', Some delta) => Pleaq r am' :: Paddq_ri r delta :: k
- end)
-(**)
- | Onegf, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- freg_of res; OK (Pnegd r :: k)
- | Oabsf, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- freg_of res; OK (Pabsd r :: k)
- | Oaddf, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- freg_of res; do r2 <- freg_of a2; OK (Paddd_ff r r2 :: k)
- | Osubf, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- freg_of res; do r2 <- freg_of a2; OK (Psubd_ff r r2 :: k)
- | Omulf, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- freg_of res; do r2 <- freg_of a2; OK (Pmuld_ff r r2 :: k)
- | Odivf, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- freg_of res; do r2 <- freg_of a2; OK (Pdivd_ff r r2 :: k)
- | Onegfs, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- freg_of res; OK (Pnegs r :: k)
- | Oabsfs, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- freg_of res; OK (Pabss r :: k)
- | Oaddfs, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- freg_of res; do r2 <- freg_of a2; OK (Padds_ff r r2 :: k)
- | Osubfs, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- freg_of res; do r2 <- freg_of a2; OK (Psubs_ff r r2 :: k)
- | Omulfs, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- freg_of res; do r2 <- freg_of a2; OK (Pmuls_ff r r2 :: k)
- | Odivfs, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- freg_of res; do r2 <- freg_of a2; OK (Pdivs_ff r r2 :: k)
- | Osingleoffloat, a1 :: nil =>
- do r <- freg_of res; do r1 <- freg_of a1; OK (Pcvtsd2ss_ff r r1 :: k)
- | Ofloatofsingle, a1 :: nil =>
- do r <- freg_of res; do r1 <- freg_of a1; OK (Pcvtss2sd_ff r r1 :: k)
- | Ointoffloat, a1 :: nil =>
- do r <- ireg_of res; do r1 <- freg_of a1; OK (Pcvttsd2si_rf r r1 :: k)
- | Ofloatofint, a1 :: nil =>
- do r <- freg_of res; do r1 <- ireg_of a1; OK (Pcvtsi2sd_fr r r1 :: k)
- | Ointofsingle, a1 :: nil =>
- do r <- ireg_of res; do r1 <- freg_of a1; OK (Pcvttss2si_rf r r1 :: k)
- | Osingleofint, a1 :: nil =>
- do r <- freg_of res; do r1 <- ireg_of a1; OK (Pcvtsi2ss_fr r r1 :: k)
- | Olongoffloat, a1 :: nil =>
- do r <- ireg_of res; do r1 <- freg_of a1; OK (Pcvttsd2sl_rf r r1 :: k)
- | Ofloatoflong, a1 :: nil =>
- do r <- freg_of res; do r1 <- ireg_of a1; OK (Pcvtsl2sd_fr r r1 :: k)
- | Olongofsingle, a1 :: nil =>
- do r <- ireg_of res; do r1 <- freg_of a1; OK (Pcvttss2sl_rf r r1 :: k)
- | Osingleoflong, a1 :: nil =>
- do r <- freg_of res; do r1 <- ireg_of a1; OK (Pcvtsl2ss_fr r r1 :: k)
- | Ocmp c, args =>
- do r <- ireg_of res;
- transl_cond c args (mk_setcc (testcond_for_condition c) r k)
- | _, _ =>
- Error(msg "Asmgen.transl_op")
- end.
-
-(** Translation of memory loads and stores *)
-
-Definition transl_load (chunk: memory_chunk)
- (addr: addressing) (args: list mreg) (dest: mreg)
- (k: code) : res code :=
- do am <- transl_addressing addr args;
- match chunk with
- | Mint8unsigned =>
- do r <- ireg_of dest; OK(Pmovzb_rm r am :: k)
- | Mint8signed =>
- do r <- ireg_of dest; OK(Pmovsb_rm r am :: k)
- | Mint16unsigned =>
- do r <- ireg_of dest; OK(Pmovzw_rm r am :: k)
- | Mint16signed =>
- do r <- ireg_of dest; OK(Pmovsw_rm r am :: k)
- | Mint32 =>
- do r <- ireg_of dest; OK(Pmovl_rm r am :: k)
- | Mint64 =>
- do r <- ireg_of dest; OK(Pmovq_rm r am :: k)
- | Mfloat32 =>
- do r <- freg_of dest; OK(Pmovss_fm r am :: k)
- | Mfloat64 =>
- do r <- freg_of dest; OK(Pmovsd_fm r am :: k)
- | _ =>
- Error (msg "Asmgen.transl_load")
- end.
-
-Definition transl_store (chunk: memory_chunk)
- (addr: addressing) (args: list mreg) (src: mreg)
- (k: code) : res code :=
- do am <- transl_addressing addr args;
- match chunk with
- | Mint8unsigned | Mint8signed =>
- do r <- ireg_of src; mk_storebyte am r k
- | Mint16unsigned | Mint16signed =>
- do r <- ireg_of src; OK(Pmovw_mr am r :: k)
- | Mint32 =>
- do r <- ireg_of src; OK(Pmovl_mr am r :: k)
- | Mint64 =>
- do r <- ireg_of src; OK(Pmovq_mr am r :: k)
- | Mfloat32 =>
- do r <- freg_of src; OK(Pmovss_mf am r :: k)
- | Mfloat64 =>
- do r <- freg_of src; OK(Pmovsd_mf am r :: k)
- | _ =>
- Error (msg "Asmgen.transl_store")
- end.
-
-(** Translation of a Mach instruction. *)
-
-Definition transl_instr (f: Mach.function) (i: Mach.instruction)
- (ax_is_parent: bool) (k: code) :=
- match i with
- | Mgetstack ofs ty dst =>
- loadind RSP ofs ty dst k
- | Msetstack src ofs ty =>
- storeind src RSP ofs ty k
- | Mgetparam ofs ty dst =>
- if ax_is_parent then
- loadind RAX ofs ty dst k
- else
- (do k1 <- loadind RAX ofs ty dst k;
- loadind RSP f.(fn_link_ofs) Tptr AX k1)
- | Mop op args res =>
- transl_op op args res k
- | Mload chunk addr args dst =>
- transl_load chunk addr args dst k
- | Mstore chunk addr args src =>
- transl_store chunk addr args src k
- | Mcall sig (inl reg) =>
- do r <- ireg_of reg; OK (Pcall_r r sig :: k)
- | Mcall sig (inr symb) =>
- OK (Pcall_s symb sig :: k)
- | Mtailcall sig (inl reg) =>
- do r <- ireg_of reg;
- OK (Pfreeframe f.(fn_stacksize) f.(fn_retaddr_ofs) f.(fn_link_ofs) ::
- Pjmp_r r sig :: k)
- | Mtailcall sig (inr symb) =>
- OK (Pfreeframe f.(fn_stacksize) f.(fn_retaddr_ofs) f.(fn_link_ofs) ::
- Pjmp_s symb sig :: k)
- | Mlabel lbl =>
- OK(Plabel lbl :: k)
- | Mgoto lbl =>
- OK(Pjmp_l lbl :: k)
- | Mcond cond args lbl =>
- transl_cond cond args (mk_jcc (testcond_for_condition cond) lbl k)
- | Mjumptable arg tbl =>
- do r <- ireg_of arg; OK (Pjmptbl r tbl :: k)
- | Mreturn =>
- OK (Pfreeframe f.(fn_stacksize) f.(fn_retaddr_ofs) f.(fn_link_ofs) ::
- Pret :: k)
- | Mbuiltin ef args res =>
- OK (Pbuiltin ef (List.map (map_builtin_arg preg_of) args) (map_builtin_res preg_of res) :: 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 AX)
- | _ => 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) (axp: bool) :=
- match il with
- | nil => OK nil
- | i1 :: il' =>
- do k <- transl_code f il' (it1_is_parent axp i1);
- transl_instr f i1 axp 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)
- (axp: bool) (k: code -> res code) :=
- match il with
- | nil => k nil
- | i1 :: il' =>
- transl_code_rec f il' (it1_is_parent axp i1)
- (fun c1 => do c2 <- transl_instr f i1 axp c1; k c2)
- end.
-
-Definition transl_code' (f: Mach.function) (il: list Mach.instruction) (axp: bool) :=
- transl_code_rec f il axp (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_retaddr_ofs) f.(fn_link_ofs) :: c)).
-
-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_fundef (f: Mach.fundef) : res Asm.fundef :=
- transf_partial_fundef transf_function f.
-
-Definition transf_program (p: Mach.program) : res Asm.program :=
- transform_partial_program transf_fundef p.
-
diff --git a/ia32/Asmgenproof.v b/ia32/Asmgenproof.v
deleted file mode 100644
index e56dc429..00000000
--- a/ia32/Asmgenproof.v
+++ /dev/null
@@ -1,916 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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 x86-64 generation: main proof. *)
-
-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 (fn_code tf) <= Ptrofs.max_unsigned.
-Proof.
- intros. monadInv H. destruct (zlt Ptrofs.max_unsigned (list_length_z (fn_code x))); monadInv 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.
-
- In passing, we also prove a "is tail" property of the generated Asm code.
-*)
-
-Section TRANSL_LABEL.
-
-Remark mk_mov_label:
- forall rd rs k c, mk_mov rd rs k = OK c -> tail_nolabel k c.
-Proof.
- unfold mk_mov; intros.
- destruct rd; try discriminate; destruct rs; TailNoLabel.
-Qed.
-Hint Resolve mk_mov_label: labels.
-
-Remark mk_shrximm_label:
- forall n k c, mk_shrximm n k = OK c -> tail_nolabel k c.
-Proof.
- intros. monadInv H; TailNoLabel.
-Qed.
-Hint Resolve mk_shrximm_label: labels.
-
-Remark mk_shrxlimm_label:
- forall n k c, mk_shrxlimm n k = OK c -> tail_nolabel k c.
-Proof.
- intros. monadInv H. destruct (Int.eq n Int.zero); TailNoLabel.
-Qed.
-Hint Resolve mk_shrxlimm_label: labels.
-
-Remark mk_intconv_label:
- forall f r1 r2 k c, mk_intconv f r1 r2 k = OK c ->
- (forall r r', nolabel (f r r')) ->
- tail_nolabel k c.
-Proof.
- unfold mk_intconv; intros. TailNoLabel.
-Qed.
-Hint Resolve mk_intconv_label: labels.
-
-Remark mk_storebyte_label:
- forall addr r k c, mk_storebyte addr r k = OK c -> tail_nolabel k c.
-Proof.
- unfold mk_storebyte; intros. TailNoLabel.
-Qed.
-Hint Resolve mk_storebyte_label: labels.
-
-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; try discriminate; destruct (preg_of dst); TailNoLabel.
-Qed.
-
-Remark storeind_label:
- forall base ofs ty src k c,
- storeind src base ofs ty k = OK c ->
- tail_nolabel k c.
-Proof.
- unfold storeind; intros. destruct ty; try discriminate; destruct (preg_of src); TailNoLabel.
-Qed.
-
-Remark mk_setcc_base_label:
- forall xc rd k,
- tail_nolabel k (mk_setcc_base xc rd k).
-Proof.
- intros. destruct xc; simpl; destruct (ireg_eq rd RAX); TailNoLabel.
-Qed.
-
-Remark mk_setcc_label:
- forall xc rd k,
- tail_nolabel k (mk_setcc xc rd k).
-Proof.
- intros. unfold mk_setcc. destruct (Archi.ptr64 || low_ireg rd).
- apply mk_setcc_base_label.
- eapply tail_nolabel_trans. apply mk_setcc_base_label. TailNoLabel.
-Qed.
-
-Remark mk_jcc_label:
- forall xc lbl' k,
- tail_nolabel k (mk_jcc xc lbl' k).
-Proof.
- intros. destruct xc; simpl; TailNoLabel.
-Qed.
-
-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 (Int.eq_dec n Int.zero); TailNoLabel.
- destruct (Int64.eq_dec n Int64.zero); TailNoLabel.
- destruct c0; simpl; TailNoLabel.
- destruct c0; simpl; TailNoLabel.
- destruct c0; simpl; TailNoLabel.
- destruct c0; simpl; 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 (Int.eq_dec n Int.zero); TailNoLabel.
- destruct (Int64.eq_dec n Int64.zero); TailNoLabel.
- destruct (Float.eq_dec n Float.zero); TailNoLabel.
- destruct (Float32.eq_dec n Float32.zero); TailNoLabel.
- destruct (normalize_addrmode_64 x) as [am' [delta|]]; TailNoLabel.
- eapply tail_nolabel_trans. eapply transl_cond_label; eauto. eapply mk_setcc_label.
-Qed.
-
-Remark transl_load_label:
- forall chunk addr args dest k c,
- transl_load chunk addr args dest k = OK c ->
- tail_nolabel k c.
-Proof.
- intros. monadInv H. destruct chunk; TailNoLabel.
-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.
- intros. monadInv H. destruct chunk; 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.
-Opaque loadind.
- unfold transl_instr; intros; destruct i; TailNoLabel.
- eapply loadind_label; eauto.
- eapply storeind_label; eauto.
- eapply loadind_label; eauto.
- eapply tail_nolabel_trans; eapply loadind_label; eauto.
- eapply transl_op_label; eauto.
- eapply transl_load_label; eauto.
- eapply transl_store_label; eauto.
- destruct s0; TailNoLabel.
- destruct s0; TailNoLabel.
- eapply tail_nolabel_trans. eapply transl_cond_label; eauto. eapply mk_jcc_label.
-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 (fn_code x))); inv EQ0.
- monadInv EQ. simpl. eapply transl_code_label; eauto. rewrite transl_code'_transl_code in EQ0; eauto.
-Qed.
-
-End TRANSL_LABEL.
-
-(** A valid branch in a piece of Mach code translates to a valid ``go to''
- transition in the generated PPC 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 (fn_code x))); inv EQ0.
- monadInv EQ. rewrite transl_code'_transl_code in EQ0.
- exists x; exists true; split; auto. unfold fn_code. repeat constructor.
-- 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 PPC code pointed by the PC register is the translation of
- the current Mach code sequence.
-- Mach register values and PPC 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)
- (AXP: ep = true -> rs#RAX = parent_sp 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#RAX = parent_sp 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]]].
- 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') ->
- 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]]]]].
- 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.
-Qed.
-
-(** We need to show that, in the simulation diagram, we cannot
- take infinitely many Mach transitions that correspond to zero
- transitions on the PPC side. Actually, all Mach transitions
- correspond to at least one Asm transition, except the
- transition from [Mach.Returnstate] to [Mach.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.
-
-(** This is the simulation diagram. We prove it by case analysis on the Mach transition. *)
-
-Theorem step_simulation:
- forall S1 t S2, Mach.step return_address_offset ge S1 t S2 ->
- forall S1' (MS: match_states S1 S1'),
- (exists S2', plus step tge S1' t S2' /\ match_states S2 S2')
- \/ (measure S2 < measure S1 /\ t = E0 /\ match_states S2 S1')%nat.
-Proof.
- induction 1; intros; inv MS.
-
-- (* Mlabel *)
- left; eapply exec_straight_steps; eauto; intros.
- monadInv TR. econstructor; split. apply exec_straight_one. simpl; eauto. auto.
- split. apply agree_nextinstr; auto. simpl; congruence.
-
-- (* Mgetstack *)
- unfold load_stack in H.
- exploit Mem.loadv_extends; eauto. intros [v' [A B]].
- rewrite (sp_val _ _ _ AG) in A.
- left; eapply exec_straight_steps; eauto. intros. simpl in TR.
- exploit loadind_correct; eauto. intros [rs' [P [Q R]]].
- exists rs'; split. eauto.
- split. eapply agree_set_mreg; eauto. congruence.
- simpl; congruence.
-
-- (* Msetstack *)
- unfold store_stack in H.
- assert (Val.lessdef (rs src) (rs0 (preg_of src))). 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. intros [rs' [P Q]].
- exists rs'; split. eauto.
- split. eapply agree_undef_regs; eauto.
- simpl; intros. rewrite Q; auto with asmgen.
-Local Transparent destroyed_by_setstack.
- destruct ty; simpl; intuition congruence.
-
-- (* 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.
- assert (DIFF: negb (mreg_eq dst AX) = true -> IR RAX <> preg_of dst).
- intros. change (IR RAX) with (preg_of AX). red; intros.
- unfold proj_sumbool in H1. destruct (mreg_eq dst AX); try discriminate.
- elim n. eapply preg_of_injective; eauto.
- destruct ep; simpl in TR.
-(* RAX contains parent *)
- exploit loadind_correct. eexact TR.
- instantiate (2 := rs0). rewrite AXP; eauto.
- intros [rs1 [P [Q R]]].
- exists rs1; split. eauto.
- split. eapply agree_set_mreg. eapply agree_set_mreg; eauto. congruence. auto.
- simpl; intros. rewrite R; auto.
-(* RAX does not contain parent *)
- monadInv TR.
- exploit loadind_correct. eexact EQ0. eauto. intros [rs1 [P [Q R]]]. simpl in Q.
- exploit loadind_correct. eexact EQ. instantiate (2 := rs1). rewrite Q. eauto.
- intros [rs2 [S [T U]]].
- exists rs2; split. eapply exec_straight_trans; eauto.
- split. eapply agree_set_mreg. eapply agree_set_mreg; eauto. congruence. auto.
- simpl; intros. rewrite U; auto.
-
-- (* Mop *)
- assert (eval_operation tge sp op 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]]].
- assert (S: Val.lessdef v (rs2 (preg_of res))) by (eapply Val.lessdef_trans; eauto).
- exists rs2; split. eauto.
- split. eapply agree_set_undef_mreg; eauto.
- simpl; congruence.
-
-- (* Mload *)
- assert (eval_addressing tge sp addr rs##args = Some a).
- rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
- exploit eval_addressing_lessdef. eapply preg_vals; eauto. eexact H1.
- intros [a' [A B]]. rewrite (sp_val _ _ _ AG) in A.
- exploit Mem.loadv_extends; eauto. intros [v' [C D]].
- left; eapply exec_straight_steps; eauto; intros. simpl in TR.
- exploit transl_load_correct; eauto. intros [rs2 [P [Q R]]].
- exists rs2; split. eauto.
- split. eapply agree_set_undef_mreg; eauto. congruence.
- simpl; congruence.
-
-- (* Mstore *)
- assert (eval_addressing tge sp addr 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))). eapply preg_val; eauto.
- exploit Mem.storev_extends; eauto. intros [m2' [C D]].
- left; eapply exec_straight_steps; eauto.
- intros. simpl in TR.
- exploit transl_store_correct; eauto. intros [rs2 [P Q]].
- exists rs2; split. eauto.
- split. eapply agree_undef_regs; eauto.
- simpl; congruence.
-
-- (* Mcall *)
- assert (f0 = f) by congruence. subst f0.
- inv AT.
- assert (NOOV: list_length_z tf.(fn_code) <= Ptrofs.max_unsigned).
- eapply transf_function_no_overflow; eauto.
- destruct ros as [rf|fid]; simpl in H; monadInv H5.
-+ (* Indirect call *)
- assert (rs rf = Vptr f' Ptrofs.zero).
- destruct (rs rf); try discriminate.
- revert H; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence.
- assert (rs0 x0 = Vptr f' Ptrofs.zero).
- exploit ireg_val; eauto. rewrite H5; intros LD; inv LD; auto.
- generalize (code_tail_next_int _ _ _ _ NOOV H6). intro CT1.
- assert (TCA: transl_code_at_pc ge (Vptr fb (Ptrofs.add ofs Ptrofs.one)) fb f c false tf x).
- econstructor; eauto.
- exploit return_address_offset_correct; eauto. intros; subst ra.
- left; econstructor; split.
- apply plus_one. eapply exec_step_internal. 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. Simplifs.
- Simplifs. 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. Simplifs.
- Simplifs. rewrite <- H2. auto.
-
-- (* Mtailcall *)
- 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.
- rewrite (sp_val _ _ _ AG) in *. unfold load_stack in *.
- exploit Mem.loadv_extends. eauto. eexact H1. auto. simpl. intros [parent' [A B]].
- exploit Mem.loadv_extends. eauto. eexact H2. auto. simpl. intros [ra' [C D]].
- exploit lessdef_parent_sp; eauto. intros. subst parent'. clear B.
- exploit lessdef_parent_ra; eauto. intros. subst ra'. clear D.
- exploit Mem.free_parallel_extends; eauto. intros [m2' [E F]].
- 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.
- generalize (code_tail_next_int _ _ _ _ NOOV H8). intro CT1.
- left; econstructor; split.
- eapply plus_left. eapply exec_step_internal. eauto.
- eapply functions_transl; eauto. eapply find_instr_tail; eauto.
- simpl. replace (chunk_of_type Tptr) with Mptr in * by (unfold Tptr, Mptr; destruct Archi.ptr64; auto).
- rewrite C. rewrite A. rewrite <- (sp_val _ _ _ AG). rewrite E. eauto.
- apply star_one. eapply exec_step_internal.
- transitivity (Val.offset_ptr rs0#PC Ptrofs.one). auto. rewrite <- H4. simpl. eauto.
- eapply functions_transl; eauto. eapply find_instr_tail; eauto.
- simpl. eauto. traceEq.
- econstructor; eauto.
- apply agree_set_other; auto. apply agree_nextinstr. apply agree_set_other; auto.
- eapply agree_change_sp; eauto. eapply parent_sp_def; eauto.
- Simplifs. rewrite Pregmap.gso; auto.
- generalize (preg_of_not_SP rf). rewrite (ireg_of_eq _ _ EQ1). congruence.
-+ (* Direct call *)
- generalize (code_tail_next_int _ _ _ _ NOOV H8). intro CT1.
- left; econstructor; split.
- eapply plus_left. eapply exec_step_internal. eauto.
- eapply functions_transl; eauto. eapply find_instr_tail; eauto.
- simpl. replace (chunk_of_type Tptr) with Mptr in * by (unfold Tptr, Mptr; destruct Archi.ptr64; auto).
- rewrite C. rewrite A. rewrite <- (sp_val _ _ _ AG). rewrite E. eauto.
- apply star_one. eapply exec_step_internal.
- transitivity (Val.offset_ptr rs0#PC Ptrofs.one). auto. rewrite <- H4. simpl. eauto.
- eapply functions_transl; eauto. eapply find_instr_tail; eauto.
- simpl. eauto. traceEq.
- econstructor; eauto.
- apply agree_set_other; auto. apply agree_nextinstr. apply agree_set_other; auto.
- eapply agree_change_sp; eauto. eapply parent_sp_def; eauto.
- rewrite Pregmap.gss. 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_nf, nextinstr. rewrite Pregmap.gss.
- rewrite undef_regs_other. 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.
- simpl; intros. intuition congruence.
- apply agree_nextinstr_nf. eapply agree_set_res; auto.
- eapply agree_undef_regs; eauto. intros; apply undef_regs_other_2; auto.
- congruence.
-
-- (* Mgoto *)
- assert (f0 = f) by congruence. subst f0.
- inv AT. monadInv H4.
- exploit find_label_goto_label; eauto. intros [tc' [rs' [GOTO [AT2 INV]]]].
- left; exists (State rs' m'); split.
- apply plus_one. econstructor; eauto.
- eapply functions_transl; eauto.
- eapply find_instr_tail; eauto.
- simpl; eauto.
- econstructor; eauto.
- eapply agree_exten; eauto with asmgen.
- congruence.
-
-- (* Mcond true *)
- assert (f0 = f) by congruence. subst f0.
- exploit eval_condition_lessdef. eapply preg_vals; eauto. eauto. eauto. intros EC.
- left; eapply exec_straight_steps_goto; eauto.
- intros. simpl in TR.
- destruct (transl_cond_correct tge tf cond args _ _ rs0 m' TR)
- as [rs' [A [B C]]].
- rewrite EC in B.
- destruct (testcond_for_condition cond); simpl in *.
-(* simple jcc *)
- exists (Pjcc c1 lbl); exists k; exists rs'.
- split. eexact A.
- split. eapply agree_exten; eauto.
- simpl. rewrite B. auto.
-(* jcc; jcc *)
- destruct (eval_testcond c1 rs') as [b1|] eqn:TC1;
- destruct (eval_testcond c2 rs') as [b2|] eqn:TC2; inv B.
- destruct b1.
- (* first jcc jumps *)
- exists (Pjcc c1 lbl); exists (Pjcc c2 lbl :: k); exists rs'.
- split. eexact A.
- split. eapply agree_exten; eauto.
- simpl. rewrite TC1. auto.
- (* second jcc jumps *)
- exists (Pjcc c2 lbl); exists k; exists (nextinstr rs').
- split. eapply exec_straight_trans. eexact A.
- eapply exec_straight_one. simpl. rewrite TC1. auto. auto.
- split. eapply agree_exten; eauto.
- intros; Simplifs.
- simpl. rewrite eval_testcond_nextinstr. rewrite TC2.
- destruct b2; auto || discriminate.
-(* jcc2 *)
- destruct (eval_testcond c1 rs') as [b1|] eqn:TC1;
- destruct (eval_testcond c2 rs') as [b2|] eqn:TC2; inv B.
- destruct (andb_prop _ _ H3). subst.
- exists (Pjcc2 c1 c2 lbl); exists k; exists rs'.
- split. eexact A.
- split. eapply agree_exten; eauto.
- simpl. rewrite TC1; rewrite TC2; auto.
-
-- (* Mcond false *)
- exploit eval_condition_lessdef. eapply preg_vals; eauto. eauto. eauto. intros EC.
- left; eapply exec_straight_steps; eauto. intros. simpl in TR.
- destruct (transl_cond_correct tge tf cond args _ _ rs0 m' TR)
- as [rs' [A [B C]]].
- rewrite EC in B.
- destruct (testcond_for_condition cond); simpl in *.
-(* simple jcc *)
- econstructor; split.
- eapply exec_straight_trans. eexact A.
- apply exec_straight_one. simpl. rewrite B. eauto. auto.
- split. apply agree_nextinstr. eapply agree_exten; eauto.
- simpl; congruence.
-(* jcc ; jcc *)
- destruct (eval_testcond c1 rs') as [b1|] eqn:TC1;
- destruct (eval_testcond c2 rs') as [b2|] eqn:TC2; inv B.
- destruct (orb_false_elim _ _ H1); subst.
- econstructor; split.
- eapply exec_straight_trans. eexact A.
- eapply exec_straight_two. simpl. rewrite TC1. eauto. auto.
- simpl. rewrite eval_testcond_nextinstr. rewrite TC2. eauto. auto. auto.
- split. apply agree_nextinstr. apply agree_nextinstr. eapply agree_exten; eauto.
- simpl; congruence.
-(* jcc2 *)
- destruct (eval_testcond c1 rs') as [b1|] eqn:TC1;
- destruct (eval_testcond c2 rs') as [b2|] eqn:TC2; inv B.
- exists (nextinstr rs'); split.
- eapply exec_straight_trans. eexact A.
- apply exec_straight_one. simpl.
- rewrite TC1; rewrite TC2.
- destruct b1. simpl in *. subst b2. auto. auto.
- auto.
- split. apply agree_nextinstr. eapply agree_exten; eauto.
- rewrite H1; congruence.
-
-- (* Mjumptable *)
- assert (f0 = f) by congruence. subst f0.
- inv AT. monadInv H6.
- exploit functions_transl; eauto. intro FN.
- generalize (transf_function_no_overflow _ _ H5); intro NOOV.
- set (rs1 := rs0 #RAX <- Vundef #RDX <- Vundef).
- exploit (find_label_goto_label f tf lbl rs1); 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. rewrite <- H9. unfold Mach.label in H0; unfold label; rewrite H0. eexact A.
- econstructor; eauto.
-Transparent destroyed_by_jumptable.
- apply agree_undef_regs with rs0; auto.
- simpl; intros. destruct H8. rewrite C by auto with asmgen. unfold rs1; Simplifs.
- congruence.
-
-- (* Mreturn *)
- 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.
- rewrite (sp_val _ _ _ AG) in *. unfold load_stack in *.
- replace (chunk_of_type Tptr) with Mptr in * by (unfold Tptr, Mptr; destruct Archi.ptr64; auto).
- exploit Mem.loadv_extends. eauto. eexact H0. auto. simpl. intros [parent' [A B]].
- exploit lessdef_parent_sp; eauto. intros. subst parent'. clear B.
- exploit Mem.loadv_extends. eauto. eexact H1. auto. simpl. intros [ra' [C D]].
- exploit lessdef_parent_ra; eauto. intros. subst ra'. clear D.
- exploit Mem.free_parallel_extends; eauto. intros [m2' [E F]].
- monadInv H6.
- exploit code_tail_next_int; eauto. intro CT1.
- left; econstructor; split.
- eapply plus_left. eapply exec_step_internal. eauto.
- eapply functions_transl; eauto. eapply find_instr_tail; eauto.
- simpl. rewrite C. rewrite A. rewrite <- (sp_val _ _ _ AG). rewrite E. eauto.
- apply star_one. eapply exec_step_internal.
- transitivity (Val.offset_ptr rs0#PC Ptrofs.one). auto. rewrite <- H3. simpl. eauto.
- eapply functions_transl; eauto. eapply find_instr_tail; eauto.
- simpl. eauto. traceEq.
- constructor; auto.
- apply agree_set_other; auto. apply agree_nextinstr. apply agree_set_other; auto.
- eapply agree_change_sp; eauto. eapply parent_sp_def; eauto.
-
-- (* 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 (fn_code x0))); inv EQ1.
- monadInv EQ0. rewrite transl_code'_transl_code in EQ1.
- unfold store_stack in *.
- exploit Mem.alloc_extends. eauto. eauto. apply Zle_refl. apply Zle_refl.
- intros [m1' [C D]].
- exploit Mem.storev_extends. eexact D. eexact H1. eauto. eauto.
- intros [m2' [F G]].
- exploit Mem.storev_extends. eexact G. eexact H2. eauto. eauto.
- intros [m3' [P Q]].
- left; econstructor; split.
- apply plus_one. econstructor; eauto.
- simpl. rewrite Ptrofs.unsigned_zero. simpl. eauto.
- simpl. rewrite C. simpl in F, P.
- replace (chunk_of_type Tptr) with Mptr in F, P by (unfold Tptr, Mptr; destruct Archi.ptr64; auto).
- rewrite (sp_val _ _ _ AG) in F. rewrite F.
- rewrite ATLR. rewrite P. eauto.
- econstructor; eauto.
- unfold nextinstr. rewrite Pregmap.gss. repeat rewrite Pregmap.gso; auto with asmgen.
- rewrite ATPC. simpl. constructor; eauto.
- unfold fn_code. eapply code_tail_next_int. simpl in g. omega.
- constructor.
- apply agree_nextinstr. eapply agree_change_sp; eauto.
-Transparent destroyed_at_function_entry.
- apply agree_undef_regs with rs0; eauto.
- simpl; intros. apply Pregmap.gso; auto with asmgen. tauto.
- congruence.
- intros. Simplifs. eapply agree_sp; eauto.
-
-- (* 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.
-
-- (* return *)
- inv STACKS. simpl in *.
- right. split. omega. split. auto.
- econstructor; eauto. rewrite ATPC; eauto. 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. reflexivity. 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. auto.
- assert (r0 = AX).
- { unfold loc_result in H1; destruct Archi.ptr64; compute in H1; congruence. }
- subst r0.
- generalize (preg_val _ _ _ AX 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).
- apply senv_preserved.
- eexact transf_initial_states.
- eexact transf_final_states.
- exact step_simulation.
-Qed.
-
-End PRESERVATION.
diff --git a/ia32/Asmgenproof1.v b/ia32/Asmgenproof1.v
deleted file mode 100644
index 05b3176a..00000000
--- a/ia32/Asmgenproof1.v
+++ /dev/null
@@ -1,1481 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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 x86-64 generation: auxiliary results. *)
-
-Require Import Coqlib.
-Require Import AST Errors Integers Floats Values Memory Globalenvs.
-Require Import Op Locations Conventions Mach Asm.
-Require Import Asmgen Asmgenproof0.
-
-Open Local Scope error_monad_scope.
-Local Transparent Archi.ptr64.
-
-(** * Correspondence between Mach registers and x86 registers *)
-
-Lemma agree_nextinstr_nf:
- forall ms sp rs,
- agree ms sp rs -> agree ms sp (nextinstr_nf rs).
-Proof.
- intros. unfold nextinstr_nf. apply agree_nextinstr.
- apply agree_undef_nondata_regs. auto.
- simpl; intros. intuition (subst r; auto).
-Qed.
-
-(** Useful properties of the PC register. *)
-
-Lemma nextinstr_nf_inv:
- forall r rs,
- match r with PC => False | CR _ => False | _ => True end ->
- (nextinstr_nf rs)#r = rs#r.
-Proof.
- intros. unfold nextinstr_nf. rewrite nextinstr_inv.
- simpl. repeat rewrite Pregmap.gso; auto;
- red; intro; subst; contradiction.
- red; intro; subst; contradiction.
-Qed.
-
-Lemma nextinstr_nf_inv1:
- forall r rs,
- data_preg r = true -> (nextinstr_nf rs)#r = rs#r.
-Proof.
- intros. apply nextinstr_nf_inv. destruct r; auto || discriminate.
-Qed.
-
-Lemma nextinstr_nf_set_preg:
- forall rs m v,
- (nextinstr_nf (rs#(preg_of m) <- v))#PC = Val.offset_ptr rs#PC Ptrofs.one.
-Proof.
- intros. unfold nextinstr_nf.
- transitivity (nextinstr (rs#(preg_of m) <- v) PC). auto.
- apply nextinstr_set_preg.
-Qed.
-
-(** Useful simplification tactic *)
-
-Ltac Simplif :=
- match goal with
- | [ |- nextinstr_nf _ _ = _ ] =>
- ((rewrite nextinstr_nf_inv by auto with asmgen)
- || (rewrite nextinstr_nf_inv1 by auto with asmgen)); auto
- | [ |- nextinstr _ _ = _ ] =>
- ((rewrite nextinstr_inv by auto with asmgen)
- || (rewrite nextinstr_inv1 by auto with asmgen)); auto
- | [ |- Pregmap.get ?x (Pregmap.set ?x _ _) = _ ] =>
- rewrite Pregmap.gss; auto
- | [ |- Pregmap.set ?x _ _ ?x = _ ] =>
- rewrite Pregmap.gss; auto
- | [ |- Pregmap.get _ (Pregmap.set _ _ _) = _ ] =>
- rewrite Pregmap.gso by (auto with asmgen); auto
- | [ |- Pregmap.set _ _ _ _ = _ ] =>
- rewrite Pregmap.gso by (auto with asmgen); auto
- end.
-
-Ltac Simplifs := repeat Simplif.
-
-(** * Correctness of x86-64 constructor functions *)
-
-Section CONSTRUCTORS.
-
-Variable ge: genv.
-Variable fn: function.
-
-(** Smart constructor for moves. *)
-
-Lemma mk_mov_correct:
- forall rd rs k c rs1 m,
- mk_mov rd rs k = OK c ->
- exists rs2,
- exec_straight ge fn c rs1 m k rs2 m
- /\ rs2#rd = rs1#rs
- /\ forall r, data_preg r = true -> r <> rd -> rs2#r = rs1#r.
-Proof.
- unfold mk_mov; intros.
- destruct rd; try (monadInv H); destruct rs; monadInv H.
-(* mov *)
- econstructor. split. apply exec_straight_one. simpl. eauto. auto.
- split. Simplifs. intros; Simplifs.
-(* movsd *)
- econstructor. split. apply exec_straight_one. simpl. eauto. auto.
- split. Simplifs. intros; Simplifs.
-Qed.
-
-(** Properties of division *)
-
-Lemma divu_modu_exists:
- forall v1 v2,
- Val.divu v1 v2 <> None \/ Val.modu v1 v2 <> None ->
- exists n d q r,
- v1 = Vint n /\ v2 = Vint d
- /\ Int.divmodu2 Int.zero n d = Some(q, r)
- /\ Val.divu v1 v2 = Some (Vint q) /\ Val.modu v1 v2 = Some (Vint r).
-Proof.
- intros v1 v2; unfold Val.divu, Val.modu.
- destruct v1; try (intuition discriminate).
- destruct v2; try (intuition discriminate).
- predSpec Int.eq Int.eq_spec i0 Int.zero ; try (intuition discriminate).
- intros _. exists i, i0, (Int.divu i i0), (Int.modu i i0); intuition auto.
- apply Int.divmodu2_divu_modu; auto.
-Qed.
-
-Lemma divs_mods_exists:
- forall v1 v2,
- Val.divs v1 v2 <> None \/ Val.mods v1 v2 <> None ->
- exists nh nl d q r,
- Val.shr v1 (Vint (Int.repr 31)) = Vint nh /\ v1 = Vint nl /\ v2 = Vint d
- /\ Int.divmods2 nh nl d = Some(q, r)
- /\ Val.divs v1 v2 = Some (Vint q) /\ Val.mods v1 v2 = Some (Vint r).
-Proof.
- intros v1 v2; unfold Val.divs, Val.mods.
- destruct v1; try (intuition discriminate).
- destruct v2; try (intuition discriminate).
- destruct (Int.eq i0 Int.zero
- || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone) eqn:OK;
- try (intuition discriminate).
- intros _.
- InvBooleans.
- exists (Int.shr i (Int.repr 31)), i, i0, (Int.divs i i0), (Int.mods i i0); intuition auto.
- rewrite Int.shr_lt_zero. apply Int.divmods2_divs_mods.
- red; intros; subst i0; rewrite Int.eq_true in H; discriminate.
- revert H0. predSpec Int.eq Int.eq_spec i (Int.repr Int.min_signed); auto.
- predSpec Int.eq Int.eq_spec i0 Int.mone; auto.
- discriminate.
-Qed.
-
-Lemma divlu_modlu_exists:
- forall v1 v2,
- Val.divlu v1 v2 <> None \/ Val.modlu v1 v2 <> None ->
- exists n d q r,
- v1 = Vlong n /\ v2 = Vlong d
- /\ Int64.divmodu2 Int64.zero n d = Some(q, r)
- /\ Val.divlu v1 v2 = Some (Vlong q) /\ Val.modlu v1 v2 = Some (Vlong r).
-Proof.
- intros v1 v2; unfold Val.divlu, Val.modlu.
- destruct v1; try (intuition discriminate).
- destruct v2; try (intuition discriminate).
- predSpec Int64.eq Int64.eq_spec i0 Int64.zero ; try (intuition discriminate).
- intros _. exists i, i0, (Int64.divu i i0), (Int64.modu i i0); intuition auto.
- apply Int64.divmodu2_divu_modu; auto.
-Qed.
-
-Lemma divls_modls_exists:
- forall v1 v2,
- Val.divls v1 v2 <> None \/ Val.modls v1 v2 <> None ->
- exists nh nl d q r,
- Val.shrl v1 (Vint (Int.repr 63)) = Vlong nh /\ v1 = Vlong nl /\ v2 = Vlong d
- /\ Int64.divmods2 nh nl d = Some(q, r)
- /\ Val.divls v1 v2 = Some (Vlong q) /\ Val.modls v1 v2 = Some (Vlong r).
-Proof.
- intros v1 v2; unfold Val.divls, Val.modls.
- destruct v1; try (intuition discriminate).
- destruct v2; try (intuition discriminate).
- destruct (Int64.eq i0 Int64.zero
- || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone) eqn:OK;
- try (intuition discriminate).
- intros _.
- InvBooleans.
- exists (Int64.shr i (Int64.repr 63)), i, i0, (Int64.divs i i0), (Int64.mods i i0); intuition auto.
- rewrite Int64.shr_lt_zero. apply Int64.divmods2_divs_mods.
- red; intros; subst i0; rewrite Int64.eq_true in H; discriminate.
- revert H0. predSpec Int64.eq Int64.eq_spec i (Int64.repr Int64.min_signed); auto.
- predSpec Int64.eq Int64.eq_spec i0 Int64.mone; auto.
- discriminate.
-Qed.
-
-(** Smart constructor for [shrx] *)
-
-Lemma mk_shrximm_correct:
- forall n k c (rs1: regset) v m,
- mk_shrximm n k = OK c ->
- Val.shrx (rs1#RAX) (Vint n) = Some v ->
- exists rs2,
- exec_straight ge fn c rs1 m k rs2 m
- /\ rs2#RAX = v
- /\ forall r, data_preg r = true -> r <> RAX -> r <> RCX -> rs2#r = rs1#r.
-Proof.
- unfold mk_shrximm; intros. inv H.
- exploit Val.shrx_shr; eauto. intros [x [y [A [B C]]]].
- inversion B; clear B; subst y; subst v; clear H0.
- set (tnm1 := Int.sub (Int.shl Int.one n) Int.one).
- set (x' := Int.add x tnm1).
- set (rs2 := nextinstr (compare_ints (Vint x) (Vint Int.zero) rs1 m)).
- set (rs3 := nextinstr (rs2#RCX <- (Vint x'))).
- set (rs4 := nextinstr (if Int.lt x Int.zero then rs3#RAX <- (Vint x') else rs3)).
- set (rs5 := nextinstr_nf (rs4#RAX <- (Val.shr rs4#RAX (Vint n)))).
- assert (rs3#RAX = Vint x). unfold rs3. Simplifs.
- assert (rs3#RCX = Vint x'). unfold rs3. Simplifs.
- exists rs5. split.
- apply exec_straight_step with rs2 m. simpl. rewrite A. simpl. rewrite Int.and_idem. auto. auto.
- apply exec_straight_step with rs3 m. simpl.
- change (rs2 RAX) with (rs1 RAX). rewrite A. simpl.
- rewrite Int.repr_unsigned. rewrite Int.add_zero_l. auto. auto.
- apply exec_straight_step with rs4 m. simpl.
- rewrite Int.lt_sub_overflow. unfold rs4. destruct (Int.lt x Int.zero); simpl; auto.
- unfold rs4. destruct (Int.lt x Int.zero); simpl; auto.
- apply exec_straight_one. auto. auto.
- split. unfold rs5. Simplifs. unfold rs4. rewrite nextinstr_inv; auto with asmgen.
- destruct (Int.lt x Int.zero). rewrite Pregmap.gss. rewrite A; auto. rewrite A; rewrite H; auto.
- intros. unfold rs5. Simplifs. unfold rs4. Simplifs.
- transitivity (rs3#r). destruct (Int.lt x Int.zero). Simplifs. auto.
- unfold rs3. Simplifs. unfold rs2. Simplifs.
- unfold compare_ints. Simplifs.
-Qed.
-
-(** Smart constructor for [shrxl] *)
-
-Lemma mk_shrxlimm_correct:
- forall n k c (rs1: regset) v m,
- mk_shrxlimm n k = OK c ->
- Val.shrxl (rs1#RAX) (Vint n) = Some v ->
- exists rs2,
- exec_straight ge fn c rs1 m k rs2 m
- /\ rs2#RAX = v
- /\ forall r, data_preg r = true -> r <> RAX -> r <> RDX -> rs2#r = rs1#r.
-Proof.
- unfold mk_shrxlimm; intros. exploit Val.shrxl_shrl_2; eauto. intros EQ.
- destruct (Int.eq n Int.zero); inv H.
-- econstructor; split. apply exec_straight_one. simpl; reflexivity. auto.
- split. Simplifs. intros; Simplifs.
-- set (v1 := Val.shrl (rs1 RAX) (Vint (Int.repr 63))) in *.
- set (v2 := Val.shrlu v1 (Vint (Int.sub (Int.repr 64) n))) in *.
- set (v3 := Val.addl (rs1 RAX) v2) in *.
- set (v4 := Val.shrl v3 (Vint n)) in *.
- set (rs2 := nextinstr_nf (rs1#RDX <- v1)).
- set (rs3 := nextinstr_nf (rs2#RDX <- v2)).
- set (rs4 := nextinstr (rs3#RAX <- v3)).
- set (rs5 := nextinstr_nf (rs4#RAX <- v4)).
- assert (X: forall v1 v2,
- Val.addl v1 (Val.addl v2 (Vlong Int64.zero)) = Val.addl v1 v2).
- { intros. destruct v0; simpl; auto; destruct v5; simpl; auto.
- rewrite Int64.add_zero; auto.
- destruct Archi.ptr64; auto. rewrite Ptrofs.add_zero; auto.
- destruct Archi.ptr64; auto. rewrite Int64.add_zero; auto.
- destruct Archi.ptr64; auto. }
- exists rs5; split.
- eapply exec_straight_trans with (rs2 := rs3).
- eapply exec_straight_two with (rs2 := rs2); reflexivity.
- eapply exec_straight_two with (rs2 := rs4).
- simpl. rewrite X. reflexivity. reflexivity. reflexivity. reflexivity.
- split. unfold rs5; Simplifs.
- intros. unfold rs5; Simplifs. unfold rs4; Simplifs. unfold rs3; Simplifs. unfold rs2; Simplifs.
-Qed.
-
-(** Smart constructor for integer conversions *)
-
-Lemma mk_intconv_correct:
- forall mk sem rd rs k c rs1 m,
- mk_intconv mk rd rs k = OK c ->
- (forall c rd rs r m,
- exec_instr ge c (mk rd rs) r m = Next (nextinstr (r#rd <- (sem r#rs))) m) ->
- exists rs2,
- exec_straight ge fn c rs1 m k rs2 m
- /\ rs2#rd = sem rs1#rs
- /\ forall r, data_preg r = true -> r <> rd -> r <> RAX -> rs2#r = rs1#r.
-Proof.
- unfold mk_intconv; intros. destruct (Archi.ptr64 || low_ireg rs); monadInv H.
- econstructor. split. apply exec_straight_one. rewrite H0. eauto. auto.
- split. Simplifs. intros. Simplifs.
- econstructor. split. eapply exec_straight_two.
- simpl. eauto. apply H0. auto. auto.
- split. Simplifs. intros. Simplifs.
-Qed.
-
-(** Smart constructor for small stores *)
-
-Lemma addressing_mentions_correct:
- forall a r (rs1 rs2: regset),
- (forall (r': ireg), r' <> r -> rs1 r' = rs2 r') ->
- addressing_mentions a r = false ->
- eval_addrmode32 ge a rs1 = eval_addrmode32 ge a rs2.
-Proof.
- intros until rs2; intro AG. unfold addressing_mentions, eval_addrmode32.
- destruct a. intros. destruct (orb_false_elim _ _ H). unfold proj_sumbool in *.
- decEq. destruct base; auto. apply AG. destruct (ireg_eq r i); congruence.
- decEq. destruct ofs as [[r' sc] | ]; auto. rewrite AG; auto. destruct (ireg_eq r r'); congruence.
-Qed.
-
-Lemma mk_storebyte_correct:
- forall addr r k c rs1 m1 m2,
- mk_storebyte addr r k = OK c ->
- Mem.storev Mint8unsigned m1 (eval_addrmode ge addr rs1) (rs1 r) = Some m2 ->
- exists rs2,
- exec_straight ge fn c rs1 m1 k rs2 m2
- /\ forall r, data_preg r = true -> preg_notin r (if Archi.ptr64 then nil else AX :: CX :: nil) -> rs2#r = rs1#r.
-Proof.
- unfold mk_storebyte; intros.
- destruct (Archi.ptr64 || low_ireg r) eqn:E.
-(* low reg *)
- monadInv H. econstructor; split. apply exec_straight_one.
- simpl. unfold exec_store. rewrite H0. eauto. auto.
- intros; Simplifs.
-(* high reg *)
- InvBooleans. rewrite H1; simpl. destruct (addressing_mentions addr RAX) eqn:E; monadInv H.
-(* RAX is mentioned. *)
- assert (r <> RCX). { red; intros; subst r; discriminate H2. }
- set (rs2 := nextinstr (rs1#RCX <- (eval_addrmode32 ge addr rs1))).
- set (rs3 := nextinstr (rs2#RAX <- (rs1 r))).
- econstructor; split.
- apply exec_straight_three with rs2 m1 rs3 m1.
- simpl. auto.
- simpl. replace (rs2 r) with (rs1 r). auto. symmetry. unfold rs2; Simplifs.
- simpl. unfold exec_store. unfold eval_addrmode; rewrite H1; simpl. rewrite Int.add_zero.
- change (rs3 RAX) with (rs1 r).
- change (rs3 RCX) with (eval_addrmode32 ge addr rs1).
- replace (Val.add (eval_addrmode32 ge addr rs1) (Vint Int.zero))
- with (eval_addrmode ge addr rs1).
- rewrite H0. eauto.
- unfold eval_addrmode in *; rewrite H1 in *.
- destruct (eval_addrmode32 ge addr rs1); simpl in H0; try discriminate.
- simpl. rewrite H1. rewrite Ptrofs.add_zero; auto.
- auto. auto. auto.
- intros. destruct H4. Simplifs. unfold rs3; Simplifs. unfold rs2; Simplifs.
-(* RAX is not mentioned *)
- set (rs2 := nextinstr (rs1#RAX <- (rs1 r))).
- econstructor; split.
- apply exec_straight_two with rs2 m1.
- simpl. auto.
- simpl. unfold exec_store. unfold eval_addrmode in *; rewrite H1 in *.
- rewrite (addressing_mentions_correct addr RAX rs2 rs1); auto.
- change (rs2 RAX) with (rs1 r). rewrite H0. eauto.
- intros. unfold rs2; Simplifs.
- auto. auto.
- intros. destruct H3. simpl. Simplifs. unfold rs2; Simplifs.
-Qed.
-
-(** Accessing slots in the stack frame *)
-
-Remark eval_addrmode_indexed:
- forall (base: ireg) ofs (rs: regset),
- match rs#base with Vptr _ _ => True | _ => False end ->
- eval_addrmode ge (Addrmode (Some base) None (inl _ (Ptrofs.unsigned ofs))) rs = Val.offset_ptr rs#base ofs.
-Proof.
- intros. destruct (rs#base) eqn:BASE; try contradiction.
- intros; unfold eval_addrmode; destruct Archi.ptr64 eqn:SF; simpl; rewrite BASE; simpl; rewrite SF; simpl.
-- apply f_equal. apply f_equal. rewrite Int64.add_zero_l. rewrite <- (Ptrofs.repr_unsigned ofs) at 2. auto with ptrofs.
-- apply f_equal. apply f_equal. rewrite Int.add_zero_l. rewrite <- (Ptrofs.repr_unsigned ofs) at 2. auto with ptrofs.
-Qed.
-
-Ltac loadind_correct_solve :=
- match goal with
- | H: Error _ = OK _ |- _ => discriminate H
- | H: OK _ = OK _ |- _ => inv H
- | H: match ?x with _ => _ end = OK _ |- _ => destruct x eqn:?; loadind_correct_solve
- | _ => idtac
- end.
-
-Lemma loadind_correct:
- forall (base: ireg) ofs ty dst k (rs: regset) c m v,
- loadind base ofs ty dst k = OK c ->
- Mem.loadv (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) = Some v ->
- exists rs',
- exec_straight ge fn c rs m k rs' m
- /\ rs'#(preg_of dst) = v
- /\ forall r, data_preg r = true -> r <> preg_of dst -> rs'#r = rs#r.
-Proof.
- unfold loadind; intros.
- set (addr := Addrmode (Some base) None (inl (ident * ptrofs) (Ptrofs.unsigned ofs))) in *.
- assert (eval_addrmode ge addr rs = Val.offset_ptr rs#base ofs).
- { apply eval_addrmode_indexed. destruct (rs base); auto || discriminate. }
- rewrite <- H1 in H0.
- exists (nextinstr_nf (rs#(preg_of dst) <- v)); split.
-- loadind_correct_solve; apply exec_straight_one; auto; simpl in *; unfold exec_load; rewrite ?Heqb, ?H0; auto.
-- intuition Simplifs.
-Qed.
-
-Lemma storeind_correct:
- forall (base: ireg) ofs ty src k (rs: regset) c m m',
- storeind src base ofs ty k = OK c ->
- Mem.storev (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) (rs#(preg_of src)) = Some m' ->
- exists rs',
- exec_straight ge fn c rs m k rs' m'
- /\ forall r, data_preg r = true -> preg_notin r (destroyed_by_setstack ty) -> rs'#r = rs#r.
-Proof.
- unfold storeind; intros.
- set (addr := Addrmode (Some base) None (inl (ident * ptrofs) (Ptrofs.unsigned ofs))) in *.
- assert (eval_addrmode ge addr rs = Val.offset_ptr rs#base ofs).
- { apply eval_addrmode_indexed. destruct (rs base); auto || discriminate. }
- rewrite <- H1 in H0.
- loadind_correct_solve; simpl in H0;
- (econstructor; split;
- [apply exec_straight_one; [simpl; unfold exec_store; rewrite ?Heqb, H0;eauto|auto]
- |simpl; intros; unfold undef_regs; repeat Simplifs]).
-Qed.
-
-(** Translation of addressing modes *)
-
-Lemma transl_addressing_mode_32_correct:
- forall addr args am (rs: regset) v,
- transl_addressing addr args = OK am ->
- eval_addressing32 ge (rs RSP) addr (List.map rs (List.map preg_of args)) = Some v ->
- Val.lessdef v (eval_addrmode32 ge am rs).
-Proof.
- assert (A: forall id ofs, Archi.ptr64 = false ->
- Val.add (Vint Int.zero) (Genv.symbol_address ge id ofs) = Genv.symbol_address ge id ofs).
- { intros. unfold Val.add; rewrite H. unfold Genv.symbol_address.
- destruct (Genv.find_symbol ge id); auto. rewrite Ptrofs.add_zero; auto. }
- assert (C: forall v i,
- Val.lessdef (Val.mul v (Vint (Int.repr i)))
- (if zeq i 1 then v else Val.mul v (Vint (Int.repr i)))).
- { intros. destruct (zeq i 1); subst; auto.
- destruct v; simpl; auto. rewrite Int.mul_one; auto. }
- unfold transl_addressing; intros.
- destruct addr; repeat (destruct args; try discriminate); simpl in H0; inv H0;
- monadInv H; try (erewrite ! ireg_of_eq by eauto); unfold eval_addrmode32.
-- simpl; rewrite Int.add_zero_l; auto.
-- rewrite Val.add_assoc. apply Val.add_lessdef; auto.
-- rewrite Val.add_permut. apply Val.add_lessdef; auto. simpl; rewrite Int.add_zero_l; auto.
-- apply Val.add_lessdef; auto. apply Val.add_lessdef; auto.
-- destruct Archi.ptr64 eqn:SF; inv H2. rewrite ! A by auto. auto.
-- destruct Archi.ptr64 eqn:SF; inv H2. erewrite ireg_of_eq by eauto.
- rewrite Val.add_commut. rewrite A by auto. auto.
-- destruct Archi.ptr64 eqn:SF; inv H2. erewrite ireg_of_eq by eauto.
- rewrite Val.add_permut. rewrite Val.add_commut. apply Val.add_lessdef; auto. rewrite A; auto.
-- destruct Archi.ptr64 eqn:SF; inv H2. simpl.
- destruct (rs RSP); simpl; auto; rewrite SF.
- rewrite Int.add_zero_l. apply Val.lessdef_same; f_equal; f_equal.
- symmetry. transitivity (Ptrofs.repr (Ptrofs.signed i)). auto with ptrofs. auto with ints.
-Qed.
-
-Lemma transl_addressing_mode_64_correct:
- forall addr args am (rs: regset) v,
- transl_addressing addr args = OK am ->
- eval_addressing64 ge (rs RSP) addr (List.map rs (List.map preg_of args)) = Some v ->
- Val.lessdef v (eval_addrmode64 ge am rs).
-Proof.
- assert (A: forall id ofs, Archi.ptr64 = true ->
- Val.addl (Vlong Int64.zero) (Genv.symbol_address ge id ofs) = Genv.symbol_address ge id ofs).
- { intros. unfold Val.addl; rewrite H. unfold Genv.symbol_address.
- destruct (Genv.find_symbol ge id); auto. rewrite Ptrofs.add_zero; auto. }
- assert (C: forall v i,
- Val.lessdef (Val.mull v (Vlong (Int64.repr i)))
- (if zeq i 1 then v else Val.mull v (Vlong (Int64.repr i)))).
- { intros. destruct (zeq i 1); subst; auto.
- destruct v; simpl; auto. rewrite Int64.mul_one; auto. }
- unfold transl_addressing; intros.
- destruct addr; repeat (destruct args; try discriminate); simpl in H0; inv H0;
- monadInv H; try (erewrite ! ireg_of_eq by eauto); unfold eval_addrmode64.
-- simpl; rewrite Int64.add_zero_l; auto.
-- rewrite Val.addl_assoc. apply Val.addl_lessdef; auto.
-- rewrite Val.addl_permut. apply Val.addl_lessdef; auto. simpl; rewrite Int64.add_zero_l; auto.
-- apply Val.addl_lessdef; auto. apply Val.addl_lessdef; auto.
-- destruct Archi.ptr64 eqn:SF; inv H2. rewrite ! A by auto. auto.
-- destruct Archi.ptr64 eqn:SF; inv H2. simpl.
- destruct (rs RSP); simpl; auto; rewrite SF.
- rewrite Int64.add_zero_l. apply Val.lessdef_same; f_equal; f_equal.
- symmetry. transitivity (Ptrofs.repr (Ptrofs.signed i)). auto with ptrofs. auto with ints.
-Qed.
-
-Lemma transl_addressing_mode_correct:
- forall addr args am (rs: regset) v,
- transl_addressing addr args = OK am ->
- eval_addressing ge (rs RSP) addr (List.map rs (List.map preg_of args)) = Some v ->
- Val.lessdef v (eval_addrmode ge am rs).
-Proof.
- unfold eval_addressing, eval_addrmode; intros. destruct Archi.ptr64.
- eapply transl_addressing_mode_64_correct; eauto.
- eapply transl_addressing_mode_32_correct; eauto.
-Qed.
-
-Lemma normalize_addrmode_32_correct:
- forall am rs, eval_addrmode32 ge (normalize_addrmode_32 am) rs = eval_addrmode32 ge am rs.
-Proof.
- intros; destruct am as [base ofs [n|r]]; simpl; auto. rewrite Int.repr_signed. auto.
-Qed.
-
-Lemma normalize_addrmode_64_correct:
- forall am rs,
- eval_addrmode64 ge am rs =
- match normalize_addrmode_64 am with
- | (am', None) => eval_addrmode64 ge am' rs
- | (am', Some delta) => Val.addl (eval_addrmode64 ge am' rs) (Vlong delta)
- end.
-Proof.
- intros; destruct am as [base ofs [n|r]]; simpl; auto.
- destruct (zeq (Int.signed (Int.repr n)) n); simpl; auto.
- rewrite ! Val.addl_assoc. apply f_equal. apply f_equal. simpl. rewrite Int64.add_zero_l; auto.
-Qed.
-
-(** Processor conditions and comparisons *)
-
-Lemma compare_ints_spec:
- forall rs v1 v2 m,
- let rs' := nextinstr (compare_ints v1 v2 rs m) in
- rs'#ZF = Val.cmpu (Mem.valid_pointer m) Ceq v1 v2
- /\ rs'#CF = Val.cmpu (Mem.valid_pointer m) Clt v1 v2
- /\ rs'#SF = Val.negative (Val.sub v1 v2)
- /\ rs'#OF = Val.sub_overflow v1 v2
- /\ (forall r, data_preg r = true -> rs'#r = rs#r).
-Proof.
- intros. unfold rs'; unfold compare_ints.
- split. auto.
- split. auto.
- split. auto.
- split. auto.
- intros. Simplifs.
-Qed.
-
-Lemma testcond_for_signed_comparison_32_correct:
- forall c v1 v2 rs m b,
- Val.cmp_bool c v1 v2 = Some b ->
- eval_testcond (testcond_for_signed_comparison c)
- (nextinstr (compare_ints v1 v2 rs m)) = Some b.
-Proof.
- intros. generalize (compare_ints_spec rs v1 v2 m).
- set (rs' := nextinstr (compare_ints v1 v2 rs m)).
- intros [A [B [C [D E]]]].
- destruct v1; destruct v2; simpl in H; inv H.
- unfold eval_testcond. rewrite A; rewrite B; rewrite C; rewrite D.
- simpl. unfold Val.cmp, Val.cmpu.
- rewrite Int.lt_sub_overflow.
- destruct c; simpl.
- destruct (Int.eq i i0); auto.
- destruct (Int.eq i i0); auto.
- destruct (Int.lt i i0); auto.
- rewrite Int.not_lt. destruct (Int.lt i i0); simpl; destruct (Int.eq i i0); auto.
- rewrite (Int.lt_not i i0). destruct (Int.lt i i0); destruct (Int.eq i i0); reflexivity.
- destruct (Int.lt i i0); reflexivity.
-Qed.
-
-Lemma testcond_for_unsigned_comparison_32_correct:
- forall c v1 v2 rs m b,
- Val.cmpu_bool (Mem.valid_pointer m) c v1 v2 = Some b ->
- eval_testcond (testcond_for_unsigned_comparison c)
- (nextinstr (compare_ints v1 v2 rs m)) = Some b.
-Proof.
- intros. generalize (compare_ints_spec rs v1 v2 m).
- set (rs' := nextinstr (compare_ints v1 v2 rs m)).
- intros [A [B [C [D E]]]].
- unfold eval_testcond. rewrite A; rewrite B. unfold Val.cmpu, Val.cmp.
- destruct v1; destruct v2; simpl in H; inv H.
-- (* int int *)
- destruct c; simpl; auto.
- destruct (Int.eq i i0); reflexivity.
- destruct (Int.eq i i0); auto.
- destruct (Int.ltu i i0); auto.
- rewrite Int.not_ltu. destruct (Int.ltu i i0); simpl; destruct (Int.eq i i0); auto.
- rewrite (Int.ltu_not i i0). destruct (Int.ltu i i0); destruct (Int.eq i i0); reflexivity.
- destruct (Int.ltu i i0); reflexivity.
-- (* int ptr *)
- unfold Val.cmpu_bool; destruct Archi.ptr64; try discriminate.
- destruct (Int.eq i Int.zero &&
- (Mem.valid_pointer m b0 (Ptrofs.unsigned i0) || Mem.valid_pointer m b0 (Ptrofs.unsigned i0 - 1))); try discriminate.
- destruct c; simpl in *; inv H1; reflexivity.
-- (* ptr int *)
- unfold Val.cmpu_bool; destruct Archi.ptr64; try discriminate.
- destruct (Int.eq i0 Int.zero &&
- (Mem.valid_pointer m b0 (Ptrofs.unsigned i) || Mem.valid_pointer m b0 (Ptrofs.unsigned i - 1))); try discriminate.
- destruct c; simpl in *; inv H1; reflexivity.
-- (* ptr ptr *)
- unfold Val.cmpu_bool; destruct Archi.ptr64; try discriminate.
- fold (Mem.weak_valid_pointer m b0 (Ptrofs.unsigned i)) in *.
- fold (Mem.weak_valid_pointer m b1 (Ptrofs.unsigned i0)) in *.
- destruct (eq_block b0 b1).
- destruct (Mem.weak_valid_pointer m b0 (Ptrofs.unsigned i) &&
- Mem.weak_valid_pointer m b1 (Ptrofs.unsigned i0)); inv H1.
- destruct c; simpl; auto.
- destruct (Ptrofs.eq i i0); auto.
- destruct (Ptrofs.eq i i0); auto.
- destruct (Ptrofs.ltu i i0); auto.
- rewrite Ptrofs.not_ltu. destruct (Ptrofs.ltu i i0); simpl; destruct (Ptrofs.eq i i0); auto.
- rewrite (Ptrofs.ltu_not i i0). destruct (Ptrofs.ltu i i0); destruct (Ptrofs.eq i i0); reflexivity.
- destruct (Ptrofs.ltu i i0); reflexivity.
- destruct (Mem.valid_pointer m b0 (Ptrofs.unsigned i) &&
- Mem.valid_pointer m b1 (Ptrofs.unsigned i0)); try discriminate.
- destruct c; simpl in *; inv H1; reflexivity.
-Qed.
-
-Lemma compare_longs_spec:
- forall rs v1 v2 m,
- let rs' := nextinstr (compare_longs v1 v2 rs m) in
- rs'#ZF = Val.maketotal (Val.cmplu (Mem.valid_pointer m) Ceq v1 v2)
- /\ rs'#CF = Val.maketotal (Val.cmplu (Mem.valid_pointer m) Clt v1 v2)
- /\ rs'#SF = Val.negativel (Val.subl v1 v2)
- /\ rs'#OF = Val.subl_overflow v1 v2
- /\ (forall r, data_preg r = true -> rs'#r = rs#r).
-Proof.
- intros. unfold rs'; unfold compare_longs.
- split. auto.
- split. auto.
- split. auto.
- split. auto.
- intros. Simplifs.
-Qed.
-
-Lemma int64_sub_overflow:
- forall x y,
- Int.xor (Int.repr (Int64.unsigned (Int64.sub_overflow x y Int64.zero)))
- (Int.repr (Int64.unsigned (Int64.negative (Int64.sub x y)))) =
- (if Int64.lt x y then Int.one else Int.zero).
-Proof.
- intros.
- transitivity (Int.repr (Int64.unsigned (if Int64.lt x y then Int64.one else Int64.zero))).
- rewrite <- (Int64.lt_sub_overflow x y).
- unfold Int64.sub_overflow, Int64.negative.
- set (s := Int64.signed x - Int64.signed y - Int64.signed Int64.zero).
- destruct (zle Int64.min_signed s && zle s Int64.max_signed);
- destruct (Int64.lt (Int64.sub x y) Int64.zero);
- auto.
- destruct (Int64.lt x y); auto.
-Qed.
-
-Lemma testcond_for_signed_comparison_64_correct:
- forall c v1 v2 rs m b,
- Val.cmpl_bool c v1 v2 = Some b ->
- eval_testcond (testcond_for_signed_comparison c)
- (nextinstr (compare_longs v1 v2 rs m)) = Some b.
-Proof.
- intros. generalize (compare_longs_spec rs v1 v2 m).
- set (rs' := nextinstr (compare_longs v1 v2 rs m)).
- intros [A [B [C [D E]]]].
- destruct v1; destruct v2; simpl in H; inv H.
- unfold eval_testcond. rewrite A; rewrite B; rewrite C; rewrite D.
- simpl; rewrite int64_sub_overflow.
- destruct c; simpl.
- destruct (Int64.eq i i0); auto.
- destruct (Int64.eq i i0); auto.
- destruct (Int64.lt i i0); auto.
- rewrite Int64.not_lt. destruct (Int64.lt i i0); simpl; destruct (Int64.eq i i0); auto.
- rewrite (Int64.lt_not i i0). destruct (Int64.lt i i0); destruct (Int64.eq i i0); reflexivity.
- destruct (Int64.lt i i0); reflexivity.
-Qed.
-
-Lemma testcond_for_unsigned_comparison_64_correct:
- forall c v1 v2 rs m b,
- Val.cmplu_bool (Mem.valid_pointer m) c v1 v2 = Some b ->
- eval_testcond (testcond_for_unsigned_comparison c)
- (nextinstr (compare_longs v1 v2 rs m)) = Some b.
-Proof.
- intros. generalize (compare_longs_spec rs v1 v2 m).
- set (rs' := nextinstr (compare_longs v1 v2 rs m)).
- intros [A [B [C [D E]]]].
- unfold eval_testcond. rewrite A; rewrite B.
- destruct v1; destruct v2; simpl in H; inv H.
-- (* int int *)
- destruct c; simpl; auto.
- destruct (Int64.eq i i0); reflexivity.
- destruct (Int64.eq i i0); auto.
- destruct (Int64.ltu i i0); auto.
- rewrite Int64.not_ltu. destruct (Int64.ltu i i0); simpl; destruct (Int64.eq i i0); auto.
- rewrite (Int64.ltu_not i i0). destruct (Int64.ltu i i0); destruct (Int64.eq i i0); reflexivity.
- destruct (Int64.ltu i i0); reflexivity.
-- (* int ptr *)
- unfold Val.cmplu; simpl; destruct Archi.ptr64; try discriminate.
- destruct (Int64.eq i Int64.zero &&
- (Mem.valid_pointer m b0 (Ptrofs.unsigned i0) || Mem.valid_pointer m b0 (Ptrofs.unsigned i0 - 1))) eqn:?; try discriminate.
- destruct c; simpl in *; inv H1; auto.
-- (* ptr int *)
- unfold Val.cmplu; simpl; destruct Archi.ptr64; try discriminate.
- destruct (Int64.eq i0 Int64.zero &&
- (Mem.valid_pointer m b0 (Ptrofs.unsigned i) || Mem.valid_pointer m b0 (Ptrofs.unsigned i - 1))) eqn:?; try discriminate.
- destruct c; simpl in *; inv H1; auto.
-- (* ptr ptr *)
- unfold Val.cmplu; simpl; destruct Archi.ptr64; try discriminate.
- fold (Mem.weak_valid_pointer m b0 (Ptrofs.unsigned i)) in *.
- fold (Mem.weak_valid_pointer m b1 (Ptrofs.unsigned i0)) in *.
- destruct (eq_block b0 b1).
- destruct (Mem.weak_valid_pointer m b0 (Ptrofs.unsigned i) &&
- Mem.weak_valid_pointer m b1 (Ptrofs.unsigned i0)); inv H1.
- destruct c; simpl; auto.
- destruct (Ptrofs.eq i i0); auto.
- destruct (Ptrofs.eq i i0); auto.
- destruct (Ptrofs.ltu i i0); auto.
- rewrite Ptrofs.not_ltu. destruct (Ptrofs.ltu i i0); simpl; destruct (Ptrofs.eq i i0); auto.
- rewrite (Ptrofs.ltu_not i i0). destruct (Ptrofs.ltu i i0); destruct (Ptrofs.eq i i0); reflexivity.
- destruct (Ptrofs.ltu i i0); reflexivity.
- destruct (Mem.valid_pointer m b0 (Ptrofs.unsigned i) &&
- Mem.valid_pointer m b1 (Ptrofs.unsigned i0)); try discriminate.
- destruct c; simpl in *; inv H1; reflexivity.
-Qed.
-
-Lemma compare_floats_spec:
- forall rs n1 n2,
- let rs' := nextinstr (compare_floats (Vfloat n1) (Vfloat n2) rs) in
- rs'#ZF = Val.of_bool (negb (Float.cmp Cne n1 n2))
- /\ rs'#CF = Val.of_bool (negb (Float.cmp Cge n1 n2))
- /\ rs'#PF = Val.of_bool (negb (Float.cmp Ceq n1 n2 || Float.cmp Clt n1 n2 || Float.cmp Cgt n1 n2))
- /\ (forall r, data_preg r = true -> rs'#r = rs#r).
-Proof.
- intros. unfold rs'; unfold compare_floats.
- split. auto.
- split. auto.
- split. auto.
- intros. Simplifs.
-Qed.
-
-Lemma compare_floats32_spec:
- forall rs n1 n2,
- let rs' := nextinstr (compare_floats32 (Vsingle n1) (Vsingle n2) rs) in
- rs'#ZF = Val.of_bool (negb (Float32.cmp Cne n1 n2))
- /\ rs'#CF = Val.of_bool (negb (Float32.cmp Cge n1 n2))
- /\ rs'#PF = Val.of_bool (negb (Float32.cmp Ceq n1 n2 || Float32.cmp Clt n1 n2 || Float32.cmp Cgt n1 n2))
- /\ (forall r, data_preg r = true -> rs'#r = rs#r).
-Proof.
- intros. unfold rs'; unfold compare_floats32.
- split. auto.
- split. auto.
- split. auto.
- intros. Simplifs.
-Qed.
-
-Definition eval_extcond (xc: extcond) (rs: regset) : option bool :=
- match xc with
- | Cond_base c =>
- eval_testcond c rs
- | Cond_and c1 c2 =>
- match eval_testcond c1 rs, eval_testcond c2 rs with
- | Some b1, Some b2 => Some (b1 && b2)
- | _, _ => None
- end
- | Cond_or c1 c2 =>
- match eval_testcond c1 rs, eval_testcond c2 rs with
- | Some b1, Some b2 => Some (b1 || b2)
- | _, _ => None
- end
- end.
-
-Definition swap_floats {A: Type} (c: comparison) (n1 n2: A) : A :=
- match c with
- | Clt | Cle => n2
- | Ceq | Cne | Cgt | Cge => n1
- end.
-
-Lemma testcond_for_float_comparison_correct:
- forall c n1 n2 rs,
- eval_extcond (testcond_for_condition (Ccompf c))
- (nextinstr (compare_floats (Vfloat (swap_floats c n1 n2))
- (Vfloat (swap_floats c n2 n1)) rs)) =
- Some(Float.cmp c n1 n2).
-Proof.
- intros.
- generalize (compare_floats_spec rs (swap_floats c n1 n2) (swap_floats c n2 n1)).
- set (rs' := nextinstr (compare_floats (Vfloat (swap_floats c n1 n2))
- (Vfloat (swap_floats c n2 n1)) rs)).
- intros [A [B [C D]]].
- unfold eval_extcond, eval_testcond. rewrite A; rewrite B; rewrite C.
- destruct c; simpl.
-(* eq *)
- rewrite Float.cmp_ne_eq.
- caseEq (Float.cmp Ceq n1 n2); intros.
- auto.
- simpl. destruct (Float.cmp Clt n1 n2 || Float.cmp Cgt n1 n2); auto.
-(* ne *)
- rewrite Float.cmp_ne_eq.
- caseEq (Float.cmp Ceq n1 n2); intros.
- auto.
- simpl. destruct (Float.cmp Clt n1 n2 || Float.cmp Cgt n1 n2); auto.
-(* lt *)
- rewrite <- (Float.cmp_swap Cge n1 n2).
- rewrite <- (Float.cmp_swap Cne n1 n2).
- simpl.
- rewrite Float.cmp_ne_eq. rewrite Float.cmp_le_lt_eq.
- caseEq (Float.cmp Clt n1 n2); intros; simpl.
- caseEq (Float.cmp Ceq n1 n2); intros; simpl.
- elimtype False. eapply Float.cmp_lt_eq_false; eauto.
- auto.
- destruct (Float.cmp Ceq n1 n2); auto.
-(* le *)
- rewrite <- (Float.cmp_swap Cge n1 n2). simpl.
- destruct (Float.cmp Cle n1 n2); auto.
-(* gt *)
- rewrite Float.cmp_ne_eq. rewrite Float.cmp_ge_gt_eq.
- caseEq (Float.cmp Cgt n1 n2); intros; simpl.
- caseEq (Float.cmp Ceq n1 n2); intros; simpl.
- elimtype False. eapply Float.cmp_gt_eq_false; eauto.
- auto.
- destruct (Float.cmp Ceq n1 n2); auto.
-(* ge *)
- destruct (Float.cmp Cge n1 n2); auto.
-Qed.
-
-Lemma testcond_for_neg_float_comparison_correct:
- forall c n1 n2 rs,
- eval_extcond (testcond_for_condition (Cnotcompf c))
- (nextinstr (compare_floats (Vfloat (swap_floats c n1 n2))
- (Vfloat (swap_floats c n2 n1)) rs)) =
- Some(negb(Float.cmp c n1 n2)).
-Proof.
- intros.
- generalize (compare_floats_spec rs (swap_floats c n1 n2) (swap_floats c n2 n1)).
- set (rs' := nextinstr (compare_floats (Vfloat (swap_floats c n1 n2))
- (Vfloat (swap_floats c n2 n1)) rs)).
- intros [A [B [C D]]].
- unfold eval_extcond, eval_testcond. rewrite A; rewrite B; rewrite C.
- destruct c; simpl.
-(* eq *)
- rewrite Float.cmp_ne_eq.
- caseEq (Float.cmp Ceq n1 n2); intros.
- auto.
- simpl. destruct (Float.cmp Clt n1 n2 || Float.cmp Cgt n1 n2); auto.
-(* ne *)
- rewrite Float.cmp_ne_eq.
- caseEq (Float.cmp Ceq n1 n2); intros.
- auto.
- simpl. destruct (Float.cmp Clt n1 n2 || Float.cmp Cgt n1 n2); auto.
-(* lt *)
- rewrite <- (Float.cmp_swap Cge n1 n2).
- rewrite <- (Float.cmp_swap Cne n1 n2).
- simpl.
- rewrite Float.cmp_ne_eq. rewrite Float.cmp_le_lt_eq.
- caseEq (Float.cmp Clt n1 n2); intros; simpl.
- caseEq (Float.cmp Ceq n1 n2); intros; simpl.
- elimtype False. eapply Float.cmp_lt_eq_false; eauto.
- auto.
- destruct (Float.cmp Ceq n1 n2); auto.
-(* le *)
- rewrite <- (Float.cmp_swap Cge n1 n2). simpl.
- destruct (Float.cmp Cle n1 n2); auto.
-(* gt *)
- rewrite Float.cmp_ne_eq. rewrite Float.cmp_ge_gt_eq.
- caseEq (Float.cmp Cgt n1 n2); intros; simpl.
- caseEq (Float.cmp Ceq n1 n2); intros; simpl.
- elimtype False. eapply Float.cmp_gt_eq_false; eauto.
- auto.
- destruct (Float.cmp Ceq n1 n2); auto.
-(* ge *)
- destruct (Float.cmp Cge n1 n2); auto.
-Qed.
-
-Lemma testcond_for_float32_comparison_correct:
- forall c n1 n2 rs,
- eval_extcond (testcond_for_condition (Ccompfs c))
- (nextinstr (compare_floats32 (Vsingle (swap_floats c n1 n2))
- (Vsingle (swap_floats c n2 n1)) rs)) =
- Some(Float32.cmp c n1 n2).
-Proof.
- intros.
- generalize (compare_floats32_spec rs (swap_floats c n1 n2) (swap_floats c n2 n1)).
- set (rs' := nextinstr (compare_floats32 (Vsingle (swap_floats c n1 n2))
- (Vsingle (swap_floats c n2 n1)) rs)).
- intros [A [B [C D]]].
- unfold eval_extcond, eval_testcond. rewrite A; rewrite B; rewrite C.
- destruct c; simpl.
-(* eq *)
- rewrite Float32.cmp_ne_eq.
- caseEq (Float32.cmp Ceq n1 n2); intros.
- auto.
- simpl. destruct (Float32.cmp Clt n1 n2 || Float32.cmp Cgt n1 n2); auto.
-(* ne *)
- rewrite Float32.cmp_ne_eq.
- caseEq (Float32.cmp Ceq n1 n2); intros.
- auto.
- simpl. destruct (Float32.cmp Clt n1 n2 || Float32.cmp Cgt n1 n2); auto.
-(* lt *)
- rewrite <- (Float32.cmp_swap Cge n1 n2).
- rewrite <- (Float32.cmp_swap Cne n1 n2).
- simpl.
- rewrite Float32.cmp_ne_eq. rewrite Float32.cmp_le_lt_eq.
- caseEq (Float32.cmp Clt n1 n2); intros; simpl.
- caseEq (Float32.cmp Ceq n1 n2); intros; simpl.
- elimtype False. eapply Float32.cmp_lt_eq_false; eauto.
- auto.
- destruct (Float32.cmp Ceq n1 n2); auto.
-(* le *)
- rewrite <- (Float32.cmp_swap Cge n1 n2). simpl.
- destruct (Float32.cmp Cle n1 n2); auto.
-(* gt *)
- rewrite Float32.cmp_ne_eq. rewrite Float32.cmp_ge_gt_eq.
- caseEq (Float32.cmp Cgt n1 n2); intros; simpl.
- caseEq (Float32.cmp Ceq n1 n2); intros; simpl.
- elimtype False. eapply Float32.cmp_gt_eq_false; eauto.
- auto.
- destruct (Float32.cmp Ceq n1 n2); auto.
-(* ge *)
- destruct (Float32.cmp Cge n1 n2); auto.
-Qed.
-
-Lemma testcond_for_neg_float32_comparison_correct:
- forall c n1 n2 rs,
- eval_extcond (testcond_for_condition (Cnotcompfs c))
- (nextinstr (compare_floats32 (Vsingle (swap_floats c n1 n2))
- (Vsingle (swap_floats c n2 n1)) rs)) =
- Some(negb(Float32.cmp c n1 n2)).
-Proof.
- intros.
- generalize (compare_floats32_spec rs (swap_floats c n1 n2) (swap_floats c n2 n1)).
- set (rs' := nextinstr (compare_floats32 (Vsingle (swap_floats c n1 n2))
- (Vsingle (swap_floats c n2 n1)) rs)).
- intros [A [B [C D]]].
- unfold eval_extcond, eval_testcond. rewrite A; rewrite B; rewrite C.
- destruct c; simpl.
-(* eq *)
- rewrite Float32.cmp_ne_eq.
- caseEq (Float32.cmp Ceq n1 n2); intros.
- auto.
- simpl. destruct (Float32.cmp Clt n1 n2 || Float32.cmp Cgt n1 n2); auto.
-(* ne *)
- rewrite Float32.cmp_ne_eq.
- caseEq (Float32.cmp Ceq n1 n2); intros.
- auto.
- simpl. destruct (Float32.cmp Clt n1 n2 || Float32.cmp Cgt n1 n2); auto.
-(* lt *)
- rewrite <- (Float32.cmp_swap Cge n1 n2).
- rewrite <- (Float32.cmp_swap Cne n1 n2).
- simpl.
- rewrite Float32.cmp_ne_eq. rewrite Float32.cmp_le_lt_eq.
- caseEq (Float32.cmp Clt n1 n2); intros; simpl.
- caseEq (Float32.cmp Ceq n1 n2); intros; simpl.
- elimtype False. eapply Float32.cmp_lt_eq_false; eauto.
- auto.
- destruct (Float32.cmp Ceq n1 n2); auto.
-(* le *)
- rewrite <- (Float32.cmp_swap Cge n1 n2). simpl.
- destruct (Float32.cmp Cle n1 n2); auto.
-(* gt *)
- rewrite Float32.cmp_ne_eq. rewrite Float32.cmp_ge_gt_eq.
- caseEq (Float32.cmp Cgt n1 n2); intros; simpl.
- caseEq (Float32.cmp Ceq n1 n2); intros; simpl.
- elimtype False. eapply Float32.cmp_gt_eq_false; eauto.
- auto.
- destruct (Float32.cmp Ceq n1 n2); auto.
-(* ge *)
- destruct (Float32.cmp Cge n1 n2); auto.
-Qed.
-
-Remark swap_floats_commut:
- forall (A B: Type) c (f: A -> B) x y, swap_floats c (f x) (f y) = f (swap_floats c x y).
-Proof.
- intros. destruct c; auto.
-Qed.
-
-Remark compare_floats_inv:
- forall vx vy rs r,
- r <> CR ZF -> r <> CR CF -> r <> CR PF -> r <> CR SF -> r <> CR OF ->
- compare_floats vx vy rs r = rs r.
-Proof.
- intros.
- assert (DFL: undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF :: nil) rs r = rs r).
- simpl. Simplifs.
- unfold compare_floats; destruct vx; destruct vy; auto. Simplifs.
-Qed.
-
-Remark compare_floats32_inv:
- forall vx vy rs r,
- r <> CR ZF -> r <> CR CF -> r <> CR PF -> r <> CR SF -> r <> CR OF ->
- compare_floats32 vx vy rs r = rs r.
-Proof.
- intros.
- assert (DFL: undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF :: nil) rs r = rs r).
- simpl. Simplifs.
- unfold compare_floats32; destruct vx; destruct vy; auto. Simplifs.
-Qed.
-
-Lemma transl_cond_correct:
- forall cond args k c rs m,
- transl_cond cond args k = OK c ->
- exists rs',
- exec_straight ge fn c rs m k rs' m
- /\ match eval_condition cond (map rs (map preg_of args)) m with
- | None => True
- | Some b => eval_extcond (testcond_for_condition cond) rs' = Some b
- end
- /\ forall r, data_preg r = true -> rs'#r = rs r.
-Proof.
- unfold transl_cond; intros.
- destruct cond; repeat (destruct args; try discriminate); monadInv H.
-- (* comp *)
- simpl. rewrite (ireg_of_eq _ _ EQ). rewrite (ireg_of_eq _ _ EQ1).
- econstructor. split. apply exec_straight_one. simpl. eauto. auto.
- split. destruct (Val.cmp_bool c0 (rs x) (rs x0)) eqn:?; auto.
- eapply testcond_for_signed_comparison_32_correct; eauto.
- intros. unfold compare_ints. Simplifs.
-- (* compu *)
- simpl. rewrite (ireg_of_eq _ _ EQ). rewrite (ireg_of_eq _ _ EQ1).
- econstructor. split. apply exec_straight_one. simpl. eauto. auto.
- split. destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs x) (rs x0)) eqn:?; auto.
- eapply testcond_for_unsigned_comparison_32_correct; eauto.
- intros. unfold compare_ints. Simplifs.
-- (* compimm *)
- simpl. rewrite (ireg_of_eq _ _ EQ). destruct (Int.eq_dec n Int.zero).
- econstructor; split. apply exec_straight_one. simpl; eauto. auto.
- split. destruct (rs x); simpl; auto. subst. rewrite Int.and_idem.
- eapply testcond_for_signed_comparison_32_correct; eauto.
- intros. unfold compare_ints. Simplifs.
- econstructor; split. apply exec_straight_one. simpl; eauto. auto.
- split. destruct (Val.cmp_bool c0 (rs x) (Vint n)) eqn:?; auto.
- eapply testcond_for_signed_comparison_32_correct; eauto.
- intros. unfold compare_ints. Simplifs.
-- (* compuimm *)
- simpl. rewrite (ireg_of_eq _ _ EQ).
- econstructor. split. apply exec_straight_one. simpl. eauto. auto.
- split. destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs x) (Vint n)) eqn:?; auto.
- eapply testcond_for_unsigned_comparison_32_correct; eauto.
- intros. unfold compare_ints. Simplifs.
-- (* compl *)
- simpl. rewrite (ireg_of_eq _ _ EQ). rewrite (ireg_of_eq _ _ EQ1).
- econstructor. split. apply exec_straight_one. simpl. eauto. auto.
- split. destruct (Val.cmpl_bool c0 (rs x) (rs x0)) eqn:?; auto.
- eapply testcond_for_signed_comparison_64_correct; eauto.
- intros. unfold compare_longs. Simplifs.
-- (* complu *)
- simpl. rewrite (ireg_of_eq _ _ EQ). rewrite (ireg_of_eq _ _ EQ1).
- econstructor. split. apply exec_straight_one. simpl. eauto. auto.
- split. destruct (Val.cmplu_bool (Mem.valid_pointer m) c0 (rs x) (rs x0)) eqn:?; auto.
- eapply testcond_for_unsigned_comparison_64_correct; eauto.
- intros. unfold compare_longs. Simplifs.
-- (* compimm *)
- simpl. rewrite (ireg_of_eq _ _ EQ). destruct (Int64.eq_dec n Int64.zero).
- econstructor; split. apply exec_straight_one. simpl; eauto. auto.
- split. destruct (rs x); simpl; auto. subst. rewrite Int64.and_idem.
- eapply testcond_for_signed_comparison_64_correct; eauto.
- intros. unfold compare_longs. Simplifs.
- econstructor; split. apply exec_straight_one. simpl; eauto. auto.
- split. destruct (Val.cmpl_bool c0 (rs x) (Vlong n)) eqn:?; auto.
- eapply testcond_for_signed_comparison_64_correct; eauto.
- intros. unfold compare_longs. Simplifs.
-- (* compuimm *)
- simpl. rewrite (ireg_of_eq _ _ EQ).
- econstructor. split. apply exec_straight_one. simpl. eauto. auto.
- split. destruct (Val.cmplu_bool (Mem.valid_pointer m) c0 (rs x) (Vlong n)) eqn:?; auto.
- eapply testcond_for_unsigned_comparison_64_correct; eauto.
- intros. unfold compare_longs. Simplifs.
-- (* compf *)
- simpl. rewrite (freg_of_eq _ _ EQ). rewrite (freg_of_eq _ _ EQ1).
- exists (nextinstr (compare_floats (swap_floats c0 (rs x) (rs x0)) (swap_floats c0 (rs x0) (rs x)) rs)).
- split. apply exec_straight_one.
- destruct c0; simpl; auto.
- unfold nextinstr. rewrite Pregmap.gss. rewrite compare_floats_inv; auto with asmgen.
- split. destruct (rs x); destruct (rs x0); simpl; auto.
- repeat rewrite swap_floats_commut. apply testcond_for_float_comparison_correct.
- intros. Simplifs. apply compare_floats_inv; auto with asmgen.
-- (* notcompf *)
- simpl. rewrite (freg_of_eq _ _ EQ). rewrite (freg_of_eq _ _ EQ1).
- exists (nextinstr (compare_floats (swap_floats c0 (rs x) (rs x0)) (swap_floats c0 (rs x0) (rs x)) rs)).
- split. apply exec_straight_one.
- destruct c0; simpl; auto.
- unfold nextinstr. rewrite Pregmap.gss. rewrite compare_floats_inv; auto with asmgen.
- split. destruct (rs x); destruct (rs x0); simpl; auto.
- repeat rewrite swap_floats_commut. apply testcond_for_neg_float_comparison_correct.
- intros. Simplifs. apply compare_floats_inv; auto with asmgen.
-- (* compfs *)
- simpl. rewrite (freg_of_eq _ _ EQ). rewrite (freg_of_eq _ _ EQ1).
- exists (nextinstr (compare_floats32 (swap_floats c0 (rs x) (rs x0)) (swap_floats c0 (rs x0) (rs x)) rs)).
- split. apply exec_straight_one.
- destruct c0; simpl; auto.
- unfold nextinstr. rewrite Pregmap.gss. rewrite compare_floats32_inv; auto with asmgen.
- split. destruct (rs x); destruct (rs x0); simpl; auto.
- repeat rewrite swap_floats_commut. apply testcond_for_float32_comparison_correct.
- intros. Simplifs. apply compare_floats32_inv; auto with asmgen.
-- (* notcompfs *)
- simpl. rewrite (freg_of_eq _ _ EQ). rewrite (freg_of_eq _ _ EQ1).
- exists (nextinstr (compare_floats32 (swap_floats c0 (rs x) (rs x0)) (swap_floats c0 (rs x0) (rs x)) rs)).
- split. apply exec_straight_one.
- destruct c0; simpl; auto.
- unfold nextinstr. rewrite Pregmap.gss. rewrite compare_floats32_inv; auto with asmgen.
- split. destruct (rs x); destruct (rs x0); simpl; auto.
- repeat rewrite swap_floats_commut. apply testcond_for_neg_float32_comparison_correct.
- intros. Simplifs. apply compare_floats32_inv; auto with asmgen.
-- (* maskzero *)
- simpl. rewrite (ireg_of_eq _ _ EQ).
- econstructor. split. apply exec_straight_one. simpl; eauto. auto.
- split. destruct (rs x); simpl; auto.
- generalize (compare_ints_spec rs (Vint (Int.and i n)) Vzero m).
- intros [A B]. rewrite A. unfold Val.cmpu; simpl. destruct (Int.eq (Int.and i n) Int.zero); auto.
- intros. unfold compare_ints. Simplifs.
-- (* masknotzero *)
- simpl. rewrite (ireg_of_eq _ _ EQ).
- econstructor. split. apply exec_straight_one. simpl; eauto. auto.
- split. destruct (rs x); simpl; auto.
- generalize (compare_ints_spec rs (Vint (Int.and i n)) Vzero m).
- intros [A B]. rewrite A. unfold Val.cmpu; simpl. destruct (Int.eq (Int.and i n) Int.zero); auto.
- intros. unfold compare_ints. Simplifs.
-Qed.
-
-Remark eval_testcond_nextinstr:
- forall c rs, eval_testcond c (nextinstr rs) = eval_testcond c rs.
-Proof.
- intros. unfold eval_testcond. repeat rewrite nextinstr_inv; auto with asmgen.
-Qed.
-
-Remark eval_testcond_set_ireg:
- forall c rs r v, eval_testcond c (rs#(IR r) <- v) = eval_testcond c rs.
-Proof.
- intros. unfold eval_testcond. repeat rewrite Pregmap.gso; auto with asmgen.
-Qed.
-
-Lemma mk_setcc_base_correct:
- forall cond rd k rs1 m,
- exists rs2,
- exec_straight ge fn (mk_setcc_base cond rd k) rs1 m k rs2 m
- /\ rs2#rd = Val.of_optbool(eval_extcond cond rs1)
- /\ forall r, data_preg r = true -> r <> RAX /\ r <> RCX -> r <> rd -> rs2#r = rs1#r.
-Proof.
- intros. destruct cond; simpl in *.
-- (* base *)
- econstructor; split.
- apply exec_straight_one. simpl; eauto. auto.
- split. Simplifs. intros; Simplifs.
-- (* or *)
- assert (Val.of_optbool
- match eval_testcond c1 rs1 with
- | Some b1 =>
- match eval_testcond c2 rs1 with
- | Some b2 => Some (b1 || b2)
- | None => None
- end
- | None => None
- end =
- Val.or (Val.of_optbool (eval_testcond c1 rs1)) (Val.of_optbool (eval_testcond c2 rs1))).
- destruct (eval_testcond c1 rs1). destruct (eval_testcond c2 rs1).
- destruct b; destruct b0; auto.
- destruct b; auto.
- auto.
- rewrite H; clear H.
- destruct (ireg_eq rd RAX).
- subst rd. econstructor; split.
- eapply exec_straight_three.
- simpl; eauto.
- simpl. rewrite eval_testcond_nextinstr. repeat rewrite eval_testcond_set_ireg. eauto.
- simpl; eauto.
- auto. auto. auto.
- intuition Simplifs.
- econstructor; split.
- eapply exec_straight_three.
- simpl; eauto.
- simpl. rewrite eval_testcond_nextinstr. repeat rewrite eval_testcond_set_ireg. eauto.
- simpl. eauto.
- auto. auto. auto.
- split. Simplifs. rewrite Val.or_commut. decEq; Simplifs.
- intros. destruct H0; Simplifs.
-- (* and *)
- assert (Val.of_optbool
- match eval_testcond c1 rs1 with
- | Some b1 =>
- match eval_testcond c2 rs1 with
- | Some b2 => Some (b1 && b2)
- | None => None
- end
- | None => None
- end =
- Val.and (Val.of_optbool (eval_testcond c1 rs1)) (Val.of_optbool (eval_testcond c2 rs1))).
- {
- destruct (eval_testcond c1 rs1). destruct (eval_testcond c2 rs1).
- destruct b; destruct b0; auto.
- destruct b; auto.
- auto.
- }
- rewrite H; clear H.
- destruct (ireg_eq rd RAX).
- subst rd. econstructor; split.
- eapply exec_straight_three.
- simpl; eauto.
- simpl. rewrite eval_testcond_nextinstr. repeat rewrite eval_testcond_set_ireg. eauto.
- simpl; eauto.
- auto. auto. auto.
- intuition Simplifs.
- econstructor; split.
- eapply exec_straight_three.
- simpl; eauto.
- simpl. rewrite eval_testcond_nextinstr. repeat rewrite eval_testcond_set_ireg. eauto.
- simpl. eauto.
- auto. auto. auto.
- split. Simplifs. rewrite Val.and_commut. decEq; Simplifs.
- intros. destruct H0; Simplifs.
-Qed.
-
-Lemma mk_setcc_correct:
- forall cond rd k rs1 m,
- exists rs2,
- exec_straight ge fn (mk_setcc cond rd k) rs1 m k rs2 m
- /\ rs2#rd = Val.of_optbool(eval_extcond cond rs1)
- /\ forall r, data_preg r = true -> r <> RAX /\ r <> RCX -> r <> rd -> rs2#r = rs1#r.
-Proof.
- intros. unfold mk_setcc. destruct (Archi.ptr64 || low_ireg rd).
-- apply mk_setcc_base_correct.
-- exploit mk_setcc_base_correct. intros [rs2 [A [B C]]].
- econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one.
- simpl. eauto. simpl. auto.
- intuition Simplifs.
-Qed.
-
-(** Translation of arithmetic operations. *)
-
-Ltac ArgsInv :=
- match goal with
- | [ H: Error _ = OK _ |- _ ] => discriminate
- | [ H: match ?args with nil => _ | _ :: _ => _ end = OK _ |- _ ] => destruct args; ArgsInv
- | [ H: bind _ _ = OK _ |- _ ] => monadInv H; ArgsInv
- | [ H: match _ with left _ => _ | right _ => assertion_failed end = OK _ |- _ ] => monadInv H; ArgsInv
- | [ H: match _ with true => _ | false => assertion_failed end = OK _ |- _ ] => monadInv H; ArgsInv
- | [ H: ireg_of _ = OK _ |- _ ] => simpl in *; rewrite (ireg_of_eq _ _ H) in *; clear H; ArgsInv
- | [ H: freg_of _ = OK _ |- _ ] => simpl in *; rewrite (freg_of_eq _ _ H) in *; clear H; ArgsInv
- | _ => idtac
- end.
-
-Ltac TranslOp :=
- econstructor; split;
- [ apply exec_straight_one; [ simpl; eauto | auto ]
- | split; [ Simplifs | intros; Simplifs ]].
-
-Lemma transl_op_correct:
- forall op args res k c (rs: regset) m v,
- transl_op op args res k = OK c ->
- eval_operation ge (rs#RSP) op (map rs (map preg_of args)) m = Some v ->
- exists rs',
- exec_straight ge fn c rs m k rs' m
- /\ Val.lessdef v rs'#(preg_of res)
- /\ forall r, data_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs' r = rs r.
-Proof.
-Transparent destroyed_by_op.
- intros until v; intros TR EV.
- assert (SAME:
- (exists rs',
- exec_straight ge fn c rs m k rs' m
- /\ rs'#(preg_of res) = v
- /\ forall r, data_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs' r = rs r) ->
- exists rs',
- exec_straight ge fn c rs m k rs' m
- /\ Val.lessdef v rs'#(preg_of res)
- /\ forall r, data_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs' r = rs r).
- {
- intros [rs' [A [B C]]]. subst v. exists rs'; auto.
- }
-
- destruct op; simpl in TR; ArgsInv; simpl in EV; try (inv EV); try (apply SAME; TranslOp; fail).
-(* move *)
- exploit mk_mov_correct; eauto. intros [rs2 [A [B C]]].
- apply SAME. exists rs2. eauto.
-(* intconst *)
- apply SAME. destruct (Int.eq_dec n Int.zero). subst n. TranslOp. TranslOp.
-(* longconst *)
- apply SAME. destruct (Int64.eq_dec n Int64.zero). subst n. TranslOp. TranslOp.
-(* floatconst *)
- apply SAME. destruct (Float.eq_dec n Float.zero). subst n. TranslOp. TranslOp.
-(* singleconst *)
- apply SAME. destruct (Float32.eq_dec n Float32.zero). subst n. TranslOp. TranslOp.
-(* cast8signed *)
- apply SAME. eapply mk_intconv_correct; eauto.
-(* cast8unsigned *)
- apply SAME. eapply mk_intconv_correct; eauto.
-(* mulhs *)
- apply SAME. TranslOp. destruct H1. Simplifs.
-(* mulhu *)
- apply SAME. TranslOp. destruct H1. Simplifs.
-(* div *)
- apply SAME.
- exploit (divs_mods_exists (rs RAX) (rs RCX)). left; congruence.
- intros (nh & nl & d & q & r & A & B & C & D & E & F).
- set (rs1 := nextinstr_nf (rs#RDX <- (Vint nh))).
- econstructor; split.
- eapply exec_straight_two with (rs2 := rs1). simpl. rewrite A. reflexivity.
- simpl. change (rs1 RAX) with (rs RAX); rewrite B.
- change (rs1 RCX) with (rs RCX); rewrite C.
- rewrite D. reflexivity. auto. auto.
- split. change (Vint q = v). congruence.
- simpl; intros. destruct H2. unfold rs1; Simplifs.
-(* divu *)
- apply SAME.
- exploit (divu_modu_exists (rs RAX) (rs RCX)). left; congruence.
- intros (n & d & q & r & B & C & D & E & F).
- set (rs1 := nextinstr_nf (rs#RDX <- Vzero)).
- econstructor; split.
- eapply exec_straight_two with (rs2 := rs1). reflexivity.
- simpl. change (rs1 RAX) with (rs RAX); rewrite B.
- change (rs1 RCX) with (rs RCX); rewrite C.
- rewrite D. reflexivity. auto. auto.
- split. change (Vint q = v). congruence.
- simpl; intros. destruct H2. unfold rs1; Simplifs.
-(* mod *)
- apply SAME.
- exploit (divs_mods_exists (rs RAX) (rs RCX)). right; congruence.
- intros (nh & nl & d & q & r & A & B & C & D & E & F).
- set (rs1 := nextinstr_nf (rs#RDX <- (Vint nh))).
- econstructor; split.
- eapply exec_straight_two with (rs2 := rs1). simpl. rewrite A. reflexivity.
- simpl. change (rs1 RAX) with (rs RAX); rewrite B.
- change (rs1 RCX) with (rs RCX); rewrite C.
- rewrite D. reflexivity. auto. auto.
- split. change (Vint r = v). congruence.
- simpl; intros. destruct H2. unfold rs1; Simplifs.
-(* modu *)
- apply SAME.
- exploit (divu_modu_exists (rs RAX) (rs RCX)). right; congruence.
- intros (n & d & q & r & B & C & D & E & F).
- set (rs1 := nextinstr_nf (rs#RDX <- Vzero)).
- econstructor; split.
- eapply exec_straight_two with (rs2 := rs1). reflexivity.
- simpl. change (rs1 RAX) with (rs RAX); rewrite B.
- change (rs1 RCX) with (rs RCX); rewrite C.
- rewrite D. reflexivity. auto. auto.
- split. change (Vint r = v). congruence.
- simpl; intros. destruct H2. unfold rs1; Simplifs.
-(* shrximm *)
- apply SAME. eapply mk_shrximm_correct; eauto.
-(* lea *)
- exploit transl_addressing_mode_32_correct; eauto. intros EA.
- TranslOp. rewrite nextinstr_inv; auto with asmgen. rewrite Pregmap.gss. rewrite normalize_addrmode_32_correct; auto.
-(* mullhs *)
- apply SAME. TranslOp. destruct H1. Simplifs.
-(* mullhu *)
- apply SAME. TranslOp. destruct H1. Simplifs.
-(* divl *)
- apply SAME.
- exploit (divls_modls_exists (rs RAX) (rs RCX)). left; congruence.
- intros (nh & nl & d & q & r & A & B & C & D & E & F).
- set (rs1 := nextinstr_nf (rs#RDX <- (Vlong nh))).
- econstructor; split.
- eapply exec_straight_two with (rs2 := rs1). simpl. rewrite A. reflexivity.
- simpl. change (rs1 RAX) with (rs RAX); rewrite B.
- change (rs1 RCX) with (rs RCX); rewrite C.
- rewrite D. reflexivity. auto. auto.
- split. change (Vlong q = v). congruence.
- simpl; intros. destruct H2. unfold rs1; Simplifs.
-(* divlu *)
- apply SAME.
- exploit (divlu_modlu_exists (rs RAX) (rs RCX)). left; congruence.
- intros (n & d & q & r & B & C & D & E & F).
- set (rs1 := nextinstr_nf (rs#RDX <- (Vlong Int64.zero))).
- econstructor; split.
- eapply exec_straight_two with (rs2 := rs1). reflexivity.
- simpl. change (rs1 RAX) with (rs RAX); rewrite B.
- change (rs1 RCX) with (rs RCX); rewrite C.
- rewrite D. reflexivity. auto. auto.
- split. change (Vlong q = v). congruence.
- simpl; intros. destruct H2. unfold rs1; Simplifs.
-(* modl *)
- apply SAME.
- exploit (divls_modls_exists (rs RAX) (rs RCX)). right; congruence.
- intros (nh & nl & d & q & r & A & B & C & D & E & F).
- set (rs1 := nextinstr_nf (rs#RDX <- (Vlong nh))).
- econstructor; split.
- eapply exec_straight_two with (rs2 := rs1). simpl. rewrite A. reflexivity.
- simpl. change (rs1 RAX) with (rs RAX); rewrite B.
- change (rs1 RCX) with (rs RCX); rewrite C.
- rewrite D. reflexivity. auto. auto.
- split. change (Vlong r = v). congruence.
- simpl; intros. destruct H2. unfold rs1; Simplifs.
-(* modlu *)
- apply SAME.
- exploit (divlu_modlu_exists (rs RAX) (rs RCX)). right; congruence.
- intros (n & d & q & r & B & C & D & E & F).
- set (rs1 := nextinstr_nf (rs#RDX <- (Vlong Int64.zero))).
- econstructor; split.
- eapply exec_straight_two with (rs2 := rs1). reflexivity.
- simpl. change (rs1 RAX) with (rs RAX); rewrite B.
- change (rs1 RCX) with (rs RCX); rewrite C.
- rewrite D. reflexivity. auto. auto.
- split. change (Vlong r = v). congruence.
- simpl; intros. destruct H2. unfold rs1; Simplifs.
-(* shrxlimm *)
- apply SAME. eapply mk_shrxlimm_correct; eauto.
-(* leal *)
- exploit transl_addressing_mode_64_correct; eauto. intros EA.
- generalize (normalize_addrmode_64_correct x rs). destruct (normalize_addrmode_64 x) as [am' [delta|]]; intros EV.
- econstructor; split. eapply exec_straight_two.
- simpl. reflexivity. simpl. reflexivity. auto. auto.
- split. rewrite nextinstr_nf_inv by auto. rewrite Pregmap.gss. rewrite nextinstr_inv by auto with asmgen.
- rewrite Pregmap.gss. rewrite <- EV; auto.
- intros; Simplifs.
- TranslOp. rewrite nextinstr_inv; auto with asmgen. rewrite Pregmap.gss; auto. rewrite <- EV; auto.
-(* intoffloat *)
- apply SAME. TranslOp. rewrite H0; auto.
-(* floatofint *)
- apply SAME. TranslOp. rewrite H0; auto.
-(* intofsingle *)
- apply SAME. TranslOp. rewrite H0; auto.
-(* singleofint *)
- apply SAME. TranslOp. rewrite H0; auto.
-(* longoffloat *)
- apply SAME. TranslOp. rewrite H0; auto.
-(* floatoflong *)
- apply SAME. TranslOp. rewrite H0; auto.
-(* longofsingle *)
- apply SAME. TranslOp. rewrite H0; auto.
-(* singleoflong *)
- apply SAME. TranslOp. rewrite H0; auto.
-(* condition *)
- exploit transl_cond_correct; eauto. intros [rs2 [P [Q R]]].
- exploit mk_setcc_correct; eauto. intros [rs3 [S [T U]]].
- exists rs3.
- split. eapply exec_straight_trans. eexact P. eexact S.
- split. rewrite T. destruct (eval_condition cond rs ## (preg_of ## args) m).
- rewrite Q. auto.
- simpl; auto.
- intros. transitivity (rs2 r); auto.
-Qed.
-
-(** Translation of memory loads. *)
-
-Lemma transl_load_correct:
- forall chunk addr args dest k c (rs: regset) m a v,
- transl_load chunk addr args dest k = OK c ->
- eval_addressing ge (rs#RSP) addr (map rs (map preg_of args)) = Some a ->
- Mem.loadv chunk m a = Some v ->
- exists rs',
- exec_straight ge fn c rs m k rs' m
- /\ rs'#(preg_of dest) = v
- /\ forall r, data_preg r = true -> r <> preg_of dest -> rs'#r = rs#r.
-Proof.
- unfold transl_load; intros. monadInv H.
- exploit transl_addressing_mode_correct; eauto. intro EA.
- assert (EA': eval_addrmode ge x rs = a). destruct a; simpl in H1; try discriminate; inv EA; auto.
- set (rs2 := nextinstr_nf (rs#(preg_of dest) <- v)).
- assert (exec_load ge chunk m x rs (preg_of dest) = Next rs2 m).
- unfold exec_load. rewrite EA'. rewrite H1. auto.
- assert (rs2 PC = Val.offset_ptr (rs PC) Ptrofs.one).
- transitivity (Val.offset_ptr ((rs#(preg_of dest) <- v) PC) Ptrofs.one).
- auto. decEq. apply Pregmap.gso; auto with asmgen.
- exists rs2. split.
- destruct chunk; ArgsInv; apply exec_straight_one; auto.
- split. unfold rs2. rewrite nextinstr_nf_inv1. Simplifs. apply preg_of_data.
- intros. unfold rs2. Simplifs.
-Qed.
-
-Lemma transl_store_correct:
- forall chunk addr args src k c (rs: regset) m a m',
- transl_store chunk addr args src k = OK c ->
- eval_addressing ge (rs#RSP) addr (map rs (map preg_of args)) = Some a ->
- Mem.storev chunk m a (rs (preg_of src)) = Some m' ->
- exists rs',
- exec_straight ge fn c rs m k rs' m'
- /\ forall r, data_preg r = true -> preg_notin r (destroyed_by_store chunk addr) -> rs'#r = rs#r.
-Proof.
- unfold transl_store; intros. monadInv H.
- exploit transl_addressing_mode_correct; eauto. intro EA.
- assert (EA': eval_addrmode ge x rs = a). destruct a; simpl in H1; try discriminate; inv EA; auto.
- rewrite <- EA' in H1. destruct chunk; ArgsInv.
-(* int8signed *)
- eapply mk_storebyte_correct; eauto.
- destruct (eval_addrmode ge x rs); simpl; auto. rewrite <- Mem.store_signed_unsigned_8; auto.
-(* int8unsigned *)
- eapply mk_storebyte_correct; eauto.
-(* int16signed *)
- econstructor; split.
- apply exec_straight_one. simpl. unfold exec_store.
- replace (Mem.storev Mint16unsigned m (eval_addrmode ge x rs) (rs x0))
- with (Mem.storev Mint16signed m (eval_addrmode ge x rs) (rs x0)).
- rewrite H1. eauto.
- destruct (eval_addrmode ge x rs); simpl; auto. rewrite Mem.store_signed_unsigned_16; auto.
- auto.
- intros. Simplifs.
-(* int16unsigned *)
- econstructor; split.
- apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto.
- intros. Simplifs.
-(* int32 *)
- econstructor; split.
- apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto.
- intros. Simplifs.
-(* int64 *)
- econstructor; split.
- apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto.
- intros. Simplifs.
-(* float32 *)
- econstructor; split.
- apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto.
- intros. Transparent destroyed_by_store. simpl in H2. simpl. Simplifs.
-(* float64 *)
- econstructor; split.
- apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto.
- intros. Simplifs.
-Qed.
-
-End CONSTRUCTORS.
diff --git a/ia32/CBuiltins.ml b/ia32/CBuiltins.ml
deleted file mode 100644
index 09303223..00000000
--- a/ia32/CBuiltins.ml
+++ /dev/null
@@ -1,94 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-(* Processor-dependent builtin C functions *)
-
-open C
-
-let (va_list_type, va_list_scalar, size_va_list) =
- if Archi.ptr64 then
- (* Actually a struct passed by reference; equivalent to 3 64-bit words *)
- (TArray(TInt(IULong, []), Some 3L, []), false, 3*8)
- else
- (* Just a pointer *)
- (TPtr(TVoid [], []), true, 4)
-
-let builtins = {
- Builtins.typedefs = [
- "__builtin_va_list", va_list_type;
- ];
- Builtins.functions = [
- (* Integer arithmetic *)
- "__builtin_bswap",
- (TInt(IUInt, []), [TInt(IUInt, [])], false);
- "__builtin_bswap64",
- (TInt(IULongLong, []), [TInt(IULongLong, [])], false);
- "__builtin_bswap32",
- (TInt(IUInt, []), [TInt(IUInt, [])], false);
- "__builtin_bswap16",
- (TInt(IUShort, []), [TInt(IUShort, [])], false);
- "__builtin_clz",
- (TInt(IInt, []), [TInt(IUInt, [])], false);
- "__builtin_clzl",
- (TInt(IInt, []), [TInt(IULong, [])], false);
- "__builtin_clzll",
- (TInt(IInt, []), [TInt(IULongLong, [])], false);
- "__builtin_ctz",
- (TInt(IInt, []), [TInt(IUInt, [])], false);
- "__builtin_ctzl",
- (TInt(IInt, []), [TInt(IULong, [])], false);
- "__builtin_ctzll",
- (TInt(IInt, []), [TInt(IULongLong, [])], false);
- (* Float arithmetic *)
- "__builtin_fsqrt",
- (TFloat(FDouble, []), [TFloat(FDouble, [])], false);
- "__builtin_fmax",
- (TFloat(FDouble, []), [TFloat(FDouble, []); TFloat(FDouble, [])], false);
- "__builtin_fmin",
- (TFloat(FDouble, []), [TFloat(FDouble, []); TFloat(FDouble, [])], false);
- "__builtin_fmadd",
- (TFloat(FDouble, []),
- [TFloat(FDouble, []); TFloat(FDouble, []); TFloat(FDouble, [])],
- false);
- "__builtin_fmsub",
- (TFloat(FDouble, []),
- [TFloat(FDouble, []); TFloat(FDouble, []); TFloat(FDouble, [])],
- false);
- "__builtin_fnmadd",
- (TFloat(FDouble, []),
- [TFloat(FDouble, []); TFloat(FDouble, []); TFloat(FDouble, [])],
- false);
- "__builtin_fnmsub",
- (TFloat(FDouble, []),
- [TFloat(FDouble, []); TFloat(FDouble, []); TFloat(FDouble, [])],
- false);
- (* Memory accesses *)
- "__builtin_read16_reversed",
- (TInt(IUShort, []), [TPtr(TInt(IUShort, [AConst]), [])], false);
- "__builtin_read32_reversed",
- (TInt(IUInt, []), [TPtr(TInt(IUInt, [AConst]), [])], false);
- "__builtin_write16_reversed",
- (TVoid [], [TPtr(TInt(IUShort, []), []); TInt(IUShort, [])], false);
- "__builtin_write32_reversed",
- (TVoid [], [TPtr(TInt(IUInt, []), []); TInt(IUInt, [])], false);
- (* no operation *)
- "__builtin_nop",
- (TVoid [], [], false);
- ]
-}
-
-(* Expand memory references inside extended asm statements. Used in C2C. *)
-
-let asm_mem_argument arg = Printf.sprintf "0(%s)" arg
diff --git a/ia32/CombineOp.v b/ia32/CombineOp.v
deleted file mode 100644
index 34c1c9cc..00000000
--- a/ia32/CombineOp.v
+++ /dev/null
@@ -1,150 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Recognition of combined operations, addressing modes and conditions
- during the [CSE] phase. *)
-
-Require Import Coqlib.
-Require Import AST Integers.
-Require Import Op CSEdomain.
-
-Definition valnum := positive.
-
-Section COMBINE.
-
-Variable get: valnum -> option rhs.
-
-Function combine_compimm_ne_0 (x: valnum) : option(condition * list valnum) :=
- match get x with
- | Some(Op (Ocmp c) ys) => Some (c, ys)
- | Some(Op (Oandimm n) ys) => Some (Cmasknotzero n, ys)
- | _ => None
- end.
-
-Function combine_compimm_eq_0 (x: valnum) : option(condition * list valnum) :=
- match get x with
- | Some(Op (Ocmp c) ys) => Some (negate_condition c, ys)
- | Some(Op (Oandimm n) ys) => Some (Cmaskzero n, ys)
- | _ => None
- end.
-
-Function combine_compimm_eq_1 (x: valnum) : option(condition * list valnum) :=
- match get x with
- | Some(Op (Ocmp c) ys) => Some (c, ys)
- | _ => None
- end.
-
-Function combine_compimm_ne_1 (x: valnum) : option(condition * list valnum) :=
- match get x with
- | Some(Op (Ocmp c) ys) => Some (negate_condition c, ys)
- | _ => None
- end.
-
-Function combine_cond (cond: condition) (args: list valnum) : option(condition * list valnum) :=
- match cond, args with
- | Ccompimm Cne n, x::nil =>
- if Int.eq_dec n Int.zero then combine_compimm_ne_0 x
- else if Int.eq_dec n Int.one then combine_compimm_ne_1 x
- else None
- | Ccompimm Ceq n, x::nil =>
- if Int.eq_dec n Int.zero then combine_compimm_eq_0 x
- else if Int.eq_dec n Int.one then combine_compimm_eq_1 x
- else None
- | Ccompuimm Cne n, x::nil =>
- if Int.eq_dec n Int.zero then combine_compimm_ne_0 x
- else if Int.eq_dec n Int.one then combine_compimm_ne_1 x
- else None
- | Ccompuimm Ceq n, x::nil =>
- if Int.eq_dec n Int.zero then combine_compimm_eq_0 x
- else if Int.eq_dec n Int.one then combine_compimm_eq_1 x
- else None
- | _, _ => None
- end.
-
-Function combine_addr_32 (addr: addressing) (args: list valnum) : option(addressing * list valnum) :=
- match addr, args with
- | Aindexed n, x::nil =>
- match get x with
- | Some(Op (Olea a) ys) =>
- match offset_addressing a n with Some a' => Some (a', ys) | None => None end
- | _ => None
- end
- | _, _ => None
- end.
-
-Function combine_addr_64 (addr: addressing) (args: list valnum) : option(addressing * list valnum) :=
- match addr, args with
- | Aindexed n, x::nil =>
- match get x with
- | Some(Op (Oleal a) ys) =>
- match offset_addressing a n with Some a' => Some (a', ys) | None => None end
- | _ => None
- end
- | _, _ => None
- end.
-
-Definition combine_addr (addr: addressing) (args: list valnum) : option(addressing * list valnum) :=
- if Archi.ptr64 then combine_addr_64 addr args else combine_addr_32 addr args.
-
-Function combine_op (op: operation) (args: list valnum) : option(operation * list valnum) :=
- match op, args with
- | Olea addr, _ =>
- match combine_addr_32 addr args with
- | Some(addr', args') => Some(Olea addr', args')
- | None => None
- end
- | Oleal addr, _ =>
- match combine_addr_64 addr args with
- | Some(addr', args') => Some(Oleal addr', args')
- | None => None
- end
- | Oandimm n, x :: nil =>
- match get x with
- | Some(Op (Oandimm m) ys) => Some(Oandimm (Int.and m n), ys)
- | _ => None
- end
- | Oorimm n, x :: nil =>
- match get x with
- | Some(Op (Oorimm m) ys) => Some(Oorimm (Int.or m n), ys)
- | _ => None
- end
- | Oxorimm n, x :: nil =>
- match get x with
- | Some(Op (Oxorimm m) ys) => Some(Oxorimm (Int.xor m n), ys)
- | _ => None
- end
- | Oandlimm n, x :: nil =>
- match get x with
- | Some(Op (Oandlimm m) ys) => Some(Oandlimm (Int64.and m n), ys)
- | _ => None
- end
- | Oorlimm n, x :: nil =>
- match get x with
- | Some(Op (Oorlimm m) ys) => Some(Oorlimm (Int64.or m n), ys)
- | _ => None
- end
- | Oxorlimm n, x :: nil =>
- match get x with
- | Some(Op (Oxorlimm m) ys) => Some(Oxorlimm (Int64.xor m n), ys)
- | _ => None
- end
- | Ocmp cond, _ =>
- match combine_cond cond args with
- | Some(cond', args') => Some(Ocmp cond', args')
- | None => None
- end
- | _, _ => None
- end.
-
-End COMBINE.
-
-
diff --git a/ia32/CombineOpproof.v b/ia32/CombineOpproof.v
deleted file mode 100644
index f59e582b..00000000
--- a/ia32/CombineOpproof.v
+++ /dev/null
@@ -1,179 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Recognition of combined operations, addressing modes and conditions
- during the [CSE] phase. *)
-
-Require Import Coqlib.
-Require Import Integers Values Memory.
-Require Import Op RTL CSEdomain.
-Require Import CombineOp.
-
-Section COMBINE.
-
-Variable ge: genv.
-Variable sp: val.
-Variable m: mem.
-Variable get: valnum -> option rhs.
-Variable valu: valnum -> val.
-Hypothesis get_sound: forall v rhs, get v = Some rhs -> rhs_eval_to valu ge sp m rhs (valu v).
-
-Lemma get_op_sound:
- forall v op vl, get v = Some (Op op vl) -> eval_operation ge sp op (map valu vl) m = Some (valu v).
-Proof.
- intros. exploit get_sound; eauto. intros REV; inv REV; auto.
-Qed.
-
-Ltac UseGetSound :=
- match goal with
- | [ H: get _ = Some _ |- _ ] =>
- let x := fresh "EQ" in (generalize (get_op_sound _ _ _ H); intros x; simpl in x; FuncInv)
- end.
-
-Lemma combine_compimm_ne_0_sound:
- forall x cond args,
- combine_compimm_ne_0 get x = Some(cond, args) ->
- eval_condition cond (map valu args) m = Val.cmp_bool Cne (valu x) (Vint Int.zero) /\
- eval_condition cond (map valu args) m = Val.cmpu_bool (Mem.valid_pointer m) Cne (valu x) (Vint Int.zero).
-Proof.
- intros until args. functional induction (combine_compimm_ne_0 get x); intros EQ; inv EQ.
- (* of cmp *)
- UseGetSound. rewrite <- H.
- destruct (eval_condition cond (map valu args) m); simpl; auto. destruct b; auto.
- (* of and *)
- UseGetSound. rewrite <- H.
- destruct v; simpl; auto.
-Qed.
-
-Lemma combine_compimm_eq_0_sound:
- forall x cond args,
- combine_compimm_eq_0 get x = Some(cond, args) ->
- eval_condition cond (map valu args) m = Val.cmp_bool Ceq (valu x) (Vint Int.zero) /\
- eval_condition cond (map valu args) m = Val.cmpu_bool (Mem.valid_pointer m) Ceq (valu x) (Vint Int.zero).
-Proof.
- intros until args. functional induction (combine_compimm_eq_0 get x); intros EQ; inv EQ.
- (* of cmp *)
- UseGetSound. rewrite <- H.
- rewrite eval_negate_condition.
- destruct (eval_condition c (map valu args) m); simpl; auto. destruct b; auto.
- (* of and *)
- UseGetSound. rewrite <- H. destruct v; auto.
-Qed.
-
-Lemma combine_compimm_eq_1_sound:
- forall x cond args,
- combine_compimm_eq_1 get x = Some(cond, args) ->
- eval_condition cond (map valu args) m = Val.cmp_bool Ceq (valu x) (Vint Int.one) /\
- eval_condition cond (map valu args) m = Val.cmpu_bool (Mem.valid_pointer m) Ceq (valu x) (Vint Int.one).
-Proof.
- intros until args. functional induction (combine_compimm_eq_1 get x); intros EQ; inv EQ.
- (* of cmp *)
- UseGetSound. rewrite <- H.
- destruct (eval_condition cond (map valu args) m); simpl; auto. destruct b; auto.
-Qed.
-
-Lemma combine_compimm_ne_1_sound:
- forall x cond args,
- combine_compimm_ne_1 get x = Some(cond, args) ->
- eval_condition cond (map valu args) m = Val.cmp_bool Cne (valu x) (Vint Int.one) /\
- eval_condition cond (map valu args) m = Val.cmpu_bool (Mem.valid_pointer m) Cne (valu x) (Vint Int.one).
-Proof.
- intros until args. functional induction (combine_compimm_ne_1 get x); intros EQ; inv EQ.
- (* of cmp *)
- UseGetSound. rewrite <- H.
- rewrite eval_negate_condition.
- destruct (eval_condition c (map valu args) m); simpl; auto. destruct b; auto.
-Qed.
-
-Theorem combine_cond_sound:
- forall cond args cond' args',
- combine_cond get cond args = Some(cond', args') ->
- eval_condition cond' (map valu args') m = eval_condition cond (map valu args) m.
-Proof.
- intros. functional inversion H; subst.
- (* compimm ne zero *)
- simpl; eapply combine_compimm_ne_0_sound; eauto.
- (* compimm ne one *)
- simpl; eapply combine_compimm_ne_1_sound; eauto.
- (* compimm eq zero *)
- simpl; eapply combine_compimm_eq_0_sound; eauto.
- (* compimm eq one *)
- simpl; eapply combine_compimm_eq_1_sound; eauto.
- (* compuimm ne zero *)
- simpl; eapply combine_compimm_ne_0_sound; eauto.
- (* compuimm ne one *)
- simpl; eapply combine_compimm_ne_1_sound; eauto.
- (* compuimm eq zero *)
- simpl; eapply combine_compimm_eq_0_sound; eauto.
- (* compuimm eq one *)
- simpl; eapply combine_compimm_eq_1_sound; eauto.
-Qed.
-
-Theorem combine_addr_32_sound:
- forall addr args addr' args',
- combine_addr_32 get addr args = Some(addr', args') ->
- eval_addressing32 ge sp addr' (map valu args') = eval_addressing32 ge sp addr (map valu args).
-Proof.
- intros. functional inversion H; subst.
- (* indexed - lea *)
- UseGetSound. simpl. unfold offset_addressing in H7. destruct (addressing_valid (offset_addressing_total a n)); inv H7.
- eapply eval_offset_addressing_total_32; eauto.
-Qed.
-
-Theorem combine_addr_64_sound:
- forall addr args addr' args',
- combine_addr_64 get addr args = Some(addr', args') ->
- eval_addressing64 ge sp addr' (map valu args') = eval_addressing64 ge sp addr (map valu args).
-Proof.
- intros. functional inversion H; subst.
- (* indexed - leal *)
- UseGetSound. simpl. unfold offset_addressing in H7. destruct (addressing_valid (offset_addressing_total a n)); inv H7.
- eapply eval_offset_addressing_total_64; eauto.
-Qed.
-
-Theorem combine_addr_sound:
- forall addr args addr' args',
- combine_addr get addr args = Some(addr', args') ->
- eval_addressing ge sp addr' (map valu args') = eval_addressing ge sp addr (map valu args).
-Proof.
- unfold combine_addr, eval_addressing; intros; destruct Archi.ptr64.
- apply combine_addr_64_sound; auto.
- apply combine_addr_32_sound; auto.
-Qed.
-
-Theorem combine_op_sound:
- forall op args op' args',
- combine_op get op args = Some(op', args') ->
- eval_operation ge sp op' (map valu args') m = eval_operation ge sp op (map valu args) m.
-Proof.
- intros. functional inversion H; subst.
-(* lea-lea *)
- simpl. eapply combine_addr_32_sound; eauto.
-(* leal-leal *)
- simpl. eapply combine_addr_64_sound; eauto.
-(* andimm - andimm *)
- UseGetSound; simpl. rewrite <- H0. rewrite Val.and_assoc. auto.
-(* orimm - orimm *)
- UseGetSound; simpl. rewrite <- H0. rewrite Val.or_assoc. auto.
-(* xorimm - xorimm *)
- UseGetSound; simpl. rewrite <- H0. rewrite Val.xor_assoc. auto.
-(* andimm - andimm *)
- UseGetSound; simpl. rewrite <- H0. rewrite Val.andl_assoc. auto.
-(* orimm - orimm *)
- UseGetSound; simpl. rewrite <- H0. rewrite Val.orl_assoc. auto.
-(* xorimm - xorimm *)
- UseGetSound; simpl. rewrite <- H0. rewrite Val.xorl_assoc. auto.
-(* cmp *)
- simpl. decEq; decEq. eapply combine_cond_sound; eauto.
-Qed.
-
-End COMBINE.
diff --git a/ia32/ConstpropOp.vp b/ia32/ConstpropOp.vp
deleted file mode 100644
index 0bf143d2..00000000
--- a/ia32/ConstpropOp.vp
+++ /dev/null
@@ -1,404 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Strength reduction for operators and conditions.
- This is the machine-dependent part of [Constprop]. *)
-
-Require Import Coqlib Compopts.
-Require Import AST Integers Floats.
-Require Import Op Registers.
-Require Import ValueDomain.
-
-(** * Converting known values to constants *)
-
-Parameter symbol_is_external: ident -> bool. (**r See [SelectOp] *)
-
-Definition Olea_ptr (a: addressing) := if Archi.ptr64 then Oleal a else Olea a.
-
-Definition const_for_result (a: aval) : option operation :=
- match a with
- | I n => Some(Ointconst n)
- | L n => if Archi.ptr64 then Some(Olongconst n) else None
- | F n => if Compopts.generate_float_constants tt then Some(Ofloatconst n) else None
- | FS n => if Compopts.generate_float_constants tt then Some(Osingleconst n) else None
- | Ptr(Gl id ofs) =>
- if symbol_is_external id then
- if Ptrofs.eq ofs Ptrofs.zero then Some (Oindirectsymbol id) else None
- else
- Some (Olea_ptr (Aglobal id ofs))
- | Ptr(Stk ofs) => Some(Olea_ptr (Ainstack ofs))
- | _ => None
- end.
-
-(** * Operator strength reduction *)
-
-(** We now define auxiliary functions for strength reduction of
- operators and addressing modes: replacing an operator with a cheaper
- one if some of its arguments are statically known. These are again
- large pattern-matchings expressed in indirect style. *)
-
-Nondetfunction cond_strength_reduction
- (cond: condition) (args: list reg) (vl: list aval) :=
- match cond, args, vl with
- | Ccomp c, r1 :: r2 :: nil, I n1 :: v2 :: nil =>
- (Ccompimm (swap_comparison c) n1, r2 :: nil)
- | Ccomp c, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
- (Ccompimm c n2, r1 :: nil)
- | Ccompu c, r1 :: r2 :: nil, I n1 :: v2 :: nil =>
- (Ccompuimm (swap_comparison c) n1, r2 :: nil)
- | Ccompu c, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
- (Ccompuimm c n2, r1 :: nil)
- | Ccompl c, r1 :: r2 :: nil, L n1 :: v2 :: nil =>
- (Ccomplimm (swap_comparison c) n1, r2 :: nil)
- | Ccompl c, r1 :: r2 :: nil, v1 :: L n2 :: nil =>
- (Ccomplimm c n2, r1 :: nil)
- | Ccomplu c, r1 :: r2 :: nil, L n1 :: v2 :: nil =>
- (Ccompluimm (swap_comparison c) n1, r2 :: nil)
- | Ccomplu c, r1 :: r2 :: nil, v1 :: L n2 :: nil =>
- (Ccompluimm c n2, r1 :: nil)
- | _, _, _ =>
- (cond, args)
- end.
-
-Definition make_cmp_base (c: condition) (args: list reg) (vl: list aval) :=
- let (c', args') := cond_strength_reduction c args vl in (Ocmp c', args').
-
-Nondetfunction make_cmp (c: condition) (args: list reg) (vl: list aval) :=
- match c, args, vl with
- | Ccompimm Ceq n, r1 :: nil, v1 :: nil =>
- if Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1) then (Omove, r1 :: nil)
- else if Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1) then (Oxorimm Int.one, r1 :: nil)
- else make_cmp_base c args vl
- | Ccompimm Cne n, r1 :: nil, v1 :: nil =>
- if Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1) then (Omove, r1 :: nil)
- else if Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1) then (Oxorimm Int.one, r1 :: nil)
- else make_cmp_base c args vl
- | _, _, _ =>
- make_cmp_base c args vl
- end.
-
-(** For addressing modes, we need to distinguish
-- reductions that produce pointers (i.e. that produce [Aglobal], [Ainstack], [Abased] and [Abasedscaled] addressing modes), which are valid only if the pointer size is right;
-- other reductions (producing [Aindexed] or [Aindexed2] modes), which are valid independently of the pointer size.
-*)
-
-Nondetfunction addr_strength_reduction_32_generic
- (addr: addressing) (args: list reg) (vl: list aval) :=
- match addr, args, vl with
- | Aindexed2 ofs, r1 :: r2 :: nil, I n1 :: v2 :: nil =>
- (Aindexed (Int.signed n1 + ofs), r2 :: nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
- (Aindexed (Int.signed n2 + ofs), r1 :: nil)
- | Aindexed2scaled sc ofs, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
- (Aindexed (Int.signed n2 * sc + ofs), r1 :: nil)
- | Aindexed2scaled sc ofs, r1 :: r2 :: nil, I n1 :: v2 :: nil =>
- (Ascaled sc (Int.signed n1 + ofs), r2 :: nil)
- | _, _ =>
- (addr, args)
- end.
-
-Nondetfunction addr_strength_reduction_32
- (addr: addressing) (args: list reg) (vl: list aval) :=
-
- if Archi.ptr64 then addr_strength_reduction_32_generic addr args vl else
-
- match addr, args, vl with
-
- | Aindexed ofs, r1 :: nil, Ptr(Gl symb n) :: nil =>
- (Aglobal symb (Ptrofs.add n (Ptrofs.repr ofs)), nil)
- | Aindexed ofs, r1 :: nil, Ptr(Stk n) :: nil =>
- (Ainstack (Ptrofs.add n (Ptrofs.repr ofs)), nil)
-
- | Aindexed2 ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: I n2 :: nil =>
- (Aglobal symb (Ptrofs.add (Ptrofs.add n1 (Ptrofs.of_int n2)) (Ptrofs.repr ofs)), nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, I n1 :: Ptr(Gl symb n2) :: nil =>
- (Aglobal symb (Ptrofs.add (Ptrofs.add n2 (Ptrofs.of_int n1)) (Ptrofs.repr ofs)), nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, Ptr(Stk n1) :: I n2 :: nil =>
- (Ainstack (Ptrofs.add (Ptrofs.add n1 (Ptrofs.of_int n2)) (Ptrofs.repr ofs)), nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, I n1 :: Ptr(Stk n2) :: nil =>
- (Ainstack (Ptrofs.add (Ptrofs.add (Ptrofs.of_int n1) n2) (Ptrofs.repr ofs)), nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: v2 :: nil =>
- (Abased symb (Ptrofs.add n1 (Ptrofs.repr ofs)), r2 :: nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, v1 :: Ptr(Gl symb n2) :: nil =>
- (Abased symb (Ptrofs.add n2 (Ptrofs.repr ofs)), r1 :: nil)
-
- | Aindexed2scaled sc ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: I n2 :: nil =>
- (Aglobal symb (Ptrofs.add (Ptrofs.add n1 (Ptrofs.of_int (Int.mul n2 (Int.repr sc)))) (Ptrofs.repr ofs)), nil)
-
- | Aindexed2scaled sc ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: v2 :: nil =>
- (Abasedscaled sc symb (Ptrofs.add n1 (Ptrofs.repr ofs)), r2 :: nil)
-
- | Abased id ofs, r1 :: nil, I n1 :: nil =>
- (Aglobal id (Ptrofs.add ofs (Ptrofs.of_int n1)), nil)
-
- | Abasedscaled sc id ofs, r1 :: nil, I n1 :: nil =>
- (Aglobal id (Ptrofs.add ofs (Ptrofs.of_int (Int.mul n1 (Int.repr sc)))), nil)
-
- | _, _ =>
- addr_strength_reduction_32_generic addr args vl
- end.
-
-Nondetfunction addr_strength_reduction_64_generic
- (addr: addressing) (args: list reg) (vl: list aval) :=
- match addr, args, vl with
- | Aindexed2 ofs, r1 :: r2 :: nil, L n1 :: v2 :: nil =>
- (Aindexed (Int64.signed n1 + ofs), r2 :: nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, v1 :: L n2 :: nil =>
- (Aindexed (Int64.signed n2 + ofs), r1 :: nil)
- | Aindexed2scaled sc ofs, r1 :: r2 :: nil, v1 :: L n2 :: nil =>
- (Aindexed (Int64.signed n2 * sc + ofs), r1 :: nil)
- | Aindexed2scaled sc ofs, r1 :: r2 :: nil, L n1 :: v2 :: nil =>
- (Ascaled sc (Int64.signed n1 + ofs), r2 :: nil)
- | _, _ =>
- (addr, args)
- end.
-
-Nondetfunction addr_strength_reduction_64
- (addr: addressing) (args: list reg) (vl: list aval) :=
-
- if negb Archi.ptr64 then addr_strength_reduction_64_generic addr args vl else
-
- match addr, args, vl with
-
- | Aindexed ofs, r1 :: nil, Ptr(Gl symb n) :: nil =>
- (Aglobal symb (Ptrofs.add n (Ptrofs.repr ofs)), nil)
- | Aindexed ofs, r1 :: nil, Ptr(Stk n) :: nil =>
- (Ainstack (Ptrofs.add n (Ptrofs.repr ofs)), nil)
-
- | Aindexed2 ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: L n2 :: nil =>
- (Aglobal symb (Ptrofs.add (Ptrofs.add n1 (Ptrofs.of_int64 n2)) (Ptrofs.repr ofs)), nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, L n1 :: Ptr(Gl symb n2) :: nil =>
- (Aglobal symb (Ptrofs.add (Ptrofs.add n2 (Ptrofs.of_int64 n1)) (Ptrofs.repr ofs)), nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, Ptr(Stk n1) :: L n2 :: nil =>
- (Ainstack (Ptrofs.add (Ptrofs.add n1 (Ptrofs.of_int64 n2)) (Ptrofs.repr ofs)), nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, L n1 :: Ptr(Stk n2) :: nil =>
- (Ainstack (Ptrofs.add (Ptrofs.add (Ptrofs.of_int64 n1) n2) (Ptrofs.repr ofs)), nil)
-
- | Aindexed2scaled sc ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: L n2 :: nil =>
- (Aglobal symb (Ptrofs.add (Ptrofs.add n1 (Ptrofs.of_int64 (Int64.mul n2 (Int64.repr sc)))) (Ptrofs.repr ofs)), nil)
-
- | _, _ =>
- addr_strength_reduction_64_generic addr args vl
- end.
-
-Definition addr_strength_reduction
- (addr: addressing) (args: list reg) (vl: list aval) :=
- if Archi.ptr64
- then addr_strength_reduction_64 addr args vl
- else addr_strength_reduction_32 addr args vl.
-
-Definition make_addimm (n: int) (r: reg) :=
- if Int.eq n Int.zero
- then (Omove, r :: nil)
- else (Olea (Aindexed (Int.signed n)), r :: nil).
-
-Definition make_shlimm (n: int) (r1 r2: reg) :=
- if Int.eq n Int.zero then (Omove, r1 :: nil)
- else if Int.ltu n Int.iwordsize then (Oshlimm n, r1 :: nil)
- else (Oshl, r1 :: r2 :: nil).
-
-Definition make_shrimm (n: int) (r1 r2: reg) :=
- if Int.eq n Int.zero then (Omove, r1 :: nil)
- else if Int.ltu n Int.iwordsize then (Oshrimm n, r1 :: nil)
- else (Oshr, r1 :: r2 :: nil).
-
-Definition make_shruimm (n: int) (r1 r2: reg) :=
- if Int.eq n Int.zero then (Omove, r1 :: nil)
- else if Int.ltu n Int.iwordsize then (Oshruimm n, r1 :: nil)
- else (Oshru, r1 :: r2 :: nil).
-
-Definition make_mulimm (n: int) (r: reg) :=
- if Int.eq n Int.zero then
- (Ointconst Int.zero, nil)
- else if Int.eq n Int.one then
- (Omove, r :: nil)
- else
- match Int.is_power2 n with
- | Some l => (Oshlimm l, r :: nil)
- | None => (Omulimm n, r :: nil)
- end.
-
-Definition make_andimm (n: int) (r: reg) (a: aval) :=
- if Int.eq n Int.zero then (Ointconst Int.zero, nil)
- else if Int.eq n Int.mone then (Omove, r :: nil)
- else if match a with Uns _ m => Int.eq (Int.zero_ext m (Int.not n)) Int.zero
- | _ => false end
- then (Omove, r :: nil)
- else (Oandimm n, r :: nil).
-
-Definition make_orimm (n: int) (r: reg) :=
- if Int.eq n Int.zero then (Omove, r :: nil)
- else if Int.eq n Int.mone then (Ointconst Int.mone, nil)
- else (Oorimm n, r :: nil).
-
-Definition make_xorimm (n: int) (r: reg) :=
- if Int.eq n Int.zero then (Omove, r :: nil)
- else if Int.eq n Int.mone then (Onot, r :: nil)
- else (Oxorimm n, r :: nil).
-
-Definition make_divimm n (r1 r2: reg) :=
- match Int.is_power2 n with
- | Some l => if Int.ltu l (Int.repr 31)
- then (Oshrximm l, r1 :: nil)
- else (Odiv, r1 :: r2 :: nil)
- | None => (Odiv, r1 :: r2 :: nil)
- end.
-
-Definition make_divuimm n (r1 r2: reg) :=
- match Int.is_power2 n with
- | Some l => (Oshruimm l, r1 :: nil)
- | None => (Odivu, r1 :: r2 :: nil)
- end.
-
-Definition make_moduimm n (r1 r2: reg) :=
- match Int.is_power2 n with
- | Some l => (Oandimm (Int.sub n Int.one), r1 :: nil)
- | None => (Omodu, r1 :: r2 :: nil)
- end.
-
-Definition make_addlimm (n: int64) (r: reg) :=
- if Int64.eq n Int64.zero
- then (Omove, r :: nil)
- else (Oleal (Aindexed (Int64.signed n)), r :: nil).
-
-Definition make_shllimm (n: int) (r1 r2: reg) :=
- if Int.eq n Int.zero then (Omove, r1 :: nil)
- else if Int.ltu n Int64.iwordsize' then (Oshllimm n, r1 :: nil)
- else (Oshll, r1 :: r2 :: nil).
-
-Definition make_shrlimm (n: int) (r1 r2: reg) :=
- if Int.eq n Int.zero then (Omove, r1 :: nil)
- else if Int.ltu n Int64.iwordsize' then (Oshrlimm n, r1 :: nil)
- else (Oshrl, r1 :: r2 :: nil).
-
-Definition make_shrluimm (n: int) (r1 r2: reg) :=
- if Int.eq n Int.zero then (Omove, r1 :: nil)
- else if Int.ltu n Int64.iwordsize' then (Oshrluimm n, r1 :: nil)
- else (Oshrlu, r1 :: r2 :: nil).
-
-Definition make_mullimm (n: int64) (r: reg) :=
- if Int64.eq n Int64.zero then
- (Olongconst Int64.zero, nil)
- else if Int64.eq n Int64.one then
- (Omove, r :: nil)
- else
- match Int64.is_power2' n with
- | Some l => (Oshllimm l, r :: nil)
- | None => (Omullimm n, r :: nil)
- end.
-
-Definition make_andlimm (n: int64) (r: reg) (a: aval) :=
- if Int64.eq n Int64.zero then (Olongconst Int64.zero, nil)
- else if Int64.eq n Int64.mone then (Omove, r :: nil)
- else (Oandlimm n, r :: nil).
-
-Definition make_orlimm (n: int64) (r: reg) :=
- if Int64.eq n Int64.zero then (Omove, r :: nil)
- else if Int64.eq n Int64.mone then (Olongconst Int64.mone, nil)
- else (Oorlimm n, r :: nil).
-
-Definition make_xorlimm (n: int64) (r: reg) :=
- if Int64.eq n Int64.zero then (Omove, r :: nil)
- else if Int64.eq n Int64.mone then (Onotl, r :: nil)
- else (Oxorlimm n, r :: nil).
-
-Definition make_divlimm n (r1 r2: reg) :=
- match Int64.is_power2' n with
- | Some l => if Int.ltu l (Int.repr 63)
- then (Oshrxlimm l, r1 :: nil)
- else (Odivl, r1 :: r2 :: nil)
- | None => (Odivl, r1 :: r2 :: nil)
- end.
-
-Definition make_divluimm n (r1 r2: reg) :=
- match Int64.is_power2' n with
- | Some l => (Oshrluimm l, r1 :: nil)
- | None => (Odivlu, r1 :: r2 :: nil)
- end.
-
-Definition make_modluimm n (r1 r2: reg) :=
- match Int64.is_power2 n with
- | Some l => (Oandlimm (Int64.sub n Int64.one), r1 :: nil)
- | None => (Omodlu, r1 :: r2 :: nil)
- end.
-
-Definition make_mulfimm (n: float) (r r1 r2: reg) :=
- if Float.eq_dec n (Float.of_int (Int.repr 2))
- then (Oaddf, r :: r :: nil)
- else (Omulf, r1 :: r2 :: nil).
-
-Definition make_mulfsimm (n: float32) (r r1 r2: reg) :=
- if Float32.eq_dec n (Float32.of_int (Int.repr 2))
- then (Oaddfs, r :: r :: nil)
- else (Omulfs, r1 :: r2 :: nil).
-
-Definition make_cast8signed (r: reg) (a: aval) :=
- if vincl a (Sgn Ptop 8) then (Omove, r :: nil) else (Ocast8signed, r :: nil).
-Definition make_cast8unsigned (r: reg) (a: aval) :=
- if vincl a (Uns Ptop 8) then (Omove, r :: nil) else (Ocast8unsigned, r :: nil).
-Definition make_cast16signed (r: reg) (a: aval) :=
- if vincl a (Sgn Ptop 16) then (Omove, r :: nil) else (Ocast16signed, r :: nil).
-Definition make_cast16unsigned (r: reg) (a: aval) :=
- if vincl a (Uns Ptop 16) then (Omove, r :: nil) else (Ocast16unsigned, r :: nil).
-
-Nondetfunction op_strength_reduction
- (op: operation) (args: list reg) (vl: list aval) :=
- match op, args, vl with
- | Ocast8signed, r1 :: nil, v1 :: nil => make_cast8signed r1 v1
- | Ocast8unsigned, r1 :: nil, v1 :: nil => make_cast8unsigned r1 v1
- | Ocast16signed, r1 :: nil, v1 :: nil => make_cast16signed r1 v1
- | Ocast16unsigned, r1 :: nil, v1 :: nil => make_cast16unsigned r1 v1
- | Osub, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_addimm (Int.neg n2) r1
- | Omul, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_mulimm n1 r2
- | Omul, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_mulimm n2 r1
- | Odiv, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_divimm n2 r1 r2
- | Odivu, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_divuimm n2 r1 r2
- | Omodu, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_moduimm n2 r1 r2
- | Oand, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_andimm n1 r2 v2
- | Oand, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_andimm n2 r1 v1
- | Oandimm n, r1 :: nil, v1 :: nil => make_andimm n r1 v1
- | Oor, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_orimm n1 r2
- | Oor, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_orimm n2 r1
- | Oxor, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_xorimm n1 r2
- | Oxor, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_xorimm n2 r1
- | Oshl, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shlimm n2 r1 r2
- | Oshr, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrimm n2 r1 r2
- | Oshru, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shruimm n2 r1 r2
- | Olea addr, args, vl =>
- let (addr', args') := addr_strength_reduction_32 addr args vl in
- (Olea addr', args')
- | Osubl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_addlimm (Int64.neg n2) r1
- | Omull, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_mullimm n1 r2
- | Omull, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_mullimm n2 r1
- | Odivl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_divlimm n2 r1 r2
- | Odivlu, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_divluimm n2 r1 r2
- | Omodlu, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_modluimm n2 r1 r2
- | Oandl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_andlimm n1 r2 v2
- | Oandl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_andlimm n2 r1 v1
- | Oandlimm n, r1 :: nil, v1 :: nil => make_andlimm n r1 v1
- | Oorl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_orlimm n1 r2
- | Oorl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_orlimm n2 r1
- | Oxorl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_xorlimm n1 r2
- | Oxorl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_xorlimm n2 r1
- | Oshll, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shllimm n2 r1 r2
- | Oshrl, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrlimm n2 r1 r2
- | Oshrlu, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrluimm n2 r1 r2
- | Oleal addr, args, vl =>
- let (addr', args') := addr_strength_reduction_64 addr args vl in
- (Oleal addr', args')
- | Ocmp c, args, vl => make_cmp c args vl
- | Omulf, r1 :: r2 :: nil, v1 :: F n2 :: nil => make_mulfimm n2 r1 r1 r2
- | Omulf, r1 :: r2 :: nil, F n1 :: v2 :: nil => make_mulfimm n1 r2 r1 r2
- | Omulfs, r1 :: r2 :: nil, v1 :: FS n2 :: nil => make_mulfsimm n2 r1 r1 r2
- | Omulfs, r1 :: r2 :: nil, FS n1 :: v2 :: nil => make_mulfsimm n1 r2 r1 r2
- | _, _, _ => (op, args)
- end.
diff --git a/ia32/ConstpropOpproof.v b/ia32/ConstpropOpproof.v
deleted file mode 100644
index 161b9579..00000000
--- a/ia32/ConstpropOpproof.v
+++ /dev/null
@@ -1,881 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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 operator strength reduction. *)
-
-Require Import Coqlib Compopts.
-Require Import Integers Floats Values Memory Globalenvs Events.
-Require Import Op Registers RTL ValueDomain.
-Require Import ConstpropOp.
-
-Section STRENGTH_REDUCTION.
-
-Variable bc: block_classification.
-Variable ge: genv.
-Hypothesis GENV: genv_match bc ge.
-Variable sp: block.
-Hypothesis STACK: bc sp = BCstack.
-Variable ae: AE.t.
-Variable e: regset.
-Variable m: mem.
-Hypothesis MATCH: ematch bc e ae.
-
-Lemma match_G:
- forall r id ofs,
- AE.get r ae = Ptr(Gl id ofs) -> Val.lessdef e#r (Genv.symbol_address ge id ofs).
-Proof.
- intros. apply vmatch_ptr_gl with bc; auto. rewrite <- H. apply MATCH.
-Qed.
-
-Lemma match_S:
- forall r ofs,
- AE.get r ae = Ptr(Stk ofs) -> Val.lessdef e#r (Vptr sp ofs).
-Proof.
- intros. apply vmatch_ptr_stk with bc; auto. rewrite <- H. apply MATCH.
-Qed.
-
-Ltac InvApproxRegs :=
- match goal with
- | [ H: _ :: _ = _ :: _ |- _ ] =>
- injection H; clear H; intros; InvApproxRegs
- | [ H: ?v = AE.get ?r ae |- _ ] =>
- generalize (MATCH r); rewrite <- H; clear H; intro; InvApproxRegs
- | _ => idtac
- end.
-
-Ltac SimplVM :=
- match goal with
- | [ H: vmatch _ ?v (I ?n) |- _ ] =>
- let E := fresh in
- assert (E: v = Vint n) by (inversion H; auto);
- rewrite E in *; clear H; SimplVM
- | [ H: vmatch _ ?v (L ?n) |- _ ] =>
- let E := fresh in
- assert (E: v = Vlong n) by (inversion H; auto);
- rewrite E in *; clear H; SimplVM
- | [ H: vmatch _ ?v (F ?n) |- _ ] =>
- let E := fresh in
- assert (E: v = Vfloat n) by (inversion H; auto);
- rewrite E in *; clear H; SimplVM
- | [ H: vmatch _ ?v (FS ?n) |- _ ] =>
- let E := fresh in
- assert (E: v = Vsingle n) by (inversion H; auto);
- rewrite E in *; clear H; SimplVM
- | [ H: vmatch _ ?v (Ptr(Gl ?id ?ofs)) |- _ ] =>
- let E := fresh in
- assert (E: Val.lessdef v (Genv.symbol_address ge id ofs)) by (eapply vmatch_ptr_gl; eauto);
- clear H; SimplVM
- | [ H: vmatch _ ?v (Ptr(Stk ?ofs)) |- _ ] =>
- let E := fresh in
- assert (E: Val.lessdef v (Vptr sp ofs)) by (eapply vmatch_ptr_stk; eauto);
- clear H; SimplVM
- | _ => idtac
- end.
-
-Lemma eval_Olea_ptr:
- forall a el,
- eval_operation ge (Vptr sp Ptrofs.zero) (Olea_ptr a) el m = eval_addressing ge (Vptr sp Ptrofs.zero) a el.
-Proof.
- unfold Olea_ptr, eval_addressing; intros. destruct Archi.ptr64; auto.
-Qed.
-
-Lemma const_for_result_correct:
- forall a op v,
- const_for_result a = Some op ->
- vmatch bc v a ->
- exists v', eval_operation ge (Vptr sp Ptrofs.zero) op nil m = Some v' /\ Val.lessdef v v'.
-Proof.
- unfold const_for_result; intros.
- destruct a; inv H; SimplVM.
-- (* integer *)
- exists (Vint n); auto.
-- (* long *)
- destruct Archi.ptr64; inv H2. exists (Vlong n); auto.
-- (* float *)
- destruct (Compopts.generate_float_constants tt); inv H2. exists (Vfloat f); auto.
-- (* single *)
- destruct (Compopts.generate_float_constants tt); inv H2. exists (Vsingle f); auto.
-- (* pointer *)
- destruct p; try discriminate; SimplVM.
- + (* global *)
- destruct (symbol_is_external id).
- * revert H2; predSpec Ptrofs.eq Ptrofs.eq_spec ofs Ptrofs.zero; intros EQ; inv EQ.
- exists (Genv.symbol_address ge id Ptrofs.zero); auto.
- * inv H2. exists (Genv.symbol_address ge id ofs); split; auto.
- rewrite eval_Olea_ptr. apply eval_addressing_Aglobal.
- + (* stack *)
- inv H2. exists (Vptr sp ofs); split; auto.
- rewrite eval_Olea_ptr. rewrite eval_addressing_Ainstack.
- simpl. rewrite Ptrofs.add_zero_l; auto.
-Qed.
-
-Lemma cond_strength_reduction_correct:
- forall cond args vl,
- vl = map (fun r => AE.get r ae) args ->
- let (cond', args') := cond_strength_reduction cond args vl in
- eval_condition cond' e##args' m = eval_condition cond e##args m.
-Proof.
- intros until vl. unfold cond_strength_reduction.
- case (cond_strength_reduction_match cond args vl); simpl; intros; InvApproxRegs; SimplVM.
-- apply Val.swap_cmp_bool.
-- auto.
-- apply Val.swap_cmpu_bool.
-- auto.
-- apply Val.swap_cmpl_bool.
-- auto.
-- apply Val.swap_cmplu_bool.
-- auto.
-- auto.
-Qed.
-
-Lemma addr_strength_reduction_32_generic_correct:
- forall addr args vl res,
- vl = map (fun r => AE.get r ae) args ->
- eval_addressing32 ge (Vptr sp Ptrofs.zero) addr e##args = Some res ->
- let (addr', args') := addr_strength_reduction_32_generic addr args vl in
- exists res', eval_addressing32 ge (Vptr sp Ptrofs.zero) addr' e##args' = Some res' /\ Val.lessdef res res'.
-Proof.
-Local Opaque Val.add.
- assert (A: forall x y, Int.repr (Int.signed x + y) = Int.add x (Int.repr y)).
- { intros; apply Int.eqm_samerepr; auto using Int.eqm_signed_unsigned with ints. }
- assert (B: forall x y z, Int.repr (Int.signed x * y + z) = Int.add (Int.mul x (Int.repr y)) (Int.repr z)).
- { intros; apply Int.eqm_samerepr; apply Int.eqm_add; auto with ints.
- unfold Int.mul; auto using Int.eqm_signed_unsigned with ints. }
- intros until res; intros VL EA.
- unfold addr_strength_reduction_32_generic; destruct (addr_strength_reduction_32_generic_match addr args vl);
- simpl in *; InvApproxRegs; SimplVM; try (inv EA).
-- econstructor; split; eauto. rewrite A, Val.add_assoc, Val.add_permut. auto.
-- econstructor; split; eauto. rewrite A, Val.add_assoc. auto.
-- Local Transparent Val.add.
- econstructor; split; eauto. simpl. rewrite B. auto.
-- econstructor; split; eauto. rewrite A, Val.add_permut. auto.
-- exists res; auto.
-Qed.
-
-Lemma addr_strength_reduction_32_correct:
- forall addr args vl res,
- vl = map (fun r => AE.get r ae) args ->
- eval_addressing32 ge (Vptr sp Ptrofs.zero) addr e##args = Some res ->
- let (addr', args') := addr_strength_reduction_32 addr args vl in
- exists res', eval_addressing32 ge (Vptr sp Ptrofs.zero) addr' e##args' = Some res' /\ Val.lessdef res res'.
-Proof.
- intros until res; intros VL EA. unfold addr_strength_reduction_32.
- destruct Archi.ptr64 eqn:SF. apply addr_strength_reduction_32_generic_correct; auto.
- assert (A: forall n, Ptrofs.of_int (Int.repr n) = Ptrofs.repr n) by auto with ptrofs.
- assert (B: forall symb ofs n,
- Genv.symbol_address ge symb (Ptrofs.add ofs (Ptrofs.repr n)) =
- Val.add (Genv.symbol_address ge symb ofs) (Vint (Int.repr n))).
- { intros. rewrite <- A. apply Genv.shift_symbol_address_32; auto. }
-Local Opaque Val.add.
- destruct (addr_strength_reduction_32_match addr args vl);
- simpl in *; InvApproxRegs; SimplVM; try (inv EA); rewrite ? SF.
-- econstructor; split; eauto. rewrite B. apply Val.add_lessdef; auto.
-- econstructor; split; eauto. rewrite Ptrofs.add_zero_l.
-Local Transparent Val.add.
- inv H0; auto. rewrite H2. simpl; rewrite SF, A. auto.
-- econstructor; split; eauto.
- unfold Ptrofs.add at 2. rewrite B.
- fold (Ptrofs.add n1 (Ptrofs.of_int n2)).
- rewrite Genv.shift_symbol_address_32 by auto.
- rewrite ! Val.add_assoc. apply Val.add_lessdef; auto.
-- econstructor; split; eauto.
- unfold Ptrofs.add at 2. rewrite B.
- fold (Ptrofs.add n2 (Ptrofs.of_int n1)).
- rewrite Genv.shift_symbol_address_32 by auto.
- rewrite ! Val.add_assoc. rewrite Val.add_permut. apply Val.add_lessdef; auto.
-- econstructor; split; eauto. rewrite Ptrofs.add_zero_l. rewrite Val.add_assoc.
- eapply Val.lessdef_trans. apply Val.add_lessdef; eauto.
- simpl. rewrite SF. rewrite Ptrofs.add_assoc. apply Val.lessdef_same; do 3 f_equal. auto with ptrofs.
-- econstructor; split; eauto. rewrite Ptrofs.add_zero_l. rewrite Val.add_assoc, Val.add_permut.
- eapply Val.lessdef_trans. apply Val.add_lessdef; eauto.
- simpl. rewrite SF. rewrite <- (Ptrofs.add_commut n2). rewrite Ptrofs.add_assoc.
- apply Val.lessdef_same; do 3 f_equal. auto with ptrofs.
-- econstructor; split; eauto. rewrite B. rewrite ! Val.add_assoc. rewrite (Val.add_commut (Vint (Int.repr ofs))).
- apply Val.add_lessdef; auto.
-- econstructor; split; eauto. rewrite B. rewrite (Val.add_commut e#r1). rewrite ! Val.add_assoc.
- rewrite (Val.add_commut (Vint (Int.repr ofs))). apply Val.add_lessdef; auto.
-- econstructor; split; eauto. rewrite B. rewrite Genv.shift_symbol_address_32 by auto.
- rewrite ! Val.add_assoc. apply Val.add_lessdef; auto.
-- econstructor; split; eauto. rewrite B. rewrite ! Val.add_assoc.
- rewrite (Val.add_commut (Vint (Int.repr ofs))). apply Val.add_lessdef; auto.
-- rewrite SF in H1; inv H1. econstructor; split; eauto.
- rewrite Genv.shift_symbol_address_32 by auto. auto.
-- rewrite SF in H1; inv H1. econstructor; split; eauto.
- rewrite Genv.shift_symbol_address_32 by auto. auto.
-- apply addr_strength_reduction_32_generic_correct; auto.
-Qed.
-
-Lemma addr_strength_reduction_64_generic_correct:
- forall addr args vl res,
- vl = map (fun r => AE.get r ae) args ->
- eval_addressing64 ge (Vptr sp Ptrofs.zero) addr e##args = Some res ->
- let (addr', args') := addr_strength_reduction_64_generic addr args vl in
- exists res', eval_addressing64 ge (Vptr sp Ptrofs.zero) addr' e##args' = Some res' /\ Val.lessdef res res'.
-Proof.
-Local Opaque Val.addl.
- assert (A: forall x y, Int64.repr (Int64.signed x + y) = Int64.add x (Int64.repr y)).
- { intros; apply Int64.eqm_samerepr; auto using Int64.eqm_signed_unsigned with ints. }
- assert (B: forall x y z, Int64.repr (Int64.signed x * y + z) = Int64.add (Int64.mul x (Int64.repr y)) (Int64.repr z)).
- { intros; apply Int64.eqm_samerepr; apply Int64.eqm_add; auto with ints.
- unfold Int64.mul; auto using Int64.eqm_signed_unsigned with ints. }
- intros until res; intros VL EA.
- unfold addr_strength_reduction_64_generic; destruct (addr_strength_reduction_64_generic_match addr args vl);
- simpl in *; InvApproxRegs; SimplVM; try (inv EA).
-- econstructor; split; eauto. rewrite A, Val.addl_assoc, Val.addl_permut. auto.
-- econstructor; split; eauto. rewrite A, Val.addl_assoc. auto.
-- Local Transparent Val.addl.
- econstructor; split; eauto. simpl. rewrite B. auto.
-- econstructor; split; eauto. rewrite A, Val.addl_permut. auto.
-- exists res; auto.
-Qed.
-
-Lemma addr_strength_reduction_64_correct:
- forall addr args vl res,
- vl = map (fun r => AE.get r ae) args ->
- eval_addressing64 ge (Vptr sp Ptrofs.zero) addr e##args = Some res ->
- let (addr', args') := addr_strength_reduction_64 addr args vl in
- exists res', eval_addressing64 ge (Vptr sp Ptrofs.zero) addr' e##args' = Some res' /\ Val.lessdef res res'.
-Proof.
- intros until res; intros VL EA. unfold addr_strength_reduction_64.
- destruct (negb Archi.ptr64) eqn:SF. apply addr_strength_reduction_64_generic_correct; auto.
- rewrite negb_false_iff in SF.
- assert (A: forall n, Ptrofs.of_int64 (Int64.repr n) = Ptrofs.repr n) by auto with ptrofs.
- assert (B: forall symb ofs n,
- Genv.symbol_address ge symb (Ptrofs.add ofs (Ptrofs.repr n)) =
- Val.addl (Genv.symbol_address ge symb ofs) (Vlong (Int64.repr n))).
- { intros. rewrite <- A. apply Genv.shift_symbol_address_64; auto. }
-Local Opaque Val.addl.
- destruct (addr_strength_reduction_64_match addr args vl);
- simpl in *; InvApproxRegs; SimplVM; try (inv EA); rewrite ? SF.
-- econstructor; split; eauto. rewrite B. apply Val.addl_lessdef; auto.
-- econstructor; split; eauto. rewrite Ptrofs.add_zero_l.
-Local Transparent Val.addl.
- inv H0; auto. rewrite H2. simpl; rewrite SF, A. auto.
-- econstructor; split; eauto.
- unfold Ptrofs.add at 2. rewrite B.
- fold (Ptrofs.add n1 (Ptrofs.of_int64 n2)).
- rewrite Genv.shift_symbol_address_64 by auto.
- rewrite ! Val.addl_assoc. apply Val.addl_lessdef; auto.
-- econstructor; split; eauto.
- unfold Ptrofs.add at 2. rewrite B.
- fold (Ptrofs.add n2 (Ptrofs.of_int64 n1)).
- rewrite Genv.shift_symbol_address_64 by auto.
- rewrite ! Val.addl_assoc. rewrite Val.addl_permut. apply Val.addl_lessdef; auto.
-- econstructor; split; eauto. rewrite Ptrofs.add_zero_l. rewrite Val.addl_assoc.
- eapply Val.lessdef_trans. apply Val.addl_lessdef; eauto.
- simpl. rewrite SF. rewrite Ptrofs.add_assoc. apply Val.lessdef_same; do 3 f_equal. auto with ptrofs.
-- econstructor; split; eauto. rewrite Ptrofs.add_zero_l. rewrite Val.addl_assoc, Val.addl_permut.
- eapply Val.lessdef_trans. apply Val.addl_lessdef; eauto.
- simpl. rewrite SF. rewrite <- (Ptrofs.add_commut n2). rewrite Ptrofs.add_assoc.
- apply Val.lessdef_same; do 3 f_equal. auto with ptrofs.
-- econstructor; split; eauto. rewrite B. rewrite Genv.shift_symbol_address_64 by auto.
- rewrite ! Val.addl_assoc. apply Val.addl_lessdef; auto.
-- apply addr_strength_reduction_64_generic_correct; auto.
-Qed.
-
-Lemma addr_strength_reduction_correct:
- forall addr args vl res,
- vl = map (fun r => AE.get r ae) args ->
- eval_addressing ge (Vptr sp Ptrofs.zero) addr e##args = Some res ->
- let (addr', args') := addr_strength_reduction addr args vl in
- exists res', eval_addressing ge (Vptr sp Ptrofs.zero) addr' e##args' = Some res' /\ Val.lessdef res res'.
-Proof.
- unfold eval_addressing, addr_strength_reduction. destruct Archi.ptr64.
- apply addr_strength_reduction_64_correct.
- apply addr_strength_reduction_32_correct.
-Qed.
-
-Lemma make_cmp_base_correct:
- forall c args vl,
- vl = map (fun r => AE.get r ae) args ->
- let (op', args') := make_cmp_base c args vl in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op' e##args' m = Some v
- /\ Val.lessdef (Val.of_optbool (eval_condition c e##args m)) v.
-Proof.
- intros. unfold make_cmp_base.
- generalize (cond_strength_reduction_correct c args vl H).
- destruct (cond_strength_reduction c args vl) as [c' args']. intros EQ.
- econstructor; split. simpl; eauto. rewrite EQ. auto.
-Qed.
-
-Lemma make_cmp_correct:
- forall c args vl,
- vl = map (fun r => AE.get r ae) args ->
- let (op', args') := make_cmp c args vl in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op' e##args' m = Some v
- /\ Val.lessdef (Val.of_optbool (eval_condition c e##args m)) v.
-Proof.
- intros c args vl.
- assert (Y: forall r, vincl (AE.get r ae) (Uns Ptop 1) = true ->
- e#r = Vundef \/ e#r = Vint Int.zero \/ e#r = Vint Int.one).
- { intros. apply vmatch_Uns_1 with bc Ptop. eapply vmatch_ge. eapply vincl_ge; eauto. apply MATCH. }
- unfold make_cmp. case (make_cmp_match c args vl); intros.
-- destruct (Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1)) eqn:E1.
- simpl in H; inv H. InvBooleans. subst n.
- exists (e#r1); split; auto. simpl.
- exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
- destruct (Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1)) eqn:E0.
- simpl in H; inv H. InvBooleans. subst n.
- exists (Val.xor e#r1 (Vint Int.one)); split; auto. simpl.
- exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
- apply make_cmp_base_correct; auto.
-- destruct (Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1)) eqn:E0.
- simpl in H; inv H. InvBooleans. subst n.
- exists (e#r1); split; auto. simpl.
- exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
- destruct (Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1)) eqn:E1.
- simpl in H; inv H. InvBooleans. subst n.
- exists (Val.xor e#r1 (Vint Int.one)); split; auto. simpl.
- exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
- apply make_cmp_base_correct; auto.
-- apply make_cmp_base_correct; auto.
-Qed.
-
-Lemma make_addimm_correct:
- forall n r,
- let (op, args) := make_addimm n r in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.add e#r (Vint n)) v.
-Proof.
- intros. unfold make_addimm.
- predSpec Int.eq Int.eq_spec n Int.zero; intros.
- subst. exists (e#r); split; auto.
- destruct (e#r); simpl; auto. rewrite Int.add_zero; auto. destruct Archi.ptr64; auto. rewrite Ptrofs.add_zero; auto.
- exists (Val.add e#r (Vint n)); split; auto. simpl. rewrite Int.repr_signed; auto.
-Qed.
-
-Lemma make_shlimm_correct:
- forall n r1 r2,
- e#r2 = Vint n ->
- let (op, args) := make_shlimm n r1 r2 in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.shl e#r1 (Vint n)) v.
-Proof.
- intros; unfold make_shlimm.
- predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int.shl_zero. auto.
- destruct (Int.ltu n Int.iwordsize).
- econstructor; split. simpl. eauto. auto.
- econstructor; split. simpl. eauto. rewrite H; auto.
-Qed.
-
-Lemma make_shrimm_correct:
- forall n r1 r2,
- e#r2 = Vint n ->
- let (op, args) := make_shrimm n r1 r2 in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.shr e#r1 (Vint n)) v.
-Proof.
- intros; unfold make_shrimm.
- predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int.shr_zero. auto.
- destruct (Int.ltu n Int.iwordsize).
- econstructor; split. simpl. eauto. auto.
- econstructor; split. simpl. eauto. rewrite H; auto.
-Qed.
-
-Lemma make_shruimm_correct:
- forall n r1 r2,
- e#r2 = Vint n ->
- let (op, args) := make_shruimm n r1 r2 in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.shru e#r1 (Vint n)) v.
-Proof.
- intros; unfold make_shruimm.
- predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int.shru_zero. auto.
- destruct (Int.ltu n Int.iwordsize).
- econstructor; split. simpl. eauto. auto.
- econstructor; split. simpl. eauto. rewrite H; auto.
-Qed.
-
-Lemma make_mulimm_correct:
- forall n r1,
- let (op, args) := make_mulimm n r1 in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.mul e#r1 (Vint n)) v.
-Proof.
- intros; unfold make_mulimm.
- predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (Vint Int.zero); split; auto. destruct (e#r1); simpl; auto. rewrite Int.mul_zero; auto.
- predSpec Int.eq Int.eq_spec n Int.one; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int.mul_one; auto.
- destruct (Int.is_power2 n) eqn:?; intros.
- rewrite (Val.mul_pow2 e#r1 _ _ Heqo). econstructor; split. simpl; eauto. auto.
- econstructor; split; eauto. auto.
-Qed.
-
-Lemma make_divimm_correct:
- forall n r1 r2 v,
- Val.divs e#r1 e#r2 = Some v ->
- e#r2 = Vint n ->
- let (op, args) := make_divimm n r1 r2 in
- exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
-Proof.
- intros; unfold make_divimm.
- destruct (Int.is_power2 n) eqn:?.
- destruct (Int.ltu i (Int.repr 31)) eqn:?.
- exists v; split; auto. simpl. eapply Val.divs_pow2; eauto. congruence.
- exists v; auto.
- exists v; auto.
-Qed.
-
-Lemma make_divuimm_correct:
- forall n r1 r2 v,
- Val.divu e#r1 e#r2 = Some v ->
- e#r2 = Vint n ->
- let (op, args) := make_divuimm n r1 r2 in
- exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
-Proof.
- intros; unfold make_divuimm.
- destruct (Int.is_power2 n) eqn:?.
- econstructor; split. simpl; eauto.
- rewrite H0 in H. erewrite Val.divu_pow2 by eauto. auto.
- exists v; auto.
-Qed.
-
-Lemma make_moduimm_correct:
- forall n r1 r2 v,
- Val.modu e#r1 e#r2 = Some v ->
- e#r2 = Vint n ->
- let (op, args) := make_moduimm n r1 r2 in
- exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
-Proof.
- intros; unfold make_moduimm.
- destruct (Int.is_power2 n) eqn:?.
- exists v; split; auto. simpl. decEq. eapply Val.modu_pow2; eauto. congruence.
- exists v; auto.
-Qed.
-
-Lemma make_andimm_correct:
- forall n r x,
- vmatch bc e#r x ->
- let (op, args) := make_andimm n r x in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.and e#r (Vint n)) v.
-Proof.
- intros; unfold make_andimm.
- predSpec Int.eq Int.eq_spec n Int.zero; intros.
- subst n. exists (Vint Int.zero); split; auto. destruct (e#r); simpl; auto. rewrite Int.and_zero; auto.
- predSpec Int.eq Int.eq_spec n Int.mone; intros.
- subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int.and_mone; auto.
- destruct (match x with Uns _ k => Int.eq (Int.zero_ext k (Int.not n)) Int.zero
- | _ => false end) eqn:UNS.
- destruct x; try congruence.
- exists (e#r); split; auto.
- inv H; auto. simpl. replace (Int.and i n) with i; auto.
- generalize (Int.eq_spec (Int.zero_ext n0 (Int.not n)) Int.zero); rewrite UNS; intro EQ.
- Int.bit_solve. destruct (zlt i0 n0).
- replace (Int.testbit n i0) with (negb (Int.testbit Int.zero i0)).
- rewrite Int.bits_zero. simpl. rewrite andb_true_r. auto.
- rewrite <- EQ. rewrite Int.bits_zero_ext by omega. rewrite zlt_true by auto.
- rewrite Int.bits_not by auto. apply negb_involutive.
- rewrite H6 by auto. auto.
- econstructor; split; eauto. auto.
-Qed.
-
-Lemma make_orimm_correct:
- forall n r,
- let (op, args) := make_orimm n r in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.or e#r (Vint n)) v.
-Proof.
- intros; unfold make_orimm.
- predSpec Int.eq Int.eq_spec n Int.zero; intros.
- subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int.or_zero; auto.
- predSpec Int.eq Int.eq_spec n Int.mone; intros.
- subst n. exists (Vint Int.mone); split; auto. destruct (e#r); simpl; auto. rewrite Int.or_mone; auto.
- econstructor; split; eauto. auto.
-Qed.
-
-Lemma make_xorimm_correct:
- forall n r,
- let (op, args) := make_xorimm n r in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.xor e#r (Vint n)) v.
-Proof.
- intros; unfold make_xorimm.
- predSpec Int.eq Int.eq_spec n Int.zero; intros.
- subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int.xor_zero; auto.
- predSpec Int.eq Int.eq_spec n Int.mone; intros.
- subst n. exists (Val.notint e#r); split; auto.
- econstructor; split; eauto. auto.
-Qed.
-
-Lemma make_addlimm_correct:
- forall n r,
- let (op, args) := make_addlimm n r in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.addl e#r (Vlong n)) v.
-Proof.
- intros. unfold make_addlimm.
- predSpec Int64.eq Int64.eq_spec n Int64.zero; intros.
- subst. exists (e#r); split; auto.
- destruct (e#r); simpl; auto. rewrite Int64.add_zero; auto. destruct Archi.ptr64; auto. rewrite Ptrofs.add_zero; auto.
- exists (Val.addl e#r (Vlong n)); split; auto. simpl. rewrite Int64.repr_signed; auto.
-Qed.
-
-Lemma make_shllimm_correct:
- forall n r1 r2,
- e#r2 = Vint n ->
- let (op, args) := make_shllimm n r1 r2 in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.shll e#r1 (Vint n)) v.
-Proof.
- intros; unfold make_shllimm.
- predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto.
- unfold Int64.shl'. rewrite Z.shiftl_0_r, Int64.repr_unsigned. auto.
- destruct (Int.ltu n Int64.iwordsize').
- econstructor; split. simpl. eauto. auto.
- econstructor; split. simpl. eauto. rewrite H; auto.
-Qed.
-
-Lemma make_shrlimm_correct:
- forall n r1 r2,
- e#r2 = Vint n ->
- let (op, args) := make_shrlimm n r1 r2 in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.shrl e#r1 (Vint n)) v.
-Proof.
- intros; unfold make_shrlimm.
- predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto.
- unfold Int64.shr'. rewrite Z.shiftr_0_r, Int64.repr_signed. auto.
- destruct (Int.ltu n Int64.iwordsize').
- econstructor; split. simpl. eauto. auto.
- econstructor; split. simpl. eauto. rewrite H; auto.
-Qed.
-
-Lemma make_shrluimm_correct:
- forall n r1 r2,
- e#r2 = Vint n ->
- let (op, args) := make_shrluimm n r1 r2 in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.shrlu e#r1 (Vint n)) v.
-Proof.
- intros; unfold make_shrluimm.
- predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto.
- unfold Int64.shru'. rewrite Z.shiftr_0_r, Int64.repr_unsigned. auto.
- destruct (Int.ltu n Int64.iwordsize').
- econstructor; split. simpl. eauto. auto.
- econstructor; split. simpl. eauto. rewrite H; auto.
-Qed.
-
-Lemma make_mullimm_correct:
- forall n r1,
- let (op, args) := make_mullimm n r1 in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.mull e#r1 (Vlong n)) v.
-Proof.
- intros; unfold make_mullimm.
- predSpec Int64.eq Int64.eq_spec n Int64.zero; intros. subst.
- exists (Vlong Int64.zero); split; auto. destruct (e#r1); simpl; auto. rewrite Int64.mul_zero; auto.
- predSpec Int64.eq Int64.eq_spec n Int64.one; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int64.mul_one; auto.
- destruct (Int64.is_power2' n) eqn:?; intros.
- exists (Val.shll e#r1 (Vint i)); split; auto.
- destruct (e#r1); simpl; auto.
- erewrite Int64.is_power2'_range by eauto.
- erewrite Int64.mul_pow2' by eauto. auto.
- econstructor; split; eauto. auto.
-Qed.
-
-Lemma make_divlimm_correct:
- forall n r1 r2 v,
- Val.divls e#r1 e#r2 = Some v ->
- e#r2 = Vlong n ->
- let (op, args) := make_divlimm n r1 r2 in
- exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
-Proof.
- intros; unfold make_divlimm.
- destruct (Int64.is_power2' n) eqn:?. destruct (Int.ltu i (Int.repr 63)) eqn:?.
- rewrite H0 in H. econstructor; split. simpl; eauto. eapply Val.divls_pow2; eauto. auto.
- exists v; auto.
- exists v; auto.
-Qed.
-
-Lemma make_divluimm_correct:
- forall n r1 r2 v,
- Val.divlu e#r1 e#r2 = Some v ->
- e#r2 = Vlong n ->
- let (op, args) := make_divluimm n r1 r2 in
- exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
-Proof.
- intros; unfold make_divluimm.
- destruct (Int64.is_power2' n) eqn:?.
- econstructor; split. simpl; eauto.
- rewrite H0 in H. destruct (e#r1); inv H. destruct (Int64.eq n Int64.zero); inv H2.
- simpl.
- erewrite Int64.is_power2'_range by eauto.
- erewrite Int64.divu_pow2' by eauto. auto.
- exists v; auto.
-Qed.
-
-Lemma make_modluimm_correct:
- forall n r1 r2 v,
- Val.modlu e#r1 e#r2 = Some v ->
- e#r2 = Vlong n ->
- let (op, args) := make_modluimm n r1 r2 in
- exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
-Proof.
- intros; unfold make_modluimm.
- destruct (Int64.is_power2 n) eqn:?.
- exists v; split; auto. simpl. decEq.
- rewrite H0 in H. destruct (e#r1); inv H. destruct (Int64.eq n Int64.zero); inv H2.
- simpl. erewrite Int64.modu_and by eauto. auto.
- exists v; auto.
-Qed.
-
-Lemma make_andlimm_correct:
- forall n r x,
- let (op, args) := make_andlimm n r x in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.andl e#r (Vlong n)) v.
-Proof.
- intros; unfold make_andlimm.
- predSpec Int64.eq Int64.eq_spec n Int64.zero; intros.
- subst n. exists (Vlong Int64.zero); split; auto. destruct (e#r); simpl; auto. rewrite Int64.and_zero; auto.
- predSpec Int64.eq Int64.eq_spec n Int64.mone; intros.
- subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int64.and_mone; auto.
- econstructor; split; eauto. auto.
-Qed.
-
-Lemma make_orlimm_correct:
- forall n r,
- let (op, args) := make_orlimm n r in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.orl e#r (Vlong n)) v.
-Proof.
- intros; unfold make_orlimm.
- predSpec Int64.eq Int64.eq_spec n Int64.zero; intros.
- subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int64.or_zero; auto.
- predSpec Int64.eq Int64.eq_spec n Int64.mone; intros.
- subst n. exists (Vlong Int64.mone); split; auto. destruct (e#r); simpl; auto. rewrite Int64.or_mone; auto.
- econstructor; split; eauto. auto.
-Qed.
-
-Lemma make_xorlimm_correct:
- forall n r,
- let (op, args) := make_xorlimm n r in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.xorl e#r (Vlong n)) v.
-Proof.
- intros; unfold make_xorlimm.
- predSpec Int64.eq Int64.eq_spec n Int64.zero; intros.
- subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int64.xor_zero; auto.
- predSpec Int64.eq Int64.eq_spec n Int64.mone; intros.
- subst n. exists (Val.notl e#r); split; auto.
- econstructor; split; eauto. auto.
-Qed.
-
-Lemma make_mulfimm_correct:
- forall n r1 r2,
- e#r2 = Vfloat n ->
- let (op, args) := make_mulfimm n r1 r1 r2 in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.mulf e#r1 e#r2) v.
-Proof.
- intros; unfold make_mulfimm.
- destruct (Float.eq_dec n (Float.of_int (Int.repr 2))); intros.
- simpl. econstructor; split. eauto. rewrite H; subst n.
- destruct (e#r1); simpl; auto. rewrite Float.mul2_add; auto.
- simpl. econstructor; split; eauto.
-Qed.
-
-Lemma make_mulfimm_correct_2:
- forall n r1 r2,
- e#r1 = Vfloat n ->
- let (op, args) := make_mulfimm n r2 r1 r2 in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.mulf e#r1 e#r2) v.
-Proof.
- intros; unfold make_mulfimm.
- destruct (Float.eq_dec n (Float.of_int (Int.repr 2))); intros.
- simpl. econstructor; split. eauto. rewrite H; subst n.
- destruct (e#r2); simpl; auto. rewrite Float.mul2_add; auto.
- rewrite Float.mul_commut; auto.
- simpl. econstructor; split; eauto.
-Qed.
-
-Lemma make_mulfsimm_correct:
- forall n r1 r2,
- e#r2 = Vsingle n ->
- let (op, args) := make_mulfsimm n r1 r1 r2 in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.mulfs e#r1 e#r2) v.
-Proof.
- intros; unfold make_mulfsimm.
- destruct (Float32.eq_dec n (Float32.of_int (Int.repr 2))); intros.
- simpl. econstructor; split. eauto. rewrite H; subst n.
- destruct (e#r1); simpl; auto. rewrite Float32.mul2_add; auto.
- simpl. econstructor; split; eauto.
-Qed.
-
-Lemma make_mulfsimm_correct_2:
- forall n r1 r2,
- e#r1 = Vsingle n ->
- let (op, args) := make_mulfsimm n r2 r1 r2 in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.mulfs e#r1 e#r2) v.
-Proof.
- intros; unfold make_mulfsimm.
- destruct (Float32.eq_dec n (Float32.of_int (Int.repr 2))); intros.
- simpl. econstructor; split. eauto. rewrite H; subst n.
- destruct (e#r2); simpl; auto. rewrite Float32.mul2_add; auto.
- rewrite Float32.mul_commut; auto.
- simpl. econstructor; split; eauto.
-Qed.
-
-Lemma make_cast8signed_correct:
- forall r x,
- vmatch bc e#r x ->
- let (op, args) := make_cast8signed r x in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.sign_ext 8 e#r) v.
-Proof.
- intros; unfold make_cast8signed. destruct (vincl x (Sgn Ptop 8)) eqn:INCL.
- exists e#r; split; auto.
- assert (V: vmatch bc e#r (Sgn Ptop 8)).
- { eapply vmatch_ge; eauto. apply vincl_ge; auto. }
- inv V; simpl; auto. rewrite is_sgn_sign_ext in H4 by auto. rewrite H4; auto.
- econstructor; split; simpl; eauto.
-Qed.
-
-Lemma make_cast8unsigned_correct:
- forall r x,
- vmatch bc e#r x ->
- let (op, args) := make_cast8unsigned r x in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.zero_ext 8 e#r) v.
-Proof.
- intros; unfold make_cast8unsigned. destruct (vincl x (Uns Ptop 8)) eqn:INCL.
- exists e#r; split; auto.
- assert (V: vmatch bc e#r (Uns Ptop 8)).
- { eapply vmatch_ge; eauto. apply vincl_ge; auto. }
- inv V; simpl; auto. rewrite is_uns_zero_ext in H4 by auto. rewrite H4; auto.
- econstructor; split; simpl; eauto.
-Qed.
-
-Lemma make_cast16signed_correct:
- forall r x,
- vmatch bc e#r x ->
- let (op, args) := make_cast16signed r x in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.sign_ext 16 e#r) v.
-Proof.
- intros; unfold make_cast16signed. destruct (vincl x (Sgn Ptop 16)) eqn:INCL.
- exists e#r; split; auto.
- assert (V: vmatch bc e#r (Sgn Ptop 16)).
- { eapply vmatch_ge; eauto. apply vincl_ge; auto. }
- inv V; simpl; auto. rewrite is_sgn_sign_ext in H4 by auto. rewrite H4; auto.
- econstructor; split; simpl; eauto.
-Qed.
-
-Lemma make_cast16unsigned_correct:
- forall r x,
- vmatch bc e#r x ->
- let (op, args) := make_cast16unsigned r x in
- exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.zero_ext 16 e#r) v.
-Proof.
- intros; unfold make_cast16unsigned. destruct (vincl x (Uns Ptop 16)) eqn:INCL.
- exists e#r; split; auto.
- assert (V: vmatch bc e#r (Uns Ptop 16)).
- { eapply vmatch_ge; eauto. apply vincl_ge; auto. }
- inv V; simpl; auto. rewrite is_uns_zero_ext in H4 by auto. rewrite H4; auto.
- econstructor; split; simpl; eauto.
-Qed.
-
-Lemma op_strength_reduction_correct:
- forall op args vl v,
- vl = map (fun r => AE.get r ae) args ->
- eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v ->
- let (op', args') := op_strength_reduction op args vl in
- exists w, eval_operation ge (Vptr sp Ptrofs.zero) op' e##args' m = Some w /\ Val.lessdef v w.
-Proof.
- intros until v; unfold op_strength_reduction;
- case (op_strength_reduction_match op args vl); simpl; intros.
-(* cast8signed *)
- InvApproxRegs; SimplVM; inv H0. apply make_cast8signed_correct; auto.
-(* cast8unsigned *)
- InvApproxRegs; SimplVM; inv H0. apply make_cast8unsigned_correct; auto.
-(* cast16signed *)
- InvApproxRegs; SimplVM; inv H0. apply make_cast16signed_correct; auto.
-(* cast16unsigned *)
- InvApproxRegs; SimplVM; inv H0. apply make_cast16unsigned_correct; auto.
-(* sub *)
- InvApproxRegs; SimplVM; inv H0. rewrite Val.sub_add_opp. apply make_addimm_correct; auto.
-(* mul *)
- rewrite Val.mul_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_mulimm_correct; auto.
- InvApproxRegs; SimplVM; inv H0. apply make_mulimm_correct; auto.
-(* divs *)
- assert (e#r2 = Vint n2). clear H0. InvApproxRegs; SimplVM; auto.
- apply make_divimm_correct; auto.
-(* divu *)
- assert (e#r2 = Vint n2). clear H0. InvApproxRegs; SimplVM; auto.
- apply make_divuimm_correct; auto.
-(* modu *)
- assert (e#r2 = Vint n2). clear H0. InvApproxRegs; SimplVM; auto.
- apply make_moduimm_correct; auto.
-(* and *)
- rewrite Val.and_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_andimm_correct; auto.
- InvApproxRegs; SimplVM; inv H0. apply make_andimm_correct; auto.
- inv H; inv H0. apply make_andimm_correct; auto.
-(* or *)
- rewrite Val.or_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_orimm_correct; auto.
- InvApproxRegs; SimplVM; inv H0. apply make_orimm_correct; auto.
-(* xor *)
- rewrite Val.xor_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_xorimm_correct; auto.
- InvApproxRegs; SimplVM; inv H0. apply make_xorimm_correct; auto.
-(* shl *)
- InvApproxRegs; SimplVM; inv H0. apply make_shlimm_correct; auto.
-(* shr *)
- InvApproxRegs; SimplVM; inv H0. apply make_shrimm_correct; auto.
-(* shru *)
- InvApproxRegs; SimplVM; inv H0. apply make_shruimm_correct; auto.
-(* lea *)
- exploit addr_strength_reduction_32_correct; eauto.
- destruct (addr_strength_reduction_32 addr args0 vl0) as [addr' args'].
- auto.
-(* subl *)
- InvApproxRegs; SimplVM; inv H0.
- replace (Val.subl e#r1 (Vlong n2)) with (Val.addl e#r1 (Vlong (Int64.neg n2))).
- apply make_addlimm_correct; auto.
- destruct (e#r1); simpl; auto.
- rewrite Int64.sub_add_opp; auto.
- destruct Archi.ptr64 eqn:SF; auto.
- rewrite Ptrofs.sub_add_opp. do 2 f_equal. auto with ptrofs.
-(* mull *)
- rewrite Val.mull_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_mullimm_correct; auto.
- InvApproxRegs; SimplVM; inv H0. apply make_mullimm_correct; auto.
-(* divl *)
- assert (e#r2 = Vlong n2). clear H0. InvApproxRegs; SimplVM; auto.
- apply make_divlimm_correct; auto.
-(* divlu *)
- assert (e#r2 = Vlong n2). clear H0. InvApproxRegs; SimplVM; auto.
- apply make_divluimm_correct; auto.
-(* modlu *)
- assert (e#r2 = Vlong n2). clear H0. InvApproxRegs; SimplVM; auto.
- apply make_modluimm_correct; auto.
-(* andl *)
- rewrite Val.andl_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_andlimm_correct; auto.
- InvApproxRegs; SimplVM; inv H0. apply make_andlimm_correct; auto.
- inv H; inv H0. apply make_andlimm_correct; auto.
-(* orl *)
- rewrite Val.orl_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_orlimm_correct; auto.
- InvApproxRegs; SimplVM; inv H0. apply make_orlimm_correct; auto.
-(* xorl *)
- rewrite Val.xorl_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_xorlimm_correct; auto.
- InvApproxRegs; SimplVM; inv H0. apply make_xorlimm_correct; auto.
-(* shll *)
- InvApproxRegs; SimplVM; inv H0. apply make_shllimm_correct; auto.
-(* shrl *)
- InvApproxRegs; SimplVM; inv H0. apply make_shrlimm_correct; auto.
-(* shrlu *)
- InvApproxRegs; SimplVM; inv H0. apply make_shrluimm_correct; auto.
-(* leal *)
- exploit addr_strength_reduction_64_correct; eauto.
- destruct (addr_strength_reduction_64 addr args0 vl0) as [addr' args'].
- auto.
-(* cond *)
- inv H0. apply make_cmp_correct; auto.
-(* mulf *)
- InvApproxRegs; SimplVM; inv H0. rewrite <- H2. apply make_mulfimm_correct; auto.
- InvApproxRegs; SimplVM; inv H0. fold (Val.mulf (Vfloat n1) e#r2).
- rewrite <- H2. apply make_mulfimm_correct_2; auto.
-(* mulfs *)
- InvApproxRegs; SimplVM; inv H0. rewrite <- H2. apply make_mulfsimm_correct; auto.
- InvApproxRegs; SimplVM; inv H0. fold (Val.mulfs (Vsingle n1) e#r2).
- rewrite <- H2. apply make_mulfsimm_correct_2; auto.
-(* default *)
- exists v; auto.
-Qed.
-
-End STRENGTH_REDUCTION.
diff --git a/ia32/Conventions1.v b/ia32/Conventions1.v
deleted file mode 100644
index dbc8b064..00000000
--- a/ia32/Conventions1.v
+++ /dev/null
@@ -1,473 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Function calling conventions and other conventions regarding the use of
- machine registers and stack slots. *)
-
-Require Import Coqlib Decidableplus.
-Require Import AST Machregs Locations.
-
-(** * Classification of machine registers *)
-
-(** Machine registers (type [mreg] in module [Locations]) are divided in
- the following groups:
-- Callee-save registers, whose value is preserved across a function call.
-- Caller-save registers that can be modified during a function call.
-
- We follow the x86-32 and x86-64 application binary interfaces (ABI)
- in our choice of callee- and caller-save registers.
-*)
-
-Definition is_callee_save (r: mreg) : bool :=
- match r with
- | AX | CX | DX => false
- | BX | BP => true
- | SI | DI => negb Archi.ptr64 (**r callee-save in 32 bits but not in 64 bits *)
- | R8 | R9 | R10 | R11 => false
- | R12 | R13 | R14 | R15 => true
- | X0 | X1 | X2 | X3 | X4 | X5 | X6 | X7 => false
- | X8 | X9 | X10 | X11 | X12 | X13 | X14 | X15 => false
- | FP0 => false
- end.
-
-Definition int_caller_save_regs :=
- if Archi.ptr64
- then AX :: CX :: DX :: SI :: DI :: R8 :: R9 :: R10 :: R11 :: nil
- else AX :: CX :: DX :: nil.
-
-Definition float_caller_save_regs :=
- if Archi.ptr64
- then X0 :: X1 :: X2 :: X3 :: X4 :: X5 :: X6 :: X7 ::
- X8 :: X9 :: X10 :: X11 :: X12 :: X13 :: X14 :: X15 :: nil
- else X0 :: X1 :: X2 :: X3 :: X4 :: X5 :: X6 :: X7 :: nil.
-
-Definition int_callee_save_regs :=
- if Archi.ptr64
- then BX :: BP :: R12 :: R13 :: R14 :: R15 :: nil
- else BX :: SI :: DI :: BP :: nil.
-
-Definition float_callee_save_regs : list mreg := nil.
-
-Definition destroyed_at_call :=
- List.filter (fun r => negb (is_callee_save r)) all_mregs.
-
-Definition dummy_int_reg := AX. (**r Used in [Regalloc]. *)
-Definition dummy_float_reg := X0. (**r Used in [Regalloc]. *)
-
-Definition is_float_reg (r: mreg) :=
- match r with
- | AX | BX | CX | DX | SI | DI | BP
- | R8 | R9 | R10 | R11 | R12 | R13 | R14 | R15 => false
- | X0 | X1 | X2 | X3 | X4 | X5 | X6 | X7
- | X8 | X9 | X10 | X11 | X12 | X13 | X14 | X15 | FP0 => true
- end.
-
-(** * Function calling conventions *)
-
-(** The functions in this section determine the locations (machine registers
- and stack slots) used to communicate arguments and results between the
- caller and the callee during function calls. These locations are functions
- of the signature of the function and of the call instruction.
- Agreement between the caller and the callee on the locations to use
- is guaranteed by our dynamic semantics for Cminor and RTL, which demand
- that the signature of the call instruction is identical to that of the
- called function.
-
- Calling conventions are largely arbitrary: they must respect the properties
- proved in this section (such as no overlapping between the locations
- of function arguments), but this leaves much liberty in choosing actual
- locations. To ensure binary interoperability of code generated by our
- compiler with libraries compiled by another compiler, we
- implement the standard x86-32 and x86-64 conventions. *)
-
-(** ** Location of function result *)
-
-(** In 32 bit mode, the result value of a function is passed back to the
- caller in registers [AX] or [DX:AX] or [FP0], depending on the type
- of the returned value. We treat a function without result as a
- function with one integer result. *)
-
-Definition loc_result_32 (s: signature) : rpair mreg :=
- match s.(sig_res) with
- | None => One AX
- | Some (Tint | Tany32) => One AX
- | Some (Tfloat | Tsingle) => One FP0
- | Some Tany64 => One X0
- | Some Tlong => Twolong DX AX
- end.
-
-(** In 64 bit mode, he result value of a function is passed back to
- the caller in registers [AX] or [X0]. *)
-
-Definition loc_result_64 (s: signature) : rpair mreg :=
- match s.(sig_res) with
- | None => One AX
- | Some (Tint | Tlong | Tany32 | Tany64) => One AX
- | Some (Tfloat | Tsingle) => One X0
- end.
-
-Definition loc_result :=
- if Archi.ptr64 then loc_result_64 else loc_result_32.
-
-(** The result registers have types compatible with that given in the signature. *)
-
-Lemma loc_result_type:
- forall sig,
- subtype (proj_sig_res sig) (typ_rpair mreg_type (loc_result sig)) = true.
-Proof.
- intros. unfold proj_sig_res, loc_result, loc_result_32, loc_result_64, mreg_type;
- destruct Archi.ptr64; destruct (sig_res sig) as [[]|]; auto.
-Qed.
-
-(** The result locations are caller-save registers *)
-
-Lemma loc_result_caller_save:
- forall (s: signature),
- forall_rpair (fun r => is_callee_save r = false) (loc_result s).
-Proof.
- intros. unfold loc_result, loc_result_32, loc_result_64, is_callee_save;
- destruct Archi.ptr64; destruct (sig_res s) as [[]|]; simpl; auto.
-Qed.
-
-(** If the result is in a pair of registers, those registers are distinct and have type [Tint] at least. *)
-
-Lemma loc_result_pair:
- forall sg,
- match loc_result sg with
- | One _ => True
- | Twolong r1 r2 =>
- r1 <> r2 /\ sg.(sig_res) = Some Tlong
- /\ subtype Tint (mreg_type r1) = true /\ subtype Tint (mreg_type r2) = true
- /\ Archi.splitlong = true
- end.
-Proof.
- intros. change Archi.splitlong with (negb Archi.ptr64).
- unfold loc_result, loc_result_32, loc_result_64, mreg_type;
- destruct Archi.ptr64; destruct (sig_res sg) as [[]|]; auto.
- split; auto. congruence.
-Qed.
-
-(** The location of the result depends only on the result part of the signature *)
-
-Lemma loc_result_exten:
- forall s1 s2, s1.(sig_res) = s2.(sig_res) -> loc_result s1 = loc_result s2.
-Proof.
- intros. unfold loc_result, loc_result_32, loc_result_64.
- destruct Archi.ptr64; rewrite H; auto.
-Qed.
-
-(** ** Location of function arguments *)
-
-(** In the x86-32 ABI, all arguments are passed on stack. (Snif.) *)
-
-Fixpoint loc_arguments_32
- (tyl: list typ) (ofs: Z) {struct tyl} : list (rpair loc) :=
- match tyl with
- | nil => nil
- | ty :: tys =>
- match ty with
- | Tlong => Twolong (S Outgoing (ofs + 1) Tint) (S Outgoing ofs Tint)
- | _ => One (S Outgoing ofs ty)
- end
- :: loc_arguments_32 tys (ofs + typesize ty)
- end.
-
-(** In the x86-64 ABI:
-- The first 6 integer arguments are passed in registers [DI], [SI], [DX], [CX], [R8], [R9].
-- The first 8 floating-point arguments are passed in registers [X0] to [X7].
-- Extra arguments are passed on the stack, in [Outgoing] slots.
- Consecutive stack slots are separated by 8 bytes, even if only 4 bytes
- of data is used in a slot.
-*)
-
-Definition int_param_regs := DI :: SI :: DX :: CX :: R8 :: R9 :: nil.
-Definition float_param_regs := X0 :: X1 :: X2 :: X3 :: X4 :: X5 :: X6 :: X7 :: nil.
-
-Fixpoint loc_arguments_64
- (tyl: list typ) (ir fr ofs: Z) {struct tyl} : list (rpair loc) :=
- match tyl with
- | nil => nil
- | (Tint | Tlong | Tany32 | Tany64) as ty :: tys =>
- match list_nth_z int_param_regs ir with
- | None =>
- One (S Outgoing ofs ty) :: loc_arguments_64 tys ir fr (ofs + 2)
- | Some ireg =>
- One (R ireg) :: loc_arguments_64 tys (ir + 1) fr ofs
- end
- | (Tfloat | Tsingle) as ty :: tys =>
- match list_nth_z float_param_regs fr with
- | None =>
- One (S Outgoing ofs ty) :: loc_arguments_64 tys ir fr (ofs + 2)
- | Some freg =>
- One (R freg) :: loc_arguments_64 tys ir (fr + 1) ofs
- end
- end.
-
-(** [loc_arguments s] returns the list of locations where to store arguments
- when calling a function with signature [s]. *)
-
-Definition loc_arguments (s: signature) : list (rpair loc) :=
- if Archi.ptr64
- then loc_arguments_64 s.(sig_args) 0 0 0
- else loc_arguments_32 s.(sig_args) 0.
-
-(** [size_arguments s] returns the number of [Outgoing] slots used
- to call a function with signature [s]. *)
-
-Fixpoint size_arguments_32
- (tyl: list typ) (ofs: Z) {struct tyl} : Z :=
- match tyl with
- | nil => ofs
- | ty :: tys => size_arguments_32 tys (ofs + typesize ty)
- end.
-
-Fixpoint size_arguments_64 (tyl: list typ) (ir fr ofs: Z) {struct tyl} : Z :=
- match tyl with
- | nil => ofs
- | (Tint | Tlong | Tany32 | Tany64) :: tys =>
- match list_nth_z int_param_regs ir with
- | None => size_arguments_64 tys ir fr (ofs + 2)
- | Some ireg => size_arguments_64 tys (ir + 1) fr ofs
- end
- | (Tfloat | Tsingle) :: tys =>
- match list_nth_z float_param_regs fr with
- | None => size_arguments_64 tys ir fr (ofs + 2)
- | Some freg => size_arguments_64 tys ir (fr + 1) ofs
- end
- end.
-
-Definition size_arguments (s: signature) : Z :=
- if Archi.ptr64
- then size_arguments_64 s.(sig_args) 0 0 0
- else size_arguments_32 s.(sig_args) 0.
-
-(** Argument locations are either caller-save registers or [Outgoing]
- stack slots at nonnegative offsets. *)
-
-Definition loc_argument_acceptable (l: loc) : Prop :=
- match l with
- | R r => is_callee_save r = false
- | S Outgoing ofs ty => ofs >= 0 /\ (typealign ty | ofs)
- | _ => False
- end.
-
-Definition loc_argument_32_charact (ofs: Z) (l: loc) : Prop :=
- match l with
- | S Outgoing ofs' ty => ofs' >= ofs /\ typealign ty = 1
- | _ => False
- end.
-
-Definition loc_argument_64_charact (ofs: Z) (l: loc) : Prop :=
- match l with
- | R r => In r int_param_regs \/ In r float_param_regs
- | S Outgoing ofs' ty => ofs' >= ofs /\ (2 | ofs')
- | _ => False
- end.
-
-Remark loc_arguments_32_charact:
- forall tyl ofs p,
- In p (loc_arguments_32 tyl ofs) -> forall_rpair (loc_argument_32_charact ofs) p.
-Proof.
- assert (X: forall ofs1 ofs2 l, loc_argument_32_charact ofs2 l -> ofs1 <= ofs2 -> loc_argument_32_charact ofs1 l).
- { destruct l; simpl; intros; auto. destruct sl; auto. intuition omega. }
- induction tyl as [ | ty tyl]; simpl loc_arguments_32; intros.
-- contradiction.
-- destruct H.
-+ destruct ty; subst p; simpl; omega.
-+ apply IHtyl in H. generalize (typesize_pos ty); intros. destruct p; simpl in *.
-* eapply X; eauto; omega.
-* destruct H; split; eapply X; eauto; omega.
-Qed.
-
-Remark loc_arguments_64_charact:
- forall tyl ir fr ofs p,
- In p (loc_arguments_64 tyl ir fr ofs) -> (2 | ofs) -> forall_rpair (loc_argument_64_charact ofs) p.
-Proof.
- assert (X: forall ofs1 ofs2 l, loc_argument_64_charact ofs2 l -> ofs1 <= ofs2 -> loc_argument_64_charact ofs1 l).
- { destruct l; simpl; intros; auto. destruct sl; auto. intuition omega. }
- assert (Y: forall ofs1 ofs2 p, forall_rpair (loc_argument_64_charact ofs2) p -> ofs1 <= ofs2 -> forall_rpair (loc_argument_64_charact ofs1) p).
- { destruct p; simpl; intuition eauto. }
- assert (Z: forall ofs, (2 | ofs) -> (2 | ofs + 2)).
- { intros. apply Z.divide_add_r; auto. apply Zdivide_refl. }
-Opaque list_nth_z.
- induction tyl; simpl loc_arguments_64; intros.
- elim H.
- assert (A: forall ty, In p
- match list_nth_z int_param_regs ir with
- | Some ireg => One (R ireg) :: loc_arguments_64 tyl (ir + 1) fr ofs
- | None => One (S Outgoing ofs ty) :: loc_arguments_64 tyl ir fr (ofs + 2)
- end ->
- forall_rpair (loc_argument_64_charact ofs) p).
- { intros. destruct (list_nth_z int_param_regs ir) as [r|] eqn:E; destruct H1.
- subst. left. eapply list_nth_z_in; eauto.
- eapply IHtyl; eauto.
- subst. split. omega. assumption.
- eapply Y; eauto. omega. }
- assert (B: forall ty, In p
- match list_nth_z float_param_regs fr with
- | Some ireg => One (R ireg) :: loc_arguments_64 tyl ir (fr + 1) ofs
- | None => One (S Outgoing ofs ty) :: loc_arguments_64 tyl ir fr (ofs + 2)
- end ->
- forall_rpair (loc_argument_64_charact ofs) p).
- { intros. destruct (list_nth_z float_param_regs fr) as [r|] eqn:E; destruct H1.
- subst. right. eapply list_nth_z_in; eauto.
- eapply IHtyl; eauto.
- subst. split. omega. assumption.
- eapply Y; eauto. omega. }
- destruct a; eauto.
-Qed.
-
-Lemma loc_arguments_acceptable:
- forall (s: signature) (p: rpair loc),
- In p (loc_arguments s) -> forall_rpair loc_argument_acceptable p.
-Proof.
- unfold loc_arguments; intros. destruct Archi.ptr64 eqn:SF.
-- (* 64 bits *)
- assert (A: forall r, In r int_param_regs -> is_callee_save r = false) by (unfold is_callee_save; rewrite SF; decide_goal).
- assert (B: forall r, In r float_param_regs -> is_callee_save r = false) by decide_goal.
- assert (X: forall l, loc_argument_64_charact 0 l -> loc_argument_acceptable l).
- { unfold loc_argument_64_charact, loc_argument_acceptable.
- destruct l as [r | [] ofs ty]; auto. intros [C|C]; auto.
- intros [C D]. split; auto. apply Zdivide_trans with 2; auto.
- exists (2 / typealign ty); destruct ty; reflexivity.
- }
- exploit loc_arguments_64_charact; eauto using Zdivide_0.
- unfold forall_rpair; destruct p; intuition auto.
-- (* 32 bits *)
- assert (X: forall l, loc_argument_32_charact 0 l -> loc_argument_acceptable l).
- { destruct l as [r | [] ofs ty]; simpl; intuition auto. rewrite H2; apply Z.divide_1_l. }
- exploit loc_arguments_32_charact; eauto.
- unfold forall_rpair; destruct p; intuition auto.
-Qed.
-
-Hint Resolve loc_arguments_acceptable: locs.
-
-(** The offsets of [Outgoing] arguments are below [size_arguments s]. *)
-
-Remark size_arguments_32_above:
- forall tyl ofs0, ofs0 <= size_arguments_32 tyl ofs0.
-Proof.
- induction tyl; simpl; intros.
- omega.
- apply Zle_trans with (ofs0 + typesize a); auto.
- generalize (typesize_pos a); omega.
-Qed.
-
-Remark size_arguments_64_above:
- forall tyl ir fr ofs0,
- ofs0 <= size_arguments_64 tyl ir fr ofs0.
-Proof.
- induction tyl; simpl; intros.
- omega.
- assert (A: ofs0 <=
- match list_nth_z int_param_regs ir with
- | Some _ => size_arguments_64 tyl (ir + 1) fr ofs0
- | None => size_arguments_64 tyl ir fr (ofs0 + 2)
- end).
- { destruct (list_nth_z int_param_regs ir); eauto.
- apply Zle_trans with (ofs0 + 2); auto. omega. }
- assert (B: ofs0 <=
- match list_nth_z float_param_regs fr with
- | Some _ => size_arguments_64 tyl ir (fr + 1) ofs0
- | None => size_arguments_64 tyl ir fr (ofs0 + 2)
- end).
- { destruct (list_nth_z float_param_regs fr); eauto.
- apply Zle_trans with (ofs0 + 2); auto. omega. }
- destruct a; auto.
-Qed.
-
-Lemma size_arguments_above:
- forall s, size_arguments s >= 0.
-Proof.
- intros; unfold size_arguments. apply Zle_ge.
- destruct Archi.ptr64; [apply size_arguments_64_above|apply size_arguments_32_above].
-Qed.
-
-Lemma loc_arguments_32_bounded:
- forall ofs ty tyl ofs0,
- In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments_32 tyl ofs0)) ->
- ofs + typesize ty <= size_arguments_32 tyl ofs0.
-Proof.
- induction tyl as [ | t l]; simpl; intros x IN.
-- contradiction.
-- rewrite in_app_iff in IN; destruct IN as [IN|IN].
-+ apply Zle_trans with (x + typesize t); [|apply size_arguments_32_above].
- Ltac decomp :=
- match goal with
- | [ H: _ \/ _ |- _ ] => destruct H; decomp
- | [ H: S _ _ _ = S _ _ _ |- _ ] => inv H
- | [ H: False |- _ ] => contradiction
- end.
- destruct t; simpl in IN; decomp; simpl; omega.
-+ apply IHl; auto.
-Qed.
-
-Lemma loc_arguments_64_bounded:
- forall ofs ty tyl ir fr ofs0,
- In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments_64 tyl ir fr ofs0)) ->
- ofs + typesize ty <= size_arguments_64 tyl ir fr ofs0.
-Proof.
- induction tyl; simpl; intros.
- contradiction.
- assert (T: forall ty0, typesize ty0 <= 2).
- { destruct ty0; simpl; omega. }
- assert (A: forall ty0,
- In (S Outgoing ofs ty) (regs_of_rpairs
- match list_nth_z int_param_regs ir with
- | Some ireg =>
- One (R ireg) :: loc_arguments_64 tyl (ir + 1) fr ofs0
- | None => One (S Outgoing ofs0 ty0) :: loc_arguments_64 tyl ir fr (ofs0 + 2)
- end) ->
- ofs + typesize ty <=
- match list_nth_z int_param_regs ir with
- | Some _ => size_arguments_64 tyl (ir + 1) fr ofs0
- | None => size_arguments_64 tyl ir fr (ofs0 + 2)
- end).
- { intros. destruct (list_nth_z int_param_regs ir); simpl in H0; destruct H0.
- - discriminate.
- - eapply IHtyl; eauto.
- - inv H0. apply Zle_trans with (ofs + 2). specialize (T ty). omega. apply size_arguments_64_above.
- - eapply IHtyl; eauto. }
- assert (B: forall ty0,
- In (S Outgoing ofs ty) (regs_of_rpairs
- match list_nth_z float_param_regs fr with
- | Some ireg =>
- One (R ireg) :: loc_arguments_64 tyl ir (fr + 1) ofs0
- | None => One (S Outgoing ofs0 ty0) :: loc_arguments_64 tyl ir fr (ofs0 + 2)
- end) ->
- ofs + typesize ty <=
- match list_nth_z float_param_regs fr with
- | Some _ => size_arguments_64 tyl ir (fr + 1) ofs0
- | None => size_arguments_64 tyl ir fr (ofs0 + 2)
- end).
- { intros. destruct (list_nth_z float_param_regs fr); simpl in H0; destruct H0.
- - discriminate.
- - eapply IHtyl; eauto.
- - inv H0. apply Zle_trans with (ofs + 2). specialize (T ty). omega. apply size_arguments_64_above.
- - eapply IHtyl; eauto. }
- destruct a; eauto.
-Qed.
-
-Lemma loc_arguments_bounded:
- forall (s: signature) (ofs: Z) (ty: typ),
- In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments s)) ->
- ofs + typesize ty <= size_arguments s.
-Proof.
- unfold loc_arguments, size_arguments; intros.
- destruct Archi.ptr64; eauto using loc_arguments_32_bounded, loc_arguments_64_bounded.
-Qed.
-
-Lemma loc_arguments_main:
- loc_arguments signature_main = nil.
-Proof.
- unfold loc_arguments; destruct Archi.ptr64; reflexivity.
-Qed.
diff --git a/ia32/Machregs.v b/ia32/Machregs.v
deleted file mode 100644
index 741081a6..00000000
--- a/ia32/Machregs.v
+++ /dev/null
@@ -1,367 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-Require Import String.
-Require Import Coqlib.
-Require Import Decidableplus.
-Require Import Maps.
-Require Import AST.
-Require Import Integers.
-Require Import Op.
-
-(** ** Machine registers *)
-
-(** The following type defines the machine registers that can be referenced
- as locations. These include:
-- Integer registers that can be allocated to RTL pseudo-registers.
-- Floating-point registers that can be allocated to RTL pseudo-registers.
-- The special [FP0] register denoting the top of the X87 float stack.
-
- The type [mreg] does not include special-purpose or reserved
- machine registers such as the stack pointer and the condition codes. *)
-
-Inductive mreg: Type :=
- (** Allocatable integer regs *)
- | AX | BX | CX | DX | SI | DI | BP
- | R8 | R9 | R10 | R11 | R12 | R13 | R14 | R15 (**r only in 64-bit mode *)
- (** Allocatable float regs *)
- | X0 | X1 | X2 | X3 | X4 | X5 | X6 | X7
- | X8 | X9 | X10 | X11 | X12 | X13 | X14 | X15 (**r only in 64-bit mode *)
- (** Special float reg *)
- | FP0.
-
-Lemma mreg_eq: forall (r1 r2: mreg), {r1 = r2} + {r1 <> r2}.
-Proof. decide equality. Defined.
-Global Opaque mreg_eq.
-
-Definition all_mregs :=
- AX :: BX :: CX :: DX :: SI :: DI :: BP
- :: R8 :: R9 :: R10 :: R11 :: R12 :: R13 :: R14 :: R15
- :: X0 :: X1 :: X2 :: X3 :: X4 :: X5 :: X6 :: X7
- :: X8 :: X9 :: X10 :: X11 :: X12 :: X13 :: X14 :: X15
- :: FP0 :: nil.
-
-Lemma all_mregs_complete:
- forall (r: mreg), In r all_mregs.
-Proof.
- assert (forall r, proj_sumbool (In_dec mreg_eq r all_mregs) = true) by (destruct r; reflexivity).
- intros. specialize (H r). InvBooleans. auto.
-Qed.
-
-Instance Decidable_eq_mreg : forall (x y: mreg), Decidable (eq x y) := Decidable_eq mreg_eq.
-
-Instance Finite_mreg : Finite mreg := {
- Finite_elements := all_mregs;
- Finite_elements_spec := all_mregs_complete
-}.
-
-Definition mreg_type (r: mreg): typ :=
- match r with
- | AX | BX | CX | DX | SI | DI | BP => if Archi.ptr64 then Tany64 else Tany32
- | R8 | R9 | R10 | R11 | R12 | R13 | R14 | R15 => Tany64
- | X0 | X1 | X2 | X3 | X4 | X5 | X6 | X7 => Tany64
- | X8 | X9 | X10 | X11 | X12 | X13 | X14 | X15 => Tany64
- | FP0 => Tany64
- end.
-
-Local Open Scope positive_scope.
-
-Module IndexedMreg <: INDEXED_TYPE.
- Definition t := mreg.
- Definition eq := mreg_eq.
- Definition index (r: mreg): positive :=
- match r with
- | AX => 1 | BX => 2 | CX => 3 | DX => 4 | SI => 5 | DI => 6 | BP => 7
- | R8 => 8 | R9 => 9 | R10 => 10 | R11 => 11 | R12 => 12 | R13 => 13 | R14 => 14 | R15 => 15
- | X0 => 16 | X1 => 17 | X2 => 18 | X3 => 19 | X4 => 20 | X5 => 21 | X6 => 22 | X7 => 23
- | X8 => 24 | X9 => 25 | X10 => 26 | X11 => 27 | X12 => 28 | X13 => 29 | X14 => 30 | X15 => 31
- | FP0 => 32
- end.
- Lemma index_inj:
- forall r1 r2, index r1 = index r2 -> r1 = r2.
- Proof.
- decide_goal.
- Qed.
-End IndexedMreg.
-
-Definition is_stack_reg (r: mreg) : bool :=
- match r with FP0 => true | _ => false end.
-
-(** ** Names of registers *)
-
-Local Open Scope string_scope.
-
-Definition register_names :=
- ("RAX", AX) :: ("RBX", BX) :: ("RCX", CX) :: ("RDX", DX) ::
- ("RSI", SI) :: ("RDI", DI) :: ("RBP", BP) ::
- ("EAX", AX) :: ("EBX", BX) :: ("ECX", CX) :: ("EDX", DX) ::
- ("ESI", SI) :: ("EDI", DI) :: ("EBP", BP) ::
- ("R8", R8) :: ("R9", R9) :: ("R10", R10) :: ("R11", R11) ::
- ("R12", R12) :: ("R13", R13) :: ("R14", R14) :: ("R15", R15) ::
- ("XMM0", X0) :: ("XMM1", X1) :: ("XMM2", X2) :: ("XMM3", X3) ::
- ("XMM4", X4) :: ("XMM5", X5) :: ("XMM6", X6) :: ("XMM7", X7) ::
- ("XMM8", X8) :: ("XMM9", X9) :: ("XMM10", X10) :: ("XMM11", X11) ::
- ("XMM12", X12) :: ("XMM13", X13) :: ("XMM14", X14) :: ("XMM15", X15) ::
- ("ST0", FP0) :: nil.
-
-Definition register_by_name (s: string) : option mreg :=
- let fix assoc (l: list (string * mreg)) : option mreg :=
- match l with
- | nil => None
- | (s1, r1) :: l' => if string_dec s s1 then Some r1 else assoc l'
- end
- in assoc register_names.
-
-(** ** Destroyed registers, preferred registers *)
-
-Definition destroyed_by_op (op: operation): list mreg :=
- match op with
- | Ocast8signed | Ocast8unsigned => AX :: nil
- | Omulhs => AX :: DX :: nil
- | Omulhu => AX :: DX :: nil
- | Odiv => AX :: DX :: nil
- | Odivu => AX :: DX :: nil
- | Omod => AX :: DX :: nil
- | Omodu => AX :: DX :: nil
- | Oshrximm _ => CX :: nil
- | Omullhs => AX :: DX :: nil
- | Omullhu => AX :: DX :: nil
- | Odivl => AX :: DX :: nil
- | Odivlu => AX :: DX :: nil
- | Omodl => AX :: DX :: nil
- | Omodlu => AX :: DX :: nil
- | Oshrxlimm _ => DX :: nil
- | Ocmp _ => AX :: CX :: nil
- | _ => nil
- end.
-
-Definition destroyed_by_load (chunk: memory_chunk) (addr: addressing): list mreg :=
- nil.
-
-Definition destroyed_by_store (chunk: memory_chunk) (addr: addressing): list mreg :=
- match chunk with
- | Mint8signed | Mint8unsigned => if Archi.ptr64 then nil else AX :: CX :: nil
- | _ => nil
- end.
-
-Definition destroyed_by_cond (cond: condition): list mreg :=
- nil.
-
-Definition destroyed_by_jumptable: list mreg :=
- AX :: DX :: nil.
-
-Fixpoint destroyed_by_clobber (cl: list string): list mreg :=
- match cl with
- | nil => nil
- | c1 :: cl =>
- match register_by_name c1 with
- | Some r => r :: destroyed_by_clobber cl
- | None => destroyed_by_clobber cl
- end
- end.
-
-Definition destroyed_by_builtin (ef: external_function): list mreg :=
- match ef with
- | EF_memcpy sz al =>
- if zle sz 32 then CX :: X7 :: nil else CX :: SI :: DI :: nil
- | EF_vstore (Mint8unsigned|Mint8signed) =>
- if Archi.ptr64 then nil else AX :: CX :: nil
- | EF_builtin name sg =>
- if string_dec name "__builtin_va_start" then AX :: nil
- else if string_dec name "__builtin_write16_reversed"
- || string_dec name "__builtin_write32_reversed"
- then CX :: DX :: nil
- else nil
- | EF_inline_asm txt sg clob => destroyed_by_clobber clob
- | _ => nil
- end.
-
-Definition destroyed_at_function_entry: list mreg :=
- (* must include [destroyed_by_setstack ty] *)
- AX :: FP0 :: nil.
-
-Definition destroyed_by_setstack (ty: typ): list mreg :=
- match ty with
- | Tfloat | Tsingle => FP0 :: nil
- | _ => nil
- end.
-
-Definition destroyed_at_indirect_call: list mreg :=
- nil.
-
-Definition temp_for_parent_frame: mreg :=
- AX.
-
-Definition mregs_for_operation (op: operation): list (option mreg) * option mreg :=
- match op with
- | Omulhs => (Some AX :: None :: nil, Some DX)
- | Omulhu => (Some AX :: None :: nil, Some DX)
- | Odiv => (Some AX :: Some CX :: nil, Some AX)
- | Odivu => (Some AX :: Some CX :: nil, Some AX)
- | Omod => (Some AX :: Some CX :: nil, Some DX)
- | Omodu => (Some AX :: Some CX :: nil, Some DX)
- | Oshl => (None :: Some CX :: nil, None)
- | Oshr => (None :: Some CX :: nil, None)
- | Oshru => (None :: Some CX :: nil, None)
- | Oshrximm _ => (Some AX :: nil, Some AX)
- | Omullhs => (Some AX :: None :: nil, Some DX)
- | Omullhu => (Some AX :: None :: nil, Some DX)
- | Odivl => (Some AX :: Some CX :: nil, Some AX)
- | Odivlu => (Some AX :: Some CX :: nil, Some AX)
- | Omodl => (Some AX :: Some CX :: nil, Some DX)
- | Omodlu => (Some AX :: Some CX :: nil, Some DX)
- | Oshll => (None :: Some CX :: nil, None)
- | Oshrl => (None :: Some CX :: nil, None)
- | Oshrlu => (None :: Some CX :: nil, None)
- | Oshrxlimm _ => (Some AX :: nil, Some AX)
- | _ => (nil, None)
- end.
-
-Definition mregs_for_builtin (ef: external_function): list (option mreg) * list (option mreg) :=
- match ef with
- | EF_memcpy sz al =>
- if zle sz 32 then (Some AX :: Some DX :: nil, nil) else (Some DI :: Some SI :: nil, nil)
- | EF_builtin name sg =>
- if string_dec name "__builtin_negl" then
- (Some DX :: Some AX :: nil, Some DX :: Some AX :: nil)
- else if string_dec name "__builtin_addl"
- || string_dec name "__builtin_subl" then
- (Some DX :: Some AX :: Some CX :: Some BX :: nil, Some DX :: Some AX :: nil)
- else if string_dec name "__builtin_mull" then
- (Some AX :: Some DX :: nil, Some DX :: Some AX :: nil)
- else if string_dec name "__builtin_va_start" then
- (Some DX :: nil, nil)
- else if (negb Archi.ptr64) && string_dec name "__builtin_bswap64" then
- (Some AX :: Some DX :: nil, Some DX :: Some AX :: nil)
- else
- (nil, nil)
- | _ => (nil, nil)
- end.
-
-Global Opaque
- destroyed_by_op destroyed_by_load destroyed_by_store
- destroyed_by_cond destroyed_by_jumptable destroyed_by_builtin
- destroyed_by_setstack destroyed_at_function_entry temp_for_parent_frame
- mregs_for_operation mregs_for_builtin.
-
-(** Two-address operations. Return [true] if the first argument and
- the result must be in the same location *and* are unconstrained
- by [mregs_for_operation]. *)
-
-Definition two_address_op (op: operation) : bool :=
- match op with
- | Omove => false
- | Ointconst _ => false
- | Olongconst _ => false
- | Ofloatconst _ => false
- | Osingleconst _ => false
- | Oindirectsymbol _ => false
- | Ocast8signed => false
- | Ocast8unsigned => false
- | Ocast16signed => false
- | Ocast16unsigned => false
- | Oneg => true
- | Osub => true
- | Omul => true
- | Omulimm _ => true
- | Omulhs => false
- | Omulhu => false
- | Odiv => false
- | Odivu => false
- | Omod => false
- | Omodu => false
- | Oand => true
- | Oandimm _ => true
- | Oor => true
- | Oorimm _ => true
- | Oxor => true
- | Oxorimm _ => true
- | Onot => true
- | Oshl => true
- | Oshlimm _ => true
- | Oshr => true
- | Oshrimm _ => true
- | Oshrximm _ => false
- | Oshru => true
- | Oshruimm _ => true
- | Ororimm _ => true
- | Oshldimm _ => true
- | Olea addr => false
- | Omakelong => true
- | Olowlong => true
- | Ohighlong => true
- | Ocast32signed => false
- | Ocast32unsigned => false
- | Onegl => true
- | Oaddlimm _ => true
- | Osubl => true
- | Omull => true
- | Omullimm _ => true
- | Omullhs => false
- | Omullhu => false
- | Odivl => false
- | Odivlu => false
- | Omodl => false
- | Omodlu => false
- | Oandl => true
- | Oandlimm _ => true
- | Oorl => true
- | Oorlimm _ => true
- | Oxorl => true
- | Oxorlimm _ => true
- | Onotl => true
- | Oshll => true
- | Oshllimm _ => true
- | Oshrl => true
- | Oshrlimm _ => true
- | Oshrxlimm _ => false
- | Oshrlu => true
- | Oshrluimm _ => true
- | Ororlimm _ => true
- | Oleal addr => false
- | Onegf => true
- | Oabsf => true
- | Oaddf => true
- | Osubf => true
- | Omulf => true
- | Odivf => true
- | Onegfs => true
- | Oabsfs => true
- | Oaddfs => true
- | Osubfs => true
- | Omulfs => true
- | Odivfs => true
- | Osingleoffloat => false
- | Ofloatofsingle => false
- | Ointoffloat => false
- | Ofloatofint => false
- | Ointofsingle => false
- | Osingleofint => false
- | Olongoffloat => false
- | Ofloatoflong => false
- | Olongofsingle => false
- | Osingleoflong => false
- | Ocmp c => false
- end.
-
-(* Constraints on constant propagation for builtins *)
-
-Definition builtin_constraints (ef: external_function) :
- list builtin_arg_constraint :=
- match ef with
- | EF_vload _ => OK_addrany :: nil
- | EF_vstore _ => OK_addrany :: OK_default :: nil
- | EF_memcpy _ _ => OK_addrany :: OK_addrany :: nil
- | EF_annot txt targs => map (fun _ => OK_all) targs
- | EF_debug kind txt targs => map (fun _ => OK_all) targs
- | _ => nil
- end.
diff --git a/ia32/Machregsaux.ml b/ia32/Machregsaux.ml
deleted file mode 100644
index 473e0602..00000000
--- a/ia32/Machregsaux.ml
+++ /dev/null
@@ -1,33 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Auxiliary functions on machine registers *)
-
-open Camlcoq
-open Machregs
-
-let register_names : (mreg, string) Hashtbl.t = Hashtbl.create 31
-
-let _ =
- List.iter
- (fun (s, r) -> Hashtbl.add register_names r (camlstring_of_coqstring s))
- Machregs.register_names
-
-let is_scratch_register r = false
-
-let name_of_register r =
- try Some (Hashtbl.find register_names r) with Not_found -> None
-
-let register_by_name s =
- Machregs.register_by_name (coqstring_uppercase_ascii_of_camlstring s)
-
-let can_reserve_register r = Conventions1.is_callee_save r
diff --git a/ia32/Machregsaux.mli b/ia32/Machregsaux.mli
deleted file mode 100644
index 9404568d..00000000
--- a/ia32/Machregsaux.mli
+++ /dev/null
@@ -1,18 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Auxiliary functions on machine registers *)
-
-val name_of_register: Machregs.mreg -> string option
-val register_by_name: string -> Machregs.mreg option
-val is_scratch_register: string -> bool
-val can_reserve_register: Machregs.mreg -> bool
diff --git a/ia32/NeedOp.v b/ia32/NeedOp.v
deleted file mode 100644
index 09013cdd..00000000
--- a/ia32/NeedOp.v
+++ /dev/null
@@ -1,254 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Neededness analysis for x86_64 operators *)
-
-Require Import Coqlib.
-Require Import AST Integers Floats Values Memory Globalenvs.
-Require Import Op NeedDomain RTL.
-
-Definition op1 (nv: nval) := nv :: nil.
-Definition op2 (nv: nval) := nv :: nv :: nil.
-
-Definition needs_of_condition (cond: condition): list nval :=
- match cond with
- | Cmaskzero n | Cmasknotzero n => op1 (maskzero n)
- | _ => nil
- end.
-
-Definition needs_of_addressing_32 (addr: addressing) (nv: nval): list nval :=
- match addr with
- | Aindexed n => op1 (modarith nv)
- | Aindexed2 n => op2 (modarith nv)
- | Ascaled sc ofs => op1 (modarith (modarith nv))
- | Aindexed2scaled sc ofs => op2 (modarith nv)
- | Aglobal s ofs => nil
- | Abased s ofs => op1 (modarith nv)
- | Abasedscaled sc s ofs => op1 (modarith (modarith nv))
- | Ainstack ofs => nil
- end.
-
-Definition needs_of_addressing_64 (addr: addressing) (nv: nval): list nval :=
- match addr with
- | Aindexed n => op1 (default nv)
- | Aindexed2 n => op2 (default nv)
- | Ascaled sc ofs => op1 (default nv)
- | Aindexed2scaled sc ofs => op2 (default nv)
- | Aglobal s ofs => nil
- | Abased s ofs => op1 (default nv)
- | Abasedscaled sc s ofs => op1 (default nv)
- | Ainstack ofs => nil
- end.
-
-Definition needs_of_addressing (addr: addressing) (nv: nval): list nval :=
- if Archi.ptr64 then needs_of_addressing_64 addr nv else needs_of_addressing_32 addr nv.
-
-Definition needs_of_operation (op: operation) (nv: nval): list nval :=
- match op with
- | Omove => op1 nv
- | Ointconst n => nil
- | Olongconst n => nil
- | Ofloatconst n => nil
- | Osingleconst n => nil
- | Oindirectsymbol id => nil
- | Ocast8signed => op1 (sign_ext 8 nv)
- | Ocast8unsigned => op1 (zero_ext 8 nv)
- | Ocast16signed => op1 (sign_ext 16 nv)
- | Ocast16unsigned => op1 (zero_ext 16 nv)
- | Oneg => op1 (modarith nv)
- | Osub => op2 (default nv)
- | Omul => op2 (modarith nv)
- | Omulimm n => op1 (modarith nv)
- | Omulhs | Omulhu | Odiv | Odivu | Omod | Omodu => op2 (default nv)
- | Oand => op2 (bitwise nv)
- | Oandimm n => op1 (andimm nv n)
- | Oor => op2 (bitwise nv)
- | Oorimm n => op1 (orimm nv n)
- | Oxor => op2 (bitwise nv)
- | Oxorimm n => op1 (bitwise nv)
- | Onot => op1 (bitwise nv)
- | Oshl => op2 (default nv)
- | Oshlimm n => op1 (shlimm nv n)
- | Oshr => op2 (default nv)
- | Oshrimm n => op1 (shrimm nv n)
- | Oshrximm n => op1 (default nv)
- | Oshru => op2 (default nv)
- | Oshruimm n => op1 (shruimm nv n)
- | Ororimm n => op1 (ror nv n)
- | Oshldimm n => op1 (default nv)
- | Olea addr => needs_of_addressing_32 addr nv
- | Omakelong => op2 (default nv)
- | Olowlong | Ohighlong => op1 (default nv)
- | Ocast32signed => op1 (default nv)
- | Ocast32unsigned => op1 (default nv)
- | Onegl => op1 (default nv)
- | Oaddlimm _ => op1 (default nv)
- | Osubl => op2 (default nv)
- | Omull => op2 (default nv)
- | Omullimm _ => op1 (default nv)
- | Omullhs | Omullhu | Odivl | Odivlu | Omodl | Omodlu => op2 (default nv)
- | Oandl => op2 (default nv)
- | Oandlimm _ => op1 (default nv)
- | Oorl => op2 (default nv)
- | Oorlimm _ => op1 (default nv)
- | Oxorl => op2 (default nv)
- | Oxorlimm _ => op1 (default nv)
- | Onotl => op1 (default nv)
- | Oshll => op2 (default nv)
- | Oshllimm _ => op1 (default nv)
- | Oshrl => op2 (default nv)
- | Oshrlimm _ => op1 (default nv)
- | Oshrxlimm n => op1 (default nv)
- | Oshrlu => op2 (default nv)
- | Oshrluimm _ => op1 (default nv)
- | Ororlimm _ => op1 (default nv)
- | Oleal addr => needs_of_addressing_64 addr nv
- | Onegf | Oabsf => op1 (default nv)
- | Oaddf | Osubf | Omulf | Odivf => op2 (default nv)
- | Onegfs | Oabsfs => op1 (default nv)
- | Oaddfs | Osubfs | Omulfs | Odivfs => op2 (default nv)
- | Osingleoffloat | Ofloatofsingle => op1 (default nv)
- | Ointoffloat | Ofloatofint | Ointofsingle | Osingleofint => op1 (default nv)
- | Olongoffloat | Ofloatoflong | Olongofsingle | Osingleoflong => op1 (default nv)
- | Ocmp c => needs_of_condition c
- end.
-
-Definition operation_is_redundant (op: operation) (nv: nval): bool :=
- match op with
- | Ocast8signed => sign_ext_redundant 8 nv
- | Ocast8unsigned => zero_ext_redundant 8 nv
- | Ocast16signed => sign_ext_redundant 16 nv
- | Ocast16unsigned => zero_ext_redundant 16 nv
- | Oandimm n => andimm_redundant nv n
- | Oorimm n => orimm_redundant nv n
- | _ => false
- end.
-
-Ltac InvAgree :=
- match goal with
- | [H: vagree_list nil _ _ |- _ ] => inv H; InvAgree
- | [H: vagree_list (_::_) _ _ |- _ ] => inv H; InvAgree
- | _ => idtac
- end.
-
-Ltac TrivialExists :=
- match goal with
- | [ |- exists v, Some ?x = Some v /\ _ ] => exists x; split; auto
- | _ => idtac
- end.
-
-Section SOUNDNESS.
-
-Variable ge: genv.
-Variable sp: block.
-Variables m m': mem.
-Hypothesis PERM: forall b ofs k p, Mem.perm m b ofs k p -> Mem.perm m' b ofs k p.
-
-Lemma needs_of_condition_sound:
- forall cond args b args',
- eval_condition cond args m = Some b ->
- vagree_list args args' (needs_of_condition cond) ->
- eval_condition cond args' m' = Some b.
-Proof.
- intros. destruct cond; simpl in H;
- try (eapply default_needs_of_condition_sound; eauto; fail);
- simpl in *; FuncInv; InvAgree.
-- eapply maskzero_sound; eauto.
-- destruct (Val.maskzero_bool v n) as [b'|] eqn:MZ; try discriminate.
- erewrite maskzero_sound; eauto.
-Qed.
-
-Lemma needs_of_addressing_32_sound:
- forall sp addr args v nv args',
- eval_addressing32 ge (Vptr sp Ptrofs.zero) addr args = Some v ->
- vagree_list args args' (needs_of_addressing_32 addr nv) ->
- exists v',
- eval_addressing32 ge (Vptr sp Ptrofs.zero) addr args' = Some v'
- /\ vagree v v' nv.
-Proof.
- unfold needs_of_addressing_32; intros.
- destruct addr; simpl in *; FuncInv; InvAgree; TrivialExists;
- auto using add_sound, mul_sound with na.
- apply add_sound; auto with na. apply add_sound; rewrite modarith_idem; auto.
- apply add_sound; auto. apply add_sound; rewrite modarith_idem; auto with na.
- apply mul_sound; rewrite modarith_idem; auto with na.
-Qed.
-
-(*
-Lemma needs_of_addressing_64_sound:
- forall sp addr args v nv args',
- eval_addressing64 ge (Vptr sp Ptrofs.zero) addr args = Some v ->
- vagree_list args args' (needs_of_addressing_64 addr nv) ->
- exists v',
- eval_addressing64 ge (Vptr sp Ptrofs.zero) addr args' = Some v'
- /\ vagree v v' nv.
-*)
-
-Lemma needs_of_operation_sound:
- forall op args v nv args',
- eval_operation ge (Vptr sp Ptrofs.zero) op args m = Some v ->
- vagree_list args args' (needs_of_operation op nv) ->
- nv <> Nothing ->
- exists v',
- eval_operation ge (Vptr sp Ptrofs.zero) op args' m' = Some v'
- /\ vagree v v' nv.
-Proof.
- unfold needs_of_operation; intros; destruct op; try (eapply default_needs_of_operation_sound; eauto; fail);
- simpl in *; FuncInv; InvAgree; TrivialExists.
-- apply sign_ext_sound; auto. compute; auto.
-- apply zero_ext_sound; auto. omega.
-- apply sign_ext_sound; auto. compute; auto.
-- apply zero_ext_sound; auto. omega.
-- apply neg_sound; auto.
-- apply mul_sound; auto.
-- apply mul_sound; auto with na.
-- apply and_sound; auto.
-- apply andimm_sound; auto.
-- apply or_sound; auto.
-- apply orimm_sound; auto.
-- apply xor_sound; auto.
-- apply xor_sound; auto with na.
-- apply notint_sound; auto.
-- apply shlimm_sound; auto.
-- apply shrimm_sound; auto.
-- apply shruimm_sound; auto.
-- apply ror_sound; auto.
-- eapply needs_of_addressing_32_sound; eauto.
-- change (eval_addressing64 ge (Vptr sp Ptrofs.zero) a args')
- with (eval_operation ge (Vptr sp Ptrofs.zero) (Oleal a) args' m').
- eapply default_needs_of_operation_sound; eauto.
- destruct a; simpl in H0; auto.
-- destruct (eval_condition cond args m) as [b|] eqn:EC; simpl in H2.
- erewrite needs_of_condition_sound by eauto.
- subst v; simpl. auto with na.
- subst v; auto with na.
-Qed.
-
-Lemma operation_is_redundant_sound:
- forall op nv arg1 args v arg1' args',
- operation_is_redundant op nv = true ->
- eval_operation ge (Vptr sp Ptrofs.zero) op (arg1 :: args) m = Some v ->
- vagree_list (arg1 :: args) (arg1' :: args') (needs_of_operation op nv) ->
- vagree v arg1' nv.
-Proof.
- intros. destruct op; simpl in *; try discriminate; inv H1; FuncInv; subst.
-- apply sign_ext_redundant_sound; auto. omega.
-- apply zero_ext_redundant_sound; auto. omega.
-- apply sign_ext_redundant_sound; auto. omega.
-- apply zero_ext_redundant_sound; auto. omega.
-- apply andimm_redundant_sound; auto.
-- apply orimm_redundant_sound; auto.
-Qed.
-
-End SOUNDNESS.
-
-
diff --git a/ia32/Op.v b/ia32/Op.v
deleted file mode 100644
index eb2fd110..00000000
--- a/ia32/Op.v
+++ /dev/null
@@ -1,1457 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Operators and addressing modes. The abstract syntax and dynamic
- semantics for the CminorSel, RTL, LTL and Mach languages depend on the
- following types, defined in this library:
-- [condition]: boolean conditions for conditional branches;
-- [operation]: arithmetic and logical operations;
-- [addressing]: addressing modes for load and store operations.
-
- These types are X86-64-specific and correspond roughly to what the
- processor can compute in one instruction. In other terms, these
- types reflect the state of the program after instruction selection.
- For a processor-independent set of operations, see the abstract
- syntax and dynamic semantics of the Cminor language.
-*)
-
-Require Import Coqlib.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Events.
-
-Set Implicit Arguments.
-
-(** Conditions (boolean-valued operators). *)
-
-Inductive condition : Type :=
- | Ccomp (c: comparison) (**r signed integer comparison *)
- | Ccompu (c: comparison) (**r unsigned integer comparison *)
- | Ccompimm (c: comparison) (n: int) (**r signed integer comparison with a constant *)
- | Ccompuimm (c: comparison) (n: int) (**r unsigned integer comparison with a constant *)
- | Ccompl (c: comparison) (**r signed 64-bit integer comparison *)
- | Ccomplu (c: comparison) (**r unsigned 64-bit integer comparison *)
- | Ccomplimm (c: comparison) (n: int64) (**r signed 64-bit integer comparison with a constant *)
- | Ccompluimm (c: comparison) (n: int64) (**r unsigned 64-bit integer comparison with a constant *)
- | Ccompf (c: comparison) (**r 64-bit floating-point comparison *)
- | Cnotcompf (c: comparison) (**r negation of a floating-point comparison *)
- | Ccompfs (c: comparison) (**r 32-bit floating-point comparison *)
- | Cnotcompfs (c: comparison) (**r negation of a floating-point comparison *)
- | Cmaskzero (n: int) (**r test [(arg & constant) == 0] *)
- | Cmasknotzero (n: int). (**r test [(arg & constant) != 0] *)
-
-(** Addressing modes. [r1], [r2], etc, are the arguments to the
- addressing. *)
-
-Inductive addressing: Type :=
- | Aindexed: Z -> addressing (**r Address is [r1 + offset] *)
- | Aindexed2: Z -> addressing (**r Address is [r1 + r2 + offset] *)
- | Ascaled: Z -> Z -> addressing (**r Address is [r1 * scale + offset] *)
- | Aindexed2scaled: Z -> Z -> addressing
- (**r Address is [r1 + r2 * scale + offset] *)
- | Aglobal: ident -> ptrofs -> addressing (**r Address is [symbol + offset] *)
- | Abased: ident -> ptrofs -> addressing (**r Address is [symbol + offset + r1] *)
- | Abasedscaled: Z -> ident -> ptrofs -> addressing (**r Address is [symbol + offset + r1 * scale] *)
- | Ainstack: ptrofs -> addressing. (**r Address is [stack_pointer + offset] *)
-
-(** Arithmetic and logical operations. In the descriptions, [rd] is the
- result of the operation and [r1], [r2], etc, are the arguments. *)
-
-Inductive operation : Type :=
- | Omove (**r [rd = r1] *)
- | Ointconst (n: int) (**r [rd] is set to the given integer constant *)
- | Olongconst (n: int64) (**r [rd] is set to the given integer constant *)
- | Ofloatconst (n: float) (**r [rd] is set to the given float constant *)
- | Osingleconst (n: float32)(**r [rd] is set to the given float constant *)
- | Oindirectsymbol (id: ident) (**r [rd] is set to the address of the symbol *)
-(*c 32-bit integer arithmetic: *)
- | Ocast8signed (**r [rd] is 8-bit sign extension of [r1] *)
- | Ocast8unsigned (**r [rd] is 8-bit zero extension of [r1] *)
- | Ocast16signed (**r [rd] is 16-bit sign extension of [r1] *)
- | Ocast16unsigned (**r [rd] is 16-bit zero extension of [r1] *)
- | Oneg (**r [rd = - r1] *)
- | Osub (**r [rd = r1 - r2] *)
- | Omul (**r [rd = r1 * r2] *)
- | Omulimm (n: int) (**r [rd = r1 * n] *)
- | Omulhs (**r [rd = high part of r1 * r2, signed] *)
- | Omulhu (**r [rd = high part of r1 * r2, unsigned] *)
- | Odiv (**r [rd = r1 / r2] (signed) *)
- | Odivu (**r [rd = r1 / r2] (unsigned) *)
- | Omod (**r [rd = r1 % r2] (signed) *)
- | Omodu (**r [rd = r1 % r2] (unsigned) *)
- | Oand (**r [rd = r1 & r2] *)
- | Oandimm (n: int) (**r [rd = r1 & n] *)
- | Oor (**r [rd = r1 | r2] *)
- | Oorimm (n: int) (**r [rd = r1 | n] *)
- | Oxor (**r [rd = r1 ^ r2] *)
- | Oxorimm (n: int) (**r [rd = r1 ^ n] *)
- | Onot (**r [rd = ~r1] *)
- | Oshl (**r [rd = r1 << r2] *)
- | Oshlimm (n: int) (**r [rd = r1 << n] *)
- | Oshr (**r [rd = r1 >> r2] (signed) *)
- | Oshrimm (n: int) (**r [rd = r1 >> n] (signed) *)
- | Oshrximm (n: int) (**r [rd = r1 / 2^n] (signed) *)
- | Oshru (**r [rd = r1 >> r2] (unsigned) *)
- | Oshruimm (n: int) (**r [rd = r1 >> n] (unsigned) *)
- | Ororimm (n: int) (**r rotate right immediate *)
- | Oshldimm (n: int) (**r [rd = r1 << n | r2 >> (32-n)] *)
- | Olea (a: addressing) (**r effective address *)
-(*c 64-bit integer arithmetic: *)
- | Omakelong (**r [rd = r1 << 32 | r2] *)
- | Olowlong (**r [rd = low-word(r1)] *)
- | Ohighlong (**r [rd = high-word(r1)] *)
- | Ocast32signed (**r [rd] is 32-bit sign extension of [r1] *)
- | Ocast32unsigned (**r [rd] is 32-bit zero extension of [r1] *)
- | Onegl (**r [rd = - r1] *)
- | Oaddlimm (n: int64) (**r [rd = r1 + n] *)
- | Osubl (**r [rd = r1 - r2] *)
- | Omull (**r [rd = r1 * r2] *)
- | Omullimm (n: int64) (**r [rd = r1 * n] *)
- | Omullhs (**r [rd = high part of r1 * r2, signed] *)
- | Omullhu (**r [rd = high part of r1 * r2, unsigned] *)
- | Odivl (**r [rd = r1 / r2] (signed) *)
- | Odivlu (**r [rd = r1 / r2] (unsigned) *)
- | Omodl (**r [rd = r1 % r2] (signed) *)
- | Omodlu (**r [rd = r1 % r2] (unsigned) *)
- | Oandl (**r [rd = r1 & r2] *)
- | Oandlimm (n: int64) (**r [rd = r1 & n] *)
- | Oorl (**r [rd = r1 | r2] *)
- | Oorlimm (n: int64) (**r [rd = r1 | n] *)
- | Oxorl (**r [rd = r1 ^ r2] *)
- | Oxorlimm (n: int64) (**r [rd = r1 ^ n] *)
- | Onotl (**r [rd = ~r1] *)
- | Oshll (**r [rd = r1 << r2] *)
- | Oshllimm (n: int) (**r [rd = r1 << n] *)
- | Oshrl (**r [rd = r1 >> r2] (signed) *)
- | Oshrlimm (n: int) (**r [rd = r1 >> n] (signed) *)
- | Oshrxlimm (n: int) (**r [rd = r1 / 2^n] (signed) *)
- | Oshrlu (**r [rd = r1 >> r2] (unsigned) *)
- | Oshrluimm (n: int) (**r [rd = r1 >> n] (unsigned) *)
- | Ororlimm (n: int) (**r rotate right immediate *)
- | Oleal (a: addressing) (**r effective address *)
-(*c Floating-point arithmetic: *)
- | Onegf (**r [rd = - r1] *)
- | Oabsf (**r [rd = abs(r1)] *)
- | Oaddf (**r [rd = r1 + r2] *)
- | Osubf (**r [rd = r1 - r2] *)
- | Omulf (**r [rd = r1 * r2] *)
- | Odivf (**r [rd = r1 / r2] *)
- | Onegfs (**r [rd = - r1] *)
- | Oabsfs (**r [rd = abs(r1)] *)
- | Oaddfs (**r [rd = r1 + r2] *)
- | Osubfs (**r [rd = r1 - r2] *)
- | Omulfs (**r [rd = r1 * r2] *)
- | Odivfs (**r [rd = r1 / r2] *)
- | Osingleoffloat (**r [rd] is [r1] truncated to single-precision float *)
- | Ofloatofsingle (**r [rd] is [r1] extended to double-precision float *)
-(*c Conversions between int and float: *)
- | Ointoffloat (**r [rd = signed_int_of_float64(r1)] *)
- | Ofloatofint (**r [rd = float64_of_signed_int(r1)] *)
- | Ointofsingle (**r [rd = signed_int_of_float32(r1)] *)
- | Osingleofint (**r [rd = float32_of_signed_int(r1)] *)
- | Olongoffloat (**r [rd = signed_long_of_float64(r1)] *)
- | Ofloatoflong (**r [rd = float64_of_signed_long(r1)] *)
- | Olongofsingle (**r [rd = signed_long_of_float32(r1)] *)
- | Osingleoflong (**r [rd = float32_of_signed_long(r1)] *)
-(*c Boolean tests: *)
- | Ocmp (cond: condition). (**r [rd = 1] if condition holds, [rd = 0] otherwise. *)
-
-(** Comparison functions (used in modules [CSE] and [Allocation]). *)
-
-Definition eq_condition (x y: condition) : {x=y} + {x<>y}.
-Proof.
- generalize Int.eq_dec Int64.eq_dec; intro.
- assert (forall (x y: comparison), {x=y}+{x<>y}). decide equality.
- decide equality.
-Defined.
-
-Definition eq_addressing (x y: addressing) : {x=y} + {x<>y}.
-Proof.
- generalize ident_eq Ptrofs.eq_dec zeq; intros.
- decide equality.
-Defined.
-
-Definition eq_operation (x y: operation): {x=y} + {x<>y}.
-Proof.
- generalize Int.eq_dec Int64.eq_dec Float.eq_dec Float32.eq_dec; intros.
- decide equality.
- apply ident_eq.
- apply eq_addressing.
- apply eq_addressing.
- apply eq_condition.
-Defined.
-
-Global Opaque eq_condition eq_addressing eq_operation.
-
-(** In addressing modes, offsets are 32-bit signed integers, even in 64-bit mode.
- The following function checks that an addressing mode is valid, i.e. that
- the offsets are in range. *)
-
-Definition offset_in_range (n: Z) : bool := zle Int.min_signed n && zle n Int.max_signed.
-
-Definition addressing_valid (a: addressing) : bool :=
- match a with
- | Aindexed n => offset_in_range n
- | Aindexed2 n => offset_in_range n
- | Ascaled sc ofs => offset_in_range ofs
- | Aindexed2scaled sc ofs => offset_in_range ofs
- | Aglobal s ofs => true
- | Abased s ofs => true
- | Abasedscaled sc s ofs => true
- | Ainstack ofs => offset_in_range (Ptrofs.signed ofs)
- end.
-
-(** * Evaluation functions *)
-
-(** Evaluation of conditions, operators and addressing modes applied
- to lists of values. Return [None] when the computation can trigger an
- error, e.g. integer division by zero. [eval_condition] returns a boolean,
- [eval_operation] and [eval_addressing] return a value. *)
-
-Definition eval_condition (cond: condition) (vl: list val) (m: mem): option bool :=
- match cond, vl with
- | Ccomp c, v1 :: v2 :: nil => Val.cmp_bool c v1 v2
- | Ccompu c, v1 :: v2 :: nil => Val.cmpu_bool (Mem.valid_pointer m) c v1 v2
- | Ccompimm c n, v1 :: nil => Val.cmp_bool c v1 (Vint n)
- | Ccompuimm c n, v1 :: nil => Val.cmpu_bool (Mem.valid_pointer m) c v1 (Vint n)
- | Ccompl c, v1 :: v2 :: nil => Val.cmpl_bool c v1 v2
- | Ccomplu c, v1 :: v2 :: nil => Val.cmplu_bool (Mem.valid_pointer m) c v1 v2
- | Ccomplimm c n, v1 :: nil => Val.cmpl_bool c v1 (Vlong n)
- | Ccompluimm c n, v1 :: nil => Val.cmplu_bool (Mem.valid_pointer m) c v1 (Vlong n)
- | Ccompf c, v1 :: v2 :: nil => Val.cmpf_bool c v1 v2
- | Cnotcompf c, v1 :: v2 :: nil => option_map negb (Val.cmpf_bool c v1 v2)
- | Ccompfs c, v1 :: v2 :: nil => Val.cmpfs_bool c v1 v2
- | Cnotcompfs c, v1 :: v2 :: nil => option_map negb (Val.cmpfs_bool c v1 v2)
- | Cmaskzero n, v1 :: nil => Val.maskzero_bool v1 n
- | Cmasknotzero n, v1 :: nil => option_map negb (Val.maskzero_bool v1 n)
- | _, _ => None
- end.
-
-Definition eval_addressing32
- (F V: Type) (genv: Genv.t F V) (sp: val)
- (addr: addressing) (vl: list val) : option val :=
- match addr, vl with
- | Aindexed n, v1::nil =>
- Some (Val.add v1 (Vint (Int.repr n)))
- | Aindexed2 n, v1::v2::nil =>
- Some (Val.add (Val.add v1 v2) (Vint (Int.repr n)))
- | Ascaled sc ofs, v1::nil =>
- Some (Val.add (Val.mul v1 (Vint (Int.repr sc))) (Vint (Int.repr ofs)))
- | Aindexed2scaled sc ofs, v1::v2::nil =>
- Some(Val.add v1 (Val.add (Val.mul v2 (Vint (Int.repr sc))) (Vint (Int.repr ofs))))
- | Aglobal s ofs, nil =>
- if Archi.ptr64 then None else Some (Genv.symbol_address genv s ofs)
- | Abased s ofs, v1::nil =>
- if Archi.ptr64 then None else Some (Val.add (Genv.symbol_address genv s ofs) v1)
- | Abasedscaled sc s ofs, v1::nil =>
- if Archi.ptr64 then None else Some (Val.add (Genv.symbol_address genv s ofs) (Val.mul v1 (Vint (Int.repr sc))))
- | Ainstack ofs, nil =>
- if Archi.ptr64 then None else Some(Val.offset_ptr sp ofs)
- | _, _ => None
- end.
-
-Definition eval_addressing64
- (F V: Type) (genv: Genv.t F V) (sp: val)
- (addr: addressing) (vl: list val) : option val :=
- match addr, vl with
- | Aindexed n, v1::nil =>
- Some (Val.addl v1 (Vlong (Int64.repr n)))
- | Aindexed2 n, v1::v2::nil =>
- Some (Val.addl (Val.addl v1 v2) (Vlong (Int64.repr n)))
- | Ascaled sc ofs, v1::nil =>
- Some (Val.addl (Val.mull v1 (Vlong (Int64.repr sc))) (Vlong (Int64.repr ofs)))
- | Aindexed2scaled sc ofs, v1::v2::nil =>
- Some(Val.addl v1 (Val.addl (Val.mull v2 (Vlong (Int64.repr sc))) (Vlong (Int64.repr ofs))))
- | Aglobal s ofs, nil =>
- if Archi.ptr64 then Some (Genv.symbol_address genv s ofs) else None
- | Ainstack ofs, nil =>
- if Archi.ptr64 then Some(Val.offset_ptr sp ofs) else None
- | _, _ => None
- end.
-
-Definition eval_addressing
- (F V: Type) (genv: Genv.t F V) (sp: val)
- (addr: addressing) (vl: list val) : option val :=
- if Archi.ptr64
- then eval_addressing64 genv sp addr vl
- else eval_addressing32 genv sp addr vl.
-
-Definition eval_operation
- (F V: Type) (genv: Genv.t F V) (sp: val)
- (op: operation) (vl: list val) (m: mem): option val :=
- match op, vl with
- | Omove, v1::nil => Some v1
- | Ointconst n, nil => Some (Vint n)
- | Olongconst n, nil => Some (Vlong n)
- | Ofloatconst n, nil => Some (Vfloat n)
- | Osingleconst n, nil => Some (Vsingle n)
- | Oindirectsymbol id, nil => Some (Genv.symbol_address genv id Ptrofs.zero)
- | Ocast8signed, v1 :: nil => Some (Val.sign_ext 8 v1)
- | Ocast8unsigned, v1 :: nil => Some (Val.zero_ext 8 v1)
- | Ocast16signed, v1 :: nil => Some (Val.sign_ext 16 v1)
- | Ocast16unsigned, v1 :: nil => Some (Val.zero_ext 16 v1)
- | Oneg, v1::nil => Some (Val.neg v1)
- | Osub, v1::v2::nil => Some (Val.sub v1 v2)
- | Omul, v1::v2::nil => Some (Val.mul v1 v2)
- | Omulimm n, v1::nil => Some (Val.mul v1 (Vint n))
- | Omulhs, v1::v2::nil => Some (Val.mulhs v1 v2)
- | Omulhu, v1::v2::nil => Some (Val.mulhu v1 v2)
- | Odiv, v1::v2::nil => Val.divs v1 v2
- | Odivu, v1::v2::nil => Val.divu v1 v2
- | Omod, v1::v2::nil => Val.mods v1 v2
- | Omodu, v1::v2::nil => Val.modu v1 v2
- | Oand, v1::v2::nil => Some(Val.and v1 v2)
- | Oandimm n, v1::nil => Some (Val.and v1 (Vint n))
- | Oor, v1::v2::nil => Some(Val.or v1 v2)
- | Oorimm n, v1::nil => Some (Val.or v1 (Vint n))
- | Oxor, v1::v2::nil => Some(Val.xor v1 v2)
- | Oxorimm n, v1::nil => Some (Val.xor v1 (Vint n))
- | Onot, v1::nil => Some(Val.notint v1)
- | Oshl, v1::v2::nil => Some (Val.shl v1 v2)
- | Oshlimm n, v1::nil => Some (Val.shl v1 (Vint n))
- | Oshr, v1::v2::nil => Some (Val.shr v1 v2)
- | Oshrimm n, v1::nil => Some (Val.shr v1 (Vint n))
- | Oshrximm n, v1::nil => Val.shrx v1 (Vint n)
- | Oshru, v1::v2::nil => Some (Val.shru v1 v2)
- | Oshruimm n, v1::nil => Some (Val.shru v1 (Vint n))
- | Ororimm n, v1::nil => Some (Val.ror v1 (Vint n))
- | Oshldimm n, v1::v2::nil => Some (Val.or (Val.shl v1 (Vint n))
- (Val.shru v2 (Vint (Int.sub Int.iwordsize n))))
- | Olea addr, _ => eval_addressing32 genv sp addr vl
- | Omakelong, v1::v2::nil => Some(Val.longofwords v1 v2)
- | Olowlong, v1::nil => Some(Val.loword v1)
- | Ohighlong, v1::nil => Some(Val.hiword v1)
- | Ocast32signed, v1 :: nil => Some (Val.longofint v1)
- | Ocast32unsigned, v1 :: nil => Some (Val.longofintu v1)
- | Onegl, v1::nil => Some (Val.negl v1)
- | Oaddlimm n, v1::nil => Some (Val.addl v1 (Vlong n))
- | Osubl, v1::v2::nil => Some (Val.subl v1 v2)
- | Omull, v1::v2::nil => Some (Val.mull v1 v2)
- | Omullimm n, v1::nil => Some (Val.mull v1 (Vlong n))
- | Omullhs, v1::v2::nil => Some (Val.mullhs v1 v2)
- | Omullhu, v1::v2::nil => Some (Val.mullhu v1 v2)
- | Odivl, v1::v2::nil => Val.divls v1 v2
- | Odivlu, v1::v2::nil => Val.divlu v1 v2
- | Omodl, v1::v2::nil => Val.modls v1 v2
- | Omodlu, v1::v2::nil => Val.modlu v1 v2
- | Oandl, v1::v2::nil => Some(Val.andl v1 v2)
- | Oandlimm n, v1::nil => Some (Val.andl v1 (Vlong n))
- | Oorl, v1::v2::nil => Some(Val.orl v1 v2)
- | Oorlimm n, v1::nil => Some (Val.orl v1 (Vlong n))
- | Oxorl, v1::v2::nil => Some(Val.xorl v1 v2)
- | Oxorlimm n, v1::nil => Some (Val.xorl v1 (Vlong n))
- | Onotl, v1::nil => Some(Val.notl v1)
- | Oshll, v1::v2::nil => Some (Val.shll v1 v2)
- | Oshllimm n, v1::nil => Some (Val.shll v1 (Vint n))
- | Oshrl, v1::v2::nil => Some (Val.shrl v1 v2)
- | Oshrlimm n, v1::nil => Some (Val.shrl v1 (Vint n))
- | Oshrxlimm n, v1::nil => Val.shrxl v1 (Vint n)
- | Oshrlu, v1::v2::nil => Some (Val.shrlu v1 v2)
- | Oshrluimm n, v1::nil => Some (Val.shrlu v1 (Vint n))
- | Ororlimm n, v1::nil => Some (Val.rorl v1 (Vint n))
- | Oleal addr, _ => eval_addressing64 genv sp addr vl
- | Onegf, v1::nil => Some(Val.negf v1)
- | Oabsf, v1::nil => Some(Val.absf v1)
- | Oaddf, v1::v2::nil => Some(Val.addf v1 v2)
- | Osubf, v1::v2::nil => Some(Val.subf v1 v2)
- | Omulf, v1::v2::nil => Some(Val.mulf v1 v2)
- | Odivf, v1::v2::nil => Some(Val.divf v1 v2)
- | Onegfs, v1::nil => Some(Val.negfs v1)
- | Oabsfs, v1::nil => Some(Val.absfs v1)
- | Oaddfs, v1::v2::nil => Some(Val.addfs v1 v2)
- | Osubfs, v1::v2::nil => Some(Val.subfs v1 v2)
- | Omulfs, v1::v2::nil => Some(Val.mulfs v1 v2)
- | Odivfs, v1::v2::nil => Some(Val.divfs v1 v2)
- | Osingleoffloat, v1::nil => Some(Val.singleoffloat v1)
- | Ofloatofsingle, v1::nil => Some(Val.floatofsingle v1)
- | Ointoffloat, v1::nil => Val.intoffloat v1
- | Ofloatofint, v1::nil => Val.floatofint v1
- | Ointofsingle, v1::nil => Val.intofsingle v1
- | Osingleofint, v1::nil => Val.singleofint v1
- | Olongoffloat, v1::nil => Val.longoffloat v1
- | Ofloatoflong, v1::nil => Val.floatoflong v1
- | Olongofsingle, v1::nil => Val.longofsingle v1
- | Osingleoflong, v1::nil => Val.singleoflong v1
- | Ocmp c, _ => Some(Val.of_optbool (eval_condition c vl m))
- | _, _ => None
- end.
-
-Remark eval_addressing_Aglobal:
- forall (F V: Type) (genv: Genv.t F V) sp id ofs,
- eval_addressing genv sp (Aglobal id ofs) nil = Some (Genv.symbol_address genv id ofs).
-Proof.
- intros. unfold eval_addressing, eval_addressing32, eval_addressing64; destruct Archi.ptr64; auto.
-Qed.
-
-Remark eval_addressing_Ainstack:
- forall (F V: Type) (genv: Genv.t F V) sp ofs,
- eval_addressing genv sp (Ainstack ofs) nil = Some (Val.offset_ptr sp ofs).
-Proof.
- intros. unfold eval_addressing, eval_addressing32, eval_addressing64; destruct Archi.ptr64; auto.
-Qed.
-
-Remark eval_addressing_Ainstack_inv:
- forall (F V: Type) (genv: Genv.t F V) sp ofs vl v,
- eval_addressing genv sp (Ainstack ofs) vl = Some v -> vl = nil /\ v = Val.offset_ptr sp ofs.
-Proof.
- unfold eval_addressing, eval_addressing32, eval_addressing64;
- intros; destruct Archi.ptr64; destruct vl; inv H; auto.
-Qed.
-
-Ltac FuncInv :=
- match goal with
- | H: (match ?x with nil => _ | _ :: _ => _ end = Some _) |- _ =>
- destruct x; simpl in H; try discriminate H; FuncInv
- | H: (match ?v with Vundef => _ | Vint _ => _ | Vfloat _ => _ | Vptr _ _ => _ end = Some _) |- _ =>
- destruct v; simpl in H; try discriminate H; FuncInv
- | H: (if Archi.ptr64 then _ else _) = Some _ |- _ =>
- destruct Archi.ptr64 eqn:?; try discriminate H; FuncInv
- | H: (Some _ = Some _) |- _ =>
- injection H; intros; clear H; FuncInv
- | H: (None = Some _) |- _ =>
- discriminate H
- | _ =>
- idtac
- end.
-
-(** * Static typing of conditions, operators and addressing modes. *)
-
-Definition type_of_condition (c: condition) : list typ :=
- match c with
- | Ccomp _ => Tint :: Tint :: nil
- | Ccompu _ => Tint :: Tint :: nil
- | Ccompimm _ _ => Tint :: nil
- | Ccompuimm _ _ => Tint :: nil
- | Ccompl _ => Tlong :: Tlong :: nil
- | Ccomplu _ => Tlong :: Tlong :: nil
- | Ccomplimm _ _ => Tlong :: nil
- | Ccompluimm _ _ => Tlong :: nil
- | Ccompf _ => Tfloat :: Tfloat :: nil
- | Cnotcompf _ => Tfloat :: Tfloat :: nil
- | Ccompfs _ => Tsingle :: Tsingle :: nil
- | Cnotcompfs _ => Tsingle :: Tsingle :: nil
- | Cmaskzero _ => Tint :: nil
- | Cmasknotzero _ => Tint :: nil
- end.
-
-Definition type_of_addressing_gen (tyA: typ) (addr: addressing): list typ :=
- match addr with
- | Aindexed _ => tyA :: nil
- | Aindexed2 _ => tyA :: tyA :: nil
- | Ascaled _ _ => tyA :: nil
- | Aindexed2scaled _ _ => tyA :: tyA :: nil
- | Aglobal _ _ => nil
- | Abased _ _ => tyA :: nil
- | Abasedscaled _ _ _ => tyA :: nil
- | Ainstack _ => nil
- end.
-
-Definition type_of_addressing := type_of_addressing_gen Tptr.
-Definition type_of_addressing32 := type_of_addressing_gen Tint.
-Definition type_of_addressing64 := type_of_addressing_gen Tlong.
-
-Definition type_of_operation (op: operation) : list typ * typ :=
- match op with
- | Omove => (nil, Tint) (* treated specially *)
- | Ointconst _ => (nil, Tint)
- | Olongconst _ => (nil, Tlong)
- | Ofloatconst f => (nil, Tfloat)
- | Osingleconst f => (nil, Tsingle)
- | Oindirectsymbol _ => (nil, Tptr)
- | Ocast8signed => (Tint :: nil, Tint)
- | Ocast8unsigned => (Tint :: nil, Tint)
- | Ocast16signed => (Tint :: nil, Tint)
- | Ocast16unsigned => (Tint :: nil, Tint)
- | Oneg => (Tint :: nil, Tint)
- | Osub => (Tint :: Tint :: nil, Tint)
- | Omul => (Tint :: Tint :: nil, Tint)
- | Omulimm _ => (Tint :: nil, Tint)
- | Omulhs => (Tint :: Tint :: nil, Tint)
- | Omulhu => (Tint :: Tint :: nil, Tint)
- | Odiv => (Tint :: Tint :: nil, Tint)
- | Odivu => (Tint :: Tint :: nil, Tint)
- | Omod => (Tint :: Tint :: nil, Tint)
- | Omodu => (Tint :: Tint :: nil, Tint)
- | Oand => (Tint :: Tint :: nil, Tint)
- | Oandimm _ => (Tint :: nil, Tint)
- | Oor => (Tint :: Tint :: nil, Tint)
- | Oorimm _ => (Tint :: nil, Tint)
- | Oxor => (Tint :: Tint :: nil, Tint)
- | Oxorimm _ => (Tint :: nil, Tint)
- | Onot => (Tint :: nil, Tint)
- | Oshl => (Tint :: Tint :: nil, Tint)
- | Oshlimm _ => (Tint :: nil, Tint)
- | Oshr => (Tint :: Tint :: nil, Tint)
- | Oshrimm _ => (Tint :: nil, Tint)
- | Oshrximm _ => (Tint :: nil, Tint)
- | Oshru => (Tint :: Tint :: nil, Tint)
- | Oshruimm _ => (Tint :: nil, Tint)
- | Ororimm _ => (Tint :: nil, Tint)
- | Oshldimm _ => (Tint :: Tint :: nil, Tint)
- | Olea addr => (type_of_addressing32 addr, Tint)
- | Omakelong => (Tint :: Tint :: nil, Tlong)
- | Olowlong => (Tlong :: nil, Tint)
- | Ohighlong => (Tlong :: nil, Tint)
- | Ocast32signed => (Tint :: nil, Tlong)
- | Ocast32unsigned => (Tint :: nil, Tlong)
- | Onegl => (Tlong :: nil, Tlong)
- | Oaddlimm _ => (Tlong :: nil, Tlong)
- | Osubl => (Tlong :: Tlong :: nil, Tlong)
- | Omull => (Tlong :: Tlong :: nil, Tlong)
- | Omullimm _ => (Tlong :: nil, Tlong)
- | Omullhs => (Tlong :: Tlong :: nil, Tlong)
- | Omullhu => (Tlong :: Tlong :: nil, Tlong)
- | Odivl => (Tlong :: Tlong :: nil, Tlong)
- | Odivlu => (Tlong :: Tlong :: nil, Tlong)
- | Omodl => (Tlong :: Tlong :: nil, Tlong)
- | Omodlu => (Tlong :: Tlong :: nil, Tlong)
- | Oandl => (Tlong :: Tlong :: nil, Tlong)
- | Oandlimm _ => (Tlong :: nil, Tlong)
- | Oorl => (Tlong :: Tlong :: nil, Tlong)
- | Oorlimm _ => (Tlong :: nil, Tlong)
- | Oxorl => (Tlong :: Tlong :: nil, Tlong)
- | Oxorlimm _ => (Tlong :: nil, Tlong)
- | Onotl => (Tlong :: nil, Tlong)
- | Oshll => (Tlong :: Tint :: nil, Tlong)
- | Oshllimm _ => (Tlong :: nil, Tlong)
- | Oshrl => (Tlong :: Tint :: nil, Tlong)
- | Oshrlimm _ => (Tlong :: nil, Tlong)
- | Oshrxlimm _ => (Tlong :: nil, Tlong)
- | Oshrlu => (Tlong :: Tint :: nil, Tlong)
- | Oshrluimm _ => (Tlong :: nil, Tlong)
- | Ororlimm _ => (Tlong :: nil, Tlong)
- | Oleal addr => (type_of_addressing64 addr, Tlong)
- | Onegf => (Tfloat :: nil, Tfloat)
- | Oabsf => (Tfloat :: nil, Tfloat)
- | Oaddf => (Tfloat :: Tfloat :: nil, Tfloat)
- | Osubf => (Tfloat :: Tfloat :: nil, Tfloat)
- | Omulf => (Tfloat :: Tfloat :: nil, Tfloat)
- | Odivf => (Tfloat :: Tfloat :: nil, Tfloat)
- | Onegfs => (Tsingle :: nil, Tsingle)
- | Oabsfs => (Tsingle :: nil, Tsingle)
- | Oaddfs => (Tsingle :: Tsingle :: nil, Tsingle)
- | Osubfs => (Tsingle :: Tsingle :: nil, Tsingle)
- | Omulfs => (Tsingle :: Tsingle :: nil, Tsingle)
- | Odivfs => (Tsingle :: Tsingle :: nil, Tsingle)
- | Osingleoffloat => (Tfloat :: nil, Tsingle)
- | Ofloatofsingle => (Tsingle :: nil, Tfloat)
- | Ointoffloat => (Tfloat :: nil, Tint)
- | Ofloatofint => (Tint :: nil, Tfloat)
- | Ointofsingle => (Tsingle :: nil, Tint)
- | Osingleofint => (Tint :: nil, Tsingle)
- | Olongoffloat => (Tfloat :: nil, Tlong)
- | Ofloatoflong => (Tlong :: nil, Tfloat)
- | Olongofsingle => (Tsingle :: nil, Tlong)
- | Osingleoflong => (Tlong :: nil, Tsingle)
- | Ocmp c => (type_of_condition c, Tint)
- end.
-
-(** Weak type soundness results for [eval_operation]:
- the result values, when defined, are always of the type predicted
- by [type_of_operation]. *)
-
-Section SOUNDNESS.
-
-Variable A V: Type.
-Variable genv: Genv.t A V.
-
-Remark type_add:
- forall v1 v2, Val.has_type (Val.add v1 v2) Tint.
-Proof.
- intros. unfold Val.has_type, Val.add. destruct Archi.ptr64, v1, v2; auto.
-Qed.
-
-Remark type_addl:
- forall v1 v2, Val.has_type (Val.addl v1 v2) Tlong.
-Proof.
- intros. unfold Val.has_type, Val.addl. destruct Archi.ptr64, v1, v2; auto.
-Qed.
-
-Lemma type_of_addressing64_sound:
- forall addr vl sp v,
- eval_addressing64 genv sp addr vl = Some v ->
- Val.has_type v Tlong.
-Proof.
- intros. destruct addr; simpl in H; FuncInv; subst; simpl; auto using type_addl.
-- unfold Genv.symbol_address; destruct (Genv.find_symbol genv i); simpl; auto.
-- destruct sp; simpl; auto.
-Qed.
-
-Lemma type_of_addressing32_sound:
- forall addr vl sp v,
- eval_addressing32 genv sp addr vl = Some v ->
- Val.has_type v Tint.
-Proof.
- intros. destruct addr; simpl in H; FuncInv; subst; simpl; auto using type_add.
-- unfold Genv.symbol_address; destruct (Genv.find_symbol genv i); simpl; auto.
-- destruct sp; simpl; auto.
-Qed.
-
-Corollary type_of_addressing_sound:
- forall addr vl sp v,
- eval_addressing genv sp addr vl = Some v ->
- Val.has_type v Tptr.
-Proof.
- unfold eval_addressing, Tptr; intros.
- destruct Archi.ptr64; eauto using type_of_addressing64_sound, type_of_addressing32_sound.
-Qed.
-
-Lemma type_of_operation_sound:
- forall op vl sp v m,
- op <> Omove ->
- eval_operation genv sp op vl m = Some v ->
- Val.has_type v (snd (type_of_operation op)).
-Proof with (try exact I).
- intros.
- destruct op; simpl in H0; FuncInv; subst; simpl.
- congruence.
- exact I.
- exact I.
- exact I.
- exact I.
- unfold Genv.symbol_address; destruct (Genv.find_symbol genv id)...
- unfold Val.has_type, Tptr; destruct Archi.ptr64; auto.
- destruct v0...
- destruct v0...
- destruct v0...
- destruct v0...
- destruct v0...
- destruct v0; destruct v1; simpl...
- unfold Val.has_type; destruct Archi.ptr64; auto.
- unfold Val.has_type; destruct Archi.ptr64; auto. destruct (eq_block b b0); auto.
- destruct v0; destruct v1...
- destruct v0...
- destruct v0; destruct v1...
- destruct v0; destruct v1...
- destruct v0; destruct v1; simpl in *; inv H0.
- destruct (Int.eq i0 Int.zero || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H2...
- destruct v0; destruct v1; simpl in *; inv H0. destruct (Int.eq i0 Int.zero); inv H2...
- destruct v0; destruct v1; simpl in *; inv H0.
- destruct (Int.eq i0 Int.zero || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H2...
- destruct v0; destruct v1; simpl in *; inv H0. destruct (Int.eq i0 Int.zero); inv H2...
- destruct v0; destruct v1...
- destruct v0...
- destruct v0; destruct v1...
- destruct v0...
- destruct v0; destruct v1...
- destruct v0...
- destruct v0...
- destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int.iwordsize)...
- destruct v0; simpl... destruct (Int.ltu n Int.iwordsize)...
- destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int.iwordsize)...
- destruct v0; simpl... destruct (Int.ltu n Int.iwordsize)...
- destruct v0; simpl in H0; try discriminate. destruct (Int.ltu n (Int.repr 31)); inv H0...
- destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int.iwordsize)...
- destruct v0; simpl... destruct (Int.ltu n Int.iwordsize)...
- destruct v0...
- destruct v0; simpl... destruct (Int.ltu n Int.iwordsize)...
- destruct v1; simpl... destruct (Int.ltu (Int.sub Int.iwordsize n) Int.iwordsize)...
- eapply type_of_addressing32_sound; eauto.
- destruct v0; destruct v1...
- destruct v0...
- destruct v0...
- destruct v0...
- destruct v0...
- destruct v0...
- destruct v0; simpl... unfold Val.has_type; destruct Archi.ptr64; auto.
- destruct v0; destruct v1; simpl...
- unfold Val.has_type; destruct Archi.ptr64; auto.
- unfold Val.has_type; destruct Archi.ptr64; simpl; auto. destruct (eq_block b b0); auto.
- destruct v0; destruct v1...
- destruct v0...
- destruct v0; destruct v1...
- destruct v0; destruct v1...
- destruct v0; destruct v1; simpl in *; inv H0.
- destruct (Int64.eq i0 Int64.zero || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2...
- destruct v0; destruct v1; simpl in *; inv H0. destruct (Int64.eq i0 Int64.zero); inv H2...
- destruct v0; destruct v1; simpl in *; inv H0.
- destruct (Int64.eq i0 Int64.zero || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2...
- destruct v0; destruct v1; simpl in *; inv H0. destruct (Int64.eq i0 Int64.zero); inv H2...
- destruct v0; destruct v1...
- destruct v0...
- destruct v0; destruct v1...
- destruct v0...
- destruct v0; destruct v1...
- destruct v0...
- destruct v0...
- destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int64.iwordsize')...
- destruct v0; simpl... destruct (Int.ltu n Int64.iwordsize')...
- destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int64.iwordsize')...
- destruct v0; simpl... destruct (Int.ltu n Int64.iwordsize')...
- destruct v0; inv H0. destruct (Int.ltu n (Int.repr 63)); inv H2...
- destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int64.iwordsize')...
- destruct v0; simpl... destruct (Int.ltu n Int64.iwordsize')...
- destruct v0...
- eapply type_of_addressing64_sound; eauto.
- destruct v0...
- destruct v0...
- destruct v0; destruct v1...
- destruct v0; destruct v1...
- destruct v0; destruct v1...
- destruct v0; destruct v1...
- destruct v0...
- destruct v0...
- destruct v0; destruct v1...
- destruct v0; destruct v1...
- destruct v0; destruct v1...
- destruct v0; destruct v1...
- destruct v0...
- destruct v0...
- destruct v0; simpl in H0; inv H0. destruct (Float.to_int f); inv H2...
- destruct v0; simpl in H0; inv H0...
- destruct v0; simpl in H0; inv H0. destruct (Float32.to_int f); inv H2...
- destruct v0; simpl in H0; inv H0...
- destruct v0; simpl in H0; inv H0. destruct (Float.to_long f); inv H2...
- destruct v0; simpl in H0; inv H0...
- destruct v0; simpl in H0; inv H0. destruct (Float32.to_long f); inv H2...
- destruct v0; simpl in H0; inv H0...
- destruct (eval_condition cond vl m); simpl... destruct b...
-Qed.
-
-End SOUNDNESS.
-
-(** * Manipulating and transforming operations *)
-
-(** Recognition of move operations. *)
-
-Definition is_move_operation
- (A: Type) (op: operation) (args: list A) : option A :=
- match op, args with
- | Omove, arg :: nil => Some arg
- | _, _ => None
- end.
-
-Lemma is_move_operation_correct:
- forall (A: Type) (op: operation) (args: list A) (a: A),
- is_move_operation op args = Some a ->
- op = Omove /\ args = a :: nil.
-Proof.
- intros until a. unfold is_move_operation; destruct op;
- try (intros; discriminate).
- destruct args. intros; discriminate.
- destruct args. intros. intuition congruence.
- intros; discriminate.
-Qed.
-
-(** [negate_condition cond] returns a condition that is logically
- equivalent to the negation of [cond]. *)
-
-Definition negate_condition (cond: condition): condition :=
- match cond with
- | Ccomp c => Ccomp(negate_comparison c)
- | Ccompu c => Ccompu(negate_comparison c)
- | Ccompimm c n => Ccompimm (negate_comparison c) n
- | Ccompuimm c n => Ccompuimm (negate_comparison c) n
- | Ccompl c => Ccompl(negate_comparison c)
- | Ccomplu c => Ccomplu(negate_comparison c)
- | Ccomplimm c n => Ccomplimm (negate_comparison c) n
- | Ccompluimm c n => Ccompluimm (negate_comparison c) n
- | Ccompf c => Cnotcompf c
- | Cnotcompf c => Ccompf c
- | Ccompfs c => Cnotcompfs c
- | Cnotcompfs c => Ccompfs c
- | Cmaskzero n => Cmasknotzero n
- | Cmasknotzero n => Cmaskzero n
- end.
-
-Lemma eval_negate_condition:
- forall cond vl m,
- eval_condition (negate_condition cond) vl m = option_map negb (eval_condition cond vl m).
-Proof.
- intros. destruct cond; simpl.
- repeat (destruct vl; auto). apply Val.negate_cmp_bool.
- repeat (destruct vl; auto). apply Val.negate_cmpu_bool.
- repeat (destruct vl; auto). apply Val.negate_cmp_bool.
- repeat (destruct vl; auto). apply Val.negate_cmpu_bool.
- repeat (destruct vl; auto). apply Val.negate_cmpl_bool.
- repeat (destruct vl; auto). apply Val.negate_cmplu_bool.
- repeat (destruct vl; auto). apply Val.negate_cmpl_bool.
- repeat (destruct vl; auto). apply Val.negate_cmplu_bool.
- repeat (destruct vl; auto).
- repeat (destruct vl; auto). destruct (Val.cmpf_bool c v v0) as [[]|]; auto.
- repeat (destruct vl; auto).
- repeat (destruct vl; auto). destruct (Val.cmpfs_bool c v v0) as [[]|]; auto.
- destruct vl; auto. destruct vl; auto.
- destruct vl; auto. destruct vl; auto. destruct (Val.maskzero_bool v n) as [[]|]; auto.
-Qed.
-
-(** Shifting stack-relative references. This is used in [Stacking]. *)
-
-Definition shift_stack_addressing (delta: Z) (addr: addressing) :=
- match addr with
- | Ainstack ofs => Ainstack (Ptrofs.add ofs (Ptrofs.repr delta))
- | _ => addr
- end.
-
-Definition shift_stack_operation (delta: Z) (op: operation) :=
- match op with
- | Olea addr => Olea (shift_stack_addressing delta addr)
- | Oleal addr => Oleal (shift_stack_addressing delta addr)
- | _ => op
- end.
-
-Lemma type_shift_stack_addressing:
- forall delta addr, type_of_addressing (shift_stack_addressing delta addr) = type_of_addressing addr.
-Proof.
- intros. destruct addr; auto.
-Qed.
-
-Lemma type_shift_stack_operation:
- forall delta op, type_of_operation (shift_stack_operation delta op) = type_of_operation op.
-Proof.
- intros. destruct op; auto; simpl; decEq; destruct a; auto.
-Qed.
-
-Lemma eval_shift_stack_addressing32:
- forall F V (ge: Genv.t F V) sp addr vl delta,
- eval_addressing32 ge (Vptr sp Ptrofs.zero) (shift_stack_addressing delta addr) vl =
- eval_addressing32 ge (Vptr sp (Ptrofs.repr delta)) addr vl.
-Proof.
- intros. destruct addr; simpl; auto.
- destruct vl; auto. destruct Archi.ptr64 eqn:SF; auto.
- do 2 f_equal. rewrite Ptrofs.add_zero_l. apply Ptrofs.add_commut.
-Qed.
-
-Lemma eval_shift_stack_addressing64:
- forall F V (ge: Genv.t F V) sp addr vl delta,
- eval_addressing64 ge (Vptr sp Ptrofs.zero) (shift_stack_addressing delta addr) vl =
- eval_addressing64 ge (Vptr sp (Ptrofs.repr delta)) addr vl.
-Proof.
- intros. destruct addr; simpl; auto.
- destruct vl; auto. destruct Archi.ptr64 eqn:SF; auto.
- do 2 f_equal. rewrite Ptrofs.add_zero_l. apply Ptrofs.add_commut.
-Qed.
-
-Lemma eval_shift_stack_addressing:
- forall F V (ge: Genv.t F V) sp addr vl delta,
- eval_addressing ge (Vptr sp Ptrofs.zero) (shift_stack_addressing delta addr) vl =
- eval_addressing ge (Vptr sp (Ptrofs.repr delta)) addr vl.
-Proof.
- intros. unfold eval_addressing.
- destruct Archi.ptr64; auto using eval_shift_stack_addressing32, eval_shift_stack_addressing64.
-Qed.
-
-Lemma eval_shift_stack_operation:
- forall F V (ge: Genv.t F V) sp op vl m delta,
- eval_operation ge (Vptr sp Ptrofs.zero) (shift_stack_operation delta op) vl m =
- eval_operation ge (Vptr sp (Ptrofs.repr delta)) op vl m.
-Proof.
- intros. destruct op; simpl; auto using eval_shift_stack_addressing32, eval_shift_stack_addressing64.
-Qed.
-
-(** Offset an addressing mode [addr] by a quantity [delta], so that
- it designates the pointer [delta] bytes past the pointer designated
- by [addr]. This may be undefined if an offset overflows, in which case
- [None] is returned. *)
-
-Definition offset_addressing_total (addr: addressing) (delta: Z) : addressing :=
- match addr with
- | Aindexed n => Aindexed (n + delta)
- | Aindexed2 n => Aindexed2 (n + delta)
- | Ascaled sc n => Ascaled sc (n + delta)
- | Aindexed2scaled sc n => Aindexed2scaled sc (n + delta)
- | Aglobal s n => Aglobal s (Ptrofs.add n (Ptrofs.repr delta))
- | Abased s n => Abased s (Ptrofs.add n (Ptrofs.repr delta))
- | Abasedscaled sc s n => Abasedscaled sc s (Ptrofs.add n (Ptrofs.repr delta))
- | Ainstack n => Ainstack (Ptrofs.add n (Ptrofs.repr delta))
- end.
-
-Definition offset_addressing (addr: addressing) (delta: Z) : option addressing :=
- let addr' := offset_addressing_total addr delta in
- if addressing_valid addr' then Some addr' else None.
-
-Lemma eval_offset_addressing_total_32:
- forall (F V: Type) (ge: Genv.t F V) sp addr args delta v,
- eval_addressing32 ge sp addr args = Some v ->
- eval_addressing32 ge sp (offset_addressing_total addr delta) args = Some(Val.add v (Vint (Int.repr delta))).
-Proof.
- assert (A: forall x y, Int.add (Int.repr x) (Int.repr y) = Int.repr (x + y)).
- { intros. apply Int.eqm_samerepr; auto with ints. }
- assert (B: forall delta, Archi.ptr64 = false -> Ptrofs.repr delta = Ptrofs.of_int (Int.repr delta)).
- { intros; symmetry; auto with ptrofs. }
- intros. destruct addr; simpl in *; FuncInv; subst; simpl.
-- rewrite <- A, ! Val.add_assoc; auto.
-- rewrite <- A, ! Val.add_assoc; auto.
-- rewrite <- A, ! Val.add_assoc; auto.
-- rewrite <- A, ! Val.add_assoc; auto.
-- rewrite B, Genv.shift_symbol_address_32 by auto. auto.
-- rewrite B, Genv.shift_symbol_address_32 by auto. rewrite ! Val.add_assoc. do 2 f_equal. apply Val.add_commut.
-- rewrite B, Genv.shift_symbol_address_32 by auto. rewrite ! Val.add_assoc. do 2 f_equal. apply Val.add_commut.
-- destruct sp; simpl; auto. rewrite Heqb. rewrite Ptrofs.add_assoc. do 4 f_equal. symmetry; auto with ptrofs.
-Qed.
-
-Lemma eval_offset_addressing_total_64:
- forall (F V: Type) (ge: Genv.t F V) sp addr args delta v,
- eval_addressing64 ge sp addr args = Some v ->
- eval_addressing64 ge sp (offset_addressing_total addr delta) args = Some(Val.addl v (Vlong (Int64.repr delta))).
-Proof.
- assert (A: forall x y, Int64.add (Int64.repr x) (Int64.repr y) = Int64.repr (x + y)).
- { intros. apply Int64.eqm_samerepr; auto with ints. }
- assert (B: forall delta, Archi.ptr64 = true -> Ptrofs.repr delta = Ptrofs.of_int64 (Int64.repr delta)).
- { intros; symmetry; auto with ptrofs. }
- intros. destruct addr; simpl in *; FuncInv; subst; simpl.
-- rewrite <- A, ! Val.addl_assoc; auto.
-- rewrite <- A, ! Val.addl_assoc; auto.
-- rewrite <- A, ! Val.addl_assoc; auto.
-- rewrite <- A, ! Val.addl_assoc; auto.
-- rewrite B, Genv.shift_symbol_address_64 by auto. auto.
-- destruct sp; simpl; auto. rewrite Heqb. rewrite Ptrofs.add_assoc. do 4 f_equal. symmetry; auto with ptrofs.
-Qed.
-
-(** The following lemma is used only in [Allocproof] in cases where [Archi.ptr64 = false]. *)
-
-Lemma eval_offset_addressing:
- forall (F V: Type) (ge: Genv.t F V) sp addr args delta addr' v,
- offset_addressing addr delta = Some addr' ->
- eval_addressing ge sp addr args = Some v ->
- Archi.ptr64 = false ->
- eval_addressing ge sp addr' args = Some(Val.add v (Vint (Int.repr delta))).
-Proof.
- intros. unfold offset_addressing in H. destruct (addressing_valid (offset_addressing_total addr delta)); inv H.
- unfold eval_addressing in *; rewrite H1 in *. apply eval_offset_addressing_total_32; auto.
-Qed.
-
-(** Operations that are so cheap to recompute that CSE should not factor them out. *)
-
-Definition is_trivial_op (op: operation) : bool :=
- match op with
- | Omove => true
- | Ointconst _ => true
- | Olongconst _ => true
- | Olea (Aglobal _ _) => true
- | Olea (Ainstack _) => true
- | Oleal (Aglobal _ _) => true
- | Oleal (Ainstack _) => true
- | _ => false
- end.
-
-(** Operations that depend on the memory state. *)
-
-Definition op_depends_on_memory (op: operation) : bool :=
- match op with
- | Ocmp (Ccompu _) => negb Archi.ptr64
- | Ocmp (Ccompuimm _ _) => negb Archi.ptr64
- | Ocmp (Ccomplu _) => Archi.ptr64
- | Ocmp (Ccompluimm _ _) => Archi.ptr64
- | _ => false
- end.
-
-Lemma op_depends_on_memory_correct:
- forall (F V: Type) (ge: Genv.t F V) sp op args m1 m2,
- op_depends_on_memory op = false ->
- eval_operation ge sp op args m1 = eval_operation ge sp op args m2.
-Proof.
- intros until m2. destruct op; simpl; try congruence.
- destruct cond; simpl; intros SF; auto; rewrite ? negb_false_iff in SF;
- unfold Val.cmpu_bool, Val.cmplu_bool; rewrite SF; reflexivity.
-Qed.
-
-(** Global variables mentioned in an operation or addressing mode *)
-
-Definition globals_addressing (addr: addressing) : list ident :=
- match addr with
- | Aglobal s n => s :: nil
- | Abased s n => s :: nil
- | Abasedscaled sc s n => s :: nil
- | _ => nil
- end.
-
-Definition globals_operation (op: operation) : list ident :=
- match op with
- | Oindirectsymbol s => s :: nil
- | Olea addr => globals_addressing addr
- | Oleal addr => globals_addressing addr
- | _ => nil
- end.
-
-(** * Invariance and compatibility properties. *)
-
-(** [eval_operation] and [eval_addressing] depend on a global environment
- for resolving references to global symbols. We show that they give
- the same results if a global environment is replaced by another that
- assigns the same addresses to the same symbols. *)
-
-Section GENV_TRANSF.
-
-Variable F1 F2 V1 V2: Type.
-Variable ge1: Genv.t F1 V1.
-Variable ge2: Genv.t F2 V2.
-Hypothesis agree_on_symbols:
- forall (s: ident), Genv.find_symbol ge2 s = Genv.find_symbol ge1 s.
-
-Lemma eval_addressing32_preserved:
- forall sp addr vl,
- eval_addressing32 ge2 sp addr vl = eval_addressing32 ge1 sp addr vl.
-Proof.
- intros.
- unfold eval_addressing32, Genv.symbol_address; destruct addr; try rewrite agree_on_symbols;
- reflexivity.
-Qed.
-
-Lemma eval_addressing64_preserved:
- forall sp addr vl,
- eval_addressing64 ge2 sp addr vl = eval_addressing64 ge1 sp addr vl.
-Proof.
- intros.
- unfold eval_addressing64, Genv.symbol_address; destruct addr; try rewrite agree_on_symbols;
- reflexivity.
-Qed.
-
-Lemma eval_addressing_preserved:
- forall sp addr vl,
- eval_addressing ge2 sp addr vl = eval_addressing ge1 sp addr vl.
-Proof.
- intros.
- unfold eval_addressing; destruct Archi.ptr64; auto using eval_addressing32_preserved, eval_addressing64_preserved.
-Qed.
-
-Lemma eval_operation_preserved:
- forall sp op vl m,
- eval_operation ge2 sp op vl m = eval_operation ge1 sp op vl m.
-Proof.
- intros.
- unfold eval_operation; destruct op; auto.
- unfold Genv.symbol_address. rewrite agree_on_symbols. auto.
- apply eval_addressing32_preserved.
- apply eval_addressing64_preserved.
-Qed.
-
-End GENV_TRANSF.
-
-(** Compatibility of the evaluation functions with value injections. *)
-
-Section EVAL_COMPAT.
-
-Variable F1 F2 V1 V2: Type.
-Variable ge1: Genv.t F1 V1.
-Variable ge2: Genv.t F2 V2.
-Variable f: meminj.
-
-Variable m1: mem.
-Variable m2: mem.
-
-Hypothesis valid_pointer_inj:
- forall b1 ofs b2 delta,
- f b1 = Some(b2, delta) ->
- Mem.valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
- Mem.valid_pointer m2 b2 (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr delta))) = true.
-
-Hypothesis weak_valid_pointer_inj:
- forall b1 ofs b2 delta,
- f b1 = Some(b2, delta) ->
- Mem.weak_valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
- Mem.weak_valid_pointer m2 b2 (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr delta))) = true.
-
-Hypothesis weak_valid_pointer_no_overflow:
- forall b1 ofs b2 delta,
- f b1 = Some(b2, delta) ->
- Mem.weak_valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
- 0 <= Ptrofs.unsigned ofs + Ptrofs.unsigned (Ptrofs.repr delta) <= Ptrofs.max_unsigned.
-
-Hypothesis valid_different_pointers_inj:
- forall b1 ofs1 b2 ofs2 b1' delta1 b2' delta2,
- b1 <> b2 ->
- Mem.valid_pointer m1 b1 (Ptrofs.unsigned ofs1) = true ->
- Mem.valid_pointer m1 b2 (Ptrofs.unsigned ofs2) = true ->
- f b1 = Some (b1', delta1) ->
- f b2 = Some (b2', delta2) ->
- b1' <> b2' \/
- Ptrofs.unsigned (Ptrofs.add ofs1 (Ptrofs.repr delta1)) <> Ptrofs.unsigned (Ptrofs.add ofs2 (Ptrofs.repr delta2)).
-
-Ltac InvInject :=
- match goal with
- | [ H: Val.inject _ (Vint _) _ |- _ ] =>
- inv H; InvInject
- | [ H: Val.inject _ (Vfloat _) _ |- _ ] =>
- inv H; InvInject
- | [ H: Val.inject _ (Vptr _ _) _ |- _ ] =>
- inv H; InvInject
- | [ H: Val.inject_list _ nil _ |- _ ] =>
- inv H; InvInject
- | [ H: Val.inject_list _ (_ :: _) _ |- _ ] =>
- inv H; InvInject
- | _ => idtac
- end.
-
-Lemma eval_condition_inj:
- forall cond vl1 vl2 b,
- Val.inject_list f vl1 vl2 ->
- eval_condition cond vl1 m1 = Some b ->
- eval_condition cond vl2 m2 = Some b.
-Proof.
- intros. destruct cond; simpl in H0; FuncInv; InvInject; simpl; auto.
-- inv H3; inv H2; simpl in H0; inv H0; auto.
-- eauto 3 using Val.cmpu_bool_inject, Mem.valid_pointer_implies.
-- inv H3; simpl in H0; inv H0; auto.
-- eauto 3 using Val.cmpu_bool_inject, Mem.valid_pointer_implies.
-- inv H3; inv H2; simpl in H0; inv H0; auto.
-- eauto 3 using Val.cmplu_bool_inject, Mem.valid_pointer_implies.
-- inv H3; simpl in H0; inv H0; auto.
-- eauto 3 using Val.cmplu_bool_inject, Mem.valid_pointer_implies.
-- inv H3; inv H2; simpl in H0; inv H0; auto.
-- inv H3; inv H2; simpl in H0; inv H0; auto.
-- inv H3; inv H2; simpl in H0; inv H0; auto.
-- inv H3; inv H2; simpl in H0; inv H0; auto.
-- inv H3; try discriminate; auto.
-- inv H3; try discriminate; auto.
-Qed.
-
-Ltac TrivialExists :=
- match goal with
- | [ |- exists v2, Some ?v1 = Some v2 /\ Val.inject _ _ v2 ] =>
- exists v1; split; auto
- | _ => idtac
- end.
-
-Lemma eval_addressing32_inj:
- forall addr sp1 vl1 sp2 vl2 v1,
- (forall id ofs,
- In id (globals_addressing addr) ->
- Val.inject f (Genv.symbol_address ge1 id ofs) (Genv.symbol_address ge2 id ofs)) ->
- Val.inject f sp1 sp2 ->
- Val.inject_list f vl1 vl2 ->
- eval_addressing32 ge1 sp1 addr vl1 = Some v1 ->
- exists v2, eval_addressing32 ge2 sp2 addr vl2 = Some v2 /\ Val.inject f v1 v2.
-Proof.
- assert (A: forall v1 v2 v1' v2', Val.inject f v1 v1' -> Val.inject f v2 v2' -> Val.inject f (Val.mul v1 v2) (Val.mul v1' v2')).
- { intros. inv H; simpl; auto. inv H0; auto. }
- intros. destruct addr; simpl in *; FuncInv; InvInject; TrivialExists; eauto using Val.add_inject, Val.offset_ptr_inject with coqlib.
-Qed.
-
-Lemma eval_addressing64_inj:
- forall addr sp1 vl1 sp2 vl2 v1,
- (forall id ofs,
- In id (globals_addressing addr) ->
- Val.inject f (Genv.symbol_address ge1 id ofs) (Genv.symbol_address ge2 id ofs)) ->
- Val.inject f sp1 sp2 ->
- Val.inject_list f vl1 vl2 ->
- eval_addressing64 ge1 sp1 addr vl1 = Some v1 ->
- exists v2, eval_addressing64 ge2 sp2 addr vl2 = Some v2 /\ Val.inject f v1 v2.
-Proof.
- assert (A: forall v1 v2 v1' v2', Val.inject f v1 v1' -> Val.inject f v2 v2' -> Val.inject f (Val.mull v1 v2) (Val.mull v1' v2')).
- { intros. inv H; simpl; auto. inv H0; auto. }
- intros. destruct addr; simpl in *; FuncInv; InvInject; TrivialExists; eauto using Val.addl_inject, Val.offset_ptr_inject with coqlib.
-Qed.
-
-Lemma eval_addressing_inj:
- forall addr sp1 vl1 sp2 vl2 v1,
- (forall id ofs,
- In id (globals_addressing addr) ->
- Val.inject f (Genv.symbol_address ge1 id ofs) (Genv.symbol_address ge2 id ofs)) ->
- Val.inject f sp1 sp2 ->
- Val.inject_list f vl1 vl2 ->
- eval_addressing ge1 sp1 addr vl1 = Some v1 ->
- exists v2, eval_addressing ge2 sp2 addr vl2 = Some v2 /\ Val.inject f v1 v2.
-Proof.
- unfold eval_addressing; intros. destruct Archi.ptr64; eauto using eval_addressing32_inj, eval_addressing64_inj.
-Qed.
-
-Lemma eval_operation_inj:
- forall op sp1 vl1 sp2 vl2 v1,
- (forall id ofs,
- In id (globals_operation op) ->
- Val.inject f (Genv.symbol_address ge1 id ofs) (Genv.symbol_address ge2 id ofs)) ->
- Val.inject f sp1 sp2 ->
- Val.inject_list f vl1 vl2 ->
- eval_operation ge1 sp1 op vl1 m1 = Some v1 ->
- exists v2, eval_operation ge2 sp2 op vl2 m2 = Some v2 /\ Val.inject f v1 v2.
-Proof.
- intros until v1; intros GL; intros. destruct op; simpl in H1; simpl; FuncInv; InvInject; TrivialExists.
- apply GL; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- apply Val.sub_inject; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int.eq i0 Int.zero || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H2. TrivialExists.
- inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int.eq i0 Int.zero); inv H2. TrivialExists.
- inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int.eq i0 Int.zero || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H2. TrivialExists.
- inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int.eq i0 Int.zero); inv H2. TrivialExists.
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
- inv H4; simpl; auto. destruct (Int.ltu n Int.iwordsize); auto.
- inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
- inv H4; simpl; auto. destruct (Int.ltu n Int.iwordsize); auto.
- inv H4; simpl in H1; try discriminate. simpl.
- destruct (Int.ltu n (Int.repr 31)); inv H1. TrivialExists.
- inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
- inv H4; simpl; auto. destruct (Int.ltu n Int.iwordsize); auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto. destruct (Int.ltu n Int.iwordsize); auto.
- inv H2; simpl; auto. destruct (Int.ltu (Int.sub Int.iwordsize n) Int.iwordsize); auto.
- eapply eval_addressing32_inj; eauto.
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- apply Val.addl_inject; auto.
- apply Val.subl_inject; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int64.eq i0 Int64.zero || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2. TrivialExists.
- inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int64.eq i0 Int64.zero); inv H2. TrivialExists.
- inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int64.eq i0 Int64.zero || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2. TrivialExists.
- inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int64.eq i0 Int64.zero); inv H2. TrivialExists.
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int64.iwordsize'); auto.
- inv H4; simpl; auto. destruct (Int.ltu n Int64.iwordsize'); auto.
- inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int64.iwordsize'); auto.
- inv H4; simpl; auto. destruct (Int.ltu n Int64.iwordsize'); auto.
- inv H4; simpl in H1; try discriminate. simpl. destruct (Int.ltu n (Int.repr 63)); inv H1. TrivialExists.
- inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int64.iwordsize'); auto.
- inv H4; simpl; auto. destruct (Int.ltu n Int64.iwordsize'); auto.
- inv H4; simpl; auto.
- eapply eval_addressing64_inj; eauto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl in H1; inv H1. simpl. destruct (Float.to_int f0); simpl in H2; inv H2.
- exists (Vint i); auto.
- inv H4; simpl in H1; inv H1. simpl. TrivialExists.
- inv H4; simpl in H1; inv H1. simpl. destruct (Float32.to_int f0); simpl in H2; inv H2.
- exists (Vint i); auto.
- inv H4; simpl in H1; inv H1. simpl. TrivialExists.
- inv H4; simpl in H1; inv H1. simpl. destruct (Float.to_long f0); simpl in H2; inv H2.
- exists (Vlong i); auto.
- inv H4; simpl in H1; inv H1. simpl. TrivialExists.
- inv H4; simpl in H1; inv H1. simpl. destruct (Float32.to_long f0); simpl in H2; inv H2.
- exists (Vlong i); auto.
- inv H4; simpl in H1; inv H1. simpl. TrivialExists.
- subst v1. destruct (eval_condition cond vl1 m1) eqn:?.
- exploit eval_condition_inj; eauto. intros EQ; rewrite EQ.
- destruct b; simpl; constructor.
- simpl; constructor.
-Qed.
-
-End EVAL_COMPAT.
-
-(** Compatibility of the evaluation functions with the ``is less defined'' relation over values. *)
-
-Section EVAL_LESSDEF.
-
-Variable F V: Type.
-Variable genv: Genv.t F V.
-
-Remark valid_pointer_extends:
- forall m1 m2, Mem.extends m1 m2 ->
- forall b1 ofs b2 delta,
- Some(b1, 0) = Some(b2, delta) ->
- Mem.valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
- Mem.valid_pointer m2 b2 (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr delta))) = true.
-Proof.
- intros. inv H0. rewrite Ptrofs.add_zero. eapply Mem.valid_pointer_extends; eauto.
-Qed.
-
-Remark weak_valid_pointer_extends:
- forall m1 m2, Mem.extends m1 m2 ->
- forall b1 ofs b2 delta,
- Some(b1, 0) = Some(b2, delta) ->
- Mem.weak_valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
- Mem.weak_valid_pointer m2 b2 (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr delta))) = true.
-Proof.
- intros. inv H0. rewrite Ptrofs.add_zero. eapply Mem.weak_valid_pointer_extends; eauto.
-Qed.
-
-Remark weak_valid_pointer_no_overflow_extends:
- forall m1 b1 ofs b2 delta,
- Some(b1, 0) = Some(b2, delta) ->
- Mem.weak_valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
- 0 <= Ptrofs.unsigned ofs + Ptrofs.unsigned (Ptrofs.repr delta) <= Ptrofs.max_unsigned.
-Proof.
- intros. inv H. rewrite Zplus_0_r. apply Ptrofs.unsigned_range_2.
-Qed.
-
-Remark valid_different_pointers_extends:
- forall m1 b1 ofs1 b2 ofs2 b1' delta1 b2' delta2,
- b1 <> b2 ->
- Mem.valid_pointer m1 b1 (Ptrofs.unsigned ofs1) = true ->
- Mem.valid_pointer m1 b2 (Ptrofs.unsigned ofs2) = true ->
- Some(b1, 0) = Some (b1', delta1) ->
- Some(b2, 0) = Some (b2', delta2) ->
- b1' <> b2' \/
- Ptrofs.unsigned(Ptrofs.add ofs1 (Ptrofs.repr delta1)) <> Ptrofs.unsigned(Ptrofs.add ofs2 (Ptrofs.repr delta2)).
-Proof.
- intros. inv H2; inv H3. auto.
-Qed.
-
-Lemma eval_condition_lessdef:
- forall cond vl1 vl2 b m1 m2,
- Val.lessdef_list vl1 vl2 ->
- Mem.extends m1 m2 ->
- eval_condition cond vl1 m1 = Some b ->
- eval_condition cond vl2 m2 = Some b.
-Proof.
- intros. eapply eval_condition_inj with (f := fun b => Some(b, 0)) (m1 := m1).
- apply valid_pointer_extends; auto.
- apply weak_valid_pointer_extends; auto.
- apply weak_valid_pointer_no_overflow_extends.
- apply valid_different_pointers_extends; auto.
- rewrite <- val_inject_list_lessdef. eauto. auto.
-Qed.
-
-Lemma eval_operation_lessdef:
- forall sp op vl1 vl2 v1 m1 m2,
- Val.lessdef_list vl1 vl2 ->
- Mem.extends m1 m2 ->
- eval_operation genv sp op vl1 m1 = Some v1 ->
- exists v2, eval_operation genv sp op vl2 m2 = Some v2 /\ Val.lessdef v1 v2.
-Proof.
- intros. rewrite val_inject_list_lessdef in H.
- assert (exists v2 : val,
- eval_operation genv sp op vl2 m2 = Some v2
- /\ Val.inject (fun b => Some(b, 0)) v1 v2).
- eapply eval_operation_inj with (m1 := m1) (sp1 := sp).
- apply valid_pointer_extends; auto.
- apply weak_valid_pointer_extends; auto.
- apply weak_valid_pointer_no_overflow_extends.
- apply valid_different_pointers_extends; auto.
- intros. apply val_inject_lessdef. auto.
- apply val_inject_lessdef; auto.
- eauto.
- auto.
- destruct H2 as [v2 [A B]]. exists v2; split; auto. rewrite val_inject_lessdef; auto.
-Qed.
-
-Lemma eval_addressing_lessdef:
- forall sp addr vl1 vl2 v1,
- Val.lessdef_list vl1 vl2 ->
- eval_addressing genv sp addr vl1 = Some v1 ->
- exists v2, eval_addressing genv sp addr vl2 = Some v2 /\ Val.lessdef v1 v2.
-Proof.
- intros. rewrite val_inject_list_lessdef in H.
- assert (exists v2 : val,
- eval_addressing genv sp addr vl2 = Some v2
- /\ Val.inject (fun b => Some(b, 0)) v1 v2).
- eapply eval_addressing_inj with (sp1 := sp).
- intros. rewrite <- val_inject_lessdef; auto.
- rewrite <- val_inject_lessdef; auto.
- eauto. auto.
- destruct H1 as [v2 [A B]]. exists v2; split; auto. rewrite val_inject_lessdef; auto.
-Qed.
-
-End EVAL_LESSDEF.
-
-(** Compatibility of the evaluation functions with memory injections. *)
-
-Section EVAL_INJECT.
-
-Variable F V: Type.
-Variable genv: Genv.t F V.
-Variable f: meminj.
-Hypothesis globals: meminj_preserves_globals genv f.
-Variable sp1: block.
-Variable sp2: block.
-Variable delta: Z.
-Hypothesis sp_inj: f sp1 = Some(sp2, delta).
-
-Remark symbol_address_inject:
- forall id ofs, Val.inject f (Genv.symbol_address genv id ofs) (Genv.symbol_address genv id ofs).
-Proof.
- intros. unfold Genv.symbol_address. destruct (Genv.find_symbol genv id) eqn:?; auto.
- exploit (proj1 globals); eauto. intros.
- econstructor; eauto. rewrite Ptrofs.add_zero; auto.
-Qed.
-
-Lemma eval_condition_inject:
- forall cond vl1 vl2 b m1 m2,
- Val.inject_list f vl1 vl2 ->
- Mem.inject f m1 m2 ->
- eval_condition cond vl1 m1 = Some b ->
- eval_condition cond vl2 m2 = Some b.
-Proof.
- intros. eapply eval_condition_inj with (f := f) (m1 := m1); eauto.
- intros; eapply Mem.valid_pointer_inject_val; eauto.
- intros; eapply Mem.weak_valid_pointer_inject_val; eauto.
- intros; eapply Mem.weak_valid_pointer_inject_no_overflow; eauto.
- intros; eapply Mem.different_pointers_inject; eauto.
-Qed.
-
-Lemma eval_addressing_inject:
- forall addr vl1 vl2 v1,
- Val.inject_list f vl1 vl2 ->
- eval_addressing genv (Vptr sp1 Ptrofs.zero) addr vl1 = Some v1 ->
- exists v2,
- eval_addressing genv (Vptr sp2 Ptrofs.zero) (shift_stack_addressing delta addr) vl2 = Some v2
- /\ Val.inject f v1 v2.
-Proof.
- intros.
- rewrite eval_shift_stack_addressing.
- eapply eval_addressing_inj with (sp1 := Vptr sp1 Ptrofs.zero); eauto.
- intros. apply symbol_address_inject.
- econstructor; eauto. rewrite Ptrofs.add_zero_l; auto.
-Qed.
-
-Lemma eval_operation_inject:
- forall op vl1 vl2 v1 m1 m2,
- Val.inject_list f vl1 vl2 ->
- Mem.inject f m1 m2 ->
- eval_operation genv (Vptr sp1 Ptrofs.zero) op vl1 m1 = Some v1 ->
- exists v2,
- eval_operation genv (Vptr sp2 Ptrofs.zero) (shift_stack_operation delta op) vl2 m2 = Some v2
- /\ Val.inject f v1 v2.
-Proof.
- intros.
- rewrite eval_shift_stack_operation. simpl.
- eapply eval_operation_inj with (sp1 := Vptr sp1 Ptrofs.zero) (m1 := m1); eauto.
- intros; eapply Mem.valid_pointer_inject_val; eauto.
- intros; eapply Mem.weak_valid_pointer_inject_val; eauto.
- intros; eapply Mem.weak_valid_pointer_inject_no_overflow; eauto.
- intros; eapply Mem.different_pointers_inject; eauto.
- intros. apply symbol_address_inject.
- econstructor; eauto. rewrite Ptrofs.add_zero_l; auto.
-Qed.
-
-End EVAL_INJECT.
diff --git a/ia32/PrintOp.ml b/ia32/PrintOp.ml
deleted file mode 100644
index faa5bb5f..00000000
--- a/ia32/PrintOp.ml
+++ /dev/null
@@ -1,169 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Pretty-printing of operators, conditions, addressing modes *)
-
-open Printf
-open Camlcoq
-open Integers
-open Op
-
-let comparison_name = function
- | Ceq -> "=="
- | Cne -> "!="
- | Clt -> "<"
- | Cle -> "<="
- | Cgt -> ">"
- | Cge -> ">="
-
-let print_condition reg pp = function
- | (Ccomp c, [r1;r2]) ->
- fprintf pp "%a %ss %a" reg r1 (comparison_name c) reg r2
- | (Ccompu c, [r1;r2]) ->
- fprintf pp "%a %su %a" reg r1 (comparison_name c) reg r2
- | (Ccompimm(c, n), [r1]) ->
- fprintf pp "%a %ss %ld" reg r1 (comparison_name c) (camlint_of_coqint n)
- | (Ccompuimm(c, n), [r1]) ->
- fprintf pp "%a %su %lu" reg r1 (comparison_name c) (camlint_of_coqint n)
- | (Ccompl c, [r1;r2]) ->
- fprintf pp "%a %sls %a" reg r1 (comparison_name c) reg r2
- | (Ccomplu c, [r1;r2]) ->
- fprintf pp "%a %slu %a" reg r1 (comparison_name c) reg r2
- | (Ccomplimm(c, n), [r1]) ->
- fprintf pp "%a %sls %Ld" reg r1 (comparison_name c) (camlint64_of_coqint n)
- | (Ccompluimm(c, n), [r1]) ->
- fprintf pp "%a %slu %Lu" reg r1 (comparison_name c) (camlint64_of_coqint n)
- | (Ccompf c, [r1;r2]) ->
- fprintf pp "%a %sf %a" reg r1 (comparison_name c) reg r2
- | (Cnotcompf c, [r1;r2]) ->
- fprintf pp "%a not(%sf) %a" reg r1 (comparison_name c) reg r2
- | (Ccompfs c, [r1;r2]) ->
- fprintf pp "%a %sfs %a" reg r1 (comparison_name c) reg r2
- | (Cnotcompfs c, [r1;r2]) ->
- fprintf pp "%a not(%sfs) %a" reg r1 (comparison_name c) reg r2
- | (Cmaskzero n, [r1]) ->
- fprintf pp "%a & 0x%lx == 0" reg r1 (camlint_of_coqint n)
- | (Cmasknotzero n, [r1]) ->
- fprintf pp "%a & 0x%lx != 0" reg r1 (camlint_of_coqint n)
- | _ ->
- fprintf pp "<bad condition>"
-
-let print_addressing reg pp = function
- | Aindexed n, [r1] ->
- fprintf pp "%a + %s" reg r1 (Z.to_string n)
- | Aindexed2 n, [r1; r2] ->
- fprintf pp "%a + %a + %s" reg r1 reg r2 (Z.to_string n)
- | Ascaled(sc,n), [r1] ->
- fprintf pp "%a * %s + %s" reg r1 (Z.to_string sc) (Z.to_string n)
- | Aindexed2scaled(sc, n), [r1; r2] ->
- fprintf pp "%a + %a * %s + %s" reg r1 reg r2 (Z.to_string sc) (Z.to_string n)
- | Aglobal(id, ofs), [] -> fprintf pp "%s + %s" (extern_atom id) (Z.to_string ofs)
- | Abased(id, ofs), [r1] -> fprintf pp "%s + %s + %a" (extern_atom id) (Z.to_string ofs) reg r1
- | Abasedscaled(sc,id, ofs), [r1] -> fprintf pp "%s + %s + %a * %ld" (extern_atom id) (Z.to_string ofs) reg r1 (camlint_of_coqint sc)
- | Ainstack ofs, [] -> fprintf pp "stack(%s)" (Z.to_string ofs)
- | _ -> fprintf pp "<bad addressing>"
-
-let print_operation reg pp = function
- | Omove, [r1] -> reg pp r1
- | Ointconst n, [] -> fprintf pp "%ld" (camlint_of_coqint n)
- | Olongconst n, [] -> fprintf pp "%LdL" (camlint64_of_coqint n)
- | Ofloatconst n, [] -> fprintf pp "%.15F" (camlfloat_of_coqfloat n)
- | Osingleconst n, [] -> fprintf pp "%.15Ff" (camlfloat_of_coqfloat32 n)
- | Oindirectsymbol id, [] -> fprintf pp "&%s" (extern_atom id)
- | Ocast8signed, [r1] -> fprintf pp "int8signed(%a)" reg r1
- | Ocast8unsigned, [r1] -> fprintf pp "int8unsigned(%a)" reg r1
- | Ocast16signed, [r1] -> fprintf pp "int16signed(%a)" reg r1
- | Ocast16unsigned, [r1] -> fprintf pp "int16unsigned(%a)" reg r1
- | Oneg, [r1] -> fprintf pp "(- %a)" reg r1
- | Osub, [r1;r2] -> fprintf pp "%a - %a" reg r1 reg r2
- | Omul, [r1;r2] -> fprintf pp "%a * %a" reg r1 reg r2
- | Omulimm n, [r1] -> fprintf pp "%a * %ld" reg r1 (camlint_of_coqint n)
- | Omulhs, [r1;r2] -> fprintf pp "mulhs(%a,%a)" reg r1 reg r2
- | Omulhu, [r1;r2] -> fprintf pp "mulhu(%a,%a)" reg r1 reg r2
- | Odiv, [r1;r2] -> fprintf pp "%a /s %a" reg r1 reg r2
- | Odivu, [r1;r2] -> fprintf pp "%a /u %a" reg r1 reg r2
- | Omod, [r1;r2] -> fprintf pp "%a %%s %a" reg r1 reg r2
- | Omodu, [r1;r2] -> fprintf pp "%a %%u %a" reg r1 reg r2
- | Oand, [r1;r2] -> fprintf pp "%a & %a" reg r1 reg r2
- | Oandimm n, [r1] -> fprintf pp "%a & %ld" reg r1 (camlint_of_coqint n)
- | Oor, [r1;r2] -> fprintf pp "%a | %a" reg r1 reg r2
- | Oorimm n, [r1] -> fprintf pp "%a | %ld" reg r1 (camlint_of_coqint n)
- | Oxor, [r1;r2] -> fprintf pp "%a ^ %a" reg r1 reg r2
- | Oxorimm n, [r1] -> fprintf pp "%a ^ %ld" reg r1 (camlint_of_coqint n)
- | Onot, [r1] -> fprintf pp "not(%a)" reg r1
- | Oshl, [r1;r2] -> fprintf pp "%a << %a" reg r1 reg r2
- | Oshlimm n, [r1] -> fprintf pp "%a << %ld" reg r1 (camlint_of_coqint n)
- | Oshr, [r1;r2] -> fprintf pp "%a >>s %a" reg r1 reg r2
- | Oshrimm n, [r1] -> fprintf pp "%a >>s %ld" reg r1 (camlint_of_coqint n)
- | Oshrximm n, [r1] -> fprintf pp "%a >>x %ld" reg r1 (camlint_of_coqint n)
- | Oshru, [r1;r2] -> fprintf pp "%a >>u %a" reg r1 reg r2
- | Oshruimm n, [r1] -> fprintf pp "%a >>u %ld" reg r1 (camlint_of_coqint n)
- | Ororimm n, [r1] -> fprintf pp "%a ror %ld" reg r1 (camlint_of_coqint n)
- | Oshldimm n, [r1;r2] -> fprintf pp "(%a, %a) << %ld" reg r1 reg r2 (camlint_of_coqint n)
- | Olea addr, args -> print_addressing reg pp (addr, args); fprintf pp " (int)"
- | Omakelong, [r1;r2] -> fprintf pp "makelong(%a,%a)" reg r1 reg r2
- | Olowlong, [r1] -> fprintf pp "lowlong(%a)" reg r1
- | Ohighlong, [r1] -> fprintf pp "highlong(%a)" reg r1
- | Ocast32signed, [r1] -> fprintf pp "long32signed(%a)" reg r1
- | Ocast32unsigned, [r1] -> fprintf pp "long32unsigned(%a)" reg r1
- | Onegl, [r1] -> fprintf pp "(-l %a)" reg r1
- | Osubl, [r1;r2] -> fprintf pp "%a -l %a" reg r1 reg r2
- | Omull, [r1;r2] -> fprintf pp "%a *l %a" reg r1 reg r2
- | Omullimm n, [r1] -> fprintf pp "%a *l %Ld" reg r1 (camlint64_of_coqint n)
- | Omullhs, [r1;r2] -> fprintf pp "mullhs(%a,%a)" reg r1 reg r2
- | Omullhu, [r1;r2] -> fprintf pp "mullhu(%a,%a)" reg r1 reg r2
- | Odivl, [r1;r2] -> fprintf pp "%a /ls %a" reg r1 reg r2
- | Odivlu, [r1;r2] -> fprintf pp "%a /lu %a" reg r1 reg r2
- | Omodl, [r1;r2] -> fprintf pp "%a %%ls %a" reg r1 reg r2
- | Omodlu, [r1;r2] -> fprintf pp "%a %%lu %a" reg r1 reg r2
- | Oandl, [r1;r2] -> fprintf pp "%a &l %a" reg r1 reg r2
- | Oandlimm n, [r1] -> fprintf pp "%a &l %Ld" reg r1 (camlint64_of_coqint n)
- | Oorl, [r1;r2] -> fprintf pp "%a |l %a" reg r1 reg r2
- | Oorlimm n, [r1] -> fprintf pp "%a |l %Ld" reg r1 (camlint64_of_coqint n)
- | Oxorl, [r1;r2] -> fprintf pp "%a ^l %a" reg r1 reg r2
- | Oxorlimm n, [r1] -> fprintf pp "%a ^l %Ld" reg r1 (camlint64_of_coqint n)
- | Onotl, [r1] -> fprintf pp "notl(%a)" reg r1
- | Oshll, [r1;r2] -> fprintf pp "%a <<l %a" reg r1 reg r2
- | Oshllimm n, [r1] -> fprintf pp "%a <<l %ld" reg r1 (camlint_of_coqint n)
- | Oshrl, [r1;r2] -> fprintf pp "%a >>ls %a" reg r1 reg r2
- | Oshrlimm n, [r1] -> fprintf pp "%a >>ls %ld" reg r1 (camlint_of_coqint n)
- | Oshrxlimm n, [r1] -> fprintf pp "%a >>lx %ld" reg r1 (camlint_of_coqint n)
- | Oshrlu, [r1;r2] -> fprintf pp "%a >>lu %a" reg r1 reg r2
- | Oshrluimm n, [r1] -> fprintf pp "%a >>lu %ld" reg r1 (camlint_of_coqint n)
- | Ororlimm n, [r1] -> fprintf pp "%a rorl %ld" reg r1 (camlint_of_coqint n)
- | Oleal addr, args -> print_addressing reg pp (addr, args); fprintf pp " (long)"
- | Onegf, [r1] -> fprintf pp "negf(%a)" reg r1
- | Oabsf, [r1] -> fprintf pp "absf(%a)" reg r1
- | Oaddf, [r1;r2] -> fprintf pp "%a +f %a" reg r1 reg r2
- | Osubf, [r1;r2] -> fprintf pp "%a -f %a" reg r1 reg r2
- | Omulf, [r1;r2] -> fprintf pp "%a *f %a" reg r1 reg r2
- | Odivf, [r1;r2] -> fprintf pp "%a /f %a" reg r1 reg r2
- | Onegfs, [r1] -> fprintf pp "negfs(%a)" reg r1
- | Oabsfs, [r1] -> fprintf pp "absfs(%a)" reg r1
- | Oaddfs, [r1;r2] -> fprintf pp "%a +fs %a" reg r1 reg r2
- | Osubfs, [r1;r2] -> fprintf pp "%a -fs %a" reg r1 reg r2
- | Omulfs, [r1;r2] -> fprintf pp "%a *fs %a" reg r1 reg r2
- | Odivfs, [r1;r2] -> fprintf pp "%a /fs %a" reg r1 reg r2
- | Osingleoffloat, [r1] -> fprintf pp "singleoffloat(%a)" reg r1
- | Ofloatofsingle, [r1] -> fprintf pp "floatofsingle(%a)" reg r1
- | Ointoffloat, [r1] -> fprintf pp "intoffloat(%a)" reg r1
- | Ofloatofint, [r1] -> fprintf pp "floatofint(%a)" reg r1
- | Ointofsingle, [r1] -> fprintf pp "intofsingle(%a)" reg r1
- | Osingleofint, [r1] -> fprintf pp "singleofint(%a)" reg r1
- | Olongoffloat, [r1] -> fprintf pp "longoffloat(%a)" reg r1
- | Ofloatoflong, [r1] -> fprintf pp "floatoflong(%a)" reg r1
- | Olongofsingle, [r1] -> fprintf pp "longofsingle(%a)" reg r1
- | Osingleoflong, [r1] -> fprintf pp "singleoflong(%a)" reg r1
- | Ocmp c, args -> print_condition reg pp (c, args)
- | _ -> fprintf pp "<bad operator>"
-
-
diff --git a/ia32/SelectLong.vp b/ia32/SelectLong.vp
deleted file mode 100644
index b213e23f..00000000
--- a/ia32/SelectLong.vp
+++ /dev/null
@@ -1,347 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Instruction selection for 64-bit integer operations *)
-
-Require Import Coqlib.
-Require Import Compopts.
-Require Import AST Integers Floats.
-Require Import Op CminorSel.
-Require Import SelectOp SplitLong.
-
-Local Open Scope cminorsel_scope.
-Local Open Scope string_scope.
-
-Section SELECT.
-
-Context {hf: helper_functions}.
-
-Definition longconst (n: int64) : expr :=
- if Archi.splitlong then SplitLong.longconst n else Eop (Olongconst n) Enil.
-
-Definition is_longconst (e: expr) :=
- if Archi.splitlong then SplitLong.is_longconst e else
- match e with
- | Eop (Olongconst n) Enil => Some n
- | _ => None
- end.
-
-Definition intoflong (e: expr) :=
- if Archi.splitlong then SplitLong.intoflong e else
- match is_longconst e with
- | Some n => Eop (Ointconst (Int.repr (Int64.unsigned n))) Enil
- | None => Eop Olowlong (e ::: Enil)
- end.
-
-Definition longofint (e: expr) :=
- if Archi.splitlong then SplitLong.longofint e else
- match is_intconst e with
- | Some n => longconst (Int64.repr (Int.signed n))
- | None => Eop Ocast32signed (e ::: Enil)
- end.
-
-Definition longofintu (e: expr) :=
- if Archi.splitlong then SplitLong.longofintu e else
- match is_intconst e with
- | Some n => longconst (Int64.repr (Int.unsigned n))
- | None => Eop Ocast32unsigned (e ::: Enil)
- end.
-
-Nondetfunction notl (e: expr) :=
- if Archi.splitlong then SplitLong.notl e else
- match e with
- | Eop (Olongconst n) Enil => longconst (Int64.not n)
- | Eop Onotl (t1:::Enil) => t1
- | _ => Eop Onotl (e:::Enil)
- end.
-
-Nondetfunction andlimm (n1: int64) (e2: expr) :=
- if Int64.eq n1 Int64.zero then longconst Int64.zero else
- if Int64.eq n1 Int64.mone then e2 else
- match e2 with
- | Eop (Olongconst n2) Enil =>
- longconst (Int64.and n1 n2)
- | Eop (Oandlimm n2) (t2:::Enil) =>
- Eop (Oandlimm (Int64.and n1 n2)) (t2:::Enil)
- | _ =>
- Eop (Oandlimm n1) (e2:::Enil)
- end.
-
-Nondetfunction andl (e1: expr) (e2: expr) :=
- if Archi.splitlong then SplitLong.andl e1 e2 else
- match e1, e2 with
- | Eop (Olongconst n1) Enil, t2 => andlimm n1 t2
- | t1, Eop (Olongconst n2) Enil => andlimm n2 t1
- | _, _ => Eop Oandl (e1:::e2:::Enil)
- end.
-
-Nondetfunction orlimm (n1: int64) (e2: expr) :=
- if Int64.eq n1 Int64.zero then e2 else
- if Int64.eq n1 Int64.mone then longconst Int64.mone else
- match e2 with
- | Eop (Olongconst n2) Enil => longconst (Int64.or n1 n2)
- | Eop (Oorlimm n2) (t2:::Enil) => Eop (Oorlimm (Int64.or n1 n2)) (t2:::Enil)
- | _ => Eop (Oorlimm n1) (e2:::Enil)
- end.
-
-Nondetfunction orl (e1: expr) (e2: expr) :=
- if Archi.splitlong then SplitLong.orl e1 e2 else
- match e1, e2 with
- | Eop (Olongconst n1) Enil, t2 => orlimm n1 t2
- | t1, Eop (Olongconst n2) Enil => orlimm n2 t1
- | Eop (Oshllimm n1) (t1:::Enil), Eop (Oshrluimm n2) (t2:::Enil) =>
- if Int.eq (Int.add n1 n2) Int64.iwordsize' && same_expr_pure t1 t2
- then Eop (Ororlimm n2) (t1:::Enil)
- else Eop Oorl (e1:::e2:::Enil)
- | Eop (Oshrluimm n2) (t2:::Enil), Eop (Oshllimm n1) (t1:::Enil) =>
- if Int.eq (Int.add n1 n2) Int64.iwordsize' && same_expr_pure t1 t2
- then Eop (Ororlimm n2) (t1:::Enil)
- else Eop Oorl (e1:::e2:::Enil)
- | _, _ =>
- Eop Oorl (e1:::e2:::Enil)
- end.
-
-Nondetfunction xorlimm (n1: int64) (e2: expr) :=
- if Int64.eq n1 Int64.zero then e2 else
- if Int64.eq n1 Int64.mone then notl e2 else
- match e2 with
- | Eop (Olongconst n2) Enil => longconst (Int64.xor n1 n2)
- | Eop (Oxorlimm n2) (t2:::Enil) => Eop (Oxorlimm (Int64.xor n1 n2)) (t2:::Enil)
- | Eop Onotl (t2:::Enil) => Eop (Oxorlimm (Int64.not n1)) (t2:::Enil)
- | _ => Eop (Oxorlimm n1) (e2:::Enil)
- end.
-
-Nondetfunction xorl (e1: expr) (e2: expr) :=
- if Archi.splitlong then SplitLong.xorl e1 e2 else
- match e1, e2 with
- | Eop (Olongconst n1) Enil, t2 => xorlimm n1 t2
- | t1, Eop (Olongconst n2) Enil => xorlimm n2 t1
- | _, _ => Eop Oxorl (e1:::e2:::Enil)
- end.
-
-Nondetfunction shllimm (e1: expr) (n: int) :=
- if Archi.splitlong then SplitLong.shllimm e1 n else
- if Int.eq n Int.zero then e1 else
- if negb (Int.ltu n Int64.iwordsize') then
- Eop Oshll (e1:::Eop (Ointconst n) Enil:::Enil)
- else
- match e1 with
- | Eop (Olongconst n1) Enil =>
- Eop (Olongconst(Int64.shl' n1 n)) Enil
- | Eop (Oshllimm n1) (t1:::Enil) =>
- if Int.ltu (Int.add n n1) Int64.iwordsize'
- then Eop (Oshllimm (Int.add n n1)) (t1:::Enil)
- else Eop (Oshllimm n) (e1:::Enil)
- | Eop (Oleal (Aindexed n1)) (t1:::Enil) =>
- if shift_is_scale n
- then Eop (Oleal (Ascaled (Int64.unsigned (Int64.shl' Int64.one n))
- (Int64.unsigned (Int64.shl' (Int64.repr n1) n)))) (t1:::Enil)
- else Eop (Oshllimm n) (e1:::Enil)
- | _ =>
- if shift_is_scale n
- then Eop (Oleal (Ascaled (Int64.unsigned (Int64.shl' Int64.one n)) 0)) (e1:::Enil)
- else Eop (Oshllimm n) (e1:::Enil)
- end.
-
-Nondetfunction shrluimm (e1: expr) (n: int) :=
- if Archi.splitlong then SplitLong.shrluimm e1 n else
- if Int.eq n Int.zero then e1 else
- if negb (Int.ltu n Int64.iwordsize') then
- Eop Oshrlu (e1:::Eop (Ointconst n) Enil:::Enil)
- else
- match e1 with
- | Eop (Olongconst n1) Enil =>
- Eop (Olongconst(Int64.shru' n1 n)) Enil
- | Eop (Oshrluimm n1) (t1:::Enil) =>
- if Int.ltu (Int.add n n1) Int64.iwordsize'
- then Eop (Oshrluimm (Int.add n n1)) (t1:::Enil)
- else Eop (Oshrluimm n) (e1:::Enil)
- | _ =>
- Eop (Oshrluimm n) (e1:::Enil)
- end.
-
-Nondetfunction shrlimm (e1: expr) (n: int) :=
- if Archi.splitlong then SplitLong.shrlimm e1 n else
- if Int.eq n Int.zero then e1 else
- if negb (Int.ltu n Int64.iwordsize') then
- Eop Oshrl (e1:::Eop (Ointconst n) Enil:::Enil)
- else
- match e1 with
- | Eop (Olongconst n1) Enil =>
- Eop (Olongconst(Int64.shr' n1 n)) Enil
- | Eop (Oshrlimm n1) (t1:::Enil) =>
- if Int.ltu (Int.add n n1) Int64.iwordsize'
- then Eop (Oshrlimm (Int.add n n1)) (t1:::Enil)
- else Eop (Oshrlimm n) (e1:::Enil)
- | _ =>
- Eop (Oshrlimm n) (e1:::Enil)
- end.
-
-Definition shll (e1: expr) (e2: expr) :=
- if Archi.splitlong then SplitLong.shll e1 e2 else
- match is_intconst e2 with
- | Some n2 => shllimm e1 n2
- | None => Eop Oshll (e1:::e2:::Enil)
- end.
-
-Definition shrl (e1: expr) (e2: expr) :=
- if Archi.splitlong then SplitLong.shrl e1 e2 else
- match is_intconst e2 with
- | Some n2 => shrlimm e1 n2
- | None => Eop Oshrl (e1:::e2:::Enil)
- end.
-
-Definition shrlu (e1: expr) (e2: expr) :=
- if Archi.splitlong then SplitLong.shrlu e1 e2 else
- match is_intconst e2 with
- | Some n2 => shrluimm e1 n2
- | _ => Eop Oshrlu (e1:::e2:::Enil)
- end.
-
-Nondetfunction addlimm (n: int64) (e: expr) :=
- if Int64.eq n Int64.zero then e else
- match e with
- | Eop (Olongconst m) Enil => longconst (Int64.add n m)
- | Eop (Oleal addr) args => Eop (Oleal (offset_addressing_total addr (Int64.signed n))) args
- | _ => Eop (Oleal (Aindexed (Int64.signed n))) (e ::: Enil)
- end.
-
-Nondetfunction addl (e1: expr) (e2: expr) :=
- if Archi.splitlong then SplitLong.addl e1 e2 else
- match e1, e2 with
- | Eop (Olongconst n1) Enil, t2 => addlimm n1 t2
- | t1, Eop (Olongconst n2) Enil => addlimm n2 t1
- | Eop (Oleal (Aindexed n1)) (t1:::Enil), Eop (Oleal (Aindexed n2)) (t2:::Enil) =>
- Eop (Oleal (Aindexed2 (n1 + n2))) (t1:::t2:::Enil)
- | Eop (Oleal (Aindexed n1)) (t1:::Enil), Eop (Oleal (Ascaled sc n2)) (t2:::Enil) =>
- Eop (Oleal (Aindexed2scaled sc (n1 + n2))) (t1:::t2:::Enil)
- | Eop (Oleal (Ascaled sc n1)) (t1:::Enil), Eop (Oleal (Aindexed n2)) (t2:::Enil) =>
- Eop (Oleal (Aindexed2scaled sc (n1 + n2))) (t2:::t1:::Enil)
- | Eop (Oleal (Ascaled sc n)) (t1:::Enil), t2 =>
- Eop (Oleal (Aindexed2scaled sc n)) (t2:::t1:::Enil)
- | t1, Eop (Oleal (Ascaled sc n)) (t2:::Enil) =>
- Eop (Oleal (Aindexed2scaled sc n)) (t1:::t2:::Enil)
- | Eop (Oleal (Aindexed n)) (t1:::Enil), t2 =>
- Eop (Oleal (Aindexed2 n)) (t1:::t2:::Enil)
- | t1, Eop (Oleal (Aindexed n)) (t2:::Enil) =>
- Eop (Oleal (Aindexed2 n)) (t1:::t2:::Enil)
- | _, _ =>
- Eop (Oleal (Aindexed2 0)) (e1:::e2:::Enil)
- end.
-
-Definition negl (e: expr) :=
- if Archi.splitlong then SplitLong.negl e else
- match is_longconst e with
- | Some n => longconst (Int64.neg n)
- | None => Eop Onegl (e ::: Enil)
- end.
-
-Nondetfunction subl (e1: expr) (e2: expr) :=
- if Archi.splitlong then SplitLong.subl e1 e2 else
- match e1, e2 with
- | t1, Eop (Olongconst n2) Enil => addlimm (Int64.neg n2) t1
- | Eop (Oleal (Aindexed n1)) (t1:::Enil), Eop (Oleal (Aindexed n2)) (t2:::Enil) =>
- addlimm (Int64.repr (n1 - n2)) (Eop Osubl (t1:::t2:::Enil))
- | Eop (Oleal (Aindexed n1)) (t1:::Enil), t2 =>
- addlimm (Int64.repr n1) (Eop Osubl (t1:::t2:::Enil))
- | t1, Eop (Oleal (Aindexed n2)) (t2:::Enil) =>
- addlimm (Int64.repr (- n2)) (Eop Osubl (t1:::t2:::Enil))
- | _, _ =>
- Eop Osubl (e1:::e2:::Enil)
- end.
-
-Definition mullimm_base (n1: int64) (e2: expr) :=
- match Int64.one_bits' n1 with
- | i :: nil =>
- shllimm e2 i
- | i :: j :: nil =>
- Elet e2 (addl (shllimm (Eletvar 0) i) (shllimm (Eletvar 0) j))
- | _ =>
- Eop (Omullimm n1) (e2:::Enil)
- end.
-
-Nondetfunction mullimm (n1: int64) (e2: expr) :=
- if Archi.splitlong then SplitLong.mullimm n1 e2
- else if Int64.eq n1 Int64.zero then longconst Int64.zero
- else if Int64.eq n1 Int64.one then e2
- else match e2 with
- | Eop (Olongconst n2) Enil => longconst (Int64.mul n1 n2)
- | Eop (Oleal (Aindexed n2)) (t2:::Enil) => addlimm (Int64.mul n1 (Int64.repr n2)) (mullimm_base n1 t2)
- | _ => mullimm_base n1 e2
- end.
-
-Nondetfunction mull (e1: expr) (e2: expr) :=
- if Archi.splitlong then SplitLong.mull e1 e2 else
- match e1, e2 with
- | Eop (Olongconst n1) Enil, t2 => mullimm n1 t2
- | t1, Eop (Olongconst n2) Enil => mullimm n2 t1
- | _, _ => Eop Omull (e1:::e2:::Enil)
- end.
-
-Definition mullhu (e1: expr) (n2: int64) :=
- if Archi.splitlong then SplitLong.mullhu e1 n2 else
- Eop Omullhu (e1 ::: longconst n2 ::: Enil).
-
-Definition mullhs (e1: expr) (n2: int64) :=
- if Archi.splitlong then SplitLong.mullhs e1 n2 else
- Eop Omullhs (e1 ::: longconst n2 ::: Enil).
-
-Definition shrxlimm (e: expr) (n: int) :=
- if Archi.splitlong then SplitLong.shrxlimm e n else
- if Int.eq n Int.zero then e else Eop (Oshrxlimm n) (e ::: Enil).
-
-Definition divlu_base (e1: expr) (e2: expr) :=
- if Archi.splitlong then SplitLong.divlu_base e1 e2 else Eop Odivlu (e1:::e2:::Enil).
-Definition modlu_base (e1: expr) (e2: expr) :=
- if Archi.splitlong then SplitLong.modlu_base e1 e2 else Eop Omodlu (e1:::e2:::Enil).
-Definition divls_base (e1: expr) (e2: expr) :=
- if Archi.splitlong then SplitLong.divls_base e1 e2 else Eop Odivl (e1:::e2:::Enil).
-Definition modls_base (e1: expr) (e2: expr) :=
- if Archi.splitlong then SplitLong.modls_base e1 e2 else Eop Omodl (e1:::e2:::Enil).
-
-Definition cmplu (c: comparison) (e1 e2: expr) :=
- if Archi.splitlong then SplitLong.cmplu c e1 e2 else
- match is_longconst e1, is_longconst e2 with
- | Some n1, Some n2 =>
- Eop (Ointconst (if Int64.cmpu c n1 n2 then Int.one else Int.zero)) Enil
- | Some n1, None => Eop (Ocmp (Ccompluimm (swap_comparison c) n1)) (e2:::Enil)
- | None, Some n2 => Eop (Ocmp (Ccompluimm c n2)) (e1:::Enil)
- | None, None => Eop (Ocmp (Ccomplu c)) (e1:::e2:::Enil)
- end.
-
-Definition cmpl (c: comparison) (e1 e2: expr) :=
- if Archi.splitlong then SplitLong.cmpl c e1 e2 else
- match is_longconst e1, is_longconst e2 with
- | Some n1, Some n2 =>
- Eop (Ointconst (if Int64.cmp c n1 n2 then Int.one else Int.zero)) Enil
- | Some n1, None => Eop (Ocmp (Ccomplimm (swap_comparison c) n1)) (e2:::Enil)
- | None, Some n2 => Eop (Ocmp (Ccomplimm c n2)) (e1:::Enil)
- | None, None => Eop (Ocmp (Ccompl c)) (e1:::e2:::Enil)
- end.
-
-Definition longoffloat (e: expr) :=
- if Archi.splitlong then SplitLong.longoffloat e else
- Eop Olongoffloat (e:::Enil).
-
-Definition floatoflong (e: expr) :=
- if Archi.splitlong then SplitLong.floatoflong e else
- Eop Ofloatoflong (e:::Enil).
-
-Definition longofsingle (e: expr) :=
- if Archi.splitlong then SplitLong.longofsingle e else
- Eop Olongofsingle (e:::Enil).
-
-Definition singleoflong (e: expr) :=
- if Archi.splitlong then SplitLong.singleoflong e else
- Eop Osingleoflong (e:::Enil).
-
-End SELECT.
diff --git a/ia32/SelectLongproof.v b/ia32/SelectLongproof.v
deleted file mode 100644
index a3d2bb19..00000000
--- a/ia32/SelectLongproof.v
+++ /dev/null
@@ -1,557 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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 of instruction selection for 64-bit integer operations *)
-
-Require Import String Coqlib Maps Integers Floats Errors.
-Require Archi.
-Require Import AST Values Memory Globalenvs Events.
-Require Import Cminor Op CminorSel.
-Require Import SelectOp SelectOpproof SplitLong SplitLongproof.
-Require Import SelectLong.
-
-Open Local Scope cminorsel_scope.
-Open Local Scope string_scope.
-
-(** * Correctness of the instruction selection functions for 64-bit operators *)
-
-Section CMCONSTR.
-
-Variable prog: program.
-Variable hf: helper_functions.
-Hypothesis HELPERS: helper_functions_declared prog hf.
-Let ge := Genv.globalenv prog.
-Variable sp: val.
-Variable e: env.
-Variable m: mem.
-
-Definition unary_constructor_sound (cstr: expr -> expr) (sem: val -> val) : Prop :=
- forall le a x,
- eval_expr ge sp e m le a x ->
- exists v, eval_expr ge sp e m le (cstr a) v /\ Val.lessdef (sem x) v.
-
-Definition binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> val -> val) : Prop :=
- forall le a x b y,
- eval_expr ge sp e m le a x ->
- eval_expr ge sp e m le b y ->
- exists v, eval_expr ge sp e m le (cstr a b) v /\ Val.lessdef (sem x y) v.
-
-Definition partial_unary_constructor_sound (cstr: expr -> expr) (sem: val -> option val) : Prop :=
- forall le a x y,
- eval_expr ge sp e m le a x ->
- sem x = Some y ->
- exists v, eval_expr ge sp e m le (cstr a) v /\ Val.lessdef y v.
-
-Definition partial_binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> val -> option val) : Prop :=
- forall le a x b y z,
- eval_expr ge sp e m le a x ->
- eval_expr ge sp e m le b y ->
- sem x y = Some z ->
- exists v, eval_expr ge sp e m le (cstr a b) v /\ Val.lessdef z v.
-
-Theorem eval_longconst:
- forall le n, eval_expr ge sp e m le (longconst n) (Vlong n).
-Proof.
- unfold longconst; intros; destruct Archi.splitlong.
- apply SplitLongproof.eval_longconst.
- EvalOp.
-Qed.
-
-Lemma is_longconst_sound:
- forall v a n le,
- is_longconst a = Some n -> eval_expr ge sp e m le a v -> v = Vlong n.
-Proof with (try discriminate).
- intros. unfold is_longconst in *. destruct Archi.splitlong.
- eapply SplitLongproof.is_longconst_sound; eauto.
- assert (a = Eop (Olongconst n) Enil).
- { destruct a... destruct o... destruct e0... congruence. }
- subst a. InvEval. auto.
-Qed.
-
-Theorem eval_intoflong: unary_constructor_sound intoflong Val.loword.
-Proof.
- unfold intoflong; destruct Archi.splitlong. apply SplitLongproof.eval_intoflong.
- red; intros. destruct (is_longconst a) as [n|] eqn:C.
-- TrivialExists. simpl. erewrite (is_longconst_sound x) by eauto. auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_longofintu: unary_constructor_sound longofintu Val.longofintu.
-Proof.
- unfold longofintu; destruct Archi.splitlong. apply SplitLongproof.eval_longofintu.
- red; intros. destruct (is_intconst a) as [n|] eqn:C.
-- econstructor; split. apply eval_longconst.
- exploit is_intconst_sound; eauto. intros; subst x. auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_longofint: unary_constructor_sound longofint Val.longofint.
-Proof.
- unfold longofint; destruct Archi.splitlong. apply SplitLongproof.eval_longofint.
- red; intros. destruct (is_intconst a) as [n|] eqn:C.
-- econstructor; split. apply eval_longconst.
- exploit is_intconst_sound; eauto. intros; subst x. auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_notl: unary_constructor_sound notl Val.notl.
-Proof.
- unfold notl; destruct Archi.splitlong. apply SplitLongproof.eval_notl.
- red; intros. destruct (notl_match a).
-- InvEval. econstructor; split. apply eval_longconst. auto.
-- InvEval. subst. exists v1; split; auto. destruct v1; simpl; auto. rewrite Int64.not_involutive; auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_andlimm: forall n, unary_constructor_sound (andlimm n) (fun v => Val.andl v (Vlong n)).
-Proof.
- unfold andlimm; intros; red; intros.
- predSpec Int64.eq Int64.eq_spec n Int64.zero.
- exists (Vlong Int64.zero); split. apply eval_longconst.
- subst. destruct x; simpl; auto. rewrite Int64.and_zero; auto.
- predSpec Int64.eq Int64.eq_spec n Int64.mone.
- exists x; split. assumption.
- subst. destruct x; simpl; auto. rewrite Int64.and_mone; auto.
- destruct (andlimm_match a); InvEval; subst.
-- econstructor; split. apply eval_longconst. simpl. rewrite Int64.and_commut; auto.
-- TrivialExists. simpl. rewrite Val.andl_assoc. rewrite Int64.and_commut; auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_andl: binary_constructor_sound andl Val.andl.
-Proof.
- unfold andl; destruct Archi.splitlong. apply SplitLongproof.eval_andl.
- red; intros. destruct (andl_match a b).
-- InvEval. rewrite Val.andl_commut. apply eval_andlimm; auto.
-- InvEval. apply eval_andlimm; auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_orlimm: forall n, unary_constructor_sound (orlimm n) (fun v => Val.orl v (Vlong n)).
-Proof.
- unfold orlimm; intros; red; intros.
- predSpec Int64.eq Int64.eq_spec n Int64.zero.
- exists x; split; auto. subst. destruct x; simpl; auto. rewrite Int64.or_zero; auto.
- predSpec Int64.eq Int64.eq_spec n Int64.mone.
- econstructor; split. apply eval_longconst. subst. destruct x; simpl; auto. rewrite Int64.or_mone; auto.
- destruct (orlimm_match a); InvEval; subst.
-- econstructor; split. apply eval_longconst. simpl. rewrite Int64.or_commut; auto.
-- TrivialExists. simpl. rewrite Val.orl_assoc. rewrite Int64.or_commut; auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_orl: binary_constructor_sound orl Val.orl.
-Proof.
- unfold orl; destruct Archi.splitlong. apply SplitLongproof.eval_orl.
- red; intros.
- assert (DEFAULT: exists v, eval_expr ge sp e m le (Eop Oorl (a:::b:::Enil)) v /\ Val.lessdef (Val.orl x y) v) by TrivialExists.
- assert (ROR: forall v n1 n2,
- Int.add n1 n2 = Int64.iwordsize' ->
- Val.lessdef (Val.orl (Val.shll v (Vint n1)) (Val.shrlu v (Vint n2)))
- (Val.rorl v (Vint n2))).
- { intros. destruct v; simpl; auto.
- destruct (Int.ltu n1 Int64.iwordsize') eqn:N1; auto.
- destruct (Int.ltu n2 Int64.iwordsize') eqn:N2; auto.
- simpl. rewrite <- Int64.or_ror'; auto. }
- destruct (orl_match a b).
-- InvEval. rewrite Val.orl_commut. apply eval_orlimm; auto.
-- InvEval. apply eval_orlimm; auto.
-- predSpec Int.eq Int.eq_spec (Int.add n1 n2) Int64.iwordsize'; auto.
- destruct (same_expr_pure t1 t2) eqn:?; auto.
- InvEval. exploit eval_same_expr; eauto. intros [EQ1 EQ2]; subst.
- exists (Val.rorl v0 (Vint n2)); split. EvalOp. apply ROR; auto.
-- predSpec Int.eq Int.eq_spec (Int.add n1 n2) Int64.iwordsize'; auto.
- destruct (same_expr_pure t1 t2) eqn:?; auto.
- InvEval. exploit eval_same_expr; eauto. intros [EQ1 EQ2]; subst.
- exists (Val.rorl v1 (Vint n2)); split. EvalOp. rewrite Val.orl_commut. apply ROR; auto.
-- apply DEFAULT.
-Qed.
-
-Theorem eval_xorlimm: forall n, unary_constructor_sound (xorlimm n) (fun v => Val.xorl v (Vlong n)).
-Proof.
- unfold xorlimm; intros; red; intros.
- predSpec Int64.eq Int64.eq_spec n Int64.zero.
- exists x; split; auto. subst. destruct x; simpl; auto. rewrite Int64.xor_zero; auto.
- predSpec Int64.eq Int64.eq_spec n Int64.mone.
- replace (Val.xorl x (Vlong n)) with (Val.notl x). apply eval_notl; auto.
- subst n. destruct x; simpl; auto.
- destruct (xorlimm_match a); InvEval; subst.
-- econstructor; split. apply eval_longconst. simpl. rewrite Int64.xor_commut; auto.
-- TrivialExists. simpl. rewrite Val.xorl_assoc. rewrite Int64.xor_commut; auto.
-- TrivialExists. simpl. destruct v1; simpl; auto. unfold Int64.not.
- rewrite Int64.xor_assoc. apply f_equal. apply f_equal. apply f_equal.
- apply Int64.xor_commut.
-- TrivialExists.
-Qed.
-
-Theorem eval_xorl: binary_constructor_sound xorl Val.xorl.
-Proof.
- unfold xorl; destruct Archi.splitlong. apply SplitLongproof.eval_xorl.
- red; intros. destruct (xorl_match a b).
-- InvEval. rewrite Val.xorl_commut. apply eval_xorlimm; auto.
-- InvEval. apply eval_xorlimm; auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_shllimm: forall n, unary_constructor_sound (fun e => shllimm e n) (fun v => Val.shll v (Vint n)).
-Proof.
- intros; unfold shllimm. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shllimm; auto.
- red; intros.
- predSpec Int.eq Int.eq_spec n Int.zero.
- exists x; split; auto. subst n; destruct x; simpl; auto.
- destruct (Int.ltu Int.zero Int64.iwordsize'); auto.
- change (Int64.shl' i Int.zero) with (Int64.shl i Int64.zero). rewrite Int64.shl_zero; auto.
- destruct (Int.ltu n Int64.iwordsize') eqn:LT; simpl.
- assert (DEFAULT: exists v, eval_expr ge sp e m le (Eop (Oshllimm n) (a:::Enil)) v
- /\ Val.lessdef (Val.shll x (Vint n)) v) by TrivialExists.
- destruct (shllimm_match a); InvEval.
-- TrivialExists. simpl; rewrite LT; auto.
-- destruct (Int.ltu (Int.add n n1) Int64.iwordsize') eqn:LT'; auto.
- subst. econstructor; split. EvalOp. simpl; eauto.
- destruct v1; simpl; auto. rewrite LT'.
- destruct (Int.ltu n1 Int64.iwordsize') eqn:LT1; auto.
- simpl; rewrite LT. rewrite Int.add_commut, Int64.shl'_shl'; auto. rewrite Int.add_commut; auto.
-- destruct (shift_is_scale n); auto.
- TrivialExists. simpl. subst x. destruct v1; simpl; auto.
- rewrite LT. rewrite ! Int64.repr_unsigned. rewrite Int64.shl'_one_two_p.
- rewrite ! Int64.shl'_mul_two_p. rewrite Int64.mul_add_distr_l. auto.
- destruct Archi.ptr64; reflexivity.
-- destruct (shift_is_scale n); auto.
- TrivialExists. simpl. destruct x; simpl; auto.
- rewrite LT. rewrite ! Int64.repr_unsigned. rewrite Int64.shl'_one_two_p.
- rewrite ! Int64.shl'_mul_two_p. rewrite Int64.add_zero. auto.
-- TrivialExists. constructor; eauto. constructor. EvalOp. simpl; eauto. constructor. auto.
-Qed.
-
-Theorem eval_shrluimm: forall n, unary_constructor_sound (fun e => shrluimm e n) (fun v => Val.shrlu v (Vint n)).
-Proof.
- intros; unfold shrluimm. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrluimm; auto.
- red; intros.
- predSpec Int.eq Int.eq_spec n Int.zero.
- exists x; split; auto. subst n; destruct x; simpl; auto.
- destruct (Int.ltu Int.zero Int64.iwordsize'); auto.
- change (Int64.shru' i Int.zero) with (Int64.shru i Int64.zero). rewrite Int64.shru_zero; auto.
- destruct (Int.ltu n Int64.iwordsize') eqn:LT; simpl.
- assert (DEFAULT: exists v, eval_expr ge sp e m le (Eop (Oshrluimm n) (a:::Enil)) v
- /\ Val.lessdef (Val.shrlu x (Vint n)) v) by TrivialExists.
- destruct (shrluimm_match a); InvEval.
-- TrivialExists. simpl; rewrite LT; auto.
-- destruct (Int.ltu (Int.add n n1) Int64.iwordsize') eqn:LT'; auto.
- subst. econstructor; split. EvalOp. simpl; eauto.
- destruct v1; simpl; auto. rewrite LT'.
- destruct (Int.ltu n1 Int64.iwordsize') eqn:LT1; auto.
- simpl; rewrite LT. rewrite Int.add_commut, Int64.shru'_shru'; auto. rewrite Int.add_commut; auto.
-- apply DEFAULT.
-- TrivialExists. constructor; eauto. constructor. EvalOp. simpl; eauto. constructor. auto.
-Qed.
-
-Theorem eval_shrlimm: forall n, unary_constructor_sound (fun e => shrlimm e n) (fun v => Val.shrl v (Vint n)).
-Proof.
- intros; unfold shrlimm. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrlimm; auto.
- red; intros.
- predSpec Int.eq Int.eq_spec n Int.zero.
- exists x; split; auto. subst n; destruct x; simpl; auto.
- destruct (Int.ltu Int.zero Int64.iwordsize'); auto.
- change (Int64.shr' i Int.zero) with (Int64.shr i Int64.zero). rewrite Int64.shr_zero; auto.
- destruct (Int.ltu n Int64.iwordsize') eqn:LT; simpl.
- assert (DEFAULT: exists v, eval_expr ge sp e m le (Eop (Oshrlimm n) (a:::Enil)) v
- /\ Val.lessdef (Val.shrl x (Vint n)) v) by TrivialExists.
- destruct (shrlimm_match a); InvEval.
-- TrivialExists. simpl; rewrite LT; auto.
-- destruct (Int.ltu (Int.add n n1) Int64.iwordsize') eqn:LT'; auto.
- subst. econstructor; split. EvalOp. simpl; eauto.
- destruct v1; simpl; auto. rewrite LT'.
- destruct (Int.ltu n1 Int64.iwordsize') eqn:LT1; auto.
- simpl; rewrite LT. rewrite Int.add_commut, Int64.shr'_shr'; auto. rewrite Int.add_commut; auto.
-- apply DEFAULT.
-- TrivialExists. constructor; eauto. constructor. EvalOp. simpl; eauto. constructor. auto.
-Qed.
-
-Theorem eval_shll: binary_constructor_sound shll Val.shll.
-Proof.
- unfold shll. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shll; auto.
- red; intros. destruct (is_intconst b) as [n2|] eqn:C.
-- exploit is_intconst_sound; eauto. intros EQ; subst y. apply eval_shllimm; auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_shrlu: binary_constructor_sound shrlu Val.shrlu.
-Proof.
- unfold shrlu. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrlu; auto.
- red; intros. destruct (is_intconst b) as [n2|] eqn:C.
-- exploit is_intconst_sound; eauto. intros EQ; subst y. apply eval_shrluimm; auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_shrl: binary_constructor_sound shrl Val.shrl.
-Proof.
- unfold shrl. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrl; auto.
- red; intros. destruct (is_intconst b) as [n2|] eqn:C.
-- exploit is_intconst_sound; eauto. intros EQ; subst y. apply eval_shrlimm; auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_negl: unary_constructor_sound negl Val.negl.
-Proof.
- unfold negl. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_negl; auto.
- red; intros. destruct (is_longconst a) as [n|] eqn:C.
-- exploit is_longconst_sound; eauto. intros EQ; subst x.
- econstructor; split. apply eval_longconst. auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_addlimm: forall n, unary_constructor_sound (addlimm n) (fun v => Val.addl v (Vlong n)).
-Proof.
- unfold addlimm; intros; red; intros.
- predSpec Int64.eq Int64.eq_spec n Int64.zero.
- subst. exists x; split; auto.
- destruct x; simpl; auto.
- rewrite Int64.add_zero; auto.
- destruct Archi.ptr64; auto. rewrite Ptrofs.add_zero; auto.
- destruct (addlimm_match a); InvEval.
-- econstructor; split. apply eval_longconst. rewrite Int64.add_commut; auto.
-- inv H. simpl in H6. TrivialExists. simpl.
- erewrite eval_offset_addressing_total_64 by eauto. rewrite Int64.repr_signed; auto.
-- TrivialExists. simpl. rewrite Int64.repr_signed; auto.
-Qed.
-
-Theorem eval_addl: binary_constructor_sound addl Val.addl.
-Proof.
- assert (A: forall x y, Int64.repr (x + y) = Int64.add (Int64.repr x) (Int64.repr y)).
- { intros; apply Int64.eqm_samerepr; auto with ints. }
- assert (B: forall id ofs n, Archi.ptr64 = true ->
- Genv.symbol_address ge id (Ptrofs.add ofs (Ptrofs.repr n)) =
- Val.addl (Genv.symbol_address ge id ofs) (Vlong (Int64.repr n))).
- { intros. replace (Ptrofs.repr n) with (Ptrofs.of_int64 (Int64.repr n)) by auto with ptrofs.
- apply Genv.shift_symbol_address_64; auto. }
- unfold addl. destruct Archi.splitlong eqn:SL.
- apply SplitLongproof.eval_addl. apply Archi.splitlong_ptr32; auto.
- red; intros; destruct (addl_match a b); InvEval.
-- rewrite Val.addl_commut. apply eval_addlimm; auto.
-- apply eval_addlimm; auto.
-- subst. TrivialExists. simpl. rewrite A, Val.addl_permut_4. auto.
-- subst. TrivialExists. simpl. rewrite A, Val.addl_assoc. decEq; decEq. rewrite Val.addl_permut. auto.
-- subst. TrivialExists. simpl. rewrite A, Val.addl_permut_4. rewrite <- Val.addl_permut. rewrite <- Val.addl_assoc. auto.
-- subst. TrivialExists. simpl. rewrite Val.addl_commut; auto.
-- subst. TrivialExists.
-- subst. TrivialExists. simpl. rewrite ! Val.addl_assoc. rewrite (Val.addl_commut y). auto.
-- subst. TrivialExists. simpl. rewrite ! Val.addl_assoc. auto.
-- TrivialExists. simpl. destruct x; destruct y; simpl; auto.
- rewrite Int64.add_zero; auto.
- destruct Archi.ptr64 eqn:SF; simpl; auto. rewrite SF. rewrite Ptrofs.add_assoc, Ptrofs.add_zero. auto.
- destruct Archi.ptr64 eqn:SF; simpl; auto. rewrite SF. rewrite Ptrofs.add_assoc, Ptrofs.add_zero. auto.
-Qed.
-
-Theorem eval_subl: binary_constructor_sound subl Val.subl.
-Proof.
- unfold subl. destruct Archi.splitlong eqn:SL.
- apply SplitLongproof.eval_subl. apply Archi.splitlong_ptr32; auto.
- red; intros; destruct (subl_match a b); InvEval.
-- rewrite Val.subl_addl_opp. apply eval_addlimm; auto.
-- subst. rewrite Val.subl_addl_l. rewrite Val.subl_addl_r.
- rewrite Val.addl_assoc. simpl. rewrite Int64.add_commut. rewrite <- Int64.sub_add_opp.
- replace (Int64.repr (n1 - n2)) with (Int64.sub (Int64.repr n1) (Int64.repr n2)).
- apply eval_addlimm; EvalOp.
- apply Int64.eqm_samerepr; auto with ints.
-- subst. rewrite Val.subl_addl_l. apply eval_addlimm; EvalOp.
-- subst. rewrite Val.subl_addl_r.
- replace (Int64.repr (-n2)) with (Int64.neg (Int64.repr n2)).
- apply eval_addlimm; EvalOp.
- apply Int64.eqm_samerepr; auto with ints.
-- TrivialExists.
-Qed.
-
-Theorem eval_mullimm_base: forall n, unary_constructor_sound (mullimm_base n) (fun v => Val.mull v (Vlong n)).
-Proof.
- intros; unfold mullimm_base. red; intros.
- generalize (Int64.one_bits'_decomp n); intros D.
- destruct (Int64.one_bits' n) as [ | i [ | j [ | ? ? ]]] eqn:B.
-- TrivialExists.
-- replace (Val.mull x (Vlong n)) with (Val.shll x (Vint i)).
- apply eval_shllimm; auto.
- simpl in D. rewrite D, Int64.add_zero. destruct x; simpl; auto.
- rewrite (Int64.one_bits'_range n) by (rewrite B; auto with coqlib).
- rewrite Int64.shl'_mul; auto.
-- set (le' := x :: le).
- assert (A0: eval_expr ge sp e m le' (Eletvar O) x) by (constructor; reflexivity).
- exploit (eval_shllimm i). eexact A0. intros (v1 & A1 & B1).
- exploit (eval_shllimm j). eexact A0. intros (v2 & A2 & B2).
- exploit (eval_addl). eexact A1. eexact A2. intros (v3 & A3 & B3).
- exists v3; split. econstructor; eauto.
- rewrite D. simpl. rewrite Int64.add_zero. destruct x; auto.
- simpl in *.
- rewrite (Int64.one_bits'_range n) in B1 by (rewrite B; auto with coqlib).
- rewrite (Int64.one_bits'_range n) in B2 by (rewrite B; auto with coqlib).
- inv B1; inv B2. simpl in B3; inv B3.
- rewrite Int64.mul_add_distr_r. rewrite <- ! Int64.shl'_mul. auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_mullimm: forall n, unary_constructor_sound (mullimm n) (fun v => Val.mull v (Vlong n)).
-Proof.
- unfold mullimm. intros; red; intros.
- destruct Archi.splitlong eqn:SL.
- eapply SplitLongproof.eval_mullimm; eauto.
- predSpec Int64.eq Int64.eq_spec n Int64.zero.
- exists (Vlong Int64.zero); split. apply eval_longconst.
- destruct x; simpl; auto. subst n; rewrite Int64.mul_zero; auto.
- predSpec Int64.eq Int64.eq_spec n Int64.one.
- exists x; split; auto.
- destruct x; simpl; auto. subst n; rewrite Int64.mul_one; auto.
- destruct (mullimm_match a); InvEval.
-- econstructor; split. apply eval_longconst. rewrite Int64.mul_commut; auto.
-- exploit (eval_mullimm_base n); eauto. intros (v2 & A2 & B2).
- exploit (eval_addlimm (Int64.mul n (Int64.repr n2))). eexact A2. intros (v3 & A3 & B3).
- exists v3; split; auto. subst x.
- destruct v1; simpl; auto.
- simpl in B2; inv B2. simpl in B3; inv B3. rewrite Int64.mul_add_distr_l.
- rewrite (Int64.mul_commut n). auto.
- destruct Archi.ptr64; auto.
-- apply eval_mullimm_base; auto.
-Qed.
-
-Theorem eval_mull: binary_constructor_sound mull Val.mull.
-Proof.
- unfold mull. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_mull; auto.
- red; intros; destruct (mull_match a b); InvEval.
-- rewrite Val.mull_commut. apply eval_mullimm; auto.
-- apply eval_mullimm; auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_mullhu:
- forall n, unary_constructor_sound (fun a => mullhu a n) (fun v => Val.mullhu v (Vlong n)).
-Proof.
- unfold mullhu; intros. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_mullhu; auto.
- red; intros. TrivialExists. constructor. eauto. constructor. apply eval_longconst. constructor. auto.
-Qed.
-
-Theorem eval_mullhs:
- forall n, unary_constructor_sound (fun a => mullhs a n) (fun v => Val.mullhs v (Vlong n)).
-Proof.
- unfold mullhs; intros. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_mullhs; auto.
- red; intros. TrivialExists. constructor. eauto. constructor. apply eval_longconst. constructor. auto.
-Qed.
-
-Theorem eval_shrxlimm:
- forall le a n x z,
- eval_expr ge sp e m le a x ->
- Val.shrxl x (Vint n) = Some z ->
- exists v, eval_expr ge sp e m le (shrxlimm a n) v /\ Val.lessdef z v.
-Proof.
- unfold shrxlimm; intros. destruct Archi.splitlong eqn:SL.
-+ eapply SplitLongproof.eval_shrxlimm; eauto. apply Archi.splitlong_ptr32; auto.
-+ predSpec Int.eq Int.eq_spec n Int.zero.
-- subst n. destruct x; simpl in H0; inv H0. econstructor; split; eauto.
- change (Int.ltu Int.zero (Int.repr 63)) with true. simpl. rewrite Int64.shrx'_zero; auto.
-- TrivialExists.
-Qed.
-
-Theorem eval_divls_base: partial_binary_constructor_sound divls_base Val.divls.
-Proof.
- unfold divls_base; red; intros. destruct Archi.splitlong eqn:SL.
- eapply SplitLongproof.eval_divls_base; eauto.
- TrivialExists.
-Qed.
-
-Theorem eval_modls_base: partial_binary_constructor_sound modls_base Val.modls.
-Proof.
- unfold modls_base; red; intros. destruct Archi.splitlong eqn:SL.
- eapply SplitLongproof.eval_modls_base; eauto.
- TrivialExists.
-Qed.
-
-Theorem eval_divlu_base: partial_binary_constructor_sound divlu_base Val.divlu.
-Proof.
- unfold divlu_base; red; intros. destruct Archi.splitlong eqn:SL.
- eapply SplitLongproof.eval_divlu_base; eauto.
- TrivialExists.
-Qed.
-
-Theorem eval_modlu_base: partial_binary_constructor_sound modlu_base Val.modlu.
-Proof.
- unfold modlu_base; red; intros. destruct Archi.splitlong eqn:SL.
- eapply SplitLongproof.eval_modlu_base; eauto.
- TrivialExists.
-Qed.
-
-Theorem eval_cmplu:
- forall c le a x b y v,
- eval_expr ge sp e m le a x ->
- eval_expr ge sp e m le b y ->
- Val.cmplu (Mem.valid_pointer m) c x y = Some v ->
- eval_expr ge sp e m le (cmplu c a b) v.
-Proof.
- unfold cmplu; intros. destruct Archi.splitlong eqn:SL.
- eapply SplitLongproof.eval_cmplu; eauto. apply Archi.splitlong_ptr32; auto.
- unfold Val.cmplu in H1.
- destruct (Val.cmplu_bool (Mem.valid_pointer m) c x y) as [vb|] eqn:C; simpl in H1; inv H1.
- destruct (is_longconst a) as [n1|] eqn:LC1; destruct (is_longconst b) as [n2|] eqn:LC2;
- try (assert (x = Vlong n1) by (eapply is_longconst_sound; eauto));
- try (assert (y = Vlong n2) by (eapply is_longconst_sound; eauto));
- subst.
-- simpl in C; inv C. EvalOp. destruct (Int64.cmpu c n1 n2); reflexivity.
-- EvalOp. simpl. rewrite Val.swap_cmplu_bool. rewrite C; auto.
-- EvalOp. simpl; rewrite C; auto.
-- EvalOp. simpl; rewrite C; auto.
-Qed.
-
-Theorem eval_cmpl:
- forall c le a x b y v,
- eval_expr ge sp e m le a x ->
- eval_expr ge sp e m le b y ->
- Val.cmpl c x y = Some v ->
- eval_expr ge sp e m le (cmpl c a b) v.
-Proof.
- unfold cmpl; intros. destruct Archi.splitlong eqn:SL.
- eapply SplitLongproof.eval_cmpl; eauto.
- unfold Val.cmpl in H1.
- destruct (Val.cmpl_bool c x y) as [vb|] eqn:C; simpl in H1; inv H1.
- destruct (is_longconst a) as [n1|] eqn:LC1; destruct (is_longconst b) as [n2|] eqn:LC2;
- try (assert (x = Vlong n1) by (eapply is_longconst_sound; eauto));
- try (assert (y = Vlong n2) by (eapply is_longconst_sound; eauto));
- subst.
-- simpl in C; inv C. EvalOp. destruct (Int64.cmp c n1 n2); reflexivity.
-- EvalOp. simpl. rewrite Val.swap_cmpl_bool. rewrite C; auto.
-- EvalOp. simpl; rewrite C; auto.
-- EvalOp. simpl; rewrite C; auto.
-Qed.
-
-Theorem eval_longoffloat: partial_unary_constructor_sound longoffloat Val.longoffloat.
-Proof.
- unfold longoffloat; red; intros. destruct Archi.splitlong eqn:SL.
- eapply SplitLongproof.eval_longoffloat; eauto.
- TrivialExists.
-Qed.
-
-Theorem eval_floatoflong: partial_unary_constructor_sound floatoflong Val.floatoflong.
-Proof.
- unfold floatoflong; red; intros. destruct Archi.splitlong eqn:SL.
- eapply SplitLongproof.eval_floatoflong; eauto.
- TrivialExists.
-Qed.
-
-Theorem eval_longofsingle: partial_unary_constructor_sound longofsingle Val.longofsingle.
-Proof.
- unfold longofsingle; red; intros. destruct Archi.splitlong eqn:SL.
- eapply SplitLongproof.eval_longofsingle; eauto.
- TrivialExists.
-Qed.
-
-Theorem eval_singleoflong: partial_unary_constructor_sound singleoflong Val.singleoflong.
-Proof.
- unfold singleoflong; red; intros. destruct Archi.splitlong eqn:SL.
- eapply SplitLongproof.eval_singleoflong; eauto.
- TrivialExists.
-Qed.
-
-End CMCONSTR.
diff --git a/ia32/SelectOp.vp b/ia32/SelectOp.vp
deleted file mode 100644
index db546d99..00000000
--- a/ia32/SelectOp.vp
+++ /dev/null
@@ -1,530 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Instruction selection for operators *)
-
-(** The instruction selection pass recognizes opportunities for using
- combined arithmetic and logical operations and addressing modes
- offered by the target processor. For instance, the expression [x + 1]
- can take advantage of the "immediate add" instruction of the processor,
- and on the PowerPC, the expression [(x >> 6) & 0xFF] can be turned
- into a "rotate and mask" instruction.
-
- This file defines functions for building CminorSel expressions and
- statements, especially expressions consisting of operator
- applications. These functions examine their arguments to choose
- cheaper forms of operators whenever possible.
-
- For instance, [add e1 e2] will return a CminorSel expression semantically
- equivalent to [Eop Oadd (e1 ::: e2 ::: Enil)], but will use a
- [Oaddimm] operator if one of the arguments is an integer constant,
- or suppress the addition altogether if one of the arguments is the
- null integer. In passing, we perform operator reassociation
- ([(e + c1) * c2] becomes [(e * c2) + (c1 * c2)]) and a small amount
- of constant propagation.
-
- On top of the "smart constructor" functions defined below,
- module [Selection] implements the actual instruction selection pass.
-*)
-
-Require Import Coqlib.
-Require Import Compopts.
-Require Import AST Integers Floats.
-Require Import Op CminorSel.
-
-Open Local Scope cminorsel_scope.
-
-(** ** Constants **)
-
-(** External oracle to determine whether a symbol should be addressed
- through [Oindirectsymbol] or can be addressed via [Oleal Aglobal].
- This is to accommodate MacOS X's limitations on references to data
- symbols imported from shared libraries. It can also help with PIC
- code under ELF. *)
-
-Parameter symbol_is_external: ident -> bool.
-
-Definition Olea_ptr (a: addressing) := if Archi.ptr64 then Oleal a else Olea a.
-
-Definition addrsymbol (id: ident) (ofs: ptrofs) :=
- if symbol_is_external id then
- if Ptrofs.eq ofs Ptrofs.zero
- then Eop (Oindirectsymbol id) Enil
- else Eop (Olea_ptr (Aindexed (Ptrofs.unsigned ofs))) (Eop (Oindirectsymbol id) Enil ::: Enil)
- else
- Eop (Olea_ptr (Aglobal id ofs)) Enil.
-
-Definition addrstack (ofs: ptrofs) :=
- Eop (Olea_ptr (Ainstack ofs)) Enil.
-
-(** ** Integer logical negation *)
-
-Nondetfunction notint (e: expr) :=
- match e with
- | Eop (Ointconst n) Enil => Eop (Ointconst (Int.not n)) Enil
- | Eop (Oxorimm n) (e1 ::: Enil) => Eop (Oxorimm (Int.not n)) (e1 ::: Enil)
- | _ => Eop Onot (e ::: Enil)
- end.
-
-(** ** Integer addition and pointer addition *)
-
-Nondetfunction addimm (n: int) (e: expr) :=
- if Int.eq n Int.zero then e else
- match e with
- | Eop (Ointconst m) Enil => Eop (Ointconst(Int.add n m)) Enil
- | Eop (Olea addr) args => Eop (Olea (offset_addressing_total addr (Int.signed n))) args
- | _ => Eop (Olea (Aindexed (Int.signed n))) (e ::: Enil)
- end.
-
-Nondetfunction add (e1: expr) (e2: expr) :=
- match e1, e2 with
- | Eop (Ointconst n1) Enil, t2 => addimm n1 t2
- | t1, Eop (Ointconst n2) Enil => addimm n2 t1
- | Eop (Olea (Aindexed n1)) (t1:::Enil), Eop (Olea (Aindexed n2)) (t2:::Enil) =>
- Eop (Olea (Aindexed2 (n1 + n2))) (t1:::t2:::Enil)
- | Eop (Olea (Aindexed n1)) (t1:::Enil), Eop (Olea (Ascaled sc n2)) (t2:::Enil) =>
- Eop (Olea (Aindexed2scaled sc (n1 + n2))) (t1:::t2:::Enil)
- | Eop (Olea (Ascaled sc n1)) (t1:::Enil), Eop (Olea (Aindexed n2)) (t2:::Enil) =>
- Eop (Olea (Aindexed2scaled sc (n1 + n2))) (t2:::t1:::Enil)
- | Eop (Olea (Aindexed n1)) (t1:::Enil), Eop (Olea (Aglobal id ofs)) Enil =>
- Eop (Olea (Abased id (Ptrofs.add ofs (Ptrofs.repr n1)))) (t1:::Enil)
- | Eop (Olea (Aglobal id ofs)) Enil, Eop (Olea (Aindexed n2)) (t2:::Enil) =>
- Eop (Olea (Abased id (Ptrofs.add ofs (Ptrofs.repr n2)))) (t2:::Enil)
- | Eop (Olea (Ascaled sc n1)) (t1:::Enil), Eop (Olea (Aglobal id ofs)) Enil =>
- Eop (Olea (Abasedscaled sc id (Ptrofs.add ofs (Ptrofs.repr n1)))) (t1:::Enil)
- | Eop (Olea (Aglobal id ofs)) Enil, Eop (Olea (Ascaled sc n2)) (t2:::Enil) =>
- Eop (Olea (Abasedscaled sc id (Ptrofs.add ofs (Ptrofs.repr n2)))) (t2:::Enil)
- | Eop (Olea (Ascaled sc n)) (t1:::Enil), t2 =>
- Eop (Olea (Aindexed2scaled sc n)) (t2:::t1:::Enil)
- | t1, Eop (Olea (Ascaled sc n)) (t2:::Enil) =>
- Eop (Olea (Aindexed2scaled sc n)) (t1:::t2:::Enil)
- | Eop (Olea (Aindexed n)) (t1:::Enil), t2 =>
- Eop (Olea (Aindexed2 n)) (t1:::t2:::Enil)
- | t1, Eop (Olea (Aindexed n)) (t2:::Enil) =>
- Eop (Olea (Aindexed2 n)) (t1:::t2:::Enil)
- | _, _ =>
- Eop (Olea (Aindexed2 0)) (e1:::e2:::Enil)
- end.
-
-(** ** Opposite *)
-
-Nondetfunction negint (e: expr) :=
- match e with
- | Eop (Ointconst n) Enil => Eop (Ointconst (Int.neg n)) Enil
- | _ => Eop Oneg (e ::: Enil)
- end.
-
-(** ** Integer and pointer subtraction *)
-
-Nondetfunction sub (e1: expr) (e2: expr) :=
- match e1, e2 with
- | t1, Eop (Ointconst n2) Enil => addimm (Int.neg n2) t1
- | Eop (Olea (Aindexed n1)) (t1:::Enil), Eop (Olea (Aindexed n2)) (t2:::Enil) =>
- addimm (Int.repr (n1 - n2)) (Eop Osub (t1:::t2:::Enil))
- | Eop (Olea (Aindexed n1)) (t1:::Enil), t2 =>
- addimm (Int.repr n1) (Eop Osub (t1:::t2:::Enil))
- | t1, Eop (Olea (Aindexed n2)) (t2:::Enil) =>
- addimm (Int.repr (-n2)) (Eop Osub (t1:::t2:::Enil))
- | _, _ =>
- Eop Osub (e1:::e2:::Enil)
- end.
-
-(** ** Immediate shifts *)
-
-Definition shift_is_scale (n: int) : bool :=
- Int.eq n (Int.repr 1) || Int.eq n (Int.repr 2) || Int.eq n (Int.repr 3).
-
-Nondetfunction shlimm (e1: expr) (n: int) :=
- if Int.eq n Int.zero then e1 else
- if negb (Int.ltu n Int.iwordsize) then
- Eop Oshl (e1:::Eop (Ointconst n) Enil:::Enil)
- else
- match e1 with
- | Eop (Ointconst n1) Enil =>
- Eop (Ointconst(Int.shl n1 n)) Enil
- | Eop (Oshlimm n1) (t1:::Enil) =>
- if Int.ltu (Int.add n n1) Int.iwordsize
- then Eop (Oshlimm (Int.add n n1)) (t1:::Enil)
- else Eop (Oshlimm n) (e1:::Enil)
- | Eop (Olea (Aindexed n1)) (t1:::Enil) =>
- if shift_is_scale n
- then Eop (Olea (Ascaled (Int.unsigned (Int.shl Int.one n))
- (Int.unsigned (Int.shl (Int.repr n1) n)))) (t1:::Enil)
- else Eop (Oshlimm n) (e1:::Enil)
- | _ =>
- if shift_is_scale n
- then Eop (Olea (Ascaled (Int.unsigned (Int.shl Int.one n)) 0)) (e1:::Enil)
- else Eop (Oshlimm n) (e1:::Enil)
- end.
-
-Nondetfunction shruimm (e1: expr) (n: int) :=
- if Int.eq n Int.zero then e1 else
- if negb (Int.ltu n Int.iwordsize) then
- Eop Oshru (e1:::Eop (Ointconst n) Enil:::Enil)
- else
- match e1 with
- | Eop (Ointconst n1) Enil =>
- Eop (Ointconst(Int.shru n1 n)) Enil
- | Eop (Oshruimm n1) (t1:::Enil) =>
- if Int.ltu (Int.add n n1) Int.iwordsize
- then Eop (Oshruimm (Int.add n n1)) (t1:::Enil)
- else Eop (Oshruimm n) (e1:::Enil)
- | _ =>
- Eop (Oshruimm n) (e1:::Enil)
- end.
-
-Nondetfunction shrimm (e1: expr) (n: int) :=
- if Int.eq n Int.zero then e1 else
- if negb (Int.ltu n Int.iwordsize) then
- Eop Oshr (e1:::Eop (Ointconst n) Enil:::Enil)
- else
- match e1 with
- | Eop (Ointconst n1) Enil =>
- Eop (Ointconst(Int.shr n1 n)) Enil
- | Eop (Oshrimm n1) (t1:::Enil) =>
- if Int.ltu (Int.add n n1) Int.iwordsize
- then Eop (Oshrimm (Int.add n n1)) (t1:::Enil)
- else Eop (Oshrimm n) (e1:::Enil)
- | _ =>
- Eop (Oshrimm n) (e1:::Enil)
- end.
-
-(** ** Integer multiply *)
-
-Definition mulimm_base (n1: int) (e2: expr) :=
- match Int.one_bits n1 with
- | i :: nil =>
- shlimm e2 i
- | i :: j :: nil =>
- Elet e2 (add (shlimm (Eletvar 0) i) (shlimm (Eletvar 0) j))
- | _ =>
- Eop (Omulimm n1) (e2:::Enil)
- end.
-
-Nondetfunction mulimm (n1: int) (e2: expr) :=
- if Int.eq n1 Int.zero then Eop (Ointconst Int.zero) Enil
- else if Int.eq n1 Int.one then e2
- else match e2 with
- | Eop (Ointconst n2) Enil => Eop (Ointconst(Int.mul n1 n2)) Enil
- | Eop (Olea (Aindexed n2)) (t2:::Enil) => addimm (Int.mul n1 (Int.repr n2)) (mulimm_base n1 t2)
- | _ => mulimm_base n1 e2
- end.
-
-Nondetfunction mul (e1: expr) (e2: expr) :=
- match e1, e2 with
- | Eop (Ointconst n1) Enil, t2 => mulimm n1 t2
- | t1, Eop (Ointconst n2) Enil => mulimm n2 t1
- | _, _ => Eop Omul (e1:::e2:::Enil)
- end.
-
-(** ** Bitwise and, or, xor *)
-
-Nondetfunction andimm (n1: int) (e2: expr) :=
- if Int.eq n1 Int.zero then Eop (Ointconst Int.zero) Enil
- else if Int.eq n1 Int.mone then e2
- else match e2 with
- | Eop (Ointconst n2) Enil =>
- Eop (Ointconst (Int.and n1 n2)) Enil
- | Eop (Oandimm n2) (t2:::Enil) =>
- Eop (Oandimm (Int.and n1 n2)) (t2:::Enil)
- | Eop Ocast8unsigned (t2:::Enil) =>
- Eop (Oandimm (Int.and n1 (Int.repr 255))) (t2:::Enil)
- | Eop Ocast16unsigned (t2:::Enil) =>
- Eop (Oandimm (Int.and n1 (Int.repr 65535))) (t2:::Enil)
- | _ =>
- Eop (Oandimm n1) (e2:::Enil)
- end.
-
-Nondetfunction and (e1: expr) (e2: expr) :=
- match e1, e2 with
- | Eop (Ointconst n1) Enil, t2 => andimm n1 t2
- | t1, Eop (Ointconst n2) Enil => andimm n2 t1
- | _, _ => Eop Oand (e1:::e2:::Enil)
- end.
-
-Nondetfunction orimm (n1: int) (e2: expr) :=
- if Int.eq n1 Int.zero then e2
- else if Int.eq n1 Int.mone then Eop (Ointconst Int.mone) Enil
- else match e2 with
- | Eop (Ointconst n2) Enil =>
- Eop (Ointconst (Int.or n1 n2)) Enil
- | Eop (Oorimm n2) (t2:::Enil) =>
- Eop (Oorimm (Int.or n1 n2)) (t2:::Enil)
- | _ =>
- Eop (Oorimm n1) (e2:::Enil)
- end.
-
-Definition same_expr_pure (e1 e2: expr) :=
- match e1, e2 with
- | Evar v1, Evar v2 => if ident_eq v1 v2 then true else false
- | _, _ => false
- end.
-
-Nondetfunction or (e1: expr) (e2: expr) :=
- match e1, e2 with
- | Eop (Ointconst n1) Enil, t2 => orimm n1 t2
- | t1, Eop (Ointconst n2) Enil => orimm n2 t1
- | Eop (Oshlimm n1) (t1:::Enil), Eop (Oshruimm n2) (t2:::Enil) =>
- if Int.eq (Int.add n1 n2) Int.iwordsize then
- if same_expr_pure t1 t2
- then Eop (Ororimm n2) (t1:::Enil)
- else Eop (Oshldimm n1) (t1:::t2:::Enil)
- else Eop Oor (e1:::e2:::Enil)
- | Eop (Oshruimm n2) (t2:::Enil), Eop (Oshlimm n1) (t1:::Enil) =>
- if Int.eq (Int.add n1 n2) Int.iwordsize then
- if same_expr_pure t1 t2
- then Eop (Ororimm n2) (t1:::Enil)
- else Eop (Oshldimm n1) (t1:::t2:::Enil)
- else Eop Oor (e1:::e2:::Enil)
- | _, _ =>
- Eop Oor (e1:::e2:::Enil)
- end.
-
-Nondetfunction xorimm (n1: int) (e2: expr) :=
- if Int.eq n1 Int.zero then e2
- else match e2 with
- | Eop (Ointconst n2) Enil =>
- Eop (Ointconst (Int.xor n1 n2)) Enil
- | Eop (Oxorimm n2) (t2:::Enil) =>
- Eop (Oxorimm (Int.xor n1 n2)) (t2:::Enil)
- | Eop Onot (t2:::Enil) =>
- Eop (Oxorimm (Int.not n1)) (t2:::Enil)
- | _ =>
- Eop (Oxorimm n1) (e2:::Enil)
- end.
-
-Nondetfunction xor (e1: expr) (e2: expr) :=
- match e1, e2 with
- | Eop (Ointconst n1) Enil, t2 => xorimm n1 t2
- | t1, Eop (Ointconst n2) Enil => xorimm n2 t1
- | _, _ => Eop Oxor (e1:::e2:::Enil)
- end.
-
-(** ** Integer division and modulus *)
-
-Definition divu_base (e1: expr) (e2: expr) := Eop Odivu (e1:::e2:::Enil).
-Definition modu_base (e1: expr) (e2: expr) := Eop Omodu (e1:::e2:::Enil).
-Definition divs_base (e1: expr) (e2: expr) := Eop Odiv (e1:::e2:::Enil).
-Definition mods_base (e1: expr) (e2: expr) := Eop Omod (e1:::e2:::Enil).
-
-Definition shrximm (e1: expr) (n2: int) :=
- if Int.eq n2 Int.zero then e1 else Eop (Oshrximm n2) (e1:::Enil).
-
-(** ** General shifts *)
-
-Nondetfunction shl (e1: expr) (e2: expr) :=
- match e2 with
- | Eop (Ointconst n2) Enil => shlimm e1 n2
- | _ => Eop Oshl (e1:::e2:::Enil)
- end.
-
-Nondetfunction shr (e1: expr) (e2: expr) :=
- match e2 with
- | Eop (Ointconst n2) Enil => shrimm e1 n2
- | _ => Eop Oshr (e1:::e2:::Enil)
- end.
-
-Nondetfunction shru (e1: expr) (e2: expr) :=
- match e2 with
- | Eop (Ointconst n2) Enil => shruimm e1 n2
- | _ => Eop Oshru (e1:::e2:::Enil)
- end.
-
-(** ** Floating-point arithmetic *)
-
-Definition negf (e: expr) := Eop Onegf (e ::: Enil).
-Definition absf (e: expr) := Eop Oabsf (e ::: Enil).
-Definition addf (e1 e2: expr) := Eop Oaddf (e1 ::: e2 ::: Enil).
-Definition subf (e1 e2: expr) := Eop Osubf (e1 ::: e2 ::: Enil).
-Definition mulf (e1 e2: expr) := Eop Omulf (e1 ::: e2 ::: Enil).
-
-Definition negfs (e: expr) := Eop Onegfs (e ::: Enil).
-Definition absfs (e: expr) := Eop Oabsfs (e ::: Enil).
-Definition addfs (e1 e2: expr) := Eop Oaddfs (e1 ::: e2 ::: Enil).
-Definition subfs (e1 e2: expr) := Eop Osubfs (e1 ::: e2 ::: Enil).
-Definition mulfs (e1 e2: expr) := Eop Omulfs (e1 ::: e2 ::: Enil).
-
-(** ** Comparisons *)
-
-Nondetfunction compimm (default: comparison -> int -> condition)
- (sem: comparison -> int -> int -> bool)
- (c: comparison) (e1: expr) (n2: int) :=
- match c, e1 with
- | c, Eop (Ointconst n1) Enil =>
- Eop (Ointconst (if sem c n1 n2 then Int.one else Int.zero)) Enil
- | Ceq, Eop (Ocmp c) el =>
- if Int.eq_dec n2 Int.zero then
- Eop (Ocmp (negate_condition c)) el
- else if Int.eq_dec n2 Int.one then
- Eop (Ocmp c) el
- else
- Eop (Ointconst Int.zero) Enil
- | Cne, Eop (Ocmp c) el =>
- if Int.eq_dec n2 Int.zero then
- Eop (Ocmp c) el
- else if Int.eq_dec n2 Int.one then
- Eop (Ocmp (negate_condition c)) el
- else
- Eop (Ointconst Int.one) Enil
- | Ceq, Eop (Oandimm n1) (t1 ::: Enil) =>
- if Int.eq_dec n2 Int.zero then
- Eop (Ocmp (Cmaskzero n1)) (t1 ::: Enil)
- else
- Eop (Ocmp (default c n2)) (e1 ::: Enil)
- | Cne, Eop (Oandimm n1) (t1 ::: Enil) =>
- if Int.eq_dec n2 Int.zero then
- Eop (Ocmp (Cmasknotzero n1)) (t1 ::: Enil)
- else
- Eop (Ocmp (default c n2)) (e1 ::: Enil)
- | _, _ =>
- Eop (Ocmp (default c n2)) (e1 ::: Enil)
- end.
-
-Nondetfunction comp (c: comparison) (e1: expr) (e2: expr) :=
- match e1, e2 with
- | Eop (Ointconst n1) Enil, t2 =>
- compimm Ccompimm Int.cmp (swap_comparison c) t2 n1
- | t1, Eop (Ointconst n2) Enil =>
- compimm Ccompimm Int.cmp c t1 n2
- | _, _ =>
- Eop (Ocmp (Ccomp c)) (e1 ::: e2 ::: Enil)
- end.
-
-Nondetfunction compu (c: comparison) (e1: expr) (e2: expr) :=
- match e1, e2 with
- | Eop (Ointconst n1) Enil, t2 =>
- compimm Ccompuimm Int.cmpu (swap_comparison c) t2 n1
- | t1, Eop (Ointconst n2) Enil =>
- compimm Ccompuimm Int.cmpu c t1 n2
- | _, _ =>
- Eop (Ocmp (Ccompu c)) (e1 ::: e2 ::: Enil)
- end.
-
-Definition compf (c: comparison) (e1: expr) (e2: expr) :=
- Eop (Ocmp (Ccompf c)) (e1 ::: e2 ::: Enil).
-
-Definition compfs (c: comparison) (e1: expr) (e2: expr) :=
- Eop (Ocmp (Ccompfs c)) (e1 ::: e2 ::: Enil).
-
-(** ** Integer conversions *)
-
-Nondetfunction cast8unsigned (e: expr) :=
- match e with
- | Eop (Ointconst n) Enil =>
- Eop (Ointconst (Int.zero_ext 8 n)) Enil
- | Eop (Oandimm n) (t:::Enil) =>
- andimm (Int.and (Int.repr 255) n) t
- | _ =>
- Eop Ocast8unsigned (e:::Enil)
- end.
-
-Nondetfunction cast8signed (e: expr) :=
- match e with
- | Eop (Ointconst n) Enil =>
- Eop (Ointconst (Int.sign_ext 8 n)) Enil
- | _ =>
- Eop Ocast8signed (e ::: Enil)
- end.
-
-Nondetfunction cast16unsigned (e: expr) :=
- match e with
- | Eop (Ointconst n) Enil =>
- Eop (Ointconst (Int.zero_ext 16 n)) Enil
- | Eop (Oandimm n) (t:::Enil) =>
- andimm (Int.and (Int.repr 65535) n) t
- | _ =>
- Eop Ocast16unsigned (e:::Enil)
- end.
-
-Nondetfunction cast16signed (e: expr) :=
- match e with
- | Eop (Ointconst n) Enil =>
- Eop (Ointconst (Int.sign_ext 16 n)) Enil
- | _ =>
- Eop Ocast16signed (e ::: Enil)
- end.
-
-(** Floating-point conversions *)
-
-Definition singleoffloat (e: expr) := Eop Osingleoffloat (e ::: Enil).
-Definition floatofsingle (e: expr) := Eop Ofloatofsingle (e ::: Enil).
-
-Definition intoffloat (e: expr) := Eop Ointoffloat (e ::: Enil).
-
-Nondetfunction floatofint (e: expr) :=
- match e with
- | Eop (Ointconst n) Enil => Eop (Ofloatconst (Float.of_int n)) Enil
- | _ => Eop Ofloatofint (e ::: Enil)
- end.
-
-Definition intuoffloat (e: expr) :=
- Elet e
- (Elet (Eop (Ofloatconst (Float.of_intu Float.ox8000_0000)) Enil)
- (Econdition (CEcond (Ccompf Clt) (Eletvar 1 ::: Eletvar 0 ::: Enil))
- (intoffloat (Eletvar 1))
- (addimm Float.ox8000_0000 (intoffloat (subf (Eletvar 1) (Eletvar 0))))))%nat.
-
-Nondetfunction floatofintu (e: expr) :=
- match e with
- | Eop (Ointconst n) Enil => Eop (Ofloatconst (Float.of_intu n)) Enil
- | _ =>
- let f := Eop (Ofloatconst (Float.of_intu Float.ox8000_0000)) Enil in
- Elet e
- (Econdition (CEcond (Ccompuimm Clt Float.ox8000_0000) (Eletvar O ::: Enil))
- (floatofint (Eletvar O))
- (addf (floatofint (addimm (Int.neg Float.ox8000_0000) (Eletvar O))) f))
- end.
-
-Definition intofsingle (e: expr) := Eop Ointofsingle (e ::: Enil).
-
-Nondetfunction singleofint (e: expr) :=
- match e with
- | Eop (Ointconst n) Enil => Eop (Osingleconst (Float32.of_int n)) Enil
- | _ => Eop Osingleofint (e ::: Enil)
- end.
-
-Definition intuofsingle (e: expr) := intuoffloat (floatofsingle e).
-
-Nondetfunction singleofintu (e: expr) :=
- match e with
- | Eop (Ointconst n) Enil => Eop (Osingleconst (Float32.of_intu n)) Enil
- | _ => singleoffloat (floatofintu e)
- end.
-
-(** ** Addressing modes *)
-
-Nondetfunction addressing (chunk: memory_chunk) (e: expr) :=
- match e with
- | Eop (Olea addr) args =>
- if (negb Archi.ptr64) && addressing_valid addr then (addr, args) else (Aindexed 0, e:::Enil)
- | Eop (Oleal addr) args =>
- if Archi.ptr64 && addressing_valid addr then (addr, args) else (Aindexed 0, e:::Enil)
- | _ => (Aindexed 0, e:::Enil)
- end.
-
-(** ** Arguments of builtins *)
-
-Nondetfunction builtin_arg (e: expr) :=
- match e with
- | Eop (Ointconst n) Enil => BA_int n
- | Eop (Olongconst n) Enil => BA_long n
- | Eop (Olea (Aglobal id ofs)) Enil => if Archi.ptr64 then BA e else BA_addrglobal id ofs
- | Eop (Olea (Ainstack ofs)) Enil => if Archi.ptr64 then BA e else BA_addrstack ofs
- | Eop (Oleal (Aglobal id ofs)) Enil => if Archi.ptr64 then BA_addrglobal id ofs else BA e
- | Eop (Oleal (Ainstack ofs)) Enil => if Archi.ptr64 then BA_addrstack ofs else BA e
- | Eop Omakelong (Eop (Ointconst h) Enil ::: Eop (Ointconst l) Enil ::: Enil) =>
- BA_long (Int64.ofwords h l)
- | Eop Omakelong (h ::: l ::: Enil) => BA_splitlong (BA h) (BA l)
- | Eload chunk (Aglobal id ofs) Enil => BA_loadglobal chunk id ofs
- | Eload chunk (Ainstack ofs) Enil => BA_loadstack chunk ofs
- | _ => BA e
- end.
diff --git a/ia32/SelectOpproof.v b/ia32/SelectOpproof.v
deleted file mode 100644
index e201d207..00000000
--- a/ia32/SelectOpproof.v
+++ /dev/null
@@ -1,958 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* 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 of instruction selection for operators *)
-
-Require Import Coqlib.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Cminor.
-Require Import Op.
-Require Import CminorSel.
-Require Import SelectOp.
-
-Open Local Scope cminorsel_scope.
-Local Transparent Archi.ptr64.
-
-(** * Useful lemmas and tactics *)
-
-(** The following are trivial lemmas and custom tactics that help
- perform backward (inversion) and forward reasoning over the evaluation
- of operator applications. *)
-
-Ltac EvalOp := eapply eval_Eop; eauto with evalexpr.
-
-Ltac InvEval1 :=
- match goal with
- | [ H: (eval_expr _ _ _ _ _ (Eop _ Enil) _) |- _ ] =>
- inv H; InvEval1
- | [ H: (eval_expr _ _ _ _ _ (Eop _ (_ ::: Enil)) _) |- _ ] =>
- inv H; InvEval1
- | [ H: (eval_expr _ _ _ _ _ (Eop _ (_ ::: _ ::: Enil)) _) |- _ ] =>
- inv H; InvEval1
- | [ H: (eval_exprlist _ _ _ _ _ Enil _) |- _ ] =>
- inv H; InvEval1
- | [ H: (eval_exprlist _ _ _ _ _ (_ ::: _) _) |- _ ] =>
- inv H; InvEval1
- | _ =>
- idtac
- end.
-
-Ltac InvEval2 :=
- match goal with
- | [ H: (eval_operation _ _ _ nil _ = Some _) |- _ ] =>
- simpl in H; inv H
- | [ H: (eval_operation _ _ _ (_ :: nil) _ = Some _) |- _ ] =>
- simpl in H; FuncInv
- | [ H: (eval_operation _ _ _ (_ :: _ :: nil) _ = Some _) |- _ ] =>
- simpl in H; FuncInv
- | [ H: (eval_operation _ _ _ (_ :: _ :: _ :: nil) _ = Some _) |- _ ] =>
- simpl in H; FuncInv
- | _ =>
- idtac
- end.
-
-Ltac InvEval := InvEval1; InvEval2; InvEval2.
-
-Ltac TrivialExists :=
- match goal with
- | [ |- exists v, _ /\ Val.lessdef ?a v ] => exists a; split; [EvalOp | auto]
- end.
-
-(** * Correctness of the smart constructors *)
-
-Section CMCONSTR.
-
-Variable ge: genv.
-Variable sp: val.
-Variable e: env.
-Variable m: mem.
-
-(** We now show that the code generated by "smart constructor" functions
- such as [SelectOp.notint] behaves as expected. Continuing the
- [notint] example, we show that if the expression [e]
- evaluates to some integer value [Vint n], then [SelectOp.notint e]
- evaluates to a value [Vint (Int.not n)] which is indeed the integer
- negation of the value of [e].
-
- All proofs follow a common pattern:
-- Reasoning by case over the result of the classification functions
- (such as [add_match] for integer addition), gathering additional
- information on the shape of the argument expressions in the non-default
- cases.
-- Inversion of the evaluations of the arguments, exploiting the additional
- information thus gathered.
-- Equational reasoning over the arithmetic operations performed,
- using the lemmas from the [Int] and [Float] modules.
-- Construction of an evaluation derivation for the expression returned
- by the smart constructor.
-*)
-
-Definition unary_constructor_sound (cstr: expr -> expr) (sem: val -> val) : Prop :=
- forall le a x,
- eval_expr ge sp e m le a x ->
- exists v, eval_expr ge sp e m le (cstr a) v /\ Val.lessdef (sem x) v.
-
-Definition binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> val -> val) : Prop :=
- forall le a x b y,
- eval_expr ge sp e m le a x ->
- eval_expr ge sp e m le b y ->
- exists v, eval_expr ge sp e m le (cstr a b) v /\ Val.lessdef (sem x y) v.
-
-Lemma eval_Olea_ptr:
- forall a el m,
- eval_operation ge sp (Olea_ptr a) el m = eval_addressing ge sp a el.
-Proof.
- unfold Olea_ptr, eval_addressing; intros. destruct Archi.ptr64; auto.
-Qed.
-
-Theorem eval_addrsymbol:
- forall le id ofs,
- exists v, eval_expr ge sp e m le (addrsymbol id ofs) v /\ Val.lessdef (Genv.symbol_address ge id ofs) v.
-Proof.
- intros. unfold addrsymbol. exists (Genv.symbol_address ge id ofs); split; auto.
- destruct (symbol_is_external id).
- predSpec Ptrofs.eq Ptrofs.eq_spec ofs Ptrofs.zero.
- subst. EvalOp.
- EvalOp. econstructor. EvalOp. simpl; eauto. econstructor.
- unfold Olea_ptr; destruct Archi.ptr64 eqn:SF; simpl.
- unfold Genv.symbol_address; destruct (Genv.find_symbol ge id); simpl; auto.
- rewrite SF. rewrite Ptrofs.add_zero_l. fold (Ptrofs.to_int64 ofs). rewrite Ptrofs.of_int64_to_int64 by auto. auto.
- unfold Genv.symbol_address; destruct (Genv.find_symbol ge id); simpl; auto.
- rewrite SF. rewrite Ptrofs.add_zero_l. fold (Ptrofs.to_int ofs). rewrite Ptrofs.of_int_to_int by auto. auto.
- EvalOp. rewrite eval_Olea_ptr. apply eval_addressing_Aglobal.
-Qed.
-
-Theorem eval_addrstack:
- forall le ofs,
- exists v, eval_expr ge sp e m le (addrstack ofs) v /\ Val.lessdef (Val.offset_ptr sp ofs) v.
-Proof.
- intros. unfold addrstack. TrivialExists. rewrite eval_Olea_ptr. apply eval_addressing_Ainstack.
-Qed.
-
-Theorem eval_notint: unary_constructor_sound notint Val.notint.
-Proof.
- unfold notint; red; intros until x. case (notint_match a); intros.
- InvEval. TrivialExists.
- InvEval. subst x. rewrite Val.not_xor. rewrite Val.xor_assoc. TrivialExists.
- TrivialExists.
-Qed.
-
-Theorem eval_addimm:
- forall n, unary_constructor_sound (addimm n) (fun x => Val.add x (Vint n)).
-Proof.
- red; unfold addimm; intros until x.
- predSpec Int.eq Int.eq_spec n Int.zero.
- subst n. intros. exists x; split; auto.
- destruct x; simpl; auto. rewrite Int.add_zero; auto. destruct Archi.ptr64; auto. rewrite Ptrofs.add_zero; auto.
- case (addimm_match a); intros; InvEval; simpl.
- TrivialExists; simpl. rewrite Int.add_commut. auto.
- inv H0. simpl in H6. TrivialExists. simpl.
- erewrite eval_offset_addressing_total_32 by eauto. rewrite Int.repr_signed; auto.
- TrivialExists. simpl. rewrite Int.repr_signed; auto.
-Qed.
-
-Theorem eval_add: binary_constructor_sound add Val.add.
-Proof.
- assert (A: forall x y, Int.repr (x + y) = Int.add (Int.repr x) (Int.repr y)).
- { intros; apply Int.eqm_samerepr; auto with ints. }
- assert (B: forall id ofs n, Archi.ptr64 = false ->
- Genv.symbol_address ge id (Ptrofs.add ofs (Ptrofs.repr n)) =
- Val.add (Genv.symbol_address ge id ofs) (Vint (Int.repr n))).
- { intros. replace (Ptrofs.repr n) with (Ptrofs.of_int (Int.repr n)) by auto with ptrofs.
- apply Genv.shift_symbol_address_32; auto. }
- red; intros until y.
- unfold add; case (add_match a b); intros; InvEval.
- rewrite Val.add_commut. apply eval_addimm; auto.
- apply eval_addimm; auto.
-- subst. TrivialExists. simpl. rewrite A, Val.add_permut_4. auto.
-- subst. TrivialExists. simpl. rewrite A, Val.add_assoc. decEq; decEq. rewrite Val.add_permut. auto.
-- subst. TrivialExists. simpl. rewrite A, Val.add_permut_4. rewrite <- Val.add_permut. rewrite <- Val.add_assoc. auto.
-- subst. TrivialExists. simpl. rewrite Heqb0. rewrite B by auto. rewrite ! Val.add_assoc.
- rewrite (Val.add_commut v1). rewrite Val.add_permut. rewrite Val.add_assoc. auto.
-- subst. TrivialExists. simpl. rewrite Heqb0. rewrite B by auto. rewrite Val.add_assoc. do 2 f_equal. apply Val.add_commut.
-- subst. TrivialExists. simpl. rewrite Heqb0. rewrite B by auto. rewrite !Val.add_assoc.
- rewrite (Val.add_commut (Vint (Int.repr n1))). rewrite Val.add_permut. do 2 f_equal. apply Val.add_commut.
-- subst. TrivialExists. simpl. rewrite Heqb0. rewrite B by auto. rewrite !Val.add_assoc.
- rewrite (Val.add_commut (Vint (Int.repr n2))). rewrite Val.add_permut. auto.
-- subst. TrivialExists. simpl. rewrite Val.add_permut. rewrite Val.add_assoc.
- decEq; decEq. apply Val.add_commut.
-- subst. TrivialExists.
-- subst. TrivialExists. simpl. repeat rewrite Val.add_assoc. decEq; decEq. apply Val.add_commut.
-- subst. TrivialExists. simpl. rewrite Val.add_assoc; auto.
-- TrivialExists. simpl. destruct x; destruct y; simpl; auto.
- rewrite Int.add_zero; auto.
- destruct Archi.ptr64 eqn:SF; simpl; auto. rewrite SF. rewrite Ptrofs.add_assoc, Ptrofs.add_zero. auto.
- destruct Archi.ptr64 eqn:SF; simpl; auto. rewrite SF. rewrite Ptrofs.add_assoc, Ptrofs.add_zero. auto.
-Qed.
-
-Theorem eval_sub: binary_constructor_sound sub Val.sub.
-Proof.
- red; intros until y.
- unfold sub; case (sub_match a b); intros; InvEval.
- rewrite Val.sub_add_opp. apply eval_addimm; auto.
- subst. rewrite Val.sub_add_l. rewrite Val.sub_add_r.
- rewrite Val.add_assoc. simpl. rewrite Int.add_commut. rewrite <- Int.sub_add_opp.
- replace (Int.repr (n1 - n2)) with (Int.sub (Int.repr n1) (Int.repr n2)).
- apply eval_addimm; EvalOp.
- apply Int.eqm_samerepr; auto with ints.
- subst. rewrite Val.sub_add_l. apply eval_addimm; EvalOp.
- subst. rewrite Val.sub_add_r. replace (Int.repr (-n2)) with (Int.neg (Int.repr n2)). apply eval_addimm; EvalOp.
- apply Int.eqm_samerepr; auto with ints.
- TrivialExists.
-Qed.
-
-Theorem eval_negint: unary_constructor_sound negint Val.neg.
-Proof.
- red; intros until x. unfold negint. case (negint_match a); intros; InvEval.
- TrivialExists.
- TrivialExists.
-Qed.
-
-Theorem eval_shlimm:
- forall n, unary_constructor_sound (fun a => shlimm a n)
- (fun x => Val.shl x (Vint n)).
-Proof.
- red; intros until x. unfold shlimm.
- predSpec Int.eq Int.eq_spec n Int.zero.
- intros; subst. exists x; split; auto. destruct x; simpl; auto. rewrite Int.shl_zero; auto.
- destruct (Int.ltu n Int.iwordsize) eqn:LT; simpl.
- destruct (shlimm_match a); intros; InvEval.
- exists (Vint (Int.shl n1 n)); split. EvalOp.
- simpl. rewrite LT. auto.
- destruct (Int.ltu (Int.add n n1) Int.iwordsize) eqn:?.
- exists (Val.shl v1 (Vint (Int.add n n1))); split. EvalOp.
- subst. destruct v1; simpl; auto.
- rewrite Heqb.
- destruct (Int.ltu n1 Int.iwordsize) eqn:?; simpl; auto.
- destruct (Int.ltu n Int.iwordsize) eqn:?; simpl; auto.
- rewrite Int.add_commut. rewrite Int.shl_shl; auto. rewrite Int.add_commut; auto.
- subst. TrivialExists. econstructor. EvalOp. simpl; eauto. constructor.
- simpl. auto.
- subst. destruct (shift_is_scale n).
- econstructor; split. EvalOp. simpl. eauto.
- rewrite ! Int.repr_unsigned.
- destruct v1; simpl; auto. rewrite LT.
- rewrite Int.shl_mul. rewrite Int.mul_add_distr_l. rewrite (Int.shl_mul (Int.repr n1)). auto.
- destruct Archi.ptr64; simpl; auto.
- TrivialExists. econstructor. EvalOp. simpl; eauto. constructor. auto.
- destruct (shift_is_scale n).
- econstructor; split. EvalOp. simpl. eauto.
- destruct x; simpl; auto. rewrite LT.
- rewrite Int.repr_unsigned. rewrite Int.add_zero. rewrite Int.shl_mul. auto.
- TrivialExists.
- intros; TrivialExists. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor.
- auto.
-Qed.
-
-Theorem eval_shruimm:
- forall n, unary_constructor_sound (fun a => shruimm a n)
- (fun x => Val.shru x (Vint n)).
-Proof.
- red; intros until x. unfold shruimm.
- predSpec Int.eq Int.eq_spec n Int.zero.
- intros; subst. exists x; split; auto. destruct x; simpl; auto. rewrite Int.shru_zero; auto.
- destruct (Int.ltu n Int.iwordsize) eqn:LT; simpl.
- destruct (shruimm_match a); intros; InvEval.
- exists (Vint (Int.shru n1 n)); split. EvalOp.
- simpl. rewrite LT; auto.
- destruct (Int.ltu (Int.add n n1) Int.iwordsize) eqn:?.
- exists (Val.shru v1 (Vint (Int.add n n1))); split. EvalOp.
- subst. destruct v1; simpl; auto.
- rewrite Heqb.
- destruct (Int.ltu n1 Int.iwordsize) eqn:?; simpl; auto.
- rewrite LT. rewrite Int.add_commut. rewrite Int.shru_shru; auto. rewrite Int.add_commut; auto.
- subst. TrivialExists. econstructor. EvalOp. simpl; eauto. constructor.
- simpl. auto.
- TrivialExists.
- intros; TrivialExists. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor.
- auto.
-Qed.
-
-Theorem eval_shrimm:
- forall n, unary_constructor_sound (fun a => shrimm a n)
- (fun x => Val.shr x (Vint n)).
-Proof.
- red; intros until x. unfold shrimm.
- predSpec Int.eq Int.eq_spec n Int.zero.
- intros; subst. exists x; split; auto. destruct x; simpl; auto. rewrite Int.shr_zero; auto.
- destruct (Int.ltu n Int.iwordsize) eqn:LT; simpl.
- destruct (shrimm_match a); intros; InvEval.
- exists (Vint (Int.shr n1 n)); split. EvalOp.
- simpl. rewrite LT; auto.
- destruct (Int.ltu (Int.add n n1) Int.iwordsize) eqn:?.
- exists (Val.shr v1 (Vint (Int.add n n1))); split. EvalOp.
- subst. destruct v1; simpl; auto.
- rewrite Heqb.
- destruct (Int.ltu n1 Int.iwordsize) eqn:?; simpl; auto.
- rewrite LT.
- rewrite Int.add_commut. rewrite Int.shr_shr; auto. rewrite Int.add_commut; auto.
- subst. TrivialExists. econstructor. EvalOp. simpl; eauto. constructor.
- simpl. auto.
- TrivialExists.
- intros; TrivialExists. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor.
- auto.
-Qed.
-
-Lemma eval_mulimm_base:
- forall n, unary_constructor_sound (mulimm_base n) (fun x => Val.mul x (Vint n)).
-Proof.
- intros; red; intros; unfold mulimm_base.
- generalize (Int.one_bits_decomp n) (Int.one_bits_range n); intros D R.
- destruct (Int.one_bits n) as [ | i l].
- TrivialExists.
- destruct l as [ | j l ].
- replace (Val.mul x (Vint n)) with (Val.shl x (Vint i)). apply eval_shlimm; auto.
- destruct x; auto; simpl. rewrite D; simpl; rewrite Int.add_zero.
- rewrite R by auto with coqlib. rewrite Int.shl_mul. auto.
- destruct l as [ | k l ].
- exploit (eval_shlimm i (x :: le) (Eletvar 0) x). constructor; auto. intros [v1 [A1 B1]].
- exploit (eval_shlimm j (x :: le) (Eletvar 0) x). constructor; auto. intros [v2 [A2 B2]].
- exploit eval_add. eexact A1. eexact A2. intros [v3 [A3 B3]].
- exists v3; split. econstructor; eauto.
- rewrite D; simpl; rewrite Int.add_zero.
- replace (Vint (Int.add (Int.shl Int.one i) (Int.shl Int.one j)))
- with (Val.add (Val.shl Vone (Vint i)) (Val.shl Vone (Vint j))).
- rewrite Val.mul_add_distr_r.
- repeat rewrite Val.shl_mul.
- apply Val.lessdef_trans with (Val.add v1 v2); auto. apply Val.add_lessdef; auto.
- simpl. rewrite ! R by auto with coqlib. auto.
- TrivialExists.
-Qed.
-
-Theorem eval_mulimm:
- forall n, unary_constructor_sound (mulimm n) (fun x => Val.mul x (Vint n)).
-Proof.
- intros; red; intros until x; unfold mulimm.
- predSpec Int.eq Int.eq_spec n Int.zero.
- intros. exists (Vint Int.zero); split. EvalOp.
- destruct x; simpl; auto. subst n. rewrite Int.mul_zero. auto.
- predSpec Int.eq Int.eq_spec n Int.one.
- intros. exists x; split; auto.
- destruct x; simpl; auto. subst n. rewrite Int.mul_one. auto.
- case (mulimm_match a); intros; InvEval.
- TrivialExists. simpl. rewrite Int.mul_commut; auto.
- subst. rewrite Val.mul_add_distr_l.
- exploit eval_mulimm_base; eauto. instantiate (1 := n). intros [v' [A1 B1]].
- exploit (eval_addimm (Int.mul n (Int.repr n2)) le (mulimm_base n t2) v'). auto. intros [v'' [A2 B2]].
- exists v''; split; auto. eapply Val.lessdef_trans. eapply Val.add_lessdef; eauto.
- rewrite Val.mul_commut; auto.
- apply eval_mulimm_base; auto.
-Qed.
-
-Theorem eval_mul: binary_constructor_sound mul Val.mul.
-Proof.
- red; intros until y.
- unfold mul; case (mul_match a b); intros; InvEval.
- rewrite Val.mul_commut. apply eval_mulimm. auto.
- apply eval_mulimm. auto.
- TrivialExists.
-Qed.
-
-Theorem eval_andimm:
- forall n, unary_constructor_sound (andimm n) (fun x => Val.and x (Vint n)).
-Proof.
- intros; red; intros until x. unfold andimm.
- predSpec Int.eq Int.eq_spec n Int.zero.
- intros. exists (Vint Int.zero); split. EvalOp.
- destruct x; simpl; auto. subst n. rewrite Int.and_zero. auto.
- predSpec Int.eq Int.eq_spec n Int.mone.
- intros. exists x; split; auto.
- destruct x; simpl; auto. subst n. rewrite Int.and_mone. auto.
- case (andimm_match a); intros; InvEval.
- TrivialExists. simpl. rewrite Int.and_commut; auto.
- subst. TrivialExists. simpl. rewrite Val.and_assoc. rewrite Int.and_commut. auto.
- subst. rewrite Val.zero_ext_and. TrivialExists. rewrite Val.and_assoc.
- rewrite Int.and_commut. auto. compute; auto.
- subst. rewrite Val.zero_ext_and. TrivialExists. rewrite Val.and_assoc.
- rewrite Int.and_commut. auto. compute; auto.
- TrivialExists.
-Qed.
-
-Theorem eval_and: binary_constructor_sound and Val.and.
-Proof.
- red; intros until y; unfold and; case (and_match a b); intros; InvEval.
- rewrite Val.and_commut. apply eval_andimm; auto.
- apply eval_andimm; auto.
- TrivialExists.
-Qed.
-
-Theorem eval_orimm:
- forall n, unary_constructor_sound (orimm n) (fun x => Val.or x (Vint n)).
-Proof.
- intros; red; intros until x. unfold orimm.
- predSpec Int.eq Int.eq_spec n Int.zero.
- intros. exists x; split. auto.
- destruct x; simpl; auto. subst n. rewrite Int.or_zero. auto.
- predSpec Int.eq Int.eq_spec n Int.mone.
- intros. exists (Vint Int.mone); split. EvalOp.
- destruct x; simpl; auto. subst n. rewrite Int.or_mone. auto.
- destruct (orimm_match a); intros; InvEval.
- TrivialExists. simpl. rewrite Int.or_commut; auto.
- subst. rewrite Val.or_assoc. simpl. rewrite Int.or_commut. TrivialExists.
- TrivialExists.
-Qed.
-
-Remark eval_same_expr:
- forall a1 a2 le v1 v2,
- same_expr_pure a1 a2 = true ->
- eval_expr ge sp e m le a1 v1 ->
- eval_expr ge sp e m le a2 v2 ->
- a1 = a2 /\ v1 = v2.
-Proof.
- intros until v2.
- destruct a1; simpl; try (intros; discriminate).
- destruct a2; simpl; try (intros; discriminate).
- case (ident_eq i i0); intros.
- subst i0. inversion H0. inversion H1. split. auto. congruence.
- discriminate.
-Qed.
-
-Remark int_add_sub_eq:
- forall x y z, Int.add x y = z -> Int.sub z x = y.
-Proof.
- intros. subst z. rewrite Int.sub_add_l. rewrite Int.sub_idem. apply Int.add_zero_l.
-Qed.
-
-Lemma eval_or: binary_constructor_sound or Val.or.
-Proof.
- red; intros until y; unfold or; case (or_match a b); intros.
-(* intconst *)
- InvEval. rewrite Val.or_commut. apply eval_orimm; auto.
- InvEval. apply eval_orimm; auto.
-(* shlimm - shruimm *)
- predSpec Int.eq Int.eq_spec (Int.add n1 n2) Int.iwordsize.
- destruct (same_expr_pure t1 t2) eqn:?.
- InvEval. exploit eval_same_expr; eauto. intros [EQ1 EQ2]; subst.
- exists (Val.ror v0 (Vint n2)); split. EvalOp.
- destruct v0; simpl; auto.
- destruct (Int.ltu n1 Int.iwordsize) eqn:?; auto.
- destruct (Int.ltu n2 Int.iwordsize) eqn:?; auto.
- simpl. rewrite <- Int.or_ror; auto.
- InvEval. exists (Val.or x y); split. EvalOp.
- simpl. erewrite int_add_sub_eq; eauto. rewrite H0; rewrite H; auto. auto.
- TrivialExists.
-(* shruimm - shlimm *)
- predSpec Int.eq Int.eq_spec (Int.add n1 n2) Int.iwordsize.
- destruct (same_expr_pure t1 t2) eqn:?.
- InvEval. exploit eval_same_expr; eauto. intros [EQ1 EQ2]; subst.
- exists (Val.ror v1 (Vint n2)); split. EvalOp.
- destruct v1; simpl; auto.
- destruct (Int.ltu n2 Int.iwordsize) eqn:?; auto.
- destruct (Int.ltu n1 Int.iwordsize) eqn:?; auto.
- simpl. rewrite Int.or_commut. rewrite <- Int.or_ror; auto.
- InvEval. exists (Val.or y x); split. EvalOp.
- simpl. erewrite int_add_sub_eq; eauto. rewrite H0; rewrite H; auto.
- rewrite Val.or_commut; auto.
- TrivialExists.
-(* default *)
- TrivialExists.
-Qed.
-
-Theorem eval_xorimm:
- forall n, unary_constructor_sound (xorimm n) (fun x => Val.xor x (Vint n)).
-Proof.
- intros; red; intros until x. unfold xorimm.
- predSpec Int.eq Int.eq_spec n Int.zero.
- intros. exists x; split. auto.
- destruct x; simpl; auto. subst n. rewrite Int.xor_zero. auto.
- destruct (xorimm_match a); intros; InvEval.
- TrivialExists. simpl. rewrite Int.xor_commut; auto.
- subst. rewrite Val.xor_assoc. simpl. rewrite Int.xor_commut. TrivialExists.
- subst. rewrite Val.not_xor. rewrite Val.xor_assoc.
- rewrite (Val.xor_commut (Vint Int.mone)). TrivialExists.
- TrivialExists.
-Qed.
-
-Theorem eval_xor: binary_constructor_sound xor Val.xor.
-Proof.
- red; intros until y; unfold xor; case (xor_match a b); intros; InvEval.
- rewrite Val.xor_commut. apply eval_xorimm; auto.
- apply eval_xorimm; auto.
- TrivialExists.
-Qed.
-
-Theorem eval_divs_base:
- forall le a b x y z,
- eval_expr ge sp e m le a x ->
- eval_expr ge sp e m le b y ->
- Val.divs x y = Some z ->
- exists v, eval_expr ge sp e m le (divs_base a b) v /\ Val.lessdef z v.
-Proof.
- intros. unfold divs_base. exists z; split. EvalOp. auto.
-Qed.
-
-Theorem eval_divu_base:
- forall le a b x y z,
- eval_expr ge sp e m le a x ->
- eval_expr ge sp e m le b y ->
- Val.divu x y = Some z ->
- exists v, eval_expr ge sp e m le (divu_base a b) v /\ Val.lessdef z v.
-Proof.
- intros. unfold divu_base. exists z; split. EvalOp. auto.
-Qed.
-
-Theorem eval_mods_base:
- forall le a b x y z,
- eval_expr ge sp e m le a x ->
- eval_expr ge sp e m le b y ->
- Val.mods x y = Some z ->
- exists v, eval_expr ge sp e m le (mods_base a b) v /\ Val.lessdef z v.
-Proof.
- intros. unfold mods_base. exists z; split. EvalOp. auto.
-Qed.
-
-Theorem eval_modu_base:
- forall le a b x y z,
- eval_expr ge sp e m le a x ->
- eval_expr ge sp e m le b y ->
- Val.modu x y = Some z ->
- exists v, eval_expr ge sp e m le (modu_base a b) v /\ Val.lessdef z v.
-Proof.
- intros. unfold modu_base. exists z; split. EvalOp. auto.
-Qed.
-
-Theorem eval_shrximm:
- forall le a n x z,
- eval_expr ge sp e m le a x ->
- Val.shrx x (Vint n) = Some z ->
- exists v, eval_expr ge sp e m le (shrximm a n) v /\ Val.lessdef z v.
-Proof.
- intros. unfold shrximm.
- predSpec Int.eq Int.eq_spec n Int.zero.
- subst n. exists x; split; auto.
- destruct x; simpl in H0; try discriminate.
- destruct (Int.ltu Int.zero (Int.repr 31)); inv H0.
- replace (Int.shrx i Int.zero) with i. auto.
- unfold Int.shrx, Int.divs. rewrite Int.shl_zero.
- change (Int.signed Int.one) with 1. rewrite Z.quot_1_r. rewrite Int.repr_signed; auto.
- econstructor; split. EvalOp. auto.
-Qed.
-
-Theorem eval_shl: binary_constructor_sound shl Val.shl.
-Proof.
- red; intros until y; unfold shl; case (shl_match b); intros.
- InvEval. apply eval_shlimm; auto.
- TrivialExists.
-Qed.
-
-Theorem eval_shr: binary_constructor_sound shr Val.shr.
-Proof.
- red; intros until y; unfold shr; case (shr_match b); intros.
- InvEval. apply eval_shrimm; auto.
- TrivialExists.
-Qed.
-
-Theorem eval_shru: binary_constructor_sound shru Val.shru.
-Proof.
- red; intros until y; unfold shru; case (shru_match b); intros.
- InvEval. apply eval_shruimm; auto.
- TrivialExists.
-Qed.
-
-Theorem eval_negf: unary_constructor_sound negf Val.negf.
-Proof.
- red; intros. TrivialExists.
-Qed.
-
-Theorem eval_absf: unary_constructor_sound absf Val.absf.
-Proof.
- red; intros. TrivialExists.
-Qed.
-
-Theorem eval_addf: binary_constructor_sound addf Val.addf.
-Proof.
- red; intros; TrivialExists.
-Qed.
-
-Theorem eval_subf: binary_constructor_sound subf Val.subf.
-Proof.
- red; intros; TrivialExists.
-Qed.
-
-Theorem eval_mulf: binary_constructor_sound mulf Val.mulf.
-Proof.
- red; intros; TrivialExists.
-Qed.
-
-Theorem eval_negfs: unary_constructor_sound negfs Val.negfs.
-Proof.
- red; intros. TrivialExists.
-Qed.
-
-Theorem eval_absfs: unary_constructor_sound absfs Val.absfs.
-Proof.
- red; intros. TrivialExists.
-Qed.
-
-Theorem eval_addfs: binary_constructor_sound addfs Val.addfs.
-Proof.
- red; intros; TrivialExists.
-Qed.
-
-Theorem eval_subfs: binary_constructor_sound subfs Val.subfs.
-Proof.
- red; intros; TrivialExists.
-Qed.
-
-Theorem eval_mulfs: binary_constructor_sound mulfs Val.mulfs.
-Proof.
- red; intros; TrivialExists.
-Qed.
-
-Section COMP_IMM.
-
-Variable default: comparison -> int -> condition.
-Variable intsem: comparison -> int -> int -> bool.
-Variable sem: comparison -> val -> val -> val.
-
-Hypothesis sem_int: forall c x y, sem c (Vint x) (Vint y) = Val.of_bool (intsem c x y).
-Hypothesis sem_undef: forall c v, sem c Vundef v = Vundef.
-Hypothesis sem_eq: forall x y, sem Ceq (Vint x) (Vint y) = Val.of_bool (Int.eq x y).
-Hypothesis sem_ne: forall x y, sem Cne (Vint x) (Vint y) = Val.of_bool (negb (Int.eq x y)).
-Hypothesis sem_default: forall c v n, sem c v (Vint n) = Val.of_optbool (eval_condition (default c n) (v :: nil) m).
-
-Lemma eval_compimm:
- forall le c a n2 x,
- eval_expr ge sp e m le a x ->
- exists v, eval_expr ge sp e m le (compimm default intsem c a n2) v
- /\ Val.lessdef (sem c x (Vint n2)) v.
-Proof.
- intros until x.
- unfold compimm; case (compimm_match c a); intros.
-(* constant *)
- InvEval. rewrite sem_int. TrivialExists. simpl. destruct (intsem c0 n1 n2); auto.
-(* eq cmp *)
- InvEval. inv H. simpl in H5. inv H5.
- destruct (Int.eq_dec n2 Int.zero). subst n2. TrivialExists.
- simpl. rewrite eval_negate_condition.
- destruct (eval_condition c0 vl m); simpl.
- unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_eq; auto.
- rewrite sem_undef; auto.
- destruct (Int.eq_dec n2 Int.one). subst n2. TrivialExists.
- simpl. destruct (eval_condition c0 vl m); simpl.
- unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_eq; auto.
- rewrite sem_undef; auto.
- exists (Vint Int.zero); split. EvalOp.
- destruct (eval_condition c0 vl m); simpl.
- unfold Vtrue, Vfalse. destruct b; rewrite sem_eq; rewrite Int.eq_false; auto.
- rewrite sem_undef; auto.
-(* ne cmp *)
- InvEval. inv H. simpl in H5. inv H5.
- destruct (Int.eq_dec n2 Int.zero). subst n2. TrivialExists.
- simpl. destruct (eval_condition c0 vl m); simpl.
- unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_ne; auto.
- rewrite sem_undef; auto.
- destruct (Int.eq_dec n2 Int.one). subst n2. TrivialExists.
- simpl. rewrite eval_negate_condition. destruct (eval_condition c0 vl m); simpl.
- unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_ne; auto.
- rewrite sem_undef; auto.
- exists (Vint Int.one); split. EvalOp.
- destruct (eval_condition c0 vl m); simpl.
- unfold Vtrue, Vfalse. destruct b; rewrite sem_ne; rewrite Int.eq_false; auto.
- rewrite sem_undef; auto.
-(* eq andimm *)
- destruct (Int.eq_dec n2 Int.zero). InvEval; subst.
- econstructor; split. EvalOp. simpl; eauto.
- destruct v1; simpl; try (rewrite sem_undef; auto). rewrite sem_eq.
- destruct (Int.eq (Int.and i n1) Int.zero); auto.
- TrivialExists. simpl. rewrite sem_default. auto.
-(* ne andimm *)
- destruct (Int.eq_dec n2 Int.zero). InvEval; subst.
- econstructor; split. EvalOp. simpl; eauto.
- destruct v1; simpl; try (rewrite sem_undef; auto). rewrite sem_ne.
- destruct (Int.eq (Int.and i n1) Int.zero); auto.
- TrivialExists. simpl. rewrite sem_default. auto.
-(* default *)
- TrivialExists. simpl. rewrite sem_default. auto.
-Qed.
-
-Hypothesis sem_swap:
- forall c x y, sem (swap_comparison c) x y = sem c y x.
-
-Lemma eval_compimm_swap:
- forall le c a n2 x,
- eval_expr ge sp e m le a x ->
- exists v, eval_expr ge sp e m le (compimm default intsem (swap_comparison c) a n2) v
- /\ Val.lessdef (sem c (Vint n2) x) v.
-Proof.
- intros. rewrite <- sem_swap. eapply eval_compimm; eauto.
-Qed.
-
-End COMP_IMM.
-
-Theorem eval_comp:
- forall c, binary_constructor_sound (comp c) (Val.cmp c).
-Proof.
- intros; red; intros until y. unfold comp; case (comp_match a b); intros; InvEval.
- eapply eval_compimm_swap; eauto.
- intros. unfold Val.cmp. rewrite Val.swap_cmp_bool; auto.
- eapply eval_compimm; eauto.
- TrivialExists.
-Qed.
-
-Theorem eval_compu:
- forall c, binary_constructor_sound (compu c) (Val.cmpu (Mem.valid_pointer m) c).
-Proof.
- intros; red; intros until y. unfold compu; case (compu_match a b); intros; InvEval.
- eapply eval_compimm_swap; eauto.
- intros. unfold Val.cmpu. rewrite Val.swap_cmpu_bool; auto.
- eapply eval_compimm; eauto.
- TrivialExists.
-Qed.
-
-Theorem eval_compf:
- forall c, binary_constructor_sound (compf c) (Val.cmpf c).
-Proof.
- intros; red; intros. unfold compf. TrivialExists.
-Qed.
-
-Theorem eval_compfs:
- forall c, binary_constructor_sound (compfs c) (Val.cmpfs c).
-Proof.
- intros; red; intros. unfold compfs. TrivialExists.
-Qed.
-
-Theorem eval_cast8signed: unary_constructor_sound cast8signed (Val.sign_ext 8).
-Proof.
- red; intros until x. unfold cast8signed. case (cast8signed_match a); intros; InvEval.
- TrivialExists.
- TrivialExists.
-Qed.
-
-Theorem eval_cast8unsigned: unary_constructor_sound cast8unsigned (Val.zero_ext 8).
-Proof.
- red; intros until x. unfold cast8unsigned. destruct (cast8unsigned_match a); intros; InvEval.
- TrivialExists.
- subst. rewrite Val.zero_ext_and. rewrite Val.and_assoc.
- rewrite Int.and_commut. apply eval_andimm; auto. compute; auto.
- TrivialExists.
-Qed.
-
-Theorem eval_cast16signed: unary_constructor_sound cast16signed (Val.sign_ext 16).
-Proof.
- red; intros until x. unfold cast16signed. case (cast16signed_match a); intros; InvEval.
- TrivialExists.
- TrivialExists.
-Qed.
-
-Theorem eval_cast16unsigned: unary_constructor_sound cast16unsigned (Val.zero_ext 16).
-Proof.
- red; intros until x. unfold cast16unsigned. destruct (cast16unsigned_match a); intros; InvEval.
- TrivialExists.
- subst. rewrite Val.zero_ext_and. rewrite Val.and_assoc.
- rewrite Int.and_commut. apply eval_andimm; auto. compute; auto.
- TrivialExists.
-Qed.
-
-Theorem eval_singleoffloat: unary_constructor_sound singleoffloat Val.singleoffloat.
-Proof.
- red; intros. unfold singleoffloat. TrivialExists.
-Qed.
-
-Theorem eval_floatofsingle: unary_constructor_sound floatofsingle Val.floatofsingle.
-Proof.
- red; intros. unfold floatofsingle. TrivialExists.
-Qed.
-
-Theorem eval_intoffloat:
- forall le a x y,
- eval_expr ge sp e m le a x ->
- Val.intoffloat x = Some y ->
- exists v, eval_expr ge sp e m le (intoffloat a) v /\ Val.lessdef y v.
-Proof.
- intros; unfold intoffloat. TrivialExists.
-Qed.
-
-Theorem eval_floatofint:
- forall le a x y,
- eval_expr ge sp e m le a x ->
- Val.floatofint x = Some y ->
- exists v, eval_expr ge sp e m le (floatofint a) v /\ Val.lessdef y v.
-Proof.
- intros until y; unfold floatofint. case (floatofint_match a); intros; InvEval.
- TrivialExists.
- TrivialExists.
-Qed.
-
-Theorem eval_intuoffloat:
- forall le a x y,
- eval_expr ge sp e m le a x ->
- Val.intuoffloat x = Some y ->
- exists v, eval_expr ge sp e m le (intuoffloat a) v /\ Val.lessdef y v.
-Proof.
- intros. destruct x; simpl in H0; try discriminate.
- destruct (Float.to_intu f) as [n|] eqn:?; simpl in H0; inv H0.
- exists (Vint n); split; auto. unfold intuoffloat.
- set (im := Int.repr Int.half_modulus).
- set (fm := Float.of_intu im).
- assert (eval_expr ge sp e m (Vfloat fm :: Vfloat f :: le) (Eletvar (S O)) (Vfloat f)).
- constructor. auto.
- assert (eval_expr ge sp e m (Vfloat fm :: Vfloat f :: le) (Eletvar O) (Vfloat fm)).
- constructor. auto.
- econstructor. eauto.
- econstructor. instantiate (1 := Vfloat fm). EvalOp.
- eapply eval_Econdition with (va := Float.cmp Clt f fm).
- eauto with evalexpr.
- destruct (Float.cmp Clt f fm) eqn:?.
- exploit Float.to_intu_to_int_1; eauto. intro EQ.
- EvalOp. simpl. rewrite EQ; auto.
- exploit Float.to_intu_to_int_2; eauto.
- change Float.ox8000_0000 with im. fold fm. intro EQ.
- set (t2 := subf (Eletvar (S O)) (Eletvar O)).
- set (t3 := intoffloat t2).
- exploit (eval_subf (Vfloat fm :: Vfloat f :: le) (Eletvar (S O)) (Vfloat f) (Eletvar O)); eauto.
- fold t2. intros [v2 [A2 B2]]. simpl in B2. inv B2.
- exploit (eval_addimm Float.ox8000_0000 (Vfloat fm :: Vfloat f :: le) t3).
- unfold t3. unfold intoffloat. EvalOp. simpl. rewrite EQ. simpl. eauto.
- intros [v4 [A4 B4]]. simpl in B4. inv B4.
- rewrite Int.sub_add_opp in A4. rewrite Int.add_assoc in A4.
- rewrite (Int.add_commut (Int.neg im)) in A4.
- rewrite Int.add_neg_zero in A4.
- rewrite Int.add_zero in A4.
- auto.
-Qed.
-
-Theorem eval_floatofintu:
- forall le a x y,
- eval_expr ge sp e m le a x ->
- Val.floatofintu x = Some y ->
- exists v, eval_expr ge sp e m le (floatofintu a) v /\ Val.lessdef y v.
-Proof.
- intros until y; unfold floatofintu. case (floatofintu_match a); intros.
- InvEval. TrivialExists.
- destruct x; simpl in H0; try discriminate. inv H0.
- exists (Vfloat (Float.of_intu i)); split; auto.
- econstructor. eauto.
- set (fm := Float.of_intu Float.ox8000_0000).
- assert (eval_expr ge sp e m (Vint i :: le) (Eletvar O) (Vint i)).
- constructor. auto.
- eapply eval_Econdition with (va := Int.ltu i Float.ox8000_0000).
- eauto with evalexpr.
- destruct (Int.ltu i Float.ox8000_0000) eqn:?.
- rewrite Float.of_intu_of_int_1; auto.
- unfold floatofint. EvalOp.
- exploit (eval_addimm (Int.neg Float.ox8000_0000) (Vint i :: le) (Eletvar 0)); eauto.
- simpl. intros [v [A B]]. inv B.
- unfold addf. EvalOp.
- constructor. unfold floatofint. EvalOp. simpl; eauto.
- constructor. EvalOp. simpl; eauto. constructor. simpl; eauto.
- fold fm. rewrite Float.of_intu_of_int_2; auto.
- rewrite Int.sub_add_opp. auto.
-Qed.
-
-Theorem eval_intofsingle:
- forall le a x y,
- eval_expr ge sp e m le a x ->
- Val.intofsingle x = Some y ->
- exists v, eval_expr ge sp e m le (intofsingle a) v /\ Val.lessdef y v.
-Proof.
- intros; unfold intofsingle. TrivialExists.
-Qed.
-
-Theorem eval_singleofint:
- forall le a x y,
- eval_expr ge sp e m le a x ->
- Val.singleofint x = Some y ->
- exists v, eval_expr ge sp e m le (singleofint a) v /\ Val.lessdef y v.
-Proof.
- intros until y; unfold singleofint. case (singleofint_match a); intros; InvEval.
- TrivialExists.
- TrivialExists.
-Qed.
-
-Theorem eval_intuofsingle:
- forall le a x y,
- eval_expr ge sp e m le a x ->
- Val.intuofsingle x = Some y ->
- exists v, eval_expr ge sp e m le (intuofsingle a) v /\ Val.lessdef y v.
-Proof.
- intros. destruct x; simpl in H0; try discriminate.
- destruct (Float32.to_intu f) as [n|] eqn:?; simpl in H0; inv H0.
- unfold intuofsingle. apply eval_intuoffloat with (Vfloat (Float.of_single f)).
- unfold floatofsingle. EvalOp.
- simpl. change (Float.of_single f) with (Float32.to_double f).
- erewrite Float32.to_intu_double; eauto. auto.
-Qed.
-
-Theorem eval_singleofintu:
- forall le a x y,
- eval_expr ge sp e m le a x ->
- Val.singleofintu x = Some y ->
- exists v, eval_expr ge sp e m le (singleofintu a) v /\ Val.lessdef y v.
-Proof.
- intros until y; unfold singleofintu. case (singleofintu_match a); intros.
- InvEval. TrivialExists.
- destruct x; simpl in H0; try discriminate. inv H0.
- exploit eval_floatofintu. eauto. simpl. reflexivity.
- intros (v & A & B).
- exists (Val.singleoffloat v); split.
- unfold singleoffloat; EvalOp.
- inv B; simpl. rewrite Float32.of_intu_double. auto.
-Qed.
-
-Theorem eval_addressing:
- forall le chunk a v b ofs,
- eval_expr ge sp e m le a v ->
- v = Vptr b ofs ->
- match addressing chunk a with (mode, args) =>
- exists vl,
- eval_exprlist ge sp e m le args vl /\
- eval_addressing ge sp mode vl = Some v
- end.
-Proof.
- intros until ofs.
- assert (A: v = Vptr b ofs -> eval_addressing ge sp (Aindexed 0) (v :: nil) = Some v).
- { intros. subst v. unfold eval_addressing.
- destruct Archi.ptr64 eqn:SF; simpl; rewrite SF; rewrite Ptrofs.add_zero; auto. }
- assert (D: forall a,
- eval_expr ge sp e m le a v ->
- v = Vptr b ofs ->
- exists vl, eval_exprlist ge sp e m le (a ::: Enil) vl
- /\ eval_addressing ge sp (Aindexed 0) vl = Some v).
- { intros. exists (v :: nil); split. constructor; auto. constructor. auto. }
- unfold addressing; case (addressing_match a); intros.
-- destruct (negb Archi.ptr64 && addressing_valid addr) eqn:E.
-+ inv H. InvBooleans. apply negb_true_iff in H. unfold eval_addressing; rewrite H.
- exists vl; auto.
-+ apply D; auto.
-- destruct (Archi.ptr64 && addressing_valid addr) eqn:E.
-+ inv H. InvBooleans. unfold eval_addressing; rewrite H.
- exists vl; auto.
-+ apply D; auto.
-- apply D; auto.
-Qed.
-
-Theorem eval_builtin_arg:
- forall a v,
- eval_expr ge sp e m nil a v ->
- CminorSel.eval_builtin_arg ge sp e m (builtin_arg a) v.
-Proof.
- intros until v. unfold builtin_arg; case (builtin_arg_match a); intros; InvEval.
-- constructor.
-- constructor.
-- destruct Archi.ptr64; inv H0. constructor.
-- destruct Archi.ptr64; inv H0. constructor.
-- destruct Archi.ptr64; inv H0. constructor.
-- destruct Archi.ptr64; inv H0. constructor.
-- simpl in H5. inv H5. constructor.
-- subst v. constructor; auto.
-- inv H. InvEval. rewrite eval_addressing_Aglobal in H6. inv H6. constructor; auto.
-- inv H. InvEval. rewrite eval_addressing_Ainstack in H6. inv H6. constructor; auto.
-- constructor; auto.
-Qed.
-
-End CMCONSTR.
diff --git a/ia32/Stacklayout.v b/ia32/Stacklayout.v
deleted file mode 100644
index 44fd43b2..00000000
--- a/ia32/Stacklayout.v
+++ /dev/null
@@ -1,148 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Machine- and ABI-dependent layout information for activation records. *)
-
-Require Import Coqlib.
-Require Import AST Memory Separation.
-Require Import Bounds.
-
-Local Open Scope sep_scope.
-
-(** The general shape of activation records is as follows,
- from bottom (lowest offsets) to top:
-- Space for outgoing arguments to function calls.
-- Back link to parent frame
-- Saved values of integer callee-save registers used by the function.
-- Saved values of float callee-save registers used by the function.
-- Local stack slots.
-- Space for the stack-allocated data declared in Cminor
-- Return address.
-*)
-
-Definition fe_ofs_arg := 0.
-
-Definition make_env (b: bounds) : frame_env :=
- let w := if Archi.ptr64 then 8 else 4 in
- let olink := align (4 * b.(bound_outgoing)) w in (* back link *)
- let ocs := olink + w in (* callee-saves *)
- let ol := align (size_callee_save_area b ocs) 8 in (* locals *)
- let ostkdata := align (ol + 4 * b.(bound_local)) 8 in (* stack data *)
- let oretaddr := align (ostkdata + b.(bound_stack_data)) w in (* return address *)
- let sz := oretaddr + w in (* total size *)
- {| fe_size := sz;
- fe_ofs_link := olink;
- fe_ofs_retaddr := oretaddr;
- fe_ofs_local := ol;
- fe_ofs_callee_save := ocs;
- fe_stack_data := ostkdata;
- fe_used_callee_save := b.(used_callee_save) |}.
-
-Lemma frame_env_separated:
- forall b sp m P,
- let fe := make_env b in
- m |= range sp 0 (fe_stack_data fe) ** range sp (fe_stack_data fe + bound_stack_data b) (fe_size fe) ** P ->
- m |= range sp (fe_ofs_local fe) (fe_ofs_local fe + 4 * bound_local b)
- ** range sp fe_ofs_arg (fe_ofs_arg + 4 * bound_outgoing b)
- ** range sp (fe_ofs_link fe) (fe_ofs_link fe + size_chunk Mptr)
- ** range sp (fe_ofs_retaddr fe) (fe_ofs_retaddr fe + size_chunk Mptr)
- ** range sp (fe_ofs_callee_save fe) (size_callee_save_area b (fe_ofs_callee_save fe))
- ** P.
-Proof.
-Local Opaque Z.add Z.mul sepconj range.
- intros; simpl.
- set (w := if Archi.ptr64 then 8 else 4).
- set (olink := align (4 * b.(bound_outgoing)) w).
- set (ocs := olink + w).
- set (ol := align (size_callee_save_area b ocs) 8).
- set (ostkdata := align (ol + 4 * b.(bound_local)) 8).
- set (oretaddr := align (ostkdata + b.(bound_stack_data)) w).
- replace (size_chunk Mptr) with w by (rewrite size_chunk_Mptr; auto).
- assert (0 < w) by (unfold w; destruct Archi.ptr64; omega).
- generalize b.(bound_local_pos) b.(bound_outgoing_pos) b.(bound_stack_data_pos); intros.
- assert (0 <= 4 * b.(bound_outgoing)) by omega.
- assert (4 * b.(bound_outgoing) <= olink) by (apply align_le; omega).
- assert (olink + w <= ocs) by (unfold ocs; omega).
- assert (ocs <= size_callee_save_area b ocs) by (apply size_callee_save_area_incr).
- assert (size_callee_save_area b ocs <= ol) by (apply align_le; omega).
- assert (ol + 4 * b.(bound_local) <= ostkdata) by (apply align_le; omega).
- assert (ostkdata + bound_stack_data b <= oretaddr) by (apply align_le; omega).
-(* Reorder as:
- outgoing
- back link
- callee-save
- local
- retaddr *)
- rewrite sep_swap12.
- rewrite sep_swap23.
- rewrite sep_swap45.
- rewrite sep_swap34.
-(* Apply range_split and range_split2 repeatedly *)
- unfold fe_ofs_arg.
- apply range_split_2. fold olink. omega. omega.
- apply range_split. omega.
- apply range_split_2. fold ol. omega. omega.
- apply range_drop_right with ostkdata. omega.
- rewrite sep_swap.
- apply range_drop_left with (ostkdata + bound_stack_data b). omega.
- rewrite sep_swap.
- exact H.
-Qed.
-
-Lemma frame_env_range:
- forall b,
- let fe := make_env b in
- 0 <= fe_stack_data fe /\ fe_stack_data fe + bound_stack_data b <= fe_size fe.
-Proof.
- intros; simpl.
- set (w := if Archi.ptr64 then 8 else 4).
- set (olink := align (4 * b.(bound_outgoing)) w).
- set (ocs := olink + w).
- set (ol := align (size_callee_save_area b ocs) 8).
- set (ostkdata := align (ol + 4 * b.(bound_local)) 8).
- set (oretaddr := align (ostkdata + b.(bound_stack_data)) w).
- assert (0 < w) by (unfold w; destruct Archi.ptr64; omega).
- generalize b.(bound_local_pos) b.(bound_outgoing_pos) b.(bound_stack_data_pos); intros.
- assert (0 <= 4 * b.(bound_outgoing)) by omega.
- assert (4 * b.(bound_outgoing) <= olink) by (apply align_le; omega).
- assert (olink + w <= ocs) by (unfold ocs; omega).
- assert (ocs <= size_callee_save_area b ocs) by (apply size_callee_save_area_incr).
- assert (size_callee_save_area b ocs <= ol) by (apply align_le; omega).
- assert (ol + 4 * b.(bound_local) <= ostkdata) by (apply align_le; omega).
- assert (ostkdata + bound_stack_data b <= oretaddr) by (apply align_le; omega).
- split. omega. omega.
-Qed.
-
-Lemma frame_env_aligned:
- forall b,
- let fe := make_env b in
- (8 | fe_ofs_arg)
- /\ (8 | fe_ofs_local fe)
- /\ (8 | fe_stack_data fe)
- /\ (align_chunk Mptr | fe_ofs_link fe)
- /\ (align_chunk Mptr | fe_ofs_retaddr fe).
-Proof.
- intros; simpl.
- set (w := if Archi.ptr64 then 8 else 4).
- set (olink := align (4 * b.(bound_outgoing)) w).
- set (ocs := olink + w).
- set (ol := align (size_callee_save_area b ocs) 8).
- set (ostkdata := align (ol + 4 * b.(bound_local)) 8).
- set (oretaddr := align (ostkdata + b.(bound_stack_data)) w).
- assert (0 < w) by (unfold w; destruct Archi.ptr64; omega).
- replace (align_chunk Mptr) with w by (rewrite align_chunk_Mptr; auto).
- split. apply Zdivide_0.
- split. apply align_divides; omega.
- split. apply align_divides; omega.
- split. apply align_divides; omega.
- apply align_divides; omega.
-Qed.
diff --git a/ia32/TargetPrinter.ml b/ia32/TargetPrinter.ml
deleted file mode 100644
index 33d47830..00000000
--- a/ia32/TargetPrinter.ml
+++ /dev/null
@@ -1,881 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-(* Printing x86-64 assembly code in asm syntax *)
-
-open Printf
-open !Datatypes
-open Camlcoq
-open Sections
-open AST
-open Asm
-open PrintAsmaux
-open Fileinfo
-
-module StringSet = Set.Make(String)
-
-(* Basic printing functions used in definition of the systems *)
-
-let int64_reg_name = function
- | RAX -> "%rax" | RBX -> "%rbx" | RCX -> "%rcx" | RDX -> "%rdx"
- | RSI -> "%rsi" | RDI -> "%rdi" | RBP -> "%rbp" | RSP -> "%rsp"
- | R8 -> "%r8" | R9 -> "%r9" | R10 -> "%r10" | R11 -> "%r11"
- | R12 -> "%r12" | R13 -> "%r13" | R14 -> "%r14" | R15 -> "%r15"
-
-let int32_reg_name = function
- | RAX -> "%eax" | RBX -> "%ebx" | RCX -> "%ecx" | RDX -> "%edx"
- | RSI -> "%esi" | RDI -> "%edi" | RBP -> "%ebp" | RSP -> "%esp"
- | R8 -> "%r8d" | R9 -> "%r9d" | R10 -> "%r10d" | R11 -> "%r11d"
- | R12 -> "%r12d" | R13 -> "%r13d" | R14 -> "%r14d" | R15 -> "%r15d"
-
-let int8_reg_name = function
- | RAX -> "%al" | RBX -> "%bl" | RCX -> "%cl" | RDX -> "%dl"
- | RSI -> "%sil" | RDI -> "%dil" | RBP -> "%bpl" | RSP -> "%spl"
- | R8 -> "%r8b" | R9 -> "%r9b" | R10 -> "%r10b" | R11 -> "%r11b"
- | R12 -> "%r12b" | R13 -> "%r13b" | R14 -> "%r14b" | R15 -> "%r15b"
-
-let int16_reg_name = function
- | RAX -> "%ax" | RBX -> "%bx" | RCX -> "%cx" | RDX -> "%dx"
- | RSI -> "%si" | RDI -> "%di" | RBP -> "%bp" | RSP -> "%sp"
- | R8 -> "%r8w" | R9 -> "%r9w" | R10 -> "%r10w" | R11 -> "%r11w"
- | R12 -> "%r12w" | R13 -> "%r13w" | R14 -> "%r14w" | R15 -> "%r15w"
-
-let float_reg_name = function
- | XMM0 -> "%xmm0" | XMM1 -> "%xmm1" | XMM2 -> "%xmm2" | XMM3 -> "%xmm3"
- | XMM4 -> "%xmm4" | XMM5 -> "%xmm5" | XMM6 -> "%xmm6" | XMM7 -> "%xmm7"
- | XMM8 -> "%xmm8" | XMM9 -> "%xmm9" | XMM10 -> "%xmm10" | XMM11 -> "%xmm11"
- | XMM12 -> "%xmm12" | XMM13 -> "%xmm13" | XMM14 -> "%xmm14" | XMM15 -> "%xmm15"
-
-let ireg8 oc r = output_string oc (int8_reg_name r)
-let ireg16 oc r = output_string oc (int16_reg_name r)
-let ireg32 oc r = output_string oc (int32_reg_name r)
-let ireg64 oc r = output_string oc (int64_reg_name r)
-let ireg = if Archi.ptr64 then ireg64 else ireg32
-let freg oc r = output_string oc (float_reg_name r)
-
-let preg oc = function
- | IR r -> ireg oc r
- | FR r -> freg oc r
- | _ -> assert false
-
-let z oc n = output_string oc (Z.to_string n)
-
-(* 32/64 bit dependencies *)
-
-let data_pointer = if Archi.ptr64 then ".quad" else ".long"
-
-(* The comment deliminiter *)
-let comment = "#"
-
-(* System dependend printer functions *)
-module type SYSTEM =
- sig
- val raw_symbol: out_channel -> string -> unit
- val symbol: out_channel -> P.t -> unit
- val label: out_channel -> int -> unit
- val name_of_section: section_name -> string
- val stack_alignment: int
- val print_align: out_channel -> int -> unit
- val print_mov_rs: out_channel -> ireg -> ident -> unit
- val print_fun_info: out_channel -> P.t -> unit
- val print_var_info: out_channel -> P.t -> unit
- val print_epilogue: out_channel -> unit
- val print_comm_decl: out_channel -> P.t -> Z.t -> int -> unit
- val print_lcomm_decl: out_channel -> P.t -> Z.t -> int -> unit
- end
-
-(* Printer functions for ELF *)
-module ELF_System : SYSTEM =
- struct
-
- let raw_symbol oc s =
- fprintf oc "%s" s
-
- let symbol = elf_symbol
-
- let label = elf_label
-
- let name_of_section = function
- | Section_text -> ".text"
- | Section_data i | Section_small_data i ->
- if i then ".data" else "COMM"
- | Section_const i | Section_small_const i ->
- if i then ".section .rodata" else "COMM"
- | Section_string -> ".section .rodata"
- | Section_literal -> ".section .rodata.cst8,\"aM\",@progbits,8"
- | Section_jumptable -> ".text"
- | Section_user(s, wr, ex) ->
- sprintf ".section \"%s\",\"a%s%s\",@progbits"
- s (if wr then "w" else "") (if ex then "x" else "")
- | Section_debug_info _ -> ".section .debug_info,\"\",@progbits"
- | Section_debug_loc -> ".section .debug_loc,\"\",@progbits"
- | Section_debug_line _ -> ".section .debug_line,\"\",@progbits"
- | Section_debug_abbrev -> ".section .debug_abbrev,\"\",@progbits"
- | Section_debug_ranges -> ".section .debug_ranges,\"\",@progbits"
- | Section_debug_str -> ".section .debug_str,\"MS\",@progbits,1"
-
- let stack_alignment = 16
-
- let print_align oc n =
- fprintf oc " .align %d\n" n
-
- let print_mov_rs oc rd id =
- if Archi.ptr64
- then fprintf oc " movq %a@GOTPCREL(%%rip), %a\n" symbol id ireg64 rd
- else fprintf oc " movl $%a, %a\n" symbol id ireg32 rd
-
- let print_fun_info = elf_print_fun_info
-
- let print_var_info = elf_print_var_info
-
- let print_epilogue _ = ()
-
- let print_comm_decl oc name sz al =
- fprintf oc " .comm %a, %s, %d\n" symbol name (Z.to_string sz) al
-
- let print_lcomm_decl oc name sz al =
- fprintf oc " .local %a\n" symbol name;
- print_comm_decl oc name sz al
-
- end
-
-(* Printer functions for MacOS *)
-module MacOS_System : SYSTEM =
- struct
-
- let raw_symbol oc s =
- fprintf oc "_%s" s
-
- let symbol oc symb =
- raw_symbol oc (extern_atom symb)
-
- let label oc lbl =
- fprintf oc "L%d" lbl
-
- let name_of_section = function
- | Section_text -> ".text"
- | Section_data i | Section_small_data i ->
- if i then ".data" else "COMM"
- | Section_const i | Section_small_const i ->
- if i then ".const" else "COMM"
- | Section_string -> ".const"
- | Section_literal -> ".literal8"
- | Section_jumptable -> ".const"
- | Section_user(s, wr, ex) ->
- sprintf ".section \"%s\", %s, %s"
- (if wr then "__DATA" else "__TEXT") s
- (if ex then "regular, pure_instructions" else "regular")
- | Section_debug_info _ -> ".section __DWARF,__debug_info,regular,debug"
- | Section_debug_loc -> ".section __DWARF,__debug_loc,regular,debug"
- | Section_debug_line _ -> ".section __DWARF,__debug_line,regular,debug"
- | Section_debug_str -> ".section __DWARF,__debug_str,regular,debug"
- | Section_debug_ranges -> ".section __DWARF,__debug_ranges,regular,debug"
- | Section_debug_abbrev -> ".section __DWARF,__debug_abbrev,regular,debug"
-
-
- let stack_alignment = 16 (* mandatory *)
-
- (* Base-2 log of a Caml integer *)
- let rec log2 n =
- assert (n > 0);
- if n = 1 then 0 else 1 + log2 (n lsr 1)
-
- let print_align oc n =
- fprintf oc " .align %d\n" (log2 n)
-
- let indirect_symbols : StringSet.t ref = ref StringSet.empty
-
- let print_mov_rs oc rd id =
- if Archi.ptr64 then begin
- fprintf oc " movq %a@GOTPCREL(%%rip), %a\n" symbol id ireg64 rd
- end else begin
- let id = extern_atom id in
- indirect_symbols := StringSet.add id !indirect_symbols;
- fprintf oc " movl L%a$non_lazy_ptr, %a\n" raw_symbol id ireg rd
- end
-
- let print_fun_info _ _ = ()
-
- let print_var_info _ _ = ()
-
- let print_epilogue oc =
- if not Archi.ptr64 then begin
- fprintf oc " .section __IMPORT,__pointers,non_lazy_symbol_pointers\n";
- StringSet.iter
- (fun s ->
- fprintf oc "L%a$non_lazy_ptr:\n" raw_symbol s;
- fprintf oc " .indirect_symbol %a\n" raw_symbol s;
- fprintf oc " .long 0\n")
- !indirect_symbols;
- indirect_symbols := StringSet.empty
- end
-
- let print_comm_decl oc name sz al =
- fprintf oc " .comm %a, %s, %d\n"
- symbol name (Z.to_string sz) (log2 al)
-
- let print_lcomm_decl oc name sz al =
- fprintf oc " .lcomm %a, %s, %d\n"
- symbol name (Z.to_string sz) (log2 al)
-
- end
-
-
-module Target(System: SYSTEM):TARGET =
- struct
- open System
- let symbol = symbol
-
-(* Basic printing functions *)
-
- let symbol_offset oc (symb, ofs) =
- symbol oc symb;
- let ofs = Z.to_int64 ofs in
- if ofs <> 0L then fprintf oc " + %Ld" ofs
-
- let addressing_gen ireg oc (Addrmode(base, shift, cst)) =
- begin match cst with
- | Coq_inl n ->
- fprintf oc "%s" (Z.to_string n)
- | Coq_inr(id, ofs) ->
- if Archi.ptr64 then begin
- (* RIP-relative addressing *)
- let ofs' = Z.to_int64 ofs in
- if ofs' = 0L
- then fprintf oc "%a(%%rip)" symbol id
- else fprintf oc "(%a + %Ld)(%%rip)" symbol id ofs'
- end else begin
- (* Absolute addressing *)
- let ofs' = Z.to_int32 ofs in
- if ofs' = 0l
- then fprintf oc "%a" symbol id
- else fprintf oc "(%a + %ld)" symbol id ofs'
- end
- end;
- begin match base, shift with
- | None, None -> ()
- | Some r1, None -> fprintf oc "(%a)" ireg r1
- | None, Some(r2,sc) -> fprintf oc "(,%a,%a)" ireg r2 z sc
- | Some r1, Some(r2,sc) -> fprintf oc "(%a,%a,%a)" ireg r1 ireg r2 z sc
- end
-
- let addressing32 = addressing_gen ireg32
- let addressing64 = addressing_gen ireg64
- let addressing = addressing_gen ireg
-
- let name_of_condition = function
- | Cond_e -> "e" | Cond_ne -> "ne"
- | Cond_b -> "b" | Cond_be -> "be" | Cond_ae -> "ae" | Cond_a -> "a"
- | Cond_l -> "l" | Cond_le -> "le" | Cond_ge -> "ge" | Cond_g -> "g"
- | Cond_p -> "p" | Cond_np -> "np"
-
- let name_of_neg_condition = function
- | Cond_e -> "ne" | Cond_ne -> "e"
- | Cond_b -> "ae" | Cond_be -> "a" | Cond_ae -> "b" | Cond_a -> "be"
- | Cond_l -> "ge" | Cond_le -> "g" | Cond_ge -> "l" | Cond_g -> "le"
- | Cond_p -> "np" | Cond_np -> "p"
-
-
-(* Names of sections *)
-
- let section oc sec =
- fprintf oc " %s\n" (name_of_section sec)
-
-(* For "abs" and "neg" FP operations *)
-
- let need_masks = ref false
-
-(* Emit .file / .loc debugging directives *)
-
- let print_file_line oc file line =
- print_file_line oc comment file line
-
-(* In 64-bit mode use RIP-relative addressing to designate labels *)
-
- let rip_rel =
- if Archi.ptr64 then "(%rip)" else ""
-
-(* Large 64-bit immediates (bigger than a 32-bit signed integer) are
- not supported by the processor. Turn them into memory operands. *)
-
- let intconst64 oc n =
- let n1 = camlint64_of_coqint n in
- let n2 = Int64.to_int32 n1 in
- if n1 = Int64.of_int32 n2 then
- (* fit in a 32-bit signed integer, can use as immediate *)
- fprintf oc "$%ld" n2
- else begin
- (* put the constant in memory and use a PC-relative memory operand *)
- let lbl = new_label() in
- float64_literals := (lbl, n1) :: !float64_literals;
- fprintf oc "%a(%%rip)" label lbl
- end
-
-
-
-(* Printing of instructions *)
-
-(* Reminder on AT&T syntax: op source, dest *)
-
- let print_instruction oc = function
- (* Moves *)
- | Pmov_rr(rd, r1) ->
- if Archi.ptr64
- then fprintf oc " movq %a, %a\n" ireg64 r1 ireg64 rd
- else fprintf oc " movl %a, %a\n" ireg32 r1 ireg32 rd
- | Pmovl_ri(rd, n) ->
- fprintf oc " movl $%ld, %a\n" (camlint_of_coqint n) ireg32 rd
- | Pmovq_ri(rd, n) ->
- let n1 = camlint64_of_coqint n in
- let n2 = Int64.to_int32 n1 in
- if n1 = Int64.of_int32 n2 then
- fprintf oc " movq $%ld, %a\n" n2 ireg64 rd
- else
- fprintf oc " movabsq $%Ld, %a\n" n1 ireg64 rd
- | Pmov_rs(rd, id) ->
- print_mov_rs oc rd id
- | Pmovl_rm(rd, a) ->
- fprintf oc " movl %a, %a\n" addressing a ireg32 rd
- | Pmovq_rm(rd, a) ->
- fprintf oc " movq %a, %a\n" addressing a ireg64 rd
- | Pmov_rm_a(rd, a) ->
- if Archi.ptr64
- then fprintf oc " movq %a, %a\n" addressing a ireg64 rd
- else fprintf oc " movl %a, %a\n" addressing a ireg32 rd
- | Pmovl_mr(a, r1) ->
- fprintf oc " movl %a, %a\n" ireg32 r1 addressing a
- | Pmovq_mr(a, r1) ->
- fprintf oc " movq %a, %a\n" ireg64 r1 addressing a
- | Pmov_mr_a(a, r1) ->
- if Archi.ptr64
- then fprintf oc " movq %a, %a\n" ireg64 r1 addressing a
- else fprintf oc " movl %a, %a\n" ireg32 r1 addressing a
- | Pmovsd_ff(rd, r1) ->
- fprintf oc " movapd %a, %a\n" freg r1 freg rd
- | Pmovsd_fi(rd, n) ->
- let b = camlint64_of_coqint (Floats.Float.to_bits n) in
- let lbl = new_label() in
- fprintf oc " movsd %a%s, %a %s %.18g\n"
- label lbl rip_rel
- freg rd comment (camlfloat_of_coqfloat n);
- float64_literals := (lbl, b) :: !float64_literals
- | Pmovsd_fm(rd, a) | Pmovsd_fm_a(rd, a) ->
- fprintf oc " movsd %a, %a\n" addressing a freg rd
- | Pmovsd_mf(a, r1) | Pmovsd_mf_a(a, r1) ->
- fprintf oc " movsd %a, %a\n" freg r1 addressing a
- | Pmovss_fi(rd, n) ->
- let b = camlint_of_coqint (Floats.Float32.to_bits n) in
- let lbl = new_label() in
- fprintf oc " movss %a%s, %a %s %.18g\n"
- label lbl rip_rel
- freg rd comment (camlfloat_of_coqfloat32 n);
- float32_literals := (lbl, b) :: !float32_literals
- | Pmovss_fm(rd, a) ->
- fprintf oc " movss %a, %a\n" addressing a freg rd
- | Pmovss_mf(a, r1) ->
- fprintf oc " movss %a, %a\n" freg r1 addressing a
- | Pfldl_m(a) ->
- fprintf oc " fldl %a\n" addressing a
- | Pfstpl_m(a) ->
- fprintf oc " fstpl %a\n" addressing a
- | Pflds_m(a) ->
- fprintf oc " flds %a\n" addressing a
- | Pfstps_m(a) ->
- fprintf oc " fstps %a\n" addressing a
- (* Moves with conversion *)
- | Pmovb_mr(a, r1) ->
- fprintf oc " movb %a, %a\n" ireg8 r1 addressing a
- | Pmovw_mr(a, r1) ->
- fprintf oc " movw %a, %a\n" ireg16 r1 addressing a
- | Pmovzb_rr(rd, r1) ->
- fprintf oc " movzbl %a, %a\n" ireg8 r1 ireg32 rd
- | Pmovzb_rm(rd, a) ->
- fprintf oc " movzbl %a, %a\n" addressing a ireg32 rd
- | Pmovsb_rr(rd, r1) ->
- fprintf oc " movsbl %a, %a\n" ireg8 r1 ireg32 rd
- | Pmovsb_rm(rd, a) ->
- fprintf oc " movsbl %a, %a\n" addressing a ireg32 rd
- | Pmovzw_rr(rd, r1) ->
- fprintf oc " movzwl %a, %a\n" ireg16 r1 ireg32 rd
- | Pmovzw_rm(rd, a) ->
- fprintf oc " movzwl %a, %a\n" addressing a ireg32 rd
- | Pmovsw_rr(rd, r1) ->
- fprintf oc " movswl %a, %a\n" ireg16 r1 ireg32 rd
- | Pmovsw_rm(rd, a) ->
- fprintf oc " movswl %a, %a\n" addressing a ireg32 rd
- | Pmovzl_rr(rd, r1) ->
- fprintf oc " movl %a, %a\n" ireg32 r1 ireg32 rd
- (* movl sets the high 32 bits of the destination to zero *)
- | Pmovsl_rr(rd, r1) ->
- fprintf oc " movslq %a, %a\n" ireg32 r1 ireg64 rd
- | Pmovls_rr(rd) ->
- () (* nothing to do *)
- | Pcvtsd2ss_ff(rd, r1) ->
- fprintf oc " cvtsd2ss %a, %a\n" freg r1 freg rd
- | Pcvtss2sd_ff(rd, r1) ->
- fprintf oc " cvtss2sd %a, %a\n" freg r1 freg rd
- | Pcvttsd2si_rf(rd, r1) ->
- fprintf oc " cvttsd2si %a, %a\n" freg r1 ireg32 rd
- | Pcvtsi2sd_fr(rd, r1) ->
- fprintf oc " cvtsi2sd %a, %a\n" ireg32 r1 freg rd
- | Pcvttss2si_rf(rd, r1) ->
- fprintf oc " cvttss2si %a, %a\n" freg r1 ireg32 rd
- | Pcvtsi2ss_fr(rd, r1) ->
- fprintf oc " cvtsi2ss %a, %a\n" ireg32 r1 freg rd
- | Pcvttsd2sl_rf(rd, r1) ->
- fprintf oc " cvttsd2si %a, %a\n" freg r1 ireg64 rd
- | Pcvtsl2sd_fr(rd, r1) ->
- fprintf oc " cvtsi2sdq %a, %a\n" ireg64 r1 freg rd
- | Pcvttss2sl_rf(rd, r1) ->
- fprintf oc " cvttss2si %a, %a\n" freg r1 ireg64 rd
- | Pcvtsl2ss_fr(rd, r1) ->
- fprintf oc " cvtsi2ssq %a, %a\n" ireg64 r1 freg rd
- (* Arithmetic and logical operations over integers *)
- | Pleal(rd, a) ->
- fprintf oc " leal %a, %a\n" addressing32 a ireg32 rd
- | Pleaq(rd, a) ->
- fprintf oc " leaq %a, %a\n" addressing64 a ireg64 rd
- | Pnegl(rd) ->
- fprintf oc " negl %a\n" ireg32 rd
- | Pnegq(rd) ->
- fprintf oc " negq %a\n" ireg64 rd
- | Paddl_ri (res,n) ->
- fprintf oc " addl $%ld, %a\n" (camlint_of_coqint n) ireg32 res
- | Paddq_ri (res,n) ->
- fprintf oc " addq %a, %a\n" intconst64 n ireg64 res
- | Psubl_rr(rd, r1) ->
- fprintf oc " subl %a, %a\n" ireg32 r1 ireg32 rd
- | Psubq_rr(rd, r1) ->
- fprintf oc " subq %a, %a\n" ireg64 r1 ireg64 rd
- | Pimull_rr(rd, r1) ->
- fprintf oc " imull %a, %a\n" ireg32 r1 ireg32 rd
- | Pimulq_rr(rd, r1) ->
- fprintf oc " imulq %a, %a\n" ireg64 r1 ireg64 rd
- | Pimull_ri(rd, n) ->
- fprintf oc " imull $%a, %a\n" coqint n ireg32 rd
- | Pimulq_ri(rd, n) ->
- fprintf oc " imulq %a, %a\n" intconst64 n ireg64 rd
- | Pimull_r(r1) ->
- fprintf oc " imull %a\n" ireg32 r1
- | Pimulq_r(r1) ->
- fprintf oc " imulq %a\n" ireg64 r1
- | Pmull_r(r1) ->
- fprintf oc " mull %a\n" ireg32 r1
- | Pmulq_r(r1) ->
- fprintf oc " mulq %a\n" ireg64 r1
- | Pcltd ->
- fprintf oc " cltd\n"
- | Pcqto ->
- fprintf oc " cqto\n";
- | Pdivl(r1) ->
- fprintf oc " divl %a\n" ireg32 r1
- | Pdivq(r1) ->
- fprintf oc " divq %a\n" ireg64 r1
- | Pidivl(r1) ->
- fprintf oc " idivl %a\n" ireg32 r1
- | Pidivq(r1) ->
- fprintf oc " idivq %a\n" ireg64 r1
- | Pandl_rr(rd, r1) ->
- fprintf oc " andl %a, %a\n" ireg32 r1 ireg32 rd
- | Pandq_rr(rd, r1) ->
- fprintf oc " andq %a, %a\n" ireg64 r1 ireg64 rd
- | Pandl_ri(rd, n) ->
- fprintf oc " andl $%a, %a\n" coqint n ireg32 rd
- | Pandq_ri(rd, n) ->
- fprintf oc " andq %a, %a\n" intconst64 n ireg64 rd
- | Porl_rr(rd, r1) ->
- fprintf oc " orl %a, %a\n" ireg32 r1 ireg32 rd
- | Porq_rr(rd, r1) ->
- fprintf oc " orq %a, %a\n" ireg64 r1 ireg64 rd
- | Porl_ri(rd, n) ->
- fprintf oc " orl $%a, %a\n" coqint n ireg32 rd
- | Porq_ri(rd, n) ->
- fprintf oc " orq %a, %a\n" intconst64 n ireg64 rd
- | Pxorl_r(rd) ->
- fprintf oc " xorl %a, %a\n" ireg32 rd ireg32 rd
- | Pxorq_r(rd) ->
- fprintf oc " xorq %a, %a\n" ireg64 rd ireg64 rd
- | Pxorl_rr(rd, r1) ->
- fprintf oc " xorl %a, %a\n" ireg32 r1 ireg32 rd
- | Pxorq_rr(rd, r1) ->
- fprintf oc " xorq %a, %a\n" ireg64 r1 ireg64 rd
- | Pxorl_ri(rd, n) ->
- fprintf oc " xorl $%a, %a\n" coqint n ireg32 rd
- | Pxorq_ri(rd, n) ->
- fprintf oc " xorq %a, %a\n" intconst64 n ireg64 rd
- | Pnotl(rd) ->
- fprintf oc " notl %a\n" ireg32 rd
- | Pnotq(rd) ->
- fprintf oc " notq %a\n" ireg64 rd
- | Psall_rcl(rd) ->
- fprintf oc " sall %%cl, %a\n" ireg32 rd
- | Psalq_rcl(rd) ->
- fprintf oc " salq %%cl, %a\n" ireg64 rd
- | Psall_ri(rd, n) ->
- fprintf oc " sall $%a, %a\n" coqint n ireg32 rd
- | Psalq_ri(rd, n) ->
- fprintf oc " salq $%a, %a\n" coqint n ireg64 rd
- | Pshrl_rcl(rd) ->
- fprintf oc " shrl %%cl, %a\n" ireg32 rd
- | Pshrq_rcl(rd) ->
- fprintf oc " shrq %%cl, %a\n" ireg64 rd
- | Pshrl_ri(rd, n) ->
- fprintf oc " shrl $%a, %a\n" coqint n ireg32 rd
- | Pshrq_ri(rd, n) ->
- fprintf oc " shrq $%a, %a\n" coqint n ireg64 rd
- | Psarl_rcl(rd) ->
- fprintf oc " sarl %%cl, %a\n" ireg32 rd
- | Psarq_rcl(rd) ->
- fprintf oc " sarq %%cl, %a\n" ireg64 rd
- | Psarl_ri(rd, n) ->
- fprintf oc " sarl $%a, %a\n" coqint n ireg32 rd
- | Psarq_ri(rd, n) ->
- fprintf oc " sarq $%a, %a\n" coqint n ireg64 rd
- | Pshld_ri(rd, r1, n) ->
- fprintf oc " shldl $%a, %a, %a\n" coqint n ireg32 r1 ireg32 rd
- | Prorl_ri(rd, n) ->
- fprintf oc " rorl $%a, %a\n" coqint n ireg32 rd
- | Prorq_ri(rd, n) ->
- fprintf oc " rorq $%a, %a\n" coqint n ireg64 rd
- | Pcmpl_rr(r1, r2) ->
- fprintf oc " cmpl %a, %a\n" ireg32 r2 ireg32 r1
- | Pcmpq_rr(r1, r2) ->
- fprintf oc " cmpq %a, %a\n" ireg64 r2 ireg64 r1
- | Pcmpl_ri(r1, n) ->
- fprintf oc " cmpl $%a, %a\n" coqint n ireg32 r1
- | Pcmpq_ri(r1, n) ->
- fprintf oc " cmpq %a, %a\n" intconst64 n ireg64 r1
- | Ptestl_rr(r1, r2) ->
- fprintf oc " testl %a, %a\n" ireg32 r2 ireg32 r1
- | Ptestq_rr(r1, r2) ->
- fprintf oc " testq %a, %a\n" ireg64 r2 ireg64 r1
- | Ptestl_ri(r1, n) ->
- fprintf oc " testl $%a, %a\n" coqint n ireg32 r1
- | Ptestq_ri(r1, n) ->
- fprintf oc " testl %a, %a\n" intconst64 n ireg64 r1
- | Pcmov(c, rd, r1) ->
- fprintf oc " cmov%s %a, %a\n" (name_of_condition c) ireg r1 ireg rd
- | Psetcc(c, rd) ->
- fprintf oc " set%s %a\n" (name_of_condition c) ireg8 rd;
- fprintf oc " movzbl %a, %a\n" ireg8 rd ireg32 rd
- (* Arithmetic operations over floats *)
- | Paddd_ff(rd, r1) ->
- fprintf oc " addsd %a, %a\n" freg r1 freg rd
- | Psubd_ff(rd, r1) ->
- fprintf oc " subsd %a, %a\n" freg r1 freg rd
- | Pmuld_ff(rd, r1) ->
- fprintf oc " mulsd %a, %a\n" freg r1 freg rd
- | Pdivd_ff(rd, r1) ->
- fprintf oc " divsd %a, %a\n" freg r1 freg rd
- | Pnegd (rd) ->
- need_masks := true;
- fprintf oc " xorpd %a%s, %a\n"
- raw_symbol "__negd_mask" rip_rel freg rd
- | Pabsd (rd) ->
- need_masks := true;
- fprintf oc " andpd %a%s, %a\n"
- raw_symbol "__absd_mask" rip_rel freg rd
- | Pcomisd_ff(r1, r2) ->
- fprintf oc " comisd %a, %a\n" freg r2 freg r1
- | Pxorpd_f (rd) ->
- fprintf oc " xorpd %a, %a\n" freg rd freg rd
- | Padds_ff(rd, r1) ->
- fprintf oc " addss %a, %a\n" freg r1 freg rd
- | Psubs_ff(rd, r1) ->
- fprintf oc " subss %a, %a\n" freg r1 freg rd
- | Pmuls_ff(rd, r1) ->
- fprintf oc " mulss %a, %a\n" freg r1 freg rd
- | Pdivs_ff(rd, r1) ->
- fprintf oc " divss %a, %a\n" freg r1 freg rd
- | Pnegs (rd) ->
- need_masks := true;
- fprintf oc " xorpd %a%s, %a\n"
- raw_symbol "__negs_mask" rip_rel freg rd
- | Pabss (rd) ->
- need_masks := true;
- fprintf oc " andpd %a%s, %a\n"
- raw_symbol "__abss_mask" rip_rel freg rd
- | Pcomiss_ff(r1, r2) ->
- fprintf oc " comiss %a, %a\n" freg r2 freg r1
- | Pxorps_f (rd) ->
- fprintf oc " xorpd %a, %a\n" freg rd freg rd
- (* Branches and calls *)
- | Pjmp_l(l) ->
- fprintf oc " jmp %a\n" label (transl_label l)
- | Pjmp_s(f, sg) ->
- fprintf oc " jmp %a\n" symbol f
- | Pjmp_r(r, sg) ->
- fprintf oc " jmp *%a\n" ireg r
- | Pjcc(c, l) ->
- let l = transl_label l in
- fprintf oc " j%s %a\n" (name_of_condition c) label l
- | Pjcc2(c1, c2, l) ->
- let l = transl_label l in
- let l' = new_label() in
- fprintf oc " j%s %a\n" (name_of_neg_condition c1) label l';
- fprintf oc " j%s %a\n" (name_of_condition c2) label l;
- fprintf oc "%a:\n" label l'
- | Pjmptbl(r, tbl) ->
- let l = new_label() in
- jumptables := (l, tbl) :: !jumptables;
- if Archi.ptr64 then begin
- let (tmp1, tmp2) =
- if r = RAX then (RDX, RAX) else (RAX, RDX) in
- fprintf oc " leaq %a(%%rip), %a\n" label l ireg tmp1;
- fprintf oc " movslq (%a, %a, 4), %a\n" ireg tmp1 ireg r ireg tmp2;
- fprintf oc " addq %a, %a\n" ireg tmp2 ireg tmp1;
- fprintf oc " jmp *%a\n" ireg tmp1
- end else begin
- fprintf oc " jmp *%a(, %a, 4)\n" label l ireg r
- end
- | Pcall_s(f, sg) ->
- fprintf oc " call %a\n" symbol f;
- if (not Archi.ptr64) && sg.sig_cc.cc_structret then
- fprintf oc " pushl %%eax\n"
- | Pcall_r(r, sg) ->
- fprintf oc " call *%a\n" ireg r;
- if (not Archi.ptr64) && sg.sig_cc.cc_structret then
- fprintf oc " pushl %%eax\n"
- | Pret ->
- if (not Archi.ptr64)
- && (!current_function_sig).sig_cc.cc_structret then begin
- fprintf oc " movl 0(%%esp), %%eax\n";
- fprintf oc " ret $4\n"
- end else begin
- fprintf oc " ret\n"
- end
- (* Instructions produced by Asmexpand *)
- | Padcl_ri (res,n) ->
- fprintf oc " adcl $%ld, %a\n" (camlint_of_coqint n) ireg32 res;
- | Padcl_rr (res,a1) ->
- fprintf oc " adcl %a, %a\n" ireg32 a1 ireg32 res;
- | Paddl_rr (res,a1) ->
- fprintf oc " addl %a, %a\n" ireg32 a1 ireg32 res;
- | Paddl_mi (addr,n) ->
- fprintf oc " addl $%ld, %a\n" (camlint_of_coqint n) addressing addr
- | Pbsfl (res,a1) ->
- fprintf oc " bsfl %a, %a\n" ireg32 a1 ireg32 res
- | Pbsfq (res,a1) ->
- fprintf oc " bsfq %a, %a\n" ireg64 a1 ireg64 res
- | Pbsrl (res,a1) ->
- fprintf oc " bsrl %a, %a\n" ireg32 a1 ireg32 res
- | Pbsrq (res,a1) ->
- fprintf oc " bsrq %a, %a\n" ireg64 a1 ireg64 res
- | Pbswap64 res ->
- fprintf oc " bswap %a\n" ireg64 res
- | Pbswap32 res ->
- fprintf oc " bswap %a\n" ireg32 res
- | Pbswap16 res ->
- fprintf oc " rolw $8, %a\n" ireg16 res
- | Pcfi_adjust sz ->
- cfi_adjust oc (camlint_of_coqint sz)
- | Pfmadd132 (res,a1,a2) ->
- fprintf oc " vfmadd132sd %a, %a, %a\n" freg a1 freg a2 freg res
- | Pfmadd213 (res,a1,a2) ->
- fprintf oc " vfmadd213sd %a, %a, %a\n" freg a1 freg a2 freg res
- | Pfmadd231 (res,a1,a2) ->
- fprintf oc " vfmadd231sd %a, %a, %a\n" freg a1 freg a2 freg res
- | Pfmsub132 (res,a1,a2) ->
- fprintf oc " vfmsub132sd %a, %a, %a\n" freg a1 freg a2 freg res
- | Pfmsub213 (res,a1,a2) ->
- fprintf oc " vfmsub213sd %a, %a, %a\n" freg a1 freg a2 freg res
- | Pfmsub231 (res,a1,a2) ->
- fprintf oc " vfmsub231sd %a, %a, %a\n" freg a1 freg a2 freg res
- | Pfnmadd132 (res,a1,a2) ->
- fprintf oc " vfnmadd132sd %a, %a, %a\n" freg a1 freg a2 freg res
- | Pfnmadd213 (res,a1,a2) ->
- fprintf oc " vfnmadd213sd %a, %a, %a\n" freg a1 freg a2 freg res
- | Pfnmadd231 (res,a1,a2) ->
- fprintf oc " vfnmadd231sd %a, %a, %a\n" freg a1 freg a2 freg res
- | Pfnmsub132 (res,a1,a2) ->
- fprintf oc " vfnmsub132sd %a, %a, %a\n" freg a1 freg a2 freg res
- | Pfnmsub213 (res,a1,a2) ->
- fprintf oc " vfnmsub213sd %a, %a, %a\n" freg a1 freg a2 freg res
- | Pfnmsub231 (res,a1,a2) ->
- fprintf oc " vfnmsub231sd %a, %a, %a\n" freg a1 freg a2 freg res
- | Pmaxsd (res,a1) ->
- fprintf oc " maxsd %a, %a\n" freg a1 freg res
- | Pminsd (res,a1) ->
- fprintf oc " minsd %a, %a\n" freg a1 freg res
- | Pmovb_rm (rd,a) ->
- fprintf oc " movb %a, %a\n" addressing a ireg8 rd
- | Pmovsq_mr(a, rs) ->
- fprintf oc " movq %a, %a\n" freg rs addressing a
- | Pmovsq_rm(rd, a) ->
- fprintf oc " movq %a, %a\n" addressing a freg rd
- | Pmovsb ->
- fprintf oc " movsb\n";
- | Pmovsw ->
- fprintf oc " movsw\n";
- | Pmovw_rm (rd, a) ->
- fprintf oc " movw %a, %a\n" addressing a ireg16 rd
- | Prep_movsl ->
- fprintf oc " rep movsl\n"
- | Psbbl_rr (res,a1) ->
- fprintf oc " sbbl %a, %a\n" ireg32 a1 ireg32 res
- | Psqrtsd (res,a1) ->
- fprintf oc " sqrtsd %a, %a\n" freg a1 freg res
- | Psubl_ri (res,n) ->
- fprintf oc " subl $%ld, %a\n" (camlint_of_coqint n) ireg32 res;
- | Psubq_ri (res,n) ->
- fprintf oc " subq %a, %a\n" intconst64 n ireg64 res;
- (* Pseudo-instructions *)
- | Plabel(l) ->
- fprintf oc "%a:\n" label (transl_label l)
- | Pallocframe(sz, ofs_ra, ofs_link)
- | Pfreeframe(sz, ofs_ra, ofs_link) ->
- assert false
- | Pbuiltin(ef, args, res) ->
- begin match ef with
- | EF_annot(txt, targs) ->
- fprintf oc "%s annotation: " comment;
- print_annot_text preg "%esp" oc (camlstring_of_coqstring txt) args
- | EF_debug(kind, txt, targs) ->
- print_debug_info comment print_file_line preg "%esp" oc
- (P.to_int kind) (extern_atom txt) args
- | EF_inline_asm(txt, sg, clob) ->
- fprintf oc "%s begin inline assembly\n\t" comment;
- print_inline_asm preg oc (camlstring_of_coqstring txt) sg args res;
- fprintf oc "%s end inline assembly\n" comment
- | _ ->
- assert false
- end
-
- let print_literal64 oc (lbl, n) =
- fprintf oc "%a: .quad 0x%Lx\n" label lbl n
- let print_literal32 oc (lbl, n) =
- fprintf oc "%a: .long 0x%lx\n" label lbl n
-
- let print_jumptable oc jmptbl =
- let print_jumptable (lbl, tbl) =
- let print_entry l =
- if Archi.ptr64 then
- fprintf oc " .long %a - %a\n" label (transl_label l) label lbl
- else
- fprintf oc " .long %a\n" label (transl_label l)
- in
- fprintf oc "%a:" label lbl;
- List.iter print_entry tbl
- in
- if !jumptables <> [] then begin
- section oc jmptbl;
- print_align oc 4;
- List.iter print_jumptable !jumptables;
- jumptables := []
- end
-
- let print_init oc = function
- | Init_int8 n ->
- fprintf oc " .byte %ld\n" (camlint_of_coqint n)
- | Init_int16 n ->
- fprintf oc " .short %ld\n" (camlint_of_coqint n)
- | Init_int32 n ->
- fprintf oc " .long %ld\n" (camlint_of_coqint n)
- | Init_int64 n ->
- fprintf oc " .quad %Ld\n" (camlint64_of_coqint n)
- | Init_float32 n ->
- fprintf oc " .long %ld %s %.18g\n"
- (camlint_of_coqint (Floats.Float32.to_bits n))
- comment (camlfloat_of_coqfloat32 n)
- | Init_float64 n ->
- fprintf oc " .quad %Ld %s %.18g\n"
- (camlint64_of_coqint (Floats.Float.to_bits n))
- comment (camlfloat_of_coqfloat n)
- | Init_space n ->
- if Z.gt n Z.zero then
- fprintf oc " .space %a\n" z n
- | Init_addrof(symb, ofs) ->
- fprintf oc " %s %a\n" data_pointer symbol_offset (symb, ofs)
-
- let print_align = print_align
-
- let print_comm_symb oc sz name align =
- if C2C.atom_is_static name
- then System.print_lcomm_decl oc name sz align
- else System.print_comm_decl oc name sz align
-
- let name_of_section = name_of_section
-
- let emit_constants oc lit =
- if !float64_literals <> [] || !float32_literals <> [] then begin
- section oc lit;
- print_align oc 8;
- List.iter (print_literal64 oc) !float64_literals;
- float64_literals := [];
- List.iter (print_literal32 oc) !float32_literals;
- float32_literals := []
- end
-
- let cfi_startproc = cfi_startproc
- let cfi_endproc = cfi_endproc
-
- let print_instructions oc fn =
- current_function_sig := fn.fn_sig;
- List.iter (print_instruction oc) fn.fn_code
-
- let print_optional_fun_info _ = ()
-
- let get_section_names name =
- match C2C.atom_sections name with
- | [t;l;j] -> (t, l, j)
- | _ -> (Section_text, Section_literal, Section_jumptable)
-
- let reset_constants = reset_constants
-
- let print_fun_info = print_fun_info
-
- let print_var_info = print_var_info
-
- let print_prologue oc =
- need_masks := false;
- if !Clflags.option_g then begin
- section oc Section_text;
- if Configuration.system <> "bsd" then cfi_section oc
- end
-
- let print_epilogue oc =
- if !need_masks then begin
- section oc (Section_const true);
- (* not Section_literal because not 8-bytes *)
- print_align oc 16;
- fprintf oc "%a: .quad 0x8000000000000000, 0\n"
- raw_symbol "__negd_mask";
- fprintf oc "%a: .quad 0x7FFFFFFFFFFFFFFF, 0xFFFFFFFFFFFFFFFF\n"
- raw_symbol "__absd_mask";
- fprintf oc "%a: .long 0x80000000, 0, 0, 0\n"
- raw_symbol "__negs_mask";
- fprintf oc "%a: .long 0x7FFFFFFF, 0xFFFFFFFF, 0xFFFFFFFF, 0xFFFFFFFF\n"
- raw_symbol "__abss_mask"
- end;
- System.print_epilogue oc;
- if !Clflags.option_g then begin
- Debug.compute_gnu_file_enum (fun f -> ignore (print_file oc f));
- section oc Section_text;
- end
-
- let comment = comment
-
- let default_falignment = 16
-
- let label = label
-
- let new_label = new_label
-
-end
-
-let sel_target () =
- let module S = (val (match Configuration.system with
- | "macosx" -> (module MacOS_System:SYSTEM)
- | "linux"
- | "bsd" -> (module ELF_System:SYSTEM)
- | _ -> invalid_arg ("System " ^ Configuration.system ^ " not supported") ):SYSTEM) in
- (module Target(S):TARGET)
diff --git a/ia32/ValueAOp.v b/ia32/ValueAOp.v
deleted file mode 100644
index 98f0dbb1..00000000
--- a/ia32/ValueAOp.v
+++ /dev/null
@@ -1,266 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, 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. *)
-(* *)
-(* *********************************************************************)
-
-Require Import Coqlib Compopts.
-Require Import AST Integers Floats Values Memory Globalenvs.
-Require Import Op RTL ValueDomain.
-
-Local Transparent Archi.ptr64.
-
-(** Value analysis for x86_64 operators *)
-
-Definition eval_static_condition (cond: condition) (vl: list aval): abool :=
- match cond, vl with
- | Ccomp c, v1 :: v2 :: nil => cmp_bool c v1 v2
- | Ccompu c, v1 :: v2 :: nil => cmpu_bool c v1 v2
- | Ccompimm c n, v1 :: nil => cmp_bool c v1 (I n)
- | Ccompuimm c n, v1 :: nil => cmpu_bool c v1 (I n)
- | Ccompl c, v1 :: v2 :: nil => cmpl_bool c v1 v2
- | Ccomplu c, v1 :: v2 :: nil => cmplu_bool c v1 v2
- | Ccomplimm c n, v1 :: nil => cmpl_bool c v1 (L n)
- | Ccompluimm c n, v1 :: nil => cmplu_bool c v1 (L n)
- | Ccompf c, v1 :: v2 :: nil => cmpf_bool c v1 v2
- | Cnotcompf c, v1 :: v2 :: nil => cnot (cmpf_bool c v1 v2)
- | Ccompfs c, v1 :: v2 :: nil => cmpfs_bool c v1 v2
- | Cnotcompfs c, v1 :: v2 :: nil => cnot (cmpfs_bool c v1 v2)
- | Cmaskzero n, v1 :: nil => maskzero v1 n
- | Cmasknotzero n, v1 :: nil => cnot (maskzero v1 n)
- | _, _ => Bnone
- end.
-
-Definition eval_static_addressing_32 (addr: addressing) (vl: list aval): aval :=
- match addr, vl with
- | Aindexed n, v1::nil => add v1 (I (Int.repr n))
- | Aindexed2 n, v1::v2::nil => add (add v1 v2) (I (Int.repr n))
- | Ascaled sc ofs, v1::nil => add (mul v1 (I (Int.repr sc))) (I (Int.repr ofs))
- | Aindexed2scaled sc ofs, v1::v2::nil => add v1 (add (mul v2 (I (Int.repr sc))) (I (Int.repr ofs)))
- | Aglobal s ofs, nil => Ptr (Gl s ofs)
- | Abased s ofs, v1::nil => add (Ptr (Gl s ofs)) v1
- | Abasedscaled sc s ofs, v1::nil => add (Ptr (Gl s ofs)) (mul v1 (I (Int.repr sc)))
- | Ainstack ofs, nil => Ptr(Stk ofs)
- | _, _ => Vbot
- end.
-
-Definition eval_static_addressing_64 (addr: addressing) (vl: list aval): aval :=
- match addr, vl with
- | Aindexed n, v1::nil => addl v1 (L (Int64.repr n))
- | Aindexed2 n, v1::v2::nil => addl (addl v1 v2) (L (Int64.repr n))
- | Ascaled sc ofs, v1::nil => addl (mull v1 (L (Int64.repr sc))) (L (Int64.repr ofs))
- | Aindexed2scaled sc ofs, v1::v2::nil => addl v1 (addl (mull v2 (L (Int64.repr sc))) (L (Int64.repr ofs)))
- | Aglobal s ofs, nil => Ptr (Gl s ofs)
- | Abased s ofs, v1::nil => addl (Ptr (Gl s ofs)) v1
- | Abasedscaled sc s ofs, v1::nil => addl (Ptr (Gl s ofs)) (mull v1 (L (Int64.repr sc)))
- | Ainstack ofs, nil => Ptr(Stk ofs)
- | _, _ => Vbot
- end.
-
-Definition eval_static_addressing (addr: addressing) (vl: list aval): aval :=
- if Archi.ptr64
- then eval_static_addressing_64 addr vl
- else eval_static_addressing_32 addr vl.
-
-Definition eval_static_operation (op: operation) (vl: list aval): aval :=
- match op, vl with
- | Omove, v1::nil => v1
- | Ointconst n, nil => I n
- | Olongconst n, nil => L n
- | Ofloatconst n, nil => if propagate_float_constants tt then F n else ntop
- | Osingleconst n, nil => if propagate_float_constants tt then FS n else ntop
- | Oindirectsymbol id, nil => Ifptr (Gl id Ptrofs.zero)
- | Ocast8signed, v1 :: nil => sign_ext 8 v1
- | Ocast8unsigned, v1 :: nil => zero_ext 8 v1
- | Ocast16signed, v1 :: nil => sign_ext 16 v1
- | Ocast16unsigned, v1 :: nil => zero_ext 16 v1
- | Oneg, v1::nil => neg v1
- | Osub, v1::v2::nil => sub v1 v2
- | Omul, v1::v2::nil => mul v1 v2
- | Omulimm n, v1::nil => mul v1 (I n)
- | Omulhs, v1::v2::nil => mulhs v1 v2
- | Omulhu, v1::v2::nil => mulhu v1 v2
- | Odiv, v1::v2::nil => divs v1 v2
- | Odivu, v1::v2::nil => divu v1 v2
- | Omod, v1::v2::nil => mods v1 v2
- | Omodu, v1::v2::nil => modu v1 v2
- | Oand, v1::v2::nil => and v1 v2
- | Oandimm n, v1::nil => and v1 (I n)
- | Oor, v1::v2::nil => or v1 v2
- | Oorimm n, v1::nil => or v1 (I n)
- | Oxor, v1::v2::nil => xor v1 v2
- | Oxorimm n, v1::nil => xor v1 (I n)
- | Onot, v1::nil => notint v1
- | Oshl, v1::v2::nil => shl v1 v2
- | Oshlimm n, v1::nil => shl v1 (I n)
- | Oshr, v1::v2::nil => shr v1 v2
- | Oshrimm n, v1::nil => shr v1 (I n)
- | Oshrximm n, v1::nil => shrx v1 (I n)
- | Oshru, v1::v2::nil => shru v1 v2
- | Oshruimm n, v1::nil => shru v1 (I n)
- | Ororimm n, v1::nil => ror v1 (I n)
- | Oshldimm n, v1::v2::nil => or (shl v1 (I n)) (shru v2 (I (Int.sub Int.iwordsize n)))
- | Olea addr, _ => eval_static_addressing_32 addr vl
- | Omakelong, v1::v2::nil => longofwords v1 v2
- | Olowlong, v1::nil => loword v1
- | Ohighlong, v1::nil => hiword v1
- | Ocast32signed, v1::nil => longofint v1
- | Ocast32unsigned, v1::nil => longofintu v1
- | Onegl, v1::nil => negl v1
- | Oaddlimm n, v1::nil => addl v1 (L n)
- | Osubl, v1::v2::nil => subl v1 v2
- | Omull, v1::v2::nil => mull v1 v2
- | Omullimm n, v1::nil => mull v1 (L n)
- | Omullhs, v1::v2::nil => mullhs v1 v2
- | Omullhu, v1::v2::nil => mullhu v1 v2
- | Odivl, v1::v2::nil => divls v1 v2
- | Odivlu, v1::v2::nil => divlu v1 v2
- | Omodl, v1::v2::nil => modls v1 v2
- | Omodlu, v1::v2::nil => modlu v1 v2
- | Oandl, v1::v2::nil => andl v1 v2
- | Oandlimm n, v1::nil => andl v1 (L n)
- | Oorl, v1::v2::nil => orl v1 v2
- | Oorlimm n, v1::nil => orl v1 (L n)
- | Oxorl, v1::v2::nil => xorl v1 v2
- | Oxorlimm n, v1::nil => xorl v1 (L n)
- | Onotl, v1::nil => notl v1
- | Oshll, v1::v2::nil => shll v1 v2
- | Oshllimm n, v1::nil => shll v1 (I n)
- | Oshrl, v1::v2::nil => shrl v1 v2
- | Oshrlimm n, v1::nil => shrl v1 (I n)
- | Oshrxlimm n, v1::nil => shrxl v1 (I n)
- | Oshrlu, v1::v2::nil => shrlu v1 v2
- | Oshrluimm n, v1::nil => shrlu v1 (I n)
- | Ororlimm n, v1::nil => rorl v1 (I n)
- | Oleal addr, _ => eval_static_addressing_64 addr vl
- | Onegf, v1::nil => negf v1
- | Oabsf, v1::nil => absf v1
- | Oaddf, v1::v2::nil => addf v1 v2
- | Osubf, v1::v2::nil => subf v1 v2
- | Omulf, v1::v2::nil => mulf v1 v2
- | Odivf, v1::v2::nil => divf v1 v2
- | Onegfs, v1::nil => negfs v1
- | Oabsfs, v1::nil => absfs v1
- | Oaddfs, v1::v2::nil => addfs v1 v2
- | Osubfs, v1::v2::nil => subfs v1 v2
- | Omulfs, v1::v2::nil => mulfs v1 v2
- | Odivfs, v1::v2::nil => divfs v1 v2
- | Osingleoffloat, v1::nil => singleoffloat v1
- | Ofloatofsingle, v1::nil => floatofsingle v1
- | Ointoffloat, v1::nil => intoffloat v1
- | Ofloatofint, v1::nil => floatofint v1
- | Ointofsingle, v1::nil => intofsingle v1
- | Osingleofint, v1::nil => singleofint v1
- | Olongoffloat, v1::nil => longoffloat v1
- | Ofloatoflong, v1::nil => floatoflong v1
- | Olongofsingle, v1::nil => longofsingle v1
- | Osingleoflong, v1::nil => singleoflong v1
- | Ocmp c, _ => of_optbool (eval_static_condition c vl)
- | _, _ => Vbot
- end.
-
-Section SOUNDNESS.
-
-Variable bc: block_classification.
-Variable ge: genv.
-Hypothesis GENV: genv_match bc ge.
-Variable sp: block.
-Hypothesis STACK: bc sp = BCstack.
-
-Theorem eval_static_condition_sound:
- forall cond vargs m aargs,
- list_forall2 (vmatch bc) vargs aargs ->
- cmatch (eval_condition cond vargs m) (eval_static_condition cond aargs).
-Proof.
- intros until aargs; intros VM. inv VM.
- destruct cond; auto with va.
- inv H0.
- destruct cond; simpl; eauto with va.
- inv H2.
- destruct cond; simpl; eauto with va.
- destruct cond; auto with va.
-Qed.
-
-Lemma symbol_address_sound:
- forall id ofs,
- vmatch bc (Genv.symbol_address ge id ofs) (Ptr (Gl id ofs)).
-Proof.
- intros; apply symbol_address_sound; apply GENV.
-Qed.
-
-Lemma symbol_address_sound_2:
- forall id ofs,
- vmatch bc (Genv.symbol_address ge id ofs) (Ifptr (Gl id ofs)).
-Proof.
- intros. unfold Genv.symbol_address. destruct (Genv.find_symbol ge id) as [b|] eqn:F.
- constructor. constructor. apply GENV; auto.
- constructor.
-Qed.
-
-Hint Resolve symbol_address_sound symbol_address_sound_2: va.
-
-Ltac InvHyps :=
- match goal with
- | [H: None = Some _ |- _ ] => discriminate
- | [H: Some _ = Some _ |- _] => inv H
- | [H1: match ?vl with nil => _ | _ :: _ => _ end = Some _ ,
- H2: list_forall2 _ ?vl _ |- _ ] => inv H2; InvHyps
- | [H: (if Archi.ptr64 then _ else _) = Some _ |- _] => destruct Archi.ptr64 eqn:?; InvHyps
- | _ => idtac
- end.
-
-Theorem eval_static_addressing_32_sound:
- forall addr vargs vres aargs,
- eval_addressing32 ge (Vptr sp Ptrofs.zero) addr vargs = Some vres ->
- list_forall2 (vmatch bc) vargs aargs ->
- vmatch bc vres (eval_static_addressing_32 addr aargs).
-Proof.
- unfold eval_addressing32, eval_static_addressing_32; intros;
- destruct addr; InvHyps; eauto with va.
- rewrite Ptrofs.add_zero_l; eauto with va.
-Qed.
-
-Theorem eval_static_addressing_64_sound:
- forall addr vargs vres aargs,
- eval_addressing64 ge (Vptr sp Ptrofs.zero) addr vargs = Some vres ->
- list_forall2 (vmatch bc) vargs aargs ->
- vmatch bc vres (eval_static_addressing_64 addr aargs).
-Proof.
- unfold eval_addressing64, eval_static_addressing_64; intros;
- destruct addr; InvHyps; eauto with va.
- rewrite Ptrofs.add_zero_l; eauto with va.
-Qed.
-
-Theorem eval_static_addressing_sound:
- forall addr vargs vres aargs,
- eval_addressing ge (Vptr sp Ptrofs.zero) addr vargs = Some vres ->
- list_forall2 (vmatch bc) vargs aargs ->
- vmatch bc vres (eval_static_addressing addr aargs).
-Proof.
- unfold eval_addressing, eval_static_addressing; intros.
- destruct Archi.ptr64; eauto using eval_static_addressing_32_sound, eval_static_addressing_64_sound.
-Qed.
-
-Theorem eval_static_operation_sound:
- forall op vargs m vres aargs,
- eval_operation ge (Vptr sp Ptrofs.zero) op vargs m = Some vres ->
- list_forall2 (vmatch bc) vargs aargs ->
- vmatch bc vres (eval_static_operation op aargs).
-Proof.
- unfold eval_operation, eval_static_operation; intros;
- destruct op; InvHyps; eauto with va.
- destruct (propagate_float_constants tt); constructor.
- destruct (propagate_float_constants tt); constructor.
- eapply eval_static_addressing_32_sound; eauto.
- eapply eval_static_addressing_64_sound; eauto.
- apply of_optbool_sound. eapply eval_static_condition_sound; eauto.
-Qed.
-
-End SOUNDNESS.
-
diff --git a/ia32/extractionMachdep.v b/ia32/extractionMachdep.v
deleted file mode 100644
index 8b395579..00000000
--- a/ia32/extractionMachdep.v
+++ /dev/null
@@ -1,31 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(* Additional extraction directives specific to the x86-64 port *)
-
-Require Archi SelectOp ConstpropOp.
-
-(* Archi *)
-
-Extract Constant Archi.ptr64 =>
- "Configuration.model = ""64"" ".
-
-(* SelectOp *)
-
-Extract Constant SelectOp.symbol_is_external =>
- "fun id -> Configuration.system = ""macosx"" && C2C.atom_is_extern id".
-
-(* ConstpropOp *)
-
-Extract Constant ConstpropOp.symbol_is_external =>
- "fun id -> Configuration.system = ""macosx"" && C2C.atom_is_extern id".
-