aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
-rw-r--r--aarch64/Op.v14
-rw-r--r--arm/Op.v14
-rw-r--r--kvx/Op.v20
-rw-r--r--powerpc/Op.v14
-rw-r--r--riscV/Op.v10
-rw-r--r--x86/Op.v14
6 files changed, 86 insertions, 0 deletions
diff --git a/aarch64/Op.v b/aarch64/Op.v
index afc25aa6..f720e545 100644
--- a/aarch64/Op.v
+++ b/aarch64/Op.v
@@ -1209,6 +1209,20 @@ Proof.
rewrite (cond_depends_on_memory_correct cond args m1 m2 H). auto.
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 eqn:OP; simpl; try congruence.
+ - intros MEM; destruct cond; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+ - intro MEM; destruct cond; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+Qed.
+
(** Global variables mentioned in an operation or addressing mode *)
Definition globals_addressing (addr: addressing) : list ident :=
diff --git a/arm/Op.v b/arm/Op.v
index 25e48ce1..6f22cece 100644
--- a/arm/Op.v
+++ b/arm/Op.v
@@ -751,6 +751,20 @@ Proof.
auto.
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 eqn:OP; simpl; try congruence.
+ - intros MEM; destruct c; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+ - intro MEM; destruct c; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+Qed.
+
(** Global variables mentioned in an operation or addressing mode *)
Definition globals_operation (op: operation) : list ident :=
diff --git a/kvx/Op.v b/kvx/Op.v
index e2ffa3e5..794bc87b 100644
--- a/kvx/Op.v
+++ b/kvx/Op.v
@@ -1238,6 +1238,26 @@ Proof.
unfold Val.cmpu_bool, Val.cmplu_bool; rewrite SF; reflexivity.
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.
+ - intros MEM; destruct cond; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+ - intros MEM; destruct c0; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+ - intros MEM; destruct c0; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+ - intros MEM; destruct c0; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+Qed.
+
(** Global variables mentioned in an operation or addressing mode *)
Definition globals_addressing (addr: addressing) : list ident :=
diff --git a/powerpc/Op.v b/powerpc/Op.v
index a0ee5bb8..4f14bfac 100644
--- a/powerpc/Op.v
+++ b/powerpc/Op.v
@@ -797,6 +797,20 @@ Proof.
auto.
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 eqn:OP; simpl; try congruence.
+ - intros MEM; destruct c; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+ - intro MEM; destruct c; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+Qed.
+
(** Global variables mentioned in an operation or addressing mode *)
Definition globals_operation (op: operation) : list ident :=
diff --git a/riscV/Op.v b/riscV/Op.v
index 14d07e0b..762c3a02 100644
--- a/riscV/Op.v
+++ b/riscV/Op.v
@@ -872,6 +872,16 @@ Proof.
unfold Val.cmpu_bool, Val.cmplu_bool; rewrite SF; reflexivity.
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.
+ intros MEM; destruct cond; repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+Qed.
+
(** Global variables mentioned in an operation or addressing mode *)
Definition globals_addressing (addr: addressing) : list ident :=
diff --git a/x86/Op.v b/x86/Op.v
index 28e6dbd8..776f9495 100644
--- a/x86/Op.v
+++ b/x86/Op.v
@@ -1037,6 +1037,20 @@ Proof.
auto.
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 eqn:OP; simpl; try congruence.
+ - intros MEM; destruct cond; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+ - intro MEM; destruct c; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+Qed.
+
(** Global variables mentioned in an operation or addressing mode *)
Definition globals_addressing (addr: addressing) : list ident :=