Merge remote-tracking branch 'origin/main' into mpz_sgn
This commit is contained in:
@ -33,6 +33,34 @@ Lemma Z_mod_add_uncarry: forall (a b m: Z),
|
||||
a + b = (a + b) mod m.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma Z_mod_3add_carry10: forall (a b c m: Z),
|
||||
m > 0 -> 0 <= a < m -> 0 <= b < m -> 0 <= c < m ->
|
||||
(a + c) mod m < c ->
|
||||
((a + c) mod m + b) mod m >= b ->
|
||||
a + b + c = ((a + c) mod m + b) mod m + m.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma Z_mod_3add_carry01: forall (a b c m: Z),
|
||||
m > 0 -> 0 <= a < m -> 0 <= b < m -> 0 <= c < m ->
|
||||
(a + c) mod m >= c ->
|
||||
((a + c) mod m + b) mod m < b ->
|
||||
a + b + c = ((a + c) mod m + b) mod m + m.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma Z_mod_3add_carry11: forall (a b c m: Z),
|
||||
m > 0 -> 0 <= a < m -> 0 <= b < m -> 0 <= c < m ->
|
||||
(a + c) mod m < c ->
|
||||
((a + c) mod m + b) mod m < b ->
|
||||
a + b + c = ((a + c) mod m + b) mod m + m * 2.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma Z_mod_3add_carry00: forall (a b c m: Z),
|
||||
m > 0 -> 0 <= a < m -> 0 <= b < m -> 0 <= c < m ->
|
||||
(a + c) mod m >= c ->
|
||||
((a + c) mod m + b) mod m >= b ->
|
||||
a + b + c = ((a + c) mod m + b) mod m.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma Z_of_nat_succ: forall (n: nat),
|
||||
Z.of_nat (S n) = Z.of_nat n + 1.
|
||||
Proof. lia. Qed.
|
||||
|
@ -89,9 +89,9 @@ Proof.
|
||||
reflexivity.
|
||||
Qed.
|
||||
|
||||
Lemma list_store_Z_compact_reverse_injection: forall l1 l2 n1 n2,
|
||||
list_store_Z_compact l1 n1 ->
|
||||
list_store_Z_compact l2 n2 ->
|
||||
Lemma list_store_Z_reverse_injection: forall l1 l2 n1 n2,
|
||||
list_store_Z l1 n1 ->
|
||||
list_store_Z l2 n2 ->
|
||||
n1 = n2 -> l1 = l2.
|
||||
Proof. Admitted.
|
||||
|
||||
@ -282,6 +282,41 @@ Proof.
|
||||
pose proof (Zlength_nonneg l1); lia.
|
||||
Qed.
|
||||
|
||||
Lemma list_store_Z_list_append: forall (l: list Z) (i: Z) (val_prefix: Z) (val_full: Z),
|
||||
0 <= i < Zlength l ->
|
||||
list_store_Z l val_full ->
|
||||
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).
|
||||
Proof.
|
||||
intros.
|
||||
assert ((sublist 0 (i + 1) l) = ((sublist 0 i l) ++ ((Znth i l 0) :: nil)))%list. {
|
||||
pose proof (sublist_split 0 (i+1) i l).
|
||||
pose proof (sublist_single i l 0).
|
||||
rewrite <-H3; try rewrite <- Zlength_correct.
|
||||
apply H2; try rewrite <- Zlength_correct.
|
||||
lia. lia. lia.
|
||||
}
|
||||
rewrite H2.
|
||||
pose proof (list_store_Z_concat (sublist 0 i l) (Znth i l 0 :: nil) (val_prefix) (Znth i l 0)).
|
||||
assert (Zlength (sublist 0 i l) = i). {
|
||||
rewrite Zlength_sublist0.
|
||||
lia.
|
||||
lia.
|
||||
}
|
||||
rewrite H4 in H3.
|
||||
apply H3.
|
||||
tauto.
|
||||
unfold list_store_Z.
|
||||
simpl.
|
||||
split.
|
||||
reflexivity.
|
||||
split; try tauto.
|
||||
apply list_within_bound_Znth.
|
||||
lia.
|
||||
unfold list_store_Z in H0.
|
||||
tauto.
|
||||
Qed.
|
||||
|
||||
Lemma list_store_Z_split: forall (l1 l2: list Z) (n: Z),
|
||||
list_store_Z (l1 ++ l2) n ->
|
||||
list_store_Z l1 (n mod UINT_MOD ^ (Zlength l1)) /\
|
||||
|
File diff suppressed because it is too large
Load Diff
@ -180,6 +180,78 @@ Proof. Admitted.
|
||||
Lemma proof_of_mpn_add_1_partial_solve_wit_7 : mpn_add_1_partial_solve_wit_7.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_safety_wit_1 : mpn_add_n_safety_wit_1.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_safety_wit_2 : mpn_add_n_safety_wit_2.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_safety_wit_3 : mpn_add_n_safety_wit_3.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_safety_wit_4 : mpn_add_n_safety_wit_4.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_safety_wit_5 : mpn_add_n_safety_wit_5.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_safety_wit_6 : mpn_add_n_safety_wit_6.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_1 : mpn_add_n_partial_solve_wit_1.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_2 : mpn_add_n_partial_solve_wit_2.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_3_pure : mpn_add_n_partial_solve_wit_3_pure.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_3 : mpn_add_n_partial_solve_wit_3.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_4 : mpn_add_n_partial_solve_wit_4.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_5 : mpn_add_n_partial_solve_wit_5.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_6_pure : mpn_add_n_partial_solve_wit_6_pure.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_6 : mpn_add_n_partial_solve_wit_6.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_7_pure : mpn_add_n_partial_solve_wit_7_pure.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_7 : mpn_add_n_partial_solve_wit_7.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_8_pure : mpn_add_n_partial_solve_wit_8_pure.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_8 : mpn_add_n_partial_solve_wit_8.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_9_pure : mpn_add_n_partial_solve_wit_9_pure.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_9 : mpn_add_n_partial_solve_wit_9.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_10 : mpn_add_n_partial_solve_wit_10.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_11 : mpn_add_n_partial_solve_wit_11.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_12 : mpn_add_n_partial_solve_wit_12.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpn_add_n_partial_solve_wit_13 : mpn_add_n_partial_solve_wit_13.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_clear_return_wit_1_3 : mpz_clear_return_wit_1_3.
|
||||
Proof. Admitted.
|
||||
|
||||
|
@ -469,7 +469,7 @@ Proof.
|
||||
assert (0 <= Znth i l_3 0 < 4294967296). {
|
||||
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.
|
||||
}
|
||||
assert (i < Zlength l_3). {
|
||||
@ -477,7 +477,7 @@ Proof.
|
||||
rewrite H17.
|
||||
tauto.
|
||||
}
|
||||
unfold list_store_Z_compact in H9.
|
||||
unfold list_store_Z in H9.
|
||||
apply list_within_bound_Znth.
|
||||
lia.
|
||||
tauto.
|
||||
@ -505,7 +505,7 @@ Proof.
|
||||
lia.
|
||||
+ 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.
|
||||
}
|
||||
|
||||
@ -539,7 +539,7 @@ Proof.
|
||||
lia.
|
||||
apply list_within_bound_Znth.
|
||||
lia.
|
||||
unfold list_store_Z_compact in H9.
|
||||
unfold list_store_Z in H9.
|
||||
tauto.
|
||||
- pose proof (Zlength_sublist0 i l'_2).
|
||||
lia.
|
||||
@ -585,7 +585,7 @@ Proof.
|
||||
assert (0 <= Znth i l_3 0 < 4294967296). {
|
||||
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.
|
||||
}
|
||||
assert (i < Zlength l_3). {
|
||||
@ -593,7 +593,7 @@ Proof.
|
||||
rewrite H17.
|
||||
tauto.
|
||||
}
|
||||
unfold list_store_Z_compact in H9.
|
||||
unfold list_store_Z in H9.
|
||||
apply list_within_bound_Znth.
|
||||
lia.
|
||||
tauto.
|
||||
@ -621,7 +621,7 @@ Proof.
|
||||
lia.
|
||||
+ 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.
|
||||
}
|
||||
|
||||
@ -655,7 +655,7 @@ Proof.
|
||||
lia.
|
||||
apply list_within_bound_Znth.
|
||||
lia.
|
||||
unfold list_store_Z_compact in H9.
|
||||
unfold list_store_Z in H9.
|
||||
tauto.
|
||||
- pose proof (Zlength_sublist0 i l'_2).
|
||||
lia.
|
||||
@ -664,10 +664,10 @@ Qed.
|
||||
Lemma proof_of_mpn_add_1_return_wit_1 : mpn_add_1_return_wit_1.
|
||||
Proof.
|
||||
pre_process.
|
||||
unfold mpd_store_Z_compact.
|
||||
unfold mpd_store_Z.
|
||||
unfold mpd_store_list.
|
||||
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.
|
||||
rewrite <-H2 in H10.
|
||||
assert (i = n_pre) by lia.
|
||||
@ -675,32 +675,33 @@ Proof.
|
||||
rewrite <- H10 in H4.
|
||||
rewrite (sublist_self l (Zlength l)) in H4; try tauto.
|
||||
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).
|
||||
apply H22 in H4; try tauto.
|
||||
apply H21 in H4; try tauto.
|
||||
rewrite H4 in H6.
|
||||
entailer!.
|
||||
Exists l.
|
||||
entailer!.
|
||||
entailer!; try rewrite H20; try tauto.
|
||||
- rewrite H10.
|
||||
entailer!.
|
||||
unfold mpd_store_Z.
|
||||
unfold mpd_store_list.
|
||||
Exists l'.
|
||||
rewrite H7.
|
||||
subst i.
|
||||
entailer!.
|
||||
rewrite H20.
|
||||
entailer!.
|
||||
apply store_uint_array_rec_def2undef.
|
||||
- rewrite <- H20. tauto.
|
||||
rewrite H10.
|
||||
entailer!.
|
||||
unfold mpd_store_Z.
|
||||
unfold mpd_store_list.
|
||||
Exists l'.
|
||||
rewrite H7.
|
||||
subst i.
|
||||
entailer!.
|
||||
rewrite H20.
|
||||
entailer!.
|
||||
apply store_uint_array_rec_def2undef.
|
||||
assert (Zlength l' = n_pre) by lia.
|
||||
rewrite <- H7.
|
||||
tauto.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpn_add_1_which_implies_wit_1 : mpn_add_1_which_implies_wit_1.
|
||||
Proof.
|
||||
pre_process.
|
||||
unfold mpd_store_Z_compact.
|
||||
unfold mpd_store_Z.
|
||||
Intros l.
|
||||
Exists l.
|
||||
unfold mpd_store_list.
|
||||
@ -807,6 +808,605 @@ Proof.
|
||||
lia.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpn_add_n_entail_wit_1 : mpn_add_n_entail_wit_1.
|
||||
Proof.
|
||||
pre_process.
|
||||
Exists l_r nil 0 0 0.
|
||||
Exists l_b_2 l_a_2.
|
||||
entailer!.
|
||||
- unfold list_store_Z.
|
||||
simpl.
|
||||
tauto.
|
||||
- rewrite sublist_nil; try lia; try tauto.
|
||||
unfold list_store_Z.
|
||||
simpl.
|
||||
tauto.
|
||||
- rewrite sublist_nil; try lia; try tauto.
|
||||
unfold list_store_Z.
|
||||
simpl.
|
||||
tauto.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpn_add_n_entail_wit_2 : mpn_add_n_entail_wit_2.
|
||||
Proof.
|
||||
pre_process.
|
||||
prop_apply (store_uint_range &("cy") cy).
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpn_add_n_entail_wit_3_1 : mpn_add_n_entail_wit_3_1.
|
||||
Proof.
|
||||
pre_process.
|
||||
rewrite replace_Znth_app_r.
|
||||
assert (l_a_3 = l_a_2). {
|
||||
pose proof (list_store_Z_reverse_injection l_a_3 l_a_2 val_a val_a).
|
||||
specialize (H37 H13 H28).
|
||||
apply H37.
|
||||
reflexivity.
|
||||
}
|
||||
subst l_a_3.
|
||||
assert (l_b_3 = l_b_2). {
|
||||
pose proof (list_store_Z_reverse_injection l_b_3 l_b_2 val_b val_b).
|
||||
specialize (H37 H14 H24).
|
||||
apply H37.
|
||||
reflexivity.
|
||||
}
|
||||
subst l_b_3.
|
||||
- Exists l_r_suffix'.
|
||||
rewrite H29.
|
||||
rewrite H18.
|
||||
assert (i - i = 0) by lia.
|
||||
rewrite H37; clear H37.
|
||||
set (partial_result_1 := (unsigned_last_nbits (Znth i l_a_2 0 + cy) 32)).
|
||||
set (partial_result_2 := (unsigned_last_nbits (partial_result_1 + Znth i l_b_2 0) 32)).
|
||||
rewrite replace_Znth_nothing; try lia.
|
||||
assert ((replace_Znth 0 partial_result_2 (a :: nil)) = partial_result_2 :: nil). {
|
||||
unfold replace_Znth.
|
||||
simpl.
|
||||
reflexivity.
|
||||
}
|
||||
rewrite H37; clear H37.
|
||||
Exists (l_r_prefix_2 ++ partial_result_2 :: nil).
|
||||
Exists (val_r_prefix_2 + partial_result_2 * (UINT_MOD ^ i)).
|
||||
Exists (val_b_prefix_2 + (Znth i l_b_2 0) * (UINT_MOD ^ i)).
|
||||
Exists (val_a_prefix_2 + (Znth i l_a_2 0) * (UINT_MOD ^ i)).
|
||||
Exists l_b_2 l_a_2.
|
||||
entailer!.
|
||||
+ assert ( (val_a_prefix_2 + Znth i l_a_2 0 * 4294967296 ^ i +(val_b_prefix_2 + Znth i l_b_2 0 * 4294967296 ^ i)) = (val_a_prefix_2 + val_b_prefix_2) + Znth i l_a_2 0 * 4294967296 ^ i + Znth i l_b_2 0 * 4294967296 ^ i).
|
||||
{
|
||||
lia.
|
||||
}
|
||||
rewrite H37; clear H37.
|
||||
rewrite <- H19.
|
||||
assert ( (Znth i l_a_2 0) + (Znth i l_b_2 0) + cy = partial_result_2 + UINT_MOD). {
|
||||
unfold unsigned_last_nbits in H4, H3.
|
||||
assert (2 ^ 32 = 4294967296). { nia. }
|
||||
rewrite H37 in H4, H3; clear H37.
|
||||
apply Z_mod_3add_carry10; try lia; try tauto;
|
||||
try unfold list_store_Z in H13, H14;
|
||||
try apply list_within_bound_Znth;
|
||||
try lia;
|
||||
try tauto.
|
||||
}
|
||||
assert ( partial_result_2 * 4294967296 ^ i + (1 + 0) * 4294967296 ^ (i + 1) = cy * 4294967296 ^ i + Znth i l_a_2 0 * 4294967296 ^ i + Znth i l_b_2 0 * 4294967296 ^ i). {
|
||||
rewrite <- Z.mul_add_distr_r.
|
||||
rewrite (Zpow_add_1 4294967296 i); try lia.
|
||||
}
|
||||
lia.
|
||||
+ pose proof (Zlength_app l_r_prefix_2 (partial_result_2 :: nil)).
|
||||
assert (Zlength (partial_result_2 :: nil) = 1). {
|
||||
unfold Zlength.
|
||||
simpl.
|
||||
reflexivity.
|
||||
}
|
||||
rewrite H38 in H37; clear H38.
|
||||
rewrite H18 in H37.
|
||||
apply H37.
|
||||
+ pose proof (list_store_Z_concat l_r_prefix_2 (partial_result_2 :: nil) val_r_prefix_2 partial_result_2).
|
||||
rewrite H18 in H37.
|
||||
apply H37.
|
||||
tauto.
|
||||
unfold list_store_Z.
|
||||
simpl.
|
||||
split.
|
||||
reflexivity.
|
||||
split.
|
||||
unfold partial_result_2.
|
||||
unfold unsigned_last_nbits.
|
||||
assert (2 ^ 32 = 4294967296). { nia. }
|
||||
rewrite H38; clear H38.
|
||||
apply Z.mod_pos_bound.
|
||||
lia.
|
||||
tauto.
|
||||
+ pose proof (list_store_Z_list_append l_b_2 i val_b_prefix_2 val_b).
|
||||
apply H37.
|
||||
lia.
|
||||
tauto.
|
||||
tauto.
|
||||
+ pose proof (list_store_Z_list_append l_a_2 i val_a_prefix_2 val_a).
|
||||
apply H37.
|
||||
lia.
|
||||
tauto.
|
||||
tauto.
|
||||
- pose proof (Zlength_sublist0 i l_r_prefix_2).
|
||||
lia.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpn_add_n_entail_wit_3_2 : mpn_add_n_entail_wit_3_2.
|
||||
Proof.
|
||||
pre_process.
|
||||
rewrite replace_Znth_app_r.
|
||||
assert (l_a_3 = l_a_2). {
|
||||
pose proof (list_store_Z_reverse_injection l_a_3 l_a_2 val_a val_a).
|
||||
specialize (H37 H13 H28).
|
||||
apply H37.
|
||||
reflexivity.
|
||||
}
|
||||
subst l_a_3.
|
||||
assert (l_b_3 = l_b_2). {
|
||||
pose proof (list_store_Z_reverse_injection l_b_3 l_b_2 val_b val_b).
|
||||
specialize (H37 H14 H24).
|
||||
apply H37.
|
||||
reflexivity.
|
||||
}
|
||||
subst l_b_3.
|
||||
- Exists l_r_suffix'.
|
||||
rewrite H29.
|
||||
rewrite H18.
|
||||
assert (i - i = 0) by lia.
|
||||
rewrite H37; clear H37.
|
||||
set (partial_result_1 := (unsigned_last_nbits (Znth i l_a_2 0 + cy) 32)).
|
||||
set (partial_result_2 := (unsigned_last_nbits (partial_result_1 + Znth i l_b_2 0) 32)).
|
||||
rewrite replace_Znth_nothing; try lia.
|
||||
assert ((replace_Znth 0 partial_result_2 (a :: nil)) = partial_result_2 :: nil). {
|
||||
unfold replace_Znth.
|
||||
simpl.
|
||||
reflexivity.
|
||||
}
|
||||
rewrite H37; clear H37.
|
||||
Exists (l_r_prefix_2 ++ partial_result_2 :: nil).
|
||||
Exists (val_r_prefix_2 + partial_result_2 * (UINT_MOD ^ i)).
|
||||
Exists (val_b_prefix_2 + (Znth i l_b_2 0) * (UINT_MOD ^ i)).
|
||||
Exists (val_a_prefix_2 + (Znth i l_a_2 0) * (UINT_MOD ^ i)).
|
||||
Exists l_b_2 l_a_2.
|
||||
entailer!.
|
||||
+ assert ( (val_a_prefix_2 + Znth i l_a_2 0 * 4294967296 ^ i +(val_b_prefix_2 + Znth i l_b_2 0 * 4294967296 ^ i)) = (val_a_prefix_2 + val_b_prefix_2) + Znth i l_a_2 0 * 4294967296 ^ i + Znth i l_b_2 0 * 4294967296 ^ i).
|
||||
{
|
||||
lia.
|
||||
}
|
||||
rewrite H37; clear H37.
|
||||
rewrite <- H19.
|
||||
assert ( (Znth i l_a_2 0) + (Znth i l_b_2 0) + cy = partial_result_2 + UINT_MOD * 2). {
|
||||
unfold unsigned_last_nbits in H4, H3.
|
||||
assert (2 ^ 32 = 4294967296). { nia. }
|
||||
rewrite H37 in H4, H3; clear H37.
|
||||
apply Z_mod_3add_carry11; try lia; try tauto;
|
||||
try unfold list_store_Z in H13, H14;
|
||||
try apply list_within_bound_Znth;
|
||||
try lia;
|
||||
try tauto.
|
||||
}
|
||||
assert ( partial_result_2 * 4294967296 ^ i + (1 + 1) * 4294967296 ^ (i + 1) = cy * 4294967296 ^ i + Znth i l_a_2 0 * 4294967296 ^ i + Znth i l_b_2 0 * 4294967296 ^ i). {
|
||||
rewrite <- Z.mul_add_distr_r.
|
||||
rewrite (Zpow_add_1 4294967296 i); try lia.
|
||||
}
|
||||
lia.
|
||||
+ pose proof (Zlength_app l_r_prefix_2 (partial_result_2 :: nil)).
|
||||
assert (Zlength (partial_result_2 :: nil) = 1). {
|
||||
unfold Zlength.
|
||||
simpl.
|
||||
reflexivity.
|
||||
}
|
||||
rewrite H38 in H37; clear H38.
|
||||
rewrite H18 in H37.
|
||||
apply H37.
|
||||
+ pose proof (list_store_Z_concat l_r_prefix_2 (partial_result_2 :: nil) val_r_prefix_2 partial_result_2).
|
||||
rewrite H18 in H37.
|
||||
apply H37.
|
||||
tauto.
|
||||
unfold list_store_Z.
|
||||
simpl.
|
||||
split.
|
||||
reflexivity.
|
||||
split.
|
||||
unfold partial_result_2.
|
||||
unfold unsigned_last_nbits.
|
||||
assert (2 ^ 32 = 4294967296). { nia. }
|
||||
rewrite H38; clear H38.
|
||||
apply Z.mod_pos_bound.
|
||||
lia.
|
||||
tauto.
|
||||
+ pose proof (list_store_Z_list_append l_b_2 i val_b_prefix_2 val_b).
|
||||
apply H37.
|
||||
lia.
|
||||
tauto.
|
||||
tauto.
|
||||
+ pose proof (list_store_Z_list_append l_a_2 i val_a_prefix_2 val_a).
|
||||
apply H37.
|
||||
lia.
|
||||
tauto.
|
||||
tauto.
|
||||
- pose proof (Zlength_sublist0 i l_r_prefix_2).
|
||||
lia.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpn_add_n_entail_wit_3_3 : mpn_add_n_entail_wit_3_3.
|
||||
Proof.
|
||||
pre_process.
|
||||
rewrite replace_Znth_app_r.
|
||||
assert (l_a_3 = l_a_2). {
|
||||
pose proof (list_store_Z_reverse_injection l_a_3 l_a_2 val_a val_a).
|
||||
specialize (H37 H13 H28).
|
||||
apply H37.
|
||||
reflexivity.
|
||||
}
|
||||
subst l_a_3.
|
||||
assert (l_b_3 = l_b_2). {
|
||||
pose proof (list_store_Z_reverse_injection l_b_3 l_b_2 val_b val_b).
|
||||
specialize (H37 H14 H24).
|
||||
apply H37.
|
||||
reflexivity.
|
||||
}
|
||||
subst l_b_3.
|
||||
- Exists l_r_suffix'.
|
||||
rewrite H29.
|
||||
rewrite H18.
|
||||
assert (i - i = 0) by lia.
|
||||
rewrite H37; clear H37.
|
||||
set (partial_result_1 := (unsigned_last_nbits (Znth i l_a_2 0 + cy) 32)).
|
||||
set (partial_result_2 := (unsigned_last_nbits (partial_result_1 + Znth i l_b_2 0) 32)).
|
||||
rewrite replace_Znth_nothing; try lia.
|
||||
assert ((replace_Znth 0 partial_result_2 (a :: nil)) = partial_result_2 :: nil). {
|
||||
unfold replace_Znth.
|
||||
simpl.
|
||||
reflexivity.
|
||||
}
|
||||
rewrite H37; clear H37.
|
||||
Exists (l_r_prefix_2 ++ partial_result_2 :: nil).
|
||||
Exists (val_r_prefix_2 + partial_result_2 * (UINT_MOD ^ i)).
|
||||
Exists (val_b_prefix_2 + (Znth i l_b_2 0) * (UINT_MOD ^ i)).
|
||||
Exists (val_a_prefix_2 + (Znth i l_a_2 0) * (UINT_MOD ^ i)).
|
||||
Exists l_b_2 l_a_2.
|
||||
entailer!.
|
||||
+ assert ( (val_a_prefix_2 + Znth i l_a_2 0 * 4294967296 ^ i +(val_b_prefix_2 + Znth i l_b_2 0 * 4294967296 ^ i)) = (val_a_prefix_2 + val_b_prefix_2) + Znth i l_a_2 0 * 4294967296 ^ i + Znth i l_b_2 0 * 4294967296 ^ i).
|
||||
{
|
||||
lia.
|
||||
}
|
||||
rewrite H37; clear H37.
|
||||
rewrite <- H19.
|
||||
assert ( (Znth i l_a_2 0) + (Znth i l_b_2 0) + cy = partial_result_2). {
|
||||
unfold unsigned_last_nbits in H4, H3.
|
||||
assert (2 ^ 32 = 4294967296). { nia. }
|
||||
rewrite H37 in H4, H3; clear H37.
|
||||
apply Z_mod_3add_carry00; try lia; try tauto;
|
||||
try unfold list_store_Z in H13, H14;
|
||||
try apply list_within_bound_Znth;
|
||||
try lia;
|
||||
try tauto.
|
||||
}
|
||||
assert ( partial_result_2 * 4294967296 ^ i + (0 + 0) * 4294967296 ^ (i + 1) = cy * 4294967296 ^ i + Znth i l_a_2 0 * 4294967296 ^ i + Znth i l_b_2 0 * 4294967296 ^ i). {
|
||||
rewrite <- Z.mul_add_distr_r.
|
||||
rewrite (Zpow_add_1 4294967296 i); try lia.
|
||||
}
|
||||
lia.
|
||||
+ pose proof (Zlength_app l_r_prefix_2 (partial_result_2 :: nil)).
|
||||
assert (Zlength (partial_result_2 :: nil) = 1). {
|
||||
unfold Zlength.
|
||||
simpl.
|
||||
reflexivity.
|
||||
}
|
||||
rewrite H38 in H37; clear H38.
|
||||
rewrite H18 in H37.
|
||||
apply H37.
|
||||
+ pose proof (list_store_Z_concat l_r_prefix_2 (partial_result_2 :: nil) val_r_prefix_2 partial_result_2).
|
||||
rewrite H18 in H37.
|
||||
apply H37.
|
||||
tauto.
|
||||
unfold list_store_Z.
|
||||
simpl.
|
||||
split.
|
||||
reflexivity.
|
||||
split.
|
||||
unfold partial_result_2.
|
||||
unfold unsigned_last_nbits.
|
||||
assert (2 ^ 32 = 4294967296). { nia. }
|
||||
rewrite H38; clear H38.
|
||||
apply Z.mod_pos_bound.
|
||||
lia.
|
||||
tauto.
|
||||
+ pose proof (list_store_Z_list_append l_b_2 i val_b_prefix_2 val_b).
|
||||
apply H37.
|
||||
lia.
|
||||
tauto.
|
||||
tauto.
|
||||
+ pose proof (list_store_Z_list_append l_a_2 i val_a_prefix_2 val_a).
|
||||
apply H37.
|
||||
lia.
|
||||
tauto.
|
||||
tauto.
|
||||
- pose proof (Zlength_sublist0 i l_r_prefix_2).
|
||||
lia.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpn_add_n_entail_wit_3_4 : mpn_add_n_entail_wit_3_4.
|
||||
Proof.
|
||||
pre_process.
|
||||
rewrite replace_Znth_app_r.
|
||||
assert (l_a_3 = l_a_2). {
|
||||
pose proof (list_store_Z_reverse_injection l_a_3 l_a_2 val_a val_a).
|
||||
specialize (H37 H13 H28).
|
||||
apply H37.
|
||||
reflexivity.
|
||||
}
|
||||
subst l_a_3.
|
||||
assert (l_b_3 = l_b_2). {
|
||||
pose proof (list_store_Z_reverse_injection l_b_3 l_b_2 val_b val_b).
|
||||
specialize (H37 H14 H24).
|
||||
apply H37.
|
||||
reflexivity.
|
||||
}
|
||||
subst l_b_3.
|
||||
- Exists l_r_suffix'.
|
||||
rewrite H29.
|
||||
rewrite H18.
|
||||
assert (i - i = 0) by lia.
|
||||
rewrite H37; clear H37.
|
||||
set (partial_result_1 := (unsigned_last_nbits (Znth i l_a_2 0 + cy) 32)).
|
||||
set (partial_result_2 := (unsigned_last_nbits (partial_result_1 + Znth i l_b_2 0) 32)).
|
||||
rewrite replace_Znth_nothing; try lia.
|
||||
assert ((replace_Znth 0 partial_result_2 (a :: nil)) = partial_result_2 :: nil). {
|
||||
unfold replace_Znth.
|
||||
simpl.
|
||||
reflexivity.
|
||||
}
|
||||
rewrite H37; clear H37.
|
||||
Exists (l_r_prefix_2 ++ partial_result_2 :: nil).
|
||||
Exists (val_r_prefix_2 + partial_result_2 * (UINT_MOD ^ i)).
|
||||
Exists (val_b_prefix_2 + (Znth i l_b_2 0) * (UINT_MOD ^ i)).
|
||||
Exists (val_a_prefix_2 + (Znth i l_a_2 0) * (UINT_MOD ^ i)).
|
||||
Exists l_b_2 l_a_2.
|
||||
entailer!.
|
||||
+ assert ( (val_a_prefix_2 + Znth i l_a_2 0 * 4294967296 ^ i +(val_b_prefix_2 + Znth i l_b_2 0 * 4294967296 ^ i)) = (val_a_prefix_2 + val_b_prefix_2) + Znth i l_a_2 0 * 4294967296 ^ i + Znth i l_b_2 0 * 4294967296 ^ i).
|
||||
{
|
||||
lia.
|
||||
}
|
||||
rewrite H37; clear H37.
|
||||
rewrite <- H19.
|
||||
assert ( (Znth i l_a_2 0) + (Znth i l_b_2 0) + cy = partial_result_2 + UINT_MOD). {
|
||||
unfold unsigned_last_nbits in H4, H3.
|
||||
assert (2 ^ 32 = 4294967296). { nia. }
|
||||
rewrite H37 in H4, H3; clear H37.
|
||||
apply Z_mod_3add_carry01; try lia; try tauto;
|
||||
try unfold list_store_Z in H13, H14;
|
||||
try apply list_within_bound_Znth;
|
||||
try lia;
|
||||
try tauto.
|
||||
}
|
||||
assert ( partial_result_2 * 4294967296 ^ i + (0 + 1) * 4294967296 ^ (i + 1) = cy * 4294967296 ^ i + Znth i l_a_2 0 * 4294967296 ^ i + Znth i l_b_2 0 * 4294967296 ^ i). {
|
||||
rewrite <- Z.mul_add_distr_r.
|
||||
rewrite (Zpow_add_1 4294967296 i); try lia.
|
||||
}
|
||||
lia.
|
||||
+ pose proof (Zlength_app l_r_prefix_2 (partial_result_2 :: nil)).
|
||||
assert (Zlength (partial_result_2 :: nil) = 1). {
|
||||
unfold Zlength.
|
||||
simpl.
|
||||
reflexivity.
|
||||
}
|
||||
rewrite H38 in H37; clear H38.
|
||||
rewrite H18 in H37.
|
||||
apply H37.
|
||||
+ pose proof (list_store_Z_concat l_r_prefix_2 (partial_result_2 :: nil) val_r_prefix_2 partial_result_2).
|
||||
rewrite H18 in H37.
|
||||
apply H37.
|
||||
tauto.
|
||||
unfold list_store_Z.
|
||||
simpl.
|
||||
split.
|
||||
reflexivity.
|
||||
split.
|
||||
unfold partial_result_2.
|
||||
unfold unsigned_last_nbits.
|
||||
assert (2 ^ 32 = 4294967296). { nia. }
|
||||
rewrite H38; clear H38.
|
||||
apply Z.mod_pos_bound.
|
||||
lia.
|
||||
tauto.
|
||||
+ pose proof (list_store_Z_list_append l_b_2 i val_b_prefix_2 val_b).
|
||||
apply H37.
|
||||
lia.
|
||||
tauto.
|
||||
tauto.
|
||||
+ pose proof (list_store_Z_list_append l_a_2 i val_a_prefix_2 val_a).
|
||||
apply H37.
|
||||
lia.
|
||||
tauto.
|
||||
tauto.
|
||||
- pose proof (Zlength_sublist0 i l_r_prefix_2).
|
||||
lia.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpn_add_n_return_wit_1 : mpn_add_n_return_wit_1.
|
||||
Proof.
|
||||
pre_process.
|
||||
assert (l_a_2 = l_a). {
|
||||
pose proof (list_store_Z_reverse_injection l_a_2 l_a val_a val_a).
|
||||
specialize (H29 H20 H5).
|
||||
apply H29.
|
||||
reflexivity.
|
||||
}
|
||||
subst l_a_2.
|
||||
assert (l_b_2 = l_b). {
|
||||
pose proof (list_store_Z_reverse_injection l_b_2 l_b val_b val_b).
|
||||
specialize (H29 H16 H6).
|
||||
apply H29.
|
||||
reflexivity.
|
||||
}
|
||||
subst l_b_2.
|
||||
assert (i = n_pre) by lia.
|
||||
Exists val_r_prefix.
|
||||
unfold mpd_store_Z.
|
||||
unfold mpd_store_list.
|
||||
Exists l_a.
|
||||
Exists l_b.
|
||||
entailer!.
|
||||
rewrite H14.
|
||||
rewrite H18.
|
||||
entailer!.
|
||||
unfold mpd_store_Z.
|
||||
Exists l_r_prefix.
|
||||
rewrite H29 in *.
|
||||
entailer!.
|
||||
unfold mpd_store_list.
|
||||
entailer!.
|
||||
rewrite H10.
|
||||
entailer!.
|
||||
apply store_uint_array_rec_def2undef.
|
||||
rewrite <- H29.
|
||||
assert (val_a_prefix = val_a). {
|
||||
rewrite <-H18 in H7.
|
||||
rewrite sublist_self in H7.
|
||||
unfold list_store_Z in H5.
|
||||
unfold list_store_Z in H7.
|
||||
lia.
|
||||
reflexivity.
|
||||
}
|
||||
rewrite <- H30; clear H30.
|
||||
assert (val_b_prefix = val_b). {
|
||||
rewrite <-H14 in H8.
|
||||
rewrite sublist_self in H8.
|
||||
unfold list_store_Z in H6.
|
||||
unfold list_store_Z in H8.
|
||||
lia.
|
||||
reflexivity.
|
||||
}
|
||||
rewrite <- H30; clear H30.
|
||||
rewrite H29.
|
||||
tauto.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpn_add_n_which_implies_wit_1 : mpn_add_n_which_implies_wit_1.
|
||||
Proof.
|
||||
pre_process.
|
||||
unfold mpd_store_Z.
|
||||
Intros l.
|
||||
Exists l.
|
||||
unfold mpd_store_list.
|
||||
entailer!.
|
||||
subst n_pre.
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpn_add_n_which_implies_wit_2 : mpn_add_n_which_implies_wit_2.
|
||||
Proof.
|
||||
pre_process.
|
||||
unfold mpd_store_Z.
|
||||
Intros l.
|
||||
Exists l.
|
||||
unfold mpd_store_list.
|
||||
entailer!.
|
||||
subst n_pre.
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpn_add_n_which_implies_wit_3 : mpn_add_n_which_implies_wit_3.
|
||||
Proof.
|
||||
pre_process.
|
||||
pose proof (store_uint_array_divide rp_pre cap_r l_r 0).
|
||||
pose proof (Zlength_nonneg l_r).
|
||||
specialize (H0 ltac:(lia) ltac:(lia)).
|
||||
destruct H0 as [H0 _].
|
||||
simpl in H0.
|
||||
entailer!.
|
||||
rewrite (sublist_nil l_r 0 0) in H0; [ | lia].
|
||||
sep_apply H0.
|
||||
entailer!.
|
||||
unfold store_uint_array, store_uint_array_rec.
|
||||
unfold store_array.
|
||||
rewrite (sublist_self l_r cap_r); [ | lia ].
|
||||
assert (rp_pre + 0 = rp_pre). { lia. }
|
||||
rewrite H2; clear H2.
|
||||
assert (cap_r - 0 = cap_r). { lia. }
|
||||
rewrite H2; clear H2.
|
||||
reflexivity.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpn_add_n_which_implies_wit_4 : mpn_add_n_which_implies_wit_4.
|
||||
Proof.
|
||||
pre_process.
|
||||
destruct l_r_suffix. {
|
||||
unfold store_uint_array_rec.
|
||||
simpl.
|
||||
entailer!.
|
||||
}
|
||||
pose proof (store_uint_array_rec_cons rp_pre i cap_r z l_r_suffix ltac:(lia)).
|
||||
sep_apply H2.
|
||||
Exists z l_r_suffix.
|
||||
entailer!.
|
||||
assert (i = 0 \/ i > 0). { lia. }
|
||||
destruct H3.
|
||||
+ subst.
|
||||
simpl.
|
||||
entailer!.
|
||||
simpl in H2.
|
||||
assert (rp_pre + 0 = rp_pre). { lia. }
|
||||
rewrite H3.
|
||||
rewrite H3 in H2.
|
||||
clear H3.
|
||||
pose proof (store_uint_array_empty rp_pre l_r_prefix).
|
||||
sep_apply H3.
|
||||
rewrite logic_equiv_andp_comm.
|
||||
rewrite logic_equiv_coq_prop_andp_sepcon.
|
||||
Intros.
|
||||
subst l_r_prefix.
|
||||
rewrite app_nil_l.
|
||||
unfold store_uint_array.
|
||||
unfold store_array.
|
||||
unfold store_array_rec.
|
||||
simpl.
|
||||
assert (rp_pre + 0 = rp_pre). { lia. }
|
||||
rewrite H4; clear H4.
|
||||
entailer!.
|
||||
+ pose proof (Aux.uint_array_rec_to_uint_array rp_pre 0 i (sublist 0 i l_r_prefix) ltac:(lia)).
|
||||
destruct H4 as [_ H4].
|
||||
assert (rp_pre + sizeof(UINT) * 0 = rp_pre). { lia. }
|
||||
rewrite H5 in H4; clear H5.
|
||||
assert (i - 0 = i). { lia. }
|
||||
rewrite H5 in H4; clear H5.
|
||||
pose proof (Aux.uint_array_rec_to_uint_array rp_pre 0 (i + 1) (sublist 0 i l_r_prefix ++ z :: nil) ltac:(lia)).
|
||||
destruct H5 as [H5 _].
|
||||
assert (i + 1 - 0 = i + 1). { lia. }
|
||||
rewrite H6 in H5; clear H6.
|
||||
assert (rp_pre + sizeof(UINT) * 0 = rp_pre). { lia. }
|
||||
rewrite H6 in H5; clear H6.
|
||||
pose proof (uint_array_rec_to_uint_array rp_pre 0 i l_r_prefix).
|
||||
specialize (H6 H).
|
||||
assert ((rp_pre + sizeof ( UINT ) * 0) = rp_pre) by lia.
|
||||
rewrite H7 in H6; clear H7.
|
||||
assert ((i-0) = i) by lia.
|
||||
rewrite H7 in H6; clear H7.
|
||||
destruct H6 as [_ H6].
|
||||
sep_apply H6.
|
||||
(* pose proof (uint_array_rec_to_uint_array rp_pre 0 (i+1) (l' ++ z :: nil)).
|
||||
assert (H_i_plus_1 : 0 <= i + 1) by lia.
|
||||
specialize (H7 H_i_plus_1); clear H_i_plus_1.
|
||||
destruct H7 as [H7 _].
|
||||
assert (i + 1 - 0 = i + 1) by lia.
|
||||
rewrite H8 in H7; clear H8.
|
||||
assert ((rp_pre + sizeof ( UINT ) * 0) = rp_pre) by lia.
|
||||
rewrite H8 in H7; clear H8.
|
||||
rewrite <-H7.
|
||||
clear H6.
|
||||
clear H7. *)
|
||||
pose proof (store_uint_array_divide_rec rp_pre (i+1) (l_r_prefix ++ z :: nil) i).
|
||||
assert (H_tmp: 0 <= i <= i+1) by lia.
|
||||
specialize (H7 H_tmp); clear H_tmp.
|
||||
rewrite <- store_uint_array_single.
|
||||
sep_apply store_uint_array_rec_divide_rev.
|
||||
entailer!.
|
||||
lia.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_clear_return_wit_1_1 : mpz_clear_return_wit_1_1.
|
||||
Proof.
|
||||
pre_process.
|
||||
|
@ -228,7 +228,7 @@ mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
||||
/*@
|
||||
With val l2 cap1 cap2
|
||||
Require
|
||||
mpd_store_Z_compact(ap, val, n, cap1) *
|
||||
mpd_store_Z(ap, val, n, cap1) *
|
||||
store_uint_array(rp, cap2, l2) &&
|
||||
Zlength(l2) == cap2 &&
|
||||
cap2 >= n &&
|
||||
@ -237,13 +237,13 @@ mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
||||
n > 0 && n <= cap1
|
||||
Ensure
|
||||
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) &&
|
||||
(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
|
||||
exists l,
|
||||
n@pre <= cap1 &&
|
||||
@ -251,7 +251,7 @@ mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
||||
cap1 <= 100000000 &&
|
||||
store_uint_array(ap@pre, n@pre, l) *
|
||||
store_undef_uint_array_rec(ap@pre, n@pre, cap1) &&
|
||||
list_store_Z_compact(l, val)
|
||||
list_store_Z(l, val)
|
||||
*/
|
||||
int i;
|
||||
//assert (n > 0);
|
||||
@ -278,7 +278,7 @@ mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
||||
/*@Inv
|
||||
exists l l' l'' val1 val2,
|
||||
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_undef_uint_array_rec(ap@pre, n@pre, cap1) &&
|
||||
list_store_Z(sublist(0, i, l), val1) &&
|
||||
@ -313,24 +313,104 @@ mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
|
||||
}
|
||||
|
||||
/* 位数相同的多精度数ap 加上多精度数bp,返回最后产生的进位 */
|
||||
/*unsigned int
|
||||
unsigned int
|
||||
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
|
||||
Require
|
||||
mpd_store_Z(ap, val_a, n, cap_a) *
|
||||
mpd_store_Z(bp, val_b, n, cap_b) *
|
||||
store_uint_array(rp, cap_r, l_r) &&
|
||||
Zlength(l_r) == cap_r &&
|
||||
cap_a <= 100000000 &&
|
||||
cap_b <= 100000000 &&
|
||||
cap_r <= 100000000 &&
|
||||
n > 0 && n <= cap_a && n <= cap_b && n <= cap_r
|
||||
Ensure
|
||||
exists val_r_out,
|
||||
mpd_store_Z(ap@pre, val_a, n@pre, cap_a) *
|
||||
mpd_store_Z(bp@pre, val_b, n@pre, cap_b) *
|
||||
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)
|
||||
*/
|
||||
{
|
||||
/*@
|
||||
mpd_store_Z(ap@pre, val_a, n@pre, cap_a)
|
||||
which implies
|
||||
exists l_a,
|
||||
n@pre <= cap_a &&
|
||||
Zlength(l_a) == n@pre &&
|
||||
cap_a <= 100000000 &&
|
||||
store_uint_array(ap@pre, n@pre, l_a) *
|
||||
store_undef_uint_array_rec(ap@pre, n@pre, cap_a) &&
|
||||
list_store_Z(l_a, val_a)
|
||||
*/
|
||||
/*@
|
||||
mpd_store_Z(bp@pre, val_b, n@pre, cap_b)
|
||||
which implies
|
||||
exists l_b,
|
||||
n@pre <= cap_b &&
|
||||
Zlength(l_b) == n@pre &&
|
||||
cap_b <= 100000000 &&
|
||||
store_uint_array(bp@pre, n@pre, l_b) *
|
||||
store_undef_uint_array_rec(bp@pre, n@pre, cap_b) &&
|
||||
list_store_Z(l_b, val_b)
|
||||
*/
|
||||
int i;
|
||||
unsigned int cy;
|
||||
|
||||
for (i = 0, cy = 0; i < n; i++)
|
||||
/*@
|
||||
store_uint_array(rp@pre, cap_r, l_r) && Zlength(l_r) == cap_r
|
||||
which implies
|
||||
store_uint_array_rec(rp@pre, 0, cap_r, l_r) * store_uint_array(rp@pre, 0, nil) &&
|
||||
Zlength(l_r) == cap_r
|
||||
*/
|
||||
i = 0;
|
||||
cy = 0;
|
||||
/*@Inv
|
||||
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 &&
|
||||
list_store_Z(l_a, val_a) &&
|
||||
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_b), val_b_prefix) &&
|
||||
list_store_Z(l_r_prefix, val_r_prefix) &&
|
||||
Zlength(l_r_prefix) == i &&
|
||||
(val_r_prefix + cy * Z::pow(UINT_MOD, i) == val_a_prefix + val_b_prefix) &&
|
||||
store_uint_array(ap@pre, n@pre, l_a) *
|
||||
store_undef_uint_array_rec(ap@pre, n@pre, cap_a) *
|
||||
store_uint_array(bp@pre, n@pre, l_b) *
|
||||
store_undef_uint_array_rec(bp@pre, n@pre, cap_b) *
|
||||
store_uint_array(rp@pre, i, l_r_prefix) *
|
||||
store_uint_array_rec(rp@pre, i, cap_r, l_r_suffix)
|
||||
*/
|
||||
while (i < n)
|
||||
{
|
||||
/*@
|
||||
Given l_a l_b l_r_prefix l_r_suffix val_a_prefix val_b_prefix val_r_prefix
|
||||
*/
|
||||
/*@ 0 <= cy && cy <= UINT_MAX by local */
|
||||
unsigned int a, b, r;
|
||||
a = ap[i]; b = bp[i];
|
||||
r = a + cy;
|
||||
cy = (r < cy);
|
||||
r += b;
|
||||
cy += (r < b);
|
||||
/*@
|
||||
0 <= i && i < n@pre && n@pre <= cap_r &&
|
||||
store_uint_array(rp@pre, i, l_r_prefix) *
|
||||
store_uint_array_rec(rp@pre, i, cap_r, l_r_suffix)
|
||||
which implies
|
||||
exists a l_r_suffix',
|
||||
l_r_suffix == cons(a, l_r_suffix') && 0 <= i && i < n@pre && n@pre <= cap_r &&
|
||||
store_uint_array_rec(rp@pre, i+1, cap_r, l_r_suffix') *
|
||||
store_uint_array(rp@pre, i+1, app(l_r_prefix, cons(a, nil)))
|
||||
*/
|
||||
rp[i] = r;
|
||||
++i;
|
||||
}
|
||||
return cy;
|
||||
}*/
|
||||
}
|
||||
|
||||
/*不同位数的多精度数相加,返回最后的进位*/
|
||||
/*unsigned int
|
||||
|
@ -84,7 +84,7 @@ void mpn_copyi (unsigned int *d, unsigned int *s, int n);
|
||||
int mpn_cmp (unsigned int *ap, unsigned int *bp, int n);
|
||||
|
||||
unsigned int mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b);
|
||||
unsigned int mpn_add_n (unsigned int *, unsigned int *, unsigned int *, int);
|
||||
unsigned int mpn_add_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n);
|
||||
unsigned int mpn_add (unsigned int *, unsigned int *, int, unsigned int *, int);
|
||||
|
||||
unsigned int mpn_sub_1 (unsigned int *, unsigned int *, int, unsigned int);
|
||||
|
Reference in New Issue
Block a user