aboutsummaryrefslogtreecommitdiffstats
path: root/mppa_k1c/Asmblock.v
diff options
context:
space:
mode:
authorCyril SIX <cyril.six@kalray.eu>2019-02-08 15:39:48 +0100
committerCyril SIX <cyril.six@kalray.eu>2019-02-08 15:39:48 +0100
commit66f236124907af065fc8396f8cefd5726a46b06a (patch)
tree02c79ca01be3b8dbc9698ed9a829255a252b685b /mppa_k1c/Asmblock.v
parent0cb0e0c239f086b766a2b4eb65f79a426db49813 (diff)
downloadcompcert-kvx-66f236124907af065fc8396f8cefd5726a46b06a.tar.gz
compcert-kvx-66f236124907af065fc8396f8cefd5726a46b06a.zip
Added indirect tailcalls
Diffstat (limited to 'mppa_k1c/Asmblock.v')
-rw-r--r--mppa_k1c/Asmblock.v3
1 files changed, 3 insertions, 0 deletions
diff --git a/mppa_k1c/Asmblock.v b/mppa_k1c/Asmblock.v
index 1600b867..fcf45bf8 100644
--- a/mppa_k1c/Asmblock.v
+++ b/mppa_k1c/Asmblock.v
@@ -224,6 +224,7 @@ Inductive cf_instruction : Type :=
(* Pgoto is for tailcalls, Pj_l is for jumping to a particular label *)
| Pgoto (l: label) (**r goto *)
+ | Pigoto (r: ireg) (**r goto from register *)
| Pj_l (l: label) (**r jump to label *)
(* Conditional branches *)
@@ -1149,6 +1150,8 @@ Definition exec_control (f: function) (oc: option control) (rs: regset) (m: mem)
Next (rs#RA <- (rs#PC) #PC <- (rs#r)) m
| Pgoto s =>
Next (rs#PC <- (Genv.symbol_address ge s Ptrofs.zero)) m
+ | Pigoto r =>
+ Next (rs#PC <- (rs#r)) m
| Pj_l l =>
goto_label f l rs m
| Pcb bt r l =>