(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) (** [hashmap_main]: function definitions *) Require Import Primitives. Import Primitives. Require Import Coq.ZArith.ZArith. Require Import List. Import ListNotations. Local Open Scope Primitives_scope. Require Export HashmapMain_Types. Import HashmapMain_Types. Require Export HashmapMain_Opaque. Import HashmapMain_Opaque. Module HashmapMain_Funs. (** [hashmap_main::hashmap::hash_key]: forward function Source: 'src/hashmap.rs', lines 27:0-27:32 *) Definition hashmap_hash_key (k : usize) : result usize := Return k. (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: loop 0: forward function Source: 'src/hashmap.rs', lines 50:4-56:5 *) 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 <- alloc_vec_Vec_push (hashmap_List_t T) slots Hashmap_List_Nil; n2 <- usize_sub n0 1%usize; hashmap_HashMap_allocate_slots_loop T n1 slots0 n2) else Return slots end . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: forward function Source: 'src/hashmap.rs', lines 50:4-50:76 *) 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_HashMap_allocate_slots_loop T n slots n0 . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new_with_capacity]: forward function Source: 'src/hashmap.rs', lines 59:4-63:13 *) Definition hashmap_HashMap_new_with_capacity (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : result (hashmap_HashMap_t T) := 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_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_main::hashmap::HashMap}::new]: forward function Source: 'src/hashmap.rs', lines 75:4-75:24 *) 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_main::hashmap::HashMap}::clear]: loop 0: merged forward/backward function (there is a single backward function, and the forward function returns ()) Source: 'src/hashmap.rs', lines 80:4-88:5 *) 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 := alloc_vec_Vec_len (hashmap_List_t T) slots in if i s< i0 then ( i1 <- usize_add i 1%usize; slots0 <- alloc_vec_Vec_index_mut_back (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (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_main::hashmap::HashMap}::clear]: merged forward/backward function (there is a single backward function, and the forward function returns ()) Source: 'src/hashmap.rs', lines 80:4-80:27 *) Definition hashmap_HashMap_clear (T : Type) (n : nat) (self : hashmap_HashMap_t T) : result (hashmap_HashMap_t T) := v <- hashmap_HashMap_clear_loop T n self.(hashmap_HashMap_slots) 0%usize; Return {| 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_main::hashmap::HashMap}::len]: forward function Source: 'src/hashmap.rs', lines 90:4-90:30 *) Definition hashmap_HashMap_len (T : Type) (self : hashmap_HashMap_t T) : result usize := Return self.(hashmap_HashMap_num_entries) . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: loop 0: forward function Source: 'src/hashmap.rs', lines 97:4-114:5 *) 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 | Hashmap_List_Cons ckey cvalue tl => if ckey s= key then Return false else hashmap_HashMap_insert_in_list_loop T n0 key value tl | Hashmap_List_Nil => Return true end end . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: forward function Source: 'src/hashmap.rs', lines 97:4-97:71 *) Definition hashmap_HashMap_insert_in_list (T : Type) (n : nat) (key : usize) (value : T) (ls : hashmap_List_t T) : result bool := hashmap_HashMap_insert_in_list_loop T n key value ls . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: loop 0: backward function 0 Source: 'src/hashmap.rs', lines 97:4-114:5 *) 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 | Hashmap_List_Cons ckey cvalue tl => if ckey s= key then Return (Hashmap_List_Cons ckey value tl) else ( 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_main::hashmap::HashMap}::insert_in_list]: backward function 0 Source: 'src/hashmap.rs', lines 97:4-97:71 *) 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_HashMap_insert_in_list_loop_back T n key value ls . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_no_resize]: merged forward/backward function (there is a single backward function, and the forward function returns ()) Source: 'src/hashmap.rs', lines 117:4-117:54 *) 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 key; let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in hash_mod <- usize_rem hash i; l <- alloc_vec_Vec_index_mut (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (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_HashMap_num_entries) 1%usize; l0 <- hashmap_HashMap_insert_in_list_back T n key value l; v <- alloc_vec_Vec_index_mut_back (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) self.(hashmap_HashMap_slots) hash_mod l0; Return {| 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_HashMap_insert_in_list_back T n key value l; v <- alloc_vec_Vec_index_mut_back (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) self.(hashmap_HashMap_slots) hash_mod l0; Return {| 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_main::hashmap::HashMap}::move_elements_from_list]: loop 0: merged forward/backward function (there is a single backward function, and the forward function returns ()) Source: 'src/hashmap.rs', lines 183:4-196:5 *) 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 | 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_main::hashmap::HashMap}::move_elements_from_list]: merged forward/backward function (there is a single backward function, and the forward function returns ()) Source: 'src/hashmap.rs', lines 183:4-183:72 *) 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_HashMap_move_elements_from_list_loop T n ntable ls . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: loop 0: merged forward/backward function (there is a single backward function, and the forward function returns ()) Source: 'src/hashmap.rs', lines 171:4-180:5 *) 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 := alloc_vec_Vec_len (hashmap_List_t T) slots in if i s< i0 then ( l <- alloc_vec_Vec_index_mut (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (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 := 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_SliceIndexUsizeSliceTInst (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_main::hashmap::HashMap}::move_elements]: merged forward/backward function (there is a single backward function, and the forward function returns ()) Source: 'src/hashmap.rs', lines 171:4-171:95 *) 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_HashMap_move_elements_loop T n ntable slots i . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::try_resize]: merged forward/backward function (there is a single backward function, and the forward function returns ()) Source: 'src/hashmap.rs', lines 140:4-140:28 *) 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 := 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_HashMap_max_load_factor) in i1 <- usize_div n1 i; if capacity s<= i1 then ( i2 <- usize_mul capacity 2%usize; ntable <- hashmap_HashMap_new_with_capacity T n i2 i i0; p <- hashmap_HashMap_move_elements T n ntable self.(hashmap_HashMap_slots) 0%usize; let (ntable0, _) := p in Return {| 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_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_main::hashmap::HashMap}::insert]: merged forward/backward function (there is a single backward function, and the forward function returns ()) Source: 'src/hashmap.rs', lines 129:4-129:48 *) Definition hashmap_HashMap_insert (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) (value : T) : result (hashmap_HashMap_t T) := 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_main::hashmap::HashMap}::contains_key_in_list]: loop 0: forward function Source: 'src/hashmap.rs', lines 206:4-219:5 *) 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 | Hashmap_List_Cons ckey t tl => if ckey s= key then Return true else hashmap_HashMap_contains_key_in_list_loop T n0 key tl | Hashmap_List_Nil => Return false end end . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: forward function Source: 'src/hashmap.rs', lines 206:4-206:68 *) 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_main::hashmap::HashMap}::contains_key]: forward function Source: 'src/hashmap.rs', lines 199:4-199:49 *) Definition hashmap_HashMap_contains_key (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : result bool := 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 <- alloc_vec_Vec_index (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) self.(hashmap_HashMap_slots) hash_mod; hashmap_HashMap_contains_key_in_list T n key l . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: loop 0: forward function Source: 'src/hashmap.rs', lines 224:4-237:5 *) 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 | Hashmap_List_Cons ckey cvalue tl => if ckey s= key then Return cvalue else hashmap_HashMap_get_in_list_loop T n0 key tl | Hashmap_List_Nil => Fail_ Failure end end . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: forward function Source: 'src/hashmap.rs', lines 224:4-224:70 *) 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_main::hashmap::HashMap}::get]: forward function Source: 'src/hashmap.rs', lines 239:4-239:55 *) 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 <- alloc_vec_Vec_index (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) self.(hashmap_HashMap_slots) hash_mod; hashmap_HashMap_get_in_list T n key l . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: loop 0: forward function Source: 'src/hashmap.rs', lines 245:4-254:5 *) 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 | Hashmap_List_Cons ckey cvalue tl => if ckey s= key then Return cvalue else hashmap_HashMap_get_mut_in_list_loop T n0 tl key | Hashmap_List_Nil => Fail_ Failure end end . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: forward function Source: 'src/hashmap.rs', lines 245:4-245:86 *) 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_main::hashmap::HashMap}::get_mut_in_list]: loop 0: backward function 0 Source: 'src/hashmap.rs', lines 245:4-254:5 *) 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 | Hashmap_List_Cons ckey cvalue tl => if ckey s= key then Return (Hashmap_List_Cons ckey ret tl) else ( 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_main::hashmap::HashMap}::get_mut_in_list]: backward function 0 Source: 'src/hashmap.rs', lines 245:4-245:86 *) 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_HashMap_get_mut_in_list_loop_back T n ls key ret . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut]: forward function Source: 'src/hashmap.rs', lines 257:4-257:67 *) 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 <- alloc_vec_Vec_index_mut (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) self.(hashmap_HashMap_slots) hash_mod; hashmap_HashMap_get_mut_in_list T n l key . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut]: backward function 0 Source: 'src/hashmap.rs', lines 257:4-257:67 *) 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 key; let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in hash_mod <- usize_rem hash i; l <- alloc_vec_Vec_index_mut (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) self.(hashmap_HashMap_slots) hash_mod; l0 <- hashmap_HashMap_get_mut_in_list_back T n l key ret; v <- alloc_vec_Vec_index_mut_back (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) self.(hashmap_HashMap_slots) hash_mod l0; Return {| 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_main::hashmap::HashMap}::remove_from_list]: loop 0: forward function Source: 'src/hashmap.rs', lines 265:4-291:5 *) 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 | Hashmap_List_Cons ckey t tl => if ckey s= key then let mv_ls := core_mem_replace (hashmap_List_t T) (Hashmap_List_Cons ckey t tl) Hashmap_List_Nil in match mv_ls with | Hashmap_List_Cons i cvalue tl0 => Return (Some cvalue) | Hashmap_List_Nil => Fail_ Failure end else hashmap_HashMap_remove_from_list_loop T n0 key tl | Hashmap_List_Nil => Return None end end . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: forward function Source: 'src/hashmap.rs', lines 265:4-265:69 *) Definition hashmap_HashMap_remove_from_list (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result (option T) := hashmap_HashMap_remove_from_list_loop T n key ls . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: loop 0: backward function 1 Source: 'src/hashmap.rs', lines 265:4-291:5 *) 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 | Hashmap_List_Cons ckey t tl => if ckey s= key then let mv_ls := core_mem_replace (hashmap_List_t T) (Hashmap_List_Cons ckey t tl) Hashmap_List_Nil in match mv_ls with | Hashmap_List_Cons i cvalue tl0 => Return tl0 | Hashmap_List_Nil => Fail_ Failure end else ( 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_main::hashmap::HashMap}::remove_from_list]: backward function 1 Source: 'src/hashmap.rs', lines 265:4-265:69 *) Definition hashmap_HashMap_remove_from_list_back (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result (hashmap_List_t T) := hashmap_HashMap_remove_from_list_loop_back T n key ls . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove]: forward function Source: 'src/hashmap.rs', lines 294:4-294:52 *) Definition hashmap_HashMap_remove (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : result (option 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 <- alloc_vec_Vec_index_mut (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (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_HashMap_num_entries) 1%usize; Return (Some x0) end . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove]: backward function 0 Source: 'src/hashmap.rs', lines 294:4-294:52 *) 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 key; let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in hash_mod <- usize_rem hash i; l <- alloc_vec_Vec_index_mut (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (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_HashMap_remove_from_list_back T n key l; v <- alloc_vec_Vec_index_mut_back (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) self.(hashmap_HashMap_slots) hash_mod l0; Return {| 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_HashMap_num_entries) 1%usize; l0 <- hashmap_HashMap_remove_from_list_back T n key l; v <- alloc_vec_Vec_index_mut_back (hashmap_List_t T) usize (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) self.(hashmap_HashMap_slots) hash_mod l0; Return {| 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 Source: 'src/hashmap.rs', lines 315:0-315:10 *) 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_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_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_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_HashMap_get u64 n hm5 128%usize; if negb (i2 s= 18%u64) then Fail_ Failure else ( 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 Source: 'src/hashmap_main.rs', lines 7:0-7:43 *) Definition insert_on_disk (n : nat) (key : usize) (value : u64) (st : state) : result (state * unit) := p <- hashmap_utils_deserialize st; let (st0, hm) := p in 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 Source: 'src/hashmap_main.rs', lines 16:0-16:13 *) Definition main : result unit := Return tt. End HashmapMain_Funs .