From ce8f5c8f67e41a74bfdf8f6d664ff4e45e9de850 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Thu, 7 Sep 2023 16:06:14 +0200 Subject: Regenerate the test files and fix a proof --- tests/coq/hashmap/Hashmap_Funs.v | 6 +----- tests/coq/hashmap/Primitives.v | 14 ++++++++++++++ 2 files changed, 15 insertions(+), 5 deletions(-) (limited to 'tests/coq/hashmap') diff --git a/tests/coq/hashmap/Hashmap_Funs.v b/tests/coq/hashmap/Hashmap_Funs.v index e950ba0b..054880d4 100644 --- a/tests/coq/hashmap/Hashmap_Funs.v +++ b/tests/coq/hashmap/Hashmap_Funs.v @@ -190,10 +190,6 @@ Definition hash_map_insert_no_resize_fwd_back |}) . -(** [core::num::u32::{8}::MAX] *) -Definition core_num_u32_max_body : result u32 := Return 4294967295%u32. -Definition core_num_u32_max_c : u32 := core_num_u32_max_body%global. - (** [hashmap::HashMap::{0}::move_elements_from_list]: loop 0: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) Fixpoint hash_map_move_elements_from_list_loop_fwd_back @@ -259,7 +255,7 @@ Definition hash_map_move_elements_fwd_back (there is a single backward function, and the forward function returns ()) *) Definition hash_map_try_resize_fwd_back (T : Type) (n : nat) (self : Hash_map_t T) : result (Hash_map_t T) := - max_usize <- scalar_cast U32 Usize core_num_u32_max_c; + max_usize <- scalar_cast U32 Usize core_u32_max; let capacity := vec_len (List_t T) self.(Hash_map_slots) in n1 <- usize_div max_usize 2%usize; let (i, i0) := self.(Hash_map_max_load_factor) in diff --git a/tests/coq/hashmap/Primitives.v b/tests/coq/hashmap/Primitives.v index 71a2d9c3..8d6c9c8d 100644 --- a/tests/coq/hashmap/Primitives.v +++ b/tests/coq/hashmap/Primitives.v @@ -394,6 +394,20 @@ Notation "x s< y" := (scalar_ltb x y) (at level 80) : Primitives_scope. Notation "x s>= y" := (scalar_geb x y) (at level 80) : Primitives_scope. Notation "x s> y" := (scalar_gtb x y) (at level 80) : Primitives_scope. +(** Constants *) +Definition core_u8_max := u8_max %u32. +Definition core_u16_max := u16_max %u32. +Definition core_u32_max := u32_max %u32. +Definition core_u64_max := u64_max %u64. +Definition core_u128_max := u64_max %u128. +Axiom core_usize_max : usize. (** TODO *) +Definition core_i8_max := i8_max %i32. +Definition core_i16_max := i16_max %i32. +Definition core_i32_max := i32_max %i32. +Definition core_i64_max := i64_max %i64. +Definition core_i128_max := i64_max %i128. +Axiom core_isize_max : isize. (** TODO *) + (*** Range *) Record range (T : Type) := mk_range { start: T; -- cgit v1.2.3 From 49117ba254679f98938223711810191c3f7d788f Mon Sep 17 00:00:00 2001 From: Son Ho Date: Fri, 27 Oct 2023 13:34:03 +0200 Subject: Regenerate the Coq test files --- tests/coq/hashmap/Hashmap_Funs.v | 514 +++++++++++++++++++++----------------- tests/coq/hashmap/Hashmap_Types.v | 30 +-- tests/coq/hashmap/Primitives.v | 405 +++++++++++++++++++++++++----- 3 files changed, 640 insertions(+), 309 deletions(-) (limited to 'tests/coq/hashmap') diff --git a/tests/coq/hashmap/Hashmap_Funs.v b/tests/coq/hashmap/Hashmap_Funs.v index 054880d4..3ca52a9f 100644 --- a/tests/coq/hashmap/Hashmap_Funs.v +++ b/tests/coq/hashmap/Hashmap_Funs.v @@ -11,98 +11,101 @@ Import Hashmap_Types. Module Hashmap_Funs. (** [hashmap::hash_key]: forward function *) -Definition hash_key_fwd (k : usize) : result usize := +Definition hash_key (k : usize) : result usize := Return k. (** [hashmap::HashMap::{0}::allocate_slots]: loop 0: forward function *) -Fixpoint hash_map_allocate_slots_loop_fwd - (T : Type) (n : nat) (slots : vec (List_t T)) (n0 : usize) : - result (vec (List_t T)) +Fixpoint hashMap_allocate_slots_loop + (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n0 : usize) : + result (alloc_vec_Vec (List_t T)) := match n with | O => Fail_ OutOfFuel | S n1 => if n0 s> 0%usize then ( - slots0 <- vec_push_back (List_t T) slots ListNil; + slots0 <- alloc_vec_Vec_push (List_t T) slots List_Nil; n2 <- usize_sub n0 1%usize; - hash_map_allocate_slots_loop_fwd T n1 slots0 n2) + hashMap_allocate_slots_loop T n1 slots0 n2) else Return slots end . (** [hashmap::HashMap::{0}::allocate_slots]: forward function *) -Definition hash_map_allocate_slots_fwd - (T : Type) (n : nat) (slots : vec (List_t T)) (n0 : usize) : - result (vec (List_t T)) +Definition hashMap_allocate_slots + (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n0 : usize) : + result (alloc_vec_Vec (List_t T)) := - hash_map_allocate_slots_loop_fwd T n slots n0 + hashMap_allocate_slots_loop T n slots n0 . (** [hashmap::HashMap::{0}::new_with_capacity]: forward function *) -Definition hash_map_new_with_capacity_fwd +Definition hashMap_new_with_capacity (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : - result (Hash_map_t T) + result (HashMap_t T) := - let v := vec_new (List_t T) in - slots <- hash_map_allocate_slots_fwd T n v 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; Return {| - Hash_map_num_entries := 0%usize; - Hash_map_max_load_factor := (max_load_dividend, max_load_divisor); - Hash_map_max_load := i0; - Hash_map_slots := slots + hashMap_num_entries := 0%usize; + hashMap_max_load_factor := (max_load_dividend, max_load_divisor); + hashMap_max_load := i0; + hashMap_slots := slots |} . (** [hashmap::HashMap::{0}::new]: forward function *) -Definition hash_map_new_fwd (T : Type) (n : nat) : result (Hash_map_t T) := - hash_map_new_with_capacity_fwd T n 32%usize 4%usize 5%usize +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::{0}::clear]: loop 0: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) -Fixpoint hash_map_clear_loop_fwd_back - (T : Type) (n : nat) (slots : vec (List_t T)) (i : usize) : - result (vec (List_t T)) +Fixpoint hashMap_clear_loop + (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (i : usize) : + result (alloc_vec_Vec (List_t T)) := match n with | O => Fail_ OutOfFuel | S n0 => - let i0 := vec_len (List_t T) slots in + let i0 := alloc_vec_Vec_len (List_t T) slots in if i s< i0 then ( i1 <- usize_add i 1%usize; - slots0 <- vec_index_mut_back (List_t T) slots i ListNil; - hash_map_clear_loop_fwd_back T n0 slots0 i1) + slots0 <- + alloc_vec_Vec_index_mut_back (List_t T) usize + (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + slots i List_Nil; + hashMap_clear_loop T n0 slots0 i1) else Return slots end . (** [hashmap::HashMap::{0}::clear]: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) -Definition hash_map_clear_fwd_back - (T : Type) (n : nat) (self : Hash_map_t T) : result (Hash_map_t T) := - v <- hash_map_clear_loop_fwd_back T n self.(Hash_map_slots) 0%usize; +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; Return {| - Hash_map_num_entries := 0%usize; - Hash_map_max_load_factor := self.(Hash_map_max_load_factor); - Hash_map_max_load := self.(Hash_map_max_load); - Hash_map_slots := v + 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::HashMap::{0}::len]: forward function *) -Definition hash_map_len_fwd (T : Type) (self : Hash_map_t T) : result usize := - Return self.(Hash_map_num_entries) +Definition hashMap_len (T : Type) (self : HashMap_t T) : result usize := + Return self.(hashMap_num_entries) . (** [hashmap::HashMap::{0}::insert_in_list]: loop 0: forward function *) -Fixpoint hash_map_insert_in_list_loop_fwd +Fixpoint hashMap_insert_in_list_loop (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : result bool := @@ -110,25 +113,25 @@ Fixpoint hash_map_insert_in_list_loop_fwd | O => Fail_ OutOfFuel | S n0 => match ls with - | ListCons ckey cvalue tl => + | List_Cons ckey cvalue tl => if ckey s= key then Return false - else hash_map_insert_in_list_loop_fwd T n0 key value tl - | ListNil => Return true + else hashMap_insert_in_list_loop T n0 key value tl + | List_Nil => Return true end end . (** [hashmap::HashMap::{0}::insert_in_list]: forward function *) -Definition hash_map_insert_in_list_fwd +Definition hashMap_insert_in_list (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : result bool := - hash_map_insert_in_list_loop_fwd T n key value ls + hashMap_insert_in_list_loop T n key value ls . (** [hashmap::HashMap::{0}::insert_in_list]: loop 0: backward function 0 *) -Fixpoint hash_map_insert_in_list_loop_back +Fixpoint hashMap_insert_in_list_loop_back (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : result (List_t T) := @@ -136,255 +139,275 @@ Fixpoint hash_map_insert_in_list_loop_back | O => Fail_ OutOfFuel | S n0 => match ls with - | ListCons ckey cvalue tl => + | List_Cons ckey cvalue tl => if ckey s= key - then Return (ListCons ckey value tl) + then Return (List_Cons ckey value tl) else ( - tl0 <- hash_map_insert_in_list_loop_back T n0 key value tl; - Return (ListCons ckey cvalue tl0)) - | ListNil => let l := ListNil in Return (ListCons key value l) + 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::{0}::insert_in_list]: backward function 0 *) -Definition hash_map_insert_in_list_back +Definition hashMap_insert_in_list_back (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : result (List_t T) := - hash_map_insert_in_list_loop_back T n key value ls + hashMap_insert_in_list_loop_back T n key value ls . (** [hashmap::HashMap::{0}::insert_no_resize]: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) -Definition hash_map_insert_no_resize_fwd_back - (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) (value : T) : - result (Hash_map_t T) +Definition hashMap_insert_no_resize + (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) : + result (HashMap_t T) := - hash <- hash_key_fwd key; - let i := vec_len (List_t T) self.(Hash_map_slots) in + hash <- hash_key key; + let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd (List_t T) self.(Hash_map_slots) hash_mod; - inserted <- hash_map_insert_in_list_fwd T n key value l; + l <- + alloc_vec_Vec_index_mut (List_t T) usize + (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + self.(hashMap_slots) hash_mod; + inserted <- hashMap_insert_in_list T n key value l; if inserted then ( - i0 <- usize_add self.(Hash_map_num_entries) 1%usize; - l0 <- hash_map_insert_in_list_back T n key value l; - v <- vec_index_mut_back (List_t T) self.(Hash_map_slots) hash_mod l0; + 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_usize_coresliceindexSliceIndexInst (List_t T)) + self.(hashMap_slots) hash_mod l0; Return {| - Hash_map_num_entries := i0; - Hash_map_max_load_factor := self.(Hash_map_max_load_factor); - Hash_map_max_load := self.(Hash_map_max_load); - Hash_map_slots := v + hashMap_num_entries := i0; + hashMap_max_load_factor := self.(hashMap_max_load_factor); + hashMap_max_load := self.(hashMap_max_load); + hashMap_slots := v |}) else ( - l0 <- hash_map_insert_in_list_back T n key value l; - v <- vec_index_mut_back (List_t T) self.(Hash_map_slots) hash_mod l0; + 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_usize_coresliceindexSliceIndexInst (List_t T)) + self.(hashMap_slots) hash_mod l0; Return {| - Hash_map_num_entries := self.(Hash_map_num_entries); - Hash_map_max_load_factor := self.(Hash_map_max_load_factor); - Hash_map_max_load := self.(Hash_map_max_load); - Hash_map_slots := v + 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::{0}::move_elements_from_list]: loop 0: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) -Fixpoint hash_map_move_elements_from_list_loop_fwd_back - (T : Type) (n : nat) (ntable : Hash_map_t T) (ls : List_t T) : - result (Hash_map_t T) +Fixpoint hashMap_move_elements_from_list_loop + (T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) : + result (HashMap_t T) := match n with | O => Fail_ OutOfFuel | S n0 => match ls with - | ListCons k v tl => - ntable0 <- hash_map_insert_no_resize_fwd_back T n0 ntable k v; - hash_map_move_elements_from_list_loop_fwd_back T n0 ntable0 tl - | ListNil => Return ntable + | 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 + | List_Nil => Return ntable end end . (** [hashmap::HashMap::{0}::move_elements_from_list]: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) -Definition hash_map_move_elements_from_list_fwd_back - (T : Type) (n : nat) (ntable : Hash_map_t T) (ls : List_t T) : - result (Hash_map_t T) +Definition hashMap_move_elements_from_list + (T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) : + result (HashMap_t T) := - hash_map_move_elements_from_list_loop_fwd_back T n ntable ls + hashMap_move_elements_from_list_loop T n ntable ls . (** [hashmap::HashMap::{0}::move_elements]: loop 0: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) -Fixpoint hash_map_move_elements_loop_fwd_back - (T : Type) (n : nat) (ntable : Hash_map_t T) (slots : vec (List_t T)) - (i : usize) : - result ((Hash_map_t T) * (vec (List_t T))) +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))) := match n with | O => Fail_ OutOfFuel | S n0 => - let i0 := vec_len (List_t T) slots in + let i0 := alloc_vec_Vec_len (List_t T) slots in if i s< i0 then ( - l <- vec_index_mut_fwd (List_t T) slots i; - let ls := mem_replace_fwd (List_t T) l ListNil in - ntable0 <- hash_map_move_elements_from_list_fwd_back T n0 ntable ls; + l <- + alloc_vec_Vec_index_mut (List_t T) usize + (core_slice_index_usize_coresliceindexSliceIndexInst (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 := mem_replace_back (List_t T) l ListNil in - slots0 <- vec_index_mut_back (List_t T) slots i l0; - hash_map_move_elements_loop_fwd_back T n0 ntable0 slots0 i1) + 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_usize_coresliceindexSliceIndexInst (List_t T)) + slots i l0; + hashMap_move_elements_loop T n0 ntable0 slots0 i1) else Return (ntable, slots) end . (** [hashmap::HashMap::{0}::move_elements]: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) -Definition hash_map_move_elements_fwd_back - (T : Type) (n : nat) (ntable : Hash_map_t T) (slots : vec (List_t T)) - (i : usize) : - result ((Hash_map_t T) * (vec (List_t T))) +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))) := - hash_map_move_elements_loop_fwd_back T n ntable slots i + hashMap_move_elements_loop T n ntable slots i . (** [hashmap::HashMap::{0}::try_resize]: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) -Definition hash_map_try_resize_fwd_back - (T : Type) (n : nat) (self : Hash_map_t T) : result (Hash_map_t T) := +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 := vec_len (List_t T) self.(Hash_map_slots) in + let capacity := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in n1 <- usize_div max_usize 2%usize; - let (i, i0) := self.(Hash_map_max_load_factor) in + let (i, i0) := self.(hashMap_max_load_factor) in i1 <- usize_div n1 i; if capacity s<= i1 then ( i2 <- usize_mul capacity 2%usize; - ntable <- hash_map_new_with_capacity_fwd T n i2 i i0; - p <- - hash_map_move_elements_fwd_back T n ntable self.(Hash_map_slots) 0%usize; + ntable <- hashMap_new_with_capacity T n i2 i i0; + p <- hashMap_move_elements T n ntable self.(hashMap_slots) 0%usize; let (ntable0, _) := p in Return {| - Hash_map_num_entries := self.(Hash_map_num_entries); - Hash_map_max_load_factor := (i, i0); - Hash_map_max_load := ntable0.(Hash_map_max_load); - Hash_map_slots := ntable0.(Hash_map_slots) + 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) |}) else Return {| - Hash_map_num_entries := self.(Hash_map_num_entries); - Hash_map_max_load_factor := (i, i0); - Hash_map_max_load := self.(Hash_map_max_load); - Hash_map_slots := self.(Hash_map_slots) + hashMap_num_entries := self.(hashMap_num_entries); + hashMap_max_load_factor := (i, i0); + hashMap_max_load := self.(hashMap_max_load); + hashMap_slots := self.(hashMap_slots) |} . (** [hashmap::HashMap::{0}::insert]: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) -Definition hash_map_insert_fwd_back - (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) (value : T) : - result (Hash_map_t T) +Definition hashMap_insert + (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) : + result (HashMap_t T) := - self0 <- hash_map_insert_no_resize_fwd_back T n self key value; - i <- hash_map_len_fwd T self0; - if i s> self0.(Hash_map_max_load) - then hash_map_try_resize_fwd_back T n self0 + 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 . (** [hashmap::HashMap::{0}::contains_key_in_list]: loop 0: forward function *) -Fixpoint hash_map_contains_key_in_list_loop_fwd +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 => match ls with - | ListCons ckey t tl => + | List_Cons ckey t tl => if ckey s= key then Return true - else hash_map_contains_key_in_list_loop_fwd T n0 key tl - | ListNil => Return false + else hashMap_contains_key_in_list_loop T n0 key tl + | List_Nil => Return false end end . (** [hashmap::HashMap::{0}::contains_key_in_list]: forward function *) -Definition hash_map_contains_key_in_list_fwd +Definition hashMap_contains_key_in_list (T : Type) (n : nat) (key : usize) (ls : List_t T) : result bool := - hash_map_contains_key_in_list_loop_fwd T n key ls + hashMap_contains_key_in_list_loop T n key ls . (** [hashmap::HashMap::{0}::contains_key]: forward function *) -Definition hash_map_contains_key_fwd - (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) : result bool := - hash <- hash_key_fwd key; - let i := vec_len (List_t T) self.(Hash_map_slots) in +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 hash_mod <- usize_rem hash i; - l <- vec_index_fwd (List_t T) self.(Hash_map_slots) hash_mod; - hash_map_contains_key_in_list_fwd T n key l + l <- + alloc_vec_Vec_index (List_t T) usize + (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + self.(hashMap_slots) hash_mod; + hashMap_contains_key_in_list T n key l . (** [hashmap::HashMap::{0}::get_in_list]: loop 0: forward function *) -Fixpoint hash_map_get_in_list_loop_fwd +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 => match ls with - | ListCons ckey cvalue tl => + | List_Cons ckey cvalue tl => if ckey s= key then Return cvalue - else hash_map_get_in_list_loop_fwd T n0 key tl - | ListNil => Fail_ Failure + else hashMap_get_in_list_loop T n0 key tl + | List_Nil => Fail_ Failure end end . (** [hashmap::HashMap::{0}::get_in_list]: forward function *) -Definition hash_map_get_in_list_fwd +Definition hashMap_get_in_list (T : Type) (n : nat) (key : usize) (ls : List_t T) : result T := - hash_map_get_in_list_loop_fwd T n key ls + hashMap_get_in_list_loop T n key ls . (** [hashmap::HashMap::{0}::get]: forward function *) -Definition hash_map_get_fwd - (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) : result T := - hash <- hash_key_fwd key; - let i := vec_len (List_t T) self.(Hash_map_slots) in +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 hash_mod <- usize_rem hash i; - l <- vec_index_fwd (List_t T) self.(Hash_map_slots) hash_mod; - hash_map_get_in_list_fwd T n key l + l <- + alloc_vec_Vec_index (List_t T) usize + (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + self.(hashMap_slots) hash_mod; + hashMap_get_in_list T n key l . (** [hashmap::HashMap::{0}::get_mut_in_list]: loop 0: forward function *) -Fixpoint hash_map_get_mut_in_list_loop_fwd +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 - | ListCons ckey cvalue tl => + | List_Cons ckey cvalue tl => if ckey s= key then Return cvalue - else hash_map_get_mut_in_list_loop_fwd T n0 tl key - | ListNil => Fail_ Failure + else hashMap_get_mut_in_list_loop T n0 tl key + | List_Nil => Fail_ Failure end end . (** [hashmap::HashMap::{0}::get_mut_in_list]: forward function *) -Definition hash_map_get_mut_in_list_fwd +Definition hashMap_get_mut_in_list (T : Type) (n : nat) (ls : List_t T) (key : usize) : result T := - hash_map_get_mut_in_list_loop_fwd T n ls key + hashMap_get_mut_in_list_loop T n ls key . (** [hashmap::HashMap::{0}::get_mut_in_list]: loop 0: backward function 0 *) -Fixpoint hash_map_get_mut_in_list_loop_back +Fixpoint hashMap_get_mut_in_list_loop_back (T : Type) (n : nat) (ls : List_t T) (key : usize) (ret : T) : result (List_t T) := @@ -392,196 +415,219 @@ Fixpoint hash_map_get_mut_in_list_loop_back | O => Fail_ OutOfFuel | S n0 => match ls with - | ListCons ckey cvalue tl => + | List_Cons ckey cvalue tl => if ckey s= key - then Return (ListCons ckey ret tl) + then Return (List_Cons ckey ret tl) else ( - tl0 <- hash_map_get_mut_in_list_loop_back T n0 tl key ret; - Return (ListCons ckey cvalue tl0)) - | ListNil => Fail_ Failure + tl0 <- hashMap_get_mut_in_list_loop_back T n0 tl key ret; + Return (List_Cons ckey cvalue tl0)) + | List_Nil => Fail_ Failure end end . (** [hashmap::HashMap::{0}::get_mut_in_list]: backward function 0 *) -Definition hash_map_get_mut_in_list_back +Definition hashMap_get_mut_in_list_back (T : Type) (n : nat) (ls : List_t T) (key : usize) (ret : T) : result (List_t T) := - hash_map_get_mut_in_list_loop_back T n ls key ret + hashMap_get_mut_in_list_loop_back T n ls key ret . (** [hashmap::HashMap::{0}::get_mut]: forward function *) -Definition hash_map_get_mut_fwd - (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) : result T := - hash <- hash_key_fwd key; - let i := vec_len (List_t T) self.(Hash_map_slots) in +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 <- vec_index_mut_fwd (List_t T) self.(Hash_map_slots) hash_mod; - hash_map_get_mut_in_list_fwd T n l key + l <- + alloc_vec_Vec_index_mut (List_t T) usize + (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + self.(hashMap_slots) hash_mod; + hashMap_get_mut_in_list T n l key . (** [hashmap::HashMap::{0}::get_mut]: backward function 0 *) -Definition hash_map_get_mut_back - (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) (ret : T) : - result (Hash_map_t T) +Definition hashMap_get_mut_back + (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (ret : T) : + result (HashMap_t T) := - hash <- hash_key_fwd key; - let i := vec_len (List_t T) self.(Hash_map_slots) in + hash <- hash_key key; + let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd (List_t T) self.(Hash_map_slots) hash_mod; - l0 <- hash_map_get_mut_in_list_back T n l key ret; - v <- vec_index_mut_back (List_t T) self.(Hash_map_slots) hash_mod l0; + l <- + alloc_vec_Vec_index_mut (List_t T) usize + (core_slice_index_usize_coresliceindexSliceIndexInst (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_usize_coresliceindexSliceIndexInst (List_t T)) + self.(hashMap_slots) hash_mod l0; Return {| - Hash_map_num_entries := self.(Hash_map_num_entries); - Hash_map_max_load_factor := self.(Hash_map_max_load_factor); - Hash_map_max_load := self.(Hash_map_max_load); - Hash_map_slots := v + 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::{0}::remove_from_list]: loop 0: forward function *) -Fixpoint hash_map_remove_from_list_loop_fwd +Fixpoint hashMap_remove_from_list_loop (T : Type) (n : nat) (key : usize) (ls : List_t T) : result (option T) := match n with | O => Fail_ OutOfFuel | S n0 => match ls with - | ListCons ckey t tl => + | List_Cons ckey t tl => if ckey s= key then - let mv_ls := mem_replace_fwd (List_t T) (ListCons ckey t tl) ListNil in + let mv_ls := core_mem_replace (List_t T) (List_Cons ckey t tl) List_Nil + in match mv_ls with - | ListCons i cvalue tl0 => Return (Some cvalue) - | ListNil => Fail_ Failure + | List_Cons i cvalue tl0 => Return (Some cvalue) + | List_Nil => Fail_ Failure end - else hash_map_remove_from_list_loop_fwd T n0 key tl - | ListNil => Return None + else hashMap_remove_from_list_loop T n0 key tl + | List_Nil => Return None end end . (** [hashmap::HashMap::{0}::remove_from_list]: forward function *) -Definition hash_map_remove_from_list_fwd +Definition hashMap_remove_from_list (T : Type) (n : nat) (key : usize) (ls : List_t T) : result (option T) := - hash_map_remove_from_list_loop_fwd T n key ls + hashMap_remove_from_list_loop T n key ls . (** [hashmap::HashMap::{0}::remove_from_list]: loop 0: backward function 1 *) -Fixpoint hash_map_remove_from_list_loop_back +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 - | ListCons ckey t tl => + | List_Cons ckey t tl => if ckey s= key then - let mv_ls := mem_replace_fwd (List_t T) (ListCons ckey t tl) ListNil in + let mv_ls := core_mem_replace (List_t T) (List_Cons ckey t tl) List_Nil + in match mv_ls with - | ListCons i cvalue tl0 => Return tl0 - | ListNil => Fail_ Failure + | List_Cons i cvalue tl0 => Return tl0 + | List_Nil => Fail_ Failure end else ( - tl0 <- hash_map_remove_from_list_loop_back T n0 key tl; - Return (ListCons ckey t tl0)) - | ListNil => Return ListNil + 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::{0}::remove_from_list]: backward function 1 *) -Definition hash_map_remove_from_list_back +Definition hashMap_remove_from_list_back (T : Type) (n : nat) (key : usize) (ls : List_t T) : result (List_t T) := - hash_map_remove_from_list_loop_back T n key ls + hashMap_remove_from_list_loop_back T n key ls . (** [hashmap::HashMap::{0}::remove]: forward function *) -Definition hash_map_remove_fwd - (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) : +Definition hashMap_remove + (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result (option T) := - hash <- hash_key_fwd key; - let i := vec_len (List_t T) self.(Hash_map_slots) in + hash <- hash_key key; + let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd (List_t T) self.(Hash_map_slots) hash_mod; - x <- hash_map_remove_from_list_fwd T n key l; + l <- + alloc_vec_Vec_index_mut (List_t T) usize + (core_slice_index_usize_coresliceindexSliceIndexInst (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.(Hash_map_num_entries) 1%usize; Return (Some x0) + _ <- usize_sub self.(hashMap_num_entries) 1%usize; Return (Some x0) end . (** [hashmap::HashMap::{0}::remove]: backward function 0 *) -Definition hash_map_remove_back - (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) : - result (Hash_map_t T) +Definition hashMap_remove_back + (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : + result (HashMap_t T) := - hash <- hash_key_fwd key; - let i := vec_len (List_t T) self.(Hash_map_slots) in + hash <- hash_key key; + let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd (List_t T) self.(Hash_map_slots) hash_mod; - x <- hash_map_remove_from_list_fwd T n key l; + l <- + alloc_vec_Vec_index_mut (List_t T) usize + (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + self.(hashMap_slots) hash_mod; + x <- hashMap_remove_from_list T n key l; match x with | None => - l0 <- hash_map_remove_from_list_back T n key l; - v <- vec_index_mut_back (List_t T) self.(Hash_map_slots) hash_mod l0; + l0 <- hashMap_remove_from_list_back T n key l; + v <- + alloc_vec_Vec_index_mut_back (List_t T) usize + (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + self.(hashMap_slots) hash_mod l0; Return {| - Hash_map_num_entries := self.(Hash_map_num_entries); - Hash_map_max_load_factor := self.(Hash_map_max_load_factor); - Hash_map_max_load := self.(Hash_map_max_load); - Hash_map_slots := v + 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.(Hash_map_num_entries) 1%usize; - l0 <- hash_map_remove_from_list_back T n key l; - v <- vec_index_mut_back (List_t T) self.(Hash_map_slots) hash_mod l0; + 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_usize_coresliceindexSliceIndexInst (List_t T)) + self.(hashMap_slots) hash_mod l0; Return {| - Hash_map_num_entries := i0; - Hash_map_max_load_factor := self.(Hash_map_max_load_factor); - Hash_map_max_load := self.(Hash_map_max_load); - Hash_map_slots := v + hashMap_num_entries := i0; + hashMap_max_load_factor := self.(hashMap_max_load_factor); + hashMap_max_load := self.(hashMap_max_load); + hashMap_slots := v |} end . (** [hashmap::test1]: forward function *) -Definition test1_fwd (n : nat) : result unit := - hm <- hash_map_new_fwd u64 n; - hm0 <- hash_map_insert_fwd_back u64 n hm 0%usize 42%u64; - hm1 <- hash_map_insert_fwd_back u64 n hm0 128%usize 18%u64; - hm2 <- hash_map_insert_fwd_back u64 n hm1 1024%usize 138%u64; - hm3 <- hash_map_insert_fwd_back u64 n hm2 1056%usize 256%u64; - i <- hash_map_get_fwd u64 n hm3 128%usize; +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; if negb (i s= 18%u64) then Fail_ Failure else ( - hm4 <- hash_map_get_mut_back u64 n hm3 1024%usize 56%u64; - i0 <- hash_map_get_fwd u64 n hm4 1024%usize; + 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) then Fail_ Failure else ( - x <- hash_map_remove_fwd u64 n hm4 1024%usize; + x <- 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 <- hash_map_remove_back u64 n hm4 1024%usize; - i1 <- hash_map_get_fwd u64 n hm5 0%usize; + hm5 <- hashMap_remove_back u64 n hm4 1024%usize; + i1 <- hashMap_get u64 n hm5 0%usize; if negb (i1 s= 42%u64) then Fail_ Failure else ( - i2 <- hash_map_get_fwd u64 n hm5 128%usize; + i2 <- hashMap_get u64 n hm5 128%usize; if negb (i2 s= 18%u64) then Fail_ Failure else ( - i3 <- hash_map_get_fwd u64 n hm5 1056%usize; + i3 <- hashMap_get u64 n hm5 1056%usize; if negb (i3 s= 256%u64) then Fail_ Failure else Return tt))) end)) . diff --git a/tests/coq/hashmap/Hashmap_Types.v b/tests/coq/hashmap/Hashmap_Types.v index dbde6be9..63d30eeb 100644 --- a/tests/coq/hashmap/Hashmap_Types.v +++ b/tests/coq/hashmap/Hashmap_Types.v @@ -10,27 +10,27 @@ Module Hashmap_Types. (** [hashmap::List] *) Inductive List_t (T : Type) := -| ListCons : usize -> T -> List_t T -> List_t T -| ListNil : List_t T +| List_Cons : usize -> T -> List_t T -> List_t T +| List_Nil : List_t T . -Arguments ListCons {T} _ _ _. -Arguments ListNil {T}. +Arguments List_Cons {T} _ _ _. +Arguments List_Nil {T}. (** [hashmap::HashMap] *) -Record Hash_map_t (T : Type) := -mkHash_map_t { - Hash_map_num_entries : usize; - Hash_map_max_load_factor : (usize * usize); - Hash_map_max_load : usize; - Hash_map_slots : vec (List_t T); +Record HashMap_t (T : Type) := +mkHashMap_t { + hashMap_num_entries : usize; + hashMap_max_load_factor : (usize * usize); + hashMap_max_load : usize; + hashMap_slots : alloc_vec_Vec (List_t T); } . -Arguments mkHash_map_t {T} _ _ _ _. -Arguments Hash_map_num_entries {T}. -Arguments Hash_map_max_load_factor {T}. -Arguments Hash_map_max_load {T}. -Arguments Hash_map_slots {T}. +Arguments mkHashMap_t {T} _ _ _ _. +Arguments hashMap_num_entries {T}. +Arguments hashMap_max_load_factor {T}. +Arguments hashMap_max_load {T}. +Arguments hashMap_slots {T}. End Hashmap_Types . diff --git a/tests/coq/hashmap/Primitives.v b/tests/coq/hashmap/Primitives.v index 8d6c9c8d..85e38f01 100644 --- a/tests/coq/hashmap/Primitives.v +++ b/tests/coq/hashmap/Primitives.v @@ -63,13 +63,15 @@ Check (if true then Return (1 + 2) else Fail_ Failure)%global = 3. (*** Misc *) - Definition string := Coq.Strings.String.string. Definition char := Coq.Strings.Ascii.ascii. Definition char_of_byte := Coq.Strings.Ascii.ascii_of_byte. -Definition mem_replace_fwd (a : Type) (x : a) (y : a) : a := x . -Definition mem_replace_back (a : Type) (x : a) (y : a) : a := y . +Definition core_mem_replace (a : Type) (x : a) (y : a) : a := x . +Definition core_mem_replace_back (a : Type) (x : a) (y : a) : a := y . + +Record mut_raw_ptr (T : Type) := { mut_raw_ptr_v : T }. +Record const_raw_ptr (T : Type) := { const_raw_ptr_v : T }. (*** Scalars *) @@ -408,12 +410,75 @@ Definition core_i64_max := i64_max %i64. Definition core_i128_max := i64_max %i128. Axiom core_isize_max : isize. (** TODO *) -(*** Range *) -Record range (T : Type) := mk_range { - start: T; - end_: T; +(*** core::ops *) + +(* Trait declaration: [core::ops::index::Index] *) +Record core_ops_index_Index (Self Idx : Type) := mk_core_ops_index_Index { + core_ops_index_Index_Output : Type; + core_ops_index_Index_index : Self -> Idx -> result core_ops_index_Index_Output; +}. +Arguments mk_core_ops_index_Index {_ _}. +Arguments core_ops_index_Index_Output {_ _}. +Arguments core_ops_index_Index_index {_ _}. + +(* Trait declaration: [core::ops::index::IndexMut] *) +Record core_ops_index_IndexMut (Self Idx : Type) := mk_core_ops_index_IndexMut { + core_ops_index_IndexMut_indexInst : core_ops_index_Index Self Idx; + core_ops_index_IndexMut_index_mut : Self -> Idx -> result core_ops_index_IndexMut_indexInst.(core_ops_index_Index_Output); + core_ops_index_IndexMut_index_mut_back : Self -> Idx -> core_ops_index_IndexMut_indexInst.(core_ops_index_Index_Output) -> result Self; +}. +Arguments mk_core_ops_index_IndexMut {_ _}. +Arguments core_ops_index_IndexMut_indexInst {_ _}. +Arguments core_ops_index_IndexMut_index_mut {_ _}. +Arguments core_ops_index_IndexMut_index_mut_back {_ _}. + +(* Trait declaration [core::ops::deref::Deref] *) +Record core_ops_deref_Deref (Self : Type) := mk_core_ops_deref_Deref { + core_ops_deref_Deref_target : Type; + core_ops_deref_Deref_deref : Self -> result core_ops_deref_Deref_target; +}. +Arguments mk_core_ops_deref_Deref {_}. +Arguments core_ops_deref_Deref_target {_}. +Arguments core_ops_deref_Deref_deref {_}. + +(* Trait declaration [core::ops::deref::DerefMut] *) +Record core_ops_deref_DerefMut (Self : Type) := mk_core_ops_deref_DerefMut { + core_ops_deref_DerefMut_derefInst : core_ops_deref_Deref Self; + core_ops_deref_DerefMut_deref_mut : Self -> result core_ops_deref_DerefMut_derefInst.(core_ops_deref_Deref_target); + core_ops_deref_DerefMut_deref_mut_back : Self -> core_ops_deref_DerefMut_derefInst.(core_ops_deref_Deref_target) -> result Self; }. -Arguments mk_range {_}. +Arguments mk_core_ops_deref_DerefMut {_}. +Arguments core_ops_deref_DerefMut_derefInst {_}. +Arguments core_ops_deref_DerefMut_deref_mut {_}. +Arguments core_ops_deref_DerefMut_deref_mut_back {_}. + +Record core_ops_range_Range (T : Type) := mk_core_ops_range_Range { + core_ops_range_Range_start : T; + core_ops_range_Range_end_ : T; +}. +Arguments mk_core_ops_range_Range {_}. +Arguments core_ops_range_Range_start {_}. +Arguments core_ops_range_Range_end_ {_}. + +(*** [alloc] *) + +Definition alloc_boxed_Box_deref (T : Type) (x : T) : result T := Return x. +Definition alloc_boxed_Box_deref_mut (T : Type) (x : T) : result T := Return x. +Definition alloc_boxed_Box_deref_mut_back (T : Type) (_ : T) (x : T) : result T := Return x. + +(* Trait instance *) +Definition alloc_boxed_Box_coreOpsDerefInst (Self : Type) : core_ops_deref_Deref Self := {| + core_ops_deref_Deref_target := Self; + core_ops_deref_Deref_deref := alloc_boxed_Box_deref Self; +|}. + +(* Trait instance *) +Definition alloc_boxed_Box_coreOpsDerefMutInst (Self : Type) : core_ops_deref_DerefMut Self := {| + core_ops_deref_DerefMut_derefInst := alloc_boxed_Box_coreOpsDerefInst Self; + core_ops_deref_DerefMut_deref_mut := alloc_boxed_Box_deref_mut Self; + core_ops_deref_DerefMut_deref_mut_back := alloc_boxed_Box_deref_mut_back Self; +|}. + (*** Arrays *) Definition array T (n : usize) := { l: list T | Z.of_nat (length l) = to_Z n}. @@ -433,51 +498,50 @@ Qed. (* TODO: finish the definitions *) Axiom mk_array : forall (T : Type) (n : usize) (l : list T), array T n. -Axiom array_index_shared : forall (T : Type) (n : usize) (x : array T n) (i : usize), result T. -Axiom array_index_mut_fwd : forall (T : Type) (n : usize) (x : array T n) (i : usize), result T. -Axiom array_index_mut_back : forall (T : Type) (n : usize) (x : array T n) (i : usize) (nx : T), result (array T n). +(* For initialization *) +Axiom array_repeat : forall (T : Type) (n : usize) (x : T), array T n. + +Axiom array_index_usize : forall (T : Type) (n : usize) (x : array T n) (i : usize), result T. +Axiom array_update_usize : forall (T : Type) (n : usize) (x : array T n) (i : usize) (nx : T), result (array T n). (*** Slice *) Definition slice T := { l: list T | Z.of_nat (length l) <= usize_max}. Axiom slice_len : forall (T : Type) (s : slice T), usize. -Axiom slice_index_shared : forall (T : Type) (x : slice T) (i : usize), result T. -Axiom slice_index_mut_fwd : forall (T : Type) (x : slice T) (i : usize), result T. -Axiom slice_index_mut_back : forall (T : Type) (x : slice T) (i : usize) (nx : T), result (slice T). +Axiom slice_index_usize : forall (T : Type) (x : slice T) (i : usize), result T. +Axiom slice_update_usize : forall (T : Type) (x : slice T) (i : usize) (nx : T), result (slice T). (*** Subslices *) -Axiom array_to_slice_shared : forall (T : Type) (n : usize) (x : array T n), result (slice T). -Axiom array_to_slice_mut_fwd : forall (T : Type) (n : usize) (x : array T n), result (slice T). -Axiom array_to_slice_mut_back : forall (T : Type) (n : usize) (x : array T n) (s : slice T), result (array T n). +Axiom array_to_slice : forall (T : Type) (n : usize) (x : array T n), result (slice T). +Axiom array_from_slice : forall (T : Type) (n : usize) (x : array T n) (s : slice T), result (array T n). + +Axiom array_subslice: forall (T : Type) (n : usize) (x : array T n) (r : core_ops_range_Range usize), result (slice T). +Axiom array_update_subslice: forall (T : Type) (n : usize) (x : array T n) (r : core_ops_range_Range usize) (ns : slice T), result (array T n). -Axiom array_subslice_shared: forall (T : Type) (n : usize) (x : array T n) (r : range usize), result (slice T). -Axiom array_subslice_mut_fwd: forall (T : Type) (n : usize) (x : array T n) (r : range usize), result (slice T). -Axiom array_subslice_mut_back: forall (T : Type) (n : usize) (x : array T n) (r : range usize) (ns : slice T), result (array T n). -Axiom slice_subslice_shared: forall (T : Type) (x : slice T) (r : range usize), result (slice T). -Axiom slice_subslice_mut_fwd: forall (T : Type) (x : slice T) (r : range usize), result (slice T). -Axiom slice_subslice_mut_back: forall (T : Type) (x : slice T) (r : range usize) (ns : slice T), result (slice T). +Axiom slice_subslice: forall (T : Type) (x : slice T) (r : core_ops_range_Range usize), result (slice T). +Axiom slice_update_subslice: forall (T : Type) (x : slice T) (r : core_ops_range_Range usize) (ns : slice T), result (slice T). (*** Vectors *) -Definition vec T := { l: list T | Z.of_nat (length l) <= usize_max }. +Definition alloc_vec_Vec T := { l: list T | Z.of_nat (length l) <= usize_max }. -Definition vec_to_list {T: Type} (v: vec T) : list T := proj1_sig v. +Definition alloc_vec_Vec_to_list {T: Type} (v: alloc_vec_Vec T) : list T := proj1_sig v. -Definition vec_length {T: Type} (v: vec T) : Z := Z.of_nat (length (vec_to_list v)). +Definition alloc_vec_Vec_length {T: Type} (v: alloc_vec_Vec T) : Z := Z.of_nat (length (alloc_vec_Vec_to_list v)). -Definition vec_new (T: Type) : vec T := (exist _ [] le_0_usize_max). +Definition alloc_vec_Vec_new (T: Type) : alloc_vec_Vec T := (exist _ [] le_0_usize_max). -Lemma vec_len_in_usize {T} (v: vec T) : usize_min <= vec_length v <= usize_max. +Lemma alloc_vec_Vec_len_in_usize {T} (v: alloc_vec_Vec T) : usize_min <= alloc_vec_Vec_length v <= usize_max. Proof. - unfold vec_length, usize_min. + unfold alloc_vec_Vec_length, usize_min. split. - lia. - apply (proj2_sig v). Qed. -Definition vec_len (T: Type) (v: vec T) : usize := - exist _ (vec_length v) (vec_len_in_usize v). +Definition alloc_vec_Vec_len (T: Type) (v: alloc_vec_Vec T) : usize := + exist _ (alloc_vec_Vec_length v) (alloc_vec_Vec_len_in_usize v). Fixpoint list_update {A} (l: list A) (n: nat) (a: A) : list A := @@ -488,50 +552,271 @@ Fixpoint list_update {A} (l: list A) (n: nat) (a: A) | S m => x :: (list_update t m a) end end. -Definition vec_bind {A B} (v: vec A) (f: list A -> result (list B)) : result (vec B) := - l <- f (vec_to_list v) ; +Definition alloc_vec_Vec_bind {A B} (v: alloc_vec_Vec A) (f: list A -> result (list B)) : result (alloc_vec_Vec B) := + l <- f (alloc_vec_Vec_to_list v) ; match sumbool_of_bool (scalar_le_max Usize (Z.of_nat (length l))) with | left H => Return (exist _ l (scalar_le_max_valid _ _ H)) | right _ => Fail_ Failure end. (* The **forward** function shouldn't be used *) -Definition vec_push_fwd (T: Type) (v: vec T) (x: T) : unit := tt. +Definition alloc_vec_Vec_push_fwd (T: Type) (v: alloc_vec_Vec T) (x: T) : unit := tt. -Definition vec_push_back (T: Type) (v: vec T) (x: T) : result (vec T) := - vec_bind v (fun l => Return (l ++ [x])). +Definition alloc_vec_Vec_push (T: Type) (v: alloc_vec_Vec T) (x: T) : result (alloc_vec_Vec T) := + alloc_vec_Vec_bind v (fun l => Return (l ++ [x])). (* The **forward** function shouldn't be used *) -Definition vec_insert_fwd (T: Type) (v: vec T) (i: usize) (x: T) : result unit := - if to_Z i +Definition alloc_vec_Vec_insert (T: Type) (v: alloc_vec_Vec T) (i: usize) (x: T) : result (alloc_vec_Vec T) := + alloc_vec_Vec_bind v (fun l => if to_Z i Return n - | None => Fail_ Failure - end. - -Definition vec_index_back (T: Type) (v: vec T) (i: usize) (x: T) : result unit := - if to_Z i Return n - | None => Fail_ Failure +(* Helper *) +Axiom alloc_vec_Vec_index_usize : forall {T : Type} (v : alloc_vec_Vec T) (i : usize) (x : T), result T. + +(* Helper *) +Axiom alloc_vec_Vec_update_usize : forall {T : Type} (v : alloc_vec_Vec T) (i : usize) (x : T), result (alloc_vec_Vec T). + +(* Trait declaration: [core::slice::index::private_slice_index::Sealed] *) +Definition core_slice_index_private_slice_index_Sealed (self : Type) := unit. + +(* Trait declaration: [core::slice::index::SliceIndex] *) +Record core_slice_index_SliceIndex (Self T : Type) := mk_core_slice_index_SliceIndex { + core_slice_index_SliceIndex_sealedInst : core_slice_index_private_slice_index_Sealed Self; + core_slice_index_SliceIndex_Output : Type; + core_slice_index_SliceIndex_get : Self -> T -> result (option core_slice_index_SliceIndex_Output); + core_slice_index_SliceIndex_get_mut : Self -> T -> result (option core_slice_index_SliceIndex_Output); + core_slice_index_SliceIndex_get_mut_back : Self -> T -> option core_slice_index_SliceIndex_Output -> result T; + core_slice_index_SliceIndex_get_unchecked : Self -> const_raw_ptr T -> result (const_raw_ptr core_slice_index_SliceIndex_Output); + core_slice_index_SliceIndex_get_unchecked_mut : Self -> mut_raw_ptr T -> result (mut_raw_ptr core_slice_index_SliceIndex_Output); + core_slice_index_SliceIndex_index : Self -> T -> result core_slice_index_SliceIndex_Output; + core_slice_index_SliceIndex_index_mut : Self -> T -> result core_slice_index_SliceIndex_Output; + core_slice_index_SliceIndex_index_mut_back : Self -> T -> core_slice_index_SliceIndex_Output -> result T; +}. +Arguments mk_core_slice_index_SliceIndex {_ _}. +Arguments core_slice_index_SliceIndex_sealedInst {_ _}. +Arguments core_slice_index_SliceIndex_Output {_ _}. +Arguments core_slice_index_SliceIndex_get {_ _}. +Arguments core_slice_index_SliceIndex_get_mut {_ _}. +Arguments core_slice_index_SliceIndex_get_mut_back {_ _}. +Arguments core_slice_index_SliceIndex_get_unchecked {_ _}. +Arguments core_slice_index_SliceIndex_get_unchecked_mut {_ _}. +Arguments core_slice_index_SliceIndex_index {_ _}. +Arguments core_slice_index_SliceIndex_index_mut {_ _}. +Arguments core_slice_index_SliceIndex_index_mut_back {_ _}. + +(* [core::slice::index::[T]::index]: forward function *) +Definition core_slice_index_Slice_index + (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)) + (s : slice T) (i : Idx) : result inst.(core_slice_index_SliceIndex_Output) := + x <- inst.(core_slice_index_SliceIndex_get) i s; + match x with + | None => Fail_ Failure + | Some x => Return x end. -Definition vec_index_mut_back (T: Type) (v: vec T) (i: usize) (x: T) : result (vec T) := - vec_bind v (fun l => - if to_Z i slice T -> result (option (slice T)). + +(* [core::slice::index::Range::get_mut]: backward function 0 *) +Axiom core_slice_index_Range_get_mut_back : + forall (T : Type), core_ops_range_Range usize -> slice T -> option (slice T) -> result (slice T). + +(* [core::slice::index::Range::get_unchecked]: forward function *) +Definition core_slice_index_Range_get_unchecked + (T : Type) : + core_ops_range_Range usize -> const_raw_ptr (slice T) -> result (const_raw_ptr (slice T)) := + (* Don't know what the model should be - for now we always fail to make + sure code which uses it fails *) + fun _ _ => Fail_ Failure. + +(* [core::slice::index::Range::get_unchecked_mut]: forward function *) +Definition core_slice_index_Range_get_unchecked_mut + (T : Type) : + core_ops_range_Range usize -> mut_raw_ptr (slice T) -> result (mut_raw_ptr (slice T)) := + (* Don't know what the model should be - for now we always fail to make + sure code which uses it fails *) + fun _ _ => Fail_ Failure. + +(* [core::slice::index::Range::index]: forward function *) +Axiom core_slice_index_Range_index : + forall (T : Type), core_ops_range_Range usize -> slice T -> result (slice T). + +(* [core::slice::index::Range::index_mut]: forward function *) +Axiom core_slice_index_Range_index_mut : + forall (T : Type), core_ops_range_Range usize -> slice T -> result (slice T). + +(* [core::slice::index::Range::index_mut]: backward function 0 *) +Axiom core_slice_index_Range_index_mut_back : + forall (T : Type), core_ops_range_Range usize -> slice T -> slice T -> result (slice T). + +(* [core::slice::index::[T]::index_mut]: forward function *) +Axiom core_slice_index_Slice_index_mut : + forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)), + slice T -> Idx -> result inst.(core_slice_index_SliceIndex_Output). + +(* [core::slice::index::[T]::index_mut]: backward function 0 *) +Axiom core_slice_index_Slice_index_mut_back : + forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)), + slice T -> Idx -> inst.(core_slice_index_SliceIndex_Output) -> result (slice T). + +(* [core::array::[T; N]::index]: forward function *) +Axiom core_array_Array_index : + forall (T Idx : Type) (N : usize) (inst : core_ops_index_Index (slice T) Idx) + (a : array T N) (i : Idx), result inst.(core_ops_index_Index_Output). + +(* [core::array::[T; N]::index_mut]: forward function *) +Axiom core_array_Array_index_mut : + forall (T Idx : Type) (N : usize) (inst : core_ops_index_IndexMut (slice T) Idx) + (a : array T N) (i : Idx), result inst.(core_ops_index_IndexMut_indexInst).(core_ops_index_Index_Output). + +(* [core::array::[T; N]::index_mut]: backward function 0 *) +Axiom core_array_Array_index_mut_back : + forall (T Idx : Type) (N : usize) (inst : core_ops_index_IndexMut (slice T) Idx) + (a : array T N) (i : Idx) (x : inst.(core_ops_index_IndexMut_indexInst).(core_ops_index_Index_Output)), result (array T N). + +(* Trait implementation: [core::slice::index::[T]] *) +Definition core_slice_index_Slice_coreopsindexIndexInst (T Idx : Type) + (inst : core_slice_index_SliceIndex Idx (slice T)) : + core_ops_index_Index (slice T) Idx := {| + core_ops_index_Index_Output := inst.(core_slice_index_SliceIndex_Output); + core_ops_index_Index_index := core_slice_index_Slice_index T Idx inst; +|}. + +(* Trait implementation: [core::slice::index::private_slice_index::Range] *) +Definition core_slice_index_private_slice_index_Range_coresliceindexprivate_slice_indexSealedInst + : core_slice_index_private_slice_index_Sealed (core_ops_range_Range usize) := tt. + +(* Trait implementation: [core::slice::index::Range] *) +Definition core_slice_index_Range_coresliceindexSliceIndexInst (T : Type) : + core_slice_index_SliceIndex (core_ops_range_Range usize) (slice T) := {| + core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_Range_coresliceindexprivate_slice_indexSealedInst; + core_slice_index_SliceIndex_Output := slice T; + core_slice_index_SliceIndex_get := core_slice_index_Range_get T; + core_slice_index_SliceIndex_get_mut := core_slice_index_Range_get_mut T; + core_slice_index_SliceIndex_get_mut_back := core_slice_index_Range_get_mut_back T; + core_slice_index_SliceIndex_get_unchecked := core_slice_index_Range_get_unchecked T; + core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_Range_get_unchecked_mut T; + core_slice_index_SliceIndex_index := core_slice_index_Range_index T; + core_slice_index_SliceIndex_index_mut := core_slice_index_Range_index_mut T; + core_slice_index_SliceIndex_index_mut_back := core_slice_index_Range_index_mut_back T; +|}. + +(* Trait implementation: [core::slice::index::[T]] *) +Definition core_slice_index_Slice_coreopsindexIndexMutInst (T Idx : Type) + (inst : core_slice_index_SliceIndex Idx (slice T)) : + core_ops_index_IndexMut (slice T) Idx := {| + core_ops_index_IndexMut_indexInst := core_slice_index_Slice_coreopsindexIndexInst T Idx inst; + core_ops_index_IndexMut_index_mut := core_slice_index_Slice_index_mut T Idx inst; + core_ops_index_IndexMut_index_mut_back := core_slice_index_Slice_index_mut_back T Idx inst; +|}. + +(* Trait implementation: [core::array::[T; N]] *) +Definition core_array_Array_coreopsindexIndexInst (T Idx : Type) (N : usize) + (inst : core_ops_index_Index (slice T) Idx) : + core_ops_index_Index (array T N) Idx := {| + core_ops_index_Index_Output := inst.(core_ops_index_Index_Output); + core_ops_index_Index_index := core_array_Array_index T Idx N inst; +|}. + +(* Trait implementation: [core::array::[T; N]] *) +Definition core_array_Array_coreopsindexIndexMutInst (T Idx : Type) (N : usize) + (inst : core_ops_index_IndexMut (slice T) Idx) : + core_ops_index_IndexMut (array T N) Idx := {| + core_ops_index_IndexMut_indexInst := core_array_Array_coreopsindexIndexInst T Idx N inst.(core_ops_index_IndexMut_indexInst); + core_ops_index_IndexMut_index_mut := core_array_Array_index_mut T Idx N inst; + core_ops_index_IndexMut_index_mut_back := core_array_Array_index_mut_back T Idx N inst; +|}. + +(* [core::slice::index::usize::get]: forward function *) +Axiom core_slice_index_usize_get : forall (T : Type), usize -> slice T -> result (option T). + +(* [core::slice::index::usize::get_mut]: forward function *) +Axiom core_slice_index_usize_get_mut : forall (T : Type), usize -> slice T -> result (option T). + +(* [core::slice::index::usize::get_mut]: backward function 0 *) +Axiom core_slice_index_usize_get_mut_back : + forall (T : Type), usize -> slice T -> option T -> result (slice T). + +(* [core::slice::index::usize::get_unchecked]: forward function *) +Axiom core_slice_index_usize_get_unchecked : + forall (T : Type), usize -> const_raw_ptr (slice T) -> result (const_raw_ptr T). + +(* [core::slice::index::usize::get_unchecked_mut]: forward function *) +Axiom core_slice_index_usize_get_unchecked_mut : + forall (T : Type), usize -> mut_raw_ptr (slice T) -> result (mut_raw_ptr T). + +(* [core::slice::index::usize::index]: forward function *) +Axiom core_slice_index_usize_index : forall (T : Type), usize -> slice T -> result T. + +(* [core::slice::index::usize::index_mut]: forward function *) +Axiom core_slice_index_usize_index_mut : forall (T : Type), usize -> slice T -> result T. + +(* [core::slice::index::usize::index_mut]: backward function 0 *) +Axiom core_slice_index_usize_index_mut_back : + forall (T : Type), usize -> slice T -> T -> result (slice T). + +(* Trait implementation: [core::slice::index::private_slice_index::usize] *) +Definition core_slice_index_private_slice_index_usize_coresliceindexprivate_slice_indexSealedInst + : core_slice_index_private_slice_index_Sealed usize := tt. + +(* Trait implementation: [core::slice::index::usize] *) +Definition core_slice_index_usize_coresliceindexSliceIndexInst (T : Type) : + core_slice_index_SliceIndex usize (slice T) := {| + core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_usize_coresliceindexprivate_slice_indexSealedInst; + core_slice_index_SliceIndex_Output := T; + core_slice_index_SliceIndex_get := core_slice_index_usize_get T; + core_slice_index_SliceIndex_get_mut := core_slice_index_usize_get_mut T; + core_slice_index_SliceIndex_get_mut_back := core_slice_index_usize_get_mut_back T; + core_slice_index_SliceIndex_get_unchecked := core_slice_index_usize_get_unchecked T; + core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_usize_get_unchecked_mut T; + core_slice_index_SliceIndex_index := core_slice_index_usize_index T; + core_slice_index_SliceIndex_index_mut := core_slice_index_usize_index_mut T; + core_slice_index_SliceIndex_index_mut_back := core_slice_index_usize_index_mut_back T; +|}. + +(* [alloc::vec::Vec::index]: forward function *) +Axiom alloc_vec_Vec_index : forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)) + (Self : alloc_vec_Vec T) (i : Idx), result inst.(core_slice_index_SliceIndex_Output). + +(* [alloc::vec::Vec::index_mut]: forward function *) +Axiom alloc_vec_Vec_index_mut : forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)) + (Self : alloc_vec_Vec T) (i : Idx), result inst.(core_slice_index_SliceIndex_Output). + +(* [alloc::vec::Vec::index_mut]: backward function 0 *) +Axiom alloc_vec_Vec_index_mut_back : + forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)) + (Self : alloc_vec_Vec T) (i : Idx) (x : inst.(core_slice_index_SliceIndex_Output)), result (alloc_vec_Vec T). + +(* Trait implementation: [alloc::vec::Vec] *) +Definition alloc_vec_Vec_coreopsindexIndexInst (T Idx : Type) + (inst : core_slice_index_SliceIndex Idx (slice T)) : + core_ops_index_Index (alloc_vec_Vec T) Idx := {| + core_ops_index_Index_Output := inst.(core_slice_index_SliceIndex_Output); + core_ops_index_Index_index := alloc_vec_Vec_index T Idx inst; +|}. + +(* Trait implementation: [alloc::vec::Vec] *) +Definition alloc_vec_Vec_coreopsindexIndexMutInst (T Idx : Type) + (inst : core_slice_index_SliceIndex Idx (slice T)) : + core_ops_index_IndexMut (alloc_vec_Vec T) Idx := {| + core_ops_index_IndexMut_indexInst := alloc_vec_Vec_coreopsindexIndexInst T Idx inst; + core_ops_index_IndexMut_index_mut := alloc_vec_Vec_index_mut T Idx inst; + core_ops_index_IndexMut_index_mut_back := alloc_vec_Vec_index_mut_back T Idx inst; +|}. + +(*** Theorems *) + +Axiom alloc_vec_Vec_index_mut_back_eq : forall {a : Type} (v : alloc_vec_Vec a) (i : usize) (x : a), + alloc_vec_Vec_index_mut_back a usize (core_slice_index_usize_coresliceindexSliceIndexInst a) v i x = + alloc_vec_Vec_update_usize v i x. End Primitives. -- cgit v1.2.3 From 49ffc966cfdbd71f8c83a3c72ab81e1bb101f420 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Thu, 9 Nov 2023 16:24:07 +0100 Subject: Regenerate the Coq test files --- tests/coq/hashmap/Hashmap_Types.v | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) (limited to 'tests/coq/hashmap') diff --git a/tests/coq/hashmap/Hashmap_Types.v b/tests/coq/hashmap/Hashmap_Types.v index 63d30eeb..8529803d 100644 --- a/tests/coq/hashmap/Hashmap_Types.v +++ b/tests/coq/hashmap/Hashmap_Types.v @@ -14,8 +14,8 @@ Inductive List_t (T : Type) := | List_Nil : List_t T . -Arguments List_Cons {T} _ _ _. -Arguments List_Nil {T}. +Arguments List_Cons { _ }. +Arguments List_Nil { _ }. (** [hashmap::HashMap] *) Record HashMap_t (T : Type) := @@ -27,10 +27,10 @@ mkHashMap_t { } . -Arguments mkHashMap_t {T} _ _ _ _. -Arguments hashMap_num_entries {T}. -Arguments hashMap_max_load_factor {T}. -Arguments hashMap_max_load {T}. -Arguments hashMap_slots {T}. +Arguments mkHashMap_t { _ }. +Arguments hashMap_num_entries { _ }. +Arguments hashMap_max_load_factor { _ }. +Arguments hashMap_max_load { _ }. +Arguments hashMap_slots { _ }. End Hashmap_Types . -- cgit v1.2.3 From 5e92ae6b361f9221f5c5f9a39ab4c28a36597a77 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 21 Nov 2023 11:40:59 +0100 Subject: Regenerate most of the test files --- tests/coq/hashmap/Hashmap_Funs.v | 106 +++++++++++++++++++-------------------- tests/coq/hashmap/Primitives.v | 88 +++++++++++++++++--------------- 2 files changed, 100 insertions(+), 94 deletions(-) (limited to 'tests/coq/hashmap') diff --git a/tests/coq/hashmap/Hashmap_Funs.v b/tests/coq/hashmap/Hashmap_Funs.v index 3ca52a9f..fbed86b5 100644 --- a/tests/coq/hashmap/Hashmap_Funs.v +++ b/tests/coq/hashmap/Hashmap_Funs.v @@ -14,7 +14,7 @@ Module Hashmap_Funs. Definition hash_key (k : usize) : result usize := Return k. -(** [hashmap::HashMap::{0}::allocate_slots]: loop 0: forward function *) +(** [hashmap::{hashmap::HashMap}::allocate_slots]: loop 0: forward function *) Fixpoint hashMap_allocate_slots_loop (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n0 : usize) : result (alloc_vec_Vec (List_t T)) @@ -31,7 +31,7 @@ Fixpoint hashMap_allocate_slots_loop end . -(** [hashmap::HashMap::{0}::allocate_slots]: forward function *) +(** [hashmap::{hashmap::HashMap}::allocate_slots]: forward function *) Definition hashMap_allocate_slots (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n0 : usize) : result (alloc_vec_Vec (List_t T)) @@ -39,7 +39,7 @@ Definition hashMap_allocate_slots hashMap_allocate_slots_loop T n slots n0 . -(** [hashmap::HashMap::{0}::new_with_capacity]: forward function *) +(** [hashmap::{hashmap::HashMap}::new_with_capacity]: forward function *) Definition hashMap_new_with_capacity (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : @@ -58,12 +58,12 @@ Definition hashMap_new_with_capacity |} . -(** [hashmap::HashMap::{0}::new]: forward function *) +(** [hashmap::{hashmap::HashMap}::new]: forward function *) 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::{0}::clear]: loop 0: merged forward/backward function +(** [hashmap::{hashmap::HashMap}::clear]: loop 0: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) Fixpoint hashMap_clear_loop (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (i : usize) : @@ -78,14 +78,14 @@ Fixpoint hashMap_clear_loop i1 <- usize_add i 1%usize; slots0 <- alloc_vec_Vec_index_mut_back (List_t T) usize - (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) - slots i List_Nil; + (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) slots i + List_Nil; hashMap_clear_loop T n0 slots0 i1) else Return slots end . -(** [hashmap::HashMap::{0}::clear]: merged forward/backward function +(** [hashmap::{hashmap::HashMap}::clear]: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) Definition hashMap_clear (T : Type) (n : nat) (self : HashMap_t T) : result (HashMap_t T) := @@ -99,12 +99,12 @@ Definition hashMap_clear |} . -(** [hashmap::HashMap::{0}::len]: forward function *) +(** [hashmap::{hashmap::HashMap}::len]: forward function *) Definition hashMap_len (T : Type) (self : HashMap_t T) : result usize := Return self.(hashMap_num_entries) . -(** [hashmap::HashMap::{0}::insert_in_list]: loop 0: forward function *) +(** [hashmap::{hashmap::HashMap}::insert_in_list]: loop 0: forward function *) Fixpoint hashMap_insert_in_list_loop (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : result bool @@ -122,7 +122,7 @@ Fixpoint hashMap_insert_in_list_loop end . -(** [hashmap::HashMap::{0}::insert_in_list]: forward function *) +(** [hashmap::{hashmap::HashMap}::insert_in_list]: forward function *) Definition hashMap_insert_in_list (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : result bool @@ -130,7 +130,7 @@ Definition hashMap_insert_in_list hashMap_insert_in_list_loop T n key value ls . -(** [hashmap::HashMap::{0}::insert_in_list]: loop 0: backward function 0 *) +(** [hashmap::{hashmap::HashMap}::insert_in_list]: loop 0: backward function 0 *) Fixpoint hashMap_insert_in_list_loop_back (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : result (List_t T) @@ -150,7 +150,7 @@ Fixpoint hashMap_insert_in_list_loop_back end . -(** [hashmap::HashMap::{0}::insert_in_list]: backward function 0 *) +(** [hashmap::{hashmap::HashMap}::insert_in_list]: backward function 0 *) Definition hashMap_insert_in_list_back (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : result (List_t T) @@ -158,7 +158,7 @@ Definition hashMap_insert_in_list_back hashMap_insert_in_list_loop_back T n key value ls . -(** [hashmap::HashMap::{0}::insert_no_resize]: merged forward/backward function +(** [hashmap::{hashmap::HashMap}::insert_no_resize]: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) Definition hashMap_insert_no_resize (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) : @@ -169,7 +169,7 @@ Definition hashMap_insert_no_resize hash_mod <- usize_rem hash i; l <- alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) self.(hashMap_slots) hash_mod; inserted <- hashMap_insert_in_list T n key value l; if inserted @@ -178,7 +178,7 @@ Definition hashMap_insert_no_resize 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_usize_coresliceindexSliceIndexInst (List_t T)) + (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) self.(hashMap_slots) hash_mod l0; Return {| @@ -191,7 +191,7 @@ Definition hashMap_insert_no_resize 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_usize_coresliceindexSliceIndexInst (List_t T)) + (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) self.(hashMap_slots) hash_mod l0; Return {| @@ -202,7 +202,7 @@ Definition hashMap_insert_no_resize |}) . -(** [hashmap::HashMap::{0}::move_elements_from_list]: loop 0: merged forward/backward function +(** [hashmap::{hashmap::HashMap}::move_elements_from_list]: loop 0: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) Fixpoint hashMap_move_elements_from_list_loop (T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) : @@ -220,7 +220,7 @@ Fixpoint hashMap_move_elements_from_list_loop end . -(** [hashmap::HashMap::{0}::move_elements_from_list]: merged forward/backward function +(** [hashmap::{hashmap::HashMap}::move_elements_from_list]: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) Definition hashMap_move_elements_from_list (T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) : @@ -229,7 +229,7 @@ Definition hashMap_move_elements_from_list hashMap_move_elements_from_list_loop T n ntable ls . -(** [hashmap::HashMap::{0}::move_elements]: loop 0: merged forward/backward function +(** [hashmap::{hashmap::HashMap}::move_elements]: loop 0: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) Fixpoint hashMap_move_elements_loop (T : Type) (n : nat) (ntable : HashMap_t T) @@ -244,22 +244,20 @@ Fixpoint hashMap_move_elements_loop then ( l <- alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) - slots i; + (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_usize_coresliceindexSliceIndexInst (List_t T)) - slots i l0; + (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) slots i l0; hashMap_move_elements_loop T n0 ntable0 slots0 i1) else Return (ntable, slots) end . -(** [hashmap::HashMap::{0}::move_elements]: merged forward/backward function +(** [hashmap::{hashmap::HashMap}::move_elements]: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) Definition hashMap_move_elements (T : Type) (n : nat) (ntable : HashMap_t T) @@ -269,7 +267,7 @@ Definition hashMap_move_elements hashMap_move_elements_loop T n ntable slots i . -(** [hashmap::HashMap::{0}::try_resize]: merged forward/backward function +(** [hashmap::{hashmap::HashMap}::try_resize]: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) Definition hashMap_try_resize (T : Type) (n : nat) (self : HashMap_t T) : result (HashMap_t T) := @@ -301,7 +299,7 @@ Definition hashMap_try_resize |} . -(** [hashmap::HashMap::{0}::insert]: merged forward/backward function +(** [hashmap::{hashmap::HashMap}::insert]: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) Definition hashMap_insert (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) : @@ -314,7 +312,7 @@ Definition hashMap_insert else Return self0 . -(** [hashmap::HashMap::{0}::contains_key_in_list]: loop 0: forward function *) +(** [hashmap::{hashmap::HashMap}::contains_key_in_list]: loop 0: forward function *) Fixpoint hashMap_contains_key_in_list_loop (T : Type) (n : nat) (key : usize) (ls : List_t T) : result bool := match n with @@ -330,13 +328,13 @@ Fixpoint hashMap_contains_key_in_list_loop end . -(** [hashmap::HashMap::{0}::contains_key_in_list]: forward function *) +(** [hashmap::{hashmap::HashMap}::contains_key_in_list]: forward function *) 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::{0}::contains_key]: forward function *) +(** [hashmap::{hashmap::HashMap}::contains_key]: forward function *) Definition hashMap_contains_key (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result bool := hash <- hash_key key; @@ -344,12 +342,12 @@ Definition hashMap_contains_key hash_mod <- usize_rem hash i; l <- alloc_vec_Vec_index (List_t T) usize - (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) self.(hashMap_slots) hash_mod; hashMap_contains_key_in_list T n key l . -(** [hashmap::HashMap::{0}::get_in_list]: loop 0: forward function *) +(** [hashmap::{hashmap::HashMap}::get_in_list]: loop 0: forward function *) Fixpoint hashMap_get_in_list_loop (T : Type) (n : nat) (key : usize) (ls : List_t T) : result T := match n with @@ -365,13 +363,13 @@ Fixpoint hashMap_get_in_list_loop end . -(** [hashmap::HashMap::{0}::get_in_list]: forward function *) +(** [hashmap::{hashmap::HashMap}::get_in_list]: forward function *) 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::{0}::get]: forward function *) +(** [hashmap::{hashmap::HashMap}::get]: forward function *) Definition hashMap_get (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result T := hash <- hash_key key; @@ -379,12 +377,12 @@ Definition hashMap_get hash_mod <- usize_rem hash i; l <- alloc_vec_Vec_index (List_t T) usize - (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) self.(hashMap_slots) hash_mod; hashMap_get_in_list T n key l . -(** [hashmap::HashMap::{0}::get_mut_in_list]: loop 0: forward function *) +(** [hashmap::{hashmap::HashMap}::get_mut_in_list]: loop 0: forward function *) Fixpoint hashMap_get_mut_in_list_loop (T : Type) (n : nat) (ls : List_t T) (key : usize) : result T := match n with @@ -400,13 +398,13 @@ Fixpoint hashMap_get_mut_in_list_loop end . -(** [hashmap::HashMap::{0}::get_mut_in_list]: forward function *) +(** [hashmap::{hashmap::HashMap}::get_mut_in_list]: forward function *) 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::{0}::get_mut_in_list]: loop 0: backward function 0 *) +(** [hashmap::{hashmap::HashMap}::get_mut_in_list]: loop 0: backward function 0 *) Fixpoint hashMap_get_mut_in_list_loop_back (T : Type) (n : nat) (ls : List_t T) (key : usize) (ret : T) : result (List_t T) @@ -426,7 +424,7 @@ Fixpoint hashMap_get_mut_in_list_loop_back end . -(** [hashmap::HashMap::{0}::get_mut_in_list]: backward function 0 *) +(** [hashmap::{hashmap::HashMap}::get_mut_in_list]: backward function 0 *) Definition hashMap_get_mut_in_list_back (T : Type) (n : nat) (ls : List_t T) (key : usize) (ret : T) : result (List_t T) @@ -434,7 +432,7 @@ Definition hashMap_get_mut_in_list_back hashMap_get_mut_in_list_loop_back T n ls key ret . -(** [hashmap::HashMap::{0}::get_mut]: forward function *) +(** [hashmap::{hashmap::HashMap}::get_mut]: forward function *) Definition hashMap_get_mut (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result T := hash <- hash_key key; @@ -442,12 +440,12 @@ Definition hashMap_get_mut hash_mod <- usize_rem hash i; l <- alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) self.(hashMap_slots) hash_mod; hashMap_get_mut_in_list T n l key . -(** [hashmap::HashMap::{0}::get_mut]: backward function 0 *) +(** [hashmap::{hashmap::HashMap}::get_mut]: backward function 0 *) Definition hashMap_get_mut_back (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (ret : T) : result (HashMap_t T) @@ -457,12 +455,12 @@ Definition hashMap_get_mut_back hash_mod <- usize_rem hash i; l <- alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + (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_usize_coresliceindexSliceIndexInst (List_t T)) + (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) self.(hashMap_slots) hash_mod l0; Return {| @@ -473,7 +471,7 @@ Definition hashMap_get_mut_back |} . -(** [hashmap::HashMap::{0}::remove_from_list]: loop 0: forward function *) +(** [hashmap::{hashmap::HashMap}::remove_from_list]: loop 0: forward function *) Fixpoint hashMap_remove_from_list_loop (T : Type) (n : nat) (key : usize) (ls : List_t T) : result (option T) := match n with @@ -495,13 +493,13 @@ Fixpoint hashMap_remove_from_list_loop end . -(** [hashmap::HashMap::{0}::remove_from_list]: forward function *) +(** [hashmap::{hashmap::HashMap}::remove_from_list]: forward function *) Definition hashMap_remove_from_list (T : Type) (n : nat) (key : usize) (ls : List_t T) : result (option T) := hashMap_remove_from_list_loop T n key ls . -(** [hashmap::HashMap::{0}::remove_from_list]: loop 0: backward function 1 *) +(** [hashmap::{hashmap::HashMap}::remove_from_list]: loop 0: backward function 1 *) Fixpoint hashMap_remove_from_list_loop_back (T : Type) (n : nat) (key : usize) (ls : List_t T) : result (List_t T) := match n with @@ -525,13 +523,13 @@ Fixpoint hashMap_remove_from_list_loop_back end . -(** [hashmap::HashMap::{0}::remove_from_list]: backward function 1 *) +(** [hashmap::{hashmap::HashMap}::remove_from_list]: backward function 1 *) 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::{0}::remove]: forward function *) +(** [hashmap::{hashmap::HashMap}::remove]: forward function *) Definition hashMap_remove (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result (option T) @@ -541,7 +539,7 @@ Definition hashMap_remove hash_mod <- usize_rem hash i; l <- alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) self.(hashMap_slots) hash_mod; x <- hashMap_remove_from_list T n key l; match x with @@ -551,7 +549,7 @@ Definition hashMap_remove end . -(** [hashmap::HashMap::{0}::remove]: backward function 0 *) +(** [hashmap::{hashmap::HashMap}::remove]: backward function 0 *) Definition hashMap_remove_back (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result (HashMap_t T) @@ -561,7 +559,7 @@ Definition hashMap_remove_back hash_mod <- usize_rem hash i; l <- alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) self.(hashMap_slots) hash_mod; x <- hashMap_remove_from_list T n key l; match x with @@ -569,7 +567,7 @@ Definition hashMap_remove_back l0 <- hashMap_remove_from_list_back T n key l; v <- alloc_vec_Vec_index_mut_back (List_t T) usize - (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) self.(hashMap_slots) hash_mod l0; Return {| @@ -583,7 +581,7 @@ Definition hashMap_remove_back l0 <- hashMap_remove_from_list_back T n key l; v <- alloc_vec_Vec_index_mut_back (List_t T) usize - (core_slice_index_usize_coresliceindexSliceIndexInst (List_t T)) + (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) self.(hashMap_slots) hash_mod l0; Return {| diff --git a/tests/coq/hashmap/Primitives.v b/tests/coq/hashmap/Primitives.v index 85e38f01..83f860b6 100644 --- a/tests/coq/hashmap/Primitives.v +++ b/tests/coq/hashmap/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. -- cgit v1.2.3 From 137cc7335e64fcb70c254e7fd2a6fa353fb43e61 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 21 Nov 2023 14:57:38 +0100 Subject: Regenerate the files --- tests/coq/hashmap/Hashmap_Funs.v | 114 +++++++++++++++++++++++++------------- tests/coq/hashmap/Hashmap_Types.v | 6 +- 2 files changed, 80 insertions(+), 40 deletions(-) (limited to 'tests/coq/hashmap') diff --git a/tests/coq/hashmap/Hashmap_Funs.v b/tests/coq/hashmap/Hashmap_Funs.v index fbed86b5..c08f7f7d 100644 --- a/tests/coq/hashmap/Hashmap_Funs.v +++ b/tests/coq/hashmap/Hashmap_Funs.v @@ -10,11 +10,13 @@ Require Export Hashmap_Types. Import Hashmap_Types. Module Hashmap_Funs. -(** [hashmap::hash_key]: forward function *) +(** [hashmap::hash_key]: forward function + Source: 'src/hashmap.rs', lines 27:0-27:32 *) Definition hash_key (k : usize) : result usize := Return k. -(** [hashmap::{hashmap::HashMap}::allocate_slots]: loop 0: forward function *) +(** [hashmap::{hashmap::HashMap}::allocate_slots]: loop 0: forward function + 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) : result (alloc_vec_Vec (List_t T)) @@ -31,7 +33,8 @@ Fixpoint hashMap_allocate_slots_loop end . -(** [hashmap::{hashmap::HashMap}::allocate_slots]: forward function *) +(** [hashmap::{hashmap::HashMap}::allocate_slots]: forward function + 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) : result (alloc_vec_Vec (List_t T)) @@ -39,7 +42,8 @@ Definition hashMap_allocate_slots hashMap_allocate_slots_loop T n slots n0 . -(** [hashmap::{hashmap::HashMap}::new_with_capacity]: forward function *) +(** [hashmap::{hashmap::HashMap}::new_with_capacity]: forward function + Source: 'src/hashmap.rs', lines 59:4-63:13 *) Definition hashMap_new_with_capacity (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : @@ -58,13 +62,15 @@ Definition hashMap_new_with_capacity |} . -(** [hashmap::{hashmap::HashMap}::new]: forward function *) +(** [hashmap::{hashmap::HashMap}::new]: forward function + 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}::clear]: loop 0: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) + (there is a single backward function, and the forward function returns ()) + 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) : result (alloc_vec_Vec (List_t T)) @@ -86,7 +92,8 @@ Fixpoint hashMap_clear_loop . (** [hashmap::{hashmap::HashMap}::clear]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) + (there is a single backward function, and the forward function returns ()) + 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; @@ -99,12 +106,14 @@ Definition hashMap_clear |} . -(** [hashmap::{hashmap::HashMap}::len]: forward function *) +(** [hashmap::{hashmap::HashMap}::len]: forward function + 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}::insert_in_list]: loop 0: forward function *) +(** [hashmap::{hashmap::HashMap}::insert_in_list]: loop 0: forward function + 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 @@ -122,7 +131,8 @@ Fixpoint hashMap_insert_in_list_loop end . -(** [hashmap::{hashmap::HashMap}::insert_in_list]: forward function *) +(** [hashmap::{hashmap::HashMap}::insert_in_list]: forward function + 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 @@ -130,7 +140,8 @@ Definition hashMap_insert_in_list hashMap_insert_in_list_loop T n key value ls . -(** [hashmap::{hashmap::HashMap}::insert_in_list]: loop 0: backward function 0 *) +(** [hashmap::{hashmap::HashMap}::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) @@ -150,7 +161,8 @@ Fixpoint hashMap_insert_in_list_loop_back end . -(** [hashmap::{hashmap::HashMap}::insert_in_list]: backward function 0 *) +(** [hashmap::{hashmap::HashMap}::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) @@ -159,7 +171,8 @@ Definition hashMap_insert_in_list_back . (** [hashmap::{hashmap::HashMap}::insert_no_resize]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) + (there is a single backward function, and the forward function returns ()) + 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) : result (HashMap_t T) @@ -203,7 +216,8 @@ Definition hashMap_insert_no_resize . (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: loop 0: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) + (there is a single backward function, and the forward function returns ()) + 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) : result (HashMap_t T) @@ -221,7 +235,8 @@ Fixpoint hashMap_move_elements_from_list_loop . (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) + (there is a single backward function, and the forward function returns ()) + 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) : result (HashMap_t T) @@ -230,7 +245,8 @@ Definition hashMap_move_elements_from_list . (** [hashmap::{hashmap::HashMap}::move_elements]: loop 0: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) + (there is a single backward function, and the forward function returns ()) + Source: 'src/hashmap.rs', lines 171:4-180:5 *) Fixpoint hashMap_move_elements_loop (T : Type) (n : nat) (ntable : HashMap_t T) (slots : alloc_vec_Vec (List_t T)) (i : usize) : @@ -258,7 +274,8 @@ Fixpoint hashMap_move_elements_loop . (** [hashmap::{hashmap::HashMap}::move_elements]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) + (there is a single backward function, and the forward function returns ()) + 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) : @@ -268,7 +285,8 @@ Definition hashMap_move_elements . (** [hashmap::{hashmap::HashMap}::try_resize]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) + (there is a single backward function, and the forward function returns ()) + 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; @@ -300,7 +318,8 @@ Definition hashMap_try_resize . (** [hashmap::{hashmap::HashMap}::insert]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) + (there is a single backward function, and the forward function returns ()) + 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) @@ -312,7 +331,8 @@ Definition hashMap_insert else Return self0 . -(** [hashmap::{hashmap::HashMap}::contains_key_in_list]: loop 0: forward function *) +(** [hashmap::{hashmap::HashMap}::contains_key_in_list]: loop 0: forward function + 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 @@ -328,13 +348,15 @@ Fixpoint hashMap_contains_key_in_list_loop end . -(** [hashmap::{hashmap::HashMap}::contains_key_in_list]: forward function *) +(** [hashmap::{hashmap::HashMap}::contains_key_in_list]: forward function + 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}::contains_key]: forward function *) +(** [hashmap::{hashmap::HashMap}::contains_key]: forward function + 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 := hash <- hash_key key; @@ -347,7 +369,8 @@ Definition hashMap_contains_key hashMap_contains_key_in_list T n key l . -(** [hashmap::{hashmap::HashMap}::get_in_list]: loop 0: forward function *) +(** [hashmap::{hashmap::HashMap}::get_in_list]: loop 0: forward function + 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 @@ -363,13 +386,15 @@ Fixpoint hashMap_get_in_list_loop end . -(** [hashmap::{hashmap::HashMap}::get_in_list]: forward function *) +(** [hashmap::{hashmap::HashMap}::get_in_list]: forward function + 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}::get]: forward function *) +(** [hashmap::{hashmap::HashMap}::get]: forward function + Source: 'src/hashmap.rs', lines 239:4-239:55 *) Definition hashMap_get (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result T := hash <- hash_key key; @@ -382,7 +407,8 @@ Definition hashMap_get hashMap_get_in_list T n key l . -(** [hashmap::{hashmap::HashMap}::get_mut_in_list]: loop 0: forward function *) +(** [hashmap::{hashmap::HashMap}::get_mut_in_list]: loop 0: forward function + 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 @@ -398,13 +424,15 @@ Fixpoint hashMap_get_mut_in_list_loop end . -(** [hashmap::{hashmap::HashMap}::get_mut_in_list]: forward function *) +(** [hashmap::{hashmap::HashMap}::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}::get_mut_in_list]: loop 0: backward function 0 *) +(** [hashmap::{hashmap::HashMap}::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) @@ -424,7 +452,8 @@ Fixpoint hashMap_get_mut_in_list_loop_back end . -(** [hashmap::{hashmap::HashMap}::get_mut_in_list]: backward function 0 *) +(** [hashmap::{hashmap::HashMap}::get_mut_in_list]: backward function 0 + 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) @@ -432,7 +461,8 @@ Definition hashMap_get_mut_in_list_back hashMap_get_mut_in_list_loop_back T n ls key ret . -(** [hashmap::{hashmap::HashMap}::get_mut]: forward function *) +(** [hashmap::{hashmap::HashMap}::get_mut]: forward function + 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; @@ -445,7 +475,8 @@ Definition hashMap_get_mut hashMap_get_mut_in_list T n l key . -(** [hashmap::{hashmap::HashMap}::get_mut]: backward function 0 *) +(** [hashmap::{hashmap::HashMap}::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) @@ -471,7 +502,8 @@ Definition hashMap_get_mut_back |} . -(** [hashmap::{hashmap::HashMap}::remove_from_list]: loop 0: forward function *) +(** [hashmap::{hashmap::HashMap}::remove_from_list]: loop 0: forward function + 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) := match n with @@ -493,13 +525,15 @@ Fixpoint hashMap_remove_from_list_loop end . -(** [hashmap::{hashmap::HashMap}::remove_from_list]: forward function *) +(** [hashmap::{hashmap::HashMap}::remove_from_list]: forward function + 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) := hashMap_remove_from_list_loop T n key ls . -(** [hashmap::{hashmap::HashMap}::remove_from_list]: loop 0: backward function 1 *) +(** [hashmap::{hashmap::HashMap}::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 @@ -523,13 +557,15 @@ Fixpoint hashMap_remove_from_list_loop_back end . -(** [hashmap::{hashmap::HashMap}::remove_from_list]: backward function 1 *) +(** [hashmap::{hashmap::HashMap}::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}::remove]: forward function *) +(** [hashmap::{hashmap::HashMap}::remove]: forward function + 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) @@ -549,7 +585,8 @@ Definition hashMap_remove end . -(** [hashmap::{hashmap::HashMap}::remove]: backward function 0 *) +(** [hashmap::{hashmap::HashMap}::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) @@ -593,7 +630,8 @@ Definition hashMap_remove_back end . -(** [hashmap::test1]: forward function *) +(** [hashmap::test1]: forward function + 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; diff --git a/tests/coq/hashmap/Hashmap_Types.v b/tests/coq/hashmap/Hashmap_Types.v index 8529803d..bfb5ae4b 100644 --- a/tests/coq/hashmap/Hashmap_Types.v +++ b/tests/coq/hashmap/Hashmap_Types.v @@ -8,7 +8,8 @@ Import ListNotations. Local Open Scope Primitives_scope. Module Hashmap_Types. -(** [hashmap::List] *) +(** [hashmap::List] + Source: 'src/hashmap.rs', lines 19:0-19:16 *) Inductive List_t (T : Type) := | List_Cons : usize -> T -> List_t T -> List_t T | List_Nil : List_t T @@ -17,7 +18,8 @@ Inductive List_t (T : Type) := Arguments List_Cons { _ }. Arguments List_Nil { _ }. -(** [hashmap::HashMap] *) +(** [hashmap::HashMap] + Source: 'src/hashmap.rs', lines 35:0-35:21 *) Record HashMap_t (T : Type) := mkHashMap_t { hashMap_num_entries : usize; -- cgit v1.2.3 From bef2bd34fcb0817f1b7d16b95122bcc3c6f05c72 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 27 Nov 2023 10:29:25 +0100 Subject: Generate a dedicated file for the external types --- tests/coq/hashmap/Hashmap_Funs.v | 6 +++--- tests/coq/hashmap/Hashmap_Types.v | 2 +- 2 files changed, 4 insertions(+), 4 deletions(-) (limited to 'tests/coq/hashmap') diff --git a/tests/coq/hashmap/Hashmap_Funs.v b/tests/coq/hashmap/Hashmap_Funs.v index c08f7f7d..64de44a6 100644 --- a/tests/coq/hashmap/Hashmap_Funs.v +++ b/tests/coq/hashmap/Hashmap_Funs.v @@ -6,8 +6,8 @@ Require Import Coq.ZArith.ZArith. Require Import List. Import ListNotations. Local Open Scope Primitives_scope. -Require Export Hashmap_Types. -Import Hashmap_Types. +Require Import Hashmap_Types. +Include Hashmap_Types. Module Hashmap_Funs. (** [hashmap::hash_key]: forward function @@ -668,4 +668,4 @@ Definition test1 (n : nat) : result unit := end)) . -End Hashmap_Funs . +End Hashmap_Funs. diff --git a/tests/coq/hashmap/Hashmap_Types.v b/tests/coq/hashmap/Hashmap_Types.v index bfb5ae4b..80a43593 100644 --- a/tests/coq/hashmap/Hashmap_Types.v +++ b/tests/coq/hashmap/Hashmap_Types.v @@ -35,4 +35,4 @@ Arguments hashMap_max_load_factor { _ }. Arguments hashMap_max_load { _ }. Arguments hashMap_slots { _ }. -End Hashmap_Types . +End Hashmap_Types. -- cgit v1.2.3 From 3b487893b2906e13b2388efc3512f2babc8514bf Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 29 Nov 2023 15:48:27 +0100 Subject: Regenerate the other files --- tests/coq/hashmap/Primitives.v | 76 ++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 76 insertions(+) (limited to 'tests/coq/hashmap') diff --git a/tests/coq/hashmap/Primitives.v b/tests/coq/hashmap/Primitives.v index 83f860b6..99ffe070 100644 --- a/tests/coq/hashmap/Primitives.v +++ b/tests/coq/hashmap/Primitives.v @@ -255,6 +255,12 @@ Definition scalar_rem {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty Definition scalar_neg {ty} (x: scalar ty) : result (scalar ty) := mk_scalar ty (-(to_Z x)). +Axiom scalar_xor : forall ty, scalar ty -> scalar ty -> scalar ty. (* TODO *) +Axiom scalar_or : forall ty, scalar ty -> scalar ty -> scalar ty. (* TODO *) +Axiom scalar_and : forall ty, scalar ty -> scalar ty -> scalar ty. (* TODO *) +Axiom scalar_shl : forall ty0 ty1, scalar ty0 -> scalar ty1 -> result (scalar ty0). (* TODO *) +Axiom scalar_shr : forall ty0 ty1, scalar ty0 -> scalar ty1 -> result (scalar ty0). (* TODO *) + (** Cast an integer from a [src_ty] to a [tgt_ty] *) (* TODO: check the semantics of casts in Rust *) Definition scalar_cast (src_ty tgt_ty : scalar_ty) (x : scalar src_ty) : result (scalar tgt_ty) := @@ -372,6 +378,76 @@ Definition u32_mul := @scalar_mul U32. Definition u64_mul := @scalar_mul U64. Definition u128_mul := @scalar_mul U128. +(** Xor *) +Definition u8_xor := @scalar_xor U8. +Definition u16_xor := @scalar_xor U16. +Definition u32_xor := @scalar_xor U32. +Definition u64_xor := @scalar_xor U64. +Definition u128_xor := @scalar_xor U128. +Definition usize_xor := @scalar_xor Usize. +Definition i8_xor := @scalar_xor I8. +Definition i16_xor := @scalar_xor I16. +Definition i32_xor := @scalar_xor I32. +Definition i64_xor := @scalar_xor I64. +Definition i128_xor := @scalar_xor I128. +Definition isize_xor := @scalar_xor Isize. + +(** Or *) +Definition u8_or := @scalar_or U8. +Definition u16_or := @scalar_or U16. +Definition u32_or := @scalar_or U32. +Definition u64_or := @scalar_or U64. +Definition u128_or := @scalar_or U128. +Definition usize_or := @scalar_or Usize. +Definition i8_or := @scalar_or I8. +Definition i16_or := @scalar_or I16. +Definition i32_or := @scalar_or I32. +Definition i64_or := @scalar_or I64. +Definition i128_or := @scalar_or I128. +Definition isize_or := @scalar_or Isize. + +(** And *) +Definition u8_and := @scalar_and U8. +Definition u16_and := @scalar_and U16. +Definition u32_and := @scalar_and U32. +Definition u64_and := @scalar_and U64. +Definition u128_and := @scalar_and U128. +Definition usize_and := @scalar_and Usize. +Definition i8_and := @scalar_and I8. +Definition i16_and := @scalar_and I16. +Definition i32_and := @scalar_and I32. +Definition i64_and := @scalar_and I64. +Definition i128_and := @scalar_and I128. +Definition isize_and := @scalar_and Isize. + +(** Shift left *) +Definition u8_shl {ty} := @scalar_shl U8 ty. +Definition u16_shl {ty} := @scalar_shl U16 ty. +Definition u32_shl {ty} := @scalar_shl U32 ty. +Definition u64_shl {ty} := @scalar_shl U64 ty. +Definition u128_shl {ty} := @scalar_shl U128 ty. +Definition usize_shl {ty} := @scalar_shl Usize ty. +Definition i8_shl {ty} := @scalar_shl I8 ty. +Definition i16_shl {ty} := @scalar_shl I16 ty. +Definition i32_shl {ty} := @scalar_shl I32 ty. +Definition i64_shl {ty} := @scalar_shl I64 ty. +Definition i128_shl {ty} := @scalar_shl I128 ty. +Definition isize_shl {ty} := @scalar_shl Isize ty. + +(** Shift right *) +Definition u8_shr {ty} := @scalar_shr U8 ty. +Definition u16_shr {ty} := @scalar_shr U16 ty. +Definition u32_shr {ty} := @scalar_shr U32 ty. +Definition u64_shr {ty} := @scalar_shr U64 ty. +Definition u128_shr {ty} := @scalar_shr U128 ty. +Definition usize_shr {ty} := @scalar_shr Usize ty. +Definition i8_shr {ty} := @scalar_shr I8 ty. +Definition i16_shr {ty} := @scalar_shr I16 ty. +Definition i32_shr {ty} := @scalar_shr I32 ty. +Definition i64_shr {ty} := @scalar_shr I64 ty. +Definition i128_shr {ty} := @scalar_shr I128 ty. +Definition isize_shr {ty} := @scalar_shr Isize ty. + (** Small utility *) Definition usize_to_nat (x: usize) : nat := Z.to_nat (to_Z x). -- cgit v1.2.3