aboutsummaryrefslogtreecommitdiffstats
path: root/src/translation/HTLgenspec.v
diff options
context:
space:
mode:
authorJames Pollard <james@pollard.dev>2020-06-18 17:47:21 +0100
committerJames Pollard <james@pollard.dev>2020-06-18 17:47:21 +0100
commitb59a2e2913aa7ad010c0652e909ae790c07c7281 (patch)
tree0942c72c352dbf73e7d63190ea04c7d7b668d9e8 /src/translation/HTLgenspec.v
parent6b20fbbeaad23724ca7fbcc10c9445f5cb94b699 (diff)
downloadvericert-kvx-b59a2e2913aa7ad010c0652e909ae790c07c7281.tar.gz
vericert-kvx-b59a2e2913aa7ad010c0652e909ae790c07c7281.zip
Enforce stack size alignment to fix proof.
Diffstat (limited to 'src/translation/HTLgenspec.v')
-rw-r--r--src/translation/HTLgenspec.v3
1 files changed, 2 insertions, 1 deletions
diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v
index 1a9144c..a916cb5 100644
--- a/src/translation/HTLgenspec.v
+++ b/src/translation/HTLgenspec.v
@@ -188,6 +188,7 @@ Inductive tr_module (f : RTL.function) : module -> Prop :=
(forall pc i, Maps.PTree.get pc f.(RTL.fn_code) = Some i ->
tr_code f.(RTL.fn_code) pc i data control fin rtrn st stk) ->
stk_len = Z.to_nat (f.(RTL.fn_stacksize) / 4) ->
+ Z.modulo (f.(RTL.fn_stacksize)) 4 = 0 ->
m = (mkmodule f.(RTL.fn_params)
data
control
@@ -451,7 +452,7 @@ Proof.
inversion s; subst.
unfold transf_module in *.
- monadInv Heqr.
+ destruct (Z.eq_dec (RTL.fn_stacksize f mod 4) 0); monadInv Heqr.
(* TODO: We should be able to fold this into the automation. *)
pose proof (create_arr_inv _ _ _ _ _ _ _ _ EQ0) as STK_LEN.