aboutsummaryrefslogtreecommitdiffstats
path: root/backend/KillUselessMoves.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-10-16 12:07:44 +0200
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-10-16 12:07:44 +0200
commitc7c1bafec40f7824da76e832ec09a628412e29da (patch)
tree868916f1b0c5cb3bd0e9ac2eeb835e61d7d052b1 /backend/KillUselessMoves.v
parenta76d23b77127fa439d7c5c60d322f355cf80c4c9 (diff)
downloadcompcert-kvx-c7c1bafec40f7824da76e832ec09a628412e29da.tar.gz
compcert-kvx-c7c1bafec40f7824da76e832ec09a628412e29da.zip
kill useless moves (not yet connected)
Diffstat (limited to 'backend/KillUselessMoves.v')
-rw-r--r--backend/KillUselessMoves.v40
1 files changed, 40 insertions, 0 deletions
diff --git a/backend/KillUselessMoves.v b/backend/KillUselessMoves.v
new file mode 100644
index 00000000..bdd7ec60
--- /dev/null
+++ b/backend/KillUselessMoves.v
@@ -0,0 +1,40 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* David Monniaux CNRS, VERIMAG *)
+(* *)
+(* Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+Require Import Coqlib Maps Errors Integers Floats Lattice Kildall.
+Require Import AST Linking.
+Require Import Memory Registers Op RTL.
+Require List.
+
+Definition transf_ros (ros: reg + ident) : reg + ident := ros.
+
+Definition transf_instr (pc: node) (instr: instruction) :=
+ match instr with
+ | Iop op args res s =>
+ if (eq_operation op Omove) && (List.list_eq_dec peq args (res :: nil))
+ then Inop s
+ else instr
+ | _ => instr
+ end.
+
+Definition transf_function (f: function) : function :=
+ {| fn_sig := f.(fn_sig);
+ fn_params := f.(fn_params);
+ fn_stacksize := f.(fn_stacksize);
+ fn_code := PTree.map transf_instr f.(fn_code);
+ fn_entrypoint := f.(fn_entrypoint) |}.
+
+Definition transf_fundef (fd: fundef) : fundef :=
+ AST.transf_fundef transf_function fd.
+
+Definition transf_program (p: program) : program :=
+ transform_program transf_fundef p.