summaryrefslogtreecommitdiff
path: root/tests/coq/hashmap/Hashmap_Funs.v
diff options
context:
space:
mode:
Diffstat (limited to 'tests/coq/hashmap/Hashmap_Funs.v')
-rw-r--r--tests/coq/hashmap/Hashmap_Funs.v503
1 files changed, 193 insertions, 310 deletions
diff --git a/tests/coq/hashmap/Hashmap_Funs.v b/tests/coq/hashmap/Hashmap_Funs.v
index 64de44a6..5cd9fe70 100644
--- a/tests/coq/hashmap/Hashmap_Funs.v
+++ b/tests/coq/hashmap/Hashmap_Funs.v
@@ -10,39 +10,39 @@ Require Import Hashmap_Types.
Include Hashmap_Types.
Module Hashmap_Funs.
-(** [hashmap::hash_key]: forward function
+(** [hashmap::hash_key]:
Source: 'src/hashmap.rs', lines 27:0-27:32 *)
Definition hash_key (k : usize) : result usize :=
Return k.
-(** [hashmap::{hashmap::HashMap<T>}::allocate_slots]: loop 0: forward function
+(** [hashmap::{hashmap::HashMap<T>}::allocate_slots]: loop 0:
Source: 'src/hashmap.rs', lines 50:4-56:5 *)
Fixpoint hashMap_allocate_slots_loop
- (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n0 : usize) :
+ (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n1 : usize) :
result (alloc_vec_Vec (List_t T))
:=
match n with
| O => Fail_ OutOfFuel
- | S n1 =>
- if n0 s> 0%usize
+ | S n2 =>
+ if n1 s> 0%usize
then (
- slots0 <- alloc_vec_Vec_push (List_t T) slots List_Nil;
- n2 <- usize_sub n0 1%usize;
- hashMap_allocate_slots_loop T n1 slots0 n2)
+ v <- alloc_vec_Vec_push (List_t T) slots List_Nil;
+ n3 <- usize_sub n1 1%usize;
+ hashMap_allocate_slots_loop T n2 v n3)
else Return slots
end
.
-(** [hashmap::{hashmap::HashMap<T>}::allocate_slots]: forward function
+(** [hashmap::{hashmap::HashMap<T>}::allocate_slots]:
Source: 'src/hashmap.rs', lines 50:4-50:76 *)
Definition hashMap_allocate_slots
- (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n0 : usize) :
+ (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n1 : usize) :
result (alloc_vec_Vec (List_t T))
:=
- hashMap_allocate_slots_loop T n slots n0
+ hashMap_allocate_slots_loop T n slots n1
.
-(** [hashmap::{hashmap::HashMap<T>}::new_with_capacity]: forward function
+(** [hashmap::{hashmap::HashMap<T>}::new_with_capacity]:
Source: 'src/hashmap.rs', lines 59:4-63:13 *)
Definition hashMap_new_with_capacity
(T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize)
@@ -52,24 +52,23 @@ Definition hashMap_new_with_capacity
let v := alloc_vec_Vec_new (List_t T) in
slots <- hashMap_allocate_slots T n v capacity;
i <- usize_mul capacity max_load_dividend;
- i0 <- usize_div i max_load_divisor;
+ i1 <- usize_div i max_load_divisor;
Return
{|
hashMap_num_entries := 0%usize;
hashMap_max_load_factor := (max_load_dividend, max_load_divisor);
- hashMap_max_load := i0;
+ hashMap_max_load := i1;
hashMap_slots := slots
|}
.
-(** [hashmap::{hashmap::HashMap<T>}::new]: forward function
+(** [hashmap::{hashmap::HashMap<T>}::new]:
Source: 'src/hashmap.rs', lines 75:4-75:24 *)
Definition hashMap_new (T : Type) (n : nat) : result (HashMap_t T) :=
hashMap_new_with_capacity T n 32%usize 4%usize 5%usize
.
-(** [hashmap::{hashmap::HashMap<T>}::clear]: loop 0: merged forward/backward function
- (there is a single backward function, and the forward function returns ())
+(** [hashmap::{hashmap::HashMap<T>}::clear]: loop 0:
Source: 'src/hashmap.rs', lines 80:4-88:5 *)
Fixpoint hashMap_clear_loop
(T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (i : usize) :
@@ -77,101 +76,73 @@ Fixpoint hashMap_clear_loop
:=
match n with
| O => Fail_ OutOfFuel
- | S n0 =>
- let i0 := alloc_vec_Vec_len (List_t T) slots in
- if i s< i0
+ | S n1 =>
+ let i1 := alloc_vec_Vec_len (List_t T) slots in
+ if i s< i1
then (
- i1 <- usize_add i 1%usize;
- slots0 <-
- alloc_vec_Vec_index_mut_back (List_t T) usize
- (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) slots i
- List_Nil;
- hashMap_clear_loop T n0 slots0 i1)
+ p <-
+ alloc_vec_Vec_index_mut (List_t T) usize
+ (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) slots i;
+ let (_, index_mut_back) := p in
+ i2 <- usize_add i 1%usize;
+ slots1 <- index_mut_back List_Nil;
+ hashMap_clear_loop T n1 slots1 i2)
else Return slots
end
.
-(** [hashmap::{hashmap::HashMap<T>}::clear]: merged forward/backward function
- (there is a single backward function, and the forward function returns ())
+(** [hashmap::{hashmap::HashMap<T>}::clear]:
Source: 'src/hashmap.rs', lines 80:4-80:27 *)
Definition hashMap_clear
(T : Type) (n : nat) (self : HashMap_t T) : result (HashMap_t T) :=
- v <- hashMap_clear_loop T n self.(hashMap_slots) 0%usize;
+ back <- hashMap_clear_loop T n self.(hashMap_slots) 0%usize;
Return
{|
hashMap_num_entries := 0%usize;
hashMap_max_load_factor := self.(hashMap_max_load_factor);
hashMap_max_load := self.(hashMap_max_load);
- hashMap_slots := v
+ hashMap_slots := back
|}
.
-(** [hashmap::{hashmap::HashMap<T>}::len]: forward function
+(** [hashmap::{hashmap::HashMap<T>}::len]:
Source: 'src/hashmap.rs', lines 90:4-90:30 *)
Definition hashMap_len (T : Type) (self : HashMap_t T) : result usize :=
Return self.(hashMap_num_entries)
.
-(** [hashmap::{hashmap::HashMap<T>}::insert_in_list]: loop 0: forward function
+(** [hashmap::{hashmap::HashMap<T>}::insert_in_list]: loop 0:
Source: 'src/hashmap.rs', lines 97:4-114:5 *)
Fixpoint hashMap_insert_in_list_loop
(T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) :
- result bool
+ result (bool * (List_t T))
:=
match n with
| O => Fail_ OutOfFuel
- | S n0 =>
+ | S n1 =>
match ls with
| List_Cons ckey cvalue tl =>
if ckey s= key
- then Return false
- else hashMap_insert_in_list_loop T n0 key value tl
- | List_Nil => Return true
+ then Return (false, List_Cons ckey value tl)
+ else (
+ p <- hashMap_insert_in_list_loop T n1 key value tl;
+ let (b, back) := p in
+ Return (b, List_Cons ckey cvalue back))
+ | List_Nil => let l := List_Nil in Return (true, List_Cons key value l)
end
end
.
-(** [hashmap::{hashmap::HashMap<T>}::insert_in_list]: forward function
+(** [hashmap::{hashmap::HashMap<T>}::insert_in_list]:
Source: 'src/hashmap.rs', lines 97:4-97:71 *)
Definition hashMap_insert_in_list
(T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) :
- result bool
+ result (bool * (List_t T))
:=
hashMap_insert_in_list_loop T n key value ls
.
-(** [hashmap::{hashmap::HashMap<T>}::insert_in_list]: loop 0: backward function 0
- Source: 'src/hashmap.rs', lines 97:4-114:5 *)
-Fixpoint hashMap_insert_in_list_loop_back
- (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) :
- result (List_t T)
- :=
- match n with
- | O => Fail_ OutOfFuel
- | S n0 =>
- match ls with
- | List_Cons ckey cvalue tl =>
- if ckey s= key
- then Return (List_Cons ckey value tl)
- else (
- tl0 <- hashMap_insert_in_list_loop_back T n0 key value tl;
- Return (List_Cons ckey cvalue tl0))
- | List_Nil => let l := List_Nil in Return (List_Cons key value l)
- end
- end
-.
-
-(** [hashmap::{hashmap::HashMap<T>}::insert_in_list]: backward function 0
- Source: 'src/hashmap.rs', lines 97:4-97:71 *)
-Definition hashMap_insert_in_list_back
- (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) :
- result (List_t T)
- :=
- hashMap_insert_in_list_loop_back T n key value ls
-.
-
-(** [hashmap::{hashmap::HashMap<T>}::insert_no_resize]: merged forward/backward function
- (there is a single backward function, and the forward function returns ())
+(** [hashmap::{hashmap::HashMap<T>}::insert_no_resize]:
Source: 'src/hashmap.rs', lines 117:4-117:54 *)
Definition hashMap_insert_no_resize
(T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) :
@@ -180,32 +151,26 @@ Definition hashMap_insert_no_resize
hash <- hash_key key;
let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in
hash_mod <- usize_rem hash i;
- l <-
+ p <-
alloc_vec_Vec_index_mut (List_t T) usize
(core_slice_index_SliceIndexUsizeSliceTInst (List_t T))
self.(hashMap_slots) hash_mod;
- inserted <- hashMap_insert_in_list T n key value l;
+ let (l, index_mut_back) := p in
+ p1 <- hashMap_insert_in_list T n key value l;
+ let (inserted, l1) := p1 in
if inserted
then (
- i0 <- usize_add self.(hashMap_num_entries) 1%usize;
- l0 <- hashMap_insert_in_list_back T n key value l;
- v <-
- alloc_vec_Vec_index_mut_back (List_t T) usize
- (core_slice_index_SliceIndexUsizeSliceTInst (List_t T))
- self.(hashMap_slots) hash_mod l0;
+ i1 <- usize_add self.(hashMap_num_entries) 1%usize;
+ v <- index_mut_back l1;
Return
{|
- hashMap_num_entries := i0;
+ hashMap_num_entries := i1;
hashMap_max_load_factor := self.(hashMap_max_load_factor);
hashMap_max_load := self.(hashMap_max_load);
hashMap_slots := v
|})
else (
- l0 <- hashMap_insert_in_list_back T n key value l;
- v <-
- alloc_vec_Vec_index_mut_back (List_t T) usize
- (core_slice_index_SliceIndexUsizeSliceTInst (List_t T))
- self.(hashMap_slots) hash_mod l0;
+ v <- index_mut_back l1;
Return
{|
hashMap_num_entries := self.(hashMap_num_entries);
@@ -215,8 +180,7 @@ Definition hashMap_insert_no_resize
|})
.
-(** [hashmap::{hashmap::HashMap<T>}::move_elements_from_list]: loop 0: merged forward/backward function
- (there is a single backward function, and the forward function returns ())
+(** [hashmap::{hashmap::HashMap<T>}::move_elements_from_list]: loop 0:
Source: 'src/hashmap.rs', lines 183:4-196:5 *)
Fixpoint hashMap_move_elements_from_list_loop
(T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) :
@@ -224,18 +188,17 @@ Fixpoint hashMap_move_elements_from_list_loop
:=
match n with
| O => Fail_ OutOfFuel
- | S n0 =>
+ | S n1 =>
match ls with
| List_Cons k v tl =>
- ntable0 <- hashMap_insert_no_resize T n0 ntable k v;
- hashMap_move_elements_from_list_loop T n0 ntable0 tl
+ hm <- hashMap_insert_no_resize T n1 ntable k v;
+ hashMap_move_elements_from_list_loop T n1 hm tl
| List_Nil => Return ntable
end
end
.
-(** [hashmap::{hashmap::HashMap<T>}::move_elements_from_list]: merged forward/backward function
- (there is a single backward function, and the forward function returns ())
+(** [hashmap::{hashmap::HashMap<T>}::move_elements_from_list]:
Source: 'src/hashmap.rs', lines 183:4-183:72 *)
Definition hashMap_move_elements_from_list
(T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) :
@@ -244,8 +207,7 @@ Definition hashMap_move_elements_from_list
hashMap_move_elements_from_list_loop T n ntable ls
.
-(** [hashmap::{hashmap::HashMap<T>}::move_elements]: loop 0: merged forward/backward function
- (there is a single backward function, and the forward function returns ())
+(** [hashmap::{hashmap::HashMap<T>}::move_elements]: loop 0:
Source: 'src/hashmap.rs', lines 171:4-180:5 *)
Fixpoint hashMap_move_elements_loop
(T : Type) (n : nat) (ntable : HashMap_t T)
@@ -254,108 +216,105 @@ Fixpoint hashMap_move_elements_loop
:=
match n with
| O => Fail_ OutOfFuel
- | S n0 =>
- let i0 := alloc_vec_Vec_len (List_t T) slots in
- if i s< i0
+ | S n1 =>
+ let i1 := alloc_vec_Vec_len (List_t T) slots in
+ if i s< i1
then (
- l <-
+ p <-
alloc_vec_Vec_index_mut (List_t T) usize
(core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) slots i;
- let ls := core_mem_replace (List_t T) l List_Nil in
- ntable0 <- hashMap_move_elements_from_list T n0 ntable ls;
- i1 <- usize_add i 1%usize;
- let l0 := core_mem_replace_back (List_t T) l List_Nil in
- slots0 <-
- alloc_vec_Vec_index_mut_back (List_t T) usize
- (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) slots i l0;
- hashMap_move_elements_loop T n0 ntable0 slots0 i1)
+ let (l, index_mut_back) := p in
+ let (ls, l1) := core_mem_replace (List_t T) l List_Nil in
+ hm <- hashMap_move_elements_from_list T n1 ntable ls;
+ i2 <- usize_add i 1%usize;
+ slots1 <- index_mut_back l1;
+ back_'a <- hashMap_move_elements_loop T n1 hm slots1 i2;
+ let (hm1, v) := back_'a in
+ Return (hm1, v))
else Return (ntable, slots)
end
.
-(** [hashmap::{hashmap::HashMap<T>}::move_elements]: merged forward/backward function
- (there is a single backward function, and the forward function returns ())
+(** [hashmap::{hashmap::HashMap<T>}::move_elements]:
Source: 'src/hashmap.rs', lines 171:4-171:95 *)
Definition hashMap_move_elements
(T : Type) (n : nat) (ntable : HashMap_t T)
(slots : alloc_vec_Vec (List_t T)) (i : usize) :
result ((HashMap_t T) * (alloc_vec_Vec (List_t T)))
:=
- hashMap_move_elements_loop T n ntable slots i
+ back_'a <- hashMap_move_elements_loop T n ntable slots i;
+ let (hm, v) := back_'a in
+ Return (hm, v)
.
-(** [hashmap::{hashmap::HashMap<T>}::try_resize]: merged forward/backward function
- (there is a single backward function, and the forward function returns ())
+(** [hashmap::{hashmap::HashMap<T>}::try_resize]:
Source: 'src/hashmap.rs', lines 140:4-140:28 *)
Definition hashMap_try_resize
(T : Type) (n : nat) (self : HashMap_t T) : result (HashMap_t T) :=
max_usize <- scalar_cast U32 Usize core_u32_max;
let capacity := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in
n1 <- usize_div max_usize 2%usize;
- let (i, i0) := self.(hashMap_max_load_factor) in
- i1 <- usize_div n1 i;
- if capacity s<= i1
+ let (i, i1) := self.(hashMap_max_load_factor) in
+ i2 <- usize_div n1 i;
+ if capacity s<= i2
then (
- i2 <- usize_mul capacity 2%usize;
- ntable <- hashMap_new_with_capacity T n i2 i i0;
+ i3 <- usize_mul capacity 2%usize;
+ ntable <- hashMap_new_with_capacity T n i3 i i1;
p <- hashMap_move_elements T n ntable self.(hashMap_slots) 0%usize;
- let (ntable0, _) := p in
+ let (ntable1, _) := p in
Return
{|
hashMap_num_entries := self.(hashMap_num_entries);
- hashMap_max_load_factor := (i, i0);
- hashMap_max_load := ntable0.(hashMap_max_load);
- hashMap_slots := ntable0.(hashMap_slots)
+ hashMap_max_load_factor := (i, i1);
+ hashMap_max_load := ntable1.(hashMap_max_load);
+ hashMap_slots := ntable1.(hashMap_slots)
|})
else
Return
{|
hashMap_num_entries := self.(hashMap_num_entries);
- hashMap_max_load_factor := (i, i0);
+ hashMap_max_load_factor := (i, i1);
hashMap_max_load := self.(hashMap_max_load);
hashMap_slots := self.(hashMap_slots)
|}
.
-(** [hashmap::{hashmap::HashMap<T>}::insert]: merged forward/backward function
- (there is a single backward function, and the forward function returns ())
+(** [hashmap::{hashmap::HashMap<T>}::insert]:
Source: 'src/hashmap.rs', lines 129:4-129:48 *)
Definition hashMap_insert
(T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) :
result (HashMap_t T)
:=
- self0 <- hashMap_insert_no_resize T n self key value;
- i <- hashMap_len T self0;
- if i s> self0.(hashMap_max_load)
- then hashMap_try_resize T n self0
- else Return self0
+ hm <- hashMap_insert_no_resize T n self key value;
+ i <- hashMap_len T hm;
+ if i s> hm.(hashMap_max_load) then hashMap_try_resize T n hm else Return hm
.
-(** [hashmap::{hashmap::HashMap<T>}::contains_key_in_list]: loop 0: forward function
+(** [hashmap::{hashmap::HashMap<T>}::contains_key_in_list]: loop 0:
Source: 'src/hashmap.rs', lines 206:4-219:5 *)
Fixpoint hashMap_contains_key_in_list_loop
(T : Type) (n : nat) (key : usize) (ls : List_t T) : result bool :=
match n with
| O => Fail_ OutOfFuel
- | S n0 =>
+ | S n1 =>
match ls with
- | List_Cons ckey t tl =>
+ | List_Cons ckey _ tl =>
if ckey s= key
then Return true
- else hashMap_contains_key_in_list_loop T n0 key tl
+ else hashMap_contains_key_in_list_loop T n1 key tl
| List_Nil => Return false
end
end
.
-(** [hashmap::{hashmap::HashMap<T>}::contains_key_in_list]: forward function
+(** [hashmap::{hashmap::HashMap<T>}::contains_key_in_list]:
Source: 'src/hashmap.rs', lines 206:4-206:68 *)
Definition hashMap_contains_key_in_list
(T : Type) (n : nat) (key : usize) (ls : List_t T) : result bool :=
hashMap_contains_key_in_list_loop T n key ls
.
-(** [hashmap::{hashmap::HashMap<T>}::contains_key]: forward function
+(** [hashmap::{hashmap::HashMap<T>}::contains_key]:
Source: 'src/hashmap.rs', lines 199:4-199:49 *)
Definition hashMap_contains_key
(T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result bool :=
@@ -369,31 +328,31 @@ Definition hashMap_contains_key
hashMap_contains_key_in_list T n key l
.
-(** [hashmap::{hashmap::HashMap<T>}::get_in_list]: loop 0: forward function
+(** [hashmap::{hashmap::HashMap<T>}::get_in_list]: loop 0:
Source: 'src/hashmap.rs', lines 224:4-237:5 *)
Fixpoint hashMap_get_in_list_loop
(T : Type) (n : nat) (key : usize) (ls : List_t T) : result T :=
match n with
| O => Fail_ OutOfFuel
- | S n0 =>
+ | S n1 =>
match ls with
| List_Cons ckey cvalue tl =>
if ckey s= key
then Return cvalue
- else hashMap_get_in_list_loop T n0 key tl
+ else hashMap_get_in_list_loop T n1 key tl
| List_Nil => Fail_ Failure
end
end
.
-(** [hashmap::{hashmap::HashMap<T>}::get_in_list]: forward function
+(** [hashmap::{hashmap::HashMap<T>}::get_in_list]:
Source: 'src/hashmap.rs', lines 224:4-224:70 *)
Definition hashMap_get_in_list
(T : Type) (n : nat) (key : usize) (ls : List_t T) : result T :=
hashMap_get_in_list_loop T n key ls
.
-(** [hashmap::{hashmap::HashMap<T>}::get]: forward function
+(** [hashmap::{hashmap::HashMap<T>}::get]:
Source: 'src/hashmap.rs', lines 239:4-239:55 *)
Definition hashMap_get
(T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result T :=
@@ -407,264 +366,188 @@ Definition hashMap_get
hashMap_get_in_list T n key l
.
-(** [hashmap::{hashmap::HashMap<T>}::get_mut_in_list]: loop 0: forward function
+(** [hashmap::{hashmap::HashMap<T>}::get_mut_in_list]: loop 0:
Source: 'src/hashmap.rs', lines 245:4-254:5 *)
Fixpoint hashMap_get_mut_in_list_loop
- (T : Type) (n : nat) (ls : List_t T) (key : usize) : result T :=
- match n with
- | O => Fail_ OutOfFuel
- | S n0 =>
- match ls with
- | List_Cons ckey cvalue tl =>
- if ckey s= key
- then Return cvalue
- else hashMap_get_mut_in_list_loop T n0 tl key
- | List_Nil => Fail_ Failure
- end
- end
-.
-
-(** [hashmap::{hashmap::HashMap<T>}::get_mut_in_list]: forward function
- Source: 'src/hashmap.rs', lines 245:4-245:86 *)
-Definition hashMap_get_mut_in_list
- (T : Type) (n : nat) (ls : List_t T) (key : usize) : result T :=
- hashMap_get_mut_in_list_loop T n ls key
-.
-
-(** [hashmap::{hashmap::HashMap<T>}::get_mut_in_list]: loop 0: backward function 0
- Source: 'src/hashmap.rs', lines 245:4-254:5 *)
-Fixpoint hashMap_get_mut_in_list_loop_back
- (T : Type) (n : nat) (ls : List_t T) (key : usize) (ret : T) :
- result (List_t T)
+ (T : Type) (n : nat) (ls : List_t T) (key : usize) :
+ result (T * (T -> result (List_t T)))
:=
match n with
| O => Fail_ OutOfFuel
- | S n0 =>
+ | S n1 =>
match ls with
| List_Cons ckey cvalue tl =>
if ckey s= key
- then Return (List_Cons ckey ret tl)
+ then
+ let back_'a := fun (ret : T) => Return (List_Cons ckey ret tl) in
+ Return (cvalue, back_'a)
else (
- tl0 <- hashMap_get_mut_in_list_loop_back T n0 tl key ret;
- Return (List_Cons ckey cvalue tl0))
+ p <- hashMap_get_mut_in_list_loop T n1 tl key;
+ let (t, back_'a) := p in
+ let back_'a1 :=
+ fun (ret : T) =>
+ tl1 <- back_'a ret; Return (List_Cons ckey cvalue tl1) in
+ Return (t, back_'a1))
| List_Nil => Fail_ Failure
end
end
.
-(** [hashmap::{hashmap::HashMap<T>}::get_mut_in_list]: backward function 0
+(** [hashmap::{hashmap::HashMap<T>}::get_mut_in_list]:
Source: 'src/hashmap.rs', lines 245:4-245:86 *)
-Definition hashMap_get_mut_in_list_back
- (T : Type) (n : nat) (ls : List_t T) (key : usize) (ret : T) :
- result (List_t T)
+Definition hashMap_get_mut_in_list
+ (T : Type) (n : nat) (ls : List_t T) (key : usize) :
+ result (T * (T -> result (List_t T)))
:=
- hashMap_get_mut_in_list_loop_back T n ls key ret
+ p <- hashMap_get_mut_in_list_loop T n ls key;
+ let (t, back_'a) := p in
+ let back_'a1 := fun (ret : T) => back_'a ret in
+ Return (t, back_'a1)
.
-(** [hashmap::{hashmap::HashMap<T>}::get_mut]: forward function
+(** [hashmap::{hashmap::HashMap<T>}::get_mut]:
Source: 'src/hashmap.rs', lines 257:4-257:67 *)
Definition hashMap_get_mut
- (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result T :=
- hash <- hash_key key;
- let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in
- hash_mod <- usize_rem hash i;
- l <-
- alloc_vec_Vec_index_mut (List_t T) usize
- (core_slice_index_SliceIndexUsizeSliceTInst (List_t T))
- self.(hashMap_slots) hash_mod;
- hashMap_get_mut_in_list T n l key
-.
-
-(** [hashmap::{hashmap::HashMap<T>}::get_mut]: backward function 0
- Source: 'src/hashmap.rs', lines 257:4-257:67 *)
-Definition hashMap_get_mut_back
- (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (ret : T) :
- result (HashMap_t T)
+ (T : Type) (n : nat) (self : HashMap_t T) (key : usize) :
+ result (T * (T -> result (HashMap_t T)))
:=
hash <- hash_key key;
let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in
hash_mod <- usize_rem hash i;
- l <-
+ p <-
alloc_vec_Vec_index_mut (List_t T) usize
(core_slice_index_SliceIndexUsizeSliceTInst (List_t T))
self.(hashMap_slots) hash_mod;
- l0 <- hashMap_get_mut_in_list_back T n l key ret;
- v <-
- alloc_vec_Vec_index_mut_back (List_t T) usize
- (core_slice_index_SliceIndexUsizeSliceTInst (List_t T))
- self.(hashMap_slots) hash_mod l0;
- Return
- {|
- hashMap_num_entries := self.(hashMap_num_entries);
- hashMap_max_load_factor := self.(hashMap_max_load_factor);
- hashMap_max_load := self.(hashMap_max_load);
- hashMap_slots := v
- |}
-.
-
-(** [hashmap::{hashmap::HashMap<T>}::remove_from_list]: loop 0: forward function
+ let (l, index_mut_back) := p in
+ p1 <- hashMap_get_mut_in_list T n l key;
+ let (t, get_mut_in_list_back) := p1 in
+ let back_'a :=
+ fun (ret : T) =>
+ l1 <- get_mut_in_list_back ret;
+ v <- index_mut_back l1;
+ Return
+ {|
+ hashMap_num_entries := self.(hashMap_num_entries);
+ hashMap_max_load_factor := self.(hashMap_max_load_factor);
+ hashMap_max_load := self.(hashMap_max_load);
+ hashMap_slots := v
+ |} in
+ Return (t, back_'a)
+.
+
+(** [hashmap::{hashmap::HashMap<T>}::remove_from_list]: loop 0:
Source: 'src/hashmap.rs', lines 265:4-291:5 *)
Fixpoint hashMap_remove_from_list_loop
- (T : Type) (n : nat) (key : usize) (ls : List_t T) : result (option T) :=
+ (T : Type) (n : nat) (key : usize) (ls : List_t T) :
+ result ((option T) * (List_t T))
+ :=
match n with
| O => Fail_ OutOfFuel
- | S n0 =>
+ | S n1 =>
match ls with
| List_Cons ckey t tl =>
if ckey s= key
then
- let mv_ls := core_mem_replace (List_t T) (List_Cons ckey t tl) List_Nil
- in
+ let (mv_ls, _) :=
+ core_mem_replace (List_t T) (List_Cons ckey t tl) List_Nil in
match mv_ls with
- | List_Cons i cvalue tl0 => Return (Some cvalue)
+ | List_Cons _ cvalue tl1 => Return (Some cvalue, tl1)
| List_Nil => Fail_ Failure
end
- else hashMap_remove_from_list_loop T n0 key tl
- | List_Nil => Return None
+ else (
+ p <- hashMap_remove_from_list_loop T n1 key tl;
+ let (o, back) := p in
+ Return (o, List_Cons ckey t back))
+ | List_Nil => Return (None, List_Nil)
end
end
.
-(** [hashmap::{hashmap::HashMap<T>}::remove_from_list]: forward function
+(** [hashmap::{hashmap::HashMap<T>}::remove_from_list]:
Source: 'src/hashmap.rs', lines 265:4-265:69 *)
Definition hashMap_remove_from_list
- (T : Type) (n : nat) (key : usize) (ls : List_t T) : result (option T) :=
+ (T : Type) (n : nat) (key : usize) (ls : List_t T) :
+ result ((option T) * (List_t T))
+ :=
hashMap_remove_from_list_loop T n key ls
.
-(** [hashmap::{hashmap::HashMap<T>}::remove_from_list]: loop 0: backward function 1
- Source: 'src/hashmap.rs', lines 265:4-291:5 *)
-Fixpoint hashMap_remove_from_list_loop_back
- (T : Type) (n : nat) (key : usize) (ls : List_t T) : result (List_t T) :=
- match n with
- | O => Fail_ OutOfFuel
- | S n0 =>
- match ls with
- | List_Cons ckey t tl =>
- if ckey s= key
- then
- let mv_ls := core_mem_replace (List_t T) (List_Cons ckey t tl) List_Nil
- in
- match mv_ls with
- | List_Cons i cvalue tl0 => Return tl0
- | List_Nil => Fail_ Failure
- end
- else (
- tl0 <- hashMap_remove_from_list_loop_back T n0 key tl;
- Return (List_Cons ckey t tl0))
- | List_Nil => Return List_Nil
- end
- end
-.
-
-(** [hashmap::{hashmap::HashMap<T>}::remove_from_list]: backward function 1
- Source: 'src/hashmap.rs', lines 265:4-265:69 *)
-Definition hashMap_remove_from_list_back
- (T : Type) (n : nat) (key : usize) (ls : List_t T) : result (List_t T) :=
- hashMap_remove_from_list_loop_back T n key ls
-.
-
-(** [hashmap::{hashmap::HashMap<T>}::remove]: forward function
+(** [hashmap::{hashmap::HashMap<T>}::remove]:
Source: 'src/hashmap.rs', lines 294:4-294:52 *)
Definition hashMap_remove
(T : Type) (n : nat) (self : HashMap_t T) (key : usize) :
- result (option T)
- :=
- hash <- hash_key key;
- let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in
- hash_mod <- usize_rem hash i;
- l <-
- alloc_vec_Vec_index_mut (List_t T) usize
- (core_slice_index_SliceIndexUsizeSliceTInst (List_t T))
- self.(hashMap_slots) hash_mod;
- x <- hashMap_remove_from_list T n key l;
- match x with
- | None => Return None
- | Some x0 =>
- _ <- usize_sub self.(hashMap_num_entries) 1%usize; Return (Some x0)
- end
-.
-
-(** [hashmap::{hashmap::HashMap<T>}::remove]: backward function 0
- Source: 'src/hashmap.rs', lines 294:4-294:52 *)
-Definition hashMap_remove_back
- (T : Type) (n : nat) (self : HashMap_t T) (key : usize) :
- result (HashMap_t T)
+ result ((option T) * (HashMap_t T))
:=
hash <- hash_key key;
let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in
hash_mod <- usize_rem hash i;
- l <-
+ p <-
alloc_vec_Vec_index_mut (List_t T) usize
(core_slice_index_SliceIndexUsizeSliceTInst (List_t T))
self.(hashMap_slots) hash_mod;
- x <- hashMap_remove_from_list T n key l;
+ let (l, index_mut_back) := p in
+ p1 <- hashMap_remove_from_list T n key l;
+ let (x, l1) := p1 in
match x with
| None =>
- l0 <- hashMap_remove_from_list_back T n key l;
- v <-
- alloc_vec_Vec_index_mut_back (List_t T) usize
- (core_slice_index_SliceIndexUsizeSliceTInst (List_t T))
- self.(hashMap_slots) hash_mod l0;
- Return
+ v <- index_mut_back l1;
+ Return (None,
{|
hashMap_num_entries := self.(hashMap_num_entries);
hashMap_max_load_factor := self.(hashMap_max_load_factor);
hashMap_max_load := self.(hashMap_max_load);
hashMap_slots := v
- |}
- | Some x0 =>
- i0 <- usize_sub self.(hashMap_num_entries) 1%usize;
- l0 <- hashMap_remove_from_list_back T n key l;
- v <-
- alloc_vec_Vec_index_mut_back (List_t T) usize
- (core_slice_index_SliceIndexUsizeSliceTInst (List_t T))
- self.(hashMap_slots) hash_mod l0;
- Return
+ |})
+ | Some x1 =>
+ i1 <- usize_sub self.(hashMap_num_entries) 1%usize;
+ v <- index_mut_back l1;
+ Return (Some x1,
{|
- hashMap_num_entries := i0;
+ hashMap_num_entries := i1;
hashMap_max_load_factor := self.(hashMap_max_load_factor);
hashMap_max_load := self.(hashMap_max_load);
hashMap_slots := v
- |}
+ |})
end
.
-(** [hashmap::test1]: forward function
+(** [hashmap::test1]:
Source: 'src/hashmap.rs', lines 315:0-315:10 *)
Definition test1 (n : nat) : result unit :=
hm <- hashMap_new u64 n;
- hm0 <- hashMap_insert u64 n hm 0%usize 42%u64;
- hm1 <- hashMap_insert u64 n hm0 128%usize 18%u64;
- hm2 <- hashMap_insert u64 n hm1 1024%usize 138%u64;
- hm3 <- hashMap_insert u64 n hm2 1056%usize 256%u64;
- i <- hashMap_get u64 n hm3 128%usize;
+ hm1 <- hashMap_insert u64 n hm 0%usize 42%u64;
+ hm2 <- hashMap_insert u64 n hm1 128%usize 18%u64;
+ hm3 <- hashMap_insert u64 n hm2 1024%usize 138%u64;
+ hm4 <- hashMap_insert u64 n hm3 1056%usize 256%u64;
+ i <- hashMap_get u64 n hm4 128%usize;
if negb (i s= 18%u64)
then Fail_ Failure
else (
- hm4 <- hashMap_get_mut_back u64 n hm3 1024%usize 56%u64;
- i0 <- hashMap_get u64 n hm4 1024%usize;
- if negb (i0 s= 56%u64)
+ p <- hashMap_get_mut u64 n hm4 1024%usize;
+ let (_, get_mut_back) := p in
+ hm5 <- get_mut_back 56%u64;
+ i1 <- hashMap_get u64 n hm5 1024%usize;
+ if negb (i1 s= 56%u64)
then Fail_ Failure
else (
- x <- hashMap_remove u64 n hm4 1024%usize;
+ p1 <- hashMap_remove u64 n hm5 1024%usize;
+ let (x, hm6) := p1 in
match x with
| None => Fail_ Failure
- | Some x0 =>
- if negb (x0 s= 56%u64)
+ | Some x1 =>
+ if negb (x1 s= 56%u64)
then Fail_ Failure
else (
- hm5 <- hashMap_remove_back u64 n hm4 1024%usize;
- i1 <- hashMap_get u64 n hm5 0%usize;
- if negb (i1 s= 42%u64)
+ i2 <- hashMap_get u64 n hm6 0%usize;
+ if negb (i2 s= 42%u64)
then Fail_ Failure
else (
- i2 <- hashMap_get u64 n hm5 128%usize;
- if negb (i2 s= 18%u64)
+ i3 <- hashMap_get u64 n hm6 128%usize;
+ if negb (i3 s= 18%u64)
then Fail_ Failure
else (
- i3 <- hashMap_get u64 n hm5 1056%usize;
- if negb (i3 s= 256%u64) then Fail_ Failure else Return tt)))
+ i4 <- hashMap_get u64 n hm6 1056%usize;
+ if negb (i4 s= 256%u64) then Fail_ Failure else Return tt)))
end))
.