aboutsummaryrefslogtreecommitdiffstats
path: root/powerpc/Asm.v
diff options
context:
space:
mode:
authorBernhard Schommer <bernhardschommer@gmail.com>2015-08-17 12:45:40 +0200
committerBernhard Schommer <bernhardschommer@gmail.com>2015-08-17 12:45:40 +0200
commit4d70c9820f5b840cfd7870395673723a3151e525 (patch)
tree57c10b973266294f4d1f31853a76e5e3c835503f /powerpc/Asm.v
parentb1a7b024933140296c7f995d20ef840687949771 (diff)
downloadcompcert-kvx-4d70c9820f5b840cfd7870395673723a3151e525.tar.gz
compcert-kvx-4d70c9820f5b840cfd7870395673723a3151e525.zip
Added builtin for the dcbi instruction.
Diffstat (limited to 'powerpc/Asm.v')
-rw-r--r--powerpc/Asm.v8
1 files changed, 4 insertions, 4 deletions
diff --git a/powerpc/Asm.v b/powerpc/Asm.v
index 4eedfba4..f1e84146 100644
--- a/powerpc/Asm.v
+++ b/powerpc/Asm.v
@@ -161,10 +161,11 @@ Inductive instruction : Type :=
| Pcmplwi: ireg -> constant -> instruction (**r same, with immediate argument *)
| Pcmpw: ireg -> ireg -> instruction (**r signed integer comparison *)
| Pcmpwi: ireg -> constant -> instruction (**r same, with immediate argument *)
- | Pcntlzw: ireg -> ireg -> instruction (**r count leading zeros *)
+ | Pcntlzw: ireg -> ireg -> instruction (**r count leading zeros *)
| Pcreqv: crbit -> crbit -> crbit -> instruction (**r not-xor between condition bits *)
| Pcror: crbit -> crbit -> crbit -> instruction (**r or between condition bits *)
| Pcrxor: crbit -> crbit -> crbit -> instruction (**r xor between condition bits *)
+ | Pdcbi: ireg -> ireg -> instruction (**r data cache invalidate *)
| Pdivw: ireg -> ireg -> ireg -> instruction (**r signed division *)
| Pdivwu: ireg -> ireg -> ireg -> instruction (**r unsigned division *)
| Peieio: instruction (**r EIEIO barrier *)
@@ -858,6 +859,7 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
| Pcntlzw _ _
| Pcreqv _ _ _
| Pcrxor _ _ _
+ | Pdcbi _ _
| Peieio
| Pfctiw _ _
| Pfctiwz _ _
@@ -1064,6 +1066,4 @@ Definition data_preg (r: preg) : bool :=
| CR0_0 => false | CR0_1 => false | CR0_2 => false | CR0_3 => false
| CARRY => false
| _ => true
- end.
-
-
+ end. \ No newline at end of file