summaryrefslogtreecommitdiff
path: root/tests/lean/Hashmap/Funs.lean
diff options
context:
space:
mode:
authorSon Ho2023-07-04 14:57:51 +0200
committerSon Ho2023-07-04 14:57:51 +0200
commit87d6f6c7c90bf7b427397d6bd2e2c70d610678e3 (patch)
treece6f570b8916db1877e505f719461241bafaed0d /tests/lean/Hashmap/Funs.lean
parent4fd17e4bb91eb46d4704643dfbfbbf0874837b07 (diff)
Reorganize the Lean tests
Diffstat (limited to 'tests/lean/Hashmap/Funs.lean')
-rw-r--r--tests/lean/Hashmap/Funs.lean474
1 files changed, 474 insertions, 0 deletions
diff --git a/tests/lean/Hashmap/Funs.lean b/tests/lean/Hashmap/Funs.lean
new file mode 100644
index 00000000..26742d5d
--- /dev/null
+++ b/tests/lean/Hashmap/Funs.lean
@@ -0,0 +1,474 @@
+-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS
+-- [hashmap]: function definitions
+import Base
+import Hashmap.Types
+open Primitives
+
+/- [hashmap::hash_key] -/
+def hash_key_fwd (k : Usize) : Result Usize :=
+ Result.ret k
+
+/- [hashmap::HashMap::{0}::allocate_slots] -/
+divergent def hash_map_allocate_slots_loop_fwd
+ (T : Type) (slots : Vec (list_t T)) (n : Usize) : Result (Vec (list_t T)) :=
+ if n > (Usize.ofInt 0 (by intlit))
+ then
+ do
+ let slots0 ← vec_push_back (list_t T) slots list_t.Nil
+ let n0 ← n - (Usize.ofInt 1 (by intlit))
+ hash_map_allocate_slots_loop_fwd T slots0 n0
+ else Result.ret slots
+
+/- [hashmap::HashMap::{0}::allocate_slots] -/
+def hash_map_allocate_slots_fwd
+ (T : Type) (slots : Vec (list_t T)) (n : Usize) : Result (Vec (list_t T)) :=
+ hash_map_allocate_slots_loop_fwd T slots n
+
+/- [hashmap::HashMap::{0}::new_with_capacity] -/
+def hash_map_new_with_capacity_fwd
+ (T : Type) (capacity : Usize) (max_load_dividend : Usize)
+ (max_load_divisor : Usize) :
+ Result (hash_map_t T)
+ :=
+ do
+ let v := vec_new (list_t T)
+ let slots ← hash_map_allocate_slots_fwd T v capacity
+ let i ← capacity * max_load_dividend
+ let i0 ← i / max_load_divisor
+ Result.ret
+ {
+ hash_map_num_entries := (Usize.ofInt 0 (by intlit)),
+ hash_map_max_load_factor := (max_load_dividend, max_load_divisor),
+ hash_map_max_load := i0,
+ hash_map_slots := slots
+ }
+
+/- [hashmap::HashMap::{0}::new] -/
+def hash_map_new_fwd (T : Type) : Result (hash_map_t T) :=
+ hash_map_new_with_capacity_fwd T (Usize.ofInt 32 (by intlit))
+ (Usize.ofInt 4 (by intlit)) (Usize.ofInt 5 (by intlit))
+
+/- [hashmap::HashMap::{0}::clear] -/
+divergent def hash_map_clear_loop_fwd_back
+ (T : Type) (slots : Vec (list_t T)) (i : Usize) : Result (Vec (list_t T)) :=
+ let i0 := vec_len (list_t T) slots
+ if i < i0
+ then
+ do
+ let i1 ← i + (Usize.ofInt 1 (by intlit))
+ let slots0 ← vec_index_mut_back (list_t T) slots i list_t.Nil
+ hash_map_clear_loop_fwd_back T slots0 i1
+ else Result.ret slots
+
+/- [hashmap::HashMap::{0}::clear] -/
+def hash_map_clear_fwd_back
+ (T : Type) (self : hash_map_t T) : Result (hash_map_t T) :=
+ do
+ let v ←
+ hash_map_clear_loop_fwd_back T self.hash_map_slots
+ (Usize.ofInt 0 (by intlit))
+ Result.ret
+ {
+ self
+ with
+ hash_map_num_entries := (Usize.ofInt 0 (by intlit)),
+ hash_map_slots := v
+ }
+
+/- [hashmap::HashMap::{0}::len] -/
+def hash_map_len_fwd (T : Type) (self : hash_map_t T) : Result Usize :=
+ Result.ret self.hash_map_num_entries
+
+/- [hashmap::HashMap::{0}::insert_in_list] -/
+divergent def hash_map_insert_in_list_loop_fwd
+ (T : Type) (key : Usize) (value : T) (ls : list_t T) : Result Bool :=
+ match h: ls with
+ | list_t.Cons ckey cvalue tl =>
+ if ckey = key
+ then Result.ret false
+ else hash_map_insert_in_list_loop_fwd T key value tl
+ | list_t.Nil => Result.ret true
+
+/- [hashmap::HashMap::{0}::insert_in_list] -/
+def hash_map_insert_in_list_fwd
+ (T : Type) (key : Usize) (value : T) (ls : list_t T) : Result Bool :=
+ hash_map_insert_in_list_loop_fwd T key value ls
+
+/- [hashmap::HashMap::{0}::insert_in_list] -/
+divergent def hash_map_insert_in_list_loop_back
+ (T : Type) (key : Usize) (value : T) (ls : list_t T) : Result (list_t T) :=
+ match h: ls with
+ | list_t.Cons ckey cvalue tl =>
+ if ckey = key
+ then Result.ret (list_t.Cons ckey value tl)
+ else
+ do
+ let tl0 ← hash_map_insert_in_list_loop_back T key value tl
+ Result.ret (list_t.Cons ckey cvalue tl0)
+ | list_t.Nil => let l := list_t.Nil
+ Result.ret (list_t.Cons key value l)
+
+/- [hashmap::HashMap::{0}::insert_in_list] -/
+def hash_map_insert_in_list_back
+ (T : Type) (key : Usize) (value : T) (ls : list_t T) : Result (list_t T) :=
+ hash_map_insert_in_list_loop_back T key value ls
+
+/- [hashmap::HashMap::{0}::insert_no_resize] -/
+def hash_map_insert_no_resize_fwd_back
+ (T : Type) (self : hash_map_t T) (key : Usize) (value : T) :
+ Result (hash_map_t T)
+ :=
+ do
+ let hash ← hash_key_fwd key
+ let i := vec_len (list_t T) self.hash_map_slots
+ let hash_mod ← hash % i
+ let l ← vec_index_mut_fwd (list_t T) self.hash_map_slots hash_mod
+ let inserted ← hash_map_insert_in_list_fwd T key value l
+ if inserted
+ then
+ do
+ let i0 ← self.hash_map_num_entries + (Usize.ofInt 1 (by intlit))
+ let l0 ← hash_map_insert_in_list_back T key value l
+ let v ← vec_index_mut_back (list_t T) self.hash_map_slots hash_mod l0
+ Result.ret
+ { self with hash_map_num_entries := i0, hash_map_slots := v }
+ else
+ do
+ let l0 ← hash_map_insert_in_list_back T key value l
+ let v ← vec_index_mut_back (list_t T) self.hash_map_slots hash_mod l0
+ Result.ret { self with hash_map_slots := v }
+
+/- [core::num::u32::{9}::MAX] -/
+def core_num_u32_max_body : Result U32 :=
+ Result.ret (U32.ofInt 4294967295 (by intlit))
+def core_num_u32_max_c : U32 := eval_global core_num_u32_max_body (by simp)
+
+/- [hashmap::HashMap::{0}::move_elements_from_list] -/
+divergent def hash_map_move_elements_from_list_loop_fwd_back
+ (T : Type) (ntable : hash_map_t T) (ls : list_t T) : Result (hash_map_t T) :=
+ match h: ls with
+ | list_t.Cons k v tl =>
+ do
+ let ntable0 ← hash_map_insert_no_resize_fwd_back T ntable k v
+ hash_map_move_elements_from_list_loop_fwd_back T ntable0 tl
+ | list_t.Nil => Result.ret ntable
+
+/- [hashmap::HashMap::{0}::move_elements_from_list] -/
+def hash_map_move_elements_from_list_fwd_back
+ (T : Type) (ntable : hash_map_t T) (ls : list_t T) : Result (hash_map_t T) :=
+ hash_map_move_elements_from_list_loop_fwd_back T ntable ls
+
+/- [hashmap::HashMap::{0}::move_elements] -/
+divergent def hash_map_move_elements_loop_fwd_back
+ (T : Type) (ntable : hash_map_t T) (slots : Vec (list_t T)) (i : Usize) :
+ Result ((hash_map_t T) × (Vec (list_t T)))
+ :=
+ let i0 := vec_len (list_t T) slots
+ if i < i0
+ then
+ do
+ let l ← vec_index_mut_fwd (list_t T) slots i
+ let ls := mem_replace_fwd (list_t T) l list_t.Nil
+ let ntable0 ← hash_map_move_elements_from_list_fwd_back T ntable ls
+ let i1 ← i + (Usize.ofInt 1 (by intlit))
+ let l0 := mem_replace_back (list_t T) l list_t.Nil
+ let slots0 ← vec_index_mut_back (list_t T) slots i l0
+ hash_map_move_elements_loop_fwd_back T ntable0 slots0 i1
+ else Result.ret (ntable, slots)
+
+/- [hashmap::HashMap::{0}::move_elements] -/
+def hash_map_move_elements_fwd_back
+ (T : Type) (ntable : hash_map_t T) (slots : Vec (list_t T)) (i : Usize) :
+ Result ((hash_map_t T) × (Vec (list_t T)))
+ :=
+ hash_map_move_elements_loop_fwd_back T ntable slots i
+
+/- [hashmap::HashMap::{0}::try_resize] -/
+def hash_map_try_resize_fwd_back
+ (T : Type) (self : hash_map_t T) : Result (hash_map_t T) :=
+ do
+ let max_usize ← Scalar.cast .Usize core_num_u32_max_c
+ let capacity := vec_len (list_t T) self.hash_map_slots
+ let n1 ← max_usize / (Usize.ofInt 2 (by intlit))
+ let (i, i0) := self.hash_map_max_load_factor
+ let i1 ← n1 / i
+ if capacity <= i1
+ then
+ do
+ let i2 ← capacity * (Usize.ofInt 2 (by intlit))
+ let ntable ← hash_map_new_with_capacity_fwd T i2 i i0
+ let (ntable0, _) ←
+ hash_map_move_elements_fwd_back T ntable self.hash_map_slots
+ (Usize.ofInt 0 (by intlit))
+ Result.ret
+ {
+ ntable0
+ with
+ hash_map_num_entries := self.hash_map_num_entries,
+ hash_map_max_load_factor := (i, i0)
+ }
+ else Result.ret { self with hash_map_max_load_factor := (i, i0) }
+
+/- [hashmap::HashMap::{0}::insert] -/
+def hash_map_insert_fwd_back
+ (T : Type) (self : hash_map_t T) (key : Usize) (value : T) :
+ Result (hash_map_t T)
+ :=
+ do
+ let self0 ← hash_map_insert_no_resize_fwd_back T self key value
+ let i ← hash_map_len_fwd T self0
+ if i > self0.hash_map_max_load
+ then hash_map_try_resize_fwd_back T self0
+ else Result.ret self0
+
+/- [hashmap::HashMap::{0}::contains_key_in_list] -/
+divergent def hash_map_contains_key_in_list_loop_fwd
+ (T : Type) (key : Usize) (ls : list_t T) : Result Bool :=
+ match h: ls with
+ | list_t.Cons ckey t tl =>
+ if ckey = key
+ then Result.ret true
+ else hash_map_contains_key_in_list_loop_fwd T key tl
+ | list_t.Nil => Result.ret false
+
+/- [hashmap::HashMap::{0}::contains_key_in_list] -/
+def hash_map_contains_key_in_list_fwd
+ (T : Type) (key : Usize) (ls : list_t T) : Result Bool :=
+ hash_map_contains_key_in_list_loop_fwd T key ls
+
+/- [hashmap::HashMap::{0}::contains_key] -/
+def hash_map_contains_key_fwd
+ (T : Type) (self : hash_map_t T) (key : Usize) : Result Bool :=
+ do
+ let hash ← hash_key_fwd key
+ let i := vec_len (list_t T) self.hash_map_slots
+ let hash_mod ← hash % i
+ let l ← vec_index_fwd (list_t T) self.hash_map_slots hash_mod
+ hash_map_contains_key_in_list_fwd T key l
+
+/- [hashmap::HashMap::{0}::get_in_list] -/
+divergent def hash_map_get_in_list_loop_fwd
+ (T : Type) (key : Usize) (ls : list_t T) : Result T :=
+ match h: ls with
+ | list_t.Cons ckey cvalue tl =>
+ if ckey = key
+ then Result.ret cvalue
+ else hash_map_get_in_list_loop_fwd T key tl
+ | list_t.Nil => Result.fail Error.panic
+
+/- [hashmap::HashMap::{0}::get_in_list] -/
+def hash_map_get_in_list_fwd
+ (T : Type) (key : Usize) (ls : list_t T) : Result T :=
+ hash_map_get_in_list_loop_fwd T key ls
+
+/- [hashmap::HashMap::{0}::get] -/
+def hash_map_get_fwd
+ (T : Type) (self : hash_map_t T) (key : Usize) : Result T :=
+ do
+ let hash ← hash_key_fwd key
+ let i := vec_len (list_t T) self.hash_map_slots
+ let hash_mod ← hash % i
+ let l ← vec_index_fwd (list_t T) self.hash_map_slots hash_mod
+ hash_map_get_in_list_fwd T key l
+
+/- [hashmap::HashMap::{0}::get_mut_in_list] -/
+divergent def hash_map_get_mut_in_list_loop_fwd
+ (T : Type) (ls : list_t T) (key : Usize) : Result T :=
+ match h: ls with
+ | list_t.Cons ckey cvalue tl =>
+ if ckey = key
+ then Result.ret cvalue
+ else hash_map_get_mut_in_list_loop_fwd T tl key
+ | list_t.Nil => Result.fail Error.panic
+
+/- [hashmap::HashMap::{0}::get_mut_in_list] -/
+def hash_map_get_mut_in_list_fwd
+ (T : Type) (ls : list_t T) (key : Usize) : Result T :=
+ hash_map_get_mut_in_list_loop_fwd T ls key
+
+/- [hashmap::HashMap::{0}::get_mut_in_list] -/
+divergent def hash_map_get_mut_in_list_loop_back
+ (T : Type) (ls : list_t T) (key : Usize) (ret0 : T) : Result (list_t T) :=
+ match h: ls with
+ | list_t.Cons ckey cvalue tl =>
+ if ckey = key
+ then Result.ret (list_t.Cons ckey ret0 tl)
+ else
+ do
+ let tl0 ← hash_map_get_mut_in_list_loop_back T tl key ret0
+ Result.ret (list_t.Cons ckey cvalue tl0)
+ | list_t.Nil => Result.fail Error.panic
+
+/- [hashmap::HashMap::{0}::get_mut_in_list] -/
+def hash_map_get_mut_in_list_back
+ (T : Type) (ls : list_t T) (key : Usize) (ret0 : T) : Result (list_t T) :=
+ hash_map_get_mut_in_list_loop_back T ls key ret0
+
+/- [hashmap::HashMap::{0}::get_mut] -/
+def hash_map_get_mut_fwd
+ (T : Type) (self : hash_map_t T) (key : Usize) : Result T :=
+ do
+ let hash ← hash_key_fwd key
+ let i := vec_len (list_t T) self.hash_map_slots
+ let hash_mod ← hash % i
+ let l ← vec_index_mut_fwd (list_t T) self.hash_map_slots hash_mod
+ hash_map_get_mut_in_list_fwd T l key
+
+/- [hashmap::HashMap::{0}::get_mut] -/
+def hash_map_get_mut_back
+ (T : Type) (self : hash_map_t T) (key : Usize) (ret0 : T) :
+ Result (hash_map_t T)
+ :=
+ do
+ let hash ← hash_key_fwd key
+ let i := vec_len (list_t T) self.hash_map_slots
+ let hash_mod ← hash % i
+ let l ← vec_index_mut_fwd (list_t T) self.hash_map_slots hash_mod
+ let l0 ← hash_map_get_mut_in_list_back T l key ret0
+ let v ← vec_index_mut_back (list_t T) self.hash_map_slots hash_mod l0
+ Result.ret { self with hash_map_slots := v }
+
+/- [hashmap::HashMap::{0}::remove_from_list] -/
+divergent def hash_map_remove_from_list_loop_fwd
+ (T : Type) (key : Usize) (ls : list_t T) : Result (Option T) :=
+ match h: ls with
+ | list_t.Cons ckey t tl =>
+ if ckey = key
+ then
+ let mv_ls :=
+ mem_replace_fwd (list_t T) (list_t.Cons ckey t tl) list_t.Nil
+ match h: mv_ls with
+ | list_t.Cons i cvalue tl0 => Result.ret (Option.some cvalue)
+ | list_t.Nil => Result.fail Error.panic
+ else hash_map_remove_from_list_loop_fwd T key tl
+ | list_t.Nil => Result.ret Option.none
+
+/- [hashmap::HashMap::{0}::remove_from_list] -/
+def hash_map_remove_from_list_fwd
+ (T : Type) (key : Usize) (ls : list_t T) : Result (Option T) :=
+ hash_map_remove_from_list_loop_fwd T key ls
+
+/- [hashmap::HashMap::{0}::remove_from_list] -/
+divergent def hash_map_remove_from_list_loop_back
+ (T : Type) (key : Usize) (ls : list_t T) : Result (list_t T) :=
+ match h: ls with
+ | list_t.Cons ckey t tl =>
+ if ckey = key
+ then
+ let mv_ls :=
+ mem_replace_fwd (list_t T) (list_t.Cons ckey t tl) list_t.Nil
+ match h: mv_ls with
+ | list_t.Cons i cvalue tl0 => Result.ret tl0
+ | list_t.Nil => Result.fail Error.panic
+ else
+ do
+ let tl0 ← hash_map_remove_from_list_loop_back T key tl
+ Result.ret (list_t.Cons ckey t tl0)
+ | list_t.Nil => Result.ret list_t.Nil
+
+/- [hashmap::HashMap::{0}::remove_from_list] -/
+def hash_map_remove_from_list_back
+ (T : Type) (key : Usize) (ls : list_t T) : Result (list_t T) :=
+ hash_map_remove_from_list_loop_back T key ls
+
+/- [hashmap::HashMap::{0}::remove] -/
+def hash_map_remove_fwd
+ (T : Type) (self : hash_map_t T) (key : Usize) : Result (Option T) :=
+ do
+ let hash ← hash_key_fwd key
+ let i := vec_len (list_t T) self.hash_map_slots
+ let hash_mod ← hash % i
+ let l ← vec_index_mut_fwd (list_t T) self.hash_map_slots hash_mod
+ let x ← hash_map_remove_from_list_fwd T key l
+ match h: x with
+ | Option.none => Result.ret Option.none
+ | Option.some x0 =>
+ do
+ let _ ← self.hash_map_num_entries - (Usize.ofInt 1 (by intlit))
+ Result.ret (Option.some x0)
+
+/- [hashmap::HashMap::{0}::remove] -/
+def hash_map_remove_back
+ (T : Type) (self : hash_map_t T) (key : Usize) : Result (hash_map_t T) :=
+ do
+ let hash ← hash_key_fwd key
+ let i := vec_len (list_t T) self.hash_map_slots
+ let hash_mod ← hash % i
+ let l ← vec_index_mut_fwd (list_t T) self.hash_map_slots hash_mod
+ let x ← hash_map_remove_from_list_fwd T key l
+ match h: x with
+ | Option.none =>
+ do
+ let l0 ← hash_map_remove_from_list_back T key l
+ let v ← vec_index_mut_back (list_t T) self.hash_map_slots hash_mod l0
+ Result.ret { self with hash_map_slots := v }
+ | Option.some x0 =>
+ do
+ let i0 ← self.hash_map_num_entries - (Usize.ofInt 1 (by intlit))
+ let l0 ← hash_map_remove_from_list_back T key l
+ let v ← vec_index_mut_back (list_t T) self.hash_map_slots hash_mod l0
+ Result.ret
+ { self with hash_map_num_entries := i0, hash_map_slots := v }
+
+/- [hashmap::test1] -/
+def test1_fwd : Result Unit :=
+ do
+ let hm ← hash_map_new_fwd U64
+ let hm0 ←
+ hash_map_insert_fwd_back U64 hm (Usize.ofInt 0 (by intlit))
+ (U64.ofInt 42 (by intlit))
+ let hm1 ←
+ hash_map_insert_fwd_back U64 hm0 (Usize.ofInt 128 (by intlit))
+ (U64.ofInt 18 (by intlit))
+ let hm2 ←
+ hash_map_insert_fwd_back U64 hm1 (Usize.ofInt 1024 (by intlit))
+ (U64.ofInt 138 (by intlit))
+ let hm3 ←
+ hash_map_insert_fwd_back U64 hm2 (Usize.ofInt 1056 (by intlit))
+ (U64.ofInt 256 (by intlit))
+ let i ← hash_map_get_fwd U64 hm3 (Usize.ofInt 128 (by intlit))
+ if not (i = (U64.ofInt 18 (by intlit)))
+ then Result.fail Error.panic
+ else
+ do
+ let hm4 ←
+ hash_map_get_mut_back U64 hm3 (Usize.ofInt 1024 (by intlit))
+ (U64.ofInt 56 (by intlit))
+ let i0 ← hash_map_get_fwd U64 hm4 (Usize.ofInt 1024 (by intlit))
+ if not (i0 = (U64.ofInt 56 (by intlit)))
+ then Result.fail Error.panic
+ else
+ do
+ let x ←
+ hash_map_remove_fwd U64 hm4 (Usize.ofInt 1024 (by intlit))
+ match h: x with
+ | Option.none => Result.fail Error.panic
+ | Option.some x0 =>
+ if not (x0 = (U64.ofInt 56 (by intlit)))
+ then Result.fail Error.panic
+ else
+ do
+ let hm5 ←
+ hash_map_remove_back U64 hm4 (Usize.ofInt 1024 (by intlit))
+ let i1 ←
+ hash_map_get_fwd U64 hm5 (Usize.ofInt 0 (by intlit))
+ if not (i1 = (U64.ofInt 42 (by intlit)))
+ then Result.fail Error.panic
+ else
+ do
+ let i2 ←
+ hash_map_get_fwd U64 hm5 (Usize.ofInt 128 (by intlit))
+ if not (i2 = (U64.ofInt 18 (by intlit)))
+ then Result.fail Error.panic
+ else
+ do
+ let i3 ←
+ hash_map_get_fwd U64 hm5
+ (Usize.ofInt 1056 (by intlit))
+ if not (i3 = (U64.ofInt 256 (by intlit)))
+ then Result.fail Error.panic
+ else Result.ret ()
+
+/- Unit test for [hashmap::test1] -/
+#assert (test1_fwd == .ret ())
+