From 4a34537a7fe33fa33d618b153832f9e750276480 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Fri, 14 Jun 2024 10:41:06 +0200 Subject: Update the code of the hashmap --- tests/coq/hashmap/Hashmap_Funs.v | 250 +++++++++++++++++++-------------------- 1 file changed, 125 insertions(+), 125 deletions(-) (limited to 'tests/coq/hashmap/Hashmap_Funs.v') diff --git a/tests/coq/hashmap/Hashmap_Funs.v b/tests/coq/hashmap/Hashmap_Funs.v index 02aa0269..d63c6f72 100644 --- a/tests/coq/hashmap/Hashmap_Funs.v +++ b/tests/coq/hashmap/Hashmap_Funs.v @@ -13,22 +13,22 @@ Include Hashmap_FunsExternal. Module Hashmap_Funs. (** [hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 37:0-37:32 *) + Source: 'tests/src/hashmap.rs', lines 38:0-38:32 *) Definition hash_key (k : usize) : result usize := Ok k. (** [hashmap::{hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 60:4-66:5 *) + Source: 'tests/src/hashmap.rs', lines 61:4-67:5 *) Fixpoint hashMap_allocate_slots_loop - (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n1 : usize) : - result (alloc_vec_Vec (List_t T)) + (T : Type) (n : nat) (slots : alloc_vec_Vec (AList_t T)) (n1 : usize) : + result (alloc_vec_Vec (AList_t T)) := match n with | O => Fail_ OutOfFuel | S n2 => if n1 s> 0%usize then ( - slots1 <- alloc_vec_Vec_push (List_t T) slots List_Nil; + slots1 <- alloc_vec_Vec_push (AList_t T) slots AList_Nil; n3 <- usize_sub n1 1%usize; hashMap_allocate_slots_loop T n2 slots1 n3) else Ok slots @@ -36,22 +36,22 @@ Fixpoint hashMap_allocate_slots_loop . (** [hashmap::{hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 60:4-60:76 *) + Source: 'tests/src/hashmap.rs', lines 61:4-61:78 *) Definition hashMap_allocate_slots - (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n1 : usize) : - result (alloc_vec_Vec (List_t T)) + (T : Type) (n : nat) (slots : alloc_vec_Vec (AList_t T)) (n1 : usize) : + result (alloc_vec_Vec (AList_t T)) := hashMap_allocate_slots_loop T n slots n1 . (** [hashmap::{hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 69:4-73:13 *) + Source: 'tests/src/hashmap.rs', lines 70:4-74:13 *) Definition hashMap_new_with_capacity (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : result (HashMap_t T) := - slots <- hashMap_allocate_slots T n (alloc_vec_Vec_new (List_t T)) capacity; + slots <- hashMap_allocate_slots T n (alloc_vec_Vec_new (AList_t T)) capacity; i <- usize_mul capacity max_load_dividend; i1 <- usize_div i max_load_divisor; Ok @@ -64,36 +64,36 @@ Definition hashMap_new_with_capacity . (** [hashmap::{hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 85:4-85:24 *) + Source: 'tests/src/hashmap.rs', lines 86:4-86: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}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 90:4-98:5 *) + Source: 'tests/src/hashmap.rs', lines 91:4-99:5 *) Fixpoint hashMap_clear_loop - (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (i : usize) : - result (alloc_vec_Vec (List_t T)) + (T : Type) (n : nat) (slots : alloc_vec_Vec (AList_t T)) (i : usize) : + result (alloc_vec_Vec (AList_t T)) := match n with | O => Fail_ OutOfFuel | S n1 => - let i1 := alloc_vec_Vec_len (List_t T) slots in + let i1 := alloc_vec_Vec_len (AList_t T) slots in if i s< i1 then ( p <- - alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) slots i; + alloc_vec_Vec_index_mut (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) slots i; let (_, index_mut_back) := p in i2 <- usize_add i 1%usize; - slots1 <- index_mut_back List_Nil; + slots1 <- index_mut_back AList_Nil; hashMap_clear_loop T n1 slots1 i2) else Ok slots end . (** [hashmap::{hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 90:4-90:27 *) + Source: 'tests/src/hashmap.rs', lines 91:4-91:27 *) Definition hashMap_clear (T : Type) (n : nat) (self : HashMap_t T) : result (HashMap_t T) := hm <- hashMap_clear_loop T n self.(hashMap_slots) 0%usize; @@ -107,62 +107,62 @@ Definition hashMap_clear . (** [hashmap::{hashmap::HashMap}::len]: - Source: 'tests/src/hashmap.rs', lines 100:4-100:30 *) + Source: 'tests/src/hashmap.rs', lines 101:4-101:30 *) Definition hashMap_len (T : Type) (self : HashMap_t T) : result usize := Ok self.(hashMap_num_entries) . (** [hashmap::{hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 107:4-124:5 *) + Source: 'tests/src/hashmap.rs', lines 108:4-125:5 *) Fixpoint hashMap_insert_in_list_loop - (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : - result (bool * (List_t T)) + (T : Type) (n : nat) (key : usize) (value : T) (ls : AList_t T) : + result (bool * (AList_t T)) := match n with | O => Fail_ OutOfFuel | S n1 => match ls with - | List_Cons ckey cvalue tl => + | AList_Cons ckey cvalue tl => if ckey s= key - then Ok (false, List_Cons ckey value tl) + then Ok (false, AList_Cons ckey value tl) else ( p <- hashMap_insert_in_list_loop T n1 key value tl; let (b, tl1) := p in - Ok (b, List_Cons ckey cvalue tl1)) - | List_Nil => Ok (true, List_Cons key value List_Nil) + Ok (b, AList_Cons ckey cvalue tl1)) + | AList_Nil => Ok (true, AList_Cons key value AList_Nil) end end . (** [hashmap::{hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 107:4-107:71 *) + Source: 'tests/src/hashmap.rs', lines 108:4-108:72 *) Definition hashMap_insert_in_list - (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : - result (bool * (List_t T)) + (T : Type) (n : nat) (key : usize) (value : T) (ls : AList_t T) : + result (bool * (AList_t T)) := hashMap_insert_in_list_loop T n key value ls . (** [hashmap::{hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 127:4-127:54 *) + Source: 'tests/src/hashmap.rs', lines 128:4-128:54 *) Definition hashMap_insert_no_resize (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) : result (HashMap_t T) := hash <- hash_key key; - let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in + let i := alloc_vec_Vec_len (AList_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; p <- - alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) + alloc_vec_Vec_index_mut (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) self.(hashMap_slots) hash_mod; - let (l, index_mut_back) := p in - p1 <- hashMap_insert_in_list T n key value l; - let (inserted, l1) := p1 in + let (a, index_mut_back) := p in + p1 <- hashMap_insert_in_list T n key value a; + let (inserted, a1) := p1 in if inserted then ( i1 <- usize_add self.(hashMap_num_entries) 1%usize; - v <- index_mut_back l1; + v <- index_mut_back a1; Ok {| hashMap_num_entries := i1; @@ -171,7 +171,7 @@ Definition hashMap_insert_no_resize hashMap_slots := v |}) else ( - v <- index_mut_back l1; + v <- index_mut_back a1; Ok {| hashMap_num_entries := self.(hashMap_num_entries); @@ -182,74 +182,74 @@ Definition hashMap_insert_no_resize . (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 193:4-206:5 *) + Source: 'tests/src/hashmap.rs', lines 194:4-207:5 *) Fixpoint hashMap_move_elements_from_list_loop - (T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) : + (T : Type) (n : nat) (ntable : HashMap_t T) (ls : AList_t T) : result (HashMap_t T) := match n with | O => Fail_ OutOfFuel | S n1 => match ls with - | List_Cons k v tl => + | AList_Cons k v tl => ntable1 <- hashMap_insert_no_resize T n1 ntable k v; hashMap_move_elements_from_list_loop T n1 ntable1 tl - | List_Nil => Ok ntable + | AList_Nil => Ok ntable end end . (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 193:4-193:72 *) + Source: 'tests/src/hashmap.rs', lines 194:4-194:73 *) Definition hashMap_move_elements_from_list - (T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) : + (T : Type) (n : nat) (ntable : HashMap_t T) (ls : AList_t T) : result (HashMap_t T) := hashMap_move_elements_from_list_loop T n ntable ls . (** [hashmap::{hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 181:4-190:5 *) + Source: 'tests/src/hashmap.rs', lines 182:4-191:5 *) Fixpoint hashMap_move_elements_loop (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))) + (slots : alloc_vec_Vec (AList_t T)) (i : usize) : + result ((HashMap_t T) * (alloc_vec_Vec (AList_t T))) := match n with | O => Fail_ OutOfFuel | S n1 => - let i1 := alloc_vec_Vec_len (List_t T) slots in + let i1 := alloc_vec_Vec_len (AList_t T) slots in if i s< i1 then ( p <- - alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) slots i; - let (l, index_mut_back) := p in - let (ls, l1) := core_mem_replace (List_t T) l List_Nil in + alloc_vec_Vec_index_mut (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) slots i; + let (a, index_mut_back) := p in + let (ls, a1) := core_mem_replace (AList_t T) a AList_Nil in ntable1 <- hashMap_move_elements_from_list T n1 ntable ls; i2 <- usize_add i 1%usize; - slots1 <- index_mut_back l1; + slots1 <- index_mut_back a1; hashMap_move_elements_loop T n1 ntable1 slots1 i2) else Ok (ntable, slots) end . (** [hashmap::{hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 181:4-181:95 *) + Source: 'tests/src/hashmap.rs', lines 182:4-182:96 *) 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))) + (slots : alloc_vec_Vec (AList_t T)) (i : usize) : + result ((HashMap_t T) * (alloc_vec_Vec (AList_t T))) := hashMap_move_elements_loop T n ntable slots i . (** [hashmap::{hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 150:4-150:28 *) + Source: 'tests/src/hashmap.rs', lines 151:4-151: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 + let capacity := alloc_vec_Vec_len (AList_t T) self.(hashMap_slots) in n1 <- usize_div max_usize 2%usize; let (i, i1) := self.(hashMap_max_load_factor) in i2 <- usize_div n1 i; @@ -277,7 +277,7 @@ Definition hashMap_try_resize . (** [hashmap::{hashmap::HashMap}::insert]: - Source: 'tests/src/hashmap.rs', lines 139:4-139:48 *) + Source: 'tests/src/hashmap.rs', lines 140:4-140:48 *) Definition hashMap_insert (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) : result (HashMap_t T) @@ -290,134 +290,134 @@ Definition hashMap_insert . (** [hashmap::{hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 216:4-229:5 *) + Source: 'tests/src/hashmap.rs', lines 217:4-230:5 *) Fixpoint hashMap_contains_key_in_list_loop - (T : Type) (n : nat) (key : usize) (ls : List_t T) : result bool := + (T : Type) (n : nat) (key : usize) (ls : AList_t T) : result bool := match n with | O => Fail_ OutOfFuel | S n1 => match ls with - | List_Cons ckey _ tl => + | AList_Cons ckey _ tl => if ckey s= key then Ok true else hashMap_contains_key_in_list_loop T n1 key tl - | List_Nil => Ok false + | AList_Nil => Ok false end end . (** [hashmap::{hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 216:4-216:68 *) + Source: 'tests/src/hashmap.rs', lines 217:4-217:69 *) Definition hashMap_contains_key_in_list - (T : Type) (n : nat) (key : usize) (ls : List_t T) : result bool := + (T : Type) (n : nat) (key : usize) (ls : AList_t T) : result bool := hashMap_contains_key_in_list_loop T n key ls . (** [hashmap::{hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 209:4-209:49 *) + Source: 'tests/src/hashmap.rs', lines 210:4-210:49 *) Definition hashMap_contains_key (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result bool := hash <- hash_key key; - let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in + let i := alloc_vec_Vec_len (AList_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; - l <- - alloc_vec_Vec_index (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) + a <- + alloc_vec_Vec_index (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) self.(hashMap_slots) hash_mod; - hashMap_contains_key_in_list T n key l + hashMap_contains_key_in_list T n key a . (** [hashmap::{hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 234:4-247:5 *) + Source: 'tests/src/hashmap.rs', lines 235:4-248:5 *) Fixpoint hashMap_get_in_list_loop - (T : Type) (n : nat) (key : usize) (ls : List_t T) : result T := + (T : Type) (n : nat) (key : usize) (ls : AList_t T) : result T := match n with | O => Fail_ OutOfFuel | S n1 => match ls with - | List_Cons ckey cvalue tl => + | AList_Cons ckey cvalue tl => if ckey s= key then Ok cvalue else hashMap_get_in_list_loop T n1 key tl - | List_Nil => Fail_ Failure + | AList_Nil => Fail_ Failure end end . (** [hashmap::{hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 234:4-234:70 *) + Source: 'tests/src/hashmap.rs', lines 235:4-235:71 *) Definition hashMap_get_in_list - (T : Type) (n : nat) (key : usize) (ls : List_t T) : result T := + (T : Type) (n : nat) (key : usize) (ls : AList_t T) : result T := hashMap_get_in_list_loop T n key ls . (** [hashmap::{hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 249:4-249:55 *) + Source: 'tests/src/hashmap.rs', lines 250:4-250:55 *) Definition hashMap_get (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 + let i := alloc_vec_Vec_len (AList_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; - l <- - alloc_vec_Vec_index (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) + a <- + alloc_vec_Vec_index (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) self.(hashMap_slots) hash_mod; - hashMap_get_in_list T n key l + hashMap_get_in_list T n key a . (** [hashmap::{hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 255:4-264:5 *) + Source: 'tests/src/hashmap.rs', lines 256:4-265:5 *) Fixpoint hashMap_get_mut_in_list_loop - (T : Type) (n : nat) (ls : List_t T) (key : usize) : - result (T * (T -> result (List_t T))) + (T : Type) (n : nat) (ls : AList_t T) (key : usize) : + result (T * (T -> result (AList_t T))) := match n with | O => Fail_ OutOfFuel | S n1 => match ls with - | List_Cons ckey cvalue tl => + | AList_Cons ckey cvalue tl => if ckey s= key then - let back := fun (ret : T) => Ok (List_Cons ckey ret tl) in + let back := fun (ret : T) => Ok (AList_Cons ckey ret tl) in Ok (cvalue, back) else ( p <- hashMap_get_mut_in_list_loop T n1 tl key; let (t, back) := p in let back1 := - fun (ret : T) => tl1 <- back ret; Ok (List_Cons ckey cvalue tl1) in + fun (ret : T) => tl1 <- back ret; Ok (AList_Cons ckey cvalue tl1) in Ok (t, back1)) - | List_Nil => Fail_ Failure + | AList_Nil => Fail_ Failure end end . (** [hashmap::{hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 255:4-255:86 *) + Source: 'tests/src/hashmap.rs', lines 256:4-256:87 *) Definition hashMap_get_mut_in_list - (T : Type) (n : nat) (ls : List_t T) (key : usize) : - result (T * (T -> result (List_t T))) + (T : Type) (n : nat) (ls : AList_t T) (key : usize) : + result (T * (T -> result (AList_t T))) := hashMap_get_mut_in_list_loop T n ls key . (** [hashmap::{hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 267:4-267:67 *) + Source: 'tests/src/hashmap.rs', lines 268:4-268:67 *) Definition hashMap_get_mut (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 + let i := alloc_vec_Vec_len (AList_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; p <- - alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) + alloc_vec_Vec_index_mut (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) self.(hashMap_slots) hash_mod; - let (l, index_mut_back) := p in - p1 <- hashMap_get_mut_in_list T n l key; + let (a, index_mut_back) := p in + p1 <- hashMap_get_mut_in_list T n a key; let (t, get_mut_in_list_back) := p1 in let back := fun (ret : T) => - l1 <- get_mut_in_list_back ret; - v <- index_mut_back l1; + a1 <- get_mut_in_list_back ret; + v <- index_mut_back a1; Ok {| hashMap_num_entries := self.(hashMap_num_entries); @@ -429,61 +429,61 @@ Definition hashMap_get_mut . (** [hashmap::{hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 275:4-301:5 *) + Source: 'tests/src/hashmap.rs', lines 276:4-302:5 *) Fixpoint hashMap_remove_from_list_loop - (T : Type) (n : nat) (key : usize) (ls : List_t T) : - result ((option T) * (List_t T)) + (T : Type) (n : nat) (key : usize) (ls : AList_t T) : + result ((option T) * (AList_t T)) := match n with | O => Fail_ OutOfFuel | S n1 => match ls with - | List_Cons ckey t tl => + | AList_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 + core_mem_replace (AList_t T) (AList_Cons ckey t tl) AList_Nil in match mv_ls with - | List_Cons _ cvalue tl1 => Ok (Some cvalue, tl1) - | List_Nil => Fail_ Failure + | AList_Cons _ cvalue tl1 => Ok (Some cvalue, tl1) + | AList_Nil => Fail_ Failure end else ( p <- hashMap_remove_from_list_loop T n1 key tl; let (o, tl1) := p in - Ok (o, List_Cons ckey t tl1)) - | List_Nil => Ok (None, List_Nil) + Ok (o, AList_Cons ckey t tl1)) + | AList_Nil => Ok (None, AList_Nil) end end . (** [hashmap::{hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 275:4-275:69 *) + Source: 'tests/src/hashmap.rs', lines 276:4-276:70 *) Definition hashMap_remove_from_list - (T : Type) (n : nat) (key : usize) (ls : List_t T) : - result ((option T) * (List_t T)) + (T : Type) (n : nat) (key : usize) (ls : AList_t T) : + result ((option T) * (AList_t T)) := hashMap_remove_from_list_loop T n key ls . (** [hashmap::{hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 304:4-304:52 *) + Source: 'tests/src/hashmap.rs', lines 305:4-305:52 *) Definition hashMap_remove (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result ((option T) * (HashMap_t T)) := hash <- hash_key key; - let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in + let i := alloc_vec_Vec_len (AList_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; p <- - alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) + alloc_vec_Vec_index_mut (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) self.(hashMap_slots) hash_mod; - let (l, index_mut_back) := p in - p1 <- hashMap_remove_from_list T n key l; - let (x, l1) := p1 in + let (a, index_mut_back) := p in + p1 <- hashMap_remove_from_list T n key a; + let (x, a1) := p1 in match x with | None => - v <- index_mut_back l1; + v <- index_mut_back a1; Ok (None, {| hashMap_num_entries := self.(hashMap_num_entries); @@ -493,7 +493,7 @@ Definition hashMap_remove |}) | Some x1 => i1 <- usize_sub self.(hashMap_num_entries) 1%usize; - v <- index_mut_back l1; + v <- index_mut_back a1; Ok (Some x1, {| hashMap_num_entries := i1; @@ -505,7 +505,7 @@ Definition hashMap_remove . (** [hashmap::insert_on_disk]: - Source: 'tests/src/hashmap.rs', lines 335:0-335:43 *) + Source: 'tests/src/hashmap.rs', lines 336:0-336:43 *) Definition insert_on_disk (n : nat) (key : usize) (value : u64) (st : state) : result (state * unit) := p <- utils_deserialize st; @@ -515,7 +515,7 @@ Definition insert_on_disk . (** [hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 350:0-350:10 *) + Source: 'tests/src/hashmap.rs', lines 351:0-351:10 *) Definition test1 (n : nat) : result unit := hm <- hashMap_new u64 n; hm1 <- hashMap_insert u64 n hm 0%usize 42%u64; -- cgit v1.2.3