aboutsummaryrefslogtreecommitdiffstats
path: root/riscV/Asm.v
diff options
context:
space:
mode:
Diffstat (limited to 'riscV/Asm.v')
-rw-r--r--riscV/Asm.v15
1 files changed, 13 insertions, 2 deletions
diff --git a/riscV/Asm.v b/riscV/Asm.v
index 4cd3b1fd..1d8fda11 100644
--- a/riscV/Asm.v
+++ b/riscV/Asm.v
@@ -344,7 +344,8 @@ Inductive instruction : Type :=
| Ploadsi (rd: freg) (f: float32) (**r load an immediate single *)
| Pbtbl (r: ireg) (tbl: list label) (**r N-way branch through a jump table *)
| Pbuiltin: external_function -> list (builtin_arg preg)
- -> builtin_res preg -> instruction. (**r built-in function (pseudo) *)
+ -> builtin_res preg -> instruction (**r built-in function (pseudo) *)
+ | Pnop : instruction. (**r nop instruction *)
(** The pseudo-instructions are the following:
@@ -985,6 +986,7 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
| Pfmsubd _ _ _ _
| Pfnmaddd _ _ _ _
| Pfnmsubd _ _ _ _
+ | Pnop
=> Stuck
end.
@@ -1013,6 +1015,15 @@ Definition preg_of (r: mreg) : preg :=
| Machregs.F28 => F28 | Machregs.F29 => F29 | Machregs.F30 => F30 | Machregs.F31 => F31
end.
+(** Undefine all registers except SP and callee-save registers *)
+
+Definition undef_caller_save_regs (rs: regset) : regset :=
+ fun r =>
+ if preg_eq r SP
+ || In_dec preg_eq r (List.map preg_of (List.filter is_callee_save all_mregs))
+ then rs r
+ else Vundef.
+
(** Extract the values of the arguments of an external call.
We exploit the calling conventions from module [Conventions], except that
we use RISC-V registers instead of locations. *)
@@ -1073,7 +1084,7 @@ Inductive step: state -> trace -> state -> Prop :=
Genv.find_funct_ptr ge b = Some (External ef) ->
external_call ef ge args m t res m' ->
extcall_arguments rs m (ef_sig ef) args ->
- rs' = (set_pair (loc_external_result (ef_sig ef) ) res rs)#PC <- (rs RA) ->
+ rs' = (set_pair (loc_external_result (ef_sig ef) ) res (undef_caller_save_regs rs))#PC <- (rs RA) ->
step (State rs m) t (State rs' m').
End RELSEM.