aboutsummaryrefslogtreecommitdiffstats
path: root/backend/Deadcode.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2015-08-22 09:46:37 +0200
committerXavier Leroy <xavier.leroy@inria.fr>2015-08-22 09:46:37 +0200
commit33dfbe7601ad16fcea5377563fa7ceb4053cb85a (patch)
tree962468d4f01ae9620441a97082c826bc6fdf8c5a /backend/Deadcode.v
parent4f187fdafdac0cf4a8b83964c89d79741dbd813e (diff)
downloadcompcert-33dfbe7601ad16fcea5377563fa7ceb4053cb85a.tar.gz
compcert-33dfbe7601ad16fcea5377563fa7ceb4053cb85a.zip
Renaming {BA,BR}_longofwords -> {BA,BR}_splitlong.
Use EF_debug instead of EF_annot for line number annotations. Introduce PrintAsmaux.print_debug_info (very incomplete). powerpc/Asmexpand: revise expand_memcpy_small.
Diffstat (limited to 'backend/Deadcode.v')
-rw-r--r--backend/Deadcode.v2
1 files changed, 1 insertions, 1 deletions
diff --git a/backend/Deadcode.v b/backend/Deadcode.v
index 32bc26fb..9bf17d1d 100644
--- a/backend/Deadcode.v
+++ b/backend/Deadcode.v
@@ -78,7 +78,7 @@ Fixpoint transfer_builtin_arg (nv: nval) (na: NA.t) (a: builtin_arg reg) : NA.t
| BA_addrstack _ | BA_addrglobal _ _ => (ne, nm)
| BA_loadstack chunk ofs => (ne, nmem_add nm (Stk ofs) (size_chunk chunk))
| BA_loadglobal chunk id ofs => (ne, nmem_add nm (Gl id ofs) (size_chunk chunk))
- | BA_longofwords hi lo =>
+ | BA_splitlong hi lo =>
transfer_builtin_arg All (transfer_builtin_arg All na hi) lo
end.