aboutsummaryrefslogtreecommitdiffstats
path: root/riscV
diff options
context:
space:
mode:
Diffstat (limited to 'riscV')
-rw-r--r--riscV/Archi.v2
-rw-r--r--riscV/Asm.v36
-rw-r--r--riscV/Asmexpand.ml42
-rw-r--r--riscV/Asmgen.v342
-rw-r--r--riscV/Asmgenproof.v50
-rw-r--r--riscV/Asmgenproof1.v479
-rw-r--r--riscV/Builtins1.v27
-rw-r--r--riscV/CBuiltins.ml8
-rw-r--r--riscV/CSE2deps.v35
-rw-r--r--riscV/CSE2depsproof.v147
-rw-r--r--riscV/ConstpropOpproof.v152
-rw-r--r--riscV/DuplicateOpcodeHeuristic.ml41
-rw-r--r--riscV/ExpansionOracle.ml1068
-rw-r--r--riscV/ExtValues.v123
-rw-r--r--riscV/Machregsaux.ml5
-rw-r--r--riscV/Machregsaux.mli2
-rw-r--r--riscV/NeedOp.v60
-rw-r--r--riscV/Op.v1054
-rw-r--r--riscV/OpWeights.ml168
l---------riscV/PrepassSchedulingOracle.ml1
l---------riscV/PrepassSchedulingOracleDeps.ml1
-rw-r--r--riscV/PrintOp.ml96
-rw-r--r--riscV/RTLpathSE_simplify.v2092
-rw-r--r--riscV/SelectLong.vp2
-rw-r--r--riscV/SelectLongproof.v55
-rw-r--r--riscV/SelectOp.vp48
-rw-r--r--riscV/SelectOpproof.v221
-rw-r--r--riscV/TargetPrinter.ml6
-rw-r--r--riscV/ValueAOp.v329
29 files changed, 6267 insertions, 425 deletions
diff --git a/riscV/Archi.v b/riscV/Archi.v
index 9e561ca8..96f34276 100644
--- a/riscV/Archi.v
+++ b/riscV/Archi.v
@@ -72,3 +72,5 @@ Global Opaque ptr64 big_endian splitlong
(** Whether to generate position-independent code or not *)
Parameter pic_code: unit -> bool.
+
+Definition has_notrap_loads := false.
diff --git a/riscV/Asm.v b/riscV/Asm.v
index a47573a2..c80c6cc2 100644
--- a/riscV/Asm.v
+++ b/riscV/Asm.v
@@ -30,6 +30,7 @@ Require Import Smallstep.
Require Import Locations.
Require Stacklayout.
Require Import Conventions.
+Require ExtValues.
(** * Abstract syntax *)
@@ -62,10 +63,10 @@ Inductive freg: Type :=
| F24: freg | F25: freg | F26: freg | F27: freg
| F28: freg | F29: freg | F30: freg | F31: freg.
-Lemma ireg_eq: forall (x y: ireg), {x=y} + {x<>y}.
+Definition ireg_eq: forall (x y: ireg), {x=y} + {x<>y}.
Proof. decide equality. Defined.
-Lemma ireg0_eq: forall (x y: ireg0), {x=y} + {x<>y}.
+Definition ireg0_eq: forall (x y: ireg0), {x=y} + {x<>y}.
Proof. decide equality. apply ireg_eq. Defined.
Lemma freg_eq: forall (x y: freg), {x=y} + {x<>y}.
@@ -255,10 +256,10 @@ Inductive instruction : Type :=
(* floating point register move *)
| Pfmv (rd: freg) (rs: freg) (**r move *)
- | Pfmvxs (rd: ireg) (rs: freg) (**r move FP single to integer register *)
- | Pfmvsx (rd: freg) (rs: ireg) (**r move integer register to FP single *)
- | Pfmvxd (rd: ireg) (rs: freg) (**r move FP double to integer register *)
- | Pfmvdx (rd: freg) (rs: ireg) (**r move integer register to FP double *)
+ | Pfmvxs (rd: ireg) (rs: freg) (**r bitwise move FP single to integer register *)
+ | Pfmvxd (rd: ireg) (rs: freg) (**r bitwise move FP double to integer register *)
+ | Pfmvsx (rd: freg) (rs: ireg) (**r bitwise move integer register to FP single *)
+ | Pfmvdx (rd: freg) (rs: ireg) (**r bitwise move integer register to FP double*)
(* 32-bit (single-precision) floating point *)
| Pfls (rd: freg) (ra: ireg) (ofs: offset) (**r load float *)
@@ -347,6 +348,7 @@ Inductive instruction : Type :=
| Pbtbl (r: ireg) (tbl: list label) (**r N-way branch through a jump table *)
| Pbuiltin: external_function -> list (builtin_arg preg)
-> builtin_res preg -> instruction (**r built-in function (pseudo) *)
+ | Pselectl (rd: ireg) (rb: ireg0) (rt: ireg0) (rf: ireg0)
| Pnop : instruction. (**r nop instruction *)
@@ -920,6 +922,17 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
Next (nextinstr (rs#d <- (Val.floatofsingle rs#s))) m
| Pfcvtsd d s =>
Next (nextinstr (rs#d <- (Val.singleoffloat rs#s))) m
+
+ | Pfmvxs d s =>
+ Next (nextinstr (rs#d <- (ExtValues.bits_of_single rs#s))) m
+ | Pfmvxd d s =>
+ Next (nextinstr (rs#d <- (ExtValues.bits_of_float rs#s))) m
+
+ | Pfmvsx d s =>
+ Next (nextinstr (rs#d <- (ExtValues.single_of_bits rs#s))) m
+ | Pfmvdx d s =>
+ Next (nextinstr (rs#d <- (ExtValues.float_of_bits rs#s))) m
+
(** Pseudo-instructions *)
| Pallocframe sz pos =>
@@ -942,6 +955,10 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
| _ => Stuck
end
end
+ | Pselectl rd rb rt rf =>
+ Next (nextinstr (rs#rd <- (ExtValues.select01_long
+ (rs###rb) (rs###rt) (rs###rf)))
+ #X31 <- Vundef) m
| Plabel lbl =>
Next (nextinstr rs) m
| Ploadsymbol rd s ofs =>
@@ -965,16 +982,12 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
end
| Pbuiltin ef args res =>
Stuck (**r treated specially below *)
+ | Pnop => Next (nextinstr rs) m (**r Pnop is used by an oracle during expansion *)
(** The following instructions and directives are not generated directly by Asmgen,
so we do not model them. *)
| Pfence
- | Pfmvxs _ _
- | Pfmvsx _ _
- | Pfmvxd _ _
- | Pfmvdx _ _
-
| Pfmins _ _ _
| Pfmaxs _ _ _
| Pfsqrts _ _
@@ -990,7 +1003,6 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
| Pfmsubd _ _ _ _
| Pfnmaddd _ _ _ _
| Pfnmsubd _ _ _ _
- | Pnop
=> Stuck
end.
diff --git a/riscV/Asmexpand.ml b/riscV/Asmexpand.ml
index dc0ec184..50dc20be 100644
--- a/riscV/Asmexpand.ml
+++ b/riscV/Asmexpand.ml
@@ -657,9 +657,49 @@ let expand_builtin_inline name args res =
raise (Error ("unrecognized builtin " ^ name))
(* Expansion of instructions *)
-
+
let expand_instruction instr =
match instr with
+ | Pselectl(rd, rb, rt, rf) ->
+ if not Archi.ptr64
+ then failwith "Pselectl not available on RV32, only on RV64"
+ else
+ if ireg0_eq rt rf then
+ begin
+ if ireg0_eq (X rd) rt then
+ begin
+ end
+ else
+ begin
+ emit (Paddl(rd, X0, rt))
+ end
+ end
+ else
+ if (ireg0_eq (X rd) rt) then
+ begin
+ emit (Psubl(X31, X0, rb));
+ emit (Pandl(X31, X X31, rt));
+ emit (Paddil(rd, rb, Int64.mone));
+ emit (Pandl(rd, X rd, rf));
+ emit (Porl(rd, X rd, X X31))
+ end
+ else
+ if (ireg0_eq (X rd) rf) then
+ begin
+ emit (Paddil(X31, rb, Int64.mone));
+ emit (Pandl(X31, X X31, rf));
+ emit (Psubl(rd, X0, rb));
+ emit (Pandl(rd, X rd, rt));
+ emit (Porl(rd, X rd, X X31))
+ end
+ else
+ begin
+ emit (Psubl(X31, X0, rb));
+ emit (Paddil(rd, rb, Int64.mone));
+ emit (Pandl(X31, X X31, rt));
+ emit (Pandl(rd, X rd, rf));
+ emit (Porl(rd, X rd, X X31))
+ end
| Pallocframe (sz, ofs) ->
let sg = get_current_function_sig() in
emit (Pmv (X30, X2));
diff --git a/riscV/Asmgen.v b/riscV/Asmgen.v
index a704ed74..3e84e950 100644
--- a/riscV/Asmgen.v
+++ b/riscV/Asmgen.v
@@ -25,6 +25,8 @@ Require Import Op Locations Mach Asm.
Local Open Scope string_scope.
Local Open Scope error_monad_scope.
+Definition time {A B: Type} (name: string) (f: A -> B) : A -> B := f.
+
(** The code generation functions take advantage of several
characteristics of the [Mach] code generated by earlier passes of the
compiler, mostly that argument and result registers are of the correct
@@ -201,8 +203,23 @@ Definition transl_cond_single (cmp: comparison) (rd: ireg) (fs1 fs2: freg) :=
| Cle => (Pfles rd fs1 fs2, true)
| Cgt => (Pflts rd fs2 fs1, true)
| Cge => (Pfles rd fs2 fs1, true)
- end.
-
+ end.
+
+(** Functions to select a special register according to the op "oreg" argument from RTL *)
+
+Definition apply_bin_oreg_ireg0 (optR: option oreg) (r1 r2: ireg0): (ireg0 * ireg0) :=
+ match optR with
+ | None => (r1, r2)
+ | Some X0_L => (X0, r1)
+ | Some X0_R => (r1, X0)
+ end.
+
+Definition get_oreg (optR: option oreg) (r: ireg0) :=
+ match optR with
+ | Some X0_L | Some X0_R => X0
+ | _ => r
+ end.
+
Definition transl_cbranch
(cond: condition) (args: list mreg) (lbl: label) (k: code) :=
match cond, args with
@@ -257,7 +274,72 @@ Definition transl_cbranch
| Cnotcompfs c, f1 :: f2 :: nil =>
do r1 <- freg_of f1; do r2 <- freg_of f2;
let (insn, normal) := transl_cond_single c X31 r1 r2 in
- OK (insn :: (if normal then Pbeqw X31 X0 lbl else Pbnew X31 X0 lbl) :: k)
+ OK (insn :: (if normal then Pbeqw X31 X0 lbl else Pbnew X31 X0 lbl) :: k)
+
+ | CEbeqw optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbeqw r1' r2' lbl :: k)
+ | CEbnew optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbnew r1' r2' lbl :: k)
+ | CEbequw optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbeqw r1' r2' lbl :: k)
+ | CEbneuw optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbnew r1' r2' lbl :: k)
+ | CEbltw optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbltw r1' r2' lbl :: k)
+ | CEbltuw optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbltuw r1' r2' lbl :: k)
+ | CEbgew optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbgew r1' r2' lbl :: k)
+ | CEbgeuw optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbgeuw r1' r2' lbl :: k)
+ | CEbeql optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbeql r1' r2' lbl :: k)
+ | CEbnel optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbnel r1' r2' lbl :: k)
+ | CEbequl optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbeql r1' r2' lbl :: k)
+ | CEbneul optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbnel r1' r2' lbl :: k)
+ | CEbltl optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbltl r1' r2' lbl :: k)
+ | CEbltul optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbltul r1' r2' lbl :: k)
+ | CEbgel optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbgel r1' r2' lbl :: k)
+ | CEbgeul optR, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ let (r1', r2') := apply_bin_oreg_ireg0 optR r1 r2 in
+ OK (Pbgeul r1' r2' lbl :: k)
| _, _ =>
Error(msg "Asmgen.transl_cond_branch")
end.
@@ -340,7 +422,7 @@ Definition transl_condimm_int64u (cmp: comparison) (rd: ireg) (r1: ireg) (n: int
match cmp with
| Clt => sltuimm64 rd r1 n k
| _ => loadimm64 X31 n (transl_cond_int64u cmp rd r1 X31 k)
- end.
+ end.
Definition transl_cond_op
(cond: condition) (rd: ireg) (args: list mreg) (k: code) :=
@@ -362,13 +444,13 @@ Definition transl_cond_op
OK (transl_cond_int64s c rd r1 r2 k)
| Ccomplu c, a1 :: a2 :: nil =>
do r1 <- ireg_of a1; do r2 <- ireg_of a2;
- OK (transl_cond_int64u c rd r1 r2 k)
+ OK (transl_cond_int64u c rd r1 r2 k)
| Ccomplimm c n, a1 :: nil =>
do r1 <- ireg_of a1;
OK (transl_condimm_int64s c rd r1 n k)
| Ccompluimm c n, a1 :: nil =>
do r1 <- ireg_of a1;
- OK (transl_condimm_int64u c rd r1 n k)
+ OK (transl_condimm_int64u c rd r1 n k)
| Ccompf c, f1 :: f2 :: nil =>
do r1 <- freg_of f1; do r2 <- freg_of f2;
let (insn, normal) := transl_cond_float c rd r1 r2 in
@@ -384,14 +466,14 @@ Definition transl_cond_op
| Cnotcompfs c, f1 :: f2 :: nil =>
do r1 <- freg_of f1; do r2 <- freg_of f2;
let (insn, normal) := transl_cond_single c rd r1 r2 in
- OK (insn :: if normal then Pxoriw rd rd Int.one :: k else k)
+ OK (insn :: if normal then Pxoriw rd rd Int.one :: k else k)
| _, _ =>
Error(msg "Asmgen.transl_cond_op")
- end.
+ 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) :=
match op, args with
@@ -503,11 +585,16 @@ Definition transl_op
OK (Psrliw rd rs n :: k)
| Oshrximm n, a1 :: nil =>
do rd <- ireg_of res; do rs <- ireg_of a1;
- OK (if Int.eq n Int.zero then Pmv rd rs :: k else
- Psraiw X31 rs (Int.repr 31) ::
- Psrliw X31 X31 (Int.sub Int.iwordsize n) ::
- Paddw X31 rs X31 ::
- Psraiw rd X31 n :: k)
+ OK (if Int.eq n Int.zero
+ then Pmv rd rs :: k
+ else if Int.eq n Int.one
+ then Psrliw X31 rs (Int.repr 31) ::
+ Paddw X31 rs X31 ::
+ Psraiw rd X31 Int.one :: k
+ else Psraiw X31 rs (Int.repr 31) ::
+ Psrliw X31 X31 (Int.sub Int.iwordsize n) ::
+ Paddw X31 rs X31 ::
+ Psraiw rd X31 n :: k)
(* [Omakelong], [Ohighlong] should not occur *)
| Olowlong, a1 :: nil =>
@@ -592,11 +679,16 @@ Definition transl_op
OK (Psrlil rd rs n :: k)
| Oshrxlimm n, a1 :: nil =>
do rd <- ireg_of res; do rs <- ireg_of a1;
- OK (if Int.eq n Int.zero then Pmv rd rs :: k else
- Psrail X31 rs (Int.repr 63) ::
- Psrlil X31 X31 (Int.sub Int64.iwordsize' n) ::
- Paddl X31 rs X31 ::
- Psrail rd X31 n :: k)
+ OK (if Int.eq n Int.zero
+ then Pmv rd rs :: k
+ else if Int.eq n Int.one
+ then Psrlil X31 rs (Int.repr 63) ::
+ Paddl X31 rs X31 ::
+ Psrail rd X31 Int.one :: k
+ else Psrail X31 rs (Int.repr 63) ::
+ Psrlil X31 X31 (Int.sub Int64.iwordsize' n) ::
+ Paddl X31 rs X31 ::
+ Psrail rd X31 n :: k)
| Onegf, a1 :: nil =>
do rd <- freg_of res; do rs <- freg_of a1;
@@ -692,11 +784,206 @@ Definition transl_op
| Osingleoflongu, a1 :: nil =>
do rd <- freg_of res; do rs <- ireg_of a1;
OK (Pfcvtslu rd rs :: k)
-
| Ocmp cmp, _ =>
do rd <- ireg_of res;
- transl_cond_op cmp rd args k
+ transl_cond_op cmp rd args k
+ (* Instructions expanded in RTL *)
+ | OEseqw optR, a1 :: a2 :: nil =>
+ do rd <- ireg_of res;
+ do rs1 <- ireg_of a1;
+ do rs2 <- ireg_of a2;
+ let (rs1', rs2') := apply_bin_oreg_ireg0 optR rs1 rs2 in
+ OK (Pseqw rd rs1' rs2' :: k)
+ | OEsnew optR, a1 :: a2 :: nil =>
+ do rd <- ireg_of res;
+ do rs1 <- ireg_of a1;
+ do rs2 <- ireg_of a2;
+ let (rs1', rs2') := apply_bin_oreg_ireg0 optR rs1 rs2 in
+ OK (Psnew rd rs1' rs2' :: k)
+ | OEsequw optR, a1 :: a2 :: nil =>
+ do rd <- ireg_of res;
+ do rs1 <- ireg_of a1;
+ do rs2 <- ireg_of a2;
+ let (rs1', rs2') := apply_bin_oreg_ireg0 optR rs1 rs2 in
+ OK (Pseqw rd rs1' rs2' :: k)
+ | OEsneuw optR, a1 :: a2 :: nil =>
+ do rd <- ireg_of res;
+ do rs1 <- ireg_of a1;
+ do rs2 <- ireg_of a2;
+ let (rs1', rs2') := apply_bin_oreg_ireg0 optR rs1 rs2 in
+ OK (Psnew rd rs1' rs2' :: k)
+ | OEsltw optR, a1 :: a2 :: nil =>
+ do rd <- ireg_of res;
+ do rs1 <- ireg_of a1;
+ do rs2 <- ireg_of a2;
+ let (rs1', rs2') := apply_bin_oreg_ireg0 optR rs1 rs2 in
+ OK (Psltw rd rs1' rs2' :: k)
+ | OEsltuw optR, a1 :: a2 :: nil =>
+ do rd <- ireg_of res;
+ do rs1 <- ireg_of a1;
+ do rs2 <- ireg_of a2;
+ let (rs1', rs2') := apply_bin_oreg_ireg0 optR rs1 rs2 in
+ OK (Psltuw rd rs1' rs2' :: k)
+ | OEsltiw n, a1 :: nil =>
+ do rd <- ireg_of res;
+ do rs <- ireg_of a1;
+ OK (Psltiw rd rs n :: k)
+ | OEsltiuw n, a1 :: nil =>
+ do rd <- ireg_of res;
+ do rs <- ireg_of a1;
+ OK (Psltiuw rd rs n :: k)
+ | OExoriw n, a1 :: nil =>
+ do rd <- ireg_of res;
+ do rs <- ireg_of a1;
+ OK (Pxoriw rd rs n :: k)
+ | OEluiw n, nil =>
+ do rd <- ireg_of res;
+ OK (Pluiw rd n :: k)
+ | OEaddiw optR n, nil =>
+ do rd <- ireg_of res;
+ let rs := get_oreg optR X0 in
+ OK (Paddiw rd rs n :: k)
+ | OEaddiw optR n, a1 :: nil =>
+ do rd <- ireg_of res;
+ do rs <- ireg_of a1;
+ let rs' := get_oreg optR rs in
+ OK (Paddiw rd rs' n :: k)
+ | OEandiw n, a1 :: nil =>
+ do rd <- ireg_of res;
+ do rs <- ireg_of a1;
+ OK (Pandiw rd rs n :: k)
+ | OEoriw n, a1 :: nil =>
+ do rd <- ireg_of res;
+ do rs <- ireg_of a1;
+ OK (Poriw rd rs n :: k)
+ | OEseql optR, a1 :: a2 :: nil =>
+ do rd <- ireg_of res;
+ do rs1 <- ireg_of a1;
+ do rs2 <- ireg_of a2;
+ let (rs1', rs2') := apply_bin_oreg_ireg0 optR rs1 rs2 in
+ OK (Pseql rd rs1' rs2' :: k)
+ | OEsnel optR, a1 :: a2 :: nil =>
+ do rd <- ireg_of res;
+ do rs1 <- ireg_of a1;
+ do rs2 <- ireg_of a2;
+ let (rs1', rs2') := apply_bin_oreg_ireg0 optR rs1 rs2 in
+ OK (Psnel rd rs1' rs2' :: k)
+ | OEsequl optR, a1 :: a2 :: nil =>
+ do rd <- ireg_of res;
+ do rs1 <- ireg_of a1;
+ do rs2 <- ireg_of a2;
+ let (rs1', rs2') := apply_bin_oreg_ireg0 optR rs1 rs2 in
+ OK (Pseql rd rs1' rs2' :: k)
+ | OEsneul optR, a1 :: a2 :: nil =>
+ do rd <- ireg_of res;
+ do rs1 <- ireg_of a1;
+ do rs2 <- ireg_of a2;
+ let (rs1', rs2') := apply_bin_oreg_ireg0 optR rs1 rs2 in
+ OK (Psnel rd rs1' rs2' :: k)
+ | OEsltl optR, a1 :: a2 :: nil =>
+ do rd <- ireg_of res;
+ do rs1 <- ireg_of a1;
+ do rs2 <- ireg_of a2;
+ let (rs1', rs2') := apply_bin_oreg_ireg0 optR rs1 rs2 in
+ OK (Psltl rd rs1' rs2' :: k)
+ | OEsltul optR, a1 :: a2 :: nil =>
+ do rd <- ireg_of res;
+ do rs1 <- ireg_of a1;
+ do rs2 <- ireg_of a2;
+ let (rs1', rs2') := apply_bin_oreg_ireg0 optR rs1 rs2 in
+ OK (Psltul rd rs1' rs2' :: k)
+ | OEsltil n, a1 :: nil =>
+ do rd <- ireg_of res;
+ do rs <- ireg_of a1;
+ OK (Psltil rd rs n :: k)
+ | OEsltiul n, a1 :: nil =>
+ do rd <- ireg_of res;
+ do rs <- ireg_of a1;
+ OK (Psltiul rd rs n :: k)
+ | OExoril n, a1 :: nil =>
+ do rd <- ireg_of res;
+ do rs <- ireg_of a1;
+ OK (Pxoril rd rs n :: k)
+ | OEluil n, nil =>
+ do rd <- ireg_of res;
+ OK (Pluil rd n :: k)
+ | OEaddil optR n, nil =>
+ do rd <- ireg_of res;
+ let rs := get_oreg optR X0 in
+ OK (Paddil rd rs n :: k)
+ | OEaddil optR n, a1 :: nil =>
+ do rd <- ireg_of res;
+ do rs <- ireg_of a1;
+ let rs' := get_oreg optR rs in
+ OK (Paddil rd rs' n :: k)
+ | OEandil n, a1 :: nil =>
+ do rd <- ireg_of res;
+ do rs <- ireg_of a1;
+ OK (Pandil rd rs n :: k)
+ | OEoril n, a1 :: nil =>
+ do rd <- ireg_of res;
+ do rs <- ireg_of a1;
+ OK (Poril rd rs n :: k)
+ | OEloadli n, nil =>
+ do rd <- ireg_of res;
+ OK (Ploadli rd n :: k)
+ | OEfeqd, f1 :: f2 :: nil =>
+ do rd <- ireg_of res;
+ do r1 <- freg_of f1;
+ do r2 <- freg_of f2;
+ OK (Pfeqd rd r1 r2 :: k)
+ | OEfltd, f1 :: f2 :: nil =>
+ do rd <- ireg_of res;
+ do r1 <- freg_of f1;
+ do r2 <- freg_of f2;
+ OK (Pfltd rd r1 r2 :: k)
+ | OEfled, f1 :: f2 :: nil =>
+ do rd <- ireg_of res;
+ do r1 <- freg_of f1;
+ do r2 <- freg_of f2;
+ OK (Pfled rd r1 r2 :: k)
+ | OEfeqs, f1 :: f2 :: nil =>
+ do rd <- ireg_of res;
+ do r1 <- freg_of f1;
+ do r2 <- freg_of f2;
+ OK (Pfeqs rd r1 r2 :: k)
+ | OEflts, f1 :: f2 :: nil =>
+ do rd <- ireg_of res;
+ do r1 <- freg_of f1;
+ do r2 <- freg_of f2;
+ OK (Pflts rd r1 r2 :: k)
+ | OEfles, f1 :: f2 :: nil =>
+ do rd <- ireg_of res;
+ do r1 <- freg_of f1;
+ do r2 <- freg_of f2;
+ OK (Pfles rd r1 r2 :: k)
+ | OEmayundef _, a1 :: a2 :: nil =>
+ do rd <- ireg_of res;
+ do r2 <- ireg_of a2;
+ if ireg_eq rd r2 then
+ OK (Pnop :: k)
+ else
+ OK (Pmv rd r2 :: k)
+
+ | Obits_of_single, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfmvxs rd rs :: k)
+ | Obits_of_float, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfmvxd rd rs :: k)
+ | Osingle_of_bits, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pfmvsx rd rs :: k)
+ | Ofloat_of_bits, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pfmvdx rd rs :: k)
+ | Oselectl, b::t::f::nil =>
+ do rd <- ireg_of res;
+ do rb <- ireg_of b;
+ do rt <- ireg_of t;
+ do rf <- ireg_of f;
+ OK (Pselectl rd rb rt rf :: k)
| _, _ =>
Error(msg "Asmgen.transl_op")
end.
@@ -770,9 +1057,13 @@ Definition transl_memory_access
Error(msg "Asmgen.transl_memory_access")
end.
-Definition transl_load (chunk: memory_chunk) (addr: addressing)
+Definition transl_load (trap : trapping_mode)
+ (chunk: memory_chunk) (addr: addressing)
(args: list mreg) (dst: mreg) (k: code) :=
- match chunk with
+ match trap with
+ | NOTRAP => Error (msg "Asmgen.transl_load non-trapping loads unsupported on Arm")
+ | TRAP =>
+ match chunk with
| Mint8signed =>
do r <- ireg_of dst;
transl_memory_access (Plb r) addr args k
@@ -799,6 +1090,7 @@ Definition transl_load (chunk: memory_chunk) (addr: addressing)
transl_memory_access (Pfld r) addr args k
| _ =>
Error (msg "Asmgen.transl_load")
+ end
end.
Definition transl_store (chunk: memory_chunk) (addr: addressing)
@@ -848,8 +1140,8 @@ Definition transl_instr (f: Mach.function) (i: Mach.instruction)
else loadind_ptr SP f.(fn_link_ofs) X30 c)
| Mop op args res =>
transl_op op args res k
- | Mload chunk addr args dst =>
- transl_load chunk addr args dst k
+ | Mload trap chunk addr args dst =>
+ transl_load trap chunk addr args dst k
| Mstore chunk addr args src =>
transl_store chunk addr args src k
| Mcall sig (inl r) =>
diff --git a/riscV/Asmgenproof.v b/riscV/Asmgenproof.v
index 798dad9f..e59c4535 100644
--- a/riscV/Asmgenproof.v
+++ b/riscV/Asmgenproof.v
@@ -173,7 +173,7 @@ Remark transl_cond_single_nolabel:
transl_cond_single c r1 r2 r3 = (insn, normal) -> nolabel insn.
Proof.
unfold transl_cond_single; intros. destruct c; inv H; exact I.
-Qed.
+ Qed.
Remark transl_cbranch_label:
forall cond args lbl k c,
@@ -211,7 +211,23 @@ Proof.
destruct normal; TailNoLabel.
- destruct (transl_cond_single c0 X31 x x0) as [insn normal] eqn:F; inv EQ2.
apply tail_nolabel_cons. eapply transl_cond_single_nolabel; eauto.
- destruct normal; TailNoLabel.
+ destruct normal; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
Qed.
Remark transl_cond_op_label:
@@ -238,7 +254,7 @@ Proof.
try (eapply tail_nolabel_trans; [apply loadimm32_label | TailNoLabel]).
apply opimm32_label; intros; exact I.
- destruct c0; simpl; TailNoLabel.
-- destruct c0; simpl; TailNoLabel.
+ - destruct c0; simpl; TailNoLabel.
- unfold transl_condimm_int64s.
destruct (Int64.eq n Int64.zero).
+ destruct c0; simpl; TailNoLabel.
@@ -254,7 +270,7 @@ Proof.
+ destruct c0; simpl; TailNoLabel.
+ destruct c0; simpl;
try (eapply tail_nolabel_trans; [apply loadimm64_label | TailNoLabel]).
- apply opimm64_label; intros; exact I.
+ apply opimm64_label; intros; exact I.
- destruct (transl_cond_float c0 r x x0) as [insn normal] eqn:F; inv EQ2.
apply tail_nolabel_cons. eapply transl_cond_float_nolabel; eauto.
destruct normal; TailNoLabel.
@@ -267,7 +283,7 @@ Proof.
- destruct (transl_cond_single c0 r x x0) as [insn normal] eqn:F; inv EQ2.
apply tail_nolabel_cons. eapply transl_cond_single_nolabel; eauto.
destruct normal; TailNoLabel.
-Qed.
+ Qed.
Remark transl_op_label:
forall op args r k c,
@@ -285,13 +301,25 @@ Opaque Int.eq.
- apply opimm32_label; intros; exact I.
- apply opimm32_label; intros; exact I.
- apply opimm32_label; intros; exact I.
-- destruct (Int.eq n Int.zero); TailNoLabel.
+- destruct (Int.eq n Int.zero); try destruct (Int.eq n Int.one); TailNoLabel.
- apply opimm64_label; intros; exact I.
- apply opimm64_label; intros; exact I.
- apply opimm64_label; intros; exact I.
- apply opimm64_label; intros; exact I.
-- destruct (Int.eq n Int.zero); TailNoLabel.
+- destruct (Int.eq n Int.zero); try destruct (Int.eq n Int.one); TailNoLabel.
- eapply transl_cond_op_label; eauto.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
+- destruct optR as [[]|]; simpl in *; TailNoLabel.
Qed.
Remark indexed_memory_access_label:
@@ -359,7 +387,7 @@ Proof.
- destruct ep. eapply loadind_label; eauto.
eapply tail_nolabel_trans. apply loadind_ptr_label. eapply loadind_label; eauto.
- eapply transl_op_label; eauto.
-- destruct m; monadInv H; eapply transl_memory_access_label; eauto; intros; exact I.
+- destruct t; (try discriminate); destruct m; monadInv H; eapply transl_memory_access_label; eauto; intros; exact I.
- destruct m; monadInv H; eapply transl_memory_access_label; eauto; intros; exact I.
- destruct s0; monadInv H; TailNoLabel.
- destruct s0; monadInv H; (eapply tail_nolabel_trans; [eapply make_epilogue_label|TailNoLabel]).
@@ -725,6 +753,12 @@ Local Transparent destroyed_by_op.
intros; auto with asmgen.
simpl; congruence.
+- (* Mload notrap *) (* isn't there a nicer way? *)
+ inv AT. simpl in *. unfold bind in *. destruct (transl_code _ _ _) in *; discriminate.
+
+- (* Mload notrap *)
+ inv AT. simpl in *. unfold bind in *. destruct (transl_code _ _ _) in *; discriminate.
+
- (* Mstore *)
assert (eval_addressing tge sp addr (map rs args) = Some a).
rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
diff --git a/riscV/Asmgenproof1.v b/riscV/Asmgenproof1.v
index 8195ce44..6abde89f 100644
--- a/riscV/Asmgenproof1.v
+++ b/riscV/Asmgenproof1.v
@@ -375,16 +375,16 @@ Proof.
rewrite <- Float32.cmp_swap. auto.
- simpl. f_equal. f_equal. f_equal. destruct (rs r2), (rs r1); auto. unfold Val.cmpfs, Val.cmpfs_bool.
rewrite <- Float32.cmp_swap. auto.
-Qed.
+ Qed.
-Remark branch_on_X31:
+(* TODO gourdinl UNUSUED ? Remark branch_on_X31:
forall normal lbl (rs: regset) m b,
rs#X31 = Val.of_bool (eqb normal b) ->
exec_instr ge fn (if normal then Pbnew X31 X0 lbl else Pbeqw X31 X0 lbl) rs m =
eval_branch fn lbl rs m (Some b).
Proof.
intros. destruct normal; simpl; rewrite H; simpl; destruct b; reflexivity.
-Qed.
+ Qed.*)
Ltac ArgsInv :=
repeat (match goal with
@@ -417,7 +417,7 @@ Proof.
{ apply eval_condition_lessdef with (map ms args) m; auto. eapply preg_vals; eauto. }
clear EVAL MEXT AG.
destruct cond; simpl in TRANSL; ArgsInv.
-- exists rs, (transl_cbranch_int32s c0 x x0 lbl).
+ - exists rs, (transl_cbranch_int32s c0 x x0 lbl).
intuition auto. constructor. apply transl_cbranch_int32s_correct; auto.
- exists rs, (transl_cbranch_int32u c0 x x0 lbl).
intuition auto. constructor. apply transl_cbranch_int32u_correct; auto.
@@ -492,7 +492,144 @@ Proof.
econstructor; econstructor.
split. constructor. apply exec_straight_one. eapply transl_cond_single_correct with (v := v); eauto. auto.
split. rewrite V; destruct normal, b; reflexivity.
- intros; Simpl.
+ intros; Simpl.
+
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32 in *;
+ destruct (rs x) eqn:EQRS; simpl in *; try congruence;
+ inv EQ2; eexists; eexists; eauto; split; constructor; auto;
+ simpl in *.
+ + rewrite EQRS;
+ assert (HB: (Int.eq Int.zero i) = b) by congruence.
+ rewrite HB; destruct b; simpl; auto.
+ + rewrite EQRS;
+ assert (HB: (Int.eq i Int.zero) = b) by congruence.
+ rewrite <- HB; destruct b; simpl; auto.
+ + rewrite EQRS;
+ destruct (rs x0); try congruence.
+ assert (HB: (Int.eq i i0) = b) by congruence.
+ rewrite <- HB; destruct b; simpl; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32 in *;
+ destruct (rs x) eqn:EQRS; simpl in *; try congruence;
+ inv EQ2; eexists; eexists; eauto; split; constructor; auto;
+ simpl in *.
+ + rewrite EQRS;
+ assert (HB: negb (Int.eq Int.zero i) = b) by congruence.
+ rewrite HB; destruct b; simpl; auto.
+ + rewrite EQRS;
+ assert (HB: negb (Int.eq i Int.zero) = b) by congruence.
+ rewrite <- HB; destruct b; simpl; auto.
+ + rewrite EQRS;
+ destruct (rs x0); try congruence.
+ assert (HB: negb (Int.eq i i0) = b) by congruence.
+ rewrite <- HB; destruct b; simpl; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32 in *; inv EQ2;
+ try (destruct (rs x); simpl in EVAL'; discriminate; fail);
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; try rewrite EVAL'; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32 in *; inv EQ2;
+ try (destruct (rs x); simpl in EVAL'; discriminate; fail);
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; try rewrite EVAL'; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32 in *; inv EQ2;
+ try (destruct (rs x); simpl in EVAL'; discriminate; fail);
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; try rewrite EVAL'; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32 in *; inv EQ2;
+ try (destruct (rs x); simpl in EVAL'; discriminate; fail);
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; try rewrite EVAL'; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32 in *; inv EQ2;
+ try (destruct (rs x); simpl in EVAL'; discriminate; fail);
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; try rewrite EVAL'; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32 in *; inv EQ2;
+ try (destruct (rs x); simpl in EVAL'; discriminate; fail);
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; try rewrite EVAL'; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32 in *; inv EQ2;
+ destruct (rs x) eqn:EQRS; simpl in *; try congruence;
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; auto.
+ + rewrite EQRS;
+ assert (HB: (Int64.eq Int64.zero i) = b) by congruence.
+ rewrite HB; destruct b; simpl; auto.
+ + rewrite EQRS;
+ assert (HB: (Int64.eq i Int64.zero) = b) by congruence.
+ rewrite <- HB; destruct b; simpl; auto.
+ + rewrite EQRS;
+ destruct (rs x0); try congruence.
+ assert (HB: (Int64.eq i i0) = b) by congruence.
+ rewrite <- HB; destruct b; simpl; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32 in *; inv EQ2;
+ destruct (rs x) eqn:EQRS; simpl in *; try congruence;
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; auto.
+ + rewrite EQRS;
+ assert (HB: negb (Int64.eq Int64.zero i) = b) by congruence.
+ rewrite HB; destruct b; simpl; auto.
+ + rewrite EQRS;
+ assert (HB: negb (Int64.eq i Int64.zero) = b) by congruence.
+ rewrite <- HB; destruct b; simpl; auto.
+ + rewrite EQRS;
+ destruct (rs x0); try congruence.
+ assert (HB: negb (Int64.eq i i0) = b) by congruence.
+ rewrite <- HB; destruct b; simpl; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32, zero64, Op.zero64 in *; inv EQ2;
+ try (destruct (rs x); simpl in EVAL'; discriminate; fail);
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; try rewrite EVAL'; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32, zero64, Op.zero64 in *; inv EQ2;
+ try (destruct (rs x); simpl in EVAL'; discriminate; fail);
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; try rewrite EVAL'; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32, zero64, Op.zero64 in *; inv EQ2;
+ try (destruct (rs x); simpl in EVAL'; discriminate; fail);
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; try rewrite EVAL'; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32, zero64, Op.zero64 in *; inv EQ2;
+ try (destruct (rs x); simpl in EVAL'; discriminate; fail);
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; try rewrite EVAL'; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32, zero64, Op.zero64 in *; inv EQ2;
+ try (destruct (rs x); simpl in EVAL'; discriminate; fail);
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; try rewrite EVAL'; auto.
+- destruct optR as [[]|];
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *;
+ unfold zero32, Op.zero32, zero64, Op.zero64 in *; inv EQ2;
+ try (destruct (rs x); simpl in EVAL'; discriminate; fail);
+ eexists; eexists; eauto; split; constructor;
+ simpl in *; try rewrite EVAL'; auto.
Qed.
Lemma transl_cbranch_correct_true:
@@ -684,18 +821,18 @@ Proof.
unfold Val.cmp; destruct (rs#r1); simpl; auto. rewrite B1.
unfold Int.lt. rewrite zlt_false. auto.
change (Int.signed (Int.repr Int.max_signed)) with Int.max_signed.
- generalize (Int.signed_range i); lia.
+ generalize (Int.signed_range i); omega.
* exploit (opimm32_correct Psltw Psltiw (Val.cmp Clt)); eauto. intros (rs1 & A1 & B1 & C1).
exists rs1; split. eexact A1. split; auto.
rewrite B1. unfold Val.cmp; simpl; destruct (rs#r1); simpl; auto.
unfold Int.lt. replace (Int.signed (Int.add n Int.one)) with (Int.signed n + 1).
destruct (zlt (Int.signed n) (Int.signed i)).
- rewrite zlt_false by lia. auto.
- rewrite zlt_true by lia. auto.
+ rewrite zlt_false by omega. auto.
+ rewrite zlt_true by omega. auto.
rewrite Int.add_signed. symmetry; apply Int.signed_repr.
assert (Int.signed n <> Int.max_signed).
{ red; intros E. elim H1. rewrite <- (Int.repr_signed n). rewrite E. auto. }
- generalize (Int.signed_range n); lia.
+ generalize (Int.signed_range n); omega.
+ apply DFL.
+ apply DFL.
Qed.
@@ -782,18 +919,18 @@ Proof.
unfold Val.cmpl; destruct (rs#r1); simpl; auto. rewrite B1.
unfold Int64.lt. rewrite zlt_false. auto.
change (Int64.signed (Int64.repr Int64.max_signed)) with Int64.max_signed.
- generalize (Int64.signed_range i); lia.
+ generalize (Int64.signed_range i); omega.
* exploit (opimm64_correct Psltl Psltil (fun v1 v2 => Val.maketotal (Val.cmpl Clt v1 v2))); eauto. intros (rs1 & A1 & B1 & C1).
exists rs1; split. eexact A1. split; auto.
rewrite B1. unfold Val.cmpl; simpl; destruct (rs#r1); simpl; auto.
unfold Int64.lt. replace (Int64.signed (Int64.add n Int64.one)) with (Int64.signed n + 1).
destruct (zlt (Int64.signed n) (Int64.signed i)).
- rewrite zlt_false by lia. auto.
- rewrite zlt_true by lia. auto.
+ rewrite zlt_false by omega. auto.
+ rewrite zlt_true by omega. auto.
rewrite Int64.add_signed. symmetry; apply Int64.signed_repr.
assert (Int64.signed n <> Int64.max_signed).
{ red; intros E. elim H1. rewrite <- (Int64.repr_signed n). rewrite E. auto. }
- generalize (Int64.signed_range n); lia.
+ generalize (Int64.signed_range n); omega.
+ apply DFL.
+ apply DFL.
Qed.
@@ -830,7 +967,7 @@ Proof.
+ apply DFL.
+ apply DFL.
+ apply DFL.
-Qed.
+ Qed.
Lemma transl_cond_op_correct:
forall cond rd args k c rs m,
@@ -858,7 +995,7 @@ Proof.
exists rs'; repeat split; eauto. rewrite MKTOT; eauto.
+ (* cmplu *)
exploit transl_cond_int64u_correct; eauto. intros (rs' & A & B & C).
- exists rs'; repeat split; eauto. rewrite B, MKTOT; eauto.
+ exists rs'; repeat split; eauto. rewrite B, MKTOT; eauto.
+ (* cmplimm *)
exploit transl_condimm_int64s_correct; eauto. instantiate (1 := x); eauto with asmgen.
intros (rs' & A & B & C).
@@ -866,7 +1003,7 @@ Proof.
+ (* cmpluimm *)
exploit transl_condimm_int64u_correct; eauto. instantiate (1 := x); eauto with asmgen.
intros (rs' & A & B & C).
- exists rs'; repeat split; eauto. rewrite MKTOT; eauto.
+ exists rs'; repeat split; eauto. rewrite MKTOT; eauto.
+ (* cmpf *)
destruct (transl_cond_float c0 rd x x0) as [insn normal] eqn:TR.
fold (Val.cmpf c0 (rs x) (rs x0)).
@@ -923,7 +1060,7 @@ Proof.
* econstructor; split.
apply exec_straight_one. eapply transl_cond_single_correct with (v := Val.notbool v); eauto. auto.
split; intros; Simpl.
-Qed.
+ Qed.
(** Some arithmetic properties. *)
@@ -964,126 +1101,226 @@ Proof.
Opaque Int.eq.
intros until c; intros TR EV.
unfold transl_op in TR; destruct op; ArgsInv; simpl in EV; SimplEval EV; try TranslOpSimpl.
-- (* move *)
- destruct (preg_of res), (preg_of m0); inv TR; TranslOpSimpl.
-- (* intconst *)
- exploit loadimm32_correct; eauto. intros (rs' & A & B & C).
- exists rs'; split; eauto. rewrite B; auto with asmgen.
-- (* longconst *)
- exploit loadimm64_correct; eauto. intros (rs' & A & B & C).
- exists rs'; split; eauto. rewrite B; auto with asmgen.
-- (* floatconst *)
- destruct (Float.eq_dec n Float.zero).
-+ subst n. econstructor; split.
- apply exec_straight_one. simpl; eauto. auto.
- split; intros; Simpl.
-+ econstructor; split.
- apply exec_straight_one. simpl; eauto. auto.
- split; intros; Simpl.
-- (* singleconst *)
- destruct (Float32.eq_dec n Float32.zero).
-+ subst n. econstructor; split.
- apply exec_straight_one. simpl; eauto. auto.
- split; intros; Simpl.
-+ econstructor; split.
- apply exec_straight_one. simpl; eauto. auto.
- split; intros; Simpl.
-- (* addrsymbol *)
- destruct (Archi.pic_code tt && negb (Ptrofs.eq ofs Ptrofs.zero)).
-+ set (rs1 := nextinstr (rs#x <- (Genv.symbol_address ge id Ptrofs.zero))).
- exploit (addptrofs_correct x x ofs k rs1 m); eauto with asmgen.
- intros (rs2 & A & B & C).
- exists rs2; split.
- apply exec_straight_step with rs1 m; auto.
- split. replace ofs with (Ptrofs.add Ptrofs.zero ofs) by (apply Ptrofs.add_zero_l).
- rewrite Genv.shift_symbol_address.
- replace (rs1 x) with (Genv.symbol_address ge id Ptrofs.zero) in B by (unfold rs1; Simpl).
- exact B.
- intros. rewrite C by eauto with asmgen. unfold rs1; Simpl.
-+ TranslOpSimpl.
-- (* stackoffset *)
- exploit addptrofs_correct. instantiate (1 := X2); auto with asmgen. intros (rs' & A & B & C).
- exists rs'; split; eauto. auto with asmgen.
-- (* cast8signed *)
- econstructor; split.
+ (* move *)
+ { destruct (preg_of res), (preg_of m0); inv TR; TranslOpSimpl. }
+ (* intconst *)
+ { exploit loadimm32_correct; eauto. intros (rs' & A & B & C).
+ exists rs'; split; eauto. rewrite B; auto with asmgen. }
+ (* longconst *)
+ { exploit loadimm64_correct; eauto. intros (rs' & A & B & C).
+ exists rs'; split; eauto. rewrite B; auto with asmgen. }
+ (* floatconst *)
+ { destruct (Float.eq_dec n Float.zero).
+ + subst n. econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl.
+ + econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl. }
+ (* singleconst *)
+ { destruct (Float32.eq_dec n Float32.zero).
+ + subst n. econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl.
+ + econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl. }
+ (* addrsymbol *)
+ { destruct (Archi.pic_code tt && negb (Ptrofs.eq ofs Ptrofs.zero)).
+ + set (rs1 := nextinstr (rs#x <- (Genv.symbol_address ge id Ptrofs.zero))).
+ exploit (addptrofs_correct x x ofs k rs1 m); eauto with asmgen.
+ intros (rs2 & A & B & C).
+ exists rs2; split.
+ apply exec_straight_step with rs1 m; auto.
+ split. replace ofs with (Ptrofs.add Ptrofs.zero ofs) by (apply Ptrofs.add_zero_l).
+ rewrite Genv.shift_symbol_address.
+ replace (rs1 x) with (Genv.symbol_address ge id Ptrofs.zero) in B by (unfold rs1; Simpl).
+ exact B.
+ intros. rewrite C by eauto with asmgen. unfold rs1; Simpl.
+ + TranslOpSimpl. }
+ (* stackoffset *)
+ { exploit addptrofs_correct. instantiate (1 := X2); auto with asmgen. intros (rs' & A & B & C).
+ exists rs'; split; eauto. auto with asmgen. }
+ (* cast8signed *)
+ { econstructor; split.
eapply exec_straight_two. simpl;eauto. simpl;eauto. auto. auto.
split; intros; Simpl.
assert (A: Int.ltu (Int.repr 24) Int.iwordsize = true) by auto.
destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A.
- apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity.
-- (* cast16signed *)
- econstructor; split.
+ apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. }
+ (* cast16signed *)
+ { econstructor; split.
eapply exec_straight_two. simpl;eauto. simpl;eauto. auto. auto.
split; intros; Simpl.
assert (A: Int.ltu (Int.repr 16) Int.iwordsize = true) by auto.
destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A.
- apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity.
-- (* addimm *)
- exploit (opimm32_correct Paddw Paddiw Val.add); auto. instantiate (1 := x0); eauto with asmgen.
+ apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. }
+ (* addimm *)
+ { exploit (opimm32_correct Paddw Paddiw Val.add); auto. instantiate (1 := x0); eauto with asmgen.
intros (rs' & A & B & C).
- exists rs'; split; eauto. rewrite B; auto with asmgen.
-- (* andimm *)
- exploit (opimm32_correct Pandw Pandiw Val.and); auto. instantiate (1 := x0); eauto with asmgen.
+ exists rs'; split; eauto. rewrite B; auto with asmgen. }
+ (* andimm *)
+ { exploit (opimm32_correct Pandw Pandiw Val.and); auto. instantiate (1 := x0); eauto with asmgen.
intros (rs' & A & B & C).
- exists rs'; split; eauto. rewrite B; auto with asmgen.
-- (* orimm *)
+ exists rs'; split; eauto. rewrite B; auto with asmgen. }
+ (* orimm *)
exploit (opimm32_correct Porw Poriw Val.or); auto. instantiate (1 := x0); eauto with asmgen.
+ { intros (rs' & A & B & C).
+ exists rs'; split; eauto. rewrite B; auto with asmgen. }
+ (* xorimm *)
+ { exploit (opimm32_correct Pxorw Pxoriw Val.xor); auto. instantiate (1 := x0); eauto with asmgen.
intros (rs' & A & B & C).
- exists rs'; split; eauto. rewrite B; auto with asmgen.
-- (* xorimm *)
- exploit (opimm32_correct Pxorw Pxoriw Val.xor); auto. instantiate (1 := x0); eauto with asmgen.
- intros (rs' & A & B & C).
- exists rs'; split; eauto. rewrite B; auto with asmgen.
-- (* shrximm *)
- clear H. exploit Val.shrx_shr_2; eauto. intros E; subst v; clear EV.
- destruct (Int.eq n Int.zero).
-+ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
- split; intros; Simpl.
-+ change (Int.repr 32) with Int.iwordsize. set (n' := Int.sub Int.iwordsize n).
- econstructor; split.
- eapply exec_straight_step. simpl; reflexivity. auto.
- eapply exec_straight_step. simpl; reflexivity. auto.
- eapply exec_straight_step. simpl; reflexivity. auto.
- apply exec_straight_one. simpl; reflexivity. auto.
- split; intros; Simpl.
-- (* longofintu *)
- econstructor; split.
+ exists rs'; split; eauto. rewrite B; auto with asmgen. }
+ (* shrximm *)
+ { destruct (Val.shrx (rs x0) (Vint n)) eqn:TOTAL; cbn.
+ {
+ exploit Val.shrx_shr_3; eauto. intros E; subst v.
+ destruct (Int.eq n Int.zero).
+ + econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl.
+ + destruct (Int.eq n Int.one).
+ * econstructor; split.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ apply exec_straight_one. simpl; reflexivity. auto.
+ split; intros; Simpl.
+ * change (Int.repr 32) with Int.iwordsize. set (n' := Int.sub Int.iwordsize n).
+ econstructor; split.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ apply exec_straight_one. simpl; reflexivity. auto.
+ split; intros; Simpl.
+ }
+ destruct (Int.eq n Int.zero).
+ + econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl.
+ + destruct (Int.eq n Int.one).
+ * econstructor; split.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ apply exec_straight_one. simpl; reflexivity. auto.
+ split; intros; Simpl.
+ * change (Int.repr 32) with Int.iwordsize. set (n' := Int.sub Int.iwordsize n).
+ econstructor; split.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ apply exec_straight_one. simpl; reflexivity. auto.
+ split; intros; Simpl. }
+ (* longofintu *)
+ { econstructor; split.
eapply exec_straight_three. simpl; eauto. simpl; eauto. simpl; eauto. auto. auto. auto.
split; intros; Simpl. destruct (rs x0); auto. simpl.
assert (A: Int.ltu (Int.repr 32) Int64.iwordsize' = true) by auto.
rewrite A; simpl. rewrite A. apply Val.lessdef_same. f_equal.
- rewrite cast32unsigned_from_cast32signed. apply Int64.zero_ext_shru_shl. compute; auto.
-- (* addlimm *)
- exploit (opimm64_correct Paddl Paddil Val.addl); auto. instantiate (1 := x0); eauto with asmgen.
+ rewrite cast32unsigned_from_cast32signed. apply Int64.zero_ext_shru_shl. compute; auto. }
+ (* addlimm *)
+ { exploit (opimm64_correct Paddl Paddil Val.addl); auto. instantiate (1 := x0); eauto with asmgen.
intros (rs' & A & B & C).
- exists rs'; split; eauto. rewrite B; auto with asmgen.
-- (* andimm *)
- exploit (opimm64_correct Pandl Pandil Val.andl); auto. instantiate (1 := x0); eauto with asmgen.
+ exists rs'; split; eauto. rewrite B; auto with asmgen. }
+ (* andimm *)
+ { exploit (opimm64_correct Pandl Pandil Val.andl); auto. instantiate (1 := x0); eauto with asmgen.
intros (rs' & A & B & C).
- exists rs'; split; eauto. rewrite B; auto with asmgen.
-- (* orimm *)
- exploit (opimm64_correct Porl Poril Val.orl); auto. instantiate (1 := x0); eauto with asmgen.
+ exists rs'; split; eauto. rewrite B; auto with asmgen. }
+ (* orimm *)
+ { exploit (opimm64_correct Porl Poril Val.orl); auto. instantiate (1 := x0); eauto with asmgen.
intros (rs' & A & B & C).
- exists rs'; split; eauto. rewrite B; auto with asmgen.
-- (* xorimm *)
- exploit (opimm64_correct Pxorl Pxoril Val.xorl); auto. instantiate (1 := x0); eauto with asmgen.
+ exists rs'; split; eauto. rewrite B; auto with asmgen. }
+ (* xorimm *)
+ { exploit (opimm64_correct Pxorl Pxoril Val.xorl); auto. instantiate (1 := x0); eauto with asmgen.
intros (rs' & A & B & C).
- exists rs'; split; eauto. rewrite B; auto with asmgen.
-- (* shrxlimm *)
- clear H. exploit Val.shrxl_shrl_2; eauto. intros E; subst v; clear EV.
- destruct (Int.eq n Int.zero).
-+ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
- split; intros; Simpl.
-+ change (Int.repr 64) with Int64.iwordsize'. set (n' := Int.sub Int64.iwordsize' n).
- econstructor; split.
- eapply exec_straight_step. simpl; reflexivity. auto.
- eapply exec_straight_step. simpl; reflexivity. auto.
- eapply exec_straight_step. simpl; reflexivity. auto.
- apply exec_straight_one. simpl; reflexivity. auto.
+ exists rs'; split; eauto. rewrite B; auto with asmgen. }
+ (* shrxlimm *)
+ { destruct (Val.shrxl (rs x0) (Vint n)) eqn:TOTAL.
+ {
+ exploit Val.shrxl_shrl_3; eauto. intros E; subst v.
+ destruct (Int.eq n Int.zero).
+ + econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl.
+ + destruct (Int.eq n Int.one).
+ * econstructor; split.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ apply exec_straight_one. simpl; reflexivity. auto.
+ split; intros; Simpl.
+
+ * change (Int.repr 64) with Int64.iwordsize'. set (n' := Int.sub Int64.iwordsize' n).
+ econstructor; split.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ apply exec_straight_one. simpl; reflexivity. auto.
+ split; intros; Simpl.
+ }
+ destruct (Int.eq n Int.zero).
+ + econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl.
+ + destruct (Int.eq n Int.one).
+ * econstructor; split.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ apply exec_straight_one. simpl; reflexivity. auto.
+ split; intros; Simpl.
+
+ * change (Int.repr 64) with Int64.iwordsize'. set (n' := Int.sub Int64.iwordsize' n).
+ econstructor; split.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ apply exec_straight_one. simpl; reflexivity. auto.
+ split; intros; Simpl. }
+ (* cond *)
+ { exploit transl_cond_op_correct; eauto. intros (rs' & A & B & C).
+ exists rs'; split. eexact A. eauto with asmgen. }
+ (* Expanded instructions from RTL *)
+ 9,10,19,20:
+ econstructor; split; try apply exec_straight_one; simpl; eauto;
+ split; intros; Simpl; try destruct (rs x0);
+ try rewrite Int64.add_commut;
+ try rewrite Int.add_commut; auto;
+ try rewrite Int64.and_commut;
+ try rewrite Int.and_commut; auto;
+ try rewrite Int64.or_commut;
+ try rewrite Int.or_commut; auto.
+ 1-16:
+ destruct optR as [[]|]; try discriminate;
+ unfold apply_bin_oreg_ireg0, apply_bin_oreg in *; try inv EQ3; try inv EQ2;
+ try destruct (Int.eq _ _) eqn:A; try inv H0;
+ try destruct (Int64.eq _ _) eqn:A; try inv H1;
+ econstructor; split; try apply exec_straight_one; simpl; eauto;
+ split; intros; Simpl;
+ try apply Int.same_if_eq in A; subst;
+ try apply Int64.same_if_eq in A; subst;
+ unfold get_sp;
+ try destruct (rs x0); auto;
+ try destruct (rs x1); auto;
+ try destruct (rs X2); auto;
+ try destruct Archi.ptr64 eqn:B;
+ try fold (Val.add (Vint Int.zero) (get_sp (rs X2)));
+ try fold (Val.addl (Vlong Int64.zero) (get_sp (rs X2)));
+ try rewrite Val.add_commut; auto;
+ try rewrite Val.addl_commut; auto;
+ try rewrite Int.add_commut; auto;
+ try rewrite Int64.add_commut; auto;
+ replace (Ptrofs.of_int Int.zero) with (Ptrofs.zero) by auto;
+ replace (Ptrofs.of_int64 Int64.zero) with (Ptrofs.zero) by auto;
+ try rewrite Ptrofs.add_zero; auto.
+ (* mayundef *)
+ { destruct (ireg_eq x x0); inv EQ2;
+ econstructor; split;
+ try apply exec_straight_one; simpl; eauto;
+ split; unfold eval_may_undef;
+ destruct mu eqn:EQMU; simpl; intros; Simpl; auto.
+ all:
+ destruct (rs (preg_of m0)) eqn:EQM0; simpl; auto;
+ destruct (rs x0); simpl; auto; Simpl;
+ try destruct (Int.ltu _ _); simpl;
+ Simpl; auto. }
+ (* select *)
+ { econstructor; split. apply exec_straight_one. simpl; eauto. auto.
split; intros; Simpl.
-- (* cond *)
- exploit transl_cond_op_correct; eauto. intros (rs' & A & B & C).
- exists rs'; split. eexact A. eauto with asmgen.
+ apply Val.lessdef_normalize. }
Qed.
(** Memory accesses *)
@@ -1302,8 +1539,8 @@ Proof.
Qed.
Lemma transl_load_correct:
- forall chunk addr args dst k c (rs: regset) m a v,
- transl_load chunk addr args dst k = OK c ->
+ forall trap chunk addr args dst k c (rs: regset) m a v,
+ transl_load trap chunk addr args dst k = OK c ->
eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some a ->
Mem.loadv chunk m a = Some v ->
exists rs',
@@ -1311,7 +1548,8 @@ Lemma transl_load_correct:
/\ rs'#(preg_of dst) = v
/\ forall r, r <> PC -> r <> X31 -> r <> preg_of dst -> rs'#r = rs#r.
Proof.
- intros until v; intros TR EV LOAD.
+ intros until v; intros TR EV LOAD.
+ destruct trap; try (simpl in *; discriminate).
assert (A: exists mk_instr,
transl_memory_access mk_instr addr args k = OK c
/\ forall base ofs rs,
@@ -1390,6 +1628,3 @@ Proof.
Qed.
End CONSTRUCTORS.
-
-
-
diff --git a/riscV/Builtins1.v b/riscV/Builtins1.v
index cd6f8cc4..6691d15c 100644
--- a/riscV/Builtins1.v
+++ b/riscV/Builtins1.v
@@ -19,16 +19,35 @@
Require Import String Coqlib.
Require Import AST Integers Floats Values.
Require Import Builtins0.
+Require ExtValues.
-Inductive platform_builtin : Type := .
+Inductive platform_builtin : Type :=
+| BI_bits_of_float
+| BI_bits_of_double
+| BI_float_of_bits
+| BI_double_of_bits.
Local Open Scope string_scope.
Definition platform_builtin_table : list (string * platform_builtin) :=
- nil.
+ ("__builtin_bits_of_float", BI_bits_of_float)
+ :: ("__builtin_bits_of_double", BI_bits_of_double)
+ :: ("__builtin_float_of_bits", BI_float_of_bits)
+ :: ("__builtin_double_of_bits", BI_double_of_bits)
+ :: nil.
Definition platform_builtin_sig (b: platform_builtin) : signature :=
- match b with end.
+ match b with
+ | BI_bits_of_float => mksignature (Tsingle :: nil) Tint cc_default
+ | BI_bits_of_double => mksignature (Tfloat :: nil) Tlong cc_default
+ | BI_float_of_bits => mksignature (Tint :: nil) Tsingle cc_default
+ | BI_double_of_bits => mksignature (Tlong :: nil) Tfloat cc_default
+ end.
Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (sig_res (platform_builtin_sig b)) :=
- match b with end.
+ match b with
+ | BI_bits_of_float => mkbuiltin_n1t Tsingle Tint Float32.to_bits
+ | BI_bits_of_double => mkbuiltin_n1t Tfloat Tlong Float.to_bits
+ | BI_float_of_bits => mkbuiltin_n1t Tint Tsingle Float32.of_bits
+ | BI_double_of_bits => mkbuiltin_n1t Tlong Tfloat Float.of_bits
+ end.
diff --git a/riscV/CBuiltins.ml b/riscV/CBuiltins.ml
index 9ff4e029..ca0dbc6d 100644
--- a/riscV/CBuiltins.ml
+++ b/riscV/CBuiltins.ml
@@ -47,6 +47,14 @@ let builtins = {
(TFloat(FDouble, []), [TFloat(FDouble, []); TFloat(FDouble, [])], false);
"__builtin_fmin",
(TFloat(FDouble, []), [TFloat(FDouble, []); TFloat(FDouble, [])], false);
+ "__builtin_bits_of_double",
+ (TInt(IULong, []), [TFloat(FDouble, [])], false);
+ "__builtin_bits_of_float",
+ (TInt(IUInt, []), [TFloat(FFloat, [])], false);
+ "__builtin_double_of_bits",
+ (TFloat(FDouble, []), [TInt(IULong, [])], false);
+ "__builtin_float_of_bits",
+ (TFloat(FFloat, []), [TInt(IUInt, [])], false);
]
}
diff --git a/riscV/CSE2deps.v b/riscV/CSE2deps.v
new file mode 100644
index 00000000..c0deacf0
--- /dev/null
+++ b/riscV/CSE2deps.v
@@ -0,0 +1,35 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* David Monniaux CNRS, VERIMAG *)
+(* *)
+(* Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+Require Import BoolEqual Coqlib.
+Require Import AST Integers Floats.
+Require Import Values Memory Globalenvs Events.
+Require Import Op.
+
+
+Definition can_swap_accesses_ofs ofsr chunkr ofsw chunkw :=
+ (0 <=? ofsw) && (ofsw <=? (Ptrofs.modulus - largest_size_chunk))
+ && (0 <=? ofsr) && (ofsr <=? (Ptrofs.modulus - largest_size_chunk))
+ && ((ofsw + size_chunk chunkw <=? ofsr) ||
+ (ofsr + size_chunk chunkr <=? ofsw)).
+
+Definition may_overlap chunk addr args chunk' addr' args' :=
+ match addr, addr', args, args' with
+ | (Aindexed ofs), (Aindexed ofs'),
+ (base :: nil), (base' :: nil) =>
+ if peq base base'
+ then negb (can_swap_accesses_ofs (Ptrofs.unsigned ofs') chunk' (Ptrofs.unsigned ofs) chunk)
+ else true
+ | (Ainstack ofs), (Ainstack ofs'), _, _ =>
+ negb (can_swap_accesses_ofs (Ptrofs.unsigned ofs') chunk' (Ptrofs.unsigned ofs) chunk)
+ | _, _, _, _ => true
+ end.
diff --git a/riscV/CSE2depsproof.v b/riscV/CSE2depsproof.v
new file mode 100644
index 00000000..cf9e62b1
--- /dev/null
+++ b/riscV/CSE2depsproof.v
@@ -0,0 +1,147 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* David Monniaux CNRS, VERIMAG *)
+(* *)
+(* Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+Require Import Coqlib Maps Errors Integers Floats Lattice Kildall.
+Require Import AST Linking.
+Require Import Memory Registers Op RTL Maps.
+
+Require Import Globalenvs Values.
+Require Import Linking Values Memory Globalenvs Events Smallstep.
+Require Import Registers Op RTL.
+Require Import CSE2 CSE2deps.
+Require Import Lia.
+
+Lemma ptrofs_size :
+ Ptrofs.wordsize = (if Archi.ptr64 then 64 else 32)%nat.
+Proof.
+ unfold Ptrofs.wordsize.
+ unfold Wordsize_Ptrofs.wordsize.
+ trivial.
+Qed.
+
+Lemma ptrofs_modulus :
+ Ptrofs.modulus = if Archi.ptr64 then 18446744073709551616 else 4294967296.
+Proof.
+ unfold Ptrofs.modulus.
+ rewrite ptrofs_size.
+ destruct Archi.ptr64; reflexivity.
+Qed.
+
+Section SOUNDNESS.
+ Variable F V : Type.
+ Variable genv: Genv.t F V.
+ Variable sp : val.
+
+Section MEMORY_WRITE.
+ Variable m m2 : mem.
+ Variable chunkw chunkr : memory_chunk.
+ Variable base : val.
+
+ Variable addrw addrr valw : val.
+ Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2.
+
+ Section INDEXED_AWAY.
+ Variable ofsw ofsr : ptrofs.
+ Hypothesis ADDRW : eval_addressing genv sp
+ (Aindexed ofsw) (base :: nil) = Some addrw.
+ Hypothesis ADDRR : eval_addressing genv sp
+ (Aindexed ofsr) (base :: nil) = Some addrr.
+
+ Lemma load_store_away1 :
+ forall RANGEW : 0 <= Ptrofs.unsigned ofsw <= Ptrofs.modulus - largest_size_chunk,
+ forall RANGER : 0 <= Ptrofs.unsigned ofsr <= Ptrofs.modulus - largest_size_chunk,
+ forall SWAPPABLE : Ptrofs.unsigned ofsw + size_chunk chunkw <= Ptrofs.unsigned ofsr
+ \/ Ptrofs.unsigned ofsr + size_chunk chunkr <= Ptrofs.unsigned ofsw,
+ Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr.
+
+ Proof.
+ intros.
+
+ pose proof (max_size_chunk chunkr) as size_chunkr_bounded.
+ pose proof (max_size_chunk chunkw) as size_chunkw_bounded.
+ unfold largest_size_chunk in *.
+
+ rewrite ptrofs_modulus in *.
+ simpl in *.
+ inv ADDRR.
+ inv ADDRW.
+ destruct base; try discriminate.
+ eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b).
+ exact STORE.
+ right.
+
+ all: try (destruct (Ptrofs.unsigned_add_either i ofsr) as [OFSR | OFSR];
+ rewrite OFSR).
+ all: try (destruct (Ptrofs.unsigned_add_either i ofsw) as [OFSW | OFSW];
+ rewrite OFSW).
+ all: try rewrite ptrofs_modulus in *.
+ all: destruct Archi.ptr64.
+
+ all: intuition lia.
+ Qed.
+
+ Theorem load_store_away :
+ can_swap_accesses_ofs (Ptrofs.unsigned ofsr) chunkr (Ptrofs.unsigned ofsw) chunkw = true ->
+ Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr.
+ Proof.
+ intro SWAP.
+ unfold can_swap_accesses_ofs in SWAP.
+ repeat rewrite andb_true_iff in SWAP.
+ repeat rewrite orb_true_iff in SWAP.
+ repeat rewrite Z.leb_le in SWAP.
+ apply load_store_away1.
+ all: tauto.
+ Qed.
+ End INDEXED_AWAY.
+End MEMORY_WRITE.
+End SOUNDNESS.
+
+
+Section SOUNDNESS.
+ Variable F V : Type.
+ Variable genv: Genv.t F V.
+ Variable sp : val.
+
+Lemma may_overlap_sound:
+ forall m m' : mem,
+ forall chunk addr args chunk' addr' args' v a a' rs,
+ (eval_addressing genv sp addr (rs ## args)) = Some a ->
+ (eval_addressing genv sp addr' (rs ## args')) = Some a' ->
+ (may_overlap chunk addr args chunk' addr' args') = false ->
+ (Mem.storev chunk m a v) = Some m' ->
+ (Mem.loadv chunk' m' a') = (Mem.loadv chunk' m a').
+Proof.
+ intros until rs.
+ intros ADDR ADDR' OVERLAP STORE.
+ destruct addr; destruct addr'; try discriminate.
+- (* Aindexed / Aindexed *)
+ destruct args as [ | base [ | ]]. 1,3: discriminate.
+ destruct args' as [ | base' [ | ]]. 1,3: discriminate.
+ simpl in OVERLAP.
+ destruct (peq base base'). 2: discriminate.
+ subst base'.
+ destruct (can_swap_accesses_ofs (Ptrofs.unsigned i0) chunk' (Ptrofs.unsigned i) chunk) eqn:SWAP.
+ 2: discriminate.
+ simpl in *.
+ eapply load_store_away with (F:=F) (V:=V) (genv:=genv) (sp:=sp); eassumption.
+
+- (* Ainstack / Ainstack *)
+ destruct args. 2: discriminate.
+ destruct args'. 2: discriminate.
+ cbn in OVERLAP.
+ destruct (can_swap_accesses_ofs (Ptrofs.unsigned i0) chunk' (Ptrofs.unsigned i) chunk) eqn:SWAP.
+ 2: discriminate.
+ cbn in *.
+ eapply load_store_away with (F:=F) (V:=V) (genv:=genv) (sp:=sp); eassumption.
+Qed.
+
+End SOUNDNESS.
diff --git a/riscV/ConstpropOpproof.v b/riscV/ConstpropOpproof.v
index 8445e55f..74dc4a05 100644
--- a/riscV/ConstpropOpproof.v
+++ b/riscV/ConstpropOpproof.v
@@ -265,52 +265,84 @@ Qed.
Lemma make_divimm_correct:
forall n r1 r2 v,
- Val.divs e#r1 e#r2 = Some v ->
+ Val.maketotal (Val.divs e#r1 e#r2) = v ->
e#r2 = Vint n ->
let (op, args) := make_divimm n r1 r2 in
exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
Proof.
intros; unfold make_divimm.
- predSpec Int.eq Int.eq_spec n Int.one; intros. subst. rewrite H0 in H.
- destruct (e#r1) eqn:?;
- try (rewrite Val.divs_one in H; exists (Vint i); split; simpl; try rewrite Heqv0; auto);
- inv H; auto.
- 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.
+ predSpec Int.eq Int.eq_spec n Int.one; intros; subst; rewrite H0.
+ { destruct (e # r1) eqn:Er1.
+ all: try (cbn; exists (e # r1); split; auto; fail).
+ rewrite Val.divs_one.
+ cbn.
+ rewrite Er1.
+ exists (Vint i); split; auto.
+ }
+ destruct (Int.is_power2 n) eqn:Power2.
+ {
+ destruct (Int.ltu i (Int.repr 31)) eqn:iLT31.
+ {
+ cbn.
+ exists (Val.maketotal (Val.shrx e # r1 (Vint i))); split; auto.
+ destruct (Val.divs e # r1 (Vint n)) eqn:DIVS; cbn; auto.
+ rewrite Val.divs_pow2 with (y:=v) (n:=n).
+ cbn.
+ all: auto.
+ }
+ exists (Val.maketotal (Val.divs e # r1 (Vint n))); split; cbn; auto; congruence.
+ }
+ exists (Val.maketotal (Val.divs e # r1 (Vint n))); split; cbn; auto; congruence.
Qed.
Lemma make_divuimm_correct:
forall n r1 r2 v,
- Val.divu e#r1 e#r2 = Some v ->
+ Val.maketotal (Val.divu e#r1 e#r2) = v ->
e#r2 = Vint n ->
let (op, args) := make_divuimm n r1 r2 in
exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
Proof.
intros; unfold make_divuimm.
- predSpec Int.eq Int.eq_spec n Int.one; intros. subst. rewrite H0 in H.
- destruct (e#r1) eqn:?;
- try (rewrite Val.divu_one in H; exists (Vint i); split; simpl; try rewrite Heqv0; auto);
- inv H; auto.
- destruct (Int.is_power2 n) eqn:?.
- econstructor; split. simpl; eauto.
- rewrite H0 in H. erewrite Val.divu_pow2 by eauto. auto.
- exists v; auto.
+ predSpec Int.eq Int.eq_spec n Int.one; intros; subst; rewrite H0.
+ { destruct (e # r1) eqn:Er1.
+ all: try (cbn; exists (e # r1); split; auto; fail).
+ rewrite Val.divu_one.
+ cbn.
+ rewrite Er1.
+ exists (Vint i); split; auto.
+ }
+ destruct (Int.is_power2 n) eqn:Power2.
+ {
+ cbn.
+ exists (Val.shru e # r1 (Vint i)); split; auto.
+ destruct (Val.divu e # r1 (Vint n)) eqn:DIVU; cbn; auto.
+ rewrite Val.divu_pow2 with (y:=v) (n:=n).
+ all: auto.
+ }
+ exists (Val.maketotal (Val.divu e # r1 (Vint n))); split; cbn; auto; congruence.
Qed.
Lemma make_moduimm_correct:
forall n r1 r2 v,
- Val.modu e#r1 e#r2 = Some v ->
+ Val.maketotal (Val.modu e#r1 e#r2) = v ->
e#r2 = Vint n ->
let (op, args) := make_moduimm n r1 r2 in
exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
Proof.
intros; unfold make_moduimm.
destruct (Int.is_power2 n) eqn:?.
- exists v; split; auto. simpl. decEq. eapply Val.modu_pow2; eauto. congruence.
- exists v; auto.
+ { destruct (Val.modu e # r1 e # r2) eqn:MODU; cbn in H.
+ { subst v0.
+ exists v; split; auto.
+ cbn. decEq. eapply Val.modu_pow2; eauto. congruence.
+ }
+ subst v.
+ eexists; split; auto.
+ cbn. reflexivity.
+ }
+ exists v; split; auto.
+ cbn.
+ congruence.
Qed.
Lemma make_andimm_correct:
@@ -444,48 +476,82 @@ Qed.
Lemma make_divlimm_correct:
forall n r1 r2 v,
- Val.divls e#r1 e#r2 = Some v ->
+ Val.maketotal (Val.divls e#r1 e#r2) = v ->
e#r2 = Vlong n ->
let (op, args) := make_divlimm n r1 r2 in
exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
Proof.
intros; unfold make_divlimm.
- destruct (Int64.is_power2' n) eqn:?. destruct (Int.ltu i (Int.repr 63)) eqn:?.
- rewrite H0 in H. econstructor; split. simpl; eauto. eapply Val.divls_pow2; eauto. auto.
- exists v; auto.
- exists v; auto.
+ destruct (Int64.is_power2' n) eqn:Power2.
+ {
+ destruct (Int.ltu i (Int.repr 63)) eqn:iLT63.
+ {
+ cbn.
+ exists (Val.maketotal (Val.shrxl e # r1 (Vint i))); split; auto.
+ rewrite H0 in H.
+ destruct (Val.divls e # r1 (Vlong n)) eqn:DIVS; cbn in H; auto.
+ {
+ subst v0.
+ rewrite Val.divls_pow2 with (y:=v) (n:=n).
+ cbn.
+ all: auto.
+ }
+ subst. auto.
+ }
+ cbn. subst. rewrite H0.
+ exists (Val.maketotal (Val.divls e # r1 (Vlong n))); split; auto.
+ }
+ cbn. subst. rewrite H0.
+ exists (Val.maketotal (Val.divls e # r1 (Vlong n))); split; auto.
Qed.
Lemma make_divluimm_correct:
forall n r1 r2 v,
- Val.divlu e#r1 e#r2 = Some v ->
+ Val.maketotal (Val.divlu e#r1 e#r2) = v ->
e#r2 = Vlong n ->
let (op, args) := make_divluimm n r1 r2 in
exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
Proof.
intros; unfold make_divluimm.
destruct (Int64.is_power2' n) eqn:?.
+ {
econstructor; split. simpl; eauto.
- rewrite H0 in H. destruct (e#r1); inv H. destruct (Int64.eq n Int64.zero); inv H2.
- simpl.
- erewrite Int64.is_power2'_range by eauto.
- erewrite Int64.divu_pow2' by eauto. auto.
- exists v; auto.
+ rewrite H0 in H. destruct (e#r1); inv H.
+ all: cbn; auto.
+ {
+ destruct (Int64.eq n Int64.zero); cbn; auto.
+ erewrite Int64.is_power2'_range by eauto.
+ erewrite Int64.divu_pow2' by eauto. auto.
+ }
+ }
+ exists v; split; auto.
+ cbn.
+ rewrite H.
+ reflexivity.
Qed.
Lemma make_modluimm_correct:
forall n r1 r2 v,
- Val.modlu e#r1 e#r2 = Some v ->
+ Val.maketotal (Val.modlu e#r1 e#r2) = v ->
e#r2 = Vlong n ->
let (op, args) := make_modluimm n r1 r2 in
exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
Proof.
intros; unfold make_modluimm.
destruct (Int64.is_power2 n) eqn:?.
- exists v; split; auto. simpl. decEq.
- rewrite H0 in H. destruct (e#r1); inv H. destruct (Int64.eq n Int64.zero); inv H2.
- simpl. erewrite Int64.modu_and by eauto. auto.
- exists v; auto.
+ {
+ econstructor; split. simpl; eauto.
+ rewrite H0 in H. destruct (e#r1); inv H.
+ all: cbn; auto.
+ {
+ destruct (Int64.eq n Int64.zero); cbn; auto.
+ erewrite Int64.modu_and by eauto. auto.
+ }
+ }
+ exists v; split; auto.
+ cbn.
+ rewrite H.
+ reflexivity.
Qed.
Lemma make_andlimm_correct:
@@ -633,14 +699,17 @@ Proof.
- (* mul 2*)
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.
+ assert (e#r2 = Vint n2). { clear H0. InvApproxRegs; SimplVM; auto. }
+ apply make_divimm_correct; auto.
+ congruence.
- (* divu *)
assert (e#r2 = Vint n2). clear H0. InvApproxRegs; SimplVM; auto.
apply make_divuimm_correct; auto.
+ congruence.
- (* modu *)
assert (e#r2 = Vint n2). clear H0. InvApproxRegs; SimplVM; auto.
apply make_moduimm_correct; auto.
+ congruence.
- (* and 1 *)
rewrite Val.and_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_andimm_correct; auto.
- (* and 2 *)
@@ -680,12 +749,15 @@ Proof.
- (* divl *)
assert (e#r2 = Vlong n2). clear H0. InvApproxRegs; SimplVM; auto.
apply make_divlimm_correct; auto.
+ congruence.
- (* divlu *)
assert (e#r2 = Vlong n2). clear H0. InvApproxRegs; SimplVM; auto.
apply make_divluimm_correct; auto.
+ congruence.
- (* modlu *)
assert (e#r2 = Vlong n2). clear H0. InvApproxRegs; SimplVM; auto.
apply make_modluimm_correct; auto.
+ congruence.
- (* andl 1 *)
rewrite Val.andl_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_andlimm_correct; auto.
- (* andl 2 *)
diff --git a/riscV/DuplicateOpcodeHeuristic.ml b/riscV/DuplicateOpcodeHeuristic.ml
new file mode 100644
index 00000000..38702e1b
--- /dev/null
+++ b/riscV/DuplicateOpcodeHeuristic.ml
@@ -0,0 +1,41 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Sylvain Boulmé Grenoble-INP, VERIMAG *)
+(* David Monniaux CNRS, VERIMAG *)
+(* Cyril Six Kalray *)
+(* *)
+(* Copyright Kalray. Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+(* open Camlcoq *)
+open Op
+open Integers
+
+let opcode_heuristic code cond ifso ifnot is_loop_header =
+ match cond with
+ | Ccompimm (c, n) | Ccompuimm (c, n) -> if n == Integers.Int.zero then (match c with
+ | Clt | Cle -> Some false
+ | Cgt | Cge -> Some true
+ | _ -> None
+ ) else None
+ | Ccomplimm (c, n) | Ccompluimm (c, n) -> if n == Integers.Int64.zero then (match c with
+ | Clt | Cle -> Some false
+ | Cgt | Cge -> Some true
+ | _ -> None
+ ) else None
+ | Ccompf c | Ccompfs c -> (match c with
+ | Ceq -> Some false
+ | Cne -> Some true
+ | _ -> None
+ )
+ | Cnotcompf c | Cnotcompfs c -> (match c with
+ | Ceq -> Some true
+ | Cne -> Some false
+ | _ -> None
+ )
+ | _ -> None
diff --git a/riscV/ExpansionOracle.ml b/riscV/ExpansionOracle.ml
new file mode 100644
index 00000000..68d4e4d2
--- /dev/null
+++ b/riscV/ExpansionOracle.ml
@@ -0,0 +1,1068 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Léo Gourdin UGA, VERIMAG *)
+(* *)
+(* Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+open RTLpathLivegenaux
+open RTLpathCommon
+open Datatypes
+open Maps
+open RTL
+open Op
+open Asmgen
+open RTLpath
+open! Integers
+open Camlcoq
+open Option
+open AST
+open DebugPrint
+
+(** Mini CSE (a dynamic numbering is applied during expansion.
+ The CSE algorithm is inspired by the "static" one used in backend/CSE.v *)
+
+(** Managing virtual registers and node index *)
+
+let reg = ref 1
+
+let node = ref 1
+
+let p2i r = P.to_int r
+
+let r2p () = P.of_int !reg
+
+let n2p () = P.of_int !node
+
+let r2pi () =
+ reg := !reg + 1;
+ r2p ()
+
+let n2pi () =
+ node := !node + 1;
+ n2p ()
+
+(** Below are the types for rhs and equations *)
+
+type rhs = Sop of operation * int list | Smove
+
+type seq = Seq of int * rhs
+
+(** This is a mini abstraction to have a simpler representation during expansion
+ - Snop will be converted to Inop
+ - (Sr r) is inserted if the value was found in register r
+ - (Sexp dest rhs args succ) represent an instruction
+ (succesor may not be defined at this point, hence the use of type option)
+ - (Sfinalcond cond args succ1 succ2 info) represents a condition (which must
+ always be the last instruction in expansion list *)
+
+type expl =
+ | Snop of P.t
+ | Sr of P.t
+ | Sexp of P.t * rhs * P.t list * node option
+ | Sfinalcond of condition * P.t list * node * node * bool option
+
+(** Record used during the "dynamic" value numbering *)
+
+type numb = {
+ mutable nnext : int; (** Next unusued value number *)
+ mutable seqs : seq list; (** equations *)
+ mutable nreg : (P.t, int) Hashtbl.t; (** mapping registers to values *)
+ mutable nval : (int, P.t list) Hashtbl.t;
+ (** reverse mapping values to registers containing it *)
+}
+
+let print_list_pos l =
+ debug "[";
+ List.iter (fun i -> debug "%d;" (p2i i)) l;
+ debug "]\n"
+
+let empty_numbering () =
+ { nnext = 1; seqs = []; nreg = Hashtbl.create 100; nval = Hashtbl.create 100 }
+
+let rec get_nvalues vn = function
+ | [] -> []
+ | r :: rs ->
+ let v =
+ match Hashtbl.find_opt !vn.nreg r with
+ | Some v ->
+ debug "getnval r=%d |-> v=%d\n" (p2i r) v;
+ v
+ | None ->
+ let n = !vn.nnext in
+ debug "getnval r=%d |-> v=%d\n" (p2i r) n;
+ !vn.nnext <- !vn.nnext + 1;
+ Hashtbl.replace !vn.nreg r n;
+ Hashtbl.replace !vn.nval n [ r ];
+ n
+ in
+ let vs = get_nvalues vn rs in
+ v :: vs
+
+let get_nval_ornil vn v =
+ match Hashtbl.find_opt !vn.nval v with None -> [] | Some l -> l
+
+let forget_reg vn rd =
+ match Hashtbl.find_opt !vn.nreg rd with
+ | Some v ->
+ debug "forget_reg: r=%d |-> v=%d\n" (p2i rd) v;
+ let old_regs = get_nval_ornil vn v in
+ debug "forget_reg: old_regs are:\n";
+ print_list_pos old_regs;
+ Hashtbl.replace !vn.nval v
+ (List.filter (fun n -> not (P.eq n rd)) old_regs)
+ | None -> debug "forget_reg: no mapping for r=%d\n" (p2i rd)
+
+let update_reg vn rd v =
+ debug "update_reg: update v=%d with r=%d\n" v (p2i rd);
+ forget_reg vn rd;
+ let old_regs = get_nval_ornil vn v in
+ Hashtbl.replace !vn.nval v (rd :: old_regs)
+
+let rec find_valnum_rhs rh = function
+ | [] -> None
+ | Seq (v, rh') :: tl -> if rh = rh' then Some v else find_valnum_rhs rh tl
+
+let set_unknown vn rd =
+ debug "set_unknown: rd=%d\n" (p2i rd);
+ forget_reg vn rd;
+ Hashtbl.remove !vn.nreg rd
+
+let set_res_unknown vn res = match res with BR r -> set_unknown vn r | _ -> ()
+
+let addrhs vn rd rh =
+ match find_valnum_rhs rh !vn.seqs with
+ | Some vres ->
+ debug "addrhs: Some v=%d\n" vres;
+ Hashtbl.replace !vn.nreg rd vres;
+ update_reg vn rd vres
+ | None ->
+ let n = !vn.nnext in
+ debug "addrhs: None v=%d\n" n;
+ !vn.nnext <- !vn.nnext + 1;
+ !vn.seqs <- Seq (n, rh) :: !vn.seqs;
+ update_reg vn rd n;
+ Hashtbl.replace !vn.nreg rd n
+
+let addsop vn v op rd =
+ debug "addsop\n";
+ if op = Omove then (
+ update_reg vn rd (List.hd v);
+ Hashtbl.replace !vn.nreg rd (List.hd v))
+ else addrhs vn rd (Sop (op, v))
+
+let rec kill_mem_operations = function
+ | (Seq (v, Sop (op, vl)) as eq) :: tl ->
+ if op_depends_on_memory op then kill_mem_operations tl
+ else eq :: kill_mem_operations tl
+ | [] -> []
+ | eq :: tl -> eq :: kill_mem_operations tl
+
+let reg_valnum vn v =
+ debug "reg_valnum: trying to find a mapping for v=%d\n" v;
+ match Hashtbl.find !vn.nval v with
+ | [] -> None
+ | r :: rs ->
+ debug "reg_valnum: found a mapping r=%d\n" (p2i r);
+ Some r
+
+let rec reg_valnums vn = function
+ | [] -> Some []
+ | v :: vs -> (
+ match (reg_valnum vn v, reg_valnums vn vs) with
+ | Some r, Some rs -> Some (r :: rs)
+ | _, _ -> None)
+
+let find_rhs vn rh =
+ match find_valnum_rhs rh !vn.seqs with
+ | None -> None
+ | Some vres -> reg_valnum vn vres
+
+(** Functions to perform the dynamic reduction during CSE *)
+
+let extract_arg l =
+ if List.length l > 0 then
+ match List.hd l with
+ | Sr r -> (r, List.tl l)
+ | Sexp (rd, _, _, _) -> (rd, l)
+ | _ -> failwith "extract_arg: final instruction arg can not be extracted"
+ else failwith "extract_arg: trying to extract on an empty list"
+
+let extract_final vn fl fdest succ =
+ if List.length fl > 0 then
+ match List.hd fl with
+ | Sr r ->
+ if not (P.eq r fdest) then (
+ let v = get_nvalues vn [ r ] in
+ addsop vn v Omove fdest;
+ Sexp (fdest, Smove, [ r ], Some succ) :: List.tl fl)
+ else Snop succ :: List.tl fl
+ | Sexp (rd, rh, args, None) ->
+ assert (rd = fdest);
+ Sexp (fdest, rh, args, Some succ) :: List.tl fl
+ | _ -> fl
+ else failwith "extract_final: trying to extract on an empty list"
+
+let addinst vn op args rd =
+ let v = get_nvalues vn args in
+ let rh = Sop (op, v) in
+ match find_rhs vn rh with
+ | Some r ->
+ debug "addinst: rhs found with r=%d\n" (p2i r);
+ Sr r
+ | None ->
+ addsop vn v op rd;
+ Sexp (rd, rh, args, None)
+
+(** Expansion functions *)
+
+type immt =
+ | Addiw
+ | Addil
+ | Andiw
+ | Andil
+ | Oriw
+ | Oril
+ | Xoriw
+ | Xoril
+ | Sltiw
+ | Sltiuw
+ | Sltil
+ | Sltiul
+
+let load_hilo32 vn dest hi lo =
+ let op1 = OEluiw hi in
+ if Int.eq lo Int.zero then [ addinst vn op1 [] dest ]
+ else
+ let r = r2pi () in
+ let op2 = OEaddiw (None, lo) in
+ let i1 = addinst vn op1 [] r in
+ let r', l = extract_arg [ i1 ] in
+ let i2 = addinst vn op2 [ r' ] dest in
+ i2 :: l
+
+let load_hilo64 vn dest hi lo =
+ let op1 = OEluil hi in
+ if Int64.eq lo Int64.zero then [ addinst vn op1 [] dest ]
+ else
+ let r = r2pi () in
+ let op2 = OEaddil (None, lo) in
+ let i1 = addinst vn op1 [] r in
+ let r', l = extract_arg [ i1 ] in
+ let i2 = addinst vn op2 [ r' ] dest in
+ i2 :: l
+
+let loadimm32 vn dest n =
+ match make_immed32 n with
+ | Imm32_single imm ->
+ let op1 = OEaddiw (Some X0_R, imm) in
+ [ addinst vn op1 [] dest ]
+ | Imm32_pair (hi, lo) -> load_hilo32 vn dest hi lo
+
+let loadimm64 vn dest n =
+ match make_immed64 n with
+ | Imm64_single imm ->
+ let op1 = OEaddil (Some X0_R, imm) in
+ [ addinst vn op1 [] dest ]
+ | Imm64_pair (hi, lo) -> load_hilo64 vn dest hi lo
+ | Imm64_large imm ->
+ let op1 = OEloadli imm in
+ [ addinst vn op1 [] dest ]
+
+let get_opimm optR imm = function
+ | Addiw -> OEaddiw (optR, imm)
+ | Andiw -> OEandiw imm
+ | Oriw -> OEoriw imm
+ | Xoriw -> OExoriw imm
+ | Sltiw -> OEsltiw imm
+ | Sltiuw -> OEsltiuw imm
+ | Addil -> OEaddil (optR, imm)
+ | Andil -> OEandil imm
+ | Oril -> OEoril imm
+ | Xoril -> OExoril imm
+ | Sltil -> OEsltil imm
+ | Sltiul -> OEsltiul imm
+
+let opimm32 vn a1 dest n optR op opimm =
+ match make_immed32 n with
+ | Imm32_single imm -> [ addinst vn (get_opimm optR imm opimm) [ a1 ] dest ]
+ | Imm32_pair (hi, lo) ->
+ let r = r2pi () in
+ let l = load_hilo32 vn r hi lo in
+ let r', l' = extract_arg l in
+ let i = addinst vn op [ a1; r' ] dest in
+ i :: l'
+
+let opimm64 vn a1 dest n optR op opimm =
+ match make_immed64 n with
+ | Imm64_single imm -> [ addinst vn (get_opimm optR imm opimm) [ a1 ] dest ]
+ | Imm64_pair (hi, lo) ->
+ let r = r2pi () in
+ let l = load_hilo64 vn r hi lo in
+ let r', l' = extract_arg l in
+ let i = addinst vn op [ a1; r' ] dest in
+ i :: l'
+ | Imm64_large imm ->
+ let r = r2pi () in
+ let op1 = OEloadli imm in
+ let i1 = addinst vn op1 [] r in
+ let r', l' = extract_arg [ i1 ] in
+ let i2 = addinst vn op [ a1; r' ] dest in
+ i2 :: l'
+
+let addimm32 vn a1 dest n optR = opimm32 vn a1 dest n optR Oadd Addiw
+
+let andimm32 vn a1 dest n = opimm32 vn a1 dest n None Oand Andiw
+
+let orimm32 vn a1 dest n = opimm32 vn a1 dest n None Oor Oriw
+
+let xorimm32 vn a1 dest n = opimm32 vn a1 dest n None Oxor Xoriw
+
+let sltimm32 vn a1 dest n = opimm32 vn a1 dest n None (OEsltw None) Sltiw
+
+let sltuimm32 vn a1 dest n = opimm32 vn a1 dest n None (OEsltuw None) Sltiuw
+
+let addimm64 vn a1 dest n optR = opimm64 vn a1 dest n optR Oaddl Addil
+
+let andimm64 vn a1 dest n = opimm64 vn a1 dest n None Oandl Andil
+
+let orimm64 vn a1 dest n = opimm64 vn a1 dest n None Oorl Oril
+
+let xorimm64 vn a1 dest n = opimm64 vn a1 dest n None Oxorl Xoril
+
+let sltimm64 vn a1 dest n = opimm64 vn a1 dest n None (OEsltl None) Sltil
+
+let sltuimm64 vn a1 dest n = opimm64 vn a1 dest n None (OEsltul None) Sltiul
+
+let is_inv_cmp = function Cle | Cgt -> true | _ -> false
+
+let make_optR is_x0 is_inv =
+ if is_x0 then if is_inv then Some X0_L else Some X0_R else None
+
+let cbranch_int32s is_x0 cmp a1 a2 info succ1 succ2 k =
+ let optR = make_optR is_x0 (is_inv_cmp cmp) in
+ match cmp with
+ | Ceq -> Sfinalcond (CEbeqw optR, [ a1; a2 ], succ1, succ2, info) :: k
+ | Cne -> Sfinalcond (CEbnew optR, [ a1; a2 ], succ1, succ2, info) :: k
+ | Clt -> Sfinalcond (CEbltw optR, [ a1; a2 ], succ1, succ2, info) :: k
+ | Cle -> Sfinalcond (CEbgew optR, [ a2; a1 ], succ1, succ2, info) :: k
+ | Cgt -> Sfinalcond (CEbltw optR, [ a2; a1 ], succ1, succ2, info) :: k
+ | Cge -> Sfinalcond (CEbgew optR, [ a1; a2 ], succ1, succ2, info) :: k
+
+let cbranch_int32u is_x0 cmp a1 a2 info succ1 succ2 k =
+ let optR = make_optR is_x0 (is_inv_cmp cmp) in
+ match cmp with
+ | Ceq -> Sfinalcond (CEbequw optR, [ a1; a2 ], succ1, succ2, info) :: k
+ | Cne -> Sfinalcond (CEbneuw optR, [ a1; a2 ], succ1, succ2, info) :: k
+ | Clt -> Sfinalcond (CEbltuw optR, [ a1; a2 ], succ1, succ2, info) :: k
+ | Cle -> Sfinalcond (CEbgeuw optR, [ a2; a1 ], succ1, succ2, info) :: k
+ | Cgt -> Sfinalcond (CEbltuw optR, [ a2; a1 ], succ1, succ2, info) :: k
+ | Cge -> Sfinalcond (CEbgeuw optR, [ a1; a2 ], succ1, succ2, info) :: k
+
+let cbranch_int64s is_x0 cmp a1 a2 info succ1 succ2 k =
+ let optR = make_optR is_x0 (is_inv_cmp cmp) in
+ match cmp with
+ | Ceq -> Sfinalcond (CEbeql optR, [ a1; a2 ], succ1, succ2, info) :: k
+ | Cne -> Sfinalcond (CEbnel optR, [ a1; a2 ], succ1, succ2, info) :: k
+ | Clt -> Sfinalcond (CEbltl optR, [ a1; a2 ], succ1, succ2, info) :: k
+ | Cle -> Sfinalcond (CEbgel optR, [ a2; a1 ], succ1, succ2, info) :: k
+ | Cgt -> Sfinalcond (CEbltl optR, [ a2; a1 ], succ1, succ2, info) :: k
+ | Cge -> Sfinalcond (CEbgel optR, [ a1; a2 ], succ1, succ2, info) :: k
+
+let cbranch_int64u is_x0 cmp a1 a2 info succ1 succ2 k =
+ let optR = make_optR is_x0 (is_inv_cmp cmp) in
+ match cmp with
+ | Ceq -> Sfinalcond (CEbequl optR, [ a1; a2 ], succ1, succ2, info) :: k
+ | Cne -> Sfinalcond (CEbneul optR, [ a1; a2 ], succ1, succ2, info) :: k
+ | Clt -> Sfinalcond (CEbltul optR, [ a1; a2 ], succ1, succ2, info) :: k
+ | Cle -> Sfinalcond (CEbgeul optR, [ a2; a1 ], succ1, succ2, info) :: k
+ | Cgt -> Sfinalcond (CEbltul optR, [ a2; a1 ], succ1, succ2, info) :: k
+ | Cge -> Sfinalcond (CEbgeul optR, [ a1; a2 ], succ1, succ2, info) :: k
+
+let cond_int32s vn is_x0 cmp a1 a2 dest =
+ let optR = make_optR is_x0 (is_inv_cmp cmp) in
+ match cmp with
+ | Ceq -> [ addinst vn (OEseqw optR) [ a1; a2 ] dest ]
+ | Cne -> [ addinst vn (OEsnew optR) [ a1; a2 ] dest ]
+ | Clt -> [ addinst vn (OEsltw optR) [ a1; a2 ] dest ]
+ | Cle ->
+ let r = r2pi () in
+ let op = OEsltw optR in
+ let i1 = addinst vn op [ a2; a1 ] r in
+ let r', l = extract_arg [ i1 ] in
+ addinst vn (OExoriw Int.one) [ r' ] dest :: l
+ | Cgt -> [ addinst vn (OEsltw optR) [ a2; a1 ] dest ]
+ | Cge ->
+ let r = r2pi () in
+ let op = OEsltw optR in
+ let i1 = addinst vn op [ a1; a2 ] r in
+ let r', l = extract_arg [ i1 ] in
+ addinst vn (OExoriw Int.one) [ r' ] dest :: l
+
+let cond_int32u vn is_x0 cmp a1 a2 dest =
+ let optR = make_optR is_x0 (is_inv_cmp cmp) in
+ match cmp with
+ | Ceq -> [ addinst vn (OEsequw optR) [ a1; a2 ] dest ]
+ | Cne -> [ addinst vn (OEsneuw optR) [ a1; a2 ] dest ]
+ | Clt -> [ addinst vn (OEsltuw optR) [ a1; a2 ] dest ]
+ | Cle ->
+ let r = r2pi () in
+ let op = OEsltuw optR in
+ let i1 = addinst vn op [ a2; a1 ] r in
+ let r', l = extract_arg [ i1 ] in
+ addinst vn (OExoriw Int.one) [ r' ] dest :: l
+ | Cgt -> [ addinst vn (OEsltuw optR) [ a2; a1 ] dest ]
+ | Cge ->
+ let r = r2pi () in
+ let op = OEsltuw optR in
+ let i1 = addinst vn op [ a1; a2 ] r in
+ let r', l = extract_arg [ i1 ] in
+ addinst vn (OExoriw Int.one) [ r' ] dest :: l
+
+let cond_int64s vn is_x0 cmp a1 a2 dest =
+ let optR = make_optR is_x0 (is_inv_cmp cmp) in
+ match cmp with
+ | Ceq -> [ addinst vn (OEseql optR) [ a1; a2 ] dest ]
+ | Cne -> [ addinst vn (OEsnel optR) [ a1; a2 ] dest ]
+ | Clt -> [ addinst vn (OEsltl optR) [ a1; a2 ] dest ]
+ | Cle ->
+ let r = r2pi () in
+ let op = OEsltl optR in
+ let i1 = addinst vn op [ a2; a1 ] r in
+ let r', l = extract_arg [ i1 ] in
+ addinst vn (OExoriw Int.one) [ r' ] dest :: l
+ | Cgt -> [ addinst vn (OEsltl optR) [ a2; a1 ] dest ]
+ | Cge ->
+ let r = r2pi () in
+ let op = OEsltl optR in
+ let i1 = addinst vn op [ a1; a2 ] r in
+ let r', l = extract_arg [ i1 ] in
+ addinst vn (OExoriw Int.one) [ r' ] dest :: l
+
+let cond_int64u vn is_x0 cmp a1 a2 dest =
+ let optR = make_optR is_x0 (is_inv_cmp cmp) in
+ match cmp with
+ | Ceq -> [ addinst vn (OEsequl optR) [ a1; a2 ] dest ]
+ | Cne -> [ addinst vn (OEsneul optR) [ a1; a2 ] dest ]
+ | Clt -> [ addinst vn (OEsltul optR) [ a1; a2 ] dest ]
+ | Cle ->
+ let r = r2pi () in
+ let op = OEsltul optR in
+ let i1 = addinst vn op [ a2; a1 ] r in
+ let r', l = extract_arg [ i1 ] in
+ addinst vn (OExoriw Int.one) [ r' ] dest :: l
+ | Cgt -> [ addinst vn (OEsltul optR) [ a2; a1 ] dest ]
+ | Cge ->
+ let r = r2pi () in
+ let op = OEsltul optR in
+ let i1 = addinst vn op [ a1; a2 ] r in
+ let r', l = extract_arg [ i1 ] in
+ addinst vn (OExoriw Int.one) [ r' ] dest :: l
+
+let is_normal_cmp = function Cne -> false | _ -> true
+
+let cond_float vn cmp f1 f2 dest =
+ match cmp with
+ | Ceq -> [ addinst vn OEfeqd [ f1; f2 ] dest ]
+ | Cne -> [ addinst vn OEfeqd [ f1; f2 ] dest ]
+ | Clt -> [ addinst vn OEfltd [ f1; f2 ] dest ]
+ | Cle -> [ addinst vn OEfled [ f1; f2 ] dest ]
+ | Cgt -> [ addinst vn OEfltd [ f2; f1 ] dest ]
+ | Cge -> [ addinst vn OEfled [ f2; f1 ] dest ]
+
+let cond_single vn cmp f1 f2 dest =
+ match cmp with
+ | Ceq -> [ addinst vn OEfeqs [ f1; f2 ] dest ]
+ | Cne -> [ addinst vn OEfeqs [ f1; f2 ] dest ]
+ | Clt -> [ addinst vn OEflts [ f1; f2 ] dest ]
+ | Cle -> [ addinst vn OEfles [ f1; f2 ] dest ]
+ | Cgt -> [ addinst vn OEflts [ f2; f1 ] dest ]
+ | Cge -> [ addinst vn OEfles [ f2; f1 ] dest ]
+
+let expanse_cbranchimm_int32s vn cmp a1 n info succ1 succ2 =
+ if Int.eq n Int.zero then cbranch_int32s true cmp a1 a1 info succ1 succ2 []
+ else
+ let r = r2pi () in
+ let l = loadimm32 vn r n in
+ let r', l' = extract_arg l in
+ cbranch_int32s false cmp a1 r' info succ1 succ2 l'
+
+let expanse_cbranchimm_int32u vn cmp a1 n info succ1 succ2 =
+ if Int.eq n Int.zero then cbranch_int32u true cmp a1 a1 info succ1 succ2 []
+ else
+ let r = r2pi () in
+ let l = loadimm32 vn r n in
+ let r', l' = extract_arg l in
+ cbranch_int32u false cmp a1 r' info succ1 succ2 l'
+
+let expanse_cbranchimm_int64s vn cmp a1 n info succ1 succ2 =
+ if Int64.eq n Int64.zero then
+ cbranch_int64s true cmp a1 a1 info succ1 succ2 []
+ else
+ let r = r2pi () in
+ let l = loadimm64 vn r n in
+ let r', l' = extract_arg l in
+ cbranch_int64s false cmp a1 r' info succ1 succ2 l'
+
+let expanse_cbranchimm_int64u vn cmp a1 n info succ1 succ2 =
+ if Int64.eq n Int64.zero then
+ cbranch_int64u true cmp a1 a1 info succ1 succ2 []
+ else
+ let r = r2pi () in
+ let l = loadimm64 vn r n in
+ let r', l' = extract_arg l in
+ cbranch_int64u false cmp a1 r' info succ1 succ2 l'
+
+let expanse_condimm_int32s vn cmp a1 n dest =
+ if Int.eq n Int.zero then cond_int32s vn true cmp a1 a1 dest
+ else
+ match cmp with
+ | Ceq | Cne ->
+ let r = r2pi () in
+ let l = xorimm32 vn a1 r n in
+ let r', l' = extract_arg l in
+ cond_int32s vn true cmp r' r' dest @ l'
+ | Clt -> sltimm32 vn a1 dest n
+ | Cle ->
+ if Int.eq n (Int.repr Int.max_signed) then
+ let l = loadimm32 vn dest Int.one in
+ let r, l' = extract_arg l in
+ addinst vn (OEmayundef MUint) [ a1; r ] dest :: l'
+ else sltimm32 vn a1 dest (Int.add n Int.one)
+ | _ ->
+ let r = r2pi () in
+ let l = loadimm32 vn r n in
+ let r', l' = extract_arg l in
+ cond_int32s vn false cmp a1 r' dest @ l'
+
+let expanse_condimm_int32u vn cmp a1 n dest =
+ if Int.eq n Int.zero then cond_int32u vn true cmp a1 a1 dest
+ else
+ match cmp with
+ | Clt -> sltuimm32 vn a1 dest n
+ | _ ->
+ let r = r2pi () in
+ let l = loadimm32 vn r n in
+ let r', l' = extract_arg l in
+ cond_int32u vn false cmp a1 r' dest @ l'
+
+let expanse_condimm_int64s vn cmp a1 n dest =
+ if Int64.eq n Int64.zero then cond_int64s vn true cmp a1 a1 dest
+ else
+ match cmp with
+ | Ceq | Cne ->
+ let r = r2pi () in
+ let l = xorimm64 vn a1 r n in
+ let r', l' = extract_arg l in
+ cond_int64s vn true cmp r' r' dest @ l'
+ | Clt -> sltimm64 vn a1 dest n
+ | Cle ->
+ if Int64.eq n (Int64.repr Int64.max_signed) then
+ let l = loadimm32 vn dest Int.one in
+ let r, l' = extract_arg l in
+ addinst vn (OEmayundef MUlong) [ a1; r ] dest :: l'
+ else sltimm64 vn a1 dest (Int64.add n Int64.one)
+ | _ ->
+ let r = r2pi () in
+ let l = loadimm64 vn r n in
+ let r', l' = extract_arg l in
+ cond_int64s vn false cmp a1 r' dest @ l'
+
+let expanse_condimm_int64u vn cmp a1 n dest =
+ if Int64.eq n Int64.zero then cond_int64u vn true cmp a1 a1 dest
+ else
+ match cmp with
+ | Clt -> sltuimm64 vn a1 dest n
+ | _ ->
+ let r = r2pi () in
+ let l = loadimm64 vn r n in
+ let r', l' = extract_arg l in
+ cond_int64u vn false cmp a1 r' dest @ l'
+
+let expanse_cond_fp vn cnot fn_cond cmp f1 f2 dest =
+ let normal = is_normal_cmp cmp in
+ let normal' = if cnot then not normal else normal in
+ let insn = fn_cond vn cmp f1 f2 dest in
+ if normal' then insn
+ else
+ let r', l = extract_arg insn in
+ addinst vn (OExoriw Int.one) [ r' ] dest :: l
+
+let expanse_cbranch_fp vn cnot fn_cond cmp f1 f2 info succ1 succ2 =
+ let r = r2pi () in
+ let normal = is_normal_cmp cmp in
+ let normal' = if cnot then not normal else normal in
+ let insn = fn_cond vn cmp f1 f2 r in
+ let r', l = extract_arg insn in
+ if normal' then
+ Sfinalcond (CEbnew (Some X0_R), [ r'; r' ], succ1, succ2, info) :: l
+ else Sfinalcond (CEbeqw (Some X0_R), [ r'; r' ], succ1, succ2, info) :: l
+
+(** Form a list containing both sources and destination regs of an instruction *)
+
+let get_regindent = function Coq_inr _ -> [] | Coq_inl r -> [ r ]
+
+let get_regs_inst = function
+ | Inop _ -> []
+ | Iop (_, args, dest, _) -> dest :: args
+ | Iload (_, _, _, args, dest, _) -> dest :: args
+ | Istore (_, _, args, src, _) -> src :: args
+ | Icall (_, t, args, dest, _) -> dest :: (get_regindent t @ args)
+ | Itailcall (_, t, args) -> get_regindent t @ args
+ | Ibuiltin (_, args, dest, _) ->
+ AST.params_of_builtin_res dest @ AST.params_of_builtin_args args
+ | Icond (_, args, _, _, _) -> args
+ | Ijumptable (arg, _) -> [ arg ]
+ | Ireturn (Some r) -> [ r ]
+ | _ -> []
+
+(** Modify pathmap according to the size of the expansion list *)
+
+let write_pathmap initial esize pm' =
+ debug "write_pathmap: initial=%d, esize=%d\n" (p2i initial) esize;
+ let path = get_some @@ PTree.get initial !pm' in
+ let npsize = Camlcoq.Nat.of_int (esize + Camlcoq.Nat.to_int path.psize) in
+ let path' =
+ {
+ psize = npsize;
+ input_regs = path.input_regs;
+ pre_output_regs = path.pre_output_regs;
+ output_regs = path.output_regs;
+ }
+ in
+ pm' := PTree.set initial path' !pm'
+
+(** Write a single instruction in the tree and update order *)
+
+let write_inst target_node inst code' new_order =
+ code' := PTree.set (P.of_int target_node) inst !code';
+ new_order := P.of_int target_node :: !new_order
+
+(** Return olds args if the CSE numbering is empty *)
+
+let get_arguments vn vals args =
+ match reg_valnums vn vals with Some args' -> args' | None -> args
+
+(** Update the code tree with the expansion list *)
+
+let rec write_tree vn exp initial current code' new_order fturn =
+ debug "wt: node is %d\n" !node;
+ let target_node, next_node =
+ if fturn then (P.to_int initial, current) else (current, current - 1)
+ in
+ match exp with
+ | Sr r :: _ ->
+ failwith "write_tree: there are still some symbolic values in the list"
+ | Sexp (rd, Sop (op, vals), args, None) :: k ->
+ let args = get_arguments vn vals args in
+ let inst = Iop (op, args, rd, P.of_int next_node) in
+ write_inst target_node inst code' new_order;
+ write_tree vn k initial next_node code' new_order false
+ | [ Snop succ ] ->
+ let inst = Inop succ in
+ write_inst target_node inst code' new_order
+ | [ Sexp (rd, Sop (op, vals), args, Some succ) ] ->
+ let args = get_arguments vn vals args in
+ let inst = Iop (op, args, rd, succ) in
+ write_inst target_node inst code' new_order
+ | [ Sexp (rd, Smove, args, Some succ) ] ->
+ let inst = Iop (Omove, args, rd, succ) in
+ write_inst target_node inst code' new_order
+ | [ Sfinalcond (cond, args, succ1, succ2, info) ] ->
+ let inst = Icond (cond, args, succ1, succ2, info) in
+ write_inst target_node inst code' new_order
+ | [] -> ()
+ | _ -> failwith "write_tree: invalid list"
+
+(** Main expansion function - TODO gourdinl to split? *)
+let expanse (sb : superblock) code pm =
+ debug "#### New superblock for expansion oracle\n";
+ let new_order = ref [] in
+ let liveins = ref sb.liveins in
+ let exp = ref [] in
+ let was_branch = ref false in
+ let was_exp = ref false in
+ let code' = ref code in
+ let pm' = ref pm in
+ let vn = ref (empty_numbering ()) in
+ Array.iter
+ (fun n ->
+ was_branch := false;
+ was_exp := false;
+ let inst = get_some @@ PTree.get n code in
+ (if !Clflags.option_fexpanse_rtlcond then
+ match inst with
+ (* Expansion of conditions - Ocmp *)
+ | Iop (Ocmp (Ccomp c), a1 :: a2 :: nil, dest, succ) ->
+ debug "Iop/Ccomp\n";
+ exp := cond_int32s vn false c a1 a2 dest;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocmp (Ccompu c), a1 :: a2 :: nil, dest, succ) ->
+ debug "Iop/Ccompu\n";
+ exp := cond_int32u vn false c a1 a2 dest;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocmp (Ccompimm (c, imm)), a1 :: nil, dest, succ) ->
+ debug "Iop/Ccompimm\n";
+ exp := expanse_condimm_int32s vn c a1 imm dest;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocmp (Ccompuimm (c, imm)), a1 :: nil, dest, succ) ->
+ debug "Iop/Ccompuimm\n";
+ exp := expanse_condimm_int32u vn c a1 imm dest;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocmp (Ccompl c), a1 :: a2 :: nil, dest, succ) ->
+ debug "Iop/Ccompl\n";
+ exp := cond_int64s vn false c a1 a2 dest;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocmp (Ccomplu c), a1 :: a2 :: nil, dest, succ) ->
+ debug "Iop/Ccomplu\n";
+ exp := cond_int64u vn false c a1 a2 dest;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocmp (Ccomplimm (c, imm)), a1 :: nil, dest, succ) ->
+ debug "Iop/Ccomplimm\n";
+ exp := expanse_condimm_int64s vn c a1 imm dest;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocmp (Ccompluimm (c, imm)), a1 :: nil, dest, succ) ->
+ debug "Iop/Ccompluimm\n";
+ exp := expanse_condimm_int64u vn c a1 imm dest;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocmp (Ccompf c), f1 :: f2 :: nil, dest, succ) ->
+ debug "Iop/Ccompf\n";
+ exp := expanse_cond_fp vn false cond_float c f1 f2 dest;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocmp (Cnotcompf c), f1 :: f2 :: nil, dest, succ) ->
+ debug "Iop/Cnotcompf\n";
+ exp := expanse_cond_fp vn true cond_float c f1 f2 dest;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocmp (Ccompfs c), f1 :: f2 :: nil, dest, succ) ->
+ debug "Iop/Ccompfs\n";
+ exp := expanse_cond_fp vn false cond_single c f1 f2 dest;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocmp (Cnotcompfs c), f1 :: f2 :: nil, dest, succ) ->
+ debug "Iop/Cnotcompfs\n";
+ exp := expanse_cond_fp vn true cond_single c f1 f2 dest;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ (* Expansion of branches - Ccomp *)
+ | Icond (Ccomp c, a1 :: a2 :: nil, succ1, succ2, info) ->
+ debug "Icond/Ccomp\n";
+ exp := cbranch_int32s false c a1 a2 info succ1 succ2 [];
+ was_branch := true;
+ was_exp := true
+ | Icond (Ccompu c, a1 :: a2 :: nil, succ1, succ2, info) ->
+ debug "Icond/Ccompu\n";
+ exp := cbranch_int32u false c a1 a2 info succ1 succ2 [];
+ was_branch := true;
+ was_exp := true
+ | Icond (Ccompimm (c, imm), a1 :: nil, succ1, succ2, info) ->
+ debug "Icond/Ccompimm\n";
+ exp := expanse_cbranchimm_int32s vn c a1 imm info succ1 succ2;
+ was_branch := true;
+ was_exp := true
+ | Icond (Ccompuimm (c, imm), a1 :: nil, succ1, succ2, info) ->
+ debug "Icond/Ccompuimm\n";
+ exp := expanse_cbranchimm_int32u vn c a1 imm info succ1 succ2;
+ was_branch := true;
+ was_exp := true
+ | Icond (Ccompl c, a1 :: a2 :: nil, succ1, succ2, info) ->
+ debug "Icond/Ccompl\n";
+ exp := cbranch_int64s false c a1 a2 info succ1 succ2 [];
+ was_branch := true;
+ was_exp := true
+ | Icond (Ccomplu c, a1 :: a2 :: nil, succ1, succ2, info) ->
+ debug "Icond/Ccomplu\n";
+ exp := cbranch_int64u false c a1 a2 info succ1 succ2 [];
+ was_branch := true;
+ was_exp := true
+ | Icond (Ccomplimm (c, imm), a1 :: nil, succ1, succ2, info) ->
+ debug "Icond/Ccomplimm\n";
+ exp := expanse_cbranchimm_int64s vn c a1 imm info succ1 succ2;
+ was_branch := true;
+ was_exp := true
+ | Icond (Ccompluimm (c, imm), a1 :: nil, succ1, succ2, info) ->
+ debug "Icond/Ccompluimm\n";
+ exp := expanse_cbranchimm_int64u vn c a1 imm info succ1 succ2;
+ was_branch := true;
+ was_exp := true
+ | Icond (Ccompf c, f1 :: f2 :: nil, succ1, succ2, info) ->
+ debug "Icond/Ccompf\n";
+ exp :=
+ expanse_cbranch_fp vn false cond_float c f1 f2 info succ1 succ2;
+ was_branch := true;
+ was_exp := true
+ | Icond (Cnotcompf c, f1 :: f2 :: nil, succ1, succ2, info) ->
+ debug "Icond/Cnotcompf\n";
+ exp := expanse_cbranch_fp vn true cond_float c f1 f2 info succ1 succ2;
+ was_branch := true;
+ was_exp := true
+ | Icond (Ccompfs c, f1 :: f2 :: nil, succ1, succ2, info) ->
+ debug "Icond/Ccompfs\n";
+ exp :=
+ expanse_cbranch_fp vn false cond_single c f1 f2 info succ1 succ2;
+ was_branch := true;
+ was_exp := true
+ | Icond (Cnotcompfs c, f1 :: f2 :: nil, succ1, succ2, info) ->
+ debug "Icond/Cnotcompfs\n";
+ exp :=
+ expanse_cbranch_fp vn true cond_single c f1 f2 info succ1 succ2;
+ was_branch := true;
+ was_exp := true
+ | _ -> ());
+ (if !Clflags.option_fexpanse_others && not !was_exp then
+ match inst with
+ | Iop (Ofloatconst f, nil, dest, succ) -> (
+ match make_immed64 (Floats.Float.to_bits f) with
+ | Imm64_single _ | Imm64_large _ -> ()
+ | Imm64_pair (hi, lo) ->
+ debug "Iop/Ofloatconst\n";
+ let r = r2pi () in
+ let l = load_hilo64 vn r hi lo in
+ let r', l' = extract_arg l in
+ exp := addinst vn Ofloat_of_bits [ r' ] dest :: l';
+ exp := extract_final vn !exp dest succ;
+ was_exp := true)
+ | Iop (Osingleconst f, nil, dest, succ) -> (
+ match make_immed32 (Floats.Float32.to_bits f) with
+ | Imm32_single imm -> ()
+ | Imm32_pair (hi, lo) ->
+ debug "Iop/Osingleconst\n";
+ let r = r2pi () in
+ let l = load_hilo32 vn r hi lo in
+ let r', l' = extract_arg l in
+ exp := addinst vn Osingle_of_bits [ r' ] dest :: l';
+ exp := extract_final vn !exp dest succ;
+ was_exp := true)
+ | Iop (Ointconst n, nil, dest, succ) ->
+ debug "Iop/Ointconst\n";
+ exp := loadimm32 vn dest n;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Olongconst n, nil, dest, succ) ->
+ debug "Iop/Olongconst\n";
+ exp := loadimm64 vn dest n;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Oaddimm n, a1 :: nil, dest, succ) ->
+ debug "Iop/Oaddimm\n";
+ exp := addimm32 vn a1 dest n None;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Oaddlimm n, a1 :: nil, dest, succ) ->
+ debug "Iop/Oaddlimm\n";
+ exp := addimm64 vn a1 dest n None;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Oandimm n, a1 :: nil, dest, succ) ->
+ debug "Iop/Oandimm\n";
+ exp := andimm32 vn a1 dest n;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Oandlimm n, a1 :: nil, dest, succ) ->
+ debug "Iop/Oandlimm\n";
+ exp := andimm64 vn a1 dest n;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Oorimm n, a1 :: nil, dest, succ) ->
+ debug "Iop/Oorimm\n";
+ exp := orimm32 vn a1 dest n;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Oorlimm n, a1 :: nil, dest, succ) ->
+ debug "Iop/Oorlimm\n";
+ exp := orimm64 vn a1 dest n;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Oxorimm n, a1 :: nil, dest, succ) ->
+ debug "Iop/Oxorimm\n";
+ exp := xorimm32 vn a1 dest n;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Oxorlimm n, a1 :: nil, dest, succ) ->
+ debug "Iop/Oxorlimm\n";
+ exp := xorimm64 vn a1 dest n;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocast8signed, a1 :: nil, dest, succ) ->
+ debug "Iop/cast8signed\n";
+ let op = Oshlimm (Int.repr (Z.of_sint 24)) in
+ let r = r2pi () in
+ let i1 = addinst vn op [ a1 ] r in
+ let r', l = extract_arg [ i1 ] in
+ exp :=
+ addinst vn (Oshrimm (Int.repr (Z.of_sint 24))) [ r' ] dest :: l;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocast16signed, a1 :: nil, dest, succ) ->
+ debug "Iop/cast16signed\n";
+ let op = Oshlimm (Int.repr (Z.of_sint 16)) in
+ let r = r2pi () in
+ let i1 = addinst vn op [ a1 ] r in
+ let r', l = extract_arg [ i1 ] in
+ exp :=
+ addinst vn (Oshrimm (Int.repr (Z.of_sint 16))) [ r' ] dest :: l;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Ocast32unsigned, a1 :: nil, dest, succ) ->
+ debug "Iop/Ocast32unsigned\n";
+ let r1 = r2pi () in
+ let r2 = r2pi () in
+ let op1 = Ocast32signed in
+ let i1 = addinst vn op1 [ a1 ] r1 in
+ let r1', l1 = extract_arg [ i1 ] in
+
+ let op2 = Oshllimm (Int.repr (Z.of_sint 32)) in
+ let i2 = addinst vn op2 [ r1' ] r2 in
+ let r2', l2 = extract_arg (i2 :: l1) in
+
+ let op3 = Oshrluimm (Int.repr (Z.of_sint 32)) in
+ exp := addinst vn op3 [ r2' ] dest :: l2;
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Oshrximm n, a1 :: nil, dest, succ) ->
+ if Int.eq n Int.zero then (
+ debug "Iop/Oshrximm1\n";
+ exp := [ addinst vn (OEmayundef (MUshrx n)) [ a1; a1 ] dest ])
+ else if Int.eq n Int.one then (
+ debug "Iop/Oshrximm2\n";
+ let r1 = r2pi () in
+ let r2 = r2pi () in
+ let op1 = Oshruimm (Int.repr (Z.of_sint 31)) in
+ let i1 = addinst vn op1 [ a1 ] r1 in
+ let r1', l1 = extract_arg [ i1 ] in
+
+ let op2 = Oadd in
+ let i2 = addinst vn op2 [ a1; r1' ] r2 in
+ let r2', l2 = extract_arg (i2 :: l1) in
+
+ let op3 = Oshrimm Int.one in
+ let i3 = addinst vn op3 [ r2' ] dest in
+ let r3, l3 = extract_arg (i3 :: l2) in
+ exp := addinst vn (OEmayundef (MUshrx n)) [ r3; r3 ] dest :: l3)
+ else (
+ debug "Iop/Oshrximm3\n";
+ let r1 = r2pi () in
+ let r2 = r2pi () in
+ let r3 = r2pi () in
+ let op1 = Oshrimm (Int.repr (Z.of_sint 31)) in
+ let i1 = addinst vn op1 [ a1 ] r1 in
+ let r1', l1 = extract_arg [ i1 ] in
+
+ let op2 = Oshruimm (Int.sub Int.iwordsize n) in
+ let i2 = addinst vn op2 [ r1' ] r2 in
+ let r2', l2 = extract_arg (i2 :: l1) in
+
+ let op3 = Oadd in
+ let i3 = addinst vn op3 [ a1; r2' ] r3 in
+ let r3', l3 = extract_arg (i3 :: l2) in
+
+ let op4 = Oshrimm n in
+ let i4 = addinst vn op4 [ r3' ] dest in
+ let r4, l4 = extract_arg (i4 :: l3) in
+ exp := addinst vn (OEmayundef (MUshrx n)) [ r4; r4 ] dest :: l4);
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | Iop (Oshrxlimm n, a1 :: nil, dest, succ) ->
+ if Int.eq n Int.zero then (
+ debug "Iop/Oshrxlimm1\n";
+ exp := [ addinst vn (OEmayundef (MUshrxl n)) [ a1; a1 ] dest ])
+ else if Int.eq n Int.one then (
+ debug "Iop/Oshrxlimm2\n";
+ let r1 = r2pi () in
+ let r2 = r2pi () in
+ let op1 = Oshrluimm (Int.repr (Z.of_sint 63)) in
+ let i1 = addinst vn op1 [ a1 ] r1 in
+ let r1', l1 = extract_arg [ i1 ] in
+
+ let op2 = Oaddl in
+ let i2 = addinst vn op2 [ a1; r1' ] r2 in
+ let r2', l2 = extract_arg (i2 :: l1) in
+
+ let op3 = Oshrlimm Int.one in
+ let i3 = addinst vn op3 [ r2' ] dest in
+ let r3, l3 = extract_arg (i3 :: l2) in
+ exp := addinst vn (OEmayundef (MUshrxl n)) [ r3; r3 ] dest :: l3)
+ else (
+ debug "Iop/Oshrxlimm3\n";
+ let r1 = r2pi () in
+ let r2 = r2pi () in
+ let r3 = r2pi () in
+ let op1 = Oshrlimm (Int.repr (Z.of_sint 63)) in
+ let i1 = addinst vn op1 [ a1 ] r1 in
+ let r1', l1 = extract_arg [ i1 ] in
+
+ let op2 = Oshrluimm (Int.sub Int64.iwordsize' n) in
+ let i2 = addinst vn op2 [ r1' ] r2 in
+ let r2', l2 = extract_arg (i2 :: l1) in
+
+ let op3 = Oaddl in
+ let i3 = addinst vn op3 [ a1; r2' ] r3 in
+ let r3', l3 = extract_arg (i3 :: l2) in
+
+ let op4 = Oshrlimm n in
+ let i4 = addinst vn op4 [ r3' ] dest in
+ let r4, l4 = extract_arg (i4 :: l3) in
+ exp := addinst vn (OEmayundef (MUshrxl n)) [ r4; r4 ] dest :: l4);
+ exp := extract_final vn !exp dest succ;
+ was_exp := true
+ | _ -> ());
+ (* Update the CSE numbering *)
+ (if not !was_exp then
+ match inst with
+ | Iop (op, args, dest, succ) ->
+ let v = get_nvalues vn args in
+ addsop vn v op dest
+ | Iload (_, _, _, _, dst, _) -> set_unknown vn dst
+ | Istore (chk, addr, args, src, s) ->
+ !vn.seqs <- kill_mem_operations !vn.seqs
+ | Icall (_, _, _, _, _) | Itailcall (_, _, _) | Ibuiltin (_, _, _, _) ->
+ vn := empty_numbering ()
+ | _ -> ());
+ (* Update code, liveins, pathmap, and order of the superblock for one expansion *)
+ if !was_exp then (
+ (if !was_branch && List.length !exp > 1 then
+ let lives = PTree.get n !liveins in
+ match lives with
+ | Some lives ->
+ let new_branch_pc = P.of_int (!node + 1) in
+ liveins := PTree.set new_branch_pc lives !liveins;
+ liveins := PTree.remove n !liveins
+ | _ -> ());
+ node := !node + List.length !exp - 1;
+ write_pathmap sb.instructions.(0) (List.length !exp - 1) pm';
+ write_tree vn (List.rev !exp) n !node code' new_order true)
+ else new_order := n :: !new_order)
+ sb.instructions;
+ sb.instructions <- Array.of_list (List.rev !new_order);
+ sb.liveins <- !liveins;
+ (!code', !pm')
+
+(** Compute the last used node and reg indexs *)
+
+let rec find_last_node_reg = function
+ | [] -> ()
+ | (pc, i) :: k ->
+ let rec traverse_list var = function
+ | [] -> ()
+ | e :: t ->
+ let e' = p2i e in
+ if e' > !var then var := e';
+ traverse_list var t
+ in
+ traverse_list node [ pc ];
+ traverse_list reg (get_regs_inst i);
+ find_last_node_reg k
diff --git a/riscV/ExtValues.v b/riscV/ExtValues.v
new file mode 100644
index 00000000..edf359ef
--- /dev/null
+++ b/riscV/ExtValues.v
@@ -0,0 +1,123 @@
+Require Import Coqlib.
+Require Import Integers.
+Require Import Values.
+Require Import Floats.
+Require Import Memory.
+Require Import Lia.
+
+Definition bits_of_float x :=
+ match x with
+ | Vfloat f => Vlong (Float.to_bits f)
+ | _ => Vundef
+ end.
+
+Definition bits_of_single x :=
+ match x with
+ | Vsingle f => Vint (Float32.to_bits f)
+ | _ => Vundef
+ end.
+
+Definition float_of_bits x :=
+ match x with
+ | Vlong f => Vfloat (Float.of_bits f)
+ | _ => Vundef
+ end.
+
+Definition single_of_bits x :=
+ match x with
+ | Vint f => Vsingle (Float32.of_bits f)
+ | _ => Vundef
+ end.
+
+Definition bitwise_select_long b vtrue vfalse :=
+ Int64.or (Int64.and (Int64.neg b) vtrue)
+ (Int64.and (Int64.sub b Int64.one) vfalse).
+
+Lemma bitwise_select_long_true :
+ forall vtrue vfalse,
+ bitwise_select_long Int64.one vtrue vfalse = vtrue.
+Proof.
+ intros. unfold bitwise_select_long. cbn.
+ change (Int64.neg Int64.one) with Int64.mone.
+ rewrite Int64.and_commut.
+ rewrite Int64.and_mone.
+ rewrite Int64.sub_idem.
+ rewrite Int64.and_commut.
+ rewrite Int64.and_zero.
+ apply Int64.or_zero.
+Qed.
+
+Lemma bitwise_select_long_false :
+ forall vtrue vfalse,
+ bitwise_select_long Int64.zero vtrue vfalse = vfalse.
+Proof.
+ intros. unfold bitwise_select_long. cbn.
+ rewrite Int64.neg_zero.
+ rewrite Int64.and_commut.
+ rewrite Int64.and_zero.
+ rewrite Int64.sub_zero_r.
+ change (Int64.neg Int64.one) with Int64.mone.
+ rewrite Int64.and_commut.
+ rewrite Int64.and_mone.
+ rewrite Int64.or_commut.
+ apply Int64.or_zero.
+Qed.
+
+Definition select01_long (vb : val) (vtrue : val) (vfalse : val) : val :=
+ match vb with
+ | (Vint b) =>
+ if Int.eq b Int.one
+ then vtrue
+ else if Int.eq b Int.zero
+ then vfalse
+ else Vundef
+ | _ => Vundef
+ end.
+
+Lemma normalize_select01:
+ forall x y z, Val.normalize (select01_long x y z) AST.Tlong = select01_long x (Val.normalize y AST.Tlong) (Val.normalize z AST.Tlong).
+Proof.
+ unfold select01_long.
+ intros.
+ destruct x; cbn; trivial.
+ destruct (Int.eq i Int.one); trivial.
+ destruct (Int.eq i Int.zero); trivial.
+Qed.
+
+Lemma select01_long_true:
+ forall vt vf,
+ select01_long Vtrue vt vf = vt.
+Proof.
+ intros. unfold select01_long. cbn.
+ rewrite Int.eq_true. reflexivity.
+Qed.
+
+Lemma select01_long_false:
+ forall vt vf,
+ select01_long Vfalse vt vf = vf.
+Proof.
+ intros. unfold select01_long. cbn.
+ rewrite Int.eq_true.
+ rewrite Int.eq_false. reflexivity.
+ cbv. discriminate.
+Qed.
+
+Lemma float_bits_normalize:
+ forall v1,
+ ExtValues.float_of_bits (Val.normalize (ExtValues.bits_of_float v1) AST.Tlong) =
+ Val.normalize v1 AST.Tfloat.
+Proof.
+ destruct v1; cbn; trivial.
+ f_equal.
+ apply Float.of_to_bits.
+Qed.
+
+Lemma single_bits_normalize:
+ forall v1,
+ ExtValues.single_of_bits (Val.normalize (ExtValues.bits_of_single v1) AST.Tint) =
+ Val.normalize v1 AST.Tsingle.
+Proof.
+ destruct v1; cbn; trivial.
+ f_equal.
+ apply Float32.of_to_bits.
+Qed.
diff --git a/riscV/Machregsaux.ml b/riscV/Machregsaux.ml
index a48749a5..840943e7 100644
--- a/riscV/Machregsaux.ml
+++ b/riscV/Machregsaux.ml
@@ -13,3 +13,8 @@
(** Auxiliary functions on machine registers *)
let is_scratch_register r = false
+
+let class_of_type = function
+ | AST.Tint | AST.Tlong -> 0
+ | AST.Tfloat | AST.Tsingle -> 1
+ | AST.Tany32 | AST.Tany64 -> assert false
diff --git a/riscV/Machregsaux.mli b/riscV/Machregsaux.mli
index f3d52849..01b0f9fd 100644
--- a/riscV/Machregsaux.mli
+++ b/riscV/Machregsaux.mli
@@ -13,3 +13,5 @@
(** Auxiliary functions on machine registers *)
val is_scratch_register: string -> bool
+
+val class_of_type: AST.typ -> int
diff --git a/riscV/NeedOp.v b/riscV/NeedOp.v
index 4070431a..6041a34d 100644
--- a/riscV/NeedOp.v
+++ b/riscV/NeedOp.v
@@ -87,6 +87,45 @@ Definition needs_of_operation (op: operation) (nv: nval): list nval :=
| Ointofsingle | Ointuofsingle | Osingleofint | Osingleofintu => op1 (default nv)
| Olongofsingle | Olonguofsingle | Osingleoflong | Osingleoflongu => op1 (default nv)
| Ocmp c => needs_of_condition c
+ | OEseqw _ => op2 (default nv)
+ | OEsnew _ => op2 (default nv)
+ | OEsequw _ => op2 (default nv)
+ | OEsneuw _ => op2 (default nv)
+ | OEsltw _ => op2 (default nv)
+ | OEsltuw _ => op2 (default nv)
+ | OEsltiw _ => op1 (default nv)
+ | OEsltiuw _ => op1 (default nv)
+ | OExoriw _ => op1 (bitwise nv)
+ | OEluiw _ => op1 (default nv)
+ | OEaddiw _ _ => op1 (default nv)
+ | OEandiw n => op1 (andimm nv n)
+ | OEoriw n => op1 (orimm nv n)
+ | OEseql _ => op2 (default nv)
+ | OEsnel _ => op2 (default nv)
+ | OEsequl _ => op2 (default nv)
+ | OEsneul _ => op2 (default nv)
+ | OEsltl _ => op2 (default nv)
+ | OEsltul _ => op2 (default nv)
+ | OEsltil _ => op1 (default nv)
+ | OEsltiul _ => op1 (default nv)
+ | OExoril _ => op1 (default nv)
+ | OEluil _ => op1 (default nv)
+ | OEaddil _ _ => op1 (default nv)
+ | OEandil _ => op1 (default nv)
+ | OEoril _ => op1 (default nv)
+ | OEloadli _ => op1 (default nv)
+ | OEmayundef _ => op2 (default nv)
+ | OEfeqd => op2 (default nv)
+ | OEfltd => op2 (default nv)
+ | OEfled => op2 (default nv)
+ | OEfeqs => op2 (default nv)
+ | OEflts => op2 (default nv)
+ | OEfles => op2 (default nv)
+ | Obits_of_single => op1 (default nv)
+ | Obits_of_float => op1 (default nv)
+ | Osingle_of_bits => op1 (default nv)
+ | Ofloat_of_bits => op1 (default nv)
+ | Oselectl => All :: nv :: nv :: nil
end.
Definition operation_is_redundant (op: operation) (nv: nval): bool :=
@@ -154,6 +193,27 @@ Proof.
- apply shlimm_sound; auto.
- apply shrimm_sound; auto.
- apply shruimm_sound; auto.
+- fold (Val.and (Vint n) v0);
+ fold (Val.and (Vint n) v2);
+ rewrite (Val.and_commut (Vint n) v0);
+ rewrite (Val.and_commut (Vint n) v2);
+ apply andimm_sound; auto.
+- fold (Val.or (Vint n) v0);
+ fold (Val.or (Vint n) v2);
+ rewrite (Val.or_commut (Vint n) v0);
+ rewrite (Val.or_commut (Vint n) v2);
+ apply orimm_sound; auto.
+- apply xor_sound; auto with na.
+- (* selectl *)
+ unfold ExtValues.select01_long.
+ destruct v0; auto with na.
+ assert (Val.lessdef (Vint i) v4) as LESSDEF by auto with na.
+ inv LESSDEF.
+ destruct (Int.eq i Int.one).
+ { apply normalize_sound; auto. }
+ destruct (Int.eq i Int.zero).
+ { apply normalize_sound; auto. }
+ cbn. auto with na.
Qed.
Lemma operation_is_redundant_sound:
diff --git a/riscV/Op.v b/riscV/Op.v
index bb04f786..9f94828f 100644
--- a/riscV/Op.v
+++ b/riscV/Op.v
@@ -32,11 +32,18 @@
Require Import BoolEqual Coqlib.
Require Import AST Integers Floats.
Require Import Values Memory Globalenvs Events.
+Require ExtValues.
Set Implicit Arguments.
(** Conditions (boolean-valued operators). *)
+(** Type to modelize the use of a special register in arith operations *)
+
+Inductive oreg: Type :=
+ | X0_L: oreg
+ | X0_R: oreg.
+
Inductive condition : Type :=
| Ccomp (c: comparison) (**r signed integer comparison *)
| Ccompu (c: comparison) (**r unsigned integer comparison *)
@@ -49,7 +56,32 @@ Inductive condition : Type :=
| Ccompf (c: comparison) (**r 64-bit floating-point comparison *)
| Cnotcompf (c: comparison) (**r negation of a floating-point comparison *)
| Ccompfs (c: comparison) (**r 32-bit floating-point comparison *)
- | Cnotcompfs (c: comparison). (**r negation of a floating-point comparison *)
+ | Cnotcompfs (c: comparison) (**r negation of a floating-point comparison *)
+ (* Expansed branches *)
+ | CEbeqw (optR: option oreg) (**r branch-if-equal signed *)
+ | CEbnew (optR: option oreg) (**r branch-if-not-equal signed *)
+ | CEbequw (optR: option oreg) (**r branch-if-equal unsigned *)
+ | CEbneuw (optR: option oreg) (**r branch-if-not-equal unsigned *)
+ | CEbltw (optR: option oreg) (**r branch-if-less signed *)
+ | CEbltuw (optR: option oreg) (**r branch-if-less unsigned *)
+ | CEbgew (optR: option oreg) (**r branch-if-greater-or-equal signed *)
+ | CEbgeuw (optR: option oreg) (**r branch-if-greater-or-equal unsigned *)
+ | CEbeql (optR: option oreg) (**r branch-if-equal signed *)
+ | CEbnel (optR: option oreg) (**r branch-if-not-equal signed *)
+ | CEbequl (optR: option oreg) (**r branch-if-equal unsigned *)
+ | CEbneul (optR: option oreg) (**r branch-if-not-equal unsigned *)
+ | CEbltl (optR: option oreg) (**r branch-if-less signed *)
+ | CEbltul (optR: option oreg) (**r branch-if-less unsigned *)
+ | CEbgel (optR: option oreg) (**r branch-if-greater-or-equal signed *)
+ | CEbgeul (optR: option oreg). (**r branch-if-greater-or-equal unsigned *)
+
+(* This type will define the eval function of a OEmayundef operation. *)
+
+Inductive mayundef: Type :=
+ | MUint: mayundef
+ | MUlong: mayundef
+ | MUshrx: int -> mayundef
+ | MUshrxl: int -> mayundef.
(** Arithmetic and logical operations. In the descriptions, [rd] is the
result of the operation and [r1], [r2], etc, are the arguments. *)
@@ -152,7 +184,47 @@ Inductive operation : Type :=
| Osingleoflong (**r [rd = float32_of_signed_long(r1)] *)
| Osingleoflongu (**r [rd = float32_of_unsigned_int(r1)] *)
(*c Boolean tests: *)
- | Ocmp (cond: condition). (**r [rd = 1] if condition holds, [rd = 0] otherwise. *)
+ | Ocmp (cond: condition) (**r [rd = 1] if condition holds, [rd = 0] otherwise. *)
+ (* Expansed conditions *)
+ | OEseqw (optR: option oreg) (**r [rd <- rs1 == rs2] signed *)
+ | OEsnew (optR: option oreg) (**r [rd <- rs1 != rs2] signed *)
+ | OEsequw (optR: option oreg) (**r [rd <- rs1 == rs2] unsigned *)
+ | OEsneuw (optR: option oreg) (**r [rd <- rs1 != rs2] unsigned *)
+ | OEsltw (optR: option oreg) (**r set-less-than *)
+ | OEsltuw (optR: option oreg) (**r set-less-than unsigned *)
+ | OEsltiw (n: int) (**r set-less-than immediate *)
+ | OEsltiuw (n: int) (**r set-less-than unsigned immediate *)
+ | OEaddiw (optR: option oreg) (n: int) (**r add immediate *)
+ | OEandiw (n: int) (**r and immediate *)
+ | OEoriw (n: int) (**r or immediate *)
+ | OExoriw (n: int) (**r xor immediate *)
+ | OEluiw (n: int) (**r load upper-immediate *)
+ | OEseql (optR: option oreg) (**r [rd <- rs1 == rs2] signed *)
+ | OEsnel (optR: option oreg) (**r [rd <- rs1 != rs2] signed *)
+ | OEsequl (optR: option oreg) (**r [rd <- rs1 == rs2] unsigned *)
+ | OEsneul (optR: option oreg) (**r [rd <- rs1 != rs2] unsigned *)
+ | OEsltl (optR: option oreg) (**r set-less-than *)
+ | OEsltul (optR: option oreg) (**r set-less-than unsigned *)
+ | OEsltil (n: int64) (**r set-less-than immediate *)
+ | OEsltiul (n: int64) (**r set-less-than unsigned immediate *)
+ | OEaddil (optR: option oreg) (n: int64) (**r add immediate *)
+ | OEandil (n: int64) (**r and immediate *)
+ | OEoril (n: int64) (**r or immediate *)
+ | OExoril (n: int64) (**r xor immediate *)
+ | OEluil (n: int64) (**r load upper-immediate *)
+ | OEloadli (n: int64) (**r load an immediate int64 *)
+ | OEmayundef (mu: mayundef)
+ | OEfeqd (**r compare equal *)
+ | OEfltd (**r compare less-than *)
+ | OEfled (**r compare less-than/equal *)
+ | OEfeqs (**r compare equal *)
+ | OEflts (**r compare less-than *)
+ | OEfles (**r compare less-than/equal *)
+ | Obits_of_single
+ | Obits_of_float
+ | Osingle_of_bits
+ | Ofloat_of_bits
+ | Oselectl.
(** Addressing modes. [r1], [r2], etc, are the arguments to the
addressing. *)
@@ -164,11 +236,15 @@ Inductive addressing: Type :=
(** Comparison functions (used in modules [CSE] and [Allocation]). *)
+Definition oreg_eq: forall (x y: oreg), {x=y} + {x<>y}.
+Proof. decide equality. Defined.
+
Definition eq_condition (x y: condition) : {x=y} + {x<>y}.
Proof.
- generalize Int.eq_dec Int64.eq_dec; intro.
+ generalize Int.eq_dec Int64.eq_dec bool_dec oreg_eq; intros.
assert (forall (x y: comparison), {x=y}+{x<>y}). decide equality.
decide equality.
+ all: destruct optR, optR0; decide equality.
Defined.
Definition eq_addressing (x y: addressing) : {x=y} + {x<>y}.
@@ -179,8 +255,9 @@ Defined.
Definition eq_operation: forall (x y: operation), {x=y} + {x<>y}.
Proof.
- generalize Int.eq_dec Int64.eq_dec Ptrofs.eq_dec Float.eq_dec Float32.eq_dec ident_eq eq_condition; intros.
+ generalize Int.eq_dec Int64.eq_dec Ptrofs.eq_dec Float.eq_dec Float32.eq_dec ident_eq eq_condition bool_dec Val.eq oreg_eq; intros.
decide equality.
+ all: try destruct optR, optR0; try decide equality.
Defined.
(* Alternate definition:
@@ -197,6 +274,44 @@ Defined.
Global Opaque eq_condition eq_addressing eq_operation.
+(** Generic function to evaluate an instruction according to the given specific register *)
+
+Definition zero32 := (Vint Int.zero).
+Definition zero64 := (Vlong Int64.zero).
+
+Definition apply_bin_oreg {B} (optR: option oreg) (sem: val -> val -> B) (v1 v2 vz: val): B :=
+ match optR with
+ | None => sem v1 v2
+ | Some X0_L => sem vz v1
+ | Some X0_R => sem v1 vz
+ end.
+
+(** Mayundef evaluation according to the above defined type *)
+
+Definition eval_may_undef (mu: mayundef) (v1 v2: val): val :=
+ match mu with
+ | MUint => match v1, v2 with
+ | Vint _, Vint _ => v2
+ | _, _ => Vundef
+ end
+ | MUlong => match v1, v2 with
+ | Vlong _, Vint _ => v2
+ | _, _ => Vundef
+ end
+ | MUshrx i =>
+ match v1, v2 with
+ | Vint _, Vint _ =>
+ if Int.ltu i (Int.repr 31) then v2 else Vundef
+ | _, _ => Vundef
+ end
+ | MUshrxl i =>
+ match v1, v2 with
+ | Vlong _, Vlong _ =>
+ if Int.ltu i (Int.repr 63) then v2 else Vundef
+ | _, _ => Vundef
+ end
+ end.
+
(** * Evaluation functions *)
(** Evaluation of conditions, operators and addressing modes applied
@@ -218,9 +333,34 @@ Definition eval_condition (cond: condition) (vl: list val) (m: mem): option bool
| 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)
+ (* Expansed branches *)
+ | CEbeqw optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmp_bool Ceq) v1 v2 zero32
+ | CEbnew optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmp_bool Cne) v1 v2 zero32
+ | CEbequw optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmpu_bool (Mem.valid_pointer m) Ceq) v1 v2 zero32
+ | CEbneuw optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmpu_bool (Mem.valid_pointer m) Cne) v1 v2 zero32
+ | CEbltw optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmp_bool Clt) v1 v2 zero32
+ | CEbltuw optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmpu_bool (Mem.valid_pointer m) Clt) v1 v2 zero32
+ | CEbgew optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmp_bool Cge) v1 v2 zero32
+ | CEbgeuw optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmpu_bool (Mem.valid_pointer m) Cge) v1 v2 zero32
+ | CEbeql optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmpl_bool Ceq) v1 v2 zero64
+ | CEbnel optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmpl_bool Cne) v1 v2 zero64
+ | CEbequl optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmplu_bool (Mem.valid_pointer m) Ceq) v1 v2 zero64
+ | CEbneul optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmplu_bool (Mem.valid_pointer m) Cne) v1 v2 zero64
+ | CEbltl optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmpl_bool Clt) v1 v2 zero64
+ | CEbltul optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmplu_bool (Mem.valid_pointer m) Clt) v1 v2 zero64
+ | CEbgel optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmpl_bool Cge) v1 v2 zero64
+ | CEbgeul optR, v1 :: v2 :: nil => apply_bin_oreg optR (Val.cmplu_bool (Mem.valid_pointer m) Cge) v1 v2 zero64
| _, _ => None
end.
+(** Assert sp is a pointer *)
+
+Definition get_sp sp :=
+ match sp with
+ | Vptr _ _ => sp
+ | _ => Vundef
+ end.
+
Definition eval_operation
(F V: Type) (genv: Genv.t F V) (sp: val)
(op: operation) (vl: list val) (m: mem): option val :=
@@ -241,10 +381,10 @@ Definition eval_operation
| Omul, v1 :: v2 :: nil => Some (Val.mul v1 v2)
| 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
+ | Odiv, v1 :: v2 :: nil => Some (Val.maketotal (Val.divs v1 v2))
+ | Odivu, v1 :: v2 :: nil => Some (Val.maketotal (Val.divu v1 v2))
+ | Omod, v1 :: v2 :: nil => Some (Val.maketotal (Val.mods v1 v2))
+ | Omodu, v1 :: v2 :: nil => Some (Val.maketotal (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)
@@ -257,7 +397,7 @@ Definition eval_operation
| Oshrimm n, v1 :: nil => Some (Val.shr v1 (Vint n))
| Oshru, v1 :: v2 :: nil => Some (Val.shru v1 v2)
| Oshruimm n, v1 :: nil => Some (Val.shru v1 (Vint n))
- | Oshrximm n, v1::nil => Val.shrx v1 (Vint n)
+ | Oshrximm n, v1::nil => Some (Val.maketotal (Val.shrx v1 (Vint n)))
| Omakelong, v1::v2::nil => Some (Val.longofwords v1 v2)
| Olowlong, v1::nil => Some (Val.loword v1)
| Ohighlong, v1::nil => Some (Val.hiword v1)
@@ -270,10 +410,10 @@ Definition eval_operation
| Omull, v1::v2::nil => Some (Val.mull v1 v2)
| Omullhs, v1::v2::nil => Some (Val.mullhs v1 v2)
| Omullhu, v1::v2::nil => Some (Val.mullhu v1 v2)
- | Odivl, v1::v2::nil => Val.divls v1 v2
- | Odivlu, v1::v2::nil => Val.divlu v1 v2
- | Omodl, v1::v2::nil => Val.modls v1 v2
- | Omodlu, v1::v2::nil => Val.modlu v1 v2
+ | Odivl, v1::v2::nil => Some (Val.maketotal (Val.divls v1 v2))
+ | Odivlu, v1::v2::nil => Some (Val.maketotal (Val.divlu v1 v2))
+ | Omodl, v1::v2::nil => Some (Val.maketotal (Val.modls v1 v2))
+ | Omodlu, v1::v2::nil => Some (Val.maketotal (Val.modlu v1 v2))
| Oandl, v1::v2::nil => Some(Val.andl v1 v2)
| Oandlimm n, v1::nil => Some (Val.andl v1 (Vlong n))
| Oorl, v1::v2::nil => Some(Val.orl v1 v2)
@@ -286,7 +426,7 @@ Definition eval_operation
| Oshrlimm n, v1::nil => Some (Val.shrl v1 (Vint n))
| Oshrlu, v1::v2::nil => Some (Val.shrlu v1 v2)
| Oshrluimm n, v1::nil => Some (Val.shrlu v1 (Vint n))
- | Oshrxlimm n, v1::nil => Val.shrxl v1 (Vint n)
+ | Oshrxlimm n, v1::nil => Some (Val.maketotal (Val.shrxl v1 (Vint n)))
| Onegf, v1::nil => Some (Val.negf v1)
| Oabsf, v1::nil => Some (Val.absf v1)
| Oaddf, v1::v2::nil => Some (Val.addf v1 v2)
@@ -301,23 +441,65 @@ Definition eval_operation
| 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
- | Ointuoffloat, v1::nil => Val.intuoffloat v1
- | Ofloatofint, v1::nil => Val.floatofint v1
- | Ofloatofintu, v1::nil => Val.floatofintu v1
- | Ointofsingle, v1::nil => Val.intofsingle v1
- | Ointuofsingle, v1::nil => Val.intuofsingle v1
- | Osingleofint, v1::nil => Val.singleofint v1
- | Osingleofintu, v1::nil => Val.singleofintu v1
- | Olongoffloat, v1::nil => Val.longoffloat v1
- | Olonguoffloat, v1::nil => Val.longuoffloat v1
- | Ofloatoflong, v1::nil => Val.floatoflong v1
- | Ofloatoflongu, v1::nil => Val.floatoflongu v1
- | Olongofsingle, v1::nil => Val.longofsingle v1
- | Olonguofsingle, v1::nil => Val.longuofsingle v1
- | Osingleoflong, v1::nil => Val.singleoflong v1
- | Osingleoflongu, v1::nil => Val.singleoflongu v1
+ | Ointoffloat, v1::nil => Some (Val.maketotal (Val.intoffloat v1))
+ | Ointuoffloat, v1::nil => Some (Val.maketotal (Val.intuoffloat v1))
+ | Ofloatofint, v1::nil => Some (Val.maketotal (Val.floatofint v1))
+ | Ofloatofintu, v1::nil => Some (Val.maketotal (Val.floatofintu v1))
+ | Ointofsingle, v1::nil => Some (Val.maketotal (Val.intofsingle v1))
+ | Ointuofsingle, v1::nil => Some (Val.maketotal (Val.intuofsingle v1))
+ | Osingleofint, v1::nil => Some (Val.maketotal (Val.singleofint v1))
+ | Osingleofintu, v1::nil => Some (Val.maketotal (Val.singleofintu v1))
+ | Olongoffloat, v1::nil => Some (Val.maketotal (Val.longoffloat v1))
+ | Olonguoffloat, v1::nil => Some (Val.maketotal (Val.longuoffloat v1))
+ | Ofloatoflong, v1::nil => Some (Val.maketotal (Val.floatoflong v1))
+ | Ofloatoflongu, v1::nil => Some (Val.maketotal (Val.floatoflongu v1))
+ | Olongofsingle, v1::nil => Some (Val.maketotal (Val.longofsingle v1))
+ | Olonguofsingle, v1::nil => Some (Val.maketotal (Val.longuofsingle v1))
+ | Osingleoflong, v1::nil => Some (Val.maketotal (Val.singleoflong v1))
+ | Osingleoflongu, v1::nil => Some (Val.maketotal (Val.singleoflongu v1))
+ | Obits_of_single, v1::nil => Some (ExtValues.bits_of_single v1)
+ | Obits_of_float, v1::nil => Some (ExtValues.bits_of_float v1)
+ | Osingle_of_bits, v1::nil => Some (ExtValues.single_of_bits v1)
+ | Ofloat_of_bits, v1::nil => Some (ExtValues.float_of_bits v1)
| Ocmp c, _ => Some (Val.of_optbool (eval_condition c vl m))
+ (* Expansed conditions *)
+ | OEseqw optR, v1::v2::nil => Some (apply_bin_oreg optR (Val.cmp Ceq) v1 v2 zero32)
+ | OEsnew optR, v1::v2::nil => Some (apply_bin_oreg optR (Val.cmp Cne) v1 v2 zero32)
+ | OEsequw optR, v1::v2::nil => Some (apply_bin_oreg optR (Val.cmpu (Mem.valid_pointer m) Ceq) v1 v2 zero32)
+ | OEsneuw optR, v1::v2::nil => Some (apply_bin_oreg optR (Val.cmpu (Mem.valid_pointer m) Cne) v1 v2 zero32)
+ | OEsltw optR, v1::v2::nil => Some (apply_bin_oreg optR (Val.cmp Clt) v1 v2 zero32)
+ | OEsltuw optR, v1::v2::nil => Some (apply_bin_oreg optR (Val.cmpu (Mem.valid_pointer m) Clt) v1 v2 zero32)
+ | OEsltiw n, v1::nil => Some (Val.cmp Clt v1 (Vint n))
+ | OEsltiuw n, v1::nil => Some (Val.cmpu (Mem.valid_pointer m) Clt v1 (Vint n))
+ | OExoriw n, v1::nil => Some (Val.xor v1 (Vint n))
+ | OEluiw n, nil => Some (Val.shl (Vint n) (Vint (Int.repr 12)))
+ | OEaddiw optR n, nil => Some (apply_bin_oreg optR Val.add (Vint n) Vundef zero32)
+ | OEaddiw optR n, v1::nil => Some (apply_bin_oreg optR Val.add v1 (Vint n) Vundef)
+ | OEandiw n, v1::nil => Some (Val.and (Vint n) v1)
+ | OEoriw n, v1::nil => Some (Val.or (Vint n) v1)
+ | OEseql optR, v1::v2::nil => Some (Val.maketotal (apply_bin_oreg optR (Val.cmpl Ceq) v1 v2 zero64))
+ | OEsnel optR, v1::v2::nil => Some (Val.maketotal (apply_bin_oreg optR (Val.cmpl Cne) v1 v2 zero64))
+ | OEsequl optR, v1::v2::nil => Some (Val.maketotal (apply_bin_oreg optR (Val.cmplu (Mem.valid_pointer m) Ceq) v1 v2 zero64))
+ | OEsneul optR, v1::v2::nil => Some (Val.maketotal (apply_bin_oreg optR (Val.cmplu (Mem.valid_pointer m) Cne) v1 v2 zero64))
+ | OEsltl optR, v1::v2::nil => Some (Val.maketotal (apply_bin_oreg optR (Val.cmpl Clt) v1 v2 zero64))
+ | OEsltul optR, v1::v2::nil => Some (Val.maketotal (apply_bin_oreg optR (Val.cmplu (Mem.valid_pointer m) Clt) v1 v2 zero64))
+ | OEsltil n, v1::nil => Some (Val.maketotal (Val.cmpl Clt v1 (Vlong n)))
+ | OEsltiul n, v1::nil => Some (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Clt v1 (Vlong n)))
+ | OExoril n, v1::nil => Some (Val.xorl v1 (Vlong n))
+ | OEluil n, nil => Some (Vlong (Int64.sign_ext 32 (Int64.shl n (Int64.repr 12))))
+ | OEaddil optR n, nil => Some (apply_bin_oreg optR Val.addl (Vlong n) Vundef zero64)
+ | OEaddil optR n, v1::nil => Some (apply_bin_oreg optR Val.addl v1 (Vlong n) Vundef)
+ | OEandil n, v1::nil => Some (Val.andl (Vlong n) v1)
+ | OEoril n, v1::nil => Some (Val.orl (Vlong n) v1)
+ | OEloadli n, nil => Some (Vlong n)
+ | OEmayundef mu, v1 :: v2 :: nil => Some (eval_may_undef mu v1 v2)
+ | OEfeqd, v1::v2::nil => Some (Val.cmpf Ceq v1 v2)
+ | OEfltd, v1::v2::nil => Some (Val.cmpf Clt v1 v2)
+ | OEfled, v1::v2::nil => Some (Val.cmpf Cle v1 v2)
+ | OEfeqs, v1::v2::nil => Some (Val.cmpfs Ceq v1 v2)
+ | OEflts, v1::v2::nil => Some (Val.cmpfs Clt v1 v2)
+ | OEfles, v1::v2::nil => Some (Val.cmpfs Cle v1 v2)
+ | Oselectl, vb::vt::vf::nil => Some (Val.normalize (ExtValues.select01_long vb vt vf) Tlong)
| _, _ => None
end.
@@ -348,9 +530,9 @@ Qed.
Ltac FuncInv :=
match goal with
| H: (match ?x with nil => _ | _ :: _ => _ end = Some _) |- _ =>
- destruct x; simpl in H; FuncInv
+ destruct x; cbn in H; FuncInv
| H: (match ?v with Vundef => _ | Vint _ => _ | Vfloat _ => _ | Vptr _ _ => _ end = Some _) |- _ =>
- destruct v; simpl in H; FuncInv
+ destruct v; cbn in H; FuncInv
| H: (if Archi.ptr64 then _ else _) = Some _ |- _ =>
destruct Archi.ptr64 eqn:?; FuncInv
| H: (Some _ = Some _) |- _ =>
@@ -377,6 +559,31 @@ Definition type_of_condition (c: condition) : list typ :=
| Cnotcompf _ => Tfloat :: Tfloat :: nil
| Ccompfs _ => Tsingle :: Tsingle :: nil
| Cnotcompfs _ => Tsingle :: Tsingle :: nil
+ | CEbeqw _ => Tint :: Tint :: nil
+ | CEbnew _ => Tint :: Tint :: nil
+ | CEbequw _ => Tint :: Tint :: nil
+ | CEbneuw _ => Tint :: Tint :: nil
+ | CEbltw _ => Tint :: Tint :: nil
+ | CEbltuw _ => Tint :: Tint :: nil
+ | CEbgew _ => Tint :: Tint :: nil
+ | CEbgeuw _ => Tint :: Tint :: nil
+ | CEbeql _ => Tlong :: Tlong :: nil
+ | CEbnel _ => Tlong :: Tlong :: nil
+ | CEbequl _ => Tlong :: Tlong :: nil
+ | CEbneul _ => Tlong :: Tlong :: nil
+ | CEbltl _ => Tlong :: Tlong :: nil
+ | CEbltul _ => Tlong :: Tlong :: nil
+ | CEbgel _ => Tlong :: Tlong :: nil
+ | CEbgeul _ => Tlong :: Tlong :: nil
+ end.
+
+(** The type of mayundef and addsp is dynamic *)
+
+Definition type_of_mayundef mu :=
+ match mu with
+ | MUint | MUshrx _ => (Tint :: Tint :: nil, Tint)
+ | MUlong => (Tlong :: Tint :: nil, Tint)
+ | MUshrxl _ => (Tlong :: Tlong :: nil, Tlong)
end.
Definition type_of_operation (op: operation) : list typ * typ :=
@@ -474,6 +681,47 @@ Definition type_of_operation (op: operation) : list typ * typ :=
| Osingleoflong => (Tlong :: nil, Tsingle)
| Osingleoflongu => (Tlong :: nil, Tsingle)
| Ocmp c => (type_of_condition c, Tint)
+ | OEseqw _ => (Tint :: Tint :: nil, Tint)
+ | OEsnew _ => (Tint :: Tint :: nil, Tint)
+ | OEsequw _ => (Tint :: Tint :: nil, Tint)
+ | OEsneuw _ => (Tint :: Tint :: nil, Tint)
+ | OEsltw _ => (Tint :: Tint :: nil, Tint)
+ | OEsltuw _ => (Tint :: Tint :: nil, Tint)
+ | OEsltiw _ => (Tint :: nil, Tint)
+ | OEsltiuw _ => (Tint :: nil, Tint)
+ | OExoriw _ => (Tint :: nil, Tint)
+ | OEluiw _ => (nil, Tint)
+ | OEaddiw None _ => (Tint :: nil, Tint)
+ | OEaddiw (Some _) _ => (nil, Tint)
+ | OEandiw _ => (Tint :: nil, Tint)
+ | OEoriw _ => (Tint :: nil, Tint)
+ | OEseql _ => (Tlong :: Tlong :: nil, Tint)
+ | OEsnel _ => (Tlong :: Tlong :: nil, Tint)
+ | OEsequl _ => (Tlong :: Tlong :: nil, Tint)
+ | OEsneul _ => (Tlong :: Tlong :: nil, Tint)
+ | OEsltl _ => (Tlong :: Tlong :: nil, Tint)
+ | OEsltul _ => (Tlong :: Tlong :: nil, Tint)
+ | OEsltil _ => (Tlong :: nil, Tint)
+ | OEsltiul _ => (Tlong :: nil, Tint)
+ | OEandil _ => (Tlong :: nil, Tlong)
+ | OEoril _ => (Tlong :: nil, Tlong)
+ | OExoril _ => (Tlong :: nil, Tlong)
+ | OEluil _ => (nil, Tlong)
+ | OEaddil None _ => (Tlong :: nil, Tlong)
+ | OEaddil (Some _) _ => (nil, Tlong)
+ | OEloadli _ => (nil, Tlong)
+ | OEmayundef mu => type_of_mayundef mu
+ | OEfeqd => (Tfloat :: Tfloat :: nil, Tint)
+ | OEfltd => (Tfloat :: Tfloat :: nil, Tint)
+ | OEfled => (Tfloat :: Tfloat :: nil, Tint)
+ | OEfeqs => (Tsingle :: Tsingle :: nil, Tint)
+ | OEflts => (Tsingle :: Tsingle :: nil, Tint)
+ | OEfles => (Tsingle :: Tsingle :: nil, Tint)
+ | Obits_of_single => (Tsingle :: nil, Tint)
+ | Obits_of_float => (Tfloat :: nil, Tlong)
+ | Osingle_of_bits => (Tint :: nil, Tsingle)
+ | Ofloat_of_bits => (Tlong :: nil, Tfloat)
+ | Oselectl => (Tint :: Tlong :: Tlong :: nil, Tlong)
end.
Definition type_of_addressing (addr: addressing) : list typ :=
@@ -504,6 +752,14 @@ Proof.
intros. unfold Val.has_type, Val.addl. destruct Archi.ptr64, v1, v2; auto.
Qed.
+Remark type_mayundef:
+ forall mu v1 v2, Val.has_type (eval_may_undef mu v1 v2) (snd (type_of_mayundef mu)).
+Proof.
+ intros. unfold eval_may_undef.
+ destruct mu eqn:EQMU, v1, v2; simpl; auto.
+ all: destruct Int.ltu; simpl; auto.
+Qed.
+
Lemma type_of_operation_sound:
forall op vl sp v m,
op <> Omove ->
@@ -513,7 +769,7 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
intros.
destruct op; simpl; simpl in H0; FuncInv; subst; simpl.
(* move *)
- - congruence.
+ - simpl in H; congruence.
(* intconst, longconst, floatconst, singleconst *)
- exact I.
- exact I.
@@ -539,15 +795,17 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
- destruct v0; destruct v1...
- destruct v0; destruct v1...
(* div, divu *)
- - 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; cbn; trivial.
+ destruct (Int.eq i0 Int.zero
+ || Int.eq i (Int.repr (-2147483648)) && Int.eq i0 Int.mone); cbn; trivial.
+ - destruct v0; destruct v1; cbn; trivial.
+ destruct (Int.eq i0 Int.zero); cbn; trivial.
(* mod, modu *)
- - 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; cbn; trivial.
+ destruct (Int.eq i0 Int.zero
+ || Int.eq i (Int.repr (-2147483648)) && Int.eq i0 Int.mone); cbn; trivial.
+ - destruct v0; destruct v1; cbn; trivial.
+ destruct (Int.eq i0 Int.zero); cbn; trivial.
(* and, andimm *)
- destruct v0; destruct v1...
- destruct v0...
@@ -567,7 +825,8 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
- destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int.iwordsize)...
- destruct v0; simpl... destruct (Int.ltu n Int.iwordsize)...
(* shrx *)
- - destruct v0; simpl in H0; try discriminate. destruct (Int.ltu n (Int.repr 31)); inv H0...
+ - destruct v0; cbn; trivial.
+ destruct (Int.ltu n (Int.repr 31)); cbn; trivial.
(* makelong, lowlong, highlong *)
- destruct v0; destruct v1...
- destruct v0...
@@ -588,15 +847,19 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
- destruct v0; destruct v1...
- destruct v0; destruct v1...
(* divl, divlu *)
- - destruct v0; destruct v1; simpl in *; inv H0.
- destruct (Int64.eq i0 Int64.zero || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2...
- - destruct v0; destruct v1; simpl in *; inv H0.
- destruct (Int64.eq i0 Int64.zero); inv H2...
+ - destruct v0; destruct v1; cbn; trivial.
+ destruct (Int64.eq i0 Int64.zero
+ || Int64.eq i (Int64.repr (-9223372036854775808)) &&
+ Int64.eq i0 Int64.mone); cbn; trivial.
+ - destruct v0; destruct v1; cbn; trivial.
+ destruct (Int64.eq i0 Int64.zero); cbn; trivial.
(* modl, modlu *)
- - destruct v0; destruct v1; simpl in *; inv H0.
- destruct (Int64.eq i0 Int64.zero || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2...
- - destruct v0; destruct v1; simpl in *; inv H0.
- destruct (Int64.eq i0 Int64.zero); inv H2...
+ - destruct v0; destruct v1; cbn; trivial.
+ destruct (Int64.eq i0 Int64.zero
+ || Int64.eq i (Int64.repr (-9223372036854775808)) &&
+ Int64.eq i0 Int64.mone); cbn; trivial.
+ - destruct v0; destruct v1; cbn; trivial.
+ destruct (Int64.eq i0 Int64.zero); cbn; trivial.
(* andl, andlimm *)
- destruct v0; destruct v1...
- destruct v0...
@@ -616,7 +879,8 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
- destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int64.iwordsize')...
- destruct v0; simpl... destruct (Int.ltu n Int64.iwordsize')...
(* shrxl *)
- - destruct v0; simpl in H0; try discriminate. destruct (Int.ltu n (Int.repr 63)); inv H0...
+ - destruct v0; cbn; trivial.
+ destruct (Int.ltu n (Int.repr 63)); cbn; trivial.
(* negf, absf *)
- destruct v0...
- destruct v0...
@@ -639,33 +903,173 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
- destruct v0...
- destruct v0...
(* intoffloat, intuoffloat *)
- - destruct v0; simpl in H0; inv H0. destruct (Float.to_int f); inv H2...
- - destruct v0; simpl in H0; inv H0. destruct (Float.to_intu f); inv H2...
+ - destruct v0; cbn; trivial.
+ destruct (Float.to_int f); cbn; trivial.
+ - destruct v0; cbn; trivial.
+ destruct (Float.to_intu f); cbn; trivial.
(* floatofint, floatofintu *)
- - destruct v0; simpl in H0; inv H0...
- - destruct v0; simpl in H0; inv H0...
+ - destruct v0; cbn; trivial.
+ - destruct v0; cbn; trivial.
(* intofsingle, intuofsingle *)
- - destruct v0; simpl in H0; inv H0. destruct (Float32.to_int f); inv H2...
- - destruct v0; simpl in H0; inv H0. destruct (Float32.to_intu f); inv H2...
+ - destruct v0; cbn; trivial.
+ destruct (Float32.to_int f); cbn; trivial.
+ - destruct v0; cbn; trivial.
+ destruct (Float32.to_intu f); cbn; trivial.
(* singleofint, singleofintu *)
- - destruct v0; simpl in H0; inv H0...
- - destruct v0; simpl in H0; inv H0...
+ - destruct v0; cbn; trivial.
+ - destruct v0; cbn; trivial.
(* longoffloat, longuoffloat *)
- - destruct v0; simpl in H0; inv H0. destruct (Float.to_long f); inv H2...
- - destruct v0; simpl in H0; inv H0. destruct (Float.to_longu f); inv H2...
+ - destruct v0; cbn; trivial.
+ destruct (Float.to_long f); cbn; trivial.
+ - destruct v0; cbn; trivial.
+ destruct (Float.to_longu f); cbn; trivial.
(* floatoflong, floatoflongu *)
- - destruct v0; simpl in H0; inv H0...
- - destruct v0; simpl in H0; inv H0...
+ - destruct v0; cbn; trivial.
+ - destruct v0; cbn; trivial.
(* longofsingle, longuofsingle *)
- - destruct v0; simpl in H0; inv H0. destruct (Float32.to_long f); inv H2...
- - destruct v0; simpl in H0; inv H0. destruct (Float32.to_longu f); inv H2...
+ - destruct v0; cbn; trivial.
+ destruct (Float32.to_long f); cbn; trivial.
+ - destruct v0; cbn; trivial.
+ destruct (Float32.to_longu f); cbn; trivial.
(* singleoflong, singleoflongu *)
- - destruct v0; simpl in H0; inv H0...
- - destruct v0; simpl in H0; inv H0...
+ - destruct v0; cbn; trivial.
+ - destruct v0; cbn; trivial.
(* cmp *)
- destruct (eval_condition cond vl m)... destruct b...
+ (* OEseqw *)
+ - destruct optR as [[]|]; simpl; unfold Val.cmp;
+ destruct Val.cmp_bool... all: destruct b...
+ (* OEsnew *)
+ - destruct optR as [[]|]; simpl; unfold Val.cmp;
+ destruct Val.cmp_bool... all: destruct b...
+ (* OEsequw *)
+ - destruct optR as [[]|]; simpl; unfold Val.cmpu;
+ destruct Val.cmpu_bool... all: destruct b...
+ (* OEsneuw *)
+ - destruct optR as [[]|]; simpl; unfold Val.cmpu;
+ destruct Val.cmpu_bool... all: destruct b...
+ (* OEsltw *)
+ - destruct optR as [[]|]; simpl; unfold Val.cmp;
+ destruct Val.cmp_bool... all: destruct b...
+ (* OEsltuw *)
+ - destruct optR as [[]|]; simpl; unfold Val.cmpu;
+ destruct Val.cmpu_bool... all: destruct b...
+ (* OEsltiw *)
+ - unfold Val.cmp; destruct Val.cmp_bool...
+ all: destruct b...
+ (* OEsltiuw *)
+ - unfold Val.cmpu; destruct Val.cmpu_bool... destruct b...
+ (* OEaddiw *)
+ - destruct optR as [[]|]; simpl in *; trivial.
+ - destruct optR as [[]|]; simpl in *; trivial;
+ apply type_add.
+ (* OEandiw *)
+ - destruct v0...
+ (* OEoriw *)
+ - destruct v0...
+ (* OExoriw *)
+ - destruct v0...
+ (* OEluiw *)
+ - destruct (Int.ltu _ _); cbn; trivial.
+ (* OEseql *)
+ - destruct optR as [[]|]; simpl; unfold Val.cmpl;
+ destruct Val.cmpl_bool... all: destruct b...
+ (* OEsnel *)
+ - destruct optR as [[]|]; simpl; unfold Val.cmpl;
+ destruct Val.cmpl_bool... all: destruct b...
+ (* OEsequl *)
+ - destruct optR as [[]|]; simpl; unfold Val.cmplu;
+ destruct Val.cmplu_bool... all: destruct b...
+ (* OEsneul *)
+ - destruct optR as [[]|]; simpl; unfold Val.cmplu;
+ destruct Val.cmplu_bool... all: destruct b...
+ (* OEsltl *)
+ - destruct optR as [[]|]; simpl; unfold Val.cmpl;
+ destruct Val.cmpl_bool... all: destruct b...
+ (* OEsltul *)
+ - destruct optR as [[]|]; simpl; unfold Val.cmplu;
+ destruct Val.cmplu_bool... all: destruct b...
+ (* OEsltil *)
+ - unfold Val.cmpl; destruct Val.cmpl_bool...
+ all: destruct b...
+ (* OEsltiul *)
+ - unfold Val.cmplu; destruct Val.cmplu_bool... destruct b...
+ (* OEaddil *)
+ - destruct optR as [[]|]; simpl in *; trivial.
+ - destruct optR as [[]|]; simpl in *; trivial;
+ apply type_addl.
+ (* OEandil *)
+ - destruct v0...
+ (* OEoril *)
+ - destruct v0...
+ (* OExoril *)
+ - destruct v0...
+ (* OEluil *)
+ - simpl; trivial.
+ (* OEloadli *)
+ - trivial.
+ (* OEmayundef *)
+ - apply type_mayundef.
+ (* OEfeqd *)
+ - destruct v0; destruct v1; cbn; auto.
+ destruct Float.cmp; cbn; auto.
+ (* OEfltd *)
+ - destruct v0; destruct v1; cbn; auto.
+ destruct Float.cmp; cbn; auto.
+ (* OEfled *)
+ - destruct v0; destruct v1; cbn; auto.
+ destruct Float.cmp; cbn; auto.
+ (* OEfeqs *)
+ - destruct v0; destruct v1; cbn; auto.
+ destruct Float32.cmp; cbn; auto.
+ (* OEflts *)
+ - destruct v0; destruct v1; cbn; auto.
+ destruct Float32.cmp; cbn; auto.
+ (* OEfles *)
+ - destruct v0; destruct v1; cbn; auto.
+ destruct Float32.cmp; cbn; auto.
+ (* Bits_of_single, float *)
+ - destruct v0; cbn; trivial.
+ - destruct v0; cbn; trivial.
+ (* single, float of bits *)
+ - destruct v0; cbn; trivial.
+ - destruct v0; cbn; trivial.
+ (* selectl *)
+ - destruct v0; cbn; trivial.
+ destruct Int.eq; cbn.
+ apply Val.normalize_type.
+ destruct Int.eq; cbn; trivial.
+ apply Val.normalize_type.
Qed.
+(* This should not be simplified to "false" because it breaks proofs elsewhere. *)
+Definition is_trapping_op (op : operation) :=
+ match op with
+ | Omove => false
+ | _ => false
+ end.
+
+Definition args_of_operation op :=
+ if eq_operation op Omove
+ then 1%nat
+ else List.length (fst (type_of_operation op)).
+
+Lemma is_trapping_op_sound:
+ forall op vl sp m,
+ is_trapping_op op = false ->
+ (List.length vl) = args_of_operation op ->
+ eval_operation genv sp op vl m <> None.
+Proof.
+ unfold args_of_operation.
+ destruct op eqn:E; destruct eq_operation; intros; simpl in *; try congruence.
+ all: try (destruct vl as [ | vh1 vl1]; try discriminate).
+ all: try (destruct vl1 as [ | vh2 vl2]; try discriminate).
+ all: try (destruct vl2 as [ | vh3 vl3]; try discriminate).
+ all: try (destruct vl3 as [ | vh4 vl4]; try discriminate).
+ all: try destruct optR as [[]|]; simpl in H0; try discriminate.
+ all: try destruct Archi.ptr64; simpl in *; try discriminate.
+ all: try destruct mu; simpl in *; try discriminate.
+Qed.
End SOUNDNESS.
(** * Manipulating and transforming operations *)
@@ -708,6 +1112,22 @@ Definition negate_condition (cond: condition): condition :=
| Cnotcompf c => Ccompf c
| Ccompfs c => Cnotcompfs c
| Cnotcompfs c => Ccompfs c
+ | CEbeqw optR => CEbnew optR
+ | CEbnew optR => CEbeqw optR
+ | CEbequw optR => CEbneuw optR
+ | CEbneuw optR => CEbequw optR
+ | CEbltw optR => CEbgew optR
+ | CEbltuw optR => CEbgeuw optR
+ | CEbgew optR => CEbltw optR
+ | CEbgeuw optR => CEbltuw optR
+ | CEbeql optR => CEbnel optR
+ | CEbnel optR => CEbeql optR
+ | CEbequl optR => CEbneul optR
+ | CEbneul optR => CEbequl optR
+ | CEbltl optR => CEbgel optR
+ | CEbltul optR => CEbgeul optR
+ | CEbgel optR => CEbltl optR
+ | CEbgeul optR => CEbltul optR
end.
Lemma eval_negate_condition:
@@ -727,6 +1147,39 @@ Proof.
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.
+
+ repeat (destruct vl; auto); replace (Cne) with (negate_comparison Ceq) by auto; destruct optR as [[]|];
+ apply Val.negate_cmp_bool.
+ repeat (destruct vl; auto); replace (Ceq) with (negate_comparison Cne) by auto; destruct optR as [[]|];
+ apply Val.negate_cmp_bool.
+ repeat (destruct vl; auto); replace (Cne) with (negate_comparison Ceq) by auto; destruct optR as [[]|];
+ apply Val.negate_cmpu_bool.
+ repeat (destruct vl; auto); replace (Ceq) with (negate_comparison Cne) by auto; destruct optR as [[]|];
+ apply Val.negate_cmpu_bool.
+ repeat (destruct vl; auto); replace (Cge) with (negate_comparison Clt) by auto; destruct optR as [[]|];
+ apply Val.negate_cmp_bool.
+ repeat (destruct vl; auto); replace (Cge) with (negate_comparison Clt) by auto; destruct optR as [[]|];
+ apply Val.negate_cmpu_bool.
+ repeat (destruct vl; auto); replace (Clt) with (negate_comparison Cge) by auto; destruct optR as [[]|];
+ apply Val.negate_cmp_bool.
+ repeat (destruct vl; auto); replace (Clt) with (negate_comparison Cge) by auto; destruct optR as [[]|];
+ apply Val.negate_cmpu_bool.
+ repeat (destruct vl; auto); replace (Cne) with (negate_comparison Ceq) by auto; destruct optR as [[]|];
+ apply Val.negate_cmpl_bool.
+ repeat (destruct vl; auto); replace (Ceq) with (negate_comparison Cne) by auto; destruct optR as [[]|];
+ apply Val.negate_cmpl_bool.
+ repeat (destruct vl; auto); replace (Cne) with (negate_comparison Ceq) by auto; destruct optR as [[]|];
+ apply Val.negate_cmplu_bool.
+ repeat (destruct vl; auto); replace (Ceq) with (negate_comparison Cne) by auto; destruct optR as [[]|];
+ apply Val.negate_cmplu_bool.
+ repeat (destruct vl; auto); replace (Cge) with (negate_comparison Clt) by auto; destruct optR as [[]|];
+ apply Val.negate_cmpl_bool.
+ repeat (destruct vl; auto); replace (Cge) with (negate_comparison Clt) by auto; destruct optR as [[]|];
+ apply Val.negate_cmplu_bool.
+ repeat (destruct vl; auto); replace (Clt) with (negate_comparison Cge) by auto; destruct optR as [[]|];
+ apply Val.negate_cmpl_bool.
+ repeat (destruct vl; auto); replace (Clt) with (negate_comparison Cge) by auto; destruct optR as [[]|];
+ apply Val.negate_cmplu_bool.
Qed.
(** Shifting stack-relative references. This is used in [Stacking]. *)
@@ -752,7 +1205,8 @@ 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.
+ intros. destruct op; auto;
+ try destruct optR as [[]|]; simpl; auto.
Qed.
Lemma eval_shift_stack_addressing:
@@ -769,7 +1223,7 @@ Lemma eval_shift_stack_operation:
eval_operation ge (Vptr sp Ptrofs.zero) (shift_stack_operation delta op) vl m =
eval_operation ge (Vptr sp (Ptrofs.repr delta)) op vl m.
Proof.
- intros. destruct op; simpl; auto. destruct vl; auto.
+ intros. destruct op eqn:E; simpl; auto; destruct vl; auto.
rewrite Ptrofs.add_zero_l, Ptrofs.add_commut; auto.
Qed.
@@ -817,23 +1271,87 @@ Definition is_trivial_op (op: operation) : bool :=
(** Operations that depend on the memory state. *)
+Definition cond_depends_on_memory (cond : condition) : bool :=
+ match cond with
+ | Ccompu _ => negb Archi.ptr64
+ | Ccompuimm _ _ => negb Archi.ptr64
+ | Ccomplu _ => Archi.ptr64
+ | Ccompluimm _ _ => Archi.ptr64
+ | CEbequw _ => negb Archi.ptr64
+ | CEbneuw _ => negb Archi.ptr64
+ | CEbltuw _ => negb Archi.ptr64
+ | CEbgeuw _ => negb Archi.ptr64
+ | CEbequl _ => Archi.ptr64
+ | CEbneul _ => Archi.ptr64
+ | CEbltul _ => Archi.ptr64
+ | CEbgeul _ => Archi.ptr64
+ | _ => false
+ end.
+
Definition op_depends_on_memory (op: operation) : bool :=
match op with
- | Ocmp (Ccompu _) => negb Archi.ptr64
- | Ocmp (Ccompuimm _ _) => negb Archi.ptr64
- | Ocmp (Ccomplu _) => Archi.ptr64
- | Ocmp (Ccompluimm _ _) => Archi.ptr64
+ | Ocmp cmp => cond_depends_on_memory cmp
+ | OEsequw _ => negb Archi.ptr64
+ | OEsneuw _ => negb Archi.ptr64
+ | OEsltiuw _ => negb Archi.ptr64
+ | OEsltuw _ => negb Archi.ptr64
+ | OEsequl _ => Archi.ptr64
+ | OEsneul _ => Archi.ptr64
+ | OEsltul _ => Archi.ptr64
+ | OEsltiul _ => Archi.ptr64
| _ => false
end.
+Lemma cond_depends_on_memory_correct:
+ forall cond args m1 m2,
+ cond_depends_on_memory cond = false ->
+ eval_condition cond args m1 = eval_condition cond args m2.
+Proof.
+ intros until m2.
+ destruct cond; cbn; try congruence.
+ all: unfold Val.cmpu_bool, Val.cmplu_bool.
+ all: destruct Archi.ptr64; cbn; intro SF; try discriminate.
+ all: reflexivity.
+Qed.
+
Lemma op_depends_on_memory_correct:
forall (F V: Type) (ge: Genv.t F V) sp op args m1 m2,
op_depends_on_memory op = false ->
eval_operation ge sp op args m1 = eval_operation ge sp op args m2.
Proof.
intros until m2. destruct op; simpl; try congruence.
- destruct cond; simpl; intros SF; auto; rewrite ? negb_false_iff in SF;
- unfold Val.cmpu_bool, Val.cmplu_bool; rewrite SF; reflexivity.
+ intro DEPEND.
+ f_equal. f_equal. apply cond_depends_on_memory_correct; trivial.
+ all: intros; repeat (destruct args; auto);
+ unfold Val.cmpu, Val.cmpu_bool, Val.cmplu, Val.cmplu_bool;
+ try destruct optR as [[]|]; simpl;
+ destruct v; try destruct v0; simpl; auto;
+ try apply negb_false_iff in H; try rewrite H; auto.
+Qed.
+
+Lemma cond_valid_pointer_eq:
+ forall cond args m1 m2,
+ (forall b z, Mem.valid_pointer m1 b z = Mem.valid_pointer m2 b z) ->
+ eval_condition cond args m1 = eval_condition cond args m2.
+Proof.
+ intros until m2. intro MEM. destruct cond eqn:COND; simpl; try congruence.
+ all: repeat (destruct args; simpl; try congruence);
+ try destruct optR as [[]|]; simpl;
+ try destruct v, v0; try rewrite !MEM; auto;
+ try erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+Qed.
+
+Lemma op_valid_pointer_eq:
+ forall (F V: Type) (ge: Genv.t F V) sp op args m1 m2,
+ (forall b z, Mem.valid_pointer m1 b z = Mem.valid_pointer m2 b z) ->
+ eval_operation ge sp op args m1 = eval_operation ge sp op args m2.
+Proof.
+ intros until m2. destruct op; simpl; try congruence.
+ intro MEM; erewrite cond_valid_pointer_eq; eauto.
+ all: intros MEM; repeat (destruct args; simpl; try congruence);
+ try destruct optR as [[]|]; simpl; try destruct v, v0; try rewrite !MEM; auto;
+ unfold Val.cmpu, Val.cmplu;
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
Qed.
(** Global variables mentioned in an operation or addressing mode *)
@@ -940,6 +1458,90 @@ Ltac InvInject :=
| _ => idtac
end.
+Lemma eval_cmpu_bool_inj': forall b c v v' v0 v0',
+ Val.inject f v v' ->
+ Val.inject f v0 v0' ->
+ Val.cmpu_bool (Mem.valid_pointer m1) c v v0 = Some b ->
+ Val.cmpu_bool (Mem.valid_pointer m2) c v' v0' = Some b.
+Proof.
+ intros.
+ eauto 3 using Val.cmpu_bool_inject, Mem.valid_pointer_implies.
+Qed.
+
+Lemma eval_cmpu_bool_inj: forall c v v' v0 v'0,
+ Val.inject f v v' ->
+ Val.inject f v0 v'0 ->
+ Val.inject f (Val.cmpu (Mem.valid_pointer m1) c v v0)
+ (Val.cmpu (Mem.valid_pointer m2) c v' v'0).
+Proof.
+ intros until v'0. intros HV1 HV2.
+ unfold Val.cmpu;
+ destruct (Val.cmpu_bool (Mem.valid_pointer m1) c _ _) eqn:?; eauto.
+ exploit eval_cmpu_bool_inj'. eapply HV1. eapply HV2. eapply Heqo.
+ intros EQ; rewrite EQ; destruct b; simpl; constructor; eauto.
+Qed.
+
+Lemma eval_cmpu_bool_inj_opt: forall c v v' v0 v'0 optR,
+ Val.inject f v v' ->
+ Val.inject f v0 v'0 ->
+ Val.inject f (apply_bin_oreg optR (Val.cmpu (Mem.valid_pointer m1) c) v v0 zero32)
+ (apply_bin_oreg optR (Val.cmpu (Mem.valid_pointer m2) c) v' v'0 zero32).
+Proof.
+ intros until optR. intros HV1 HV2.
+ destruct optR as [[]|]; simpl; unfold zero32, Val.cmpu;
+ destruct (Val.cmpu_bool (Mem.valid_pointer m1) c _ _) eqn:?; eauto;
+ assert (HVI: Val.inject f (Vint Int.zero) (Vint Int.zero)) by apply Val.inject_int.
+ + exploit eval_cmpu_bool_inj'. eapply HVI. eapply HV1. eapply Heqo.
+ intros EQ; rewrite EQ; destruct b; simpl; constructor; eauto.
+ + exploit eval_cmpu_bool_inj'. eapply HV1. eapply HVI. eapply Heqo.
+ intros EQ; rewrite EQ; destruct b; simpl; constructor; eauto.
+ + exploit eval_cmpu_bool_inj'. eapply HV1. instantiate (1:=v'0).
+ eauto. eapply Heqo.
+ intros EQ; rewrite EQ; destruct b; simpl; constructor; eauto.
+Qed.
+
+Lemma eval_cmplu_bool_inj': forall b c v v' v0 v0',
+ Val.inject f v v' ->
+ Val.inject f v0 v0' ->
+ Val.cmplu_bool (Mem.valid_pointer m1) c v v0 = Some b ->
+ Val.cmplu_bool (Mem.valid_pointer m2) c v' v0' = Some b.
+Proof.
+ intros.
+ eauto 3 using Val.cmplu_bool_inject, Mem.valid_pointer_implies.
+Qed.
+
+Lemma eval_cmplu_bool_inj: forall c v v' v0 v'0,
+ Val.inject f v v' ->
+ Val.inject f v0 v'0 ->
+ Val.inject f (Val.maketotal (Val.cmplu (Mem.valid_pointer m1) c v v0))
+ (Val.maketotal (Val.cmplu (Mem.valid_pointer m2) c v' v'0)).
+Proof.
+ intros until v'0. intros HV1 HV2.
+ unfold Val.cmplu;
+ destruct (Val.cmplu_bool (Mem.valid_pointer m1) c _ _) eqn:?; eauto.
+ exploit eval_cmplu_bool_inj'. eapply HV1. eapply HV2. eapply Heqo.
+ intros EQ; rewrite EQ; destruct b; simpl; constructor; eauto.
+Qed.
+
+Lemma eval_cmplu_bool_inj_opt: forall c v v' v0 v'0 optR,
+ Val.inject f v v' ->
+ Val.inject f v0 v'0 ->
+ Val.inject f (Val.maketotal (apply_bin_oreg optR (Val.cmplu (Mem.valid_pointer m1) c) v v0 zero64))
+ (Val.maketotal (apply_bin_oreg optR (Val.cmplu (Mem.valid_pointer m2) c) v' v'0 zero64)).
+Proof.
+ intros until optR. intros HV1 HV2.
+ destruct optR as [[]|]; simpl; unfold zero64, Val.cmplu;
+ destruct (Val.cmplu_bool (Mem.valid_pointer m1) c _ _) eqn:?; eauto;
+ assert (HVI: Val.inject f (Vlong Int64.zero) (Vlong Int64.zero)) by apply Val.inject_long.
+ + exploit eval_cmplu_bool_inj'. eapply HVI. eapply HV1. eapply Heqo.
+ intros EQ; rewrite EQ; destruct b; simpl; constructor; eauto.
+ + exploit eval_cmplu_bool_inj'. eapply HV1. eapply HVI. eapply Heqo.
+ intros EQ; rewrite EQ; destruct b; simpl; constructor; eauto.
+ + exploit eval_cmplu_bool_inj'. eapply HV1. instantiate (1:=v'0).
+ eauto. eapply Heqo.
+ intros EQ; rewrite EQ; destruct b; simpl; constructor; eauto.
+Qed.
+
Lemma eval_condition_inj:
forall cond vl1 vl2 b,
Val.inject_list f vl1 vl2 ->
@@ -947,6 +1549,9 @@ Lemma eval_condition_inj:
eval_condition cond vl2 m2 = Some b.
Proof.
intros. destruct cond; simpl in H0; FuncInv; InvInject; simpl; auto.
+ all: assert (HVI32: Val.inject f (Vint Int.zero) (Vint Int.zero)) by apply Val.inject_int;
+ assert (HVI64: Val.inject f (Vlong Int64.zero) (Vlong Int64.zero)) by apply Val.inject_long;
+ try unfold zero32, zero64.
- 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.
@@ -959,6 +1564,38 @@ Proof.
- 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.
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
+ inv H3; inv H2; simpl in H0; inv H0; auto.
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
+ inv H3; inv H2; simpl in H0; inv H0; auto.
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
+ eapply eval_cmpu_bool_inj'; eauto.
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
+ eapply eval_cmpu_bool_inj'; eauto.
+- destruct optR as [[]|]; simpl;
+ inv H3; inv H2; simpl in H0; inv H0; auto.
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
+ eapply eval_cmpu_bool_inj'; eauto.
+- destruct optR as [[]|]; simpl;
+ inv H3; inv H2; simpl in H0; inv H0; auto.
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
+ eapply eval_cmpu_bool_inj'; eauto.
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
+ inv H3; inv H2; simpl in H0; inv H0; auto.
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
+ inv H3; inv H2; simpl in H0; inv H0; auto.
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
+ eapply eval_cmplu_bool_inj'; eauto.
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
+ eapply eval_cmplu_bool_inj'; eauto.
+- destruct optR as [[]|]; simpl;
+ inv H3; inv H2; simpl in H0; inv H0; auto.
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
+ eapply eval_cmplu_bool_inj'; eauto.
+- destruct optR as [[]|]; simpl;
+ inv H3; inv H2; simpl in H0; inv H0; auto.
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
+ eapply eval_cmplu_bool_inj'; eauto.
Qed.
Ltac TrivialExists :=
@@ -997,19 +1634,29 @@ Proof.
- inv H4; inv H2; simpl; auto.
- inv H4; inv H2; simpl; auto.
(* div, divu *)
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ - inv H4; inv H2; cbn.
+ all: try apply Val.val_inject_undef.
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.
+ || Int.eq i (Int.repr (-2147483648)) && Int.eq i0 Int.mone); cbn.
+ apply Val.val_inject_undef.
+ apply Val.inject_int.
+ - inv H4; inv H2; cbn.
+ all: try apply Val.val_inject_undef.
+ destruct (Int.eq i0 Int.zero); cbn.
+ apply Val.val_inject_undef.
+ apply Val.inject_int.
(* mod, modu *)
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ - inv H4; inv H2; cbn.
+ all: try apply Val.val_inject_undef.
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.
+ || Int.eq i (Int.repr (-2147483648)) && Int.eq i0 Int.mone); cbn.
+ apply Val.val_inject_undef.
+ apply Val.inject_int.
+ - inv H4; inv H2; cbn.
+ all: try apply Val.val_inject_undef.
+ destruct (Int.eq i0 Int.zero); cbn.
+ apply Val.val_inject_undef.
+ apply Val.inject_int.
(* and, andimm *)
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
@@ -1029,8 +1676,10 @@ Proof.
- inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
- inv H4; simpl; auto. destruct (Int.ltu n Int.iwordsize); auto.
(* shrx *)
- - inv H4; simpl in H1; try discriminate. simpl.
- destruct (Int.ltu n (Int.repr 31)); inv H1. TrivialExists.
+ - inv H4; cbn; try apply Val.val_inject_undef.
+ destruct (Int.ltu n (Int.repr 31)); cbn.
+ apply Val.inject_int.
+ apply Val.val_inject_undef.
(* makelong, highlong, lowlong *)
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
@@ -1049,19 +1698,31 @@ Proof.
- inv H4; inv H2; simpl; auto.
- inv H4; inv H2; simpl; auto.
(* divl, divlu *)
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ - inv H4; inv H2; cbn.
+ all: try apply Val.val_inject_undef.
destruct (Int64.eq i0 Int64.zero
- || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2.
- TrivialExists.
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int64.eq i0 Int64.zero); inv H2. TrivialExists.
+ || Int64.eq i (Int64.repr (-9223372036854775808)) &&
+ Int64.eq i0 Int64.mone); cbn.
+ apply Val.val_inject_undef.
+ apply Val.inject_long.
+ - inv H4; inv H2; cbn.
+ all: try apply Val.val_inject_undef.
+ destruct (Int64.eq i0 Int64.zero); cbn.
+ apply Val.val_inject_undef.
+ apply Val.inject_long.
(* modl, modlu *)
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ - inv H4; inv H2; cbn.
+ all: try apply Val.val_inject_undef.
destruct (Int64.eq i0 Int64.zero
- || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2.
- TrivialExists.
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int64.eq i0 Int64.zero); inv H2. TrivialExists.
+ || Int64.eq i (Int64.repr (-9223372036854775808)) &&
+ Int64.eq i0 Int64.mone); cbn.
+ apply Val.val_inject_undef.
+ apply Val.inject_long.
+ - inv H4; inv H2; cbn.
+ all: try apply Val.val_inject_undef.
+ destruct (Int64.eq i0 Int64.zero); cbn.
+ apply Val.val_inject_undef.
+ apply Val.inject_long.
(* andl, andlimm *)
- inv H4; inv H2; simpl; auto.
- inv H4; simpl; auto.
@@ -1081,8 +1742,10 @@ Proof.
- inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int64.iwordsize'); auto.
- inv H4; simpl; auto. destruct (Int.ltu n Int64.iwordsize'); auto.
(* shrx *)
- - inv H4; simpl in H1; try discriminate. simpl.
- destruct (Int.ltu n (Int.repr 63)); inv H1. TrivialExists.
+ - inv H4; cbn; try apply Val.val_inject_undef.
+ destruct (Int.ltu n (Int.repr 63)); cbn.
+ apply Val.inject_long.
+ apply Val.val_inject_undef.
(* negf, absf *)
- inv H4; simpl; auto.
- inv H4; simpl; auto.
@@ -1105,42 +1768,145 @@ Proof.
- inv H4; simpl; auto.
- inv H4; simpl; auto.
(* intoffloat, intuoffloat *)
- - 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. destruct (Float.to_intu f0); simpl in H2; inv H2.
- exists (Vint i); auto.
+ - inv H4; cbn; auto.
+ destruct (Float.to_int f0); cbn; auto.
+ - inv H4; cbn; auto.
+ destruct (Float.to_intu f0); cbn; auto.
(* floatofint, floatofintu *)
- - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
- - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
+ - inv H4; cbn; auto.
+ - inv H4; cbn; auto.
(* intofsingle, intuofsingle *)
- - 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. destruct (Float32.to_intu f0); simpl in H2; inv H2.
- exists (Vint i); auto.
+ - inv H4; cbn; auto.
+ destruct (Float32.to_int f0); cbn; auto.
+ - inv H4; cbn; auto.
+ destruct (Float32.to_intu f0); cbn; auto.
(* singleofint, singleofintu *)
- - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
- - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
+ - inv H4; cbn; auto.
+ - inv H4; cbn; auto.
(* longoffloat, longuoffloat *)
- - inv H4; simpl in H1; inv H1. simpl. destruct (Float.to_long f0); simpl in H2; inv H2.
- exists (Vlong i); auto.
- - inv H4; simpl in H1; inv H1. simpl. destruct (Float.to_longu f0); simpl in H2; inv H2.
- exists (Vlong i); auto.
+ - inv H4; cbn; auto.
+ destruct (Float.to_long f0); cbn; auto.
+ - inv H4; cbn; auto.
+ destruct (Float.to_longu f0); cbn; auto.
(* floatoflong, floatoflongu *)
- - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
- - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
+ - inv H4; cbn; auto.
+ - inv H4; cbn; auto.
(* longofsingle, longuofsingle *)
- - inv H4; simpl in H1; inv H1. simpl. destruct (Float32.to_long f0); simpl in H2; inv H2.
- exists (Vlong i); auto.
- - inv H4; simpl in H1; inv H1. simpl. destruct (Float32.to_longu f0); simpl in H2; inv H2.
- exists (Vlong i); auto.
+ - inv H4; cbn; auto.
+ destruct (Float32.to_long f0); cbn; auto.
+ - inv H4; cbn; auto.
+ destruct (Float32.to_longu f0); cbn; auto.
(* singleoflong, singleoflongu *)
- - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
- - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
+ - inv H4; cbn; auto.
+ - inv H4; cbn; auto.
(* cmp *)
- subst v1. destruct (eval_condition cond vl1 m1) eqn:?.
exploit eval_condition_inj; eauto. intros EQ; rewrite EQ.
destruct b; simpl; constructor.
simpl; constructor.
+ (* OEseqw *)
+ - destruct optR as [[]|]; simpl; unfold zero32, Val.cmp;
+ inv H4; inv H2; simpl; try destruct (Int.eq _ _); simpl; cbn; auto;
+ try apply Val.inject_int.
+ (* OEsnew *)
+ - destruct optR as [[]|]; simpl; unfold zero32, Val.cmp;
+ inv H4; inv H2; simpl; try destruct (Int.eq _ _); simpl; cbn; auto;
+ try apply Val.inject_int.
+ (* OEsequw *)
+ - apply eval_cmpu_bool_inj_opt; auto.
+ (* OEsneuw *)
+ - apply eval_cmpu_bool_inj_opt; auto.
+ (* OEsltw *)
+ - destruct optR as [[]|]; simpl; unfold zero32, Val.cmp;
+ inv H4; inv H2; simpl; try destruct (Int.lt _ _); simpl; cbn; auto;
+ try apply Val.inject_int.
+ (* OEsltuw *)
+ - apply eval_cmpu_bool_inj_opt; auto.
+ (* OEsltiw *)
+ - inv H4; simpl; cbn; auto; try destruct (Int.lt _ _); apply Val.inject_int.
+ (* OEsltiuw *)
+ - apply eval_cmpu_bool_inj; auto.
+ (* OEaddiw *)
+ - destruct optR as [[]|]; auto; simpl.
+ rewrite Int.add_zero_l; auto.
+ rewrite Int.add_commut, Int.add_zero_l; auto.
+ - destruct optR as [[]|]; auto; simpl;
+ eapply Val.add_inject; auto.
+ (* OEandiw *)
+ - inv H4; cbn; auto.
+ (* OEoriw *)
+ - inv H4; cbn; auto.
+ (* OExoriw *)
+ - inv H4; simpl; auto.
+ (* OEluiw *)
+ - destruct (Int.ltu _ _); auto.
+ (* OEseql *)
+ - destruct optR as [[]|]; simpl; unfold zero64, Val.cmpl;
+ inv H4; inv H2; simpl; try destruct (Int64.eq _ _); simpl; cbn; auto;
+ try apply Val.inject_int.
+ (* OEsnel *)
+ - destruct optR as [[]|]; simpl; unfold zero64, Val.cmpl;
+ inv H4; inv H2; simpl; try destruct (Int64.eq _ _); simpl; cbn; auto;
+ try apply Val.inject_int.
+ (* OEsequl *)
+ - apply eval_cmplu_bool_inj_opt; auto.
+ (* OEsneul *)
+ - apply eval_cmplu_bool_inj_opt; auto.
+ (* OEsltl *)
+ - destruct optR as [[]|]; simpl; unfold zero64, Val.cmpl;
+ inv H4; inv H2; simpl; try destruct (Int64.lt _ _); simpl; cbn; auto;
+ try apply Val.inject_int.
+ (* OEsltul *)
+ - apply eval_cmplu_bool_inj_opt; auto.
+ (* OEsltil *)
+ - inv H4; simpl; cbn; auto; try destruct (Int64.lt _ _); apply Val.inject_int.
+ (* OEsltiul *)
+ - apply eval_cmplu_bool_inj; auto.
+ (* OEaddil *)
+ - destruct optR as [[]|]; auto; simpl.
+ rewrite Int64.add_zero_l; auto.
+ rewrite Int64.add_commut, Int64.add_zero_l; auto.
+ - destruct optR as [[]|]; auto; simpl;
+ eapply Val.addl_inject; auto.
+ (* OEandil *)
+ - inv H4; cbn; auto.
+ (* OEoril *)
+ - inv H4; cbn; auto.
+ (* OExoril *)
+ - inv H4; simpl; auto.
+ (* OEmayundef *)
+ - destruct mu; inv H4; inv H2; simpl; auto;
+ try destruct (Int.ltu _ _); simpl; auto.
+ all: eapply Val.inject_ptr; eauto.
+ (* OEfeqd *)
+ - inv H4; inv H2; cbn; simpl; auto.
+ destruct Float.cmp; unfold Vtrue, Vfalse; cbn; auto.
+ (* OEfltd *)
+ - inv H4; inv H2; cbn; simpl; auto.
+ destruct Float.cmp; unfold Vtrue, Vfalse; cbn; auto.
+ (* OEfled *)
+ - inv H4; inv H2; cbn; simpl; auto.
+ destruct Float.cmp; unfold Vtrue, Vfalse; cbn; auto.
+ (* OEfeqs *)
+ - inv H4; inv H2; cbn; simpl; auto.
+ destruct Float32.cmp; unfold Vtrue, Vfalse; cbn; auto.
+ (* OEflts *)
+ - inv H4; inv H2; cbn; simpl; auto.
+ destruct Float32.cmp; unfold Vtrue, Vfalse; cbn; auto.
+ (* OEfles *)
+ - inv H4; inv H2; cbn; simpl; auto.
+ destruct Float32.cmp; unfold Vtrue, Vfalse; cbn; auto.
+ (* Bits_of_single, double *)
+ - inv H4; simpl; auto.
+ - inv H4; simpl; auto.
+ (* single, double of bits *)
+ - inv H4; simpl; auto.
+ - inv H4; simpl; auto.
+ (* selectl *)
+ - inv H4; trivial. cbn.
+ destruct (Int.eq i Int.one).
+ + auto using Val.normalize_inject.
+ + destruct (Int.eq i Int.zero); cbn; auto using Val.normalize_inject.
Qed.
Lemma eval_addressing_inj:
@@ -1159,6 +1925,20 @@ Proof.
apply Val.offset_ptr_inject; auto.
Qed.
+Lemma eval_addressing_inj_none:
+ forall addr sp1 vl1 sp2 vl2,
+ (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 = None ->
+ eval_addressing ge2 sp2 addr vl2 = None.
+Proof.
+ intros until vl2. intros Hglobal Hinjsp Hinjvl.
+ destruct addr; simpl in *;
+ inv Hinjvl; trivial; try discriminate; inv H0; trivial; try discriminate; inv H2; trivial; try discriminate.
+Qed.
End EVAL_COMPAT.
(** Compatibility of the evaluation functions with the ``is less defined'' relation over values. *)
@@ -1265,6 +2045,18 @@ Proof.
destruct H1 as [v2 [A B]]. exists v2; split; auto. rewrite val_inject_lessdef; auto.
Qed.
+Lemma eval_addressing_lessdef_none:
+ forall sp addr vl1 vl2,
+ Val.lessdef_list vl1 vl2 ->
+ eval_addressing genv sp addr vl1 = None ->
+ eval_addressing genv sp addr vl2 = None.
+Proof.
+ intros until vl2. intros Hlessdef Heval1.
+ destruct addr; simpl in *;
+ inv Hlessdef; trivial; try discriminate;
+ inv H0; trivial; try discriminate;
+ inv H2; trivial; try discriminate.
+Qed.
End EVAL_LESSDEF.
(** Compatibility of the evaluation functions with memory injections. *)
@@ -1317,6 +2109,20 @@ Proof.
econstructor; eauto. rewrite Ptrofs.add_zero_l; auto.
Qed.
+
+Lemma eval_addressing_inject_none:
+ forall addr vl1 vl2,
+ Val.inject_list f vl1 vl2 ->
+ eval_addressing genv (Vptr sp1 Ptrofs.zero) addr vl1 = None ->
+ eval_addressing genv (Vptr sp2 Ptrofs.zero) (shift_stack_addressing delta addr) vl2 = None.
+Proof.
+ intros.
+ rewrite eval_shift_stack_addressing.
+ eapply eval_addressing_inj_none with (sp1 := Vptr sp1 Ptrofs.zero); eauto.
+ intros. apply symbol_address_inject.
+ econstructor; eauto. rewrite Ptrofs.add_zero_l; auto.
+Qed.
+
Lemma eval_operation_inject:
forall op vl1 vl2 v1 m1 m2,
Val.inject_list f vl1 vl2 ->
@@ -1358,4 +2164,4 @@ Definition builtin_arg_ok
match ba with
| (BA _ | BA_splitlong (BA _) (BA _)) => true
| _ => builtin_arg_ok_1 ba c
- end.
+ end.
diff --git a/riscV/OpWeights.ml b/riscV/OpWeights.ml
new file mode 100644
index 00000000..0a1d9ad4
--- /dev/null
+++ b/riscV/OpWeights.ml
@@ -0,0 +1,168 @@
+open Op
+open PrepassSchedulingOracleDeps
+
+module Rocket = struct
+ (* Attempt at modeling the Rocket core *)
+
+ let resource_bounds = [| 1 |]
+
+ let nr_non_pipelined_units = 1
+
+ (* divider *)
+
+ let latency_of_op (op : operation) (nargs : int) =
+ match op with
+ | Omul | Omulhs | Omulhu | Omull | Omullhs | Omullhu -> 4
+ | Onegf -> 1 (*r [rd = - r1] *)
+ | Oabsf (*r [rd = abs(r1)] *)
+ | Oaddf (*r [rd = r1 + r2] *)
+ | Osubf (*r [rd = r1 - r2] *)
+ | Omulf ->
+ 6 (*r [rd = r1 * r2] *)
+ | Onegfs -> 1 (*r [rd = - r1] *)
+ | Oabsfs (*r [rd = abs(r1)] *)
+ | Oaddfs (*r [rd = r1 + r2] *)
+ | Osubfs (*r [rd = r1 - r2] *)
+ | Omulfs ->
+ 4 (*r [rd = r1 * r2] *)
+ | Osingleoffloat (*r [rd] is [r1] truncated to single-precision float *)
+ | Ofloatofsingle (*r [rd] is [r1] extended to double-precision float *)
+ (*c Conversions between int and float: *)
+ | Ofloatconst _ | Osingleconst _
+ | Ointoffloat (*r [rd = signed_int_of_float64(r1)] *)
+ | Ointuoffloat (*r [rd = unsigned_int_of_float64(r1)] *)
+ | Ofloatofint (*r [rd = float64_of_signed_int(r1)] *)
+ | Ofloatofintu (*r [rd = float64_of_unsigned_int(r1)] *)
+ | Ointofsingle (*r [rd = signed_int_of_float32(r1)] *)
+ | Ointuofsingle (*r [rd = unsigned_int_of_float32(r1)] *)
+ | Osingleofint (*r [rd = float32_of_signed_int(r1)] *)
+ | Osingleofintu (*r [rd = float32_of_unsigned_int(r1)] *)
+ | Olongoffloat (*r [rd = signed_long_of_float64(r1)] *)
+ | Olonguoffloat (*r [rd = unsigned_long_of_float64(r1)] *)
+ | Ofloatoflong (*r [rd = float64_of_signed_long(r1)] *)
+ | Ofloatoflongu (*r [rd = float64_of_unsigned_long(r1)] *)
+ | Olongofsingle (*r [rd = signed_long_of_float32(r1)] *)
+ | Olonguofsingle (*r [rd = unsigned_long_of_float32(r1)] *)
+ | Osingleoflong (*r [rd = float32_of_signed_long(r1)] *)
+ | Osingleoflongu ->
+ 2 (*r [rd = float32_of_unsigned_int(r1)] *)
+ | OEfeqd | OEfltd | OEfeqs | OEflts | OEfles | OEfled | Obits_of_single
+ | Obits_of_float | Osingle_of_bits | Ofloat_of_bits ->
+ 2
+ | OEloadli _ -> 2
+ | Odiv | Odivu | Odivl | Odivlu -> 16
+ | Odivfs -> 35
+ | Odivf -> 50
+ | Ocmp cond -> (
+ match cond with
+ | Ccomp _ | Ccompu _ | Ccompimm _ | Ccompuimm _ | Ccompl _ | Ccomplu _
+ | Ccomplimm _ | Ccompluimm _ | CEbeqw _ | CEbnew _ | CEbequw _
+ | CEbneuw _ | CEbltw _ | CEbltuw _ | CEbgew _ | CEbgeuw _ | CEbeql _
+ | CEbnel _ | CEbequl _ | CEbneul _ | CEbltl _ | CEbltul _ | CEbgel _
+ | CEbgeul _ ->
+ 1
+ | Ccompf _ | Cnotcompf _ -> 2
+ | Ccompfs _ | Cnotcompfs _ -> 2)
+ | OEmayundef _ -> 0
+ | _ -> 1
+
+ let resources_of_op (op : operation) (nargs : int) = resource_bounds
+
+ let non_pipelined_resources_of_op (op : operation) (nargs : int) =
+ match op with
+ | Odiv | Odivu -> [| 29 |]
+ | Odivfs -> [| 20 |]
+ | Odivl | Odivlu | Odivf -> [| 50 |]
+ | _ -> [| -1 |]
+
+ let resources_of_cond (cond : condition) (nargs : int) = resource_bounds
+
+ let latency_of_load trap chunk (addr : addressing) (nargs : int) = 3
+
+ let latency_of_call _ _ = 6
+
+ let resources_of_load trap chunk addressing nargs = resource_bounds
+
+ let resources_of_store chunk addressing nargs = resource_bounds
+
+ let resources_of_call _ _ = resource_bounds
+
+ let resources_of_builtin _ = resource_bounds
+end
+
+module SweRV_EH1 = struct
+ (* Attempt at modeling SweRV EH1
+ [| issues ; LSU ; multiplier |] *)
+ let resource_bounds = [| 2; 1; 1 |]
+
+ let nr_non_pipelined_units = 1
+
+ (* divider *)
+
+ let latency_of_op (op : operation) (nargs : int) =
+ match op with
+ | Omul | Omulhs | Omulhu | Omull | Omullhs | Omullhu -> 3
+ | Odiv | Odivu | Odivl | Odivlu -> 16
+ | _ -> 1
+
+ let resources_of_op (op : operation) (nargs : int) =
+ match op with
+ | Omul | Omulhs | Omulhu | Omull | Omullhs | Omullhu -> [| 1; 0; 1 |]
+ | Odiv | Odivu | Odivl | Odivlu -> [| 0; 0; 0 |]
+ | _ -> [| 1; 0; 0 |]
+
+ let non_pipelined_resources_of_op (op : operation) (nargs : int) =
+ match op with
+ | Odiv | Odivu -> [| 29 |]
+ | Odivfs -> [| 20 |]
+ | Odivl | Odivlu | Odivf -> [| 50 |]
+ | _ -> [| -1 |]
+
+ let resources_of_cond (cond : condition) (nargs : int) = [| 1; 0; 0 |]
+
+ let latency_of_load trap chunk (addr : addressing) (nargs : int) = 3
+
+ let latency_of_call _ _ = 6
+
+ let resources_of_load trap chunk addressing nargs = [| 1; 1; 0 |]
+
+ let resources_of_store chunk addressing nargs = [| 1; 1; 0 |]
+
+ let resources_of_call _ _ = resource_bounds
+
+ let resources_of_builtin _ = resource_bounds
+end
+
+let get_opweights () : opweights =
+ match !Clflags.option_mtune with
+ | "rocket" | "" ->
+ {
+ pipelined_resource_bounds = Rocket.resource_bounds;
+ nr_non_pipelined_units = Rocket.nr_non_pipelined_units;
+ latency_of_op = Rocket.latency_of_op;
+ resources_of_op = Rocket.resources_of_op;
+ non_pipelined_resources_of_op = Rocket.non_pipelined_resources_of_op;
+ latency_of_load = Rocket.latency_of_load;
+ resources_of_load = Rocket.resources_of_load;
+ resources_of_store = Rocket.resources_of_store;
+ resources_of_cond = Rocket.resources_of_cond;
+ latency_of_call = Rocket.latency_of_call;
+ resources_of_call = Rocket.resources_of_call;
+ resources_of_builtin = Rocket.resources_of_builtin;
+ }
+ | "SweRV_EH1" | "EH1" ->
+ {
+ pipelined_resource_bounds = SweRV_EH1.resource_bounds;
+ nr_non_pipelined_units = SweRV_EH1.nr_non_pipelined_units;
+ latency_of_op = SweRV_EH1.latency_of_op;
+ resources_of_op = SweRV_EH1.resources_of_op;
+ non_pipelined_resources_of_op = SweRV_EH1.non_pipelined_resources_of_op;
+ latency_of_load = SweRV_EH1.latency_of_load;
+ resources_of_load = SweRV_EH1.resources_of_load;
+ resources_of_store = SweRV_EH1.resources_of_store;
+ resources_of_cond = SweRV_EH1.resources_of_cond;
+ latency_of_call = SweRV_EH1.latency_of_call;
+ resources_of_call = SweRV_EH1.resources_of_call;
+ resources_of_builtin = SweRV_EH1.resources_of_builtin;
+ }
+ | xxx -> failwith (Printf.sprintf "unknown -mtune: %s" xxx)
diff --git a/riscV/PrepassSchedulingOracle.ml b/riscV/PrepassSchedulingOracle.ml
new file mode 120000
index 00000000..912e9ffa
--- /dev/null
+++ b/riscV/PrepassSchedulingOracle.ml
@@ -0,0 +1 @@
+../aarch64/PrepassSchedulingOracle.ml \ No newline at end of file
diff --git a/riscV/PrepassSchedulingOracleDeps.ml b/riscV/PrepassSchedulingOracleDeps.ml
new file mode 120000
index 00000000..1e955b85
--- /dev/null
+++ b/riscV/PrepassSchedulingOracleDeps.ml
@@ -0,0 +1 @@
+../aarch64/PrepassSchedulingOracleDeps.ml \ No newline at end of file
diff --git a/riscV/PrintOp.ml b/riscV/PrintOp.ml
index 9ec474b3..0d47192a 100644
--- a/riscV/PrintOp.ml
+++ b/riscV/PrintOp.ml
@@ -30,6 +30,21 @@ let comparison_name = function
| Cgt -> ">"
| Cge -> ">="
+let mu_name pp = function
+ | MUint -> fprintf pp "MUint"
+ | MUlong -> fprintf pp "MUlong"
+ | MUshrx i -> fprintf pp "MUshrx(%ld)" (camlint_of_coqint i)
+ | MUshrxl i -> fprintf pp "MUshrxl(%ld)" (camlint_of_coqint i)
+
+let get_optR_s c reg pp r1 r2 = function
+ | None -> fprintf pp "(%a %s %a)" reg r1 (comparison_name c) reg r2
+ | Some X0_L -> fprintf pp "(X0 %s %a)" (comparison_name c) reg r1
+ | Some X0_R -> fprintf pp "(%a %s X0)" reg r1 (comparison_name c)
+
+let get_optR_a pp = function
+ | None -> failwith "PrintOp: None in get_optR_a instruction (problem with RTL expansions?)"
+ | Some X0_L | Some X0_R -> fprintf pp "X0"
+
let print_condition reg pp = function
| (Ccomp c, [r1;r2]) ->
fprintf pp "%a %ss %a" reg r1 (comparison_name c) reg r2
@@ -55,15 +70,47 @@ let print_condition reg pp = function
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
+ | (CEbeqw optR, [r1;r2]) ->
+ fprintf pp "CEbeqw"; (get_optR_s Ceq reg pp r1 r2 optR)
+ | (CEbnew optR, [r1;r2]) ->
+ fprintf pp "CEbnew"; (get_optR_s Cne reg pp r1 r2 optR)
+ | (CEbequw optR, [r1;r2]) ->
+ fprintf pp "CEbequw"; (get_optR_s Ceq reg pp r1 r2 optR)
+ | (CEbneuw optR, [r1;r2]) ->
+ fprintf pp "CEbneuw"; (get_optR_s Cne reg pp r1 r2 optR)
+ | (CEbltw optR, [r1;r2]) ->
+ fprintf pp "CEbltw"; (get_optR_s Clt reg pp r1 r2 optR)
+ | (CEbltuw optR, [r1;r2]) ->
+ fprintf pp "CEbltuw"; (get_optR_s Clt reg pp r1 r2 optR)
+ | (CEbgew optR, [r1;r2]) ->
+ fprintf pp "CEbgew"; (get_optR_s Cge reg pp r1 r2 optR)
+ | (CEbgeuw optR, [r1;r2]) ->
+ fprintf pp "CEbgeuw"; (get_optR_s Cge reg pp r1 r2 optR)
+ | (CEbeql optR, [r1;r2]) ->
+ fprintf pp "CEbeql"; (get_optR_s Ceq reg pp r1 r2 optR)
+ | (CEbnel optR, [r1;r2]) ->
+ fprintf pp "CEbnel"; (get_optR_s Cne reg pp r1 r2 optR)
+ | (CEbequl optR, [r1;r2]) ->
+ fprintf pp "CEbequl"; (get_optR_s Ceq reg pp r1 r2 optR)
+ | (CEbneul optR, [r1;r2]) ->
+ fprintf pp "CEbneul"; (get_optR_s Cne reg pp r1 r2 optR)
+ | (CEbltl optR, [r1;r2]) ->
+ fprintf pp "CEbltl"; (get_optR_s Clt reg pp r1 r2 optR)
+ | (CEbltul optR, [r1;r2]) ->
+ fprintf pp "CEbltul"; (get_optR_s Clt reg pp r1 r2 optR)
+ | (CEbgel optR, [r1;r2]) ->
+ fprintf pp "CEbgel"; (get_optR_s Cge reg pp r1 r2 optR)
+ | (CEbgeul optR, [r1;r2]) ->
+ fprintf pp "CEbgeul"; (get_optR_s Cge reg pp r1 r2 optR)
| _ ->
fprintf pp "<bad condition>"
let print_operation reg pp = function
| Omove, [r1] -> reg pp r1
- | Ointconst n, [] -> fprintf pp "%ld" (camlint_of_coqint n)
- | Olongconst n, [] -> fprintf pp "%LdL" (camlint64_of_coqint n)
- | Ofloatconst n, [] -> fprintf pp "%F" (camlfloat_of_coqfloat n)
- | Osingleconst n, [] -> fprintf pp "%Ff" (camlfloat_of_coqfloat32 n)
+ | Ointconst n, [] -> fprintf pp "Ointconst(%ld)" (camlint_of_coqint n)
+ | Olongconst n, [] -> fprintf pp "Olongconst(%LdL)" (camlint64_of_coqint n)
+ | Ofloatconst n, [] -> fprintf pp "Ofloatconst(%F)" (camlfloat_of_coqfloat n)
+ | Osingleconst n, [] -> fprintf pp "Osingleconst(%Ff)" (camlfloat_of_coqfloat32 n)
| Oaddrsymbol(id, ofs), [] ->
fprintf pp "\"%s\" + %Ld" (extern_atom id) (camlint64_of_ptrofs ofs)
| Oaddrstack ofs, [] ->
@@ -156,6 +203,47 @@ let print_operation reg pp = function
| Osingleoflong, [r1] -> fprintf pp "singleoflong(%a)" reg r1
| Osingleoflongu, [r1] -> fprintf pp "singleoflongu(%a)" reg r1
| Ocmp c, args -> print_condition reg pp (c, args)
+ | OEseqw optR, [r1;r2] -> fprintf pp "OEseqw"; (get_optR_s Ceq reg pp r1 r2 optR)
+ | OEsnew optR, [r1;r2] -> fprintf pp "OEsnew"; (get_optR_s Cne reg pp r1 r2 optR)
+ | OEsequw optR, [r1;r2] -> fprintf pp "OEsequw"; (get_optR_s Ceq reg pp r1 r2 optR)
+ | OEsneuw optR, [r1;r2] -> fprintf pp "OEsneuw"; (get_optR_s Cne reg pp r1 r2 optR)
+ | OEsltw optR, [r1;r2] -> fprintf pp "OEsltw"; (get_optR_s Clt reg pp r1 r2 optR)
+ | OEsltuw optR, [r1;r2] -> fprintf pp "OEsltuw"; (get_optR_s Clt reg pp r1 r2 optR)
+ | OEsltiw n, [r1] -> fprintf pp "OEsltiw(%a,%ld)" reg r1 (camlint_of_coqint n)
+ | OEsltiuw n, [r1] -> fprintf pp "OEsltiuw(%a,%ld)" reg r1 (camlint_of_coqint n)
+ | OExoriw n, [r1] -> fprintf pp "OExoriw(%a,%ld)" reg r1 (camlint_of_coqint n)
+ | OEluiw n, _ -> fprintf pp "OEluiw(%ld)" (camlint_of_coqint n)
+ | OEaddiw (optR, n), [] -> fprintf pp "OEaddiw(%a,%ld)" get_optR_a optR (camlint_of_coqint n)
+ | OEaddiw (optR, n), [r1] -> fprintf pp "OEaddiw(%a,%ld)" reg r1 (camlint_of_coqint n)
+ | OEandiw n, [r1] -> fprintf pp "OEandiw(%a,%ld)" reg r1 (camlint_of_coqint n)
+ | OEoriw n, [r1] -> fprintf pp "OEoriw(%a,%ld)" reg r1 (camlint_of_coqint n)
+ | OEseql optR, [r1;r2] -> fprintf pp "OEseql"; (get_optR_s Ceq reg pp r1 r2 optR)
+ | OEsnel optR, [r1;r2] -> fprintf pp "OEsnel"; (get_optR_s Cne reg pp r1 r2 optR)
+ | OEsequl optR, [r1;r2] -> fprintf pp "OEsequl"; (get_optR_s Ceq reg pp r1 r2 optR)
+ | OEsneul optR, [r1;r2] -> fprintf pp "OEsneul"; (get_optR_s Cne reg pp r1 r2 optR)
+ | OEsltl optR, [r1;r2] -> fprintf pp "OEsltl"; (get_optR_s Clt reg pp r1 r2 optR)
+ | OEsltul optR, [r1;r2] -> fprintf pp "OEsltul"; (get_optR_s Clt reg pp r1 r2 optR)
+ | OEsltil n, [r1] -> fprintf pp "OEsltil(%a,%ld)" reg r1 (camlint_of_coqint n)
+ | OEsltiul n, [r1] -> fprintf pp "OEsltiul(%a,%ld)" reg r1 (camlint_of_coqint n)
+ | OExoril n, [r1] -> fprintf pp "OExoril(%a,%ld)" reg r1 (camlint_of_coqint n)
+ | OEluil n, _ -> fprintf pp "OEluil(%ld)" (camlint_of_coqint n)
+ | OEaddil (optR, n), [] -> fprintf pp "OEaddil(%a,%ld)" get_optR_a optR (camlint_of_coqint n)
+ | OEaddil (optR, n), [r1] -> fprintf pp "OEaddil(%a,%ld)" reg r1 (camlint_of_coqint n)
+ | OEandil n, [r1] -> fprintf pp "OEandil(%a,%ld)" reg r1 (camlint_of_coqint n)
+ | OEoril n, [r1] -> fprintf pp "OEoril(%a,%ld)" reg r1 (camlint_of_coqint n)
+ | OEloadli n, _ -> fprintf pp "OEloadli(%ld)" (camlint_of_coqint n)
+ | OEmayundef mu, [r1;r2] -> fprintf pp "OEmayundef (%a,%a,%a)" mu_name mu reg r1 reg r2
+ | OEfeqd, [r1;r2] -> fprintf pp "OEfeqd(%a,%s,%a)" reg r1 (comparison_name Ceq) reg r2
+ | OEfltd, [r1;r2] -> fprintf pp "OEfltd(%a,%s,%a)" reg r1 (comparison_name Clt) reg r2
+ | OEfled, [r1;r2] -> fprintf pp "OEfled(%a,%s,%a)" reg r1 (comparison_name Cle) reg r2
+ | OEfeqs, [r1;r2] -> fprintf pp "OEfeqs(%a,%s,%a)" reg r1 (comparison_name Ceq) reg r2
+ | OEflts, [r1;r2] -> fprintf pp "OEflts(%a,%s,%a)" reg r1 (comparison_name Clt) reg r2
+ | OEfles, [r1;r2] -> fprintf pp "OEfles(%a,%s,%a)" reg r1 (comparison_name Cle) reg r2
+ | Obits_of_single, [r1] -> fprintf pp "bits_of_single(%a)" reg r1
+ | Obits_of_float, [r1] -> fprintf pp "bits_of_float(%a)" reg r1
+ | Osingle_of_bits, [r1] -> fprintf pp "single_of_bits(%a)" reg r1
+ | Ofloat_of_bits, [r1] -> fprintf pp "float_of_bits(%a)" reg r1
+ | Oselectl, [rb;rt;rf] -> fprintf pp "selectl(b:%a, t:%a, f:%a)" reg rb reg rt reg rf
| _ -> fprintf pp "<bad operator>"
let print_addressing reg pp = function
diff --git a/riscV/RTLpathSE_simplify.v b/riscV/RTLpathSE_simplify.v
new file mode 100644
index 00000000..2739bc5d
--- /dev/null
+++ b/riscV/RTLpathSE_simplify.v
@@ -0,0 +1,2092 @@
+Require Import Coqlib Floats Values Memory.
+Require Import Integers.
+Require Import Op Registers.
+Require Import RTLpathSE_theory.
+Require Import RTLpathSE_simu_specs.
+Require Import Asmgen Asmgenproof1.
+Require Import Lia.
+
+(** Useful functions for conditions/branches expansion *)
+
+Definition is_inv_cmp_int (cmp: comparison) : bool :=
+ match cmp with | Cle | Cgt => true | _ => false end.
+
+Definition is_inv_cmp_float (cmp: comparison) : bool :=
+ match cmp with | Cge | Cgt => true | _ => false end.
+
+Definition make_optR (is_x0 is_inv: bool) : option oreg :=
+ if is_x0 then
+ (if is_inv then Some (X0_L)
+ else Some (X0_R))
+ else None.
+
+(** Functions to manage lists of "fake" values *)
+
+Definition make_lhsv_cmp (is_inv: bool) (hv1 hv2: hsval) : list_hsval :=
+ let (hvfirst, hvsec) := if is_inv then (hv1, hv2) else (hv2, hv1) in
+ let lhsv := fScons hvfirst fSnil in
+ fScons hvsec lhsv.
+
+Definition make_lhsv_single (hvs: hsval) : list_hsval :=
+ fScons hvs fSnil.
+
+(** * Expansion functions *)
+
+(** ** Immediate loads *)
+
+Definition load_hilo32 (hi lo: int) :=
+ if Int.eq lo Int.zero then
+ fSop (OEluiw hi) fSnil
+ else
+ let hvs := fSop (OEluiw hi) fSnil in
+ let hl := make_lhsv_single hvs in
+ fSop (OEaddiw None lo) hl.
+
+Definition load_hilo64 (hi lo: int64) :=
+ if Int64.eq lo Int64.zero then
+ fSop (OEluil hi) fSnil
+ else
+ let hvs := fSop (OEluil hi) fSnil in
+ let hl := make_lhsv_single hvs in
+ fSop (OEaddil None lo) hl.
+
+Definition loadimm32 (n: int) :=
+ match make_immed32 n with
+ | Imm32_single imm =>
+ fSop (OEaddiw (Some X0_R) imm) fSnil
+ | Imm32_pair hi lo => load_hilo32 hi lo
+ end.
+
+Definition loadimm64 (n: int64) :=
+ match make_immed64 n with
+ | Imm64_single imm =>
+ fSop (OEaddil (Some X0_R) imm) fSnil
+ | Imm64_pair hi lo => load_hilo64 hi lo
+ | Imm64_large imm => fSop (OEloadli imm) fSnil
+ end.
+
+Definition opimm32 (hv1: hsval) (n: int) (op: operation) (opimm: int -> operation) :=
+ match make_immed32 n with
+ | Imm32_single imm =>
+ let hl := make_lhsv_single hv1 in
+ fSop (opimm imm) hl
+ | Imm32_pair hi lo =>
+ let hvs := load_hilo32 hi lo in
+ let hl := make_lhsv_cmp false hv1 hvs in
+ fSop op hl
+ end.
+
+Definition opimm64 (hv1: hsval) (n: int64) (op: operation) (opimm: int64 -> operation) :=
+ match make_immed64 n with
+ | Imm64_single imm =>
+ let hl := make_lhsv_single hv1 in
+ fSop (opimm imm) hl
+ | Imm64_pair hi lo =>
+ let hvs := load_hilo64 hi lo in
+ let hl := make_lhsv_cmp false hv1 hvs in
+ fSop op hl
+ | Imm64_large imm =>
+ let hvs := fSop (OEloadli imm) fSnil in
+ let hl := make_lhsv_cmp false hv1 hvs in
+ fSop op hl
+ end.
+
+Definition addimm32 (hv1: hsval) (n: int) (or: option oreg) := opimm32 hv1 n Oadd (OEaddiw or).
+Definition andimm32 (hv1: hsval) (n: int) := opimm32 hv1 n Oand OEandiw.
+Definition orimm32 (hv1: hsval) (n: int) := opimm32 hv1 n Oor OEoriw.
+Definition xorimm32 (hv1: hsval) (n: int) := opimm32 hv1 n Oxor OExoriw.
+Definition sltimm32 (hv1: hsval) (n: int) := opimm32 hv1 n (OEsltw None) OEsltiw.
+Definition sltuimm32 (hv1: hsval) (n: int) := opimm32 hv1 n (OEsltuw None) OEsltiuw.
+Definition addimm64 (hv1: hsval) (n: int64) (or: option oreg) := opimm64 hv1 n Oaddl (OEaddil or).
+Definition andimm64 (hv1: hsval) (n: int64) := opimm64 hv1 n Oandl OEandil.
+Definition orimm64 (hv1: hsval) (n: int64) := opimm64 hv1 n Oorl OEoril.
+Definition xorimm64 (hv1: hsval) (n: int64) := opimm64 hv1 n Oxorl OExoril.
+Definition sltimm64 (hv1: hsval) (n: int64) := opimm64 hv1 n (OEsltl None) OEsltil.
+Definition sltuimm64 (hv1: hsval) (n: int64) := opimm64 hv1 n (OEsltul None) OEsltiul.
+
+(** ** Comparisons intructions *)
+
+Definition cond_int32s (cmp: comparison) (lhsv: list_hsval) (optR: option oreg) :=
+ match cmp with
+ | Ceq => fSop (OEseqw optR) lhsv
+ | Cne => fSop (OEsnew optR) lhsv
+ | Clt | Cgt => fSop (OEsltw optR) lhsv
+ | Cle | Cge =>
+ let hvs := (fSop (OEsltw optR) lhsv) in
+ let hl := make_lhsv_single hvs in
+ fSop (OExoriw Int.one) hl
+ end.
+
+Definition cond_int32u (cmp: comparison) (lhsv: list_hsval) (optR: option oreg) :=
+ match cmp with
+ | Ceq => fSop (OEsequw optR) lhsv
+ | Cne => fSop (OEsneuw optR) lhsv
+ | Clt | Cgt => fSop (OEsltuw optR) lhsv
+ | Cle | Cge =>
+ let hvs := (fSop (OEsltuw optR) lhsv) in
+ let hl := make_lhsv_single hvs in
+ fSop (OExoriw Int.one) hl
+ end.
+
+Definition cond_int64s (cmp: comparison) (lhsv: list_hsval) (optR: option oreg) :=
+ match cmp with
+ | Ceq => fSop (OEseql optR) lhsv
+ | Cne => fSop (OEsnel optR) lhsv
+ | Clt | Cgt => fSop (OEsltl optR) lhsv
+ | Cle | Cge =>
+ let hvs := (fSop (OEsltl optR) lhsv) in
+ let hl := make_lhsv_single hvs in
+ fSop (OExoriw Int.one) hl
+ end.
+
+Definition cond_int64u (cmp: comparison) (lhsv: list_hsval) (optR: option oreg) :=
+ match cmp with
+ | Ceq => fSop (OEsequl optR) lhsv
+ | Cne => fSop (OEsneul optR) lhsv
+ | Clt | Cgt => fSop (OEsltul optR) lhsv
+ | Cle | Cge =>
+ let hvs := (fSop (OEsltul optR) lhsv) in
+ let hl := make_lhsv_single hvs in
+ fSop (OExoriw Int.one) hl
+ end.
+
+Definition expanse_condimm_int32s (cmp: comparison) (hv1: hsval) (n: int) :=
+ let is_inv := is_inv_cmp_int cmp in
+ if Int.eq n Int.zero then
+ let optR := make_optR true is_inv in
+ let hl := make_lhsv_cmp is_inv hv1 hv1 in
+ cond_int32s cmp hl optR
+ else
+ match cmp with
+ | Ceq | Cne =>
+ let optR := make_optR true is_inv in
+ let hvs := xorimm32 hv1 n in
+ let hl := make_lhsv_cmp false hvs hvs in
+ cond_int32s cmp hl optR
+ | Clt => sltimm32 hv1 n
+ | Cle =>
+ if Int.eq n (Int.repr Int.max_signed) then
+ let hvs := loadimm32 Int.one in
+ let hl := make_lhsv_cmp false hv1 hvs in
+ fSop (OEmayundef MUint) hl
+ else sltimm32 hv1 (Int.add n Int.one)
+ | _ =>
+ let optR := make_optR false is_inv in
+ let hvs := loadimm32 n in
+ let hl := make_lhsv_cmp is_inv hv1 hvs in
+ cond_int32s cmp hl optR
+ end.
+
+Definition expanse_condimm_int32u (cmp: comparison) (hv1: hsval) (n: int) :=
+ let is_inv := is_inv_cmp_int cmp in
+ if Int.eq n Int.zero then
+ let optR := make_optR true is_inv in
+ let hl := make_lhsv_cmp is_inv hv1 hv1 in
+ cond_int32u cmp hl optR
+ else
+ match cmp with
+ | Clt => sltuimm32 hv1 n
+ | _ =>
+ let optR := make_optR false is_inv in
+ let hvs := loadimm32 n in
+ let hl := make_lhsv_cmp is_inv hv1 hvs in
+ cond_int32u cmp hl optR
+ end.
+
+Definition expanse_condimm_int64s (cmp: comparison) (hv1: hsval) (n: int64) :=
+ let is_inv := is_inv_cmp_int cmp in
+ if Int64.eq n Int64.zero then
+ let optR := make_optR true is_inv in
+ let hl := make_lhsv_cmp is_inv hv1 hv1 in
+ cond_int64s cmp hl optR
+ else
+ match cmp with
+ | Ceq | Cne =>
+ let optR := make_optR true is_inv in
+ let hvs := xorimm64 hv1 n in
+ let hl := make_lhsv_cmp false hvs hvs in
+ cond_int64s cmp hl optR
+ | Clt => sltimm64 hv1 n
+ | Cle =>
+ if Int64.eq n (Int64.repr Int64.max_signed) then
+ let hvs := loadimm32 Int.one in
+ let hl := make_lhsv_cmp false hv1 hvs in
+ fSop (OEmayundef MUlong) hl
+ else sltimm64 hv1 (Int64.add n Int64.one)
+ | _ =>
+ let optR := make_optR false is_inv in
+ let hvs := loadimm64 n in
+ let hl := make_lhsv_cmp is_inv hv1 hvs in
+ cond_int64s cmp hl optR
+ end.
+
+Definition expanse_condimm_int64u (cmp: comparison) (hv1: hsval) (n: int64) :=
+ let is_inv := is_inv_cmp_int cmp in
+ if Int64.eq n Int64.zero then
+ let optR := make_optR true is_inv in
+ let hl := make_lhsv_cmp is_inv hv1 hv1 in
+ cond_int64u cmp hl optR
+ else
+ match cmp with
+ | Clt => sltuimm64 hv1 n
+ | _ =>
+ let optR := make_optR false is_inv in
+ let hvs := loadimm64 n in
+ let hl := make_lhsv_cmp is_inv hv1 hvs in
+ cond_int64u cmp hl optR
+ end.
+
+Definition cond_float (cmp: comparison) (lhsv: list_hsval) :=
+ match cmp with
+ | Ceq | Cne => fSop OEfeqd lhsv
+ | Clt | Cgt => fSop OEfltd lhsv
+ | Cle | Cge => fSop OEfled lhsv
+ end.
+
+Definition cond_single (cmp: comparison) (lhsv: list_hsval) :=
+ match cmp with
+ | Ceq | Cne => fSop OEfeqs lhsv
+ | Clt | Cgt => fSop OEflts lhsv
+ | Cle | Cge => fSop OEfles lhsv
+ end.
+
+Definition is_normal_cmp cmp :=
+ match cmp with | Cne => false | _ => true end.
+
+Definition expanse_cond_fp (cnot: bool) fn_cond cmp (lhsv: list_hsval) :=
+ let normal := is_normal_cmp cmp in
+ let normal' := if cnot then negb normal else normal in
+ let hvs := fn_cond cmp lhsv in
+ let hl := make_lhsv_single hvs in
+ if normal' then hvs else fSop (OExoriw Int.one) hl.
+
+(** ** Branches instructions *)
+
+Definition transl_cbranch_int32s (cmp: comparison) (optR: option oreg) :=
+ match cmp with
+ | Ceq => CEbeqw optR
+ | Cne => CEbnew optR
+ | Clt => CEbltw optR
+ | Cle => CEbgew optR
+ | Cgt => CEbltw optR
+ | Cge => CEbgew optR
+ end.
+
+Definition transl_cbranch_int32u (cmp: comparison) (optR: option oreg) :=
+ match cmp with
+ | Ceq => CEbequw optR
+ | Cne => CEbneuw optR
+ | Clt => CEbltuw optR
+ | Cle => CEbgeuw optR
+ | Cgt => CEbltuw optR
+ | Cge => CEbgeuw optR
+ end.
+
+Definition transl_cbranch_int64s (cmp: comparison) (optR: option oreg) :=
+ match cmp with
+ | Ceq => CEbeql optR
+ | Cne => CEbnel optR
+ | Clt => CEbltl optR
+ | Cle => CEbgel optR
+ | Cgt => CEbltl optR
+ | Cge => CEbgel optR
+ end.
+
+Definition transl_cbranch_int64u (cmp: comparison) (optR: option oreg) :=
+ match cmp with
+ | Ceq => CEbequl optR
+ | Cne => CEbneul optR
+ | Clt => CEbltul optR
+ | Cle => CEbgeul optR
+ | Cgt => CEbltul optR
+ | Cge => CEbgeul optR
+ end.
+
+Definition expanse_cbranch_fp (cnot: bool) fn_cond cmp (lhsv: list_hsval) : (condition * list_hsval) :=
+ let normal := is_normal_cmp cmp in
+ let normal' := if cnot then negb normal else normal in
+ let hvs := fn_cond cmp lhsv in
+ let hl := make_lhsv_cmp false hvs hvs in
+ if normal' then ((CEbnew (Some X0_R)), hl) else ((CEbeqw (Some X0_R)), hl).
+
+(** * Target simplifications using "fake" values *)
+
+Definition target_op_simplify (op: operation) (lr: list reg) (hst: hsistate_local): option hsval :=
+ match op, lr with
+ | Ocmp (Ccomp c), a1 :: a2 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ let hv2 := fsi_sreg_get hst a2 in
+ let is_inv := is_inv_cmp_int c in
+ let optR := make_optR false is_inv in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (cond_int32s c lhsv optR)
+ | Ocmp (Ccompu c), a1 :: a2 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ let hv2 := fsi_sreg_get hst a2 in
+ let is_inv := is_inv_cmp_int c in
+ let optR := make_optR false is_inv in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (cond_int32u c lhsv optR)
+ | Ocmp (Ccompimm c imm), a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ Some (expanse_condimm_int32s c hv1 imm)
+ | Ocmp (Ccompuimm c imm), a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ Some (expanse_condimm_int32u c hv1 imm)
+ | Ocmp (Ccompl c), a1 :: a2 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ let hv2 := fsi_sreg_get hst a2 in
+ let is_inv := is_inv_cmp_int c in
+ let optR := make_optR false is_inv in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (cond_int64s c lhsv optR)
+ | Ocmp (Ccomplu c), a1 :: a2 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ let hv2 := fsi_sreg_get hst a2 in
+ let is_inv := is_inv_cmp_int c in
+ let optR := make_optR false is_inv in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (cond_int64u c lhsv optR)
+ | Ocmp (Ccomplimm c imm), a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ Some (expanse_condimm_int64s c hv1 imm)
+ | Ocmp (Ccompluimm c imm), a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ Some (expanse_condimm_int64u c hv1 imm)
+ | Ocmp (Ccompf c), f1 :: f2 :: nil =>
+ let hv1 := fsi_sreg_get hst f1 in
+ let hv2 := fsi_sreg_get hst f2 in
+ let is_inv := is_inv_cmp_float c in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (expanse_cond_fp false cond_float c lhsv)
+ | Ocmp (Cnotcompf c), f1 :: f2 :: nil =>
+ let hv1 := fsi_sreg_get hst f1 in
+ let hv2 := fsi_sreg_get hst f2 in
+ let is_inv := is_inv_cmp_float c in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (expanse_cond_fp true cond_float c lhsv)
+ | Ocmp (Ccompfs c), f1 :: f2 :: nil =>
+ let hv1 := fsi_sreg_get hst f1 in
+ let hv2 := fsi_sreg_get hst f2 in
+ let is_inv := is_inv_cmp_float c in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (expanse_cond_fp false cond_single c lhsv)
+ | Ocmp (Cnotcompfs c), f1 :: f2 :: nil =>
+ let hv1 := fsi_sreg_get hst f1 in
+ let hv2 := fsi_sreg_get hst f2 in
+ let is_inv := is_inv_cmp_float c in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (expanse_cond_fp true cond_single c lhsv)
+ | Ofloatconst f, nil =>
+ let hvs := loadimm64 (Float.to_bits f) in
+ let hl := make_lhsv_single hvs in
+ Some (fSop (Ofloat_of_bits) hl)
+ | Osingleconst f, nil =>
+ let hvs := loadimm32 (Float32.to_bits f) in
+ let hl := make_lhsv_single hvs in
+ Some (fSop (Osingle_of_bits) hl)
+ | Ointconst n, nil =>
+ Some (loadimm32 n)
+ | Olongconst n, nil =>
+ Some (loadimm64 n)
+ | Oaddimm n, a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ Some (addimm32 hv1 n None)
+ | Oaddlimm n, a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ Some (addimm64 hv1 n None)
+ | Oandimm n, a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ Some (andimm32 hv1 n)
+ | Oandlimm n, a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ Some (andimm64 hv1 n)
+ | Oorimm n, a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ Some (orimm32 hv1 n)
+ | Oorlimm n, a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ Some (orimm64 hv1 n)
+ | Oxorimm n, a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ Some (xorimm32 hv1 n)
+ | Oxorlimm n, a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ Some (xorimm64 hv1 n)
+ | Ocast8signed, a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ let hl := make_lhsv_single hv1 in
+ let hvs := fSop (Oshlimm (Int.repr 24)) hl in
+ let hl' := make_lhsv_single hvs in
+ Some (fSop (Oshrimm (Int.repr 24)) hl')
+ | Ocast16signed, a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ let hl := make_lhsv_single hv1 in
+ let hvs := fSop (Oshlimm (Int.repr 16)) hl in
+ let hl' := make_lhsv_single hvs in
+ Some (fSop (Oshrimm (Int.repr 16)) hl')
+ | Ocast32unsigned, a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ let hl := make_lhsv_single hv1 in
+ let cast32s_s := fSop Ocast32signed hl in
+ let cast32s_l := make_lhsv_single cast32s_s in
+ let sllil_s := fSop (Oshllimm (Int.repr 32)) cast32s_l in
+ let sllil_l := make_lhsv_single sllil_s in
+ Some (fSop (Oshrluimm (Int.repr 32)) sllil_l)
+ | Oshrximm n, a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ let hl := make_lhsv_single hv1 in
+ if Int.eq n Int.zero then
+ let lhl := make_lhsv_cmp false hv1 hv1 in
+ Some (fSop (OEmayundef (MUshrx n)) lhl)
+ else
+ if Int.eq n Int.one then
+ let srliw_s := fSop (Oshruimm (Int.repr 31)) hl in
+ let srliw_l := make_lhsv_cmp false hv1 srliw_s in
+ let addw_s := fSop Oadd srliw_l in
+ let addw_l := make_lhsv_single addw_s in
+ let sraiw_s := fSop (Oshrimm Int.one) addw_l in
+ let sraiw_l := make_lhsv_cmp false sraiw_s sraiw_s in
+ Some (fSop (OEmayundef (MUshrx n)) sraiw_l)
+ else
+ let sraiw_s := fSop (Oshrimm (Int.repr 31)) hl in
+ let sraiw_l := make_lhsv_single sraiw_s in
+ let srliw_s := fSop (Oshruimm (Int.sub Int.iwordsize n)) sraiw_l in
+ let srliw_l := make_lhsv_cmp false hv1 srliw_s in
+ let addw_s := fSop Oadd srliw_l in
+ let addw_l := make_lhsv_single addw_s in
+ let sraiw_s' := fSop (Oshrimm n) addw_l in
+ let sraiw_l' := make_lhsv_cmp false sraiw_s' sraiw_s' in
+ Some (fSop (OEmayundef (MUshrx n)) sraiw_l')
+ | Oshrxlimm n, a1 :: nil =>
+ let hv1 := fsi_sreg_get hst a1 in
+ let hl := make_lhsv_single hv1 in
+ if Int.eq n Int.zero then
+ let lhl := make_lhsv_cmp false hv1 hv1 in
+ Some (fSop (OEmayundef (MUshrxl n)) lhl)
+ else
+ if Int.eq n Int.one then
+ let srlil_s := fSop (Oshrluimm (Int.repr 63)) hl in
+ let srlil_l := make_lhsv_cmp false hv1 srlil_s in
+ let addl_s := fSop Oaddl srlil_l in
+ let addl_l := make_lhsv_single addl_s in
+ let srail_s := fSop (Oshrlimm Int.one) addl_l in
+ let srail_l := make_lhsv_cmp false srail_s srail_s in
+ Some (fSop (OEmayundef (MUshrxl n)) srail_l)
+ else
+ let srail_s := fSop (Oshrlimm (Int.repr 63)) hl in
+ let srail_l := make_lhsv_single srail_s in
+ let srlil_s := fSop (Oshrluimm (Int.sub Int64.iwordsize' n)) srail_l in
+ let srlil_l := make_lhsv_cmp false hv1 srlil_s in
+ let addl_s := fSop Oaddl srlil_l in
+ let addl_l := make_lhsv_single addl_s in
+ let srail_s' := fSop (Oshrlimm n) addl_l in
+ let srail_l' := make_lhsv_cmp false srail_s' srail_s' in
+ Some (fSop (OEmayundef (MUshrxl n)) srail_l')
+ | _, _ => None
+ end.
+
+Definition target_cbranch_expanse (prev: hsistate_local) (cond: condition) (args: list reg) : option (condition * list_hsval) :=
+ match cond, args with
+ | (Ccomp c), (a1 :: a2 :: nil) =>
+ let is_inv := is_inv_cmp_int c in
+ let cond := transl_cbranch_int32s c (make_optR false is_inv) in
+ let hv1 := fsi_sreg_get prev a1 in
+ let hv2 := fsi_sreg_get prev a2 in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (cond, lhsv)
+ | (Ccompu c), (a1 :: a2 :: nil) =>
+ let is_inv := is_inv_cmp_int c in
+ let cond := transl_cbranch_int32u c (make_optR false is_inv) in
+ let hv1 := fsi_sreg_get prev a1 in
+ let hv2 := fsi_sreg_get prev a2 in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (cond, lhsv)
+ | (Ccompimm c n), (a1 :: nil) =>
+ let is_inv := is_inv_cmp_int c in
+ let hv1 := fsi_sreg_get prev a1 in
+ (if Int.eq n Int.zero then
+ let lhsv := make_lhsv_cmp is_inv hv1 hv1 in
+ let cond := transl_cbranch_int32s c (make_optR true is_inv) in
+ Some (cond, lhsv)
+ else
+ let hvs := loadimm32 n in
+ let lhsv := make_lhsv_cmp is_inv hv1 hvs in
+ let cond := transl_cbranch_int32s c (make_optR false is_inv) in
+ Some (cond, lhsv))
+ | (Ccompuimm c n), (a1 :: nil) =>
+ let is_inv := is_inv_cmp_int c in
+ let hv1 := fsi_sreg_get prev a1 in
+ (if Int.eq n Int.zero then
+ let lhsv := make_lhsv_cmp is_inv hv1 hv1 in
+ let cond := transl_cbranch_int32u c (make_optR true is_inv) in
+ Some (cond, lhsv)
+ else
+ let hvs := loadimm32 n in
+ let lhsv := make_lhsv_cmp is_inv hv1 hvs in
+ let cond := transl_cbranch_int32u c (make_optR false is_inv) in
+ Some (cond, lhsv))
+ | (Ccompl c), (a1 :: a2 :: nil) =>
+ let is_inv := is_inv_cmp_int c in
+ let cond := transl_cbranch_int64s c (make_optR false is_inv) in
+ let hv1 := fsi_sreg_get prev a1 in
+ let hv2 := fsi_sreg_get prev a2 in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (cond, lhsv)
+ | (Ccomplu c), (a1 :: a2 :: nil) =>
+ let is_inv := is_inv_cmp_int c in
+ let cond := transl_cbranch_int64u c (make_optR false is_inv) in
+ let hv1 := fsi_sreg_get prev a1 in
+ let hv2 := fsi_sreg_get prev a2 in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (cond, lhsv)
+ | (Ccomplimm c n), (a1 :: nil) =>
+ let is_inv := is_inv_cmp_int c in
+ let hv1 := fsi_sreg_get prev a1 in
+ (if Int64.eq n Int64.zero then
+ let lhsv := make_lhsv_cmp is_inv hv1 hv1 in
+ let cond := transl_cbranch_int64s c (make_optR true is_inv) in
+ Some (cond, lhsv)
+ else
+ let hvs := loadimm64 n in
+ let lhsv := make_lhsv_cmp is_inv hv1 hvs in
+ let cond := transl_cbranch_int64s c (make_optR false is_inv) in
+ Some (cond, lhsv))
+ | (Ccompluimm c n), (a1 :: nil) =>
+ let is_inv := is_inv_cmp_int c in
+ let hv1 := fsi_sreg_get prev a1 in
+ (if Int64.eq n Int64.zero then
+ let lhsv := make_lhsv_cmp is_inv hv1 hv1 in
+ let cond := transl_cbranch_int64u c (make_optR true is_inv) in
+ Some (cond, lhsv)
+ else
+ let hvs := loadimm64 n in
+ let lhsv := make_lhsv_cmp is_inv hv1 hvs in
+ let cond := transl_cbranch_int64u c (make_optR false is_inv) in
+ Some (cond, lhsv))
+ | (Ccompf c), (f1 :: f2 :: nil) =>
+ let hv1 := fsi_sreg_get prev f1 in
+ let hv2 := fsi_sreg_get prev f2 in
+ let is_inv := is_inv_cmp_float c in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (expanse_cbranch_fp false cond_float c lhsv)
+ | (Cnotcompf c), (f1 :: f2 :: nil) =>
+ let hv1 := fsi_sreg_get prev f1 in
+ let hv2 := fsi_sreg_get prev f2 in
+ let is_inv := is_inv_cmp_float c in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (expanse_cbranch_fp true cond_float c lhsv)
+ | (Ccompfs c), (f1 :: f2 :: nil) =>
+ let hv1 := fsi_sreg_get prev f1 in
+ let hv2 := fsi_sreg_get prev f2 in
+ let is_inv := is_inv_cmp_float c in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (expanse_cbranch_fp false cond_single c lhsv)
+ | (Cnotcompfs c), (f1 :: f2 :: nil) =>
+ let hv1 := fsi_sreg_get prev f1 in
+ let hv2 := fsi_sreg_get prev f2 in
+ let is_inv := is_inv_cmp_float c in
+ let lhsv := make_lhsv_cmp is_inv hv1 hv2 in
+ Some (expanse_cbranch_fp true cond_single c lhsv)
+ | _, _ => None
+ end.
+
+(** * Auxiliary lemmas on comparisons *)
+
+(** ** Signed ints *)
+
+Lemma xor_neg_ltle_cmp: forall v1 v2,
+ Some (Val.xor (Val.cmp Clt v1 v2) (Vint Int.one)) =
+ Some (Val.of_optbool (Val.cmp_bool Cle v2 v1)).
+Proof.
+ intros. eapply f_equal.
+ destruct v1, v2; simpl; try congruence.
+ unfold Val.cmp; simpl;
+ try rewrite Int.eq_sym;
+ try destruct (Int.eq _ _); try destruct (Int.lt _ _) eqn:ELT ; simpl;
+ try rewrite Int.xor_one_one; try rewrite Int.xor_zero_one;
+ auto.
+Qed.
+
+(** ** Unsigned ints *)
+
+Lemma xor_neg_ltle_cmpu: forall mptr v1 v2,
+ Some (Val.xor (Val.cmpu (Mem.valid_pointer mptr) Clt v1 v2) (Vint Int.one)) =
+ Some (Val.of_optbool (Val.cmpu_bool (Mem.valid_pointer mptr) Cle v2 v1)).
+Proof.
+ intros. eapply f_equal.
+ destruct v1, v2; simpl; try congruence.
+ unfold Val.cmpu; simpl;
+ try rewrite Int.eq_sym;
+ try destruct (Int.eq _ _); try destruct (Int.ltu _ _) eqn:ELT ; simpl;
+ try rewrite Int.xor_one_one; try rewrite Int.xor_zero_one;
+ auto.
+ 1,2:
+ unfold Val.cmpu, Val.cmpu_bool;
+ destruct Archi.ptr64; try destruct (_ && _); try destruct (_ || _);
+ try destruct (eq_block _ _); auto.
+ unfold Val.cmpu, Val.cmpu_bool; simpl;
+ destruct Archi.ptr64; try destruct (_ || _); simpl; auto;
+ destruct (eq_block b b0); destruct (eq_block b0 b);
+ try congruence;
+ try destruct (_ || _); simpl; try destruct (Ptrofs.ltu _ _);
+ simpl; auto;
+ repeat destruct (_ && _); simpl; auto.
+Qed.
+
+Remark ltu_12_wordsize:
+ Int.ltu (Int.repr 12) Int.iwordsize = true.
+Proof.
+ unfold Int.iwordsize, Int.zwordsize. simpl.
+ unfold Int.ltu. apply zlt_true.
+ rewrite !Int.unsigned_repr; try cbn; try lia.
+Qed.
+
+(** ** Signed longs *)
+
+Lemma xor_neg_ltle_cmpl: forall v1 v2,
+ Some (Val.xor (Val.maketotal (Val.cmpl Clt v1 v2)) (Vint Int.one)) =
+ Some (Val.of_optbool (Val.cmpl_bool Cle v2 v1)).
+Proof.
+ intros. eapply f_equal.
+ destruct v1, v2; simpl; try congruence.
+ destruct (Int64.lt _ _); auto.
+Qed.
+
+Lemma xor_neg_ltge_cmpl: forall v1 v2,
+ Some (Val.xor (Val.maketotal (Val.cmpl Clt v1 v2)) (Vint Int.one)) =
+ Some (Val.of_optbool (Val.cmpl_bool Cge v1 v2)).
+Proof.
+ intros. eapply f_equal.
+ destruct v1, v2; simpl; try congruence.
+ destruct (Int64.lt _ _); auto.
+Qed.
+
+Lemma xorl_zero_eq_cmpl: forall c v1 v2,
+ c = Ceq \/ c = Cne ->
+ Some
+ (Val.maketotal
+ (option_map Val.of_bool
+ (Val.cmpl_bool c (Val.xorl v1 v2) (Vlong Int64.zero)))) =
+ Some (Val.of_optbool (Val.cmpl_bool c v1 v2)).
+Proof.
+ intros. destruct c; inv H; try discriminate;
+ destruct v1, v2; simpl; auto;
+ destruct (Int64.eq i i0) eqn:EQ0.
+ 1,3:
+ apply Int64.same_if_eq in EQ0; subst;
+ rewrite Int64.xor_idem;
+ rewrite Int64.eq_true; trivial.
+ 1,2:
+ destruct (Int64.eq (Int64.xor i i0) Int64.zero) eqn:EQ1; simpl; try congruence;
+ rewrite Int64.xor_is_zero in EQ1; congruence.
+Qed.
+
+Lemma cmp_ltle_add_one: forall v n,
+ Int.eq n (Int.repr Int.max_signed) = false ->
+ Some (Val.of_optbool (Val.cmp_bool Clt v (Vint (Int.add n Int.one)))) =
+ Some (Val.of_optbool (Val.cmp_bool Cle v (Vint n))).
+Proof.
+ intros v n EQMAX. unfold Val.cmp_bool; destruct v; simpl; auto.
+ unfold Int.lt. replace (Int.signed (Int.add n Int.one)) with (Int.signed n + 1).
+ destruct (zlt (Int.signed n) (Int.signed i)).
+ rewrite zlt_false by lia. auto.
+ rewrite zlt_true by lia. auto.
+ rewrite Int.add_signed. symmetry; apply Int.signed_repr.
+ specialize (Int.eq_spec n (Int.repr Int.max_signed)).
+ rewrite EQMAX; simpl; intros.
+ assert (Int.signed n <> Int.max_signed).
+ { red; intros E. elim H. rewrite <- (Int.repr_signed n). rewrite E. auto. }
+ generalize (Int.signed_range n); lia.
+Qed.
+
+Lemma cmpl_ltle_add_one: forall v n,
+ Int64.eq n (Int64.repr Int64.max_signed) = false ->
+ Some (Val.of_optbool (Val.cmpl_bool Clt v (Vlong (Int64.add n Int64.one)))) =
+ Some (Val.of_optbool (Val.cmpl_bool Cle v (Vlong n))).
+Proof.
+ intros v n EQMAX. unfold Val.cmpl_bool; destruct v; simpl; auto.
+ unfold Int64.lt. replace (Int64.signed (Int64.add n Int64.one)) with (Int64.signed n + 1).
+ destruct (zlt (Int64.signed n) (Int64.signed i)).
+ rewrite zlt_false by lia. auto.
+ rewrite zlt_true by lia. auto.
+ rewrite Int64.add_signed. symmetry; apply Int64.signed_repr.
+ specialize (Int64.eq_spec n (Int64.repr Int64.max_signed)).
+ rewrite EQMAX; simpl; intros.
+ assert (Int64.signed n <> Int64.max_signed).
+ { red; intros E. elim H. rewrite <- (Int64.repr_signed n). rewrite E. auto. }
+ generalize (Int64.signed_range n); lia.
+Qed.
+
+Remark lt_maxsgn_false_int: forall i,
+ Int.lt (Int.repr Int.max_signed) i = false.
+Proof.
+ intros; unfold Int.lt.
+ specialize Int.signed_range with i; intros.
+ rewrite zlt_false; auto. destruct H.
+ rewrite Int.signed_repr; try (cbn; lia).
+ apply Z.le_ge. trivial.
+Qed.
+
+Remark lt_maxsgn_false_long: forall i,
+ Int64.lt (Int64.repr Int64.max_signed) i = false.
+Proof.
+ intros; unfold Int64.lt.
+ specialize Int64.signed_range with i; intros.
+ rewrite zlt_false; auto. destruct H.
+ rewrite Int64.signed_repr; try (cbn; lia).
+ apply Z.le_ge. trivial.
+Qed.
+
+(** ** Unsigned longs *)
+
+Lemma xor_neg_ltle_cmplu: forall mptr v1 v2,
+ Some (Val.xor (Val.maketotal (Val.cmplu (Mem.valid_pointer mptr) Clt v1 v2)) (Vint Int.one)) =
+ Some (Val.of_optbool (Val.cmplu_bool (Mem.valid_pointer mptr) Cle v2 v1)).
+Proof.
+ intros. eapply f_equal.
+ destruct v1, v2; simpl; try congruence.
+ destruct (Int64.ltu _ _); auto.
+ 1,2: unfold Val.cmplu; simpl; auto;
+ destruct (Archi.ptr64); simpl;
+ try destruct (eq_block _ _); simpl;
+ try destruct (_ && _); simpl;
+ try destruct (Ptrofs.cmpu _ _);
+ try destruct cmp; simpl; auto.
+ unfold Val.cmplu; simpl;
+ destruct Archi.ptr64; try destruct (_ || _); simpl; auto;
+ destruct (eq_block b b0); destruct (eq_block b0 b);
+ try congruence;
+ try destruct (_ || _); simpl; try destruct (Ptrofs.ltu _ _);
+ simpl; auto;
+ repeat destruct (_ && _); simpl; auto.
+Qed.
+
+Lemma xor_neg_ltge_cmplu: forall mptr v1 v2,
+ Some (Val.xor (Val.maketotal (Val.cmplu (Mem.valid_pointer mptr) Clt v1 v2)) (Vint Int.one)) =
+ Some (Val.of_optbool (Val.cmplu_bool (Mem.valid_pointer mptr) Cge v1 v2)).
+Proof.
+ intros. eapply f_equal.
+ destruct v1, v2; simpl; try congruence.
+ destruct (Int64.ltu _ _); auto.
+ 1,2: unfold Val.cmplu; simpl; auto;
+ destruct (Archi.ptr64); simpl;
+ try destruct (eq_block _ _); simpl;
+ try destruct (_ && _); simpl;
+ try destruct (Ptrofs.cmpu _ _);
+ try destruct cmp; simpl; auto.
+ unfold Val.cmplu; simpl;
+ destruct Archi.ptr64; try destruct (_ || _); simpl; auto;
+ destruct (eq_block b b0); destruct (eq_block b0 b);
+ try congruence;
+ try destruct (_ || _); simpl; try destruct (Ptrofs.ltu _ _);
+ simpl; auto;
+ repeat destruct (_ && _); simpl; auto.
+Qed.
+
+(** ** Floats *)
+
+Lemma xor_neg_eqne_cmpf: forall v1 v2,
+ Some (Val.xor (Val.cmpf Ceq v1 v2) (Vint Int.one)) =
+ Some (Val.of_optbool (Val.cmpf_bool Cne v1 v2)).
+Proof.
+ intros. eapply f_equal.
+ destruct v1, v2; simpl; try congruence;
+ unfold Val.cmpf; simpl.
+ rewrite Float.cmp_ne_eq.
+ destruct (Float.cmp _ _ _); simpl; auto.
+Qed.
+
+(** ** Singles *)
+
+Lemma xor_neg_eqne_cmpfs: forall v1 v2,
+ Some (Val.xor (Val.cmpfs Ceq v1 v2) (Vint Int.one)) =
+ Some (Val.of_optbool (Val.cmpfs_bool Cne v1 v2)).
+Proof.
+ intros. eapply f_equal.
+ destruct v1, v2; simpl; try congruence;
+ unfold Val.cmpfs; simpl.
+ rewrite Float32.cmp_ne_eq.
+ destruct (Float32.cmp _ _ _); simpl; auto.
+Qed.
+
+(** ** More useful lemmas *)
+
+Lemma xor_neg_optb: forall v,
+ Some (Val.xor (Val.of_optbool (option_map negb v))
+ (Vint Int.one)) = Some (Val.of_optbool v).
+Proof.
+ intros.
+ destruct v; simpl; trivial.
+ destruct b; simpl; auto.
+Qed.
+
+Lemma xor_neg_optb': forall v,
+ Some (Val.xor (Val.of_optbool v) (Vint Int.one)) =
+ Some (Val.of_optbool (option_map negb v)).
+Proof.
+ intros.
+ destruct v; simpl; trivial.
+ destruct b; simpl; auto.
+Qed.
+
+Lemma optbool_mktotal: forall v,
+ Val.maketotal (option_map Val.of_bool v) =
+ Val.of_optbool v.
+Proof.
+ intros.
+ destruct v; simpl; auto.
+Qed.
+
+(* TODO gourdinl move to common/Values ? *)
+Theorem swap_cmpf_bool:
+ forall c x y,
+ Val.cmpf_bool (swap_comparison c) x y = Val.cmpf_bool c y x.
+Proof.
+ destruct x; destruct y; simpl; auto. rewrite Float.cmp_swap. auto.
+Qed.
+
+Theorem swap_cmpfs_bool:
+ forall c x y,
+ Val.cmpfs_bool (swap_comparison c) x y = Val.cmpfs_bool c y x.
+Proof.
+ destruct x; destruct y; simpl; auto. rewrite Float32.cmp_swap. auto.
+Qed.
+
+(** * Intermediates lemmas on each expanded instruction *)
+
+Lemma simplify_ccomp_correct ge sp hst st c r r0 rs0 m0 v v0: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (OKv1 : seval_sval ge sp (si_sreg st r) rs0 m0 = Some v)
+ (OKv2 : seval_sval ge sp (si_sreg st r0) rs0 m0 = Some v0),
+ seval_sval ge sp
+ (hsval_proj
+ (cond_int32s c
+ (make_lhsv_cmp (is_inv_cmp_int c) (fsi_sreg_get hst r)
+ (fsi_sreg_get hst r0)) None)) rs0 m0 =
+ Some (Val.of_optbool (Val.cmp_bool c v v0)).
+Proof.
+ intros.
+ unfold cond_int32s in *; destruct c; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ rewrite OKv1, OKv2; trivial;
+ unfold Val.cmp.
+ - apply xor_neg_ltle_cmp.
+ - replace (Clt) with (swap_comparison Cgt) by auto;
+ rewrite Val.swap_cmp_bool; trivial.
+ - replace (Clt) with (negate_comparison Cge) by auto;
+ rewrite Val.negate_cmp_bool.
+ rewrite xor_neg_optb; trivial.
+Qed.
+
+Lemma simplify_ccompu_correct ge sp hst st c r r0 rs0 m m0 v v0: forall
+ (SMEM : forall (m : mem) (b : Values.block) (ofs : Z),
+ seval_smem ge sp (si_smem st) rs0 m0 = Some m ->
+ Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs)
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (OKv1 : seval_sval ge sp (si_sreg st r) rs0 m0 = Some v)
+ (OKv2 : seval_sval ge sp (si_sreg st r0) rs0 m0 = Some v0)
+ (OK2 : seval_smem ge sp (si_smem st) rs0 m0 = Some m),
+ seval_sval ge sp
+ (hsval_proj
+ (cond_int32u c
+ (make_lhsv_cmp (is_inv_cmp_int c) (fsi_sreg_get hst r)
+ (fsi_sreg_get hst r0)) None)) rs0 m0 =
+ Some (Val.of_optbool (Val.cmpu_bool (Mem.valid_pointer m) c v v0)).
+Proof.
+ intros.
+ erewrite (cmpu_bool_valid_pointer_eq (Mem.valid_pointer m) (Mem.valid_pointer m0)).
+ 2: eauto.
+ unfold cond_int32u in *; destruct c; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ rewrite OKv1, OKv2; trivial;
+ unfold Val.cmpu.
+ - apply xor_neg_ltle_cmpu.
+ - replace (Clt) with (swap_comparison Cgt) by auto;
+ rewrite Val.swap_cmpu_bool; trivial.
+ - replace (Clt) with (negate_comparison Cge) by auto;
+ rewrite Val.negate_cmpu_bool.
+ rewrite xor_neg_optb; trivial.
+Qed.
+
+Lemma simplify_ccompimm_correct ge sp hst st c r n rs0 m m0 v: forall
+ (SMEM : forall (m : mem) (b : Values.block) (ofs : Z),
+ seval_smem ge sp (si_smem st) rs0 m0 = Some m ->
+ Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs)
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (OKv1 : seval_sval ge sp (si_sreg st r) rs0 m0 = Some v)
+ (OK2 : seval_smem ge sp (si_smem st) rs0 m0 = Some m),
+ seval_sval ge sp
+ (hsval_proj (expanse_condimm_int32s c (fsi_sreg_get hst r) n)) rs0 m0 =
+ Some (Val.of_optbool (Val.cmp_bool c v (Vint n))).
+Proof.
+ intros.
+ unfold expanse_condimm_int32s, cond_int32s in *; destruct c;
+ intros; destruct (Int.eq n Int.zero) eqn:EQIMM; simpl;
+ try apply Int.same_if_eq in EQIMM; subst;
+ unfold loadimm32, sltimm32, xorimm32, opimm32, load_hilo32;
+ try erewrite !fsi_sreg_get_correct; eauto;
+ try rewrite OKv1;
+ unfold Val.cmp, zero32.
+ all:
+ try apply xor_neg_ltle_cmp;
+ try apply xor_neg_ltge_cmp; trivial.
+ 4:
+ try destruct (Int.eq n (Int.repr Int.max_signed)) eqn:EQMAX; subst;
+ try apply Int.same_if_eq in EQMAX; subst; simpl.
+ 4:
+ intros; try (specialize make_immed32_sound with (Int.one);
+ destruct (make_immed32 Int.one) eqn:EQMKI_A1); intros; simpl.
+ 6:
+ intros; try (specialize make_immed32_sound with (Int.add n Int.one);
+ destruct (make_immed32 (Int.add n Int.one)) eqn:EQMKI_A2); intros; simpl.
+ 1,2,3,8,9:
+ intros; try (specialize make_immed32_sound with (n);
+ destruct (make_immed32 n) eqn:EQMKI); intros; simpl.
+ all:
+ try destruct (Int.eq lo Int.zero) eqn:EQLO32;
+ try apply Int.same_if_eq in EQLO32; subst;
+ try erewrite fSop_correct; eauto; simpl;
+ try erewrite !fsi_sreg_get_correct; eauto;
+ try rewrite OKv1;
+ try rewrite OK2;
+ try rewrite (Int.add_commut _ Int.zero), Int.add_zero_l in H; subst;
+ unfold Val.cmp, eval_may_undef, zero32, Val.add; simpl;
+ destruct v; auto.
+ all:
+ try rewrite ltu_12_wordsize;
+ try rewrite <- H;
+ try (apply cmp_ltle_add_one; auto);
+ try rewrite Int.add_commut, Int.add_zero_l in *;
+ try rewrite Int.add_commut;
+ try rewrite <- H; try rewrite cmp_ltle_add_one;
+ try rewrite Int.add_zero_l;
+ try (
+ simpl; trivial;
+ try rewrite Int.xor_is_zero;
+ try destruct (Int.lt _ _) eqn:EQLT; trivial;
+ try rewrite lt_maxsgn_false_int in EQLT;
+ simpl; trivial; try discriminate; fail).
+Qed.
+
+Lemma simplify_ccompuimm_correct ge sp hst st c r n rs0 m m0 v: forall
+ (SMEM : forall (m : mem) (b : Values.block) (ofs : Z),
+ seval_smem ge sp (si_smem st) rs0 m0 = Some m ->
+ Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs)
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (OKv1 : seval_sval ge sp (si_sreg st r) rs0 m0 = Some v)
+ (OK2 : seval_smem ge sp (si_smem st) rs0 m0 = Some m),
+ seval_sval ge sp
+ (hsval_proj (expanse_condimm_int32u c (fsi_sreg_get hst r) n)) rs0 m0 =
+ Some (Val.of_optbool (Val.cmpu_bool (Mem.valid_pointer m) c v (Vint n))).
+Proof.
+ intros.
+ assert (HMEM: Val.cmpu_bool (Mem.valid_pointer m) c v (Vint n) =
+ Val.cmpu_bool (Mem.valid_pointer m0) c v (Vint n)).
+ erewrite (cmpu_bool_valid_pointer_eq (Mem.valid_pointer m) (Mem.valid_pointer m0)); eauto.
+ unfold expanse_condimm_int32u, cond_int32u in *; destruct c;
+ intros; destruct (Int.eq n Int.zero) eqn:EQIMM; simpl;
+ try apply Int.same_if_eq in EQIMM; subst;
+ unfold loadimm32, sltuimm32, opimm32, load_hilo32;
+ try erewrite !fsi_sreg_get_correct; eauto;
+ try rewrite OKv1; trivial;
+ try rewrite xor_neg_ltle_cmpu;
+ unfold Val.cmpu, zero32.
+ all:
+ try (specialize make_immed32_sound with n;
+ destruct (make_immed32 n) eqn:EQMKI);
+ try destruct (Int.eq lo Int.zero) eqn:EQLO;
+ try apply Int.same_if_eq in EQLO; subst;
+ intros; subst;
+ try erewrite fSop_correct; eauto; simpl;
+ try erewrite !fsi_sreg_get_correct; eauto;
+ try rewrite OKv1;
+ try rewrite OK2;
+ rewrite HMEM;
+ unfold eval_may_undef, Val.cmpu;
+ destruct v; simpl; auto;
+ try rewrite EQIMM; try destruct (Archi.ptr64) eqn:EQARCH; simpl;
+ try rewrite ltu_12_wordsize; trivial;
+ try rewrite Int.add_commut, Int.add_zero_l in *;
+ try rewrite Int.add_zero_l;
+ try destruct (Int.ltu _ _) eqn:EQLTU; simpl;
+ try rewrite EQLTU; simpl; try rewrite EQIMM;
+ try rewrite EQARCH; trivial.
+Qed.
+
+Lemma simplify_ccompl_correct ge sp hst st c r r0 rs0 m0 v v0: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (OKv1 : seval_sval ge sp (si_sreg st r) rs0 m0 = Some v)
+ (OKv2 : seval_sval ge sp (si_sreg st r0) rs0 m0 = Some v0),
+ seval_sval ge sp
+ (hsval_proj
+ (cond_int64s c
+ (make_lhsv_cmp (is_inv_cmp_int c) (fsi_sreg_get hst r)
+ (fsi_sreg_get hst r0)) None)) rs0 m0 =
+ Some (Val.of_optbool (Val.cmpl_bool c v v0)).
+Proof.
+ intros.
+ unfold cond_int64s in *; destruct c; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ rewrite OKv1, OKv2; trivial;
+ unfold Val.cmpl.
+ 1,2,3: rewrite optbool_mktotal; trivial.
+ - apply xor_neg_ltle_cmpl.
+ - replace (Clt) with (swap_comparison Cgt) by auto;
+ rewrite Val.swap_cmpl_bool; trivial.
+ rewrite optbool_mktotal; trivial.
+ - apply xor_neg_ltge_cmpl.
+Qed.
+
+Lemma simplify_ccomplu_correct ge sp hst st c r r0 rs0 m m0 v v0: forall
+ (SMEM : forall (m : mem) (b : Values.block) (ofs : Z),
+ seval_smem ge sp (si_smem st) rs0 m0 = Some m ->
+ Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs)
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (OKv1 : seval_sval ge sp (si_sreg st r) rs0 m0 = Some v)
+ (OKv2 : seval_sval ge sp (si_sreg st r0) rs0 m0 = Some v0)
+ (OK2 : seval_smem ge sp (si_smem st) rs0 m0 = Some m),
+ seval_sval ge sp
+ (hsval_proj
+ (cond_int64u c
+ (make_lhsv_cmp (is_inv_cmp_int c) (fsi_sreg_get hst r)
+ (fsi_sreg_get hst r0)) None)) rs0 m0 =
+ Some (Val.of_optbool (Val.cmplu_bool (Mem.valid_pointer m) c v v0)).
+Proof.
+ intros.
+ erewrite (cmplu_bool_valid_pointer_eq (Mem.valid_pointer m) (Mem.valid_pointer m0)).
+ 2: eauto.
+ unfold cond_int64u in *; destruct c; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ rewrite OKv1, OKv2; trivial;
+ unfold Val.cmplu.
+ 1,2,3: rewrite optbool_mktotal; trivial.
+ - apply xor_neg_ltle_cmplu.
+ - replace (Clt) with (swap_comparison Cgt) by auto;
+ rewrite Val.swap_cmplu_bool; trivial.
+ rewrite optbool_mktotal; trivial.
+ - apply xor_neg_ltge_cmplu.
+Qed.
+
+Lemma simplify_ccomplimm_correct ge sp hst st c r n rs0 m m0 v: forall
+ (SMEM : forall (m : mem) (b : Values.block) (ofs : Z),
+ seval_smem ge sp (si_smem st) rs0 m0 = Some m ->
+ Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs)
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (OKv1 : seval_sval ge sp (si_sreg st r) rs0 m0 = Some v)
+ (OK2 : seval_smem ge sp (si_smem st) rs0 m0 = Some m),
+ seval_sval ge sp
+ (hsval_proj (expanse_condimm_int64s c (fsi_sreg_get hst r) n)) rs0 m0 =
+ Some (Val.of_optbool (Val.cmpl_bool c v (Vlong n))).
+Proof.
+ intros.
+ unfold expanse_condimm_int64s, cond_int64s in *; destruct c;
+ intros; destruct (Int64.eq n Int64.zero) eqn:EQIMM; simpl;
+ try apply Int64.same_if_eq in EQIMM; subst;
+ unfold loadimm32, loadimm64, sltimm64, xorimm64, opimm64, load_hilo32, load_hilo64;
+ try erewrite !fsi_sreg_get_correct; eauto;
+ try rewrite OKv1;
+ unfold Val.cmpl, zero64.
+ all:
+ try apply xor_neg_ltle_cmpl;
+ try apply xor_neg_ltge_cmpl;
+ try rewrite optbool_mktotal; trivial.
+ 4:
+ try destruct (Int64.eq n (Int64.repr Int64.max_signed)) eqn:EQMAX; subst;
+ try apply Int64.same_if_eq in EQMAX; subst; simpl.
+ 4:
+ intros; try (specialize make_immed32_sound with (Int.one);
+ destruct (make_immed32 Int.one) eqn:EQMKI_A1); intros; simpl.
+ 6:
+ intros; try (specialize make_immed64_sound with (Int64.add n Int64.one);
+ destruct (make_immed64 (Int64.add n Int64.one)) eqn:EQMKI_A2); intros; simpl.
+ 1,2,3,9,10:
+ intros; try (specialize make_immed64_sound with (n);
+ destruct (make_immed64 n) eqn:EQMKI); intros; simpl.
+ all:
+ try destruct (Int.eq lo Int.zero) eqn:EQLO32;
+ try apply Int.same_if_eq in EQLO32; subst;
+ try destruct (Int64.eq lo Int64.zero) eqn:EQLO64;
+ try apply Int64.same_if_eq in EQLO64; subst;
+ try erewrite fSop_correct; eauto; simpl;
+ try erewrite !fsi_sreg_get_correct; eauto;
+ try rewrite OKv1;
+ try rewrite OK2;
+ try rewrite (Int64.add_commut _ Int64.zero), Int64.add_zero_l in H; subst;
+ try fold (Val.cmpl Clt v (Vlong imm));
+ try rewrite xor_neg_ltge_cmpl; trivial;
+ try rewrite xor_neg_ltle_cmpl; trivial;
+ unfold Val.cmpl, Val.addl;
+ try rewrite xorl_zero_eq_cmpl; trivial;
+ try rewrite optbool_mktotal; trivial;
+ unfold eval_may_undef, zero32, Val.add; simpl;
+ destruct v; auto.
+ 1,2,3,4,5,6,7,8,9,10,11,12:
+ try rewrite <- optbool_mktotal; trivial;
+ try rewrite Int64.add_commut, Int64.add_zero_l in *;
+ try fold (Val.cmpl Clt (Vlong i) (Vlong imm));
+ try fold (Val.cmpl Clt (Vlong i) (Vlong (Int64.sign_ext 32 (Int64.shl hi (Int64.repr 12)))));
+ try fold (Val.cmpl Clt (Vlong i) (Vlong (Int64.add (Int64.sign_ext 32 (Int64.shl hi (Int64.repr 12))) lo)));
+ try rewrite xor_neg_ltge_cmpl; trivial;
+ try rewrite xor_neg_ltle_cmpl; trivial.
+ 6:
+ rewrite <- H;
+ try apply cmpl_ltle_add_one; auto.
+ all:
+ try rewrite <- H;
+ try apply cmpl_ltle_add_one; auto;
+ try rewrite <- cmpl_ltle_add_one; auto;
+ try rewrite ltu_12_wordsize;
+ try rewrite Int.add_commut, Int.add_zero_l in *;
+ try rewrite Int64.add_commut, Int64.add_zero_l in *;
+ try rewrite Int64.add_zero_l;
+ simpl; try rewrite lt_maxsgn_false_long;
+ try (rewrite <- H; trivial; fail);
+ simpl; trivial.
+Qed.
+
+Lemma simplify_ccompluimm_correct ge sp hst st c r n rs0 m m0 v: forall
+ (SMEM : forall (m : mem) (b : Values.block) (ofs : Z),
+ seval_smem ge sp (si_smem st) rs0 m0 = Some m ->
+ Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs)
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (OKv1 : seval_sval ge sp (si_sreg st r) rs0 m0 = Some v)
+ (OK2 : seval_smem ge sp (si_smem st) rs0 m0 = Some m),
+ seval_sval ge sp
+ (hsval_proj (expanse_condimm_int64u c (fsi_sreg_get hst r) n)) rs0 m0 =
+ Some (Val.of_optbool (Val.cmplu_bool (Mem.valid_pointer m) c v (Vlong n))).
+Proof.
+ intros.
+ assert (HMEM: Val.cmplu_bool (Mem.valid_pointer m) c v (Vlong n) =
+ Val.cmplu_bool (Mem.valid_pointer m0) c v (Vlong n)).
+ erewrite (cmplu_bool_valid_pointer_eq (Mem.valid_pointer m) (Mem.valid_pointer m0)); eauto.
+ unfold expanse_condimm_int64u, cond_int64u in *; destruct c;
+ intros; destruct (Int64.eq n Int64.zero) eqn:EQIMM; simpl;
+ unfold loadimm64, sltuimm64, opimm64, load_hilo64;
+ try erewrite !fsi_sreg_get_correct; eauto;
+ try rewrite OKv1;
+ unfold Val.cmplu, zero64.
+ (* Simplify make immediate and decompose subcases *)
+ all:
+ try (specialize make_immed64_sound with n;
+ destruct (make_immed64 n) eqn:EQMKI);
+ try destruct (Int64.eq lo Int64.zero) eqn:EQLO;
+ try erewrite fSop_correct; eauto; simpl;
+ try erewrite !fsi_sreg_get_correct; eauto;
+ try rewrite OKv1;
+ try rewrite OK2;
+ rewrite HMEM.
+ (* Ceq, Cne, Clt = itself *)
+ all: intros; try apply Int64.same_if_eq in EQIMM; subst; trivial.
+ (* Cle = xor (Clt) *)
+ all: try apply xor_neg_ltle_cmplu; trivial.
+ (* Others subcases with swap/negation *)
+ all:
+ unfold Val.cmplu, eval_may_undef, zero64, Val.addl;
+ try apply Int64.same_if_eq in EQLO; subst;
+ try rewrite Int64.add_commut, Int64.add_zero_l in *; trivial;
+ try rewrite Int64.add_zero_l;
+ try (rewrite <- xor_neg_ltle_cmplu; unfold Val.cmplu;
+ trivial; fail);
+ try (replace (Clt) with (swap_comparison Cgt) by auto;
+ rewrite Val.swap_cmplu_bool; trivial; fail);
+ try (replace (Clt) with (negate_comparison Cge) by auto;
+ rewrite Val.negate_cmplu_bool; rewrite xor_neg_optb; trivial; fail);
+ try rewrite optbool_mktotal; trivial.
+ all:
+ try destruct v; simpl; auto;
+ try destruct (Archi.ptr64); simpl;
+ try rewrite EQIMM;
+ try rewrite HMEM; trivial;
+ try destruct (Int64.ltu _ _);
+ try rewrite <- xor_neg_ltge_cmplu; unfold Val.cmplu;
+ try rewrite <- optbool_mktotal; trivial.
+Qed.
+
+Lemma simplify_ccompf_correct ge sp hst st c r r0 rs0 m0 v v0: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (OKv1 : seval_sval ge sp (si_sreg st r) rs0 m0 = Some v)
+ (OKv2 : seval_sval ge sp (si_sreg st r0) rs0 m0 = Some v0),
+ seval_sval ge sp
+ (hsval_proj
+ (expanse_cond_fp false cond_float c
+ (make_lhsv_cmp (is_inv_cmp_float c) (fsi_sreg_get hst r)
+ (fsi_sreg_get hst r0)))) rs0 m0 =
+ Some (Val.of_optbool (Val.cmpf_bool c v v0)).
+Proof.
+ intros.
+ unfold expanse_cond_fp in *; destruct c; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ rewrite OKv1, OKv2; trivial;
+ unfold Val.cmpf.
+ - apply xor_neg_eqne_cmpf.
+ - replace (Clt) with (swap_comparison Cgt) by auto;
+ rewrite swap_cmpf_bool; trivial.
+ - replace (Cle) with (swap_comparison Cge) by auto;
+ rewrite swap_cmpf_bool; trivial.
+Qed.
+
+Lemma simplify_cnotcompf_correct ge sp hst st c r r0 rs0 m0 v v0: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (OKv1 : seval_sval ge sp (si_sreg st r) rs0 m0 = Some v)
+ (OKv2 : seval_sval ge sp (si_sreg st r0) rs0 m0 = Some v0),
+ seval_sval ge sp
+ (hsval_proj
+ (expanse_cond_fp true cond_float c
+ (make_lhsv_cmp (is_inv_cmp_float c) (fsi_sreg_get hst r)
+ (fsi_sreg_get hst r0)))) rs0 m0 =
+ Some (Val.of_optbool (option_map negb (Val.cmpf_bool c v v0))).
+Proof.
+ intros.
+ unfold expanse_cond_fp in *; destruct c; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ rewrite OKv1, OKv2; trivial;
+ unfold Val.cmpf.
+ 1,3,4: apply xor_neg_optb'.
+ all: destruct v, v0; simpl; trivial.
+ rewrite Float.cmp_ne_eq; rewrite negb_involutive; trivial.
+ 1: replace (Clt) with (swap_comparison Cgt) by auto; rewrite <- Float.cmp_swap; simpl.
+ 2: replace (Cle) with (swap_comparison Cge) by auto; rewrite <- Float.cmp_swap; simpl.
+ all: destruct (Float.cmp _ _ _); trivial.
+Qed.
+
+Lemma simplify_ccompfs_correct ge sp hst st c r r0 rs0 m0 v v0: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (OKv1 : seval_sval ge sp (si_sreg st r) rs0 m0 = Some v)
+ (OKv2 : seval_sval ge sp (si_sreg st r0) rs0 m0 = Some v0),
+ seval_sval ge sp
+ (hsval_proj
+ (expanse_cond_fp false cond_single c
+ (make_lhsv_cmp (is_inv_cmp_float c) (fsi_sreg_get hst r)
+ (fsi_sreg_get hst r0)))) rs0 m0 =
+ Some (Val.of_optbool (Val.cmpfs_bool c v v0)).
+Proof.
+ intros.
+ unfold expanse_cond_fp in *; destruct c; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ rewrite OKv1, OKv2; trivial;
+ unfold Val.cmpfs.
+ - apply xor_neg_eqne_cmpfs.
+ - replace (Clt) with (swap_comparison Cgt) by auto;
+ rewrite swap_cmpfs_bool; trivial.
+ - replace (Cle) with (swap_comparison Cge) by auto;
+ rewrite swap_cmpfs_bool; trivial.
+Qed.
+
+Lemma simplify_cnotcompfs_correct ge sp hst st c r r0 rs0 m0 v v0: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (OKv1 : seval_sval ge sp (si_sreg st r) rs0 m0 = Some v)
+ (OKv2 : seval_sval ge sp (si_sreg st r0) rs0 m0 = Some v0),
+ seval_sval ge sp
+ (hsval_proj
+ (expanse_cond_fp true cond_single c
+ (make_lhsv_cmp (is_inv_cmp_float c) (fsi_sreg_get hst r)
+ (fsi_sreg_get hst r0)))) rs0 m0 =
+ Some (Val.of_optbool (option_map negb (Val.cmpfs_bool c v v0))).
+Proof.
+ intros.
+ unfold expanse_cond_fp in *; destruct c; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ rewrite OKv1, OKv2; trivial;
+ unfold Val.cmpfs.
+ 1,3,4: apply xor_neg_optb'.
+ all: destruct v, v0; simpl; trivial.
+ rewrite Float32.cmp_ne_eq; rewrite negb_involutive; trivial.
+ 1: replace (Clt) with (swap_comparison Cgt) by auto; rewrite <- Float32.cmp_swap; simpl.
+ 2: replace (Cle) with (swap_comparison Cge) by auto; rewrite <- Float32.cmp_swap; simpl.
+ all: destruct (Float32.cmp _ _ _); trivial.
+Qed.
+
+Lemma simplify_floatconst_correct ge sp rs0 m0 args m n fsv lr st: forall
+ (H : match lr with
+ | nil =>
+ Some
+ (fSop Ofloat_of_bits
+ (make_lhsv_single (loadimm64 (Float.to_bits n))))
+ | _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Ofloatconst n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ unfold loadimm64, load_hilo64; simpl;
+ specialize make_immed64_sound with (Float.to_bits n);
+ destruct (make_immed64 (Float.to_bits n)) eqn:EQMKI; intros;
+ try destruct (Int64.eq lo Int64.zero) eqn:EQLO;
+ simpl.
+ - try rewrite Int64.add_commut, Int64.add_zero_l; inv H;
+ try rewrite Float.of_to_bits; trivial.
+ - apply Int64.same_if_eq in EQLO; subst.
+ try rewrite Int64.add_commut, Int64.add_zero_l in H.
+ rewrite <- H; try rewrite Float.of_to_bits; trivial.
+ - rewrite <- H; try rewrite Float.of_to_bits; trivial.
+ - rewrite <- H; try rewrite Float.of_to_bits; trivial.
+Qed.
+
+Lemma simplify_singleconst_correct ge sp rs0 m0 args m n fsv lr st: forall
+ (H : match lr with
+ | nil =>
+ Some
+ (fSop Osingle_of_bits
+ (make_lhsv_single (loadimm32 (Float32.to_bits n))))
+ | _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Osingleconst n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ unfold loadimm32, load_hilo32; simpl;
+ specialize make_immed32_sound with (Float32.to_bits n);
+ destruct (make_immed32 (Float32.to_bits n)) eqn:EQMKI; intros;
+ try destruct (Int.eq lo Int.zero) eqn:EQLO;
+ simpl.
+ { try rewrite Int.add_commut, Int.add_zero_l; inv H;
+ try rewrite Float32.of_to_bits; trivial. }
+ all:
+ try apply Int.same_if_eq in EQLO; subst;
+ try rewrite Int.add_commut, Int.add_zero_l in H; simpl;
+ rewrite ltu_12_wordsize; simpl; try rewrite <- H;
+ try rewrite Float32.of_to_bits; trivial.
+Qed.
+
+Lemma simplify_addimm_correct ge sp rs0 m0 lr n hst fsv st args m: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (H : match lr with
+ | nil => None
+ | a1 :: nil => Some (addimm32 (fsi_sreg_get hst a1) n None)
+ | a1 :: _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Oaddimm n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ unfold addimm32, opimm32, load_hilo32, make_lhsv_cmp; simpl;
+ specialize make_immed32_sound with (n);
+ destruct (make_immed32 (n)) eqn:EQMKI; intros; simpl;
+ try destruct (Int.eq lo Int.zero) eqn:EQLO; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ destruct (seval_sval ge sp (si_sreg st p) rs0 m0) eqn:OKv1; try congruence; inv H1.
+ fold (Val.add (Vint imm) v); rewrite Val.add_commut; trivial.
+ all:
+ try apply Int.same_if_eq in EQLO; subst;
+ try rewrite Int.add_commut, Int.add_zero_l;
+ try rewrite ltu_12_wordsize; trivial.
+Qed.
+
+Lemma simplify_addlimm_correct ge sp rs0 m0 lr n hst fsv st args m: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (H : match lr with
+ | nil => None
+ | a1 :: nil => Some (addimm64 (fsi_sreg_get hst a1) n None)
+ | a1 :: _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Oaddlimm n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ unfold addimm64, opimm64, load_hilo64, make_lhsv_cmp; simpl;
+ specialize make_immed64_sound with (n);
+ destruct (make_immed64 (n)) eqn:EQMKI; intros; simpl;
+ try destruct (Int64.eq lo Int64.zero) eqn:EQLO; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ destruct (seval_sval ge sp (si_sreg st p) rs0 m0) eqn:OKv1; try congruence; inv H1.
+ fold (Val.addl (Vlong imm) v); rewrite Val.addl_commut; trivial.
+ all:
+ try apply Int64.same_if_eq in EQLO; subst;
+ try rewrite Int64.add_commut, Int64.add_zero_l;
+ try rewrite Int64.add_commut;
+ try rewrite ltu_12_wordsize; trivial.
+Qed.
+
+Lemma simplify_andimm_correct ge sp rs0 m0 lr n hst fsv st args m: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (H : match lr with
+ | nil => None
+ | a1 :: nil => Some (andimm32 (fsi_sreg_get hst a1) n)
+ | a1 :: _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Oandimm n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ unfold andimm32, opimm32, load_hilo32, make_lhsv_cmp; simpl;
+ specialize make_immed32_sound with (n);
+ destruct (make_immed32 (n)) eqn:EQMKI; intros; simpl;
+ try destruct (Int.eq lo Int.zero) eqn:EQLO; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ destruct (seval_sval ge sp (si_sreg st p) rs0 m0) eqn:OKv1; try congruence; inv H1.
+ fold (Val.and (Vint imm) v); rewrite Val.and_commut; trivial.
+ all:
+ try apply Int.same_if_eq in EQLO; subst;
+ try rewrite Int.add_commut, Int.add_zero_l;
+ try rewrite ltu_12_wordsize; trivial.
+Qed.
+
+Lemma simplify_andlimm_correct ge sp rs0 m0 lr n hst fsv st args m: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (H : match lr with
+ | nil => None
+ | a1 :: nil => Some (andimm64 (fsi_sreg_get hst a1) n)
+ | a1 :: _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Oandlimm n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ unfold andimm64, opimm64, load_hilo64, make_lhsv_cmp; simpl;
+ specialize make_immed64_sound with (n);
+ destruct (make_immed64 (n)) eqn:EQMKI; intros; simpl;
+ try destruct (Int64.eq lo Int64.zero) eqn:EQLO; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ destruct (seval_sval ge sp (si_sreg st p) rs0 m0) eqn:OKv1; try congruence; inv H1.
+ fold (Val.andl (Vlong imm) v); rewrite Val.andl_commut; trivial.
+ all:
+ try apply Int64.same_if_eq in EQLO; subst;
+ try rewrite Int64.add_commut, Int64.add_zero_l;
+ try rewrite Int64.add_commut;
+ try rewrite ltu_12_wordsize; trivial.
+Qed.
+
+Lemma simplify_orimm_correct ge sp rs0 m0 lr n hst fsv st args m: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (H : match lr with
+ | nil => None
+ | a1 :: nil => Some (orimm32 (fsi_sreg_get hst a1) n)
+ | a1 :: _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Oorimm n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ unfold orimm32, opimm32, load_hilo32, make_lhsv_cmp; simpl;
+ specialize make_immed32_sound with (n);
+ destruct (make_immed32 (n)) eqn:EQMKI; intros; simpl;
+ try destruct (Int.eq lo Int.zero) eqn:EQLO; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ destruct (seval_sval ge sp (si_sreg st p) rs0 m0) eqn:OKv1; try congruence; inv H1.
+ fold (Val.or (Vint imm) v); rewrite Val.or_commut; trivial.
+ all:
+ try apply Int.same_if_eq in EQLO; subst;
+ try rewrite Int.add_commut, Int.add_zero_l;
+ try rewrite ltu_12_wordsize; trivial.
+Qed.
+
+Lemma simplify_orlimm_correct ge sp rs0 m0 lr n hst fsv st args m: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (H : match lr with
+ | nil => None
+ | a1 :: nil => Some (orimm64 (fsi_sreg_get hst a1) n)
+ | a1 :: _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Oorlimm n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ unfold orimm64, opimm64, load_hilo64, make_lhsv_cmp; simpl;
+ specialize make_immed64_sound with (n);
+ destruct (make_immed64 (n)) eqn:EQMKI; intros; simpl;
+ try destruct (Int64.eq lo Int64.zero) eqn:EQLO; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ destruct (seval_sval ge sp (si_sreg st p) rs0 m0) eqn:OKv1; try congruence; inv H1.
+ fold (Val.orl (Vlong imm) v); rewrite Val.orl_commut; trivial.
+ all:
+ try apply Int64.same_if_eq in EQLO; subst;
+ try rewrite Int64.add_commut, Int64.add_zero_l;
+ try rewrite Int64.add_commut;
+ try rewrite ltu_12_wordsize; trivial.
+Qed.
+
+Lemma simplify_xorimm_correct ge sp rs0 m0 lr n hst fsv st args m: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (H : match lr with
+ | nil => None
+ | a1 :: nil => Some (xorimm32 (fsi_sreg_get hst a1) n)
+ | a1 :: _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Oxorimm n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ unfold xorimm32, opimm32, load_hilo32, make_lhsv_cmp; simpl;
+ specialize make_immed32_sound with (n);
+ destruct (make_immed32 (n)) eqn:EQMKI; intros; simpl;
+ try destruct (Int.eq lo Int.zero) eqn:EQLO; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ destruct (seval_sval ge sp (si_sreg st p) rs0 m0) eqn:OKv1; try congruence; inv H1.
+ fold (Val.xor (Vint imm) v); rewrite Val.xor_commut; trivial.
+ all:
+ try apply Int.same_if_eq in EQLO; subst;
+ try rewrite Int.add_commut, Int.add_zero_l;
+ try rewrite ltu_12_wordsize; trivial.
+Qed.
+
+Lemma simplify_xorlimm_correct ge sp rs0 m0 lr n hst fsv st args m: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (H : match lr with
+ | nil => None
+ | a1 :: nil => Some (xorimm64 (fsi_sreg_get hst a1) n)
+ | a1 :: _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Oxorlimm n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ unfold xorimm64, opimm64, load_hilo64, make_lhsv_cmp; simpl;
+ specialize make_immed64_sound with (n);
+ destruct (make_immed64 (n)) eqn:EQMKI; intros; simpl;
+ try destruct (Int64.eq lo Int64.zero) eqn:EQLO; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ destruct (seval_sval ge sp (si_sreg st p) rs0 m0) eqn:OKv1; try congruence; inv H1.
+ fold (Val.xorl (Vlong imm) v); rewrite Val.xorl_commut; trivial.
+ all:
+ try apply Int64.same_if_eq in EQLO; subst;
+ try rewrite Int64.add_commut, Int64.add_zero_l;
+ try rewrite Int64.add_commut;
+ try rewrite ltu_12_wordsize; trivial.
+Qed.
+
+Lemma simplify_intconst_correct ge sp rs0 m0 args m n fsv lr st: forall
+ (H : match lr with
+ | nil => Some (loadimm32 n)
+ | _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Ointconst n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ unfold loadimm32, load_hilo32, make_lhsv_single; simpl;
+ specialize make_immed32_sound with (n);
+ destruct (make_immed32 (n)) eqn:EQMKI; intros; simpl;
+ try destruct (Int.eq lo Int.zero) eqn:EQLO; simpl;
+ try apply Int.same_if_eq in EQLO; subst;
+ try rewrite Int.add_commut, Int.add_zero_l;
+ try rewrite ltu_12_wordsize; try rewrite H; trivial.
+Qed.
+
+Lemma simplify_longconst_correct ge sp rs0 m0 args m n fsv lr st: forall
+ (H : match lr with
+ | nil => Some (loadimm64 n)
+ | _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Olongconst n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ unfold loadimm64, load_hilo64, make_lhsv_single; simpl;
+ specialize make_immed64_sound with (n);
+ destruct (make_immed64 (n)) eqn:EQMKI; intros; simpl;
+ try destruct (Int64.eq lo Int64.zero) eqn:EQLO; simpl;
+ try apply Int64.same_if_eq in EQLO; subst;
+ try rewrite Int64.add_commut, Int64.add_zero_l;
+ try rewrite Int64.add_commut;
+ try rewrite ltu_12_wordsize; try rewrite H; trivial.
+Qed.
+
+Lemma simplify_cast8signed_correct ge sp rs0 m0 lr hst fsv st args m: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (H : match lr with
+ | nil => None
+ | a1 :: nil =>
+ Some
+ (fSop (Oshrimm (Int.repr 24))
+ (make_lhsv_single
+ (fSop (Oshlimm (Int.repr 24))
+ (make_lhsv_single (fsi_sreg_get hst a1)))))
+ | a1 :: _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp Ocast8signed args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ destruct (seval_sval ge sp (si_sreg st p) rs0 m0) eqn:OKv1; try congruence; inv H1.
+ unfold Val.shr, Val.shl, Val.sign_ext;
+ destruct v; simpl; auto.
+ assert (A: Int.ltu (Int.repr 24) Int.iwordsize = true) by auto.
+ rewrite A. rewrite Int.sign_ext_shr_shl; simpl; trivial. cbn; lia.
+Qed.
+
+Lemma simplify_cast16signed_correct ge sp rs0 m0 lr hst fsv st args m: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (H : match lr with
+ | nil => None
+ | a1 :: nil =>
+ Some
+ (fSop (Oshrimm (Int.repr 16))
+ (make_lhsv_single
+ (fSop (Oshlimm (Int.repr 16))
+ (make_lhsv_single (fsi_sreg_get hst a1)))))
+ | a1 :: _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp Ocast16signed args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ destruct (seval_sval ge sp (si_sreg st p) rs0 m0) eqn:OKv1; try congruence; inv H1.
+ unfold Val.shr, Val.shl, Val.sign_ext;
+ destruct v; simpl; auto.
+ assert (A: Int.ltu (Int.repr 16) Int.iwordsize = true) by auto.
+ rewrite A. rewrite Int.sign_ext_shr_shl; simpl; trivial. cbn; lia.
+Qed.
+
+Lemma simplify_shrximm_correct ge sp rs0 m0 lr hst fsv st args m n: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (H : match lr with
+ | nil => None
+ | a1 :: nil =>
+ if Int.eq n Int.zero
+ then
+ Some
+ (fSop (OEmayundef (MUshrx n))
+ (make_lhsv_cmp false (fsi_sreg_get hst a1)
+ (fsi_sreg_get hst a1)))
+ else
+ if Int.eq n Int.one
+ then
+ Some
+ (fSop (OEmayundef (MUshrx n))
+ (make_lhsv_cmp false
+ (fSop (Oshrimm Int.one)
+ (make_lhsv_single
+ (fSop Oadd
+ (make_lhsv_cmp false (fsi_sreg_get hst a1)
+ (fSop (Oshruimm (Int.repr 31))
+ (make_lhsv_single (fsi_sreg_get hst a1)))))))
+ (fSop (Oshrimm Int.one)
+ (make_lhsv_single
+ (fSop Oadd
+ (make_lhsv_cmp false (fsi_sreg_get hst a1)
+ (fSop (Oshruimm (Int.repr 31))
+ (make_lhsv_single (fsi_sreg_get hst a1)))))))))
+ else
+ Some
+ (fSop (OEmayundef (MUshrx n))
+ (make_lhsv_cmp false
+ (fSop (Oshrimm n)
+ (make_lhsv_single
+ (fSop Oadd
+ (make_lhsv_cmp false (fsi_sreg_get hst a1)
+ (fSop (Oshruimm (Int.sub Int.iwordsize n))
+ (make_lhsv_single
+ (fSop (Oshrimm (Int.repr 31))
+ (make_lhsv_single
+ (fsi_sreg_get hst a1)))))))))
+ (fSop (Oshrimm n)
+ (make_lhsv_single
+ (fSop Oadd
+ (make_lhsv_cmp false (fsi_sreg_get hst a1)
+ (fSop (Oshruimm (Int.sub Int.iwordsize n))
+ (make_lhsv_single
+ (fSop (Oshrimm (Int.repr 31))
+ (make_lhsv_single
+ (fsi_sreg_get hst a1)))))))))))
+ | a1 :: _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Oshrximm n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence).
+ assert (A: Int.ltu Int.zero (Int.repr 31) = true) by auto.
+ assert (B: Int.ltu (Int.repr 31) Int.iwordsize = true) by auto.
+ assert (C: Int.ltu Int.one Int.iwordsize = true) by auto.
+ destruct (Int.eq n Int.zero) eqn:EQ0;
+ destruct (Int.eq n Int.one) eqn:EQ1.
+ { apply Int.same_if_eq in EQ0.
+ apply Int.same_if_eq in EQ1; subst. discriminate. }
+ all:
+ simpl in OK1; inv OK1; inv H; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ destruct (seval_sval ge sp (si_sreg st p) rs0 m0) eqn:OKv1; try congruence; inv H1;
+ destruct (Val.shrx v (Vint n)) eqn:TOTAL; cbn;
+ unfold eval_may_undef.
+ 2,4,6:
+ unfold Val.shrx in TOTAL;
+ destruct v; simpl in TOTAL; simpl; try congruence;
+ try rewrite B; simpl; try rewrite C; simpl;
+ try destruct (Val.shr _ _);
+ destruct (Int.ltu n (Int.repr 31)); try congruence.
+ - destruct v; simpl in TOTAL; try congruence;
+ apply Int.same_if_eq in EQ0; subst;
+ rewrite A, Int.shrx_zero in TOTAL;
+ [auto | cbn; lia].
+ - apply Int.same_if_eq in EQ1; subst;
+ unfold Val.shr, Val.shru, Val.shrx, Val.add; simpl;
+ destruct v; simpl in *; try discriminate; trivial.
+ rewrite B, C.
+ rewrite Int.shrx1_shr in TOTAL; auto.
+ - exploit Val.shrx_shr_2; eauto. rewrite EQ0.
+ intros; subst.
+ destruct v; simpl in *; try discriminate; trivial.
+ rewrite B in *.
+ destruct Int.ltu eqn:EQN0 in TOTAL; try discriminate.
+ simpl in *.
+ destruct Int.ltu eqn:EQN1 in TOTAL; try discriminate.
+ replace Int.iwordsize with (Int.repr 32) in * by auto.
+ rewrite !EQN1. simpl in *.
+ destruct Int.ltu eqn:EQN2 in TOTAL; try discriminate.
+ rewrite !EQN2. rewrite EQN0.
+ reflexivity.
+Qed.
+
+Lemma simplify_shrxlimm_correct ge sp rs0 m0 lr hst fsv st args m n: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (H : match lr with
+ | nil => None
+ | a1 :: nil =>
+ if Int.eq n Int.zero
+ then
+ Some
+ (fSop (OEmayundef (MUshrxl n))
+ (make_lhsv_cmp false (fsi_sreg_get hst a1)
+ (fsi_sreg_get hst a1)))
+ else
+ if Int.eq n Int.one
+ then
+ Some
+ (fSop (OEmayundef (MUshrxl n))
+ (make_lhsv_cmp false
+ (fSop (Oshrlimm Int.one)
+ (make_lhsv_single
+ (fSop Oaddl
+ (make_lhsv_cmp false (fsi_sreg_get hst a1)
+ (fSop (Oshrluimm (Int.repr 63))
+ (make_lhsv_single (fsi_sreg_get hst a1)))))))
+ (fSop (Oshrlimm Int.one)
+ (make_lhsv_single
+ (fSop Oaddl
+ (make_lhsv_cmp false (fsi_sreg_get hst a1)
+ (fSop (Oshrluimm (Int.repr 63))
+ (make_lhsv_single (fsi_sreg_get hst a1)))))))))
+ else
+ Some
+ (fSop (OEmayundef (MUshrxl n))
+ (make_lhsv_cmp false
+ (fSop (Oshrlimm n)
+ (make_lhsv_single
+ (fSop Oaddl
+ (make_lhsv_cmp false (fsi_sreg_get hst a1)
+ (fSop (Oshrluimm (Int.sub Int64.iwordsize' n))
+ (make_lhsv_single
+ (fSop (Oshrlimm (Int.repr 63))
+ (make_lhsv_single
+ (fsi_sreg_get hst a1)))))))))
+ (fSop (Oshrlimm n)
+ (make_lhsv_single
+ (fSop Oaddl
+ (make_lhsv_cmp false (fsi_sreg_get hst a1)
+ (fSop (Oshrluimm (Int.sub Int64.iwordsize' n))
+ (make_lhsv_single
+ (fSop (Oshrlimm (Int.repr 63))
+ (make_lhsv_single
+ (fsi_sreg_get hst a1)))))))))))
+ | a1 :: _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp (Oshrxlimm n) args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence).
+ assert (A: Int.ltu Int.zero (Int.repr 63) = true) by auto.
+ assert (B: Int.ltu (Int.repr 63) Int64.iwordsize' = true) by auto.
+ assert (C: Int.ltu Int.one Int64.iwordsize' = true) by auto.
+ destruct (Int.eq n Int.zero) eqn:EQ0;
+ destruct (Int.eq n Int.one) eqn:EQ1.
+ { apply Int.same_if_eq in EQ0.
+ apply Int.same_if_eq in EQ1; subst. discriminate. }
+ all:
+ simpl in OK1; inv OK1; inv H; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ destruct (seval_sval ge sp (si_sreg st p) rs0 m0) eqn:OKv1; try congruence; inv H1;
+ destruct (Val.shrxl v (Vint n)) eqn:TOTAL; cbn;
+ unfold eval_may_undef.
+ 2,4,6:
+ unfold Val.shrxl in TOTAL;
+ destruct v; simpl in TOTAL; simpl; try congruence;
+ try rewrite B; simpl; try rewrite C; simpl;
+ try destruct (Val.shrl _ _);
+ destruct (Int.ltu n (Int.repr 63)); try congruence.
+ - destruct v; simpl in TOTAL; try congruence;
+ apply Int.same_if_eq in EQ0; subst;
+ rewrite A, Int64.shrx'_zero in *.
+ assumption.
+ - apply Int.same_if_eq in EQ1; subst;
+ unfold Val.shrl, Val.shrlu, Val.shrxl, Val.addl; simpl;
+ destruct v; simpl in *; try discriminate; trivial.
+ rewrite B, C.
+ rewrite Int64.shrx'1_shr' in TOTAL; auto.
+ - exploit Val.shrxl_shrl_2; eauto. rewrite EQ0.
+ intros; subst.
+ destruct v; simpl in *; try discriminate; trivial.
+ rewrite B in *.
+ destruct Int.ltu eqn:EQN0 in TOTAL; try discriminate.
+ simpl in *.
+ destruct Int.ltu eqn:EQN1 in TOTAL; try discriminate.
+ replace Int64.iwordsize' with (Int.repr 64) in * by auto.
+ rewrite !EQN1. simpl in *.
+ destruct Int.ltu eqn:EQN2 in TOTAL; try discriminate.
+ rewrite !EQN2. rewrite EQN0.
+ reflexivity.
+Qed.
+
+Lemma simplify_cast32unsigned_correct ge sp rs0 m0 lr hst fsv st args m: forall
+ (SREG: forall r: positive,
+ hsi_sreg_eval ge sp hst r rs0 m0 =
+ seval_sval ge sp (si_sreg st r) rs0 m0)
+ (H : match lr with
+ | nil => None
+ | a1 :: nil =>
+ Some
+ (fSop (Oshrluimm (Int.repr 32))
+ (make_lhsv_single
+ (fSop (Oshllimm (Int.repr 32))
+ (make_lhsv_single
+ (fSop Ocast32signed
+ (make_lhsv_single (fsi_sreg_get hst a1)))))))
+ | a1 :: _ :: _ => None
+ end = Some fsv)
+ (OK1 : seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 =
+ eval_operation ge sp Ocast32unsigned args m.
+Proof.
+ intros.
+ repeat (destruct lr; simpl; try congruence);
+ simpl in OK1; inv OK1; inv H; simpl;
+ erewrite !fsi_sreg_get_correct; eauto;
+ destruct (seval_sval ge sp (si_sreg st p) rs0 m0) eqn:OKv1; try congruence; inv H1.
+ unfold Val.shrlu, Val.shll, Val.longofint, Val.longofintu.
+ destruct v; simpl; auto.
+ assert (A: Int.ltu (Int.repr 32) Int64.iwordsize' = true) by auto.
+ rewrite A. rewrite Int64.shru'_shl'; auto.
+ replace (Int.ltu (Int.repr 32) (Int.repr 32)) with (false) by auto.
+ rewrite cast32unsigned_from_cast32signed.
+ replace Int64.zwordsize with 64 by auto.
+ rewrite Int.unsigned_repr; cbn; try lia.
+ replace (Int.sub (Int.repr 32) (Int.repr 32)) with (Int.zero) by auto.
+ rewrite Int64.shru'_zero. reflexivity.
+Qed.
+
+(** * Main proof of simplification *)
+
+Lemma target_op_simplify_correct op lr hst fsv ge sp rs0 m0 st args m: forall
+ (H: target_op_simplify op lr hst = Some fsv)
+ (REF: hsilocal_refines ge sp rs0 m0 hst st)
+ (OK0: hsok_local ge sp rs0 m0 hst)
+ (OK1: seval_list_sval ge sp (list_sval_inj (map (si_sreg st) lr)) rs0 m0 = Some args)
+ (OK2: seval_smem ge sp (si_smem st) rs0 m0 = Some m),
+ seval_sval ge sp (hsval_proj fsv) rs0 m0 = eval_operation ge sp op args m.
+Proof.
+ unfold target_op_simplify; simpl.
+ intros H (LREF & SREF & SREG & SMEM) ? ? ?.
+ destruct op; try congruence.
+ eapply simplify_intconst_correct; eauto.
+ eapply simplify_longconst_correct; eauto.
+ eapply simplify_floatconst_correct; eauto.
+ eapply simplify_singleconst_correct; eauto.
+ eapply simplify_cast8signed_correct; eauto.
+ eapply simplify_cast16signed_correct; eauto.
+ eapply simplify_addimm_correct; eauto.
+ eapply simplify_andimm_correct; eauto.
+ eapply simplify_orimm_correct; eauto.
+ eapply simplify_xorimm_correct; eauto.
+ eapply simplify_shrximm_correct; eauto.
+ eapply simplify_cast32unsigned_correct; eauto.
+ eapply simplify_addlimm_correct; eauto.
+ eapply simplify_andlimm_correct; eauto.
+ eapply simplify_orlimm_correct; eauto.
+ eapply simplify_xorlimm_correct; eauto.
+ eapply simplify_shrxlimm_correct; eauto.
+ (* Ocmp expansions *)
+ destruct cond; repeat (destruct lr; simpl; try congruence);
+ simpl in OK1;
+ try (destruct (seval_sval ge sp (si_sreg st r) rs0 m0) eqn:OKv1; try congruence);
+ try (destruct (seval_sval ge sp (si_sreg st r0) rs0 m0) eqn:OKv2; try congruence);
+ inv H; inv OK1.
+ - eapply simplify_ccomp_correct; eauto.
+ - eapply simplify_ccompu_correct; eauto.
+ - eapply simplify_ccompimm_correct; eauto.
+ - eapply simplify_ccompuimm_correct; eauto.
+ - eapply simplify_ccompl_correct; eauto.
+ - eapply simplify_ccomplu_correct; eauto.
+ - eapply simplify_ccomplimm_correct; eauto.
+ - eapply simplify_ccompluimm_correct; eauto.
+ - eapply simplify_ccompf_correct; eauto.
+ - eapply simplify_cnotcompf_correct; eauto.
+ - eapply simplify_ccompfs_correct; eauto.
+ - eapply simplify_cnotcompfs_correct; eauto.
+Qed.
+
+Lemma target_cbranch_expanse_correct hst c l ge sp rs0 m0 st c' l': forall
+ (TARGET: target_cbranch_expanse hst c l = Some (c', l'))
+ (LREF : hsilocal_refines ge sp rs0 m0 hst st)
+ (OK: hsok_local ge sp rs0 m0 hst),
+ seval_condition ge sp c' (hsval_list_proj l') (si_smem st) rs0 m0 =
+ seval_condition ge sp c (list_sval_inj (map (si_sreg st) l)) (si_smem st) rs0 m0.
+Proof.
+ unfold target_cbranch_expanse, seval_condition; simpl.
+ intros H (LREF & SREF & SREG & SMEM) ?.
+ destruct c; try congruence;
+ repeat (destruct l; simpl in H; try congruence).
+ 1,2,5,6:
+ destruct c; inv H; simpl;
+ try erewrite !fsi_sreg_get_correct; eauto;
+ try (destruct (seval_smem ge sp (si_smem st) rs0 m0) eqn:OKmem; try congruence);
+ try (destruct (seval_sval ge sp (si_sreg st r) rs0 m0) eqn:OKv1; try congruence);
+ try (destruct (seval_sval ge sp (si_sreg st r0) rs0 m0) eqn:OKv2; try congruence);
+ try replace (Cle) with (swap_comparison Cge) by auto;
+ try replace (Clt) with (swap_comparison Cgt) by auto;
+ try rewrite Val.swap_cmp_bool; trivial;
+ try rewrite Val.swap_cmpu_bool; trivial;
+ try rewrite Val.swap_cmpl_bool; trivial;
+ try rewrite Val.swap_cmplu_bool; trivial.
+ 1,2,3,4:
+ try destruct (Int.eq n Int.zero) eqn: EQIMM;
+ try apply Int.same_if_eq in EQIMM;
+ try destruct (Int64.eq n Int64.zero) eqn: EQIMM;
+ try apply Int64.same_if_eq in EQIMM;
+ destruct c; inv H; simpl;
+ try erewrite !fsi_sreg_get_correct; eauto;
+ try (destruct (seval_smem ge sp (si_smem st) rs0 m0) eqn:OKmem; try congruence);
+ try (destruct (seval_sval ge sp (si_sreg st r) rs0 m0) eqn:OKv1; try congruence);
+ try (destruct (seval_sval ge sp (si_sreg st r0) rs0 m0) eqn:OKv2; try congruence);
+ unfold loadimm32, load_hilo32, Val.cmp, Val.cmpu, zero32;
+ unfold loadimm64, load_hilo64, Val.cmpl, Val.cmplu, zero64;
+ intros; try (specialize make_immed32_sound with (n);
+ destruct (make_immed32 n) eqn:EQMKI); intros; simpl;
+ intros; try (specialize make_immed64_sound with (n);
+ destruct (make_immed64 n) eqn:EQMKI); intros; simpl;
+ try rewrite EQLO; simpl;
+ try destruct (Int.eq lo Int.zero) eqn:EQLO;
+ try destruct (Int64.eq lo Int64.zero) eqn:EQLO;
+ try apply Int.same_if_eq in EQLO; simpl; trivial;
+ try apply Int64.same_if_eq in EQLO; simpl; trivial;
+ unfold eval_may_undef;
+ try erewrite !fsi_sreg_get_correct; eauto;
+ try rewrite OKv1; simpl; trivial;
+ try destruct v; try rewrite H;
+ try rewrite ltu_12_wordsize; try rewrite EQLO;
+ try rewrite Int.add_commut, Int.add_zero_l;
+ try rewrite Int64.add_commut, Int64.add_zero_l;
+ try rewrite Int64.add_commut;
+ try rewrite Int.add_zero_l; try rewrite Int64.add_zero_l;
+ auto; simpl;
+ try rewrite H in EQIMM;
+ try rewrite EQLO in EQIMM;
+ try rewrite Int.add_commut, Int.add_zero_l in EQIMM;
+ try rewrite Int64.add_commut, Int64.add_zero_l in EQIMM;
+ try rewrite EQIMM; simpl;
+ try destruct (Archi.ptr64); trivial.
+
+ 1,2,3,4:
+ destruct c; inv H; simpl;
+ try erewrite !fsi_sreg_get_correct; eauto;
+ try (destruct (seval_smem ge sp (si_smem st) rs0 m0) eqn:OKmem; try congruence);
+ try (destruct (seval_sval ge sp (si_sreg st r) rs0 m0) eqn:OKv1; try congruence);
+ try (destruct (seval_sval ge sp (si_sreg st r0) rs0 m0) eqn:OKv2; try congruence);
+ unfold zero32, zero64, Val.cmpf, Val.cmpfs;
+ destruct v, v0; simpl; trivial;
+ try rewrite Float.cmp_ne_eq;
+ try rewrite Float32.cmp_ne_eq;
+ try rewrite <- Float.cmp_swap; simpl;
+ try rewrite <- Float32.cmp_swap; simpl;
+ try destruct (Float.cmp _ _); simpl;
+ try destruct (Float32.cmp _ _); simpl;
+ try rewrite Int.eq_true; simpl;
+ try rewrite Int.eq_false; try apply Int.one_not_zero;
+ simpl; trivial.
+Qed.
+Global Opaque target_op_simplify.
+Global Opaque target_cbranch_expanse.
diff --git a/riscV/SelectLong.vp b/riscV/SelectLong.vp
index b3e07bf5..0ccc4725 100644
--- a/riscV/SelectLong.vp
+++ b/riscV/SelectLong.vp
@@ -21,7 +21,7 @@ Require Import Coqlib.
Require Import Compopts.
Require Import AST Integers Floats.
Require Import Op CminorSel.
-Require Import SelectOp SplitLong.
+Require Import OpHelpers SelectOp SplitLong.
Local Open Scope cminorsel_scope.
Local Open Scope string_scope.
diff --git a/riscV/SelectLongproof.v b/riscV/SelectLongproof.v
index 3794e050..0fc578bf 100644
--- a/riscV/SelectLongproof.v
+++ b/riscV/SelectLongproof.v
@@ -21,6 +21,7 @@ Require Import String Coqlib Maps Integers Floats Errors.
Require Archi.
Require Import AST Values Memory Globalenvs Events.
Require Import Cminor Op CminorSel.
+Require Import OpHelpers OpHelpersproof.
Require Import SelectOp SelectOpproof SplitLong SplitLongproof.
Require Import SelectLong.
@@ -454,6 +455,10 @@ Proof.
unfold divls_base; red; intros. destruct Archi.splitlong eqn:SL.
eapply SplitLongproof.eval_divls_base; eauto.
TrivialExists.
+ cbn.
+ rewrite H1.
+ cbn.
+ trivial.
Qed.
Theorem eval_modls_base: partial_binary_constructor_sound modls_base Val.modls.
@@ -461,6 +466,10 @@ Proof.
unfold modls_base; red; intros. destruct Archi.splitlong eqn:SL.
eapply SplitLongproof.eval_modls_base; eauto.
TrivialExists.
+ cbn.
+ rewrite H1.
+ cbn.
+ trivial.
Qed.
Theorem eval_divlu_base: partial_binary_constructor_sound divlu_base Val.divlu.
@@ -468,6 +477,10 @@ Proof.
unfold divlu_base; red; intros. destruct Archi.splitlong eqn:SL.
eapply SplitLongproof.eval_divlu_base; eauto.
TrivialExists.
+ cbn.
+ rewrite H1.
+ cbn.
+ trivial.
Qed.
Theorem eval_modlu_base: partial_binary_constructor_sound modlu_base Val.modlu.
@@ -475,6 +488,10 @@ Proof.
unfold modlu_base; red; intros. destruct Archi.splitlong eqn:SL.
eapply SplitLongproof.eval_modlu_base; eauto.
TrivialExists.
+ cbn.
+ rewrite H1.
+ cbn.
+ trivial.
Qed.
Theorem eval_shrxlimm:
@@ -489,33 +506,9 @@ Proof.
- subst n. destruct x; simpl in H0; inv H0. econstructor; split; eauto.
change (Int.ltu Int.zero (Int.repr 63)) with true. simpl. rewrite Int64.shrx'_zero; auto.
- TrivialExists.
-(*
- intros. unfold shrxlimm. destruct Archi.splitlong eqn:SL.
-+ eapply SplitLongproof.eval_shrxlimm; eauto using Archi.splitlong_ptr32.
-+ destruct x; simpl in H0; try discriminate.
- destruct (Int.ltu n (Int.repr 63)) eqn:LTU; inv H0.
- predSpec Int.eq Int.eq_spec n Int.zero.
- - subst n. exists (Vlong i); split; auto. rewrite Int64.shrx'_zero. auto.
- - assert (NZ: Int.unsigned n <> 0).
- { intro EQ; elim H0. rewrite <- (Int.repr_unsigned n). rewrite EQ; auto. }
- assert (LT: 0 <= Int.unsigned n < 63) by (apply Int.ltu_inv in LTU; assumption).
- assert (LTU2: Int.ltu (Int.sub Int64.iwordsize' n) Int64.iwordsize' = true).
- { unfold Int.ltu; apply zlt_true.
- unfold Int.sub. change (Int.unsigned Int64.iwordsize') with 64.
- rewrite Int.unsigned_repr. lia.
- assert (64 < Int.max_unsigned) by reflexivity. lia. }
- assert (X: eval_expr ge sp e m le
- (Eop (Oshrlimm (Int.repr (Int64.zwordsize - 1))) (a ::: Enil))
- (Vlong (Int64.shr' i (Int.repr (Int64.zwordsize - 1))))).
- { EvalOp. }
- assert (Y: eval_expr ge sp e m le (shrxlimm_inner a n)
- (Vlong (Int64.shru' (Int64.shr' i (Int.repr (Int64.zwordsize - 1))) (Int.sub Int64.iwordsize' n)))).
- { EvalOp. simpl. rewrite LTU2. auto. }
- TrivialExists.
- constructor. EvalOp. simpl; eauto. constructor.
- simpl. unfold Int.ltu; rewrite zlt_true. rewrite Int64.shrx'_shr_2 by auto. reflexivity.
- change (Int.unsigned Int64.iwordsize') with 64; lia.
-*)
+ cbn.
+ rewrite H0.
+ reflexivity.
Qed.
Theorem eval_cmplu:
@@ -565,6 +558,7 @@ Proof.
unfold longoffloat; red; intros. destruct Archi.splitlong eqn:SL.
eapply SplitLongproof.eval_longoffloat; eauto.
TrivialExists.
+ cbn; rewrite H0; reflexivity.
Qed.
Theorem eval_longuoffloat: partial_unary_constructor_sound longuoffloat Val.longuoffloat.
@@ -572,6 +566,7 @@ Proof.
unfold longuoffloat; red; intros. destruct Archi.splitlong eqn:SL.
eapply SplitLongproof.eval_longuoffloat; eauto.
TrivialExists.
+ cbn; rewrite H0; reflexivity.
Qed.
Theorem eval_floatoflong: partial_unary_constructor_sound floatoflong Val.floatoflong.
@@ -579,6 +574,7 @@ Proof.
unfold floatoflong; red; intros. destruct Archi.splitlong eqn:SL.
eapply SplitLongproof.eval_floatoflong; eauto.
TrivialExists.
+ cbn; rewrite H0; reflexivity.
Qed.
Theorem eval_floatoflongu: partial_unary_constructor_sound floatoflongu Val.floatoflongu.
@@ -586,6 +582,7 @@ Proof.
unfold floatoflongu; red; intros. destruct Archi.splitlong eqn:SL.
eapply SplitLongproof.eval_floatoflongu; eauto.
TrivialExists.
+ cbn; rewrite H0; reflexivity.
Qed.
Theorem eval_longofsingle: partial_unary_constructor_sound longofsingle Val.longofsingle.
@@ -593,6 +590,7 @@ Proof.
unfold longofsingle; red; intros. destruct Archi.splitlong eqn:SL.
eapply SplitLongproof.eval_longofsingle; eauto.
TrivialExists.
+ cbn; rewrite H0; reflexivity.
Qed.
Theorem eval_longuofsingle: partial_unary_constructor_sound longuofsingle Val.longuofsingle.
@@ -600,6 +598,7 @@ Proof.
unfold longuofsingle; red; intros. destruct Archi.splitlong eqn:SL.
eapply SplitLongproof.eval_longuofsingle; eauto.
TrivialExists.
+ cbn; rewrite H0; reflexivity.
Qed.
Theorem eval_singleoflong: partial_unary_constructor_sound singleoflong Val.singleoflong.
@@ -607,6 +606,7 @@ Proof.
unfold singleoflong; red; intros. destruct Archi.splitlong eqn:SL.
eapply SplitLongproof.eval_singleoflong; eauto.
TrivialExists.
+ cbn; rewrite H0; reflexivity.
Qed.
Theorem eval_singleoflongu: partial_unary_constructor_sound singleoflongu Val.singleoflongu.
@@ -614,6 +614,7 @@ Proof.
unfold singleoflongu; red; intros. destruct Archi.splitlong eqn:SL.
eapply SplitLongproof.eval_singleoflongu; eauto.
TrivialExists.
+ cbn; rewrite H0; reflexivity.
Qed.
End CMCONSTR.
diff --git a/riscV/SelectOp.vp b/riscV/SelectOp.vp
index 99806006..9932aaf8 100644
--- a/riscV/SelectOp.vp
+++ b/riscV/SelectOp.vp
@@ -419,9 +419,39 @@ Definition floatofsingle (e: expr) := Eop Ofloatofsingle (e ::: Enil).
(** ** Selection *)
+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.
+
Definition select (ty: typ) (cond: condition) (args: exprlist) (e1 e2: expr)
- : option expr
- := None.
+ : option expr :=
+ if same_expr_pure e1 e2
+ then Some e1
+ else
+ if Archi.ptr64 then
+ match ty with
+ | Tlong => Some (Eop Oselectl
+ ((Eop (Ocmp cond) args) ::: e1 ::: e2 ::: Enil))
+ | Tint => Some (Eop Olowlong ((Eop Oselectl
+ ((Eop (Ocmp cond) args) :::
+ (Eop Ocast32signed (e1 ::: Enil)) :::
+ (Eop Ocast32signed (e2 ::: Enil)) ::: Enil)) ::: Enil))
+ | Tfloat => Some (Eop Ofloat_of_bits ((Eop Oselectl
+ ((Eop (Ocmp cond) args) :::
+ (Eop Obits_of_float (e1 ::: Enil)) :::
+ (Eop Obits_of_float (e2 ::: Enil)) ::: Enil)) ::: Enil))
+ | Tsingle => Some
+ (Eop Osingle_of_bits
+ ((Eop Olowlong ((Eop Oselectl
+ ((Eop (Ocmp cond) args) :::
+ (Eop Ocast32signed ((Eop Obits_of_single (e1 ::: Enil)) ::: Enil)) :::
+ (Eop Ocast32signed ((Eop Obits_of_single (e2 ::: Enil)) ::: Enil))
+ ::: Enil)) ::: Enil)) ::: Enil))
+ | _ => None
+ end
+ else None.
(** ** Recognition of addressing modes for load and store operations *)
@@ -452,7 +482,19 @@ Nondetfunction builtin_arg (e: expr) :=
| _ => BA e
end.
+(* floats *)
+Definition divf_base (e1: expr) (e2: expr) :=
+ Eop Odivf (e1 ::: e2 ::: Enil).
+
+Definition divfs_base (e1: expr) (e2: expr) :=
+ Eop Odivfs (e1 ::: e2 ::: Enil).
+
(** Platform-specific known builtins *)
Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr :=
- None.
+ match b with
+ | BI_bits_of_float => Some (Eop Obits_of_single args)
+ | BI_bits_of_double => Some (Eop Obits_of_float args)
+ | BI_float_of_bits => Some (Eop Osingle_of_bits args)
+ | BI_double_of_bits => Some (Eop Ofloat_of_bits args)
+ end.
diff --git a/riscV/SelectOpproof.v b/riscV/SelectOpproof.v
index b0b4b794..f450fe6c 100644
--- a/riscV/SelectOpproof.v
+++ b/riscV/SelectOpproof.v
@@ -22,6 +22,9 @@ Require Import AST Integers Floats.
Require Import Values Memory Builtins Globalenvs.
Require Import Cminor Op CminorSel.
Require Import SelectOp.
+Require Import OpHelpers.
+Require Import OpHelpersproof.
+Require Import Lia.
Local Open Scope cminorsel_scope.
@@ -73,8 +76,10 @@ Ltac TrivialExists :=
(** * Correctness of the smart constructors *)
Section CMCONSTR.
-
-Variable ge: genv.
+Variable prog: program.
+Variable hf: helper_functions.
+Hypothesis HELPERS: helper_functions_declared prog hf.
+Let ge := Genv.globalenv prog.
Variable sp: val.
Variable e: env.
Variable m: mem.
@@ -502,7 +507,12 @@ Theorem eval_divs_base:
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.
+ intros. unfold divs_base. exists z; split. EvalOp.
+ 2: apply Val.lessdef_refl.
+ cbn.
+ rewrite H1.
+ cbn.
+ trivial.
Qed.
Theorem eval_mods_base:
@@ -512,7 +522,12 @@ Theorem eval_mods_base:
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.
+ intros. unfold mods_base. exists z; split. EvalOp.
+ 2: apply Val.lessdef_refl.
+ cbn.
+ rewrite H1.
+ cbn.
+ trivial.
Qed.
Theorem eval_divu_base:
@@ -522,7 +537,12 @@ Theorem eval_divu_base:
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.
+ intros. unfold divu_base. exists z; split. EvalOp.
+ 2: apply Val.lessdef_refl.
+ cbn.
+ rewrite H1.
+ cbn.
+ trivial.
Qed.
Theorem eval_modu_base:
@@ -532,7 +552,12 @@ Theorem eval_modu_base:
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.
+ intros. unfold modu_base. exists z; split. EvalOp.
+ 2: apply Val.lessdef_refl.
+ cbn.
+ rewrite H1.
+ cbn.
+ trivial.
Qed.
Theorem eval_shrximm:
@@ -549,34 +574,12 @@ Proof.
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.
-(*
- intros. destruct x; simpl in H0; try discriminate.
- destruct (Int.ltu n (Int.repr 31)) eqn:LTU; inv H0.
- unfold shrximm.
- predSpec Int.eq Int.eq_spec n Int.zero.
- - subst n. exists (Vint i); split; auto.
- unfold Int.shrx, Int.divs. rewrite Z.quot_1_r. rewrite Int.repr_signed. auto.
- - assert (NZ: Int.unsigned n <> 0).
- { intro EQ; elim H0. rewrite <- (Int.repr_unsigned n). rewrite EQ; auto. }
- assert (LT: 0 <= Int.unsigned n < 31) by (apply Int.ltu_inv in LTU; assumption).
- assert (LTU2: Int.ltu (Int.sub Int.iwordsize n) Int.iwordsize = true).
- { unfold Int.ltu; apply zlt_true.
- unfold Int.sub. change (Int.unsigned Int.iwordsize) with 32.
- rewrite Int.unsigned_repr. lia.
- assert (32 < Int.max_unsigned) by reflexivity. lia. }
- assert (X: eval_expr ge sp e m le
- (Eop (Oshrimm (Int.repr (Int.zwordsize - 1))) (a ::: Enil))
- (Vint (Int.shr i (Int.repr (Int.zwordsize - 1))))).
- { EvalOp. }
- assert (Y: eval_expr ge sp e m le (shrximm_inner a n)
- (Vint (Int.shru (Int.shr i (Int.repr (Int.zwordsize - 1))) (Int.sub Int.iwordsize n)))).
- { EvalOp. simpl. rewrite LTU2. auto. }
- TrivialExists.
- constructor. EvalOp. simpl; eauto. constructor.
- simpl. unfold Int.ltu; rewrite zlt_true. rewrite Int.shrx_shr_2 by auto. reflexivity.
- change (Int.unsigned Int.iwordsize) with 32; lia.
-*)
+ econstructor; split. EvalOp.
+ cbn.
+ rewrite H0.
+ cbn.
+ reflexivity.
+ apply Val.lessdef_refl.
Qed.
Theorem eval_shl: binary_constructor_sound shl Val.shl.
@@ -786,6 +789,7 @@ Theorem eval_intoffloat:
exists v, eval_expr ge sp e m le (intoffloat a) v /\ Val.lessdef y v.
Proof.
intros; unfold intoffloat. TrivialExists.
+ cbn. rewrite H0. reflexivity.
Qed.
Theorem eval_intuoffloat:
@@ -795,6 +799,7 @@ Theorem eval_intuoffloat:
exists v, eval_expr ge sp e m le (intuoffloat a) v /\ Val.lessdef y v.
Proof.
intros; unfold intuoffloat. TrivialExists.
+ cbn. rewrite H0. reflexivity.
Qed.
Theorem eval_floatofintu:
@@ -806,6 +811,7 @@ Proof.
intros until y; unfold floatofintu. case (floatofintu_match a); intros.
InvEval. simpl in H0. TrivialExists.
TrivialExists.
+ cbn. rewrite H0. reflexivity.
Qed.
Theorem eval_floatofint:
@@ -817,6 +823,7 @@ Proof.
intros until y; unfold floatofint. case (floatofint_match a); intros.
InvEval. simpl in H0. TrivialExists.
TrivialExists.
+ cbn. rewrite H0. reflexivity.
Qed.
Theorem eval_intofsingle:
@@ -826,6 +833,7 @@ Theorem eval_intofsingle:
exists v, eval_expr ge sp e m le (intofsingle a) v /\ Val.lessdef y v.
Proof.
intros; unfold intofsingle. TrivialExists.
+ cbn. rewrite H0. reflexivity.
Qed.
Theorem eval_singleofint:
@@ -835,6 +843,7 @@ Theorem eval_singleofint:
exists v, eval_expr ge sp e m le (singleofint a) v /\ Val.lessdef y v.
Proof.
intros; unfold singleofint; TrivialExists.
+ cbn. rewrite H0. reflexivity.
Qed.
Theorem eval_intuofsingle:
@@ -844,6 +853,7 @@ Theorem eval_intuofsingle:
exists v, eval_expr ge sp e m le (intuofsingle a) v /\ Val.lessdef y v.
Proof.
intros; unfold intuofsingle. TrivialExists.
+ cbn. rewrite H0. reflexivity.
Qed.
Theorem eval_singleofintu:
@@ -853,6 +863,7 @@ Theorem eval_singleofintu:
exists v, eval_expr ge sp e m le (singleofintu a) v /\ Val.lessdef y v.
Proof.
intros; unfold intuofsingle. TrivialExists.
+ cbn. rewrite H0. reflexivity.
Qed.
Theorem eval_singleoffloat: unary_constructor_sound singleoffloat Val.singleoffloat.
@@ -865,6 +876,71 @@ Proof.
red; intros. unfold floatofsingle. TrivialExists.
Qed.
+Lemma mod_small_negative:
+ forall a modulus,
+ modulus > 0 -> -modulus < a < 0 -> a mod modulus = a + modulus.
+Proof.
+ intros.
+ replace (a mod modulus) with ((a + modulus) mod modulus).
+ apply Z.mod_small.
+ lia.
+ rewrite <- Zplus_mod_idemp_r.
+ rewrite Z.mod_same by lia.
+ rewrite Z.add_0_r.
+ reflexivity.
+Qed.
+
+Remark normalize_low_long: forall
+ (PTR64 : Archi.ptr64 = true) v1,
+ Val.loword (Val.normalize (Val.longofint v1) Tlong) = Val.normalize v1 Tint.
+Proof.
+ intros.
+ destruct v1; cbn; try rewrite PTR64; trivial.
+ f_equal.
+ unfold Int64.loword.
+ unfold Int.signed.
+ destruct zlt.
+ { rewrite Int64.int_unsigned_repr.
+ apply Int.repr_unsigned.
+ }
+ pose proof (Int.unsigned_range i).
+ rewrite Int64.unsigned_repr_eq.
+ replace ((Int.unsigned i - Int.modulus) mod Int64.modulus)
+ with (Int64.modulus + Int.unsigned i - Int.modulus).
+ {
+ rewrite <- (Int.repr_unsigned i) at 2.
+ apply Int.eqm_samerepr.
+ unfold Int.eqm, eqmod.
+ change Int.modulus with 4294967296 in *.
+ change Int64.modulus with 18446744073709551616 in *.
+ exists 4294967295.
+ lia.
+ }
+ { rewrite mod_small_negative.
+ lia.
+ constructor.
+ constructor.
+ change Int.modulus with 4294967296 in *.
+ change Int.half_modulus with 2147483648 in *.
+ change Int64.modulus with 18446744073709551616 in *.
+ lia.
+ lia.
+ }
+Qed.
+
+Lemma same_expr_pure_correct:
+ forall le a1 a2 v1 v2
+ (PURE : same_expr_pure a1 a2 = true)
+ (EVAL1 : eval_expr ge sp e m le a1 v1)
+ (EVAL2 : eval_expr ge sp e m le a2 v2),
+ v1 = v2.
+Proof.
+ intros.
+ destruct a1; destruct a2; cbn in *; try discriminate.
+ inv EVAL1. inv EVAL2.
+ destruct (ident_eq i i0); congruence.
+Qed.
+
Theorem eval_select:
forall le ty cond al vl a1 v1 a2 v2 a b,
select ty cond al a1 a2 = Some a ->
@@ -876,7 +952,56 @@ Theorem eval_select:
eval_expr ge sp e m le a v
/\ Val.lessdef (Val.select (Some b) v1 v2 ty) v.
Proof.
- unfold select; intros; discriminate.
+ unfold select; intros.
+ pose proof (same_expr_pure_correct le a1 a2 v1 v2) as PURE.
+ destruct (same_expr_pure a1 a2).
+ { rewrite <- PURE by auto.
+ inv H.
+ exists v1. split. assumption.
+ unfold Val.select.
+ destruct b; apply Val.lessdef_normalize.
+ }
+ clear PURE.
+ destruct Archi.ptr64 eqn:PTR64.
+ 2: discriminate.
+ destruct ty; cbn in *; try discriminate.
+ - (* Tint *)
+ inv H. TrivialExists.
+ + cbn. repeat econstructor; eassumption.
+ + cbn. f_equal. rewrite ExtValues.normalize_select01.
+ rewrite H3. destruct b.
+ * rewrite ExtValues.select01_long_true. apply normalize_low_long; assumption.
+ * rewrite ExtValues.select01_long_false. apply normalize_low_long; assumption.
+
+ - (* Tfloat *)
+ inv H. TrivialExists.
+ + cbn. repeat econstructor; eassumption.
+ + cbn. f_equal. rewrite ExtValues.normalize_select01.
+ rewrite H3. destruct b.
+ * rewrite ExtValues.select01_long_true.
+ apply ExtValues.float_bits_normalize.
+ * rewrite ExtValues.select01_long_false.
+ apply ExtValues.float_bits_normalize.
+
+ - (* Tlong *)
+ inv H. TrivialExists.
+ + cbn. repeat econstructor; eassumption.
+ + cbn. f_equal. rewrite ExtValues.normalize_select01.
+ rewrite H3. destruct b.
+ * rewrite ExtValues.select01_long_true. reflexivity.
+ * rewrite ExtValues.select01_long_false. reflexivity.
+
+ - (* Tsingle *)
+ inv H. TrivialExists.
+ + cbn. repeat econstructor; eassumption.
+ + cbn. f_equal. rewrite ExtValues.normalize_select01.
+ rewrite H3. destruct b.
+ * rewrite ExtValues.select01_long_true.
+ rewrite normalize_low_long by assumption.
+ apply ExtValues.single_bits_normalize.
+ * rewrite ExtValues.select01_long_false.
+ rewrite normalize_low_long by assumption.
+ apply ExtValues.single_bits_normalize.
Qed.
Theorem eval_addressing:
@@ -929,6 +1054,27 @@ Proof.
- constructor; auto.
Qed.
+(* floating-point division without HELPERS *)
+Theorem eval_divf_base:
+ forall le a b x 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 (divf_base a b) v /\ Val.lessdef (Val.divf x y) v.
+Proof.
+ intros; unfold divf_base.
+ TrivialExists.
+Qed.
+
+Theorem eval_divfs_base:
+ forall le a b x 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 (divfs_base a b) v /\ Val.lessdef (Val.divfs x y) v.
+Proof.
+ intros; unfold divfs_base.
+ TrivialExists.
+Qed.
+
(** Platform-specific known builtins *)
Theorem eval_platform_builtin:
@@ -938,7 +1084,10 @@ Theorem eval_platform_builtin:
platform_builtin_sem bf vl = Some v ->
exists v', eval_expr ge sp e m le a v' /\ Val.lessdef v v'.
Proof.
- intros. discriminate.
+ destruct bf; intros until le; intro Heval.
+ all: try (inversion Heval; subst a; clear Heval;
+ exists v; split; trivial;
+ repeat (try econstructor; try eassumption)).
Qed.
End CMCONSTR.
diff --git a/riscV/TargetPrinter.ml b/riscV/TargetPrinter.ml
index d8137f84..aab6b9b8 100644
--- a/riscV/TargetPrinter.ml
+++ b/riscV/TargetPrinter.ml
@@ -107,7 +107,9 @@ module Target : TARGET =
let name_of_section = function
| Section_text -> ".text"
- | Section_data i | Section_small_data i ->
+ | Section_data(i, true) ->
+ failwith "_Thread_local unsupported on this platform"
+ | Section_data(i, false) | Section_small_data i ->
variable_section ~sec:".data" ~bss:".bss" i
| Section_const i | Section_small_const i ->
variable_section ~sec:".section .rodata" i
@@ -527,6 +529,8 @@ module Target : TARGET =
fprintf oc " fcvt.s.d %a, %a\n" freg fd freg fs
(* Pseudo-instructions expanded in Asmexpand *)
+ | Pselectl(_, _, _, _) ->
+ assert false
| Pallocframe(sz, ofs) ->
assert false
| Pfreeframe(sz, ofs) ->
diff --git a/riscV/ValueAOp.v b/riscV/ValueAOp.v
index 5670b5fe..e0314c6a 100644
--- a/riscV/ValueAOp.v
+++ b/riscV/ValueAOp.v
@@ -13,9 +13,46 @@
Require Import Coqlib Compopts.
Require Import AST Integers Floats Values Memory Globalenvs.
Require Import Op RTL ValueDomain.
+Require Import Zbits Lia.
(** Value analysis for RISC V operators *)
+Definition zero32 := (I Int.zero).
+Definition zero64 := (L Int64.zero).
+
+(** Functions to select a special register (see Op.v) *)
+
+Definition apply_bin_oreg {B} (optR: option oreg) (sem: aval -> aval -> B) (v1 v2 vz: aval): B :=
+ match optR with
+ | None => sem v1 v2
+ | Some X0_L => sem vz v1
+ | Some X0_R => sem v1 vz
+ end.
+
+Definition eval_may_undef (mu: mayundef) (v1 v2: aval): aval :=
+ match mu with
+ | MUint => match v1, v2 with
+ | I _, I _ => v2
+ | _, _ => Ifptr Ptop
+ end
+ | MUlong => match v1, v2 with
+ | L _, I _ => v2
+ | _, _ => Ifptr Ptop
+ end
+ | MUshrx i =>
+ match v1, v2 with
+ | I _, I _ =>
+ if Int.ltu i (Int.repr 31) then v2 else Ifptr Ptop
+ | _, _ => Ifptr Ptop
+ end
+ | MUshrxl i =>
+ match v1, v2 with
+ | L _, L _ =>
+ if Int.ltu i (Int.repr 63) then v2 else Ifptr Ptop
+ | _, _ => Ifptr Ptop
+ end
+ end.
+
Definition eval_static_condition (cond: condition) (vl: list aval): abool :=
match cond, vl with
| Ccomp c, v1 :: v2 :: nil => cmp_bool c v1 v2
@@ -30,6 +67,22 @@ Definition eval_static_condition (cond: condition) (vl: list aval): abool :=
| 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)
+ | CEbeqw optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmp_bool Ceq) v1 v2 zero32
+ | CEbnew optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmp_bool Cne) v1 v2 zero32
+ | CEbequw optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmpu_bool Ceq) v1 v2 zero32
+ | CEbneuw optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmpu_bool Cne) v1 v2 zero32
+ | CEbltw optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmp_bool Clt) v1 v2 zero32
+ | CEbltuw optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmpu_bool Clt) v1 v2 zero32
+ | CEbgew optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmp_bool Cge) v1 v2 zero32
+ | CEbgeuw optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmpu_bool Cge) v1 v2 zero32
+ | CEbeql optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmpl_bool Ceq) v1 v2 zero64
+ | CEbnel optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmpl_bool Cne) v1 v2 zero64
+ | CEbequl optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmplu_bool Ceq) v1 v2 zero64
+ | CEbneul optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmplu_bool Cne) v1 v2 zero64
+ | CEbltl optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmpl_bool Clt) v1 v2 zero64
+ | CEbltul optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmplu_bool Clt) v1 v2 zero64
+ | CEbgel optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmpl_bool Cge) v1 v2 zero64
+ | CEbgeul optR, v1 :: v2 :: nil => apply_bin_oreg optR (cmplu_bool Cge) v1 v2 zero64
| _, _ => Bnone
end.
@@ -41,6 +94,39 @@ Definition eval_static_addressing (addr: addressing) (vl: list aval): aval :=
| _, _ => Vbot
end.
+Definition bits_of_single (v : aval) : aval :=
+ match v with
+ | FS f => I (Float32.to_bits f)
+ | _ => ntop1 v
+ end.
+
+Definition bits_of_float (v : aval) : aval :=
+ match v with
+ | F f => L (Float.to_bits f)
+ | _ => ntop1 v
+ end.
+
+Definition single_of_bits (v : aval) : aval :=
+ match v with
+ | I f => FS (Float32.of_bits f)
+ | _ => ntop1 v
+ end.
+
+Definition float_of_bits (v : aval) : aval :=
+ match v with
+ | L f => F (Float.of_bits f)
+ | _ => ntop1 v
+ end.
+
+Definition select01_long (vb : aval) (vt : aval) (vf : aval) :=
+ match vb with
+ | I b =>
+ if Int.eq b Int.one then add_undef vt
+ else if Int.eq b Int.zero then add_undef vf
+ else add_undef (vlub vt vf)
+ | _ => add_undef (vlub vt vf)
+ end.
+
Definition eval_static_operation (op: operation) (vl: list aval): aval :=
match op, vl with
| Omove, v1::nil => v1
@@ -59,10 +145,10 @@ Definition eval_static_operation (op: operation) (vl: list aval): aval :=
| Omul, v1::v2::nil => mul v1 v2
| 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
+ | Odiv, v1::v2::nil => divs_total v1 v2
+ | Odivu, v1::v2::nil => divu_total v1 v2
+ | Omod, v1::v2::nil => mods_total v1 v2
+ | Omodu, v1::v2::nil => modu_total v1 v2
| Oand, v1::v2::nil => and v1 v2
| Oandimm n, v1::nil => and v1 (I n)
| Oor, v1::v2::nil => or v1 v2
@@ -88,10 +174,10 @@ Definition eval_static_operation (op: operation) (vl: list aval): aval :=
| Omull, v1::v2::nil => mull v1 v2
| Omullhs, v1::v2::nil => mullhs v1 v2
| Omullhu, v1::v2::nil => mullhu v1 v2
- | Odivl, v1::v2::nil => divls v1 v2
- | Odivlu, v1::v2::nil => divlu v1 v2
- | Omodl, v1::v2::nil => modls v1 v2
- | Omodlu, v1::v2::nil => modlu v1 v2
+ | Odivl, v1::v2::nil => divls_total v1 v2
+ | Odivlu, v1::v2::nil => divlu_total v1 v2
+ | Omodl, v1::v2::nil => modls_total v1 v2
+ | Omodlu, v1::v2::nil => modlu_total v1 v2
| Oandl, v1::v2::nil => andl v1 v2
| Oandlimm n, v1::nil => andl v1 (L n)
| Oorl, v1::v2::nil => orl v1 v2
@@ -119,23 +205,64 @@ Definition eval_static_operation (op: operation) (vl: list aval): aval :=
| Odivfs, v1::v2::nil => divfs v1 v2
| Osingleoffloat, v1::nil => singleoffloat v1
| Ofloatofsingle, v1::nil => floatofsingle v1
- | Ointoffloat, v1::nil => intoffloat v1
- | Ointuoffloat, v1::nil => intuoffloat v1
+ | Ointoffloat, v1::nil => intoffloat_total v1
+ | Ointuoffloat, v1::nil => intuoffloat_total v1
| Ofloatofint, v1::nil => floatofint v1
| Ofloatofintu, v1::nil => floatofintu v1
- | Ointofsingle, v1::nil => intofsingle v1
- | Ointuofsingle, v1::nil => intuofsingle v1
+ | Ointofsingle, v1::nil => intofsingle_total v1
+ | Ointuofsingle, v1::nil => intuofsingle_total v1
| Osingleofint, v1::nil => singleofint v1
| Osingleofintu, v1::nil => singleofintu v1
- | Olongoffloat, v1::nil => longoffloat v1
- | Olonguoffloat, v1::nil => longuoffloat v1
+ | Olongoffloat, v1::nil => longoffloat_total v1
+ | Olonguoffloat, v1::nil => longuoffloat_total v1
| Ofloatoflong, v1::nil => floatoflong v1
| Ofloatoflongu, v1::nil => floatoflongu v1
- | Olongofsingle, v1::nil => longofsingle v1
- | Olonguofsingle, v1::nil => longuofsingle v1
+ | Olongofsingle, v1::nil => longofsingle_total v1
+ | Olonguofsingle, v1::nil => longuofsingle_total v1
| Osingleoflong, v1::nil => singleoflong v1
| Osingleoflongu, v1::nil => singleoflongu v1
| Ocmp c, _ => of_optbool (eval_static_condition c vl)
+ | OEseqw optR, v1::v2::nil => of_optbool (apply_bin_oreg optR (cmp_bool Ceq) v1 v2 zero32)
+ | OEsnew optR, v1::v2::nil => of_optbool (apply_bin_oreg optR (cmp_bool Cne) v1 v2 zero32)
+ | OEsequw optR, v1::v2::nil => of_optbool (apply_bin_oreg optR (cmpu_bool Ceq) v1 v2 zero32)
+ | OEsneuw optR, v1::v2::nil => of_optbool (apply_bin_oreg optR (cmpu_bool Cne) v1 v2 zero32)
+ | OEsltw optR, v1::v2::nil => of_optbool (apply_bin_oreg optR (cmp_bool Clt) v1 v2 zero32)
+ | OEsltuw optR, v1::v2::nil => of_optbool (apply_bin_oreg optR (cmpu_bool Clt) v1 v2 zero32)
+ | OEsltiw n, v1::nil => of_optbool (cmp_bool Clt v1 (I n))
+ | OEsltiuw n, v1::nil => of_optbool (cmpu_bool Clt v1 (I n))
+ | OExoriw n, v1::nil => xor v1 (I n)
+ | OEluiw n, nil => shl (I n) (I (Int.repr 12))
+ | OEaddiw optR n, nil => apply_bin_oreg optR add (I n) (Ifptr Ptop) zero32
+ | OEaddiw optR n, v1::nil => apply_bin_oreg optR add v1 (I n) (Ifptr Ptop)
+ | OEandiw n, v1::nil => and (I n) v1
+ | OEoriw n, v1::nil => or (I n) v1
+ | OEseql optR, v1::v2::nil => of_optbool (apply_bin_oreg optR (cmpl_bool Ceq) v1 v2 zero64)
+ | OEsnel optR, v1::v2::nil => of_optbool (apply_bin_oreg optR (cmpl_bool Cne) v1 v2 zero64)
+ | OEsequl optR, v1::v2::nil => of_optbool (apply_bin_oreg optR (cmplu_bool Ceq) v1 v2 zero64)
+ | OEsneul optR, v1::v2::nil => of_optbool (apply_bin_oreg optR (cmplu_bool Cne) v1 v2 zero64)
+ | OEsltl optR, v1::v2::nil => of_optbool (apply_bin_oreg optR (cmpl_bool Clt) v1 v2 zero64)
+ | OEsltul optR, v1::v2::nil => of_optbool (apply_bin_oreg optR (cmplu_bool Clt) v1 v2 zero64)
+ | OEsltil n, v1::nil => of_optbool (cmpl_bool Clt v1 (L n))
+ | OEsltiul n, v1::nil => of_optbool (cmplu_bool Clt v1 (L n))
+ | OEandil n, v1::nil => andl (L n) v1
+ | OEoril n, v1::nil => orl (L n) v1
+ | OExoril n, v1::nil => xorl v1 (L n)
+ | OEluil n, nil => sign_ext 32 (shll (L n) (L (Int64.repr 12)))
+ | OEaddil optR n, nil => apply_bin_oreg optR addl (L n) (Ifptr Ptop) zero64
+ | OEaddil optR n, v1::nil => apply_bin_oreg optR addl v1 (L n) (Ifptr Ptop)
+ | OEloadli n, nil => L (n)
+ | OEmayundef mu, v1 :: v2 :: nil => eval_may_undef mu v1 v2
+ | OEfeqd, v1::v2::nil => of_optbool (cmpf_bool Ceq v1 v2)
+ | OEfltd, v1::v2::nil => of_optbool (cmpf_bool Clt v1 v2)
+ | OEfled, v1::v2::nil => of_optbool (cmpf_bool Cle v1 v2)
+ | OEfeqs, v1::v2::nil => of_optbool (cmpfs_bool Ceq v1 v2)
+ | OEflts, v1::v2::nil => of_optbool (cmpfs_bool Clt v1 v2)
+ | OEfles, v1::v2::nil => of_optbool (cmpfs_bool Cle v1 v2)
+ | Obits_of_single, v1::nil => bits_of_single v1
+ | Obits_of_float, v1::nil => bits_of_float v1
+ | Osingle_of_bits, v1::nil => single_of_bits v1
+ | Ofloat_of_bits, v1::nil => float_of_bits v1
+ | Oselectl, vb::vt::vf::nil => select01_long vb vt vf
| _, _ => Vbot
end.
@@ -147,6 +274,75 @@ Hypothesis GENV: genv_match bc ge.
Variable sp: block.
Hypothesis STACK: bc sp = BCstack.
+Lemma bits_of_single_sound:
+ forall v x, vmatch bc v x -> vmatch bc (ExtValues.bits_of_single v) (bits_of_single x).
+Proof.
+ unfold ExtValues.bits_of_single; intros. inv H; cbn; constructor.
+Qed.
+
+Lemma bits_of_float_sound:
+ forall v x, vmatch bc v x -> vmatch bc (ExtValues.bits_of_float v) (bits_of_float x).
+Proof.
+ unfold ExtValues.bits_of_float; intros. inv H; cbn; constructor.
+Qed.
+
+Lemma single_of_bits_sound:
+ forall v x, vmatch bc v x -> vmatch bc (ExtValues.single_of_bits v) (single_of_bits x).
+Proof.
+ unfold ExtValues.bits_of_single; intros. inv H; cbn; constructor.
+Qed.
+
+Lemma float_of_bits_sound:
+ forall v x, vmatch bc v x -> vmatch bc (ExtValues.float_of_bits v) (float_of_bits x).
+Proof.
+ unfold ExtValues.bits_of_float; intros. inv H; cbn; constructor.
+Qed.
+
+
+Lemma select01_long_sound:
+ forall vb xb vt xt vf xf
+ (MATCH_b : vmatch bc vb xb)
+ (MATCH_t : vmatch bc vt xt)
+ (MATCH_f : vmatch bc vf xf),
+ vmatch bc (Val.normalize (ExtValues.select01_long vb vt vf) Tlong)
+ (select01_long xb xt xf).
+Proof.
+ intros.
+ inv MATCH_b; cbn; try apply add_undef_undef.
+ - destruct (Int.eq i Int.one). { apply add_undef_normalize; trivial. }
+ destruct (Int.eq i Int.zero). { apply add_undef_normalize; trivial. }
+ cbn. apply add_undef_undef.
+ - destruct (Int.eq i Int.one).
+ { apply add_undef_normalize.
+ apply vmatch_lub_l.
+ trivial. }
+ destruct (Int.eq i Int.zero).
+ { apply add_undef_normalize.
+ apply vmatch_lub_r.
+ trivial. }
+ cbn. apply add_undef_undef.
+ - destruct (Int.eq i Int.one).
+ { apply add_undef_normalize.
+ apply vmatch_lub_l.
+ trivial. }
+ destruct (Int.eq i Int.zero).
+ { apply add_undef_normalize.
+ apply vmatch_lub_r.
+ trivial. }
+ cbn. apply add_undef_undef.
+ - destruct (Int.eq i Int.one).
+ { apply add_undef_normalize.
+ apply vmatch_lub_l.
+ trivial. }
+ destruct (Int.eq i Int.zero).
+ { apply add_undef_normalize.
+ apply vmatch_lub_r.
+ trivial. }
+ cbn. apply add_undef_undef.
+Qed.
+
+Hint Resolve bits_of_single_sound bits_of_float_sound single_of_bits_sound float_of_bits_sound select01_long_sound : va.
+
Theorem eval_static_condition_sound:
forall cond vargs m aargs,
list_forall2 (vmatch bc) vargs aargs ->
@@ -158,7 +354,9 @@ Proof.
destruct cond; simpl; eauto with va.
inv H2.
destruct cond; simpl; eauto with va.
- destruct cond; auto with va.
+ 17: destruct cond; simpl; eauto with va.
+ all: destruct optR as [[]|]; unfold apply_bin_oreg, Op.apply_bin_oreg;
+ unfold zero32, Op.zero32, zero64, Op.zero64; eauto with va.
Qed.
Lemma symbol_address_sound:
@@ -200,6 +398,70 @@ Proof.
rewrite Ptrofs.add_zero_l; eauto with va.
Qed.
+Lemma of_optbool_maketotal_sound:
+ forall ob ab, cmatch ob ab -> vmatch bc (Val.maketotal (option_map Val.of_bool ob)) (of_optbool ab).
+Proof.
+ intros.
+ assert (DEFAULT: vmatch bc (Val.maketotal (option_map Val.of_bool ob)) (Uns Pbot 1)).
+ {
+ destruct ob; simpl; auto with va.
+ destruct b; constructor; try lia.
+ change 1 with (usize Int.one). apply is_uns_usize.
+ red; intros. apply Int.bits_zero.
+ }
+ inv H; auto. simpl. destruct b; constructor.
+Qed.
+
+Lemma eval_cmpu_sound c: forall a1 b1 a0 b0 optR m,
+ c = Ceq \/ c = Cne \/ c = Clt->
+ vmatch bc a1 b1 ->
+ vmatch bc a0 b0 ->
+ vmatch bc (Op.apply_bin_oreg optR (Val.cmpu (Mem.valid_pointer m) c) a1 a0 Op.zero32)
+ (of_optbool (apply_bin_oreg optR (cmpu_bool c) b1 b0 zero32)).
+Proof.
+ intros.
+ destruct optR as [[]|]; unfold Op.apply_bin_oreg, apply_bin_oreg;
+ apply of_optbool_sound; unfold Op.zero32, zero32; eauto with va.
+Qed.
+
+Lemma eval_cmplu_sound c: forall a1 b1 a0 b0 optR m,
+ c = Ceq \/ c = Cne \/ c = Clt->
+ vmatch bc a1 b1 ->
+ vmatch bc a0 b0 ->
+ vmatch bc
+ (Val.maketotal
+ (Op.apply_bin_oreg optR (Val.cmplu (Mem.valid_pointer m) c) a1 a0
+ Op.zero64))
+ (of_optbool (apply_bin_oreg optR (cmplu_bool c) b1 b0 zero64)).
+Proof.
+ intros.
+ destruct optR as [[]|]; unfold Op.apply_bin_oreg, apply_bin_oreg;
+ apply of_optbool_maketotal_sound; unfold Op.zero64, zero64; eauto with va.
+Qed.
+
+Lemma eval_cmp_sound: forall a1 b1 a0 b0 optR cmp,
+ vmatch bc a1 b1 ->
+ vmatch bc a0 b0 ->
+ vmatch bc (Op.apply_bin_oreg optR (Val.cmp cmp) a1 a0 Op.zero32)
+ (of_optbool (apply_bin_oreg optR (cmp_bool cmp) b1 b0 zero32)).
+Proof.
+ intros.
+ destruct optR as [[]|]; unfold Op.apply_bin_oreg, apply_bin_oreg;
+ apply of_optbool_sound; unfold Op.zero32, zero32; eauto with va.
+Qed.
+
+Lemma eval_cmpl_sound: forall a1 b1 a0 b0 optR cmp,
+ vmatch bc a1 b1 ->
+ vmatch bc a0 b0 ->
+ vmatch bc
+ (Val.maketotal (Op.apply_bin_oreg optR (Val.cmpl cmp) a1 a0 Op.zero64))
+ (of_optbool (apply_bin_oreg optR (cmpl_bool cmp) b1 b0 zero64)).
+Proof.
+ intros.
+ destruct optR as [[]|]; unfold Op.apply_bin_oreg, apply_bin_oreg;
+ apply of_optbool_maketotal_sound; unfold Op.zero64, zero64; eauto with va.
+Qed.
+
Theorem eval_static_operation_sound:
forall op vargs m vres aargs,
eval_operation ge (Vptr sp Ptrofs.zero) op vargs m = Some vres ->
@@ -212,6 +474,39 @@ Proof.
destruct (propagate_float_constants tt); constructor.
rewrite Ptrofs.add_zero_l; eauto with va.
apply of_optbool_sound. eapply eval_static_condition_sound; eauto.
+
+ 3,4,6: apply eval_cmpu_sound; auto.
+ 1,2,3: apply eval_cmp_sound; auto.
+ unfold Val.cmp; apply of_optbool_sound; eauto with va.
+ unfold Val.cmpu; apply of_optbool_sound; eauto with va.
+
+ { destruct optR as [[]|]; simpl; eauto with va. }
+ { destruct optR as [[]|];
+ unfold apply_bin_oreg, Op.apply_bin_oreg; eauto with va. }
+ { fold (Val.and (Vint n) a1); eauto with va. }
+ { fold (Val.or (Vint n) a1); eauto with va. }
+ { simpl; try destruct (Int.ltu _ _); eauto with va; unfold ntop1;
+ try apply vmatch_ifptr_undef. }
+ 9: { destruct optR as [[]|]; simpl; eauto with va. }
+ 9: { destruct optR as [[]|];
+ unfold apply_bin_oreg, Op.apply_bin_oreg; eauto with va. }
+ 9: { fold (Val.andl (Vlong n) a1); eauto with va. }
+ 9: { fold (Val.orl (Vlong n) a1); eauto with va. }
+ 9: { simpl; unfold ntop1, sign_ext, Int64.sign_ext, sgn; simpl;
+ apply vmatch_ifptr_l. }
+
+ 1,10: simpl; eauto with va.
+ 10:
+ unfold Op.eval_may_undef, eval_may_undef; destruct mu;
+ inv H1; inv H0; eauto with va;
+ try destruct (Int.ltu _ _); simpl;
+ try eapply vmatch_ifptr_p, pmatch_top'; eauto with va.
+
+ 4,5,7: apply eval_cmplu_sound; auto.
+ 1,3,4: apply eval_cmpl_sound; auto.
+ 2: { unfold Val.cmpl; apply of_optbool_maketotal_sound; eauto with va. }
+ 2: { unfold Val.cmplu; apply of_optbool_maketotal_sound; eauto with va. }
+ all: unfold Val.cmpf; apply of_optbool_sound; eauto with va.
Qed.
End SOUNDNESS.