theory IICF_Array_List32
imports
"Refine_Imperative_HOL.IICF_List"
Separation_Logic_Imperative_HOL.Array_Blit
Array_UInt
WB_Word_Assn
begin
type_synonym 'a array_list32 = "'a Heap.array × uint32"
definition "is_array_list32 l ≡ λ(a,n). ∃⇩Al'. a ↦⇩a l' * ↑(nat_of_uint32 n ≤ length l' ∧ l = take (nat_of_uint32 n) l' ∧ length l'>0 ∧ nat_of_uint32 n ≤ uint32_max ∧ length l' ≤ uint32_max)"
lemma is_array_list32_prec[safe_constraint_rules]: "precise is_array_list32"
unfolding is_array_list32_def[abs_def]
apply(rule preciseI)
apply(simp split: prod.splits)
using preciseD snga_prec by fastforce
definition "arl32_empty ≡ do {
a ← Array.new initial_capacity default;
return (a,0)
}"
definition "arl32_empty_sz init_cap ≡ do {
a ← Array.new (min uint32_max (max init_cap minimum_capacity)) default;
return (a,0)
}"
definition uint32_max_uint32 :: uint32 where
‹uint32_max_uint32 = 2 ^32 - 1›
definition "arl32_append ≡ λ(a,n) x. do {
len ← length_u_code a;
if n<len then do {
a ← Array_upd_u n x a;
return (a,n+1)
} else do {
let newcap = (if len < uint32_max_uint32 >> 1 then 2 * len else uint32_max_uint32);
a ← array_grow a (nat_of_uint32 newcap) default;
a ← Array_upd_u n x a;
return (a,n+1)
}
}"
definition "arl32_copy ≡ λ(a,n). do {
a ← array_copy a;
return (a,n)
}"
definition arl32_length :: "'a::heap array_list32 ⇒ uint32 Heap" where
"arl32_length ≡ λ(a,n). return (n)"
definition arl32_is_empty :: "'a::heap array_list32 ⇒ bool Heap" where
"arl32_is_empty ≡ λ(a,n). return (n=0)"
definition arl32_last :: "'a::heap array_list32 ⇒ 'a Heap" where
"arl32_last ≡ λ(a,n). do {
nth_u_code a (n - 1)
}"
definition arl32_butlast :: "'a::heap array_list32 ⇒ 'a array_list32 Heap" where
"arl32_butlast ≡ λ(a,n). do {
let n = n - 1;
len ← length_u_code a;
if (n*4 < len ∧ nat_of_uint32 n*2≥minimum_capacity) then do {
a ← array_shrink a (nat_of_uint32 n*2);
return (a,n)
} else
return (a,n)
}"
definition arl32_get :: "'a::heap array_list32 ⇒ uint32 ⇒ 'a Heap" where
"arl32_get ≡ λ(a,n) i. nth_u_code a i"
definition arl32_set :: "'a::heap array_list32 ⇒ uint32 ⇒ 'a ⇒ 'a array_list32 Heap" where
"arl32_set ≡ λ(a,n) i x. do { a ← heap_array_set_u a i x; return (a,n)}"
lemma arl32_empty_rule[sep_heap_rules]: "< emp > arl32_empty <is_array_list32 []>"
by (sep_auto simp: arl32_empty_def is_array_list32_def initial_capacity_def uint32_max_def)
lemma arl32_empty_sz_rule[sep_heap_rules]: "< emp > arl32_empty_sz N <is_array_list32 []>"
by (sep_auto simp: arl32_empty_sz_def is_array_list32_def minimum_capacity_def uint32_max_def)
lemma arl32_copy_rule[sep_heap_rules]: "< is_array_list32 l a > arl32_copy a <λr. is_array_list32 l a * is_array_list32 l r>"
by (sep_auto simp: arl32_copy_def is_array_list32_def)
lemma nat_of_uint32_shiftl: ‹nat_of_uint32 (xs >> a) = nat_of_uint32 xs >> a›
by transfer (auto simp: unat_shiftr nat_shifl_div)
lemma [simp]: ‹nat_of_uint32 uint32_max_uint32 = uint32_max›
by (auto simp: nat_of_uint32_mult_le nat_of_uint32_shiftl uint32_max_uint32_def uint32_max_def)[]
lemma ‹2 * (uint32_max div 2) = uint32_max - 1›
by (auto simp: nat_of_uint32_mult_le nat_of_uint32_shiftl uint32_max_uint32_def uint32_max_def)[]
lemma arl32_append_rule[sep_heap_rules]:
assumes ‹length l < uint32_max›
shows "< is_array_list32 l a >
arl32_append a x
<λa. is_array_list32 (l@[x]) a >⇩t"
proof -
have [simp]: ‹⋀x1 x2 y ys.
x2 < uint32_of_nat ys ⟹
nat_of_uint32 x2 ≤ ys ⟹
ys ≤ uint32_max ⟹ nat_of_uint32 x2 < ys›
by (metis nat_of_uint32_less_iff nat_of_uint32_uint32_of_nat_id)
have [simp]: ‹⋀x2 ys. x2 < uint32_of_nat (Suc (ys)) ⟹
Suc (ys) ≤ uint32_max ⟹
nat_of_uint32 (x2 + 1) = 1 + nat_of_uint32 x2›
by (smt ab_semigroup_add_class.add.commute le_neq_implies_less less_or_eq_imp_le
less_trans_Suc linorder_neqE_nat nat_of_uint32_012(3) nat_of_uint32_add
nat_of_uint32_less_iff nat_of_uint32_uint32_of_nat_id not_less_eq plus_1_eq_Suc)
have [dest]: ‹⋀x2a x2 ys. x2 < uint32_of_nat (Suc (ys)) ⟹
Suc (ys) ≤ uint32_max ⟹
nat_of_uint32 x2 = Suc x2a ⟹Suc x2a ≤ ys›
by (metis less_Suc_eq_le nat_of_uint32_less_iff nat_of_uint32_uint32_of_nat_id)
have [simp]: ‹⋀ys. ys ≤ uint32_max ⟹
uint32_of_nat ys ≤ uint32_max_uint32 >> Suc 0 ⟹
nat_of_uint32 (2 * uint32_of_nat ys) = 2 * ys›
by (subst (asm) nat_of_uint32_le_iff[symmetric])
(auto simp: nat_of_uint32_uint32_of_nat_id uint32_max_uint32_def uint32_max_def nat_of_uint32_shiftl
nat_of_uint32_mult_le)
have [simp]: ‹⋀ys. ys ≤ uint32_max ⟹
uint32_of_nat ys ≤ uint32_max_uint32 >> Suc 0 ⟷ ys ≤ uint32_max div 2›
by (subst nat_of_uint32_le_iff[symmetric])
(auto simp: nat_of_uint32_uint32_of_nat_id uint32_max_uint32_def uint32_max_def nat_of_uint32_shiftl
nat_of_uint32_mult_le)
have [simp]: ‹⋀ys. ys ≤ uint32_max ⟹
uint32_of_nat ys < uint32_max_uint32 >> Suc 0 ⟷ ys < uint32_max div 2›
by (subst nat_of_uint32_less_iff[symmetric])
(auto simp: nat_of_uint32_uint32_of_nat_id uint32_max_uint32_def uint32_max_def nat_of_uint32_shiftl
nat_of_uint32_mult_le)
show ?thesis
using assms
apply (sep_auto
simp: arl32_append_def is_array_list32_def take_update_last neq_Nil_conv nat_of_uint32_mult_le
length_u_code_def min_def nat_of_uint32_add nat_of_uint32_uint32_of_nat_id
take_Suc_conv_app_nth list_update_append nat_of_uint32_0_iff
split: if_split
split: prod.splits nat.split)
apply (subst Array_upd_u_def)
apply (sep_auto
simp: arl32_append_def is_array_list32_def take_update_last neq_Nil_conv nat_of_uint32_mult_le
length_u_code_def min_def nat_of_uint32_add nat_of_uint32_uint32_of_nat_id
take_Suc_conv_app_nth list_update_append
split: if_split
split: prod.splits nat.split)
apply (sep_auto
simp: arl32_append_def is_array_list32_def take_update_last neq_Nil_conv nat_of_uint32_mult_le
length_u_code_def min_def nat_of_uint32_add nat_of_uint32_uint32_of_nat_id
take_Suc_conv_app_nth list_update_append
split: if_split
split: prod.splits nat.split)
apply (sep_auto
simp: arl32_append_def is_array_list32_def take_update_last neq_Nil_conv nat_of_uint32_mult_le
length_u_code_def min_def nat_of_uint32_add nat_of_uint32_uint32_of_nat_id
take_Suc_conv_app_nth list_update_append
split: if_split
split: prod.splits nat.split)
apply (subst Array_upd_u_def)
apply (sep_auto
simp: arl32_append_def is_array_list32_def take_update_last neq_Nil_conv nat_of_uint32_mult_le
length_u_code_def min_def nat_of_uint32_add nat_of_uint32_uint32_of_nat_id
take_Suc_conv_app_nth list_update_append
split: if_split
split: prod.splits nat.split)
apply (sep_auto
simp: arl32_append_def is_array_list32_def take_update_last neq_Nil_conv nat_of_uint32_mult_le
length_u_code_def min_def nat_of_uint32_add nat_of_uint32_uint32_of_nat_id
take_Suc_conv_app_nth list_update_append
split: if_split
split: prod.splits nat.split)
apply (sep_auto
simp: arl32_append_def is_array_list32_def take_update_last neq_Nil_conv nat_of_uint32_mult_le
length_u_code_def min_def nat_of_uint32_add nat_of_uint32_uint32_of_nat_id
take_Suc_conv_app_nth list_update_append
split: if_split
split: prod.splits nat.split)
apply (subst Array_upd_u_def)
apply (rule frame_rule)
apply (rule upd_rule)
apply (sep_auto
simp: arl32_append_def is_array_list32_def take_update_last neq_Nil_conv nat_of_uint32_mult_le
length_u_code_def min_def nat_of_uint32_add nat_of_uint32_uint32_of_nat_id
take_Suc_conv_app_nth list_update_append nat_of_uint32_0_iff
split: if_splits
split: prod.splits nat.split)
apply (sep_auto
simp: arl32_append_def is_array_list32_def take_update_last neq_Nil_conv nat_of_uint32_mult_le
length_u_code_def min_def nat_of_uint32_add nat_of_uint32_uint32_of_nat_id
take_Suc_conv_app_nth list_update_append
split: if_splits
split: prod.splits nat.split)
done
qed
lemma arl32_length_rule[sep_heap_rules]: "
<is_array_list32 l a>
arl32_length a
<λr. is_array_list32 l a * ↑(nat_of_uint32 r=length l)>"
by (sep_auto simp: arl32_length_def is_array_list32_def)
lemma arl32_is_empty_rule[sep_heap_rules]: "
<is_array_list32 l a>
arl32_is_empty a
<λr. is_array_list32 l a * ↑(r⟷(l=[]))>"
by (sep_auto simp: arl32_is_empty_def nat_of_uint32_0_iff is_array_list32_def)
lemma nat_of_uint32_ge_minus:
‹ai ≥ bi ⟹
nat_of_uint32 (ai - bi) = nat_of_uint32 ai - nat_of_uint32 bi›
apply transfer
unfolding unat_def
by (subst uint_sub_lem[THEN iffD1])
(auto simp: unat_def uint_nonnegative nat_diff_distrib word_le_def[symmetric] intro: leI)
lemma arl32_last_rule[sep_heap_rules]: "
l≠[] ⟹
<is_array_list32 l a>
arl32_last a
<λr. is_array_list32 l a * ↑(r=last l)>"
by (sep_auto simp: arl32_last_def is_array_list32_def nth_u_code_def Array.nth'_def last_take_nth_conv
nat_of_integer_integer_of_nat nat_of_uint32_ge_minus nat_of_uint32_le_iff[symmetric]
simp flip: nat_of_uint32_code)
lemma arl32_get_rule[sep_heap_rules]: "
i<length l ⟹ (i', i) ∈ uint32_nat_rel ⟹
<is_array_list32 l a>
arl32_get a i'
<λr. is_array_list32 l a * ↑(r=l!i)>"
by (sep_auto simp: arl32_get_def nth_u_code_def is_array_list32_def uint32_nat_rel_def
Array.nth'_def br_def split: prod.split simp flip: nat_of_uint32_code)
lemma arl32_set_rule[sep_heap_rules]: "
i<length l ⟹ (i', i) ∈ uint32_nat_rel ⟹
<is_array_list32 l a>
arl32_set a i' x
<is_array_list32 (l[i:=x])>"
by (sep_auto simp: arl32_set_def is_array_list32_def heap_array_set_u_def uint32_nat_rel_def
heap_array_set'_u_def br_def Array.upd'_def split: prod.split simp flip: nat_of_uint32_code)
definition "arl32_assn A ≡ hr_comp is_array_list32 (⟨the_pure A⟩list_rel)"
lemmas [safe_constraint_rules] = CN_FALSEI[of is_pure "arl32_assn A" for A]
lemma arl32_assn_comp: "is_pure A ⟹ hr_comp (arl32_assn A) (⟨B⟩list_rel) = arl32_assn (hr_comp A B)"
unfolding arl32_assn_def
by (auto simp: hr_comp_the_pure hr_comp_assoc list_rel_compp)
lemma arl32_assn_comp': "hr_comp (arl32_assn id_assn) (⟨B⟩list_rel) = arl32_assn (pure B)"
by (simp add: arl32_assn_comp)
context
notes [fcomp_norm_unfold] = arl32_assn_def[symmetric] arl32_assn_comp'
notes [intro!] = hfrefI hn_refineI[THEN hn_refine_preI]
notes [simp] = pure_def hn_ctxt_def invalid_assn_def
begin
lemma arl32_empty_hnr_aux: "(uncurry0 arl32_empty,uncurry0 (RETURN op_list_empty)) ∈ unit_assn⇧k →⇩a is_array_list32"
by sep_auto
sepref_decl_impl (no_register) arl32_empty: arl32_empty_hnr_aux .
lemma arl32_empty_sz_hnr_aux: "(uncurry0 (arl32_empty_sz N),uncurry0 (RETURN op_list_empty)) ∈ unit_assn⇧k →⇩a is_array_list32"
by sep_auto
sepref_decl_impl (no_register) arl32_empty_sz: arl32_empty_sz_hnr_aux .
definition "op_arl32_empty ≡ op_list_empty"
definition "op_arl32_empty_sz (N::nat) ≡ op_list_empty"
lemma arl32_copy_hnr_aux: "(arl32_copy,RETURN o op_list_copy) ∈ is_array_list32⇧k →⇩a is_array_list32"
by sep_auto
sepref_decl_impl arl32_copy: arl32_copy_hnr_aux .
lemma arl32_append_hnr_aux: "(uncurry arl32_append,uncurry (RETURN oo op_list_append)) ∈ [λ(xs, x). length xs < uint32_max]⇩a (is_array_list32⇧d *⇩a id_assn⇧k) → is_array_list32"
by sep_auto
sepref_decl_impl arl32_append: arl32_append_hnr_aux
unfolding fref_param1 by (auto intro!: frefI nres_relI simp: list_rel_imp_same_length)
lemma arl32_length_hnr_aux: "(arl32_length,RETURN o op_list_length) ∈ is_array_list32⇧k →⇩a uint32_nat_assn"
by (sep_auto simp: uint32_nat_rel_def br_def)
sepref_decl_impl arl32_length: arl32_length_hnr_aux .
lemma arl32_is_empty_hnr_aux: "(arl32_is_empty,RETURN o op_list_is_empty) ∈ is_array_list32⇧k →⇩a bool_assn"
by sep_auto
sepref_decl_impl arl32_is_empty: arl32_is_empty_hnr_aux .
lemma arl32_last_hnr_aux: "(arl32_last,RETURN o op_list_last) ∈ [pre_list_last]⇩a is_array_list32⇧k → id_assn"
by sep_auto
sepref_decl_impl arl32_last: arl32_last_hnr_aux .
lemma arl32_get_hnr_aux: "(uncurry arl32_get,uncurry (RETURN oo op_list_get)) ∈ [λ(l,i). i<length l]⇩a (is_array_list32⇧k *⇩a uint32_nat_assn⇧k) → id_assn"
by sep_auto
sepref_decl_impl arl32_get: arl32_get_hnr_aux .
lemma arl32_set_hnr_aux: "(uncurry2 arl32_set,uncurry2 (RETURN ooo op_list_set)) ∈ [λ((l,i),_). i<length l]⇩a (is_array_list32⇧d *⇩a uint32_nat_assn⇧k *⇩a id_assn⇧k) → is_array_list32"
by sep_auto
sepref_decl_impl arl32_set: arl32_set_hnr_aux .
sepref_definition arl32_swap is "uncurry2 mop_list_swap" :: "((arl32_assn id_assn)⇧d *⇩a uint32_nat_assn⇧k *⇩a uint32_nat_assn⇧k →⇩a arl32_assn id_assn)"
unfolding gen_mop_list_swap[abs_def]
by sepref
sepref_decl_impl (ismop) arl32_swap: arl32_swap.refine .
end
interpretation arl32: list_custom_empty "arl32_assn A" arl32_empty op_arl32_empty
apply unfold_locales
apply (rule arl32_empty_hnr)
by (auto simp: op_arl32_empty_def)
lemma [def_pat_rules]: "op_arl32_empty_sz$N ≡ UNPROTECT (op_arl32_empty_sz N)" by simp
interpretation arl32_sz: list_custom_empty "arl32_assn A" "arl32_empty_sz N" "PR_CONST (op_arl32_empty_sz N)"
apply unfold_locales
apply (rule arl32_empty_sz_hnr)
by (auto simp: op_arl32_empty_sz_def)
definition arl32_to_arl_conv where
‹arl32_to_arl_conv S = S›
definition arl32_to_arl :: ‹'a array_list32 ⇒ 'a array_list› where
‹arl32_to_arl = (λ(xs, n). (xs, nat_of_uint32 n))›
lemma arl32_to_arl_hnr[sepref_fr_rules]:
‹(return o arl32_to_arl, RETURN o arl32_to_arl_conv) ∈ (arl32_assn R)⇧d →⇩a arl_assn R›
by (sepref_to_hoare)
(sep_auto simp: arl32_to_arl_def arl32_to_arl_conv_def arl_assn_def arl32_assn_def is_array_list32_def
is_array_list_def hr_comp_def)
definition arl32_take where
‹arl32_take n = (λ(xs, _). (xs, n))›
lemma arl32_take[sepref_fr_rules]:
‹(uncurry (return oo arl32_take), uncurry (RETURN oo take)) ∈
[λ(n, xs). n ≤ length xs]⇩a uint32_nat_assn⇧k *⇩a (arl32_assn R)⇧d → arl32_assn R›
by (sepref_to_hoare)
(sep_auto simp: arl32_assn_def arl32_take_def is_array_list32_def hr_comp_def
uint32_nat_rel_def br_def list_rel_def list_all2_conv_all_nth)
definition arl32_butlast_nonresizing :: ‹'a array_list32 ⇒ 'a array_list32› where
‹arl32_butlast_nonresizing = (λ(xs, a). (xs, a - 1))›
lemma butlast32_nonresizing_hnr[sepref_fr_rules]:
‹(return o arl32_butlast_nonresizing, RETURN o butlast_nonresizing) ∈
[λxs. xs ≠ []]⇩a (arl32_assn R)⇧d → arl32_assn R›
proof -
have [simp]: ‹nat_of_uint32 (b - 1) = nat_of_uint32 b - 1›
if
‹x ≠ []› and
‹(take (nat_of_uint32 b) l', x) ∈ ⟨the_pure R⟩list_rel›
for x :: ‹'b list› and a :: ‹'a array› and b :: ‹uint32› and l' :: ‹'a list› and aa :: ‹Heap.heap› and ba :: ‹nat set›
by (metis less_one list_rel_pres_neq_nil nat_of_uint32_012(3) nat_of_uint32_less_iff
nat_of_uint32_notle_minus take_eq_Nil that)
show ?thesis
by sepref_to_hoare
(sep_auto simp: arl32_butlast_nonresizing_def arl32_assn_def hr_comp_def
is_array_list32_def butlast_take list_rel_imp_same_length nat_of_uint32_ge_minus
dest:
list_rel_butlast[of ‹take _ _›]
simp flip: nat_of_uint32_le_iff)
qed
end