feat(cmp4): modied certain annotations for mpn_cmp and proved correctness of mpn_cmp4.
This commit is contained in:
@ -112,6 +112,67 @@ Proof.
|
|||||||
reflexivity.
|
reflexivity.
|
||||||
Qed.
|
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),
|
Lemma store_array_rec_false: forall x storeA lo hi (l: list Z),
|
||||||
lo > hi ->
|
lo > hi ->
|
||||||
store_array_rec storeA x lo hi l |-- [| False |].
|
store_array_rec storeA x lo hi l |-- [| False |].
|
||||||
|
@ -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 :=
|
Definition mpd_store_Z (ptr: addr) (n: Z) (size: Z) (cap: Z): Assertion :=
|
||||||
EX data,
|
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,
|
Lemma list_store_Z_injection: forall l1 l2 n1 n2,
|
||||||
list_store_Z l1 n1 ->
|
list_store_Z l1 n1 ->
|
||||||
@ -450,6 +477,95 @@ Proof.
|
|||||||
lia.
|
lia.
|
||||||
Qed.
|
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.
|
End Internal.
|
||||||
|
|
||||||
Record bigint_ent: Type := {
|
Record bigint_ent: Type := {
|
||||||
|
@ -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 :=
|
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)) ,
|
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_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) ,
|
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) |]
|
[| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) |]
|
&& [| (n >= 0) |]
|
||||||
&& [| ((-1) <= n) |]
|
&& [| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) |]
|
&& [| (n >= 0) |]
|
||||||
&& [| ((-1) <= n) |]
|
&& [| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) |]
|
&& [| (n >= 0) |]
|
||||||
&& [| ((-1) <= n) |]
|
&& [| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) |]
|
&& [| (n >= 0) |]
|
||||||
&& [| ((-1) <= n) |]
|
&& [| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) |]
|
[| (n < 0) |]
|
||||||
&& [| ((-1) <= n) |]
|
&& [| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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 :=
|
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)) ,
|
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_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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 )) |]
|
[| ((-1) <= (n_pre - 1 )) |]
|
||||||
&& [| ((n_pre - 1 ) < n_pre) |]
|
&& [| ((n_pre - 1 ) < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist (((n_pre - 1 ) + 1 )) (n_pre) (l1)) = (sublist (((n_pre - 1 ) + 1 )) (n_pre) (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_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) |]
|
&& [| (n >= 0) |]
|
||||||
&& [| ((-1) <= n) |]
|
&& [| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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 )) |]
|
[| ((-1) <= (n - 1 )) |]
|
||||||
&& [| ((n - 1 ) < n_pre) |]
|
&& [| ((n - 1 ) < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist (((n - 1 ) + 1 )) (n_pre) (l1)) = (sublist (((n - 1 ) + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist (((n - 1 ) + 1 )) (n_pre) (l1)) = (sublist (((n - 1 ) + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) |]
|
&& [| (n >= 0) |]
|
||||||
&& [| ((-1) <= n) |]
|
&& [| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) |]
|
([| (val1 < val2) |]
|
||||||
&& [| ((-1) = (-1)) |]
|
&& [| ((-1) = (-1)) |]
|
||||||
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
|
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
|
||||||
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
|
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
|
||||||
||
|
||
|
||||||
([| (val1 = val2) |]
|
([| (val1 = val2) |]
|
||||||
&& [| ((-1) = 0) |]
|
&& [| ((-1) = 0) |]
|
||||||
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
|
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
|
||||||
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
|
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
|
||||||
||
|
||
|
||||||
([| (val1 > val2) |]
|
([| (val1 > val2) |]
|
||||||
&& [| ((-1) = 1) |]
|
&& [| ((-1) = 1) |]
|
||||||
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
|
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
|
||||||
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
|
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
|
||||||
.
|
.
|
||||||
|
|
||||||
Definition mpn_cmp_return_wit_1_2 :=
|
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) |]
|
&& [| (n >= 0) |]
|
||||||
&& [| ((-1) <= n) |]
|
&& [| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) |]
|
([| (val1 < val2) |]
|
||||||
&& [| (1 = (-1)) |]
|
&& [| (1 = (-1)) |]
|
||||||
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
|
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
|
||||||
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
|
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
|
||||||
||
|
||
|
||||||
([| (val1 = val2) |]
|
([| (val1 = val2) |]
|
||||||
&& [| (1 = 0) |]
|
&& [| (1 = 0) |]
|
||||||
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
|
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
|
||||||
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
|
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
|
||||||
||
|
||
|
||||||
([| (val1 > val2) |]
|
([| (val1 > val2) |]
|
||||||
&& [| (1 = 1) |]
|
&& [| (1 = 1) |]
|
||||||
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
|
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
|
||||||
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
|
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
|
||||||
.
|
.
|
||||||
|
|
||||||
Definition mpn_cmp_return_wit_2 :=
|
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) |]
|
[| (n < 0) |]
|
||||||
&& [| ((-1) <= n) |]
|
&& [| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) |]
|
([| (val1 < val2) |]
|
||||||
&& [| (0 = (-1)) |]
|
&& [| (0 = (-1)) |]
|
||||||
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
|
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
|
||||||
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
|
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
|
||||||
||
|
||
|
||||||
([| (val1 = val2) |]
|
([| (val1 = val2) |]
|
||||||
&& [| (0 = 0) |]
|
&& [| (0 = 0) |]
|
||||||
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
|
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
|
||||||
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
|
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
|
||||||
||
|
||
|
||||||
([| (val1 > val2) |]
|
([| (val1 > val2) |]
|
||||||
&& [| (0 = 1) |]
|
&& [| (0 = 1) |]
|
||||||
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
|
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
|
||||||
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
|
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 ))
|
||||||
.
|
.
|
||||||
|
|
||||||
Definition mpn_cmp_partial_solve_wit_1 :=
|
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) ,
|
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 <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (cap1 <= 100000000) |]
|
||||||
&& [| (cap2 <= 100000000) |]
|
&& [| (cap2 <= 100000000) |]
|
||||||
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
|
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
|
||||||
** (mpd_store_Z bp_pre val2 n_pre cap2 )
|
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 )
|
||||||
|--
|
|--
|
||||||
[| (0 < n_pre) |]
|
[| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (cap1 <= 100000000) |]
|
||||||
&& [| (cap2 <= 100000000) |]
|
&& [| (cap2 <= 100000000) |]
|
||||||
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
|
&& (mpd_store_Z_compact ap_pre val1 n_pre cap1 )
|
||||||
** (mpd_store_Z bp_pre val2 n_pre cap2 )
|
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 )
|
||||||
.
|
.
|
||||||
|
|
||||||
Definition mpn_cmp_partial_solve_wit_2 :=
|
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) |]
|
[| (n >= 0) |]
|
||||||
&& [| ((-1) <= n) |]
|
&& [| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) |]
|
[| (n >= 0) |]
|
||||||
&& [| ((-1) <= n) |]
|
&& [| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) |]
|
[| (n >= 0) |]
|
||||||
&& [| ((-1) <= n) |]
|
&& [| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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) |]
|
[| (n >= 0) |]
|
||||||
&& [| ((-1) <= n) |]
|
&& [| ((-1) <= n) |]
|
||||||
&& [| (n < n_pre) |]
|
&& [| (n < n_pre) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
&& [| ((sublist ((n + 1 )) (n_pre) (l1)) = (sublist ((n + 1 )) (n_pre) (l2))) |]
|
||||||
&& [| (list_store_Z l1 val1 ) |]
|
&& [| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& [| (0 < n_pre) |]
|
&& [| (0 <= n_pre) |]
|
||||||
&& [| (n_pre <= cap1) |]
|
&& [| (n_pre <= cap1) |]
|
||||||
&& [| (n_pre <= cap2) |]
|
&& [| (n_pre <= cap2) |]
|
||||||
&& [| (cap1 <= 100000000) |]
|
&& [| (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 :=
|
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) ,
|
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_compact ap_pre val1 n_pre cap1 )
|
||||||
** (mpd_store_Z bp_pre val2 n_pre cap2 )
|
** (mpd_store_Z_compact bp_pre val2 n_pre cap2 )
|
||||||
|--
|
|--
|
||||||
EX (l2: (@list Z)) (l1: (@list Z)) ,
|
EX (l2: (@list Z)) (l1: (@list Z)) ,
|
||||||
[| (list_store_Z l1 val1 ) |]
|
[| (list_store_Z_compact l1 val1 ) |]
|
||||||
&& [| (list_store_Z l2 val2 ) |]
|
&& [| (list_store_Z_compact l2 val2 ) |]
|
||||||
&& [| (n_pre = (Zlength (l1))) |]
|
&& [| (n_pre = (Zlength (l1))) |]
|
||||||
&& [| (n_pre = (Zlength (l2))) |]
|
&& [| (n_pre = (Zlength (l2))) |]
|
||||||
&& (store_uint_array ap_pre n_pre l1 )
|
&& (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 )
|
** (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.
|
Module Type VC_Correct.
|
||||||
|
|
||||||
Axiom proof_of_gmp_abs_safety_wit_1 : gmp_abs_safety_wit_1.
|
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_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_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_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.
|
End VC_Correct.
|
||||||
|
@ -93,3 +93,27 @@ Proof. Admitted.
|
|||||||
Lemma proof_of_mpn_cmp_partial_solve_wit_3 : mpn_cmp_partial_solve_wit_3.
|
Lemma proof_of_mpn_cmp_partial_solve_wit_3 : mpn_cmp_partial_solve_wit_3.
|
||||||
Proof. Admitted.
|
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.
|
||||||
|
|
||||||
|
@ -198,9 +198,9 @@ Lemma proof_of_mpn_cmp_return_wit_1_1 : mpn_cmp_return_wit_1_1.
|
|||||||
Proof.
|
Proof.
|
||||||
pre_process.
|
pre_process.
|
||||||
entailer!.
|
entailer!.
|
||||||
repeat rewrite <-derivable1_orp_intros1.
|
Left; Left.
|
||||||
entailer!.
|
entailer!.
|
||||||
+ unfold mpd_store_Z.
|
+ unfold mpd_store_Z_compact.
|
||||||
Exists l1 l2.
|
Exists l1 l2.
|
||||||
unfold mpd_store_list.
|
unfold mpd_store_list.
|
||||||
entailer!.
|
entailer!.
|
||||||
@ -208,7 +208,7 @@ Proof.
|
|||||||
entailer!.
|
entailer!.
|
||||||
+ assert (Znth n l1 0 < Znth n l2 0). { lia. }
|
+ assert (Znth n l1 0 < Znth n l2 0). { lia. }
|
||||||
clear H H0.
|
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.
|
- rewrite <-H6, <-H7.
|
||||||
tauto.
|
tauto.
|
||||||
- lia.
|
- lia.
|
||||||
@ -217,13 +217,13 @@ Qed.
|
|||||||
Lemma proof_of_mpn_cmp_return_wit_1_2 : mpn_cmp_return_wit_1_2.
|
Lemma proof_of_mpn_cmp_return_wit_1_2 : mpn_cmp_return_wit_1_2.
|
||||||
Proof.
|
Proof.
|
||||||
pre_process.
|
pre_process.
|
||||||
rewrite <-derivable1_orp_intros2.
|
Right.
|
||||||
entailer!.
|
entailer!.
|
||||||
+ unfold mpd_store_Z, mpd_store_list.
|
+ unfold mpd_store_Z_compact, mpd_store_list.
|
||||||
Exists l1 l2.
|
Exists l1 l2.
|
||||||
rewrite <-H6, <-H7.
|
rewrite <-H6, <-H7.
|
||||||
entailer!.
|
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 <-H6, <-H7 in H18.
|
||||||
rewrite H8 in H18.
|
rewrite H8 in H18.
|
||||||
specialize (H18 ltac:(reflexivity) ltac:(lia)).
|
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.
|
Lemma proof_of_mpn_cmp_which_implies_wit_1 : mpn_cmp_which_implies_wit_1.
|
||||||
Proof.
|
Proof.
|
||||||
pre_process.
|
pre_process.
|
||||||
unfold mpd_store_Z.
|
unfold mpd_store_Z_compact.
|
||||||
Intros l1 l2.
|
Intros l1 l2.
|
||||||
Exists l2 l1.
|
Exists l2 l1.
|
||||||
unfold mpd_store_list.
|
unfold mpd_store_list.
|
||||||
entailer!.
|
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!.
|
entailer!.
|
||||||
Qed.
|
Qed.
|
@ -102,26 +102,26 @@ mpn_cmp (unsigned int *ap, unsigned int *bp, int n)
|
|||||||
/*@
|
/*@
|
||||||
With cap1 cap2 val1 val2
|
With cap1 cap2 val1 val2
|
||||||
Require
|
Require
|
||||||
mpd_store_Z(ap, val1, n, cap1) *
|
mpd_store_Z_compact(ap, val1, n, cap1) *
|
||||||
mpd_store_Z(bp, val2, n, cap2) &&
|
mpd_store_Z_compact(bp, val2, n, cap2) &&
|
||||||
0 < n && n <= cap1 && n <= cap2 &&
|
0 <= n && n <= cap1 && n <= cap2 &&
|
||||||
cap1 <= 100000000 && cap2 <= 100000000
|
cap1 <= 100000000 && cap2 <= 100000000
|
||||||
Ensure
|
Ensure
|
||||||
(val1 > val2 && __return == 1 ||
|
(val1 > val2 && __return == 1 ||
|
||||||
val1 == val2 && __return == 0 ||
|
val1 == val2 && __return == 0 ||
|
||||||
val1 < val2 && __return == -1) &&
|
val1 < val2 && __return == -1) &&
|
||||||
mpd_store_Z(ap@pre, val1, n@pre, cap1) *
|
mpd_store_Z_compact(ap@pre, val1, n@pre, cap1) *
|
||||||
mpd_store_Z(bp@pre, val2, n@pre, cap2)
|
mpd_store_Z_compact(bp@pre, val2, n@pre, cap2)
|
||||||
*/
|
*/
|
||||||
{
|
{
|
||||||
/*@
|
/*@
|
||||||
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)
|
||||||
which implies
|
which implies
|
||||||
exists l1 l2,
|
exists l1 l2,
|
||||||
store_uint_array(ap@pre, n@pre, l1) * store_uint_array(bp@pre, n@pre, l2) *
|
store_uint_array(ap@pre, n@pre, l1) * store_uint_array(bp@pre, n@pre, l2) *
|
||||||
store_undef_uint_array_rec(ap@pre, n@pre + 1, cap1) *
|
store_undef_uint_array_rec(ap@pre, n@pre + 1, cap1) *
|
||||||
store_undef_uint_array_rec(bp@pre, n@pre + 1, cap2) &&
|
store_undef_uint_array_rec(bp@pre, n@pre + 1, cap2) &&
|
||||||
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)
|
n@pre == Zlength(l1) && n@pre == Zlength(l2)
|
||||||
*/
|
*/
|
||||||
/*@
|
/*@
|
||||||
@ -133,7 +133,7 @@ mpn_cmp (unsigned int *ap, unsigned int *bp, int n)
|
|||||||
store_uint_array(ap@pre, n@pre, l1) * store_uint_array(bp@pre, n@pre, l2) *
|
store_uint_array(ap@pre, n@pre, l1) * store_uint_array(bp@pre, n@pre, l2) *
|
||||||
store_undef_uint_array_rec(ap@pre, n@pre + 1, cap1) *
|
store_undef_uint_array_rec(ap@pre, n@pre + 1, cap1) *
|
||||||
store_undef_uint_array_rec(bp@pre, n@pre + 1, cap2) &&
|
store_undef_uint_array_rec(bp@pre, n@pre + 1, cap2) &&
|
||||||
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) &&
|
n@pre == Zlength(l1) && n@pre == Zlength(l2) &&
|
||||||
sublist(n + 1, n@pre, l1) == sublist(n + 1, n@pre, l2)
|
sublist(n + 1, n@pre, l1) == sublist(n + 1, n@pre, l2)
|
||||||
*/
|
*/
|
||||||
@ -157,24 +157,45 @@ mpn_cmp (unsigned int *ap, unsigned int *bp, int n)
|
|||||||
/*处理位数不同的情况*/
|
/*处理位数不同的情况*/
|
||||||
static int
|
static int
|
||||||
mpn_cmp4 (unsigned int *ap, int an, unsigned int *bp, int bn)
|
mpn_cmp4 (unsigned int *ap, int an, unsigned int *bp, int bn)
|
||||||
|
/*@
|
||||||
|
With cap1 cap2 val1 val2
|
||||||
|
Require
|
||||||
|
mpd_store_Z_compact(ap, val1, an, cap1) *
|
||||||
|
mpd_store_Z_compact(bp, val2, bn, cap2) &&
|
||||||
|
an >= 0 && bn >= 0 && an <= cap1 && bn <= cap2 &&
|
||||||
|
cap1 <= 100000000 && cap2 <= 100000000
|
||||||
|
Ensure
|
||||||
|
(val1 > val2 && __return == 1 ||
|
||||||
|
val1 == val2 && __return == 0 ||
|
||||||
|
val1 < val2 && __return == -1) &&
|
||||||
|
mpd_store_Z_compact(ap@pre, val1, an@pre, cap1) *
|
||||||
|
mpd_store_Z_compact(bp@pre, val2, bn@pre, cap2)
|
||||||
|
*/
|
||||||
{
|
{
|
||||||
if (an != bn)
|
if (an != bn)
|
||||||
return an < bn ? -1 : 1;
|
return an < bn ? -1 : 1;
|
||||||
else
|
else {
|
||||||
|
/*@
|
||||||
|
an@pre == bn@pre && bn@pre <= cap2
|
||||||
|
which implies
|
||||||
|
an@pre <= cap2
|
||||||
|
*/
|
||||||
return mpn_cmp (ap, bp, an);
|
return mpn_cmp (ap, bp, an);
|
||||||
|
}
|
||||||
}
|
}
|
||||||
|
|
||||||
|
|
||||||
/*返回非0的位数*/
|
/*返回非0的位数*/
|
||||||
static int
|
/*static int
|
||||||
mpn_normalized_size (unsigned int *xp, int n)
|
mpn_normalized_size (unsigned int *xp, int n)
|
||||||
{
|
{
|
||||||
while (n > 0 && xp[n-1] == 0)
|
while (n > 0 && xp[n-1] == 0)
|
||||||
--n;
|
--n;
|
||||||
return n;
|
return n;
|
||||||
}
|
}*/
|
||||||
|
|
||||||
/* 多精度数ap 加上单精度数b,返回最后产生的进位 */
|
/* 多精度数ap 加上单精度数b,返回最后产生的进位 */
|
||||||
unsigned int
|
/*unsigned int
|
||||||
mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
||||||
{
|
{
|
||||||
int i;
|
int i;
|
||||||
@ -183,17 +204,17 @@ mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
|||||||
do
|
do
|
||||||
{
|
{
|
||||||
unsigned int r = ap[i] + b;
|
unsigned int r = ap[i] + b;
|
||||||
/* Carry out */
|
// Carry out
|
||||||
b = (r < b);
|
b = (r < b);
|
||||||
rp[i] = r;
|
rp[i] = r;
|
||||||
}
|
}
|
||||||
while (++i < n);
|
while (++i < n);
|
||||||
|
|
||||||
return b;
|
return b;
|
||||||
}
|
}*/
|
||||||
|
|
||||||
/* 位数相同的多精度数ap 加上多精度数bp,返回最后产生的进位 */
|
/* 位数相同的多精度数ap 加上多精度数bp,返回最后产生的进位 */
|
||||||
unsigned int
|
/*unsigned int
|
||||||
mpn_add_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n)
|
mpn_add_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n)
|
||||||
{
|
{
|
||||||
int i;
|
int i;
|
||||||
@ -210,10 +231,10 @@ mpn_add_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n)
|
|||||||
rp[i] = r;
|
rp[i] = r;
|
||||||
}
|
}
|
||||||
return cy;
|
return cy;
|
||||||
}
|
}*/
|
||||||
|
|
||||||
/*不同位数的多精度数相加,返回最后的进位*/
|
/*不同位数的多精度数相加,返回最后的进位*/
|
||||||
unsigned int
|
/*unsigned int
|
||||||
mpn_add (unsigned int *rp, unsigned int *ap, int an, unsigned int *bp, int bn)
|
mpn_add (unsigned int *rp, unsigned int *ap, int an, unsigned int *bp, int bn)
|
||||||
{
|
{
|
||||||
unsigned int cy;
|
unsigned int cy;
|
||||||
@ -222,9 +243,9 @@ mpn_add (unsigned int *rp, unsigned int *ap, int an, unsigned int *bp, int bn)
|
|||||||
if (an > bn)
|
if (an > bn)
|
||||||
cy = mpn_add_1 (rp + bn, ap + bn, an - bn, cy);
|
cy = mpn_add_1 (rp + bn, ap + bn, an - bn, cy);
|
||||||
return cy;
|
return cy;
|
||||||
}
|
}*/
|
||||||
|
|
||||||
unsigned int
|
/*unsigned int
|
||||||
mpn_sub_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
mpn_sub_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
||||||
{
|
{
|
||||||
int i;
|
int i;
|
||||||
@ -233,7 +254,7 @@ mpn_sub_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
|||||||
do
|
do
|
||||||
{
|
{
|
||||||
unsigned int a = ap[i];
|
unsigned int a = ap[i];
|
||||||
/* Carry out */
|
// Carry out
|
||||||
unsigned int cy = a < b;
|
unsigned int cy = a < b;
|
||||||
rp[i] = a - b;
|
rp[i] = a - b;
|
||||||
b = cy;
|
b = cy;
|
||||||
@ -241,9 +262,9 @@ mpn_sub_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
|||||||
while (++i < n);
|
while (++i < n);
|
||||||
|
|
||||||
return b;
|
return b;
|
||||||
}
|
}*/
|
||||||
|
|
||||||
unsigned int
|
/*unsigned int
|
||||||
mpn_sub_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n)
|
mpn_sub_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n)
|
||||||
{
|
{
|
||||||
int i;
|
int i;
|
||||||
@ -259,9 +280,9 @@ mpn_sub_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n)
|
|||||||
rp[i] = a - b;
|
rp[i] = a - b;
|
||||||
}
|
}
|
||||||
return cy;
|
return cy;
|
||||||
}
|
}*/
|
||||||
|
|
||||||
unsigned int
|
/*unsigned int
|
||||||
mpn_sub (unsigned int *rp, unsigned int *ap, int an, unsigned int *bp, int bn)
|
mpn_sub (unsigned int *rp, unsigned int *ap, int an, unsigned int *bp, int bn)
|
||||||
{
|
{
|
||||||
unsigned int cy;
|
unsigned int cy;
|
||||||
@ -270,18 +291,18 @@ mpn_sub (unsigned int *rp, unsigned int *ap, int an, unsigned int *bp, int bn)
|
|||||||
if (an > bn)
|
if (an > bn)
|
||||||
cy = mpn_sub_1 (rp + bn, ap + bn, an - bn, cy);
|
cy = mpn_sub_1 (rp + bn, ap + bn, an - bn, cy);
|
||||||
return cy;
|
return cy;
|
||||||
}
|
}*/
|
||||||
|
|
||||||
/* MPZ interface */
|
/* MPZ interface */
|
||||||
|
|
||||||
void
|
/*void
|
||||||
mpz_clear (mpz_t r)
|
mpz_clear (mpz_t r)
|
||||||
{
|
{
|
||||||
if (r->_mp_alloc)
|
if (r->_mp_alloc)
|
||||||
gmp_free_limbs (r->_mp_d, r->_mp_alloc);
|
gmp_free_limbs (r->_mp_d, r->_mp_alloc);
|
||||||
}
|
}*/
|
||||||
|
|
||||||
static unsigned int *
|
/*static unsigned int *
|
||||||
mpz_realloc (mpz_t r, int size)
|
mpz_realloc (mpz_t r, int size)
|
||||||
{
|
{
|
||||||
size = gmp_max (size, 1);
|
size = gmp_max (size, 1);
|
||||||
@ -296,31 +317,31 @@ mpz_realloc (mpz_t r, int size)
|
|||||||
r->_mp_size = 0;
|
r->_mp_size = 0;
|
||||||
|
|
||||||
return r->_mp_d;
|
return r->_mp_d;
|
||||||
}
|
}*/
|
||||||
|
|
||||||
/* Realloc for an mpz_t WHAT if it has less than NEEDED limbs. */
|
/* Realloc for an mpz_t WHAT if it has less than NEEDED limbs. */
|
||||||
unsigned int *mrz_realloc_if(mpz_t z,int n) {
|
/*unsigned int *mrz_realloc_if(mpz_t z,int n) {
|
||||||
return n > z->_mp_alloc ? mpz_realloc(z, n) : z->_mp_d;
|
return n > z->_mp_alloc ? mpz_realloc(z, n) : z->_mp_d;
|
||||||
}
|
}*/
|
||||||
|
|
||||||
/* MPZ comparisons and the like. */
|
/* MPZ comparisons and the like. */
|
||||||
int
|
/*int
|
||||||
mpz_sgn (const mpz_t u)
|
mpz_sgn (const mpz_t u)
|
||||||
{
|
{
|
||||||
return gmp_cmp (u->_mp_size, 0);
|
return gmp_cmp (u->_mp_size, 0);
|
||||||
}
|
}*/
|
||||||
|
|
||||||
void
|
/*void
|
||||||
mpz_swap (mpz_t u, mpz_t v)
|
mpz_swap (mpz_t u, mpz_t v)
|
||||||
{
|
{
|
||||||
int_swap (u->_mp_alloc, v->_mp_alloc);
|
int_swap (u->_mp_alloc, v->_mp_alloc);
|
||||||
unsigned int *_swap(u->_mp_d, v->_mp_d);
|
unsigned int *_swap(u->_mp_d, v->_mp_d);
|
||||||
int_swap (u->_mp_size, v->_mp_size);
|
int_swap (u->_mp_size, v->_mp_size);
|
||||||
}
|
}*/
|
||||||
|
|
||||||
/* MPZ addition and subtraction */
|
/* MPZ addition and subtraction */
|
||||||
|
|
||||||
static int
|
/*static int
|
||||||
mpz_abs_add (mpz_t r, const mpz_t a, const mpz_t b)
|
mpz_abs_add (mpz_t r, const mpz_t a, const mpz_t b)
|
||||||
{
|
{
|
||||||
int an = gmp_abs (a->_mp_size);
|
int an = gmp_abs (a->_mp_size);
|
||||||
@ -340,9 +361,9 @@ mpz_abs_add (mpz_t r, const mpz_t a, const mpz_t b)
|
|||||||
rp[an] = cy;
|
rp[an] = cy;
|
||||||
|
|
||||||
return an + cy;
|
return an + cy;
|
||||||
}
|
}*/
|
||||||
|
|
||||||
static int
|
/*static int
|
||||||
mpz_abs_sub (mpz_t r, const mpz_t a, const mpz_t b)
|
mpz_abs_sub (mpz_t r, const mpz_t a, const mpz_t b)
|
||||||
{
|
{
|
||||||
int an = gmp_abs (a->_mp_size);
|
int an = gmp_abs (a->_mp_size);
|
||||||
@ -365,9 +386,9 @@ mpz_abs_sub (mpz_t r, const mpz_t a, const mpz_t b)
|
|||||||
}
|
}
|
||||||
else
|
else
|
||||||
return 0;
|
return 0;
|
||||||
}
|
}*/
|
||||||
|
|
||||||
void
|
/*void
|
||||||
mpz_add (mpz_t r, const mpz_t a, const mpz_t b)
|
mpz_add (mpz_t r, const mpz_t a, const mpz_t b)
|
||||||
{
|
{
|
||||||
int rn;
|
int rn;
|
||||||
@ -378,9 +399,9 @@ mpz_add (mpz_t r, const mpz_t a, const mpz_t b)
|
|||||||
rn = mpz_abs_sub (r, a, b);
|
rn = mpz_abs_sub (r, a, b);
|
||||||
|
|
||||||
r->_mp_size = a->_mp_size >= 0 ? rn : - rn;
|
r->_mp_size = a->_mp_size >= 0 ? rn : - rn;
|
||||||
}
|
}*/
|
||||||
|
|
||||||
void
|
/*void
|
||||||
mpz_sub (mpz_t r, const mpz_t a, const mpz_t b)
|
mpz_sub (mpz_t r, const mpz_t a, const mpz_t b)
|
||||||
{
|
{
|
||||||
int rn;
|
int rn;
|
||||||
@ -391,4 +412,4 @@ mpz_sub (mpz_t r, const mpz_t a, const mpz_t b)
|
|||||||
rn = mpz_abs_add (r, a, b);
|
rn = mpz_abs_add (r, a, b);
|
||||||
|
|
||||||
r->_mp_size = a->_mp_size >= 0 ? rn : - rn;
|
r->_mp_size = a->_mp_size >= 0 ? rn : - rn;
|
||||||
}
|
}*/
|
||||||
|
@ -62,6 +62,9 @@ void mpz_set (mpz_t, const mpz_t);
|
|||||||
Extern Coq (Zabs : Z -> Z)
|
Extern Coq (Zabs : Z -> Z)
|
||||||
(Zmax : Z -> Z -> Z)
|
(Zmax : Z -> Z -> Z)
|
||||||
(mpd_store_Z : Z -> Z -> Z -> Z -> Assertion)
|
(mpd_store_Z : Z -> Z -> Z -> Z -> Assertion)
|
||||||
|
(mpd_store_Z_compact: Z -> Z -> Z -> Z -> Assertion)
|
||||||
(mpd_store_list : Z -> list Z -> Z -> Assertion)
|
(mpd_store_list : Z -> list Z -> Z -> Assertion)
|
||||||
(list_store_Z : list Z -> Z -> Prop)
|
(list_store_Z : list Z -> Z -> Prop)
|
||||||
|
(list_store_Z_compact: list Z -> Z -> Prop)
|
||||||
|
(last: list Z -> Z -> Z)
|
||||||
*/
|
*/
|
Reference in New Issue
Block a user