aboutsummaryrefslogtreecommitdiffstats
path: root/src/hls/RTLPar.v
diff options
context:
space:
mode:
authorYann Herklotz <git@yannherklotz.com>2022-04-08 23:13:53 +0100
committerYann Herklotz <git@yannherklotz.com>2022-04-08 23:13:53 +0100
commit34ea2d39acfc9b0a368ae22e47b77f53dbcbfb76 (patch)
treee20022d1fd3574efd5a8ee64386a4257e560c58f /src/hls/RTLPar.v
parent337d9e45bb6b96ec89f905cf0192d732c7bd53ff (diff)
downloadvericert-34ea2d39acfc9b0a368ae22e47b77f53dbcbfb76.tar.gz
vericert-34ea2d39acfc9b0a368ae22e47b77f53dbcbfb76.zip
Add intermediate files
Diffstat (limited to 'src/hls/RTLPar.v')
-rw-r--r--src/hls/RTLPar.v25
1 files changed, 14 insertions, 11 deletions
diff --git a/src/hls/RTLPar.v b/src/hls/RTLPar.v
index 2c60116..d769b43 100644
--- a/src/hls/RTLPar.v
+++ b/src/hls/RTLPar.v
@@ -52,16 +52,19 @@ Section RELSEM.
Inductive step: state -> trace -> state -> Prop :=
| exec_bblock:
- forall s (f: function) sp pc rs rs' m m' t s' bb pr pr',
- f.(fn_code)!pc = Some bb ->
- step_instr_block sp (mk_instr_state rs pr m) bb.(bb_body)
+ forall s f sp pc rs rs' m m' bb pr pr',
+ step_instr_block sp (mk_instr_state rs pr m) bb
(mk_instr_state rs' pr' m') ->
- step_cf_instr ge (State s f sp pc (mk_bblock nil bb.(bb_exit)) rs' pr' m') t s' ->
- step (State s f sp pc bb rs pr m) t s'
+ step (State s f sp pc bb rs pr m) E0 (JumpState s f sp pc rs' pr' m')
+ | exec_jumpstate :
+ forall s f sp pc rs pr m block t st,
+ f.(fn_code) ! pc = Some block ->
+ step_cf_instr ge (JumpState s f sp pc rs pr m) block.(bb_exit) t st ->
+ step (JumpState s f sp pc rs pr m) t st
| exec_function_internal:
- forall s f args m m' stk bb,
+ forall s f args m m' stk bb cf,
Mem.alloc m 0 f.(fn_stacksize) = (m', stk) ->
- f.(fn_code) ! (f.(fn_entrypoint)) = Some bb ->
+ f.(fn_code) ! (f.(fn_entrypoint)) = Some (mk_bblock bb cf) ->
step (Callstate s (Internal f) args m)
E0 (State s
f
@@ -75,11 +78,11 @@ Section RELSEM.
forall s ef args res t m m',
external_call ef ge args m t res m' ->
step (Callstate s (External ef) args m)
- t (Returnstate s res m')
+ t (Returnstate s res m')
| exec_return:
- forall res f sp pc rs s vres m pr bb,
- step (Returnstate (Stackframe res f sp pc bb rs pr :: s) vres m)
- E0 (State s f sp pc bb (rs#res <- vres) pr m).
+ forall res f sp pc rs s vres m pr,
+ step (Returnstate (Stackframe res f sp pc rs pr :: s) vres m)
+ E0 (State s f sp pc nil (rs#res <- vres) pr m).
End RELSEM.