From c7c1bafec40f7824da76e832ec09a628412e29da Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 16 Oct 2020 12:07:44 +0200 Subject: kill useless moves (not yet connected) --- backend/KillUselessMoves.v | 40 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 40 insertions(+) create mode 100644 backend/KillUselessMoves.v (limited to 'backend/KillUselessMoves.v') 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. -- cgit