aboutsummaryrefslogtreecommitdiffstats
path: root/ia32
diff options
context:
space:
mode:
Diffstat (limited to 'ia32')
-rw-r--r--ia32/Archi.v45
-rw-r--r--ia32/Asm.v1017
-rw-r--r--ia32/AsmToJSON.ml17
-rw-r--r--ia32/AsmToJSON.mli13
-rw-r--r--ia32/Asmexpand.ml462
-rw-r--r--ia32/Asmgen.v591
-rw-r--r--ia32/Asmgenproof.v897
-rw-r--r--ia32/Asmgenproof1.v1195
-rw-r--r--ia32/CBuiltins.ml87
-rw-r--r--ia32/CombineOp.v117
-rw-r--r--ia32/CombineOpproof.v153
-rw-r--r--ia32/ConstpropOp.vp227
-rw-r--r--ia32/ConstpropOpproof.v543
-rw-r--r--ia32/Conventions1.v240
-rw-r--r--ia32/Machregs.v297
-rw-r--r--ia32/Machregsaux.ml33
-rw-r--r--ia32/Machregsaux.mli18
-rw-r--r--ia32/NeedOp.v194
-rw-r--r--ia32/Op.v1075
-rw-r--r--ia32/PrintOp.ml128
-rw-r--r--ia32/SelectOp.vp523
-rw-r--r--ia32/SelectOpproof.v917
-rw-r--r--ia32/Stacklayout.v142
-rw-r--r--ia32/TargetPrinter.ml765
-rw-r--r--ia32/ValueAOp.v194
-rw-r--r--ia32/extractionMachdep.v20
26 files changed, 0 insertions, 9910 deletions
diff --git a/ia32/Archi.v b/ia32/Archi.v
deleted file mode 100644
index ded460d2..00000000
--- a/ia32/Archi.v
+++ /dev/null
@@ -1,45 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* Jacques-Henri Jourdan, 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Architecture-dependent parameters for IA32 *)
-
-Require Import ZArith.
-Require Import Fappli_IEEE.
-Require Import Fappli_IEEE_bits.
-
-Definition big_endian := false.
-
-Notation align_int64 := 4%Z (only parsing).
-Notation align_float64 := 4%Z (only parsing).
-
-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 big_endian
- 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 b4fc950b..00000000
--- a/ia32/Asm.v
+++ /dev/null
@@ -1,1017 +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. *)
-(* *)
-(* *********************************************************************)
-
-(** Abstract syntax and semantics for IA32 assembly language *)
-
-Require Import Coqlib.
-Require Import Maps.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import Events.
-Require Import Globalenvs.
-Require Import Smallstep.
-Require Import Locations.
-Require Import Stacklayout.
-Require Import Conventions.
-
-(** * Abstract syntax *)
-
-(** ** Registers. *)
-
-(** Integer registers. *)
-
-Inductive ireg: Type :=
- | EAX: ireg | EBX: ireg | ECX: ireg | EDX: ireg
- | ESI: ireg | EDI: ireg | EBP: ireg | ESP: ireg.
-
-(** Floating-point registers, i.e. SSE2 registers *)
-
-Inductive freg: Type :=
- | XMM0: freg | XMM1: freg | XMM2: freg | XMM3: freg
- | XMM4: freg | XMM5: freg | XMM6: freg | XMM7: freg.
-
-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 := ESP (only parsing).
-
-(** ** Instruction set. *)
-
-Definition label := positive.
-
-(** General form of an addressing mode. *)
-
-Inductive addrmode: Type :=
- | Addrmode (base: option ireg)
- (ofs: option (ireg * int))
- (const: int + ident * int).
-
-(** 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:
-- [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] (32-bit int) *)
- | Pmov_ri (rd: ireg) (n: int)
- | Pmov_ra (rd: ireg) (id: ident)
- | Pmov_rm (rd: ireg) (a: addrmode)
- | Pmov_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 *)
- | Pxchg_rr (r1: ireg) (r2: ireg) (**r register-register exchange *)
- (** 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)
- | 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 *)
- (** Integer arithmetic *)
- | Plea (rd: ireg) (a: addrmode)
- | Pneg (rd: ireg)
- | Psub_rr (rd: ireg) (r1: ireg)
- | Pimul_rr (rd: ireg) (r1: ireg)
- | Pimul_ri (rd: ireg) (n: int)
- | Pimul_r (r1: ireg)
- | Pmul_r (r1: ireg)
- | Pcltd
- | Pdiv (r1: ireg)
- | Pidiv (r1: ireg)
- | Pand_rr (rd: ireg) (r1: ireg)
- | Pand_ri (rd: ireg) (n: int)
- | Por_rr (rd: ireg) (r1: ireg)
- | Por_ri (rd: ireg) (n: int)
- | Pxor_r (rd: ireg) (**r [xor] with self = set to zero *)
- | Pxor_rr (rd: ireg) (r1: ireg)
- | Pxor_ri (rd: ireg) (n: int)
- | Pnot (rd: ireg)
- | Psal_rcl (rd: ireg)
- | Psal_ri (rd: ireg) (n: int)
- | Pshr_rcl (rd: ireg)
- | Pshr_ri (rd: ireg) (n: int)
- | Psar_rcl (rd: ireg)
- | Psar_ri (rd: ireg) (n: int)
- | Pshld_ri (rd: ireg) (r1: ireg) (n: int)
- | Pror_ri (rd: ireg) (n: int)
- | Pcmp_rr (r1 r2: ireg)
- | Pcmp_ri (r1: ireg) (n: int)
- | Ptest_rr (r1 r2: ireg)
- | Ptest_ri (r1: ireg) (n: int)
- | 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 [Many32] chunk *)
- | Pmov_mr_a (a: addrmode) (rs: ireg) (**r like [Pmov_mr], using [Many32] 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: int)
- | Pfreeframe(sz: Z)(ofs_ra ofs_link: int)
- | Pbuiltin(ef: external_function)(args: list (builtin_arg preg))(res: builtin_res preg)
- (** Instructions not generated by [Asmgen] *)
- | Padc_ri (rd: ireg) (n: int)
- | Padc_rr (rd: ireg) (r2: ireg)
- | Padd_mi (a: addrmode) (n: int)
- | Padd_ri (rd: ireg) (n: int)
- | Padd_rr (rd: ireg) (r2: ireg)
- | Pbsf (rd: ireg) (r1: ireg)
- | Pbsr (rd: ireg) (r1: ireg)
- | Pbswap (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)
- | Pmovq_mr (a: addrmode) (rs: freg)
- | Pmovq_rm (rd: freg) (a: addrmode)
- | Pmovsb
- | Pmovsw
- | Pmovw_rm (rd: ireg) (ad: addrmode)
- | Prep_movsl
- | Psbb_rr (rd: ireg) (r2: ireg)
- | Psqrtsd (rd: freg) (r1: freg)
- | Psub_ri (rd: ireg) (n: int).
-
-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_addrmode (a: addrmode) (rs: regset) : val :=
- match a with Addrmode base ofs const =>
- Val.add (match base with
- | None => Vzero
- | Some r => rs r
- end)
- (Val.add (match ofs with
- | None => Vzero
- | Some(r, sc) =>
- if Int.eq sc Int.one then rs r else Val.mul (rs r) (Vint sc)
- end)
- (match const with
- | inl ofs => Vint ofs
- | inr(id, ofs) => Genv.symbol_address ge id ofs
- end))
- end.
-
-(** 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.
-
-(** 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.add rs#PC Vone).
-
-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 (Int.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
- | Pmov_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Vint n))) m
- | Pmov_ra rd id =>
- Next (nextinstr_nf (rs#rd <- (Genv.symbol_address ge id Int.zero))) m
- | Pmov_rm rd a =>
- exec_load Mint32 m a rs rd
- | Pmov_mr a r1 =>
- exec_store Mint32 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)
- | Pxchg_rr r1 r2 =>
- Next (nextinstr (rs#r1 <- (rs r2) #r2 <- (rs r1))) m
- (** 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
- | 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
- (** Integer arithmetic *)
- | Plea rd a =>
- Next (nextinstr (rs#rd <- (eval_addrmode a rs))) m
- | Pneg rd =>
- Next (nextinstr_nf (rs#rd <- (Val.neg rs#rd))) m
- | Psub_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.sub rs#rd rs#r1))) m
- | Pimul_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.mul rs#rd rs#r1))) m
- | Pimul_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.mul rs#rd (Vint n)))) m
- | Pimul_r r1 =>
- Next (nextinstr_nf (rs#EAX <- (Val.mul rs#EAX rs#r1)
- #EDX <- (Val.mulhs rs#EAX rs#r1))) m
- | Pmul_r r1 =>
- Next (nextinstr_nf (rs#EAX <- (Val.mul rs#EAX rs#r1)
- #EDX <- (Val.mulhu rs#EAX rs#r1))) m
- | Pcltd =>
- Next (nextinstr_nf (rs#EDX <- (Val.shr rs#EAX (Vint (Int.repr 31))))) m
- | Pdiv r1 =>
- match rs#EDX, rs#EAX, rs#r1 with
- | Vint nh, Vint nl, Vint d =>
- match Int.divmodu2 nh nl d with
- | Some(q, r) => Next (nextinstr_nf (rs#EAX <- (Vint q) #EDX <- (Vint r))) m
- | None => Stuck
- end
- | _, _, _ => Stuck
- end
- | Pidiv r1 =>
- match rs#EDX, rs#EAX, rs#r1 with
- | Vint nh, Vint nl, Vint d =>
- match Int.divmods2 nh nl d with
- | Some(q, r) => Next (nextinstr_nf (rs#EAX <- (Vint q) #EDX <- (Vint r))) m
- | None => Stuck
- end
- | _, _, _ => Stuck
- end
- | Pand_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.and rs#rd rs#r1))) m
- | Pand_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.and rs#rd (Vint n)))) m
- | Por_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.or rs#rd rs#r1))) m
- | Por_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.or rs#rd (Vint n)))) m
- | Pxor_r rd =>
- Next (nextinstr_nf (rs#rd <- Vzero)) m
- | Pxor_rr rd r1 =>
- Next (nextinstr_nf (rs#rd <- (Val.xor rs#rd rs#r1))) m
- | Pxor_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.xor rs#rd (Vint n)))) m
- | Pnot rd =>
- Next (nextinstr_nf (rs#rd <- (Val.notint rs#rd))) m
- | Psal_rcl rd =>
- Next (nextinstr_nf (rs#rd <- (Val.shl rs#rd rs#ECX))) m
- | Psal_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.shl rs#rd (Vint n)))) m
- | Pshr_rcl rd =>
- Next (nextinstr_nf (rs#rd <- (Val.shru rs#rd rs#ECX))) m
- | Pshr_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.shru rs#rd (Vint n)))) m
- | Psar_rcl rd =>
- Next (nextinstr_nf (rs#rd <- (Val.shr rs#rd rs#ECX))) m
- | Psar_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.shr 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
- | Pror_ri rd n =>
- Next (nextinstr_nf (rs#rd <- (Val.ror rs#rd (Vint n)))) m
- | Pcmp_rr r1 r2 =>
- Next (nextinstr (compare_ints (rs r1) (rs r2) rs m)) m
- | Pcmp_ri r1 n =>
- Next (nextinstr (compare_ints (rs r1) (Vint n) rs m)) m
- | Ptest_rr r1 r2 =>
- Next (nextinstr (compare_ints (Val.and (rs r1) (rs r2)) Vzero rs m)) m
- | Ptest_ri r1 n =>
- Next (nextinstr (compare_ints (Val.and (rs r1) (Vint n)) Vzero 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 Int.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 m
- end
- | _ => Stuck
- end
- | Pcall_s id sg =>
- Next (rs#RA <- (Val.add rs#PC Vone) #PC <- (Genv.symbol_address ge id Int.zero)) m
- | Pcall_r r sg =>
- Next (rs#RA <- (Val.add rs#PC Vone) #PC <- (rs r)) m
- | Pret =>
- Next (rs#PC <- (rs#RA)) m
- (** Saving and restoring registers *)
- | Pmov_rm_a rd a =>
- exec_load Many32 m a rs rd
- | Pmov_mr_a a r1 =>
- exec_store 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 Int.zero in
- match Mem.storev Mint32 m1 (Val.add sp (Vint ofs_link)) rs#ESP with
- | None => Stuck
- | Some m2 =>
- match Mem.storev Mint32 m2 (Val.add sp (Vint ofs_ra)) rs#RA with
- | None => Stuck
- | Some m3 => Next (nextinstr (rs #EDX <- (rs#ESP) #ESP <- sp)) m3
- end
- end
- | Pfreeframe sz ofs_ra ofs_link =>
- match Mem.loadv Mint32 m (Val.add rs#ESP (Vint ofs_ra)) with
- | None => Stuck
- | Some ra =>
- match Mem.loadv Mint32 m (Val.add rs#ESP (Vint ofs_link)) with
- | None => Stuck
- | Some sp =>
- match rs#ESP with
- | Vptr stk ofs =>
- match Mem.free m stk 0 sz with
- | None => Stuck
- | Some m' => Next (nextinstr (rs#ESP <- 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. *)
- | Padc_ri _ _
- | Padc_rr _ _
- | Padd_mi _ _
- | Padd_ri _ _
- | Padd_rr _ _
- | Pbsf _ _
- | Pbsr _ _
- | Pbswap _
- | Pbswap16 _
- | Pcfi_adjust _
- | Pfmadd132 _ _ _
- | Pfmadd213 _ _ _
- | Pfmadd231 _ _ _
- | Pfmsub132 _ _ _
- | Pfmsub213 _ _ _
- | Pfmsub231 _ _ _
- | Pfnmadd132 _ _ _
- | Pfnmadd213 _ _ _
- | Pfnmadd231 _ _ _
- | Pfnmsub132 _ _ _
- | Pfnmsub213 _ _ _
- | Pfnmsub231 _ _ _
- | Pmaxsd _ _
- | Pminsd _ _
- | Pmovb_rm _ _
- | Pmovq_rm _ _
- | Pmovq_mr _ _
- | Pmovsb
- | Pmovsw
- | Pmovw_rm _ _
- | Prep_movsl
- | Psbb_rr _ _
- | Psqrtsd _ _
- | Psub_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 EAX
- | BX => IR EBX
- | CX => IR ECX
- | DX => IR EDX
- | SI => IR ESI
- | DI => IR EDI
- | BP => IR EBP
- | X0 => FR XMM0
- | X1 => FR XMM1
- | X2 => FR XMM2
- | X3 => FR XMM3
- | X4 => FR XMM4
- | X5 => FR XMM5
- | X6 => FR XMM6
- | X7 => FR XMM7
- | 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.add (rs (IR ESP)) (Vint (Int.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 (Int.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 (Int.unsigned ofs) f.(fn_code) = Some (Pbuiltin ef args res) ->
- eval_builtin_args ge rs (rs ESP) 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 Int.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) Int.zero)
- # RA <- Vzero
- # ESP <- Vzero in
- initial_state p (State rs0 m0).
-
-Inductive final_state: state -> int -> Prop :=
- | final_state_intro: forall rs m r,
- rs#PC = Vzero ->
- rs#EAX = 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 *)
- inv H. unfold Vzero in H0. red; intros; red; intros. inv H; congruence.
-- (* 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 6a64221e..00000000
--- a/ia32/Asmexpand.ml
+++ /dev/null
@@ -1,462 +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
-open Integers
-
-exception Error of string
-
-(* Useful constants and helper functions *)
-
-let _0 = Int.zero
-let _1 = Int.one
-let _2 = coqint_of_camlint 2l
-let _4 = coqint_of_camlint 4l
-let _8 = coqint_of_camlint 8l
-
-let stack_alignment () =
- if Configuration.system = "macosx" then 16
- else 8
-
-(* SP adjustment to allocate or free a stack frame *)
-
-let int32_align n a =
- if n >= 0l
- then Int32.logand (Int32.add n (Int32.of_int (a-1))) (Int32.of_int (-a))
- else Int32.logand n (Int32.of_int (-a))
-
-let sp_adjustment sz =
- let sz = camlint_of_coqint sz in
- (* Preserve proper alignment of the stack *)
- let sz = int32_align sz (stack_alignment ()) in
- (* The top 4 bytes have already been allocated by the "call" instruction. *)
- let sz = Int32.sub sz 4l in
- sz
-
-
-(* 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")
-
-(* Translate a builtin argument into an addressing mode *)
-
-let addressing_of_builtin_arg = function
- | BA (IR r) -> Addrmode(Some r, None, Coq_inl Integers.Int.zero)
- | BA_addrstack ofs -> Addrmode(Some ESP, None, Coq_inl ofs)
- | BA_addrglobal(id, ofs) -> Addrmode(None, None, Coq_inr(id, ofs))
- | _ -> assert false
-
-(* Operations on addressing modes *)
-
-let offset_addressing (Addrmode(base, ofs, cst)) delta =
- Addrmode(base, ofs,
- match cst with
- | Coq_inl n -> Coq_inl(Int.add n delta)
- | Coq_inr(id, n) -> Coq_inr(id, Int.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))
-
-(* 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 && !Clflags.option_ffpu then begin
- emit (Pmovq_rm (XMM7, src));
- emit (Pmovq_mr (dst, XMM7));
- copy (offset_addressing src _8) (offset_addressing dst _8) (sz - 8)
- end else if sz >= 4 then begin
- emit (Pmov_rm (ECX, src));
- emit (Pmov_mr (dst, ECX));
- copy (offset_addressing src _4) (offset_addressing dst _4) (sz - 4)
- end else if sz >= 2 then begin
- emit (Pmovw_rm (ECX, src));
- emit (Pmovw_mr (dst, ECX));
- copy (offset_addressing src _2) (offset_addressing dst _2) (sz - 2)
- end else if sz >= 1 then begin
- emit (Pmovb_rm (ECX, src));
- emit (Pmovb_mr (dst, ECX));
- copy (offset_addressing src _1) (offset_addressing dst _1) (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 ESI) then emit (Plea (ESI, addressing_of_builtin_arg src));
- if dst <> BA (IR EDI) then emit (Plea (EDI, addressing_of_builtin_arg dst));
- emit (Pmov_ri (ECX,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 (Pmov_rm (res,addr))
- | Mint64, BR_splitlong(BR(IR res1), BR(IR res2)) ->
- let addr' = offset_addressing addr _4 in
- if not (Asmgen.addressing_mentions addr res2) then begin
- emit (Pmov_rm (res2,addr));
- emit (Pmov_rm (res1,addr'))
- end else begin
- emit (Pmov_rm (res1,addr'));
- emit (Pmov_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 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 (Pmov_mr (addr,src))
- | Mint64, BA_splitlong(BA(IR src1), BA(IR src2)) ->
- let addr' = offset_addressing addr _4 in
- emit (Pmov_mr (addr,src2));
- emit (Pmov_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 EAX then ECX else EAX)
- | _ -> assert false
-
-(* Handling of varargs *)
-
-let expand_builtin_va_start r =
- if not (is_current_function_variadic ()) then
- invalid_arg "Fatal error: va_start used in non-vararg function";
- let ofs = coqint_of_camlint
- Int32.(add (add !PrintAsmaux.current_function_stacksize 4l)
- (mul 4l (Z.to_int32 (Conventions1.size_arguments
- (get_current_function_sig ()))))) in
- emit (Pmov_mr (linear_addr r _0, ESP));
- emit (Padd_mi (linear_addr r _0, ofs))
-
-(* 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 (Pbswap res)
- | "__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 (Pbsr (res,a1));
- emit (Pxor_ri(res,coqint_of_camlint 31l))
- | "__builtin_clzll", [BA_splitlong(BA (IR ah), BA (IR al))], BR(IR res) ->
- let lbl1 = new_label() in
- let lbl2 = new_label() in
- emit (Ptest_rr(ah, ah));
- emit (Pjcc(Cond_e, lbl1));
- emit (Pbsr(res, ah));
- emit (Pxor_ri(res, coqint_of_camlint 31l));
- emit (Pjmp_l lbl2);
- emit (Plabel lbl1);
- emit (Pbsr(res, al));
- emit (Pxor_ri(res, coqint_of_camlint 63l));
- emit (Plabel lbl2)
- | ("__builtin_ctz" | "__builtin_ctzl"), [BA(IR a1)], BR(IR res) ->
- emit (Pbsf (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 (Ptest_rr(al, al));
- emit (Pjcc(Cond_e, lbl1));
- emit (Pbsf(res, al));
- emit (Pjmp_l lbl2);
- emit (Plabel lbl1);
- emit (Pbsf(res, ah));
- emit (Padd_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 = EDX && al = EAX && rh = EDX && rl = EAX);
- emit (Pneg EAX);
- emit (Padc_ri (EDX,_0));
- emit (Pneg EDX)
- | "__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 = EDX && al = EAX && bh = ECX && bl = EBX && rh = EDX && rl = EAX);
- emit (Padd_rr (EAX,EBX));
- emit (Padc_rr (EDX,ECX))
- | "__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 = EDX && al = EAX && bh = ECX && bl = EBX && rh = EDX && rl = EAX);
- emit (Psub_rr (EAX,EBX));
- emit (Psbb_rr (EDX,ECX))
- | "__builtin_mull", [BA(IR a); BA(IR b)],
- BR_splitlong(BR(IR rh), BR(IR rl)) ->
- assert (a = EAX && b = EDX && rh = EDX && rl = EAX);
- emit (Pmul_r EDX)
- (* 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 (Pmov_rm (res, linear_addr a1 _0));
- emit (Pbswap res)
- | "__builtin_write16_reversed", [BA(IR a1); BA(IR a2)], _ ->
- let tmp = if a1 = ECX then EDX else ECX in
- if a2 <> tmp then
- emit (Pmov_rr (tmp,a2));
- emit (Pbswap16 tmp);
- emit (Pmovw_mr (linear_addr a1 _0, tmp))
- | "__builtin_write32_reversed", [BA(IR a1); BA(IR a2)], _ ->
- let tmp = if a1 = ECX then EDX else ECX in
- if a2 <> tmp then
- emit (Pmov_rr (tmp,a2));
- emit (Pbswap tmp);
- emit (Pmov_mr (linear_addr a1 _0, tmp))
- (* Vararg stuff *)
- | "__builtin_va_start", [BA(IR a)], _ ->
- expand_builtin_va_start a
- (* Synchronization *)
- | "__builtin_membar", [], _ ->
- ()
- (* no operation *)
- | "__builtin_nop", [], _ ->
- emit (Pxchg_rr (EAX,EAX))
- (* Catch-all *)
- | _ ->
- raise (Error ("unrecognized builtin " ^ name))
-
-(* Expansion of instructions *)
-
-let expand_instruction instr =
- match instr with
- | Pallocframe (sz, ofs_ra, ofs_link) ->
- let sz = sp_adjustment sz in
- let addr = linear_addr ESP (coqint_of_camlint (Int32.add sz 4l)) in
- let addr' = linear_addr ESP ofs_link in
- let sz' = coqint_of_camlint sz in
- emit (Psub_ri (ESP,sz'));
- emit (Pcfi_adjust sz');
- emit (Plea (EDX,addr));
- emit (Pmov_mr (addr',EDX));
- PrintAsmaux.current_function_stacksize := sz
- | Pfreeframe(sz, ofs_ra, ofs_link) ->
- let sz = sp_adjustment sz in
- emit (Padd_ri (ESP,coqint_of_camlint sz))
- | 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
- (Int32.to_int (camlint_of_coqint sz))
- (Int32.to_int (camlint_of_coqint 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 = function
- | EAX -> 0
- | EBX -> 3
- | ECX -> 1
- | EDX -> 2
- | ESI -> 6
- | EDI -> 7
- | EBP -> 5
- | ESP -> 4
-
-let float_reg_to_dwarf = function
- | XMM0 -> 21
- | XMM1 -> 22
- | XMM2 -> 23
- | XMM3 -> 24
- | XMM4 -> 25
- | XMM5 -> 26
- | XMM6 -> 27
- | XMM7 -> 28
-
-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 1d718c26..00000000
--- a/ia32/Asmgen.v
+++ /dev/null
@@ -1,591 +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. *)
-(* *)
-(* *********************************************************************)
-
-(** Translation from Mach to IA32 Asm. *)
-
-Require Import Coqlib Errors.
-Require Import Integers Floats AST 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 various 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 (Ptest_rr EAX EAX ::
- Plea ECX (Addrmode (Some EAX) None (inl _ p)) ::
- Pcmov Cond_l EAX ECX ::
- Psar_ri EAX n :: k).
-
-Definition low_ireg (r: ireg) : bool :=
- match r with
- | EAX | EBX | ECX | EDX => true
- | ESI | EDI | EBP | ESP => false
- end.
-
-Definition mk_intconv (mk: ireg -> ireg -> instruction) (rd rs: ireg) (k: code) :=
- if low_ireg rs then
- OK (mk rd rs :: k)
- else
- OK (Pmov_rr EAX rs :: mk rd EAX :: 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_smallstore (sto: addrmode -> ireg ->instruction)
- (addr: addrmode) (rs: ireg) (k: code) :=
- if low_ireg rs then
- OK (sto addr rs :: k)
- else if addressing_mentions addr EAX then
- OK (Plea ECX addr :: Pmov_rr EAX rs ::
- sto (Addrmode (Some ECX) None (inl _ Int.zero)) EAX :: k)
- else
- OK (Pmov_rr EAX rs :: sto addr EAX :: k).
-
-(** Accessing slots in the stack frame. *)
-
-Definition loadind (base: ireg) (ofs: int) (ty: typ) (dst: mreg) (k: code) :=
- match ty, preg_of dst with
- | Tint, IR r => OK (Pmov_rm r (Addrmode (Some base) None (inl _ ofs)) :: k)
- | Tsingle, FR r => OK (Pmovss_fm r (Addrmode (Some base) None (inl _ ofs)) :: k)
- | Tsingle, ST0 => OK (Pflds_m (Addrmode (Some base) None (inl _ ofs)) :: k)
- | Tfloat, FR r => OK (Pmovsd_fm r (Addrmode (Some base) None (inl _ ofs)) :: k)
- | Tfloat, ST0 => OK (Pfldl_m (Addrmode (Some base) None (inl _ ofs)) :: k)
- | Tany32, IR r => OK (Pmov_rm_a r (Addrmode (Some base) None (inl _ ofs)) :: k)
- | Tany64, FR r => OK (Pmovsd_fm_a r (Addrmode (Some base) None (inl _ ofs)) :: k)
- | _, _ => Error (msg "Asmgen.loadind")
- end.
-
-Definition storeind (src: mreg) (base: ireg) (ofs: int) (ty: typ) (k: code) :=
- match ty, preg_of src with
- | Tint, IR r => OK (Pmov_mr (Addrmode (Some base) None (inl _ ofs)) r :: k)
- | Tsingle, FR r => OK (Pmovss_mf (Addrmode (Some base) None (inl _ ofs)) r :: k)
- | Tsingle, ST0 => OK (Pfstps_m (Addrmode (Some base) None (inl _ ofs)) :: k)
- | Tfloat, FR r => OK (Pmovsd_mf (Addrmode (Some base) None (inl _ ofs)) r :: k)
- | Tfloat, ST0 => OK (Pfstpl_m (Addrmode (Some base) None (inl _ ofs)) :: k)
- | Tany32, IR r => OK (Pmov_mr_a (Addrmode (Some base) None (inl _ ofs)) r :: k)
- | Tany64, FR r => OK (Pmovsd_mf_a (Addrmode (Some base) None (inl _ ofs)) 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, Int.one)) (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 ESP) None (inl _ n))
- | _, _ =>
- Error(msg "Asmgen.transl_addressing")
- 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 (Pcmp_rr r1 r2 :: k)
- | Ccompu c, a1 :: a2 :: nil =>
- do r1 <- ireg_of a1; do r2 <- ireg_of a2; OK (Pcmp_rr r1 r2 :: k)
- | Ccompimm c n, a1 :: nil =>
- do r1 <- ireg_of a1;
- OK (if Int.eq_dec n Int.zero then Ptest_rr r1 r1 :: k else Pcmp_ri r1 n :: k)
- | Ccompuimm c n, a1 :: nil =>
- do r1 <- ireg_of a1; OK (Pcmp_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 (Ptest_ri r1 n :: k)
- | Cmasknotzero n, a1 :: nil =>
- do r1 <- ireg_of a1; OK (Ptest_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)
- | 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 EAX
- then Psetcc c1 EAX :: Psetcc c2 ECX :: Pand_rr EAX ECX :: k
- else Psetcc c1 EAX :: Psetcc c2 rd :: Pand_rr rd EAX :: k
- | Cond_or c1 c2 =>
- if ireg_eq rd EAX
- then Psetcc c1 EAX :: Psetcc c2 ECX :: Por_rr EAX ECX :: k
- else Psetcc c1 EAX :: Psetcc c2 rd :: Por_rr rd EAX :: k
- end.
-
-Definition mk_setcc (cond: extcond) (rd: ireg) (k: code) :=
- if low_ireg rd
- then mk_setcc_base cond rd k
- else mk_setcc_base cond EAX (Pmov_rr rd EAX :: 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 Pxor_r r else Pmov_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_ra 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; mk_intconv Pmovsw_rr r r1 k
- | Ocast16unsigned, a1 :: nil =>
- do r1 <- ireg_of a1; do r <- ireg_of res; mk_intconv Pmovzw_rr r r1 k
- | Oneg, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pneg r :: k)
- | Osub, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Psub_rr r r2 :: k)
- | Omul, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Pimul_rr r r2 :: k)
- | Omulimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pimul_ri r n :: k)
- | Omulhs, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq res DX);
- do r2 <- ireg_of a2; OK (Pimul_r r2 :: k)
- | Omulhu, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq res DX);
- do r2 <- ireg_of a2; OK (Pmul_r r2 :: k)
- | Odiv, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq a2 CX);
- assertion (mreg_eq res AX);
- OK(Pcltd :: Pidiv ECX :: k)
- | Odivu, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq a2 CX);
- assertion (mreg_eq res AX);
- OK(Pxor_r EDX :: Pdiv ECX :: k)
- | Omod, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq a2 CX);
- assertion (mreg_eq res DX);
- OK(Pcltd :: Pidiv ECX :: k)
- | Omodu, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 AX);
- assertion (mreg_eq a2 CX);
- assertion (mreg_eq res DX);
- OK(Pxor_r EDX :: Pdiv ECX :: k)
- | Oand, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Pand_rr r r2 :: k)
- | Oandimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pand_ri r n :: k)
- | Oor, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Por_rr r r2 :: k)
- | Oorimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Por_ri r n :: k)
- | Oxor, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; do r2 <- ireg_of a2; OK (Pxor_rr r r2 :: k)
- | Oxorimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pxor_ri r n :: k)
- | Onot, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pnot r :: k)
- | Oshl, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- assertion (mreg_eq a2 CX);
- do r <- ireg_of res; OK (Psal_rcl r :: k)
- | Oshlimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Psal_ri r n :: k)
- | Oshr, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- assertion (mreg_eq a2 CX);
- do r <- ireg_of res; OK (Psar_rcl r :: k)
- | Oshrimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Psar_ri r n :: k)
- | Oshru, a1 :: a2 :: nil =>
- assertion (mreg_eq a1 res);
- assertion (mreg_eq a2 CX);
- do r <- ireg_of res; OK (Pshr_rcl r :: k)
- | Oshruimm n, a1 :: nil =>
- assertion (mreg_eq a1 res);
- do r <- ireg_of res; OK (Pshr_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 (Pror_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 (Plea r am :: k)
- | 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)
- | 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(Pmov_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_smallstore Pmovb_mr am r k
- | Mint16unsigned | Mint16signed =>
- do r <- ireg_of src; OK(Pmovw_mr am r :: k)
- | Mint32 =>
- do r <- ireg_of src; OK(Pmov_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)
- (edx_is_parent: bool) (k: code) :=
- match i with
- | Mgetstack ofs ty dst =>
- loadind ESP ofs ty dst k
- | Msetstack src ofs ty =>
- storeind src ESP ofs ty k
- | Mgetparam ofs ty dst =>
- if edx_is_parent then
- loadind EDX ofs ty dst k
- else
- (do k1 <- loadind EDX ofs ty dst k;
- loadind ESP f.(fn_link_ofs) Tint DX 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 DX)
- | _ => false
- end.
-
-(** This is the naive definition that we no longer use because it
- is not tail-recursive. It is kept as specification. *)
-
-Fixpoint transl_code (f: Mach.function) (il: list Mach.instruction) (it1p: bool) :=
- match il with
- | nil => OK nil
- | i1 :: il' =>
- do k <- transl_code f il' (it1_is_parent it1p i1);
- transl_instr f i1 it1p k
- end.
-
-(** This is an equivalent definition in continuation-passing style
- that runs in constant stack space. *)
-
-Fixpoint transl_code_rec (f: Mach.function) (il: list Mach.instruction)
- (it1p: bool) (k: code -> res code) :=
- match il with
- | nil => k nil
- | i1 :: il' =>
- transl_code_rec f il' (it1_is_parent it1p i1)
- (fun c1 => do c2 <- transl_instr f i1 it1p c1; k c2)
- end.
-
-Definition transl_code' (f: Mach.function) (il: list Mach.instruction) (it1p: bool) :=
- transl_code_rec f il it1p (fun c => OK c).
-
-(** Translation of a whole function. Note that we must check
- that the generated code contains less than [2^32] instructions,
- otherwise the offset part of the [PC] code pointer could wrap
- around, leading to incorrect executions. *)
-
-Definition transl_function (f: Mach.function) :=
- do c <- transl_code' f f.(Mach.fn_code) true;
- OK (mkfunction f.(Mach.fn_sig)
- (Pallocframe f.(fn_stacksize) f.(fn_retaddr_ofs) f.(fn_link_ofs) :: c)).
-
-Definition transf_function (f: Mach.function) : res Asm.function :=
- do tf <- transl_function f;
- if zlt Int.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 c498b601..00000000
--- a/ia32/Asmgenproof.v
+++ /dev/null
@@ -1,897 +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 proof for x86 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) <= Int.max_unsigned.
-Proof.
- intros. monadInv H. destruct (zlt Int.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_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_smallstore_label:
- forall f addr r k c, mk_smallstore f addr r k = OK c ->
- (forall r addr, nolabel (f r addr)) ->
- tail_nolabel k c.
-Proof.
- unfold mk_smallstore; intros. TailNoLabel.
-Qed.
-Hint Resolve mk_smallstore_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 EAX); TailNoLabel.
-Qed.
-
-Remark mk_setcc_label:
- forall xc rd k,
- tail_nolabel k (mk_setcc xc rd k).
-Proof.
- intros. unfold mk_setcc. destruct (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 i Int.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 i Int.zero); TailNoLabel.
- destruct (Float.eq_dec f Float.zero); TailNoLabel.
- destruct (Float32.eq_dec f Float32.zero); 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 Int.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 (Int.repr pos'))).
- split. unfold goto_label. rewrite P. rewrite H1. auto.
- split. rewrite Pregmap.gss. constructor; auto.
- rewrite Int.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 Int.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)
- (DXP: ep = true -> rs#EDX = 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 Int.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#EDX = 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 DX) = true -> IR EDX <> preg_of dst).
- intros. change (IR EDX) with (preg_of DX). red; intros.
- unfold proj_sumbool in H1. destruct (mreg_eq dst DX); try discriminate.
- elim n. eapply preg_of_injective; eauto.
- destruct ep; simpl in TR.
-(* EDX contains parent *)
- exploit loadind_correct. eexact TR.
- instantiate (2 := rs0). rewrite DXP; 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.
-(* EDX 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) <= Int.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' Int.zero).
- destruct (rs rf); try discriminate.
- revert H; predSpec Int.eq Int.eq_spec i Int.zero; intros; congruence.
- assert (rs0 x0 = Vptr f' Int.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 (Int.add ofs Int.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 (Int.add ofs Int.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) <= Int.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' Int.zero).
- destruct (rs rf); try discriminate.
- revert H; predSpec Int.eq Int.eq_spec i Int.zero; intros; congruence.
- assert (rs0 x0 = Vptr f' Int.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. rewrite C. rewrite A. rewrite <- (sp_val _ _ _ AG). rewrite E. eauto.
- apply star_one. eapply exec_step_internal.
- transitivity (Val.add rs0#PC Vone). 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. rewrite C. rewrite A. rewrite <- (sp_val _ _ _ AG). rewrite E. eauto.
- apply star_one. eapply exec_step_internal.
- transitivity (Val.add rs0#PC Vone). 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.
- exploit find_label_goto_label; 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. eauto.
- econstructor; eauto.
-Transparent destroyed_by_jumptable.
- simpl. eapply agree_exten; eauto. intros. rewrite C; auto with asmgen.
- congruence.
-
-- (* Mreturn *)
- assert (f0 = f) by congruence. subst f0.
- inv AT.
- assert (NOOV: list_length_z tf.(fn_code) <= Int.max_unsigned).
- eapply transf_function_no_overflow; eauto.
- rewrite (sp_val _ _ _ AG) in *. unfold load_stack in *.
- 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.add rs0#PC Vone). 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 Int.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 Int.unsigned_zero. simpl. eauto.
- simpl. rewrite C. simpl in F. rewrite (sp_val _ _ _ AG) in F. rewrite F.
- simpl in P. 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) Int.zero)
- with (Vptr fb Int.zero).
- econstructor; eauto.
- constructor.
- apply Mem.extends_refl.
- split. auto. simpl. unfold Vzero; 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.
- compute in H1. inv H1.
- 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 9703d419..00000000
--- a/ia32/Asmgenproof1.v
+++ /dev/null
@@ -1,1195 +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 proof for IA32 generation: auxiliary results. *)
-
-Require Import Coqlib.
-Require Import AST.
-Require Import Errors.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Op.
-Require Import Locations.
-Require Import Mach.
-Require Import Asm.
-Require Import Asmgen.
-Require Import Asmgenproof0.
-Require Import Conventions.
-
-Open Local Scope error_monad_scope.
-
-(** * Correspondence between Mach registers and IA32 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.add rs#PC Vone.
-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 IA32 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.
-(* movd *)
- 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.
-
-(** 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#EAX) (Vint n) = Some v ->
- exists rs2,
- exec_straight ge fn c rs1 m k rs2 m
- /\ rs2#EAX = v
- /\ forall r, data_preg r = true -> r <> EAX -> r <> ECX -> 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#ECX <- (Vint x'))).
- set (rs4 := nextinstr (if Int.lt x Int.zero then rs3#EAX <- (Vint x') else rs3)).
- set (rs5 := nextinstr_nf (rs4#EAX <- (Val.shr rs4#EAX (Vint n)))).
- assert (rs3#EAX = Vint x). unfold rs3. Simplifs.
- assert (rs3#ECX = 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 EAX) with (rs1 EAX). rewrite A. simpl.
- rewrite (Int.add_commut Int.zero tnm1). rewrite Int.add_zero. 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 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 <> EAX -> rs2#r = rs1#r.
-Proof.
- unfold mk_intconv; intros. destruct (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_addrmode ge a rs1 = eval_addrmode ge a rs2.
-Proof.
- intros until rs2; intro AG. unfold addressing_mentions, eval_addrmode.
- 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_smallstore_correct:
- forall chunk sto addr r k c rs1 m1 m2,
- mk_smallstore sto addr r k = OK c ->
- Mem.storev chunk m1 (eval_addrmode ge addr rs1) (rs1 r) = Some m2 ->
- (forall c r addr rs m,
- exec_instr ge c (sto addr r) rs m = exec_store ge chunk m addr rs r nil) ->
- exists rs2,
- exec_straight ge fn c rs1 m1 k rs2 m2
- /\ forall r, data_preg r = true -> r <> EAX /\ r <> ECX -> rs2#r = rs1#r.
-Proof.
- unfold mk_smallstore; intros.
- remember (low_ireg r) as low. destruct low.
-(* low reg *)
- monadInv H. econstructor; split. apply exec_straight_one. rewrite H1.
- unfold exec_store. rewrite H0. eauto. auto.
- intros; Simplifs.
-(* high reg *)
- remember (addressing_mentions addr EAX) as mentions. destruct mentions; monadInv H.
-(* EAX is mentioned. *)
- assert (r <> ECX). red; intros; subst r; discriminate.
- set (rs2 := nextinstr (rs1#ECX <- (eval_addrmode ge addr rs1))).
- set (rs3 := nextinstr (rs2#EAX <- (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.
- rewrite H1. unfold exec_store. simpl. rewrite Int.add_zero.
- change (rs3 EAX) with (rs1 r).
- change (rs3 ECX) with (eval_addrmode ge addr rs1).
- replace (Val.add (eval_addrmode ge addr rs1) (Vint Int.zero))
- with (eval_addrmode ge addr rs1).
- rewrite H0. eauto.
- destruct (eval_addrmode ge addr rs1); simpl in H0; try discriminate.
- simpl. rewrite Int.add_zero; auto.
- auto. auto. auto.
- intros. destruct H3. Simplifs. unfold rs3; Simplifs. unfold rs2; Simplifs.
-(* EAX is not mentioned *)
- set (rs2 := nextinstr (rs1#EAX <- (rs1 r))).
- econstructor; split.
- apply exec_straight_two with rs2 m1.
- simpl. auto.
- rewrite H1. unfold exec_store.
- rewrite (addressing_mentions_correct addr EAX rs2 rs1); auto.
- change (rs2 EAX) with (rs1 r). rewrite H0. eauto.
- intros. unfold rs2; Simplifs.
- auto. auto.
- intros. destruct H2. simpl. Simplifs. unfold rs2; Simplifs.
-Qed.
-
-(** Accessing slots in the stack frame *)
-
-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.add rs#base (Vint 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 * int) ofs)) in *.
- assert (eval_addrmode ge addr rs = Val.add rs#base (Vint ofs)).
- unfold addr. simpl. rewrite Int.add_commut; rewrite Int.add_zero; auto.
- exists (nextinstr_nf (rs#(preg_of dst) <- v)); split.
-- destruct ty; try discriminate; destruct (preg_of dst); inv H; simpl in H0;
- apply exec_straight_one; auto; simpl; unfold exec_load; rewrite H1, 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.add rs#base (Vint 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.
-Local Transparent destroyed_by_setstack.
- unfold storeind; intros.
- set (addr := Addrmode (Some base) None (inl (ident * int) ofs)) in *.
- assert (eval_addrmode ge addr rs = Val.add rs#base (Vint ofs)).
- unfold addr. simpl. rewrite Int.add_commut; rewrite Int.add_zero; auto.
- destruct ty; try discriminate; destruct (preg_of src); inv H; simpl in H0;
- (econstructor; split;
- [apply exec_straight_one; [simpl; unfold exec_store; rewrite H1, H0; eauto|auto]
- |simpl; intros; unfold undef_regs; repeat Simplifs]).
-Qed.
-
-(** Translation of addressing modes *)
-
-Lemma transl_addressing_mode_correct:
- forall addr args am (rs: regset) v,
- transl_addressing addr args = OK am ->
- eval_addressing ge (rs ESP) addr (List.map rs (List.map preg_of args)) = Some v ->
- Val.lessdef v (eval_addrmode ge am rs).
-Proof.
- assert (A: forall n, Int.add Int.zero n = n).
- intros. rewrite Int.add_commut. apply Int.add_zero.
- assert (B: forall n i, (if Int.eq i Int.one then Vint n else Vint (Int.mul n i)) = Vint (Int.mul n i)).
- intros. predSpec Int.eq Int.eq_spec i Int.one.
- subst i. rewrite Int.mul_one. auto. auto.
- assert (C: forall v i,
- Val.lessdef (Val.mul v (Vint i))
- (if Int.eq i Int.one then v else Val.mul v (Vint i))).
- intros. predSpec Int.eq Int.eq_spec i Int.one.
- subst i. destruct v; simpl; auto. rewrite Int.mul_one; auto.
- destruct v; simpl; auto.
- unfold transl_addressing; intros.
- destruct addr; repeat (destruct args; try discriminate); simpl in H0; inv H0.
-(* indexed *)
- monadInv H. rewrite (ireg_of_eq _ _ EQ). simpl. rewrite A; auto.
-(* indexed2 *)
- monadInv H. rewrite (ireg_of_eq _ _ EQ); rewrite (ireg_of_eq _ _ EQ1). simpl.
- rewrite Val.add_assoc; auto.
-(* scaled *)
- monadInv H. rewrite (ireg_of_eq _ _ EQ). unfold eval_addrmode.
- rewrite Val.add_permut. simpl. rewrite A. apply Val.add_lessdef; auto.
-(* indexed2scaled *)
- monadInv H. rewrite (ireg_of_eq _ _ EQ); rewrite (ireg_of_eq _ _ EQ1); simpl.
- apply Val.add_lessdef; auto. apply Val.add_lessdef; auto.
-(* global *)
- inv H. simpl. unfold Genv.symbol_address.
- destruct (Genv.find_symbol ge i); simpl; auto. repeat rewrite Int.add_zero. auto.
-(* based *)
- monadInv H. rewrite (ireg_of_eq _ _ EQ). simpl.
- unfold Genv.symbol_address. destruct (Genv.find_symbol ge i); simpl; auto.
- rewrite Int.add_zero. rewrite Val.add_commut. auto.
-(* basedscaled *)
- monadInv H. rewrite (ireg_of_eq _ _ EQ). unfold eval_addrmode.
- rewrite (Val.add_commut Vzero). rewrite Val.add_assoc. rewrite Val.add_permut.
- apply Val.add_lessdef; auto. destruct (rs x); simpl; auto. rewrite B. simpl.
- rewrite Int.add_zero. auto.
-(* instack *)
- inv H; simpl. rewrite A; 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 int_signed_eq:
- forall x y, Int.eq x y = zeq (Int.signed x) (Int.signed y).
-Proof.
- intros. unfold Int.eq. unfold proj_sumbool.
- destruct (zeq (Int.unsigned x) (Int.unsigned y));
- destruct (zeq (Int.signed x) (Int.signed y)); auto.
- elim n. unfold Int.signed. rewrite e; auto.
- elim n. apply Int.eqm_small_eq; auto with ints.
- eapply Int.eqm_trans. apply Int.eqm_sym. apply Int.eqm_signed_unsigned.
- rewrite e. apply Int.eqm_signed_unsigned.
-Qed.
-
-Lemma int_not_lt:
- forall x y, negb (Int.lt y x) = (Int.lt x y || Int.eq x y).
-Proof.
- intros. unfold Int.lt. rewrite int_signed_eq. unfold proj_sumbool.
- destruct (zlt (Int.signed y) (Int.signed x)).
- rewrite zlt_false. rewrite zeq_false. auto. omega. omega.
- destruct (zeq (Int.signed x) (Int.signed y)).
- rewrite zlt_false. auto. omega.
- rewrite zlt_true. auto. omega.
-Qed.
-
-Lemma int_lt_not:
- forall x y, Int.lt y x = negb (Int.lt x y) && negb (Int.eq x y).
-Proof.
- intros. rewrite <- negb_orb. rewrite <- int_not_lt. rewrite negb_involutive. auto.
-Qed.
-
-Lemma int_not_ltu:
- forall x y, negb (Int.ltu y x) = (Int.ltu x y || Int.eq x y).
-Proof.
- intros. unfold Int.ltu, Int.eq.
- destruct (zlt (Int.unsigned y) (Int.unsigned x)).
- rewrite zlt_false. rewrite zeq_false. auto. omega. omega.
- destruct (zeq (Int.unsigned x) (Int.unsigned y)).
- rewrite zlt_false. auto. omega.
- rewrite zlt_true. auto. omega.
-Qed.
-
-Lemma int_ltu_not:
- forall x y, Int.ltu y x = negb (Int.ltu x y) && negb (Int.eq x y).
-Proof.
- intros. rewrite <- negb_orb. rewrite <- int_not_ltu. rewrite negb_involutive. auto.
-Qed.
-
-Lemma testcond_for_signed_comparison_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_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 *)
- destruct (Int.eq i Int.zero &&
- (Mem.valid_pointer m b0 (Int.unsigned i0) || Mem.valid_pointer m b0 (Int.unsigned i0 - 1))) eqn:?; try discriminate.
- destruct c; simpl in *; inv H1.
- rewrite Heqb1; reflexivity.
- rewrite Heqb1; reflexivity.
-(* ptr int *)
- destruct (Int.eq i0 Int.zero &&
- (Mem.valid_pointer m b0 (Int.unsigned i) || Mem.valid_pointer m b0 (Int.unsigned i - 1))) eqn:?; try discriminate.
- destruct c; simpl in *; inv H1.
- rewrite Heqb1; reflexivity.
- rewrite Heqb1; reflexivity.
-(* ptr ptr *)
- simpl.
- fold (Mem.weak_valid_pointer m b0 (Int.unsigned i)) in *.
- fold (Mem.weak_valid_pointer m b1 (Int.unsigned i0)) in *.
- destruct (eq_block b0 b1).
- destruct (Mem.weak_valid_pointer m b0 (Int.unsigned i) &&
- Mem.weak_valid_pointer m b1 (Int.unsigned i0)); inversion H1.
- 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.
- destruct (Mem.valid_pointer m b0 (Int.unsigned i) &&
- Mem.valid_pointer m b1 (Int.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_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_correct; eauto.
- intros. unfold compare_ints. Simplifs.
-(* compimm *)
- simpl. rewrite (ireg_of_eq _ _ EQ). destruct (Int.eq_dec i 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_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 i)) eqn:?; auto.
- eapply testcond_for_signed_comparison_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 i)) eqn:?; auto.
- eapply testcond_for_unsigned_comparison_correct; eauto.
- intros. unfold compare_ints. 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 i0 i)) Vzero m).
- intros [A B]. rewrite A. unfold Val.cmpu; simpl. destruct (Int.eq (Int.and i0 i) 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 i0 i)) Vzero m).
- intros [A B]. rewrite A. unfold Val.cmpu; simpl. destruct (Int.eq (Int.and i0 i) 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 <> EAX /\ r <> ECX -> 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 EAX).
- 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 EAX).
- 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 <> EAX /\ r <> ECX -> r <> rd -> rs2#r = rs1#r.
-Proof.
- intros. unfold mk_setcc. destruct (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#ESP) 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 i Int.zero). subst i. TranslOp. TranslOp.
-(* floatconst *)
- apply SAME. destruct (Float.eq_dec f Float.zero). subst f. TranslOp. TranslOp.
-(* singleconst *)
- apply SAME. destruct (Float32.eq_dec f Float32.zero). subst f. TranslOp. TranslOp.
-(* cast8signed *)
- apply SAME. eapply mk_intconv_correct; eauto.
-(* cast8unsigned *)
- apply SAME. eapply mk_intconv_correct; eauto.
-(* cast16signed *)
- apply SAME. eapply mk_intconv_correct; eauto.
-(* cast16unsigned *)
- 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 EAX) (rs ECX)). left; congruence.
- intros (nh & nl & d & q & r & A & B & C & D & E & F).
- set (rs1 := nextinstr_nf (rs#EDX <- (Vint nh))).
- econstructor; split.
- eapply exec_straight_two with (rs2 := rs1). simpl. rewrite A. reflexivity.
- simpl. change (rs1 EAX) with (rs EAX); rewrite B.
- change (rs1 ECX) with (rs ECX); 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 EAX) (rs ECX)). left; congruence.
- intros (n & d & q & r & B & C & D & E & F).
- set (rs1 := nextinstr_nf (rs#EDX <- Vzero)).
- econstructor; split.
- eapply exec_straight_two with (rs2 := rs1). reflexivity.
- simpl. change (rs1 EAX) with (rs EAX); rewrite B.
- change (rs1 ECX) with (rs ECX); 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 EAX) (rs ECX)). right; congruence.
- intros (nh & nl & d & q & r & A & B & C & D & E & F).
- set (rs1 := nextinstr_nf (rs#EDX <- (Vint nh))).
- econstructor; split.
- eapply exec_straight_two with (rs2 := rs1). simpl. rewrite A. reflexivity.
- simpl. change (rs1 EAX) with (rs EAX); rewrite B.
- change (rs1 ECX) with (rs ECX); 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 EAX) (rs ECX)). right; congruence.
- intros (n & d & q & r & B & C & D & E & F).
- set (rs1 := nextinstr_nf (rs#EDX <- Vzero)).
- econstructor; split.
- eapply exec_straight_two with (rs2 := rs1). reflexivity.
- simpl. change (rs1 EAX) with (rs EAX); rewrite B.
- change (rs1 ECX) with (rs ECX); 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_correct; eauto. intros EA.
- TranslOp. rewrite nextinstr_inv; auto with asmgen. rewrite Pregmap.gss; 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.
-(* 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 c0 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#ESP) 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.add (rs PC) Vone).
- transitivity (Val.add ((rs#(preg_of dest) <- v) PC) Vone).
- 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#ESP) 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_smallstore_correct; eauto.
- intros. simpl. unfold exec_store.
- destruct (eval_addrmode ge addr0 rs0); simpl; auto. rewrite Mem.store_signed_unsigned_8; auto.
-(* int8unsigned *)
- eapply mk_smallstore_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.
-(* 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 79a839f3..00000000
--- a/ia32/CBuiltins.ml
+++ /dev/null
@@ -1,87 +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 builtins = {
- Builtins.typedefs = [
- "__builtin_va_list", TPtr(TVoid [], [])
- ];
- Builtins.functions = [
- (* Integer arithmetic *)
- "__builtin_bswap",
- (TInt(IUInt, []), [TInt(IUInt, [])], 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);
- ]
-}
-
-let size_va_list = 4
-let va_list_scalar = true
-
-(* 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 cdd16071..00000000
--- a/ia32/CombineOp.v
+++ /dev/null
@@ -1,117 +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. *)
-(* *)
-(* *********************************************************************)
-
-(** Recognition of combined operations, addressing modes and conditions
- during the [CSE] phase. *)
-
-Require Import Coqlib.
-Require Import AST.
-Require Import Integers.
-Require Import Op.
-Require Import 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 (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) => Some(offset_addressing_total a n, ys)
- | _ => None
- end
- | _, _ => None
- end.
-
-Function combine_op (op: operation) (args: list valnum) : option(operation * list valnum) :=
- match op, args with
- | Olea addr, _ =>
- match combine_addr addr args with
- | Some(addr', args') => Some(Olea 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
- | 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 8f600054..00000000
--- a/ia32/CombineOpproof.v
+++ /dev/null
@@ -1,153 +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. *)
-(* *)
-(* *********************************************************************)
-
-(** Recognition of combined operations, addressing modes and conditions
- during the [CSE] phase. *)
-
-Require Import Coqlib.
-Require Import Integers.
-Require Import Values.
-Require Import Memory.
-Require Import Op.
-Require Import RTL.
-Require Import 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_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.
- intros. functional inversion H; subst.
- (* indexed - lea *)
- UseGetSound. simpl. eapply eval_offset_addressing_total; eauto.
-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_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.
-(* 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 a3de748c..00000000
--- a/ia32/ConstpropOp.vp
+++ /dev/null
@@ -1,227 +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. *)
-(* *)
-(* *********************************************************************)
-
-(** Strength reduction for operators and conditions.
- This is the machine-dependent part of [Constprop]. *)
-
-Require Import Coqlib.
-Require Import Compopts.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Op.
-Require Import Registers.
-Require Import ValueDomain.
-
-(** * 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)
- | _, _, _ =>
- (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.
-
-Nondetfunction addr_strength_reduction
- (addr: addressing) (args: list reg) (vl: list aval) :=
- match addr, args, vl with
-
- | Aindexed ofs, r1 :: nil, Ptr(Gl symb n) :: nil =>
- (Aglobal symb (Int.add n ofs), nil)
- | Aindexed ofs, r1 :: nil, Ptr(Stk n) :: nil =>
- (Ainstack (Int.add n ofs), nil)
-
- | Aindexed2 ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: I n2 :: nil =>
- (Aglobal symb (Int.add (Int.add n1 n2) ofs), nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, I n1 :: Ptr(Gl symb n2) :: nil =>
- (Aglobal symb (Int.add (Int.add n1 n2) ofs), nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, Ptr(Stk n1) :: I n2 :: nil =>
- (Ainstack (Int.add (Int.add n1 n2) ofs), nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, I n1 :: Ptr(Stk n2) :: nil =>
- (Ainstack (Int.add (Int.add n1 n2) ofs), nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: v2 :: nil =>
- (Abased symb (Int.add n1 ofs), r2 :: nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, v1 :: Ptr(Gl symb n2) :: nil =>
- (Abased symb (Int.add n2 ofs), r1 :: nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, I n1 :: v2 :: nil =>
- (Aindexed (Int.add n1 ofs), r2 :: nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
- (Aindexed (Int.add n2 ofs), r1 :: nil)
-
- | Aindexed2scaled sc ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: I n2 :: nil =>
- (Aglobal symb (Int.add (Int.add n1 (Int.mul n2 sc)) ofs), nil)
- | Aindexed2scaled sc ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: v2 :: nil =>
- (Abasedscaled sc symb (Int.add n1 ofs), r2 :: nil)
- | Aindexed2scaled sc ofs, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
- (Aindexed (Int.add (Int.mul n2 sc) ofs), r1 :: nil)
-
- | Abased id ofs, r1 :: nil, I n1 :: nil =>
- (Aglobal id (Int.add ofs n1), nil)
-
- | Abasedscaled sc id ofs, r1 :: nil, I n1 :: nil =>
- (Aglobal id (Int.add ofs (Int.mul sc n1)), nil)
-
- | _, _ =>
- (addr, args)
- end.
-
-Definition make_addimm (n: int) (r: reg) :=
- if Int.eq n Int.zero
- then (Omove, r :: nil)
- else (Olea (Aindexed 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_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 addr args vl in
- (Olea 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 3dfb8ccf..00000000
--- a/ia32/ConstpropOpproof.v
+++ /dev/null
@@ -1,543 +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 proof for operator strength reduction. *)
-
-Require Import Coqlib.
-Require Import Compopts.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Events.
-Require Import Op.
-Require Import Registers.
-Require Import RTL.
-Require Import ValueDomain.
-Require Import ConstpropOp.
-
-(** We now show that strength reduction over operators and addressing
- modes preserve semantics: the strength-reduced operations and
- addressings evaluate to the same values as the original ones if the
- actual arguments match the static approximations used for strength
- reduction. *)
-
-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 (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 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.
-- 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 Int.zero) addr e##args = Some res ->
- let (addr', args') := addr_strength_reduction addr args vl in
- exists res', eval_addressing ge (Vptr sp Int.zero) addr' e##args' = Some res' /\ Val.lessdef res res'.
-Proof.
- intros until res. unfold addr_strength_reduction.
- destruct (addr_strength_reduction_match addr args vl); simpl;
- intros VL EA; InvApproxRegs; SimplVM; try (inv EA).
-- rewrite Genv.shift_symbol_address. econstructor; split. eauto. apply Val.add_lessdef; auto.
-- econstructor; split; eauto. rewrite Int.add_zero_l.
- change (Vptr sp (Int.add n ofs)) with (Val.add (Vptr sp n) (Vint ofs)). apply Val.add_lessdef; auto.
-- econstructor; split; eauto. rewrite Int.add_assoc. rewrite Genv.shift_symbol_address.
- rewrite Val.add_assoc. apply Val.add_lessdef; auto.
-- econstructor; split; eauto.
- fold (Val.add (Vint n1) e#r2). rewrite (Val.add_commut (Vint n1)).
- rewrite Genv.shift_symbol_address. apply Val.add_lessdef; auto.
- rewrite Int.add_commut. rewrite Genv.shift_symbol_address. apply Val.add_lessdef; auto.
-- econstructor; split; eauto. rewrite Int.add_zero_l. rewrite Int.add_assoc.
- change (Vptr sp (Int.add n1 (Int.add n2 ofs)))
- with (Val.add (Vptr sp n1) (Vint (Int.add n2 ofs))).
- rewrite Val.add_assoc. apply Val.add_lessdef; auto.
-- econstructor; split; eauto. rewrite Int.add_zero_l.
- fold (Val.add (Vint n1) e#r2). rewrite (Int.add_commut n1).
- change (Vptr sp (Int.add (Int.add n2 n1) ofs))
- with (Val.add (Val.add (Vint n1) (Vptr sp n2)) (Vint ofs)).
- apply Val.add_lessdef; auto. apply Val.add_lessdef; auto.
-- econstructor; split; eauto. rewrite Genv.shift_symbol_address.
- rewrite ! Val.add_assoc. apply Val.add_lessdef; auto.
- rewrite Val.add_commut. apply Val.add_lessdef; auto.
-- econstructor; split; eauto. rewrite Genv.shift_symbol_address.
- rewrite (Val.add_commut e#r1). rewrite ! Val.add_assoc.
- apply Val.add_lessdef; auto. rewrite Val.add_commut. apply Val.add_lessdef; auto.
-- fold (Val.add (Vint n1) e#r2). econstructor; split; eauto.
- rewrite (Val.add_commut (Vint n1)). rewrite Val.add_assoc.
- apply Val.add_lessdef; eauto.
-- econstructor; split; eauto. rewrite ! Val.add_assoc.
- apply Val.add_lessdef; eauto.
-- econstructor; split; eauto. rewrite Int.add_assoc.
- rewrite Genv.shift_symbol_address. apply Val.add_lessdef; auto.
-- econstructor; split; eauto.
- rewrite Genv.shift_symbol_address. rewrite ! Val.add_assoc. apply Val.add_lessdef; auto.
- rewrite Val.add_commut; auto.
-- econstructor; split; eauto.
-- econstructor; split; eauto. rewrite Genv.shift_symbol_address. auto.
-- econstructor; split; eauto. rewrite Genv.shift_symbol_address. rewrite Int.mul_commut; auto.
-- econstructor; eauto.
-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 Int.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 Int.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 Int.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.
- exists (Val.add e#r (Vint n)); 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 Int.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 Int.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 Int.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 Int.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 Int.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 Int.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 Int.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 Int.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 Int.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 Int.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_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 Int.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 Int.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 Int.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 Int.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 Int.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 Int.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 Int.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 Int.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 Int.zero) op e##args m = Some v ->
- let (op', args') := op_strength_reduction op args vl in
- exists w, eval_operation ge (Vptr sp Int.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_correct; eauto.
- destruct (addr_strength_reduction 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 08a86815..00000000
--- a/ia32/Conventions1.v
+++ /dev/null
@@ -1,240 +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. *)
-(* *)
-(* *********************************************************************)
-
-(** Function calling conventions and other conventions regarding the use of
- machine registers and stack slots. *)
-
-Require Import Coqlib.
-Require Import Decidableplus.
-Require Import AST.
-Require Import Events.
-Require Import 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 application binary interface (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 | SI | DI | BP => true
- | X0 | X1 | X2 | X3 | X4 | X5 | X6 | X7 => false
- | FP0 => false
- end.
-
-Definition int_caller_save_regs := AX :: CX :: DX :: nil.
-
-Definition float_caller_save_regs := X0 :: X1 :: X2 :: X3 :: X4 :: X5 :: X6 :: X7 :: nil.
-
-Definition int_callee_save_regs := BX :: SI :: DI :: BP :: nil.
-
-Definition float_callee_save_regs : list mreg := nil.
-
-Definition destroyed_at_call :=
- 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]. *)
-
-(** * 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 conventions. *)
-
-(** ** Location of function result *)
-
-(** 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 (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.
-
-(** 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. 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. 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
- end.
-Proof.
- intros; unfold loc_result; destruct (sig_res sg) as [[]|]; auto. intuition congruence.
-Qed.
-
-(** ** Location of function arguments *)
-
-(** All arguments are passed on stack. (Snif.) *)
-
-Fixpoint loc_arguments_rec
- (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_rec tys (ofs + typesize ty)
- 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) :=
- loc_arguments_rec s.(sig_args) 0.
-
-(** [size_arguments s] returns the number of [Outgoing] slots used
- to call a function with signature [s]. *)
-
-Fixpoint size_arguments_rec
- (tyl: list typ) (ofs: Z) {struct tyl} : Z :=
- match tyl with
- | nil => ofs
- | ty :: tys => size_arguments_rec tys (ofs + typesize ty)
- end.
-
-Definition size_arguments (s: signature) : Z :=
- size_arguments_rec s.(sig_args) 0.
-
-(** Argument locations are either caller-save registers or [Outgoing]
- stack slots at nonnegative offsets. *)
-
-Definition loc_argument_acceptable (l: loc) : Prop :=
- match l with
- | R r => is_callee_save r = false
- | S Outgoing ofs ty => ofs >= 0 /\ (typealign ty | ofs)
- | _ => False
- end.
-
-Definition loc_argument_charact (ofs: Z) (l: loc) : Prop :=
- match l with
- | S Outgoing ofs' ty => ofs' >= ofs /\ typealign ty = 1
- | _ => False
- end.
-
-Remark loc_arguments_rec_charact:
- forall tyl ofs p,
- In p (loc_arguments_rec tyl ofs) -> forall_rpair (loc_argument_charact ofs) p.
-Proof.
- assert (X: forall ofs1 ofs2 l, loc_argument_charact ofs2 l -> ofs1 <= ofs2 -> loc_argument_charact ofs1 l).
- { destruct l; simpl; intros; auto. destruct sl; auto. intuition omega. }
- induction tyl as [ | ty tyl]; simpl loc_arguments_rec; 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.
-
-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.
- exploit loc_arguments_rec_charact; eauto.
- assert (X: forall l, loc_argument_charact 0 l -> loc_argument_acceptable l).
- { destruct l as [r | [] ofs ty]; simpl; intuition auto. rewrite H2; apply Z.divide_1_l. }
- destruct p; simpl; intuition auto.
-Qed.
-
-Hint Resolve loc_arguments_acceptable: locs.
-
-(** The offsets of [Outgoing] arguments are below [size_arguments s]. *)
-
-Remark size_arguments_rec_above:
- forall tyl ofs0, ofs0 <= size_arguments_rec tyl ofs0.
-Proof.
- induction tyl; simpl; intros.
- omega.
- apply Zle_trans with (ofs0 + typesize a); auto.
- generalize (typesize_pos a); omega.
-Qed.
-
-Lemma size_arguments_above:
- forall s, size_arguments s >= 0.
-Proof.
- intros; unfold size_arguments. apply Zle_ge.
- apply size_arguments_rec_above.
-Qed.
-
-Lemma loc_arguments_bounded:
- forall (s: signature) (ofs: Z) (ty: typ),
- In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments s)) ->
- ofs + typesize ty <= size_arguments s.
-Proof.
- intros until ty. unfold loc_arguments, size_arguments. generalize (sig_args s) 0.
- induction l 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_rec_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_main:
- loc_arguments signature_main = nil.
-Proof.
- reflexivity.
-Qed.
diff --git a/ia32/Machregs.v b/ia32/Machregs.v
deleted file mode 100644
index 3a6ae674..00000000
--- a/ia32/Machregs.v
+++ /dev/null
@@ -1,297 +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: mreg | BX: mreg | CX: mreg | DX: mreg | SI: mreg | DI: mreg | BP: mreg
- (** Allocatable float regs *)
- | X0: mreg | X1: mreg | X2: mreg | X3: mreg
- | X4: mreg | X5: mreg | X6: mreg | X7: mreg
- (** Special float reg *)
- | FP0: mreg (**r top of x87 FP stack *).
-
-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
- :: X0 :: X1 :: X2 :: X3 :: X4 :: X5 :: X6 :: X7
- :: 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 => Tany32
- | X0 | X1 | X2 | X3 | X4 | X5 | X6 | X7 | 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
- | X0 => 8 | X1 => 9 | X2 => 10 | X3 => 11
- | X4 => 12 | X5 => 13 | X6 => 14 | X7 => 15
- | FP0 => 16
- 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 :=
- ("EAX", AX) :: ("EBX", BX) :: ("ECX", CX) :: ("EDX", DX) ::
- ("ESI", SI) :: ("EDI", DI) :: ("EBP", BP) ::
- ("XMM0", X0) :: ("XMM1", X1) :: ("XMM2", X2) :: ("XMM3", X3) ::
- ("XMM4", X4) :: ("XMM5", X5) :: ("XMM6", X6) :: ("XMM7", X7) ::
- ("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 | Ocast16signed | Ocast16unsigned => 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
- | 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 => AX :: CX :: nil
- | _ => nil
- end.
-
-Definition destroyed_by_cond (cond: condition): list mreg :=
- nil.
-
-Definition destroyed_by_jumptable: list mreg :=
- 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) => AX :: CX :: nil
- | EF_builtin name sg =>
- 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] *)
- DX :: FP0 :: nil.
-
-Definition destroyed_at_indirect_call: list mreg :=
- nil.
-
-Definition destroyed_by_setstack (ty: typ): list mreg :=
- match ty with
- | Tfloat | Tsingle => FP0 :: nil
- | _ => nil
- end.
-
-Definition temp_for_parent_frame: mreg :=
- DX.
-
-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)
- | _ => (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
- (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_at_indirect_call
- 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
- | 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
- | 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
- | Omakelong => false
- | Olowlong => false
- | Ohighlong => 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 07eec160..00000000
--- a/ia32/NeedOp.v
+++ /dev/null
@@ -1,194 +0,0 @@
-Require Import Coqlib.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Op.
-Require Import NeedDomain.
-Require Import RTL.
-
-(** Neededness analysis for IA32 operators *)
-
-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 (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_operation (op: operation) (nv: nval): list nval :=
- match op with
- | Omove => op1 nv
- | Ointconst 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 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)
- | Omakelong => op2 (default nv)
- | Olowlong | Ohighlong => 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 i) as [b'|] eqn:MZ; try discriminate.
- erewrite maskzero_sound; eauto.
-Qed.
-
-Lemma needs_of_addressing_sound:
- forall (ge: genv) sp addr args v nv args',
- eval_addressing ge (Vptr sp Int.zero) addr args = Some v ->
- vagree_list args args' (needs_of_addressing addr nv) ->
- exists v',
- eval_addressing ge (Vptr sp Int.zero) addr args' = Some v'
- /\ vagree v v' nv.
-Proof.
- unfold needs_of_addressing; 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_operation_sound:
- forall op args v nv args',
- eval_operation ge (Vptr sp Int.zero) op args m = Some v ->
- vagree_list args args' (needs_of_operation op nv) ->
- nv <> Nothing ->
- exists v',
- eval_operation ge (Vptr sp Int.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_sound; eauto.
-- destruct (eval_condition c 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 Int.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 f21d7c6a..00000000
--- a/ia32/Op.v
+++ /dev/null
@@ -1,1075 +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 IA32-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: comparison -> condition (**r signed integer comparison *)
- | Ccompu: comparison -> condition (**r unsigned integer comparison *)
- | Ccompimm: comparison -> int -> condition (**r signed integer comparison with a constant *)
- | Ccompuimm: comparison -> int -> condition (**r unsigned integer comparison with a constant *)
- | Ccompf: comparison -> condition (**r 64-bit floating-point comparison *)
- | Cnotcompf: comparison -> condition (**r negation of a floating-point comparison *)
- | Ccompfs: comparison -> condition (**r 32-bit floating-point comparison *)
- | Cnotcompfs: comparison -> condition (**r negation of a floating-point comparison *)
- | Cmaskzero: int -> condition (**r test [(arg & constant) == 0] *)
- | Cmasknotzero: int -> condition. (**r test [(arg & constant) != 0] *)
-
-(** Addressing modes. [r1], [r2], etc, are the arguments to the
- addressing. *)
-
-Inductive addressing: Type :=
- | Aindexed: int -> addressing (**r Address is [r1 + offset] *)
- | Aindexed2: int -> addressing (**r Address is [r1 + r2 + offset] *)
- | Ascaled: int -> int -> addressing (**r Address is [r1 * scale + offset] *)
- | Aindexed2scaled: int -> int -> addressing
- (**r Address is [r1 + r2 * scale + offset] *)
- | Aglobal: ident -> int -> addressing (**r Address is [symbol + offset] *)
- | Abased: ident -> int -> addressing (**r Address is [symbol + offset + r1] *)
- | Abasedscaled: int -> ident -> int -> addressing (**r Address is [symbol + offset + r1 * scale] *)
- | Ainstack: int -> 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: operation (**r [rd = r1] *)
- | Ointconst: int -> operation (**r [rd] is set to the given integer constant *)
- | Ofloatconst: float -> operation (**r [rd] is set to the given float constant *)
- | Osingleconst: float32 -> operation (**r [rd] is set to the given float constant *)
- | Oindirectsymbol: ident -> operation (**r [rd] is set to the address of the symbol *)
-(*c Integer arithmetic: *)
- | Ocast8signed: operation (**r [rd] is 8-bit sign extension of [r1] *)
- | Ocast8unsigned: operation (**r [rd] is 8-bit zero extension of [r1] *)
- | Ocast16signed: operation (**r [rd] is 16-bit sign extension of [r1] *)
- | Ocast16unsigned: operation (**r [rd] is 16-bit zero extension of [r1] *)
- | Oneg: operation (**r [rd = - r1] *)
- | Osub: operation (**r [rd = r1 - r2] *)
- | Omul: operation (**r [rd = r1 * r2] *)
- | Omulimm: int -> operation (**r [rd = r1 * n] *)
- | Omulhs: operation (**r [rd = high part of r1 * r2, signed] *)
- | Omulhu: operation (**r [rd = high part of r1 * r2, unsigned] *)
- | Odiv: operation (**r [rd = r1 / r2] (signed) *)
- | Odivu: operation (**r [rd = r1 / r2] (unsigned) *)
- | Omod: operation (**r [rd = r1 % r2] (signed) *)
- | Omodu: operation (**r [rd = r1 % r2] (unsigned) *)
- | Oand: operation (**r [rd = r1 & r2] *)
- | Oandimm: int -> operation (**r [rd = r1 & n] *)
- | Oor: operation (**r [rd = r1 | r2] *)
- | Oorimm: int -> operation (**r [rd = r1 | n] *)
- | Oxor: operation (**r [rd = r1 ^ r2] *)
- | Oxorimm: int -> operation (**r [rd = r1 ^ n] *)
- | Onot: operation (**r [rd = ~r1] *)
- | Oshl: operation (**r [rd = r1 << r2] *)
- | Oshlimm: int -> operation (**r [rd = r1 << n] *)
- | Oshr: operation (**r [rd = r1 >> r2] (signed) *)
- | Oshrimm: int -> operation (**r [rd = r1 >> n] (signed) *)
- | Oshrximm: int -> operation (**r [rd = r1 / 2^n] (signed) *)
- | Oshru: operation (**r [rd = r1 >> r2] (unsigned) *)
- | Oshruimm: int -> operation (**r [rd = r1 >> n] (unsigned) *)
- | Ororimm: int -> operation (**r rotate right immediate *)
- | Oshldimm: int -> operation (**r [rd = r1 << n | r2 >> (32-n)] *)
- | Olea: addressing -> operation (**r effective address *)
-(*c Floating-point arithmetic: *)
- | Onegf: operation (**r [rd = - r1] *)
- | Oabsf: operation (**r [rd = abs(r1)] *)
- | Oaddf: operation (**r [rd = r1 + r2] *)
- | Osubf: operation (**r [rd = r1 - r2] *)
- | Omulf: operation (**r [rd = r1 * r2] *)
- | Odivf: operation (**r [rd = r1 / r2] *)
- | Onegfs: operation (**r [rd = - r1] *)
- | Oabsfs: operation (**r [rd = abs(r1)] *)
- | Oaddfs: operation (**r [rd = r1 + r2] *)
- | Osubfs: operation (**r [rd = r1 - r2] *)
- | Omulfs: operation (**r [rd = r1 * r2] *)
- | Odivfs: operation (**r [rd = r1 / r2] *)
- | Osingleoffloat: operation (**r [rd] is [r1] truncated to single-precision float *)
- | Ofloatofsingle: operation (**r [rd] is [r1] extended to double-precision float *)
-(*c Conversions between int and float: *)
- | Ointoffloat: operation (**r [rd = signed_int_of_float64(r1)] *)
- | Ofloatofint: operation (**r [rd = float64_of_signed_int(r1)] *)
- | Ointofsingle: operation (**r [rd = signed_int_of_float32(r1)] *)
- | Osingleofint: operation (**r [rd = float32_of_signed_int(r1)] *)
-(*c Manipulating 64-bit integers: *)
- | Omakelong: operation (**r [rd = r1 << 32 | r2] *)
- | Olowlong: operation (**r [rd = low-word(r1)] *)
- | Ohighlong: operation (**r [rd = high-word(r1)] *)
-(*c Boolean tests: *)
- | Ocmp: condition -> operation. (**r [rd = 1] if condition holds, [rd = 0] otherwise. *)
-
-(** Derived operators. *)
-
-Definition Oaddrsymbol (id: ident) (ofs: int) : operation := Olea (Aglobal id ofs).
-Definition Oaddrstack (ofs: int) : operation := Olea (Ainstack ofs).
-Definition Oaddimm (n: int) : operation := Olea (Aindexed n).
-
-(** Comparison functions (used in modules [CSE] and [Allocation]). *)
-
-Definition eq_condition (x y: condition) : {x=y} + {x<>y}.
-Proof.
- generalize Int.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 Int.eq_dec; intro.
- assert (forall (x y: ident), {x=y}+{x<>y}). exact peq.
- decide equality.
-Defined.
-
-Definition eq_operation (x y: operation): {x=y} + {x<>y}.
-Proof.
- generalize Int.eq_dec; intro.
- generalize Float.eq_dec; intro.
- generalize Float32.eq_dec; intro.
- generalize Int64.eq_dec; intro.
- decide equality.
- apply peq.
- apply eq_addressing.
- apply eq_condition.
-Defined.
-
-Global Opaque eq_condition eq_addressing eq_operation.
-
-(** * 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)
- | 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_addressing
- (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 n))
- | Aindexed2 n, v1::v2::nil =>
- Some (Val.add (Val.add v1 v2) (Vint n))
- | Ascaled sc ofs, v1::nil =>
- Some (Val.add (Val.mul v1 (Vint sc)) (Vint ofs))
- | Aindexed2scaled sc ofs, v1::v2::nil =>
- Some(Val.add v1 (Val.add (Val.mul v2 (Vint sc)) (Vint ofs)))
- | Aglobal s ofs, nil =>
- Some (Genv.symbol_address genv s ofs)
- | Abased s ofs, v1::nil =>
- Some (Val.add (Genv.symbol_address genv s ofs) v1)
- | Abasedscaled sc s ofs, v1::nil =>
- Some (Val.add (Genv.symbol_address genv s ofs) (Val.mul v1 (Vint sc)))
- | Ainstack ofs, nil =>
- Some(Val.add sp (Vint ofs))
- | _, _ => None
- end.
-
-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)
- | Ofloatconst n, nil => Some (Vfloat n)
- | Osingleconst n, nil => Some (Vsingle n)
- | Oindirectsymbol id, nil => Some (Genv.symbol_address genv id Int.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_addressing 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
- | Omakelong, v1::v2::nil => Some(Val.longofwords v1 v2)
- | Olowlong, v1::nil => Some(Val.loword v1)
- | Ohighlong, v1::nil => Some(Val.hiword v1)
- | Ocmp c, _ => Some(Val.of_optbool (eval_condition c vl m))
- | _, _ => None
- end.
-
-Ltac FuncInv :=
- match goal with
- | H: (match ?x with nil => _ | _ :: _ => _ end = Some _) |- _ =>
- destruct x; simpl in H; try discriminate; FuncInv
- | H: (match ?v with Vundef => _ | Vint _ => _ | Vfloat _ => _ | Vptr _ _ => _ end = Some _) |- _ =>
- destruct v; simpl in H; try discriminate; FuncInv
- | H: (Some _ = Some _) |- _ =>
- injection H; intros; clear H; FuncInv
- | _ =>
- 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
- | 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 (addr: addressing) : list typ :=
- match addr with
- | Aindexed _ => Tint :: nil
- | Aindexed2 _ => Tint :: Tint :: nil
- | Ascaled _ _ => Tint :: nil
- | Aindexed2scaled _ _ => Tint :: Tint :: nil
- | Aglobal _ _ => nil
- | Abased _ _ => Tint :: nil
- | Abasedscaled _ _ _ => Tint :: nil
- | Ainstack _ => nil
- end.
-
-Definition type_of_operation (op: operation) : list typ * typ :=
- match op with
- | Omove => (nil, Tint) (* treated specially *)
- | Ointconst _ => (nil, Tint)
- | Ofloatconst f => (nil, Tfloat)
- | Osingleconst f => (nil, Tsingle)
- | Oindirectsymbol _ => (nil, Tint)
- | 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_addressing addr, Tint)
- | 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)
- | Omakelong => (Tint :: Tint :: nil, Tlong)
- | Olowlong => (Tlong :: nil, Tint)
- | Ohighlong => (Tlong :: nil, Tint)
- | 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.
-
-Lemma type_of_addressing_sound:
- forall addr vl sp v,
- eval_addressing genv sp addr vl = Some v ->
- Val.has_type v Tint.
-Proof with (try exact I).
- intros. destruct addr; simpl in H; FuncInv; subst; simpl.
- destruct v0...
- destruct v0... destruct v1... destruct v1...
- destruct v0...
- destruct v0... destruct v1... destruct v1...
- unfold Genv.symbol_address; destruct (Genv.find_symbol genv i)...
- unfold Genv.symbol_address; destruct (Genv.find_symbol genv i)...
- unfold Genv.symbol_address; destruct (Genv.find_symbol genv i)... destruct v0...
- destruct v0...
- unfold Genv.symbol_address; destruct (Genv.find_symbol genv i0)... destruct v0...
- destruct sp...
-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.
- unfold Genv.symbol_address; destruct (Genv.find_symbol genv i)...
- destruct v0...
- destruct v0...
- destruct v0...
- destruct v0...
- destruct v0...
- destruct v0; destruct v1... simpl. destruct (eq_block b b0)...
- 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 i Int.iwordsize)...
- destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int.iwordsize)...
- destruct v0; simpl... destruct (Int.ltu i Int.iwordsize)...
- destruct v0; simpl in H0; try discriminate. destruct (Int.ltu i (Int.repr 31)); inv H0...
- destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int.iwordsize)...
- destruct v0; simpl... destruct (Int.ltu i Int.iwordsize)...
- destruct v0...
- destruct v0; simpl... destruct (Int.ltu i Int.iwordsize)...
- destruct v1; simpl... destruct (Int.ltu (Int.sub Int.iwordsize i) Int.iwordsize)...
- eapply type_of_addressing_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; destruct v1...
- destruct v0...
- destruct v0...
- destruct (eval_condition c 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
- | 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).
- 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 i) as [[]|]; auto.
-Qed.
-
-(** Shifting stack-relative references. This is used in [Stacking]. *)
-
-Definition shift_stack_addressing (delta: int) (addr: addressing) :=
- match addr with
- | Ainstack ofs => Ainstack (Int.add delta ofs)
- | _ => addr
- end.
-
-Definition shift_stack_operation (delta: int) (op: operation) :=
- match op with
- | Olea addr => Olea (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. apply type_shift_stack_addressing.
-Qed.
-
-Lemma eval_shift_stack_addressing:
- forall F V (ge: Genv.t F V) sp addr vl delta,
- eval_addressing ge sp (shift_stack_addressing delta addr) vl =
- eval_addressing ge (Val.add sp (Vint delta)) addr vl.
-Proof.
- intros. destruct addr; simpl; auto.
- rewrite Val.add_assoc. simpl. auto.
-Qed.
-
-Lemma eval_shift_stack_operation:
- forall F V (ge: Genv.t F V) sp op vl m delta,
- eval_operation ge sp (shift_stack_operation delta op) vl m =
- eval_operation ge (Val.add sp (Vint delta)) op vl m.
-Proof.
- intros. destruct op; simpl; auto.
- apply eval_shift_stack_addressing.
-Qed.
-
-(** Offset an addressing mode [addr] by a quantity [delta], so that
- it designates the pointer [delta] bytes past the pointer designated
- by [addr]. On PowerPC and ARM, this may be undefined, in which case
- [None] is returned. On IA32, it is always defined, but we keep the
- same interface. *)
-
-Definition offset_addressing_total (addr: addressing) (delta: int) : addressing :=
- match addr with
- | Aindexed n => Aindexed (Int.add n delta)
- | Aindexed2 n => Aindexed2 (Int.add n delta)
- | Ascaled sc n => Ascaled sc (Int.add n delta)
- | Aindexed2scaled sc n => Aindexed2scaled sc (Int.add n delta)
- | Aglobal s n => Aglobal s (Int.add n delta)
- | Abased s n => Abased s (Int.add n delta)
- | Abasedscaled sc s n => Abasedscaled sc s (Int.add n delta)
- | Ainstack n => Ainstack (Int.add n delta)
- end.
-
-Definition offset_addressing (addr: addressing) (delta: int) : option addressing :=
- Some(offset_addressing_total addr delta).
-
-Lemma eval_offset_addressing_total:
- forall (F V: Type) (ge: Genv.t F V) sp addr args delta v,
- eval_addressing ge sp addr args = Some v ->
- eval_addressing ge sp (offset_addressing_total addr delta) args =
- Some(Val.add v (Vint delta)).
-Proof.
- intros. destruct addr; simpl in *; FuncInv; subst.
- rewrite Val.add_assoc; auto.
- rewrite !Val.add_assoc; auto.
- rewrite !Val.add_assoc; auto.
- rewrite !Val.add_assoc; auto.
- unfold Genv.symbol_address. destruct (Genv.find_symbol ge i); auto.
- unfold Genv.symbol_address. destruct (Genv.find_symbol ge i); auto.
- rewrite Val.add_assoc. rewrite Val.add_permut. rewrite Val.add_commut. auto.
- unfold Genv.symbol_address. destruct (Genv.find_symbol ge i0); auto.
- rewrite Val.add_assoc. rewrite Val.add_permut. rewrite Val.add_commut. auto.
- rewrite Val.add_assoc. auto.
-Qed.
-
-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 ->
- eval_addressing ge sp addr' args = Some(Val.add v (Vint delta)).
-Proof.
- intros. unfold offset_addressing in H; inv H.
- eapply eval_offset_addressing_total; eauto.
-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
- | Olea (Aglobal _ _) => true
- | Olea (Ainstack _) => true
- | _ => false
- end.
-
-(** Operations that depend on the memory state. *)
-
-Definition op_depends_on_memory (op: operation) : bool :=
- match op with
- | Ocmp (Ccompu _) => true
- | Ocmp (Ccompuimm _ _) => true
- | _ => 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 c; simpl; auto; congruence.
-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
- | _ => 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_addressing_preserved:
- forall sp addr vl,
- eval_addressing ge2 sp addr vl = eval_addressing ge1 sp addr vl.
-Proof.
- intros.
- unfold eval_addressing, Genv.symbol_address; destruct addr; try rewrite agree_on_symbols;
- reflexivity.
-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_addressing_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 (Int.unsigned ofs) = true ->
- Mem.valid_pointer m2 b2 (Int.unsigned (Int.add ofs (Int.repr delta))) = true.
-
-Hypothesis weak_valid_pointer_inj:
- forall b1 ofs b2 delta,
- f b1 = Some(b2, delta) ->
- Mem.weak_valid_pointer m1 b1 (Int.unsigned ofs) = true ->
- Mem.weak_valid_pointer m2 b2 (Int.unsigned (Int.add ofs (Int.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 (Int.unsigned ofs) = true ->
- 0 <= Int.unsigned ofs + Int.unsigned (Int.repr delta) <= Int.max_unsigned.
-
-Hypothesis valid_different_pointers_inj:
- forall b1 ofs1 b2 ofs2 b1' delta1 b2' delta2,
- b1 <> b2 ->
- Mem.valid_pointer m1 b1 (Int.unsigned ofs1) = true ->
- Mem.valid_pointer m1 b2 (Int.unsigned ofs2) = true ->
- f b1 = Some (b1', delta1) ->
- f b2 = Some (b2', delta2) ->
- b1' <> b2' \/
- Int.unsigned (Int.add ofs1 (Int.repr delta1)) <> Int.unsigned (Int.add ofs2 (Int.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.
- 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_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.
- intros. destruct addr; simpl in H2; simpl; FuncInv; InvInject; TrivialExists.
- apply Values.Val.add_inject; auto.
- apply Values.Val.add_inject; auto. apply Values.Val.add_inject; auto.
- apply Values.Val.add_inject; auto. inv H5; simpl; auto.
- apply Values.Val.add_inject; auto. apply Values.Val.add_inject; auto. inv H3; simpl; auto.
- apply H; simpl; auto.
- apply Values.Val.add_inject; auto. apply H; simpl; auto.
- apply Values.Val.add_inject; auto. apply H; simpl; auto. inv H5; simpl; auto.
- apply Values.Val.add_inject; auto.
-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.
- inv H4; inv H2; simpl; auto. econstructor; eauto.
- rewrite Int.sub_add_l. auto.
- destruct (eq_block b1 b0); auto. subst. rewrite H1 in H0. inv H0. rewrite dec_eq_true.
- rewrite Int.sub_shifted. 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 i Int.iwordsize); auto.
- inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
- inv H4; simpl; auto. destruct (Int.ltu i Int.iwordsize); auto.
- inv H4; simpl in H1; try discriminate. simpl.
- destruct (Int.ltu i (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 i Int.iwordsize); auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto. destruct (Int.ltu i Int.iwordsize); auto.
- inv H2; simpl; auto. destruct (Int.ltu (Int.sub Int.iwordsize i) Int.iwordsize); auto.
- eapply eval_addressing_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; inv H2; simpl; auto.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
- subst v1. destruct (eval_condition c 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 (Int.unsigned ofs) = true ->
- Mem.valid_pointer m2 b2 (Int.unsigned (Int.add ofs (Int.repr delta))) = true.
-Proof.
- intros. inv H0. rewrite Int.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 (Int.unsigned ofs) = true ->
- Mem.weak_valid_pointer m2 b2 (Int.unsigned (Int.add ofs (Int.repr delta))) = true.
-Proof.
- intros. inv H0. rewrite Int.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 (Int.unsigned ofs) = true ->
- 0 <= Int.unsigned ofs + Int.unsigned (Int.repr delta) <= Int.max_unsigned.
-Proof.
- intros. inv H. rewrite Zplus_0_r. apply Int.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 (Int.unsigned ofs1) = true ->
- Mem.valid_pointer m1 b2 (Int.unsigned ofs2) = true ->
- Some(b1, 0) = Some (b1', delta1) ->
- Some(b2, 0) = Some (b2', delta2) ->
- b1' <> b2' \/
- Int.unsigned(Int.add ofs1 (Int.repr delta1)) <> Int.unsigned(Int.add ofs2 (Int.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 Int.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 Int.zero) addr vl1 = Some v1 ->
- exists v2,
- eval_addressing genv (Vptr sp2 Int.zero) (shift_stack_addressing (Int.repr delta) addr) vl2 = Some v2
- /\ Val.inject f v1 v2.
-Proof.
- intros.
- rewrite eval_shift_stack_addressing. simpl.
- eapply eval_addressing_inj with (sp1 := Vptr sp1 Int.zero); eauto.
- intros. apply symbol_address_inject.
-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 Int.zero) op vl1 m1 = Some v1 ->
- exists v2,
- eval_operation genv (Vptr sp2 Int.zero) (shift_stack_operation (Int.repr 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 Int.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.
-Qed.
-
-End EVAL_INJECT.
diff --git a/ia32/PrintOp.ml b/ia32/PrintOp.ml
deleted file mode 100644
index 2a80e3d4..00000000
--- a/ia32/PrintOp.ml
+++ /dev/null
@@ -1,128 +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 %ld" reg r1 (comparison_name c) (camlint_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 + %ld" reg r1 (camlint_of_coqint n)
- | Aindexed2 n, [r1; r2] ->
- fprintf pp "%a + %a + %ld" reg r1 reg r2 (camlint_of_coqint n)
- | Ascaled(sc,n), [r1] ->
- fprintf pp "%a * %ld + %ld" reg r1 (camlint_of_coqint sc) (camlint_of_coqint n)
- | Aindexed2scaled(sc, n), [r1; r2] ->
- fprintf pp "%a + %a * %ld + %ld" reg r1 reg r2 (camlint_of_coqint sc) (camlint_of_coqint n)
- | Aglobal(id, ofs), [] -> fprintf pp "%s + %ld" (extern_atom id) (camlint_of_coqint ofs)
- | Abased(id, ofs), [r1] -> fprintf pp "%s + %ld + %a" (extern_atom id) (camlint_of_coqint ofs) reg r1
- | Abasedscaled(sc,id, ofs), [r1] -> fprintf pp "%s + %ld + %a * %ld" (extern_atom id) (camlint_of_coqint ofs) reg r1 (camlint_of_coqint sc)
- | Ainstack ofs, [] -> fprintf pp "stack(%ld)" (camlint_of_coqint 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)
- | 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)
- | 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)
- | 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)
- | 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
- | 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
- | Onot, [r1] -> fprintf pp "not(%a)" reg r1
- | Omulhs, [r1;r2] -> fprintf pp "mulhs(%a,%a)" reg r1 reg r2
- | Omulhu, [r1;r2] -> fprintf pp "mulhu(%a,%a)" reg r1 reg r2
- | Ocmp c, args -> print_condition reg pp (c, args)
- | _ -> fprintf pp "<bad operator>"
-
-
diff --git a/ia32/SelectOp.vp b/ia32/SelectOp.vp
deleted file mode 100644
index bc331b9c..00000000
--- a/ia32/SelectOp.vp
+++ /dev/null
@@ -1,523 +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. *)
-(* *)
-(* *********************************************************************)
-
-(** 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.
-Require Import Integers.
-Require Import Floats.
-Require Import Op.
-Require Import CminorSel.
-
-Open Local Scope cminorsel_scope.
-
-(** ** Constants **)
-
-(** External oracle to determine whether a symbol is external and must
- be addressed through [Oaddrsymbol], or is local and can be addressed
- through [Olea Aglobal]. This is to accommodate MacOS X's limitations
- on references to data symbols imported from shared libraries. *)
-
-Parameter symbol_is_external: ident -> bool.
-
-Definition addrsymbol (id: ident) (ofs: int) :=
- if symbol_is_external id then
- if Int.eq ofs Int.zero
- then Eop (Oindirectsymbol id) Enil
- else Eop (Olea (Aindexed ofs)) (Eop (Oindirectsymbol id) Enil ::: Enil)
- else
- Eop (Olea (Aglobal id ofs)) Enil.
-
-Definition addrstack (ofs: int) :=
- Eop (Olea (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 n)) args
- | _ => Eop (Olea (Aindexed 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 (Int.add n1 n2))) (t1:::t2:::Enil)
- | Eop (Olea (Aindexed n1)) (t1:::Enil), Eop (Olea (Ascaled sc n2)) (t2:::Enil) =>
- Eop (Olea (Aindexed2scaled sc (Int.add n1 n2))) (t1:::t2:::Enil)
- | Eop (Olea (Ascaled sc n1)) (t1:::Enil), Eop (Olea (Aindexed n2)) (t2:::Enil) =>
- Eop (Olea (Aindexed2scaled sc (Int.add n1 n2))) (t2:::t1:::Enil)
- | Eop (Olea (Aindexed n1)) (t1:::Enil), Eop (Olea (Aglobal id ofs)) Enil =>
- Eop (Olea (Abased id (Int.add ofs n1))) (t1:::Enil)
- | Eop (Olea (Aglobal id ofs)) Enil, Eop (Olea (Aindexed n2)) (t2:::Enil) =>
- Eop (Olea (Abased id (Int.add ofs n2))) (t2:::Enil)
- | Eop (Olea (Ascaled sc n1)) (t1:::Enil), Eop (Olea (Aglobal id ofs)) Enil =>
- Eop (Olea (Abasedscaled sc id (Int.add ofs n1))) (t1:::Enil)
- | Eop (Olea (Aglobal id ofs)) Enil, Eop (Olea (Ascaled sc n2)) (t2:::Enil) =>
- Eop (Olea (Abasedscaled sc id (Int.add ofs 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 Int.zero)) (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.sub n1 n2) (Eop Osub (t1:::t2:::Enil))
- | Eop (Olea (Aindexed n1)) (t1:::Enil), t2 =>
- addimm n1 (Eop Osub (t1:::t2:::Enil))
- | t1, Eop (Olea (Aindexed n2)) (t2:::Enil) =>
- addimm (Int.neg 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.shl Int.one n) (Int.shl n1 n))) (t1:::Enil)
- else Eop (Oshlimm n) (e1:::Enil)
- | _ =>
- if shift_is_scale n
- then Eop (Olea (Ascaled (Int.shl Int.one n) Int.zero)) (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 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 => (addr, args)
- | _ => (Aindexed Int.zero, e:::Enil)
- end.
-
-(** ** Arguments of builtins *)
-
-Nondetfunction builtin_arg (e: expr) :=
- match e with
- | Eop (Ointconst n) Enil => BA_int n
- | Eop (Olea (Aglobal id ofs)) Enil => BA_addrglobal id ofs
- | Eop (Olea (Ainstack ofs)) Enil => BA_addrstack ofs
- | 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 bcfc13c9..00000000
--- a/ia32/SelectOpproof.v
+++ /dev/null
@@ -1,917 +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.
-
-(** * 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.
-
-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 Int.eq Int.eq_spec ofs Int.zero.
- subst. EvalOp.
- EvalOp. econstructor. EvalOp. simpl; eauto. econstructor. simpl.
- unfold Genv.symbol_address. destruct (Genv.find_symbol ge id); auto.
- simpl. rewrite Int.add_commut. rewrite Int.add_zero. auto.
- EvalOp.
-Qed.
-
-Theorem eval_addrstack:
- forall le ofs,
- exists v, eval_expr ge sp e m le (addrstack ofs) v /\ Val.lessdef (Val.add sp (Vint ofs)) v.
-Proof.
- intros. unfold addrstack. econstructor; split.
- EvalOp. simpl; eauto.
- auto.
-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. rewrite Int.add_zero. auto.
- case (addimm_match a); intros; InvEval; simpl.
- TrivialExists; simpl. rewrite Int.add_commut. auto.
- inv H0. simpl in H6. TrivialExists. simpl. eapply eval_offset_addressing_total; eauto.
- TrivialExists.
-Qed.
-
-Theorem eval_add: binary_constructor_sound add Val.add.
-Proof.
- 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 Val.add_permut_4. auto.
- subst. TrivialExists. simpl. rewrite Val.add_assoc. decEq; decEq. rewrite Val.add_permut. auto.
- subst. TrivialExists. simpl. rewrite Val.add_permut_4. rewrite <- Val.add_permut. rewrite <- Val.add_assoc. auto.
- subst. TrivialExists. simpl. rewrite Genv.shift_symbol_address.
- rewrite Val.add_commut. rewrite Val.add_assoc. decEq. decEq. apply Val.add_commut.
- subst. TrivialExists. simpl. rewrite Genv.shift_symbol_address. rewrite Val.add_assoc.
- decEq; decEq. apply Val.add_commut.
- subst. TrivialExists. simpl. rewrite Genv.shift_symbol_address. rewrite Val.add_commut.
- rewrite Val.add_assoc. decEq; decEq. apply Val.add_commut.
- subst. TrivialExists. simpl. rewrite Genv.shift_symbol_address.
- rewrite Val.add_assoc. decEq; decEq. apply Val.add_commut.
- 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.
-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.
- apply eval_addimm; EvalOp.
- subst. rewrite Val.sub_add_l. apply eval_addimm; EvalOp.
- subst. rewrite Val.sub_add_r. apply eval_addimm; EvalOp.
- TrivialExists.
-Qed.
-
-Theorem eval_negint: unary_constructor_sound negint (fun v => Val.sub Vzero v).
-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.
- destruct v1; simpl; auto. rewrite LT.
- rewrite Int.shl_mul. rewrite Int.mul_add_distr_l. rewrite (Int.shl_mul n1). 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.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).
- generalize (Int.one_bits_range n).
- destruct (Int.one_bits n).
- intros. TrivialExists.
- destruct l.
- intros. rewrite H1. simpl.
- rewrite Int.add_zero.
- replace (Vint (Int.shl Int.one i)) with (Val.shl Vone (Vint i)). rewrite Val.shl_mul.
- apply eval_shlimm. auto. simpl. rewrite H0; auto with coqlib.
- destruct l.
- intros. rewrite H1. simpl.
- exploit (eval_shlimm i (x :: le) (Eletvar 0) x). constructor; auto. intros [v1 [A1 B1]].
- exploit (eval_shlimm i0 (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 Int.add_zero.
- replace (Vint (Int.add (Int.shl Int.one i) (Int.shl Int.one i0)))
- with (Val.add (Val.shl Vone (Vint i)) (Val.shl Vone (Vint i0))).
- 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. repeat rewrite H0; auto with coqlib.
- intros. 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 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 v. unfold addressing; case (addressing_match a); intros; InvEval.
- inv H. exists vl; auto.
- exists (v :: nil); split. constructor; auto. constructor. subst; simpl. rewrite Int.add_zero; 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.
-- constructor.
-- simpl in H5. inv H5. constructor.
-- subst v. constructor; auto.
-- inv H. InvEval. simpl in H6; inv H6. constructor; auto.
-- inv H. InvEval. simpl 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 f19f036c..00000000
--- a/ia32/Stacklayout.v
+++ /dev/null
@@ -1,142 +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. *)
-(* *)
-(* *********************************************************************)
-
-(** Machine- and ABI-dependent layout information for activation records. *)
-
-Require Import Coqlib.
-Require Import Memory Separation.
-Require Import Bounds.
-
-(** 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.
-
-(** Computation of the frame environment from the bounds of the current
- function. *)
-
-Definition make_env (b: bounds) : frame_env :=
- let olink := 4 * b.(bound_outgoing) in (* back link *)
- let ocs := olink + 4 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)) 4 in (* return address *)
- let sz := oretaddr + 4 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) |}.
-
-(** Separation property *)
-
-Local Open Scope sep_scope.
-
-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 + 4)
- ** range sp (fe_ofs_retaddr fe) (fe_ofs_retaddr fe + 4)
- ** 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 (olink := 4 * b.(bound_outgoing)).
- set (ocs := olink + 4).
- 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)) 4).
- generalize b.(bound_local_pos) b.(bound_outgoing_pos) b.(bound_stack_data_pos); intros.
- assert (0 <= olink) by (unfold olink; omega).
- assert (olink + 4 <= 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. 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 (olink := 4 * b.(bound_outgoing)).
- set (ocs := olink + 4).
- 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)) 4).
- generalize b.(bound_local_pos) b.(bound_outgoing_pos) b.(bound_stack_data_pos); intros.
- assert (0 <= olink) by (unfold olink; omega).
- assert (olink + 4 <= 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)
- /\ (4 | fe_ofs_link fe)
- /\ (4 | fe_ofs_retaddr fe).
-Proof.
- intros; simpl.
- set (olink := 4 * b.(bound_outgoing)).
- set (ocs := olink + 4).
- 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)) 4).
- split. apply Zdivide_0.
- split. apply align_divides; omega.
- split. apply align_divides; omega.
- split. apply Z.divide_factor_l.
- apply align_divides; omega.
-Qed.
diff --git a/ia32/TargetPrinter.ml b/ia32/TargetPrinter.ml
deleted file mode 100644
index 4ffb701b..00000000
--- a/ia32/TargetPrinter.ml
+++ /dev/null
@@ -1,765 +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 IA32 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 int_reg_name = function
- | EAX -> "%eax" | EBX -> "%ebx" | ECX -> "%ecx" | EDX -> "%edx"
- | ESI -> "%esi" | EDI -> "%edi" | EBP -> "%ebp" | ESP -> "%esp"
-
-let int8_reg_name = function
- | EAX -> "%al" | EBX -> "%bl" | ECX -> "%cl" | EDX -> "%dl"
- | _ -> assert false
-
-let high_int8_reg_name = function
- | EAX -> "%ah" | EBX -> "%bh" | ECX -> "%ch" | EDX -> "%dh"
- | _ -> assert false
-
-let int16_reg_name = function
- | EAX -> "%ax" | EBX -> "%bx" | ECX -> "%cx" | EDX -> "%dx"
- | ESI -> "%si" | EDI -> "%di" | EBP -> "%bp" | ESP -> "%sp"
-
-let float_reg_name = function
- | XMM0 -> "%xmm0" | XMM1 -> "%xmm1" | XMM2 -> "%xmm2" | XMM3 -> "%xmm3"
- | XMM4 -> "%xmm4" | XMM5 -> "%xmm5" | XMM6 -> "%xmm6" | XMM7 -> "%xmm7"
-
-let ireg oc r = output_string oc (int_reg_name r)
-let ireg8 oc r = output_string oc (int8_reg_name r)
-let high_ireg8 oc r = output_string oc (high_int8_reg_name r)
-let ireg16 oc r = output_string oc (int16_reg_name r)
-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
-
-(* 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_ra: 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 cygwin *)
-module Cygwin_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 ".section .rdata,\"dr\"" else "COMM"
- | Section_string -> ".section .rdata,\"dr\""
- | Section_literal -> ".section .rdata,\"dr\""
- | Section_jumptable -> ".text"
- | Section_user(s, wr, ex) ->
- sprintf ".section \"%s\", \"%s\"\n"
- s (if ex then "xr" else if wr then "d" else "dr")
- | Section_debug_info _ -> ".section .debug_info,\"dr\""
- | Section_debug_loc -> ".section .debug_loc,\"dr\""
- | Section_debug_line _ -> ".section .debug_line,\"dr\""
- | Section_debug_abbrev -> ".section .debug_abbrev,\"dr\""
- | Section_debug_ranges -> ".section .debug_ranges,\"dr\""
- | Section_debug_str-> assert false (* Should not be used *)
-
- let stack_alignment = 8 (* minimum is 4, 8 is better for perfs *)
-
- let print_align oc n =
- fprintf oc " .align %d\n" n
-
- let print_mov_ra oc rd id =
- fprintf oc " movl $%a, %a\n" symbol id ireg rd
-
- let print_fun_info _ _ = ()
-
- let 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 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 = 8 (* minimum is 4, 8 is better for perfs *)
-
- let print_align oc n =
- fprintf oc " .align %d\n" n
-
- let print_mov_ra oc rd id =
- fprintf oc " movl $%a, %a\n" symbol id ireg 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_ra oc rd id =
- 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
-
- let print_fun_info _ _ = ()
-
- let print_var_info _ _ = ()
-
- let print_epilogue oc =
- 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
-
- 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;
- if ofs <> 0l then fprintf oc " + %ld" ofs
-
-
- let addressing oc (Addrmode(base, shift, cst)) =
- begin match cst with
- | Coq_inl n ->
- let n = camlint_of_coqint n in
- fprintf oc "%ld" n
- | Coq_inr(id, ofs) ->
- let ofs = camlint_of_coqint ofs in
- if ofs = 0l
- then symbol oc id
- else fprintf oc "(%a + %ld)" symbol id ofs
- 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 coqint sc
- | Some r1, Some(r2,sc) -> fprintf oc "(%a,%a,%a)" ireg r1 ireg r2 coqint sc
- end
-
- 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)
-
-(* Emit a align directive *)
-
- let need_masks = ref false
-
-(* Emit .file / .loc debugging directives *)
-
- let print_file_line oc file line =
- print_file_line oc comment file line
-
-
-
-(* Built-in functions *)
-
-(* 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. *)
-
-(* Printing of instructions *)
-
-(* Reminder on AT&T syntax: op source, dest *)
-
- let print_instruction oc = function
- (* Moves *)
- | Pmov_rr(rd, r1) ->
- fprintf oc " movl %a, %a\n" ireg r1 ireg rd
- | Pmov_ri(rd, n) ->
- fprintf oc " movl $%ld, %a\n" (camlint_of_coqint n) ireg rd
- | Pmov_ra(rd, id) ->
- print_mov_ra oc rd id
- | Pmov_rm(rd, a) | Pmov_rm_a(rd, a) ->
- fprintf oc " movl %a, %a\n" addressing a ireg rd
- | Pmov_mr(a, r1) | Pmov_mr_a(a, r1) ->
- fprintf oc " movl %a, %a\n" ireg 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, %a %s %.18g\n" label lbl 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, %a %s %.18g\n" label lbl 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
- | Pxchg_rr(r1, r2) ->
- fprintf oc " xchgl %a, %a\n" ireg r1 ireg r2
- (* 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 ireg rd
- | Pmovzb_rm(rd, a) ->
- fprintf oc " movzbl %a, %a\n" addressing a ireg rd
- | Pmovsb_rr(rd, r1) ->
- fprintf oc " movsbl %a, %a\n" ireg8 r1 ireg rd
- | Pmovsb_rm(rd, a) ->
- fprintf oc " movsbl %a, %a\n" addressing a ireg rd
- | Pmovzw_rr(rd, r1) ->
- fprintf oc " movzwl %a, %a\n" ireg16 r1 ireg rd
- | Pmovzw_rm(rd, a) ->
- fprintf oc " movzwl %a, %a\n" addressing a ireg rd
- | Pmovsw_rr(rd, r1) ->
- fprintf oc " movswl %a, %a\n" ireg16 r1 ireg rd
- | Pmovsw_rm(rd, a) ->
- fprintf oc " movswl %a, %a\n" addressing a ireg rd
- | 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 ireg rd
- | Pcvtsi2sd_fr(rd, r1) ->
- fprintf oc " cvtsi2sd %a, %a\n" ireg r1 freg rd
- | Pcvttss2si_rf(rd, r1) ->
- fprintf oc " cvttss2si %a, %a\n" freg r1 ireg rd
- | Pcvtsi2ss_fr(rd, r1) ->
- fprintf oc " cvtsi2ss %a, %a\n" ireg r1 freg rd
- (* Arithmetic and logical operations over integers *)
- | Plea(rd, a) ->
- fprintf oc " leal %a, %a\n" addressing a ireg rd
- | Pneg(rd) ->
- fprintf oc " negl %a\n" ireg rd
- | Psub_rr(rd, r1) ->
- fprintf oc " subl %a, %a\n" ireg r1 ireg rd
- | Pimul_rr(rd, r1) ->
- fprintf oc " imull %a, %a\n" ireg r1 ireg rd
- | Pimul_ri(rd, n) ->
- fprintf oc " imull $%a, %a\n" coqint n ireg rd
- | Pimul_r(r1) ->
- fprintf oc " imull %a\n" ireg r1
- | Pmul_r(r1) ->
- fprintf oc " mull %a\n" ireg r1
- | Pcltd ->
- fprintf oc " cltd\n"
- | Pdiv(r1) ->
- fprintf oc " divl %a\n" ireg r1
- | Pidiv(r1) ->
- fprintf oc " idivl %a\n" ireg r1
- | Pand_rr(rd, r1) ->
- fprintf oc " andl %a, %a\n" ireg r1 ireg rd
- | Pand_ri(rd, n) ->
- fprintf oc " andl $%a, %a\n" coqint n ireg rd
- | Por_rr(rd, r1) ->
- fprintf oc " orl %a, %a\n" ireg r1 ireg rd
- | Por_ri(rd, n) ->
- fprintf oc " orl $%a, %a\n" coqint n ireg rd
- | Pxor_r(rd) ->
- fprintf oc " xorl %a, %a\n" ireg rd ireg rd
- | Pxor_rr(rd, r1) ->
- fprintf oc " xorl %a, %a\n" ireg r1 ireg rd
- | Pxor_ri(rd, n) ->
- fprintf oc " xorl $%a, %a\n" coqint n ireg rd
- | Pnot(rd) ->
- fprintf oc " notl %a\n" ireg rd
- | Psal_rcl(rd) ->
- fprintf oc " sall %%cl, %a\n" ireg rd
- | Psal_ri(rd, n) ->
- fprintf oc " sall $%a, %a\n" coqint n ireg rd
- | Pshr_rcl(rd) ->
- fprintf oc " shrl %%cl, %a\n" ireg rd
- | Pshr_ri(rd, n) ->
- fprintf oc " shrl $%a, %a\n" coqint n ireg rd
- | Psar_rcl(rd) ->
- fprintf oc " sarl %%cl, %a\n" ireg rd
- | Psar_ri(rd, n) ->
- fprintf oc " sarl $%a, %a\n" coqint n ireg rd
- | Pshld_ri(rd, r1, n) ->
- fprintf oc " shldl $%a, %a, %a\n" coqint n ireg r1 ireg rd
- | Pror_ri(rd, n) ->
- fprintf oc " rorl $%a, %a\n" coqint n ireg rd
- | Pcmp_rr(r1, r2) ->
- fprintf oc " cmpl %a, %a\n" ireg r2 ireg r1
- | Pcmp_ri(r1, n) ->
- fprintf oc " cmpl $%a, %a\n" coqint n ireg r1
- | Ptest_rr(r1, r2) ->
- fprintf oc " testl %a, %a\n" ireg r2 ireg r1
- | Ptest_ri(r1, n) ->
- fprintf oc " testl $%a, %a\n" coqint n ireg 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 ireg 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, %a\n" raw_symbol "__negd_mask" freg rd
- | Pabsd (rd) ->
- need_masks := true;
- fprintf oc " andpd %a, %a\n" raw_symbol "__absd_mask" 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, %a\n" raw_symbol "__negs_mask" freg rd
- | Pabss (rd) ->
- need_masks := true;
- fprintf oc " andpd %a, %a\n" raw_symbol "__abss_mask" 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) ->
- assert (not sg.sig_cc.cc_structret);
- fprintf oc " jmp %a\n" symbol f
- | Pjmp_r(r, sg) ->
- assert (not sg.sig_cc.cc_structret);
- 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
- fprintf oc " jmp *%a(, %a, 4)\n" label l ireg r;
- jumptables := (l, tbl) :: !jumptables
- | Pcall_s(f, sg) ->
- fprintf oc " call %a\n" symbol f;
- if sg.sig_cc.cc_structret then
- fprintf oc " pushl %%eax\n"
- | Pcall_r(r, sg) ->
- fprintf oc " call *%a\n" ireg r;
- if sg.sig_cc.cc_structret then
- fprintf oc " pushl %%eax\n"
- | Pret ->
- if (!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 *)
- | Padc_ri (res,n) ->
- fprintf oc " adcl $%ld, %a\n" (camlint_of_coqint n) ireg res;
- | Padc_rr (res,a1) ->
- fprintf oc " adcl %a, %a\n" ireg a1 ireg res;
- | Padd_ri (res,n) ->
- fprintf oc " addl $%ld, %a\n" (camlint_of_coqint n) ireg res
- | Padd_rr (res,a1) ->
- fprintf oc " addl %a, %a\n" ireg a1 ireg res;
- | Padd_mi (addr,n) ->
- fprintf oc " addl $%ld, %a\n" (camlint_of_coqint n) addressing addr
- | Pbsf (res,a1) ->
- fprintf oc " bsfl %a, %a\n" ireg a1 ireg res
- | Pbsr (res,a1) ->
- fprintf oc " bsrl %a, %a\n" ireg a1 ireg res
- | Pbswap res ->
- fprintf oc " bswap %a\n" ireg 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
- | Pmovq_mr(a, rs) ->
- fprintf oc " movq %a, %a\n" freg rs addressing a
- | Pmovq_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"
- | Psbb_rr (res,a1) ->
- fprintf oc " sbbl %a, %a\n" ireg a1 ireg res
- | Psqrtsd (res,a1) ->
- fprintf oc " sqrtsd %a, %a\n" freg a1 freg res
- | Psub_ri (res,n) ->
- fprintf oc " subl $%ld, %a\n" (camlint_of_coqint n) ireg 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 oc (lbl, tbl) =
- fprintf oc "%a:" label lbl;
- List.iter
- (fun l -> fprintf oc " .long %a\n" label (transl_label l))
- tbl in
- if !jumptables <> [] then begin
- section oc jmptbl;
- print_align oc 4;
- List.iter (print_jumptable oc) !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 %s\n" (Z.to_string n)
- | Init_addrof(symb, ofs) ->
- fprintf oc " .long %a\n"
- symbol_offset (symb, camlint_of_coqint 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)
- | "cygwin" -> (module Cygwin_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 ad18c4f6..00000000
--- a/ia32/ValueAOp.v
+++ /dev/null
@@ -1,194 +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 Coqlib.
-Require Import Compopts.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Op.
-Require Import ValueDomain.
-Require Import RTL.
-
-(** Value analysis for IA32 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)
- | 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 (addr: addressing) (vl: list aval): aval :=
- match addr, vl with
- | Aindexed n, v1::nil => add v1 (I n)
- | Aindexed2 n, v1::v2::nil => add (add v1 v2) (I n)
- | Ascaled sc ofs, v1::nil => add (mul v1 (I sc)) (I ofs)
- | Aindexed2scaled sc ofs, v1::v2::nil => add v1 (add (mul v2 (I sc)) (I 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 sc))
- | Ainstack ofs, nil => Ptr(Stk ofs)
- | _, _ => Vbot
- end.
-
-Definition eval_static_operation (op: operation) (vl: list aval): aval :=
- match op, vl with
- | Omove, v1::nil => v1
- | Ointconst n, nil => I 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 Int.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 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
- | Omakelong, v1::v2::nil => longofwords v1 v2
- | Olowlong, v1::nil => loword v1
- | Ohighlong, v1::nil => hiword 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
- | _ => idtac
- end.
-
-Theorem eval_static_addressing_sound:
- forall addr vargs vres aargs,
- eval_addressing ge (Vptr sp Int.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 addr; InvHyps; eauto with va.
- rewrite Int.add_zero_l; auto with va.
-Qed.
-
-Theorem eval_static_operation_sound:
- forall op vargs m vres aargs,
- eval_operation ge (Vptr sp Int.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_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 3c6ee2e0..00000000
--- a/ia32/extractionMachdep.v
+++ /dev/null
@@ -1,20 +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 IA32 port *)
-
-Require SelectOp.
-
-(* SelectOp *)
-
-Extract Constant SelectOp.symbol_is_external =>
- "fun id -> Configuration.system = ""macosx"" && C2C.atom_is_extern id".