feat(mpz_clear & mpz_realloc): Proved correctness of mpz_clear and mpz_realloc
This commit is contained in:
@ -597,20 +597,10 @@ Record bigint_ent: Type := {
|
||||
sign: Prop;
|
||||
}.
|
||||
|
||||
Definition store_bigint_ent (x: addr) (n: bigint_ent): Assertion :=
|
||||
EX size,
|
||||
&(x # "__mpz_struct" ->ₛ "_mp_size") # Int |-> size &&
|
||||
([| size < 0 |] && [| sign n |] && [| size = -Zlength (data n) |] ||
|
||||
[| size >= 0 |] && [| ~(sign n) |] && [| size = Zlength (data n) |]) **
|
||||
&(x # "__mpz_struct" ->ₛ "_mp_alloc") # Int |-> cap n **
|
||||
EX p,
|
||||
&(x # "__mpz_struct" ->ₛ "_mp_d") # Ptr |-> p **
|
||||
Internal.mpd_store_list p (data n) (cap n).
|
||||
|
||||
Definition bigint_ent_store_Z (n: bigint_ent) (x: Z): Assertion :=
|
||||
[| sign n |] && [| Internal.list_store_Z (data n) (-x) |] ||
|
||||
[| ~(sign n) |] && [| Internal.list_store_Z (data n) x |].
|
||||
|
||||
Definition store_Z (x: addr) (n: Z): Assertion :=
|
||||
EX ent,
|
||||
store_bigint_ent x ent && bigint_ent_store_Z ent n.
|
||||
EX (ptr: addr) (cap size: Z),
|
||||
(([| size < 0 |] && [| n < 0 |] && Internal.mpd_store_Z_compact ptr (-n) (-size) cap) ||
|
||||
([| size >= 0 |] && [| n >= 0 |] && Internal.mpd_store_Z_compact ptr n size cap)) **
|
||||
&(x # "__mpz_struct" ->ₛ "_mp_size") # Int |-> size **
|
||||
&(x # "__mpz_struct" ->ₛ "_mp_alloc") # Int |-> cap **
|
||||
&(x # "__mpz_struct" ->ₛ "_mp_d") # Ptr |-> ptr.
|
1068
projects/lib/gmp_goal.v
Normal file → Executable file
1068
projects/lib/gmp_goal.v
Normal file → Executable file
File diff suppressed because it is too large
Load Diff
0
projects/lib/gmp_goal_check.v
Normal file → Executable file
0
projects/lib/gmp_goal_check.v
Normal file → Executable file
63
projects/lib/gmp_proof_auto.v
Normal file → Executable file
63
projects/lib/gmp_proof_auto.v
Normal file → Executable file
@ -141,3 +141,66 @@ Proof. Admitted.
|
||||
Lemma proof_of_mpn_normalized_size_partial_solve_wit_2 : mpn_normalized_size_partial_solve_wit_2.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_clear_return_wit_1_3 : mpz_clear_return_wit_1_3.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_clear_return_wit_1_4 : mpz_clear_return_wit_1_4.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_clear_partial_solve_wit_1 : mpz_clear_partial_solve_wit_1.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_clear_partial_solve_wit_2 : mpz_clear_partial_solve_wit_2.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_clear_partial_solve_wit_3 : mpz_clear_partial_solve_wit_3.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_safety_wit_1 : mpz_realloc_safety_wit_1.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_safety_wit_2 : mpz_realloc_safety_wit_2.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_safety_wit_3 : mpz_realloc_safety_wit_3.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_safety_wit_4 : mpz_realloc_safety_wit_4.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_safety_wit_5 : mpz_realloc_safety_wit_5.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_safety_wit_6 : mpz_realloc_safety_wit_6.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_1 : mpz_realloc_partial_solve_wit_1.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_2 : mpz_realloc_partial_solve_wit_2.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_3 : mpz_realloc_partial_solve_wit_3.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_4 : mpz_realloc_partial_solve_wit_4.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_5 : mpz_realloc_partial_solve_wit_5.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_6 : mpz_realloc_partial_solve_wit_6.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_7 : mpz_realloc_partial_solve_wit_7.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_8 : mpz_realloc_partial_solve_wit_8.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_9 : mpz_realloc_partial_solve_wit_9.
|
||||
Proof. Admitted.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_10 : mpz_realloc_partial_solve_wit_10.
|
||||
Proof. Admitted.
|
||||
|
||||
|
297
projects/lib/gmp_proof_manual.v
Normal file → Executable file
297
projects/lib/gmp_proof_manual.v
Normal file → Executable file
@ -410,4 +410,299 @@ Proof.
|
||||
entailer!.
|
||||
+ rewrite sublist_self; try lia.
|
||||
tauto.
|
||||
Qed.
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma proof_of_mpz_clear_return_wit_1_1 : mpz_clear_return_wit_1_1.
|
||||
Proof.
|
||||
pre_process.
|
||||
Exists ptr_2 cap_2 size_2.
|
||||
entailer!.
|
||||
unfold mpd_store_Z_compact.
|
||||
Intros data.
|
||||
unfold mpd_store_list.
|
||||
subst.
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_clear_return_wit_1_2 : mpz_clear_return_wit_1_2.
|
||||
Proof.
|
||||
pre_process.
|
||||
Exists ptr_2 cap_2 size_2.
|
||||
entailer!.
|
||||
unfold mpd_store_Z_compact.
|
||||
Intros data.
|
||||
unfold mpd_store_list.
|
||||
entailer!.
|
||||
assert (size_2 = 0). {
|
||||
pose proof (Zlength_nonneg data).
|
||||
lia.
|
||||
}
|
||||
rewrite H6 in *.
|
||||
rewrite <-H3 in *.
|
||||
unfold store_uint_array, store_undef_uint_array_rec.
|
||||
unfold store_array.
|
||||
assert (cap_2 - 0 = 0). { lia. }
|
||||
rewrite H7; clear H7.
|
||||
pose proof (Zlength_nil_inv data ltac:(lia)).
|
||||
rewrite H7 in *; clear H7.
|
||||
simpl.
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_clear_return_wit_1_3 : mpz_clear_return_wit_1_3.
|
||||
Proof.
|
||||
pre_process.
|
||||
Exists ptr_2 cap_2 size_2.
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma proof_of_mpz_clear_return_wit_1_4 : mpz_clear_return_wit_1_4.
|
||||
Proof.
|
||||
pre_process.
|
||||
Exists ptr_2 cap_2 size_2.
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_clear_which_implies_wit_1 : mpz_clear_which_implies_wit_1.
|
||||
Proof.
|
||||
pre_process.
|
||||
unfold store_Z.
|
||||
Intros ptr cap size.
|
||||
entailer!.
|
||||
rewrite orp_sepcon_left.
|
||||
Split.
|
||||
+ Right.
|
||||
Exists ptr cap size.
|
||||
entailer!.
|
||||
+ Left.
|
||||
Exists ptr cap size.
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_return_wit_1_1 : mpz_realloc_return_wit_1_1.
|
||||
Proof.
|
||||
pre_process.
|
||||
Right.
|
||||
Exists retval_3 retval_2.
|
||||
entailer!.
|
||||
unfold Zmax in *.
|
||||
pose proof (Z.le_max_l size_pre 1).
|
||||
lia.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_return_wit_1_2 : mpz_realloc_return_wit_1_2.
|
||||
Proof.
|
||||
pre_process.
|
||||
Left.
|
||||
Exists retval_3 retval_2.
|
||||
entailer!.
|
||||
unfold Zmax in *.
|
||||
pose proof (Z.le_max_l size_pre 1).
|
||||
lia.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_return_wit_1_3 : mpz_realloc_return_wit_1_3.
|
||||
Proof.
|
||||
pre_process.
|
||||
Right.
|
||||
Exists retval_3 retval_2.
|
||||
entailer!.
|
||||
+ subst.
|
||||
unfold mpd_store_Z_compact.
|
||||
Intros data.
|
||||
Exists data.
|
||||
unfold mpd_store_list, store_undef_uint_array_rec.
|
||||
entailer!.
|
||||
- assert (Zlength data = 0). {
|
||||
pose proof (Zlength_nonneg data).
|
||||
lia.
|
||||
}
|
||||
rewrite H8 in *.
|
||||
simpl.
|
||||
entailer!.
|
||||
pose proof (Zlength_nil_inv data H8).
|
||||
repeat subst.
|
||||
unfold store_uint_array, store_array; simpl; entailer!.
|
||||
unfold store_undef_uint_array, store_undef_array.
|
||||
rewrite Z.sub_0_r.
|
||||
reflexivity.
|
||||
- unfold Zmax in *.
|
||||
assert (size_pre < 1 \/ size_pre >= 1). { lia. }
|
||||
destruct H8.
|
||||
* rewrite (Z.max_r size_pre 1 ltac:(lia)); lia.
|
||||
* rewrite (Z.max_l size_pre 1 ltac:(lia)); lia.
|
||||
+ pose proof (Z.le_max_l size_pre 1).
|
||||
unfold Zmax in *.
|
||||
lia.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_return_wit_1_4 : mpz_realloc_return_wit_1_4.
|
||||
Proof.
|
||||
pre_process.
|
||||
Left.
|
||||
Exists retval_3 retval_2.
|
||||
entailer!.
|
||||
+ subst.
|
||||
unfold mpd_store_Z_compact, mpd_store_list.
|
||||
Intros data.
|
||||
Exists data.
|
||||
assert (Zlength data = 0). {
|
||||
pose proof (Zlength_nonneg data).
|
||||
lia.
|
||||
}
|
||||
rewrite H8 in *; clear H2.
|
||||
pose proof (Zlength_nil_inv data H8).
|
||||
rewrite H2 in *; clear H2 H8.
|
||||
unfold store_uint_array, store_array.
|
||||
simpl.
|
||||
entailer!.
|
||||
+ pose proof (Z.le_max_l size_pre 1).
|
||||
unfold Zmax in *.
|
||||
lia.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_return_wit_1_5 : mpz_realloc_return_wit_1_5.
|
||||
Proof.
|
||||
pre_process.
|
||||
Left.
|
||||
Exists retval_3 retval_2.
|
||||
entailer!.
|
||||
+ subst.
|
||||
unfold mpd_store_Z_compact, mpd_store_list.
|
||||
Intros data.
|
||||
Exists data.
|
||||
unfold store_uint_array, store_array.
|
||||
simpl.
|
||||
entailer!.
|
||||
+ pose proof (Z.le_max_l size_pre 1).
|
||||
unfold Zmax in *.
|
||||
lia.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_return_wit_1_6 : mpz_realloc_return_wit_1_6.
|
||||
Proof.
|
||||
pre_process.
|
||||
Right.
|
||||
Exists retval_3 retval_2.
|
||||
subst.
|
||||
entailer!.
|
||||
+ unfold mpd_store_Z_compact, mpd_store_list.
|
||||
Intros data; Exists data.
|
||||
unfold store_uint_array, store_array.
|
||||
assert (Zlength data = 0). {
|
||||
pose proof (Zlength_nonneg data).
|
||||
lia.
|
||||
}
|
||||
rewrite H8 in *; clear H2.
|
||||
pose proof (Zlength_nil_inv data H8).
|
||||
rewrite H2 in *; clear H2 H8.
|
||||
unfold store_undef_uint_array, store_undef_uint_array_rec, store_undef_array.
|
||||
subst.
|
||||
simpl.
|
||||
entailer!.
|
||||
- rewrite Z.sub_0_r.
|
||||
entailer!.
|
||||
- pose proof (Z.le_max_r size_pre 1).
|
||||
simpl in H.
|
||||
unfold Zmax in *.
|
||||
lia.
|
||||
+ unfold Zmax in *.
|
||||
pose proof (Z.le_max_l size_pre 1).
|
||||
lia.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_return_wit_1_7 : mpz_realloc_return_wit_1_7.
|
||||
Proof.
|
||||
pre_process.
|
||||
Left.
|
||||
Exists retval_3 retval_2.
|
||||
subst.
|
||||
unfold Zmax in *.
|
||||
rewrite (Z.abs_neq old ltac:(lia)) in H.
|
||||
pose proof (Z.le_max_l size_pre 1).
|
||||
unfold mpd_store_Z_compact.
|
||||
Intros data; entailer!.
|
||||
unfold mpd_store_list.
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_return_wit_1_8 : mpz_realloc_return_wit_1_8.
|
||||
Proof.
|
||||
pre_process.
|
||||
Right.
|
||||
Exists retval_3 retval_2.
|
||||
subst.
|
||||
unfold Zmax in *.
|
||||
rewrite (Z.abs_eq old ltac:(lia)) in H.
|
||||
pose proof (Z.le_max_l size_pre 1).
|
||||
unfold mpd_store_Z_compact.
|
||||
Intros data; entailer!.
|
||||
unfold mpd_store_list.
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_3_pure : mpz_realloc_partial_solve_wit_3_pure.
|
||||
Proof.
|
||||
pre_process.
|
||||
unfold Zmax in *.
|
||||
pose proof (Z.le_max_l size_pre 1).
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_4_pure : mpz_realloc_partial_solve_wit_4_pure.
|
||||
Proof.
|
||||
pre_process.
|
||||
unfold Zmax in *.
|
||||
pose proof (Z.le_max_l size_pre 1).
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_5_pure : mpz_realloc_partial_solve_wit_5_pure.
|
||||
Proof.
|
||||
pre_process.
|
||||
unfold Zmax in *.
|
||||
pose proof (Z.le_max_l size_pre 1).
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_6_pure : mpz_realloc_partial_solve_wit_6_pure.
|
||||
Proof.
|
||||
pre_process.
|
||||
unfold Zmax in *.
|
||||
pose proof (Z.le_max_l size_pre 1).
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_7_pure : mpz_realloc_partial_solve_wit_7_pure.
|
||||
Proof.
|
||||
pre_process.
|
||||
unfold mpd_store_Z_compact, mpd_store_list.
|
||||
Intros data.
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_8_pure : mpz_realloc_partial_solve_wit_8_pure.
|
||||
Proof.
|
||||
pre_process.
|
||||
unfold mpd_store_Z_compact, mpd_store_list.
|
||||
Intros data.
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_9_pure : mpz_realloc_partial_solve_wit_9_pure.
|
||||
Proof.
|
||||
pre_process.
|
||||
unfold mpd_store_Z_compact, mpd_store_list.
|
||||
Intros data.
|
||||
entailer!.
|
||||
Qed.
|
||||
|
||||
Lemma proof_of_mpz_realloc_partial_solve_wit_10_pure : mpz_realloc_partial_solve_wit_10_pure.
|
||||
Proof.
|
||||
pre_process.
|
||||
unfold mpd_store_Z_compact, mpd_store_list.
|
||||
Intros data.
|
||||
entailer!.
|
||||
Qed.
|
||||
|
Reference in New Issue
Block a user