aboutsummaryrefslogtreecommitdiffstats
path: root/riscV/NeedOp.v
diff options
context:
space:
mode:
authorCyril SIX <cyril.six@kalray.eu>2021-06-01 14:58:05 +0200
committerCyril SIX <cyril.six@kalray.eu>2021-06-01 14:58:05 +0200
commit75a2885f610e1d6e91df8e2386a4a4559b615bb9 (patch)
tree4456fa4a9876bc73d3d08b6dad1ad8f2f836578c /riscV/NeedOp.v
parent22504b61be43e5d4a97db621ce3c1785618bbaf0 (diff)
parentc44fc24eb6111c177d1d6fc973a366ebf646202b (diff)
downloadcompcert-kvx-75a2885f610e1d6e91df8e2386a4a4559b615bb9.tar.gz
compcert-kvx-75a2885f610e1d6e91df8e2386a4a4559b615bb9.zip
Merge remote-tracking branch 'origin/kvx-work' into merge_master_8.13.1
Diffstat (limited to 'riscV/NeedOp.v')
-rw-r--r--riscV/NeedOp.v21
1 files changed, 18 insertions, 3 deletions
diff --git a/riscV/NeedOp.v b/riscV/NeedOp.v
index fe000976..6041a34d 100644
--- a/riscV/NeedOp.v
+++ b/riscV/NeedOp.v
@@ -96,8 +96,10 @@ Definition needs_of_operation (op: operation) (nv: nval): list nval :=
| OEsltiw _ => op1 (default nv)
| OEsltiuw _ => op1 (default nv)
| OExoriw _ => op1 (bitwise nv)
- | OEluiw _ _ => op1 (default nv)
- | OEaddiwr0 _ _ => op1 (default nv) (* TODO gourdinl modarith impossible? *)
+ | 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)
@@ -108,8 +110,11 @@ Definition needs_of_operation (op: operation) (nv: nval): list nval :=
| OEsltiul _ => op1 (default nv)
| OExoril _ => op1 (default nv)
| OEluil _ => op1 (default nv)
- | OEaddilr0 _ => op1 (default nv) (* TODO gourdinl modarith impossible? *)
+ | 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)
@@ -188,6 +193,16 @@ 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.