aboutsummaryrefslogtreecommitdiffstats
path: root/riscV/Op.v
diff options
context:
space:
mode:
Diffstat (limited to 'riscV/Op.v')
-rw-r--r--riscV/Op.v296
1 files changed, 152 insertions, 144 deletions
diff --git a/riscV/Op.v b/riscV/Op.v
index 2ceffd4a..a8ff3666 100644
--- a/riscV/Op.v
+++ b/riscV/Op.v
@@ -38,6 +38,11 @@ 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 *)
@@ -52,22 +57,22 @@ Inductive condition : Type :=
| Ccompfs (c: comparison) (**r 32-bit floating-point comparison *)
| Cnotcompfs (c: comparison) (**r negation of a floating-point comparison *)
(* Expansed branches *)
- | CEbeqw (optR0: option bool) (**r branch-if-equal signed *)
- | CEbnew (optR0: option bool) (**r branch-if-not-equal signed *)
- | CEbequw (optR0: option bool) (**r branch-if-equal unsigned *)
- | CEbneuw (optR0: option bool) (**r branch-if-not-equal unsigned *)
- | CEbltw (optR0: option bool) (**r branch-if-less signed *)
- | CEbltuw (optR0: option bool) (**r branch-if-less unsigned *)
- | CEbgew (optR0: option bool) (**r branch-if-greater-or-equal signed *)
- | CEbgeuw (optR0: option bool) (**r branch-if-greater-or-equal unsigned *)
- | CEbeql (optR0: option bool) (**r branch-if-equal signed *)
- | CEbnel (optR0: option bool) (**r branch-if-not-equal signed *)
- | CEbequl (optR0: option bool) (**r branch-if-equal unsigned *)
- | CEbneul (optR0: option bool) (**r branch-if-not-equal unsigned *)
- | CEbltl (optR0: option bool) (**r branch-if-less signed *)
- | CEbltul (optR0: option bool) (**r branch-if-less unsigned *)
- | CEbgel (optR0: option bool) (**r branch-if-greater-or-equal signed *)
- | CEbgeul (optR0: option bool). (**r branch-if-greater-or-equal unsigned *)
+ | 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 :=
@@ -76,7 +81,7 @@ Inductive mayundef: Type :=
| MUshrx: int -> mayundef
| MUshrxl: int -> mayundef.
-(* This allow to have a single RTL constructor to perform an arith operation between an immediate and X0 *)
+(* Type for allowing a single RTL constructor to perform an arith operation between an immediate and X0 *)
Inductive opimm: Type :=
| OPimmADD: int -> opimm
| OPimmADDL: int64 -> opimm.
@@ -185,12 +190,12 @@ Inductive operation : Type :=
| Ocmp (cond: condition) (**r [rd = 1] if condition holds, [rd = 0] otherwise. *)
(* Expansed conditions *)
| OEimmR0 (opi: opimm)
- | OEseqw (optR0: option bool) (**r [rd <- rs1 == rs2] signed *)
- | OEsnew (optR0: option bool) (**r [rd <- rs1 != rs2] signed *)
- | OEsequw (optR0: option bool) (**r [rd <- rs1 == rs2] unsigned *)
- | OEsneuw (optR0: option bool) (**r [rd <- rs1 != rs2] unsigned *)
- | OEsltw (optR0: option bool) (**r set-less-than *)
- | OEsltuw (optR0: option bool) (**r set-less-than unsigned *)
+ | 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 (n: int) (**r add immediate *)
@@ -198,12 +203,12 @@ Inductive operation : Type :=
| OEoriw (n: int) (**r or immediate *)
| OExoriw (n: int) (**r xor immediate *)
| OEluiw (n: int) (**r load upper-immediate *)
- | OEseql (optR0: option bool) (**r [rd <- rs1 == rs2] signed *)
- | OEsnel (optR0: option bool) (**r [rd <- rs1 != rs2] signed *)
- | OEsequl (optR0: option bool) (**r [rd <- rs1 == rs2] unsigned *)
- | OEsneul (optR0: option bool) (**r [rd <- rs1 != rs2] unsigned *)
- | OEsltl (optR0: option bool) (**r set-less-than *)
- | OEsltul (optR0: option bool) (**r set-less-than unsigned *)
+ | 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 (n: int64) (**r add immediate *)
@@ -235,12 +240,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 bool_dec; intros.
+ 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 optR0, optR1; decide equality.
+ all: destruct optR, optR0; decide equality.
Defined.
Definition eq_addressing (x y: addressing) : {x=y} + {x<>y}.
@@ -251,9 +259,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 bool_dec Val.eq; 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 optR0, optR1; try decide equality.
+ all: try destruct optR, optR0; try decide equality.
Defined.
(* Alternate definition:
@@ -273,11 +281,11 @@ Global Opaque eq_condition eq_addressing eq_operation.
Definition zero32 := (Vint Int.zero).
Definition zero64 := (Vlong Int64.zero).
-Definition apply_bin_r0 {B} (optR0: option bool) (sem: val -> val -> B) (v1 v2 vz: val): B :=
- match optR0 with
+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 true => sem vz v1
- | Some false => sem v1 vz
+ | Some X0_L => sem vz v1
+ | Some X0_R => sem v1 vz
end.
Definition eval_may_undef (mu: mayundef) (v1 v2: val): val :=
@@ -332,22 +340,22 @@ Definition eval_condition (cond: condition) (vl: list val) (m: mem): option bool
| 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 optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmp_bool Ceq) v1 v2 zero32
- | CEbnew optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmp_bool Cne) v1 v2 zero32
- | CEbequw optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmpu_bool (Mem.valid_pointer m) Ceq) v1 v2 zero32
- | CEbneuw optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmpu_bool (Mem.valid_pointer m) Cne) v1 v2 zero32
- | CEbltw optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmp_bool Clt) v1 v2 zero32
- | CEbltuw optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmpu_bool (Mem.valid_pointer m) Clt) v1 v2 zero32
- | CEbgew optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmp_bool Cge) v1 v2 zero32
- | CEbgeuw optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmpu_bool (Mem.valid_pointer m) Cge) v1 v2 zero32
- | CEbeql optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmpl_bool Ceq) v1 v2 zero64
- | CEbnel optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmpl_bool Cne) v1 v2 zero64
- | CEbequl optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmplu_bool (Mem.valid_pointer m) Ceq) v1 v2 zero64
- | CEbneul optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmplu_bool (Mem.valid_pointer m) Cne) v1 v2 zero64
- | CEbltl optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmpl_bool Clt) v1 v2 zero64
- | CEbltul optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmplu_bool (Mem.valid_pointer m) Clt) v1 v2 zero64
- | CEbgel optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmpl_bool Cge) v1 v2 zero64
- | CEbgeul optR0, v1 :: v2 :: nil => apply_bin_r0 optR0 (Val.cmplu_bool (Mem.valid_pointer m) Cge) v1 v2 zero64
+ | 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.
@@ -454,12 +462,12 @@ Definition eval_operation
| Ocmp c, _ => Some (Val.of_optbool (eval_condition c vl m))
(* Expansed conditions *)
| OEimmR0 opi, nil => Some (eval_opimmR0 opi)
- | OEseqw optR0, v1::v2::nil => Some (apply_bin_r0 optR0 (Val.cmp Ceq) v1 v2 zero32)
- | OEsnew optR0, v1::v2::nil => Some (apply_bin_r0 optR0 (Val.cmp Cne) v1 v2 zero32)
- | OEsequw optR0, v1::v2::nil => Some (apply_bin_r0 optR0 (Val.cmpu (Mem.valid_pointer m) Ceq) v1 v2 zero32)
- | OEsneuw optR0, v1::v2::nil => Some (apply_bin_r0 optR0 (Val.cmpu (Mem.valid_pointer m) Cne) v1 v2 zero32)
- | OEsltw optR0, v1::v2::nil => Some (apply_bin_r0 optR0 (Val.cmp Clt) v1 v2 zero32)
- | OEsltuw optR0, v1::v2::nil => Some (apply_bin_r0 optR0 (Val.cmpu (Mem.valid_pointer m) Clt) v1 v2 zero32)
+ | 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))
@@ -467,12 +475,12 @@ Definition eval_operation
| OEaddiw n, v1::nil => Some (Val.add (Vint n) v1)
| OEandiw n, v1::nil => Some (Val.and (Vint n) v1)
| OEoriw n, v1::nil => Some (Val.or (Vint n) v1)
- | OEseql optR0, v1::v2::nil => Some (Val.maketotal (apply_bin_r0 optR0 (Val.cmpl Ceq) v1 v2 zero64))
- | OEsnel optR0, v1::v2::nil => Some (Val.maketotal (apply_bin_r0 optR0 (Val.cmpl Cne) v1 v2 zero64))
- | OEsequl optR0, v1::v2::nil => Some (Val.maketotal (apply_bin_r0 optR0 (Val.cmplu (Mem.valid_pointer m) Ceq) v1 v2 zero64))
- | OEsneul optR0, v1::v2::nil => Some (Val.maketotal (apply_bin_r0 optR0 (Val.cmplu (Mem.valid_pointer m) Cne) v1 v2 zero64))
- | OEsltl optR0, v1::v2::nil => Some (Val.maketotal (apply_bin_r0 optR0 (Val.cmpl Clt) v1 v2 zero64))
- | OEsltul optR0, v1::v2::nil => Some (Val.maketotal (apply_bin_r0 optR0 (Val.cmplu (Mem.valid_pointer m) Clt) v1 v2 zero64))
+ | 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))
@@ -924,22 +932,22 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
(* OEimmR0 *)
- destruct opi; unfold eval_opimmR0; simpl; auto.
(* OEseqw *)
- - destruct optR0 as [[]|]; simpl; unfold Val.cmp;
+ - destruct optR as [[]|]; simpl; unfold Val.cmp;
destruct Val.cmp_bool... all: destruct b...
(* OEsnew *)
- - destruct optR0 as [[]|]; simpl; unfold Val.cmp;
+ - destruct optR as [[]|]; simpl; unfold Val.cmp;
destruct Val.cmp_bool... all: destruct b...
(* OEsequw *)
- - destruct optR0 as [[]|]; simpl; unfold Val.cmpu;
+ - destruct optR as [[]|]; simpl; unfold Val.cmpu;
destruct Val.cmpu_bool... all: destruct b...
(* OEsneuw *)
- - destruct optR0 as [[]|]; simpl; unfold Val.cmpu;
+ - destruct optR as [[]|]; simpl; unfold Val.cmpu;
destruct Val.cmpu_bool... all: destruct b...
(* OEsltw *)
- - destruct optR0 as [[]|]; simpl; unfold Val.cmp;
+ - destruct optR as [[]|]; simpl; unfold Val.cmp;
destruct Val.cmp_bool... all: destruct b...
(* OEsltuw *)
- - destruct optR0 as [[]|]; simpl; unfold Val.cmpu;
+ - destruct optR as [[]|]; simpl; unfold Val.cmpu;
destruct Val.cmpu_bool... all: destruct b...
(* OEsltiw *)
- unfold Val.cmp; destruct Val.cmp_bool...
@@ -957,22 +965,22 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
(* OEluiw *)
- destruct (Int.ltu _ _); cbn; trivial.
(* OEseql *)
- - destruct optR0 as [[]|]; simpl; unfold Val.cmpl;
+ - destruct optR as [[]|]; simpl; unfold Val.cmpl;
destruct Val.cmpl_bool... all: destruct b...
(* OEsnel *)
- - destruct optR0 as [[]|]; simpl; unfold Val.cmpl;
+ - destruct optR as [[]|]; simpl; unfold Val.cmpl;
destruct Val.cmpl_bool... all: destruct b...
(* OEsequl *)
- - destruct optR0 as [[]|]; simpl; unfold Val.cmplu;
+ - destruct optR as [[]|]; simpl; unfold Val.cmplu;
destruct Val.cmplu_bool... all: destruct b...
(* OEsneul *)
- - destruct optR0 as [[]|]; simpl; unfold Val.cmplu;
+ - destruct optR as [[]|]; simpl; unfold Val.cmplu;
destruct Val.cmplu_bool... all: destruct b...
(* OEsltl *)
- - destruct optR0 as [[]|]; simpl; unfold Val.cmpl;
+ - destruct optR as [[]|]; simpl; unfold Val.cmpl;
destruct Val.cmpl_bool... all: destruct b...
(* OEsltul *)
- - destruct optR0 as [[]|]; simpl; unfold Val.cmplu;
+ - destruct optR as [[]|]; simpl; unfold Val.cmplu;
destruct Val.cmplu_bool... all: destruct b...
(* OEsltil *)
- unfold Val.cmpl; destruct Val.cmpl_bool...
@@ -1092,22 +1100,22 @@ Definition negate_condition (cond: condition): condition :=
| Cnotcompf c => Ccompf c
| Ccompfs c => Cnotcompfs c
| Cnotcompfs c => Ccompfs c
- | CEbeqw optR0 => CEbnew optR0
- | CEbnew optR0 => CEbeqw optR0
- | CEbequw optR0 => CEbneuw optR0
- | CEbneuw optR0 => CEbequw optR0
- | CEbltw optR0 => CEbgew optR0
- | CEbltuw optR0 => CEbgeuw optR0
- | CEbgew optR0 => CEbltw optR0
- | CEbgeuw optR0 => CEbltuw optR0
- | CEbeql optR0 => CEbnel optR0
- | CEbnel optR0 => CEbeql optR0
- | CEbequl optR0 => CEbneul optR0
- | CEbneul optR0 => CEbequl optR0
- | CEbltl optR0 => CEbgel optR0
- | CEbltul optR0 => CEbgeul optR0
- | CEbgel optR0 => CEbltl optR0
- | CEbgeul optR0 => CEbltul optR0
+ | 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:
@@ -1128,37 +1136,37 @@ Proof.
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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ 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 optR0 as [[]|];
+ repeat (destruct vl; auto); replace (Clt) with (negate_comparison Cge) by auto; destruct optR as [[]|];
apply Val.negate_cmplu_bool.
Qed.
@@ -1303,7 +1311,7 @@ Proof.
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 optR0 as [[]|]; simpl;
+ try destruct optR as [[]|]; simpl;
destruct v; try destruct v0; simpl; auto;
try apply negb_false_iff in H; try rewrite H; auto.
Qed.
@@ -1315,7 +1323,7 @@ Lemma cond_valid_pointer_eq:
Proof.
intros until m2. intro MEM. destruct cond eqn:COND; simpl; try congruence.
all: repeat (destruct args; simpl; try congruence);
- try destruct optR0 as [[]|]; simpl;
+ 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.
@@ -1328,7 +1336,7 @@ 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 optR0 as [[]|]; simpl; try destruct v, v0; try rewrite !MEM; auto;
+ 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.
@@ -1460,14 +1468,14 @@ Proof.
intros EQ; rewrite EQ; destruct b; simpl; constructor; eauto.
Qed.
-Lemma eval_cmpu_bool_inj_opt: forall c v v' v0 v'0 optR0,
+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_r0 optR0 (Val.cmpu (Mem.valid_pointer m1) c) v v0 zero32)
- (apply_bin_r0 optR0 (Val.cmpu (Mem.valid_pointer m2) c) v' v'0 zero32).
+ 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 optR0. intros HV1 HV2.
- destruct optR0 as [[]|]; simpl; unfold zero32, Val.cmpu;
+ 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.
@@ -1501,14 +1509,14 @@ Proof.
intros EQ; rewrite EQ; destruct b; simpl; constructor; eauto.
Qed.
-Lemma eval_cmplu_bool_inj_opt: forall c v v' v0 v'0 optR0,
+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_r0 optR0 (Val.cmplu (Mem.valid_pointer m1) c) v v0 zero64))
- (Val.maketotal (apply_bin_r0 optR0 (Val.cmplu (Mem.valid_pointer m2) c) v' v'0 zero64)).
+ 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 optR0. intros HV1 HV2.
- destruct optR0 as [[]|]; simpl; unfold zero64, Val.cmplu;
+ 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.
@@ -1541,37 +1549,37 @@ 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 optR0 as [[]|]; unfold apply_bin_r0 in *;
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
inv H3; inv H2; simpl in H0; inv H0; auto.
-- destruct optR0 as [[]|]; unfold apply_bin_r0 in *;
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
inv H3; inv H2; simpl in H0; inv H0; auto.
-- destruct optR0 as [[]|]; unfold apply_bin_r0 in *;
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
eapply eval_cmpu_bool_inj'; eauto.
-- destruct optR0 as [[]|]; unfold apply_bin_r0 in *;
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
eapply eval_cmpu_bool_inj'; eauto.
-- destruct optR0 as [[]|]; simpl;
+- destruct optR as [[]|]; simpl;
inv H3; inv H2; simpl in H0; inv H0; auto.
-- destruct optR0 as [[]|]; unfold apply_bin_r0 in *;
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
eapply eval_cmpu_bool_inj'; eauto.
-- destruct optR0 as [[]|]; simpl;
+- destruct optR as [[]|]; simpl;
inv H3; inv H2; simpl in H0; inv H0; auto.
-- destruct optR0 as [[]|]; unfold apply_bin_r0 in *;
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
eapply eval_cmpu_bool_inj'; eauto.
-- destruct optR0 as [[]|]; unfold apply_bin_r0 in *;
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
inv H3; inv H2; simpl in H0; inv H0; auto.
-- destruct optR0 as [[]|]; unfold apply_bin_r0 in *;
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
inv H3; inv H2; simpl in H0; inv H0; auto.
-- destruct optR0 as [[]|]; unfold apply_bin_r0 in *;
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
eapply eval_cmplu_bool_inj'; eauto.
-- destruct optR0 as [[]|]; unfold apply_bin_r0 in *;
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
eapply eval_cmplu_bool_inj'; eauto.
-- destruct optR0 as [[]|]; simpl;
+- destruct optR as [[]|]; simpl;
inv H3; inv H2; simpl in H0; inv H0; auto.
-- destruct optR0 as [[]|]; unfold apply_bin_r0 in *;
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
eapply eval_cmplu_bool_inj'; eauto.
-- destruct optR0 as [[]|]; simpl;
+- destruct optR as [[]|]; simpl;
inv H3; inv H2; simpl in H0; inv H0; auto.
-- destruct optR0 as [[]|]; unfold apply_bin_r0 in *;
+- destruct optR as [[]|]; unfold apply_bin_oreg in *;
eapply eval_cmplu_bool_inj'; eauto.
Qed.
@@ -1784,11 +1792,11 @@ Proof.
(* OEimmR0 *)
- destruct opi; unfold eval_opimmR0; simpl; auto.
(* OEseqw *)
- - destruct optR0 as [[]|]; simpl; unfold zero32, Val.cmp;
+ - 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 optR0 as [[]|]; simpl; unfold zero32, Val.cmp;
+ - 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 *)
@@ -1796,7 +1804,7 @@ Proof.
(* OEsneuw *)
- apply eval_cmpu_bool_inj_opt; auto.
(* OEsltw *)
- - destruct optR0 as [[]|]; simpl; unfold zero32, Val.cmp;
+ - 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 *)
@@ -1818,11 +1826,11 @@ Proof.
(* OEluiw *)
- destruct (Int.ltu _ _); auto.
(* OEseql *)
- - destruct optR0 as [[]|]; simpl; unfold zero64, Val.cmpl;
+ - 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 optR0 as [[]|]; simpl; unfold zero64, Val.cmpl;
+ - 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 *)
@@ -1830,7 +1838,7 @@ Proof.
(* OEsneul *)
- apply eval_cmplu_bool_inj_opt; auto.
(* OEsltl *)
- - destruct optR0 as [[]|]; simpl; unfold zero64, Val.cmpl;
+ - 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 *)