feat(cmp4): modied certain annotations for mpn_cmp and proved correctness of mpn_cmp4.

This commit is contained in:
xiaoh105
2025-06-12 12:37:01 +08:00
parent 36204b8877
commit f7432dca84
7 changed files with 788 additions and 166 deletions

View File

@ -112,6 +112,67 @@ Proof.
reflexivity.
Qed.
Lemma list_last_cons: forall (a: Z) (l: list Z),
l <> nil ->
last (a :: l) 0 = last l 0.
Proof.
intros.
destruct l.
+ contradiction.
+ simpl.
reflexivity.
Qed.
Lemma list_last_to_Znth: forall (l: list Z),
l <> nil ->
last l 0 = Znth ((Zlength l) - 1) l 0.
Proof.
intros.
induction l.
+ auto.
+ destruct l.
- simpl.
rewrite Znth0_cons.
lia.
- pose proof (@nil_cons Z z l).
specialize (IHl ltac:(auto)); clear H0.
rewrite list_last_cons; [ | pose proof (@nil_cons Z z l); auto ].
rewrite IHl.
pose proof (Zlength_cons a (z :: l)).
unfold Z.succ in H0; rewrite H0; clear H0.
pose proof (Zlength_nonneg l).
pose proof (Zlength_cons z l); unfold Z.succ in H1.
pose proof (Znth_cons (Zlength (z :: l)) a (z :: l) 0 ltac:(lia)).
assert (Zlength (z :: l) + 1 - 1 = Zlength (z :: l)). { lia. }
rewrite H3; clear H3.
rewrite H2.
reflexivity.
Qed.
Lemma Zlength_removelast: forall (l: list Z),
l <> [] ->
Zlength (removelast l) = Zlength l - 1.
Proof.
intros.
induction l.
+ contradiction.
+ destruct l.
- simpl.
rewrite Zlength_nil.
lia.
- pose proof (@nil_cons Z z l).
specialize (IHl ltac:(auto)).
assert (removelast (a :: z :: l) = a :: removelast(z :: l)). {
simpl.
reflexivity.
}
rewrite H1; clear H1.
repeat rewrite Zlength_cons; unfold Z.succ.
rewrite IHl.
rewrite Zlength_cons; unfold Z.succ.
lia.
Qed.
Lemma store_array_rec_false: forall x storeA lo hi (l: list Z),
lo > hi ->
store_array_rec storeA x lo hi l |-- [| False |].

View File

