summaryrefslogtreecommitdiff
path: root/tests/lean/hashmap/Hashmap/Funs.lean
diff options
context:
space:
mode:
authorSon Ho2023-03-07 08:41:57 +0100
committerSon HO2023-06-04 21:44:33 +0200
commitfeb60683216a6d9193d6353605560c6c80a1ab41 (patch)
tree222f61e4c5cbcd166e81d82350afc54b002774df /tests/lean/hashmap/Hashmap/Funs.lean
parentb4bad8df4eabb17c71dfa7b24d79d62fc06d0a70 (diff)
Make minor modifications and regenerate the Lean files
Diffstat (limited to '')
-rw-r--r--tests/lean/hashmap/Hashmap/Funs.lean400
1 files changed, 200 insertions, 200 deletions
diff --git a/tests/lean/hashmap/Hashmap/Funs.lean b/tests/lean/hashmap/Hashmap/Funs.lean
index 3bdd9dd2..0b3708fa 100644
--- a/tests/lean/hashmap/Hashmap/Funs.lean
+++ b/tests/lean/hashmap/Hashmap/Funs.lean
@@ -5,42 +5,42 @@ import Hashmap.Types
import Hashmap.Clauses.Clauses
/- [hashmap::hash_key] -/
-def hash_key_fwd (k : USize) : result USize :=
- result.ret k
+def hash_key_fwd (k : USize) : Result USize :=
+ Result.ret k
/- [hashmap::HashMap::{0}::allocate_slots] -/
def hash_map_allocate_slots_loop_fwd
- (T : Type) (slots : vec (list_t T)) (n : USize) :
- (result (vec (list_t T)))
+ (T : Type) (slots : Vec (list_t T)) (n : USize) :
+ (Result (Vec (list_t T)))
:=
- if n > (USize.ofNatCore 0 (by intlit))
+ if h: n > (USize.ofNatCore 0 (by intlit))
then
do
- let slots0 <- vec_push_back (list_t T) slots list_t.ListNil
- let n0 <- USize.checked_sub n (USize.ofNatCore 1 (by intlit))
+ let slots0 ← vec_push_back (list_t T) slots list_t.ListNil
+ let n0 ← USize.checked_sub n (USize.ofNatCore 1 (by intlit))
hash_map_allocate_slots_loop_fwd T slots0 n0
- else result.ret slots
+ else Result.ret slots
termination_by hash_map_allocate_slots_loop_fwd slots n =>
hash_map_allocate_slots_loop_terminates T slots n
decreasing_by hash_map_allocate_slots_loop_decreases slots n
/- [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)) :=
+ (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)
+ 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 <- USize.checked_mul capacity max_load_dividend
- let i0 <- USize.checked_div i max_load_divisor
- result.ret
+ let slots ← hash_map_allocate_slots_fwd T v capacity
+ let i ← USize.checked_mul capacity max_load_dividend
+ let i0 ← USize.checked_div i max_load_divisor
+ Result.ret
{
hash_map_num_entries := (USize.ofNatCore 0 (by intlit)),
hash_map_max_load_factor := (max_load_dividend, max_load_divisor),
@@ -49,35 +49,35 @@ def hash_map_new_with_capacity_fwd
}
/- [hashmap::HashMap::{0}::new] -/
-def hash_map_new_fwd (T : Type) : result (hash_map_t T) :=
+def hash_map_new_fwd (T : Type) : Result (hash_map_t T) :=
hash_map_new_with_capacity_fwd T (USize.ofNatCore 32 (by intlit))
(USize.ofNatCore 4 (by intlit)) (USize.ofNatCore 5 (by intlit))
/- [hashmap::HashMap::{0}::clear] -/
def hash_map_clear_loop_fwd_back
- (T : Type) (slots : vec (list_t T)) (i : USize) :
- (result (vec (list_t T)))
+ (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
+ if h: i < i0
then
do
- let i1 <- USize.checked_add i (USize.ofNatCore 1 (by intlit))
- let slots0 <- vec_index_mut_back (list_t T) slots i list_t.ListNil
+ let i1 ← USize.checked_add i (USize.ofNatCore 1 (by intlit))
+ let slots0 ← vec_index_mut_back (list_t T) slots i list_t.ListNil
hash_map_clear_loop_fwd_back T slots0 i1
- else result.ret slots
+ else Result.ret slots
termination_by hash_map_clear_loop_fwd_back slots i =>
hash_map_clear_loop_terminates T slots i
decreasing_by hash_map_clear_loop_decreases slots i
/- [hashmap::HashMap::{0}::clear] -/
def hash_map_clear_fwd_back
- (T : Type) (self : hash_map_t T) : result (hash_map_t T) :=
+ (T : Type) (self : hash_map_t T) : Result (hash_map_t T) :=
do
- let v <-
+ let v ←
hash_map_clear_loop_fwd_back T self.hash_map_slots
(USize.ofNatCore 0 (by intlit))
- result.ret
+ Result.ret
{
hash_map_num_entries := (USize.ofNatCore 0 (by intlit)),
hash_map_max_load_factor := self.hash_map_max_load_factor,
@@ -86,69 +86,69 @@ def hash_map_clear_fwd_back
}
/- [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
+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] -/
def hash_map_insert_in_list_loop_fwd
- (T : Type) (key : USize) (value : T) (ls : list_t T) : (result Bool) :=
- match ls with
+ (T : Type) (key : USize) (value : T) (ls : list_t T) : (Result Bool) :=
+ match h: ls with
| list_t.ListCons ckey cvalue tl =>
- if ckey = key
- then result.ret false
+ if h: ckey = key
+ then Result.ret false
else hash_map_insert_in_list_loop_fwd T key value tl
- | list_t.ListNil => result.ret true
+ | list_t.ListNil => Result.ret true
termination_by hash_map_insert_in_list_loop_fwd key value ls =>
hash_map_insert_in_list_loop_terminates T key value ls
decreasing_by hash_map_insert_in_list_loop_decreases key value ls
/- [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 :=
+ (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] -/
def hash_map_insert_in_list_loop_back
- (T : Type) (key : USize) (value : T) (ls : list_t T) : (result (list_t T)) :=
- match ls with
+ (T : Type) (key : USize) (value : T) (ls : list_t T) : (Result (list_t T)) :=
+ match h: ls with
| list_t.ListCons ckey cvalue tl =>
- if ckey = key
- then result.ret (list_t.ListCons ckey value tl)
+ if h: ckey = key
+ then Result.ret (list_t.ListCons ckey value tl)
else
do
- let tl0 <- hash_map_insert_in_list_loop_back T key value tl
- result.ret (list_t.ListCons ckey cvalue tl0)
+ let tl0 ← hash_map_insert_in_list_loop_back T key value tl
+ Result.ret (list_t.ListCons ckey cvalue tl0)
| list_t.ListNil =>
let l := list_t.ListNil
- result.ret (list_t.ListCons key value l)
+ Result.ret (list_t.ListCons key value l)
termination_by hash_map_insert_in_list_loop_back key value ls =>
hash_map_insert_in_list_loop_terminates T key value ls
decreasing_by hash_map_insert_in_list_loop_decreases key value ls
/- [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) :=
+ (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)
+ Result (hash_map_t T)
:=
do
- let hash <- hash_key_fwd key
+ let hash ← hash_key_fwd key
let i := vec_len (list_t T) self.hash_map_slots
- let hash_mod <- USize.checked_rem 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
+ let hash_mod ← USize.checked_rem 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 h: inserted
then
do
- let i0 <- USize.checked_add self.hash_map_num_entries
+ let i0 ← USize.checked_add self.hash_map_num_entries
(USize.ofNatCore 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
+ 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
{
hash_map_num_entries := i0,
hash_map_max_load_factor := self.hash_map_max_load_factor,
@@ -157,9 +157,9 @@ def hash_map_insert_no_resize_fwd_back
}
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
+ 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
{
hash_map_num_entries := self.hash_map_num_entries,
hash_map_max_load_factor := self.hash_map_max_load_factor,
@@ -168,76 +168,76 @@ def hash_map_insert_no_resize_fwd_back
}
/- [core::num::u32::{9}::MAX] -/
-def core_num_u32_max_body : result UInt32 :=
- result.ret (UInt32.ofNatCore 4294967295 (by intlit))
+def core_num_u32_max_body : Result UInt32 :=
+ Result.ret (UInt32.ofNatCore 4294967295 (by intlit))
def core_num_u32_max_c : UInt32 := eval_global core_num_u32_max_body (by simp)
/- [hashmap::HashMap::{0}::move_elements_from_list] -/
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))
+ (Result (hash_map_t T))
:=
- match ls with
+ match h: ls with
| list_t.ListCons k v tl =>
do
- let ntable0 <- hash_map_insert_no_resize_fwd_back T ntable k v
+ 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.ListNil => result.ret ntable
+ | list_t.ListNil => Result.ret ntable
termination_by hash_map_move_elements_from_list_loop_fwd_back ntable ls =>
hash_map_move_elements_from_list_loop_terminates T ntable ls
decreasing_by hash_map_move_elements_from_list_loop_decreases ntable ls
/- [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) :=
+ (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] -/
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))))
+ (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
+ if h: i < i0
then
do
- let l <- vec_index_mut_fwd (list_t T) slots i
+ let l ← vec_index_mut_fwd (list_t T) slots i
let ls := mem_replace_fwd (list_t T) l list_t.ListNil
- let ntable0 <- hash_map_move_elements_from_list_fwd_back T ntable ls
- let i1 <- USize.checked_add i (USize.ofNatCore 1 (by intlit))
+ let ntable0 ← hash_map_move_elements_from_list_fwd_back T ntable ls
+ let i1 ← USize.checked_add i (USize.ofNatCore 1 (by intlit))
let l0 := mem_replace_back (list_t T) l list_t.ListNil
- let slots0 <- vec_index_mut_back (list_t T) slots i l0
+ 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)
+ else Result.ret (ntable, slots)
termination_by hash_map_move_elements_loop_fwd_back ntable slots i =>
hash_map_move_elements_loop_terminates T ntable slots i
decreasing_by hash_map_move_elements_loop_decreases ntable slots i
/- [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)))
+ (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) :=
+ (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 max_usize ← scalar_cast USize core_num_u32_max_c
let capacity := vec_len (list_t T) self.hash_map_slots
- let n1 <- USize.checked_div max_usize (USize.ofNatCore 2 (by intlit))
+ let n1 ← USize.checked_div max_usize (USize.ofNatCore 2 (by intlit))
let (i, i0) := self.hash_map_max_load_factor
- let i1 <- USize.checked_div n1 i
- if capacity <= i1
+ let i1 ← USize.checked_div n1 i
+ if h: capacity <= i1
then
do
- let i2 <- USize.checked_mul capacity (USize.ofNatCore 2 (by intlit))
- let ntable <- hash_map_new_with_capacity_fwd T i2 i i0
- let (ntable0, _) <-
+ let i2 ← USize.checked_mul capacity (USize.ofNatCore 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.ofNatCore 0 (by intlit))
- result.ret
+ Result.ret
{
hash_map_num_entries := self.hash_map_num_entries,
hash_map_max_load_factor := (i, i0),
@@ -245,7 +245,7 @@ def hash_map_try_resize_fwd_back
hash_map_slots := ntable0.hash_map_slots
}
else
- result.ret
+ Result.ret
{
hash_map_num_entries := self.hash_map_num_entries,
hash_map_max_load_factor := (i, i0),
@@ -256,133 +256,133 @@ def hash_map_try_resize_fwd_back
/- [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)
+ 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
+ let self0 ← hash_map_insert_no_resize_fwd_back T self key value
+ let i ← hash_map_len_fwd T self0
+ if h: i > self0.hash_map_max_load
then hash_map_try_resize_fwd_back T self0
- else result.ret self0
+ else Result.ret self0
/- [hashmap::HashMap::{0}::contains_key_in_list] -/
def hash_map_contains_key_in_list_loop_fwd
- (T : Type) (key : USize) (ls : list_t T) : (result Bool) :=
- match ls with
+ (T : Type) (key : USize) (ls : list_t T) : (Result Bool) :=
+ match h: ls with
| list_t.ListCons ckey t tl =>
- if ckey = key
- then result.ret true
+ if h: ckey = key
+ then Result.ret true
else hash_map_contains_key_in_list_loop_fwd T key tl
- | list_t.ListNil => result.ret false
+ | list_t.ListNil => Result.ret false
termination_by hash_map_contains_key_in_list_loop_fwd key ls =>
hash_map_contains_key_in_list_loop_terminates T key ls
decreasing_by hash_map_contains_key_in_list_loop_decreases key ls
/- [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 :=
+ (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 :=
+ (T : Type) (self : hash_map_t T) (key : USize) : Result Bool :=
do
- let hash <- hash_key_fwd key
+ let hash ← hash_key_fwd key
let i := vec_len (list_t T) self.hash_map_slots
- let hash_mod <- USize.checked_rem hash i
- let l <- vec_index_fwd (list_t T) self.hash_map_slots hash_mod
+ let hash_mod ← USize.checked_rem 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] -/
def hash_map_get_in_list_loop_fwd
- (T : Type) (key : USize) (ls : list_t T) : (result T) :=
- match ls with
+ (T : Type) (key : USize) (ls : list_t T) : (Result T) :=
+ match h: ls with
| list_t.ListCons ckey cvalue tl =>
- if ckey = key
- then result.ret cvalue
+ if h: ckey = key
+ then Result.ret cvalue
else hash_map_get_in_list_loop_fwd T key tl
- | list_t.ListNil => result.fail error.panic
+ | list_t.ListNil => Result.fail Error.panic
termination_by hash_map_get_in_list_loop_fwd key ls =>
hash_map_get_in_list_loop_terminates T key ls
decreasing_by hash_map_get_in_list_loop_decreases key ls
/- [hashmap::HashMap::{0}::get_in_list] -/
def hash_map_get_in_list_fwd
- (T : Type) (key : USize) (ls : list_t T) : result T :=
+ (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 :=
+ (T : Type) (self : hash_map_t T) (key : USize) : Result T :=
do
- let hash <- hash_key_fwd key
+ let hash ← hash_key_fwd key
let i := vec_len (list_t T) self.hash_map_slots
- let hash_mod <- USize.checked_rem hash i
- let l <- vec_index_fwd (list_t T) self.hash_map_slots hash_mod
+ let hash_mod ← USize.checked_rem 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] -/
def hash_map_get_mut_in_list_loop_fwd
- (T : Type) (ls : list_t T) (key : USize) : (result T) :=
- match ls with
+ (T : Type) (ls : list_t T) (key : USize) : (Result T) :=
+ match h: ls with
| list_t.ListCons ckey cvalue tl =>
- if ckey = key
- then result.ret cvalue
+ if h: ckey = key
+ then Result.ret cvalue
else hash_map_get_mut_in_list_loop_fwd T tl key
- | list_t.ListNil => result.fail error.panic
+ | list_t.ListNil => Result.fail Error.panic
termination_by hash_map_get_mut_in_list_loop_fwd ls key =>
hash_map_get_mut_in_list_loop_terminates T ls key
decreasing_by hash_map_get_mut_in_list_loop_decreases ls key
/- [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 :=
+ (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] -/
def hash_map_get_mut_in_list_loop_back
- (T : Type) (ls : list_t T) (key : USize) (ret0 : T) : (result (list_t T)) :=
- match ls with
+ (T : Type) (ls : list_t T) (key : USize) (ret0 : T) : (Result (list_t T)) :=
+ match h: ls with
| list_t.ListCons ckey cvalue tl =>
- if ckey = key
- then result.ret (list_t.ListCons ckey ret0 tl)
+ if h: ckey = key
+ then Result.ret (list_t.ListCons ckey ret0 tl)
else
do
- let tl0 <- hash_map_get_mut_in_list_loop_back T tl key ret0
- result.ret (list_t.ListCons ckey cvalue tl0)
- | list_t.ListNil => result.fail error.panic
+ let tl0 ← hash_map_get_mut_in_list_loop_back T tl key ret0
+ Result.ret (list_t.ListCons ckey cvalue tl0)
+ | list_t.ListNil => Result.fail Error.panic
termination_by hash_map_get_mut_in_list_loop_back ls key ret0 =>
hash_map_get_mut_in_list_loop_terminates T ls key
decreasing_by hash_map_get_mut_in_list_loop_decreases ls key
/- [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) :=
+ (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 :=
+ (T : Type) (self : hash_map_t T) (key : USize) : Result T :=
do
- let hash <- hash_key_fwd key
+ let hash ← hash_key_fwd key
let i := vec_len (list_t T) self.hash_map_slots
- let hash_mod <- USize.checked_rem hash i
- let l <- vec_index_mut_fwd (list_t T) self.hash_map_slots hash_mod
+ let hash_mod ← USize.checked_rem 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)
+ Result (hash_map_t T)
:=
do
- let hash <- hash_key_fwd key
+ let hash ← hash_key_fwd key
let i := vec_len (list_t T) self.hash_map_slots
- let hash_mod <- USize.checked_rem 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
+ let hash_mod ← USize.checked_rem 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
{
hash_map_num_entries := self.hash_map_num_entries,
hash_map_max_load_factor := self.hash_map_max_load_factor,
@@ -392,85 +392,85 @@ def hash_map_get_mut_back
/- [hashmap::HashMap::{0}::remove_from_list] -/
def hash_map_remove_from_list_loop_fwd
- (T : Type) (key : USize) (ls : list_t T) : (result (Option T)) :=
- match ls with
+ (T : Type) (key : USize) (ls : list_t T) : (Result (Option T)) :=
+ match h: ls with
| list_t.ListCons ckey t tl =>
- if ckey = key
+ if h: ckey = key
then
let mv_ls :=
mem_replace_fwd (list_t T) (list_t.ListCons ckey t tl) list_t.ListNil
- match mv_ls with
- | list_t.ListCons i cvalue tl0 => result.ret (Option.some cvalue)
- | list_t.ListNil => result.fail error.panic
+ match h: mv_ls with
+ | list_t.ListCons i cvalue tl0 => Result.ret (Option.some cvalue)
+ | list_t.ListNil => Result.fail Error.panic
else hash_map_remove_from_list_loop_fwd T key tl
- | list_t.ListNil => result.ret Option.none
+ | list_t.ListNil => Result.ret Option.none
termination_by hash_map_remove_from_list_loop_fwd key ls =>
hash_map_remove_from_list_loop_terminates T key ls
decreasing_by hash_map_remove_from_list_loop_decreases key ls
/- [hashmap::HashMap::{0}::remove_from_list] -/
def hash_map_remove_from_list_fwd
- (T : Type) (key : USize) (ls : list_t T) : result (Option T) :=
+ (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] -/
def hash_map_remove_from_list_loop_back
- (T : Type) (key : USize) (ls : list_t T) : (result (list_t T)) :=
- match ls with
+ (T : Type) (key : USize) (ls : list_t T) : (Result (list_t T)) :=
+ match h: ls with
| list_t.ListCons ckey t tl =>
- if ckey = key
+ if h: ckey = key
then
let mv_ls :=
mem_replace_fwd (list_t T) (list_t.ListCons ckey t tl) list_t.ListNil
- match mv_ls with
- | list_t.ListCons i cvalue tl0 => result.ret tl0
- | list_t.ListNil => result.fail error.panic
+ match h: mv_ls with
+ | list_t.ListCons i cvalue tl0 => Result.ret tl0
+ | list_t.ListNil => Result.fail Error.panic
else
do
- let tl0 <- hash_map_remove_from_list_loop_back T key tl
- result.ret (list_t.ListCons ckey t tl0)
- | list_t.ListNil => result.ret list_t.ListNil
+ let tl0 ← hash_map_remove_from_list_loop_back T key tl
+ Result.ret (list_t.ListCons ckey t tl0)
+ | list_t.ListNil => Result.ret list_t.ListNil
termination_by hash_map_remove_from_list_loop_back key ls =>
hash_map_remove_from_list_loop_terminates T key ls
decreasing_by hash_map_remove_from_list_loop_decreases key ls
/- [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) :=
+ (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) :=
+ (T : Type) (self : hash_map_t T) (key : USize) : Result (Option T) :=
do
- let hash <- hash_key_fwd key
+ let hash ← hash_key_fwd key
let i := vec_len (list_t T) self.hash_map_slots
- let hash_mod <- USize.checked_rem 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 x with
- | Option.none => result.ret Option.none
+ let hash_mod ← USize.checked_rem 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 _ <- USize.checked_sub self.hash_map_num_entries
+ let _ ← USize.checked_sub self.hash_map_num_entries
(USize.ofNatCore 1 (by intlit))
- result.ret (Option.some x0)
+ 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) :=
+ (T : Type) (self : hash_map_t T) (key : USize) : Result (hash_map_t T) :=
do
- let hash <- hash_key_fwd key
+ let hash ← hash_key_fwd key
let i := vec_len (list_t T) self.hash_map_slots
- let hash_mod <- USize.checked_rem 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 x with
+ let hash_mod ← USize.checked_rem 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
+ 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
{
hash_map_num_entries := self.hash_map_num_entries,
hash_map_max_load_factor := self.hash_map_max_load_factor,
@@ -479,11 +479,11 @@ def hash_map_remove_back
}
| Option.some x0 =>
do
- let i0 <- USize.checked_sub self.hash_map_num_entries
+ let i0 ← USize.checked_sub self.hash_map_num_entries
(USize.ofNatCore 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
+ 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
{
hash_map_num_entries := i0,
hash_map_max_load_factor := self.hash_map_max_load_factor,
@@ -492,67 +492,67 @@ def hash_map_remove_back
}
/- [hashmap::test1] -/
-def test1_fwd : result Unit :=
+def test1_fwd : Result Unit :=
do
- let hm <- hash_map_new_fwd UInt64
- let hm0 <-
+ let hm ← hash_map_new_fwd UInt64
+ let hm0 ←
hash_map_insert_fwd_back UInt64 hm (USize.ofNatCore 0 (by intlit))
(UInt64.ofNatCore 42 (by intlit))
- let hm1 <-
+ let hm1 ←
hash_map_insert_fwd_back UInt64 hm0 (USize.ofNatCore 128 (by intlit))
(UInt64.ofNatCore 18 (by intlit))
- let hm2 <-
+ let hm2 ←
hash_map_insert_fwd_back UInt64 hm1 (USize.ofNatCore 1024 (by intlit))
(UInt64.ofNatCore 138 (by intlit))
- let hm3 <-
+ let hm3 ←
hash_map_insert_fwd_back UInt64 hm2 (USize.ofNatCore 1056 (by intlit))
(UInt64.ofNatCore 256 (by intlit))
- let i <- hash_map_get_fwd UInt64 hm3 (USize.ofNatCore 128 (by intlit))
- if not (i = (UInt64.ofNatCore 18 (by intlit)))
- then result.fail error.panic
+ let i ← hash_map_get_fwd UInt64 hm3 (USize.ofNatCore 128 (by intlit))
+ if h: not (i = (UInt64.ofNatCore 18 (by intlit)))
+ then Result.fail Error.panic
else
do
- let hm4 <-
+ let hm4 ←
hash_map_get_mut_back UInt64 hm3 (USize.ofNatCore 1024 (by intlit))
(UInt64.ofNatCore 56 (by intlit))
- let i0 <-
+ let i0 ←
hash_map_get_fwd UInt64 hm4 (USize.ofNatCore 1024 (by intlit))
- if not (i0 = (UInt64.ofNatCore 56 (by intlit)))
- then result.fail error.panic
+ if h: not (i0 = (UInt64.ofNatCore 56 (by intlit)))
+ then Result.fail Error.panic
else
do
- let x <-
+ let x ←
hash_map_remove_fwd UInt64 hm4 (USize.ofNatCore 1024 (by intlit))
- match x with
- | Option.none => result.fail error.panic
+ match h: x with
+ | Option.none => Result.fail Error.panic
| Option.some x0 =>
- if not (x0 = (UInt64.ofNatCore 56 (by intlit)))
- then result.fail error.panic
+ if h: not (x0 = (UInt64.ofNatCore 56 (by intlit)))
+ then Result.fail Error.panic
else
do
- let hm5 <-
+ let hm5 ←
hash_map_remove_back UInt64 hm4
(USize.ofNatCore 1024 (by intlit))
- let i1 <-
+ let i1 ←
hash_map_get_fwd UInt64 hm5 (USize.ofNatCore 0 (by intlit))
- if not (i1 = (UInt64.ofNatCore 42 (by intlit)))
- then result.fail error.panic
+ if h: not (i1 = (UInt64.ofNatCore 42 (by intlit)))
+ then Result.fail Error.panic
else
do
- let i2 <-
+ let i2 ←
hash_map_get_fwd UInt64 hm5
(USize.ofNatCore 128 (by intlit))
- if not (i2 = (UInt64.ofNatCore 18 (by intlit)))
- then result.fail error.panic
+ if h: not (i2 = (UInt64.ofNatCore 18 (by intlit)))
+ then Result.fail Error.panic
else
do
- let i3 <-
+ let i3 ←
hash_map_get_fwd UInt64 hm5
(USize.ofNatCore 1056 (by intlit))
- if not (i3 = (UInt64.ofNatCore 256 (by intlit)))
- then result.fail error.panic
- else result.ret ()
+ if h: not (i3 = (UInt64.ofNatCore 256 (by intlit)))
+ then Result.fail Error.panic
+ else Result.ret ()
/- Unit test for [hashmap::test1] -/
-#assert (test1_fwd = .ret ())
+#assert (test1_fwd == .ret ())