feat(cmp): Proved correctness of mpn_cmp.

This commit is contained in:
xiaoh105
2025-06-11 16:54:36 +08:00
parent 4c0b0e98fa
commit 36204b8877
7 changed files with 824 additions and 18 deletions

View File

@ -101,6 +101,17 @@ Proof.
reflexivity. reflexivity.
Qed. Qed.
Lemma list_app_single_l: forall (l: list Z) (a: Z),
([a] ++ l)%list = a :: l.
Proof.
intros.
induction l.
+ simpl; reflexivity.
+ rewrite list_app_cons.
simpl.
reflexivity.
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 |].

View File

@ -49,6 +49,19 @@ 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 |]. mpd_store_list ptr data cap && [| list_store_Z data n |] && [| size = Zlength data |].
Lemma list_store_Z_injection: forall l1 l2 n1 n2,
list_store_Z l1 n1 ->
list_store_Z l2 n2 ->
l1 = l2 -> n1 = n2.
Proof.
intros.
unfold list_store_Z in *.
destruct H, H0.
rewrite <-H, <-H0.
rewrite <-H1.
reflexivity.
Qed.
Lemma __list_within_bound_concat_r: forall (l1: list Z) (a: Z), Lemma __list_within_bound_concat_r: forall (l1: list Z) (a: Z),
list_within_bound l1 -> list_within_bound l1 ->
0 <= a < UINT_MOD -> 0 <= a < UINT_MOD ->
@ -350,6 +363,93 @@ Proof.
reflexivity. reflexivity.
Qed. Qed.
Lemma list_store_Z_cmp: forall l1 l2 n1 n2 i,
0 <= i < Zlength l1 ->
Zlength l1 = Zlength l2 ->
list_store_Z l1 n1 ->
list_store_Z 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.
assert (Zlength l1 = Z.of_nat (Datatypes.length l1)). { apply Zlength_correct. }
pose proof (sublist_split 0 (Zlength l1) i l1 ltac:(lia) ltac:(lia)).
assert (Zlength l2 = Z.of_nat (Datatypes.length l2)). { apply Zlength_correct. }
pose proof (sublist_split 0 (Zlength l2) i l2 ltac:(lia) ltac:(lia)).
clear H5 H7.
rewrite (sublist_self l1 (Zlength l1) ltac:(lia)) in H6.
rewrite (sublist_self l2 (Zlength l2) ltac:(lia)) in H8.
rewrite H6 in H1.
rewrite H8 in H2.
apply list_store_Z_split in H1, H2.
remember (Zlength l1) as n eqn:Hn.
assert (Zlength l2 = n). { lia. }
rewrite H5 in *.
rewrite Zlength_sublist0 in H1, H2; try lia.
destruct H1, H2.
remember (n1 mod UINT_MOD ^ i) as n1_lo eqn:Hn1_lo.
remember (n1 / UINT_MOD ^ i) as n1_hi eqn:Hn1_hi.
remember (n2 mod UINT_MOD ^ i) as n2_lo eqn:Hn2_lo.
remember (n2 / UINT_MOD ^ i) as n2_hi eqn:Hn2_hi.
remember (sublist 0 i l1) as l1_lo eqn:Hl1_lo.
remember (sublist i n l1) as l1_hi eqn:Hl1_hi.
remember (sublist 0 i l2) as l2_lo eqn:Hl2_lo.
remember (sublist i n l2) as l2_hi eqn:Hl2_hi.
assert (n1_lo - n2_lo < UINT_MOD ^ i). {
pose proof (list_store_Z_bound l1_lo n1_lo H1).
pose proof (list_store_Z_bound l2_lo n2_lo H2).
rewrite Hl1_lo, Zlength_sublist0 in H10; lia.
}
assert (n2_hi - n1_hi >= 1). {
assert (Zlength l1_hi >= 1 /\ Zlength l2_hi >= 1). {
pose proof (Zlength_sublist i n l1 ltac:(lia)).
pose proof (Zlength_sublist i n l2 ltac:(lia)).
rewrite <-Hl1_hi in H11.
rewrite <-Hl2_hi in H12.
lia.
}
destruct H11.
destruct l1_hi, l2_hi; try rewrite Zlength_nil in *; try lia.
unfold list_store_Z in H7, H9.
destruct H7, H9.
simpl in H7, H9.
assert (forall a (l0 l: list Z),
a :: l0 = sublist i n l ->
Zlength l = n ->
l0 = sublist (i + 1) n l /\ a = Znth i l 0
). {
intros.
assert (n = Z.of_nat(Datatypes.length l)). { rewrite <-Zlength_correct. lia. }
pose proof (sublist_split i n (i + 1) l ltac:(lia) ltac:(lia)).
rewrite (sublist_single i l 0) in H18; try lia.
rewrite <-H15 in H18.
rewrite Aux.list_app_single_l in H18.
injection H18; intros.
rewrite H19, H20.
split; reflexivity.
}
pose proof (H15 z0 l2_hi l2 Hl2_hi ltac:(lia)).
specialize (H15 z l1_hi l1 Hl1_hi ltac:(lia)).
destruct H15, H16.
rewrite H15, H17 in H7.
rewrite H16, H18 in H9.
rewrite <-H3 in H9.
lia.
}
pose proof (Zmod_eq_full n1 (UINT_MOD ^ i) ltac:(lia)).
pose proof (Zmod_eq_full n2 (UINT_MOD ^ i) ltac:(lia)).
rewrite <-Hn1_lo, <-Hn1_hi in H12.
rewrite <-Hn2_lo, <-Hn2_hi in H13.
assert (n2_hi * UINT_MOD ^ i - n1_hi * UINT_MOD ^ i >= UINT_MOD ^ i). {
rewrite <-Z.mul_sub_distr_r.
pose proof (Zmult_ge_compat_r (n2_hi - n1_hi) 1 (UINT_MOD ^ i) ltac:(lia) ltac:(lia)).
rewrite Z.mul_1_l in H14.
lia.
}
lia.
Qed.
End Internal. End Internal.
Record bigint_ent: Type := { Record bigint_ent: Type := {

View File

@ -9,8 +9,8 @@ Require Import Coq.Sorting.Permutation.
From AUXLib Require Import int_auto Axioms Feq Idents List_lemma VMap. From AUXLib Require Import int_auto Axioms Feq Idents List_lemma VMap.
Require Import SetsClass.SetsClass. Import SetsNotation. Require Import SetsClass.SetsClass. Import SetsNotation.
From SimpleC.SL Require Import Mem SeparationLogic. From SimpleC.SL Require Import Mem SeparationLogic.
Require Import GmpLib.GmpNumber. Import Internal.
Require Import Logic.LogicGenerator.demo932.Interface. Require Import Logic.LogicGenerator.demo932.Interface.
Require Import GmpLib.GmpNumber. Import Internal.
Local Open Scope Z_scope. Local Open Scope Z_scope.
Local Open Scope sets. Local Open Scope sets.
Local Open Scope string. Local Open Scope string.
@ -591,6 +591,562 @@ forall (cap2: Z) (l': (@list Z)) (l: (@list Z)) (i: Z) (n: Z) (d: Z) ,
** (store_uint_array d (i + 1 ) (app ((sublist (0) (i) (l))) ((cons (a) (nil)))) ) ** (store_uint_array d (i + 1 ) (app ((sublist (0) (i) (l))) ((cons (a) (nil)))) )
. .
(*----- Function mpn_cmp -----*)
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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (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 bp_pre (n_pre + 1 ) cap2 )
** ((( &( "n" ) )) # Int |-> n_pre)
** ((( &( "bp" ) )) # Ptr |-> bp_pre)
** ((( &( "ap" ) )) # Ptr |-> ap_pre)
|--
[| ((n_pre - 1 ) <= INT_MAX) |]
&& [| ((INT_MIN) <= (n_pre - 1 )) |]
.
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 ) |]
&& [| (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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& ((( &( "n" ) )) # Int |-> n)
** (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 bp_pre (n_pre + 1 ) cap2 )
** ((( &( "bp" ) )) # Ptr |-> bp_pre)
** ((( &( "ap" ) )) # Ptr |-> ap_pre)
|--
[| (0 <= INT_MAX) |]
&& [| ((INT_MIN) <= 0) |]
.
Definition mpn_cmp_safety_wit_3 :=
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) ,
[| ((Znth n l1 0) <= (Znth n l2 0)) |]
&& [| ((Znth n l1 0) <> (Znth n l2 0)) |]
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (store_uint_array bp_pre n_pre l2 )
** (store_uint_array ap_pre n_pre l1 )
** ((( &( "n" ) )) # Int |-> n)
** (store_undef_uint_array_rec ap_pre (n_pre + 1 ) cap1 )
** (store_undef_uint_array_rec bp_pre (n_pre + 1 ) cap2 )
** ((( &( "bp" ) )) # Ptr |-> bp_pre)
** ((( &( "ap" ) )) # Ptr |-> ap_pre)
|--
[| (1 <> (INT_MIN)) |]
.
Definition mpn_cmp_safety_wit_4 :=
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) ,
[| ((Znth n l1 0) <= (Znth n l2 0)) |]
&& [| ((Znth n l1 0) <> (Znth n l2 0)) |]
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (store_uint_array bp_pre n_pre l2 )
** (store_uint_array ap_pre n_pre l1 )
** ((( &( "n" ) )) # Int |-> n)
** (store_undef_uint_array_rec ap_pre (n_pre + 1 ) cap1 )
** (store_undef_uint_array_rec bp_pre (n_pre + 1 ) cap2 )
** ((( &( "bp" ) )) # Ptr |-> bp_pre)
** ((( &( "ap" ) )) # Ptr |-> ap_pre)
|--
[| (1 <= INT_MAX) |]
&& [| ((INT_MIN) <= 1) |]
.
Definition mpn_cmp_safety_wit_5 :=
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) ,
[| ((Znth n l1 0) > (Znth n l2 0)) |]
&& [| ((Znth n l1 0) <> (Znth n l2 0)) |]
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (store_uint_array bp_pre n_pre l2 )
** (store_uint_array ap_pre n_pre l1 )
** ((( &( "n" ) )) # Int |-> n)
** (store_undef_uint_array_rec ap_pre (n_pre + 1 ) cap1 )
** (store_undef_uint_array_rec bp_pre (n_pre + 1 ) cap2 )
** ((( &( "bp" ) )) # Ptr |-> bp_pre)
** ((( &( "ap" ) )) # Ptr |-> ap_pre)
|--
[| (1 <= INT_MAX) |]
&& [| ((INT_MIN) <= 1) |]
.
Definition mpn_cmp_safety_wit_6 :=
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) ,
[| ((Znth n l1 0) = (Znth n l2 0)) |]
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (store_uint_array bp_pre n_pre l2 )
** (store_uint_array ap_pre n_pre l1 )
** ((( &( "n" ) )) # Int |-> n)
** (store_undef_uint_array_rec ap_pre (n_pre + 1 ) cap1 )
** (store_undef_uint_array_rec bp_pre (n_pre + 1 ) cap2 )
** ((( &( "bp" ) )) # Ptr |-> bp_pre)
** ((( &( "ap" ) )) # Ptr |-> ap_pre)
|--
[| ((n - 1 ) <= INT_MAX) |]
&& [| ((INT_MIN) <= (n - 1 )) |]
.
Definition mpn_cmp_safety_wit_7 :=
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) ,
[| (n < 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& ((( &( "n" ) )) # Int |-> n)
** (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 bp_pre (n_pre + 1 ) cap2 )
** ((( &( "bp" ) )) # Ptr |-> bp_pre)
** ((( &( "ap" ) )) # Ptr |-> ap_pre)
|--
[| (0 <= INT_MAX) |]
&& [| ((INT_MIN) <= 0) |]
.
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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (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 bp_pre (n_pre + 1 ) cap2 )
|--
[| ((-1) <= (n_pre - 1 )) |]
&& [| ((n_pre - 1 ) < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (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 bp_pre (n_pre + 1 ) cap2 )
.
Definition mpn_cmp_entail_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) ,
[| ((Znth n l1 0) = (Znth n l2 0)) |]
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (store_uint_array bp_pre n_pre l2 )
** (store_uint_array ap_pre n_pre l1 )
** (store_undef_uint_array_rec ap_pre (n_pre + 1 ) cap1 )
** (store_undef_uint_array_rec bp_pre (n_pre + 1 ) cap2 )
|--
[| ((-1) <= (n - 1 )) |]
&& [| ((n - 1 ) < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (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 bp_pre (n_pre + 1 ) cap2 )
.
Definition mpn_cmp_return_wit_1_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)) (n: Z) ,
[| ((Znth n l1 0) <= (Znth n l2 0)) |]
&& [| ((Znth n l1 0) <> (Znth n l2 0)) |]
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (store_uint_array bp_pre n_pre l2 )
** (store_uint_array ap_pre n_pre l1 )
** (store_undef_uint_array_rec ap_pre (n_pre + 1 ) cap1 )
** (store_undef_uint_array_rec bp_pre (n_pre + 1 ) cap2 )
|--
([| (val1 < val2) |]
&& [| ((-1) = (-1)) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z 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 ))
||
([| (val1 > val2) |]
&& [| ((-1) = 1) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
.
Definition mpn_cmp_return_wit_1_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) ,
[| ((Znth n l1 0) > (Znth n l2 0)) |]
&& [| ((Znth n l1 0) <> (Znth n l2 0)) |]
&& [| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (store_uint_array bp_pre n_pre l2 )
** (store_uint_array ap_pre n_pre l1 )
** (store_undef_uint_array_rec ap_pre (n_pre + 1 ) cap1 )
** (store_undef_uint_array_rec bp_pre (n_pre + 1 ) cap2 )
|--
([| (val1 < val2) |]
&& [| (1 = (-1)) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z 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 ))
||
([| (val1 > val2) |]
&& [| (1 = 1) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z bp_pre val2 n_pre cap2 ))
.
Definition mpn_cmp_return_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) ,
[| (n < 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (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 bp_pre (n_pre + 1 ) cap2 )
|--
([| (val1 < val2) |]
&& [| (0 = (-1)) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z 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 ))
||
([| (val1 > val2) |]
&& [| (0 = 1) |]
&& (mpd_store_Z ap_pre val1 n_pre cap1 )
** (mpd_store_Z 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) |]
&& [| (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 )
|--
[| (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 )
.
Definition mpn_cmp_partial_solve_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) ,
[| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (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 bp_pre (n_pre + 1 ) cap2 )
|--
[| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (((ap_pre + (n * sizeof(UINT) ) )) # UInt |-> (Znth n l1 0))
** (store_uint_array_missing_i_rec ap_pre n 0 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 bp_pre (n_pre + 1 ) cap2 )
.
Definition mpn_cmp_partial_solve_wit_3 :=
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) ,
[| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (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 bp_pre (n_pre + 1 ) cap2 )
|--
[| (n >= 0) |]
&& [| ((-1) <= n) |]
&& [| (n < n_pre) |]
&& [| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z 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 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (l2))) |]
&& [| (0 < n_pre) |]
&& [| (n_pre <= cap1) |]
&& [| (n_pre <= cap2) |]
&& [| (cap1 <= 100000000) |]
&& [| (cap2 <= 100000000) |]
&& (((bp_pre + (n * sizeof(UINT) ) )) # UInt |-> (Znth n l2 0))
** (store_uint_array_missing_i_rec bp_pre n 0 n_pre l2 )
** (store_uint_array ap_pre n_pre l1 )
** (store_undef_uint_array_rec ap_pre (n_pre + 1 ) cap1 )
** (store_undef_uint_array_rec bp_pre (n_pre + 1 ) cap2 )
.
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 )
|--
EX (l2: (@list Z)) (l1: (@list Z)) ,
[| (list_store_Z l1 val1 ) |]
&& [| (list_store_Z l2 val2 ) |]
&& [| (n_pre = (Zlength (l1))) |]
&& [| (n_pre = (Zlength (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 bp_pre (n_pre + 1 ) cap2 )
.
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.
@ -620,5 +1176,21 @@ Axiom proof_of_mpn_copyi_partial_solve_wit_5 : mpn_copyi_partial_solve_wit_5.
Axiom proof_of_mpn_copyi_which_implies_wit_1 : mpn_copyi_which_implies_wit_1. Axiom proof_of_mpn_copyi_which_implies_wit_1 : mpn_copyi_which_implies_wit_1.
Axiom proof_of_mpn_copyi_which_implies_wit_2 : mpn_copyi_which_implies_wit_2. Axiom proof_of_mpn_copyi_which_implies_wit_2 : mpn_copyi_which_implies_wit_2.
Axiom proof_of_mpn_copyi_which_implies_wit_3 : mpn_copyi_which_implies_wit_3. Axiom proof_of_mpn_copyi_which_implies_wit_3 : mpn_copyi_which_implies_wit_3.
Axiom proof_of_mpn_cmp_safety_wit_1 : mpn_cmp_safety_wit_1.
Axiom proof_of_mpn_cmp_safety_wit_2 : mpn_cmp_safety_wit_2.
Axiom proof_of_mpn_cmp_safety_wit_3 : mpn_cmp_safety_wit_3.
Axiom proof_of_mpn_cmp_safety_wit_4 : mpn_cmp_safety_wit_4.
Axiom proof_of_mpn_cmp_safety_wit_5 : mpn_cmp_safety_wit_5.
Axiom proof_of_mpn_cmp_safety_wit_6 : mpn_cmp_safety_wit_6.
Axiom proof_of_mpn_cmp_safety_wit_7 : mpn_cmp_safety_wit_7.
Axiom proof_of_mpn_cmp_entail_wit_1 : mpn_cmp_entail_wit_1.
Axiom proof_of_mpn_cmp_entail_wit_2 : mpn_cmp_entail_wit_2.
Axiom proof_of_mpn_cmp_return_wit_1_1 : mpn_cmp_return_wit_1_1.
Axiom proof_of_mpn_cmp_return_wit_1_2 : mpn_cmp_return_wit_1_2.
Axiom proof_of_mpn_cmp_return_wit_2 : mpn_cmp_return_wit_2.
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.
End VC_Correct. End VC_Correct.

View File

@ -1,6 +1,6 @@
From Require Import gmp_goal gmp_proof_auto gmp_proof_manual. From Require Import gmp_goal gmp_proof_auto gmp_proof_manual_tmp.
Module VC_Correctness : VC_Correct. Module VC_Correctness : VC_Correct.
Include gmp_proof_auto. Include gmp_proof_auto.
Include gmp_proof_manual. Include gmp_proof_manual_tmp.
End VC_Correctness. End VC_Correctness.

View File

@ -63,3 +63,33 @@ Proof. Admitted.
Lemma proof_of_mpn_copyi_partial_solve_wit_5 : mpn_copyi_partial_solve_wit_5. Lemma proof_of_mpn_copyi_partial_solve_wit_5 : mpn_copyi_partial_solve_wit_5.
Proof. Admitted. Proof. Admitted.
Lemma proof_of_mpn_cmp_safety_wit_1 : mpn_cmp_safety_wit_1.
Proof. Admitted.
Lemma proof_of_mpn_cmp_safety_wit_2 : mpn_cmp_safety_wit_2.
Proof. Admitted.
Lemma proof_of_mpn_cmp_safety_wit_3 : mpn_cmp_safety_wit_3.
Proof. Admitted.
Lemma proof_of_mpn_cmp_safety_wit_4 : mpn_cmp_safety_wit_4.
Proof. Admitted.
Lemma proof_of_mpn_cmp_safety_wit_5 : mpn_cmp_safety_wit_5.
Proof. Admitted.
Lemma proof_of_mpn_cmp_safety_wit_6 : mpn_cmp_safety_wit_6.
Proof. Admitted.
Lemma proof_of_mpn_cmp_safety_wit_7 : mpn_cmp_safety_wit_7.
Proof. Admitted.
Lemma proof_of_mpn_cmp_partial_solve_wit_1 : mpn_cmp_partial_solve_wit_1.
Proof. Admitted.
Lemma proof_of_mpn_cmp_partial_solve_wit_2 : mpn_cmp_partial_solve_wit_2.
Proof. Admitted.
Lemma proof_of_mpn_cmp_partial_solve_wit_3 : mpn_cmp_partial_solve_wit_3.
Proof. Admitted.

View File

@ -150,3 +150,94 @@ Proof.
rewrite <-H4. rewrite <-H4.
sep_apply store_uint_array_rec_tail_merge; [ reflexivity | lia ]. sep_apply store_uint_array_rec_tail_merge; [ reflexivity | lia ].
Qed. Qed.
Lemma proof_of_mpn_cmp_safety_wit_1 : mpn_cmp_safety_wit_1.
Proof.
pre_process.
Qed.
Lemma proof_of_mpn_cmp_entail_wit_1 : mpn_cmp_entail_wit_1.
Proof.
pre_process.
entailer!.
assert (n_pre - 1 + 1 = n_pre). { lia. }
rewrite H8; clear H8.
pose proof (Zlength_sublist n_pre n_pre l1 ltac:(lia)).
pose proof (Zlength_nil_inv (sublist n_pre n_pre l1) ltac:(lia)).
rewrite H9.
pose proof (Zlength_sublist n_pre n_pre l2 ltac:(lia)).
pose proof (Zlength_nil_inv (sublist n_pre n_pre l2) ltac:(lia)).
rewrite H11.
reflexivity.
Qed.
Lemma proof_of_mpn_cmp_entail_wit_2 : mpn_cmp_entail_wit_2.
Proof.
pre_process.
entailer!; try lia.
assert (n - 1 + 1 = n). { lia. }
rewrite H17; clear H17.
assert (n_pre <= Z.of_nat (Datatypes.length l1)). {
rewrite <-Zlength_correct.
lia.
}
assert (n_pre <= Z.of_nat (Datatypes.length l2)). {
rewrite <-Zlength_correct.
lia.
}
rewrite (sublist_split n n_pre (n + 1) l1 ltac:(lia) ltac:(lia)).
rewrite (sublist_split n n_pre (n + 1) l2 ltac:(lia) ltac:(lia)).
rewrite (sublist_single n l1 0 ltac:(lia)).
rewrite (sublist_single n l2 0 ltac:(lia)).
rewrite H.
rewrite H7.
reflexivity.
Qed.
Lemma proof_of_mpn_cmp_return_wit_1_1 : mpn_cmp_return_wit_1_1.
Proof.
pre_process.
entailer!.
repeat rewrite <-derivable1_orp_intros1.
entailer!.
+ unfold mpd_store_Z.
Exists l1 l2.
unfold mpd_store_list.
entailer!.
rewrite <-H6, <-H7.
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).
- rewrite <-H6, <-H7.
tauto.
- lia.
Qed.
Lemma proof_of_mpn_cmp_return_wit_1_2 : mpn_cmp_return_wit_1_2.
Proof.
pre_process.
rewrite <-derivable1_orp_intros2.
entailer!.
+ unfold mpd_store_Z, 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).
rewrite <-H6, <-H7 in H18.
rewrite H8 in H18.
specialize (H18 ltac:(reflexivity) ltac:(lia)).
lia.
Qed.
Lemma proof_of_mpn_cmp_which_implies_wit_1 : mpn_cmp_which_implies_wit_1.
Proof.
pre_process.
unfold mpd_store_Z.
Intros l1 l2.
Exists l2 l1.
unfold mpd_store_list.
entailer!.
rewrite <-H4, <-H6.
entailer!.
Qed.

View File

@ -104,44 +104,46 @@ mpn_cmp (unsigned int *ap, unsigned int *bp, int n)
Require Require
mpd_store_Z(ap, val1, n, cap1) * mpd_store_Z(ap, val1, n, cap1) *
mpd_store_Z(bp, val2, n, cap2) && mpd_store_Z(bp, val2, n, cap2) &&
n <= cap1 && n <= cap2 0 < n && n <= cap1 && n <= cap2 &&
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(bp@pre, val2, n@pre, cap2)
*/ */
{ {
/*@ /*@
mpd_store_Z(ap, val1, n, cap1) * mpd_store_Z(bp, val2, n, cap2) mpd_store_Z(ap@pre, val1, n@pre, cap1) * mpd_store_Z(bp@pre, val2, n@pre, cap2)
which implies which implies
exists l1 l2, exists l1 l2,
mpd_store_list(ap, l1, cap1) * mpd_store_list(bp, l2, cap2) && 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(bp@pre, n@pre + 1, cap2) &&
list_store_Z(l1, val1) && list_store_Z(l2, val2) && list_store_Z(l1, val1) && list_store_Z(l2, val2) &&
n == Zlength(l1) && n == Zlength(l2) n@pre == Zlength(l1) && n@pre == Zlength(l2)
*/ */
/*@ /*@
Given l1 l2 Given l1 l2
*/ */
--n; --n;
/*@Inv /*@Inv
mpd_store_list(ap, l1, cap1) * mpd_store_list(bp, l2, cap2) && -1 <= n && n < n@pre &&
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(bp@pre, n@pre + 1, cap2) &&
list_store_Z(l1, val1) && list_store_Z(l2, val2) && list_store_Z(l1, val1) && list_store_Z(l2, val2) &&
n@pre == Zlength(l1) && n@pre == Zlength(l2) && n@pre == Zlength(l1) && n@pre == Zlength(l2) &&
sublist(n, n@pre, l1) == sublist(n, n@pre, l2) sublist(n + 1, n@pre, l1) == sublist(n + 1, n@pre, l2)
*/ */
while (n >= 0) while (n >= 0)
{ {
/*@
mpd_store_list(ap, l1, cap1) * mpd_store_list(bp, l2, cap2)
which implies
store_uint_array(ap, n, l1) * store_uint_array(bp, n, l2) *
store_undef_uint_array(ap, n + 1, cap1) * store_uint_array(bp, n + 1, cap2) &&
*/
if (ap[n] != bp[n]) if (ap[n] != bp[n])
return ap[n] > bp[n] ? 1 : -1; return ap[n] > bp[n] ? 1 : -1;
--n; --n;
} }
// Note: The parser cannot parse "--n" in loop so we paraphrased it. // Note: The parser cannot parse "--n" in loop condition so we paraphrased it.
/* /*
while (--n >= 0) while (--n >= 0)
{ {