aboutsummaryrefslogtreecommitdiffstats
path: root/src/translation/HTLgenspec.v
diff options
context:
space:
mode:
authorJames Pollard <james@pollard.dev>2020-06-23 23:00:08 +0100
committerJames Pollard <james@pollard.dev>2020-06-23 23:01:44 +0100
commitec0fa1ac249a8eeb0df9700c50a3e6c4f1b540f2 (patch)
treee168031be13a7e4ed5f7b245c03b18fc48d8b49c /src/translation/HTLgenspec.v
parentb70f007eae5886990a8ffc1e7b94294702e238f8 (diff)
downloadvericert-kvx-ec0fa1ac249a8eeb0df9700c50a3e6c4f1b540f2.tar.gz
vericert-kvx-ec0fa1ac249a8eeb0df9700c50a3e6c4f1b540f2.zip
Normalise entire expression to avoid overflow issues.
Diffstat (limited to 'src/translation/HTLgenspec.v')
-rw-r--r--src/translation/HTLgenspec.v7
1 files changed, 4 insertions, 3 deletions
diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v
index 528c662..42de96b 100644
--- a/src/translation/HTLgenspec.v
+++ b/src/translation/HTLgenspec.v
@@ -189,7 +189,7 @@ Inductive tr_module (f : RTL.function) : module -> Prop :=
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 ->
- 0 <= f.(RTL.fn_stacksize) ->
+ 0 <= f.(RTL.fn_stacksize) < Integers.Ptrofs.modulus ->
m = (mkmodule f.(RTL.fn_params)
data
control
@@ -454,7 +454,8 @@ Proof.
unfold transf_module in *.
unfold stack_correct in *.
- destruct (0 <=? RTL.fn_stacksize f) eqn:STACK_BOUND;
+ destruct (0 <=? RTL.fn_stacksize f) eqn:STACK_BOUND_LOW;
+ destruct (RTL.fn_stacksize f <? Integers.Ptrofs.modulus) eqn:STACK_BOUND_HIGH;
destruct (RTL.fn_stacksize f mod 4 =? 0) eqn:STACK_ALIGN;
simplify;
monadInv Heqr.
@@ -463,7 +464,7 @@ Proof.
pose proof (create_arr_inv _ _ _ _ _ _ _ _ EQ0) as STK_LEN.
rewrite <- STK_LEN.
- econstructor; simpl; trivial.
+ econstructor; simpl; auto.
intros.
inv_incr.
assert (EQ3D := EQ3).