aboutsummaryrefslogtreecommitdiffstats
path: root/backend/Selection.v
diff options
context:
space:
mode:
authorBernhard Schommer <bernhardschommer@gmail.com>2023-02-13 11:59:41 +0100
committerXavier Leroy <xavier.leroy@college-de-france.fr>2023-02-20 09:43:38 +0100
commit9d06ed97f493f89ab41f1dc3962840eadacd3865 (patch)
tree5d4c5df92b19de4400d0dae39b2e7a79a7c5bf42 /backend/Selection.v
parent243c0f7ce4e8b6c22ddd930bd470ab74187e5256 (diff)
downloadcompcert-9d06ed97f493f89ab41f1dc3962840eadacd3865.tar.gz
compcert-9d06ed97f493f89ab41f1dc3962840eadacd3865.zip
Ignore self assign in if-conversion
A self assign `x = x` can be classified as a skip in `classify_stmt`. This helps with the if-conversion of some code that the previous passes generates, e.g. for nested C conditional expressions. Closes: #466
Diffstat (limited to 'backend/Selection.v')
-rw-r--r--backend/Selection.v6
1 files changed, 5 insertions, 1 deletions
diff --git a/backend/Selection.v b/backend/Selection.v
index ca1f1619..e4be867f 100644
--- a/backend/Selection.v
+++ b/backend/Selection.v
@@ -348,7 +348,11 @@ Inductive stmt_class : Type :=
Fixpoint classify_stmt (s: Cminor.stmt) : stmt_class :=
match s with
| Cminor.Sskip => SCskip
- | Cminor.Sassign id a => SCassign id a
+ | Cminor.Sassign id a =>
+ match a with
+ | Cminor.Evar id2 => if ident_eq id id2 then SCskip else SCassign id a
+ | _ => SCassign id a
+ end
| Cminor.Sbuiltin None (EF_debug _ _ _) _ => SCskip
| Cminor.Sseq s1 s2 =>
match classify_stmt s1, classify_stmt s2 with