@ -47,7 +47,34 @@ Definition list_store_Z (data: list Z) (n: Z): Prop :=
Definition mpd_store_Z (ptr: addr) (n: Z) (size: Z) (cap: Z): Assertion :=
EX data,
mpd_store_list ptr data cap && [| list_store_Z data n |] && [| size = Zlength data |].
[| list_store_Z data n |] && [| size = Zlength data |] && mpd_store_list ptr data cap.
Definition list_store_Z_compact (data: list Z) (n: Z): Prop :=
list_to_Z data = n /\ last data 0 >= 1 /\ list_within_bound data.
Definition mpd_store_Z_compact (ptr: addr) (n size cap: Z): Assertion :=
EX data,
[| list_store_Z_compact data n |] && [| size = Zlength data |] && mpd_store_list ptr data cap.
Lemma list_store_Z_normal_to_compact: forall (data: list Z) (n: Z),
list_store_Z data n ->
last data 0 >= 1 ->
list_store_Z_compact data n.
Proof.
unfold list_store_Z_compact, list_store_Z.
intros.
tauto.
Qed.
Lemma list_store_Z_compact_to_normal: forall data n,
list_store_Z_compact data n ->
list_store_Z data n.
Proof.
intros.
unfold list_store_Z_compact in H.
unfold list_store_Z.
tauto.
Qed.
Lemma list_store_Z_injection: forall l1 l2 n1 n2,
list_store_Z l1 n1 ->
@ -450,6 +477,95 @@ Proof.
lia.
Qed.
Lemma list_store_Z_compact_cmp: forall l1 l2 n1 n2 i,
0 <= i < Zlength l1 ->
Zlength l1 = Zlength l2 ->
list_store_Z_compact l1 n1 ->
list_store_Z_compact l2 n2 ->
sublist (i + 1) (Zlength l1) l1 = sublist (i + 1) (Zlength l2) l2 ->
Znth i l1 0 < Znth i l2 0 ->
n1 < n2.
Proof.
intros.
apply list_store_Z_compact_to_normal in H1, H2.
apply (list_store_Z_cmp l1 l2 n1 n2 i); tauto.
Qed.
Lemma list_store_Z_cmp_length: forall l1 l2 n1 n2,
Zlength l1 < Zlength l2 ->
list_store_Z_compact l1 n1 ->
list_store_Z_compact l2 n2 ->
n1 < n2.
Proof.
intros.
unfold list_store_Z_compact in *.
destruct H0, H1, H2, H3.
assert (list_store_Z l1 n1 /\ list_store_Z l2 n2). {
unfold list_store_Z.
tauto.
}
clear H0 H1 H4 H5.
destruct H6.
assert (l2 <> []). {
pose proof (Zlength_nonneg l1).
assert (Zlength l2 >= 1). { lia. }
destruct l2.
+ rewrite Zlength_nil in H5.
lia.
+ pose proof (@nil_cons Z z l2).
auto.
}
pose proof (list_store_Z_bound l2 n2 H1) as Hn2.
pose proof (@app_removelast_last Z l2 0 H4).
rewrite H5 in H1.
apply list_store_Z_split in H1; destruct H1.
rewrite (Aux.Zlength_removelast l2) in H1, H6; try auto.
remember (Zlength l2 - 1) as n eqn:Hn.
unfold list_store_Z in H6; destruct H6.
simpl in H6.
assert (n2 >= UINT_MOD ^ n). {
assert (n2 / UINT_MOD ^ n >= 1). { lia. }
pose proof (Zmult_ge_compat_r (n2 / UINT_MOD ^ n) 1 (UINT_MOD ^ n) ltac:(lia) ltac:(lia)).
rewrite Z.mul_1_l in H9.
assert (n2 >= n2 / UINT_MOD ^ n * UINT_MOD ^ n). {
assert (n >= 0). {
destruct l2.
+ contradiction.
+ rewrite Zlength_cons in Hn; unfold Z.succ in Hn.
pose proof (Zlength_nonneg l2).
lia.
}
pose proof (Zmod_eq n2 (UINT_MOD ^ n) ltac:(lia)).
assert (n2 mod UINT_MOD ^ n >= 0). {
pose proof (Z.mod_bound_pos n2 (UINT_MOD ^ n) ltac:(lia) ltac:(lia)).
lia.
}
lia.
}
lia.
}
assert (Zlength l1 <= n). { lia. }
pose proof (list_store_Z_bound l1 n1 H0).
assert (UINT_MOD ^ n >= UINT_MOD ^ (Zlength l1)). {
assert (Zlength l1 = n \/ Zlength l1 < n). { lia. }
destruct H11.
+ rewrite H11.
lia.
+ assert (n >= 0). {
destruct l2.
+ contradiction.
+ rewrite Zlength_cons in Hn; unfold Z.succ in Hn.
pose proof (Zlength_nonneg l2).
lia.
}
pose proof (Z.pow_lt_mono_r_iff UINT_MOD (Zlength l1) n ltac:(lia) ltac:(lia)).
destruct H13 as [H13 _].
specialize (H13 H11).
lia.
}
lia.
Qed.
End Internal.
Record bigint_ent: Type := {

View File

@ -595,11 +595,11 @@ forall (cap2: Z) (l': (@list Z)) (l: (@list Z)) (i: Z) (n: Z) (d: Z) ,
Definition mpn_cmp_safety_wit_1 :=
forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (l1: (@list Z)) ,
[| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
[| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -620,16 +620,16 @@ Definition mpn_cmp_safety_wit_2 :=
forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (l1: (@list Z)) (n: Z) ,
[| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -653,16 +653,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -685,16 +685,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -718,16 +718,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -750,16 +750,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -781,16 +781,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
[| (n < 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -809,11 +809,11 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
Definition mpn_cmp_entail_wit_1 :=
forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (l1: (@list Z)) ,
[| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
[| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -825,16 +825,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
|--
[| ((-1) <= (n_pre - 1 )) |]
&& [| ((n_pre - 1 ) < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist (((n_pre - 1 ) + 1 )) (n_pre) (l1)) = (sublist (((n_pre - 1 ) + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -851,16 +851,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -872,16 +872,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
|--
[| ((-1) <= (n - 1 )) |]
&& [| ((n - 1 ) < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist (((n - 1 ) + 1 )) (n_pre) (l1)) = (sublist (((n - 1 ) + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -899,16 +899,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -920,18 +920,18 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
|--
([| (val1 < val2) |]
&& [| ((-1) = (-1)) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
||
([| (val1 = val2) |]
&& [| ((-1) = 0) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
||
([| (val1 > val2) |]
&& [| ((-1) = 1) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
.
Definition mpn_cmp_return_wit_1_2 :=
@ -941,16 +941,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -962,18 +962,18 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
|--
([| (val1 < val2) |]
&& [| (1 = (-1)) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
||
([| (val1 = val2) |]
&& [| (1 = 0) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
||
([| (val1 > val2) |]
&& [| (1 = 1) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
.
Definition mpn_cmp_return_wit_2 :=
@ -981,16 +981,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
[| (n < 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -1002,37 +1002,37 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
|--
([| (val1 < val2) |]
&& [| (0 = (-1)) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
||
([| (val1 = val2) |]
&& [| (0 = 0) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
||
([| (val1 > val2) |]
&& [| (0 = 1) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
.
Definition mpn_cmp_partial_solve_wit_1 :=
forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) ,
[| (0 < n_pre) |]
[| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 )
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 )
|--
[| (0 < n_pre) |]
[| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 )
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 )
.
Definition mpn_cmp_partial_solve_wit_2 :=
@ -1040,16 +1040,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
[| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -1062,16 +1062,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
[| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -1088,16 +1088,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
[| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -1110,16 +1110,16 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
[| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (0 <= n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
@ -1133,12 +1133,12 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
Definition mpn_cmp_which_implies_wit_1 :=
forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) ,
(mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 )
(mpd_store_Z_compact ap_pre val1 n_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 )
|--
EX (l2: (@list Z)) (l1: (@list Z)) ,
[| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
[| (list_store_Z_compact l1 val1 ) |]
&& [| (list_store_Z_compact l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& (store_uint_array ap_pre n_pre l1 )
@ -1147,6 +1147,334 @@ forall (n_pre: Z) (bp_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z
** (store_undef_uint_array_rec bp_pre (n_pre + 1 ) cap2 )
.
(*----- Function mpn_cmp4 -----*)
Definition mpn_cmp4_safety_wit_1 :=
forall (bn_pre: Z) (bp_pre: Z) (an_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) ,
[| (an_pre >= bn_pre) |]
&& [| (an_pre <> bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& ((( &( "bn" ) )) # Int |-> bn_pre)
** ((( &( "bp" ) )) # Ptr |-> bp_pre)
** ((( &( "an" ) )) # Int |-> an_pre)
** ((( &( "ap" ) )) # Ptr |-> ap_pre)
** (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 )
|--
[| (1 <= INT_MAX) |]
&& [| ((INT_MIN) <= 1) |]
.
Definition mpn_cmp4_safety_wit_2 :=
forall (bn_pre: Z) (bp_pre: Z) (an_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) ,
[| (an_pre < bn_pre) |]
&& [| (an_pre <> bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& ((( &( "bn" ) )) # Int |-> bn_pre)
** ((( &( "bp" ) )) # Ptr |-> bp_pre)
** ((( &( "an" ) )) # Int |-> an_pre)
** ((( &( "ap" ) )) # Ptr |-> ap_pre)
** (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 )
|--
[| (1 <> (INT_MIN)) |]
.
Definition mpn_cmp4_safety_wit_3 :=
forall (bn_pre: Z) (bp_pre: Z) (an_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) ,
[| (an_pre < bn_pre) |]
&& [| (an_pre <> bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& ((( &( "bn" ) )) # Int |-> bn_pre)
** ((( &( "bp" ) )) # Ptr |-> bp_pre)
** ((( &( "an" ) )) # Int |-> an_pre)
** ((( &( "ap" ) )) # Ptr |-> ap_pre)
** (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 )
|--
[| (1 <= INT_MAX) |]
&& [| ((INT_MIN) <= 1) |]
.
Definition mpn_cmp4_return_wit_1_1 :=
forall (bn_pre: Z) (bp_pre: Z) (an_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) ,
[| (an_pre >= bn_pre) |]
&& [| (an_pre <> bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 )
|--
([| (val1 < val2) |]
&& [| (1 = (-1)) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
||
([| (val1 = val2) |]
&& [| (1 = 0) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
||
([| (val1 > val2) |]
&& [| (1 = 1) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
.
Definition mpn_cmp4_return_wit_1_2 :=
forall (bn_pre: Z) (bp_pre: Z) (an_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) ,
[| (an_pre < bn_pre) |]
&& [| (an_pre <> bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 )
|--
([| (val1 < val2) |]
&& [| ((-1) = (-1)) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
||
([| (val1 = val2) |]
&& [| ((-1) = 0) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
||
([| (val1 > val2) |]
&& [| ((-1) = 1) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
.
Definition mpn_cmp4_return_wit_2_1 :=
forall (bn_pre: Z) (bp_pre: Z) (an_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) (retval: Z) ,
[| (val1 > val2) |]
&& [| (retval = 1) |]
&& [| (an_pre <= cap2) |]
&& [| (an_pre = bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 an_pre cap2 )
|--
([| (val1 < val2) |]
&& [| (retval = (-1)) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
||
([| (val1 = val2) |]
&& [| (retval = 0) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
||
([| (val1 > val2) |]
&& [| (retval = 1) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
.
Definition mpn_cmp4_return_wit_2_2 :=
forall (bn_pre: Z) (bp_pre: Z) (an_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) (retval: Z) ,
[| (val1 = val2) |]
&& [| (retval = 0) |]
&& [| (an_pre <= cap2) |]
&& [| (an_pre = bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 an_pre cap2 )
|--
([| (val1 < val2) |]
&& [| (retval = (-1)) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
||
([| (val1 = val2) |]
&& [| (retval = 0) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
||
([| (val1 > val2) |]
&& [| (retval = 1) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
.
Definition mpn_cmp4_return_wit_2_3 :=
forall (bn_pre: Z) (bp_pre: Z) (an_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) (retval: Z) ,
[| (val1 < val2) |]
&& [| (retval = (-1)) |]
&& [| (an_pre <= cap2) |]
&& [| (an_pre = bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 an_pre cap2 )
|--
([| (val1 < val2) |]
&& [| (retval = (-1)) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
||
([| (val1 = val2) |]
&& [| (retval = 0) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
||
([| (val1 > val2) |]
&& [| (retval = 1) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 ))
.
Definition mpn_cmp4_partial_solve_wit_1_pure :=
forall (bn_pre: Z) (bp_pre: Z) (an_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) ,
[| (an_pre = bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& ((( &( "bn" ) )) # Int |-> bn_pre)
** ((( &( "bp" ) )) # Ptr |-> bp_pre)
** ((( &( "an" ) )) # Int |-> an_pre)
** ((( &( "ap" ) )) # Ptr |-> ap_pre)
** (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 )
|--
[| (an_pre = bn_pre) |]
&& [| (bn_pre <= cap2) |]
.
Definition mpn_cmp4_partial_solve_wit_1_aux :=
forall (bn_pre: Z) (bp_pre: Z) (an_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) ,
[| (an_pre = bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 )
|--
[| (an_pre = bn_pre) |]
&& [| (bn_pre <= cap2) |]
&& [| (an_pre = bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 )
.
Definition mpn_cmp4_partial_solve_wit_1 := mpn_cmp4_partial_solve_wit_1_pure -> mpn_cmp4_partial_solve_wit_1_aux.
Definition mpn_cmp4_partial_solve_wit_2_pure :=
forall (bn_pre: Z) (bp_pre: Z) (an_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) ,
[| (an_pre <= cap2) |]
&& [| (an_pre = bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& ((( &( "bn" ) )) # Int |-> bn_pre)
** ((( &( "bp" ) )) # Ptr |-> bp_pre)
** ((( &( "an" ) )) # Int |-> an_pre)
** ((( &( "ap" ) )) # Ptr |-> ap_pre)
** (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 )
|--
[| (0 <= an_pre) |]
&& [| (an_pre <= cap1) |]
&& [| (an_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
.
Definition mpn_cmp4_partial_solve_wit_2_aux :=
forall (bn_pre: Z) (bp_pre: Z) (an_pre: Z) (ap_pre: Z) (val2: Z) (val1: Z) (cap2: Z) (cap1: Z) ,
[| (an_pre <= cap2) |]
&& [| (an_pre = bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 bn_pre cap2 )
|--
[| (0 <= an_pre) |]
&& [| (an_pre <= cap1) |]
&& [| (an_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& [| (an_pre <= cap2) |]
&& [| (an_pre = bn_pre) |]
&& [| (an_pre >= 0) |]
&& [| (bn_pre >= 0) |]
&& [| (an_pre <= cap1) |]
&& [| (bn_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (mpd_store_Z_compact ap_pre val1 an_pre cap1 )
** (mpd_store_Z_compact bp_pre val2 an_pre cap2 )
.
Definition mpn_cmp4_partial_solve_wit_2 := mpn_cmp4_partial_solve_wit_2_pure -> mpn_cmp4_partial_solve_wit_2_aux.
Definition mpn_cmp4_which_implies_wit_1 :=
forall (bn_pre: Z) (an_pre: Z) (cap2: Z) ,
[| (an_pre = bn_pre) |]
&& [| (bn_pre <= cap2) |]
&& emp
|--
[| (an_pre <= cap2) |]
&& emp
.
Module Type VC_Correct.
Axiom proof_of_gmp_abs_safety_wit_1 : gmp_abs_safety_wit_1.
@ -1192,5 +1520,18 @@ Axiom proof_of_mpn_cmp_partial_solve_wit_1 : mpn_cmp_partial_solve_wit_1.
Axiom proof_of_mpn_cmp_partial_solve_wit_2 : mpn_cmp_partial_solve_wit_2.
Axiom proof_of_mpn_cmp_partial_solve_wit_3 : mpn_cmp_partial_solve_wit_3.
Axiom proof_of_mpn_cmp_which_implies_wit_1 : mpn_cmp_which_implies_wit_1.
Axiom proof_of_mpn_cmp4_safety_wit_1 : mpn_cmp4_safety_wit_1.
Axiom proof_of_mpn_cmp4_safety_wit_2 : mpn_cmp4_safety_wit_2.
Axiom proof_of_mpn_cmp4_safety_wit_3 : mpn_cmp4_safety_wit_3.
Axiom proof_of_mpn_cmp4_return_wit_1_1 : mpn_cmp4_return_wit_1_1.
Axiom proof_of_mpn_cmp4_return_wit_1_2 : mpn_cmp4_return_wit_1_2.
Axiom proof_of_mpn_cmp4_return_wit_2_1 : mpn_cmp4_return_wit_2_1.
Axiom proof_of_mpn_cmp4_return_wit_2_2 : mpn_cmp4_return_wit_2_2.
Axiom proof_of_mpn_cmp4_return_wit_2_3 : mpn_cmp4_return_wit_2_3.
Axiom proof_of_mpn_cmp4_partial_solve_wit_1_pure : mpn_cmp4_partial_solve_wit_1_pure.
Axiom proof_of_mpn_cmp4_partial_solve_wit_1 : mpn_cmp4_partial_solve_wit_1.
Axiom proof_of_mpn_cmp4_partial_solve_wit_2_pure : mpn_cmp4_partial_solve_wit_2_pure.
Axiom proof_of_mpn_cmp4_partial_solve_wit_2 : mpn_cmp4_partial_solve_wit_2.
Axiom proof_of_mpn_cmp4_which_implies_wit_1 : mpn_cmp4_which_implies_wit_1.
End VC_Correct.

View File

@ -93,3 +93,27 @@ Proof. Admitted.
Lemma proof_of_mpn_cmp_partial_solve_wit_3 : mpn_cmp_partial_solve_wit_3.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_safety_wit_1 : mpn_cmp4_safety_wit_1.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_safety_wit_2 : mpn_cmp4_safety_wit_2.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_safety_wit_3 : mpn_cmp4_safety_wit_3.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_partial_solve_wit_1_pure : mpn_cmp4_partial_solve_wit_1_pure.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_partial_solve_wit_1 : mpn_cmp4_partial_solve_wit_1.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_partial_solve_wit_2_pure : mpn_cmp4_partial_solve_wit_2_pure.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_partial_solve_wit_2 : mpn_cmp4_partial_solve_wit_2.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_which_implies_wit_1 : mpn_cmp4_which_implies_wit_1.
Proof. Admitted.

View File

@ -198,9 +198,9 @@ Lemma proof_of_mpn_cmp_return_wit_1_1 : mpn_cmp_return_wit_1_1.
Proof.
pre_process.
entailer!.
repeat rewrite <-derivable1_orp_intros1.
Left; Left.
entailer!.
+ unfold mpd_store_Z.
+ unfold mpd_store_Z_compact.
Exists l1 l2.
unfold mpd_store_list.
entailer!.
@ -208,7 +208,7 @@ Proof.
entailer!.
+ assert (Znth n l1 0 < Znth n l2 0). { lia. }
clear H H0.
apply (list_store_Z_cmp l1 l2 val1 val2 n ltac:(lia) ltac:(lia) H4 H5).
apply (list_store_Z_compact_cmp l1 l2 val1 val2 n ltac:(lia) ltac:(lia) H4 H5).
- rewrite <-H6, <-H7.
tauto.
- lia.
@ -217,13 +217,13 @@ Qed.
Lemma proof_of_mpn_cmp_return_wit_1_2 : mpn_cmp_return_wit_1_2.
Proof.
pre_process.
rewrite <-derivable1_orp_intros2.
Right.
entailer!.
+ unfold mpd_store_Z, mpd_store_list.
+ unfold mpd_store_Z_compact, mpd_store_list.
Exists l1 l2.
rewrite <-H6, <-H7.
entailer!.
+ pose proof (list_store_Z_cmp l2 l1 val2 val1 n ltac:(lia) ltac:(lia) H5 H4).
+ pose proof (list_store_Z_compact_cmp l2 l1 val2 val1 n ltac:(lia) ltac:(lia) H5 H4).
rewrite <-H6, <-H7 in H18.
rewrite H8 in H18.
specialize (H18 ltac:(reflexivity) ltac:(lia)).
@ -233,11 +233,67 @@ Qed.
Lemma proof_of_mpn_cmp_which_implies_wit_1 : mpn_cmp_which_implies_wit_1.
Proof.
pre_process.
unfold mpd_store_Z.
unfold mpd_store_Z_compact.
Intros l1 l2.
Exists l2 l1.
unfold mpd_store_list.
entailer!.
rewrite <-H4, <-H6.
rewrite <-H0, <-H2.
entailer!.
Qed.
Lemma proof_of_mpn_cmp4_return_wit_1_1 : mpn_cmp4_return_wit_1_1.
Proof.
pre_process.
Right.
unfold mpd_store_Z_compact.
Intros l1 l2.
Exists l1 l2.
entailer!.
pose proof (list_store_Z_cmp_length l2 l1 val2 val1 ltac:(lia) H9 H7).
lia.
Qed.
Lemma proof_of_mpn_cmp4_return_wit_1_2 : mpn_cmp4_return_wit_1_2.
Proof.
pre_process.
Left; Left.
unfold mpd_store_Z_compact.
entailer!.
Intros l1 l2.
Exists l1 l2.
entailer!.
pose proof (list_store_Z_cmp_length l1 l2 val1 val2 ltac:(lia) H7 H9).
lia.
Qed.
Lemma proof_of_mpn_cmp4_return_wit_2_1 : mpn_cmp4_return_wit_2_1.
Proof.
pre_process.
Right.
unfold mpd_store_Z_compact.
Intros l1 l2.
Exists l1 l2.
entailer!.
Qed.
Lemma proof_of_mpn_cmp4_return_wit_2_2 : mpn_cmp4_return_wit_2_2.
Proof.
pre_process.
Left; Right.
unfold mpd_store_Z_compact.
Intros l1 l2.
Exists l1 l2.
entailer!.
Qed.
Lemma proof_of_mpn_cmp4_return_wit_2_3 : mpn_cmp4_return_wit_2_3.
Proof.
pre_process.
Left; Left.
unfold mpd_store_Z_compact.
Intros l1 l2.
Exists l1 l2.
entailer!.
Qed.