summaryrefslogtreecommitdiff
path: root/tests/coq/hashmap_on_disk
diff options
context:
space:
mode:
Diffstat (limited to 'tests/coq/hashmap_on_disk')
-rw-r--r--tests/coq/hashmap_on_disk/HashmapMain_Funs.v182
-rw-r--r--tests/coq/hashmap_on_disk/HashmapMain_Opaque.v6
-rw-r--r--tests/coq/hashmap_on_disk/HashmapMain_Types.v6
-rw-r--r--tests/coq/hashmap_on_disk/Primitives.v88
4 files changed, 167 insertions, 115 deletions
diff --git a/tests/coq/hashmap_on_disk/HashmapMain_Funs.v b/tests/coq/hashmap_on_disk/HashmapMain_Funs.v
index eac78186..46d3ee29 100644
--- a/tests/coq/hashmap_on_disk/HashmapMain_Funs.v
+++ b/tests/coq/hashmap_on_disk/HashmapMain_Funs.v
@@ -12,11 +12,13 @@ Require Export HashmapMain_Opaque.
Import HashmapMain_Opaque.
Module HashmapMain_Funs.
-(** [hashmap_main::hashmap::hash_key]: forward function *)
+(** [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::{0}::allocate_slots]: loop 0: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
:
@@ -34,7 +36,8 @@ Fixpoint hashmap_HashMap_allocate_slots_loop
end
.
-(** [hashmap_main::hashmap::HashMap::{0}::allocate_slots]: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
:
@@ -43,7 +46,8 @@ Definition hashmap_HashMap_allocate_slots
hashmap_HashMap_allocate_slots_loop T n slots n0
.
-(** [hashmap_main::hashmap::HashMap::{0}::new_with_capacity]: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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) :
@@ -62,14 +66,16 @@ Definition hashmap_HashMap_new_with_capacity
|}
.
-(** [hashmap_main::hashmap::HashMap::{0}::new]: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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::{0}::clear]: loop 0: merged forward/backward function
- (there is a single backward function, and the forward function returns ()) *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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))
@@ -83,15 +89,16 @@ Fixpoint hashmap_HashMap_clear_loop
i1 <- usize_add i 1%usize;
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;
+ (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::{0}::clear]: merged forward/backward function
- (there is a single backward function, and the forward function returns ()) *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -106,13 +113,15 @@ Definition hashmap_HashMap_clear
|}
.
-(** [hashmap_main::hashmap::HashMap::{0}::len]: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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::{0}::insert_in_list]: loop 0: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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
@@ -130,7 +139,8 @@ Fixpoint hashmap_HashMap_insert_in_list_loop
end
.
-(** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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
@@ -138,7 +148,8 @@ Definition hashmap_HashMap_insert_in_list
hashmap_HashMap_insert_in_list_loop T n key value ls
.
-(** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: loop 0: backward function 0 *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -159,7 +170,8 @@ Fixpoint hashmap_HashMap_insert_in_list_loop_back
end
.
-(** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: backward function 0 *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -167,8 +179,9 @@ Definition hashmap_HashMap_insert_in_list_back
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 ()) *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -178,7 +191,7 @@ Definition hashmap_HashMap_insert_no_resize
hash_mod <- usize_rem hash i;
l <-
alloc_vec_Vec_index_mut (hashmap_List_t T) usize
- (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ (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
@@ -187,8 +200,8 @@ Definition hashmap_HashMap_insert_no_resize
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_usize_coresliceindexSliceIndexInst (hashmap_List_t
- T)) self.(hashmap_HashMap_slots) hash_mod l0;
+ (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T))
+ self.(hashmap_HashMap_slots) hash_mod l0;
Return
{|
hashmap_HashMap_num_entries := i0;
@@ -201,8 +214,8 @@ Definition hashmap_HashMap_insert_no_resize
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_usize_coresliceindexSliceIndexInst (hashmap_List_t
- T)) self.(hashmap_HashMap_slots) hash_mod l0;
+ (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);
@@ -213,8 +226,9 @@ Definition hashmap_HashMap_insert_no_resize
|})
.
-(** [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 ()) *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -231,8 +245,9 @@ Fixpoint hashmap_HashMap_move_elements_from_list_loop
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 ()) *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -240,8 +255,9 @@ Definition hashmap_HashMap_move_elements_from_list
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 ()) *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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) :
@@ -255,23 +271,24 @@ Fixpoint hashmap_HashMap_move_elements_loop
then (
l <-
alloc_vec_Vec_index_mut (hashmap_List_t T) usize
- (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t
- T)) slots i;
+ (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_usize_coresliceindexSliceIndexInst (hashmap_List_t
- T)) slots i l0;
+ (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::{0}::move_elements]: merged forward/backward function
- (there is a single backward function, and the forward function returns ()) *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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) :
@@ -280,8 +297,9 @@ Definition hashmap_HashMap_move_elements
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 ()) *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -317,8 +335,9 @@ Definition hashmap_HashMap_try_resize
|}
.
-(** [hashmap_main::hashmap::HashMap::{0}::insert]: merged forward/backward function
- (there is a single backward function, and the forward function returns ()) *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -330,7 +349,8 @@ Definition hashmap_HashMap_insert
else Return self0
.
-(** [hashmap_main::hashmap::HashMap::{0}::contains_key_in_list]: loop 0: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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
@@ -346,13 +366,15 @@ Fixpoint hashmap_HashMap_contains_key_in_list_loop
end
.
-(** [hashmap_main::hashmap::HashMap::{0}::contains_key_in_list]: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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::{0}::contains_key]: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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
@@ -362,12 +384,13 @@ Definition hashmap_HashMap_contains_key
hash_mod <- usize_rem hash i;
l <-
alloc_vec_Vec_index (hashmap_List_t T) usize
- (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ (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::{0}::get_in_list]: loop 0: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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
@@ -383,13 +406,15 @@ Fixpoint hashmap_HashMap_get_in_list_loop
end
.
-(** [hashmap_main::hashmap::HashMap::{0}::get_in_list]: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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::{0}::get]: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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;
@@ -397,12 +422,13 @@ Definition hashmap_HashMap_get
hash_mod <- usize_rem hash i;
l <-
alloc_vec_Vec_index (hashmap_List_t T) usize
- (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ (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::{0}::get_mut_in_list]: loop 0: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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
@@ -418,13 +444,15 @@ Fixpoint hashmap_HashMap_get_mut_in_list_loop
end
.
-(** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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::{0}::get_mut_in_list]: loop 0: backward function 0 *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -444,7 +472,8 @@ Fixpoint hashmap_HashMap_get_mut_in_list_loop_back
end
.
-(** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: backward function 0 *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -452,7 +481,8 @@ Definition hashmap_HashMap_get_mut_in_list_back
hashmap_HashMap_get_mut_in_list_loop_back T n ls key ret
.
-(** [hashmap_main::hashmap::HashMap::{0}::get_mut]: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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;
@@ -460,12 +490,13 @@ Definition hashmap_HashMap_get_mut
hash_mod <- usize_rem hash i;
l <-
alloc_vec_Vec_index_mut (hashmap_List_t T) usize
- (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ (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::{0}::get_mut]: backward function 0 *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -475,12 +506,12 @@ Definition hashmap_HashMap_get_mut_back
hash_mod <- usize_rem hash i;
l <-
alloc_vec_Vec_index_mut (hashmap_List_t T) usize
- (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ (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_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T))
self.(hashmap_HashMap_slots) hash_mod l0;
Return
{|
@@ -491,7 +522,8 @@ Definition hashmap_HashMap_get_mut_back
|}
.
-(** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: loop 0: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -516,7 +548,8 @@ Fixpoint hashmap_HashMap_remove_from_list_loop
end
.
-(** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -524,7 +557,8 @@ Definition hashmap_HashMap_remove_from_list
hashmap_HashMap_remove_from_list_loop T n key ls
.
-(** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: loop 0: backward function 1 *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -551,7 +585,8 @@ Fixpoint hashmap_HashMap_remove_from_list_loop_back
end
.
-(** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: backward function 1 *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -559,7 +594,8 @@ Definition hashmap_HashMap_remove_from_list_back
hashmap_HashMap_remove_from_list_loop_back T n key ls
.
-(** [hashmap_main::hashmap::HashMap::{0}::remove]: forward function *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -569,7 +605,7 @@ Definition hashmap_HashMap_remove
hash_mod <- usize_rem hash i;
l <-
alloc_vec_Vec_index_mut (hashmap_List_t T) usize
- (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ (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
@@ -579,7 +615,8 @@ Definition hashmap_HashMap_remove
end
.
-(** [hashmap_main::hashmap::HashMap::{0}::remove]: backward function 0 *)
+(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap<T>}::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)
@@ -589,7 +626,7 @@ Definition hashmap_HashMap_remove_back
hash_mod <- usize_rem hash i;
l <-
alloc_vec_Vec_index_mut (hashmap_List_t T) usize
- (core_slice_index_usize_coresliceindexSliceIndexInst (hashmap_List_t T))
+ (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
@@ -597,8 +634,8 @@ Definition hashmap_HashMap_remove_back
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_usize_coresliceindexSliceIndexInst (hashmap_List_t
- T)) self.(hashmap_HashMap_slots) hash_mod l0;
+ (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);
@@ -612,8 +649,8 @@ Definition hashmap_HashMap_remove_back
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_usize_coresliceindexSliceIndexInst (hashmap_List_t
- T)) self.(hashmap_HashMap_slots) hash_mod l0;
+ (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T))
+ self.(hashmap_HashMap_slots) hash_mod l0;
Return
{|
hashmap_HashMap_num_entries := i0;
@@ -625,7 +662,8 @@ Definition hashmap_HashMap_remove_back
end
.
-(** [hashmap_main::hashmap::test1]: forward function *)
+(** [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;
@@ -662,7 +700,8 @@ Definition hashmap_test1 (n : nat) : result unit :=
end))
.
-(** [hashmap_main::insert_on_disk]: forward function *)
+(** [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;
@@ -673,7 +712,8 @@ Definition insert_on_disk
Return (st1, tt)
.
-(** [hashmap_main::main]: forward function *)
+(** [hashmap_main::main]: forward function
+ Source: 'src/hashmap_main.rs', lines 16:0-16:13 *)
Definition main : result unit :=
Return tt.
diff --git a/tests/coq/hashmap_on_disk/HashmapMain_Opaque.v b/tests/coq/hashmap_on_disk/HashmapMain_Opaque.v
index 5e376239..a0e9003d 100644
--- a/tests/coq/hashmap_on_disk/HashmapMain_Opaque.v
+++ b/tests/coq/hashmap_on_disk/HashmapMain_Opaque.v
@@ -10,12 +10,14 @@ Require Export HashmapMain_Types.
Import HashmapMain_Types.
Module HashmapMain_Opaque.
-(** [hashmap_main::hashmap_utils::deserialize]: forward function *)
+(** [hashmap_main::hashmap_utils::deserialize]: forward function
+ Source: 'src/hashmap_utils.rs', lines 10:0-10:43 *)
Axiom hashmap_utils_deserialize
: state -> result (state * (hashmap_HashMap_t u64))
.
-(** [hashmap_main::hashmap_utils::serialize]: forward function *)
+(** [hashmap_main::hashmap_utils::serialize]: forward function
+ Source: 'src/hashmap_utils.rs', lines 5:0-5:42 *)
Axiom hashmap_utils_serialize
: hashmap_HashMap_t u64 -> state -> result (state * unit)
.
diff --git a/tests/coq/hashmap_on_disk/HashmapMain_Types.v b/tests/coq/hashmap_on_disk/HashmapMain_Types.v
index 95e5f35b..039b7e72 100644
--- a/tests/coq/hashmap_on_disk/HashmapMain_Types.v
+++ b/tests/coq/hashmap_on_disk/HashmapMain_Types.v
@@ -8,7 +8,8 @@ Import ListNotations.
Local Open Scope Primitives_scope.
Module HashmapMain_Types.
-(** [hashmap_main::hashmap::List] *)
+(** [hashmap_main::hashmap::List]
+ Source: 'src/hashmap.rs', lines 19:0-19:16 *)
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
@@ -17,7 +18,8 @@ Inductive hashmap_List_t (T : Type) :=
Arguments Hashmap_List_Cons { _ }.
Arguments Hashmap_List_Nil { _ }.
-(** [hashmap_main::hashmap::HashMap] *)
+(** [hashmap_main::hashmap::HashMap]
+ Source: 'src/hashmap.rs', lines 35:0-35:21 *)
Record hashmap_HashMap_t (T : Type) :=
mkhashmap_HashMap_t {
hashmap_HashMap_num_entries : usize;
diff --git a/tests/coq/hashmap_on_disk/Primitives.v b/tests/coq/hashmap_on_disk/Primitives.v
index 85e38f01..83f860b6 100644
--- a/tests/coq/hashmap_on_disk/Primitives.v
+++ b/tests/coq/hashmap_on_disk/Primitives.v
@@ -467,14 +467,14 @@ 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 := {|
+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;
+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;
|}.
@@ -576,7 +576,7 @@ Definition alloc_vec_Vec_insert (T: Type) (v: alloc_vec_Vec T) (i: usize) (x: T)
else Fail_ Failure).
(* Helper *)
-Axiom alloc_vec_Vec_index_usize : forall {T : Type} (v : alloc_vec_Vec T) (i : usize) (x : T), result T.
+Axiom alloc_vec_Vec_index_usize : forall {T : Type} (v : alloc_vec_Vec T) (i : usize), 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).
@@ -620,18 +620,18 @@ Definition core_slice_index_Slice_index
end.
(* [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)).
+Axiom core_slice_index_RangeUsize_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 :
+Axiom core_slice_index_RangeUsize_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 :
+Axiom core_slice_index_RangeUsize_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
+Definition core_slice_index_RangeUsize_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
@@ -639,7 +639,7 @@ Definition core_slice_index_Range_get_unchecked
fun _ _ => Fail_ Failure.
(* [core::slice::index::Range::get_unchecked_mut]: forward function *)
-Definition core_slice_index_Range_get_unchecked_mut
+Definition core_slice_index_RangeUsize_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
@@ -647,15 +647,15 @@ Definition core_slice_index_Range_get_unchecked_mut
fun _ _ => Fail_ Failure.
(* [core::slice::index::Range::index]: forward function *)
-Axiom core_slice_index_Range_index :
+Axiom core_slice_index_RangeUsize_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 :
+Axiom core_slice_index_RangeUsize_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 :
+Axiom core_slice_index_RangeUsize_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 *)
@@ -683,44 +683,44 @@ 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
+Definition core_slice_index_private_slice_index_SealedRangeUsizeInst
: 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) :
+Definition core_slice_index_SliceIndexRangeUsizeSliceTInst (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_sealedInst := core_slice_index_private_slice_index_SealedRangeUsizeInst;
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;
+ core_slice_index_SliceIndex_get := core_slice_index_RangeUsize_get T;
+ core_slice_index_SliceIndex_get_mut := core_slice_index_RangeUsize_get_mut T;
+ core_slice_index_SliceIndex_get_mut_back := core_slice_index_RangeUsize_get_mut_back T;
+ core_slice_index_SliceIndex_get_unchecked := core_slice_index_RangeUsize_get_unchecked T;
+ core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_RangeUsize_get_unchecked_mut T;
+ core_slice_index_SliceIndex_index := core_slice_index_RangeUsize_index T;
+ core_slice_index_SliceIndex_index_mut := core_slice_index_RangeUsize_index_mut T;
+ core_slice_index_SliceIndex_index_mut_back := core_slice_index_RangeUsize_index_mut_back T;
+|}.
+
+(* Trait implementation: [core::slice::index::[T]] *)
+Definition core_ops_index_IndexSliceTIInst (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::[T]] *)
-Definition core_slice_index_Slice_coreopsindexIndexMutInst (T Idx : Type)
+Definition core_ops_index_IndexMutSliceTIInst (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_indexInst := core_ops_index_IndexSliceTIInst 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)
+Definition core_ops_index_IndexArrayInst (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);
@@ -728,10 +728,10 @@ Definition core_array_Array_coreopsindexIndexInst (T Idx : Type) (N : usize)
|}.
(* Trait implementation: [core::array::[T; N]] *)
-Definition core_array_Array_coreopsindexIndexMutInst (T Idx : Type) (N : usize)
+Definition core_ops_index_IndexMutArrayInst (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_indexInst := core_ops_index_IndexArrayInst 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;
|}.
@@ -765,13 +765,13 @@ 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
+Definition core_slice_index_private_slice_index_SealedUsizeInst
: core_slice_index_private_slice_index_Sealed usize := tt.
(* Trait implementation: [core::slice::index::usize] *)
-Definition core_slice_index_usize_coresliceindexSliceIndexInst (T : Type) :
+Definition core_slice_index_SliceIndexUsizeSliceTInst (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_sealedInst := core_slice_index_private_slice_index_SealedUsizeInst;
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;
@@ -815,8 +815,16 @@ Definition alloc_vec_Vec_coreopsindexIndexMutInst (T Idx : Type)
(*** Theorems *)
+Axiom alloc_vec_Vec_index_eq : forall {a : Type} (v : alloc_vec_Vec a) (i : usize) (x : a),
+ alloc_vec_Vec_index a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i =
+ alloc_vec_Vec_index_usize v i.
+
+Axiom alloc_vec_Vec_index_mut_eq : forall {a : Type} (v : alloc_vec_Vec a) (i : usize) (x : a),
+ alloc_vec_Vec_index_mut a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i =
+ alloc_vec_Vec_index_usize v i.
+
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_index_mut_back a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i x =
alloc_vec_Vec_update_usize v i x.
End Primitives.