feat(mpn_copyi): Proved correctness of mpn_copyi and other simple util functions.
This commit is contained in:
4
Makefile
4
Makefile
@ -37,7 +37,9 @@ 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
|
GMP_FILES = \
|
||||||
|
./projects/lib/GmpAux.v ./projects/lib/GmpNumber.v \
|
||||||
|
./projects/lib/gmp_goal.v
|
||||||
|
|
||||||
FILES = $(PV_FILES) \
|
FILES = $(PV_FILES) \
|
||||||
$(GMP_FILES)
|
$(GMP_FILES)
|
||||||
|
@ -1,7 +1,8 @@
|
|||||||
/*@ Extern Coq (sum : list Z -> Z)
|
/*@ Extern Coq (sum : list Z -> Z)
|
||||||
(sublist : {A} -> Z -> Z -> list A -> list A)
|
(sublist : {A} -> Z -> Z -> list A -> list A)
|
||||||
(store_int_array : Z -> Z -> list Z -> Assertion)
|
(store_int_array : Z -> Z -> list Z -> Assertion)
|
||||||
(store_uint_array : Z -> Z -> list Z -> Assertion)(store_uint_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_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_uint_array_missing_i_rec: Z -> Z -> Z -> Z -> list Z -> Assertion)
|
||||||
(store_array_box: Z -> Z -> list Z -> Assertion)
|
(store_array_box: Z -> Z -> list Z -> Assertion)
|
||||||
|
@ -100,5 +100,91 @@ Proof.
|
|||||||
rewrite IHl1.
|
rewrite IHl1.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
Qed.
|
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 store_uarray_rec_equals_store_uarray: 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.
|
||||||
|
|
||||||
End Aux.
|
End Aux.
|
@ -21,11 +21,12 @@ 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 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) + 1) cap.
|
||||||
|
|
||||||
@ -46,7 +47,7 @@ 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 |].
|
mpd_store_list ptr data cap && [| list_store_Z data n |] && [| size = Zlength data |].
|
||||||
|
|
||||||
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 ->
|
||||||
|
624
projects/lib/gmp_goal.v
Normal file
624
projects/lib/gmp_goal.v
Normal file
@ -0,0 +1,624 @@
|
|||||||
|
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.
|
||||||
|
Require Import GmpLib.GmpNumber. Import Internal.
|
||||||
|
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.
|
||||||
|
|
||||||
|
Definition Zmax := Z.max.
|
||||||
|
|
||||||
|
(*----- Function gmp_abs -----*)
|
||||||
|
|
||||||
|
Definition gmp_abs_safety_wit_1 :=
|
||||||
|
forall (x_pre: Z) ,
|
||||||
|
[| (x_pre < 0) |]
|
||||||
|
&& [| (INT_MIN < x_pre) |]
|
||||||
|
&& [| (x_pre <= INT_MAX) |]
|
||||||
|
&& ((( &( "x" ) )) # Int |-> x_pre)
|
||||||
|
|--
|
||||||
|
[| (x_pre <> (INT_MIN)) |]
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition gmp_abs_return_wit_1_1 :=
|
||||||
|
forall (x_pre: Z) ,
|
||||||
|
[| (x_pre < 0) |]
|
||||||
|
&& [| (INT_MIN < x_pre) |]
|
||||||
|
&& [| (x_pre <= INT_MAX) |]
|
||||||
|
&& emp
|
||||||
|
|--
|
||||||
|
[| ((-x_pre) = (Zabs (x_pre))) |]
|
||||||
|
&& emp
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition gmp_abs_return_wit_1_2 :=
|
||||||
|
forall (x_pre: Z) ,
|
||||||
|
[| (x_pre >= 0) |]
|
||||||
|
&& [| (INT_MIN < x_pre) |]
|
||||||
|
&& [| (x_pre <= INT_MAX) |]
|
||||||
|
&& emp
|
||||||
|
|--
|
||||||
|
[| (x_pre = (Zabs (x_pre))) |]
|
||||||
|
&& emp
|
||||||
|
.
|
||||||
|
|
||||||
|
(*----- Function gmp_max -----*)
|
||||||
|
|
||||||
|
Definition gmp_max_return_wit_1_1 :=
|
||||||
|
forall (b_pre: Z) (a_pre: Z) ,
|
||||||
|
[| (a_pre <= b_pre) |]
|
||||||
|
&& emp
|
||||||
|
|--
|
||||||
|
[| (b_pre = (Zmax (a_pre) (b_pre))) |]
|
||||||
|
&& emp
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition gmp_max_return_wit_1_2 :=
|
||||||
|
forall (b_pre: Z) (a_pre: Z) ,
|
||||||
|
[| (a_pre > b_pre) |]
|
||||||
|
&& emp
|
||||||
|
|--
|
||||||
|
[| (a_pre = (Zmax (a_pre) (b_pre))) |]
|
||||||
|
&& emp
|
||||||
|
.
|
||||||
|
|
||||||
|
(*----- Function gmp_cmp -----*)
|
||||||
|
|
||||||
|
Definition gmp_cmp_safety_wit_1 :=
|
||||||
|
forall (b_pre: Z) (a_pre: Z) ,
|
||||||
|
[| (a_pre >= b_pre) |]
|
||||||
|
&& [| (a_pre <= b_pre) |]
|
||||||
|
&& ((( &( "b" ) )) # Int |-> b_pre)
|
||||||
|
** ((( &( "a" ) )) # Int |-> a_pre)
|
||||||
|
|--
|
||||||
|
[| ((0 - 0 ) <= INT_MAX) |]
|
||||||
|
&& [| ((INT_MIN) <= (0 - 0 )) |]
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition gmp_cmp_safety_wit_2 :=
|
||||||
|
forall (b_pre: Z) (a_pre: Z) ,
|
||||||
|
[| (a_pre < b_pre) |]
|
||||||
|
&& [| (a_pre <= b_pre) |]
|
||||||
|
&& ((( &( "b" ) )) # Int |-> b_pre)
|
||||||
|
** ((( &( "a" ) )) # Int |-> a_pre)
|
||||||
|
|--
|
||||||
|
[| ((0 - 1 ) <= INT_MAX) |]
|
||||||
|
&& [| ((INT_MIN) <= (0 - 1 )) |]
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition gmp_cmp_safety_wit_3 :=
|
||||||
|
forall (b_pre: Z) (a_pre: Z) ,
|
||||||
|
[| (a_pre >= b_pre) |]
|
||||||
|
&& [| (a_pre > b_pre) |]
|
||||||
|
&& ((( &( "b" ) )) # Int |-> b_pre)
|
||||||
|
** ((( &( "a" ) )) # Int |-> a_pre)
|
||||||
|
|--
|
||||||
|
[| ((1 - 0 ) <= INT_MAX) |]
|
||||||
|
&& [| ((INT_MIN) <= (1 - 0 )) |]
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition gmp_cmp_safety_wit_4 :=
|
||||||
|
forall (b_pre: Z) (a_pre: Z) ,
|
||||||
|
[| (a_pre < b_pre) |]
|
||||||
|
&& [| (a_pre > b_pre) |]
|
||||||
|
&& ((( &( "b" ) )) # Int |-> b_pre)
|
||||||
|
** ((( &( "a" ) )) # Int |-> a_pre)
|
||||||
|
|--
|
||||||
|
[| False |]
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition gmp_cmp_return_wit_1_1 :=
|
||||||
|
forall (b_pre: Z) (a_pre: Z) ,
|
||||||
|
[| (a_pre >= b_pre) |]
|
||||||
|
&& [| (a_pre > b_pre) |]
|
||||||
|
&& emp
|
||||||
|
|--
|
||||||
|
([| (a_pre < b_pre) |]
|
||||||
|
&& [| ((1 - 0 ) = (-1)) |]
|
||||||
|
&& emp)
|
||||||
|
||
|
||||||
|
([| (a_pre = b_pre) |]
|
||||||
|
&& [| ((1 - 0 ) = 0) |]
|
||||||
|
&& emp)
|
||||||
|
||
|
||||||
|
([| (a_pre > b_pre) |]
|
||||||
|
&& [| ((1 - 0 ) = 1) |]
|
||||||
|
&& emp)
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition gmp_cmp_return_wit_1_2 :=
|
||||||
|
forall (b_pre: Z) (a_pre: Z) ,
|
||||||
|
[| (a_pre < b_pre) |]
|
||||||
|
&& [| (a_pre <= b_pre) |]
|
||||||
|
&& emp
|
||||||
|
|--
|
||||||
|
([| (a_pre < b_pre) |]
|
||||||
|
&& [| ((0 - 1 ) = (-1)) |]
|
||||||
|
&& emp)
|
||||||
|
||
|
||||||
|
([| (a_pre = b_pre) |]
|
||||||
|
&& [| ((0 - 1 ) = 0) |]
|
||||||
|
&& emp)
|
||||||
|
||
|
||||||
|
([| (a_pre > b_pre) |]
|
||||||
|
&& [| ((0 - 1 ) = 1) |]
|
||||||
|
&& emp)
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition gmp_cmp_return_wit_1_3 :=
|
||||||
|
forall (b_pre: Z) (a_pre: Z) ,
|
||||||
|
[| (a_pre >= b_pre) |]
|
||||||
|
&& [| (a_pre <= b_pre) |]
|
||||||
|
&& emp
|
||||||
|
|--
|
||||||
|
([| (a_pre < b_pre) |]
|
||||||
|
&& [| ((0 - 0 ) = (-1)) |]
|
||||||
|
&& emp)
|
||||||
|
||
|
||||||
|
([| (a_pre = b_pre) |]
|
||||||
|
&& [| ((0 - 0 ) = 0) |]
|
||||||
|
&& emp)
|
||||||
|
||
|
||||||
|
([| (a_pre > b_pre) |]
|
||||||
|
&& [| ((0 - 0 ) = 1) |]
|
||||||
|
&& emp)
|
||||||
|
.
|
||||||
|
|
||||||
|
(*----- Function mpn_copyi -----*)
|
||||||
|
|
||||||
|
Definition mpn_copyi_safety_wit_1 :=
|
||||||
|
forall (n_pre: Z) (s_pre: Z) (d_pre: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (val: Z) (l: (@list Z)) ,
|
||||||
|
[| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& ((( &( "i" ) )) # Int |->_)
|
||||||
|
** ((( &( "d" ) )) # Ptr |-> d_pre)
|
||||||
|
** (store_uint_array_rec d_pre 0 cap2 l2 )
|
||||||
|
** (store_uint_array d_pre 0 nil )
|
||||||
|
** ((( &( "n" ) )) # Int |-> n_pre)
|
||||||
|
** ((( &( "s" ) )) # Ptr |-> s_pre)
|
||||||
|
** (store_uint_array s_pre n_pre l )
|
||||||
|
** (store_undef_uint_array_rec s_pre (n_pre + 1 ) cap1 )
|
||||||
|
|--
|
||||||
|
[| (0 <= INT_MAX) |]
|
||||||
|
&& [| ((INT_MIN) <= 0) |]
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_safety_wit_2 :=
|
||||||
|
forall (n_pre: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (val: Z) (l: (@list Z)) (l': (@list Z)) (d: Z) (s: Z) (l_2: (@list Z)) (n: Z) (i: Z) (a: Z) (l2': (@list Z)) ,
|
||||||
|
[| (l' = (cons (a) (l2'))) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (n <= cap2) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (0 <= i) |]
|
||||||
|
&& [| (i <= n) |]
|
||||||
|
&& [| ((Zlength (l_2)) = n) |]
|
||||||
|
&& [| (list_store_Z l_2 val ) |]
|
||||||
|
&& [| (n <= cap1) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (store_uint_array d (i + 1 ) (replace_Znth (i) ((Znth i l_2 0)) ((app ((sublist (0) (i) (l_2))) ((cons (a) (nil)))))) )
|
||||||
|
** (store_uint_array s n l_2 )
|
||||||
|
** ((( &( "i" ) )) # Int |-> i)
|
||||||
|
** ((( &( "n" ) )) # Int |-> n)
|
||||||
|
** ((( &( "d" ) )) # Ptr |-> d)
|
||||||
|
** (store_uint_array_rec d (i + 1 ) cap2 l2' )
|
||||||
|
** ((( &( "s" ) )) # Ptr |-> s)
|
||||||
|
** (store_undef_uint_array_rec s (n + 1 ) cap1 )
|
||||||
|
|--
|
||||||
|
[| ((i + 1 ) <= INT_MAX) |]
|
||||||
|
&& [| ((INT_MIN) <= (i + 1 )) |]
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_entail_wit_1 :=
|
||||||
|
forall (n_pre: Z) (s_pre: Z) (d_pre: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (val: Z) (l_2: (@list Z)) ,
|
||||||
|
[| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l_2)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l_2 val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (store_uint_array_rec d_pre 0 cap2 l2 )
|
||||||
|
** (store_uint_array d_pre 0 nil )
|
||||||
|
** (store_uint_array s_pre n_pre l_2 )
|
||||||
|
** (store_undef_uint_array_rec s_pre (n_pre + 1 ) cap1 )
|
||||||
|
|--
|
||||||
|
EX (l': (@list Z)) (l: (@list Z)) ,
|
||||||
|
[| (0 <= 0) |]
|
||||||
|
&& [| (0 <= n_pre) |]
|
||||||
|
&& [| ((Zlength (l)) = n_pre) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l_2)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l_2 val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (store_uint_array s_pre n_pre l )
|
||||||
|
** (store_undef_uint_array_rec s_pre (n_pre + 1 ) cap1 )
|
||||||
|
** (store_uint_array d_pre 0 (sublist (0) (0) (l)) )
|
||||||
|
** (store_uint_array_rec d_pre 0 cap2 l' )
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_entail_wit_2 :=
|
||||||
|
forall (n_pre: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (val: Z) (l_2: (@list Z)) (l'_2: (@list Z)) (d: Z) (s: Z) (l_3: (@list Z)) (n: Z) (i: Z) (a: Z) (l2': (@list Z)) ,
|
||||||
|
[| (l'_2 = (cons (a) (l2'))) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (n <= cap2) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (0 <= i) |]
|
||||||
|
&& [| (i <= n) |]
|
||||||
|
&& [| ((Zlength (l_3)) = n) |]
|
||||||
|
&& [| (list_store_Z l_3 val ) |]
|
||||||
|
&& [| (n <= cap1) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l_2)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l_2 val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (store_uint_array d (i + 1 ) (replace_Znth (i) ((Znth i l_3 0)) ((app ((sublist (0) (i) (l_3))) ((cons (a) (nil)))))) )
|
||||||
|
** (store_uint_array s n l_3 )
|
||||||
|
** (store_uint_array_rec d (i + 1 ) cap2 l2' )
|
||||||
|
** (store_undef_uint_array_rec s (n + 1 ) cap1 )
|
||||||
|
|--
|
||||||
|
EX (l': (@list Z)) (l: (@list Z)) ,
|
||||||
|
[| (0 <= (i + 1 )) |]
|
||||||
|
&& [| ((i + 1 ) <= n) |]
|
||||||
|
&& [| ((Zlength (l)) = n) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| (n <= cap1) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l_2)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l_2 val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (store_uint_array s n l )
|
||||||
|
** (store_undef_uint_array_rec s (n + 1 ) cap1 )
|
||||||
|
** (store_uint_array d (i + 1 ) (sublist (0) ((i + 1 )) (l)) )
|
||||||
|
** (store_uint_array_rec d (i + 1 ) cap2 l' )
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_return_wit_1 :=
|
||||||
|
forall (n_pre: Z) (s_pre: Z) (d_pre: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (val: Z) (l_2: (@list Z)) (l': (@list Z)) (d: Z) (s: Z) (l: (@list Z)) (n: Z) (i: Z) ,
|
||||||
|
[| (i >= n) |]
|
||||||
|
&& [| (0 <= i) |]
|
||||||
|
&& [| (i <= n) |]
|
||||||
|
&& [| ((Zlength (l)) = n) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| (n <= cap1) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l_2)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l_2 val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (store_uint_array s n l )
|
||||||
|
** (store_undef_uint_array_rec s (n + 1 ) cap1 )
|
||||||
|
** (store_uint_array d i (sublist (0) (i) (l)) )
|
||||||
|
** (store_uint_array_rec d i cap2 l' )
|
||||||
|
|--
|
||||||
|
(mpd_store_Z s_pre val n_pre cap1 )
|
||||||
|
** (mpd_store_Z d_pre val n_pre cap2 )
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_partial_solve_wit_1 :=
|
||||||
|
forall (n_pre: Z) (s_pre: Z) (d_pre: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (val: Z) ,
|
||||||
|
[| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (mpd_store_Z s_pre val n_pre cap1 )
|
||||||
|
** (store_uint_array d_pre cap2 l2 )
|
||||||
|
|--
|
||||||
|
[| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (mpd_store_Z s_pre val n_pre cap1 )
|
||||||
|
** (store_uint_array d_pre cap2 l2 )
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_partial_solve_wit_2_pure :=
|
||||||
|
forall (n_pre: Z) (s_pre: Z) (d_pre: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (val: Z) (l: (@list Z)) ,
|
||||||
|
[| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& ((( &( "n" ) )) # Int |-> n_pre)
|
||||||
|
** ((( &( "s" ) )) # Ptr |-> s_pre)
|
||||||
|
** (store_uint_array s_pre n_pre l )
|
||||||
|
** (store_undef_uint_array_rec s_pre (n_pre + 1 ) cap1 )
|
||||||
|
** ((( &( "d" ) )) # Ptr |-> d_pre)
|
||||||
|
** (store_uint_array d_pre cap2 l2 )
|
||||||
|
|--
|
||||||
|
[| ((Zlength (l2)) = cap2) |]
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_partial_solve_wit_2_aux :=
|
||||||
|
forall (n_pre: Z) (s_pre: Z) (d_pre: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (val: Z) (l: (@list Z)) ,
|
||||||
|
[| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (store_uint_array s_pre n_pre l )
|
||||||
|
** (store_undef_uint_array_rec s_pre (n_pre + 1 ) cap1 )
|
||||||
|
** (store_uint_array d_pre cap2 l2 )
|
||||||
|
|--
|
||||||
|
[| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (store_uint_array d_pre cap2 l2 )
|
||||||
|
** (store_uint_array s_pre n_pre l )
|
||||||
|
** (store_undef_uint_array_rec s_pre (n_pre + 1 ) cap1 )
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_partial_solve_wit_2 := mpn_copyi_partial_solve_wit_2_pure -> mpn_copyi_partial_solve_wit_2_aux.
|
||||||
|
|
||||||
|
Definition mpn_copyi_partial_solve_wit_3_pure :=
|
||||||
|
forall (n_pre: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (val: Z) (l_2: (@list Z)) (l': (@list Z)) (d: Z) (s: Z) (l: (@list Z)) (n: Z) (i: Z) ,
|
||||||
|
[| (i < n) |]
|
||||||
|
&& [| (0 <= i) |]
|
||||||
|
&& [| (i <= n) |]
|
||||||
|
&& [| ((Zlength (l)) = n) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| (n <= cap1) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l_2)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l_2 val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& ((( &( "i" ) )) # Int |-> i)
|
||||||
|
** ((( &( "n" ) )) # Int |-> n)
|
||||||
|
** ((( &( "s" ) )) # Ptr |-> s)
|
||||||
|
** (store_uint_array s n l )
|
||||||
|
** (store_undef_uint_array_rec s (n + 1 ) cap1 )
|
||||||
|
** ((( &( "d" ) )) # Ptr |-> d)
|
||||||
|
** (store_uint_array d i (sublist (0) (i) (l)) )
|
||||||
|
** (store_uint_array_rec d i cap2 l' )
|
||||||
|
|--
|
||||||
|
[| (0 <= i) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (n <= cap2) |]
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_partial_solve_wit_3_aux :=
|
||||||
|
forall (n_pre: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (val: Z) (l: (@list Z)) (l': (@list Z)) (d: Z) (s: Z) (l_2: (@list Z)) (n: Z) (i: Z) ,
|
||||||
|
[| (i < n) |]
|
||||||
|
&& [| (0 <= i) |]
|
||||||
|
&& [| (i <= n) |]
|
||||||
|
&& [| ((Zlength (l_2)) = n) |]
|
||||||
|
&& [| (list_store_Z l_2 val ) |]
|
||||||
|
&& [| (n <= cap1) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (store_uint_array s n l_2 )
|
||||||
|
** (store_undef_uint_array_rec s (n + 1 ) cap1 )
|
||||||
|
** (store_uint_array d i (sublist (0) (i) (l_2)) )
|
||||||
|
** (store_uint_array_rec d i cap2 l' )
|
||||||
|
|--
|
||||||
|
[| (0 <= i) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (n <= cap2) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (0 <= i) |]
|
||||||
|
&& [| (i <= n) |]
|
||||||
|
&& [| ((Zlength (l_2)) = n) |]
|
||||||
|
&& [| (list_store_Z l_2 val ) |]
|
||||||
|
&& [| (n <= cap1) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (store_uint_array_rec d i cap2 l' )
|
||||||
|
** (store_uint_array d i (sublist (0) (i) (l_2)) )
|
||||||
|
** (store_uint_array s n l_2 )
|
||||||
|
** (store_undef_uint_array_rec s (n + 1 ) cap1 )
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_partial_solve_wit_3 := mpn_copyi_partial_solve_wit_3_pure -> mpn_copyi_partial_solve_wit_3_aux.
|
||||||
|
|
||||||
|
Definition mpn_copyi_partial_solve_wit_4 :=
|
||||||
|
forall (n_pre: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (val: Z) (l: (@list Z)) (l': (@list Z)) (d: Z) (s: Z) (l_2: (@list Z)) (n: Z) (i: Z) (a: Z) (l2': (@list Z)) ,
|
||||||
|
[| (l' = (cons (a) (l2'))) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (n <= cap2) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (0 <= i) |]
|
||||||
|
&& [| (i <= n) |]
|
||||||
|
&& [| ((Zlength (l_2)) = n) |]
|
||||||
|
&& [| (list_store_Z l_2 val ) |]
|
||||||
|
&& [| (n <= cap1) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (store_uint_array_rec d (i + 1 ) cap2 l2' )
|
||||||
|
** (store_uint_array d (i + 1 ) (app ((sublist (0) (i) (l_2))) ((cons (a) (nil)))) )
|
||||||
|
** (store_uint_array s n l_2 )
|
||||||
|
** (store_undef_uint_array_rec s (n + 1 ) cap1 )
|
||||||
|
|--
|
||||||
|
[| (l' = (cons (a) (l2'))) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (n <= cap2) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (0 <= i) |]
|
||||||
|
&& [| (i <= n) |]
|
||||||
|
&& [| ((Zlength (l_2)) = n) |]
|
||||||
|
&& [| (list_store_Z l_2 val ) |]
|
||||||
|
&& [| (n <= cap1) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (((s + (i * sizeof(UINT) ) )) # UInt |-> (Znth i l_2 0))
|
||||||
|
** (store_uint_array_missing_i_rec s i 0 n l_2 )
|
||||||
|
** (store_uint_array_rec d (i + 1 ) cap2 l2' )
|
||||||
|
** (store_uint_array d (i + 1 ) (app ((sublist (0) (i) (l_2))) ((cons (a) (nil)))) )
|
||||||
|
** (store_undef_uint_array_rec s (n + 1 ) cap1 )
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_partial_solve_wit_5 :=
|
||||||
|
forall (n_pre: Z) (cap2: Z) (cap1: Z) (l2: (@list Z)) (val: Z) (l: (@list Z)) (l': (@list Z)) (d: Z) (s: Z) (l_2: (@list Z)) (n: Z) (i: Z) (a: Z) (l2': (@list Z)) ,
|
||||||
|
[| (l' = (cons (a) (l2'))) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (n <= cap2) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (0 <= i) |]
|
||||||
|
&& [| (i <= n) |]
|
||||||
|
&& [| ((Zlength (l_2)) = n) |]
|
||||||
|
&& [| (list_store_Z l_2 val ) |]
|
||||||
|
&& [| (n <= cap1) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (store_uint_array s n l_2 )
|
||||||
|
** (store_uint_array_rec d (i + 1 ) cap2 l2' )
|
||||||
|
** (store_uint_array d (i + 1 ) (app ((sublist (0) (i) (l_2))) ((cons (a) (nil)))) )
|
||||||
|
** (store_undef_uint_array_rec s (n + 1 ) cap1 )
|
||||||
|
|--
|
||||||
|
[| (l' = (cons (a) (l2'))) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (n <= cap2) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (0 <= i) |]
|
||||||
|
&& [| (i <= n) |]
|
||||||
|
&& [| ((Zlength (l_2)) = n) |]
|
||||||
|
&& [| (list_store_Z l_2 val ) |]
|
||||||
|
&& [| (n <= cap1) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (n_pre <= cap1) |]
|
||||||
|
&& [| ((Zlength (l)) = n_pre) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& [| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& [| (cap2 >= n_pre) |]
|
||||||
|
&& (((d + (i * sizeof(UINT) ) )) # UInt |->_)
|
||||||
|
** (store_uint_array_missing_i_rec d i 0 (i + 1 ) (app ((sublist (0) (i) (l_2))) ((cons (a) (nil)))) )
|
||||||
|
** (store_uint_array s n l_2 )
|
||||||
|
** (store_uint_array_rec d (i + 1 ) cap2 l2' )
|
||||||
|
** (store_undef_uint_array_rec s (n + 1 ) cap1 )
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_which_implies_wit_1 :=
|
||||||
|
forall (cap1: Z) (val: Z) (n: Z) (s: Z) ,
|
||||||
|
(mpd_store_Z s val n cap1 )
|
||||||
|
|--
|
||||||
|
EX (l: (@list Z)) ,
|
||||||
|
[| (n <= cap1) |]
|
||||||
|
&& [| ((Zlength (l)) = n) |]
|
||||||
|
&& [| (cap1 <= 100000000) |]
|
||||||
|
&& [| (list_store_Z l val ) |]
|
||||||
|
&& (store_uint_array s n l )
|
||||||
|
** (store_undef_uint_array_rec s (n + 1 ) cap1 )
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_which_implies_wit_2 :=
|
||||||
|
forall (cap2: Z) (l2: (@list Z)) (d: Z) ,
|
||||||
|
[| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& (store_uint_array d cap2 l2 )
|
||||||
|
|--
|
||||||
|
[| ((Zlength (l2)) = cap2) |]
|
||||||
|
&& (store_uint_array_rec d 0 cap2 l2 )
|
||||||
|
** (store_uint_array d 0 nil )
|
||||||
|
.
|
||||||
|
|
||||||
|
Definition mpn_copyi_which_implies_wit_3 :=
|
||||||
|
forall (cap2: Z) (l': (@list Z)) (l: (@list Z)) (i: Z) (n: Z) (d: Z) ,
|
||||||
|
[| (0 <= i) |]
|
||||||
|
&& [| (i < n) |]
|
||||||
|
&& [| (n <= cap2) |]
|
||||||
|
&& (store_uint_array_rec d i cap2 l' )
|
||||||
|
** (store_uint_array d i (sublist (0) (i) (l)) )
|
||||||
|
|--
|
||||||
|
EX (a: Z) (l2': (@list Z)) ,
|
||||||
|
[| (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)))) )
|
||||||
|
.
|
||||||
|
|
||||||
|
Module Type VC_Correct.
|
||||||
|
|
||||||
|
Axiom proof_of_gmp_abs_safety_wit_1 : gmp_abs_safety_wit_1.
|
||||||
|
Axiom proof_of_gmp_abs_return_wit_1_1 : gmp_abs_return_wit_1_1.
|
||||||
|
Axiom proof_of_gmp_abs_return_wit_1_2 : gmp_abs_return_wit_1_2.
|
||||||
|
Axiom proof_of_gmp_max_return_wit_1_1 : gmp_max_return_wit_1_1.
|
||||||
|
Axiom proof_of_gmp_max_return_wit_1_2 : gmp_max_return_wit_1_2.
|
||||||
|
Axiom proof_of_gmp_cmp_safety_wit_1 : gmp_cmp_safety_wit_1.
|
||||||
|
Axiom proof_of_gmp_cmp_safety_wit_2 : gmp_cmp_safety_wit_2.
|
||||||
|
Axiom proof_of_gmp_cmp_safety_wit_3 : gmp_cmp_safety_wit_3.
|
||||||
|
Axiom proof_of_gmp_cmp_safety_wit_4 : gmp_cmp_safety_wit_4.
|
||||||
|
Axiom proof_of_gmp_cmp_return_wit_1_1 : gmp_cmp_return_wit_1_1.
|
||||||
|
Axiom proof_of_gmp_cmp_return_wit_1_2 : gmp_cmp_return_wit_1_2.
|
||||||
|
Axiom proof_of_gmp_cmp_return_wit_1_3 : gmp_cmp_return_wit_1_3.
|
||||||
|
Axiom proof_of_mpn_copyi_safety_wit_1 : mpn_copyi_safety_wit_1.
|
||||||
|
Axiom proof_of_mpn_copyi_safety_wit_2 : mpn_copyi_safety_wit_2.
|
||||||
|
Axiom proof_of_mpn_copyi_entail_wit_1 : mpn_copyi_entail_wit_1.
|
||||||
|
Axiom proof_of_mpn_copyi_entail_wit_2 : mpn_copyi_entail_wit_2.
|
||||||
|
Axiom proof_of_mpn_copyi_return_wit_1 : mpn_copyi_return_wit_1.
|
||||||
|
Axiom proof_of_mpn_copyi_partial_solve_wit_1 : mpn_copyi_partial_solve_wit_1.
|
||||||
|
Axiom proof_of_mpn_copyi_partial_solve_wit_2_pure : mpn_copyi_partial_solve_wit_2_pure.
|
||||||
|
Axiom proof_of_mpn_copyi_partial_solve_wit_2 : mpn_copyi_partial_solve_wit_2.
|
||||||
|
Axiom proof_of_mpn_copyi_partial_solve_wit_3_pure : mpn_copyi_partial_solve_wit_3_pure.
|
||||||
|
Axiom proof_of_mpn_copyi_partial_solve_wit_3 : mpn_copyi_partial_solve_wit_3.
|
||||||
|
Axiom proof_of_mpn_copyi_partial_solve_wit_4 : mpn_copyi_partial_solve_wit_4.
|
||||||
|
Axiom proof_of_mpn_copyi_partial_solve_wit_5 : mpn_copyi_partial_solve_wit_5.
|
||||||
|
Axiom proof_of_mpn_copyi_which_implies_wit_1 : mpn_copyi_which_implies_wit_1.
|
||||||
|
Axiom proof_of_mpn_copyi_which_implies_wit_2 : mpn_copyi_which_implies_wit_2.
|
||||||
|
Axiom proof_of_mpn_copyi_which_implies_wit_3 : mpn_copyi_which_implies_wit_3.
|
||||||
|
|
||||||
|
End VC_Correct.
|
6
projects/lib/gmp_goal_check.v
Normal file
6
projects/lib/gmp_goal_check.v
Normal file
@ -0,0 +1,6 @@
|
|||||||
|
From Require Import gmp_goal gmp_proof_auto gmp_proof_manual.
|
||||||
|
|
||||||
|
Module VC_Correctness : VC_Correct.
|
||||||
|
Include gmp_proof_auto.
|
||||||
|
Include gmp_proof_manual.
|
||||||
|
End VC_Correctness.
|
65
projects/lib/gmp_proof_auto.v
Normal file
65
projects/lib/gmp_proof_auto.v
Normal file
@ -0,0 +1,65 @@
|
|||||||
|
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.
|
||||||
|
|
152
projects/lib/gmp_proof_manual.v
Normal file
152
projects/lib/gmp_proof_manual.v
Normal file
@ -0,0 +1,152 @@
|
|||||||
|
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.store_uarray_rec_equals_store_uarray 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.store_uarray_rec_equals_store_uarray 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.
|
@ -5,7 +5,7 @@
|
|||||||
|
|
||||||
int gmp_abs(int x)
|
int gmp_abs(int x)
|
||||||
/*@
|
/*@
|
||||||
Require emp
|
Require INT_MIN < x && x <= INT_MAX
|
||||||
Ensure __return == Zabs(x)
|
Ensure __return == Zabs(x)
|
||||||
*/
|
*/
|
||||||
{
|
{
|
||||||
@ -25,9 +25,10 @@ int gmp_cmp(int a, int b)
|
|||||||
/*@
|
/*@
|
||||||
Require emp
|
Require emp
|
||||||
Ensure
|
Ensure
|
||||||
a > b && __return == 1 ||
|
emp *
|
||||||
|
(a > b && __return == 1 ||
|
||||||
a == b && __return == 0 ||
|
a == b && __return == 0 ||
|
||||||
a < b && __return == -1
|
a < b && __return == -1)
|
||||||
*/
|
*/
|
||||||
{
|
{
|
||||||
return (a > b) - (a < b);
|
return (a > b) - (a < b);
|
||||||
@ -39,25 +40,60 @@ int gmp_cmp(int a, int b)
|
|||||||
void
|
void
|
||||||
mpn_copyi (unsigned int *d, unsigned int *s, int n)
|
mpn_copyi (unsigned int *d, unsigned int *s, int n)
|
||||||
/*@
|
/*@
|
||||||
With val data cap1 cap2
|
With val l2 cap1 cap2
|
||||||
Require
|
Require
|
||||||
mpd_store_Z(s, val, n, cap1) *
|
mpd_store_Z(s, val, n, cap1) *
|
||||||
store_uint_array(d, cap2, data)
|
store_uint_array(d, cap2, l2) &&
|
||||||
|
Zlength(l2) == cap2 &&
|
||||||
|
cap2 >= n
|
||||||
Ensure
|
Ensure
|
||||||
mpd_store_Z(s, val, n, cap1) *
|
mpd_store_Z(s, val, n, cap1) *
|
||||||
mpd_store_Z(d, val, n, cap2)
|
mpd_store_Z(d, val, n, cap2)
|
||||||
*/
|
*/
|
||||||
{
|
{
|
||||||
/*
|
/*@
|
||||||
mpd_store_Z(s, val, n, cap1)
|
mpd_store_Z(s, val, n, cap1)
|
||||||
which implies
|
which implies
|
||||||
exists l,
|
exists l,
|
||||||
store_uint_array(s, n, l) **
|
n <= cap1 &&
|
||||||
store_undef
|
Zlength(l) == n &&
|
||||||
|
cap1 <= 100000000 &&
|
||||||
|
store_uint_array(s, n, l) *
|
||||||
|
store_undef_uint_array_rec(s, n + 1, 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;
|
int i;
|
||||||
for (i = 0; i < n; i++)
|
/*@ Inv
|
||||||
d[i] = s[i];
|
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 + 1, 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 */
|
||||||
@ -66,19 +102,53 @@ mpn_cmp (unsigned int *ap, unsigned int *bp, int n)
|
|||||||
/*@
|
/*@
|
||||||
With cap1 cap2 val1 val2
|
With cap1 cap2 val1 val2
|
||||||
Require
|
Require
|
||||||
mpd_store_Z(ap, val1, n, cap1) **
|
mpd_store_Z(ap, val1, n, cap1) *
|
||||||
mpd_store_Z(bp, val2, n, cap2)
|
mpd_store_Z(bp, val2, n, cap2) &&
|
||||||
|
n <= cap1 && n <= cap2
|
||||||
Ensure
|
Ensure
|
||||||
val1 > val2 && __return == 1 ||
|
val1 > val2 && __return == 1 ||
|
||||||
val1 == val2 && __return == 0 ||
|
val1 == val2 && __return == 0 ||
|
||||||
val1 < val2 && __return == -1
|
val1 < val2 && __return == -1
|
||||||
*/
|
*/
|
||||||
{
|
{
|
||||||
|
/*@
|
||||||
|
mpd_store_Z(ap, val1, n, cap1) * mpd_store_Z(bp, val2, n, cap2)
|
||||||
|
which implies
|
||||||
|
exists l1 l2,
|
||||||
|
mpd_store_list(ap, l1, cap1) * mpd_store_list(bp, l2, cap2) &&
|
||||||
|
list_store_Z(l1, val1) && list_store_Z(l2, val2) &&
|
||||||
|
n == Zlength(l1) && n == Zlength(l2)
|
||||||
|
*/
|
||||||
|
/*@
|
||||||
|
Given l1 l2
|
||||||
|
*/
|
||||||
|
--n;
|
||||||
|
/*@Inv
|
||||||
|
mpd_store_list(ap, l1, cap1) * mpd_store_list(bp, l2, cap2) &&
|
||||||
|
list_store_Z(l1, val1) && list_store_Z(l2, val2) &&
|
||||||
|
n@pre == Zlength(l1) && n@pre == Zlength(l2) &&
|
||||||
|
sublist(n, n@pre, l1) == sublist(n, n@pre, l2)
|
||||||
|
*/
|
||||||
|
while (n >= 0)
|
||||||
|
{
|
||||||
|
/*@
|
||||||
|
mpd_store_list(ap, l1, cap1) * mpd_store_list(bp, l2, cap2)
|
||||||
|
which implies
|
||||||
|
store_uint_array(ap, n, l1) * store_uint_array(bp, n, l2) *
|
||||||
|
store_undef_uint_array(ap, n + 1, cap1) * store_uint_array(bp, n + 1, cap2) &&
|
||||||
|
*/
|
||||||
|
if (ap[n] != bp[n])
|
||||||
|
return ap[n] > bp[n] ? 1 : -1;
|
||||||
|
--n;
|
||||||
|
}
|
||||||
|
// Note: The parser cannot parse "--n" in loop 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;
|
||||||
}
|
}
|
||||||
|
|
||||||
|
@ -36,7 +36,7 @@ gmp_free_limbs (unsigned int *old, int size);
|
|||||||
|
|
||||||
void mpn_copyi (unsigned int *d, unsigned int *s, int n);
|
void mpn_copyi (unsigned int *d, unsigned int *s, int n);
|
||||||
|
|
||||||
int mpn_cmp (unsigned int *, unsigned int *, int);
|
int mpn_cmp (unsigned int *ap, unsigned int *bp, int n);
|
||||||
|
|
||||||
unsigned int mpn_add_1 (unsigned int *, unsigned int *, int, unsigned int);
|
unsigned int mpn_add_1 (unsigned int *, unsigned int *, int, unsigned int);
|
||||||
unsigned int mpn_add_n (unsigned int *, unsigned int *, unsigned int *, int);
|
unsigned int mpn_add_n (unsigned int *, unsigned int *, unsigned int *, int);
|
||||||
@ -59,7 +59,9 @@ 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)
|
Extern Coq (Zabs : Z -> Z)
|
||||||
(Zmax: Z -> Z -> Z)
|
(Zmax : Z -> Z -> Z)
|
||||||
(mpd_store_Z: Z -> Z -> Z -> Z -> Assertion)
|
(mpd_store_Z : Z -> Z -> Z -> Z -> Assertion)
|
||||||
|
(mpd_store_list : Z -> list Z -> Z -> Assertion)
|
||||||
|
(list_store_Z : list Z -> Z -> Prop)
|
||||||
*/
|
*/
|
@ -9,7 +9,7 @@
|
|||||||
id : 1
|
id : 1
|
||||||
priority : core(1)
|
priority : core(1)
|
||||||
left : store_uint_array(?p, ?n, ?l) at 0
|
left : store_uint_array(?p, ?n, ?l) at 0
|
||||||
right : data_at(?p + (?i * sizeof(I32)), I32, ?v) at 1
|
right : data_at(?p + (?i * sizeof(U32)), U32, ?v) at 1
|
||||||
check : infer(0 <= i);
|
check : infer(0 <= i);
|
||||||
infer(i < n);
|
infer(i < n);
|
||||||
action : right_erase(1);
|
action : right_erase(1);
|
||||||
@ -20,7 +20,7 @@ action : right_erase(1);
|
|||||||
id : 2
|
id : 2
|
||||||
priority : post(1)
|
priority : post(1)
|
||||||
left : store_uint_array_missing_i_rec(?p, ?i, 0, ?n, ?l) at 0
|
left : store_uint_array_missing_i_rec(?p, ?i, 0, ?n, ?l) at 0
|
||||||
data_at(p + i * sizeof(I32), I32, l[i]) at 1
|
data_at(p + i * sizeof(U32), U32, l[i]) at 1
|
||||||
check : infer(0 <= i);
|
check : infer(0 <= i);
|
||||||
infer(i < n);
|
infer(i < n);
|
||||||
action : left_erase(1);
|
action : left_erase(1);
|
||||||
@ -30,7 +30,7 @@ action : left_erase(1);
|
|||||||
id : 3
|
id : 3
|
||||||
priority : post(3)
|
priority : post(3)
|
||||||
left : store_uint_array_missing_i_rec(?p, ?i, 0, ?n, ?l) at 0
|
left : store_uint_array_missing_i_rec(?p, ?i, 0, ?n, ?l) at 0
|
||||||
data_at(p + i * sizeof(I32), I32, ?v) at 1
|
data_at(p + i * sizeof(U32), U32, ?v) at 1
|
||||||
check : infer(0 <= i);
|
check : infer(0 <= i);
|
||||||
infer(i < n);
|
infer(i < n);
|
||||||
action : left_erase(1);
|
action : left_erase(1);
|
||||||
@ -65,7 +65,7 @@ action : right_erase(1);
|
|||||||
id : 7
|
id : 7
|
||||||
priority : core(1)
|
priority : core(1)
|
||||||
left : store_uint_array_rec(?p, ?x, ?y, ?l) at 0
|
left : store_uint_array_rec(?p, ?x, ?y, ?l) at 0
|
||||||
right : data_at(?p + (?i * sizeof(I32)), I32, ?v) at 1
|
right : data_at(?p + (?i * sizeof(U32)), U32, ?v) at 1
|
||||||
check : infer(x <= i);
|
check : infer(x <= i);
|
||||||
infer(i < y);
|
infer(i < y);
|
||||||
action : right_erase(1);
|
action : right_erase(1);
|
||||||
@ -108,7 +108,7 @@ action : right_erase(0);
|
|||||||
id : 11
|
id : 11
|
||||||
priority : post(1)
|
priority : post(1)
|
||||||
left : store_uint_array_missing_i_rec(?p, ?i, ?x, ?y, ?l) at 0
|
left : store_uint_array_missing_i_rec(?p, ?i, ?x, ?y, ?l) at 0
|
||||||
data_at(p + ?i * sizeof(I32), I32, l[?i - ?x]) at 1
|
data_at(p + ?i * sizeof(U32), U32, l[?i - ?x]) at 1
|
||||||
check : infer(x <= i);
|
check : infer(x <= i);
|
||||||
infer(i < y);
|
infer(i < y);
|
||||||
action : left_erase(1);
|
action : left_erase(1);
|
||||||
@ -118,7 +118,7 @@ action : left_erase(1);
|
|||||||
id : 12
|
id : 12
|
||||||
priority : post(4)
|
priority : post(4)
|
||||||
left : store_uint_array_missing_i_rec(?p, ?i, ?x, ?y, ?l) at 0
|
left : store_uint_array_missing_i_rec(?p, ?i, ?x, ?y, ?l) at 0
|
||||||
data_at(p + ?i * sizeof(I32), I32, ?v) at 1
|
data_at(p + ?i * sizeof(U32), U32, ?v) at 1
|
||||||
check : infer(x <= i);
|
check : infer(x <= i);
|
||||||
infer(i < y);
|
infer(i < y);
|
||||||
action : left_erase(1);
|
action : left_erase(1);
|
||||||
|
Reference in New Issue
Block a user