aboutsummaryrefslogtreecommitdiffstats
path: root/aarch64/CSE2deps.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-03-03 14:57:20 +0100
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-03-03 14:57:20 +0100
commita2b5e7c85dbbc6a27d941dcd931b36c4aa747fb5 (patch)
tree9312909603179f69e560bdcdcd112060d4b341f7 /aarch64/CSE2deps.v
parent091e00ed16d4189c27a05ad7056eab47bd29f5b7 (diff)
downloadcompcert-kvx-a2b5e7c85dbbc6a27d941dcd931b36c4aa747fb5.tar.gz
compcert-kvx-a2b5e7c85dbbc6a27d941dcd931b36c4aa747fb5.zip
aarch64
Diffstat (limited to 'aarch64/CSE2deps.v')
-rw-r--r--aarch64/CSE2deps.v20
1 files changed, 20 insertions, 0 deletions
diff --git a/aarch64/CSE2deps.v b/aarch64/CSE2deps.v
new file mode 100644
index 00000000..90b514a2
--- /dev/null
+++ b/aarch64/CSE2deps.v
@@ -0,0 +1,20 @@
+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 (Int64.unsigned ofs') chunk' (Int64.unsigned ofs) chunk)
+ else true | _, _, _, _ => true
+ end.