This commit is contained in:
2025-06-13 12:50:34 +00:00
13 changed files with 3596 additions and 263 deletions

View File

@ -15,7 +15,9 @@ SETS_FLAG = -R $(SETS_DIR) SetsClass
COMPCERT_FLAG = -R $(COMPCERT_DIR) compcert.lib COMPCERT_FLAG = -R $(COMPCERT_DIR) compcert.lib
DEP_FLAG = -R $(PV_DIR) PV -R $(SETS_DIR) SetsClass -R $(COMPCERT_DIR) compcert.lib GMP_FLAG = $(shell grep -E '^-([RQ])' _CoqProject)
DEP_FLAG = $(GMP_FLAG)
SETS_FILE_NAMES = \ SETS_FILE_NAMES = \
SetsClass.v SetsDomain.v SetElement.v RelsDomain.v SetElementProperties.v SetProd.v SetsClass_AxiomFree.v SetsDomain_Classic.v SetsClass.v SetsDomain.v SetElement.v RelsDomain.v SetElementProperties.v SetProd.v SetsClass_AxiomFree.v SetsDomain_Classic.v
@ -35,21 +37,20 @@ PV_FILE_NAMES = \
PV_FILES=$(PV_FILE_NAMES:%.v=$(PV_DIR)/%.v) PV_FILES=$(PV_FILE_NAMES:%.v=$(PV_DIR)/%.v)
GMP_FILES = \
./projects/lib/GmpAux.v ./projects/lib/GmpNumber.v \
./projects/lib/gmp_goal.v
FILES = $(PV_FILES) \ FILES = $(PV_FILES) \
$(SETS_FILES) \ $(GMP_FILES)
$(COMPCERT_FILES)
$(SETS_FILES:%.v=%.vo): %.vo: %.v
@echo COQC $<
@$(COQC) $(SETS_FLAG) $<
$(COMPCERT_FILES:%.v=%.vo): %.vo: %.v
@echo COQC $<
@$(COQC) $(COMPCERT_FLAG) $<
$(PV_FILES:%.v=%.vo): %.vo: %.v $(PV_FILES:%.v=%.vo): %.vo: %.v
@echo COQC $(<F) @echo COQC $(<F)
@$(COQC) $(PV_FLAG) $< @$(COQC) $(GMP_FLAG) $<
$(GMP_FILES:%.v=%.vo): %.vo: %.v
@echo COQC $(<F)
@$(COQC) $(GMP_FLAG) $<
all: $(FILES:%.v=%.vo) all: $(FILES:%.v=%.vo)

View File

@ -0,0 +1,126 @@
//test curent rules
//max rule id:30
#include "int_array_def.h"
#include "verification_list.h"
#include "verification_stdlib.h"
id : 1
priority : core(1)
left : store_int_array(?p, ?n, ?l) at 0
right : data_at(?p + (?i * sizeof(I32)), I32, ?v) at 1
check : infer(0 <= i);
infer(i < n);
action : right_erase(1);
left_erase(0);
left_add(store_int_array_missing_i_rec(p, i, 0, n, l));
right_add(v == l[i]);
id : 2
priority : post(1)
left : store_int_array_missing_i_rec(?p, ?i, 0, ?n, ?l) at 0
data_at(p + i * sizeof(I32), I32, l[i]) at 1
check : infer(0 <= i);
infer(i < n);
action : left_erase(1);
left_erase(0);
left_add(store_int_array(p, n, l));
id : 3
priority : post(3)
left : store_int_array_missing_i_rec(?p, ?i, 0, ?n, ?l) at 0
data_at(p + i * sizeof(I32), I32, ?v) at 1
check : infer(0 <= i);
infer(i < n);
action : left_erase(1);
left_erase(0);
left_add(store_int_array(p, n, replace_Znth{Z}(i, v, l)));
id : 4
priority : core(1)
left : store_int_array(?p, ?n, ?l1) at 0
right : store_int_array(p, n, ?l2) at 1
action : right_erase(1);
left_erase(0);
right_add(l1 == l2);
id : 5
priority : core(1)
left : store_int_array_missing_i_rec(?p, ?i, ?v, ?n, ?l) at 0
right : store_int_array_missing_i_rec(p, i, v, n, l) at 1
action : left_erase(0);
right_erase(1);
id : 6
priority : core(1)
left : store_int_array_rec(?p, ?x, ?y, ?l1) at 0
right : store_int_array_rec(?p, ?x, ?y, ?l2) at 1
action : right_erase(1);
left_erase(0);
right_add(l1 == l2);
id : 7
priority : core(1)
left : store_int_array_rec(?p, ?x, ?y, ?l) at 0
right : data_at(?p + (?i * sizeof(I32)), I32, ?v) at 1
check : infer(x <= i);
infer(i < y);
action : right_erase(1);
left_erase(0);
left_add(store_int_array_missing_i_rec(p, i, x, y, l));
right_add(v == l[i - x]);
id : 8
priority : core(1)
left : store_int_array_rec(?p, ?x, ?y, ?l1) at 0
store_int_array_rec(?p, ?y, ?z, ?l2) at 1
right : store_int_array_rec(?p, ?x, ?z, ?l3) at 2
check : infer(y <= z);
infer(x <= y);
action : right_erase(2);
left_erase(0);
left_erase(1);
right_add(l3 == app{Z}(l1, l2));
id : 9
priority : core(1)
left : store_int_array_rec(?p, ?x, ?z, ?l3) at 2
right : store_int_array_rec(?p, ?x, ?y, ?l1) at 0
store_int_array_rec(?p, ?y, ?z, ?l2) at 1
check : infer(y <= z);
infer(x <= y);
action : left_erase(2);
right_erase(0);
right_erase(1);
right_add(l3 == app{Z}(l1, l2));
right_add(Zlength{Z}(l1) == y - x);
id : 10
priority : core(1)
right : store_int_array_rec(?p, ?x, ?x, ?l) at 0
action : right_erase(0);
right_add(l == nil{Z});
id : 11
priority : post(1)
left : store_int_array_missing_i_rec(?p, ?i, ?x, ?y, ?l) at 0
data_at(p + ?i * sizeof(I32), I32, l[?i - ?x]) at 1
check : infer(x <= i);
infer(i < y);
action : left_erase(1);
left_erase(0);
left_add(store_int_array_rec(p, x, y, l));
id : 12
priority : post(4)
left : store_int_array_missing_i_rec(?p, ?i, ?x, ?y, ?l) at 0
data_at(p + ?i * sizeof(I32), I32, ?v) at 1
check : infer(x <= i);
infer(i < y);
action : left_erase(1);
left_erase(0);
left_add(store_int_array_rec(p, x, y, replace_Znth{Z}(i - x, v, l)));

19
projects/int_array_def.h Normal file
View File

@ -0,0 +1,19 @@
/*@ Extern Coq (sum : list Z -> Z)
(sublist : {A} -> Z -> Z -> list A -> list A)
(store_int_array : Z -> Z -> list Z -> Assertion)
(store_uint_array : Z -> Z -> list Z -> Assertion)
(store_undef_uint_array_rec : Z -> Z -> Z -> Assertion)
(store_int_array_missing_i_rec: Z -> Z -> Z -> Z -> list Z -> Assertion)
(store_uint_array_missing_i_rec: Z -> Z -> Z -> Z -> list Z -> Assertion)
(store_array_box: Z -> Z -> list Z -> Assertion)
(store_array_box_array: Z -> list Z -> Assertion)
(store_int_array_rec: Z -> Z -> Z -> list Z -> Assertion)
(store_uint_array_rec: Z -> Z -> Z -> list Z -> Assertion)
(Znth: {A} -> Z -> list A -> A -> A)
(replace_Znth: {A} -> Z -> A -> list A -> list A)
(zeros: Z -> list Z)
*/
/*@ include strategies "int_array.strategies" */
/*@ include strategies "uint_array.strategies" */

345
projects/lib/GmpAux.v Executable file
View File

