diff options
author | Son Ho | 2023-07-05 15:17:58 +0200 |
---|---|---|
committer | Son Ho | 2023-07-05 15:17:58 +0200 |
commit | 5ca36bfc50083a01af2b7ae5f75993a520757ef5 (patch) | |
tree | 11660b73a27ad2e0807a18ac9286a1e87c560050 | |
parent | c07721dedb2cfe4c726c42606e623395cdfe5b80 (diff) |
Simplify the names used in Primitives.lean
-rw-r--r-- | backends/lean/Base/Primitives.lean | 37 | ||||
-rw-r--r-- | compiler/Extract.ml | 49 | ||||
-rw-r--r-- | tests/lean/BetreeMain/Funs.lean | 6 | ||||
-rw-r--r-- | tests/lean/External/Funs.lean | 4 | ||||
-rw-r--r-- | tests/lean/Hashmap/Funs.lean | 62 | ||||
-rw-r--r-- | tests/lean/HashmapMain/Funs.lean | 64 | ||||
-rw-r--r-- | tests/lean/Loops/Funs.lean | 12 | ||||
-rw-r--r-- | tests/lean/NoNestedBorrows.lean | 4 | ||||
-rw-r--r-- | tests/lean/lake-manifest.json | 2 |
9 files changed, 130 insertions, 110 deletions
diff --git a/backends/lean/Base/Primitives.lean b/backends/lean/Base/Primitives.lean index 117f76a2..808c1461 100644 --- a/backends/lean/Base/Primitives.lean +++ b/backends/lean/Base/Primitives.lean @@ -564,15 +564,17 @@ macro_rules def Vec (α : Type u) := { l : List α // List.length l ≤ Usize.max } -def vec_new (α : Type u): Vec α := ⟨ [], by apply Scalar.cMax_suffices .Usize; simp ⟩ +def Vec.new (α : Type u): Vec α := ⟨ [], by apply Scalar.cMax_suffices .Usize; simp ⟩ -def vec_len (α : Type u) (v : Vec α) : Usize := +def Vec.len (α : Type u) (v : Vec α) : Usize := let ⟨ v, l ⟩ := v Usize.ofIntCore (List.length v) (by simp [Scalar.min]) l -def vec_push_fwd (α : Type u) (_ : Vec α) (_ : α) : Unit := () +-- This shouldn't be used +def Vec.push_fwd (α : Type u) (_ : Vec α) (_ : α) : Unit := () -def vec_push_back (α : Type u) (v : Vec α) (x : α) : Result (Vec α) +-- This is actually the backward function +def Vec.push (α : Type u) (v : Vec α) (x : α) : Result (Vec α) := let nlen := List.length v.val + 1 if h : nlen ≤ U32.max || nlen ≤ Usize.max then @@ -588,13 +590,15 @@ def vec_push_back (α : Type u) (v : Vec α) (x : α) : Result (Vec α) else fail maximumSizeExceeded -def vec_insert_fwd (α : Type u) (v: Vec α) (i: Usize) (_: α): Result Unit := +-- This shouldn't be used +def Vec.insert_fwd (α : Type u) (v: Vec α) (i: Usize) (_: α): Result Unit := if i.val < List.length v.val then .ret () else .fail arrayOutOfBounds -def vec_insert_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α) := +-- This is actually the backward function +def Vec.insert (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α) := if i.val < List.length v.val then -- TODO: maybe we should redefine a list library which uses integers -- (instead of natural numbers) @@ -607,7 +611,7 @@ def vec_insert_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α else .fail arrayOutOfBounds -def vec_index_to_fin {α : Type u} {v: Vec α} {i: Usize} (h : i.val < List.length v.val) : +def Vec.index_to_fin {α : Type u} {v: Vec α} {i: Usize} (h : i.val < List.length v.val) : Fin (List.length v.val) := let j := i.val.toNat let h: j < List.length v.val := by @@ -616,29 +620,30 @@ def vec_index_to_fin {α : Type u} {v: Vec α} {i: Usize} (h : i.val < List.leng assumption ⟨j, h⟩ -def vec_index_fwd (α : Type u) (v: Vec α) (i: Usize): Result α := +def Vec.index (α : Type u) (v: Vec α) (i: Usize): Result α := if h: i.val < List.length v.val then - let i := vec_index_to_fin h + let i := Vec.index_to_fin h .ret (List.get v.val i) else .fail arrayOutOfBounds -def vec_index_back (α : Type u) (v: Vec α) (i: Usize) (_: α): Result Unit := +-- This shouldn't be used +def Vec.index_back (α : Type u) (v: Vec α) (i: Usize) (_: α): Result Unit := if i.val < List.length v.val then .ret () else .fail arrayOutOfBounds -def vec_index_mut_fwd (α : Type u) (v: Vec α) (i: Usize): Result α := +def Vec.index_mut (α : Type u) (v: Vec α) (i: Usize): Result α := if h: i.val < List.length v.val then - let i := vec_index_to_fin h + let i := Vec.index_to_fin h .ret (List.get v.val i) else .fail arrayOutOfBounds -def vec_index_mut_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α) := +def Vec.index_mut_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec α) := if h: i.val < List.length v.val then - let i := vec_index_to_fin h + let i := Vec.index_to_fin h .ret ⟨ List.set v.val i x, by have h: List.length v.val ≤ Usize.max := v.property simp [*] at * @@ -651,8 +656,8 @@ def vec_index_mut_back (α : Type u) (v: Vec α) (i: Usize) (x: α): Result (Vec -- MISC -- ---------- -@[simp] def mem_replace_fwd (a : Type) (x : a) (_ : a) : a := x -@[simp] def mem_replace_back (a : Type) (_ : a) (y : a) : a := y +@[simp] def mem.replace_fwd (a : Type) (x : a) (_ : a) : a := x +@[simp] def mem.replace_back (a : Type) (_ : a) (y : a) : a := y /-- Aeneas-translated function -- useful to reduce non-recursive definitions. Use with `simp [ aeneas ]` -/ diff --git a/compiler/Extract.ml b/compiler/Extract.ml index 821bf4f7..180ca7ca 100644 --- a/compiler/Extract.ml +++ b/compiler/Extract.ml @@ -300,23 +300,40 @@ let assumed_variants () : (assumed_ty * VariantId.id * string) list = (Option, option_none_id, "NONE"); ] -let assumed_llbc_functions : +let assumed_llbc_functions () : (A.assumed_fun_id * T.RegionGroupId.id option * string) list = let rg0 = Some T.RegionGroupId.zero in - [ - (Replace, None, "mem_replace_fwd"); - (Replace, rg0, "mem_replace_back"); - (VecNew, None, "vec_new"); - (VecPush, None, "vec_push_fwd") (* Shouldn't be used *); - (VecPush, rg0, "vec_push_back"); - (VecInsert, None, "vec_insert_fwd") (* Shouldn't be used *); - (VecInsert, rg0, "vec_insert_back"); - (VecLen, None, "vec_len"); - (VecIndex, None, "vec_index_fwd"); - (VecIndex, rg0, "vec_index_back") (* shouldn't be used *); - (VecIndexMut, None, "vec_index_mut_fwd"); - (VecIndexMut, rg0, "vec_index_mut_back"); - ] + match !backend with + | FStar | Coq | HOL4 -> + [ + (Replace, None, "mem_replace_fwd"); + (Replace, rg0, "mem_replace_back"); + (VecNew, None, "vec_new"); + (VecPush, None, "vec_push_fwd") (* Shouldn't be used *); + (VecPush, rg0, "vec_push_back"); + (VecInsert, None, "vec_insert_fwd") (* Shouldn't be used *); + (VecInsert, rg0, "vec_insert_back"); + (VecLen, None, "vec_len"); + (VecIndex, None, "vec_index_fwd"); + (VecIndex, rg0, "vec_index_back") (* shouldn't be used *); + (VecIndexMut, None, "vec_index_mut_fwd"); + (VecIndexMut, rg0, "vec_index_mut_back"); + ] + | Lean -> + [ + (Replace, None, "mem.replace_fwd"); + (Replace, rg0, "mem.replace_back"); + (VecNew, None, "Vec.new"); + (VecPush, None, "Vec.push_fwd") (* Shouldn't be used *); + (VecPush, rg0, "Vec.push"); + (VecInsert, None, "Vec.insert_fwd") (* Shouldn't be used *); + (VecInsert, rg0, "Vec.insert"); + (VecLen, None, "Vec.len"); + (VecIndex, None, "Vec.index"); + (VecIndex, rg0, "Vec.index_back") (* shouldn't be used *); + (VecIndexMut, None, "Vec.index_mut"); + (VecIndexMut, rg0, "Vec.index_mut_back"); + ] let assumed_pure_functions () : (pure_assumed_fun_id * string) list = match !backend with @@ -344,7 +361,7 @@ let names_map_init () : names_map_init = assumed_adts = assumed_adts (); assumed_structs; assumed_variants = assumed_variants (); - assumed_llbc_functions; + assumed_llbc_functions = assumed_llbc_functions (); assumed_pure_functions = assumed_pure_functions (); } diff --git a/tests/lean/BetreeMain/Funs.lean b/tests/lean/BetreeMain/Funs.lean index 3a678c71..074fff5e 100644 --- a/tests/lean/BetreeMain/Funs.lean +++ b/tests/lean/BetreeMain/Funs.lean @@ -124,13 +124,13 @@ divergent def betree.List.split_at_fwd /- [betree_main::betree::List::{1}::push_front] -/ def betree.List.push_front_fwd_back (T : Type) (self : betree.List T) (x : T) : Result (betree.List T) := - let tl := mem_replace_fwd (betree.List T) self betree.List.Nil + let tl := mem.replace_fwd (betree.List T) self betree.List.Nil let l := tl Result.ret (betree.List.Cons x l) /- [betree_main::betree::List::{1}::pop_front] -/ def betree.List.pop_front_fwd (T : Type) (self : betree.List T) : Result T := - let ls := mem_replace_fwd (betree.List T) self betree.List.Nil + let ls := mem.replace_fwd (betree.List T) self betree.List.Nil match ls with | betree.List.Cons x tl => Result.ret x | betree.List.Nil => Result.fail Error.panic @@ -138,7 +138,7 @@ def betree.List.pop_front_fwd (T : Type) (self : betree.List T) : Result T := /- [betree_main::betree::List::{1}::pop_front] -/ def betree.List.pop_front_back (T : Type) (self : betree.List T) : Result (betree.List T) := - let ls := mem_replace_fwd (betree.List T) self betree.List.Nil + let ls := mem.replace_fwd (betree.List T) self betree.List.Nil match ls with | betree.List.Cons x tl => Result.ret tl | betree.List.Nil => Result.fail Error.panic diff --git a/tests/lean/External/Funs.lean b/tests/lean/External/Funs.lean index 10efc3db..3fc21d91 100644 --- a/tests/lean/External/Funs.lean +++ b/tests/lean/External/Funs.lean @@ -36,8 +36,8 @@ def test_new_non_zero_u32_fwd /- [external::test_vec] -/ def test_vec_fwd : Result Unit := do - let v := vec_new U32 - let _ ← vec_push_back U32 v (U32.ofInt 0 (by intlit)) + let v := Vec.new U32 + let _ ← Vec.push U32 v (U32.ofInt 0 (by intlit)) Result.ret () /- Unit test for [external::test_vec] -/ diff --git a/tests/lean/Hashmap/Funs.lean b/tests/lean/Hashmap/Funs.lean index 4f16688b..48038bfb 100644 --- a/tests/lean/Hashmap/Funs.lean +++ b/tests/lean/Hashmap/Funs.lean @@ -15,7 +15,7 @@ divergent def HashMap.allocate_slots_loop_fwd if n > (Usize.ofInt 0 (by intlit)) then do - let slots0 ← vec_push_back (List T) slots List.Nil + let slots0 ← Vec.push (List T) slots List.Nil let n0 ← n - (Usize.ofInt 1 (by intlit)) HashMap.allocate_slots_loop_fwd T slots0 n0 else Result.ret slots @@ -32,7 +32,7 @@ def HashMap.new_with_capacity_fwd Result (HashMap T) := do - let v := vec_new (List T) + let v := Vec.new (List T) let slots ← HashMap.allocate_slots_fwd T v capacity let i ← capacity * max_load_dividend let i0 ← i / max_load_divisor @@ -52,12 +52,12 @@ def HashMap.new_fwd (T : Type) : Result (HashMap T) := /- [hashmap::HashMap::{0}::clear] -/ divergent def HashMap.clear_loop_fwd_back (T : Type) (slots : Vec (List T)) (i : Usize) : Result (Vec (List T)) := - let i0 := vec_len (List T) slots + let i0 := Vec.len (List T) slots if i < i0 then do let i1 ← i + (Usize.ofInt 1 (by intlit)) - let slots0 ← vec_index_mut_back (List T) slots i List.Nil + let slots0 ← Vec.index_mut_back (List T) slots i List.Nil HashMap.clear_loop_fwd_back T slots0 i1 else Result.ret slots @@ -121,22 +121,22 @@ def HashMap.insert_no_resize_fwd_back := do let hash ← hash_key_fwd key - let i := vec_len (List T) self.hash_map_slots + let i := Vec.len (List T) self.hash_map_slots let hash_mod ← hash % i - let l ← vec_index_mut_fwd (List T) self.hash_map_slots hash_mod + let l ← Vec.index_mut (List T) self.hash_map_slots hash_mod let inserted ← HashMap.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 ← HashMap.insert_in_list_back T key value l - let v ← vec_index_mut_back (List T) self.hash_map_slots hash_mod l0 + let v ← Vec.index_mut_back (List 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 ← HashMap.insert_in_list_back T key value l - let v ← vec_index_mut_back (List T) self.hash_map_slots hash_mod l0 + let v ← Vec.index_mut_back (List T) self.hash_map_slots hash_mod l0 Result.ret { self with hash_map_slots := v } /- [core::num::u32::{9}::MAX] -/ @@ -164,16 +164,16 @@ divergent def HashMap.move_elements_loop_fwd_back (T : Type) (ntable : HashMap T) (slots : Vec (List T)) (i : Usize) : Result ((HashMap T) × (Vec (List T))) := - let i0 := vec_len (List T) slots + let i0 := Vec.len (List T) slots if i < i0 then do - let l ← vec_index_mut_fwd (List T) slots i - let ls := mem_replace_fwd (List T) l List.Nil + let l ← Vec.index_mut (List T) slots i + let ls := mem.replace_fwd (List T) l List.Nil let ntable0 ← HashMap.move_elements_from_list_fwd_back T ntable ls let i1 ← i + (Usize.ofInt 1 (by intlit)) - let l0 := mem_replace_back (List T) l List.Nil - let slots0 ← vec_index_mut_back (List T) slots i l0 + let l0 := mem.replace_back (List T) l List.Nil + let slots0 ← Vec.index_mut_back (List T) slots i l0 HashMap.move_elements_loop_fwd_back T ntable0 slots0 i1 else Result.ret (ntable, slots) @@ -189,7 +189,7 @@ def HashMap.try_resize_fwd_back (T : Type) (self : HashMap T) : Result (HashMap T) := do let max_usize ← Scalar.cast .Usize core_num_u32_max_c - let capacity := vec_len (List T) self.hash_map_slots + let capacity := Vec.len (List 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 @@ -242,9 +242,9 @@ def HashMap.contains_key_fwd (T : Type) (self : HashMap T) (key : Usize) : Result Bool := do let hash ← hash_key_fwd key - let i := vec_len (List T) self.hash_map_slots + let i := Vec.len (List T) self.hash_map_slots let hash_mod ← hash % i - let l ← vec_index_fwd (List T) self.hash_map_slots hash_mod + let l ← Vec.index (List T) self.hash_map_slots hash_mod HashMap.contains_key_in_list_fwd T key l /- [hashmap::HashMap::{0}::get_in_list] -/ @@ -266,9 +266,9 @@ def HashMap.get_in_list_fwd def HashMap.get_fwd (T : Type) (self : HashMap T) (key : Usize) : Result T := do let hash ← hash_key_fwd key - let i := vec_len (List T) self.hash_map_slots + let i := Vec.len (List T) self.hash_map_slots let hash_mod ← hash % i - let l ← vec_index_fwd (List T) self.hash_map_slots hash_mod + let l ← Vec.index (List T) self.hash_map_slots hash_mod HashMap.get_in_list_fwd T key l /- [hashmap::HashMap::{0}::get_mut_in_list] -/ @@ -309,9 +309,9 @@ def HashMap.get_mut_fwd (T : Type) (self : HashMap T) (key : Usize) : Result T := do let hash ← hash_key_fwd key - let i := vec_len (List T) self.hash_map_slots + let i := Vec.len (List T) self.hash_map_slots let hash_mod ← hash % i - let l ← vec_index_mut_fwd (List T) self.hash_map_slots hash_mod + let l ← Vec.index_mut (List T) self.hash_map_slots hash_mod HashMap.get_mut_in_list_fwd T l key /- [hashmap::HashMap::{0}::get_mut] -/ @@ -321,11 +321,11 @@ def HashMap.get_mut_back := do let hash ← hash_key_fwd key - let i := vec_len (List T) self.hash_map_slots + let i := Vec.len (List T) self.hash_map_slots let hash_mod ← hash % i - let l ← vec_index_mut_fwd (List T) self.hash_map_slots hash_mod + let l ← Vec.index_mut (List T) self.hash_map_slots hash_mod let l0 ← HashMap.get_mut_in_list_back T l key ret0 - let v ← vec_index_mut_back (List T) self.hash_map_slots hash_mod l0 + let v ← Vec.index_mut_back (List T) self.hash_map_slots hash_mod l0 Result.ret { self with hash_map_slots := v } /- [hashmap::HashMap::{0}::remove_from_list] -/ @@ -335,7 +335,7 @@ divergent def HashMap.remove_from_list_loop_fwd | List.Cons ckey t tl => if ckey = key then - let mv_ls := mem_replace_fwd (List T) (List.Cons ckey t tl) List.Nil + let mv_ls := mem.replace_fwd (List T) (List.Cons ckey t tl) List.Nil match mv_ls with | List.Cons i cvalue tl0 => Result.ret (Option.some cvalue) | List.Nil => Result.fail Error.panic @@ -354,7 +354,7 @@ divergent def HashMap.remove_from_list_loop_back | List.Cons ckey t tl => if ckey = key then - let mv_ls := mem_replace_fwd (List T) (List.Cons ckey t tl) List.Nil + let mv_ls := mem.replace_fwd (List T) (List.Cons ckey t tl) List.Nil match mv_ls with | List.Cons i cvalue tl0 => Result.ret tl0 | List.Nil => Result.fail Error.panic @@ -374,9 +374,9 @@ def HashMap.remove_fwd (T : Type) (self : HashMap T) (key : Usize) : Result (Option T) := do let hash ← hash_key_fwd key - let i := vec_len (List T) self.hash_map_slots + let i := Vec.len (List T) self.hash_map_slots let hash_mod ← hash % i - let l ← vec_index_mut_fwd (List T) self.hash_map_slots hash_mod + let l ← Vec.index_mut (List T) self.hash_map_slots hash_mod let x ← HashMap.remove_from_list_fwd T key l match x with | Option.none => Result.ret Option.none @@ -390,21 +390,21 @@ def HashMap.remove_back (T : Type) (self : HashMap T) (key : Usize) : Result (HashMap T) := do let hash ← hash_key_fwd key - let i := vec_len (List T) self.hash_map_slots + let i := Vec.len (List T) self.hash_map_slots let hash_mod ← hash % i - let l ← vec_index_mut_fwd (List T) self.hash_map_slots hash_mod + let l ← Vec.index_mut (List T) self.hash_map_slots hash_mod let x ← HashMap.remove_from_list_fwd T key l match x with | Option.none => do let l0 ← HashMap.remove_from_list_back T key l - let v ← vec_index_mut_back (List T) self.hash_map_slots hash_mod l0 + let v ← Vec.index_mut_back (List 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 ← HashMap.remove_from_list_back T key l - let v ← vec_index_mut_back (List T) self.hash_map_slots hash_mod l0 + let v ← Vec.index_mut_back (List T) self.hash_map_slots hash_mod l0 Result.ret { self with hash_map_num_entries := i0, hash_map_slots := v } diff --git a/tests/lean/HashmapMain/Funs.lean b/tests/lean/HashmapMain/Funs.lean index fd556878..1a741a2d 100644 --- a/tests/lean/HashmapMain/Funs.lean +++ b/tests/lean/HashmapMain/Funs.lean @@ -18,7 +18,7 @@ divergent def hashmap.HashMap.allocate_slots_loop_fwd if n > (Usize.ofInt 0 (by intlit)) then do - let slots0 ← vec_push_back (hashmap.List T) slots hashmap.List.Nil + let slots0 ← Vec.push (hashmap.List T) slots hashmap.List.Nil let n0 ← n - (Usize.ofInt 1 (by intlit)) hashmap.HashMap.allocate_slots_loop_fwd T slots0 n0 else Result.ret slots @@ -37,7 +37,7 @@ def hashmap.HashMap.new_with_capacity_fwd Result (hashmap.HashMap T) := do - let v := vec_new (hashmap.List T) + let v := Vec.new (hashmap.List T) let slots ← hashmap.HashMap.allocate_slots_fwd T v capacity let i ← capacity * max_load_dividend let i0 ← i / max_load_divisor @@ -60,13 +60,13 @@ divergent def hashmap.HashMap.clear_loop_fwd_back (T : Type) (slots : Vec (hashmap.List T)) (i : Usize) : Result (Vec (hashmap.List T)) := - let i0 := vec_len (hashmap.List T) slots + let i0 := Vec.len (hashmap.List T) slots if i < i0 then do let i1 ← i + (Usize.ofInt 1 (by intlit)) let slots0 ← - vec_index_mut_back (hashmap.List T) slots i hashmap.List.Nil + Vec.index_mut_back (hashmap.List T) slots i hashmap.List.Nil hashmap.HashMap.clear_loop_fwd_back T slots0 i1 else Result.ret slots @@ -136,10 +136,10 @@ def hashmap.HashMap.insert_no_resize_fwd_back := do let hash ← hashmap.hash_key_fwd key - let i := vec_len (hashmap.List T) self.hashmap_hash_map_slots + let i := Vec.len (hashmap.List T) self.hashmap_hash_map_slots let hash_mod ← hash % i let l ← - vec_index_mut_fwd (hashmap.List T) self.hashmap_hash_map_slots hash_mod + Vec.index_mut (hashmap.List T) self.hashmap_hash_map_slots hash_mod let inserted ← hashmap.HashMap.insert_in_list_fwd T key value l if inserted then @@ -148,7 +148,7 @@ def hashmap.HashMap.insert_no_resize_fwd_back (Usize.ofInt 1 (by intlit)) let l0 ← hashmap.HashMap.insert_in_list_back T key value l let v ← - vec_index_mut_back (hashmap.List T) self.hashmap_hash_map_slots + Vec.index_mut_back (hashmap.List T) self.hashmap_hash_map_slots hash_mod l0 Result.ret { @@ -160,7 +160,7 @@ def hashmap.HashMap.insert_no_resize_fwd_back do let l0 ← hashmap.HashMap.insert_in_list_back T key value l let v ← - vec_index_mut_back (hashmap.List T) self.hashmap_hash_map_slots + Vec.index_mut_back (hashmap.List T) self.hashmap_hash_map_slots hash_mod l0 Result.ret { self with hashmap_hash_map_slots := v } @@ -194,17 +194,17 @@ divergent def hashmap.HashMap.move_elements_loop_fwd_back (i : Usize) : Result ((hashmap.HashMap T) × (Vec (hashmap.List T))) := - let i0 := vec_len (hashmap.List T) slots + let i0 := Vec.len (hashmap.List T) slots if i < i0 then do - let l ← vec_index_mut_fwd (hashmap.List T) slots i - let ls := mem_replace_fwd (hashmap.List T) l hashmap.List.Nil + let l ← Vec.index_mut (hashmap.List T) slots i + let ls := mem.replace_fwd (hashmap.List T) l hashmap.List.Nil let ntable0 ← hashmap.HashMap.move_elements_from_list_fwd_back T ntable ls let i1 ← i + (Usize.ofInt 1 (by intlit)) - let l0 := mem_replace_back (hashmap.List T) l hashmap.List.Nil - let slots0 ← vec_index_mut_back (hashmap.List T) slots i l0 + let l0 := mem.replace_back (hashmap.List T) l hashmap.List.Nil + let slots0 ← Vec.index_mut_back (hashmap.List T) slots i l0 hashmap.HashMap.move_elements_loop_fwd_back T ntable0 slots0 i1 else Result.ret (ntable, slots) @@ -221,7 +221,7 @@ def hashmap.HashMap.try_resize_fwd_back (T : Type) (self : hashmap.HashMap T) : Result (hashmap.HashMap T) := do let max_usize ← Scalar.cast .Usize core_num_u32_max_c - let capacity := vec_len (hashmap.List T) self.hashmap_hash_map_slots + let capacity := Vec.len (hashmap.List T) self.hashmap_hash_map_slots let n1 ← max_usize / (Usize.ofInt 2 (by intlit)) let (i, i0) := self.hashmap_hash_map_max_load_factor let i1 ← n1 / i @@ -274,10 +274,9 @@ def hashmap.HashMap.contains_key_fwd (T : Type) (self : hashmap.HashMap T) (key : Usize) : Result Bool := do let hash ← hashmap.hash_key_fwd key - let i := vec_len (hashmap.List T) self.hashmap_hash_map_slots + let i := Vec.len (hashmap.List T) self.hashmap_hash_map_slots let hash_mod ← hash % i - let l ← - vec_index_fwd (hashmap.List T) self.hashmap_hash_map_slots hash_mod + let l ← Vec.index (hashmap.List T) self.hashmap_hash_map_slots hash_mod hashmap.HashMap.contains_key_in_list_fwd T key l /- [hashmap_main::hashmap::HashMap::{0}::get_in_list] -/ @@ -300,10 +299,9 @@ def hashmap.HashMap.get_fwd (T : Type) (self : hashmap.HashMap T) (key : Usize) : Result T := do let hash ← hashmap.hash_key_fwd key - let i := vec_len (hashmap.List T) self.hashmap_hash_map_slots + let i := Vec.len (hashmap.List T) self.hashmap_hash_map_slots let hash_mod ← hash % i - let l ← - vec_index_fwd (hashmap.List T) self.hashmap_hash_map_slots hash_mod + let l ← Vec.index (hashmap.List T) self.hashmap_hash_map_slots hash_mod hashmap.HashMap.get_in_list_fwd T key l /- [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list] -/ @@ -348,10 +346,10 @@ def hashmap.HashMap.get_mut_fwd (T : Type) (self : hashmap.HashMap T) (key : Usize) : Result T := do let hash ← hashmap.hash_key_fwd key - let i := vec_len (hashmap.List T) self.hashmap_hash_map_slots + let i := Vec.len (hashmap.List T) self.hashmap_hash_map_slots let hash_mod ← hash % i let l ← - vec_index_mut_fwd (hashmap.List T) self.hashmap_hash_map_slots hash_mod + Vec.index_mut (hashmap.List T) self.hashmap_hash_map_slots hash_mod hashmap.HashMap.get_mut_in_list_fwd T l key /- [hashmap_main::hashmap::HashMap::{0}::get_mut] -/ @@ -361,13 +359,13 @@ def hashmap.HashMap.get_mut_back := do let hash ← hashmap.hash_key_fwd key - let i := vec_len (hashmap.List T) self.hashmap_hash_map_slots + let i := Vec.len (hashmap.List T) self.hashmap_hash_map_slots let hash_mod ← hash % i let l ← - vec_index_mut_fwd (hashmap.List T) self.hashmap_hash_map_slots hash_mod + Vec.index_mut (hashmap.List T) self.hashmap_hash_map_slots hash_mod let l0 ← hashmap.HashMap.get_mut_in_list_back T l key ret0 let v ← - vec_index_mut_back (hashmap.List T) self.hashmap_hash_map_slots hash_mod + Vec.index_mut_back (hashmap.List T) self.hashmap_hash_map_slots hash_mod l0 Result.ret { self with hashmap_hash_map_slots := v } @@ -379,7 +377,7 @@ divergent def hashmap.HashMap.remove_from_list_loop_fwd if ckey = key then let mv_ls := - mem_replace_fwd (hashmap.List T) (hashmap.List.Cons ckey t tl) + mem.replace_fwd (hashmap.List T) (hashmap.List.Cons ckey t tl) hashmap.List.Nil match mv_ls with | hashmap.List.Cons i cvalue tl0 => Result.ret (Option.some cvalue) @@ -400,7 +398,7 @@ divergent def hashmap.HashMap.remove_from_list_loop_back if ckey = key then let mv_ls := - mem_replace_fwd (hashmap.List T) (hashmap.List.Cons ckey t tl) + mem.replace_fwd (hashmap.List T) (hashmap.List.Cons ckey t tl) hashmap.List.Nil match mv_ls with | hashmap.List.Cons i cvalue tl0 => Result.ret tl0 @@ -421,10 +419,10 @@ def hashmap.HashMap.remove_fwd (T : Type) (self : hashmap.HashMap T) (key : Usize) : Result (Option T) := do let hash ← hashmap.hash_key_fwd key - let i := vec_len (hashmap.List T) self.hashmap_hash_map_slots + let i := Vec.len (hashmap.List T) self.hashmap_hash_map_slots let hash_mod ← hash % i let l ← - vec_index_mut_fwd (hashmap.List T) self.hashmap_hash_map_slots hash_mod + Vec.index_mut (hashmap.List T) self.hashmap_hash_map_slots hash_mod let x ← hashmap.HashMap.remove_from_list_fwd T key l match x with | Option.none => Result.ret Option.none @@ -441,17 +439,17 @@ def hashmap.HashMap.remove_back := do let hash ← hashmap.hash_key_fwd key - let i := vec_len (hashmap.List T) self.hashmap_hash_map_slots + let i := Vec.len (hashmap.List T) self.hashmap_hash_map_slots let hash_mod ← hash % i let l ← - vec_index_mut_fwd (hashmap.List T) self.hashmap_hash_map_slots hash_mod + Vec.index_mut (hashmap.List T) self.hashmap_hash_map_slots hash_mod let x ← hashmap.HashMap.remove_from_list_fwd T key l match x with | Option.none => do let l0 ← hashmap.HashMap.remove_from_list_back T key l let v ← - vec_index_mut_back (hashmap.List T) self.hashmap_hash_map_slots + Vec.index_mut_back (hashmap.List T) self.hashmap_hash_map_slots hash_mod l0 Result.ret { self with hashmap_hash_map_slots := v } | Option.some x0 => @@ -460,7 +458,7 @@ def hashmap.HashMap.remove_back (Usize.ofInt 1 (by intlit)) let l0 ← hashmap.HashMap.remove_from_list_back T key l let v ← - vec_index_mut_back (hashmap.List T) self.hashmap_hash_map_slots + Vec.index_mut_back (hashmap.List T) self.hashmap_hash_map_slots hash_mod l0 Result.ret { diff --git a/tests/lean/Loops/Funs.lean b/tests/lean/Loops/Funs.lean index 383bc819..6e6eef3b 100644 --- a/tests/lean/Loops/Funs.lean +++ b/tests/lean/Loops/Funs.lean @@ -54,12 +54,12 @@ def sum_with_shared_borrows_fwd (max : U32) : Result U32 := /- [loops::clear] -/ divergent def clear_loop_fwd_back (v : Vec U32) (i : Usize) : Result (Vec U32) := - let i0 := vec_len U32 v + let i0 := Vec.len U32 v if i < i0 then do let i1 ← i + (Usize.ofInt 1 (by intlit)) - let v0 ← vec_index_mut_back U32 v i (U32.ofInt 0 (by intlit)) + let v0 ← Vec.index_mut_back U32 v i (U32.ofInt 0 (by intlit)) clear_loop_fwd_back v0 i1 else Result.ret v @@ -145,7 +145,7 @@ divergent def get_elem_mut_loop_fwd /- [loops::get_elem_mut] -/ def get_elem_mut_fwd (slots : Vec (List Usize)) (x : Usize) : Result Usize := do - let l ← vec_index_mut_fwd (List Usize) slots (Usize.ofInt 0 (by intlit)) + let l ← Vec.index_mut (List Usize) slots (Usize.ofInt 0 (by intlit)) get_elem_mut_loop_fwd x l /- [loops::get_elem_mut] -/ @@ -167,9 +167,9 @@ def get_elem_mut_back Result (Vec (List Usize)) := do - let l ← vec_index_mut_fwd (List Usize) slots (Usize.ofInt 0 (by intlit)) + let l ← Vec.index_mut (List Usize) slots (Usize.ofInt 0 (by intlit)) let l0 ← get_elem_mut_loop_back x l ret0 - vec_index_mut_back (List Usize) slots (Usize.ofInt 0 (by intlit)) l0 + Vec.index_mut_back (List Usize) slots (Usize.ofInt 0 (by intlit)) l0 /- [loops::get_elem_shared] -/ divergent def get_elem_shared_loop_fwd @@ -185,7 +185,7 @@ divergent def get_elem_shared_loop_fwd def get_elem_shared_fwd (slots : Vec (List Usize)) (x : Usize) : Result Usize := do - let l ← vec_index_fwd (List Usize) slots (Usize.ofInt 0 (by intlit)) + let l ← Vec.index (List Usize) slots (Usize.ofInt 0 (by intlit)) get_elem_shared_loop_fwd x l /- [loops::id_mut] -/ diff --git a/tests/lean/NoNestedBorrows.lean b/tests/lean/NoNestedBorrows.lean index cdbbcd67..e4fa7612 100644 --- a/tests/lean/NoNestedBorrows.lean +++ b/tests/lean/NoNestedBorrows.lean @@ -321,7 +321,7 @@ divergent def list_rev_aux_fwd /- [no_nested_borrows::list_rev] -/ def list_rev_fwd_back (T : Type) (l : List T) : Result (List T) := - let li := mem_replace_fwd (List T) l List.Nil + let li := mem.replace_fwd (List T) l List.Nil list_rev_aux_fwd T li List.Nil /- [no_nested_borrows::test_list_functions] -/ @@ -514,7 +514,7 @@ def test_weird_borrows1_fwd : Result Unit := /- [no_nested_borrows::test_mem_replace] -/ def test_mem_replace_fwd_back (px : U32) : Result U32 := - let y := mem_replace_fwd U32 px (U32.ofInt 1 (by intlit)) + let y := mem.replace_fwd U32 px (U32.ofInt 1 (by intlit)) if not (y = (U32.ofInt 0 (by intlit))) then Result.fail Error.panic else Result.ret (U32.ofInt 2 (by intlit)) diff --git a/tests/lean/lake-manifest.json b/tests/lean/lake-manifest.json index 637bda23..75f7010a 100644 --- a/tests/lean/lake-manifest.json +++ b/tests/lean/lake-manifest.json @@ -11,7 +11,7 @@ {"git": {"url": "https://github.com/leanprover-community/mathlib4.git", "subDir?": null, - "rev": "cc5d11f24e1b92db65ec3389bb5142f4b2d7670e", + "rev": "bb4fb766e41dd3a64197263ec132c7f9c4b50065", "name": "mathlib", "inputRev?": null}}, {"git": |