aboutsummaryrefslogtreecommitdiffstats
path: root/backend/CSE2.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-01-28 13:25:12 +0100
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-01-28 13:25:12 +0100
commit47bcd22f7e1febf10bd0629c1774b7ab39fac872 (patch)
tree86515f5dd3565447a470b36ab904dcf27d57681b /backend/CSE2.v
parent76049f12161c0eeeeec8841d2cc07d6601f39b4f (diff)
downloadcompcert-kvx-47bcd22f7e1febf10bd0629c1774b7ab39fac872.tar.gz
compcert-kvx-47bcd22f7e1febf10bd0629c1774b7ab39fac872.zip
CSE2 now works for expressions
Diffstat (limited to 'backend/CSE2.v')
-rw-r--r--backend/CSE2.v11
1 files changed, 3 insertions, 8 deletions
diff --git a/backend/CSE2.v b/backend/CSE2.v
index 0bd5bf81..5b850cbb 100644
--- a/backend/CSE2.v
+++ b/backend/CSE2.v
@@ -343,8 +343,8 @@ Definition apply_instr instr (rel : RELATION.t) : RB.t :=
| Ijumptable _ _ => Some rel
| Istore _ _ _ _ _ => Some (kill_mem rel)
| Iop op args dst _ => Some (gen_oper op dst args rel)
- | Iload _ _ _ dst _
- | Icall _ _ _ dst _ => Some (kill_reg dst rel)
+ | Iload _ _ _ dst _ => Some (kill_reg dst rel)
+ | Icall _ _ _ dst _ => Some (kill_reg dst (kill_mem rel))
| Ibuiltin _ _ res _ => Some (RELATION.top) (* TODO (kill_builtin_res res x) *)
| Itailcall _ _ _ | Ireturn _ => RB.bot
end.
@@ -377,11 +377,6 @@ Definition subst_arg (fmap : option (PMap.t RB.t)) (pc : node) (x : reg) : reg :
Definition subst_args fmap pc := List.map (subst_arg fmap pc).
-Definition gen_move src dst s :=
- if peq src dst
- then Inop s
- else Iop Omove (src::nil) dst s.
-
(* Transform *)
Definition find_op_in_fmap fmap pc op args :=
match fmap with
@@ -399,7 +394,7 @@ Definition transf_instr (fmap : option (PMap.t RB.t))
| Iop op args dst s =>
match find_op_in_fmap fmap pc op args with
| None => Iop op (subst_args fmap pc args) dst s
- | Some src => gen_move src dst s
+ | Some src => Iop Omove (src::nil) dst s
end
| Iload chunk addr args dst s =>
Iload chunk addr (subst_args fmap pc args) dst s