@ -0,0 +1,345 @@
Require Import Coq.ZArith.ZArith.
Require Import Coq.Bool.Bool.
Require Import Coq.Lists.List.
Require Import Coq.Classes.RelationClasses.
Require Import Coq.Classes.Morphisms.
Require Import Coq.micromega.Psatz.
Require Import Permutation.
Require Import String.
From AUXLib Require Import int_auto Axioms Feq Idents List_lemma VMap.
Require Import SetsClass.SetsClass. Import SetsNotation.
From SimpleC.SL Require Import CommonAssertion Mem SeparationLogic IntLib.
Require Import Logic.LogicGenerator.demo932.Interface.
Local Open Scope Z_scope.
Local Open Scope sets.
Import ListNotations.
Local Open Scope list.
Require Import String.
Local Open Scope string.
Import naive_C_Rules.
Local Open Scope sac.
Module Aux.
Lemma Z_of_nat_succ: forall (n: nat),
Z.of_nat (S n) = Z.of_nat n + 1.
Proof. lia. Qed.
Lemma Zpow_add_1: forall (a b: Z),
a >= 0 -> b >= 0 ->
a ^ (b + 1) = a ^ b * a.
Proof.
intros.
rewrite (Z.pow_add_r a b 1); lia.
Qed.
Lemma Zmul_mod_cancel: forall (n a b: Z),
n >= 0 -> a > 0 -> b >= 0 ->
(n * a) mod (a ^ (b + 1)) = a * (n mod (a ^ b)).
Proof.
intros.
pose proof (Z_div_mod_eq_full n (a ^ b)).
pose proof (Z.mod_bound_pos n (a ^ b) ltac:(lia) ltac:(nia)).
remember (n / a ^ b) as q eqn:Hq.
remember (n mod a ^ b) as rem eqn:Hrem.
rewrite H2.
rewrite Z.mul_add_distr_r.
rewrite (Z.mul_comm (a ^ b) q); rewrite <-Z.mul_assoc.
rewrite <-Zpow_add_1; try lia.
assert (0 <= rem * a < a ^ (b + 1)). {
rewrite Zpow_add_1; try lia.
nia.
}
rewrite <-(Zmod_unique_full (q * a ^ (b + 1) + rem * a) (a ^ (b + 1)) q (rem * a)).
+ lia.
+ unfold Remainder.
lia.
+ lia.
Qed.
Lemma Zdiv_mod_pow: forall (n a b: Z),
a > 0 -> b >= 0 -> n >= 0 ->
(n / a) mod (a ^ b) = (n mod (a ^ (b + 1))) / a.
Proof.
intros.
pose proof (Z_div_mod_eq_full n (a ^ (b + 1))).
remember (n / (a ^ (b + 1))) as q eqn:Hq.
remember (n mod a ^ (b + 1)) as rem eqn:Hrem.
assert (n / a = a ^ b * q + rem / a). {
rewrite H2.
rewrite Zpow_add_1; try lia.
pose proof (Z_div_plus_full_l (a ^ b * q) a rem ltac:(lia)).
assert (a ^ b * q * a + rem = a ^ b * a * q + rem). { lia. }
rewrite H4 in H3.
tauto.
}
apply Znumtheory.Zdivide_mod_minus.
+ pose proof (Z.mod_bound_pos n (a ^ (b + 1)) ltac:(lia) (Z.pow_pos_nonneg a (b + 1) ltac:(lia) ltac:(lia))).
rewrite <-Hrem in H4.
rewrite Zpow_add_1 in H4; try lia.
pose proof (Z.div_lt_upper_bound rem a (a ^ b) ltac:(lia) ltac:(lia)).
split; try lia.
apply (Z_div_pos rem a ltac:(lia) ltac:(lia)).
+ unfold Z.divide.
exists q.
lia.
Qed.
Lemma list_app_cons: forall (l1 l2: list Z) (a: Z),
app l1 (a :: l2) = app (app l1 (a :: nil)) l2.
Proof.
intros.
revert a l2.
induction l1.
+ intros.
rewrite app_nil_l.
reflexivity.
+ intros.
simpl in *.
specialize (IHl1 a0 l2).
rewrite IHl1.
reflexivity.
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 list_last_cons: forall (a: Z) (l: list Z),
l <> nil ->
last (a :: l) 1 = last l 1.
Proof.
intros.
destruct l.
+ contradiction.
+ simpl.
reflexivity.
Qed.
Lemma list_last_to_Znth: forall (l: list Z),
l <> nil ->
last l 1 = Znth ((Zlength l) - 1) l 0.
Proof.
intros.
induction l.
+ contradiction.
+ destruct l.
- simpl.
rewrite Znth0_cons.
lia.
- pose proof (@nil_cons Z z l).
specialize (IHl ltac:(auto)); clear H0.
rewrite list_last_cons; [ | pose proof (@nil_cons Z z l); auto ].
rewrite IHl.
pose proof (Zlength_cons a (z :: l)).
unfold Z.succ in H0; rewrite H0; clear H0.
pose proof (Zlength_nonneg l).
pose proof (Zlength_cons z l); unfold Z.succ in H1.
pose proof (Znth_cons (Zlength (z :: l)) a (z :: l) 0 ltac:(lia)).
assert (Zlength (z :: l) + 1 - 1 = Zlength (z :: l)). { lia. }
rewrite H3; clear H3.
rewrite H2.
reflexivity.
Qed.
Lemma Zlength_removelast: forall (l: list Z),
l <> [] ->
Zlength (removelast l) = Zlength l - 1.
Proof.
intros.
induction l.
+ contradiction.
+ destruct l.
- simpl.
rewrite Zlength_nil.
lia.
- pose proof (@nil_cons Z z l).
specialize (IHl ltac:(auto)).
assert (removelast (a :: z :: l) = a :: removelast(z :: l)). {
simpl.
reflexivity.
}
rewrite H1; clear H1.
repeat rewrite Zlength_cons; unfold Z.succ.
rewrite IHl.
rewrite Zlength_cons; unfold Z.succ.
lia.
Qed.
Lemma store_array_rec_false: forall x storeA lo hi (l: list Z),
lo > hi ->
store_array_rec storeA x lo hi l |-- [| False |].
Proof.
intros.
revert x storeA lo hi H.
induction l; intros.
+ simpl.
entailer!.
+ simpl.
specialize (IHl x storeA (lo + 1) hi ltac:(lia)).
sep_apply IHl.
entailer!.
Qed.
Lemma store_array_rec_empty: forall x storeA lo (l: list Z),
store_array_rec storeA x lo lo l |-- emp && [| l = nil |].
Proof.
intros.
destruct l.
+ simpl.
entailer!.
+ simpl.
sep_apply store_array_rec_false; [ entailer! | lia ].
Qed.
Lemma store_uint_array_rec_false: forall x lo hi l,
lo > hi ->
store_uint_array_rec x lo hi l |-- [| False |].
Proof.
intros.
unfold store_uint_array_rec.
sep_apply store_array_rec_false; [ entailer! | lia ].
Qed.
Lemma store_uint_array_rec_empty: forall x lo l,
store_uint_array_rec x lo lo l |-- emp && [| l = nil |].
Proof.
induction l.
+ unfold store_uint_array_rec.
simpl.
entailer!.
+ pose proof (store_uint_array_rec_false x (lo + 1) lo l ltac:(lia)).
unfold store_uint_array_rec in *.
simpl in *.
sep_apply H.
entailer!.
Qed.
Lemma store_uint_array_empty: forall x l,
store_uint_array x 0 l |-- emp && [| l = nil |].
Proof.
intros x l.
revert x.
induction l; intros.
+ unfold store_uint_array, store_array.
simpl.
entailer!.
+ unfold store_uint_array, store_array.
simpl.
sep_apply store_array_rec_false; [ entailer! | lia ].
Qed.
Lemma uint_array_rec_to_uint_array: forall x lo hi l,
lo <= hi ->
store_uint_array_rec x lo hi l --||--
store_uint_array (x + sizeof(UINT) * lo) (hi - lo) l.
Proof.
intros.
unfold store_uint_array_rec, store_uint_array, store_array.
pose proof (store_array_rec_base x 0 lo hi l (sizeof(UINT))
store_uint
(fun (x: addr) (lo a: Z) =>
(x + lo * sizeof(UINT)) # UInt |-> a) ltac:(reflexivity)).
assert (x + sizeof(UINT) * lo = x + lo * sizeof(UINT)). { lia. }
rewrite H1; clear H1.
assert (0 + lo = lo). { lia. }
repeat rewrite H1 in H0; clear H1.
destruct H0.
split.
+ sep_apply H0.
entailer!.
+ sep_apply H1.
entailer!.
Qed.
Lemma store_uint_array_single: forall x n a,
store_uint_array_rec x n (n + 1) [a] --||--
(x + n * sizeof(UINT)) # UInt |-> a.
Proof.
intros.
unfold store_uint_array_rec.
simpl.
split; entailer!.
Qed.
Lemma store_undef_uint_array_single: forall x n a,
(x + n * sizeof(UINT)) # UInt |-> a |--
store_undef_uint_array_rec x n (n + 1).
Proof.
intros.
unfold store_undef_uint_array_rec.
assert (Z.to_nat (n + 1 - n) = S O). { lia. }
rewrite H; clear H.
simpl.
entailer!.
sep_apply store_uint_undef_store_uint.
entailer!.
Qed.
Lemma store_uint_array_single_to_undef: forall x n a,
store_uint_array_rec x n (n + 1) [a] |--
store_undef_uint_array_rec x n (n + 1).
Proof.
intros.
pose proof (store_uint_array_single x n a).
destruct H as [H _].
sep_apply H.
sep_apply store_undef_uint_array_single.
entailer!.
Qed.
Lemma store_uint_array_divide_rec: forall x n l m,
0 <= m <= n ->
Zlength l = n ->
store_uint_array x n l --||--
store_uint_array x m (sublist 0 m l) **
store_uint_array_rec x m n (sublist m n l).
Proof.
intros.
pose proof (store_uint_array_divide x n l m H H0).
pose proof (uint_array_rec_to_uint_array x m n (sublist m n l) ltac:(lia)).
destruct H2 .
destruct H1.
rewrite Z.mul_comm in H2, H3.
rewrite H3 in H1.
rewrite <-H2 in H4.
clear H2 H3.
split; tauto.
Qed.
Lemma store_undef_uint_array_rec_divide: forall x l mid r,
0 <= l <= r ->
l <= mid <= r ->
store_undef_uint_array_rec x l r --||--
store_undef_uint_array_rec x l mid ** store_undef_uint_array_rec x mid r.
Proof.
intros.
unfold store_undef_uint_array_rec.
pose proof (store_undef_array_divide (x + l * sizeof(UINT)) (r - l) (mid - l) (sizeof(UINT))
(fun x => x # UInt |->_)
(fun x0 lo => (x0 + lo * sizeof(UINT)) # UInt |->_) ltac:(lia) ltac:(reflexivity)).
unfold store_undef_array in H1.
pose (fun (l: Z) (n: Z) (len: nat) => store_undef_array_rec_base x 0 l n len (sizeof(UINT))
(fun x => x # UInt |->_)
(fun x0 lo => (x0 + lo * sizeof(UINT)) # UInt |->_) ltac:(reflexivity)).
rewrite <-(l0 l r (Z.to_nat (r - l))) in H1.
rewrite <-(l0 l mid (Z.to_nat (mid - l))) in H1.
assert (r - l - (mid - l) = r - mid). { lia. }
rewrite H2 in H1; clear H2.
rewrite <-Z.add_assoc, <-Z.mul_add_distr_r in H1.
assert (l + (mid - l) = mid). { lia. }
rewrite H2 in H1; clear H2.
rewrite <-(l0 mid r (Z.to_nat (r - mid))) in H1.
repeat rewrite Z.add_0_l in H1.
clear l0.
rewrite H1.
split; entailer!.
Qed.
End Aux.

View File

@ -9,6 +9,7 @@ Require Import String.
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 CommonAssertion Mem SeparationLogic IntLib. From SimpleC.SL Require Import CommonAssertion Mem SeparationLogic IntLib.
Require Import GmpLib.GmpAux.
Require Import Logic.LogicGenerator.demo932.Interface. Require Import Logic.LogicGenerator.demo932.Interface.
Local Open Scope Z_scope. Local Open Scope Z_scope.
Local Open Scope sets. Local Open Scope sets.
@ -20,97 +21,14 @@ Import naive_C_Rules.
Local Open Scope sac. Local Open Scope sac.
Notation "'UINT_MOD'" := (4294967296). Notation "'UINT_MOD'" := (4294967296).
Notation "'LENGTH_MAX'" := (100000000).
Module Aux.
Lemma Z_of_nat_succ: forall (n: nat),
Z.of_nat (S n) = Z.of_nat n + 1.
Proof. lia. Qed.
Lemma Zpow_add_1: forall (a b: Z),
a >= 0 -> b >= 0 ->
a ^ (b + 1) = a ^ b * a.
Proof.
intros.
rewrite (Z.pow_add_r a b 1); lia.
Qed.
Lemma Zmul_mod_cancel: forall (n a b: Z),
n >= 0 -> a > 0 -> b >= 0 ->
(n * a) mod (a ^ (b + 1)) = a * (n mod (a ^ b)).
Proof.
intros.
pose proof (Z_div_mod_eq_full n (a ^ b)).
pose proof (Z.mod_bound_pos n (a ^ b) ltac:(lia) ltac:(nia)).
remember (n / a ^ b) as q eqn:Hq.
remember (n mod a ^ b) as rem eqn:Hrem.
rewrite H2.
rewrite Z.mul_add_distr_r.
rewrite (Z.mul_comm (a ^ b) q); rewrite <-Z.mul_assoc.
rewrite <-Zpow_add_1; try lia.
assert (0 <= rem * a < a ^ (b + 1)). {
rewrite Zpow_add_1; try lia.
nia.
}
rewrite <-(Zmod_unique_full (q * a ^ (b + 1) + rem * a) (a ^ (b + 1)) q (rem * a)).
+ lia.
+ unfold Remainder.
lia.
+ lia.
Qed.
Lemma Zdiv_mod_pow: forall (n a b: Z),
a > 0 -> b >= 0 -> n >= 0 ->
(n / a) mod (a ^ b) = (n mod (a ^ (b + 1))) / a.
Proof.
intros.
pose proof (Z_div_mod_eq_full n (a ^ (b + 1))).
remember (n / (a ^ (b + 1))) as q eqn:Hq.
remember (n mod a ^ (b + 1)) as rem eqn:Hrem.
assert (n / a = a ^ b * q + rem / a). {
rewrite H2.
rewrite Zpow_add_1; try lia.
pose proof (Z_div_plus_full_l (a ^ b * q) a rem ltac:(lia)).
assert (a ^ b * q * a + rem = a ^ b * a * q + rem). { lia. }
rewrite H4 in H3.
tauto.
}
apply Znumtheory.Zdivide_mod_minus.
+ pose proof (Z.mod_bound_pos n (a ^ (b + 1)) ltac:(lia) (Z.pow_pos_nonneg a (b + 1) ltac:(lia) ltac:(lia))).
rewrite <-Hrem in H4.
rewrite Zpow_add_1 in H4; try lia.
pose proof (Z.div_lt_upper_bound rem a (a ^ b) ltac:(lia) ltac:(lia)).
split; try lia.
apply (Z_div_pos rem a ltac:(lia) ltac:(lia)).
+ unfold Z.divide.
exists q.
lia.
Qed.
Lemma list_app_cons: forall (l1 l2: list Z) (a: Z),
app l1 (a :: l2) = app (app l1 (a :: nil)) l2.
Proof.
intros.
revert a l2.
induction l1.
+ intros.
rewrite app_nil_l.
reflexivity.
+ intros.
simpl in *.
specialize (IHl1 a0 l2).
rewrite IHl1.
reflexivity.
Qed.
End Aux.
Module Internal. Module Internal.
Definition mpd_store_list (ptr: addr) (data: list Z) (cap: Z): Assertion := Definition mpd_store_list (ptr: addr) (data: list Z) (cap: Z): Assertion :=
[| Zlength data <= cap |] && [| Zlength data <= cap |] && [| cap <= LENGTH_MAX |] &&
store_uint_array ptr (Zlength data) data && store_uint_array ptr (Zlength data) data **
store_undef_uint_array_rec ptr ((Zlength data) + 1) cap. store_undef_uint_array_rec ptr (Zlength data) cap.
Fixpoint list_to_Z (data: list Z): Z := Fixpoint list_to_Z (data: list Z): Z :=
match data with match data with
@ -129,7 +47,47 @@ Definition list_store_Z (data: list Z) (n: Z): Prop :=
Definition mpd_store_Z (ptr: addr) (n: Z) (size: Z) (cap: Z): Assertion := Definition mpd_store_Z (ptr: addr) (n: Z) (size: Z) (cap: Z): Assertion :=
EX data, EX data,
mpd_store_list ptr data cap && [| list_store_Z data n|] && [| size = Zlength data |]. [| list_store_Z data n |] && [| size = Zlength data |] && mpd_store_list ptr data cap.
Definition list_store_Z_compact (data: list Z) (n: Z): Prop :=
list_to_Z data = n /\ last data 1 >= 1 /\ list_within_bound data.
Definition mpd_store_Z_compact (ptr: addr) (n size cap: Z): Assertion :=
EX data,
[| list_store_Z_compact data n |] && [| size = Zlength data |] && mpd_store_list ptr data cap.
Lemma list_store_Z_normal_to_compact: forall (data: list Z) (n: Z),
list_store_Z data n ->
last data 1 >= 1 ->
list_store_Z_compact data n.
Proof.
unfold list_store_Z_compact, list_store_Z.
intros.
tauto.
Qed.
Lemma list_store_Z_compact_to_normal: forall data n,
list_store_Z_compact data n ->
list_store_Z data n.
Proof.
intros.
unfold list_store_Z_compact in H.
unfold list_store_Z.
tauto.
Qed.
Lemma list_store_Z_injection: forall l1 l2 n1 n2,
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 ->
@ -164,6 +122,29 @@ Proof.
tauto. tauto.
Qed. Qed.
Lemma list_within_bound_Znth: forall (l: list Z) (i: Z),
0 <= i < Zlength l ->
list_within_bound l ->
0 <= Znth i l 0 < UINT_MOD.
Proof.
intros.
revert i H.
induction l; intros.
+ rewrite Zlength_nil in H.
lia.
+ assert (i = 0 \/ i > 0). { lia. }
destruct H1.
- rewrite H1.
rewrite (Znth0_cons a l 0).
simpl in H0.
lia.
- rewrite Znth_cons; try lia.
simpl in H0; destruct H0.
rewrite Zlength_cons in H; unfold Z.succ in H.
specialize (IHl H2 (i - 1) ltac:(lia)).
lia.
Qed.
Lemma __list_within_bound_split_r: forall (l1: list Z) (a: Z), Lemma __list_within_bound_split_r: forall (l1: list Z) (a: Z),
list_within_bound (l1 ++ [a]) -> list_within_bound (l1 ++ [a]) ->
list_within_bound l1 /\ 0 <= a < UINT_MOD. list_within_bound l1 /\ 0 <= a < UINT_MOD.
@ -371,6 +352,243 @@ Proof.
- pose proof (Zlength_nonneg l1); lia. - pose proof (Zlength_nonneg l1); lia.
Qed. Qed.
Lemma list_store_Z_nth: forall (l: list Z) (n: Z) (i: Z),
0 <= i < Zlength l ->
list_store_Z l n ->
Znth i l 0 = (n / (UINT_MOD ^ i)) mod UINT_MOD.
Proof.
intros.
revert n i H H0.
induction l; intros.
+ rewrite Zlength_nil in H.
lia.
+ rewrite Zlength_cons in H; unfold Z.succ in H.
assert (i = 0 \/ i > 0). { lia. }
destruct H1.
- pose proof (list_store_Z_split [a] l n).
assert (a :: l = app [a] l). { auto. }
rewrite <-H3 in H2; clear H3.
specialize (H2 H0); destruct H2.
unfold list_store_Z, list_to_Z in H2.
unfold Zlength, Zlength_aux, Z.succ in H2.
destruct H2.
rewrite (Aux.Zpow_add_1 UINT_MOD 0) in H2; try lia.
rewrite Z.pow_0_r, Z.mul_1_l in H2.
simpl in H2.
rewrite H1; rewrite Znth0_cons.
rewrite Z.pow_0_r, Z.div_1_r.
lia.
- rewrite Znth_cons; [ | lia ].
unfold list_store_Z in H0; destruct H0.
simpl in H0.
simpl in H2.
unfold list_store_Z in IHl.
assert (list_to_Z l = (n - a) / UINT_MOD /\ list_within_bound l). {
rewrite (Z.div_unique_exact (n - a) UINT_MOD (list_to_Z l)); try lia; try tauto.
}
specialize (IHl ((n - a) / UINT_MOD) (i - 1) ltac:(lia) H3).
rewrite IHl.
assert ((n - a) / UINT_MOD / (UINT_MOD ^ (i - 1)) = (n - a) / (UINT_MOD ^ 1 * UINT_MOD ^ (i - 1))). {
rewrite Zdiv_Zdiv; try lia.
rewrite Z.pow_1_r.
reflexivity.
}
rewrite H4.
rewrite <-Z.pow_add_r; try lia.
assert (n / UINT_MOD ^ i = (n - a) / UINT_MOD ^ i). {
assert (i = 1 + (i - 1)). { lia. }
rewrite H5.
rewrite (Z.pow_add_r UINT_MOD 1 (i - 1)); try lia.
repeat rewrite <-Zdiv_Zdiv; try lia.
repeat rewrite Z.pow_1_r.
rewrite <-(Zdiv_unique_full n UINT_MOD (list_to_Z l) a); try lia.
+ destruct H3.
rewrite <-H3.
reflexivity.
+ unfold Remainder; lia.
}
rewrite H5.
assert (1 + (i - 1) = i). { lia. }
rewrite H6.
reflexivity.
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.
Lemma list_store_Z_compact_cmp: forall l1 l2 n1 n2 i,
0 <= i < Zlength l1 ->
Zlength l1 = Zlength l2 ->
list_store_Z_compact l1 n1 ->
list_store_Z_compact l2 n2 ->
sublist (i + 1) (Zlength l1) l1 = sublist (i + 1) (Zlength l2) l2 ->
Znth i l1 0 < Znth i l2 0 ->
n1 < n2.
Proof.
intros.
apply list_store_Z_compact_to_normal in H1, H2.
apply (list_store_Z_cmp l1 l2 n1 n2 i); tauto.
Qed.
Lemma list_store_Z_cmp_length: forall l1 l2 n1 n2,
Zlength l1 < Zlength l2 ->
list_store_Z_compact l1 n1 ->
list_store_Z_compact l2 n2 ->
n1 < n2.
Proof.
intros.
unfold list_store_Z_compact in *.
destruct H0, H1, H2, H3.
assert (list_store_Z l1 n1 /\ list_store_Z l2 n2). {
unfold list_store_Z.
tauto.
}
clear H0 H1 H4 H5.
destruct H6.
assert (l2 <> []). {
pose proof (Zlength_nonneg l1).
assert (Zlength l2 >= 1). { lia. }
destruct l2.
+ rewrite Zlength_nil in H5.
lia.
+ pose proof (@nil_cons Z z l2).
auto.
}
pose proof (list_store_Z_bound l2 n2 H1) as Hn2.
pose proof (@app_removelast_last Z l2 1 H4).
rewrite H5 in H1.
apply list_store_Z_split in H1; destruct H1.
rewrite (Aux.Zlength_removelast l2) in H1, H6; try auto.
remember (Zlength l2 - 1) as n eqn:Hn.
unfold list_store_Z in H6; destruct H6.
simpl in H6.
assert (n2 >= UINT_MOD ^ n). {
assert (n2 / UINT_MOD ^ n >= 1). { lia. }
pose proof (Zmult_ge_compat_r (n2 / UINT_MOD ^ n) 1 (UINT_MOD ^ n) ltac:(lia) ltac:(lia)).
rewrite Z.mul_1_l in H9.
assert (n2 >= n2 / UINT_MOD ^ n * UINT_MOD ^ n). {
assert (n >= 0). {
destruct l2.
+ contradiction.
+ rewrite Zlength_cons in Hn; unfold Z.succ in Hn.
pose proof (Zlength_nonneg l2).
lia.
}
pose proof (Zmod_eq n2 (UINT_MOD ^ n) ltac:(lia)).
assert (n2 mod UINT_MOD ^ n >= 0). {
pose proof (Z.mod_bound_pos n2 (UINT_MOD ^ n) ltac:(lia) ltac:(lia)).
lia.
}
lia.
}
lia.
}
assert (Zlength l1 <= n). { lia. }
pose proof (list_store_Z_bound l1 n1 H0).
assert (UINT_MOD ^ n >= UINT_MOD ^ (Zlength l1)). {
assert (Zlength l1 = n \/ Zlength l1 < n). { lia. }
destruct H11.
+ rewrite H11.
lia.
+ assert (n >= 0). {
destruct l2.
+ contradiction.
+ rewrite Zlength_cons in Hn; unfold Z.succ in Hn.
pose proof (Zlength_nonneg l2).
lia.
}
pose proof (Z.pow_lt_mono_r_iff UINT_MOD (Zlength l1) n ltac:(lia) ltac:(lia)).
destruct H13 as [H13 _].
specialize (H13 H11).
lia.
}
lia.
Qed.
End Internal. End Internal.
Record bigint_ent: Type := { Record bigint_ent: Type := {

View File

@ -1,63 +0,0 @@
Require Import Coq.ZArith.ZArith.
Require Import Coq.Bool.Bool.
Require Import Coq.Lists.List.
Require Import Coq.Classes.RelationClasses.
Require Import Coq.Classes.Morphisms.
Require Import Coq.micromega.Psatz.
Require Import Permutation.
Require Import String.
From AUXLib Require Import int_auto Axioms Feq Idents List_lemma VMap.
Require Import SetsClass.SetsClass. Import SetsNotation.
From SimpleC.SL Require Import CommonAssertion Mem SeparationLogic IntLib.
Require Import Logic.LogicGenerator.demo932.Interface.
Local Open Scope Z_scope.
Local Open Scope sets.
Import ListNotations.
Local Open Scope list.
Require Import String.
Local Open Scope string.
Import naive_C_Rules.
Local Open Scope sac.
Module Aux.
Lemma Zpow_add_1: forall (a b: Z),
a >= 0 -> b >= 0 ->
a ^ (b + 1) = a ^ b * a.
Proof.
intros.
rewrite (Z.pow_add_r a b 1); lia.
Qed.
Lemma Z_of_nat_succ: forall (n: nat),
Z.of_nat (S n) = Z.of_nat n + 1.
Lemma Zdiv_mod_pow: forall (n a b: Z),
a > 0 -> b >= 0 -> n >= 0 ->
(n / a) mod (a ^ b) = (n mod (a ^ (b + 1))) / a.
Proof.
intros.
pose proof (Z_div_mod_eq_full n (a ^ (b + 1))).
remember (n / (a ^ (b + 1))) as q eqn:Hq.
remember (n mod a ^ (b + 1)) as rem eqn:Hrem.
assert (n / a = a ^ b * q + rem / a). {
rewrite H2.
rewrite Zpow_add_1; try lia.
pose proof (Z_div_plus_full_l (a ^ b * q) a rem ltac:(lia)).
assert (a ^ b * q * a + rem = a ^ b * a * q + rem). { lia. }
rewrite H4 in H3.
tauto.
}
apply Znumtheory.Zdivide_mod_minus.
+ pose proof (Z.mod_bound_pos n (a ^ (b + 1)) ltac:(lia) (Z.pow_pos_nonneg a (b + 1) ltac:(lia) ltac:(lia))).
rewrite <-Hrem in H4.
rewrite Zpow_add_1 in H4; try lia.
pose proof (Z.div_lt_upper_bound rem a (a ^ b) ltac:(lia) ltac:(lia)).
split; try lia.
apply (Z_div_pos rem a ltac:(lia) ltac:(lia)).
+ unfold Z.divide.
exists q.
lia.
Qed.
End Aux.

1827
projects/lib/gmp_goal.v Normal file

File diff suppressed because it is too large Load Diff

View File

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

View File

@ -0,0 +1,143 @@
Require Import Coq.ZArith.ZArith.
Require Import Coq.Bool.Bool.
Require Import Coq.Strings.String.
Require Import Coq.Lists.List.
Require Import Coq.Classes.RelationClasses.
Require Import Coq.Classes.Morphisms.
Require Import Coq.micromega.Psatz.
Require Import Coq.Sorting.Permutation.
From AUXLib Require Import int_auto Axioms Feq Idents List_lemma VMap.
Require Import SetsClass.SetsClass. Import SetsNotation.
From SimpleC.SL Require Import Mem SeparationLogic.
From Require Import gmp_goal.
Require Import Logic.LogicGenerator.demo932.Interface.
Local Open Scope Z_scope.
Local Open Scope sets.
Local Open Scope string.
Local Open Scope list.
Import naive_C_Rules.
Local Open Scope sac.
Lemma proof_of_gmp_abs_safety_wit_1 : gmp_abs_safety_wit_1.
Proof. Admitted.
Lemma proof_of_gmp_cmp_safety_wit_1 : gmp_cmp_safety_wit_1.
Proof. Admitted.
Lemma proof_of_gmp_cmp_safety_wit_2 : gmp_cmp_safety_wit_2.
Proof. Admitted.
Lemma proof_of_gmp_cmp_safety_wit_3 : gmp_cmp_safety_wit_3.
Proof. Admitted.
Lemma proof_of_gmp_cmp_safety_wit_4 : gmp_cmp_safety_wit_4.
Proof. Admitted.
Lemma proof_of_gmp_cmp_return_wit_1_1 : gmp_cmp_return_wit_1_1.
Proof. Admitted.
Lemma proof_of_gmp_cmp_return_wit_1_3 : gmp_cmp_return_wit_1_3.
Proof. Admitted.
Lemma proof_of_mpn_copyi_safety_wit_1 : mpn_copyi_safety_wit_1.
Proof. Admitted.
Lemma proof_of_mpn_copyi_safety_wit_2 : mpn_copyi_safety_wit_2.
Proof. Admitted.
Lemma proof_of_mpn_copyi_partial_solve_wit_1 : mpn_copyi_partial_solve_wit_1.
Proof. Admitted.
Lemma proof_of_mpn_copyi_partial_solve_wit_2_pure : mpn_copyi_partial_solve_wit_2_pure.
Proof. Admitted.
Lemma proof_of_mpn_copyi_partial_solve_wit_2 : mpn_copyi_partial_solve_wit_2.
Proof. Admitted.
Lemma proof_of_mpn_copyi_partial_solve_wit_3 : mpn_copyi_partial_solve_wit_3.
Proof. Admitted.
Lemma proof_of_mpn_copyi_partial_solve_wit_4 : mpn_copyi_partial_solve_wit_4.
Proof. Admitted.
Lemma proof_of_mpn_copyi_partial_solve_wit_5 : mpn_copyi_partial_solve_wit_5.
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.
Lemma proof_of_mpn_cmp4_safety_wit_1 : mpn_cmp4_safety_wit_1.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_safety_wit_2 : mpn_cmp4_safety_wit_2.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_safety_wit_3 : mpn_cmp4_safety_wit_3.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_partial_solve_wit_1_pure : mpn_cmp4_partial_solve_wit_1_pure.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_partial_solve_wit_1 : mpn_cmp4_partial_solve_wit_1.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_partial_solve_wit_2_pure : mpn_cmp4_partial_solve_wit_2_pure.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_partial_solve_wit_2 : mpn_cmp4_partial_solve_wit_2.
Proof. Admitted.
Lemma proof_of_mpn_cmp4_which_implies_wit_1 : mpn_cmp4_which_implies_wit_1.
Proof. Admitted.
Lemma proof_of_mpn_normalized_size_safety_wit_1 : mpn_normalized_size_safety_wit_1.
Proof. Admitted.
Lemma proof_of_mpn_normalized_size_safety_wit_2 : mpn_normalized_size_safety_wit_2.
Proof. Admitted.
Lemma proof_of_mpn_normalized_size_safety_wit_3 : mpn_normalized_size_safety_wit_3.
Proof. Admitted.
Lemma proof_of_mpn_normalized_size_safety_wit_4 : mpn_normalized_size_safety_wit_4.
Proof. Admitted.
Lemma proof_of_mpn_normalized_size_safety_wit_5 : mpn_normalized_size_safety_wit_5.
Proof. Admitted.
Lemma proof_of_mpn_normalized_size_entail_wit_1 : mpn_normalized_size_entail_wit_1.
Proof. Admitted.
Lemma proof_of_mpn_normalized_size_partial_solve_wit_1 : mpn_normalized_size_partial_solve_wit_1.
Proof. Admitted.
Lemma proof_of_mpn_normalized_size_partial_solve_wit_2 : mpn_normalized_size_partial_solve_wit_2.
Proof. Admitted.

View File

@ -0,0 +1,413 @@
Require Import Coq.ZArith.ZArith.
Require Import Coq.Bool.Bool.
Require Import Coq.Strings.String.
Require Import Coq.Lists.List.
Require Import Coq.Classes.RelationClasses.
Require Import Coq.Classes.Morphisms.
Require Import Coq.micromega.Psatz.
Require Import Coq.Sorting.Permutation.
From AUXLib Require Import int_auto Axioms Feq Idents List_lemma VMap.
Require Import SetsClass.SetsClass. Import SetsNotation.
From SimpleC.SL Require Import Mem SeparationLogic.
From GmpLib Require Import gmp_goal.
Require Import GmpLib.GmpNumber. Import Internal.
Require Import GmpLib.GmpAux.
Require Import Logic.LogicGenerator.demo932.Interface.
Local Open Scope Z_scope.
Local Open Scope sets.
Local Open Scope string.
Local Open Scope list.
Import naive_C_Rules.
Local Open Scope sac.
Lemma proof_of_gmp_abs_return_wit_1_1 : gmp_abs_return_wit_1_1.
Proof. pre_process. Qed.
Lemma proof_of_gmp_abs_return_wit_1_2 : gmp_abs_return_wit_1_2.
Proof. pre_process. Qed.
Lemma proof_of_gmp_max_return_wit_1_1 : gmp_max_return_wit_1_1.
Proof.
pre_process.
entailer!.
unfold Zmax.
rewrite Z.max_r; lia.
Qed.
Lemma proof_of_gmp_max_return_wit_1_2 : gmp_max_return_wit_1_2.
Proof.
pre_process.
entailer!.
unfold Zmax.
rewrite Z.max_l; lia.
Qed.
Lemma proof_of_gmp_cmp_return_wit_1_2 : gmp_cmp_return_wit_1_2.
Proof.
pre_process.
repeat rewrite <-derivable1_orp_intros1.
entailer!.
Qed.
Lemma proof_of_mpn_copyi_entail_wit_1 : mpn_copyi_entail_wit_1.
Proof.
pre_process.
Exists l2 l_2.
entailer!.
pose proof (Zlength_nonneg l_2).
lia.
Qed.
Lemma proof_of_mpn_copyi_entail_wit_2 : mpn_copyi_entail_wit_2.
Proof.
pre_process.
Exists l2' l_3.
entailer!.
rewrite replace_Znth_app_r.
+ rewrite Zlength_sublist0; [ | lia ].
assert (i - i = 0). { lia. }
rewrite H15; clear H15.
assert (replace_Znth 0 (Znth i l_3 0) (a :: nil) = sublist i (i + 1) l_3). {
unfold replace_Znth, Z.to_nat, replace_nth.
rewrite (sublist_single i l_3 0); [ reflexivity | ].
rewrite <-Zlength_correct; lia.
}
rewrite H15; clear H15.
rewrite replace_Znth_nothing.
- rewrite <-sublist_split; try lia; try reflexivity.
rewrite <-Zlength_correct; lia.
- pose proof (Zlength_sublist0 i l_3 ltac:(lia)).
lia.
+ pose proof (Zlength_sublist0 i l_3); lia.
Qed.
Lemma proof_of_mpn_copyi_which_implies_wit_1 : mpn_copyi_which_implies_wit_1.
Proof.
pre_process.
unfold mpd_store_Z.
Intros l.
Exists l.
unfold mpd_store_list.
entailer!.
subst.
entailer!.
Qed.
Lemma proof_of_mpn_copyi_which_implies_wit_2 : mpn_copyi_which_implies_wit_2.
Proof.
pre_process.
pose proof (store_uint_array_divide d cap2 l2 0).
pose proof (Zlength_nonneg l2).
specialize (H0 ltac:(lia) ltac:(lia)).
destruct H0 as [H0 _].
simpl in H0.
entailer!.
rewrite (sublist_nil l2 0 0) in H0; [ | lia].
sep_apply H0.
entailer!.
unfold store_uint_array, store_uint_array_rec.
unfold store_array.
rewrite (sublist_self l2 cap2); [ | lia ].
assert (d + 0 = d). { lia. }
rewrite H2; clear H2.
assert (cap2 - 0 = cap2). { lia. }
rewrite H2; clear H2.
reflexivity.
Qed.
Lemma proof_of_mpn_copyi_which_implies_wit_3 : mpn_copyi_which_implies_wit_3.
Proof.
pre_process.
destruct l'. {
unfold store_uint_array_rec.
simpl.
entailer!.
}
pose proof (store_uint_array_rec_cons d i cap2 z l' ltac:(lia)).
sep_apply H2.
Exists z l'.
entailer!.
assert (i = 0 \/ i > 0). { lia. }
destruct H3.
+ subst.
unfold store_uint_array, store_array.
simpl.
entailer!.
+ pose proof (Aux.uint_array_rec_to_uint_array d 0 i (sublist 0 i l) ltac:(lia)).
destruct H4 as [_ H4].
assert (d + sizeof(UINT) * 0 = d). { lia. }
rewrite H5 in H4; clear H5.
assert (i - 0 = i). { lia. }
rewrite H5 in H4; clear H5.
sep_apply H4; clear H4.
pose proof (Aux.uint_array_rec_to_uint_array d 0 (i + 1) (sublist 0 i l ++ z :: nil) ltac:(lia)).
destruct H4 as [H4 _].
assert (i + 1 - 0 = i + 1). { lia. }
rewrite H5 in H4; clear H5.
assert (d + sizeof(UINT) * 0 = d). { lia. }
rewrite H5 in H4; clear H5.
rewrite <-H4.
sep_apply store_uint_array_rec_tail_merge; [ reflexivity | lia ].
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!.
Left; Left.
entailer!.
+ unfold mpd_store_Z_compact.
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_compact_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.
Right.
entailer!.
+ unfold mpd_store_Z_compact, mpd_store_list.
Exists l1 l2.
rewrite <-H6, <-H7.
entailer!.
+ pose proof (list_store_Z_compact_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_compact.
Intros l1 l2.
Exists l2 l1.
unfold mpd_store_list.
entailer!.
rewrite <-H0, <-H2.
entailer!.
Qed.
Lemma proof_of_mpn_cmp4_return_wit_1_1 : mpn_cmp4_return_wit_1_1.
Proof.
pre_process.
Right.
unfold mpd_store_Z_compact.
Intros l1 l2.
Exists l1 l2.
entailer!.
pose proof (list_store_Z_cmp_length l2 l1 val2 val1 ltac:(lia) H9 H7).
lia.
Qed.
Lemma proof_of_mpn_cmp4_return_wit_1_2 : mpn_cmp4_return_wit_1_2.
Proof.
pre_process.
Left; Left.
unfold mpd_store_Z_compact.
entailer!.
Intros l1 l2.
Exists l1 l2.
entailer!.
pose proof (list_store_Z_cmp_length l1 l2 val1 val2 ltac:(lia) H7 H9).
lia.
Qed.
Lemma proof_of_mpn_cmp4_return_wit_2_1 : mpn_cmp4_return_wit_2_1.
Proof.
pre_process.
Right.
unfold mpd_store_Z_compact.
Intros l1 l2.
Exists l1 l2.
entailer!.
Qed.
Lemma proof_of_mpn_cmp4_return_wit_2_2 : mpn_cmp4_return_wit_2_2.
Proof.
pre_process.
Left; Right.
unfold mpd_store_Z_compact.
Intros l1 l2.
Exists l1 l2.
entailer!.
Qed.
Lemma proof_of_mpn_cmp4_return_wit_2_3 : mpn_cmp4_return_wit_2_3.
Proof.
pre_process.
Left; Left.
unfold mpd_store_Z_compact.
Intros l1 l2.
Exists l1 l2.
entailer!.
Qed.
Lemma proof_of_mpn_normalized_size_entail_wit_2 : mpn_normalized_size_entail_wit_2.
Proof.
pre_process.
entailer!.
+ pose proof (store_uint_array_divide_rec
xp_pre n (sublist 0 n l) (n - 1) ltac:(lia)).
rewrite (Zlength_sublist0 n l ltac:(lia)) in H12.
specialize (H12 ltac:(lia)).
destruct H12 as [H12 _].
rewrite H12; clear H12.
rewrite (sublist_sublist00 (n - 1) n l ltac:(lia)).
rewrite (sublist_sublist0 n n (n - 1) l ltac:(lia) ltac:(lia)).
pose proof (Aux.uint_array_rec_to_uint_array xp_pre 0 (n - 1) (sublist 0 (n - 1) l) ltac:(lia)).
destruct H12 as [H12 _].
rewrite Z.mul_0_r, Z.add_0_r, Z.sub_0_r in H12.
rewrite H12; clear H12.
entailer!.
assert (n - 1 < Z.of_nat (Datatypes.length l)). {
rewrite <-Zlength_correct.
lia.
}
pose proof (sublist_single (n - 1) l 0 ltac:(lia)).
clear H12.
pose proof (Aux.store_uint_array_single_to_undef xp_pre (n - 1) (Znth (n - 1) l 0)).
assert (n - 1 + 1 = n). { lia. }
rewrite H14 in H12, H13; clear H14.
rewrite H13, H12; clear H13 H12.
pose proof (Aux.store_undef_uint_array_rec_divide xp_pre (n - 1) n cap ltac:(lia) ltac:(lia)).
rewrite <-H12.
entailer!.
+ assert (n <= Z.of_nat (Datatypes.length l)). {
rewrite <-Zlength_correct.
lia.
}
pose proof (sublist_split 0 n (n - 1) l ltac:(lia) ltac:(lia)).
clear H12.
rewrite H13 in H6.
apply (list_store_Z_split) in H6; destruct H6.
assert (Z.of_nat (Datatypes.length l) = Zlength l). {
rewrite (Zlength_correct l); reflexivity.
}
pose proof (sublist_single (n - 1) l 0 ltac:(lia)).
assert (n - 1 + 1 = n). { lia. }
rewrite H16 in H15; clear H16.
rewrite H15 in H12.
unfold list_store_Z in H12; destruct H12.
simpl in H12.
rewrite Znth_sublist0 in H; try lia.
rewrite H in H12.
rewrite (Zlength_sublist0 (n - 1) l) in *; try lia.
pose proof (Z_div_mod_eq_full val (UINT_MOD ^ (n - 1))).
rewrite <-H12, Z.mul_0_r, Z.add_0_l in H17.
rewrite <-H17 in H6.
tauto.
Qed.
Lemma proof_of_mpn_normalized_size_return_wit_1_1 : mpn_normalized_size_return_wit_1_1.
Proof.
pre_process.
assert (n = 0). { lia. }
clear H H0.
rewrite H11 in *.
unfold mpd_store_Z_compact, mpd_store_list.
Exists nil.
entailer!.
+ rewrite Zlength_nil.
lia.
+ unfold list_store_Z_compact.
simpl.
rewrite sublist_nil in H5; try lia.
unfold list_store_Z in H5; simpl in H5.
destruct H5.
lia.
Qed.
Lemma proof_of_mpn_normalized_size_return_wit_1_2 : mpn_normalized_size_return_wit_1_2.
Proof.
pre_process.
unfold mpd_store_Z_compact, mpd_store_list.
Exists (sublist 0 n l).
entailer!.
+ rewrite Zlength_sublist0; try lia.
entailer!.
+ rewrite Zlength_sublist0; lia.
+ rewrite Zlength_sublist0; lia.
+ unfold list_store_Z_compact.
unfold list_store_Z in H6.
destruct H6.
rewrite Aux.list_last_to_Znth.
- rewrite Zlength_sublist0; try lia.
repeat split; try tauto.
pose proof (list_within_bound_Znth (sublist 0 n l) (n - 1)).
rewrite Zlength_sublist0 in H13; try lia.
specialize (H13 ltac:(lia) H12).
lia.
- assert (sublist 0 n l = nil \/ sublist 0 n l <> nil). { tauto. }
destruct H13.
* pose proof (Zlength_sublist0 n l ltac:(lia)).
rewrite H13 in H14.
rewrite Zlength_nil in H14.
lia.
* tauto.
Qed.
Lemma proof_of_mpn_normalized_size_which_implies_wit_1 : mpn_normalized_size_which_implies_wit_1.
Proof.
pre_process.
unfold mpd_store_Z.
Intros l.
Exists l.
unfold mpd_store_list.
entailer!.
+ rewrite H0.
rewrite sublist_self; try lia.
entailer!.
+ rewrite sublist_self; try lia.
tauto.
Qed.

View File

@ -1,14 +1,36 @@
#include "mini-gmp.h" #include "mini-gmp.h"
#include "verification_stdlib.h"
#include "verification_list.h"
#include "int_array_def.h"
mp_size_t gmp_abs(mp_size_t x) { int gmp_abs(int x)
/*@
Require INT_MIN < x && x <= INT_MAX
Ensure __return == Zabs(x)
*/
{
return x >= 0 ? x : -x; return x >= 0 ? x : -x;
} }
mp_size_t gmp_max(mp_size_t a, mp_size_t b) { int gmp_max(int a, int b)
/*@
Require emp
Ensure __return == Zmax(a, b)
*/
{
return a > b ? a : b; return a > b ? a : b;
} }
int gmp_cmp(mp_size_t a, mp_size_t b) { int gmp_cmp(int a, int b)
/*@
Require emp
Ensure
emp *
(a > b && __return == 1 ||
a == b && __return == 0 ||
a < b && __return == -1)
*/
{
return (a > b) - (a < b); return (a > b) - (a < b);
} }
@ -16,73 +38,219 @@ int gmp_cmp(mp_size_t a, mp_size_t b) {
/* 从 低地址向高地址 顺序复制 */ /* 从 低地址向高地址 顺序复制 */
void void
mpn_copyi (mp_ptr d, mp_srcptr s, mp_size_t n) mpn_copyi (unsigned int *d, unsigned int *s, int n)
/*@
With val l2 cap1 cap2
Require
mpd_store_Z(s, val, n, cap1) *
store_uint_array(d, cap2, l2) &&
Zlength(l2) == cap2 &&
cap2 >= n
Ensure
mpd_store_Z(s, val, n, cap1) *
mpd_store_Z(d, val, n, cap2)
*/
{ {
mp_size_t i; /*@
for (i = 0; i < n; i++) mpd_store_Z(s, val, n, cap1)
d[i] = s[i]; which implies
exists l,
n <= cap1 &&
Zlength(l) == n &&
cap1 <= 100000000 &&
store_uint_array(s, n, l) *
store_undef_uint_array_rec(s, n, cap1) &&
list_store_Z(l, val)
*/
/*@
store_uint_array(d, cap2, l2) && Zlength(l2) == cap2
which implies
store_uint_array_rec(d, 0, cap2, l2) * store_uint_array(d, 0, nil) &&
Zlength(l2) == cap2
*/
int i;
/*@ Inv
exists l l',
0 <= i && i <= n && Zlength(l) == n &&
list_store_Z(l, val) && n <= cap1 &&
store_uint_array(s, n, l) *
store_undef_uint_array_rec(s, n, cap1) *
store_uint_array(d, i, sublist(0, i, l)) *
store_uint_array_rec(d, i, cap2, l')
*/
for (i = 0; i < n; i++) {
/*@
Given l l'
*/
/*@
0 <= i && i < n && n <= cap2 &&
store_uint_array_rec(d, i, cap2, l') *
store_uint_array(d, i, sublist(0, i, l))
which implies
exists a l2',
l' == cons(a, l2') && i < n && n <= cap2 &&
store_uint_array_rec(d, i + 1, cap2, l2') *
store_uint_array(d, i + 1, app(sublist(0, i, l), cons(a, nil)))
*/
d[i] = s[i];
}
} }
/* 大于返回1小于返回-1等于返回0 */ /* 大于返回1小于返回-1等于返回0 */
int int
mpn_cmp (mp_srcptr ap, mp_srcptr bp, mp_size_t n) mpn_cmp (unsigned int *ap, unsigned int *bp, int n)
/*@
With cap1 cap2 val1 val2
Require
mpd_store_Z_compact(ap, val1, n, cap1) *
mpd_store_Z_compact(bp, val2, n, cap2) &&
0 <= n && n <= cap1 && n <= cap2 &&
cap1 <= 100000000 && cap2 <= 100000000
Ensure
(val1 > val2 && __return == 1 ||
val1 == val2 && __return == 0 ||
val1 < val2 && __return == -1) &&
mpd_store_Z_compact(ap@pre, val1, n@pre, cap1) *
mpd_store_Z_compact(bp@pre, val2, n@pre, cap2)
*/
{ {
/*@
mpd_store_Z_compact(ap@pre, val1, n@pre, cap1) * mpd_store_Z_compact(bp@pre, val2, n@pre, cap2)
which implies
exists l1 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, cap1) *
store_undef_uint_array_rec(bp@pre, n@pre, cap2) &&
list_store_Z_compact(l1, val1) && list_store_Z_compact(l2, val2) &&
n@pre == Zlength(l1) && n@pre == Zlength(l2)
*/
/*@
Given l1 l2
*/
--n;
/*@Inv
-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, cap1) *
store_undef_uint_array_rec(bp@pre, n@pre, cap2) &&
list_store_Z_compact(l1, val1) && list_store_Z_compact(l2, val2) &&
n@pre == Zlength(l1) && n@pre == Zlength(l2) &&
sublist(n + 1, n@pre, l1) == sublist(n + 1, n@pre, l2)
*/
while (n >= 0)
{
if (ap[n] != bp[n])
return ap[n] > bp[n] ? 1 : -1;
--n;
}
// Note: The parser cannot parse "--n" in loop condition so we paraphrased it.
/*
while (--n >= 0) while (--n >= 0)
{ {
if (ap[n] != bp[n]) if (ap[n] != bp[n])
return ap[n] > bp[n] ? 1 : -1; return ap[n] > bp[n] ? 1 : -1;
} }
*/
return 0; return 0;
} }
/*处理位数不同的情况*/ /*处理位数不同的情况*/
static int static int
mpn_cmp4 (mp_srcptr ap, mp_size_t an, mp_srcptr bp, mp_size_t bn) mpn_cmp4 (unsigned int *ap, int an, unsigned int *bp, int bn)
/*@
With cap1 cap2 val1 val2
Require
mpd_store_Z_compact(ap, val1, an, cap1) *
mpd_store_Z_compact(bp, val2, bn, cap2) &&
an >= 0 && bn >= 0 && an <= cap1 && bn <= cap2 &&
cap1 <= 100000000 && cap2 <= 100000000
Ensure
(val1 > val2 && __return == 1 ||
val1 == val2 && __return == 0 ||
val1 < val2 && __return == -1) &&
mpd_store_Z_compact(ap@pre, val1, an@pre, cap1) *
mpd_store_Z_compact(bp@pre, val2, bn@pre, cap2)
*/
{ {
if (an != bn) if (an != bn)
return an < bn ? -1 : 1; return an < bn ? -1 : 1;
else else {
/*@
an@pre == bn@pre && bn@pre <= cap2
which implies
an@pre <= cap2
*/
return mpn_cmp (ap, bp, an); return mpn_cmp (ap, bp, an);
}
} }
/*返回非0的位数*/ /*返回非0的位数*/
static mp_size_t static int
mpn_normalized_size (mp_srcptr xp, mp_size_t n) mpn_normalized_size (unsigned int *xp, int n)
/*@
With cap val
Require
mpd_store_Z(xp, val, n, cap) &&
0 <= n && n <= cap && cap <= 100000000
Ensure
0 <= __return && __return <= cap &&
mpd_store_Z_compact(xp@pre, val, __return, cap)
*/
{ {
/*@
mpd_store_Z(xp@pre, val, n, cap)
which implies
exists l,
list_store_Z(sublist(0, n, l), val) &&
Zlength(l) == n &&
store_uint_array(xp@pre, n, sublist(0, n, l)) *
store_undef_uint_array_rec(xp@pre, n, cap)
*/
/*@
Given l
*/
/*@Inv
n >= 0 && n <= n@pre &&
n@pre >= 0 && n@pre <= cap && cap <= 100000000 &&
list_store_Z(sublist(0, n, l), val) &&
store_uint_array(xp@pre, n, sublist(0, n, l)) *
store_undef_uint_array_rec(xp@pre, n, cap)
*/
while (n > 0 && xp[n-1] == 0) while (n > 0 && xp[n-1] == 0)
--n; --n;
return n; return n;
} }
/* 多精度数ap 加上单精度数b返回最后产生的进位 */ /* 多精度数ap 加上单精度数b返回最后产生的进位 */
mp_limb_t /*unsigned int
mpn_add_1 (mp_ptr rp, mp_srcptr ap, mp_size_t n, mp_limb_t b) mpn_add_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
{ {
mp_size_t i; int i;
//assert (n > 0); //assert (n > 0);
i = 0; i = 0;
do do
{ {
mp_limb_t r = ap[i] + b; unsigned int r = ap[i] + b;
/* Carry out */ // Carry out
b = (r < b); b = (r < b);
rp[i] = r; rp[i] = r;
} }
while (++i < n); while (++i < n);
return b; return b;
} }*/
/* 位数相同的多精度数ap 加上多精度数bp返回最后产生的进位 */ /* 位数相同的多精度数ap 加上多精度数bp返回最后产生的进位 */
mp_limb_t /*unsigned int
mpn_add_n (mp_ptr rp, mp_srcptr ap, mp_srcptr bp, mp_size_t n) mpn_add_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n)
{ {
mp_size_t i; int i;
mp_limb_t cy; unsigned int cy;
for (i = 0, cy = 0; i < n; i++) for (i = 0, cy = 0; i < n; i++)
{ {
mp_limb_t a, b, r; unsigned int a, b, r;
a = ap[i]; b = bp[i]; a = ap[i]; b = bp[i];
r = a + cy; r = a + cy;
cy = (r < cy); cy = (r < cy);
@ -91,48 +259,48 @@ mpn_add_n (mp_ptr rp, mp_srcptr ap, mp_srcptr bp, mp_size_t n)
rp[i] = r; rp[i] = r;
} }
return cy; return cy;
} }*/
/*不同位数的多精度数相加,返回最后的进位*/ /*不同位数的多精度数相加,返回最后的进位*/
mp_limb_t /*unsigned int
mpn_add (mp_ptr rp, mp_srcptr ap, mp_size_t an, mp_srcptr bp, mp_size_t bn) mpn_add (unsigned int *rp, unsigned int *ap, int an, unsigned int *bp, int bn)
{ {
mp_limb_t cy; unsigned int cy;
//assert (an >= bn); //assert (an >= bn);
cy = mpn_add_n (rp, ap, bp, bn); cy = mpn_add_n (rp, ap, bp, bn);
if (an > bn) if (an > bn)
cy = mpn_add_1 (rp + bn, ap + bn, an - bn, cy); cy = mpn_add_1 (rp + bn, ap + bn, an - bn, cy);
return cy; return cy;
} }*/
mp_limb_t /*unsigned int
mpn_sub_1 (mp_ptr rp, mp_srcptr ap, mp_size_t n, mp_limb_t b) mpn_sub_1 (unsigned int *rp, unsigned int *ap, int n, unsigned int b)
{ {
mp_size_t i; int i;
//assert (n > 0); //assert (n > 0);
i = 0; i = 0;
do do
{ {
mp_limb_t a = ap[i]; unsigned int a = ap[i];
/* Carry out */ // Carry out
mp_limb_t cy = a < b; unsigned int cy = a < b;
rp[i] = a - b; rp[i] = a - b;
b = cy; b = cy;
} }
while (++i < n); while (++i < n);
return b; return b;
} }*/
mp_limb_t /*unsigned int
mpn_sub_n (mp_ptr rp, mp_srcptr ap, mp_srcptr bp, mp_size_t n) mpn_sub_n (unsigned int *rp, unsigned int *ap, unsigned int *bp, int n)
{ {
mp_size_t i; int i;
mp_limb_t cy; unsigned int cy;
for (i = 0, cy = 0; i < n; i++) for (i = 0, cy = 0; i < n; i++)
{ {
mp_limb_t a, b; unsigned int a, b;
a = ap[i]; b = bp[i]; a = ap[i]; b = bp[i];
b += cy; b += cy;
cy = (b < cy); cy = (b < cy);
@ -140,30 +308,30 @@ mpn_sub_n (mp_ptr rp, mp_srcptr ap, mp_srcptr bp, mp_size_t n)
rp[i] = a - b; rp[i] = a - b;
} }
return cy; return cy;
} }*/
mp_limb_t /*unsigned int
mpn_sub (mp_ptr rp, mp_srcptr ap, mp_size_t an, mp_srcptr bp, mp_size_t bn) mpn_sub (unsigned int *rp, unsigned int *ap, int an, unsigned int *bp, int bn)
{ {
mp_limb_t cy; unsigned int cy;
//assert (an >= bn); //assert (an >= bn);
cy = mpn_sub_n (rp, ap, bp, bn); cy = mpn_sub_n (rp, ap, bp, bn);
if (an > bn) if (an > bn)
cy = mpn_sub_1 (rp + bn, ap + bn, an - bn, cy); cy = mpn_sub_1 (rp + bn, ap + bn, an - bn, cy);
return cy; return cy;
} }*/
/* MPZ interface */ /* MPZ interface */
void /*void
mpz_clear (mpz_t r) mpz_clear (mpz_t r)
{ {
if (r->_mp_alloc) if (r->_mp_alloc)
gmp_free_limbs (r->_mp_d, r->_mp_alloc); gmp_free_limbs (r->_mp_d, r->_mp_alloc);
} }*/
static mp_ptr /*static unsigned int *
mpz_realloc (mpz_t r, mp_size_t size) mpz_realloc (mpz_t r, int size)
{ {
size = gmp_max (size, 1); size = gmp_max (size, 1);
@ -177,42 +345,42 @@ mpz_realloc (mpz_t r, mp_size_t size)
r->_mp_size = 0; r->_mp_size = 0;
return r->_mp_d; return r->_mp_d;
} }*/
/* Realloc for an mpz_t WHAT if it has less than NEEDED limbs. */ /* Realloc for an mpz_t WHAT if it has less than NEEDED limbs. */
mp_ptr mrz_realloc_if(mpz_t z,mp_size_t n) { /*unsigned int *mrz_realloc_if(mpz_t z,int n) {
return n > z->_mp_alloc ? mpz_realloc(z, n) : z->_mp_d; return n > z->_mp_alloc ? mpz_realloc(z, n) : z->_mp_d;
} }*/
/* MPZ comparisons and the like. */ /* MPZ comparisons and the like. */
int /*int
mpz_sgn (const mpz_t u) mpz_sgn (const mpz_t u)
{ {
return gmp_cmp (u->_mp_size, 0); return gmp_cmp (u->_mp_size, 0);
} }*/
void /*void
mpz_swap (mpz_t u, mpz_t v) mpz_swap (mpz_t u, mpz_t v)
{ {
mp_size_t_swap (u->_mp_alloc, v->_mp_alloc); int_swap (u->_mp_alloc, v->_mp_alloc);
mp_ptr_swap(u->_mp_d, v->_mp_d); unsigned int *_swap(u->_mp_d, v->_mp_d);
mp_size_t_swap (u->_mp_size, v->_mp_size); int_swap (u->_mp_size, v->_mp_size);
} }*/
/* MPZ addition and subtraction */ /* MPZ addition and subtraction */
static mp_size_t /*static int
mpz_abs_add (mpz_t r, const mpz_t a, const mpz_t b) mpz_abs_add (mpz_t r, const mpz_t a, const mpz_t b)
{ {
mp_size_t an = gmp_abs (a->_mp_size); int an = gmp_abs (a->_mp_size);
mp_size_t bn = gmp_abs (b->_mp_size); int bn = gmp_abs (b->_mp_size);
mp_ptr rp; unsigned int *rp;
mp_limb_t cy; unsigned int cy;
if (an < bn) if (an < bn)
{ {
mpz_srcptr_swap (a, b); mpz_srcptr_swap (a, b);
mp_size_t_swap (an, bn); int_swap (an, bn);
} }
rp = mrz_realloc_if (r, an + 1); rp = mrz_realloc_if (r, an + 1);
@ -221,15 +389,15 @@ mpz_abs_add (mpz_t r, const mpz_t a, const mpz_t b)
rp[an] = cy; rp[an] = cy;
return an + cy; return an + cy;
} }*/
static mp_size_t /*static int
mpz_abs_sub (mpz_t r, const mpz_t a, const mpz_t b) mpz_abs_sub (mpz_t r, const mpz_t a, const mpz_t b)
{ {
mp_size_t an = gmp_abs (a->_mp_size); int an = gmp_abs (a->_mp_size);
mp_size_t bn = gmp_abs (b->_mp_size); int bn = gmp_abs (b->_mp_size);
int cmp; int cmp;
mp_ptr rp; unsigned int *rp;
cmp = mpn_cmp4 (a->_mp_d, an, b->_mp_d, bn); cmp = mpn_cmp4 (a->_mp_d, an, b->_mp_d, bn);
if (cmp > 0) if (cmp > 0)
@ -246,12 +414,12 @@ mpz_abs_sub (mpz_t r, const mpz_t a, const mpz_t b)
} }
else else
return 0; return 0;
} }*/
void /*void
mpz_add (mpz_t r, const mpz_t a, const mpz_t b) mpz_add (mpz_t r, const mpz_t a, const mpz_t b)
{ {
mp_size_t rn; int rn;
if ( (a->_mp_size ^ b->_mp_size) >= 0) if ( (a->_mp_size ^ b->_mp_size) >= 0)
rn = mpz_abs_add (r, a, b); rn = mpz_abs_add (r, a, b);
@ -259,12 +427,12 @@ mpz_add (mpz_t r, const mpz_t a, const mpz_t b)
rn = mpz_abs_sub (r, a, b); rn = mpz_abs_sub (r, a, b);
r->_mp_size = a->_mp_size >= 0 ? rn : - rn; r->_mp_size = a->_mp_size >= 0 ? rn : - rn;
} }*/
void /*void
mpz_sub (mpz_t r, const mpz_t a, const mpz_t b) mpz_sub (mpz_t r, const mpz_t a, const mpz_t b)
{ {
mp_size_t rn; int rn;
if ( (a->_mp_size ^ b->_mp_size) >= 0) if ( (a->_mp_size ^ b->_mp_size) >= 0)
rn = mpz_abs_sub (r, a, b); rn = mpz_abs_sub (r, a, b);
@ -272,4 +440,4 @@ mpz_sub (mpz_t r, const mpz_t a, const mpz_t b)
rn = mpz_abs_add (r, a, b); rn = mpz_abs_add (r, a, b);
r->_mp_size = a->_mp_size >= 0 ? rn : - rn; r->_mp_size = a->_mp_size >= 0 ? rn : - rn;
} }*/

View File

@ -1,10 +1,3 @@
typedef unsigned long mp_limb_t;
typedef long mp_size_t;
typedef unsigned long mp_bitcnt_t;
typedef mp_limb_t *mp_ptr;
typedef const mp_limb_t *mp_srcptr;
typedef struct typedef struct
{ {
int _mp_alloc; /* Number of *limbs* allocated and pointed int _mp_alloc; /* Number of *limbs* allocated and pointed
@ -12,7 +5,7 @@ typedef struct
int _mp_size; /* abs(_mp_size) is the number of limbs the int _mp_size; /* abs(_mp_size) is the number of limbs the
last field points to. If _mp_size is last field points to. If _mp_size is
negative this is a negative number. */ negative this is a negative number. */
mp_limb_t *_mp_d; /* Pointer to the limbs. */ unsigned int *_mp_d; /* Pointer to the limbs. */
} __mpz_struct; } __mpz_struct;
/* mpz_t is an array type that contains a single element of __mpz_struct, acting as a reference. */ /* mpz_t is an array type that contains a single element of __mpz_struct, acting as a reference. */
@ -23,35 +16,35 @@ typedef const __mpz_struct *mpz_srcptr;
/* BEGIN Given Functions */ /* BEGIN Given Functions */
/* Swap functions. */ /* Swap functions. */
void mp_size_t_swap(mp_size_t x, mp_size_t y); void int_swap(int x, int y);
void mp_ptr_swap(mp_ptr x, mp_ptr y); void mp_ptr_swap(unsigned int *x, unsigned int *y);
void mpz_srcptr_swap(mpz_srcptr x, mpz_srcptr y); void mpz_srcptr_swap(mpz_srcptr x, mpz_srcptr y);
/* Memory allocation functions. */ /* Memory allocation functions. */
static mp_ptr static unsigned int *
gmp_alloc_limbs (mp_size_t size); gmp_alloc_limbs (int size);
static mp_ptr static unsigned int *
gmp_realloc_limbs (mp_ptr old, mp_size_t old_size, mp_size_t size); gmp_realloc_limbs (unsigned int *old, int old_size, int size);
static void static void
gmp_free_limbs (mp_ptr old, mp_size_t size); gmp_free_limbs (unsigned int *old, int size);
/* END Given Functions */ /* END Given Functions */
void mpn_copyi (mp_ptr, mp_srcptr, mp_size_t); void mpn_copyi (unsigned int *d, unsigned int *s, int n);
int mpn_cmp (mp_srcptr, mp_srcptr, mp_size_t); int mpn_cmp (unsigned int *ap, unsigned int *bp, int n);
mp_limb_t mpn_add_1 (mp_ptr, mp_srcptr, mp_size_t, mp_limb_t); unsigned int mpn_add_1 (unsigned int *, unsigned int *, int, unsigned int);
mp_limb_t mpn_add_n (mp_ptr, mp_srcptr, mp_srcptr, mp_size_t); unsigned int mpn_add_n (unsigned int *, unsigned int *, unsigned int *, int);
mp_limb_t mpn_add (mp_ptr, mp_srcptr, mp_size_t, mp_srcptr, mp_size_t); unsigned int mpn_add (unsigned int *, unsigned int *, int, unsigned int *, int);
mp_limb_t mpn_sub_1 (mp_ptr, mp_srcptr, mp_size_t, mp_limb_t); unsigned int mpn_sub_1 (unsigned int *, unsigned int *, int, unsigned int);
mp_limb_t mpn_sub_n (mp_ptr, mp_srcptr, mp_srcptr, mp_size_t); unsigned int mpn_sub_n (unsigned int *, unsigned int *, unsigned int *, int);
mp_limb_t mpn_sub (mp_ptr, mp_srcptr, mp_size_t, mp_srcptr, mp_size_t); unsigned int mpn_sub (unsigned int *, unsigned int *, int, unsigned int *, int);
void mpz_clear (mpz_t); void mpz_clear (mpz_t);
@ -64,3 +57,14 @@ void mpz_add (mpz_t, const mpz_t, const mpz_t);
void mpz_sub (mpz_t, const mpz_t, const mpz_t); void mpz_sub (mpz_t, const mpz_t, const mpz_t);
void mpz_set (mpz_t, const mpz_t); void mpz_set (mpz_t, const mpz_t);
/*@
Extern Coq (Zabs : Z -> Z)
(Zmax : Z -> Z -> Z)
(mpd_store_Z : Z -> Z -> Z -> Z -> Assertion)
(mpd_store_Z_compact: Z -> Z -> Z -> Z -> Assertion)
(mpd_store_list : Z -> list Z -> Z -> Assertion)
(list_store_Z : list Z -> Z -> Prop)
(list_store_Z_compact: list Z -> Z -> Prop)
(last: list Z -> Z -> Z)
*/

View File

@ -0,0 +1,126 @@
//test curent rules
//max rule id:30
#include "int_array_def.h"
#include "verification_list.h"
#include "verification_stdlib.h"
id : 1
priority : core(1)
left : store_uint_array(?p, ?n, ?l) at 0
right : data_at(?p + (?i * sizeof(U32)), U32, ?v) at 1
check : infer(0 <= i);
infer(i < n);
action : right_erase(1);
left_erase(0);
left_add(store_uint_array_missing_i_rec(p, i, 0, n, l));
right_add(v == l[i]);
id : 2
priority : post(1)
left : store_uint_array_missing_i_rec(?p, ?i, 0, ?n, ?l) at 0
data_at(p + i * sizeof(U32), U32, l[i]) at 1
check : infer(0 <= i);
infer(i < n);
action : left_erase(1);
left_erase(0);
left_add(store_uint_array(p, n, l));
id : 3
priority : post(3)
left : store_uint_array_missing_i_rec(?p, ?i, 0, ?n, ?l) at 0
data_at(p + i * sizeof(U32), U32, ?v) at 1
check : infer(0 <= i);
infer(i < n);
action : left_erase(1);
left_erase(0);
left_add(store_uint_array(p, n, replace_Znth{Z}(i, v, l)));
id : 4
priority : core(1)
left : store_uint_array(?p, ?n, ?l1) at 0
right : store_uint_array(p, n, ?l2) at 1
action : right_erase(1);
left_erase(0);
right_add(l1 == l2);
id : 5
priority : core(1)
left : store_uint_array_missing_i_rec(?p, ?i, ?v, ?n, ?l) at 0
right : store_uint_array_missing_i_rec(p, i, v, n, l) at 1
action : left_erase(0);
right_erase(1);
id : 6
priority : core(1)
left : store_uint_array_rec(?p, ?x, ?y, ?l1) at 0
right : store_uint_array_rec(?p, ?x, ?y, ?l2) at 1
action : right_erase(1);
left_erase(0);
right_add(l1 == l2);
id : 7
priority : core(1)
left : store_uint_array_rec(?p, ?x, ?y, ?l) at 0
right : data_at(?p + (?i * sizeof(U32)), U32, ?v) at 1
check : infer(x <= i);
infer(i < y);
action : right_erase(1);
left_erase(0);
left_add(store_uint_array_missing_i_rec(p, i, x, y, l));
right_add(v == l[i - x]);
id : 8
priority : core(1)
left : store_uint_array_rec(?p, ?x, ?y, ?l1) at 0
store_uint_array_rec(?p, ?y, ?z, ?l2) at 1
right : store_uint_array_rec(?p, ?x, ?z, ?l3) at 2
check : infer(y <= z);
infer(x <= y);
action : right_erase(2);
left_erase(0);
left_erase(1);
right_add(l3 == app{Z}(l1, l2));
id : 9
priority : core(1)
left : store_uint_array_rec(?p, ?x, ?z, ?l3) at 2
right : store_uint_array_rec(?p, ?x, ?y, ?l1) at 0
store_uint_array_rec(?p, ?y, ?z, ?l2) at 1
check : infer(y <= z);
infer(x <= y);
action : left_erase(2);
right_erase(0);
right_erase(1);
right_add(l3 == app{Z}(l1, l2));
right_add(Zlength{Z}(l1) == y - x);
id : 10
priority : core(1)
right : store_uint_array_rec(?p, ?x, ?x, ?l) at 0
action : right_erase(0);
right_add(l == nil{Z});
id : 11
priority : post(1)
left : store_uint_array_missing_i_rec(?p, ?i, ?x, ?y, ?l) at 0
data_at(p + ?i * sizeof(U32), U32, l[?i - ?x]) at 1
check : infer(x <= i);
infer(i < y);
action : left_erase(1);
left_erase(0);
left_add(store_uint_array_rec(p, x, y, l));
id : 12
priority : post(4)
left : store_uint_array_missing_i_rec(?p, ?i, ?x, ?y, ?l) at 0
data_at(p + ?i * sizeof(U32), U32, ?v) at 1
check : infer(x <= i);
infer(i < y);
action : left_erase(1);
left_erase(0);
left_add(store_uint_array_rec(p, x, y, replace_Znth{Z}(i - x, v, l)));