Skip to content
Snippets Groups Projects
Commit 8fab551a authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Fix compilation with Coq 8.10 (lia is weaker).

parent 834105d1
No related branches found
No related tags found
No related merge requests found
...@@ -195,11 +195,12 @@ Section proof. ...@@ -195,11 +195,12 @@ Section proof.
}}}. }}}.
Proof. Proof.
iIntros (Hn Φ) "Hf HΦ". wp_lam. wp_alloc l as "Hl"; first done. iIntros (Hn Φ) "Hf HΦ". wp_lam. wp_alloc l as "Hl"; first done.
wp_apply (wp_array_init_loop _ _ _ 0 n (Z.to_nat n) wp_apply (wp_array_init_loop _ _ _ 0 n (Z.to_nat n) with "[Hl $Hf] [HΦ]").
with "[Hl $Hf] [HΦ]"); first lia. { by rewrite /= Z2Nat.id; last lia. }
{ by rewrite loc_add_0. } { by rewrite loc_add_0. }
iIntros "!>" (vs). iDestruct 1 as (?) "[Hl Hvs]". wp_pures. iIntros "!>" (vs). iDestruct 1 as (Hlen) "[Hl Hvs]". wp_pures.
iApply ("HΦ" $! _ vs). iSplit; [iPureIntro; lia|]. iApply ("HΦ" $! _ vs). iSplit.
{ iPureIntro. by rewrite Hlen Z2Nat.id; last lia. }
rewrite loc_add_0. iFrame. rewrite loc_add_0. iFrame.
Qed. Qed.
Lemma twp_array_init stk E n f : Lemma twp_array_init stk E n f :
...@@ -213,11 +214,12 @@ Section proof. ...@@ -213,11 +214,12 @@ Section proof.
}]]. }]].
Proof. Proof.
iIntros (Hn Φ) "Hf HΦ". wp_lam. wp_alloc l as "Hl"; first done. iIntros (Hn Φ) "Hf HΦ". wp_lam. wp_alloc l as "Hl"; first done.
wp_apply (twp_array_init_loop _ _ _ 0 n (Z.to_nat n) wp_apply (twp_array_init_loop _ _ _ 0 n (Z.to_nat n) with "[Hl $Hf] [HΦ]").
with "[Hl $Hf] [HΦ]"); first lia. { by rewrite /= Z2Nat.id; last lia. }
{ by rewrite loc_add_0. } { by rewrite loc_add_0. }
iIntros (vs). iDestruct 1 as (?) "[Hl Hvs]". wp_pures. iIntros (vs). iDestruct 1 as (Hlen) "[Hl Hvs]". wp_pures.
iApply ("HΦ" $! _ vs). iSplit; [iPureIntro; lia|]. iApply ("HΦ" $! _ vs). iSplit.
{ iPureIntro. by rewrite Hlen Z2Nat.id; last lia. }
rewrite loc_add_0. iFrame. rewrite loc_add_0. iFrame.
Qed. Qed.
End array_init. End array_init.
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment