diff options
47 files changed, 733 insertions, 523 deletions
diff --git a/backends/lean/Base/Arith/Int.lean b/backends/lean/Base/Arith/Int.lean index 1d3e82be..068d6f2f 100644 --- a/backends/lean/Base/Arith/Int.lean +++ b/backends/lean/Base/Arith/Int.lean @@ -21,6 +21,11 @@ class HasIntProp (a : Sort u) where prop_ty : a → Prop prop : ∀ x:a, prop_ty x +/- Terms that induces predicates: if we can find the term `x`, we can introduce `concl` in the context. -/ +class HasIntPred {a: Sort u} (x: a) where + concl : Prop + prop : concl + /- Proposition with implications: if we find P we can introduce Q in the context -/ class PropHasImp (x : Sort u) where concl : Prop @@ -106,36 +111,52 @@ def collectInstancesFromMainCtx (k : Expr → MetaM (Option Expr)) : Tactic.Tact let hs := HashSet.empty -- Explore the declarations let decls ← ctx.getDecls - let hs ← decls.foldlM (fun hs d => collectInstances k hs d.toExpr) hs + let hs ← decls.foldlM (fun hs d => do + -- Collect instances over all subexpressions in the context. + -- Note that we explore the *type* of the local declarations: if we have + -- for instance `h : A ∧ B` in the context, the expression itself is simply + -- `h`; the information we are interested in is its type. + let ty ← Lean.Meta.inferType d.toExpr + collectInstances k hs ty + ) hs -- Also explore the goal collectInstances k hs (← Tactic.getMainTarget) -- Helper -def lookupProp (fName : String) (className : Name) (e : Expr) : MetaM (Option Expr) := do +def lookupProp (fName : String) (className : Name) (e : Expr) + (instantiateClassFn : Expr → MetaM (Array Expr)) + (instantiateProjectionFn : Expr → MetaM (Array Expr)) : MetaM (Option Expr) := do trace[Arith] fName -- TODO: do we need Lean.observing? -- This actually eliminates the error messages trace[Arith] m!"{fName}: {e}" Lean.observing? do trace[Arith] m!"{fName}: observing: {e}" - let ty ← Lean.Meta.inferType e - let hasProp ← mkAppM className #[ty] + let hasProp ← mkAppM className (← instantiateClassFn e) let hasPropInst ← trySynthInstance hasProp match hasPropInst with | LOption.some i => trace[Arith] "Found {fName} instance" let i_prop ← mkProjection i (Name.mkSimple "prop") - some (← mkAppM' i_prop #[e]) + some (← mkAppM' i_prop (← instantiateProjectionFn e)) | _ => none -- Return an instance of `HasIntProp` for `e` if it has some def lookupHasIntProp (e : Expr) : MetaM (Option Expr) := - lookupProp "lookupHasIntProp" ``HasIntProp e + lookupProp "lookupHasIntProp" ``HasIntProp e (fun e => do pure #[← Lean.Meta.inferType e]) (fun e => pure #[e]) -- Collect the instances of `HasIntProp` for the subexpressions in the context def collectHasIntPropInstancesFromMainCtx : Tactic.TacticM (HashSet Expr) := do collectInstancesFromMainCtx lookupHasIntProp +-- Return an instance of `HasIntPred` for `e` if it has some +def lookupHasIntPred (e : Expr) : MetaM (Option Expr) := + lookupProp "lookupHasIntPred" ``HasIntPred e (fun term => pure #[term]) (fun _ => pure #[]) + +-- Collect the instances of `HasIntPred` for the subexpressions in the context +def collectHasIntPredInstancesFromMainCtx : Tactic.TacticM (HashSet Expr) := do + collectInstancesFromMainCtx lookupHasIntPred + -- Return an instance of `PropHasImp` for `e` if it has some def lookupPropHasImp (e : Expr) : MetaM (Option Expr) := do trace[Arith] m!"lookupPropHasImp: {e}" @@ -193,6 +214,13 @@ def introHasIntPropInstances : Tactic.TacticM (Array Expr) := do elab "intro_has_int_prop_instances" : tactic => do let _ ← introHasIntPropInstances +def introHasIntPredInstances : Tactic.TacticM (Array Expr) := do + trace[Arith] "Introducing the HasIntPred instances" + introInstances ``HasIntPred.concl lookupHasIntPred + +elab "intro_has_int_pred_instances" : tactic => do + let _ ← introHasIntPredInstances + def introPropHasImpInstances : Tactic.TacticM (Array Expr) := do trace[Arith] "Introducing the PropHasImp instances" introInstances ``PropHasImp.concl lookupPropHasImp @@ -208,6 +236,8 @@ def intTacPreprocess (extraPreprocess : Tactic.TacticM Unit) : Tactic.TacticM U Tactic.withMainContext do -- Introduce the instances of `HasIntProp` let _ ← introHasIntPropInstances + -- Introduce the instances of `HasIntPred` + let _ ← introHasIntPredInstances -- Introduce the instances of `PropHasImp` let _ ← introPropHasImpInstances -- Extra preprocessing diff --git a/backends/lean/Base/IList/IList.lean b/backends/lean/Base/IList/IList.lean index a1897191..96843f55 100644 --- a/backends/lean/Base/IList/IList.lean +++ b/backends/lean/Base/IList/IList.lean @@ -18,9 +18,11 @@ theorem len_pos : 0 ≤ (ls : List α).len := by induction ls <;> simp [*] omega -instance (a : Type u) : Arith.HasIntProp (List a) where - prop_ty := λ ls => 0 ≤ ls.len - prop := λ ls => ls.len_pos +instance (l: List a) : Arith.HasIntPred (l.len) where + concl := 0 ≤ l.len + prop := l.len_pos + +example (l: List a): 0 ≤ l.len := by scalar_tac -- Remark: if i < 0, then the result is none def indexOpt (ls : List α) (i : Int) : Option α := diff --git a/backends/lean/lake-manifest.json b/backends/lean/lake-manifest.json index de2e22cd..aa2349c2 100644 --- a/backends/lean/lake-manifest.json +++ b/backends/lean/lake-manifest.json @@ -4,7 +4,7 @@ [{"url": "https://github.com/leanprover-community/batteries", "type": "git", "subDir": null, - "rev": "f96a34401de084c73c787ecb45b85d4fb47bb981", + "rev": "47e4cc5c5800c07d9bf232173c9941fa5bf68589", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -22,7 +22,7 @@ {"url": "https://github.com/leanprover-community/aesop", "type": "git", "subDir": null, - "rev": "7e3bd939c6badfcb1e607c0fddb509548baafd05", + "rev": "882561b77bd2aaa98bd8665a56821062bdb3034c", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -49,7 +49,7 @@ {"url": "https://github.com/leanprover-community/import-graph.git", "type": "git", "subDir": null, - "rev": "7983e959f8f4a79313215720de3ef1eca2d6d474", + "rev": "1588be870b9c76fe62286e8f42f0b4dafa154c96", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -58,7 +58,7 @@ {"url": "https://github.com/leanprover-community/mathlib4.git", "type": "git", "subDir": null, - "rev": "659d35143a857ceb5ba7c02a0e1530a1c7aec70c", + "rev": "77e1ea0a339a4663eced9cacc3a46eb45f967b51", "name": "mathlib", "manifestFile": "lake-manifest.json", "inputRev": null, diff --git a/backends/lean/lean-toolchain b/backends/lean/lean-toolchain index 0ba3faf8..29c0cea4 100644 --- a/backends/lean/lean-toolchain +++ b/backends/lean/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.9.0-rc1 +leanprover/lean4:v4.9.0-rc2 diff --git a/compiler/Extract.ml b/compiler/Extract.ml index 4acf3f99..b1adb936 100644 --- a/compiler/Extract.ml +++ b/compiler/Extract.ml @@ -1448,6 +1448,10 @@ let extract_fun_decl_gen (ctx : extraction_ctx) (fmt : F.formatter) (* Open two boxes for the definition, so that whenever possible it gets printed on * one line and indents are correct *) F.pp_open_hvbox fmt 0; + (* Print the attributes *) + if def.backend_attributes.reducible && backend () = Lean then ( + F.pp_print_string fmt "@[reducible]"; + F.pp_print_space fmt ()); F.pp_open_vbox fmt ctx.indent_incr; (* For HOL4: we may need to put parentheses around the definition *) let parenthesize = backend () = HOL4 && decl_is_not_last_from_group kind in diff --git a/compiler/Pure.ml b/compiler/Pure.ml index d07b8cfa..f7445575 100644 --- a/compiler/Pure.ml +++ b/compiler/Pure.ml @@ -1077,11 +1077,18 @@ type fun_body = { type item_kind = A.item_kind [@@deriving show] +(** Attributes to add to the generated code *) +type backend_attributes = { + reducible : bool; (** Lean "reducible" attribute *) +} +[@@deriving show] + type fun_decl = { def_id : FunDeclId.id; is_local : bool; span : span; kind : item_kind; + backend_attributes : backend_attributes; num_loops : int; (** The number of loops in the parent forward function (basically the number of loops appearing in the original Rust functions, unless some loops are diff --git a/compiler/PureMicroPasses.ml b/compiler/PureMicroPasses.ml index b0cba250..8b95f729 100644 --- a/compiler/PureMicroPasses.ml +++ b/compiler/PureMicroPasses.ml @@ -1498,6 +1498,7 @@ let decompose_loops (_ctx : trans_ctx) (def : fun_decl) : is_local = def.is_local; span = loop.span; kind = def.kind; + backend_attributes = def.backend_attributes; num_loops; loop_id = Some loop.loop_id; llbc_name = def.llbc_name; @@ -2277,6 +2278,52 @@ let filter_loop_inputs (ctx : trans_ctx) (transl : pure_fun_translation list) : (* Return *) transl +(** Update the [reducible] attribute. + + For now we mark a function as reducible when its body is only a call to a loop + function. This situation often happens for simple functions whose body contains + a loop: we introduce an intermediate function for the loop body, and the + translation of the function itself simply calls the loop body. By marking + the function as reducible, we allow tactics like [simp] or [progress] to + see through the definition. + *) +let compute_reducible (_ctx : trans_ctx) (transl : pure_fun_translation list) : + pure_fun_translation list = + let update_one (trans : pure_fun_translation) : pure_fun_translation = + match trans.f.body with + | None -> trans + | Some body -> ( + (* Check if the body is exactly a call to a loop function. + Note that we check that the arguments are exactly the input + variables - otherwise we may not want the call to be reducible; + for instance when using the [progress] tactic we might want to + use a more specialized specification theorem. *) + let app, args = destruct_apps body.body in + match app.e with + | Qualif + { + id = FunOrOp (Fun (FromLlbc (FunId fid, Some _lp_id))); + generics = _; + } + when fid = FRegular trans.f.def_id -> + if + List.length body.inputs = List.length args + && List.for_all + (fun ((var, arg) : var * texpression) -> + match arg.e with + | Var var_id -> var_id = var.id + | _ -> false) + (List.combine body.inputs args) + then + let f = + { trans.f with backend_attributes = { reducible = true } } + in + { trans with f } + else trans + | _ -> trans) + in + List.map update_one transl + (** Apply all the micro-passes to a function. As loops are initially directly integrated into the function definition, @@ -2337,4 +2384,7 @@ let apply_passes_to_pure_fun_translations (ctx : trans_ctx) (* Filter the useless inputs in the loop functions (loops are initially parameterized by *all* the symbolic values in the context, because they may access any of them). *) - filter_loop_inputs ctx transl + let transl = filter_loop_inputs ctx transl in + + (* Update the "reducible" attribute *) + compute_reducible ctx transl diff --git a/compiler/SymbolicToPure.ml b/compiler/SymbolicToPure.ml index 87f1128d..ad61ddd1 100644 --- a/compiler/SymbolicToPure.ml +++ b/compiler/SymbolicToPure.ml @@ -1744,11 +1744,11 @@ and aloan_content_to_consumed (ctx : bs_ctx) (ectx : C.eval_ctx) (* Ignore *) None -and aborrow_content_to_consumed (_ctx : bs_ctx) (bc : V.aborrow_content) : +and aborrow_content_to_consumed (ctx : bs_ctx) (bc : V.aborrow_content) : texpression option = match bc with | V.AMutBorrow (_, _, _) | ASharedBorrow (_, _) | AIgnoredMutBorrow (_, _) -> - craise __FILE__ __LINE__ _ctx.span "Unreachable" + craise __FILE__ __LINE__ ctx.span "Unreachable" | AEndedMutBorrow (_, _) -> (* We collect consumed values: ignore *) None @@ -3894,12 +3894,14 @@ let translate_fun_decl (ctx : bs_ctx) (body : S.expression option) : fun_decl = let loop_id = None in (* Assemble the declaration *) + let backend_attributes = { reducible = false } in let def : fun_decl = { def_id; is_local = def.is_local; span = def.item_meta.span; kind = def.kind; + backend_attributes; num_loops; loop_id; llbc_name; diff --git a/compiler/Translate.ml b/compiler/Translate.ml index 2bc9bb25..23c0782a 100644 --- a/compiler/Translate.ml +++ b/compiler/Translate.ml @@ -927,6 +927,15 @@ let extract_file (config : gen_config) (ctx : gen_ctx) (fi : extract_file_info) List.iter (fun m -> Printf.fprintf out "import %s\n" m) fi.custom_includes; (* Always open the Primitives namespace *) Printf.fprintf out "open Primitives\n"; + (* It happens that we generate duplicated namespaces, like `betree.betree`. + We deactivate the linter for this, because otherwise it leads to too much + noise. *) + Printf.fprintf out "set_option linter.dupNamespace false\n"; + (* The mathlib linter generates warnings when we use hash commands like `#assert`: + we deactivate this linter. *) + Printf.fprintf out "set_option linter.hashCommand false\n"; + (* Definitions often contain unused variables: deactivate the corresponding linter *) + Printf.fprintf out "set_option linter.unusedVariables false\n"; (* If we are inside the namespace: declare it *) if fi.in_namespace then Printf.fprintf out "\nnamespace %s\n" fi.namespace; (* We might need to open the namespace *) diff --git a/tests/coq/hashmap/Hashmap_Funs.v b/tests/coq/hashmap/Hashmap_Funs.v index 02aa0269..d63c6f72 100644 --- a/tests/coq/hashmap/Hashmap_Funs.v +++ b/tests/coq/hashmap/Hashmap_Funs.v @@ -13,22 +13,22 @@ Include Hashmap_FunsExternal. Module Hashmap_Funs. (** [hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 37:0-37:32 *) + Source: 'tests/src/hashmap.rs', lines 38:0-38:32 *) Definition hash_key (k : usize) : result usize := Ok k. (** [hashmap::{hashmap::HashMap<T>}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 60:4-66:5 *) + Source: 'tests/src/hashmap.rs', lines 61:4-67:5 *) Fixpoint hashMap_allocate_slots_loop - (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n1 : usize) : - result (alloc_vec_Vec (List_t T)) + (T : Type) (n : nat) (slots : alloc_vec_Vec (AList_t T)) (n1 : usize) : + result (alloc_vec_Vec (AList_t T)) := match n with | O => Fail_ OutOfFuel | S n2 => if n1 s> 0%usize then ( - slots1 <- alloc_vec_Vec_push (List_t T) slots List_Nil; + slots1 <- alloc_vec_Vec_push (AList_t T) slots AList_Nil; n3 <- usize_sub n1 1%usize; hashMap_allocate_slots_loop T n2 slots1 n3) else Ok slots @@ -36,22 +36,22 @@ Fixpoint hashMap_allocate_slots_loop . (** [hashmap::{hashmap::HashMap<T>}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 60:4-60:76 *) + Source: 'tests/src/hashmap.rs', lines 61:4-61:78 *) Definition hashMap_allocate_slots - (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n1 : usize) : - result (alloc_vec_Vec (List_t T)) + (T : Type) (n : nat) (slots : alloc_vec_Vec (AList_t T)) (n1 : usize) : + result (alloc_vec_Vec (AList_t T)) := hashMap_allocate_slots_loop T n slots n1 . (** [hashmap::{hashmap::HashMap<T>}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 69:4-73:13 *) + Source: 'tests/src/hashmap.rs', lines 70:4-74:13 *) Definition hashMap_new_with_capacity (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : result (HashMap_t T) := - slots <- hashMap_allocate_slots T n (alloc_vec_Vec_new (List_t T)) capacity; + slots <- hashMap_allocate_slots T n (alloc_vec_Vec_new (AList_t T)) capacity; i <- usize_mul capacity max_load_dividend; i1 <- usize_div i max_load_divisor; Ok @@ -64,36 +64,36 @@ Definition hashMap_new_with_capacity . (** [hashmap::{hashmap::HashMap<T>}::new]: - Source: 'tests/src/hashmap.rs', lines 85:4-85:24 *) + Source: 'tests/src/hashmap.rs', lines 86:4-86: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<T>}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 90:4-98:5 *) + Source: 'tests/src/hashmap.rs', lines 91:4-99: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)) + (T : Type) (n : nat) (slots : alloc_vec_Vec (AList_t T)) (i : usize) : + result (alloc_vec_Vec (AList_t T)) := match n with | O => Fail_ OutOfFuel | S n1 => - let i1 := alloc_vec_Vec_len (List_t T) slots in + let i1 := alloc_vec_Vec_len (AList_t T) slots in if i s< i1 then ( p <- - alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) slots i; + alloc_vec_Vec_index_mut (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) slots i; let (_, index_mut_back) := p in i2 <- usize_add i 1%usize; - slots1 <- index_mut_back List_Nil; + slots1 <- index_mut_back AList_Nil; hashMap_clear_loop T n1 slots1 i2) else Ok slots end . (** [hashmap::{hashmap::HashMap<T>}::clear]: - Source: 'tests/src/hashmap.rs', lines 90:4-90:27 *) + Source: 'tests/src/hashmap.rs', lines 91:4-91:27 *) Definition hashMap_clear (T : Type) (n : nat) (self : HashMap_t T) : result (HashMap_t T) := hm <- hashMap_clear_loop T n self.(hashMap_slots) 0%usize; @@ -107,62 +107,62 @@ Definition hashMap_clear . (** [hashmap::{hashmap::HashMap<T>}::len]: - Source: 'tests/src/hashmap.rs', lines 100:4-100:30 *) + Source: 'tests/src/hashmap.rs', lines 101:4-101:30 *) Definition hashMap_len (T : Type) (self : HashMap_t T) : result usize := Ok self.(hashMap_num_entries) . (** [hashmap::{hashmap::HashMap<T>}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 107:4-124:5 *) + Source: 'tests/src/hashmap.rs', lines 108:4-125:5 *) Fixpoint hashMap_insert_in_list_loop - (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : - result (bool * (List_t T)) + (T : Type) (n : nat) (key : usize) (value : T) (ls : AList_t T) : + result (bool * (AList_t T)) := match n with | O => Fail_ OutOfFuel | S n1 => match ls with - | List_Cons ckey cvalue tl => + | AList_Cons ckey cvalue tl => if ckey s= key - then Ok (false, List_Cons ckey value tl) + then Ok (false, AList_Cons ckey value tl) else ( p <- hashMap_insert_in_list_loop T n1 key value tl; let (b, tl1) := p in - Ok (b, List_Cons ckey cvalue tl1)) - | List_Nil => Ok (true, List_Cons key value List_Nil) + Ok (b, AList_Cons ckey cvalue tl1)) + | AList_Nil => Ok (true, AList_Cons key value AList_Nil) end end . (** [hashmap::{hashmap::HashMap<T>}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 107:4-107:71 *) + Source: 'tests/src/hashmap.rs', lines 108:4-108:72 *) Definition hashMap_insert_in_list - (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : - result (bool * (List_t T)) + (T : Type) (n : nat) (key : usize) (value : T) (ls : AList_t T) : + result (bool * (AList_t T)) := hashMap_insert_in_list_loop T n key value ls . (** [hashmap::{hashmap::HashMap<T>}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 127:4-127:54 *) + Source: 'tests/src/hashmap.rs', lines 128:4-128:54 *) Definition hashMap_insert_no_resize (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) : result (HashMap_t T) := hash <- hash_key key; - let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in + let i := alloc_vec_Vec_len (AList_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; p <- - alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) + alloc_vec_Vec_index_mut (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) self.(hashMap_slots) hash_mod; - let (l, index_mut_back) := p in - p1 <- hashMap_insert_in_list T n key value l; - let (inserted, l1) := p1 in + let (a, index_mut_back) := p in + p1 <- hashMap_insert_in_list T n key value a; + let (inserted, a1) := p1 in if inserted then ( i1 <- usize_add self.(hashMap_num_entries) 1%usize; - v <- index_mut_back l1; + v <- index_mut_back a1; Ok {| hashMap_num_entries := i1; @@ -171,7 +171,7 @@ Definition hashMap_insert_no_resize hashMap_slots := v |}) else ( - v <- index_mut_back l1; + v <- index_mut_back a1; Ok {| hashMap_num_entries := self.(hashMap_num_entries); @@ -182,74 +182,74 @@ Definition hashMap_insert_no_resize . (** [hashmap::{hashmap::HashMap<T>}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 193:4-206:5 *) + Source: 'tests/src/hashmap.rs', lines 194:4-207:5 *) Fixpoint hashMap_move_elements_from_list_loop - (T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) : + (T : Type) (n : nat) (ntable : HashMap_t T) (ls : AList_t T) : result (HashMap_t T) := match n with | O => Fail_ OutOfFuel | S n1 => match ls with - | List_Cons k v tl => + | AList_Cons k v tl => ntable1 <- hashMap_insert_no_resize T n1 ntable k v; hashMap_move_elements_from_list_loop T n1 ntable1 tl - | List_Nil => Ok ntable + | AList_Nil => Ok ntable end end . (** [hashmap::{hashmap::HashMap<T>}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 193:4-193:72 *) + Source: 'tests/src/hashmap.rs', lines 194:4-194:73 *) Definition hashMap_move_elements_from_list - (T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) : + (T : Type) (n : nat) (ntable : HashMap_t T) (ls : AList_t T) : result (HashMap_t T) := hashMap_move_elements_from_list_loop T n ntable ls . (** [hashmap::{hashmap::HashMap<T>}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 181:4-190:5 *) + Source: 'tests/src/hashmap.rs', lines 182:4-191:5 *) 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))) + (slots : alloc_vec_Vec (AList_t T)) (i : usize) : + result ((HashMap_t T) * (alloc_vec_Vec (AList_t T))) := match n with | O => Fail_ OutOfFuel | S n1 => - let i1 := alloc_vec_Vec_len (List_t T) slots in + let i1 := alloc_vec_Vec_len (AList_t T) slots in if i s< i1 then ( p <- - alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) slots i; - let (l, index_mut_back) := p in - let (ls, l1) := core_mem_replace (List_t T) l List_Nil in + alloc_vec_Vec_index_mut (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) slots i; + let (a, index_mut_back) := p in + let (ls, a1) := core_mem_replace (AList_t T) a AList_Nil in ntable1 <- hashMap_move_elements_from_list T n1 ntable ls; i2 <- usize_add i 1%usize; - slots1 <- index_mut_back l1; + slots1 <- index_mut_back a1; hashMap_move_elements_loop T n1 ntable1 slots1 i2) else Ok (ntable, slots) end . (** [hashmap::{hashmap::HashMap<T>}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 181:4-181:95 *) + Source: 'tests/src/hashmap.rs', lines 182:4-182:96 *) 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))) + (slots : alloc_vec_Vec (AList_t T)) (i : usize) : + result ((HashMap_t T) * (alloc_vec_Vec (AList_t T))) := hashMap_move_elements_loop T n ntable slots i . (** [hashmap::{hashmap::HashMap<T>}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 150:4-150:28 *) + Source: 'tests/src/hashmap.rs', lines 151:4-151: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; - let capacity := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in + let capacity := alloc_vec_Vec_len (AList_t T) self.(hashMap_slots) in n1 <- usize_div max_usize 2%usize; let (i, i1) := self.(hashMap_max_load_factor) in i2 <- usize_div n1 i; @@ -277,7 +277,7 @@ Definition hashMap_try_resize . (** [hashmap::{hashmap::HashMap<T>}::insert]: - Source: 'tests/src/hashmap.rs', lines 139:4-139:48 *) + Source: 'tests/src/hashmap.rs', lines 140:4-140:48 *) Definition hashMap_insert (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) : result (HashMap_t T) @@ -290,134 +290,134 @@ Definition hashMap_insert . (** [hashmap::{hashmap::HashMap<T>}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 216:4-229:5 *) + Source: 'tests/src/hashmap.rs', lines 217:4-230:5 *) Fixpoint hashMap_contains_key_in_list_loop - (T : Type) (n : nat) (key : usize) (ls : List_t T) : result bool := + (T : Type) (n : nat) (key : usize) (ls : AList_t T) : result bool := match n with | O => Fail_ OutOfFuel | S n1 => match ls with - | List_Cons ckey _ tl => + | AList_Cons ckey _ tl => if ckey s= key then Ok true else hashMap_contains_key_in_list_loop T n1 key tl - | List_Nil => Ok false + | AList_Nil => Ok false end end . (** [hashmap::{hashmap::HashMap<T>}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 216:4-216:68 *) + Source: 'tests/src/hashmap.rs', lines 217:4-217:69 *) Definition hashMap_contains_key_in_list - (T : Type) (n : nat) (key : usize) (ls : List_t T) : result bool := + (T : Type) (n : nat) (key : usize) (ls : AList_t T) : result bool := hashMap_contains_key_in_list_loop T n key ls . (** [hashmap::{hashmap::HashMap<T>}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 209:4-209:49 *) + Source: 'tests/src/hashmap.rs', lines 210:4-210:49 *) 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 + let i := alloc_vec_Vec_len (AList_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; - l <- - alloc_vec_Vec_index (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) + a <- + alloc_vec_Vec_index (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) self.(hashMap_slots) hash_mod; - hashMap_contains_key_in_list T n key l + hashMap_contains_key_in_list T n key a . (** [hashmap::{hashmap::HashMap<T>}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 234:4-247:5 *) + Source: 'tests/src/hashmap.rs', lines 235:4-248:5 *) Fixpoint hashMap_get_in_list_loop - (T : Type) (n : nat) (key : usize) (ls : List_t T) : result T := + (T : Type) (n : nat) (key : usize) (ls : AList_t T) : result T := match n with | O => Fail_ OutOfFuel | S n1 => match ls with - | List_Cons ckey cvalue tl => + | AList_Cons ckey cvalue tl => if ckey s= key then Ok cvalue else hashMap_get_in_list_loop T n1 key tl - | List_Nil => Fail_ Failure + | AList_Nil => Fail_ Failure end end . (** [hashmap::{hashmap::HashMap<T>}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 234:4-234:70 *) + Source: 'tests/src/hashmap.rs', lines 235:4-235:71 *) Definition hashMap_get_in_list - (T : Type) (n : nat) (key : usize) (ls : List_t T) : result T := + (T : Type) (n : nat) (key : usize) (ls : AList_t T) : result T := hashMap_get_in_list_loop T n key ls . (** [hashmap::{hashmap::HashMap<T>}::get]: - Source: 'tests/src/hashmap.rs', lines 249:4-249:55 *) + Source: 'tests/src/hashmap.rs', lines 250:4-250:55 *) 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 + let i := alloc_vec_Vec_len (AList_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; - l <- - alloc_vec_Vec_index (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) + a <- + alloc_vec_Vec_index (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) self.(hashMap_slots) hash_mod; - hashMap_get_in_list T n key l + hashMap_get_in_list T n key a . (** [hashmap::{hashmap::HashMap<T>}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 255:4-264:5 *) + Source: 'tests/src/hashmap.rs', lines 256:4-265:5 *) Fixpoint hashMap_get_mut_in_list_loop - (T : Type) (n : nat) (ls : List_t T) (key : usize) : - result (T * (T -> result (List_t T))) + (T : Type) (n : nat) (ls : AList_t T) (key : usize) : + result (T * (T -> result (AList_t T))) := match n with | O => Fail_ OutOfFuel | S n1 => match ls with - | List_Cons ckey cvalue tl => + | AList_Cons ckey cvalue tl => if ckey s= key then - let back := fun (ret : T) => Ok (List_Cons ckey ret tl) in + let back := fun (ret : T) => Ok (AList_Cons ckey ret tl) in Ok (cvalue, back) else ( p <- hashMap_get_mut_in_list_loop T n1 tl key; let (t, back) := p in let back1 := - fun (ret : T) => tl1 <- back ret; Ok (List_Cons ckey cvalue tl1) in + fun (ret : T) => tl1 <- back ret; Ok (AList_Cons ckey cvalue tl1) in Ok (t, back1)) - | List_Nil => Fail_ Failure + | AList_Nil => Fail_ Failure end end . (** [hashmap::{hashmap::HashMap<T>}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 255:4-255:86 *) + Source: 'tests/src/hashmap.rs', lines 256:4-256:87 *) Definition hashMap_get_mut_in_list - (T : Type) (n : nat) (ls : List_t T) (key : usize) : - result (T * (T -> result (List_t T))) + (T : Type) (n : nat) (ls : AList_t T) (key : usize) : + result (T * (T -> result (AList_t T))) := hashMap_get_mut_in_list_loop T n ls key . (** [hashmap::{hashmap::HashMap<T>}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 267:4-267:67 *) + Source: 'tests/src/hashmap.rs', lines 268:4-268:67 *) Definition hashMap_get_mut (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result (T * (T -> result (HashMap_t T))) := hash <- hash_key key; - let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in + let i := alloc_vec_Vec_len (AList_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; p <- - alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) + alloc_vec_Vec_index_mut (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) self.(hashMap_slots) hash_mod; - let (l, index_mut_back) := p in - p1 <- hashMap_get_mut_in_list T n l key; + let (a, index_mut_back) := p in + p1 <- hashMap_get_mut_in_list T n a key; let (t, get_mut_in_list_back) := p1 in let back := fun (ret : T) => - l1 <- get_mut_in_list_back ret; - v <- index_mut_back l1; + a1 <- get_mut_in_list_back ret; + v <- index_mut_back a1; Ok {| hashMap_num_entries := self.(hashMap_num_entries); @@ -429,61 +429,61 @@ Definition hashMap_get_mut . (** [hashmap::{hashmap::HashMap<T>}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 275:4-301:5 *) + Source: 'tests/src/hashmap.rs', lines 276:4-302:5 *) Fixpoint hashMap_remove_from_list_loop - (T : Type) (n : nat) (key : usize) (ls : List_t T) : - result ((option T) * (List_t T)) + (T : Type) (n : nat) (key : usize) (ls : AList_t T) : + result ((option T) * (AList_t T)) := match n with | O => Fail_ OutOfFuel | S n1 => match ls with - | List_Cons ckey t tl => + | AList_Cons ckey t tl => if ckey s= key then let (mv_ls, _) := - core_mem_replace (List_t T) (List_Cons ckey t tl) List_Nil in + core_mem_replace (AList_t T) (AList_Cons ckey t tl) AList_Nil in match mv_ls with - | List_Cons _ cvalue tl1 => Ok (Some cvalue, tl1) - | List_Nil => Fail_ Failure + | AList_Cons _ cvalue tl1 => Ok (Some cvalue, tl1) + | AList_Nil => Fail_ Failure end else ( p <- hashMap_remove_from_list_loop T n1 key tl; let (o, tl1) := p in - Ok (o, List_Cons ckey t tl1)) - | List_Nil => Ok (None, List_Nil) + Ok (o, AList_Cons ckey t tl1)) + | AList_Nil => Ok (None, AList_Nil) end end . (** [hashmap::{hashmap::HashMap<T>}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 275:4-275:69 *) + Source: 'tests/src/hashmap.rs', lines 276:4-276:70 *) Definition hashMap_remove_from_list - (T : Type) (n : nat) (key : usize) (ls : List_t T) : - result ((option T) * (List_t T)) + (T : Type) (n : nat) (key : usize) (ls : AList_t T) : + result ((option T) * (AList_t T)) := hashMap_remove_from_list_loop T n key ls . (** [hashmap::{hashmap::HashMap<T>}::remove]: - Source: 'tests/src/hashmap.rs', lines 304:4-304:52 *) + Source: 'tests/src/hashmap.rs', lines 305:4-305:52 *) Definition hashMap_remove (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result ((option T) * (HashMap_t T)) := hash <- hash_key key; - let i := alloc_vec_Vec_len (List_t T) self.(hashMap_slots) in + let i := alloc_vec_Vec_len (AList_t T) self.(hashMap_slots) in hash_mod <- usize_rem hash i; p <- - alloc_vec_Vec_index_mut (List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (List_t T)) + alloc_vec_Vec_index_mut (AList_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (AList_t T)) self.(hashMap_slots) hash_mod; - let (l, index_mut_back) := p in - p1 <- hashMap_remove_from_list T n key l; - let (x, l1) := p1 in + let (a, index_mut_back) := p in + p1 <- hashMap_remove_from_list T n key a; + let (x, a1) := p1 in match x with | None => - v <- index_mut_back l1; + v <- index_mut_back a1; Ok (None, {| hashMap_num_entries := self.(hashMap_num_entries); @@ -493,7 +493,7 @@ Definition hashMap_remove |}) | Some x1 => i1 <- usize_sub self.(hashMap_num_entries) 1%usize; - v <- index_mut_back l1; + v <- index_mut_back a1; Ok (Some x1, {| hashMap_num_entries := i1; @@ -505,7 +505,7 @@ Definition hashMap_remove . (** [hashmap::insert_on_disk]: - Source: 'tests/src/hashmap.rs', lines 335:0-335:43 *) + Source: 'tests/src/hashmap.rs', lines 336:0-336:43 *) Definition insert_on_disk (n : nat) (key : usize) (value : u64) (st : state) : result (state * unit) := p <- utils_deserialize st; @@ -515,7 +515,7 @@ Definition insert_on_disk . (** [hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 350:0-350:10 *) + Source: 'tests/src/hashmap.rs', lines 351:0-351:10 *) Definition test1 (n : nat) : result unit := hm <- hashMap_new u64 n; hm1 <- hashMap_insert u64 n hm 0%usize 42%u64; diff --git a/tests/coq/hashmap/Hashmap_FunsExternal_Template.v b/tests/coq/hashmap/Hashmap_FunsExternal_Template.v index c58b021d..ca198a8d 100644 --- a/tests/coq/hashmap/Hashmap_FunsExternal_Template.v +++ b/tests/coq/hashmap/Hashmap_FunsExternal_Template.v @@ -12,11 +12,11 @@ Include Hashmap_Types. Module Hashmap_FunsExternal_Template. (** [hashmap::utils::deserialize]: - Source: 'tests/src/hashmap.rs', lines 330:4-330:47 *) + Source: 'tests/src/hashmap.rs', lines 331:4-331:47 *) Axiom utils_deserialize : state -> result (state * (HashMap_t u64)). (** [hashmap::utils::serialize]: - Source: 'tests/src/hashmap.rs', lines 325:4-325:46 *) + Source: 'tests/src/hashmap.rs', lines 326:4-326:46 *) Axiom utils_serialize : HashMap_t u64 -> state -> result (state * unit). End Hashmap_FunsExternal_Template. diff --git a/tests/coq/hashmap/Hashmap_Types.v b/tests/coq/hashmap/Hashmap_Types.v index 452d56f8..070d6dfb 100644 --- a/tests/coq/hashmap/Hashmap_Types.v +++ b/tests/coq/hashmap/Hashmap_Types.v @@ -10,24 +10,24 @@ Require Import Hashmap_TypesExternal. Include Hashmap_TypesExternal. Module Hashmap_Types. -(** [hashmap::List] - Source: 'tests/src/hashmap.rs', lines 29:0-29:16 *) -Inductive List_t (T : Type) := -| List_Cons : usize -> T -> List_t T -> List_t T -| List_Nil : List_t T +(** [hashmap::AList] + Source: 'tests/src/hashmap.rs', lines 30:0-30:17 *) +Inductive AList_t (T : Type) := +| AList_Cons : usize -> T -> AList_t T -> AList_t T +| AList_Nil : AList_t T . -Arguments List_Cons { _ }. -Arguments List_Nil { _ }. +Arguments AList_Cons { _ }. +Arguments AList_Nil { _ }. (** [hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 45:0-45:21 *) + Source: 'tests/src/hashmap.rs', lines 46:0-46:21 *) 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); + hashMap_slots : alloc_vec_Vec (AList_t T); } . diff --git a/tests/fstar/hashmap/Hashmap.Clauses.Template.fst b/tests/fstar/hashmap/Hashmap.Clauses.Template.fst index 57003613..888cd4ec 100644 --- a/tests/fstar/hashmap/Hashmap.Clauses.Template.fst +++ b/tests/fstar/hashmap/Hashmap.Clauses.Template.fst @@ -7,65 +7,65 @@ open Hashmap.Types #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [hashmap::{hashmap::HashMap<T>}::allocate_slots]: decreases clause - Source: 'tests/src/hashmap.rs', lines 60:4-66:5 *) + Source: 'tests/src/hashmap.rs', lines 61:4-67:5 *) unfold let hashMap_allocate_slots_loop_decreases (t : Type0) - (slots : alloc_vec_Vec (list_t t)) (n : usize) : nat = + (slots : alloc_vec_Vec (aList_t t)) (n : usize) : nat = admit () (** [hashmap::{hashmap::HashMap<T>}::clear]: decreases clause - Source: 'tests/src/hashmap.rs', lines 90:4-98:5 *) + Source: 'tests/src/hashmap.rs', lines 91:4-99:5 *) unfold -let hashMap_clear_loop_decreases (t : Type0) (slots : alloc_vec_Vec (list_t t)) - (i : usize) : nat = +let hashMap_clear_loop_decreases (t : Type0) + (slots : alloc_vec_Vec (aList_t t)) (i : usize) : nat = admit () (** [hashmap::{hashmap::HashMap<T>}::insert_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 107:4-124:5 *) + Source: 'tests/src/hashmap.rs', lines 108:4-125:5 *) unfold let hashMap_insert_in_list_loop_decreases (t : Type0) (key : usize) (value : t) - (ls : list_t t) : nat = + (ls : aList_t t) : nat = admit () (** [hashmap::{hashmap::HashMap<T>}::move_elements_from_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 193:4-206:5 *) + Source: 'tests/src/hashmap.rs', lines 194:4-207:5 *) unfold let hashMap_move_elements_from_list_loop_decreases (t : Type0) - (ntable : hashMap_t t) (ls : list_t t) : nat = + (ntable : hashMap_t t) (ls : aList_t t) : nat = admit () (** [hashmap::{hashmap::HashMap<T>}::move_elements]: decreases clause - Source: 'tests/src/hashmap.rs', lines 181:4-190:5 *) + Source: 'tests/src/hashmap.rs', lines 182:4-191:5 *) unfold let hashMap_move_elements_loop_decreases (t : Type0) (ntable : hashMap_t t) - (slots : alloc_vec_Vec (list_t t)) (i : usize) : nat = + (slots : alloc_vec_Vec (aList_t t)) (i : usize) : nat = admit () (** [hashmap::{hashmap::HashMap<T>}::contains_key_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 216:4-229:5 *) + Source: 'tests/src/hashmap.rs', lines 217:4-230:5 *) unfold let hashMap_contains_key_in_list_loop_decreases (t : Type0) (key : usize) - (ls : list_t t) : nat = + (ls : aList_t t) : nat = admit () (** [hashmap::{hashmap::HashMap<T>}::get_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 234:4-247:5 *) + Source: 'tests/src/hashmap.rs', lines 235:4-248:5 *) unfold let hashMap_get_in_list_loop_decreases (t : Type0) (key : usize) - (ls : list_t t) : nat = + (ls : aList_t t) : nat = admit () (** [hashmap::{hashmap::HashMap<T>}::get_mut_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 255:4-264:5 *) + Source: 'tests/src/hashmap.rs', lines 256:4-265:5 *) unfold -let hashMap_get_mut_in_list_loop_decreases (t : Type0) (ls : list_t t) +let hashMap_get_mut_in_list_loop_decreases (t : Type0) (ls : aList_t t) (key : usize) : nat = admit () (** [hashmap::{hashmap::HashMap<T>}::remove_from_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 275:4-301:5 *) + Source: 'tests/src/hashmap.rs', lines 276:4-302:5 *) unfold let hashMap_remove_from_list_loop_decreases (t : Type0) (key : usize) - (ls : list_t t) : nat = + (ls : aList_t t) : nat = admit () diff --git a/tests/fstar/hashmap/Hashmap.Clauses.fst b/tests/fstar/hashmap/Hashmap.Clauses.fst index 6c699d05..3a389b94 100644 --- a/tests/fstar/hashmap/Hashmap.Clauses.fst +++ b/tests/fstar/hashmap/Hashmap.Clauses.fst @@ -9,53 +9,53 @@ open Hashmap.Types (** [hashmap::HashMap::allocate_slots]: decreases clause *) unfold let hashMap_allocate_slots_loop_decreases (t : Type0) - (slots : alloc_vec_Vec (list_t t)) (n : usize) : nat = n + (slots : alloc_vec_Vec (aList_t t)) (n : usize) : nat = n (** [hashmap::HashMap::clear]: decreases clause *) unfold -let hashMap_clear_loop_decreases (t : Type0) (slots : alloc_vec_Vec (list_t t)) +let hashMap_clear_loop_decreases (t : Type0) (slots : alloc_vec_Vec (aList_t t)) (i : usize) : nat = if i < length slots then length slots - i else 0 (** [hashmap::HashMap::insert_in_list]: decreases clause *) unfold let hashMap_insert_in_list_loop_decreases (t : Type0) (key : usize) (value : t) - (ls : list_t t) : list_t t = + (ls : aList_t t) : aList_t t = ls (** [hashmap::HashMap::move_elements_from_list]: decreases clause *) unfold let hashMap_move_elements_from_list_loop_decreases (t : Type0) - (ntable : hashMap_t t) (ls : list_t t) : list_t t = + (ntable : hashMap_t t) (ls : aList_t t) : aList_t t = ls (** [hashmap::HashMap::move_elements]: decreases clause *) unfold let hashMap_move_elements_loop_decreases (t : Type0) (ntable : hashMap_t t) - (slots : alloc_vec_Vec (list_t t)) (i : usize) : nat = + (slots : alloc_vec_Vec (aList_t t)) (i : usize) : nat = if i < length slots then length slots - i else 0 (** [hashmap::HashMap::contains_key_in_list]: decreases clause *) unfold let hashMap_contains_key_in_list_loop_decreases (t : Type0) (key : usize) - (ls : list_t t) : list_t t = + (ls : aList_t t) : aList_t t = ls (** [hashmap::HashMap::get_in_list]: decreases clause *) unfold -let hashMap_get_in_list_loop_decreases (t : Type0) (key : usize) (ls : list_t t) : - list_t t = +let hashMap_get_in_list_loop_decreases (t : Type0) (key : usize) (ls : aList_t t) : + aList_t t = ls (** [hashmap::HashMap::get_mut_in_list]: decreases clause *) unfold -let hashMap_get_mut_in_list_loop_decreases (t : Type0) (ls : list_t t) - (key : usize) : list_t t = +let hashMap_get_mut_in_list_loop_decreases (t : Type0) (ls : aList_t t) + (key : usize) : aList_t t = ls (** [hashmap::HashMap::remove_from_list]: decreases clause *) unfold let hashMap_remove_from_list_loop_decreases (t : Type0) (key : usize) - (ls : list_t t) : list_t t = + (ls : aList_t t) : aList_t t = ls diff --git a/tests/fstar/hashmap/Hashmap.Funs.fst b/tests/fstar/hashmap/Hashmap.Funs.fst index fb77c7ef..0569c32a 100644 --- a/tests/fstar/hashmap/Hashmap.Funs.fst +++ b/tests/fstar/hashmap/Hashmap.Funs.fst @@ -9,41 +9,41 @@ include Hashmap.Clauses #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 37:0-37:32 *) + Source: 'tests/src/hashmap.rs', lines 38:0-38:32 *) let hash_key (k : usize) : result usize = Ok k (** [hashmap::{hashmap::HashMap<T>}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 60:4-66:5 *) + Source: 'tests/src/hashmap.rs', lines 61:4-67:5 *) let rec hashMap_allocate_slots_loop - (t : Type0) (slots : alloc_vec_Vec (list_t t)) (n : usize) : - Tot (result (alloc_vec_Vec (list_t t))) + (t : Type0) (slots : alloc_vec_Vec (aList_t t)) (n : usize) : + Tot (result (alloc_vec_Vec (aList_t t))) (decreases (hashMap_allocate_slots_loop_decreases t slots n)) = if n > 0 then - let* slots1 = alloc_vec_Vec_push (list_t t) slots List_Nil in + let* slots1 = alloc_vec_Vec_push (aList_t t) slots AList_Nil in let* n1 = usize_sub n 1 in hashMap_allocate_slots_loop t slots1 n1 else Ok slots (** [hashmap::{hashmap::HashMap<T>}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 60:4-60:76 *) + Source: 'tests/src/hashmap.rs', lines 61:4-61:78 *) let hashMap_allocate_slots - (t : Type0) (slots : alloc_vec_Vec (list_t t)) (n : usize) : - result (alloc_vec_Vec (list_t t)) + (t : Type0) (slots : alloc_vec_Vec (aList_t t)) (n : usize) : + result (alloc_vec_Vec (aList_t t)) = hashMap_allocate_slots_loop t slots n (** [hashmap::{hashmap::HashMap<T>}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 69:4-73:13 *) + Source: 'tests/src/hashmap.rs', lines 70:4-74:13 *) let hashMap_new_with_capacity (t : Type0) (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : result (hashMap_t t) = - let* slots = hashMap_allocate_slots t (alloc_vec_Vec_new (list_t t)) capacity - in + let* slots = + hashMap_allocate_slots t (alloc_vec_Vec_new (aList_t t)) capacity in let* i = usize_mul capacity max_load_dividend in let* i1 = usize_div i max_load_divisor in Ok @@ -55,141 +55,141 @@ let hashMap_new_with_capacity } (** [hashmap::{hashmap::HashMap<T>}::new]: - Source: 'tests/src/hashmap.rs', lines 85:4-85:24 *) + Source: 'tests/src/hashmap.rs', lines 86:4-86:24 *) let hashMap_new (t : Type0) : result (hashMap_t t) = hashMap_new_with_capacity t 32 4 5 (** [hashmap::{hashmap::HashMap<T>}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 90:4-98:5 *) + Source: 'tests/src/hashmap.rs', lines 91:4-99:5 *) let rec hashMap_clear_loop - (t : Type0) (slots : alloc_vec_Vec (list_t t)) (i : usize) : - Tot (result (alloc_vec_Vec (list_t t))) + (t : Type0) (slots : alloc_vec_Vec (aList_t t)) (i : usize) : + Tot (result (alloc_vec_Vec (aList_t t))) (decreases (hashMap_clear_loop_decreases t slots i)) = - let i1 = alloc_vec_Vec_len (list_t t) slots in + let i1 = alloc_vec_Vec_len (aList_t t) slots in if i < i1 then let* (_, index_mut_back) = - alloc_vec_Vec_index_mut (list_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (list_t t)) slots i in + alloc_vec_Vec_index_mut (aList_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (aList_t t)) slots i in let* i2 = usize_add i 1 in - let* slots1 = index_mut_back List_Nil in + let* slots1 = index_mut_back AList_Nil in hashMap_clear_loop t slots1 i2 else Ok slots (** [hashmap::{hashmap::HashMap<T>}::clear]: - Source: 'tests/src/hashmap.rs', lines 90:4-90:27 *) + Source: 'tests/src/hashmap.rs', lines 91:4-91:27 *) let hashMap_clear (t : Type0) (self : hashMap_t t) : result (hashMap_t t) = let* hm = hashMap_clear_loop t self.slots 0 in Ok { self with num_entries = 0; slots = hm } (** [hashmap::{hashmap::HashMap<T>}::len]: - Source: 'tests/src/hashmap.rs', lines 100:4-100:30 *) + Source: 'tests/src/hashmap.rs', lines 101:4-101:30 *) let hashMap_len (t : Type0) (self : hashMap_t t) : result usize = Ok self.num_entries (** [hashmap::{hashmap::HashMap<T>}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 107:4-124:5 *) + Source: 'tests/src/hashmap.rs', lines 108:4-125:5 *) let rec hashMap_insert_in_list_loop - (t : Type0) (key : usize) (value : t) (ls : list_t t) : - Tot (result (bool & (list_t t))) + (t : Type0) (key : usize) (value : t) (ls : aList_t t) : + Tot (result (bool & (aList_t t))) (decreases (hashMap_insert_in_list_loop_decreases t key value ls)) = begin match ls with - | List_Cons ckey cvalue tl -> + | AList_Cons ckey cvalue tl -> if ckey = key - then Ok (false, List_Cons ckey value tl) + then Ok (false, AList_Cons ckey value tl) else let* (b, tl1) = hashMap_insert_in_list_loop t key value tl in - Ok (b, List_Cons ckey cvalue tl1) - | List_Nil -> Ok (true, List_Cons key value List_Nil) + Ok (b, AList_Cons ckey cvalue tl1) + | AList_Nil -> Ok (true, AList_Cons key value AList_Nil) end (** [hashmap::{hashmap::HashMap<T>}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 107:4-107:71 *) + Source: 'tests/src/hashmap.rs', lines 108:4-108:72 *) let hashMap_insert_in_list - (t : Type0) (key : usize) (value : t) (ls : list_t t) : - result (bool & (list_t t)) + (t : Type0) (key : usize) (value : t) (ls : aList_t t) : + result (bool & (aList_t t)) = hashMap_insert_in_list_loop t key value ls (** [hashmap::{hashmap::HashMap<T>}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 127:4-127:54 *) + Source: 'tests/src/hashmap.rs', lines 128:4-128:54 *) let hashMap_insert_no_resize (t : Type0) (self : hashMap_t t) (key : usize) (value : t) : result (hashMap_t t) = let* hash = hash_key key in - let i = alloc_vec_Vec_len (list_t t) self.slots in + let i = alloc_vec_Vec_len (aList_t t) self.slots in let* hash_mod = usize_rem hash i in - let* (l, index_mut_back) = - alloc_vec_Vec_index_mut (list_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (list_t t)) self.slots + let* (a, index_mut_back) = + alloc_vec_Vec_index_mut (aList_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (aList_t t)) self.slots hash_mod in - let* (inserted, l1) = hashMap_insert_in_list t key value l in + let* (inserted, a1) = hashMap_insert_in_list t key value a in if inserted then let* i1 = usize_add self.num_entries 1 in - let* v = index_mut_back l1 in + let* v = index_mut_back a1 in Ok { self with num_entries = i1; slots = v } - else let* v = index_mut_back l1 in Ok { self with slots = v } + else let* v = index_mut_back a1 in Ok { self with slots = v } (** [hashmap::{hashmap::HashMap<T>}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 193:4-206:5 *) + Source: 'tests/src/hashmap.rs', lines 194:4-207:5 *) let rec hashMap_move_elements_from_list_loop - (t : Type0) (ntable : hashMap_t t) (ls : list_t t) : + (t : Type0) (ntable : hashMap_t t) (ls : aList_t t) : Tot (result (hashMap_t t)) (decreases (hashMap_move_elements_from_list_loop_decreases t ntable ls)) = begin match ls with - | List_Cons k v tl -> + | AList_Cons k v tl -> let* ntable1 = hashMap_insert_no_resize t ntable k v in hashMap_move_elements_from_list_loop t ntable1 tl - | List_Nil -> Ok ntable + | AList_Nil -> Ok ntable end (** [hashmap::{hashmap::HashMap<T>}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 193:4-193:72 *) + Source: 'tests/src/hashmap.rs', lines 194:4-194:73 *) let hashMap_move_elements_from_list - (t : Type0) (ntable : hashMap_t t) (ls : list_t t) : result (hashMap_t t) = + (t : Type0) (ntable : hashMap_t t) (ls : aList_t t) : result (hashMap_t t) = hashMap_move_elements_from_list_loop t ntable ls (** [hashmap::{hashmap::HashMap<T>}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 181:4-190:5 *) + Source: 'tests/src/hashmap.rs', lines 182:4-191:5 *) let rec hashMap_move_elements_loop - (t : Type0) (ntable : hashMap_t t) (slots : alloc_vec_Vec (list_t t)) + (t : Type0) (ntable : hashMap_t t) (slots : alloc_vec_Vec (aList_t t)) (i : usize) : - Tot (result ((hashMap_t t) & (alloc_vec_Vec (list_t t)))) + Tot (result ((hashMap_t t) & (alloc_vec_Vec (aList_t t)))) (decreases (hashMap_move_elements_loop_decreases t ntable slots i)) = - let i1 = alloc_vec_Vec_len (list_t t) slots in + let i1 = alloc_vec_Vec_len (aList_t t) slots in if i < i1 then - let* (l, index_mut_back) = - alloc_vec_Vec_index_mut (list_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (list_t t)) slots i in - let (ls, l1) = core_mem_replace (list_t t) l List_Nil in + let* (a, index_mut_back) = + alloc_vec_Vec_index_mut (aList_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (aList_t t)) slots i in + let (ls, a1) = core_mem_replace (aList_t t) a AList_Nil in let* ntable1 = hashMap_move_elements_from_list t ntable ls in let* i2 = usize_add i 1 in - let* slots1 = index_mut_back l1 in + let* slots1 = index_mut_back a1 in hashMap_move_elements_loop t ntable1 slots1 i2 else Ok (ntable, slots) (** [hashmap::{hashmap::HashMap<T>}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 181:4-181:95 *) + Source: 'tests/src/hashmap.rs', lines 182:4-182:96 *) let hashMap_move_elements - (t : Type0) (ntable : hashMap_t t) (slots : alloc_vec_Vec (list_t t)) + (t : Type0) (ntable : hashMap_t t) (slots : alloc_vec_Vec (aList_t t)) (i : usize) : - result ((hashMap_t t) & (alloc_vec_Vec (list_t t))) + result ((hashMap_t t) & (alloc_vec_Vec (aList_t t))) = hashMap_move_elements_loop t ntable slots i (** [hashmap::{hashmap::HashMap<T>}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 150:4-150:28 *) + Source: 'tests/src/hashmap.rs', lines 151:4-151:28 *) let hashMap_try_resize (t : Type0) (self : hashMap_t t) : result (hashMap_t t) = let* max_usize = scalar_cast U32 Usize core_u32_max in - let capacity = alloc_vec_Vec_len (list_t t) self.slots in + let capacity = alloc_vec_Vec_len (aList_t t) self.slots in let* n1 = usize_div max_usize 2 in let (i, i1) = self.max_load_factor in let* i2 = usize_div n1 i in @@ -205,7 +205,7 @@ let hashMap_try_resize else Ok { self with max_load_factor = (i, i1) } (** [hashmap::{hashmap::HashMap<T>}::insert]: - Source: 'tests/src/hashmap.rs', lines 139:4-139:48 *) + Source: 'tests/src/hashmap.rs', lines 140:4-140:48 *) let hashMap_insert (t : Type0) (self : hashMap_t t) (key : usize) (value : t) : result (hashMap_t t) @@ -215,169 +215,169 @@ let hashMap_insert if i > self1.max_load then hashMap_try_resize t self1 else Ok self1 (** [hashmap::{hashmap::HashMap<T>}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 216:4-229:5 *) + Source: 'tests/src/hashmap.rs', lines 217:4-230:5 *) let rec hashMap_contains_key_in_list_loop - (t : Type0) (key : usize) (ls : list_t t) : + (t : Type0) (key : usize) (ls : aList_t t) : Tot (result bool) (decreases (hashMap_contains_key_in_list_loop_decreases t key ls)) = begin match ls with - | List_Cons ckey _ tl -> + | AList_Cons ckey _ tl -> if ckey = key then Ok true else hashMap_contains_key_in_list_loop t key tl - | List_Nil -> Ok false + | AList_Nil -> Ok false end (** [hashmap::{hashmap::HashMap<T>}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 216:4-216:68 *) + Source: 'tests/src/hashmap.rs', lines 217:4-217:69 *) let hashMap_contains_key_in_list - (t : Type0) (key : usize) (ls : list_t t) : result bool = + (t : Type0) (key : usize) (ls : aList_t t) : result bool = hashMap_contains_key_in_list_loop t key ls (** [hashmap::{hashmap::HashMap<T>}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 209:4-209:49 *) + Source: 'tests/src/hashmap.rs', lines 210:4-210:49 *) let hashMap_contains_key (t : Type0) (self : hashMap_t t) (key : usize) : result bool = let* hash = hash_key key in - let i = alloc_vec_Vec_len (list_t t) self.slots in + let i = alloc_vec_Vec_len (aList_t t) self.slots in let* hash_mod = usize_rem hash i in - let* l = - alloc_vec_Vec_index (list_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (list_t t)) self.slots + let* a = + alloc_vec_Vec_index (aList_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (aList_t t)) self.slots hash_mod in - hashMap_contains_key_in_list t key l + hashMap_contains_key_in_list t key a (** [hashmap::{hashmap::HashMap<T>}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 234:4-247:5 *) + Source: 'tests/src/hashmap.rs', lines 235:4-248:5 *) let rec hashMap_get_in_list_loop - (t : Type0) (key : usize) (ls : list_t t) : + (t : Type0) (key : usize) (ls : aList_t t) : Tot (result t) (decreases (hashMap_get_in_list_loop_decreases t key ls)) = begin match ls with - | List_Cons ckey cvalue tl -> + | AList_Cons ckey cvalue tl -> if ckey = key then Ok cvalue else hashMap_get_in_list_loop t key tl - | List_Nil -> Fail Failure + | AList_Nil -> Fail Failure end (** [hashmap::{hashmap::HashMap<T>}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 234:4-234:70 *) -let hashMap_get_in_list (t : Type0) (key : usize) (ls : list_t t) : result t = + Source: 'tests/src/hashmap.rs', lines 235:4-235:71 *) +let hashMap_get_in_list (t : Type0) (key : usize) (ls : aList_t t) : result t = hashMap_get_in_list_loop t key ls (** [hashmap::{hashmap::HashMap<T>}::get]: - Source: 'tests/src/hashmap.rs', lines 249:4-249:55 *) + Source: 'tests/src/hashmap.rs', lines 250:4-250:55 *) let hashMap_get (t : Type0) (self : hashMap_t t) (key : usize) : result t = let* hash = hash_key key in - let i = alloc_vec_Vec_len (list_t t) self.slots in + let i = alloc_vec_Vec_len (aList_t t) self.slots in let* hash_mod = usize_rem hash i in - let* l = - alloc_vec_Vec_index (list_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (list_t t)) self.slots + let* a = + alloc_vec_Vec_index (aList_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (aList_t t)) self.slots hash_mod in - hashMap_get_in_list t key l + hashMap_get_in_list t key a (** [hashmap::{hashmap::HashMap<T>}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 255:4-264:5 *) + Source: 'tests/src/hashmap.rs', lines 256:4-265:5 *) let rec hashMap_get_mut_in_list_loop - (t : Type0) (ls : list_t t) (key : usize) : - Tot (result (t & (t -> result (list_t t)))) + (t : Type0) (ls : aList_t t) (key : usize) : + Tot (result (t & (t -> result (aList_t t)))) (decreases (hashMap_get_mut_in_list_loop_decreases t ls key)) = begin match ls with - | List_Cons ckey cvalue tl -> + | AList_Cons ckey cvalue tl -> if ckey = key - then let back = fun ret -> Ok (List_Cons ckey ret tl) in Ok (cvalue, back) + then let back = fun ret -> Ok (AList_Cons ckey ret tl) in Ok (cvalue, back) else let* (x, back) = hashMap_get_mut_in_list_loop t tl key in let back1 = - fun ret -> let* tl1 = back ret in Ok (List_Cons ckey cvalue tl1) in + fun ret -> let* tl1 = back ret in Ok (AList_Cons ckey cvalue tl1) in Ok (x, back1) - | List_Nil -> Fail Failure + | AList_Nil -> Fail Failure end (** [hashmap::{hashmap::HashMap<T>}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 255:4-255:86 *) + Source: 'tests/src/hashmap.rs', lines 256:4-256:87 *) let hashMap_get_mut_in_list - (t : Type0) (ls : list_t t) (key : usize) : - result (t & (t -> result (list_t t))) + (t : Type0) (ls : aList_t t) (key : usize) : + result (t & (t -> result (aList_t t))) = hashMap_get_mut_in_list_loop t ls key (** [hashmap::{hashmap::HashMap<T>}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 267:4-267:67 *) + Source: 'tests/src/hashmap.rs', lines 268:4-268:67 *) let hashMap_get_mut (t : Type0) (self : hashMap_t t) (key : usize) : result (t & (t -> result (hashMap_t t))) = let* hash = hash_key key in - let i = alloc_vec_Vec_len (list_t t) self.slots in + let i = alloc_vec_Vec_len (aList_t t) self.slots in let* hash_mod = usize_rem hash i in - let* (l, index_mut_back) = - alloc_vec_Vec_index_mut (list_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (list_t t)) self.slots + let* (a, index_mut_back) = + alloc_vec_Vec_index_mut (aList_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (aList_t t)) self.slots hash_mod in - let* (x, get_mut_in_list_back) = hashMap_get_mut_in_list t l key in + let* (x, get_mut_in_list_back) = hashMap_get_mut_in_list t a key in let back = fun ret -> - let* l1 = get_mut_in_list_back ret in - let* v = index_mut_back l1 in + let* a1 = get_mut_in_list_back ret in + let* v = index_mut_back a1 in Ok { self with slots = v } in Ok (x, back) (** [hashmap::{hashmap::HashMap<T>}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 275:4-301:5 *) + Source: 'tests/src/hashmap.rs', lines 276:4-302:5 *) let rec hashMap_remove_from_list_loop - (t : Type0) (key : usize) (ls : list_t t) : - Tot (result ((option t) & (list_t t))) + (t : Type0) (key : usize) (ls : aList_t t) : + Tot (result ((option t) & (aList_t t))) (decreases (hashMap_remove_from_list_loop_decreases t key ls)) = begin match ls with - | List_Cons ckey x tl -> + | AList_Cons ckey x tl -> if ckey = key then let (mv_ls, _) = - core_mem_replace (list_t t) (List_Cons ckey x tl) List_Nil in + core_mem_replace (aList_t t) (AList_Cons ckey x tl) AList_Nil in begin match mv_ls with - | List_Cons _ cvalue tl1 -> Ok (Some cvalue, tl1) - | List_Nil -> Fail Failure + | AList_Cons _ cvalue tl1 -> Ok (Some cvalue, tl1) + | AList_Nil -> Fail Failure end else let* (o, tl1) = hashMap_remove_from_list_loop t key tl in - Ok (o, List_Cons ckey x tl1) - | List_Nil -> Ok (None, List_Nil) + Ok (o, AList_Cons ckey x tl1) + | AList_Nil -> Ok (None, AList_Nil) end (** [hashmap::{hashmap::HashMap<T>}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 275:4-275:69 *) + Source: 'tests/src/hashmap.rs', lines 276:4-276:70 *) let hashMap_remove_from_list - (t : Type0) (key : usize) (ls : list_t t) : - result ((option t) & (list_t t)) + (t : Type0) (key : usize) (ls : aList_t t) : + result ((option t) & (aList_t t)) = hashMap_remove_from_list_loop t key ls (** [hashmap::{hashmap::HashMap<T>}::remove]: - Source: 'tests/src/hashmap.rs', lines 304:4-304:52 *) + Source: 'tests/src/hashmap.rs', lines 305:4-305:52 *) let hashMap_remove (t : Type0) (self : hashMap_t t) (key : usize) : result ((option t) & (hashMap_t t)) = let* hash = hash_key key in - let i = alloc_vec_Vec_len (list_t t) self.slots in + let i = alloc_vec_Vec_len (aList_t t) self.slots in let* hash_mod = usize_rem hash i in - let* (l, index_mut_back) = - alloc_vec_Vec_index_mut (list_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (list_t t)) self.slots + let* (a, index_mut_back) = + alloc_vec_Vec_index_mut (aList_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (aList_t t)) self.slots hash_mod in - let* (x, l1) = hashMap_remove_from_list t key l in + let* (x, a1) = hashMap_remove_from_list t key a in begin match x with - | None -> let* v = index_mut_back l1 in Ok (None, { self with slots = v }) + | None -> let* v = index_mut_back a1 in Ok (None, { self with slots = v }) | Some x1 -> let* i1 = usize_sub self.num_entries 1 in - let* v = index_mut_back l1 in + let* v = index_mut_back a1 in Ok (Some x1, { self with num_entries = i1; slots = v }) end (** [hashmap::insert_on_disk]: - Source: 'tests/src/hashmap.rs', lines 335:0-335:43 *) + Source: 'tests/src/hashmap.rs', lines 336:0-336:43 *) let insert_on_disk (key : usize) (value : u64) (st : state) : result (state & unit) = let* (st1, hm) = utils_deserialize st in @@ -385,7 +385,7 @@ let insert_on_disk utils_serialize hm1 st1 (** [hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 350:0-350:10 *) + Source: 'tests/src/hashmap.rs', lines 351:0-351:10 *) let test1 : result unit = let* hm = hashMap_new u64 in let* hm1 = hashMap_insert u64 hm 0 42 in diff --git a/tests/fstar/hashmap/Hashmap.FunsExternal.fsti b/tests/fstar/hashmap/Hashmap.FunsExternal.fsti index f2f305e6..9362150d 100644 --- a/tests/fstar/hashmap/Hashmap.FunsExternal.fsti +++ b/tests/fstar/hashmap/Hashmap.FunsExternal.fsti @@ -7,10 +7,10 @@ include Hashmap.Types #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [hashmap::utils::deserialize]: - Source: 'tests/src/hashmap.rs', lines 330:4-330:47 *) + Source: 'tests/src/hashmap.rs', lines 331:4-331:47 *) val utils_deserialize : state -> result (state & (hashMap_t u64)) (** [hashmap::utils::serialize]: - Source: 'tests/src/hashmap.rs', lines 325:4-325:46 *) + Source: 'tests/src/hashmap.rs', lines 326:4-326:46 *) val utils_serialize : hashMap_t u64 -> state -> result (state & unit) diff --git a/tests/fstar/hashmap/Hashmap.Types.fst b/tests/fstar/hashmap/Hashmap.Types.fst index 818cb9d1..a10bd16c 100644 --- a/tests/fstar/hashmap/Hashmap.Types.fst +++ b/tests/fstar/hashmap/Hashmap.Types.fst @@ -6,19 +6,19 @@ include Hashmap.TypesExternal #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" -(** [hashmap::List] - Source: 'tests/src/hashmap.rs', lines 29:0-29:16 *) -type list_t (t : Type0) = -| List_Cons : usize -> t -> list_t t -> list_t t -| List_Nil : list_t t +(** [hashmap::AList] + Source: 'tests/src/hashmap.rs', lines 30:0-30:17 *) +type aList_t (t : Type0) = +| AList_Cons : usize -> t -> aList_t t -> aList_t t +| AList_Nil : aList_t t (** [hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 45:0-45:21 *) + Source: 'tests/src/hashmap.rs', lines 46:0-46:21 *) type hashMap_t (t : Type0) = { num_entries : usize; max_load_factor : (usize & usize); max_load : usize; - slots : alloc_vec_Vec (list_t t); + slots : alloc_vec_Vec (aList_t t); } diff --git a/tests/lean/Arrays.lean b/tests/lean/Arrays.lean index 9748919e..bb97d5c4 100644 --- a/tests/lean/Arrays.lean +++ b/tests/lean/Arrays.lean @@ -2,6 +2,9 @@ -- [arrays] import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace arrays diff --git a/tests/lean/Betree/Funs.lean b/tests/lean/Betree/Funs.lean index 7d8b4714..4592e91c 100644 --- a/tests/lean/Betree/Funs.lean +++ b/tests/lean/Betree/Funs.lean @@ -4,6 +4,9 @@ import Base import Betree.Types import Betree.FunsExternal open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace betree @@ -248,6 +251,7 @@ divergent def betree.Node.lookup_first_message_for_key_loop /- [betree::betree::{betree::betree::Node#5}::lookup_first_message_for_key]: Source: 'src/betree.rs', lines 792:4-795:34 -/ +@[reducible] def betree.Node.lookup_first_message_for_key (key : U64) (msgs : betree.List (U64 × betree.Message)) : Result ((betree.List (U64 × betree.Message)) × (betree.List (U64 × @@ -272,6 +276,7 @@ divergent def betree.Node.lookup_in_bindings_loop /- [betree::betree::{betree::betree::Node#5}::lookup_in_bindings]: Source: 'src/betree.rs', lines 649:4-649:84 -/ +@[reducible] def betree.Node.lookup_in_bindings (key : U64) (bindings : betree.List (U64 × U64)) : Result (Option U64) := betree.Node.lookup_in_bindings_loop key bindings @@ -307,6 +312,7 @@ divergent def betree.Node.apply_upserts_loop /- [betree::betree::{betree::betree::Node#5}::apply_upserts]: Source: 'src/betree.rs', lines 820:4-820:94 -/ +@[reducible] def betree.Node.apply_upserts (msgs : betree.List (U64 × betree.Message)) (prev : Option U64) (key : U64) : @@ -411,6 +417,7 @@ divergent def betree.Node.filter_messages_for_key_loop /- [betree::betree::{betree::betree::Node#5}::filter_messages_for_key]: Source: 'src/betree.rs', lines 683:4-683:77 -/ +@[reducible] def betree.Node.filter_messages_for_key (key : U64) (msgs : betree.List (U64 × betree.Message)) : Result (betree.List (U64 × betree.Message)) @@ -443,6 +450,7 @@ divergent def betree.Node.lookup_first_message_after_key_loop /- [betree::betree::{betree::betree::Node#5}::lookup_first_message_after_key]: Source: 'src/betree.rs', lines 694:4-697:34 -/ +@[reducible] def betree.Node.lookup_first_message_after_key (key : U64) (msgs : betree.List (U64 × betree.Message)) : Result ((betree.List (U64 × betree.Message)) × (betree.List (U64 × @@ -531,6 +539,7 @@ divergent def betree.Node.apply_messages_to_internal_loop /- [betree::betree::{betree::betree::Node#5}::apply_messages_to_internal]: Source: 'src/betree.rs', lines 518:4-521:5 -/ +@[reducible] def betree.Node.apply_messages_to_internal (msgs : betree.List (U64 × betree.Message)) (new_msgs : betree.List (U64 × betree.Message)) : @@ -563,6 +572,7 @@ divergent def betree.Node.lookup_mut_in_bindings_loop /- [betree::betree::{betree::betree::Node#5}::lookup_mut_in_bindings]: Source: 'src/betree.rs', lines 664:4-667:32 -/ +@[reducible] def betree.Node.lookup_mut_in_bindings (key : U64) (bindings : betree.List (U64 × U64)) : Result ((betree.List (U64 × U64)) × (betree.List (U64 × U64) → Result @@ -627,6 +637,7 @@ divergent def betree.Node.apply_messages_to_leaf_loop /- [betree::betree::{betree::betree::Node#5}::apply_messages_to_leaf]: Source: 'src/betree.rs', lines 463:4-466:5 -/ +@[reducible] def betree.Node.apply_messages_to_leaf (bindings : betree.List (U64 × U64)) (new_msgs : betree.List (U64 × betree.Message)) : diff --git a/tests/lean/Betree/FunsExternal_Template.lean b/tests/lean/Betree/FunsExternal_Template.lean index 014f0d83..2fb40ebc 100644 --- a/tests/lean/Betree/FunsExternal_Template.lean +++ b/tests/lean/Betree/FunsExternal_Template.lean @@ -4,6 +4,9 @@ import Base import Betree.Types open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false open betree /- [betree::betree_utils::load_internal_node]: diff --git a/tests/lean/Betree/Types.lean b/tests/lean/Betree/Types.lean index 3b46c00c..9e7c505b 100644 --- a/tests/lean/Betree/Types.lean +++ b/tests/lean/Betree/Types.lean @@ -3,6 +3,9 @@ import Base import Betree.TypesExternal open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace betree diff --git a/tests/lean/Betree/TypesExternal_Template.lean b/tests/lean/Betree/TypesExternal_Template.lean index 12fce657..f97517a6 100644 --- a/tests/lean/Betree/TypesExternal_Template.lean +++ b/tests/lean/Betree/TypesExternal_Template.lean @@ -3,6 +3,9 @@ -- This is a template file: rename it to "TypesExternal.lean" and fill the holes. import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false /- The state type used in the state-error monad -/ axiom State : Type diff --git a/tests/lean/Bitwise.lean b/tests/lean/Bitwise.lean index 23ec66b4..7f347661 100644 --- a/tests/lean/Bitwise.lean +++ b/tests/lean/Bitwise.lean @@ -2,6 +2,9 @@ -- [bitwise] import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace bitwise diff --git a/tests/lean/Constants.lean b/tests/lean/Constants.lean index ecb91c16..fff74646 100644 --- a/tests/lean/Constants.lean +++ b/tests/lean/Constants.lean @@ -2,6 +2,9 @@ -- [constants] import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace constants diff --git a/tests/lean/Demo/Demo.lean b/tests/lean/Demo/Demo.lean index 48ac2062..a557cf73 100644 --- a/tests/lean/Demo/Demo.lean +++ b/tests/lean/Demo/Demo.lean @@ -2,6 +2,9 @@ -- [demo] import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace demo @@ -112,6 +115,7 @@ divergent def list_nth_mut1_loop /- [demo::list_nth_mut1]: Source: 'tests/src/demo.rs', lines 71:0-71:77 -/ +@[reducible] def list_nth_mut1 (T : Type) (l : CList T) (i : U32) : Result (T × (T → Result (CList T))) diff --git a/tests/lean/External/Funs.lean b/tests/lean/External/Funs.lean index 1b1d5cdf..cd1883e5 100644 --- a/tests/lean/External/Funs.lean +++ b/tests/lean/External/Funs.lean @@ -4,6 +4,9 @@ import Base import External.Types import External.FunsExternal open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace external diff --git a/tests/lean/External/FunsExternal_Template.lean b/tests/lean/External/FunsExternal_Template.lean index 870a79c0..51050b21 100644 --- a/tests/lean/External/FunsExternal_Template.lean +++ b/tests/lean/External/FunsExternal_Template.lean @@ -4,6 +4,9 @@ import Base import External.Types open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false open external /- [core::cell::{core::cell::Cell<T>#10}::get]: diff --git a/tests/lean/External/Types.lean b/tests/lean/External/Types.lean index 836ddff0..50446e1c 100644 --- a/tests/lean/External/Types.lean +++ b/tests/lean/External/Types.lean @@ -3,6 +3,9 @@ import Base import External.TypesExternal open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace external diff --git a/tests/lean/External/TypesExternal_Template.lean b/tests/lean/External/TypesExternal_Template.lean index 24687d83..2cfbcc80 100644 --- a/tests/lean/External/TypesExternal_Template.lean +++ b/tests/lean/External/TypesExternal_Template.lean @@ -3,6 +3,9 @@ -- This is a template file: rename it to "TypesExternal.lean" and fill the holes. import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false /- [core::cell::Cell] Source: '/rustc/65ea825f4021eaf77f1b25139969712d65b435a4/library/core/src/cell.rs', lines 294:0-294:26 diff --git a/tests/lean/Hashmap/Funs.lean b/tests/lean/Hashmap/Funs.lean index 612e1848..7972b715 100644 --- a/tests/lean/Hashmap/Funs.lean +++ b/tests/lean/Hashmap/Funs.lean @@ -4,45 +4,49 @@ import Base import Hashmap.Types import Hashmap.FunsExternal open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace hashmap /- [hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 37:0-37:32 -/ + Source: 'tests/src/hashmap.rs', lines 38:0-38:32 -/ def hash_key (k : Usize) : Result Usize := Result.ok k /- [hashmap::{hashmap::HashMap<T>}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 60:4-66:5 -/ + Source: 'tests/src/hashmap.rs', lines 61:4-67:5 -/ divergent def HashMap.allocate_slots_loop - (T : Type) (slots : alloc.vec.Vec (List T)) (n : Usize) : - Result (alloc.vec.Vec (List T)) + (T : Type) (slots : alloc.vec.Vec (AList T)) (n : Usize) : + Result (alloc.vec.Vec (AList T)) := if n > 0#usize then do - let slots1 ← alloc.vec.Vec.push (List T) slots List.Nil + let slots1 ← alloc.vec.Vec.push (AList T) slots AList.Nil let n1 ← n - 1#usize HashMap.allocate_slots_loop T slots1 n1 else Result.ok slots /- [hashmap::{hashmap::HashMap<T>}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 60:4-60:76 -/ + Source: 'tests/src/hashmap.rs', lines 61:4-61:78 -/ +@[reducible] def HashMap.allocate_slots - (T : Type) (slots : alloc.vec.Vec (List T)) (n : Usize) : - Result (alloc.vec.Vec (List T)) + (T : Type) (slots : alloc.vec.Vec (AList T)) (n : Usize) : + Result (alloc.vec.Vec (AList T)) := HashMap.allocate_slots_loop T slots n /- [hashmap::{hashmap::HashMap<T>}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 69:4-73:13 -/ + Source: 'tests/src/hashmap.rs', lines 70:4-74:13 -/ def HashMap.new_with_capacity (T : Type) (capacity : Usize) (max_load_dividend : Usize) (max_load_divisor : Usize) : Result (HashMap T) := do - let slots ← HashMap.allocate_slots T (alloc.vec.Vec.new (List T)) capacity + let slots ← HashMap.allocate_slots T (alloc.vec.Vec.new (AList T)) capacity let i ← capacity * max_load_dividend let i1 ← i / max_load_divisor Result.ok @@ -54,141 +58,145 @@ def HashMap.new_with_capacity } /- [hashmap::{hashmap::HashMap<T>}::new]: - Source: 'tests/src/hashmap.rs', lines 85:4-85:24 -/ + Source: 'tests/src/hashmap.rs', lines 86:4-86:24 -/ def HashMap.new (T : Type) : Result (HashMap T) := HashMap.new_with_capacity T 32#usize 4#usize 5#usize /- [hashmap::{hashmap::HashMap<T>}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 90:4-98:5 -/ + Source: 'tests/src/hashmap.rs', lines 91:4-99:5 -/ divergent def HashMap.clear_loop - (T : Type) (slots : alloc.vec.Vec (List T)) (i : Usize) : - Result (alloc.vec.Vec (List T)) + (T : Type) (slots : alloc.vec.Vec (AList T)) (i : Usize) : + Result (alloc.vec.Vec (AList T)) := - let i1 := alloc.vec.Vec.len (List T) slots + let i1 := alloc.vec.Vec.len (AList T) slots if i < i1 then do let (_, index_mut_back) ← - alloc.vec.Vec.index_mut (List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (List T)) slots i + alloc.vec.Vec.index_mut (AList T) Usize + (core.slice.index.SliceIndexUsizeSliceTInst (AList T)) slots i let i2 ← i + 1#usize - let slots1 ← index_mut_back List.Nil + let slots1 ← index_mut_back AList.Nil HashMap.clear_loop T slots1 i2 else Result.ok slots /- [hashmap::{hashmap::HashMap<T>}::clear]: - Source: 'tests/src/hashmap.rs', lines 90:4-90:27 -/ + Source: 'tests/src/hashmap.rs', lines 91:4-91:27 -/ def HashMap.clear (T : Type) (self : HashMap T) : Result (HashMap T) := do let hm ← HashMap.clear_loop T self.slots 0#usize Result.ok { self with num_entries := 0#usize, slots := hm } /- [hashmap::{hashmap::HashMap<T>}::len]: - Source: 'tests/src/hashmap.rs', lines 100:4-100:30 -/ + Source: 'tests/src/hashmap.rs', lines 101:4-101:30 -/ def HashMap.len (T : Type) (self : HashMap T) : Result Usize := Result.ok self.num_entries /- [hashmap::{hashmap::HashMap<T>}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 107:4-124:5 -/ + Source: 'tests/src/hashmap.rs', lines 108:4-125:5 -/ divergent def HashMap.insert_in_list_loop - (T : Type) (key : Usize) (value : T) (ls : List T) : - Result (Bool × (List T)) + (T : Type) (key : Usize) (value : T) (ls : AList T) : + Result (Bool × (AList T)) := match ls with - | List.Cons ckey cvalue tl => + | AList.Cons ckey cvalue tl => if ckey = key - then Result.ok (false, List.Cons ckey value tl) + then Result.ok (false, AList.Cons ckey value tl) else do let (b, tl1) ← HashMap.insert_in_list_loop T key value tl - Result.ok (b, List.Cons ckey cvalue tl1) - | List.Nil => Result.ok (true, List.Cons key value List.Nil) + Result.ok (b, AList.Cons ckey cvalue tl1) + | AList.Nil => Result.ok (true, AList.Cons key value AList.Nil) /- [hashmap::{hashmap::HashMap<T>}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 107:4-107:71 -/ + Source: 'tests/src/hashmap.rs', lines 108:4-108:72 -/ +@[reducible] def HashMap.insert_in_list - (T : Type) (key : Usize) (value : T) (ls : List T) : - Result (Bool × (List T)) + (T : Type) (key : Usize) (value : T) (ls : AList T) : + Result (Bool × (AList T)) := HashMap.insert_in_list_loop T key value ls /- [hashmap::{hashmap::HashMap<T>}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 127:4-127:54 -/ + Source: 'tests/src/hashmap.rs', lines 128:4-128:54 -/ def HashMap.insert_no_resize (T : Type) (self : HashMap T) (key : Usize) (value : T) : Result (HashMap T) := do let hash ← hash_key key - let i := alloc.vec.Vec.len (List T) self.slots + let i := alloc.vec.Vec.len (AList T) self.slots let hash_mod ← hash % i - let (l, index_mut_back) ← - alloc.vec.Vec.index_mut (List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (List T)) self.slots hash_mod - let (inserted, l1) ← HashMap.insert_in_list T key value l + let (a, index_mut_back) ← + alloc.vec.Vec.index_mut (AList T) Usize + (core.slice.index.SliceIndexUsizeSliceTInst (AList T)) self.slots + hash_mod + let (inserted, a1) ← HashMap.insert_in_list T key value a if inserted then do let i1 ← self.num_entries + 1#usize - let v ← index_mut_back l1 + let v ← index_mut_back a1 Result.ok { self with num_entries := i1, slots := v } else do - let v ← index_mut_back l1 + let v ← index_mut_back a1 Result.ok { self with slots := v } /- [hashmap::{hashmap::HashMap<T>}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 193:4-206:5 -/ + Source: 'tests/src/hashmap.rs', lines 194:4-207:5 -/ divergent def HashMap.move_elements_from_list_loop - (T : Type) (ntable : HashMap T) (ls : List T) : Result (HashMap T) := + (T : Type) (ntable : HashMap T) (ls : AList T) : Result (HashMap T) := match ls with - | List.Cons k v tl => + | AList.Cons k v tl => do let ntable1 ← HashMap.insert_no_resize T ntable k v HashMap.move_elements_from_list_loop T ntable1 tl - | List.Nil => Result.ok ntable + | AList.Nil => Result.ok ntable /- [hashmap::{hashmap::HashMap<T>}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 193:4-193:72 -/ + Source: 'tests/src/hashmap.rs', lines 194:4-194:73 -/ +@[reducible] def HashMap.move_elements_from_list - (T : Type) (ntable : HashMap T) (ls : List T) : Result (HashMap T) := + (T : Type) (ntable : HashMap T) (ls : AList T) : Result (HashMap T) := HashMap.move_elements_from_list_loop T ntable ls /- [hashmap::{hashmap::HashMap<T>}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 181:4-190:5 -/ + Source: 'tests/src/hashmap.rs', lines 182:4-191:5 -/ divergent def HashMap.move_elements_loop - (T : Type) (ntable : HashMap T) (slots : alloc.vec.Vec (List T)) (i : Usize) + (T : Type) (ntable : HashMap T) (slots : alloc.vec.Vec (AList T)) (i : Usize) : - Result ((HashMap T) × (alloc.vec.Vec (List T))) + Result ((HashMap T) × (alloc.vec.Vec (AList T))) := - let i1 := alloc.vec.Vec.len (List T) slots + let i1 := alloc.vec.Vec.len (AList T) slots if i < i1 then do - let (l, index_mut_back) ← - alloc.vec.Vec.index_mut (List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (List T)) slots i - let (ls, l1) := core.mem.replace (List T) l List.Nil + let (a, index_mut_back) ← + alloc.vec.Vec.index_mut (AList T) Usize + (core.slice.index.SliceIndexUsizeSliceTInst (AList T)) slots i + let (ls, a1) := core.mem.replace (AList T) a AList.Nil let ntable1 ← HashMap.move_elements_from_list T ntable ls let i2 ← i + 1#usize - let slots1 ← index_mut_back l1 + let slots1 ← index_mut_back a1 HashMap.move_elements_loop T ntable1 slots1 i2 else Result.ok (ntable, slots) /- [hashmap::{hashmap::HashMap<T>}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 181:4-181:95 -/ + Source: 'tests/src/hashmap.rs', lines 182:4-182:96 -/ +@[reducible] def HashMap.move_elements - (T : Type) (ntable : HashMap T) (slots : alloc.vec.Vec (List T)) (i : Usize) + (T : Type) (ntable : HashMap T) (slots : alloc.vec.Vec (AList T)) (i : Usize) : - Result ((HashMap T) × (alloc.vec.Vec (List T))) + Result ((HashMap T) × (alloc.vec.Vec (AList T))) := HashMap.move_elements_loop T ntable slots i /- [hashmap::{hashmap::HashMap<T>}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 150:4-150:28 -/ + Source: 'tests/src/hashmap.rs', lines 151:4-151:28 -/ def HashMap.try_resize (T : Type) (self : HashMap T) : Result (HashMap T) := do let max_usize ← Scalar.cast .Usize core_u32_max - let capacity := alloc.vec.Vec.len (List T) self.slots + let capacity := alloc.vec.Vec.len (AList T) self.slots let n1 ← max_usize / 2#usize let (i, i1) := self.max_load_factor let i2 ← n1 / i @@ -208,7 +216,7 @@ def HashMap.try_resize (T : Type) (self : HashMap T) : Result (HashMap T) := else Result.ok { self with max_load_factor := (i, i1) } /- [hashmap::{hashmap::HashMap<T>}::insert]: - Source: 'tests/src/hashmap.rs', lines 139:4-139:48 -/ + Source: 'tests/src/hashmap.rs', lines 140:4-140:48 -/ def HashMap.insert (T : Type) (self : HashMap T) (key : Usize) (value : T) : Result (HashMap T) @@ -221,74 +229,78 @@ def HashMap.insert else Result.ok self1 /- [hashmap::{hashmap::HashMap<T>}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 216:4-229:5 -/ + Source: 'tests/src/hashmap.rs', lines 217:4-230:5 -/ divergent def HashMap.contains_key_in_list_loop - (T : Type) (key : Usize) (ls : List T) : Result Bool := + (T : Type) (key : Usize) (ls : AList T) : Result Bool := match ls with - | List.Cons ckey _ tl => + | AList.Cons ckey _ tl => if ckey = key then Result.ok true else HashMap.contains_key_in_list_loop T key tl - | List.Nil => Result.ok false + | AList.Nil => Result.ok false /- [hashmap::{hashmap::HashMap<T>}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 216:4-216:68 -/ + Source: 'tests/src/hashmap.rs', lines 217:4-217:69 -/ +@[reducible] def HashMap.contains_key_in_list - (T : Type) (key : Usize) (ls : List T) : Result Bool := + (T : Type) (key : Usize) (ls : AList T) : Result Bool := HashMap.contains_key_in_list_loop T key ls /- [hashmap::{hashmap::HashMap<T>}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 209:4-209:49 -/ + Source: 'tests/src/hashmap.rs', lines 210:4-210:49 -/ def HashMap.contains_key (T : Type) (self : HashMap T) (key : Usize) : Result Bool := do let hash ← hash_key key - let i := alloc.vec.Vec.len (List T) self.slots + let i := alloc.vec.Vec.len (AList T) self.slots let hash_mod ← hash % i - let l ← - alloc.vec.Vec.index (List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (List T)) self.slots hash_mod - HashMap.contains_key_in_list T key l + let a ← + alloc.vec.Vec.index (AList T) Usize + (core.slice.index.SliceIndexUsizeSliceTInst (AList T)) self.slots + hash_mod + HashMap.contains_key_in_list T key a /- [hashmap::{hashmap::HashMap<T>}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 234:4-247:5 -/ + Source: 'tests/src/hashmap.rs', lines 235:4-248:5 -/ divergent def HashMap.get_in_list_loop - (T : Type) (key : Usize) (ls : List T) : Result T := + (T : Type) (key : Usize) (ls : AList T) : Result T := match ls with - | List.Cons ckey cvalue tl => + | AList.Cons ckey cvalue tl => if ckey = key then Result.ok cvalue else HashMap.get_in_list_loop T key tl - | List.Nil => Result.fail .panic + | AList.Nil => Result.fail .panic /- [hashmap::{hashmap::HashMap<T>}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 234:4-234:70 -/ -def HashMap.get_in_list (T : Type) (key : Usize) (ls : List T) : Result T := + Source: 'tests/src/hashmap.rs', lines 235:4-235:71 -/ +@[reducible] +def HashMap.get_in_list (T : Type) (key : Usize) (ls : AList T) : Result T := HashMap.get_in_list_loop T key ls /- [hashmap::{hashmap::HashMap<T>}::get]: - Source: 'tests/src/hashmap.rs', lines 249:4-249:55 -/ + Source: 'tests/src/hashmap.rs', lines 250:4-250:55 -/ def HashMap.get (T : Type) (self : HashMap T) (key : Usize) : Result T := do let hash ← hash_key key - let i := alloc.vec.Vec.len (List T) self.slots + let i := alloc.vec.Vec.len (AList T) self.slots let hash_mod ← hash % i - let l ← - alloc.vec.Vec.index (List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (List T)) self.slots hash_mod - HashMap.get_in_list T key l + let a ← + alloc.vec.Vec.index (AList T) Usize + (core.slice.index.SliceIndexUsizeSliceTInst (AList T)) self.slots + hash_mod + HashMap.get_in_list T key a /- [hashmap::{hashmap::HashMap<T>}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 255:4-264:5 -/ + Source: 'tests/src/hashmap.rs', lines 256:4-265:5 -/ divergent def HashMap.get_mut_in_list_loop - (T : Type) (ls : List T) (key : Usize) : - Result (T × (T → Result (List T))) + (T : Type) (ls : AList T) (key : Usize) : + Result (T × (T → Result (AList T))) := match ls with - | List.Cons ckey cvalue tl => + | AList.Cons ckey cvalue tl => if ckey = key then - let back := fun ret => Result.ok (List.Cons ckey ret tl) + let back := fun ret => Result.ok (AList.Cons ckey ret tl) Result.ok (cvalue, back) else do @@ -297,92 +309,96 @@ divergent def HashMap.get_mut_in_list_loop fun ret => do let tl1 ← back ret - Result.ok (List.Cons ckey cvalue tl1) + Result.ok (AList.Cons ckey cvalue tl1) Result.ok (t, back1) - | List.Nil => Result.fail .panic + | AList.Nil => Result.fail .panic /- [hashmap::{hashmap::HashMap<T>}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 255:4-255:86 -/ + Source: 'tests/src/hashmap.rs', lines 256:4-256:87 -/ +@[reducible] def HashMap.get_mut_in_list - (T : Type) (ls : List T) (key : Usize) : - Result (T × (T → Result (List T))) + (T : Type) (ls : AList T) (key : Usize) : + Result (T × (T → Result (AList T))) := HashMap.get_mut_in_list_loop T ls key /- [hashmap::{hashmap::HashMap<T>}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 267:4-267:67 -/ + Source: 'tests/src/hashmap.rs', lines 268:4-268:67 -/ def HashMap.get_mut (T : Type) (self : HashMap T) (key : Usize) : Result (T × (T → Result (HashMap T))) := do let hash ← hash_key key - let i := alloc.vec.Vec.len (List T) self.slots + let i := alloc.vec.Vec.len (AList T) self.slots let hash_mod ← hash % i - let (l, index_mut_back) ← - alloc.vec.Vec.index_mut (List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (List T)) self.slots hash_mod - let (t, get_mut_in_list_back) ← HashMap.get_mut_in_list T l key + let (a, index_mut_back) ← + alloc.vec.Vec.index_mut (AList T) Usize + (core.slice.index.SliceIndexUsizeSliceTInst (AList T)) self.slots + hash_mod + let (t, get_mut_in_list_back) ← HashMap.get_mut_in_list T a key let back := fun ret => do - let l1 ← get_mut_in_list_back ret - let v ← index_mut_back l1 + let a1 ← get_mut_in_list_back ret + let v ← index_mut_back a1 Result.ok { self with slots := v } Result.ok (t, back) /- [hashmap::{hashmap::HashMap<T>}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 275:4-301:5 -/ + Source: 'tests/src/hashmap.rs', lines 276:4-302:5 -/ divergent def HashMap.remove_from_list_loop - (T : Type) (key : Usize) (ls : List T) : Result ((Option T) × (List T)) := + (T : Type) (key : Usize) (ls : AList T) : Result ((Option T) × (AList T)) := match ls with - | List.Cons ckey t tl => + | AList.Cons ckey t tl => if ckey = key then let (mv_ls, _) := - core.mem.replace (List T) (List.Cons ckey t tl) List.Nil + core.mem.replace (AList T) (AList.Cons ckey t tl) AList.Nil match mv_ls with - | List.Cons _ cvalue tl1 => Result.ok (some cvalue, tl1) - | List.Nil => Result.fail .panic + | AList.Cons _ cvalue tl1 => Result.ok (some cvalue, tl1) + | AList.Nil => Result.fail .panic else do let (o, tl1) ← HashMap.remove_from_list_loop T key tl - Result.ok (o, List.Cons ckey t tl1) - | List.Nil => Result.ok (none, List.Nil) + Result.ok (o, AList.Cons ckey t tl1) + | AList.Nil => Result.ok (none, AList.Nil) /- [hashmap::{hashmap::HashMap<T>}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 275:4-275:69 -/ + Source: 'tests/src/hashmap.rs', lines 276:4-276:70 -/ +@[reducible] def HashMap.remove_from_list - (T : Type) (key : Usize) (ls : List T) : Result ((Option T) × (List T)) := + (T : Type) (key : Usize) (ls : AList T) : Result ((Option T) × (AList T)) := HashMap.remove_from_list_loop T key ls /- [hashmap::{hashmap::HashMap<T>}::remove]: - Source: 'tests/src/hashmap.rs', lines 304:4-304:52 -/ + Source: 'tests/src/hashmap.rs', lines 305:4-305:52 -/ def HashMap.remove (T : Type) (self : HashMap T) (key : Usize) : Result ((Option T) × (HashMap T)) := do let hash ← hash_key key - let i := alloc.vec.Vec.len (List T) self.slots + let i := alloc.vec.Vec.len (AList T) self.slots let hash_mod ← hash % i - let (l, index_mut_back) ← - alloc.vec.Vec.index_mut (List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (List T)) self.slots hash_mod - let (x, l1) ← HashMap.remove_from_list T key l + let (a, index_mut_back) ← + alloc.vec.Vec.index_mut (AList T) Usize + (core.slice.index.SliceIndexUsizeSliceTInst (AList T)) self.slots + hash_mod + let (x, a1) ← HashMap.remove_from_list T key a match x with | none => do - let v ← index_mut_back l1 + let v ← index_mut_back a1 Result.ok (none, { self with slots := v }) | some x1 => do let i1 ← self.num_entries - 1#usize - let v ← index_mut_back l1 + let v ← index_mut_back a1 Result.ok (some x1, { self with num_entries := i1, slots := v }) /- [hashmap::insert_on_disk]: - Source: 'tests/src/hashmap.rs', lines 335:0-335:43 -/ + Source: 'tests/src/hashmap.rs', lines 336:0-336:43 -/ def insert_on_disk (key : Usize) (value : U64) (st : State) : Result (State × Unit) := do @@ -391,7 +407,7 @@ def insert_on_disk utils.serialize hm1 st1 /- [hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 350:0-350:10 -/ + Source: 'tests/src/hashmap.rs', lines 351:0-351:10 -/ def test1 : Result Unit := do let hm ← HashMap.new U64 diff --git a/tests/lean/Hashmap/FunsExternal_Template.lean b/tests/lean/Hashmap/FunsExternal_Template.lean index 2af57d10..ea5ceed3 100644 --- a/tests/lean/Hashmap/FunsExternal_Template.lean +++ b/tests/lean/Hashmap/FunsExternal_Template.lean @@ -4,13 +4,16 @@ import Base import Hashmap.Types open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false open hashmap /- [hashmap::utils::deserialize]: - Source: 'tests/src/hashmap.rs', lines 330:4-330:47 -/ + Source: 'tests/src/hashmap.rs', lines 331:4-331:47 -/ axiom utils.deserialize : State → Result (State × (HashMap U64)) /- [hashmap::utils::serialize]: - Source: 'tests/src/hashmap.rs', lines 325:4-325:46 -/ + Source: 'tests/src/hashmap.rs', lines 326:4-326:46 -/ axiom utils.serialize : HashMap U64 → State → Result (State × Unit) diff --git a/tests/lean/Hashmap/Properties.lean b/tests/lean/Hashmap/Properties.lean index 5be778e7..29b5834b 100644 --- a/tests/lean/Hashmap/Properties.lean +++ b/tests/lean/Hashmap/Properties.lean @@ -9,7 +9,7 @@ namespace List -- TODO: rewrite rule: match x == y with ... -> if x = y then ... else ... ? (actually doesn't work because of sugar) -- TODO: move? @[simp] -def lookup' {α : Type} (ls: _root_.List (Usize × α)) (key: Usize) : Option α := +def lookup' {α : Type} (ls: List (Usize × α)) (key: Usize) : Option α := match ls with | [] => none | (k, x) :: tl => if k = key then some x else lookup' tl key @@ -18,29 +18,27 @@ end List namespace hashmap -namespace List +namespace AList -def v {α : Type} (ls: List α) : _root_.List (Usize × α) := +def v {α : Type} (ls: AList α) : List (Usize × α) := match ls with | Nil => [] | Cons k x tl => (k, x) :: v tl -@[simp] theorem v_nil (α : Type) : (Nil : List α).v = [] := by rfl -@[simp] theorem v_cons {α : Type} k x (tl: List α) : (Cons k x tl).v = (k, x) :: v tl := by rfl +@[simp] theorem v_nil (α : Type) : (Nil : AList α).v = [] := by rfl +@[simp] theorem v_cons {α : Type} k x (tl: AList α) : (Cons k x tl).v = (k, x) :: v tl := by rfl @[simp] -abbrev lookup {α : Type} (ls: List α) (key: Usize) : Option α := +abbrev lookup {α : Type} (ls: AList α) (key: Usize) : Option α := ls.v.lookup' key @[simp] -abbrev len {α : Type} (ls : List α) : Int := ls.v.len +abbrev len {α : Type} (ls : AList α) : Int := ls.v.len -end List +end AList namespace HashMap -abbrev Core.List := _root_.List - namespace List end List @@ -55,7 +53,7 @@ theorem match_lawful_beq [BEq α] [LawfulBEq α] [DecidableEq α] (x y : α) : (x == y) = (if x = y then true else false) := by split <;> simp_all -def distinct_keys (ls : Core.List (Usize × α)) := ls.pairwise_rel (λ x y => x.fst ≠ y.fst) +def distinct_keys (ls : List (Usize × α)) := ls.pairwise_rel (λ x y => x.fst ≠ y.fst) def hash_mod_key (k : Usize) (l : Int) : Int := match hash_key k with @@ -66,33 +64,33 @@ def hash_mod_key (k : Usize) (l : Int) : Int := theorem hash_mod_key_eq : hash_mod_key k l = k.val % l := by simp [hash_mod_key, hash_key] -def slot_s_inv_hash (l i : Int) (ls : Core.List (Usize × α)) : Prop := +def slot_s_inv_hash (l i : Int) (ls : List (Usize × α)) : Prop := ls.allP (λ (k, _) => hash_mod_key k l = i) @[simp] -def slot_s_inv (l i : Int) (ls : Core.List (Usize × α)) : Prop := +def slot_s_inv (l i : Int) (ls : List (Usize × α)) : Prop := distinct_keys ls ∧ slot_s_inv_hash l i ls -def slot_t_inv (l i : Int) (s : List α) : Prop := slot_s_inv l i s.v +def slot_t_inv (l i : Int) (s : AList α) : Prop := slot_s_inv l i s.v -- Interpret the hashmap as a list of lists -def v (hm : HashMap α) : Core.List (Core.List (Usize × α)) := - hm.slots.val.map List.v +def v (hm : HashMap α) : List (List (Usize × α)) := + hm.slots.val.map AList.v -- Interpret the hashmap as an associative list -def al_v (hm : HashMap α) : Core.List (Usize × α) := +def al_v (hm : HashMap α) : List (Usize × α) := hm.v.flatten -- TODO: automatic derivation -instance : Inhabited (List α) where +instance : Inhabited (AList α) where default := .Nil @[simp] -def slots_s_inv (s : Core.List (List α)) : Prop := +def slots_s_inv (s : List (AList α)) : Prop := ∀ (i : Int), 0 ≤ i → i < s.len → slot_t_inv s.len i (s.index i) -def slots_t_inv (s : alloc.vec.Vec (List α)) : Prop := +def slots_t_inv (s : alloc.vec.Vec (AList α)) : Prop := slots_s_inv s.v @[simp] @@ -117,7 +115,7 @@ attribute [-simp] Bool.exists_bool -- of heart beats set_option maxHeartbeats 1000000 -theorem insert_in_list_spec_aux {α : Type} (l : Int) (key: Usize) (value: α) (l0: List α) +theorem insert_in_list_spec_aux {α : Type} (l : Int) (key: Usize) (value: α) (l0: AList α) (hinv : slot_s_inv_hash l (hash_mod_key key l) l0.v) (hdk : distinct_keys l0.v) : ∃ b l1, @@ -144,7 +142,7 @@ theorem insert_in_list_spec_aux {α : Type} (l : Int) (key: Usize) (value: α) ( simp [insert_in_list] rw [insert_in_list_loop] simp (config := {contextual := true}) - [List.v, slot_s_inv_hash, distinct_keys, List.pairwise_rel] + [AList.v, slot_s_inv_hash, distinct_keys, List.pairwise_rel] | Cons k v tl0 => if h: k = key then rw [insert_in_list] @@ -153,7 +151,7 @@ theorem insert_in_list_spec_aux {α : Type} (l : Int) (key: Usize) (value: α) ( exists false; simp -- TODO: why do we need to do this? split_conjs . intros; simp [*] - . simp [List.v, slot_s_inv_hash] at * + . simp [AList.v, slot_s_inv_hash] at * simp [*] . simp [*, distinct_keys] at * apply hdk @@ -162,17 +160,16 @@ theorem insert_in_list_spec_aux {α : Type} (l : Int) (key: Usize) (value: α) ( rw [insert_in_list] rw [insert_in_list_loop] simp [h] - have : slot_s_inv_hash l (hash_mod_key key l) (List.v tl0) := by - simp_all [List.v, slot_s_inv_hash] - have : distinct_keys (List.v tl0) := by + have : slot_s_inv_hash l (hash_mod_key key l) (AList.v tl0) := by + simp_all [AList.v, slot_s_inv_hash] + have : distinct_keys (AList.v tl0) := by simp [distinct_keys] at hdk simp [hdk, distinct_keys] - progress keep heq as ⟨ b, tl1 .. ⟩ - simp only [insert_in_list] at heq - have : slot_s_inv_hash l (hash_mod_key key l) (List.v (List.Cons k v tl1)) := by - simp [List.v, slot_s_inv_hash] at * + progress as ⟨ b, tl1 .. ⟩ + have : slot_s_inv_hash l (hash_mod_key key l) (AList.v (AList.Cons k v tl1)) := by + simp [AList.v, slot_s_inv_hash] at * simp [*] - have : distinct_keys ((k, v) :: List.v tl1) := by + have : distinct_keys ((k, v) :: AList.v tl1) := by simp [distinct_keys] at * simp [*] -- TODO: canonize addition by default? @@ -180,7 +177,7 @@ theorem insert_in_list_spec_aux {α : Type} (l : Int) (key: Usize) (value: α) ( simp_all [Int.add_assoc, Int.add_comm, Int.add_left_comm] @[pspec] -theorem insert_in_list_spec {α : Type} (l : Int) (key: Usize) (value: α) (l0: List α) +theorem insert_in_list_spec {α : Type} (l : Int) (key: Usize) (value: α) (l0: AList α) (hinv : slot_s_inv_hash l (hash_mod_key key l) l0.v) (hdk : distinct_keys l0.v) : ∃ b l1, @@ -203,7 +200,7 @@ theorem insert_in_list_spec {α : Type} (l : Int) (key: Usize) (value: α) (l0: exists l1 @[simp] -def slots_t_lookup (s : Core.List (List α)) (k : Usize) : Option α := +def slots_t_lookup (s : List (AList α)) (k : Usize) : Option α := let i := hash_mod_key k s.len let slot := s.index i slot.lookup k @@ -254,7 +251,7 @@ theorem insert_no_resize_spec {α : Type} (hm : HashMap α) (key : Usize) (value rw [insert_no_resize] -- Simplify. Note that this also simplifies some function calls, like array index simp [hash_key, bind_tc_ok] - have _ : (alloc.vec.Vec.len (List α) hm.slots).val ≠ 0 := by + have _ : (alloc.vec.Vec.len (AList α) hm.slots).val ≠ 0 := by intro simp_all [inv] progress as ⟨ hash_mod, hhm ⟩ @@ -344,11 +341,11 @@ theorem insert_no_resize_spec {α : Type} (hm : HashMap α) (key : Usize) (value simp only [lookup, List.lookup, len_s, al_v, HashMap.v, slots_t_lookup] at * -- We have to do a case disjunction simp_all - simp [_root_.List.update_map_eq] + simp [List.update_map_eq] -- TODO: dependent rewrites - have _ : key.val % hm.slots.val.len < (List.map List.v hm.slots.val).len := by + have _ : key.val % hm.slots.val.len < (List.map AList.v hm.slots.val).len := by simp [*] - simp [_root_.List.len_flatten_update_eq, *] + simp [List.len_flatten_update_eq, *] split <;> rename_i heq <;> simp [heq] at hlen <;> @@ -371,7 +368,7 @@ theorem insert_no_resize_spec {α : Type} (hm : HashMap α) (key : Usize) (value have _ := hinv.right.left i hipos (by simp_all) simp [hhm, h_veq, nhm_eq] at * -- TODO: annoying, we do that because simp_all fails below -- We need a case disjunction - if h_ieq : i = key.val % _root_.List.len hm.slots.val then + if h_ieq : i = key.val % List.len hm.slots.val then -- TODO: simp_all fails: "(deterministic) timeout at 'whnf'" -- Also, it is annoying to do this kind -- of rewritings by hand. We could have a different simp diff --git a/tests/lean/Hashmap/Types.lean b/tests/lean/Hashmap/Types.lean index b4301106..6f5d99a5 100644 --- a/tests/lean/Hashmap/Types.lean +++ b/tests/lean/Hashmap/Types.lean @@ -3,21 +3,24 @@ import Base import Hashmap.TypesExternal open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace hashmap -/- [hashmap::List] - Source: 'tests/src/hashmap.rs', lines 29:0-29:16 -/ -inductive List (T : Type) := -| Cons : Usize → T → List T → List T -| Nil : List T +/- [hashmap::AList] + Source: 'tests/src/hashmap.rs', lines 30:0-30:17 -/ +inductive AList (T : Type) := +| Cons : Usize → T → AList T → AList T +| Nil : AList T /- [hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 45:0-45:21 -/ + Source: 'tests/src/hashmap.rs', lines 46:0-46:21 -/ structure HashMap (T : Type) where num_entries : Usize max_load_factor : (Usize × Usize) max_load : Usize - slots : alloc.vec.Vec (List T) + slots : alloc.vec.Vec (AList T) end hashmap diff --git a/tests/lean/Hashmap/TypesExternal_Template.lean b/tests/lean/Hashmap/TypesExternal_Template.lean index 03c3d157..b6f24513 100644 --- a/tests/lean/Hashmap/TypesExternal_Template.lean +++ b/tests/lean/Hashmap/TypesExternal_Template.lean @@ -3,6 +3,9 @@ -- This is a template file: rename it to "TypesExternal.lean" and fill the holes. import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false /- The state type used in the state-error monad -/ axiom State : Type diff --git a/tests/lean/InfiniteLoop.lean b/tests/lean/InfiniteLoop.lean index 9eb8e22c..b5986ed8 100644 --- a/tests/lean/InfiniteLoop.lean +++ b/tests/lean/InfiniteLoop.lean @@ -2,6 +2,9 @@ -- [infinite_loop] import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace infinite_loop @@ -19,7 +22,7 @@ divergent def foo_loop : Result Unit := /- [infinite_loop::foo]: Source: 'tests/src/infinite-loop.rs', lines 6:0-6:8 -/ -def foo : Result Unit := - foo_loop +@[reducible] def foo : Result Unit := + foo_loop end infinite_loop diff --git a/tests/lean/Issue194RecursiveStructProjector.lean b/tests/lean/Issue194RecursiveStructProjector.lean index 4eb23934..730d1aa6 100644 --- a/tests/lean/Issue194RecursiveStructProjector.lean +++ b/tests/lean/Issue194RecursiveStructProjector.lean @@ -2,6 +2,9 @@ -- [issue_194_recursive_struct_projector] import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace issue_194_recursive_struct_projector diff --git a/tests/lean/Loops.lean b/tests/lean/Loops.lean index 199605d5..e676336e 100644 --- a/tests/lean/Loops.lean +++ b/tests/lean/Loops.lean @@ -2,6 +2,9 @@ -- [loops] import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace loops @@ -110,6 +113,7 @@ divergent def list_mem_loop (x : U32) (ls : List U32) : Result Bool := /- [loops::list_mem]: Source: 'tests/src/loops.rs', lines 80:0-80:52 -/ +@[reducible] def list_mem (x : U32) (ls : List U32) : Result Bool := list_mem_loop x ls @@ -136,6 +140,7 @@ divergent def list_nth_mut_loop_loop /- [loops::list_nth_mut_loop]: Source: 'tests/src/loops.rs', lines 92:0-92:71 -/ +@[reducible] def list_nth_mut_loop (T : Type) (ls : List T) (i : U32) : Result (T × (T → Result (List T))) := list_nth_mut_loop_loop T ls i @@ -155,6 +160,7 @@ divergent def list_nth_shared_loop_loop /- [loops::list_nth_shared_loop]: Source: 'tests/src/loops.rs', lines 105:0-105:66 -/ +@[reducible] def list_nth_shared_loop (T : Type) (ls : List T) (i : U32) : Result T := list_nth_shared_loop_loop T ls i @@ -316,6 +322,7 @@ divergent def list_nth_mut_loop_pair_loop /- [loops::list_nth_mut_loop_pair]: Source: 'tests/src/loops.rs', lines 188:0-192:27 -/ +@[reducible] def list_nth_mut_loop_pair (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T)) × (T → Result (List T))) @@ -340,6 +347,7 @@ divergent def list_nth_shared_loop_pair_loop /- [loops::list_nth_shared_loop_pair]: Source: 'tests/src/loops.rs', lines 212:0-216:19 -/ +@[reducible] def list_nth_shared_loop_pair (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result (T × T) := list_nth_shared_loop_pair_loop T ls0 ls1 i @@ -376,6 +384,7 @@ divergent def list_nth_mut_loop_pair_merge_loop /- [loops::list_nth_mut_loop_pair_merge]: Source: 'tests/src/loops.rs', lines 237:0-241:27 -/ +@[reducible] def list_nth_mut_loop_pair_merge (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × ((T × T) → Result ((List T) × (List T)))) @@ -401,6 +410,7 @@ divergent def list_nth_shared_loop_pair_merge_loop /- [loops::list_nth_shared_loop_pair_merge]: Source: 'tests/src/loops.rs', lines 255:0-259:19 -/ +@[reducible] def list_nth_shared_loop_pair_merge (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result (T × T) := list_nth_shared_loop_pair_merge_loop T ls0 ls1 i @@ -433,6 +443,7 @@ divergent def list_nth_mut_shared_loop_pair_loop /- [loops::list_nth_mut_shared_loop_pair]: Source: 'tests/src/loops.rs', lines 273:0-277:23 -/ +@[reducible] def list_nth_mut_shared_loop_pair (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -467,6 +478,7 @@ divergent def list_nth_mut_shared_loop_pair_merge_loop /- [loops::list_nth_mut_shared_loop_pair_merge]: Source: 'tests/src/loops.rs', lines 292:0-296:23 -/ +@[reducible] def list_nth_mut_shared_loop_pair_merge (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -501,6 +513,7 @@ divergent def list_nth_shared_mut_loop_pair_loop /- [loops::list_nth_shared_mut_loop_pair]: Source: 'tests/src/loops.rs', lines 311:0-315:23 -/ +@[reducible] def list_nth_shared_mut_loop_pair (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -535,6 +548,7 @@ divergent def list_nth_shared_mut_loop_pair_merge_loop /- [loops::list_nth_shared_mut_loop_pair_merge]: Source: 'tests/src/loops.rs', lines 330:0-334:23 -/ +@[reducible] def list_nth_shared_mut_loop_pair_merge (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) diff --git a/tests/lean/Matches.lean b/tests/lean/Matches.lean index 9233841b..8d592a85 100644 --- a/tests/lean/Matches.lean +++ b/tests/lean/Matches.lean @@ -2,6 +2,9 @@ -- [matches] import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace matches diff --git a/tests/lean/NoNestedBorrows.lean b/tests/lean/NoNestedBorrows.lean index f0438050..090ca2b2 100644 --- a/tests/lean/NoNestedBorrows.lean +++ b/tests/lean/NoNestedBorrows.lean @@ -2,6 +2,9 @@ -- [no_nested_borrows] import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace no_nested_borrows diff --git a/tests/lean/Paper.lean b/tests/lean/Paper.lean index 03b96903..b1aef703 100644 --- a/tests/lean/Paper.lean +++ b/tests/lean/Paper.lean @@ -2,6 +2,9 @@ -- [paper] import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace paper diff --git a/tests/lean/PoloniusList.lean b/tests/lean/PoloniusList.lean index ed279d58..489871ba 100644 --- a/tests/lean/PoloniusList.lean +++ b/tests/lean/PoloniusList.lean @@ -2,6 +2,9 @@ -- [polonius_list] import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace polonius_list diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean index 0dd692fe..51aba5bf 100644 --- a/tests/lean/Traits.lean +++ b/tests/lean/Traits.lean @@ -2,6 +2,9 @@ -- [traits] import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace traits diff --git a/tests/lean/lake-manifest.json b/tests/lean/lake-manifest.json index b1c3c9dc..94b2f080 100644 --- a/tests/lean/lake-manifest.json +++ b/tests/lean/lake-manifest.json @@ -4,7 +4,7 @@ [{"url": "https://github.com/leanprover-community/batteries", "type": "git", "subDir": null, - "rev": "f96a34401de084c73c787ecb45b85d4fb47bb981", + "rev": "47e4cc5c5800c07d9bf232173c9941fa5bf68589", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -22,7 +22,7 @@ {"url": "https://github.com/leanprover-community/aesop", "type": "git", "subDir": null, - "rev": "7e3bd939c6badfcb1e607c0fddb509548baafd05", + "rev": "882561b77bd2aaa98bd8665a56821062bdb3034c", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -49,7 +49,7 @@ {"url": "https://github.com/leanprover-community/import-graph.git", "type": "git", "subDir": null, - "rev": "7983e959f8f4a79313215720de3ef1eca2d6d474", + "rev": "1588be870b9c76fe62286e8f42f0b4dafa154c96", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -58,7 +58,7 @@ {"url": "https://github.com/leanprover-community/mathlib4.git", "type": "git", "subDir": null, - "rev": "c9f95f42b82684937a26ba395ebf7f25a81734ca", + "rev": "77e1ea0a339a4663eced9cacc3a46eb45f967b51", "name": "mathlib", "manifestFile": "lake-manifest.json", "inputRev": null, diff --git a/tests/lean/lean-toolchain b/tests/lean/lean-toolchain index 0ba3faf8..29c0cea4 100644 --- a/tests/lean/lean-toolchain +++ b/tests/lean/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.9.0-rc1 +leanprover/lean4:v4.9.0-rc2 diff --git a/tests/lean/misc/MutuallyRecursiveTraits.lean b/tests/lean/misc/MutuallyRecursiveTraits.lean index b0fcb9e9..05871009 100644 --- a/tests/lean/misc/MutuallyRecursiveTraits.lean +++ b/tests/lean/misc/MutuallyRecursiveTraits.lean @@ -2,5 +2,8 @@ -- [mutually_recursive_traits] import Base open Primitives +set_option linter.dupNamespace false +set_option linter.hashCommand false +set_option linter.unusedVariables false namespace mutually_recursive_traits diff --git a/tests/src/hashmap.rs b/tests/src/hashmap.rs index 7dda2404..9ff448db 100644 --- a/tests/src/hashmap.rs +++ b/tests/src/hashmap.rs @@ -26,8 +26,9 @@ use std::vec::Vec; pub type Key = usize; // TODO: make this generic pub type Hash = usize; -pub enum List<T> { - Cons(Key, T, Box<List<T>>), +/// Associative list +pub enum AList<T> { + Cons(Key, T, Box<AList<T>>), Nil, } @@ -51,15 +52,15 @@ pub struct HashMap<T> { /// gives the threshold at which to resize the table. max_load: usize, /// The table itself - slots: Vec<List<T>>, + slots: Vec<AList<T>>, } impl<T> HashMap<T> { /// Allocate a vector of slots of a given size. /// We would need a loop, but can't use loops for now... - fn allocate_slots(mut slots: Vec<List<T>>, mut n: usize) -> Vec<List<T>> { + fn allocate_slots(mut slots: Vec<AList<T>>, mut n: usize) -> Vec<AList<T>> { while n > 0 { - slots.push(List::Nil); + slots.push(AList::Nil); n -= 1; } slots @@ -92,7 +93,7 @@ impl<T> HashMap<T> { let slots = &mut self.slots; let mut i = 0; while i < slots.len() { - slots[i] = List::Nil; + slots[i] = AList::Nil; i += 1; } } @@ -104,14 +105,14 @@ impl<T> HashMap<T> { /// Insert in a list. /// Return `true` if we inserted an element, `false` if we simply updated /// a value. - fn insert_in_list(key: Key, value: T, mut ls: &mut List<T>) -> bool { + fn insert_in_list(key: Key, value: T, mut ls: &mut AList<T>) -> bool { loop { match ls { - List::Nil => { - *ls = List::Cons(key, value, Box::new(List::Nil)); + AList::Nil => { + *ls = AList::Cons(key, value, Box::new(AList::Nil)); return true; } - List::Cons(ckey, cvalue, tl) => { + AList::Cons(ckey, cvalue, tl) => { if *ckey == key { *cvalue = value; return false; @@ -178,10 +179,10 @@ impl<T> HashMap<T> { /// Auxiliary function called by [try_resize] to move all the elements /// from the table to a new table - fn move_elements<'a>(ntable: &'a mut HashMap<T>, slots: &'a mut Vec<List<T>>, mut i: usize) { + fn move_elements<'a>(ntable: &'a mut HashMap<T>, slots: &'a mut Vec<AList<T>>, mut i: usize) { while i < slots.len() { // Move the elements out of the slot i - let ls = std::mem::replace(&mut slots[i], List::Nil); + let ls = std::mem::replace(&mut slots[i], AList::Nil); // Move all those elements to the new table HashMap::move_elements_from_list(ntable, ls); // Do the same for slot i+1 @@ -190,12 +191,12 @@ impl<T> HashMap<T> { } /// Auxiliary function. - fn move_elements_from_list(ntable: &mut HashMap<T>, mut ls: List<T>) { + fn move_elements_from_list(ntable: &mut HashMap<T>, mut ls: AList<T>) { // As long as there are elements in the list, move them loop { match ls { - List::Nil => return, // We're done - List::Cons(k, v, tl) => { + AList::Nil => return, // We're done + AList::Cons(k, v, tl) => { // Insert the element in the new table ntable.insert_no_resize(k, v); // Move the elements out of the tail @@ -213,11 +214,11 @@ impl<T> HashMap<T> { } /// Returns `true` if the list contains a value for the specified key. - pub fn contains_key_in_list(key: &Key, mut ls: &List<T>) -> bool { + pub fn contains_key_in_list(key: &Key, mut ls: &AList<T>) -> bool { loop { match ls { - List::Nil => return false, - List::Cons(ckey, _, tl) => { + AList::Nil => return false, + AList::Cons(ckey, _, tl) => { if *ckey == *key { return true; } else { @@ -231,11 +232,11 @@ impl<T> HashMap<T> { /// We don't support borrows inside of enumerations for now, so we /// can't return an option... /// TODO: add support for that - fn get_in_list<'a, 'k>(key: &'k Key, mut ls: &'a List<T>) -> &'a T { + fn get_in_list<'a, 'k>(key: &'k Key, mut ls: &'a AList<T>) -> &'a T { loop { match ls { - List::Nil => panic!(), - List::Cons(ckey, cvalue, tl) => { + AList::Nil => panic!(), + AList::Cons(ckey, cvalue, tl) => { if *ckey == *key { return cvalue; } else { @@ -252,8 +253,8 @@ impl<T> HashMap<T> { HashMap::get_in_list(key, &self.slots[hash_mod]) } - pub fn get_mut_in_list<'a, 'k>(mut ls: &'a mut List<T>, key: &'k Key) -> &'a mut T { - while let List::Cons(ckey, cvalue, tl) = ls { + pub fn get_mut_in_list<'a, 'k>(mut ls: &'a mut AList<T>, key: &'k Key) -> &'a mut T { + while let AList::Cons(ckey, cvalue, tl) = ls { if *ckey == *key { return cvalue; } else { @@ -272,20 +273,20 @@ impl<T> HashMap<T> { /// Remove an element from the list. /// Return the removed element. - fn remove_from_list(key: &Key, mut ls: &mut List<T>) -> Option<T> { + fn remove_from_list(key: &Key, mut ls: &mut AList<T>) -> Option<T> { loop { match ls { - List::Nil => return None, + AList::Nil => return None, // We have to use a guard and split the Cons case into two // branches, otherwise the borrow checker is not happy. - List::Cons(ckey, _, _) if *ckey == *key => { + AList::Cons(ckey, _, _) if *ckey == *key => { // We have to move under borrows, so we need to use // [std::mem::replace] in several steps. // Retrieve the tail - let mv_ls = std::mem::replace(ls, List::Nil); + let mv_ls = std::mem::replace(ls, AList::Nil); match mv_ls { - List::Nil => unreachable!(), - List::Cons(_, cvalue, tl) => { + AList::Nil => unreachable!(), + AList::Cons(_, cvalue, tl) => { // Make the list equal to its tail *ls = *tl; // Return the dropped value @@ -293,7 +294,7 @@ impl<T> HashMap<T> { } } } - List::Cons(_, _, tl) => { + AList::Cons(_, _, tl) => { ls = tl; } } diff --git a/tests/src/mutually-recursive-traits.lean.out b/tests/src/mutually-recursive-traits.lean.out index e2ed3abc..9b3b0737 100644 --- a/tests/src/mutually-recursive-traits.lean.out +++ b/tests/src/mutually-recursive-traits.lean.out @@ -12,6 +12,6 @@ Uncaught exception: Raised at Aeneas__Errors.craise_opt_span in file "Errors.ml", line 46, characters 4-76 Called from Stdlib__List.iter in file "list.ml", line 110, characters 12-15 Called from Aeneas__Translate.extract_definitions in file "Translate.ml", line 836, characters 2-52 -Called from Aeneas__Translate.extract_file in file "Translate.ml", line 954, characters 2-36 -Called from Aeneas__Translate.translate_crate in file "Translate.ml", line 1503, characters 5-42 +Called from Aeneas__Translate.extract_file in file "Translate.ml", line 963, characters 2-36 +Called from Aeneas__Translate.translate_crate in file "Translate.ml", line 1512, characters 5-42 Called from Dune__exe__Main in file "Main.ml", line 314, characters 14-66 |