summaryrefslogtreecommitdiff
path: root/tests/coq/hashmap_on_disk
diff options
context:
space:
mode:
authorSon Ho2023-10-27 13:34:03 +0200
committerSon Ho2023-10-27 13:34:03 +0200
commit49117ba254679f98938223711810191c3f7d788f (patch)
tree2a29b2be5becbb90104c552b77b4f9d3b3fa2546 /tests/coq/hashmap_on_disk
parent1c55079dac0c51e45f7d77c67bfdbdb8c52ed192 (diff)
Regenerate the Coq test files
Diffstat (limited to 'tests/coq/hashmap_on_disk')
-rw-r--r--tests/coq/hashmap_on_disk/HashmapMain_Funs.v604
-rw-r--r--tests/coq/hashmap_on_disk/HashmapMain_Opaque.v8
-rw-r--r--tests/coq/hashmap_on_disk/HashmapMain_Types.v32
-rw-r--r--tests/coq/hashmap_on_disk/Primitives.v405
4 files changed, 675 insertions, 374 deletions
diff --git a/tests/coq/hashmap_on_disk/HashmapMain_Funs.v b/tests/coq/hashmap_on_disk/HashmapMain_Funs.v
index a85adbf2..eac78186 100644
--- a/tests/coq/hashmap_on_disk/HashmapMain_Funs.v
+++ b/tests/coq/hashmap_on_disk/HashmapMain_Funs.v
@@ -13,652 +13,668 @@ Import HashmapMain_Opaque.
Module HashmapMain_Funs.
(** [hashmap_main::hashmap::hash_key]: forward function *)
-Definition hashmap_hash_key_fwd (k : usize) : result usize :=
+Definition hashmap_hash_key (k : usize) : result usize :=
Return k.
(** [hashmap_main::hashmap::HashMap::{0}::allocate_slots]: loop 0: forward function *)
-Fixpoint hashmap_hash_map_allocate_slots_loop_fwd
- (T : Type) (n : nat) (slots : vec (Hashmap_list_t T)) (n0 : usize) :
- result (vec (Hashmap_list_t T))
+Fixpoint hashmap_HashMap_allocate_slots_loop
+ (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (n0 : usize)
+ :
+ result (alloc_vec_Vec (hashmap_List_t T))
:=
match n with
| O => Fail_ OutOfFuel
| S n1 =>
if n0 s> 0%usize
then (
- slots0 <- vec_push_back (Hashmap_list_t T) slots HashmapListNil;
+ slots0 <- alloc_vec_Vec_push (hashmap_List_t T) slots Hashmap_List_Nil;
n2 <- usize_sub n0 1%usize;
- hashmap_hash_map_allocate_slots_loop_fwd T n1 slots0 n2)
+ hashmap_HashMap_allocate_slots_loop T n1 slots0 n2)
else Return slots
end
.
(** [hashmap_main::hashmap::HashMap::{0}::allocate_slots]: forward function *)
-Definition hashmap_hash_map_allocate_slots_fwd
- (T : Type) (n : nat) (slots : vec (Hashmap_list_t T)) (n0 : usize) :
- result (vec (Hashmap_list_t T))
+Definition hashmap_HashMap_allocate_slots
+ (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (n0 : usize)
+ :
+ result (alloc_vec_Vec (hashmap_List_t T))
:=
- hashmap_hash_map_allocate_slots_loop_fwd T n slots n0
+ hashmap_HashMap_allocate_slots_loop T n slots n0
.
(** [hashmap_main::hashmap::HashMap::{0}::new_with_capacity]: forward function *)
-Definition hashmap_hash_map_new_with_capacity_fwd
+Definition hashmap_HashMap_new_with_capacity
(T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize)
(max_load_divisor : usize) :
- result (Hashmap_hash_map_t T)
+ result (hashmap_HashMap_t T)
:=
- let v := vec_new (Hashmap_list_t T) in
- slots <- hashmap_hash_map_allocate_slots_fwd T n v capacity;
+ let v := alloc_vec_Vec_new (hashmap_List_t T) in
+ slots <- hashmap_HashMap_allocate_slots T n v capacity;
i <- usize_mul capacity max_load_dividend;
i0 <- usize_div i max_load_divisor;
Return
{|
- Hashmap_hash_map_num_entries := 0%usize;
- Hashmap_hash_map_max_load_factor := (max_load_dividend, max_load_divisor);
- Hashmap_hash_map_max_load := i0;
- Hashmap_hash_map_slots := slots
+ hashmap_HashMap_num_entries := 0%usize;
+ hashmap_HashMap_max_load_factor := (max_load_dividend, max_load_divisor);
+ hashmap_HashMap_max_load := i0;
+ hashmap_HashMap_slots := slots
|}
.
(** [hashmap_main::hashmap::HashMap::{0}::new]: forward function *)
-Definition hashmap_hash_map_new_fwd
- (T : Type) (n : nat) : result (Hashmap_hash_map_t T) :=
- hashmap_hash_map_new_with_capacity_fwd T n 32%usize 4%usize 5%usize
+Definition hashmap_HashMap_new
+ (T : Type) (n : nat) : result (hashmap_HashMap_t T) :=
+ hashmap_HashMap_new_with_capacity T n 32%usize 4%usize 5%usize
.
(** [hashmap_main::hashmap::HashMap::{0}::clear]: loop 0: merged forward/backward function
(there is a single backward function, and the forward function returns ()) *)
-Fixpoint hashmap_hash_map_clear_loop_fwd_back
- (T : Type) (n : nat) (slots : vec (Hashmap_list_t T)) (i : usize) :
- result (vec (Hashmap_list_t T))
+Fixpoint hashmap_HashMap_clear_loop
+ (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) :
+ result (alloc_vec_Vec (hashmap_List_t T))
:=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
- let i0 := vec_len (Hashmap_list_t T) slots in
+ let i0 := alloc_vec_Vec_len (hashmap_List_t T) slots in
if i s< i0
then (
i1 <- usize_add i 1%usize;
- slots0 <- vec_index_mut_back (Hashmap_list_t T) slots i HashmapListNil;
- hashmap_hash_map_clear_loop_fwd_back T n0 slots0 i1)
+ slots0 <-
+ alloc_vec_Vec_index_mut_back (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t
+ T)) slots i Hashmap_List_Nil;
+ hashmap_HashMap_clear_loop T n0 slots0 i1)
else Return slots
end
.
(** [hashmap_main::hashmap::HashMap::{0}::clear]: merged forward/backward function
(there is a single backward function, and the forward function returns ()) *)
-Definition hashmap_hash_map_clear_fwd_back
- (T : Type) (n : nat) (self : Hashmap_hash_map_t T) :
- result (Hashmap_hash_map_t T)
+Definition hashmap_HashMap_clear
+ (T : Type) (n : nat) (self : hashmap_HashMap_t T) :
+ result (hashmap_HashMap_t T)
:=
- v <-
- hashmap_hash_map_clear_loop_fwd_back T n self.(Hashmap_hash_map_slots)
- 0%usize;
+ v <- hashmap_HashMap_clear_loop T n self.(hashmap_HashMap_slots) 0%usize;
Return
{|
- Hashmap_hash_map_num_entries := 0%usize;
- Hashmap_hash_map_max_load_factor :=
- self.(Hashmap_hash_map_max_load_factor);
- Hashmap_hash_map_max_load := self.(Hashmap_hash_map_max_load);
- Hashmap_hash_map_slots := v
+ hashmap_HashMap_num_entries := 0%usize;
+ hashmap_HashMap_max_load_factor := self.(hashmap_HashMap_max_load_factor);
+ hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load);
+ hashmap_HashMap_slots := v
|}
.
(** [hashmap_main::hashmap::HashMap::{0}::len]: forward function *)
-Definition hashmap_hash_map_len_fwd
- (T : Type) (self : Hashmap_hash_map_t T) : result usize :=
- Return self.(Hashmap_hash_map_num_entries)
+Definition hashmap_HashMap_len
+ (T : Type) (self : hashmap_HashMap_t T) : result usize :=
+ Return self.(hashmap_HashMap_num_entries)
.
(** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: loop 0: forward function *)
-Fixpoint hashmap_hash_map_insert_in_list_loop_fwd
- (T : Type) (n : nat) (key : usize) (value : T) (ls : Hashmap_list_t T) :
+Fixpoint hashmap_HashMap_insert_in_list_loop
+ (T : Type) (n : nat) (key : usize) (value : T) (ls : hashmap_List_t T) :
result bool
:=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | HashmapListCons ckey cvalue tl =>
+ | Hashmap_List_Cons ckey cvalue tl =>
if ckey s= key
then Return false
- else hashmap_hash_map_insert_in_list_loop_fwd T n0 key value tl
- | HashmapListNil => Return true
+ else hashmap_HashMap_insert_in_list_loop T n0 key value tl
+ | Hashmap_List_Nil => Return true
end
end
.
(** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: forward function *)
-Definition hashmap_hash_map_insert_in_list_fwd
- (T : Type) (n : nat) (key : usize) (value : T) (ls : Hashmap_list_t T) :
+Definition hashmap_HashMap_insert_in_list
+ (T : Type) (n : nat) (key : usize) (value : T) (ls : hashmap_List_t T) :
result bool
:=
- hashmap_hash_map_insert_in_list_loop_fwd T n key value ls
+ hashmap_HashMap_insert_in_list_loop T n key value ls
.
(** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: loop 0: backward function 0 *)
-Fixpoint hashmap_hash_map_insert_in_list_loop_back
- (T : Type) (n : nat) (key : usize) (value : T) (ls : Hashmap_list_t T) :
- result (Hashmap_list_t T)
+Fixpoint hashmap_HashMap_insert_in_list_loop_back
+ (T : Type) (n : nat) (key : usize) (value : T) (ls : hashmap_List_t T) :
+ result (hashmap_List_t T)
:=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | HashmapListCons ckey cvalue tl =>
+ | Hashmap_List_Cons ckey cvalue tl =>
if ckey s= key
- then Return (HashmapListCons ckey value tl)
+ then Return (Hashmap_List_Cons ckey value tl)
else (
- tl0 <- hashmap_hash_map_insert_in_list_loop_back T n0 key value tl;
- Return (HashmapListCons ckey cvalue tl0))
- | HashmapListNil =>
- let l := HashmapListNil in Return (HashmapListCons key value l)
+ tl0 <- hashmap_HashMap_insert_in_list_loop_back T n0 key value tl;
+ Return (Hashmap_List_Cons ckey cvalue tl0))
+ | Hashmap_List_Nil =>
+ let l := Hashmap_List_Nil in Return (Hashmap_List_Cons key value l)
end
end
.
(** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: backward function 0 *)
-Definition hashmap_hash_map_insert_in_list_back
- (T : Type) (n : nat) (key : usize) (value : T) (ls : Hashmap_list_t T) :
- result (Hashmap_list_t T)
+Definition hashmap_HashMap_insert_in_list_back
+ (T : Type) (n : nat) (key : usize) (value : T) (ls : hashmap_List_t T) :
+ result (hashmap_List_t T)
:=
- hashmap_hash_map_insert_in_list_loop_back T n key value ls
+ hashmap_HashMap_insert_in_list_loop_back T n key value ls
.
(** [hashmap_main::hashmap::HashMap::{0}::insert_no_resize]: merged forward/backward function
(there is a single backward function, and the forward function returns ()) *)
-Definition hashmap_hash_map_insert_no_resize_fwd_back
- (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) (value : T)
- :
- result (Hashmap_hash_map_t T)
+Definition hashmap_HashMap_insert_no_resize
+ (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) (value : T) :
+ result (hashmap_HashMap_t T)
:=
- hash <- hashmap_hash_key_fwd key;
- let i := vec_len (Hashmap_list_t T) self.(Hashmap_hash_map_slots) in
+ hash <- hashmap_hash_key key;
+ let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in
hash_mod <- usize_rem hash i;
l <-
- vec_index_mut_fwd (Hashmap_list_t T) self.(Hashmap_hash_map_slots) hash_mod;
- inserted <- hashmap_hash_map_insert_in_list_fwd T n key value l;
+ alloc_vec_Vec_index_mut (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ self.(hashmap_HashMap_slots) hash_mod;
+ inserted <- hashmap_HashMap_insert_in_list T n key value l;
if inserted
then (
- i0 <- usize_add self.(Hashmap_hash_map_num_entries) 1%usize;
- l0 <- hashmap_hash_map_insert_in_list_back T n key value l;
+ i0 <- usize_add self.(hashmap_HashMap_num_entries) 1%usize;
+ l0 <- hashmap_HashMap_insert_in_list_back T n key value l;
v <-
- vec_index_mut_back (Hashmap_list_t T) self.(Hashmap_hash_map_slots)
- hash_mod l0;
+ alloc_vec_Vec_index_mut_back (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t
+ T)) self.(hashmap_HashMap_slots) hash_mod l0;
Return
{|
- Hashmap_hash_map_num_entries := i0;
- Hashmap_hash_map_max_load_factor :=
- self.(Hashmap_hash_map_max_load_factor);
- Hashmap_hash_map_max_load := self.(Hashmap_hash_map_max_load);
- Hashmap_hash_map_slots := v
+ hashmap_HashMap_num_entries := i0;
+ hashmap_HashMap_max_load_factor :=
+ self.(hashmap_HashMap_max_load_factor);
+ hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load);
+ hashmap_HashMap_slots := v
|})
else (
- l0 <- hashmap_hash_map_insert_in_list_back T n key value l;
+ l0 <- hashmap_HashMap_insert_in_list_back T n key value l;
v <-
- vec_index_mut_back (Hashmap_list_t T) self.(Hashmap_hash_map_slots)
- hash_mod l0;
+ alloc_vec_Vec_index_mut_back (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t
+ T)) self.(hashmap_HashMap_slots) hash_mod l0;
Return
{|
- Hashmap_hash_map_num_entries := self.(Hashmap_hash_map_num_entries);
- Hashmap_hash_map_max_load_factor :=
- self.(Hashmap_hash_map_max_load_factor);
- Hashmap_hash_map_max_load := self.(Hashmap_hash_map_max_load);
- Hashmap_hash_map_slots := v
+ hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries);
+ hashmap_HashMap_max_load_factor :=
+ self.(hashmap_HashMap_max_load_factor);
+ hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load);
+ hashmap_HashMap_slots := v
|})
.
(** [hashmap_main::hashmap::HashMap::{0}::move_elements_from_list]: loop 0: merged forward/backward function
(there is a single backward function, and the forward function returns ()) *)
-Fixpoint hashmap_hash_map_move_elements_from_list_loop_fwd_back
- (T : Type) (n : nat) (ntable : Hashmap_hash_map_t T) (ls : Hashmap_list_t T)
- :
- result (Hashmap_hash_map_t T)
+Fixpoint hashmap_HashMap_move_elements_from_list_loop
+ (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) (ls : hashmap_List_t T) :
+ result (hashmap_HashMap_t T)
:=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | HashmapListCons k v tl =>
- ntable0 <- hashmap_hash_map_insert_no_resize_fwd_back T n0 ntable k v;
- hashmap_hash_map_move_elements_from_list_loop_fwd_back T n0 ntable0 tl
- | HashmapListNil => Return ntable
+ | Hashmap_List_Cons k v tl =>
+ ntable0 <- hashmap_HashMap_insert_no_resize T n0 ntable k v;
+ hashmap_HashMap_move_elements_from_list_loop T n0 ntable0 tl
+ | Hashmap_List_Nil => Return ntable
end
end
.
(** [hashmap_main::hashmap::HashMap::{0}::move_elements_from_list]: merged forward/backward function
(there is a single backward function, and the forward function returns ()) *)
-Definition hashmap_hash_map_move_elements_from_list_fwd_back
- (T : Type) (n : nat) (ntable : Hashmap_hash_map_t T) (ls : Hashmap_list_t T)
- :
- result (Hashmap_hash_map_t T)
+Definition hashmap_HashMap_move_elements_from_list
+ (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) (ls : hashmap_List_t T) :
+ result (hashmap_HashMap_t T)
:=
- hashmap_hash_map_move_elements_from_list_loop_fwd_back T n ntable ls
+ hashmap_HashMap_move_elements_from_list_loop T n ntable ls
.
(** [hashmap_main::hashmap::HashMap::{0}::move_elements]: loop 0: merged forward/backward function
(there is a single backward function, and the forward function returns ()) *)
-Fixpoint hashmap_hash_map_move_elements_loop_fwd_back
- (T : Type) (n : nat) (ntable : Hashmap_hash_map_t T)
- (slots : vec (Hashmap_list_t T)) (i : usize) :
- result ((Hashmap_hash_map_t T) * (vec (Hashmap_list_t T)))
+Fixpoint hashmap_HashMap_move_elements_loop
+ (T : Type) (n : nat) (ntable : hashmap_HashMap_t T)
+ (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) :
+ result ((hashmap_HashMap_t T) * (alloc_vec_Vec (hashmap_List_t T)))
:=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
- let i0 := vec_len (Hashmap_list_t T) slots in
+ let i0 := alloc_vec_Vec_len (hashmap_List_t T) slots in
if i s< i0
then (
- l <- vec_index_mut_fwd (Hashmap_list_t T) slots i;
- let ls := mem_replace_fwd (Hashmap_list_t T) l HashmapListNil in
- ntable0 <-
- hashmap_hash_map_move_elements_from_list_fwd_back T n0 ntable ls;
+ l <-
+ alloc_vec_Vec_index_mut (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t
+ T)) slots i;
+ let ls := core_mem_replace (hashmap_List_t T) l Hashmap_List_Nil in
+ ntable0 <- hashmap_HashMap_move_elements_from_list T n0 ntable ls;
i1 <- usize_add i 1%usize;
- let l0 := mem_replace_back (Hashmap_list_t T) l HashmapListNil in
- slots0 <- vec_index_mut_back (Hashmap_list_t T) slots i l0;
- hashmap_hash_map_move_elements_loop_fwd_back T n0 ntable0 slots0 i1)
+ let l0 := core_mem_replace_back (hashmap_List_t T) l Hashmap_List_Nil in
+ slots0 <-
+ alloc_vec_Vec_index_mut_back (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t
+ T)) slots i l0;
+ hashmap_HashMap_move_elements_loop T n0 ntable0 slots0 i1)
else Return (ntable, slots)
end
.
(** [hashmap_main::hashmap::HashMap::{0}::move_elements]: merged forward/backward function
(there is a single backward function, and the forward function returns ()) *)
-Definition hashmap_hash_map_move_elements_fwd_back
- (T : Type) (n : nat) (ntable : Hashmap_hash_map_t T)
- (slots : vec (Hashmap_list_t T)) (i : usize) :
- result ((Hashmap_hash_map_t T) * (vec (Hashmap_list_t T)))
+Definition hashmap_HashMap_move_elements
+ (T : Type) (n : nat) (ntable : hashmap_HashMap_t T)
+ (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) :
+ result ((hashmap_HashMap_t T) * (alloc_vec_Vec (hashmap_List_t T)))
:=
- hashmap_hash_map_move_elements_loop_fwd_back T n ntable slots i
+ hashmap_HashMap_move_elements_loop T n ntable slots i
.
(** [hashmap_main::hashmap::HashMap::{0}::try_resize]: merged forward/backward function
(there is a single backward function, and the forward function returns ()) *)
-Definition hashmap_hash_map_try_resize_fwd_back
- (T : Type) (n : nat) (self : Hashmap_hash_map_t T) :
- result (Hashmap_hash_map_t T)
+Definition hashmap_HashMap_try_resize
+ (T : Type) (n : nat) (self : hashmap_HashMap_t T) :
+ result (hashmap_HashMap_t T)
:=
max_usize <- scalar_cast U32 Usize core_u32_max;
- let capacity := vec_len (Hashmap_list_t T) self.(Hashmap_hash_map_slots) in
+ let capacity :=
+ alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in
n1 <- usize_div max_usize 2%usize;
- let (i, i0) := self.(Hashmap_hash_map_max_load_factor) in
+ let (i, i0) := self.(hashmap_HashMap_max_load_factor) in
i1 <- usize_div n1 i;
if capacity s<= i1
then (
i2 <- usize_mul capacity 2%usize;
- ntable <- hashmap_hash_map_new_with_capacity_fwd T n i2 i i0;
+ ntable <- hashmap_HashMap_new_with_capacity T n i2 i i0;
p <-
- hashmap_hash_map_move_elements_fwd_back T n ntable
- self.(Hashmap_hash_map_slots) 0%usize;
+ hashmap_HashMap_move_elements T n ntable self.(hashmap_HashMap_slots)
+ 0%usize;
let (ntable0, _) := p in
Return
{|
- Hashmap_hash_map_num_entries := self.(Hashmap_hash_map_num_entries);
- Hashmap_hash_map_max_load_factor := (i, i0);
- Hashmap_hash_map_max_load := ntable0.(Hashmap_hash_map_max_load);
- Hashmap_hash_map_slots := ntable0.(Hashmap_hash_map_slots)
+ hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries);
+ hashmap_HashMap_max_load_factor := (i, i0);
+ hashmap_HashMap_max_load := ntable0.(hashmap_HashMap_max_load);
+ hashmap_HashMap_slots := ntable0.(hashmap_HashMap_slots)
|})
else
Return
{|
- Hashmap_hash_map_num_entries := self.(Hashmap_hash_map_num_entries);
- Hashmap_hash_map_max_load_factor := (i, i0);
- Hashmap_hash_map_max_load := self.(Hashmap_hash_map_max_load);
- Hashmap_hash_map_slots := self.(Hashmap_hash_map_slots)
+ hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries);
+ hashmap_HashMap_max_load_factor := (i, i0);
+ hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load);
+ hashmap_HashMap_slots := self.(hashmap_HashMap_slots)
|}
.
(** [hashmap_main::hashmap::HashMap::{0}::insert]: merged forward/backward function
(there is a single backward function, and the forward function returns ()) *)
-Definition hashmap_hash_map_insert_fwd_back
- (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) (value : T)
- :
- result (Hashmap_hash_map_t T)
+Definition hashmap_HashMap_insert
+ (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) (value : T) :
+ result (hashmap_HashMap_t T)
:=
- self0 <- hashmap_hash_map_insert_no_resize_fwd_back T n self key value;
- i <- hashmap_hash_map_len_fwd T self0;
- if i s> self0.(Hashmap_hash_map_max_load)
- then hashmap_hash_map_try_resize_fwd_back T n self0
+ self0 <- hashmap_HashMap_insert_no_resize T n self key value;
+ i <- hashmap_HashMap_len T self0;
+ if i s> self0.(hashmap_HashMap_max_load)
+ then hashmap_HashMap_try_resize T n self0
else Return self0
.
(** [hashmap_main::hashmap::HashMap::{0}::contains_key_in_list]: loop 0: forward function *)
-Fixpoint hashmap_hash_map_contains_key_in_list_loop_fwd
- (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) : result bool :=
+Fixpoint hashmap_HashMap_contains_key_in_list_loop
+ (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result bool :=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | HashmapListCons ckey t tl =>
+ | Hashmap_List_Cons ckey t tl =>
if ckey s= key
then Return true
- else hashmap_hash_map_contains_key_in_list_loop_fwd T n0 key tl
- | HashmapListNil => Return false
+ else hashmap_HashMap_contains_key_in_list_loop T n0 key tl
+ | Hashmap_List_Nil => Return false
end
end
.
(** [hashmap_main::hashmap::HashMap::{0}::contains_key_in_list]: forward function *)
-Definition hashmap_hash_map_contains_key_in_list_fwd
- (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) : result bool :=
- hashmap_hash_map_contains_key_in_list_loop_fwd T n key ls
+Definition hashmap_HashMap_contains_key_in_list
+ (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result bool :=
+ hashmap_HashMap_contains_key_in_list_loop T n key ls
.
(** [hashmap_main::hashmap::HashMap::{0}::contains_key]: forward function *)
-Definition hashmap_hash_map_contains_key_fwd
- (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) :
+Definition hashmap_HashMap_contains_key
+ (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) :
result bool
:=
- hash <- hashmap_hash_key_fwd key;
- let i := vec_len (Hashmap_list_t T) self.(Hashmap_hash_map_slots) in
+ hash <- hashmap_hash_key key;
+ let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in
hash_mod <- usize_rem hash i;
- l <- vec_index_fwd (Hashmap_list_t T) self.(Hashmap_hash_map_slots) hash_mod;
- hashmap_hash_map_contains_key_in_list_fwd T n key l
+ l <-
+ alloc_vec_Vec_index (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ self.(hashmap_HashMap_slots) hash_mod;
+ hashmap_HashMap_contains_key_in_list T n key l
.
(** [hashmap_main::hashmap::HashMap::{0}::get_in_list]: loop 0: forward function *)
-Fixpoint hashmap_hash_map_get_in_list_loop_fwd
- (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) : result T :=
+Fixpoint hashmap_HashMap_get_in_list_loop
+ (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result T :=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | HashmapListCons ckey cvalue tl =>
+ | Hashmap_List_Cons ckey cvalue tl =>
if ckey s= key
then Return cvalue
- else hashmap_hash_map_get_in_list_loop_fwd T n0 key tl
- | HashmapListNil => Fail_ Failure
+ else hashmap_HashMap_get_in_list_loop T n0 key tl
+ | Hashmap_List_Nil => Fail_ Failure
end
end
.
(** [hashmap_main::hashmap::HashMap::{0}::get_in_list]: forward function *)
-Definition hashmap_hash_map_get_in_list_fwd
- (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) : result T :=
- hashmap_hash_map_get_in_list_loop_fwd T n key ls
+Definition hashmap_HashMap_get_in_list
+ (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result T :=
+ hashmap_HashMap_get_in_list_loop T n key ls
.
(** [hashmap_main::hashmap::HashMap::{0}::get]: forward function *)
-Definition hashmap_hash_map_get_fwd
- (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) :
- result T
- :=
- hash <- hashmap_hash_key_fwd key;
- let i := vec_len (Hashmap_list_t T) self.(Hashmap_hash_map_slots) in
+Definition hashmap_HashMap_get
+ (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : result T :=
+ hash <- hashmap_hash_key key;
+ let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in
hash_mod <- usize_rem hash i;
- l <- vec_index_fwd (Hashmap_list_t T) self.(Hashmap_hash_map_slots) hash_mod;
- hashmap_hash_map_get_in_list_fwd T n key l
+ l <-
+ alloc_vec_Vec_index (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ self.(hashmap_HashMap_slots) hash_mod;
+ hashmap_HashMap_get_in_list T n key l
.
(** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: loop 0: forward function *)
-Fixpoint hashmap_hash_map_get_mut_in_list_loop_fwd
- (T : Type) (n : nat) (ls : Hashmap_list_t T) (key : usize) : result T :=
+Fixpoint hashmap_HashMap_get_mut_in_list_loop
+ (T : Type) (n : nat) (ls : hashmap_List_t T) (key : usize) : result T :=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | HashmapListCons ckey cvalue tl =>
+ | Hashmap_List_Cons ckey cvalue tl =>
if ckey s= key
then Return cvalue
- else hashmap_hash_map_get_mut_in_list_loop_fwd T n0 tl key
- | HashmapListNil => Fail_ Failure
+ else hashmap_HashMap_get_mut_in_list_loop T n0 tl key
+ | Hashmap_List_Nil => Fail_ Failure
end
end
.
(** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: forward function *)
-Definition hashmap_hash_map_get_mut_in_list_fwd
- (T : Type) (n : nat) (ls : Hashmap_list_t T) (key : usize) : result T :=
- hashmap_hash_map_get_mut_in_list_loop_fwd T n ls key
+Definition hashmap_HashMap_get_mut_in_list
+ (T : Type) (n : nat) (ls : hashmap_List_t T) (key : usize) : result T :=
+ hashmap_HashMap_get_mut_in_list_loop T n ls key
.
(** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: loop 0: backward function 0 *)
-Fixpoint hashmap_hash_map_get_mut_in_list_loop_back
- (T : Type) (n : nat) (ls : Hashmap_list_t T) (key : usize) (ret : T) :
- result (Hashmap_list_t T)
+Fixpoint hashmap_HashMap_get_mut_in_list_loop_back
+ (T : Type) (n : nat) (ls : hashmap_List_t T) (key : usize) (ret : T) :
+ result (hashmap_List_t T)
:=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | HashmapListCons ckey cvalue tl =>
+ | Hashmap_List_Cons ckey cvalue tl =>
if ckey s= key
- then Return (HashmapListCons ckey ret tl)
+ then Return (Hashmap_List_Cons ckey ret tl)
else (
- tl0 <- hashmap_hash_map_get_mut_in_list_loop_back T n0 tl key ret;
- Return (HashmapListCons ckey cvalue tl0))
- | HashmapListNil => Fail_ Failure
+ tl0 <- hashmap_HashMap_get_mut_in_list_loop_back T n0 tl key ret;
+ Return (Hashmap_List_Cons ckey cvalue tl0))
+ | Hashmap_List_Nil => Fail_ Failure
end
end
.
(** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: backward function 0 *)
-Definition hashmap_hash_map_get_mut_in_list_back
- (T : Type) (n : nat) (ls : Hashmap_list_t T) (key : usize) (ret : T) :
- result (Hashmap_list_t T)
+Definition hashmap_HashMap_get_mut_in_list_back
+ (T : Type) (n : nat) (ls : hashmap_List_t T) (key : usize) (ret : T) :
+ result (hashmap_List_t T)
:=
- hashmap_hash_map_get_mut_in_list_loop_back T n ls key ret
+ hashmap_HashMap_get_mut_in_list_loop_back T n ls key ret
.
(** [hashmap_main::hashmap::HashMap::{0}::get_mut]: forward function *)
-Definition hashmap_hash_map_get_mut_fwd
- (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) :
- result T
- :=
- hash <- hashmap_hash_key_fwd key;
- let i := vec_len (Hashmap_list_t T) self.(Hashmap_hash_map_slots) in
+Definition hashmap_HashMap_get_mut
+ (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : result T :=
+ hash <- hashmap_hash_key key;
+ let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in
hash_mod <- usize_rem hash i;
l <-
- vec_index_mut_fwd (Hashmap_list_t T) self.(Hashmap_hash_map_slots) hash_mod;
- hashmap_hash_map_get_mut_in_list_fwd T n l key
+ alloc_vec_Vec_index_mut (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ self.(hashmap_HashMap_slots) hash_mod;
+ hashmap_HashMap_get_mut_in_list T n l key
.
(** [hashmap_main::hashmap::HashMap::{0}::get_mut]: backward function 0 *)
-Definition hashmap_hash_map_get_mut_back
- (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) (ret : T) :
- result (Hashmap_hash_map_t T)
+Definition hashmap_HashMap_get_mut_back
+ (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) (ret : T) :
+ result (hashmap_HashMap_t T)
:=
- hash <- hashmap_hash_key_fwd key;
- let i := vec_len (Hashmap_list_t T) self.(Hashmap_hash_map_slots) in
+ hash <- hashmap_hash_key key;
+ let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in
hash_mod <- usize_rem hash i;
l <-
- vec_index_mut_fwd (Hashmap_list_t T) self.(Hashmap_hash_map_slots) hash_mod;
- l0 <- hashmap_hash_map_get_mut_in_list_back T n l key ret;
+ alloc_vec_Vec_index_mut (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ self.(hashmap_HashMap_slots) hash_mod;
+ l0 <- hashmap_HashMap_get_mut_in_list_back T n l key ret;
v <-
- vec_index_mut_back (Hashmap_list_t T) self.(Hashmap_hash_map_slots)
- hash_mod l0;
+ alloc_vec_Vec_index_mut_back (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ self.(hashmap_HashMap_slots) hash_mod l0;
Return
{|
- Hashmap_hash_map_num_entries := self.(Hashmap_hash_map_num_entries);
- Hashmap_hash_map_max_load_factor :=
- self.(Hashmap_hash_map_max_load_factor);
- Hashmap_hash_map_max_load := self.(Hashmap_hash_map_max_load);
- Hashmap_hash_map_slots := v
+ hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries);
+ hashmap_HashMap_max_load_factor := self.(hashmap_HashMap_max_load_factor);
+ hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load);
+ hashmap_HashMap_slots := v
|}
.
(** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: loop 0: forward function *)
-Fixpoint hashmap_hash_map_remove_from_list_loop_fwd
- (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) :
+Fixpoint hashmap_HashMap_remove_from_list_loop
+ (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) :
result (option T)
:=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | HashmapListCons ckey t tl =>
+ | Hashmap_List_Cons ckey t tl =>
if ckey s= key
then
let mv_ls :=
- mem_replace_fwd (Hashmap_list_t T) (HashmapListCons ckey t tl)
- HashmapListNil in
+ core_mem_replace (hashmap_List_t T) (Hashmap_List_Cons ckey t tl)
+ Hashmap_List_Nil in
match mv_ls with
- | HashmapListCons i cvalue tl0 => Return (Some cvalue)
- | HashmapListNil => Fail_ Failure
+ | Hashmap_List_Cons i cvalue tl0 => Return (Some cvalue)
+ | Hashmap_List_Nil => Fail_ Failure
end
- else hashmap_hash_map_remove_from_list_loop_fwd T n0 key tl
- | HashmapListNil => Return None
+ else hashmap_HashMap_remove_from_list_loop T n0 key tl
+ | Hashmap_List_Nil => Return None
end
end
.
(** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: forward function *)
-Definition hashmap_hash_map_remove_from_list_fwd
- (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) :
+Definition hashmap_HashMap_remove_from_list
+ (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) :
result (option T)
:=
- hashmap_hash_map_remove_from_list_loop_fwd T n key ls
+ hashmap_HashMap_remove_from_list_loop T n key ls
.
(** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: loop 0: backward function 1 *)
-Fixpoint hashmap_hash_map_remove_from_list_loop_back
- (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) :
- result (Hashmap_list_t T)
+Fixpoint hashmap_HashMap_remove_from_list_loop_back
+ (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) :
+ result (hashmap_List_t T)
:=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | HashmapListCons ckey t tl =>
+ | Hashmap_List_Cons ckey t tl =>
if ckey s= key
then
let mv_ls :=
- mem_replace_fwd (Hashmap_list_t T) (HashmapListCons ckey t tl)
- HashmapListNil in
+ core_mem_replace (hashmap_List_t T) (Hashmap_List_Cons ckey t tl)
+ Hashmap_List_Nil in
match mv_ls with
- | HashmapListCons i cvalue tl0 => Return tl0
- | HashmapListNil => Fail_ Failure
+ | Hashmap_List_Cons i cvalue tl0 => Return tl0
+ | Hashmap_List_Nil => Fail_ Failure
end
else (
- tl0 <- hashmap_hash_map_remove_from_list_loop_back T n0 key tl;
- Return (HashmapListCons ckey t tl0))
- | HashmapListNil => Return HashmapListNil
+ tl0 <- hashmap_HashMap_remove_from_list_loop_back T n0 key tl;
+ Return (Hashmap_List_Cons ckey t tl0))
+ | Hashmap_List_Nil => Return Hashmap_List_Nil
end
end
.
(** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: backward function 1 *)
-Definition hashmap_hash_map_remove_from_list_back
- (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) :
- result (Hashmap_list_t T)
+Definition hashmap_HashMap_remove_from_list_back
+ (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) :
+ result (hashmap_List_t T)
:=
- hashmap_hash_map_remove_from_list_loop_back T n key ls
+ hashmap_HashMap_remove_from_list_loop_back T n key ls
.
(** [hashmap_main::hashmap::HashMap::{0}::remove]: forward function *)
-Definition hashmap_hash_map_remove_fwd
- (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) :
+Definition hashmap_HashMap_remove
+ (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) :
result (option T)
:=
- hash <- hashmap_hash_key_fwd key;
- let i := vec_len (Hashmap_list_t T) self.(Hashmap_hash_map_slots) in
+ hash <- hashmap_hash_key key;
+ let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in
hash_mod <- usize_rem hash i;
l <-
- vec_index_mut_fwd (Hashmap_list_t T) self.(Hashmap_hash_map_slots) hash_mod;
- x <- hashmap_hash_map_remove_from_list_fwd T n key l;
+ alloc_vec_Vec_index_mut (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ self.(hashmap_HashMap_slots) hash_mod;
+ x <- hashmap_HashMap_remove_from_list T n key l;
match x with
| None => Return None
| Some x0 =>
- _ <- usize_sub self.(Hashmap_hash_map_num_entries) 1%usize;
- Return (Some x0)
+ _ <- usize_sub self.(hashmap_HashMap_num_entries) 1%usize; Return (Some x0)
end
.
(** [hashmap_main::hashmap::HashMap::{0}::remove]: backward function 0 *)
-Definition hashmap_hash_map_remove_back
- (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) :
- result (Hashmap_hash_map_t T)
+Definition hashmap_HashMap_remove_back
+ (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) :
+ result (hashmap_HashMap_t T)
:=
- hash <- hashmap_hash_key_fwd key;
- let i := vec_len (Hashmap_list_t T) self.(Hashmap_hash_map_slots) in
+ hash <- hashmap_hash_key key;
+ let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in
hash_mod <- usize_rem hash i;
l <-
- vec_index_mut_fwd (Hashmap_list_t T) self.(Hashmap_hash_map_slots) hash_mod;
- x <- hashmap_hash_map_remove_from_list_fwd T n key l;
+ alloc_vec_Vec_index_mut (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ self.(hashmap_HashMap_slots) hash_mod;
+ x <- hashmap_HashMap_remove_from_list T n key l;
match x with
| None =>
- l0 <- hashmap_hash_map_remove_from_list_back T n key l;
+ l0 <- hashmap_HashMap_remove_from_list_back T n key l;
v <-
- vec_index_mut_back (Hashmap_list_t T) self.(Hashmap_hash_map_slots)
- hash_mod l0;
+ alloc_vec_Vec_index_mut_back (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t
+ T)) self.(hashmap_HashMap_slots) hash_mod l0;
Return
{|
- Hashmap_hash_map_num_entries := self.(Hashmap_hash_map_num_entries);
- Hashmap_hash_map_max_load_factor :=
- self.(Hashmap_hash_map_max_load_factor);
- Hashmap_hash_map_max_load := self.(Hashmap_hash_map_max_load);
- Hashmap_hash_map_slots := v
+ hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries);
+ hashmap_HashMap_max_load_factor :=
+ self.(hashmap_HashMap_max_load_factor);
+ hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load);
+ hashmap_HashMap_slots := v
|}
| Some x0 =>
- i0 <- usize_sub self.(Hashmap_hash_map_num_entries) 1%usize;
- l0 <- hashmap_hash_map_remove_from_list_back T n key l;
+ i0 <- usize_sub self.(hashmap_HashMap_num_entries) 1%usize;
+ l0 <- hashmap_HashMap_remove_from_list_back T n key l;
v <-
- vec_index_mut_back (Hashmap_list_t T) self.(Hashmap_hash_map_slots)
- hash_mod l0;
+ alloc_vec_Vec_index_mut_back (hashmap_List_t T) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t
+ T)) self.(hashmap_HashMap_slots) hash_mod l0;
Return
{|
- Hashmap_hash_map_num_entries := i0;
- Hashmap_hash_map_max_load_factor :=
- self.(Hashmap_hash_map_max_load_factor);
- Hashmap_hash_map_max_load := self.(Hashmap_hash_map_max_load);
- Hashmap_hash_map_slots := v
+ hashmap_HashMap_num_entries := i0;
+ hashmap_HashMap_max_load_factor :=
+ self.(hashmap_HashMap_max_load_factor);
+ hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load);
+ hashmap_HashMap_slots := v
|}
end
.
(** [hashmap_main::hashmap::test1]: forward function *)
-Definition hashmap_test1_fwd (n : nat) : result unit :=
- hm <- hashmap_hash_map_new_fwd u64 n;
- hm0 <- hashmap_hash_map_insert_fwd_back u64 n hm 0%usize 42%u64;
- hm1 <- hashmap_hash_map_insert_fwd_back u64 n hm0 128%usize 18%u64;
- hm2 <- hashmap_hash_map_insert_fwd_back u64 n hm1 1024%usize 138%u64;
- hm3 <- hashmap_hash_map_insert_fwd_back u64 n hm2 1056%usize 256%u64;
- i <- hashmap_hash_map_get_fwd u64 n hm3 128%usize;
+Definition hashmap_test1 (n : nat) : result unit :=
+ hm <- hashmap_HashMap_new u64 n;
+ hm0 <- hashmap_HashMap_insert u64 n hm 0%usize 42%u64;
+ hm1 <- hashmap_HashMap_insert u64 n hm0 128%usize 18%u64;
+ hm2 <- hashmap_HashMap_insert u64 n hm1 1024%usize 138%u64;
+ hm3 <- hashmap_HashMap_insert u64 n hm2 1056%usize 256%u64;
+ i <- hashmap_HashMap_get u64 n hm3 128%usize;
if negb (i s= 18%u64)
then Fail_ Failure
else (
- hm4 <- hashmap_hash_map_get_mut_back u64 n hm3 1024%usize 56%u64;
- i0 <- hashmap_hash_map_get_fwd u64 n hm4 1024%usize;
+ hm4 <- hashmap_HashMap_get_mut_back u64 n hm3 1024%usize 56%u64;
+ i0 <- hashmap_HashMap_get u64 n hm4 1024%usize;
if negb (i0 s= 56%u64)
then Fail_ Failure
else (
- x <- hashmap_hash_map_remove_fwd u64 n hm4 1024%usize;
+ x <- hashmap_HashMap_remove u64 n hm4 1024%usize;
match x with
| None => Fail_ Failure
| Some x0 =>
if negb (x0 s= 56%u64)
then Fail_ Failure
else (
- hm5 <- hashmap_hash_map_remove_back u64 n hm4 1024%usize;
- i1 <- hashmap_hash_map_get_fwd u64 n hm5 0%usize;
+ hm5 <- hashmap_HashMap_remove_back u64 n hm4 1024%usize;
+ i1 <- hashmap_HashMap_get u64 n hm5 0%usize;
if negb (i1 s= 42%u64)
then Fail_ Failure
else (
- i2 <- hashmap_hash_map_get_fwd u64 n hm5 128%usize;
+ i2 <- hashmap_HashMap_get u64 n hm5 128%usize;
if negb (i2 s= 18%u64)
then Fail_ Failure
else (
- i3 <- hashmap_hash_map_get_fwd u64 n hm5 1056%usize;
+ i3 <- hashmap_HashMap_get u64 n hm5 1056%usize;
if negb (i3 s= 256%u64) then Fail_ Failure else Return tt)))
end))
.
(** [hashmap_main::insert_on_disk]: forward function *)
-Definition insert_on_disk_fwd
+Definition insert_on_disk
(n : nat) (key : usize) (value : u64) (st : state) : result (state * unit) :=
- p <- hashmap_utils_deserialize_fwd st;
+ p <- hashmap_utils_deserialize st;
let (st0, hm) := p in
- hm0 <- hashmap_hash_map_insert_fwd_back u64 n hm key value;
- p0 <- hashmap_utils_serialize_fwd hm0 st0;
+ hm0 <- hashmap_HashMap_insert u64 n hm key value;
+ p0 <- hashmap_utils_serialize hm0 st0;
let (st1, _) := p0 in
Return (st1, tt)
.
(** [hashmap_main::main]: forward function *)
-Definition main_fwd : result unit :=
+Definition main : result unit :=
Return tt.
-(** Unit test for [hashmap_main::main] *)
-Check (main_fwd )%return.
-
End HashmapMain_Funs .
diff --git a/tests/coq/hashmap_on_disk/HashmapMain_Opaque.v b/tests/coq/hashmap_on_disk/HashmapMain_Opaque.v
index 2d17cc29..5e376239 100644
--- a/tests/coq/hashmap_on_disk/HashmapMain_Opaque.v
+++ b/tests/coq/hashmap_on_disk/HashmapMain_Opaque.v
@@ -11,13 +11,13 @@ Import HashmapMain_Types.
Module HashmapMain_Opaque.
(** [hashmap_main::hashmap_utils::deserialize]: forward function *)
-Axiom hashmap_utils_deserialize_fwd
- : state -> result (state * (Hashmap_hash_map_t u64))
+Axiom hashmap_utils_deserialize
+ : state -> result (state * (hashmap_HashMap_t u64))
.
(** [hashmap_main::hashmap_utils::serialize]: forward function *)
-Axiom hashmap_utils_serialize_fwd
- : Hashmap_hash_map_t u64 -> state -> result (state * unit)
+Axiom hashmap_utils_serialize
+ : hashmap_HashMap_t u64 -> state -> result (state * unit)
.
End HashmapMain_Opaque .
diff --git a/tests/coq/hashmap_on_disk/HashmapMain_Types.v b/tests/coq/hashmap_on_disk/HashmapMain_Types.v
index 36aaaf25..466119f8 100644
--- a/tests/coq/hashmap_on_disk/HashmapMain_Types.v
+++ b/tests/coq/hashmap_on_disk/HashmapMain_Types.v
@@ -9,29 +9,29 @@ Local Open Scope Primitives_scope.
Module HashmapMain_Types.
(** [hashmap_main::hashmap::List] *)
-Inductive Hashmap_list_t (T : Type) :=
-| HashmapListCons : usize -> T -> Hashmap_list_t T -> Hashmap_list_t T
-| HashmapListNil : Hashmap_list_t T
+Inductive hashmap_List_t (T : Type) :=
+| Hashmap_List_Cons : usize -> T -> hashmap_List_t T -> hashmap_List_t T
+| Hashmap_List_Nil : hashmap_List_t T
.
-Arguments HashmapListCons {T} _ _ _.
-Arguments HashmapListNil {T}.
+Arguments Hashmap_List_Cons {T} _ _ _.
+Arguments Hashmap_List_Nil {T}.
(** [hashmap_main::hashmap::HashMap] *)
-Record Hashmap_hash_map_t (T : Type) :=
-mkHashmap_hash_map_t {
- Hashmap_hash_map_num_entries : usize;
- Hashmap_hash_map_max_load_factor : (usize * usize);
- Hashmap_hash_map_max_load : usize;
- Hashmap_hash_map_slots : vec (Hashmap_list_t T);
+Record hashmap_HashMap_t (T : Type) :=
+mkhashmap_HashMap_t {
+ hashmap_HashMap_num_entries : usize;
+ hashmap_HashMap_max_load_factor : (usize * usize);
+ hashmap_HashMap_max_load : usize;
+ hashmap_HashMap_slots : alloc_vec_Vec (hashmap_List_t T);
}
.
-Arguments mkHashmap_hash_map_t {T} _ _ _ _.
-Arguments Hashmap_hash_map_num_entries {T}.
-Arguments Hashmap_hash_map_max_load_factor {T}.
-Arguments Hashmap_hash_map_max_load {T}.
-Arguments Hashmap_hash_map_slots {T}.
+Arguments mkhashmap_HashMap_t {T} _ _ _ _.
+Arguments hashmap_HashMap_num_entries {T}.
+Arguments hashmap_HashMap_max_load_factor {T}.
+Arguments hashmap_HashMap_max_load {T}.
+Arguments hashmap_HashMap_slots {T}.
(** The state type used in the state-error monad *)
Axiom state : Type.
diff --git a/tests/coq/hashmap_on_disk/Primitives.v b/tests/coq/hashmap_on_disk/Primitives.v
index 8d6c9c8d..85e38f01 100644
--- a/tests/coq/hashmap_on_disk/Primitives.v
+++ b/tests/coq/hashmap_on_disk/Primitives.v
@@ -63,13 +63,15 @@ Check (if true then Return (1 + 2) else Fail_ Failure)%global = 3.
(*** Misc *)
-
Definition string := Coq.Strings.String.string.
Definition char := Coq.Strings.Ascii.ascii.
Definition char_of_byte := Coq.Strings.Ascii.ascii_of_byte.
-Definition mem_replace_fwd (a : Type) (x : a) (y : a) : a := x .
-Definition mem_replace_back (a : Type) (x : a) (y : a) : a := y .
+Definition core_mem_replace (a : Type) (x : a) (y : a) : a := x .
+Definition core_mem_replace_back (a : Type) (x : a) (y : a) : a := y .
+
+Record mut_raw_ptr (T : Type) := { mut_raw_ptr_v : T }.
+Record const_raw_ptr (T : Type) := { const_raw_ptr_v : T }.
(*** Scalars *)
@@ -408,12 +410,75 @@ Definition core_i64_max := i64_max %i64.
Definition core_i128_max := i64_max %i128.
Axiom core_isize_max : isize. (** TODO *)
-(*** Range *)
-Record range (T : Type) := mk_range {
- start: T;
- end_: T;
+(*** core::ops *)
+
+(* Trait declaration: [core::ops::index::Index] *)
+Record core_ops_index_Index (Self Idx : Type) := mk_core_ops_index_Index {
+ core_ops_index_Index_Output : Type;
+ core_ops_index_Index_index : Self -> Idx -> result core_ops_index_Index_Output;
+}.
+Arguments mk_core_ops_index_Index {_ _}.
+Arguments core_ops_index_Index_Output {_ _}.
+Arguments core_ops_index_Index_index {_ _}.
+
+(* Trait declaration: [core::ops::index::IndexMut] *)
+Record core_ops_index_IndexMut (Self Idx : Type) := mk_core_ops_index_IndexMut {
+ core_ops_index_IndexMut_indexInst : core_ops_index_Index Self Idx;
+ core_ops_index_IndexMut_index_mut : Self -> Idx -> result core_ops_index_IndexMut_indexInst.(core_ops_index_Index_Output);
+ core_ops_index_IndexMut_index_mut_back : Self -> Idx -> core_ops_index_IndexMut_indexInst.(core_ops_index_Index_Output) -> result Self;
+}.
+Arguments mk_core_ops_index_IndexMut {_ _}.
+Arguments core_ops_index_IndexMut_indexInst {_ _}.
+Arguments core_ops_index_IndexMut_index_mut {_ _}.
+Arguments core_ops_index_IndexMut_index_mut_back {_ _}.
+
+(* Trait declaration [core::ops::deref::Deref] *)
+Record core_ops_deref_Deref (Self : Type) := mk_core_ops_deref_Deref {
+ core_ops_deref_Deref_target : Type;
+ core_ops_deref_Deref_deref : Self -> result core_ops_deref_Deref_target;
+}.
+Arguments mk_core_ops_deref_Deref {_}.
+Arguments core_ops_deref_Deref_target {_}.
+Arguments core_ops_deref_Deref_deref {_}.
+
+(* Trait declaration [core::ops::deref::DerefMut] *)
+Record core_ops_deref_DerefMut (Self : Type) := mk_core_ops_deref_DerefMut {
+ core_ops_deref_DerefMut_derefInst : core_ops_deref_Deref Self;
+ core_ops_deref_DerefMut_deref_mut : Self -> result core_ops_deref_DerefMut_derefInst.(core_ops_deref_Deref_target);
+ core_ops_deref_DerefMut_deref_mut_back : Self -> core_ops_deref_DerefMut_derefInst.(core_ops_deref_Deref_target) -> result Self;
}.
-Arguments mk_range {_}.
+Arguments mk_core_ops_deref_DerefMut {_}.
+Arguments core_ops_deref_DerefMut_derefInst {_}.
+Arguments core_ops_deref_DerefMut_deref_mut {_}.
+Arguments core_ops_deref_DerefMut_deref_mut_back {_}.
+
+Record core_ops_range_Range (T : Type) := mk_core_ops_range_Range {
+ core_ops_range_Range_start : T;
+ core_ops_range_Range_end_ : T;
+}.
+Arguments mk_core_ops_range_Range {_}.
+Arguments core_ops_range_Range_start {_}.
+Arguments core_ops_range_Range_end_ {_}.
+
+(*** [alloc] *)
+
+Definition alloc_boxed_Box_deref (T : Type) (x : T) : result T := Return x.
+Definition alloc_boxed_Box_deref_mut (T : Type) (x : T) : result T := Return x.
+Definition alloc_boxed_Box_deref_mut_back (T : Type) (_ : T) (x : T) : result T := Return x.
+
+(* Trait instance *)
+Definition alloc_boxed_Box_coreOpsDerefInst (Self : Type) : core_ops_deref_Deref Self := {|
+ core_ops_deref_Deref_target := Self;
+ core_ops_deref_Deref_deref := alloc_boxed_Box_deref Self;
+|}.
+
+(* Trait instance *)
+Definition alloc_boxed_Box_coreOpsDerefMutInst (Self : Type) : core_ops_deref_DerefMut Self := {|
+ core_ops_deref_DerefMut_derefInst := alloc_boxed_Box_coreOpsDerefInst Self;
+ core_ops_deref_DerefMut_deref_mut := alloc_boxed_Box_deref_mut Self;
+ core_ops_deref_DerefMut_deref_mut_back := alloc_boxed_Box_deref_mut_back Self;
+|}.
+
(*** Arrays *)
Definition array T (n : usize) := { l: list T | Z.of_nat (length l) = to_Z n}.
@@ -433,51 +498,50 @@ Qed.
(* TODO: finish the definitions *)
Axiom mk_array : forall (T : Type) (n : usize) (l : list T), array T n.
-Axiom array_index_shared : forall (T : Type) (n : usize) (x : array T n) (i : usize), result T.
-Axiom array_index_mut_fwd : forall (T : Type) (n : usize) (x : array T n) (i : usize), result T.
-Axiom array_index_mut_back : forall (T : Type) (n : usize) (x : array T n) (i : usize) (nx : T), result (array T n).
+(* For initialization *)
+Axiom array_repeat : forall (T : Type) (n : usize) (x : T), array T n.
+
+Axiom array_index_usize : forall (T : Type) (n : usize) (x : array T n) (i : usize), result T.
+Axiom array_update_usize : forall (T : Type) (n : usize) (x : array T n) (i : usize) (nx : T), result (array T n).
(*** Slice *)
Definition slice T := { l: list T | Z.of_nat (length l) <= usize_max}.
Axiom slice_len : forall (T : Type) (s : slice T), usize.
-Axiom slice_index_shared : forall (T : Type) (x : slice T) (i : usize), result T.
-Axiom slice_index_mut_fwd : forall (T : Type) (x : slice T) (i : usize), result T.
-Axiom slice_index_mut_back : forall (T : Type) (x : slice T) (i : usize) (nx : T), result (slice T).
+Axiom slice_index_usize : forall (T : Type) (x : slice T) (i : usize), result T.
+Axiom slice_update_usize : forall (T : Type) (x : slice T) (i : usize) (nx : T), result (slice T).
(*** Subslices *)
-Axiom array_to_slice_shared : forall (T : Type) (n : usize) (x : array T n), result (slice T).
-Axiom array_to_slice_mut_fwd : forall (T : Type) (n : usize) (x : array T n), result (slice T).
-Axiom array_to_slice_mut_back : forall (T : Type) (n : usize) (x : array T n) (s : slice T), result (array T n).
+Axiom array_to_slice : forall (T : Type) (n : usize) (x : array T n), result (slice T).
+Axiom array_from_slice : forall (T : Type) (n : usize) (x : array T n) (s : slice T), result (array T n).
+
+Axiom array_subslice: forall (T : Type) (n : usize) (x : array T n) (r : core_ops_range_Range usize), result (slice T).
+Axiom array_update_subslice: forall (T : Type) (n : usize) (x : array T n) (r : core_ops_range_Range usize) (ns : slice T), result (array T n).
-Axiom array_subslice_shared: forall (T : Type) (n : usize) (x : array T n) (r : range usize), result (slice T).
-Axiom array_subslice_mut_fwd: forall (T : Type) (n : usize) (x : array T n) (r : range usize), result (slice T).
-Axiom array_subslice_mut_back: forall (T : Type) (n : usize) (x : array T n) (r : range usize) (ns : slice T), result (array T n).
-Axiom slice_subslice_shared: forall (T : Type) (x : slice T) (r : range usize), result (slice T).
-Axiom slice_subslice_mut_fwd: forall (T : Type) (x : slice T) (r : range usize), result (slice T).
-Axiom slice_subslice_mut_back: forall (T : Type) (x : slice T) (r : range usize) (ns : slice T), result (slice T).
+Axiom slice_subslice: forall (T : Type) (x : slice T) (r : core_ops_range_Range usize), result (slice T).
+Axiom slice_update_subslice: forall (T : Type) (x : slice T) (r : core_ops_range_Range usize) (ns : slice T), result (slice T).
(*** Vectors *)
-Definition vec T := { l: list T | Z.of_nat (length l) <= usize_max }.
+Definition alloc_vec_Vec T := { l: list T | Z.of_nat (length l) <= usize_max }.
-Definition vec_to_list {T: Type} (v: vec T) : list T := proj1_sig v.
+Definition alloc_vec_Vec_to_list {T: Type} (v: alloc_vec_Vec T) : list T := proj1_sig v.
-Definition vec_length {T: Type} (v: vec T) : Z := Z.of_nat (length (vec_to_list v)).
+Definition alloc_vec_Vec_length {T: Type} (v: alloc_vec_Vec T) : Z := Z.of_nat (length (alloc_vec_Vec_to_list v)).
-Definition vec_new (T: Type) : vec T := (exist _ [] le_0_usize_max).
+Definition alloc_vec_Vec_new (T: Type) : alloc_vec_Vec T := (exist _ [] le_0_usize_max).
-Lemma vec_len_in_usize {T} (v: vec T) : usize_min <= vec_length v <= usize_max.
+Lemma alloc_vec_Vec_len_in_usize {T} (v: alloc_vec_Vec T) : usize_min <= alloc_vec_Vec_length v <= usize_max.
Proof.
- unfold vec_length, usize_min.
+ unfold alloc_vec_Vec_length, usize_min.
split.
- lia.
- apply (proj2_sig v).
Qed.
-Definition vec_len (T: Type) (v: vec T) : usize :=
- exist _ (vec_length v) (vec_len_in_usize v).
+Definition alloc_vec_Vec_len (T: Type) (v: alloc_vec_Vec T) : usize :=
+ exist _ (alloc_vec_Vec_length v) (alloc_vec_Vec_len_in_usize v).
Fixpoint list_update {A} (l: list A) (n: nat) (a: A)
: list A :=
@@ -488,50 +552,271 @@ Fixpoint list_update {A} (l: list A) (n: nat) (a: A)
| S m => x :: (list_update t m a)
end end.
-Definition vec_bind {A B} (v: vec A) (f: list A -> result (list B)) : result (vec B) :=
- l <- f (vec_to_list v) ;
+Definition alloc_vec_Vec_bind {A B} (v: alloc_vec_Vec A) (f: list A -> result (list B)) : result (alloc_vec_Vec B) :=
+ l <- f (alloc_vec_Vec_to_list v) ;
match sumbool_of_bool (scalar_le_max Usize (Z.of_nat (length l))) with
| left H => Return (exist _ l (scalar_le_max_valid _ _ H))
| right _ => Fail_ Failure
end.
(* The **forward** function shouldn't be used *)
-Definition vec_push_fwd (T: Type) (v: vec T) (x: T) : unit := tt.
+Definition alloc_vec_Vec_push_fwd (T: Type) (v: alloc_vec_Vec T) (x: T) : unit := tt.
-Definition vec_push_back (T: Type) (v: vec T) (x: T) : result (vec T) :=
- vec_bind v (fun l => Return (l ++ [x])).
+Definition alloc_vec_Vec_push (T: Type) (v: alloc_vec_Vec T) (x: T) : result (alloc_vec_Vec T) :=
+ alloc_vec_Vec_bind v (fun l => Return (l ++ [x])).
(* The **forward** function shouldn't be used *)
-Definition vec_insert_fwd (T: Type) (v: vec T) (i: usize) (x: T) : result unit :=
- if to_Z i <? vec_length v then Return tt else Fail_ Failure.
+Definition alloc_vec_Vec_insert_fwd (T: Type) (v: alloc_vec_Vec T) (i: usize) (x: T) : result unit :=
+ if to_Z i <? alloc_vec_Vec_length v then Return tt else Fail_ Failure.
-Definition vec_insert_back (T: Type) (v: vec T) (i: usize) (x: T) : result (vec T) :=
- vec_bind v (fun l =>
+Definition alloc_vec_Vec_insert (T: Type) (v: alloc_vec_Vec T) (i: usize) (x: T) : result (alloc_vec_Vec T) :=
+ alloc_vec_Vec_bind v (fun l =>
if to_Z i <? Z.of_nat (length l)
then Return (list_update l (usize_to_nat i) x)
else Fail_ Failure).
-(* The **backward** function shouldn't be used *)
-Definition vec_index_fwd (T: Type) (v: vec T) (i: usize) : result T :=
- match nth_error (vec_to_list v) (usize_to_nat i) with
- | Some n => Return n
- | None => Fail_ Failure
- end.
-
-Definition vec_index_back (T: Type) (v: vec T) (i: usize) (x: T) : result unit :=
- if to_Z i <? vec_length v then Return tt else Fail_ Failure.
-
-(* The **backward** function shouldn't be used *)
-Definition vec_index_mut_fwd (T: Type) (v: vec T) (i: usize) : result T :=
- match nth_error (vec_to_list v) (usize_to_nat i) with
- | Some n => Return n
- | None => Fail_ Failure
+(* Helper *)
+Axiom alloc_vec_Vec_index_usize : forall {T : Type} (v : alloc_vec_Vec T) (i : usize) (x : T), result T.
+
+(* Helper *)
+Axiom alloc_vec_Vec_update_usize : forall {T : Type} (v : alloc_vec_Vec T) (i : usize) (x : T), result (alloc_vec_Vec T).
+
+(* Trait declaration: [core::slice::index::private_slice_index::Sealed] *)
+Definition core_slice_index_private_slice_index_Sealed (self : Type) := unit.
+
+(* Trait declaration: [core::slice::index::SliceIndex] *)
+Record core_slice_index_SliceIndex (Self T : Type) := mk_core_slice_index_SliceIndex {
+ core_slice_index_SliceIndex_sealedInst : core_slice_index_private_slice_index_Sealed Self;
+ core_slice_index_SliceIndex_Output : Type;
+ core_slice_index_SliceIndex_get : Self -> T -> result (option core_slice_index_SliceIndex_Output);
+ core_slice_index_SliceIndex_get_mut : Self -> T -> result (option core_slice_index_SliceIndex_Output);
+ core_slice_index_SliceIndex_get_mut_back : Self -> T -> option core_slice_index_SliceIndex_Output -> result T;
+ core_slice_index_SliceIndex_get_unchecked : Self -> const_raw_ptr T -> result (const_raw_ptr core_slice_index_SliceIndex_Output);
+ core_slice_index_SliceIndex_get_unchecked_mut : Self -> mut_raw_ptr T -> result (mut_raw_ptr core_slice_index_SliceIndex_Output);
+ core_slice_index_SliceIndex_index : Self -> T -> result core_slice_index_SliceIndex_Output;
+ core_slice_index_SliceIndex_index_mut : Self -> T -> result core_slice_index_SliceIndex_Output;
+ core_slice_index_SliceIndex_index_mut_back : Self -> T -> core_slice_index_SliceIndex_Output -> result T;
+}.
+Arguments mk_core_slice_index_SliceIndex {_ _}.
+Arguments core_slice_index_SliceIndex_sealedInst {_ _}.
+Arguments core_slice_index_SliceIndex_Output {_ _}.
+Arguments core_slice_index_SliceIndex_get {_ _}.
+Arguments core_slice_index_SliceIndex_get_mut {_ _}.
+Arguments core_slice_index_SliceIndex_get_mut_back {_ _}.
+Arguments core_slice_index_SliceIndex_get_unchecked {_ _}.
+Arguments core_slice_index_SliceIndex_get_unchecked_mut {_ _}.
+Arguments core_slice_index_SliceIndex_index {_ _}.
+Arguments core_slice_index_SliceIndex_index_mut {_ _}.
+Arguments core_slice_index_SliceIndex_index_mut_back {_ _}.
+
+(* [core::slice::index::[T]::index]: forward function *)
+Definition core_slice_index_Slice_index
+ (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T))
+ (s : slice T) (i : Idx) : result inst.(core_slice_index_SliceIndex_Output) :=
+ x <- inst.(core_slice_index_SliceIndex_get) i s;
+ match x with
+ | None => Fail_ Failure
+ | Some x => Return x
end.
-Definition vec_index_mut_back (T: Type) (v: vec T) (i: usize) (x: T) : result (vec T) :=
- vec_bind v (fun l =>
- if to_Z i <? Z.of_nat (length l)
- then Return (list_update l (usize_to_nat i) x)
- else Fail_ Failure).
+(* [core::slice::index::Range:::get]: forward function *)
+Axiom core_slice_index_Range_get : forall (T : Type) (i : core_ops_range_Range usize) (s : slice T), result (option (slice T)).
+
+(* [core::slice::index::Range::get_mut]: forward function *)
+Axiom core_slice_index_Range_get_mut :
+ forall (T : Type), core_ops_range_Range usize -> slice T -> result (option (slice T)).
+
+(* [core::slice::index::Range::get_mut]: backward function 0 *)
+Axiom core_slice_index_Range_get_mut_back :
+ forall (T : Type), core_ops_range_Range usize -> slice T -> option (slice T) -> result (slice T).
+
+(* [core::slice::index::Range::get_unchecked]: forward function *)
+Definition core_slice_index_Range_get_unchecked
+ (T : Type) :
+ core_ops_range_Range usize -> const_raw_ptr (slice T) -> result (const_raw_ptr (slice T)) :=
+ (* Don't know what the model should be - for now we always fail to make
+ sure code which uses it fails *)
+ fun _ _ => Fail_ Failure.
+
+(* [core::slice::index::Range::get_unchecked_mut]: forward function *)
+Definition core_slice_index_Range_get_unchecked_mut
+ (T : Type) :
+ core_ops_range_Range usize -> mut_raw_ptr (slice T) -> result (mut_raw_ptr (slice T)) :=
+ (* Don't know what the model should be - for now we always fail to make
+ sure code which uses it fails *)
+ fun _ _ => Fail_ Failure.
+
+(* [core::slice::index::Range::index]: forward function *)
+Axiom core_slice_index_Range_index :
+ forall (T : Type), core_ops_range_Range usize -> slice T -> result (slice T).
+
+(* [core::slice::index::Range::index_mut]: forward function *)
+Axiom core_slice_index_Range_index_mut :
+ forall (T : Type), core_ops_range_Range usize -> slice T -> result (slice T).
+
+(* [core::slice::index::Range::index_mut]: backward function 0 *)
+Axiom core_slice_index_Range_index_mut_back :
+ forall (T : Type), core_ops_range_Range usize -> slice T -> slice T -> result (slice T).
+
+(* [core::slice::index::[T]::index_mut]: forward function *)
+Axiom core_slice_index_Slice_index_mut :
+ forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)),
+ slice T -> Idx -> result inst.(core_slice_index_SliceIndex_Output).
+
+(* [core::slice::index::[T]::index_mut]: backward function 0 *)
+Axiom core_slice_index_Slice_index_mut_back :
+ forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)),
+ slice T -> Idx -> inst.(core_slice_index_SliceIndex_Output) -> result (slice T).
+
+(* [core::array::[T; N]::index]: forward function *)
+Axiom core_array_Array_index :
+ forall (T Idx : Type) (N : usize) (inst : core_ops_index_Index (slice T) Idx)
+ (a : array T N) (i : Idx), result inst.(core_ops_index_Index_Output).
+
+(* [core::array::[T; N]::index_mut]: forward function *)
+Axiom core_array_Array_index_mut :
+ forall (T Idx : Type) (N : usize) (inst : core_ops_index_IndexMut (slice T) Idx)
+ (a : array T N) (i : Idx), result inst.(core_ops_index_IndexMut_indexInst).(core_ops_index_Index_Output).
+
+(* [core::array::[T; N]::index_mut]: backward function 0 *)
+Axiom core_array_Array_index_mut_back :
+ forall (T Idx : Type) (N : usize) (inst : core_ops_index_IndexMut (slice T) Idx)
+ (a : array T N) (i : Idx) (x : inst.(core_ops_index_IndexMut_indexInst).(core_ops_index_Index_Output)), result (array T N).
+
+(* Trait implementation: [core::slice::index::[T]] *)
+Definition core_slice_index_Slice_coreopsindexIndexInst (T Idx : Type)
+ (inst : core_slice_index_SliceIndex Idx (slice T)) :
+ core_ops_index_Index (slice T) Idx := {|
+ core_ops_index_Index_Output := inst.(core_slice_index_SliceIndex_Output);
+ core_ops_index_Index_index := core_slice_index_Slice_index T Idx inst;
+|}.
+
+(* Trait implementation: [core::slice::index::private_slice_index::Range] *)
+Definition core_slice_index_private_slice_index_Range_coresliceindexprivate_slice_indexSealedInst
+ : core_slice_index_private_slice_index_Sealed (core_ops_range_Range usize) := tt.
+
+(* Trait implementation: [core::slice::index::Range] *)
+Definition core_slice_index_Range_coresliceindexSliceIndexInst (T : Type) :
+ core_slice_index_SliceIndex (core_ops_range_Range usize) (slice T) := {|
+ core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_Range_coresliceindexprivate_slice_indexSealedInst;
+ core_slice_index_SliceIndex_Output := slice T;
+ core_slice_index_SliceIndex_get := core_slice_index_Range_get T;
+ core_slice_index_SliceIndex_get_mut := core_slice_index_Range_get_mut T;
+ core_slice_index_SliceIndex_get_mut_back := core_slice_index_Range_get_mut_back T;
+ core_slice_index_SliceIndex_get_unchecked := core_slice_index_Range_get_unchecked T;
+ core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_Range_get_unchecked_mut T;
+ core_slice_index_SliceIndex_index := core_slice_index_Range_index T;
+ core_slice_index_SliceIndex_index_mut := core_slice_index_Range_index_mut T;
+ core_slice_index_SliceIndex_index_mut_back := core_slice_index_Range_index_mut_back T;
+|}.
+
+(* Trait implementation: [core::slice::index::[T]] *)
+Definition core_slice_index_Slice_coreopsindexIndexMutInst (T Idx : Type)
+ (inst : core_slice_index_SliceIndex Idx (slice T)) :
+ core_ops_index_IndexMut (slice T) Idx := {|
+ core_ops_index_IndexMut_indexInst := core_slice_index_Slice_coreopsindexIndexInst T Idx inst;
+ core_ops_index_IndexMut_index_mut := core_slice_index_Slice_index_mut T Idx inst;
+ core_ops_index_IndexMut_index_mut_back := core_slice_index_Slice_index_mut_back T Idx inst;
+|}.
+
+(* Trait implementation: [core::array::[T; N]] *)
+Definition core_array_Array_coreopsindexIndexInst (T Idx : Type) (N : usize)
+ (inst : core_ops_index_Index (slice T) Idx) :
+ core_ops_index_Index (array T N) Idx := {|
+ core_ops_index_Index_Output := inst.(core_ops_index_Index_Output);
+ core_ops_index_Index_index := core_array_Array_index T Idx N inst;
+|}.
+
+(* Trait implementation: [core::array::[T; N]] *)
+Definition core_array_Array_coreopsindexIndexMutInst (T Idx : Type) (N : usize)
+ (inst : core_ops_index_IndexMut (slice T) Idx) :
+ core_ops_index_IndexMut (array T N) Idx := {|
+ core_ops_index_IndexMut_indexInst := core_array_Array_coreopsindexIndexInst T Idx N inst.(core_ops_index_IndexMut_indexInst);
+ core_ops_index_IndexMut_index_mut := core_array_Array_index_mut T Idx N inst;
+ core_ops_index_IndexMut_index_mut_back := core_array_Array_index_mut_back T Idx N inst;
+|}.
+
+(* [core::slice::index::usize::get]: forward function *)
+Axiom core_slice_index_usize_get : forall (T : Type), usize -> slice T -> result (option T).
+
+(* [core::slice::index::usize::get_mut]: forward function *)
+Axiom core_slice_index_usize_get_mut : forall (T : Type), usize -> slice T -> result (option T).
+
+(* [core::slice::index::usize::get_mut]: backward function 0 *)
+Axiom core_slice_index_usize_get_mut_back :
+ forall (T : Type), usize -> slice T -> option T -> result (slice T).
+
+(* [core::slice::index::usize::get_unchecked]: forward function *)
+Axiom core_slice_index_usize_get_unchecked :
+ forall (T : Type), usize -> const_raw_ptr (slice T) -> result (const_raw_ptr T).
+
+(* [core::slice::index::usize::get_unchecked_mut]: forward function *)
+Axiom core_slice_index_usize_get_unchecked_mut :
+ forall (T : Type), usize -> mut_raw_ptr (slice T) -> result (mut_raw_ptr T).
+
+(* [core::slice::index::usize::index]: forward function *)
+Axiom core_slice_index_usize_index : forall (T : Type), usize -> slice T -> result T.
+
+(* [core::slice::index::usize::index_mut]: forward function *)
+Axiom core_slice_index_usize_index_mut : forall (T : Type), usize -> slice T -> result T.
+
+(* [core::slice::index::usize::index_mut]: backward function 0 *)
+Axiom core_slice_index_usize_index_mut_back :
+ forall (T : Type), usize -> slice T -> T -> result (slice T).
+
+(* Trait implementation: [core::slice::index::private_slice_index::usize] *)
+Definition core_slice_index_private_slice_index_usize_coresliceindexprivate_slice_indexSealedInst
+ : core_slice_index_private_slice_index_Sealed usize := tt.
+
+(* Trait implementation: [core::slice::index::usize] *)
+Definition core_slice_index_usize_coresliceindexSliceIndexInst (T : Type) :
+ core_slice_index_SliceIndex usize (slice T) := {|
+ core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_usize_coresliceindexprivate_slice_indexSealedInst;
+ core_slice_index_SliceIndex_Output := T;
+ core_slice_index_SliceIndex_get := core_slice_index_usize_get T;
+ core_slice_index_SliceIndex_get_mut := core_slice_index_usize_get_mut T;
+ core_slice_index_SliceIndex_get_mut_back := core_slice_index_usize_get_mut_back T;
+ core_slice_index_SliceIndex_get_unchecked := core_slice_index_usize_get_unchecked T;
+ core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_usize_get_unchecked_mut T;
+ core_slice_index_SliceIndex_index := core_slice_index_usize_index T;
+ core_slice_index_SliceIndex_index_mut := core_slice_index_usize_index_mut T;
+ core_slice_index_SliceIndex_index_mut_back := core_slice_index_usize_index_mut_back T;
+|}.
+
+(* [alloc::vec::Vec::index]: forward function *)
+Axiom alloc_vec_Vec_index : forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T))
+ (Self : alloc_vec_Vec T) (i : Idx), result inst.(core_slice_index_SliceIndex_Output).
+
+(* [alloc::vec::Vec::index_mut]: forward function *)
+Axiom alloc_vec_Vec_index_mut : forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T))
+ (Self : alloc_vec_Vec T) (i : Idx), result inst.(core_slice_index_SliceIndex_Output).
+
+(* [alloc::vec::Vec::index_mut]: backward function 0 *)
+Axiom alloc_vec_Vec_index_mut_back :
+ forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T))
+ (Self : alloc_vec_Vec T) (i : Idx) (x : inst.(core_slice_index_SliceIndex_Output)), result (alloc_vec_Vec T).
+
+(* Trait implementation: [alloc::vec::Vec] *)
+Definition alloc_vec_Vec_coreopsindexIndexInst (T Idx : Type)
+ (inst : core_slice_index_SliceIndex Idx (slice T)) :
+ core_ops_index_Index (alloc_vec_Vec T) Idx := {|
+ core_ops_index_Index_Output := inst.(core_slice_index_SliceIndex_Output);
+ core_ops_index_Index_index := alloc_vec_Vec_index T Idx inst;
+|}.
+
+(* Trait implementation: [alloc::vec::Vec] *)
+Definition alloc_vec_Vec_coreopsindexIndexMutInst (T Idx : Type)
+ (inst : core_slice_index_SliceIndex Idx (slice T)) :
+ core_ops_index_IndexMut (alloc_vec_Vec T) Idx := {|
+ core_ops_index_IndexMut_indexInst := alloc_vec_Vec_coreopsindexIndexInst T Idx inst;
+ core_ops_index_IndexMut_index_mut := alloc_vec_Vec_index_mut T Idx inst;
+ core_ops_index_IndexMut_index_mut_back := alloc_vec_Vec_index_mut_back T Idx inst;
+|}.
+
+(*** Theorems *)
+
+Axiom alloc_vec_Vec_index_mut_back_eq : forall {a : Type} (v : alloc_vec_Vec a) (i : usize) (x : a),
+ alloc_vec_Vec_index_mut_back a usize (core_slice_index_usize_coresliceindexSliceIndexInst a) v i x =
+ alloc_vec_Vec_update_usize v i x.
End Primitives.