finish refactoring
This commit is contained in:
@ -89,9 +89,9 @@ Proof.
|
|||||||
reflexivity.
|
reflexivity.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Lemma list_store_Z_compact_reverse_injection: forall l1 l2 n1 n2,
|
Lemma list_store_Z_reverse_injection: forall l1 l2 n1 n2,
|
||||||
list_store_Z_compact l1 n1 ->
|
list_store_Z l1 n1 ->
|
||||||
list_store_Z_compact l2 n2 ->
|
list_store_Z l2 n2 ->
|
||||||
n1 = n2 -> l1 = l2.
|
n1 = n2 -> l1 = l2.
|
||||||
Proof. Admitted.
|
Proof. Admitted.
|
||||||
|
|
||||||
@ -284,7 +284,7 @@ Qed.
|
|||||||
|
|
||||||
Lemma list_store_Z_list_append: forall (l: list Z) (i: Z) (val_prefix: Z) (val_full: Z),
|
Lemma list_store_Z_list_append: forall (l: list Z) (i: Z) (val_prefix: Z) (val_full: Z),
|
||||||
0 <= i < Zlength l ->
|
0 <= i < Zlength l ->
|
||||||
list_store_Z_compact l val_full ->
|
list_store_Z l val_full ->
|
||||||
list_store_Z (sublist 0 i l) val_prefix ->
|
list_store_Z (sublist 0 i l) val_prefix ->
|
||||||
list_store_Z (sublist 0 (i+1) l) (val_prefix + Znth i l 0 * UINT_MOD ^ i).
|
list_store_Z (sublist 0 (i+1) l) (val_prefix + Znth i l 0 * UINT_MOD ^ i).
|
||||||
Proof.
|
Proof.
|
||||||
@ -313,7 +313,7 @@ Proof.
|
|||||||
split; try tauto.
|
split; try tauto.
|
||||||
apply list_within_bound_Znth.
|
apply list_within_bound_Znth.
|
||||||
lia.
|
lia.
|
||||||
unfold list_store_Z_compact in H0.
|
unfold list_store_Z in H0.
|
||||||
tauto.
|
tauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
|
File diff suppressed because it is too large
Load Diff
@ -469,7 +469,7 @@ Proof.
|
|||||||
assert (0 <= Znth i l_3 0 < 4294967296). {
|
assert (0 <= Znth i l_3 0 < 4294967296). {
|
||||||
assert (l_2=l_3).
|
assert (l_2=l_3).
|
||||||
{
|
{
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_2 l_3 val val).
|
pose proof (list_store_Z_reverse_injection l_2 l_3 val val).
|
||||||
apply H30 in H9; try tauto.
|
apply H30 in H9; try tauto.
|
||||||
}
|
}
|
||||||
assert (i < Zlength l_3). {
|
assert (i < Zlength l_3). {
|
||||||
@ -477,7 +477,7 @@ Proof.
|
|||||||
rewrite H17.
|
rewrite H17.
|
||||||
tauto.
|
tauto.
|
||||||
}
|
}
|
||||||
unfold list_store_Z_compact in H9.
|
unfold list_store_Z in H9.
|
||||||
apply list_within_bound_Znth.
|
apply list_within_bound_Znth.
|
||||||
lia.
|
lia.
|
||||||
tauto.
|
tauto.
|
||||||
@ -505,7 +505,7 @@ Proof.
|
|||||||
lia.
|
lia.
|
||||||
+ assert (l_2=l_3).
|
+ assert (l_2=l_3).
|
||||||
{
|
{
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_2 l_3 val val).
|
pose proof (list_store_Z_reverse_injection l_2 l_3 val val).
|
||||||
apply H28 in H9; try tauto.
|
apply H28 in H9; try tauto.
|
||||||
}
|
}
|
||||||
|
|
||||||
@ -539,7 +539,7 @@ Proof.
|
|||||||
lia.
|
lia.
|
||||||
apply list_within_bound_Znth.
|
apply list_within_bound_Znth.
|
||||||
lia.
|
lia.
|
||||||
unfold list_store_Z_compact in H9.
|
unfold list_store_Z in H9.
|
||||||
tauto.
|
tauto.
|
||||||
- pose proof (Zlength_sublist0 i l'_2).
|
- pose proof (Zlength_sublist0 i l'_2).
|
||||||
lia.
|
lia.
|
||||||
@ -585,7 +585,7 @@ Proof.
|
|||||||
assert (0 <= Znth i l_3 0 < 4294967296). {
|
assert (0 <= Znth i l_3 0 < 4294967296). {
|
||||||
assert (l_2=l_3).
|
assert (l_2=l_3).
|
||||||
{
|
{
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_2 l_3 val val).
|
pose proof (list_store_Z_reverse_injection l_2 l_3 val val).
|
||||||
apply H30 in H9; try tauto.
|
apply H30 in H9; try tauto.
|
||||||
}
|
}
|
||||||
assert (i < Zlength l_3). {
|
assert (i < Zlength l_3). {
|
||||||
@ -593,7 +593,7 @@ Proof.
|
|||||||
rewrite H17.
|
rewrite H17.
|
||||||
tauto.
|
tauto.
|
||||||
}
|
}
|
||||||
unfold list_store_Z_compact in H9.
|
unfold list_store_Z in H9.
|
||||||
apply list_within_bound_Znth.
|
apply list_within_bound_Znth.
|
||||||
lia.
|
lia.
|
||||||
tauto.
|
tauto.
|
||||||
@ -621,7 +621,7 @@ Proof.
|
|||||||
lia.
|
lia.
|
||||||
+ assert (l_2=l_3).
|
+ assert (l_2=l_3).
|
||||||
{
|
{
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_2 l_3 val val).
|
pose proof (list_store_Z_reverse_injection l_2 l_3 val val).
|
||||||
apply H28 in H9; try tauto.
|
apply H28 in H9; try tauto.
|
||||||
}
|
}
|
||||||
|
|
||||||
@ -655,7 +655,7 @@ Proof.
|
|||||||
lia.
|
lia.
|
||||||
apply list_within_bound_Znth.
|
apply list_within_bound_Znth.
|
||||||
lia.
|
lia.
|
||||||
unfold list_store_Z_compact in H9.
|
unfold list_store_Z in H9.
|
||||||
tauto.
|
tauto.
|
||||||
- pose proof (Zlength_sublist0 i l'_2).
|
- pose proof (Zlength_sublist0 i l'_2).
|
||||||
lia.
|
lia.
|
||||||
@ -664,10 +664,10 @@ Qed.
|
|||||||
Lemma proof_of_mpn_add_1_return_wit_1 : mpn_add_1_return_wit_1.
|
Lemma proof_of_mpn_add_1_return_wit_1 : mpn_add_1_return_wit_1.
|
||||||
Proof.
|
Proof.
|
||||||
pre_process.
|
pre_process.
|
||||||
unfold mpd_store_Z_compact.
|
unfold mpd_store_Z.
|
||||||
unfold mpd_store_list.
|
unfold mpd_store_list.
|
||||||
Exists val2.
|
Exists val2.
|
||||||
pose proof (list_store_Z_compact_reverse_injection l l_2 val val).
|
pose proof (list_store_Z_reverse_injection l l_2 val val).
|
||||||
apply H19 in H2; try tauto.
|
apply H19 in H2; try tauto.
|
||||||
rewrite <-H2 in H10.
|
rewrite <-H2 in H10.
|
||||||
assert (i = n_pre) by lia.
|
assert (i = n_pre) by lia.
|
||||||
@ -675,15 +675,14 @@ Proof.
|
|||||||
rewrite <- H10 in H4.
|
rewrite <- H10 in H4.
|
||||||
rewrite (sublist_self l (Zlength l)) in H4; try tauto.
|
rewrite (sublist_self l (Zlength l)) in H4; try tauto.
|
||||||
rewrite <-H2 in H12.
|
rewrite <-H2 in H12.
|
||||||
assert (list_store_Z l val). { apply list_store_Z_compact_to_normal. tauto. }
|
|
||||||
pose proof (list_store_Z_injection l l val1 val).
|
pose proof (list_store_Z_injection l l val1 val).
|
||||||
apply H22 in H4; try tauto.
|
apply H21 in H4; try tauto.
|
||||||
rewrite H4 in H6.
|
rewrite H4 in H6.
|
||||||
entailer!.
|
entailer!.
|
||||||
Exists l.
|
Exists l.
|
||||||
entailer!.
|
entailer!.
|
||||||
entailer!; try rewrite H20; try tauto.
|
entailer!; try rewrite H20; try tauto.
|
||||||
- rewrite H10.
|
rewrite H10.
|
||||||
entailer!.
|
entailer!.
|
||||||
unfold mpd_store_Z.
|
unfold mpd_store_Z.
|
||||||
unfold mpd_store_list.
|
unfold mpd_store_list.
|
||||||
@ -694,13 +693,15 @@ Proof.
|
|||||||
rewrite H20.
|
rewrite H20.
|
||||||
entailer!.
|
entailer!.
|
||||||
apply store_uint_array_rec_def2undef.
|
apply store_uint_array_rec_def2undef.
|
||||||
- rewrite <- H20. tauto.
|
assert (Zlength l' = n_pre) by lia.
|
||||||
|
rewrite <- H7.
|
||||||
|
tauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Lemma proof_of_mpn_add_1_which_implies_wit_1 : mpn_add_1_which_implies_wit_1.
|
Lemma proof_of_mpn_add_1_which_implies_wit_1 : mpn_add_1_which_implies_wit_1.
|
||||||
Proof.
|
Proof.
|
||||||
pre_process.
|
pre_process.
|
||||||
unfold mpd_store_Z_compact.
|
unfold mpd_store_Z.
|
||||||
Intros l.
|
Intros l.
|
||||||
Exists l.
|
Exists l.
|
||||||
unfold mpd_store_list.
|
unfold mpd_store_list.
|
||||||
@ -838,14 +839,14 @@ Proof.
|
|||||||
pre_process.
|
pre_process.
|
||||||
rewrite replace_Znth_app_r.
|
rewrite replace_Znth_app_r.
|
||||||
assert (l_a_3 = l_a_2). {
|
assert (l_a_3 = l_a_2). {
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_a_3 l_a_2 val_a val_a).
|
pose proof (list_store_Z_reverse_injection l_a_3 l_a_2 val_a val_a).
|
||||||
specialize (H37 H13 H28).
|
specialize (H37 H13 H28).
|
||||||
apply H37.
|
apply H37.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
}
|
}
|
||||||
subst l_a_3.
|
subst l_a_3.
|
||||||
assert (l_b_3 = l_b_2). {
|
assert (l_b_3 = l_b_2). {
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_b_3 l_b_2 val_b val_b).
|
pose proof (list_store_Z_reverse_injection l_b_3 l_b_2 val_b val_b).
|
||||||
specialize (H37 H14 H24).
|
specialize (H37 H14 H24).
|
||||||
apply H37.
|
apply H37.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
@ -882,7 +883,7 @@ Proof.
|
|||||||
assert (2 ^ 32 = 4294967296). { nia. }
|
assert (2 ^ 32 = 4294967296). { nia. }
|
||||||
rewrite H37 in H4, H3; clear H37.
|
rewrite H37 in H4, H3; clear H37.
|
||||||
apply Z_mod_3add_carry10; try lia; try tauto;
|
apply Z_mod_3add_carry10; try lia; try tauto;
|
||||||
try unfold list_store_Z_compact in H13, H14;
|
try unfold list_store_Z in H13, H14;
|
||||||
try apply list_within_bound_Znth;
|
try apply list_within_bound_Znth;
|
||||||
try lia;
|
try lia;
|
||||||
try tauto.
|
try tauto.
|
||||||
@ -936,14 +937,14 @@ Proof.
|
|||||||
pre_process.
|
pre_process.
|
||||||
rewrite replace_Znth_app_r.
|
rewrite replace_Znth_app_r.
|
||||||
assert (l_a_3 = l_a_2). {
|
assert (l_a_3 = l_a_2). {
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_a_3 l_a_2 val_a val_a).
|
pose proof (list_store_Z_reverse_injection l_a_3 l_a_2 val_a val_a).
|
||||||
specialize (H37 H13 H28).
|
specialize (H37 H13 H28).
|
||||||
apply H37.
|
apply H37.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
}
|
}
|
||||||
subst l_a_3.
|
subst l_a_3.
|
||||||
assert (l_b_3 = l_b_2). {
|
assert (l_b_3 = l_b_2). {
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_b_3 l_b_2 val_b val_b).
|
pose proof (list_store_Z_reverse_injection l_b_3 l_b_2 val_b val_b).
|
||||||
specialize (H37 H14 H24).
|
specialize (H37 H14 H24).
|
||||||
apply H37.
|
apply H37.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
@ -980,7 +981,7 @@ Proof.
|
|||||||
assert (2 ^ 32 = 4294967296). { nia. }
|
assert (2 ^ 32 = 4294967296). { nia. }
|
||||||
rewrite H37 in H4, H3; clear H37.
|
rewrite H37 in H4, H3; clear H37.
|
||||||
apply Z_mod_3add_carry11; try lia; try tauto;
|
apply Z_mod_3add_carry11; try lia; try tauto;
|
||||||
try unfold list_store_Z_compact in H13, H14;
|
try unfold list_store_Z in H13, H14;
|
||||||
try apply list_within_bound_Znth;
|
try apply list_within_bound_Znth;
|
||||||
try lia;
|
try lia;
|
||||||
try tauto.
|
try tauto.
|
||||||
@ -1034,14 +1035,14 @@ Proof.
|
|||||||
pre_process.
|
pre_process.
|
||||||
rewrite replace_Znth_app_r.
|
rewrite replace_Znth_app_r.
|
||||||
assert (l_a_3 = l_a_2). {
|
assert (l_a_3 = l_a_2). {
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_a_3 l_a_2 val_a val_a).
|
pose proof (list_store_Z_reverse_injection l_a_3 l_a_2 val_a val_a).
|
||||||
specialize (H37 H13 H28).
|
specialize (H37 H13 H28).
|
||||||
apply H37.
|
apply H37.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
}
|
}
|
||||||
subst l_a_3.
|
subst l_a_3.
|
||||||
assert (l_b_3 = l_b_2). {
|
assert (l_b_3 = l_b_2). {
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_b_3 l_b_2 val_b val_b).
|
pose proof (list_store_Z_reverse_injection l_b_3 l_b_2 val_b val_b).
|
||||||
specialize (H37 H14 H24).
|
specialize (H37 H14 H24).
|
||||||
apply H37.
|
apply H37.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
@ -1078,7 +1079,7 @@ Proof.
|
|||||||
assert (2 ^ 32 = 4294967296). { nia. }
|
assert (2 ^ 32 = 4294967296). { nia. }
|
||||||
rewrite H37 in H4, H3; clear H37.
|
rewrite H37 in H4, H3; clear H37.
|
||||||
apply Z_mod_3add_carry00; try lia; try tauto;
|
apply Z_mod_3add_carry00; try lia; try tauto;
|
||||||
try unfold list_store_Z_compact in H13, H14;
|
try unfold list_store_Z in H13, H14;
|
||||||
try apply list_within_bound_Znth;
|
try apply list_within_bound_Znth;
|
||||||
try lia;
|
try lia;
|
||||||
try tauto.
|
try tauto.
|
||||||
@ -1132,14 +1133,14 @@ Proof.
|
|||||||
pre_process.
|
pre_process.
|
||||||
rewrite replace_Znth_app_r.
|
rewrite replace_Znth_app_r.
|
||||||
assert (l_a_3 = l_a_2). {
|
assert (l_a_3 = l_a_2). {
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_a_3 l_a_2 val_a val_a).
|
pose proof (list_store_Z_reverse_injection l_a_3 l_a_2 val_a val_a).
|
||||||
specialize (H37 H13 H28).
|
specialize (H37 H13 H28).
|
||||||
apply H37.
|
apply H37.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
}
|
}
|
||||||
subst l_a_3.
|
subst l_a_3.
|
||||||
assert (l_b_3 = l_b_2). {
|
assert (l_b_3 = l_b_2). {
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_b_3 l_b_2 val_b val_b).
|
pose proof (list_store_Z_reverse_injection l_b_3 l_b_2 val_b val_b).
|
||||||
specialize (H37 H14 H24).
|
specialize (H37 H14 H24).
|
||||||
apply H37.
|
apply H37.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
@ -1176,7 +1177,7 @@ Proof.
|
|||||||
assert (2 ^ 32 = 4294967296). { nia. }
|
assert (2 ^ 32 = 4294967296). { nia. }
|
||||||
rewrite H37 in H4, H3; clear H37.
|
rewrite H37 in H4, H3; clear H37.
|
||||||
apply Z_mod_3add_carry01; try lia; try tauto;
|
apply Z_mod_3add_carry01; try lia; try tauto;
|
||||||
try unfold list_store_Z_compact in H13, H14;
|
try unfold list_store_Z in H13, H14;
|
||||||
try apply list_within_bound_Znth;
|
try apply list_within_bound_Znth;
|
||||||
try lia;
|
try lia;
|
||||||
try tauto.
|
try tauto.
|
||||||
@ -1229,14 +1230,14 @@ Lemma proof_of_mpn_add_n_return_wit_1 : mpn_add_n_return_wit_1.
|
|||||||
Proof.
|
Proof.
|
||||||
pre_process.
|
pre_process.
|
||||||
assert (l_a_2 = l_a). {
|
assert (l_a_2 = l_a). {
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_a_2 l_a val_a val_a).
|
pose proof (list_store_Z_reverse_injection l_a_2 l_a val_a val_a).
|
||||||
specialize (H29 H20 H5).
|
specialize (H29 H20 H5).
|
||||||
apply H29.
|
apply H29.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
}
|
}
|
||||||
subst l_a_2.
|
subst l_a_2.
|
||||||
assert (l_b_2 = l_b). {
|
assert (l_b_2 = l_b). {
|
||||||
pose proof (list_store_Z_compact_reverse_injection l_b_2 l_b val_b val_b).
|
pose proof (list_store_Z_reverse_injection l_b_2 l_b val_b val_b).
|
||||||
specialize (H29 H16 H6).
|
specialize (H29 H16 H6).
|
||||||
apply H29.
|
apply H29.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
@ -1244,12 +1245,12 @@ Proof.
|
|||||||
subst l_b_2.
|
subst l_b_2.
|
||||||
assert (i = n_pre) by lia.
|
assert (i = n_pre) by lia.
|
||||||
Exists val_r_prefix.
|
Exists val_r_prefix.
|
||||||
unfold mpd_store_Z_compact.
|
unfold mpd_store_Z.
|
||||||
unfold mpd_store_list.
|
unfold mpd_store_list.
|
||||||
Exists l_a.
|
Exists l_a.
|
||||||
Exists l_b.
|
Exists l_b.
|
||||||
entailer!.
|
entailer!.
|
||||||
+ rewrite H14.
|
rewrite H14.
|
||||||
rewrite H18.
|
rewrite H18.
|
||||||
entailer!.
|
entailer!.
|
||||||
unfold mpd_store_Z.
|
unfold mpd_store_Z.
|
||||||
@ -1261,38 +1262,33 @@ Proof.
|
|||||||
rewrite H10.
|
rewrite H10.
|
||||||
entailer!.
|
entailer!.
|
||||||
apply store_uint_array_rec_def2undef.
|
apply store_uint_array_rec_def2undef.
|
||||||
+ rewrite <- H29.
|
rewrite <- H29.
|
||||||
assert (val_a_prefix = val_a). {
|
assert (val_a_prefix = val_a). {
|
||||||
assert (i = Zlength l_a). {
|
rewrite <-H18 in H7.
|
||||||
lia.
|
|
||||||
}
|
|
||||||
rewrite H30 in H7.
|
|
||||||
rewrite sublist_self in H7.
|
rewrite sublist_self in H7.
|
||||||
unfold list_store_Z_compact in H5.
|
unfold list_store_Z in H5.
|
||||||
unfold list_store_Z in H7.
|
unfold list_store_Z in H7.
|
||||||
lia.
|
lia.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
}
|
}
|
||||||
rewrite <- H30; clear H30.
|
rewrite <- H30; clear H30.
|
||||||
assert (val_b_prefix = val_b). {
|
assert (val_b_prefix = val_b). {
|
||||||
assert (i = Zlength l_b). {
|
rewrite <-H14 in H8.
|
||||||
lia.
|
|
||||||
}
|
|
||||||
rewrite H30 in H8.
|
|
||||||
rewrite sublist_self in H8.
|
rewrite sublist_self in H8.
|
||||||
unfold list_store_Z_compact in H6.
|
unfold list_store_Z in H6.
|
||||||
unfold list_store_Z in H8.
|
unfold list_store_Z in H8.
|
||||||
lia.
|
lia.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
}
|
}
|
||||||
rewrite <- H30; clear H30.
|
rewrite <- H30; clear H30.
|
||||||
|
rewrite H29.
|
||||||
tauto.
|
tauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Lemma proof_of_mpn_add_n_which_implies_wit_1 : mpn_add_n_which_implies_wit_1.
|
Lemma proof_of_mpn_add_n_which_implies_wit_1 : mpn_add_n_which_implies_wit_1.
|
||||||
Proof.
|
Proof.
|
||||||
pre_process.
|
pre_process.
|
||||||
unfold mpd_store_Z_compact.
|
unfold mpd_store_Z.
|
||||||
Intros l.
|
Intros l.
|
||||||
Exists l.
|
Exists l.
|
||||||
unfold mpd_store_list.
|
unfold mpd_store_list.
|
||||||
@ -1304,7 +1300,7 @@ Qed.
|
|||||||
Lemma proof_of_mpn_add_n_which_implies_wit_2 : mpn_add_n_which_implies_wit_2.
|
Lemma proof_of_mpn_add_n_which_implies_wit_2 : mpn_add_n_which_implies_wit_2.
|
||||||
Proof.
|
Proof.
|
||||||
pre_process.
|
pre_process.
|
||||||
unfold mpd_store_Z_compact.
|
unfold mpd_store_Z.
|
||||||
Intros l.
|
Intros l.
|
||||||
Exists l.
|
Exists l.
|
||||||
unfold mpd_store_list.
|
unfold mpd_store_list.
|
||||||
|
@ -228,7 +228,7 @@ mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
|||||||
/*@
|
/*@
|
||||||
With val l2 cap1 cap2
|
With val l2 cap1 cap2
|
||||||
Require
|
Require
|
||||||
mpd_store_Z_compact(ap, val, n, cap1) *
|
mpd_store_Z(ap, val, n, cap1) *
|
||||||
store_uint_array(rp, cap2, l2) &&
|
store_uint_array(rp, cap2, l2) &&
|
||||||
Zlength(l2) == cap2 &&
|
Zlength(l2) == cap2 &&
|
||||||
cap2 >= n &&
|
cap2 >= n &&
|
||||||
@ -237,13 +237,13 @@ mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
|||||||
n > 0 && n <= cap1
|
n > 0 && n <= cap1
|
||||||
Ensure
|
Ensure
|
||||||
exists val',
|
exists val',
|
||||||
mpd_store_Z_compact(ap@pre, val, n@pre, cap1) *
|
mpd_store_Z(ap@pre, val, n@pre, cap1) *
|
||||||
mpd_store_Z(rp@pre, val', n@pre, cap2) &&
|
mpd_store_Z(rp@pre, val', n@pre, cap2) &&
|
||||||
(val' + __return * Z::pow(UINT_MOD, n@pre) == val + b@pre)
|
(val' + __return * Z::pow(UINT_MOD, n@pre) == val + b@pre)
|
||||||
*/
|
*/
|
||||||
{
|
{
|
||||||
/*@
|
/*@
|
||||||
mpd_store_Z_compact(ap@pre, val, n@pre, cap1)
|
mpd_store_Z(ap@pre, val, n@pre, cap1)
|
||||||
which implies
|
which implies
|
||||||
exists l,
|
exists l,
|
||||||
n@pre <= cap1 &&
|
n@pre <= cap1 &&
|
||||||
@ -251,7 +251,7 @@ mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
|||||||
cap1 <= 100000000 &&
|
cap1 <= 100000000 &&
|
||||||
store_uint_array(ap@pre, n@pre, l) *
|
store_uint_array(ap@pre, n@pre, l) *
|
||||||
store_undef_uint_array_rec(ap@pre, n@pre, cap1) &&
|
store_undef_uint_array_rec(ap@pre, n@pre, cap1) &&
|
||||||
list_store_Z_compact(l, val)
|
list_store_Z(l, val)
|
||||||
*/
|
*/
|
||||||
int i;
|
int i;
|
||||||
//assert (n > 0);
|
//assert (n > 0);
|
||||||
@ -278,7 +278,7 @@ mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
|||||||
/*@Inv
|
/*@Inv
|
||||||
exists l l' l'' val1 val2,
|
exists l l' l'' val1 val2,
|
||||||
0 <= i && i <= n@pre &&
|
0 <= i && i <= n@pre &&
|
||||||
list_store_Z_compact(l, val) && n@pre <= cap1 &&
|
list_store_Z(l, val) && n@pre <= cap1 &&
|
||||||
store_uint_array(ap@pre, n@pre, l) *
|
store_uint_array(ap@pre, n@pre, l) *
|
||||||
store_undef_uint_array_rec(ap@pre, n@pre, cap1) &&
|
store_undef_uint_array_rec(ap@pre, n@pre, cap1) &&
|
||||||
list_store_Z(sublist(0, i, l), val1) &&
|
list_store_Z(sublist(0, i, l), val1) &&
|
||||||
@ -318,8 +318,8 @@ mpn_add_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n)
|
|||||||
/*@
|
/*@
|
||||||
With cap_a cap_b cap_r val_a val_b l_r
|
With cap_a cap_b cap_r val_a val_b l_r
|
||||||
Require
|
Require
|
||||||
mpd_store_Z_compact(ap, val_a, n, cap_a) *
|
mpd_store_Z(ap, val_a, n, cap_a) *
|
||||||
mpd_store_Z_compact(bp, val_b, n, cap_b) *
|
mpd_store_Z(bp, val_b, n, cap_b) *
|
||||||
store_uint_array(rp, cap_r, l_r) &&
|
store_uint_array(rp, cap_r, l_r) &&
|
||||||
Zlength(l_r) == cap_r &&
|
Zlength(l_r) == cap_r &&
|
||||||
cap_a <= 100000000 &&
|
cap_a <= 100000000 &&
|
||||||
@ -328,14 +328,14 @@ mpn_add_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n)
|
|||||||
n > 0 && n <= cap_a && n <= cap_b && n <= cap_r
|
n > 0 && n <= cap_a && n <= cap_b && n <= cap_r
|
||||||
Ensure
|
Ensure
|
||||||
exists val_r_out,
|
exists val_r_out,
|
||||||
mpd_store_Z_compact(ap@pre, val_a, n@pre, cap_a) *
|
mpd_store_Z(ap@pre, val_a, n@pre, cap_a) *
|
||||||
mpd_store_Z_compact(bp@pre, val_b, n@pre, cap_b) *
|
mpd_store_Z(bp@pre, val_b, n@pre, cap_b) *
|
||||||
mpd_store_Z(rp@pre, val_r_out, n@pre, cap_r) &&
|
mpd_store_Z(rp@pre, val_r_out, n@pre, cap_r) &&
|
||||||
(val_r_out + __return * Z::pow(UINT_MOD, n@pre) == val_a + val_b)
|
(val_r_out + __return * Z::pow(UINT_MOD, n@pre) == val_a + val_b)
|
||||||
*/
|
*/
|
||||||
{
|
{
|
||||||
/*@
|
/*@
|
||||||
mpd_store_Z_compact(ap@pre, val_a, n@pre, cap_a)
|
mpd_store_Z(ap@pre, val_a, n@pre, cap_a)
|
||||||
which implies
|
which implies
|
||||||
exists l_a,
|
exists l_a,
|
||||||
n@pre <= cap_a &&
|
n@pre <= cap_a &&
|
||||||
@ -343,10 +343,10 @@ mpn_add_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n)
|
|||||||
cap_a <= 100000000 &&
|
cap_a <= 100000000 &&
|
||||||
store_uint_array(ap@pre, n@pre, l_a) *
|
store_uint_array(ap@pre, n@pre, l_a) *
|
||||||
store_undef_uint_array_rec(ap@pre, n@pre, cap_a) &&
|
store_undef_uint_array_rec(ap@pre, n@pre, cap_a) &&
|
||||||
list_store_Z_compact(l_a, val_a)
|
list_store_Z(l_a, val_a)
|
||||||
*/
|
*/
|
||||||
/*@
|
/*@
|
||||||
mpd_store_Z_compact(bp@pre, val_b, n@pre, cap_b)
|
mpd_store_Z(bp@pre, val_b, n@pre, cap_b)
|
||||||
which implies
|
which implies
|
||||||
exists l_b,
|
exists l_b,
|
||||||
n@pre <= cap_b &&
|
n@pre <= cap_b &&
|
||||||
@ -354,7 +354,7 @@ mpn_add_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n)
|
|||||||
cap_b <= 100000000 &&
|
cap_b <= 100000000 &&
|
||||||
store_uint_array(bp@pre, n@pre, l_b) *
|
store_uint_array(bp@pre, n@pre, l_b) *
|
||||||
store_undef_uint_array_rec(bp@pre, n@pre, cap_b) &&
|
store_undef_uint_array_rec(bp@pre, n@pre, cap_b) &&
|
||||||
list_store_Z_compact(l_b, val_b)
|
list_store_Z(l_b, val_b)
|
||||||
*/
|
*/
|
||||||
int i;
|
int i;
|
||||||
unsigned int cy;
|
unsigned int cy;
|
||||||
@ -370,8 +370,8 @@ mpn_add_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n)
|
|||||||
/*@Inv
|
/*@Inv
|
||||||
exists l_a l_b l_r_prefix l_r_suffix val_a_prefix val_b_prefix val_r_prefix,
|
exists l_a l_b l_r_prefix l_r_suffix val_a_prefix val_b_prefix val_r_prefix,
|
||||||
0 <= i && i <= n@pre && n@pre <= cap_a && n@pre <= cap_b && n@pre <= cap_r &&
|
0 <= i && i <= n@pre && n@pre <= cap_a && n@pre <= cap_b && n@pre <= cap_r &&
|
||||||
list_store_Z_compact(l_a, val_a) &&
|
list_store_Z(l_a, val_a) &&
|
||||||
list_store_Z_compact(l_b, val_b) &&
|
list_store_Z(l_b, val_b) &&
|
||||||
list_store_Z(sublist(0, i, l_a), val_a_prefix) &&
|
list_store_Z(sublist(0, i, l_a), val_a_prefix) &&
|
||||||
list_store_Z(sublist(0, i, l_b), val_b_prefix) &&
|
list_store_Z(sublist(0, i, l_b), val_b_prefix) &&
|
||||||
list_store_Z(l_r_prefix, val_r_prefix) &&
|
list_store_Z(l_r_prefix, val_r_prefix) &&
|
||||||
|
Reference in New Issue
Block a user