From 71834958a958523a4881d822e729af1ddd78c9df Mon Sep 17 00:00:00 2001 From: Aymeric Fromherz Date: Thu, 23 May 2024 10:56:28 +0200 Subject: Add printing of projectors for recursive structs in Lean backend --- compiler/ExtractTypes.ml | 133 +++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 130 insertions(+), 3 deletions(-) diff --git a/compiler/ExtractTypes.ml b/compiler/ExtractTypes.ml index 6a6067de..4272e7a9 100644 --- a/compiler/ExtractTypes.ml +++ b/compiler/ExtractTypes.ml @@ -1662,6 +1662,127 @@ let extract_type_decl_coq_arguments (ctx : extraction_ctx) (fmt : F.formatter) (* Add breaks to insert new lines between definitions *) F.pp_print_break fmt 0 0 +(** Auxiliary function. + + Generate field projectors in Lean. + + Recursive structs are defined as inductives in Lean. + Field projectors allow to retrieve the facilities provided by + Lean structures. + *) +let extract_type_decl_lean_record_field_projectors (ctx : extraction_ctx) + (fmt : F.formatter) (kind : decl_kind) (decl : type_decl) : unit = + sanity_check __FILE__ __LINE__ (!backend = Lean) decl.meta; + match decl.kind with + | Opaque | Enum _ -> () + | Struct fields -> + (* Records are extracted as inductives only if they are recursive *) + let is_rec = decl_is_from_rec_group kind in + if is_rec then + (* Add the type params *) + let ctx, type_params, cg_params, trait_clauses = + ctx_add_generic_params decl.meta decl.llbc_name decl.llbc_generics + decl.generics ctx + in + (* Record_var will be the ADT argument to the projector *) + let ctx, record_var = ctx_add_var decl.meta "x" (VarId.of_int 0) ctx in + (* Field_var will be the variable in the constructor that is returned by the projector *) + let ctx, field_var = ctx_add_var decl.meta "x" (VarId.of_int 1) ctx in + (* Name of the ADT *) + let def_name = ctx_get_local_type decl.meta decl.def_id ctx in + (* Name of the ADT constructor. As we are in the struct case, we only have + one constructor *) + let cons_name = ctx_get_struct decl.meta (TAdtId decl.def_id) ctx in + + let extract_field_proj (field_id : FieldId.id) (_ : field) : unit = + F.pp_print_space fmt (); + (* Box for the projector definition *) + F.pp_open_hvbox fmt 0; + (* Box for the [def ADT.proj ... :=] *) + F.pp_open_hovbox fmt ctx.indent_incr; + F.pp_print_string fmt "def"; + F.pp_print_space fmt (); + + (* Print the function name. In Lean, the syntax ADT.proj will + allow us to call x.proj for any x of type ADT *) + let field_name = + ctx_get_field decl.meta (TAdtId decl.def_id) field_id ctx + in + F.pp_print_string fmt def_name; + F.pp_print_string fmt "."; + F.pp_print_string fmt field_name; + + (* Print the generics *) + let as_implicits = true in + extract_generic_params decl.meta ctx fmt TypeDeclId.Set.empty + ~as_implicits decl.generics type_params cg_params trait_clauses; + + (* Print the record parameter as "(x : ADT)" *) + F.pp_print_space fmt (); + F.pp_print_string fmt "("; + F.pp_print_string fmt record_var; + F.pp_print_space fmt (); + F.pp_print_string fmt ":"; + F.pp_print_space fmt (); + F.pp_print_string fmt def_name; + List.iter + (fun p -> + F.pp_print_space fmt (); + F.pp_print_string fmt p) + type_params; + F.pp_print_string fmt ")"; + + F.pp_print_space fmt (); + F.pp_print_string fmt ":="; + + (* Close the box for the [def ADT.proj ... :=] *) + F.pp_close_box fmt (); + F.pp_print_space fmt (); + + (* Open a box for the whole match *) + F.pp_open_hvbox fmt 0; + + (* Open a box for the [match ... with] *) + F.pp_open_hovbox fmt ctx.indent_incr; + F.pp_print_string fmt "match"; + F.pp_print_space fmt (); + F.pp_print_string fmt record_var; + F.pp_print_space fmt (); + F.pp_print_string fmt "with"; + (* Close the box for the [match ... with] *) + F.pp_close_box fmt (); + + (* Open a box for the branch *) + F.pp_open_hovbox fmt ctx.indent_incr; + (* Print the match branch *) + F.pp_print_space fmt (); + F.pp_print_string fmt "|"; + F.pp_print_space fmt (); + F.pp_print_string fmt cons_name; + FieldId.iteri + (fun id _ -> + F.pp_print_space fmt (); + if field_id = id then F.pp_print_string fmt field_var + else F.pp_print_string fmt "_") + fields; + F.pp_print_space fmt (); + F.pp_print_string fmt "=>"; + F.pp_print_space fmt (); + F.pp_print_string fmt field_var; + (* Close the box for the branch *) + F.pp_close_box fmt (); + + (* Close the box for the whole match *) + F.pp_close_box fmt (); + + (* Close the box for projector definition *) + F.pp_close_box fmt (); + (* Add breaks to insert new lines between definitions *) + F.pp_print_break fmt 0 0 + in + + FieldId.iteri extract_field_proj fields + (** Auxiliary function. Generate field projectors in Coq. @@ -1669,7 +1790,7 @@ let extract_type_decl_coq_arguments (ctx : extraction_ctx) (fmt : F.formatter) Sometimes we extract records as inductives in Coq: when this happens we have to define the field projectors afterwards. *) -let extract_type_decl_record_field_projectors (ctx : extraction_ctx) +let extract_type_decl_coq_record_field_projectors (ctx : extraction_ctx) (fmt : F.formatter) (kind : decl_kind) (decl : type_decl) : unit = sanity_check __FILE__ __LINE__ (!backend = Coq) decl.meta; match decl.kind with @@ -1826,7 +1947,13 @@ let extract_type_decl_record_field_projectors (ctx : extraction_ctx) let extract_type_decl_extra_info (ctx : extraction_ctx) (fmt : F.formatter) (kind : decl_kind) (decl : type_decl) : unit = match !backend with - | FStar | Lean | HOL4 -> () + | FStar | HOL4 -> () + | Lean -> + if + not + (TypesUtils.type_decl_from_decl_id_is_tuple_struct + ctx.trans_ctx.type_ctx.type_infos decl.def_id) + then extract_type_decl_lean_record_field_projectors ctx fmt kind decl | Coq -> if not @@ -1834,7 +1961,7 @@ let extract_type_decl_extra_info (ctx : extraction_ctx) (fmt : F.formatter) ctx.trans_ctx.type_ctx.type_infos decl.def_id) then ( extract_type_decl_coq_arguments ctx fmt kind decl; - extract_type_decl_record_field_projectors ctx fmt kind decl) + extract_type_decl_coq_record_field_projectors ctx fmt kind decl) (** Extract the state type declaration. *) let extract_state_type (fmt : F.formatter) (ctx : extraction_ctx) -- cgit v1.2.3 From e660a4d575ba55bd23c64b5bd7b4e435be82d7ed Mon Sep 17 00:00:00 2001 From: Aymeric Fromherz Date: Thu, 23 May 2024 10:58:47 +0200 Subject: Improve formatting of Lean struct projectors --- compiler/ExtractTypes.ml | 8 ++++++-- 1 file changed, 6 insertions(+), 2 deletions(-) diff --git a/compiler/ExtractTypes.ml b/compiler/ExtractTypes.ml index 4272e7a9..8520b0fb 100644 --- a/compiler/ExtractTypes.ml +++ b/compiler/ExtractTypes.ml @@ -1696,8 +1696,11 @@ let extract_type_decl_lean_record_field_projectors (ctx : extraction_ctx) let extract_field_proj (field_id : FieldId.id) (_ : field) : unit = F.pp_print_space fmt (); - (* Box for the projector definition *) + (* Outer box for the projector definition *) F.pp_open_hvbox fmt 0; + (* Inner box for the projector definition *) + F.pp_open_hvbox fmt ctx.indent_incr; + (* Box for the [def ADT.proj ... :=] *) F.pp_open_hovbox fmt ctx.indent_incr; F.pp_print_string fmt "def"; @@ -1775,7 +1778,8 @@ let extract_type_decl_lean_record_field_projectors (ctx : extraction_ctx) (* Close the box for the whole match *) F.pp_close_box fmt (); - (* Close the box for projector definition *) + F.pp_close_box fmt (); + (* Close the outer box for projector definition *) F.pp_close_box fmt (); (* Add breaks to insert new lines between definitions *) F.pp_print_break fmt 0 0 -- cgit v1.2.3 From 7238177f5c7cff15f924e13c01d5b3b4802daf77 Mon Sep 17 00:00:00 2001 From: Aymeric Fromherz Date: Thu, 23 May 2024 11:33:46 +0200 Subject: Do not expand field projector for recursive structs to a let binding in Lean --- compiler/SymbolicToPure.ml | 5 ----- 1 file changed, 5 deletions(-) diff --git a/compiler/SymbolicToPure.ml b/compiler/SymbolicToPure.ml index 6c925bcd..351f5cf2 100644 --- a/compiler/SymbolicToPure.ml +++ b/compiler/SymbolicToPure.ml @@ -2902,14 +2902,9 @@ and translate_ExpandAdt_one_branch (sv : V.symbolic_value) - if the ADT is an enumeration (which must have exactly one branch) - if we forbid using field projectors. *) - let is_rec_def = - T.TypeDeclId.Set.mem adt_id ctx.type_ctx.recursive_decls - in let use_let_with_cons = is_enum || !Config.dont_use_field_projectors - (* TODO: for now, we don't have field projectors over recursive ADTs in Lean. *) - || (!Config.backend = Lean && is_rec_def) (* Also, there is a special case when the ADT is to be extracted as a tuple, because it is a structure with unnamed fields. Some backends like Lean have projectors for tuples (like so: `x.3`), but others -- cgit v1.2.3 From b4cc9c908935eb95c761b664e75ce065fb97d85c Mon Sep 17 00:00:00 2001 From: Aymeric Fromherz Date: Thu, 23 May 2024 12:04:09 +0200 Subject: Regenerate Lean files for betree following extraction changes --- tests/lean/BetreeMain/Funs.lean | 81 +++++++++++++++++----------------------- tests/lean/BetreeMain/Types.lean | 12 ++++++ 2 files changed, 46 insertions(+), 47 deletions(-) diff --git a/tests/lean/BetreeMain/Funs.lean b/tests/lean/BetreeMain/Funs.lean index 7cc52159..f6fda6db 100644 --- a/tests/lean/BetreeMain/Funs.lean +++ b/tests/lean/BetreeMain/Funs.lean @@ -250,16 +250,15 @@ mutual divergent def betree.Internal.lookup_in_children (self : betree.Internal) (key : U64) (st : State) : Result (State × ((Option U64) × betree.Internal)) := - let ⟨ i, i1, n, n1 ⟩ := self - if key < i1 + if key < self.pivot then do - let (st1, (o, n2)) ← betree.Node.lookup n key st - Result.ok (st1, (o, betree.Internal.mk i i1 n2 n1)) + let (st1, (o, n)) ← betree.Node.lookup self.left key st + Result.ok (st1, (o, betree.Internal.mk self.id self.pivot n self.right)) else do - let (st1, (o, n2)) ← betree.Node.lookup n1 key st - Result.ok (st1, (o, betree.Internal.mk i i1 n n2)) + let (st1, (o, n)) ← betree.Node.lookup self.right key st + Result.ok (st1, (o, betree.Internal.mk self.id self.pivot self.left n)) /- [betree_main::betree::{betree_main::betree::Node#5}::lookup]: Source: 'src/betree.rs', lines 709:4-709:58 -/ @@ -270,8 +269,7 @@ divergent def betree.Node.lookup match self with | betree.Node.Internal node => do - let ⟨ i, i1, n, n1 ⟩ := node - let (st1, msgs) ← betree.load_internal_node i st + let (st1, msgs) ← betree.load_internal_node node.id st let (pending, lookup_first_message_for_key_back) ← betree.Node.lookup_first_message_for_key key msgs match pending with @@ -281,8 +279,7 @@ divergent def betree.Node.lookup then do let (st2, (o, node1)) ← - betree.Internal.lookup_in_children (betree.Internal.mk i i1 n n1) key - st1 + betree.Internal.lookup_in_children node key st1 let _ ← lookup_first_message_for_key_back (betree.List.Cons (k, msg) l) Result.ok (st2, (o, betree.Node.Internal node1)) @@ -293,33 +290,26 @@ divergent def betree.Node.lookup let _ ← lookup_first_message_for_key_back (betree.List.Cons (k, betree.Message.Insert v) l) - Result.ok (st1, (some v, betree.Node.Internal (betree.Internal.mk i - i1 n n1))) + Result.ok (st1, (some v, betree.Node.Internal node)) | betree.Message.Delete => do let _ ← lookup_first_message_for_key_back (betree.List.Cons (k, betree.Message.Delete) l) - Result.ok (st1, (none, betree.Node.Internal (betree.Internal.mk i i1 - n n1))) + Result.ok (st1, (none, betree.Node.Internal node)) | betree.Message.Upsert ufs => do let (st2, (v, node1)) ← - betree.Internal.lookup_in_children (betree.Internal.mk i i1 n n1) - key st1 + betree.Internal.lookup_in_children node key st1 let (v1, pending1) ← betree.Node.apply_upserts (betree.List.Cons (k, betree.Message.Upsert ufs) l) v key - let ⟨ i2, i3, n2, n3 ⟩ := node1 let msgs1 ← lookup_first_message_for_key_back pending1 - let (st3, _) ← betree.store_internal_node i2 msgs1 st2 - Result.ok (st3, (some v1, betree.Node.Internal (betree.Internal.mk i2 - i3 n2 n3))) + let (st3, _) ← betree.store_internal_node node1.id msgs1 st2 + Result.ok (st3, (some v1, betree.Node.Internal node1)) | betree.List.Nil => do - let (st2, (o, node1)) ← - betree.Internal.lookup_in_children (betree.Internal.mk i i1 n n1) key - st1 + let (st2, (o, node1)) ← betree.Internal.lookup_in_children node key st1 let _ ← lookup_first_message_for_key_back betree.List.Nil Result.ok (st2, (o, betree.Node.Internal node1)) | betree.Node.Leaf node => @@ -541,34 +531,36 @@ mutual divergent def betree.Internal.flush × betree.NodeIdCounter))) := do - let ⟨ i, i1, n, n1 ⟩ := self - let p ← betree.ListPairU64T.partition_at_pivot betree.Message content i1 + let p ← + betree.ListPairU64T.partition_at_pivot betree.Message content self.pivot let (msgs_left, msgs_right) := p let len_left ← betree.List.len (U64 × betree.Message) msgs_left if len_left >= params.min_flush_size then do let (st1, p1) ← - betree.Node.apply_messages n params node_id_cnt msgs_left st - let (n2, node_id_cnt1) := p1 + betree.Node.apply_messages self.left params node_id_cnt msgs_left st + let (n, node_id_cnt1) := p1 let len_right ← betree.List.len (U64 × betree.Message) msgs_right if len_right >= params.min_flush_size then do let (st2, p2) ← - betree.Node.apply_messages n1 params node_id_cnt1 msgs_right st1 - let (n3, node_id_cnt2) := p2 - Result.ok (st2, (betree.List.Nil, (betree.Internal.mk i i1 n2 n3, - node_id_cnt2))) + betree.Node.apply_messages self.right params node_id_cnt1 msgs_right + st1 + let (n1, node_id_cnt2) := p2 + Result.ok (st2, (betree.List.Nil, (betree.Internal.mk self.id self.pivot + n n1, node_id_cnt2))) else - Result.ok (st1, (msgs_right, (betree.Internal.mk i i1 n2 n1, - node_id_cnt1))) + Result.ok (st1, (msgs_right, (betree.Internal.mk self.id self.pivot n + self.right, node_id_cnt1))) else do let (st1, p1) ← - betree.Node.apply_messages n1 params node_id_cnt msgs_right st - let (n2, node_id_cnt1) := p1 - Result.ok (st1, (msgs_left, (betree.Internal.mk i i1 n n2, node_id_cnt1))) + betree.Node.apply_messages self.right params node_id_cnt msgs_right st + let (n, node_id_cnt1) := p1 + Result.ok (st1, (msgs_left, (betree.Internal.mk self.id self.pivot + self.left n, node_id_cnt1))) /- [betree_main::betree::{betree_main::betree::Node#5}::apply_messages]: Source: 'src/betree.rs', lines 588:4-593:5 -/ @@ -581,26 +573,21 @@ divergent def betree.Node.apply_messages match self with | betree.Node.Internal node => do - let ⟨ i, i1, n, n1 ⟩ := node - let (st1, content) ← betree.load_internal_node i st + let (st1, content) ← betree.load_internal_node node.id st let content1 ← betree.Node.apply_messages_to_internal content msgs let num_msgs ← betree.List.len (U64 × betree.Message) content1 if num_msgs >= params.min_flush_size then do let (st2, (content2, p)) ← - betree.Internal.flush (betree.Internal.mk i i1 n n1) params node_id_cnt - content1 st1 + betree.Internal.flush node params node_id_cnt content1 st1 let (node1, node_id_cnt1) := p - let ⟨ i2, i3, n2, n3 ⟩ := node1 - let (st3, _) ← betree.store_internal_node i2 content2 st2 - Result.ok (st3, (betree.Node.Internal (betree.Internal.mk i2 i3 n2 n3), - node_id_cnt1)) + let (st3, _) ← betree.store_internal_node node1.id content2 st2 + Result.ok (st3, (betree.Node.Internal node1, node_id_cnt1)) else do - let (st2, _) ← betree.store_internal_node i content1 st1 - Result.ok (st2, (betree.Node.Internal (betree.Internal.mk i i1 n n1), - node_id_cnt)) + let (st2, _) ← betree.store_internal_node node.id content1 st1 + Result.ok (st2, (betree.Node.Internal node, node_id_cnt)) | betree.Node.Leaf node => do let (st1, content) ← betree.load_leaf_node node.id st diff --git a/tests/lean/BetreeMain/Types.lean b/tests/lean/BetreeMain/Types.lean index 877508f6..3f115fe4 100644 --- a/tests/lean/BetreeMain/Types.lean +++ b/tests/lean/BetreeMain/Types.lean @@ -46,6 +46,18 @@ inductive betree.Node := end +def betree.Internal.id (x : betree.Internal) := + match x with | betree.Internal.mk x1 _ _ _ => x1 + +def betree.Internal.pivot (x : betree.Internal) := + match x with | betree.Internal.mk _ x1 _ _ => x1 + +def betree.Internal.left (x : betree.Internal) := + match x with | betree.Internal.mk _ _ x1 _ => x1 + +def betree.Internal.right (x : betree.Internal) := + match x with | betree.Internal.mk _ _ _ x1 => x1 + /- [betree_main::betree::Params] Source: 'src/betree.rs', lines 187:0-187:13 -/ structure betree.Params where -- cgit v1.2.3 From e76ef95ddda11d4220e1f0fd863b0df568de95bc Mon Sep 17 00:00:00 2001 From: Aymeric Fromherz Date: Thu, 23 May 2024 16:41:54 +0200 Subject: Add simp, reducible attributes to generated Lean projectors --- compiler/ExtractTypes.ml | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/compiler/ExtractTypes.ml b/compiler/ExtractTypes.ml index 8520b0fb..70a4d000 100644 --- a/compiler/ExtractTypes.ml +++ b/compiler/ExtractTypes.ml @@ -1701,6 +1701,16 @@ let extract_type_decl_lean_record_field_projectors (ctx : extraction_ctx) (* Inner box for the projector definition *) F.pp_open_hvbox fmt ctx.indent_incr; + (* Box for the attributes *) + F.pp_open_vbox fmt 0; + (* Annotate the projectors with both simp and reducible. + The first one allows to automatically unfold when calling simp in proofs. + The second ensures that projectors will interact well with the unifier *) + F.pp_print_string fmt "@[simp, reducible]"; + F.pp_print_break fmt 0 0; + (* Close box for the attributes *) + F.pp_close_box fmt (); + (* Box for the [def ADT.proj ... :=] *) F.pp_open_hovbox fmt ctx.indent_incr; F.pp_print_string fmt "def"; -- cgit v1.2.3 From 765cb792916c1c69f864a6cf59a49c504ad603a2 Mon Sep 17 00:00:00 2001 From: Aymeric Fromherz Date: Thu, 23 May 2024 16:46:18 +0200 Subject: Regenerate Lean files for betree --- tests/lean/BetreeMain/Types.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/tests/lean/BetreeMain/Types.lean b/tests/lean/BetreeMain/Types.lean index 3f115fe4..e79da43f 100644 --- a/tests/lean/BetreeMain/Types.lean +++ b/tests/lean/BetreeMain/Types.lean @@ -46,15 +46,19 @@ inductive betree.Node := end +@[simp, reducible] def betree.Internal.id (x : betree.Internal) := match x with | betree.Internal.mk x1 _ _ _ => x1 +@[simp, reducible] def betree.Internal.pivot (x : betree.Internal) := match x with | betree.Internal.mk _ x1 _ _ => x1 +@[simp, reducible] def betree.Internal.left (x : betree.Internal) := match x with | betree.Internal.mk _ _ x1 _ => x1 +@[simp, reducible] def betree.Internal.right (x : betree.Internal) := match x with | betree.Internal.mk _ _ _ x1 => x1 -- cgit v1.2.3 From d9d0b7cc27abecfefcb22b46333ecdeb9ec397fa Mon Sep 17 00:00:00 2001 From: Son Ho Date: Fri, 24 May 2024 15:03:37 +0200 Subject: Downgrade the version of dune --- tests/test_runner/dune-project | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/tests/test_runner/dune-project b/tests/test_runner/dune-project index c614e923..dc352bd0 100644 --- a/tests/test_runner/dune-project +++ b/tests/test_runner/dune-project @@ -1,4 +1,4 @@ -(lang dune 3.12) +(lang dune 3.7) (name aeneas_test_runner) -- cgit v1.2.3 From 3ff6d93822fe5b2e233d4b12b88b38839c8533c5 Mon Sep 17 00:00:00 2001 From: Ryan Lahfa Date: Wed, 15 May 2024 15:42:18 +0200 Subject: feat: add small pieces of max theory `0#ty` is neutral for `max` for unsigned integers. Without the `Fact` instances, those theorems are not as automatic as they could be. Signed-off-by: Ryan Lahfa --- backends/lean/Base/Primitives/Scalar.lean | 39 +++++++++++++++++++++++++++++++ 1 file changed, 39 insertions(+) diff --git a/backends/lean/Base/Primitives/Scalar.lean b/backends/lean/Base/Primitives/Scalar.lean index f08b8dec..8fb067e1 100644 --- a/backends/lean/Base/Primitives/Scalar.lean +++ b/backends/lean/Base/Primitives/Scalar.lean @@ -146,6 +146,25 @@ def ScalarTy.isSigned (ty : ScalarTy) : Bool := | U64 | U128 => false +-- FIXME(chore): bulk prove them via macro? +instance : Fact (¬ ScalarTy.isSigned .Usize) where + out := by decide + +instance : Fact (¬ ScalarTy.isSigned .U8) where + out := by decide + +instance : Fact (¬ ScalarTy.isSigned .U16) where + out := by decide + +instance : Fact (¬ ScalarTy.isSigned .U32) where + out := by decide + +instance : Fact (¬ ScalarTy.isSigned .U64) where + out := by decide + +instance : Fact (¬ ScalarTy.isSigned .U128) where + out := by decide + def Scalar.smin (ty : ScalarTy) : Int := match ty with @@ -386,6 +405,10 @@ theorem Scalar.tryMk_eq (ty : ScalarTy) (x : Int) : simp [tryMk, ofOption, tryMkOpt] split_ifs <;> simp +instance (ty: ScalarTy) : InBounds ty 0 where + hInBounds := by + induction ty <;> simp [Scalar.cMax, Scalar.cMin, Scalar.max, Scalar.min] <;> decide + def Scalar.neg {ty : ScalarTy} (x : Scalar ty) : Result (Scalar ty) := Scalar.tryMk ty (- x.val) -- Our custom remainder operation, which satisfies the semantics of Rust @@ -1431,6 +1454,22 @@ theorem coe_max {ty: ScalarTy} (a b: Scalar ty): ↑(Max.max a b) = (Max.max ( refine' absurd _ (lt_irrefl a) exact lt_of_le_of_lt (by assumption) ((Scalar.lt_equiv _ _).2 (by assumption)) +-- Max theory +-- TODO: do the min theory later on. + +theorem Scalar.zero_le_unsigned {ty} (s: ¬ ty.isSigned) (x: Scalar ty): Scalar.ofInt 0 ≤ x := by + apply (Scalar.le_equiv _ _).2 + convert x.hmin + cases ty <;> simp [ScalarTy.isSigned] at s <;> simp [Scalar.min] + +@[simp] +theorem Scalar.max_unsigned_left_zero_eq {ty} [s: Fact (¬ ty.isSigned)] (x: Scalar ty): + Max.max (Scalar.ofInt 0) x = x := max_eq_right (Scalar.zero_le_unsigned s.out x) + +@[simp] +theorem Scalar.max_unsigned_right_zero_eq {ty} [s: Fact (¬ ty.isSigned)] (x: Scalar ty): + Max.max x (Scalar.ofInt 0) = x := max_eq_left (Scalar.zero_le_unsigned s.out x) + -- Leading zeros def core.num.Usize.leading_zeros (x : Usize) : U32 := sorry def core.num.U8.leading_zeros (x : U8) : U32 := sorry -- cgit v1.2.3 From 7935e74a9cedd93e885ab546d5513ea6c31db5ad Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Thu, 23 May 2024 17:01:03 +0200 Subject: runner: Strongly typed Backend enum --- flake.nix | 1 + tests/test_runner/dune | 6 ++ tests/test_runner/run_test.ml | 144 ++++++++++++++++++++++++------------------ 3 files changed, 89 insertions(+), 62 deletions(-) diff --git a/flake.nix b/flake.nix index caa29fae..68378954 100644 --- a/flake.nix +++ b/flake.nix @@ -80,6 +80,7 @@ OCAMLPARAM = "_,warn-error=+A"; # Turn all warnings into errors. propagatedBuildInputs = (with ocamlPackages; [ core_unix + ppx_deriving ]); }; diff --git a/tests/test_runner/dune b/tests/test_runner/dune index 7caf661f..65b0c5fe 100644 --- a/tests/test_runner/dune +++ b/tests/test_runner/dune @@ -1,4 +1,10 @@ (executable (public_name test_runner) (libraries core_unix.sys_unix unix) + (preprocess + (pps ppx_deriving.show ppx_deriving.ord ppx_sexp_conv)) (name run_test)) + +(env + (dev + (flags :standard -warn-error -5@8-11-14-32-33-20-21-26-27-39))) diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index 8aa6347c..2c411c95 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -5,22 +5,38 @@ type runner_env = { llbc_dir : string; } -(* The data for a specific test to run aeneas on *) -type aeneas_test_case = { - name : string; - backend : string; - subdir : string; - extra_aeneas_options : string list; -} +module Backend = struct + type t = Coq | Lean | FStar | HOL4 [@@deriving ord, sexp] + + (* TODO: reactivate HOL4 once traits are parameterized by their associated types *) + let all = [ Coq; Lean; FStar ] + + let of_string = function + | "coq" -> Coq + | "lean" -> Lean + | "fstar" -> FStar + | "hol4" -> HOL4 + | backend -> failwith ("Unknown backend: `" ^ backend ^ "`") + + let to_string = function + | Coq -> "coq" + | Lean -> "lean" + | FStar -> "fstar" + | HOL4 -> "hol4" +end + +module BackendMap = Map.Make (Backend) type input_kind = SingleFile | Crate -(* The data for a specific test to generate llbc for *) -type charon_test_case = { - kind : input_kind; +(* The data for a specific test input *) +type test_case = { name : string; path : string; - extra_charon_options : string list; + input_kind : input_kind; + charon_options : string list; + aeneas_options : string list BackendMap.t; + subdir : string BackendMap.t; } let concat_path = List.fold_left Filename.concat "" @@ -43,21 +59,26 @@ let run_command args = () (* Run Aeneas on a specific input with the given options *) -let run_aeneas (env : runner_env) (case : aeneas_test_case) = - let input_file = concat_path [ env.llbc_dir; case.name ] ^ ".llbc" in - let dest_dir = concat_path [ "tests"; case.backend; case.subdir ] in +let run_aeneas (env : runner_env) (case : test_case) (backend : Backend.t) = + (* FIXME: remove this special case *) + let test_name = if case.name = "betree" then "betree_main" else case.name in + let input_file = concat_path [ env.llbc_dir; test_name ] ^ ".llbc" in + let subdir = BackendMap.find backend case.subdir in + let aeneas_options = BackendMap.find backend case.aeneas_options in + let backend_str = Backend.to_string backend in + let dest_dir = concat_path [ "tests"; backend_str; subdir ] in let args = [| - env.aeneas_path; input_file; "-dest"; dest_dir; "-backend"; case.backend; + env.aeneas_path; input_file; "-dest"; dest_dir; "-backend"; backend_str; |] in - let args = Array.append args (Array.of_list case.extra_aeneas_options) in + let args = Array.append args (Array.of_list aeneas_options) in (* Run Aeneas *) run_command args (* Run Charon on a specific input with the given options *) -let run_charon (env : runner_env) (case : charon_test_case) = - match case.kind with +let run_charon (env : runner_env) (case : test_case) = + match case.input_kind with | SingleFile -> let args = [| @@ -71,16 +92,14 @@ let run_charon (env : runner_env) (case : charon_test_case) = env.llbc_dir; |] in - let args = Array.append args (Array.of_list case.extra_charon_options) in + let args = Array.append args (Array.of_list case.charon_options) in (* Run Charon on the rust file *) run_command args | Crate -> ( match Sys.getenv_opt "IN_CI" with | None -> let args = [| env.charon_path; "--dest"; env.llbc_dir |] in - let args = - Array.append args (Array.of_list case.extra_charon_options) - in + let args = Array.append args (Array.of_list case.charon_options) in (* Run Charon inside the crate *) let old_pwd = Unix.getcwd () in Unix.chdir case.path; @@ -96,12 +115,13 @@ let run_charon (env : runner_env) (case : charon_test_case) = (* File-specific options *) let aeneas_options_for_test backend test_name = + let backend = Backend.to_string backend in (* TODO: reactivate -test-trans-units for hashmap and hashmap_main *) let use_fuel = match (backend, test_name) with | ( "coq", - ( "arrays" | "betree_main" | "demo" | "hashmap" | "hashmap_main" - | "loops" ) ) -> + ("arrays" | "betree" | "demo" | "hashmap" | "hashmap_main" | "loops") ) + -> true | "fstar", "demo" -> true | _ -> false @@ -112,8 +132,7 @@ let aeneas_options_for_test backend test_name = backend = "fstar" && match test_name with - | "arrays" | "betree_main" | "hashmap" | "hashmap_main" | "loops" | "traits" - -> + | "arrays" | "betree" | "hashmap" | "hashmap_main" | "loops" | "traits" -> true | _ -> false in @@ -125,7 +144,7 @@ let aeneas_options_for_test backend test_name = let extra_options = match (backend, test_name) with - | _, "betree_main" -> + | _, "betree" -> [ "-backward-no-state-update"; "-test-trans-units"; @@ -172,11 +191,12 @@ let charon_options_for_test test_name = (* The subdirectory in which to store the outputs. *) (* This reproduces the file layout that was set by the old Makefile. FIXME: cleanup *) let test_subdir backend test_name = + let backend = Backend.to_string backend in match (backend, test_name) with | "lean", "demo" -> "Demo" | "lean", _ -> "." | _, ("arrays" | "demo" | "hashmap" | "traits") -> test_name - | _, "betree_main" -> "betree" + | _, "betree" -> "betree" | _, "hashmap_main" -> "hashmap_on_disk" | "hol4", _ -> "misc-" ^ test_name | ( _, @@ -185,49 +205,49 @@ let test_subdir backend test_name = "misc" | _ -> test_name +let build_test (path : string) : test_case = + let name = Filename.remove_extension (Filename.basename path) in + let input_kind = + if Sys_unix.is_file_exn path then SingleFile + else if Sys_unix.is_directory_exn path then Crate + else failwith ("`" ^ path ^ "` is not a file or a directory.") + in + let charon_options = charon_options_for_test name in + let subdir = + List.fold_left + (fun map backend -> + let subdir = test_subdir backend name in + BackendMap.add backend subdir map) + BackendMap.empty Backend.all + in + let aeneas_options = + List.fold_left + (fun map backend -> + let opts = aeneas_options_for_test backend name in + BackendMap.add backend opts map) + BackendMap.empty Backend.all + in + { path; name; input_kind; charon_options; subdir; aeneas_options } + let () = match Array.to_list Sys.argv with (* Ad-hoc argument passing for now. *) | _exe_path :: charon_path :: aeneas_path :: llbc_dir :: test_path :: aeneas_options -> let runner_env = { charon_path; aeneas_path; llbc_dir } in - let test_name = Filename.remove_extension (Filename.basename test_path) in - - let charon_kind = - if Sys_unix.is_file_exn test_path then SingleFile - else if Sys_unix.is_directory_exn test_path then Crate - else failwith ("`" ^ test_path ^ "` is not a file or a directory.") - in - let extra_charon_options = charon_options_for_test test_name in - let charon_case = + let test_case = build_test test_path in + let test_case = { - path = test_path; - name = test_name; - kind = charon_kind; - extra_charon_options; + test_case with + aeneas_options = + BackendMap.map (List.append aeneas_options) test_case.aeneas_options; } in - (* Generate the llbc file *) - run_charon runner_env charon_case; - (* FIXME: remove this special case *) - let test_name = - if test_name = "betree" then "betree_main" else test_name - in - (* TODO: reactivate HOL4 once traits are parameterized by their associated types *) - let backends = [ "coq"; "lean"; "fstar" ] in + (* Generate the llbc file *) + run_charon runner_env test_case; + (* Process the llbc file for the each backend *) List.iter - (fun backend -> - let subdir = test_subdir backend test_name in - let extra_aeneas_options = - List.append - (aeneas_options_for_test backend test_name) - aeneas_options - in - let aeneas_case = - { name = test_name; backend; subdir; extra_aeneas_options } - in - (* Process the llbc file for the current backend *) - run_aeneas runner_env aeneas_case) - backends + (fun backend -> run_aeneas runner_env test_case backend) + Backend.all | _ -> failwith "Incorrect options passed to test runner" -- cgit v1.2.3 From 03f9d1767b19fe68c35a5adcf936ac9f39ab7882 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Thu, 23 May 2024 17:29:05 +0200 Subject: runner: define an Input module --- tests/test_runner/run_test.ml | 182 +++++++++++++++++++++--------------------- 1 file changed, 92 insertions(+), 90 deletions(-) diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index 2c411c95..e168069b 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -27,18 +27,6 @@ end module BackendMap = Map.Make (Backend) -type input_kind = SingleFile | Crate - -(* The data for a specific test input *) -type test_case = { - name : string; - path : string; - input_kind : input_kind; - charon_options : string list; - aeneas_options : string list BackendMap.t; - subdir : string BackendMap.t; -} - let concat_path = List.fold_left Filename.concat "" let run_command args = @@ -58,61 +46,6 @@ let run_command args = let _ = Unix.waitpid [] pid in () -(* Run Aeneas on a specific input with the given options *) -let run_aeneas (env : runner_env) (case : test_case) (backend : Backend.t) = - (* FIXME: remove this special case *) - let test_name = if case.name = "betree" then "betree_main" else case.name in - let input_file = concat_path [ env.llbc_dir; test_name ] ^ ".llbc" in - let subdir = BackendMap.find backend case.subdir in - let aeneas_options = BackendMap.find backend case.aeneas_options in - let backend_str = Backend.to_string backend in - let dest_dir = concat_path [ "tests"; backend_str; subdir ] in - let args = - [| - env.aeneas_path; input_file; "-dest"; dest_dir; "-backend"; backend_str; - |] - in - let args = Array.append args (Array.of_list aeneas_options) in - (* Run Aeneas *) - run_command args - -(* Run Charon on a specific input with the given options *) -let run_charon (env : runner_env) (case : test_case) = - match case.input_kind with - | SingleFile -> - let args = - [| - env.charon_path; - "--no-cargo"; - "--input"; - case.path; - "--crate"; - case.name; - "--dest"; - env.llbc_dir; - |] - in - let args = Array.append args (Array.of_list case.charon_options) in - (* Run Charon on the rust file *) - run_command args - | Crate -> ( - match Sys.getenv_opt "IN_CI" with - | None -> - let args = [| env.charon_path; "--dest"; env.llbc_dir |] in - let args = Array.append args (Array.of_list case.charon_options) in - (* Run Charon inside the crate *) - let old_pwd = Unix.getcwd () in - Unix.chdir case.path; - run_command args; - Unix.chdir old_pwd - | Some _ -> - (* Crates with dependencies must be generated separately in CI. We skip - here and trust that CI takes care to generate the necessary llbc - file. *) - print_endline - "Warn: IN_CI is set; we skip generating llbc files for whole crates" - ) - (* File-specific options *) let aeneas_options_for_test backend test_name = let backend = Backend.to_string backend in @@ -205,29 +138,98 @@ let test_subdir backend test_name = "misc" | _ -> test_name -let build_test (path : string) : test_case = - let name = Filename.remove_extension (Filename.basename path) in - let input_kind = - if Sys_unix.is_file_exn path then SingleFile - else if Sys_unix.is_directory_exn path then Crate - else failwith ("`" ^ path ^ "` is not a file or a directory.") - in - let charon_options = charon_options_for_test name in - let subdir = - List.fold_left - (fun map backend -> - let subdir = test_subdir backend name in - BackendMap.add backend subdir map) - BackendMap.empty Backend.all - in - let aeneas_options = - List.fold_left - (fun map backend -> - let opts = aeneas_options_for_test backend name in - BackendMap.add backend opts map) - BackendMap.empty Backend.all +(* The data for a specific test input *) +module Input = struct + type kind = SingleFile | Crate + + type t = { + name : string; + path : string; + kind : kind; + charon_options : string list; + aeneas_options : string list BackendMap.t; + subdir : string BackendMap.t; + } + + let build (path : string) : t = + let name = Filename.remove_extension (Filename.basename path) in + let kind = + if Sys_unix.is_file_exn path then SingleFile + else if Sys_unix.is_directory_exn path then Crate + else failwith ("`" ^ path ^ "` is not a file or a directory.") + in + let charon_options = charon_options_for_test name in + let subdir = + List.fold_left + (fun map backend -> + let subdir = test_subdir backend name in + BackendMap.add backend subdir map) + BackendMap.empty Backend.all + in + let aeneas_options = + List.fold_left + (fun map backend -> + let opts = aeneas_options_for_test backend name in + BackendMap.add backend opts map) + BackendMap.empty Backend.all + in + { path; name; kind; charon_options; subdir; aeneas_options } +end + +(* Run Aeneas on a specific input with the given options *) +let run_aeneas (env : runner_env) (case : Input.t) (backend : Backend.t) = + (* FIXME: remove this special case *) + let test_name = if case.name = "betree" then "betree_main" else case.name in + let input_file = concat_path [ env.llbc_dir; test_name ] ^ ".llbc" in + let subdir = BackendMap.find backend case.subdir in + let aeneas_options = BackendMap.find backend case.aeneas_options in + let backend_str = Backend.to_string backend in + let dest_dir = concat_path [ "tests"; backend_str; subdir ] in + let args = + [| + env.aeneas_path; input_file; "-dest"; dest_dir; "-backend"; backend_str; + |] in - { path; name; input_kind; charon_options; subdir; aeneas_options } + let args = Array.append args (Array.of_list aeneas_options) in + (* Run Aeneas *) + run_command args + +(* Run Charon on a specific input with the given options *) +let run_charon (env : runner_env) (case : Input.t) = + match case.kind with + | SingleFile -> + let args = + [| + env.charon_path; + "--no-cargo"; + "--input"; + case.path; + "--crate"; + case.name; + "--dest"; + env.llbc_dir; + |] + in + let args = Array.append args (Array.of_list case.charon_options) in + (* Run Charon on the rust file *) + run_command args + | Crate -> ( + match Sys.getenv_opt "IN_CI" with + | None -> + let args = [| env.charon_path; "--dest"; env.llbc_dir |] in + let args = Array.append args (Array.of_list case.charon_options) in + (* Run Charon inside the crate *) + let old_pwd = Unix.getcwd () in + Unix.chdir case.path; + run_command args; + Unix.chdir old_pwd + | Some _ -> + (* Crates with dependencies must be generated separately in CI. We skip + here and trust that CI takes care to generate the necessary llbc + file. *) + print_endline + "Warn: IN_CI is set; we skip generating llbc files for whole crates" + ) let () = match Array.to_list Sys.argv with @@ -235,7 +237,7 @@ let () = | _exe_path :: charon_path :: aeneas_path :: llbc_dir :: test_path :: aeneas_options -> let runner_env = { charon_path; aeneas_path; llbc_dir } in - let test_case = build_test test_path in + let test_case = Input.build test_path in let test_case = { test_case with -- cgit v1.2.3 From c0a6916549b55fbae2226e4c34af45e500928645 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Fri, 24 May 2024 14:56:26 +0200 Subject: Fix running individual tests --- Makefile | 10 +++++++--- 1 file changed, 7 insertions(+), 3 deletions(-) diff --git a/Makefile b/Makefile index b3dc9f2b..ca3bd66c 100644 --- a/Makefile +++ b/Makefile @@ -47,7 +47,12 @@ build-test-verify: build test verify # Build the project, without formatting the code .PHONY: build-dev +ifdef IN_CI +build-dev: + @true +else build-dev: build-bin build-lib build-bin-dir doc +endif .PHONY: build-bin build-bin: check-charon @@ -166,14 +171,13 @@ test-all: $(INPUTS_LIST) ifdef IN_CI # In CI we do extra sanity checks. test-%: AENEAS_OPTIONS += -checks -else -test-%: check-charon endif # Translate the given rust file to available backends. The test runner decides # which backends to use and sets test-specific options. +# Note: the tests have the fulle file name: `test-arrays.rs`, `test-loops.rs`, `test-betree`. .PHONY: test-% -test-%: +test-%: build-dev $(TEST_RUNNER_EXE) $(CHARON_EXE) $(AENEAS_EXE) $(LLBC_DIR) $(INPUTS_DIR)/"$*" $(AENEAS_OPTIONS) echo "# Test $* done" -- cgit v1.2.3 From 4d3778bea3112168645efc03308056ec341abb5f Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Fri, 24 May 2024 15:47:20 +0200 Subject: runner: Pass options in special comments --- Makefile | 4 - tests/src/arrays.rs | 3 + tests/src/bitwise.rs | 1 + tests/src/constants.rs | 2 + tests/src/demo.rs | 1 + tests/src/external.rs | 3 + tests/src/hashmap.rs | 8 ++ tests/src/hashmap_main.rs | 6 ++ tests/src/hashmap_utils.rs | 1 + tests/src/loops.rs | 3 + tests/src/nested_borrows.rs | 2 + tests/src/no_nested_borrows.rs | 2 + tests/src/paper.rs | 2 + tests/src/polonius_list.rs | 2 + tests/src/traits.rs | 1 + tests/test_runner/dune | 2 +- tests/test_runner/run_test.ml | 189 ++++++++++++++++++++++++----------------- 17 files changed, 150 insertions(+), 82 deletions(-) diff --git a/Makefile b/Makefile index ca3bd66c..bda88c74 100644 --- a/Makefile +++ b/Makefile @@ -159,10 +159,6 @@ verify: INPUTS_LIST = $(wildcard $(INPUTS_DIR)/*) # Remove the directory prefix, replace with `test-` INPUTS_LIST := $(subst $(INPUTS_DIR)/,test-,$(INPUTS_LIST)) -# Remove some tests we don't want to run. -# FIXME: move this information to the file itself -INPUTS_LIST := $(filter-out test-hashmap_utils.rs,$(INPUTS_LIST)) -INPUTS_LIST := $(filter-out test-nested_borrows.rs,$(INPUTS_LIST)) # Run all the tests we found. .PHONY: test-all diff --git a/tests/src/arrays.rs b/tests/src/arrays.rs index 1f094541..ddad2ad3 100644 --- a/tests/src/arrays.rs +++ b/tests/src/arrays.rs @@ -1,3 +1,6 @@ +//@ [coq] aeneas-args=-use-fuel +//@ [fstar] aeneas-args=-decreases-clauses -template-clauses +//@ [fstar] aeneas-args=-split-files //! Exercise the translation of arrays, with features supported by Eurydice pub enum AB { diff --git a/tests/src/bitwise.rs b/tests/src/bitwise.rs index 9f48cb04..15962047 100644 --- a/tests/src/bitwise.rs +++ b/tests/src/bitwise.rs @@ -1,3 +1,4 @@ +//@ aeneas-args=-test-trans-units //! Exercise the bitwise operations pub fn shift_u32(a: u32) -> u32 { diff --git a/tests/src/constants.rs b/tests/src/constants.rs index 83904eed..925c62b1 100644 --- a/tests/src/constants.rs +++ b/tests/src/constants.rs @@ -1,3 +1,5 @@ +//@ charon-args=--no-code-duplication +//@ aeneas-args=-test-trans-units //! Tests with constants // Integers diff --git a/tests/src/demo.rs b/tests/src/demo.rs index bc74cc8b..b9bb7ca2 100644 --- a/tests/src/demo.rs +++ b/tests/src/demo.rs @@ -1,3 +1,4 @@ +//@ [coq,fstar] aeneas-args=-use-fuel #![allow(clippy::needless_lifetimes)] /* Simple functions */ diff --git a/tests/src/external.rs b/tests/src/external.rs index 521749d6..ddd5539f 100644 --- a/tests/src/external.rs +++ b/tests/src/external.rs @@ -1,3 +1,6 @@ +//@ charon-args=--no-code-duplication +//@ aeneas-args=-state -split-files +//@ aeneas-args=-test-trans-units //! This module uses external types and functions use std::cell::Cell; diff --git a/tests/src/hashmap.rs b/tests/src/hashmap.rs index 58d22acd..4552e4f2 100644 --- a/tests/src/hashmap.rs +++ b/tests/src/hashmap.rs @@ -1,3 +1,11 @@ +//@ [coq] aeneas-args=-use-fuel +//@ aeneas-args=-split-files +//@ [fstar] aeneas-args=-decreases-clauses -template-clauses +//@ [lean] aeneas-args=-no-gen-lib-entry +// ^ the `-no-gen-lib-entry` is because we add a custom import in the Hashmap.lean file: we do not +// want to overwrite it. +// TODO: reactivate -test-trans-units + //! A hashmap implementation. //! //! Current limitations: diff --git a/tests/src/hashmap_main.rs b/tests/src/hashmap_main.rs index 45dfa6e2..0c827844 100644 --- a/tests/src/hashmap_main.rs +++ b/tests/src/hashmap_main.rs @@ -1,3 +1,9 @@ +//@ charon-args=--opaque=hashmap_utils +//@ aeneas-args=-state -split-files +//@ [coq] aeneas-args=-use-fuel +//@ [fstar] aeneas-args=-decreases-clauses -template-clauses +// Possible to add `--no-code-duplication` if we use the optimized MIR +// TODO: reactivate -test-trans-units mod hashmap; mod hashmap_utils; diff --git a/tests/src/hashmap_utils.rs b/tests/src/hashmap_utils.rs index cd7b481f..33de49e1 100644 --- a/tests/src/hashmap_utils.rs +++ b/tests/src/hashmap_utils.rs @@ -1,3 +1,4 @@ +//@ skip use crate::hashmap::*; /// Serialize a hash map - we don't have traits, so we fix the type of the diff --git a/tests/src/loops.rs b/tests/src/loops.rs index 2f71d75b..8692c60e 100644 --- a/tests/src/loops.rs +++ b/tests/src/loops.rs @@ -1,3 +1,6 @@ +//@ [coq] aeneas-args=-use-fuel +//@ [fstar] aeneas-args=-decreases-clauses -template-clauses +//@ [fstar] aeneas-args=-split-files use std::vec::Vec; /// No borrows diff --git a/tests/src/nested_borrows.rs b/tests/src/nested_borrows.rs index 94db0cec..d4d8cf73 100644 --- a/tests/src/nested_borrows.rs +++ b/tests/src/nested_borrows.rs @@ -1,3 +1,5 @@ +//@ skip +//@ charon-args=--no-code-duplication //! This module contains functions with nested borrows in their signatures. pub fn id_mut_mut<'a, 'b, T>(x: &'a mut &'b mut T) -> &'a mut &'b mut T { diff --git a/tests/src/no_nested_borrows.rs b/tests/src/no_nested_borrows.rs index 9a7604e6..78163371 100644 --- a/tests/src/no_nested_borrows.rs +++ b/tests/src/no_nested_borrows.rs @@ -1,3 +1,5 @@ +//@ charon-args=--no-code-duplication +//@ aeneas-args=-test-trans-units //! This module doesn't contain **functions which use nested borrows in their //! signatures**, and doesn't contain functions with loops. diff --git a/tests/src/paper.rs b/tests/src/paper.rs index 156f13ab..07453098 100644 --- a/tests/src/paper.rs +++ b/tests/src/paper.rs @@ -1,3 +1,5 @@ +//@ charon-args=--no-code-duplication +//@ aeneas-args=-test-trans-units //! The examples from the ICFP submission, all in one place. // 2.1 diff --git a/tests/src/polonius_list.rs b/tests/src/polonius_list.rs index 8c64110d..a8d51e40 100644 --- a/tests/src/polonius_list.rs +++ b/tests/src/polonius_list.rs @@ -1,3 +1,5 @@ +//@ charon-args=--polonius +//@ aeneas-args=-test-trans-units #![allow(dead_code)] pub enum List { diff --git a/tests/src/traits.rs b/tests/src/traits.rs index 27c90586..fd50db8c 100644 --- a/tests/src/traits.rs +++ b/tests/src/traits.rs @@ -1,3 +1,4 @@ +//@ [fstar] aeneas-args=-decreases-clauses -template-clauses pub trait BoolTrait { // Required method fn get_bool(&self) -> bool; diff --git a/tests/test_runner/dune b/tests/test_runner/dune index 65b0c5fe..e8b29d66 100644 --- a/tests/test_runner/dune +++ b/tests/test_runner/dune @@ -1,6 +1,6 @@ (executable (public_name test_runner) - (libraries core_unix.sys_unix unix) + (libraries core_unix.sys_unix re unix) (preprocess (pps ppx_deriving.show ppx_deriving.ord ppx_sexp_conv)) (name run_test)) diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index e168069b..25efbcfd 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -1,3 +1,15 @@ +(* Convenience functions *) +let map_while (f : 'a -> 'b option) (input : 'a list) : 'b list = + let _, result = + List.fold_left + (fun (continue, out) a -> + if continue then + match f a with None -> (false, out) | Some b -> (true, b :: out) + else (continue, out)) + (true, []) input + in + List.rev result + (* Paths to use for tests *) type runner_env = { charon_path : string; @@ -48,78 +60,30 @@ let run_command args = (* File-specific options *) let aeneas_options_for_test backend test_name = - let backend = Backend.to_string backend in - (* TODO: reactivate -test-trans-units for hashmap and hashmap_main *) - let use_fuel = - match (backend, test_name) with - | ( "coq", - ("arrays" | "betree" | "demo" | "hashmap" | "hashmap_main" | "loops") ) - -> - true - | "fstar", "demo" -> true - | _ -> false - in - let options = if use_fuel then "-use-fuel" :: [] else [] in - - let decrease_template_clauses = - backend = "fstar" - && - match test_name with - | "arrays" | "betree" | "hashmap" | "hashmap_main" | "loops" | "traits" -> - true - | _ -> false - in - let options = - if decrease_template_clauses then - "-decreases-clauses" :: "-template-clauses" :: options - else options - in - - let extra_options = - match (backend, test_name) with - | _, "betree" -> - [ - "-backward-no-state-update"; - "-test-trans-units"; - "-state"; - "-split-files"; - ] - | _, "bitwise" -> [ "-test-trans-units" ] - | _, "constants" -> [ "-test-trans-units" ] - | _, "external" -> [ "-test-trans-units"; "-state"; "-split-files" ] - | _, "hashmap_main" -> [ "-state"; "-split-files" ] - | _, "no_nested_borrows" -> [ "-test-trans-units" ] - | _, "paper" -> [ "-test-trans-units" ] - | _, "polonius_list" -> [ "-test-trans-units" ] - | "fstar", "arrays" -> [ "-split-files" ] - | "fstar", "loops" -> [ "-split-files" ] - (* We add a custom import in the Hashmap.lean file: we do not want to overwrite it *) - | "lean", "hashmap" -> [ "-split-files"; "-no-gen-lib-entry" ] - | _, "hashmap" -> [ "-split-files" ] - | _ -> [] - in - let options = List.append extra_options options in - options + if test_name = "betree" then + let options = + [ + "-backward-no-state-update"; + "-test-trans-units"; + "-state"; + "-split-files"; + ] + in + let extra_options = + match backend with + | Backend.Coq -> [ "-use-fuel" ] + | Backend.FStar -> [ "-decreases-clauses"; "-template-clauses" ] + | _ -> [] + in + List.append extra_options options + else [] (* File-specific options *) let charon_options_for_test test_name = - (* Possible to add `--no-code-duplication` for `hashmap_main` if we use the optimized MIR *) - let no_code_dup = - match test_name with - | "constants" | "external" | "nested_borrows" | "no_nested_borrows" - | "paper" -> - [ "--no-code-duplication" ] - | _ -> [] - in - let extra_options = - match test_name with - | "betree" -> - [ "--polonius"; "--opaque=betree_utils"; "--crate"; "betree_main" ] - | "hashmap_main" -> [ "--opaque=hashmap_utils" ] - | "polonius_list" -> [ "--polonius" ] - | _ -> [] - in - List.append no_code_dup extra_options + match test_name with + | "betree" -> + [ "--polonius"; "--opaque=betree_utils"; "--crate"; "betree_main" ] + | _ -> [] (* The subdirectory in which to store the outputs. *) (* This reproduces the file layout that was set by the old Makefile. FIXME: cleanup *) @@ -141,16 +105,66 @@ let test_subdir backend test_name = (* The data for a specific test input *) module Input = struct type kind = SingleFile | Crate + type action = Normal | Skip | KnownFailure type t = { name : string; path : string; kind : kind; + action : action; charon_options : string list; aeneas_options : string list BackendMap.t; subdir : string BackendMap.t; } + (* Parse lines that start `//@`. Each of them modifies the options we use for the test. + Supported comments: + - `skip`: don't process the file; + - `known-failure`: TODO; + - `charon-args=...`: extra arguments to pass to charon; + - `aeneas-args=...`: extra arguments to pass to aeneas; + - `[backend,..]...`: where each `backend` is the name of a backend supported by + aeneas; this sets options for these backends only. Only supported for `aeneas-args`. + *) + let apply_special_comment comment input = + let comment = String.trim comment in + (* Parse the backends if any *) + let re = Re.compile (Re.Pcre.re "^\\[([a-zA-Z,]+)\\](.*)$") in + let comment, (backends : Backend.t list) = + match Re.exec_opt re comment with + | Some groups -> + let backends = Re.Group.get groups 1 in + let backends = String.split_on_char ',' backends in + let backends = List.map Backend.of_string backends in + let rest = Re.Group.get groups 2 in + (String.trim rest, backends) + | None -> (comment, Backend.all) + in + (* Parse the other options *) + let charon_args = Core.String.chop_prefix comment ~prefix:"charon-args=" in + let aeneas_args = Core.String.chop_prefix comment ~prefix:"aeneas-args=" in + + if comment = "skip" then { input with action = Skip } + else if comment = "known-failure" then { input with action = KnownFailure } + else if Option.is_some charon_args then + let args = Option.get charon_args in + let args = String.split_on_char ' ' args in + { input with charon_options = List.append input.charon_options args } + else if Option.is_some aeneas_args then + let args = Option.get aeneas_args in + let args = String.split_on_char ' ' args in + let add_args opts = List.append opts args in + { + input with + aeneas_options = + List.fold_left + (fun map backend -> + BackendMap.update backend (Option.map add_args) map) + input.aeneas_options backends; + } + else failwith ("Unrecognized special comment: `" ^ comment ^ "`") + + (* Given a path to a rust file or crate, gather the details and options about how to build the test. *) let build (path : string) : t = let name = Filename.remove_extension (Filename.basename path) in let kind = @@ -158,6 +172,7 @@ module Input = struct else if Sys_unix.is_directory_exn path then Crate else failwith ("`" ^ path ^ "` is not a file or a directory.") in + let action = Normal in let charon_options = charon_options_for_test name in let subdir = List.fold_left @@ -173,7 +188,23 @@ module Input = struct BackendMap.add backend opts map) BackendMap.empty Backend.all in - { path; name; kind; charon_options; subdir; aeneas_options } + let input = + { path; name; kind; action; charon_options; subdir; aeneas_options } + in + match kind with + | SingleFile -> + let file_lines = Core.In_channel.read_lines path in + (* Extract the special lines. Stop at the first non-special line. *) + let special_comments = + map_while + (fun line -> Core.String.chop_prefix line ~prefix:"//@") + file_lines + in + (* Apply the changes from the special lines to our input. *) + List.fold_left + (fun input comment -> apply_special_comment comment input) + input special_comments + | Crate -> input end (* Run Aeneas on a specific input with the given options *) @@ -235,7 +266,7 @@ let () = match Array.to_list Sys.argv with (* Ad-hoc argument passing for now. *) | _exe_path :: charon_path :: aeneas_path :: llbc_dir :: test_path - :: aeneas_options -> + :: aeneas_options -> ( let runner_env = { charon_path; aeneas_path; llbc_dir } in let test_case = Input.build test_path in let test_case = @@ -246,10 +277,14 @@ let () = } in - (* Generate the llbc file *) - run_charon runner_env test_case; - (* Process the llbc file for the each backend *) - List.iter - (fun backend -> run_aeneas runner_env test_case backend) - Backend.all + match test_case.action with + | Skip -> () + | Normal -> + (* Generate the llbc file *) + run_charon runner_env test_case; + (* Process the llbc file for the each backend *) + List.iter + (fun backend -> run_aeneas runner_env test_case backend) + Backend.all + | KnownFailure -> failwith "KnownFailure is unimplemented") | _ -> failwith "Incorrect options passed to test runner" -- cgit v1.2.3 From 37e8a0f5ff7d964eb9525fef765b38e44f79302b Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Fri, 24 May 2024 15:48:27 +0200 Subject: Update test outputs --- tests/coq/arrays/Arrays.v | 126 ++++++++-------- tests/coq/demo/Demo.v | 32 ++-- tests/coq/hashmap/Hashmap_Funs.v | 60 ++++---- tests/coq/hashmap/Hashmap_Types.v | 4 +- tests/coq/hashmap_on_disk/HashmapMain_Funs.v | 64 ++++---- .../HashmapMain_FunsExternal_Template.v | 4 +- tests/coq/hashmap_on_disk/HashmapMain_Types.v | 4 +- tests/coq/misc/Bitwise.v | 10 +- tests/coq/misc/Constants.v | 62 ++++---- tests/coq/misc/External_Funs.v | 4 +- tests/coq/misc/Loops.v | 98 ++++++------ tests/coq/misc/NoNestedBorrows.v | 126 ++++++++-------- tests/coq/misc/Paper.v | 18 +-- tests/coq/misc/PoloniusList.v | 4 +- tests/coq/traits/Traits.v | 166 ++++++++++----------- tests/fstar/arrays/Arrays.Clauses.Template.fst | 10 +- tests/fstar/arrays/Arrays.Funs.fst | 124 +++++++-------- tests/fstar/arrays/Arrays.Types.fst | 2 +- tests/fstar/demo/Demo.fst | 32 ++-- tests/fstar/hashmap/Hashmap.Clauses.Template.fst | 18 +-- tests/fstar/hashmap/Hashmap.Funs.fst | 60 ++++---- tests/fstar/hashmap/Hashmap.Types.fst | 4 +- .../HashmapMain.Clauses.Template.fst | 18 +-- tests/fstar/hashmap_on_disk/HashmapMain.Funs.fst | 64 ++++---- .../hashmap_on_disk/HashmapMain.FunsExternal.fsti | 4 +- tests/fstar/hashmap_on_disk/HashmapMain.Types.fst | 4 +- tests/fstar/misc/Bitwise.fst | 10 +- tests/fstar/misc/Constants.fst | 62 ++++---- tests/fstar/misc/External.Funs.fst | 4 +- tests/fstar/misc/Loops.Clauses.Template.fst | 46 +++--- tests/fstar/misc/Loops.Funs.fst | 96 ++++++------ tests/fstar/misc/Loops.Types.fst | 2 +- tests/fstar/misc/NoNestedBorrows.fst | 126 ++++++++-------- tests/fstar/misc/Paper.fst | 18 +-- tests/fstar/misc/PoloniusList.fst | 4 +- tests/fstar/traits/Traits.fst | 166 ++++++++++----------- tests/lean/Arrays.lean | 126 ++++++++-------- tests/lean/Bitwise.lean | 10 +- tests/lean/Constants.lean | 62 ++++---- tests/lean/Demo/Demo.lean | 32 ++-- tests/lean/External/Funs.lean | 4 +- tests/lean/Hashmap/Funs.lean | 60 ++++---- tests/lean/Hashmap/Types.lean | 4 +- tests/lean/HashmapMain/Funs.lean | 64 ++++---- tests/lean/HashmapMain/FunsExternal_Template.lean | 4 +- tests/lean/HashmapMain/Types.lean | 4 +- tests/lean/Loops.lean | 98 ++++++------ tests/lean/NoNestedBorrows.lean | 126 ++++++++-------- tests/lean/Paper.lean | 18 +-- tests/lean/PoloniusList.lean | 4 +- tests/lean/Traits.lean | 166 ++++++++++----------- 51 files changed, 1219 insertions(+), 1219 deletions(-) diff --git a/tests/coq/arrays/Arrays.v b/tests/coq/arrays/Arrays.v index 35dea58c..371e4a12 100644 --- a/tests/coq/arrays/Arrays.v +++ b/tests/coq/arrays/Arrays.v @@ -9,23 +9,23 @@ Local Open Scope Primitives_scope. Module Arrays. (** [arrays::AB] - Source: 'tests/src/arrays.rs', lines 3:0-3:11 *) + Source: 'tests/src/arrays.rs', lines 6:0-6:11 *) Inductive AB_t := | AB_A : AB_t | AB_B : AB_t. (** [arrays::incr]: - Source: 'tests/src/arrays.rs', lines 8:0-8:24 *) + Source: 'tests/src/arrays.rs', lines 11:0-11:24 *) Definition incr (x : u32) : result u32 := u32_add x 1%u32. (** [arrays::array_to_shared_slice_]: - Source: 'tests/src/arrays.rs', lines 16:0-16:53 *) + Source: 'tests/src/arrays.rs', lines 19:0-19:53 *) Definition array_to_shared_slice_ (T : Type) (s : array T 32%usize) : result (slice T) := array_to_slice T 32%usize s . (** [arrays::array_to_mut_slice_]: - Source: 'tests/src/arrays.rs', lines 21:0-21:58 *) + Source: 'tests/src/arrays.rs', lines 24:0-24:58 *) Definition array_to_mut_slice_ (T : Type) (s : array T 32%usize) : result ((slice T) * (slice T -> result (array T 32%usize))) @@ -34,44 +34,44 @@ Definition array_to_mut_slice_ . (** [arrays::array_len]: - Source: 'tests/src/arrays.rs', lines 25:0-25:40 *) + Source: 'tests/src/arrays.rs', lines 28:0-28:40 *) Definition array_len (T : Type) (s : array T 32%usize) : result usize := s1 <- array_to_slice T 32%usize s; Ok (slice_len T s1) . (** [arrays::shared_array_len]: - Source: 'tests/src/arrays.rs', lines 29:0-29:48 *) + Source: 'tests/src/arrays.rs', lines 32:0-32:48 *) Definition shared_array_len (T : Type) (s : array T 32%usize) : result usize := s1 <- array_to_slice T 32%usize s; Ok (slice_len T s1) . (** [arrays::shared_slice_len]: - Source: 'tests/src/arrays.rs', lines 33:0-33:44 *) + Source: 'tests/src/arrays.rs', lines 36:0-36:44 *) Definition shared_slice_len (T : Type) (s : slice T) : result usize := Ok (slice_len T s) . (** [arrays::index_array_shared]: - Source: 'tests/src/arrays.rs', lines 37:0-37:57 *) + Source: 'tests/src/arrays.rs', lines 40:0-40:57 *) Definition index_array_shared (T : Type) (s : array T 32%usize) (i : usize) : result T := array_index_usize T 32%usize s i . (** [arrays::index_array_u32]: - Source: 'tests/src/arrays.rs', lines 44:0-44:53 *) + Source: 'tests/src/arrays.rs', lines 47:0-47:53 *) Definition index_array_u32 (s : array u32 32%usize) (i : usize) : result u32 := array_index_usize u32 32%usize s i . (** [arrays::index_array_copy]: - Source: 'tests/src/arrays.rs', lines 48:0-48:45 *) + Source: 'tests/src/arrays.rs', lines 51:0-51:45 *) Definition index_array_copy (x : array u32 32%usize) : result u32 := array_index_usize u32 32%usize x 0%usize . (** [arrays::index_mut_array]: - Source: 'tests/src/arrays.rs', lines 52:0-52:62 *) + Source: 'tests/src/arrays.rs', lines 55:0-55:62 *) Definition index_mut_array (T : Type) (s : array T 32%usize) (i : usize) : result (T * (T -> result (array T 32%usize))) @@ -80,13 +80,13 @@ Definition index_mut_array . (** [arrays::index_slice]: - Source: 'tests/src/arrays.rs', lines 56:0-56:46 *) + Source: 'tests/src/arrays.rs', lines 59:0-59:46 *) Definition index_slice (T : Type) (s : slice T) (i : usize) : result T := slice_index_usize T s i . (** [arrays::index_mut_slice]: - Source: 'tests/src/arrays.rs', lines 60:0-60:58 *) + Source: 'tests/src/arrays.rs', lines 63:0-63:58 *) Definition index_mut_slice (T : Type) (s : slice T) (i : usize) : result (T * (T -> result (slice T))) @@ -95,7 +95,7 @@ Definition index_mut_slice . (** [arrays::slice_subslice_shared_]: - Source: 'tests/src/arrays.rs', lines 64:0-64:70 *) + Source: 'tests/src/arrays.rs', lines 67:0-67:70 *) Definition slice_subslice_shared_ (x : slice u32) (y : usize) (z : usize) : result (slice u32) := core_slice_index_Slice_index u32 (core_ops_range_Range usize) @@ -104,7 +104,7 @@ Definition slice_subslice_shared_ . (** [arrays::slice_subslice_mut_]: - Source: 'tests/src/arrays.rs', lines 68:0-68:75 *) + Source: 'tests/src/arrays.rs', lines 71:0-71:75 *) Definition slice_subslice_mut_ (x : slice u32) (y : usize) (z : usize) : result ((slice u32) * (slice u32 -> result (slice u32))) @@ -118,14 +118,14 @@ Definition slice_subslice_mut_ . (** [arrays::array_to_slice_shared_]: - Source: 'tests/src/arrays.rs', lines 72:0-72:54 *) + Source: 'tests/src/arrays.rs', lines 75:0-75:54 *) Definition array_to_slice_shared_ (x : array u32 32%usize) : result (slice u32) := array_to_slice u32 32%usize x . (** [arrays::array_to_slice_mut_]: - Source: 'tests/src/arrays.rs', lines 76:0-76:59 *) + Source: 'tests/src/arrays.rs', lines 79:0-79:59 *) Definition array_to_slice_mut_ (x : array u32 32%usize) : result ((slice u32) * (slice u32 -> result (array u32 32%usize))) @@ -134,7 +134,7 @@ Definition array_to_slice_mut_ . (** [arrays::array_subslice_shared_]: - Source: 'tests/src/arrays.rs', lines 80:0-80:74 *) + Source: 'tests/src/arrays.rs', lines 83:0-83:74 *) Definition array_subslice_shared_ (x : array u32 32%usize) (y : usize) (z : usize) : result (slice u32) := core_array_Array_index u32 (core_ops_range_Range usize) 32%usize @@ -144,7 +144,7 @@ Definition array_subslice_shared_ . (** [arrays::array_subslice_mut_]: - Source: 'tests/src/arrays.rs', lines 84:0-84:79 *) + Source: 'tests/src/arrays.rs', lines 87:0-87:79 *) Definition array_subslice_mut_ (x : array u32 32%usize) (y : usize) (z : usize) : result ((slice u32) * (slice u32 -> result (array u32 32%usize))) @@ -159,19 +159,19 @@ Definition array_subslice_mut_ . (** [arrays::index_slice_0]: - Source: 'tests/src/arrays.rs', lines 88:0-88:38 *) + Source: 'tests/src/arrays.rs', lines 91:0-91:38 *) Definition index_slice_0 (T : Type) (s : slice T) : result T := slice_index_usize T s 0%usize . (** [arrays::index_array_0]: - Source: 'tests/src/arrays.rs', lines 92:0-92:42 *) + Source: 'tests/src/arrays.rs', lines 95:0-95:42 *) Definition index_array_0 (T : Type) (s : array T 32%usize) : result T := array_index_usize T 32%usize s 0%usize . (** [arrays::index_index_array]: - Source: 'tests/src/arrays.rs', lines 103:0-103:71 *) + Source: 'tests/src/arrays.rs', lines 106:0-106:71 *) Definition index_index_array (s : array (array u32 32%usize) 32%usize) (i : usize) (j : usize) : result u32 @@ -181,7 +181,7 @@ Definition index_index_array . (** [arrays::update_update_array]: - Source: 'tests/src/arrays.rs', lines 114:0-114:70 *) + Source: 'tests/src/arrays.rs', lines 117:0-117:70 *) Definition update_update_array (s : array (array u32 32%usize) 32%usize) (i : usize) (j : usize) : result unit @@ -196,46 +196,46 @@ Definition update_update_array . (** [arrays::array_local_deep_copy]: - Source: 'tests/src/arrays.rs', lines 118:0-118:43 *) + Source: 'tests/src/arrays.rs', lines 121:0-121:43 *) Definition array_local_deep_copy (x : array u32 32%usize) : result unit := Ok tt . (** [arrays::take_array]: - Source: 'tests/src/arrays.rs', lines 122:0-122:30 *) + Source: 'tests/src/arrays.rs', lines 125:0-125:30 *) Definition take_array (a : array u32 2%usize) : result unit := Ok tt. (** [arrays::take_array_borrow]: - Source: 'tests/src/arrays.rs', lines 123:0-123:38 *) + Source: 'tests/src/arrays.rs', lines 126:0-126:38 *) Definition take_array_borrow (a : array u32 2%usize) : result unit := Ok tt. (** [arrays::take_slice]: - Source: 'tests/src/arrays.rs', lines 124:0-124:28 *) + Source: 'tests/src/arrays.rs', lines 127:0-127:28 *) Definition take_slice (s : slice u32) : result unit := Ok tt. (** [arrays::take_mut_slice]: - Source: 'tests/src/arrays.rs', lines 125:0-125:36 *) + Source: 'tests/src/arrays.rs', lines 128:0-128:36 *) Definition take_mut_slice (s : slice u32) : result (slice u32) := Ok s. (** [arrays::const_array]: - Source: 'tests/src/arrays.rs', lines 127:0-127:32 *) + Source: 'tests/src/arrays.rs', lines 130:0-130:32 *) Definition const_array : result (array u32 2%usize) := Ok (mk_array u32 2%usize [ 0%u32; 0%u32 ]) . (** [arrays::const_slice]: - Source: 'tests/src/arrays.rs', lines 131:0-131:20 *) + Source: 'tests/src/arrays.rs', lines 134:0-134:20 *) Definition const_slice : result unit := _ <- array_to_slice u32 2%usize (mk_array u32 2%usize [ 0%u32; 0%u32 ]); Ok tt . (** [arrays::take_all]: - Source: 'tests/src/arrays.rs', lines 141:0-141:17 *) + Source: 'tests/src/arrays.rs', lines 144:0-144:17 *) Definition take_all : result unit := _ <- take_array (mk_array u32 2%usize [ 0%u32; 0%u32 ]); _ <- take_array (mk_array u32 2%usize [ 0%u32; 0%u32 ]); @@ -250,32 +250,32 @@ Definition take_all : result unit := . (** [arrays::index_array]: - Source: 'tests/src/arrays.rs', lines 155:0-155:38 *) + Source: 'tests/src/arrays.rs', lines 158:0-158:38 *) Definition index_array (x : array u32 2%usize) : result u32 := array_index_usize u32 2%usize x 0%usize . (** [arrays::index_array_borrow]: - Source: 'tests/src/arrays.rs', lines 158:0-158:46 *) + Source: 'tests/src/arrays.rs', lines 161:0-161:46 *) Definition index_array_borrow (x : array u32 2%usize) : result u32 := array_index_usize u32 2%usize x 0%usize . (** [arrays::index_slice_u32_0]: - Source: 'tests/src/arrays.rs', lines 162:0-162:42 *) + Source: 'tests/src/arrays.rs', lines 165:0-165:42 *) Definition index_slice_u32_0 (x : slice u32) : result u32 := slice_index_usize u32 x 0%usize . (** [arrays::index_mut_slice_u32_0]: - Source: 'tests/src/arrays.rs', lines 166:0-166:50 *) + Source: 'tests/src/arrays.rs', lines 169:0-169:50 *) Definition index_mut_slice_u32_0 (x : slice u32) : result (u32 * (slice u32)) := i <- slice_index_usize u32 x 0%usize; Ok (i, x) . (** [arrays::index_all]: - Source: 'tests/src/arrays.rs', lines 170:0-170:25 *) + Source: 'tests/src/arrays.rs', lines 173:0-173:25 *) Definition index_all : result u32 := i <- index_array (mk_array u32 2%usize [ 0%u32; 0%u32 ]); i1 <- index_array (mk_array u32 2%usize [ 0%u32; 0%u32 ]); @@ -295,7 +295,7 @@ Definition index_all : result u32 := . (** [arrays::update_array]: - Source: 'tests/src/arrays.rs', lines 184:0-184:36 *) + Source: 'tests/src/arrays.rs', lines 187:0-187:36 *) Definition update_array (x : array u32 2%usize) : result unit := p <- array_index_mut_usize u32 2%usize x 0%usize; let (_, index_mut_back) := p in @@ -304,7 +304,7 @@ Definition update_array (x : array u32 2%usize) : result unit := . (** [arrays::update_array_mut_borrow]: - Source: 'tests/src/arrays.rs', lines 187:0-187:48 *) + Source: 'tests/src/arrays.rs', lines 190:0-190:48 *) Definition update_array_mut_borrow (x : array u32 2%usize) : result (array u32 2%usize) := p <- array_index_mut_usize u32 2%usize x 0%usize; @@ -313,7 +313,7 @@ Definition update_array_mut_borrow . (** [arrays::update_mut_slice]: - Source: 'tests/src/arrays.rs', lines 190:0-190:38 *) + Source: 'tests/src/arrays.rs', lines 193:0-193:38 *) Definition update_mut_slice (x : slice u32) : result (slice u32) := p <- slice_index_mut_usize u32 x 0%usize; let (_, index_mut_back) := p in @@ -321,7 +321,7 @@ Definition update_mut_slice (x : slice u32) : result (slice u32) := . (** [arrays::update_all]: - Source: 'tests/src/arrays.rs', lines 194:0-194:19 *) + Source: 'tests/src/arrays.rs', lines 197:0-197:19 *) Definition update_all : result unit := _ <- update_array (mk_array u32 2%usize [ 0%u32; 0%u32 ]); _ <- update_array (mk_array u32 2%usize [ 0%u32; 0%u32 ]); @@ -334,7 +334,7 @@ Definition update_all : result unit := . (** [arrays::range_all]: - Source: 'tests/src/arrays.rs', lines 205:0-205:18 *) + Source: 'tests/src/arrays.rs', lines 208:0-208:18 *) Definition range_all : result unit := p <- core_array_Array_index_mut u32 (core_ops_range_Range usize) 4%usize @@ -352,31 +352,31 @@ Definition range_all : result unit := . (** [arrays::deref_array_borrow]: - Source: 'tests/src/arrays.rs', lines 214:0-214:46 *) + Source: 'tests/src/arrays.rs', lines 217:0-217:46 *) Definition deref_array_borrow (x : array u32 2%usize) : result u32 := array_index_usize u32 2%usize x 0%usize . (** [arrays::deref_array_mut_borrow]: - Source: 'tests/src/arrays.rs', lines 219:0-219:54 *) + Source: 'tests/src/arrays.rs', lines 222:0-222:54 *) Definition deref_array_mut_borrow (x : array u32 2%usize) : result (u32 * (array u32 2%usize)) := i <- array_index_usize u32 2%usize x 0%usize; Ok (i, x) . (** [arrays::take_array_t]: - Source: 'tests/src/arrays.rs', lines 227:0-227:31 *) + Source: 'tests/src/arrays.rs', lines 230:0-230:31 *) Definition take_array_t (a : array AB_t 2%usize) : result unit := Ok tt. (** [arrays::non_copyable_array]: - Source: 'tests/src/arrays.rs', lines 229:0-229:27 *) + Source: 'tests/src/arrays.rs', lines 232:0-232:27 *) Definition non_copyable_array : result unit := take_array_t (mk_array AB_t 2%usize [ AB_A; AB_B ]) . (** [arrays::sum]: loop 0: - Source: 'tests/src/arrays.rs', lines 242:0-250:1 *) + Source: 'tests/src/arrays.rs', lines 245:0-253:1 *) Fixpoint sum_loop (n : nat) (s : slice u32) (sum1 : u32) (i : usize) : result u32 := match n with @@ -394,13 +394,13 @@ Fixpoint sum_loop . (** [arrays::sum]: - Source: 'tests/src/arrays.rs', lines 242:0-242:28 *) + Source: 'tests/src/arrays.rs', lines 245:0-245:28 *) Definition sum (n : nat) (s : slice u32) : result u32 := sum_loop n s 0%u32 0%usize . (** [arrays::sum2]: loop 0: - Source: 'tests/src/arrays.rs', lines 252:0-261:1 *) + Source: 'tests/src/arrays.rs', lines 255:0-264:1 *) Fixpoint sum2_loop (n : nat) (s : slice u32) (s2 : slice u32) (sum1 : u32) (i : usize) : result u32 @@ -422,7 +422,7 @@ Fixpoint sum2_loop . (** [arrays::sum2]: - Source: 'tests/src/arrays.rs', lines 252:0-252:41 *) + Source: 'tests/src/arrays.rs', lines 255:0-255:41 *) Definition sum2 (n : nat) (s : slice u32) (s2 : slice u32) : result u32 := let i := slice_len u32 s in let i1 := slice_len u32 s2 in @@ -430,7 +430,7 @@ Definition sum2 (n : nat) (s : slice u32) (s2 : slice u32) : result u32 := . (** [arrays::f0]: - Source: 'tests/src/arrays.rs', lines 263:0-263:11 *) + Source: 'tests/src/arrays.rs', lines 266:0-266:11 *) Definition f0 : result unit := p <- array_to_slice_mut u32 2%usize (mk_array u32 2%usize [ 1%u32; 2%u32 ]); let (s, to_slice_mut_back) := p in @@ -442,7 +442,7 @@ Definition f0 : result unit := . (** [arrays::f1]: - Source: 'tests/src/arrays.rs', lines 268:0-268:11 *) + Source: 'tests/src/arrays.rs', lines 271:0-271:11 *) Definition f1 : result unit := p <- array_index_mut_usize u32 2%usize (mk_array u32 2%usize [ 1%u32; 2%u32 ]) @@ -453,12 +453,12 @@ Definition f1 : result unit := . (** [arrays::f2]: - Source: 'tests/src/arrays.rs', lines 273:0-273:17 *) + Source: 'tests/src/arrays.rs', lines 276:0-276:17 *) Definition f2 (i : u32) : result unit := Ok tt. (** [arrays::f4]: - Source: 'tests/src/arrays.rs', lines 282:0-282:54 *) + Source: 'tests/src/arrays.rs', lines 285:0-285:54 *) Definition f4 (x : array u32 32%usize) (y : usize) (z : usize) : result (slice u32) := core_array_Array_index u32 (core_ops_range_Range usize) 32%usize @@ -468,7 +468,7 @@ Definition f4 . (** [arrays::f3]: - Source: 'tests/src/arrays.rs', lines 275:0-275:18 *) + Source: 'tests/src/arrays.rs', lines 278:0-278:18 *) Definition f3 (n : nat) : result u32 := i <- array_index_usize u32 2%usize (mk_array u32 2%usize [ 1%u32; 2%u32 ]) @@ -481,18 +481,18 @@ Definition f3 (n : nat) : result u32 := . (** [arrays::SZ] - Source: 'tests/src/arrays.rs', lines 286:0-286:19 *) + Source: 'tests/src/arrays.rs', lines 289:0-289:19 *) Definition sz_body : result usize := Ok 32%usize. Definition sz : usize := sz_body%global. (** [arrays::f5]: - Source: 'tests/src/arrays.rs', lines 289:0-289:31 *) + Source: 'tests/src/arrays.rs', lines 292:0-292:31 *) Definition f5 (x : array u32 32%usize) : result u32 := array_index_usize u32 32%usize x 0%usize . (** [arrays::ite]: - Source: 'tests/src/arrays.rs', lines 294:0-294:12 *) + Source: 'tests/src/arrays.rs', lines 297:0-297:12 *) Definition ite : result unit := p <- array_to_slice_mut u32 2%usize (mk_array u32 2%usize [ 0%u32; 0%u32 ]); let (s, to_slice_mut_back) := p in @@ -508,7 +508,7 @@ Definition ite : result unit := . (** [arrays::zero_slice]: loop 0: - Source: 'tests/src/arrays.rs', lines 303:0-310:1 *) + Source: 'tests/src/arrays.rs', lines 306:0-313:1 *) Fixpoint zero_slice_loop (n : nat) (a : slice u8) (i : usize) (len : usize) : result (slice u8) := match n with @@ -526,13 +526,13 @@ Fixpoint zero_slice_loop . (** [arrays::zero_slice]: - Source: 'tests/src/arrays.rs', lines 303:0-303:31 *) + Source: 'tests/src/arrays.rs', lines 306:0-306:31 *) Definition zero_slice (n : nat) (a : slice u8) : result (slice u8) := let len := slice_len u8 a in zero_slice_loop n a 0%usize len . (** [arrays::iter_mut_slice]: loop 0: - Source: 'tests/src/arrays.rs', lines 312:0-318:1 *) + Source: 'tests/src/arrays.rs', lines 315:0-321:1 *) Fixpoint iter_mut_slice_loop (n : nat) (len : usize) (i : usize) : result unit := match n with @@ -545,13 +545,13 @@ Fixpoint iter_mut_slice_loop . (** [arrays::iter_mut_slice]: - Source: 'tests/src/arrays.rs', lines 312:0-312:35 *) + Source: 'tests/src/arrays.rs', lines 315:0-315:35 *) Definition iter_mut_slice (n : nat) (a : slice u8) : result (slice u8) := let len := slice_len u8 a in _ <- iter_mut_slice_loop n len 0%usize; Ok a . (** [arrays::sum_mut_slice]: loop 0: - Source: 'tests/src/arrays.rs', lines 320:0-328:1 *) + Source: 'tests/src/arrays.rs', lines 323:0-331:1 *) Fixpoint sum_mut_slice_loop (n : nat) (a : slice u32) (i : usize) (s : u32) : result u32 := match n with @@ -569,7 +569,7 @@ Fixpoint sum_mut_slice_loop . (** [arrays::sum_mut_slice]: - Source: 'tests/src/arrays.rs', lines 320:0-320:42 *) + Source: 'tests/src/arrays.rs', lines 323:0-323:42 *) Definition sum_mut_slice (n : nat) (a : slice u32) : result (u32 * (slice u32)) := i <- sum_mut_slice_loop n a 0%usize 0%u32; Ok (i, a) diff --git a/tests/coq/demo/Demo.v b/tests/coq/demo/Demo.v index 8d8f840d..1cccbeda 100644 --- a/tests/coq/demo/Demo.v +++ b/tests/coq/demo/Demo.v @@ -9,7 +9,7 @@ Local Open Scope Primitives_scope. Module Demo. (** [demo::choose]: - Source: 'tests/src/demo.rs', lines 5:0-5:70 *) + Source: 'tests/src/demo.rs', lines 6:0-6:70 *) Definition choose (T : Type) (b : bool) (x : T) (y : T) : result (T * (T -> result (T * T))) := if b @@ -18,30 +18,30 @@ Definition choose . (** [demo::mul2_add1]: - Source: 'tests/src/demo.rs', lines 13:0-13:31 *) + Source: 'tests/src/demo.rs', lines 14:0-14:31 *) Definition mul2_add1 (x : u32) : result u32 := i <- u32_add x x; u32_add i 1%u32 . (** [demo::use_mul2_add1]: - Source: 'tests/src/demo.rs', lines 17:0-17:43 *) + Source: 'tests/src/demo.rs', lines 18:0-18:43 *) Definition use_mul2_add1 (x : u32) (y : u32) : result u32 := i <- mul2_add1 x; u32_add i y . (** [demo::incr]: - Source: 'tests/src/demo.rs', lines 21:0-21:31 *) + Source: 'tests/src/demo.rs', lines 22:0-22:31 *) Definition incr (x : u32) : result u32 := u32_add x 1%u32. (** [demo::use_incr]: - Source: 'tests/src/demo.rs', lines 25:0-25:17 *) + Source: 'tests/src/demo.rs', lines 26:0-26:17 *) Definition use_incr : result unit := x <- incr 0%u32; x1 <- incr x; _ <- incr x1; Ok tt . (** [demo::CList] - Source: 'tests/src/demo.rs', lines 34:0-34:17 *) + Source: 'tests/src/demo.rs', lines 35:0-35:17 *) Inductive CList_t (T : Type) := | CList_CCons : T -> CList_t T -> CList_t T | CList_CNil : CList_t T @@ -51,7 +51,7 @@ Arguments CList_CCons { _ }. Arguments CList_CNil { _ }. (** [demo::list_nth]: - Source: 'tests/src/demo.rs', lines 39:0-39:56 *) + Source: 'tests/src/demo.rs', lines 40:0-40:56 *) Fixpoint list_nth (T : Type) (n : nat) (l : CList_t T) (i : u32) : result T := match n with | O => Fail_ OutOfFuel @@ -65,7 +65,7 @@ Fixpoint list_nth (T : Type) (n : nat) (l : CList_t T) (i : u32) : result T := . (** [demo::list_nth_mut]: - Source: 'tests/src/demo.rs', lines 54:0-54:68 *) + Source: 'tests/src/demo.rs', lines 55:0-55:68 *) Fixpoint list_nth_mut (T : Type) (n : nat) (l : CList_t T) (i : u32) : result (T * (T -> result (CList_t T))) @@ -91,7 +91,7 @@ Fixpoint list_nth_mut . (** [demo::list_nth_mut1]: loop 0: - Source: 'tests/src/demo.rs', lines 69:0-78:1 *) + Source: 'tests/src/demo.rs', lines 70:0-79:1 *) Fixpoint list_nth_mut1_loop (T : Type) (n : nat) (l : CList_t T) (i : u32) : result (T * (T -> result (CList_t T))) @@ -116,7 +116,7 @@ Fixpoint list_nth_mut1_loop . (** [demo::list_nth_mut1]: - Source: 'tests/src/demo.rs', lines 69:0-69:77 *) + Source: 'tests/src/demo.rs', lines 70:0-70:77 *) Definition list_nth_mut1 (T : Type) (n : nat) (l : CList_t T) (i : u32) : result (T * (T -> result (CList_t T))) @@ -125,7 +125,7 @@ Definition list_nth_mut1 . (** [demo::i32_id]: - Source: 'tests/src/demo.rs', lines 80:0-80:28 *) + Source: 'tests/src/demo.rs', lines 81:0-81:28 *) Fixpoint i32_id (n : nat) (i : i32) : result i32 := match n with | O => Fail_ OutOfFuel @@ -137,7 +137,7 @@ Fixpoint i32_id (n : nat) (i : i32) : result i32 := . (** [demo::list_tail]: - Source: 'tests/src/demo.rs', lines 88:0-88:64 *) + Source: 'tests/src/demo.rs', lines 89:0-89:64 *) Fixpoint list_tail (T : Type) (n : nat) (l : CList_t T) : result ((CList_t T) * (CList_t T -> result (CList_t T))) @@ -159,7 +159,7 @@ Fixpoint list_tail . (** Trait declaration: [demo::Counter] - Source: 'tests/src/demo.rs', lines 97:0-97:17 *) + Source: 'tests/src/demo.rs', lines 98:0-98:17 *) Record Counter_t (Self : Type) := mkCounter_t { Counter_t_incr : Self -> result (usize * Self); }. @@ -168,19 +168,19 @@ Arguments mkCounter_t { _ }. Arguments Counter_t_incr { _ }. (** [demo::{(demo::Counter for usize)}::incr]: - Source: 'tests/src/demo.rs', lines 102:4-102:31 *) + Source: 'tests/src/demo.rs', lines 103:4-103:31 *) Definition counterUsize_incr (self : usize) : result (usize * usize) := self1 <- usize_add self 1%usize; Ok (self, self1) . (** Trait implementation: [demo::{(demo::Counter for usize)}] - Source: 'tests/src/demo.rs', lines 101:0-101:22 *) + Source: 'tests/src/demo.rs', lines 102:0-102:22 *) Definition CounterUsize : Counter_t usize := {| Counter_t_incr := counterUsize_incr; |}. (** [demo::use_counter]: - Source: 'tests/src/demo.rs', lines 109:0-109:59 *) + Source: 'tests/src/demo.rs', lines 110:0-110:59 *) Definition use_counter (T : Type) (counterInst : Counter_t T) (cnt : T) : result (usize * T) := counterInst.(Counter_t_incr) cnt diff --git a/tests/coq/hashmap/Hashmap_Funs.v b/tests/coq/hashmap/Hashmap_Funs.v index 1f2b2b22..778b9d56 100644 --- a/tests/coq/hashmap/Hashmap_Funs.v +++ b/tests/coq/hashmap/Hashmap_Funs.v @@ -11,12 +11,12 @@ Include Hashmap_Types. Module Hashmap_Funs. (** [hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 27:0-27:32 *) + Source: 'tests/src/hashmap.rs', lines 35:0-35:32 *) Definition hash_key (k : usize) : result usize := Ok k. (** [hashmap::{hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 50:4-56:5 *) + Source: 'tests/src/hashmap.rs', lines 58:4-64: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)) @@ -34,7 +34,7 @@ Fixpoint hashMap_allocate_slots_loop . (** [hashmap::{hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 50:4-50:76 *) + Source: 'tests/src/hashmap.rs', lines 58:4-58:76 *) Definition hashMap_allocate_slots (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n1 : usize) : result (alloc_vec_Vec (List_t T)) @@ -43,7 +43,7 @@ Definition hashMap_allocate_slots . (** [hashmap::{hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 59:4-63:13 *) + Source: 'tests/src/hashmap.rs', lines 67:4-71:13 *) Definition hashMap_new_with_capacity (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : @@ -62,13 +62,13 @@ Definition hashMap_new_with_capacity . (** [hashmap::{hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 75:4-75:24 *) + Source: 'tests/src/hashmap.rs', lines 83:4-83:24 *) Definition hashMap_new (T : Type) (n : nat) : result (HashMap_t T) := hashMap_new_with_capacity T n 32%usize 4%usize 5%usize . (** [hashmap::{hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 80:4-88:5 *) + Source: 'tests/src/hashmap.rs', lines 88:4-96: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)) @@ -91,7 +91,7 @@ Fixpoint hashMap_clear_loop . (** [hashmap::{hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 80:4-80:27 *) + Source: 'tests/src/hashmap.rs', lines 88:4-88: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; @@ -105,13 +105,13 @@ Definition hashMap_clear . (** [hashmap::{hashmap::HashMap}::len]: - Source: 'tests/src/hashmap.rs', lines 90:4-90:30 *) + Source: 'tests/src/hashmap.rs', lines 98:4-98:30 *) Definition hashMap_len (T : Type) (self : HashMap_t T) : result usize := Ok self.(hashMap_num_entries) . (** [hashmap::{hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 97:4-114:5 *) + Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) Fixpoint hashMap_insert_in_list_loop (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : result (bool * (List_t T)) @@ -133,7 +133,7 @@ Fixpoint hashMap_insert_in_list_loop . (** [hashmap::{hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 97:4-97:71 *) + Source: 'tests/src/hashmap.rs', lines 105:4-105:71 *) Definition hashMap_insert_in_list (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : result (bool * (List_t T)) @@ -142,7 +142,7 @@ Definition hashMap_insert_in_list . (** [hashmap::{hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 117:4-117:54 *) + Source: 'tests/src/hashmap.rs', lines 125:4-125:54 *) Definition hashMap_insert_no_resize (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) : result (HashMap_t T) @@ -180,7 +180,7 @@ Definition hashMap_insert_no_resize . (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 183:4-196:5 *) + Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) Fixpoint hashMap_move_elements_from_list_loop (T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) : result (HashMap_t T) @@ -198,7 +198,7 @@ Fixpoint hashMap_move_elements_from_list_loop . (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 183:4-183:72 *) + Source: 'tests/src/hashmap.rs', lines 191:4-191:72 *) Definition hashMap_move_elements_from_list (T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) : result (HashMap_t T) @@ -207,7 +207,7 @@ Definition hashMap_move_elements_from_list . (** [hashmap::{hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 171:4-180:5 *) + Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) Fixpoint hashMap_move_elements_loop (T : Type) (n : nat) (ntable : HashMap_t T) (slots : alloc_vec_Vec (List_t T)) (i : usize) : @@ -233,7 +233,7 @@ Fixpoint hashMap_move_elements_loop . (** [hashmap::{hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 171:4-171:95 *) + Source: 'tests/src/hashmap.rs', lines 179:4-179:95 *) Definition hashMap_move_elements (T : Type) (n : nat) (ntable : HashMap_t T) (slots : alloc_vec_Vec (List_t T)) (i : usize) : @@ -243,7 +243,7 @@ Definition hashMap_move_elements . (** [hashmap::{hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 140:4-140:28 *) + Source: 'tests/src/hashmap.rs', lines 148:4-148: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; @@ -275,7 +275,7 @@ Definition hashMap_try_resize . (** [hashmap::{hashmap::HashMap}::insert]: - Source: 'tests/src/hashmap.rs', lines 129:4-129:48 *) + Source: 'tests/src/hashmap.rs', lines 137:4-137:48 *) Definition hashMap_insert (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) : result (HashMap_t T) @@ -288,7 +288,7 @@ Definition hashMap_insert . (** [hashmap::{hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 206:4-219:5 *) + Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) Fixpoint hashMap_contains_key_in_list_loop (T : Type) (n : nat) (key : usize) (ls : List_t T) : result bool := match n with @@ -305,14 +305,14 @@ Fixpoint hashMap_contains_key_in_list_loop . (** [hashmap::{hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 206:4-206:68 *) + Source: 'tests/src/hashmap.rs', lines 214:4-214:68 *) Definition hashMap_contains_key_in_list (T : Type) (n : nat) (key : usize) (ls : List_t T) : result bool := hashMap_contains_key_in_list_loop T n key ls . (** [hashmap::{hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 199:4-199:49 *) + Source: 'tests/src/hashmap.rs', lines 207:4-207:49 *) Definition hashMap_contains_key (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result bool := hash <- hash_key key; @@ -326,7 +326,7 @@ Definition hashMap_contains_key . (** [hashmap::{hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 224:4-237:5 *) + Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) Fixpoint hashMap_get_in_list_loop (T : Type) (n : nat) (key : usize) (ls : List_t T) : result T := match n with @@ -341,14 +341,14 @@ Fixpoint hashMap_get_in_list_loop . (** [hashmap::{hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 224:4-224:70 *) + Source: 'tests/src/hashmap.rs', lines 232:4-232:70 *) Definition hashMap_get_in_list (T : Type) (n : nat) (key : usize) (ls : List_t T) : result T := hashMap_get_in_list_loop T n key ls . (** [hashmap::{hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 239:4-239:55 *) + Source: 'tests/src/hashmap.rs', lines 247:4-247:55 *) Definition hashMap_get (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result T := hash <- hash_key key; @@ -362,7 +362,7 @@ Definition hashMap_get . (** [hashmap::{hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 245:4-254:5 *) + Source: 'tests/src/hashmap.rs', lines 253:4-262: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))) @@ -388,7 +388,7 @@ Fixpoint hashMap_get_mut_in_list_loop . (** [hashmap::{hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 245:4-245:86 *) + Source: 'tests/src/hashmap.rs', lines 253:4-253:86 *) Definition hashMap_get_mut_in_list (T : Type) (n : nat) (ls : List_t T) (key : usize) : result (T * (T -> result (List_t T))) @@ -397,7 +397,7 @@ Definition hashMap_get_mut_in_list . (** [hashmap::{hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 257:4-257:67 *) + Source: 'tests/src/hashmap.rs', lines 265:4-265:67 *) Definition hashMap_get_mut (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result (T * (T -> result (HashMap_t T))) @@ -427,7 +427,7 @@ Definition hashMap_get_mut . (** [hashmap::{hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 265:4-291:5 *) + Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) Fixpoint hashMap_remove_from_list_loop (T : Type) (n : nat) (key : usize) (ls : List_t T) : result ((option T) * (List_t T)) @@ -455,7 +455,7 @@ Fixpoint hashMap_remove_from_list_loop . (** [hashmap::{hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:69 *) + Source: 'tests/src/hashmap.rs', lines 273:4-273:69 *) Definition hashMap_remove_from_list (T : Type) (n : nat) (key : usize) (ls : List_t T) : result ((option T) * (List_t T)) @@ -464,7 +464,7 @@ Definition hashMap_remove_from_list . (** [hashmap::{hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 294:4-294:52 *) + Source: 'tests/src/hashmap.rs', lines 302:4-302:52 *) Definition hashMap_remove (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result ((option T) * (HashMap_t T)) @@ -503,7 +503,7 @@ Definition hashMap_remove . (** [hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 315:0-315:10 *) + Source: 'tests/src/hashmap.rs', lines 323:0-323: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_Types.v b/tests/coq/hashmap/Hashmap_Types.v index 2a9f7ddb..8a8137d5 100644 --- a/tests/coq/hashmap/Hashmap_Types.v +++ b/tests/coq/hashmap/Hashmap_Types.v @@ -9,7 +9,7 @@ Local Open Scope Primitives_scope. Module Hashmap_Types. (** [hashmap::List] - Source: 'tests/src/hashmap.rs', lines 19:0-19:16 *) + Source: 'tests/src/hashmap.rs', lines 27:0-27:16 *) Inductive List_t (T : Type) := | List_Cons : usize -> T -> List_t T -> List_t T | List_Nil : List_t T @@ -19,7 +19,7 @@ Arguments List_Cons { _ }. Arguments List_Nil { _ }. (** [hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 35:0-35:21 *) + Source: 'tests/src/hashmap.rs', lines 43:0-43:21 *) Record HashMap_t (T : Type) := mkHashMap_t { hashMap_num_entries : usize; diff --git a/tests/coq/hashmap_on_disk/HashmapMain_Funs.v b/tests/coq/hashmap_on_disk/HashmapMain_Funs.v index facd84ea..f6467d5a 100644 --- a/tests/coq/hashmap_on_disk/HashmapMain_Funs.v +++ b/tests/coq/hashmap_on_disk/HashmapMain_Funs.v @@ -13,12 +13,12 @@ Include HashmapMain_FunsExternal. Module HashmapMain_Funs. (** [hashmap_main::hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 27:0-27:32 *) + Source: 'tests/src/hashmap.rs', lines 35:0-35:32 *) Definition hashmap_hash_key (k : usize) : result usize := Ok k. (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 50:4-56:5 *) + Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) Fixpoint hashmap_HashMap_allocate_slots_loop (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (n1 : usize) : @@ -37,7 +37,7 @@ Fixpoint hashmap_HashMap_allocate_slots_loop . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 50:4-50:76 *) + Source: 'tests/src/hashmap.rs', lines 58:4-58:76 *) Definition hashmap_HashMap_allocate_slots (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (n1 : usize) : @@ -47,7 +47,7 @@ Definition hashmap_HashMap_allocate_slots . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 59:4-63:13 *) + Source: 'tests/src/hashmap.rs', lines 67:4-71:13 *) Definition hashmap_HashMap_new_with_capacity (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : @@ -68,14 +68,14 @@ Definition hashmap_HashMap_new_with_capacity . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 75:4-75:24 *) + Source: 'tests/src/hashmap.rs', lines 83:4-83:24 *) Definition hashmap_HashMap_new (T : Type) (n : nat) : result (hashmap_HashMap_t T) := hashmap_HashMap_new_with_capacity T n 32%usize 4%usize 5%usize . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 80:4-88:5 *) + Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) Fixpoint hashmap_HashMap_clear_loop (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) : result (alloc_vec_Vec (hashmap_List_t T)) @@ -99,7 +99,7 @@ Fixpoint hashmap_HashMap_clear_loop . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 80:4-80:27 *) + Source: 'tests/src/hashmap.rs', lines 88:4-88:27 *) Definition hashmap_HashMap_clear (T : Type) (n : nat) (self : hashmap_HashMap_t T) : result (hashmap_HashMap_t T) @@ -115,14 +115,14 @@ Definition hashmap_HashMap_clear . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::len]: - Source: 'tests/src/hashmap.rs', lines 90:4-90:30 *) + Source: 'tests/src/hashmap.rs', lines 98:4-98:30 *) Definition hashmap_HashMap_len (T : Type) (self : hashmap_HashMap_t T) : result usize := Ok self.(hashmap_HashMap_num_entries) . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 97:4-114:5 *) + Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) Fixpoint hashmap_HashMap_insert_in_list_loop (T : Type) (n : nat) (key : usize) (value : T) (ls : hashmap_List_t T) : result (bool * (hashmap_List_t T)) @@ -145,7 +145,7 @@ Fixpoint hashmap_HashMap_insert_in_list_loop . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 97:4-97:71 *) + Source: 'tests/src/hashmap.rs', lines 105:4-105:71 *) Definition hashmap_HashMap_insert_in_list (T : Type) (n : nat) (key : usize) (value : T) (ls : hashmap_List_t T) : result (bool * (hashmap_List_t T)) @@ -154,7 +154,7 @@ Definition hashmap_HashMap_insert_in_list . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 117:4-117:54 *) + Source: 'tests/src/hashmap.rs', lines 125:4-125:54 *) Definition hashmap_HashMap_insert_no_resize (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) (value : T) : result (hashmap_HashMap_t T) @@ -194,7 +194,7 @@ Definition hashmap_HashMap_insert_no_resize . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 183:4-196:5 *) + Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) Fixpoint hashmap_HashMap_move_elements_from_list_loop (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) (ls : hashmap_List_t T) : result (hashmap_HashMap_t T) @@ -212,7 +212,7 @@ Fixpoint hashmap_HashMap_move_elements_from_list_loop . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 183:4-183:72 *) + Source: 'tests/src/hashmap.rs', lines 191:4-191:72 *) Definition hashmap_HashMap_move_elements_from_list (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) (ls : hashmap_List_t T) : result (hashmap_HashMap_t T) @@ -221,7 +221,7 @@ Definition hashmap_HashMap_move_elements_from_list . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 171:4-180:5 *) + Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) Fixpoint hashmap_HashMap_move_elements_loop (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) : @@ -248,7 +248,7 @@ Fixpoint hashmap_HashMap_move_elements_loop . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 171:4-171:95 *) + Source: 'tests/src/hashmap.rs', lines 179:4-179:95 *) Definition hashmap_HashMap_move_elements (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) : @@ -258,7 +258,7 @@ Definition hashmap_HashMap_move_elements . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 140:4-140:28 *) + Source: 'tests/src/hashmap.rs', lines 148:4-148:28 *) Definition hashmap_HashMap_try_resize (T : Type) (n : nat) (self : hashmap_HashMap_t T) : result (hashmap_HashMap_t T) @@ -295,7 +295,7 @@ Definition hashmap_HashMap_try_resize . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert]: - Source: 'tests/src/hashmap.rs', lines 129:4-129:48 *) + Source: 'tests/src/hashmap.rs', lines 137:4-137:48 *) Definition hashmap_HashMap_insert (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) (value : T) : result (hashmap_HashMap_t T) @@ -308,7 +308,7 @@ Definition hashmap_HashMap_insert . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 206:4-219:5 *) + Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) Fixpoint hashmap_HashMap_contains_key_in_list_loop (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result bool := match n with @@ -325,14 +325,14 @@ Fixpoint hashmap_HashMap_contains_key_in_list_loop . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 206:4-206:68 *) + Source: 'tests/src/hashmap.rs', lines 214:4-214:68 *) Definition hashmap_HashMap_contains_key_in_list (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result bool := hashmap_HashMap_contains_key_in_list_loop T n key ls . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 199:4-199:49 *) + Source: 'tests/src/hashmap.rs', lines 207:4-207:49 *) Definition hashmap_HashMap_contains_key (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : result bool @@ -348,7 +348,7 @@ Definition hashmap_HashMap_contains_key . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 224:4-237:5 *) + Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) Fixpoint hashmap_HashMap_get_in_list_loop (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result T := match n with @@ -365,14 +365,14 @@ Fixpoint hashmap_HashMap_get_in_list_loop . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 224:4-224:70 *) + Source: 'tests/src/hashmap.rs', lines 232:4-232:70 *) Definition hashmap_HashMap_get_in_list (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result T := hashmap_HashMap_get_in_list_loop T n key ls . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 239:4-239:55 *) + Source: 'tests/src/hashmap.rs', lines 247:4-247:55 *) Definition hashmap_HashMap_get (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : result T := hash <- hashmap_hash_key key; @@ -386,7 +386,7 @@ Definition hashmap_HashMap_get . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 245:4-254:5 *) + Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) Fixpoint hashmap_HashMap_get_mut_in_list_loop (T : Type) (n : nat) (ls : hashmap_List_t T) (key : usize) : result (T * (T -> result (hashmap_List_t T))) @@ -413,7 +413,7 @@ Fixpoint hashmap_HashMap_get_mut_in_list_loop . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 245:4-245:86 *) + Source: 'tests/src/hashmap.rs', lines 253:4-253:86 *) Definition hashmap_HashMap_get_mut_in_list (T : Type) (n : nat) (ls : hashmap_List_t T) (key : usize) : result (T * (T -> result (hashmap_List_t T))) @@ -422,7 +422,7 @@ Definition hashmap_HashMap_get_mut_in_list . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 257:4-257:67 *) + Source: 'tests/src/hashmap.rs', lines 265:4-265:67 *) Definition hashmap_HashMap_get_mut (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : result (T * (T -> result (hashmap_HashMap_t T))) @@ -453,7 +453,7 @@ Definition hashmap_HashMap_get_mut . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 265:4-291:5 *) + Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) Fixpoint hashmap_HashMap_remove_from_list_loop (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result ((option T) * (hashmap_List_t T)) @@ -482,7 +482,7 @@ Fixpoint hashmap_HashMap_remove_from_list_loop . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:69 *) + Source: 'tests/src/hashmap.rs', lines 273:4-273:69 *) Definition hashmap_HashMap_remove_from_list (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result ((option T) * (hashmap_List_t T)) @@ -491,7 +491,7 @@ Definition hashmap_HashMap_remove_from_list . (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 294:4-294:52 *) + Source: 'tests/src/hashmap.rs', lines 302:4-302:52 *) Definition hashmap_HashMap_remove (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : result ((option T) * (hashmap_HashMap_t T)) @@ -532,7 +532,7 @@ Definition hashmap_HashMap_remove . (** [hashmap_main::hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 315:0-315:10 *) + Source: 'tests/src/hashmap.rs', lines 323:0-323:10 *) Definition hashmap_test1 (n : nat) : result unit := hm <- hashmap_HashMap_new u64 n; hm1 <- hashmap_HashMap_insert u64 n hm 0%usize 42%u64; @@ -572,7 +572,7 @@ Definition hashmap_test1 (n : nat) : result unit := . (** [hashmap_main::insert_on_disk]: - Source: 'tests/src/hashmap_main.rs', lines 7:0-7:43 *) + Source: 'tests/src/hashmap_main.rs', lines 13:0-13:43 *) Definition insert_on_disk (n : nat) (key : usize) (value : u64) (st : state) : result (state * unit) := p <- hashmap_utils_deserialize st; @@ -582,7 +582,7 @@ Definition insert_on_disk . (** [hashmap_main::main]: - Source: 'tests/src/hashmap_main.rs', lines 16:0-16:13 *) + Source: 'tests/src/hashmap_main.rs', lines 22:0-22:13 *) Definition main : result unit := Ok tt. diff --git a/tests/coq/hashmap_on_disk/HashmapMain_FunsExternal_Template.v b/tests/coq/hashmap_on_disk/HashmapMain_FunsExternal_Template.v index f9e3f747..66835e8c 100644 --- a/tests/coq/hashmap_on_disk/HashmapMain_FunsExternal_Template.v +++ b/tests/coq/hashmap_on_disk/HashmapMain_FunsExternal_Template.v @@ -12,13 +12,13 @@ Include HashmapMain_Types. Module HashmapMain_FunsExternal_Template. (** [hashmap_main::hashmap_utils::deserialize]: - Source: 'tests/src/hashmap_utils.rs', lines 10:0-10:43 *) + Source: 'tests/src/hashmap_utils.rs', lines 11:0-11:43 *) Axiom hashmap_utils_deserialize : state -> result (state * (hashmap_HashMap_t u64)) . (** [hashmap_main::hashmap_utils::serialize]: - Source: 'tests/src/hashmap_utils.rs', lines 5:0-5:42 *) + Source: 'tests/src/hashmap_utils.rs', lines 6:0-6:42 *) Axiom hashmap_utils_serialize : hashmap_HashMap_t u64 -> state -> result (state * unit) . diff --git a/tests/coq/hashmap_on_disk/HashmapMain_Types.v b/tests/coq/hashmap_on_disk/HashmapMain_Types.v index 26029c9d..5656bd9c 100644 --- a/tests/coq/hashmap_on_disk/HashmapMain_Types.v +++ b/tests/coq/hashmap_on_disk/HashmapMain_Types.v @@ -11,7 +11,7 @@ Include HashmapMain_TypesExternal. Module HashmapMain_Types. (** [hashmap_main::hashmap::List] - Source: 'tests/src/hashmap.rs', lines 19:0-19:16 *) + Source: 'tests/src/hashmap.rs', lines 27:0-27:16 *) Inductive hashmap_List_t (T : Type) := | Hashmap_List_Cons : usize -> T -> hashmap_List_t T -> hashmap_List_t T | Hashmap_List_Nil : hashmap_List_t T @@ -21,7 +21,7 @@ Arguments Hashmap_List_Cons { _ }. Arguments Hashmap_List_Nil { _ }. (** [hashmap_main::hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 35:0-35:21 *) + Source: 'tests/src/hashmap.rs', lines 43:0-43:21 *) Record hashmap_HashMap_t (T : Type) := mkhashmap_HashMap_t { hashmap_HashMap_num_entries : usize; diff --git a/tests/coq/misc/Bitwise.v b/tests/coq/misc/Bitwise.v index 539141e5..d0dbfd51 100644 --- a/tests/coq/misc/Bitwise.v +++ b/tests/coq/misc/Bitwise.v @@ -9,29 +9,29 @@ Local Open Scope Primitives_scope. Module Bitwise. (** [bitwise::shift_u32]: - Source: 'tests/src/bitwise.rs', lines 3:0-3:31 *) + Source: 'tests/src/bitwise.rs', lines 4:0-4:31 *) Definition shift_u32 (a : u32) : result u32 := t <- u32_shr a 16%usize; u32_shl t 16%usize . (** [bitwise::shift_i32]: - Source: 'tests/src/bitwise.rs', lines 10:0-10:31 *) + Source: 'tests/src/bitwise.rs', lines 11:0-11:31 *) Definition shift_i32 (a : i32) : result i32 := t <- i32_shr a 16%isize; i32_shl t 16%isize . (** [bitwise::xor_u32]: - Source: 'tests/src/bitwise.rs', lines 17:0-17:37 *) + Source: 'tests/src/bitwise.rs', lines 18:0-18:37 *) Definition xor_u32 (a : u32) (b : u32) : result u32 := Ok (u32_xor a b). (** [bitwise::or_u32]: - Source: 'tests/src/bitwise.rs', lines 21:0-21:36 *) + Source: 'tests/src/bitwise.rs', lines 22:0-22:36 *) Definition or_u32 (a : u32) (b : u32) : result u32 := Ok (u32_or a b). (** [bitwise::and_u32]: - Source: 'tests/src/bitwise.rs', lines 25:0-25:37 *) + Source: 'tests/src/bitwise.rs', lines 26:0-26:37 *) Definition and_u32 (a : u32) (b : u32) : result u32 := Ok (u32_and a b). diff --git a/tests/coq/misc/Constants.v b/tests/coq/misc/Constants.v index a5aba663..c3ecdb83 100644 --- a/tests/coq/misc/Constants.v +++ b/tests/coq/misc/Constants.v @@ -9,37 +9,37 @@ Local Open Scope Primitives_scope. Module Constants. (** [constants::X0] - Source: 'tests/src/constants.rs', lines 5:0-5:17 *) + Source: 'tests/src/constants.rs', lines 7:0-7:17 *) Definition x0_body : result u32 := Ok 0%u32. Definition x0 : u32 := x0_body%global. (** [constants::X1] - Source: 'tests/src/constants.rs', lines 7:0-7:17 *) + Source: 'tests/src/constants.rs', lines 9:0-9:17 *) Definition x1_body : result u32 := Ok core_u32_max. Definition x1 : u32 := x1_body%global. (** [constants::X2] - Source: 'tests/src/constants.rs', lines 10:0-10:17 *) + Source: 'tests/src/constants.rs', lines 12:0-12:17 *) Definition x2_body : result u32 := Ok 3%u32. Definition x2 : u32 := x2_body%global. (** [constants::incr]: - Source: 'tests/src/constants.rs', lines 17:0-17:32 *) + Source: 'tests/src/constants.rs', lines 19:0-19:32 *) Definition incr (n : u32) : result u32 := u32_add n 1%u32. (** [constants::X3] - Source: 'tests/src/constants.rs', lines 15:0-15:17 *) + Source: 'tests/src/constants.rs', lines 17:0-17:17 *) Definition x3_body : result u32 := incr 32%u32. Definition x3 : u32 := x3_body%global. (** [constants::mk_pair0]: - Source: 'tests/src/constants.rs', lines 23:0-23:51 *) + Source: 'tests/src/constants.rs', lines 25:0-25:51 *) Definition mk_pair0 (x : u32) (y1 : u32) : result (u32 * u32) := Ok (x, y1). (** [constants::Pair] - Source: 'tests/src/constants.rs', lines 36:0-36:23 *) + Source: 'tests/src/constants.rs', lines 38:0-38:23 *) Record Pair_t (T1 T2 : Type) := mkPair_t { pair_x : T1; pair_y : T2; }. Arguments mkPair_t { _ _ }. @@ -47,130 +47,130 @@ Arguments pair_x { _ _ }. Arguments pair_y { _ _ }. (** [constants::mk_pair1]: - Source: 'tests/src/constants.rs', lines 27:0-27:55 *) + Source: 'tests/src/constants.rs', lines 29:0-29:55 *) Definition mk_pair1 (x : u32) (y1 : u32) : result (Pair_t u32 u32) := Ok {| pair_x := x; pair_y := y1 |} . (** [constants::P0] - Source: 'tests/src/constants.rs', lines 31:0-31:24 *) + Source: 'tests/src/constants.rs', lines 33:0-33:24 *) Definition p0_body : result (u32 * u32) := mk_pair0 0%u32 1%u32. Definition p0 : (u32 * u32) := p0_body%global. (** [constants::P1] - Source: 'tests/src/constants.rs', lines 32:0-32:28 *) + Source: 'tests/src/constants.rs', lines 34:0-34:28 *) Definition p1_body : result (Pair_t u32 u32) := mk_pair1 0%u32 1%u32. Definition p1 : Pair_t u32 u32 := p1_body%global. (** [constants::P2] - Source: 'tests/src/constants.rs', lines 33:0-33:24 *) + Source: 'tests/src/constants.rs', lines 35:0-35:24 *) Definition p2_body : result (u32 * u32) := Ok (0%u32, 1%u32). Definition p2 : (u32 * u32) := p2_body%global. (** [constants::P3] - Source: 'tests/src/constants.rs', lines 34:0-34:28 *) + Source: 'tests/src/constants.rs', lines 36:0-36:28 *) Definition p3_body : result (Pair_t u32 u32) := Ok {| pair_x := 0%u32; pair_y := 1%u32 |} . Definition p3 : Pair_t u32 u32 := p3_body%global. (** [constants::Wrap] - Source: 'tests/src/constants.rs', lines 49:0-49:18 *) + Source: 'tests/src/constants.rs', lines 51:0-51:18 *) Record Wrap_t (T : Type) := mkWrap_t { wrap_value : T; }. Arguments mkWrap_t { _ }. Arguments wrap_value { _ }. (** [constants::{constants::Wrap}::new]: - Source: 'tests/src/constants.rs', lines 54:4-54:41 *) + Source: 'tests/src/constants.rs', lines 56:4-56:41 *) Definition wrap_new (T : Type) (value : T) : result (Wrap_t T) := Ok {| wrap_value := value |} . (** [constants::Y] - Source: 'tests/src/constants.rs', lines 41:0-41:22 *) + Source: 'tests/src/constants.rs', lines 43:0-43:22 *) Definition y_body : result (Wrap_t i32) := wrap_new i32 2%i32. Definition y : Wrap_t i32 := y_body%global. (** [constants::unwrap_y]: - Source: 'tests/src/constants.rs', lines 43:0-43:30 *) + Source: 'tests/src/constants.rs', lines 45:0-45:30 *) Definition unwrap_y : result i32 := Ok y.(wrap_value). (** [constants::YVAL] - Source: 'tests/src/constants.rs', lines 47:0-47:19 *) + Source: 'tests/src/constants.rs', lines 49:0-49:19 *) Definition yval_body : result i32 := unwrap_y. Definition yval : i32 := yval_body%global. (** [constants::get_z1::Z1] - Source: 'tests/src/constants.rs', lines 62:4-62:17 *) + Source: 'tests/src/constants.rs', lines 64:4-64:17 *) Definition get_z1_z1_body : result i32 := Ok 3%i32. Definition get_z1_z1 : i32 := get_z1_z1_body%global. (** [constants::get_z1]: - Source: 'tests/src/constants.rs', lines 61:0-61:28 *) + Source: 'tests/src/constants.rs', lines 63:0-63:28 *) Definition get_z1 : result i32 := Ok get_z1_z1. (** [constants::add]: - Source: 'tests/src/constants.rs', lines 66:0-66:39 *) + Source: 'tests/src/constants.rs', lines 68:0-68:39 *) Definition add (a : i32) (b : i32) : result i32 := i32_add a b. (** [constants::Q1] - Source: 'tests/src/constants.rs', lines 74:0-74:17 *) + Source: 'tests/src/constants.rs', lines 76:0-76:17 *) Definition q1_body : result i32 := Ok 5%i32. Definition q1 : i32 := q1_body%global. (** [constants::Q2] - Source: 'tests/src/constants.rs', lines 75:0-75:17 *) + Source: 'tests/src/constants.rs', lines 77:0-77:17 *) Definition q2_body : result i32 := Ok q1. Definition q2 : i32 := q2_body%global. (** [constants::Q3] - Source: 'tests/src/constants.rs', lines 76:0-76:17 *) + Source: 'tests/src/constants.rs', lines 78:0-78:17 *) Definition q3_body : result i32 := add q2 3%i32. Definition q3 : i32 := q3_body%global. (** [constants::get_z2]: - Source: 'tests/src/constants.rs', lines 70:0-70:28 *) + Source: 'tests/src/constants.rs', lines 72:0-72:28 *) Definition get_z2 : result i32 := i <- get_z1; i1 <- add i q3; add q1 i1. (** [constants::S1] - Source: 'tests/src/constants.rs', lines 80:0-80:18 *) + Source: 'tests/src/constants.rs', lines 82:0-82:18 *) Definition s1_body : result u32 := Ok 6%u32. Definition s1 : u32 := s1_body%global. (** [constants::S2] - Source: 'tests/src/constants.rs', lines 81:0-81:18 *) + Source: 'tests/src/constants.rs', lines 83:0-83:18 *) Definition s2_body : result u32 := incr s1. Definition s2 : u32 := s2_body%global. (** [constants::S3] - Source: 'tests/src/constants.rs', lines 82:0-82:29 *) + Source: 'tests/src/constants.rs', lines 84:0-84:29 *) Definition s3_body : result (Pair_t u32 u32) := Ok p3. Definition s3 : Pair_t u32 u32 := s3_body%global. (** [constants::S4] - Source: 'tests/src/constants.rs', lines 83:0-83:29 *) + Source: 'tests/src/constants.rs', lines 85:0-85:29 *) Definition s4_body : result (Pair_t u32 u32) := mk_pair1 7%u32 8%u32. Definition s4 : Pair_t u32 u32 := s4_body%global. (** [constants::V] - Source: 'tests/src/constants.rs', lines 86:0-86:31 *) + Source: 'tests/src/constants.rs', lines 88:0-88:31 *) Record V_t (T : Type) (N : usize) := mkV_t { v_x : array T N; }. Arguments mkV_t { _ _ }. Arguments v_x { _ _ }. (** [constants::{constants::V#1}::LEN] - Source: 'tests/src/constants.rs', lines 91:4-91:24 *) + Source: 'tests/src/constants.rs', lines 93:4-93:24 *) Definition v_len_body (T : Type) (N : usize) : result usize := Ok N. Definition v_len (T : Type) (N : usize) : usize := (v_len_body T N)%global. (** [constants::use_v]: - Source: 'tests/src/constants.rs', lines 94:0-94:42 *) + Source: 'tests/src/constants.rs', lines 96:0-96:42 *) Definition use_v (T : Type) (N : usize) : result usize := Ok (v_len T N). diff --git a/tests/coq/misc/External_Funs.v b/tests/coq/misc/External_Funs.v index d0de80d9..18586012 100644 --- a/tests/coq/misc/External_Funs.v +++ b/tests/coq/misc/External_Funs.v @@ -20,14 +20,14 @@ Definition core_marker_CopyU32 : core_marker_Copy_t u32 := {| |}. (** [external::use_get]: - Source: 'tests/src/external.rs', lines 5:0-5:37 *) + Source: 'tests/src/external.rs', lines 8:0-8:37 *) Definition use_get (rc : core_cell_Cell_t u32) (st : state) : result (state * u32) := core_cell_Cell_get u32 core_marker_CopyU32 rc st . (** [external::incr]: - Source: 'tests/src/external.rs', lines 9:0-9:31 *) + Source: 'tests/src/external.rs', lines 12:0-12:31 *) Definition incr (rc : core_cell_Cell_t u32) (st : state) : result (state * (core_cell_Cell_t u32)) diff --git a/tests/coq/misc/Loops.v b/tests/coq/misc/Loops.v index bf0a8bc1..bc8708f4 100644 --- a/tests/coq/misc/Loops.v +++ b/tests/coq/misc/Loops.v @@ -9,7 +9,7 @@ Local Open Scope Primitives_scope. Module Loops. (** [loops::sum]: loop 0: - Source: 'tests/src/loops.rs', lines 4:0-14:1 *) + Source: 'tests/src/loops.rs', lines 7:0-17:1 *) Fixpoint sum_loop (n : nat) (max : u32) (i : u32) (s : u32) : result u32 := match n with | O => Fail_ OutOfFuel @@ -21,13 +21,13 @@ Fixpoint sum_loop (n : nat) (max : u32) (i : u32) (s : u32) : result u32 := . (** [loops::sum]: - Source: 'tests/src/loops.rs', lines 4:0-4:27 *) + Source: 'tests/src/loops.rs', lines 7:0-7:27 *) Definition sum (n : nat) (max : u32) : result u32 := sum_loop n max 0%u32 0%u32 . (** [loops::sum_with_mut_borrows]: loop 0: - Source: 'tests/src/loops.rs', lines 19:0-31:1 *) + Source: 'tests/src/loops.rs', lines 22:0-34:1 *) Fixpoint sum_with_mut_borrows_loop (n : nat) (max : u32) (i : u32) (s : u32) : result u32 := match n with @@ -43,13 +43,13 @@ Fixpoint sum_with_mut_borrows_loop . (** [loops::sum_with_mut_borrows]: - Source: 'tests/src/loops.rs', lines 19:0-19:44 *) + Source: 'tests/src/loops.rs', lines 22:0-22:44 *) Definition sum_with_mut_borrows (n : nat) (max : u32) : result u32 := sum_with_mut_borrows_loop n max 0%u32 0%u32 . (** [loops::sum_with_shared_borrows]: loop 0: - Source: 'tests/src/loops.rs', lines 34:0-48:1 *) + Source: 'tests/src/loops.rs', lines 37:0-51:1 *) Fixpoint sum_with_shared_borrows_loop (n : nat) (max : u32) (i : u32) (s : u32) : result u32 := match n with @@ -65,13 +65,13 @@ Fixpoint sum_with_shared_borrows_loop . (** [loops::sum_with_shared_borrows]: - Source: 'tests/src/loops.rs', lines 34:0-34:47 *) + Source: 'tests/src/loops.rs', lines 37:0-37:47 *) Definition sum_with_shared_borrows (n : nat) (max : u32) : result u32 := sum_with_shared_borrows_loop n max 0%u32 0%u32 . (** [loops::sum_array]: loop 0: - Source: 'tests/src/loops.rs', lines 50:0-58:1 *) + Source: 'tests/src/loops.rs', lines 53:0-61:1 *) Fixpoint sum_array_loop (N : usize) (n : nat) (a : array u32 N) (i : usize) (s : u32) : result u32 := match n with @@ -88,13 +88,13 @@ Fixpoint sum_array_loop . (** [loops::sum_array]: - Source: 'tests/src/loops.rs', lines 50:0-50:52 *) + Source: 'tests/src/loops.rs', lines 53:0-53:52 *) Definition sum_array (N : usize) (n : nat) (a : array u32 N) : result u32 := sum_array_loop N n a 0%usize 0%u32 . (** [loops::clear]: loop 0: - Source: 'tests/src/loops.rs', lines 62:0-68:1 *) + Source: 'tests/src/loops.rs', lines 65:0-71:1 *) Fixpoint clear_loop (n : nat) (v : alloc_vec_Vec u32) (i : usize) : result (alloc_vec_Vec u32) := match n with @@ -115,14 +115,14 @@ Fixpoint clear_loop . (** [loops::clear]: - Source: 'tests/src/loops.rs', lines 62:0-62:30 *) + Source: 'tests/src/loops.rs', lines 65:0-65:30 *) Definition clear (n : nat) (v : alloc_vec_Vec u32) : result (alloc_vec_Vec u32) := clear_loop n v 0%usize . (** [loops::List] - Source: 'tests/src/loops.rs', lines 70:0-70:16 *) + Source: 'tests/src/loops.rs', lines 73:0-73:16 *) Inductive List_t (T : Type) := | List_Cons : T -> List_t T -> List_t T | List_Nil : List_t T @@ -132,7 +132,7 @@ Arguments List_Cons { _ }. Arguments List_Nil { _ }. (** [loops::list_mem]: loop 0: - Source: 'tests/src/loops.rs', lines 76:0-85:1 *) + Source: 'tests/src/loops.rs', lines 79:0-88:1 *) Fixpoint list_mem_loop (n : nat) (x : u32) (ls : List_t u32) : result bool := match n with | O => Fail_ OutOfFuel @@ -145,13 +145,13 @@ Fixpoint list_mem_loop (n : nat) (x : u32) (ls : List_t u32) : result bool := . (** [loops::list_mem]: - Source: 'tests/src/loops.rs', lines 76:0-76:52 *) + Source: 'tests/src/loops.rs', lines 79:0-79:52 *) Definition list_mem (n : nat) (x : u32) (ls : List_t u32) : result bool := list_mem_loop n x ls . (** [loops::list_nth_mut_loop]: loop 0: - Source: 'tests/src/loops.rs', lines 88:0-98:1 *) + Source: 'tests/src/loops.rs', lines 91:0-101:1 *) Fixpoint list_nth_mut_loop_loop (T : Type) (n : nat) (ls : List_t T) (i : u32) : result (T * (T -> result (List_t T))) @@ -175,7 +175,7 @@ Fixpoint list_nth_mut_loop_loop . (** [loops::list_nth_mut_loop]: - Source: 'tests/src/loops.rs', lines 88:0-88:71 *) + Source: 'tests/src/loops.rs', lines 91:0-91:71 *) Definition list_nth_mut_loop (T : Type) (n : nat) (ls : List_t T) (i : u32) : result (T * (T -> result (List_t T))) @@ -184,7 +184,7 @@ Definition list_nth_mut_loop . (** [loops::list_nth_shared_loop]: loop 0: - Source: 'tests/src/loops.rs', lines 101:0-111:1 *) + Source: 'tests/src/loops.rs', lines 104:0-114:1 *) Fixpoint list_nth_shared_loop_loop (T : Type) (n : nat) (ls : List_t T) (i : u32) : result T := match n with @@ -201,14 +201,14 @@ Fixpoint list_nth_shared_loop_loop . (** [loops::list_nth_shared_loop]: - Source: 'tests/src/loops.rs', lines 101:0-101:66 *) + Source: 'tests/src/loops.rs', lines 104:0-104:66 *) Definition list_nth_shared_loop (T : Type) (n : nat) (ls : List_t T) (i : u32) : result T := list_nth_shared_loop_loop T n ls i . (** [loops::get_elem_mut]: loop 0: - Source: 'tests/src/loops.rs', lines 113:0-127:1 *) + Source: 'tests/src/loops.rs', lines 116:0-130:1 *) Fixpoint get_elem_mut_loop (n : nat) (x : usize) (ls : List_t usize) : result (usize * (usize -> result (List_t usize))) @@ -233,7 +233,7 @@ Fixpoint get_elem_mut_loop . (** [loops::get_elem_mut]: - Source: 'tests/src/loops.rs', lines 113:0-113:73 *) + Source: 'tests/src/loops.rs', lines 116:0-116:73 *) Definition get_elem_mut (n : nat) (slots : alloc_vec_Vec (List_t usize)) (x : usize) : result (usize * (usize -> result (alloc_vec_Vec (List_t usize)))) @@ -249,7 +249,7 @@ Definition get_elem_mut . (** [loops::get_elem_shared]: loop 0: - Source: 'tests/src/loops.rs', lines 129:0-143:1 *) + Source: 'tests/src/loops.rs', lines 132:0-146:1 *) Fixpoint get_elem_shared_loop (n : nat) (x : usize) (ls : List_t usize) : result usize := match n with @@ -263,7 +263,7 @@ Fixpoint get_elem_shared_loop . (** [loops::get_elem_shared]: - Source: 'tests/src/loops.rs', lines 129:0-129:68 *) + Source: 'tests/src/loops.rs', lines 132:0-132:68 *) Definition get_elem_shared (n : nat) (slots : alloc_vec_Vec (List_t usize)) (x : usize) : result usize @@ -275,7 +275,7 @@ Definition get_elem_shared . (** [loops::id_mut]: - Source: 'tests/src/loops.rs', lines 145:0-145:50 *) + Source: 'tests/src/loops.rs', lines 148:0-148:50 *) Definition id_mut (T : Type) (ls : List_t T) : result ((List_t T) * (List_t T -> result (List_t T))) @@ -284,12 +284,12 @@ Definition id_mut . (** [loops::id_shared]: - Source: 'tests/src/loops.rs', lines 149:0-149:45 *) + Source: 'tests/src/loops.rs', lines 152:0-152:45 *) Definition id_shared (T : Type) (ls : List_t T) : result (List_t T) := Ok ls. (** [loops::list_nth_mut_loop_with_id]: loop 0: - Source: 'tests/src/loops.rs', lines 154:0-165:1 *) + Source: 'tests/src/loops.rs', lines 157:0-168:1 *) Fixpoint list_nth_mut_loop_with_id_loop (T : Type) (n : nat) (i : u32) (ls : List_t T) : result (T * (T -> result (List_t T))) @@ -313,7 +313,7 @@ Fixpoint list_nth_mut_loop_with_id_loop . (** [loops::list_nth_mut_loop_with_id]: - Source: 'tests/src/loops.rs', lines 154:0-154:75 *) + Source: 'tests/src/loops.rs', lines 157:0-157:75 *) Definition list_nth_mut_loop_with_id (T : Type) (n : nat) (ls : List_t T) (i : u32) : result (T * (T -> result (List_t T))) @@ -327,7 +327,7 @@ Definition list_nth_mut_loop_with_id . (** [loops::list_nth_shared_loop_with_id]: loop 0: - Source: 'tests/src/loops.rs', lines 168:0-179:1 *) + Source: 'tests/src/loops.rs', lines 171:0-182:1 *) Fixpoint list_nth_shared_loop_with_id_loop (T : Type) (n : nat) (i : u32) (ls : List_t T) : result T := match n with @@ -345,14 +345,14 @@ Fixpoint list_nth_shared_loop_with_id_loop . (** [loops::list_nth_shared_loop_with_id]: - Source: 'tests/src/loops.rs', lines 168:0-168:70 *) + Source: 'tests/src/loops.rs', lines 171:0-171:70 *) Definition list_nth_shared_loop_with_id (T : Type) (n : nat) (ls : List_t T) (i : u32) : result T := ls1 <- id_shared T ls; list_nth_shared_loop_with_id_loop T n i ls1 . (** [loops::list_nth_mut_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 184:0-205:1 *) + Source: 'tests/src/loops.rs', lines 187:0-208:1 *) Fixpoint list_nth_mut_loop_pair_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T)) * (T -> result (List_t T))) @@ -386,7 +386,7 @@ Fixpoint list_nth_mut_loop_pair_loop . (** [loops::list_nth_mut_loop_pair]: - Source: 'tests/src/loops.rs', lines 184:0-188:27 *) + Source: 'tests/src/loops.rs', lines 187:0-191:27 *) Definition list_nth_mut_loop_pair (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T)) * (T -> result (List_t T))) @@ -395,7 +395,7 @@ Definition list_nth_mut_loop_pair . (** [loops::list_nth_shared_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 208:0-229:1 *) + Source: 'tests/src/loops.rs', lines 211:0-232:1 *) Fixpoint list_nth_shared_loop_pair_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result (T * T) @@ -419,7 +419,7 @@ Fixpoint list_nth_shared_loop_pair_loop . (** [loops::list_nth_shared_loop_pair]: - Source: 'tests/src/loops.rs', lines 208:0-212:19 *) + Source: 'tests/src/loops.rs', lines 211:0-215:19 *) Definition list_nth_shared_loop_pair (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result (T * T) @@ -428,7 +428,7 @@ Definition list_nth_shared_loop_pair . (** [loops::list_nth_mut_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 233:0-248:1 *) + Source: 'tests/src/loops.rs', lines 236:0-251:1 *) Fixpoint list_nth_mut_loop_pair_merge_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * ((T * T) -> result ((List_t T) * (List_t T)))) @@ -464,7 +464,7 @@ Fixpoint list_nth_mut_loop_pair_merge_loop . (** [loops::list_nth_mut_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 233:0-237:27 *) + Source: 'tests/src/loops.rs', lines 236:0-240:27 *) Definition list_nth_mut_loop_pair_merge (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * ((T * T) -> result ((List_t T) * (List_t T)))) @@ -473,7 +473,7 @@ Definition list_nth_mut_loop_pair_merge . (** [loops::list_nth_shared_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 251:0-266:1 *) + Source: 'tests/src/loops.rs', lines 254:0-269:1 *) Fixpoint list_nth_shared_loop_pair_merge_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result (T * T) @@ -498,7 +498,7 @@ Fixpoint list_nth_shared_loop_pair_merge_loop . (** [loops::list_nth_shared_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 251:0-255:19 *) + Source: 'tests/src/loops.rs', lines 254:0-258:19 *) Definition list_nth_shared_loop_pair_merge (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result (T * T) @@ -507,7 +507,7 @@ Definition list_nth_shared_loop_pair_merge . (** [loops::list_nth_mut_shared_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 269:0-284:1 *) + Source: 'tests/src/loops.rs', lines 272:0-287:1 *) Fixpoint list_nth_mut_shared_loop_pair_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -538,7 +538,7 @@ Fixpoint list_nth_mut_shared_loop_pair_loop . (** [loops::list_nth_mut_shared_loop_pair]: - Source: 'tests/src/loops.rs', lines 269:0-273:23 *) + Source: 'tests/src/loops.rs', lines 272:0-276:23 *) Definition list_nth_mut_shared_loop_pair (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -547,7 +547,7 @@ Definition list_nth_mut_shared_loop_pair . (** [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 288:0-303:1 *) + Source: 'tests/src/loops.rs', lines 291:0-306:1 *) Fixpoint list_nth_mut_shared_loop_pair_merge_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -578,7 +578,7 @@ Fixpoint list_nth_mut_shared_loop_pair_merge_loop . (** [loops::list_nth_mut_shared_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 288:0-292:23 *) + Source: 'tests/src/loops.rs', lines 291:0-295:23 *) Definition list_nth_mut_shared_loop_pair_merge (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -587,7 +587,7 @@ Definition list_nth_mut_shared_loop_pair_merge . (** [loops::list_nth_shared_mut_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 307:0-322:1 *) + Source: 'tests/src/loops.rs', lines 310:0-325:1 *) Fixpoint list_nth_shared_mut_loop_pair_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -618,7 +618,7 @@ Fixpoint list_nth_shared_mut_loop_pair_loop . (** [loops::list_nth_shared_mut_loop_pair]: - Source: 'tests/src/loops.rs', lines 307:0-311:23 *) + Source: 'tests/src/loops.rs', lines 310:0-314:23 *) Definition list_nth_shared_mut_loop_pair (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -627,7 +627,7 @@ Definition list_nth_shared_mut_loop_pair . (** [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 326:0-341:1 *) + Source: 'tests/src/loops.rs', lines 329:0-344:1 *) Fixpoint list_nth_shared_mut_loop_pair_merge_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -658,7 +658,7 @@ Fixpoint list_nth_shared_mut_loop_pair_merge_loop . (** [loops::list_nth_shared_mut_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 326:0-330:23 *) + Source: 'tests/src/loops.rs', lines 329:0-333:23 *) Definition list_nth_shared_mut_loop_pair_merge (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -667,7 +667,7 @@ Definition list_nth_shared_mut_loop_pair_merge . (** [loops::ignore_input_mut_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 345:0-349:1 *) + Source: 'tests/src/loops.rs', lines 348:0-352:1 *) Fixpoint ignore_input_mut_borrow_loop (n : nat) (i : u32) : result unit := match n with | O => Fail_ OutOfFuel @@ -679,14 +679,14 @@ Fixpoint ignore_input_mut_borrow_loop (n : nat) (i : u32) : result unit := . (** [loops::ignore_input_mut_borrow]: - Source: 'tests/src/loops.rs', lines 345:0-345:56 *) + Source: 'tests/src/loops.rs', lines 348:0-348:56 *) Definition ignore_input_mut_borrow (n : nat) (_a : u32) (i : u32) : result u32 := _ <- ignore_input_mut_borrow_loop n i; Ok _a . (** [loops::incr_ignore_input_mut_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 353:0-358:1 *) + Source: 'tests/src/loops.rs', lines 356:0-361:1 *) Fixpoint incr_ignore_input_mut_borrow_loop (n : nat) (i : u32) : result unit := match n with | O => Fail_ OutOfFuel @@ -698,14 +698,14 @@ Fixpoint incr_ignore_input_mut_borrow_loop (n : nat) (i : u32) : result unit := . (** [loops::incr_ignore_input_mut_borrow]: - Source: 'tests/src/loops.rs', lines 353:0-353:60 *) + Source: 'tests/src/loops.rs', lines 356:0-356:60 *) Definition incr_ignore_input_mut_borrow (n : nat) (a : u32) (i : u32) : result u32 := a1 <- u32_add a 1%u32; _ <- incr_ignore_input_mut_borrow_loop n i; Ok a1 . (** [loops::ignore_input_shared_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 362:0-366:1 *) + Source: 'tests/src/loops.rs', lines 365:0-369:1 *) Fixpoint ignore_input_shared_borrow_loop (n : nat) (i : u32) : result unit := match n with | O => Fail_ OutOfFuel @@ -717,7 +717,7 @@ Fixpoint ignore_input_shared_borrow_loop (n : nat) (i : u32) : result unit := . (** [loops::ignore_input_shared_borrow]: - Source: 'tests/src/loops.rs', lines 362:0-362:59 *) + Source: 'tests/src/loops.rs', lines 365:0-365:59 *) Definition ignore_input_shared_borrow (n : nat) (_a : u32) (i : u32) : result u32 := _ <- ignore_input_shared_borrow_loop n i; Ok _a diff --git a/tests/coq/misc/NoNestedBorrows.v b/tests/coq/misc/NoNestedBorrows.v index a83347a7..434b820c 100644 --- a/tests/coq/misc/NoNestedBorrows.v +++ b/tests/coq/misc/NoNestedBorrows.v @@ -9,7 +9,7 @@ Local Open Scope Primitives_scope. Module NoNestedBorrows. (** [no_nested_borrows::Pair] - Source: 'tests/src/no_nested_borrows.rs', lines 4:0-4:23 *) + Source: 'tests/src/no_nested_borrows.rs', lines 6:0-6:23 *) Record Pair_t (T1 T2 : Type) := mkPair_t { pair_x : T1; pair_y : T2; }. Arguments mkPair_t { _ _ }. @@ -17,7 +17,7 @@ Arguments pair_x { _ _ }. Arguments pair_y { _ _ }. (** [no_nested_borrows::List] - Source: 'tests/src/no_nested_borrows.rs', lines 9:0-9:16 *) + Source: 'tests/src/no_nested_borrows.rs', lines 11:0-11:16 *) Inductive List_t (T : Type) := | List_Cons : T -> List_t T -> List_t T | List_Nil : List_t T @@ -27,25 +27,25 @@ Arguments List_Cons { _ }. Arguments List_Nil { _ }. (** [no_nested_borrows::One] - Source: 'tests/src/no_nested_borrows.rs', lines 20:0-20:16 *) + Source: 'tests/src/no_nested_borrows.rs', lines 22:0-22:16 *) Inductive One_t (T1 : Type) := | One_One : T1 -> One_t T1. Arguments One_One { _ }. (** [no_nested_borrows::EmptyEnum] - Source: 'tests/src/no_nested_borrows.rs', lines 26:0-26:18 *) + Source: 'tests/src/no_nested_borrows.rs', lines 28:0-28:18 *) Inductive EmptyEnum_t := | EmptyEnum_Empty : EmptyEnum_t. (** [no_nested_borrows::Enum] - Source: 'tests/src/no_nested_borrows.rs', lines 32:0-32:13 *) + Source: 'tests/src/no_nested_borrows.rs', lines 34:0-34:13 *) Inductive Enum_t := | Enum_Variant1 : Enum_t | Enum_Variant2 : Enum_t. (** [no_nested_borrows::EmptyStruct] - Source: 'tests/src/no_nested_borrows.rs', lines 39:0-39:22 *) + Source: 'tests/src/no_nested_borrows.rs', lines 41:0-41:22 *) Definition EmptyStruct_t : Type := unit. (** [no_nested_borrows::Sum] - Source: 'tests/src/no_nested_borrows.rs', lines 41:0-41:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 43:0-43:20 *) Inductive Sum_t (T1 T2 : Type) := | Sum_Left : T1 -> Sum_t T1 T2 | Sum_Right : T2 -> Sum_t T1 T2 @@ -55,22 +55,22 @@ Arguments Sum_Left { _ _ }. Arguments Sum_Right { _ _ }. (** [no_nested_borrows::cast_u32_to_i32]: - Source: 'tests/src/no_nested_borrows.rs', lines 46:0-46:37 *) + Source: 'tests/src/no_nested_borrows.rs', lines 48:0-48:37 *) Definition cast_u32_to_i32 (x : u32) : result i32 := scalar_cast U32 I32 x. (** [no_nested_borrows::cast_bool_to_i32]: - Source: 'tests/src/no_nested_borrows.rs', lines 50:0-50:39 *) + Source: 'tests/src/no_nested_borrows.rs', lines 52:0-52:39 *) Definition cast_bool_to_i32 (x : bool) : result i32 := scalar_cast_bool I32 x. (** [no_nested_borrows::cast_bool_to_bool]: - Source: 'tests/src/no_nested_borrows.rs', lines 55:0-55:41 *) + Source: 'tests/src/no_nested_borrows.rs', lines 57:0-57:41 *) Definition cast_bool_to_bool (x : bool) : result bool := Ok x. (** [no_nested_borrows::test2]: - Source: 'tests/src/no_nested_borrows.rs', lines 60:0-60:14 *) + Source: 'tests/src/no_nested_borrows.rs', lines 62:0-62:14 *) Definition test2 : result unit := _ <- u32_add 23%u32 44%u32; Ok tt. @@ -78,13 +78,13 @@ Definition test2 : result unit := Check (test2 )%return. (** [no_nested_borrows::get_max]: - Source: 'tests/src/no_nested_borrows.rs', lines 72:0-72:37 *) + Source: 'tests/src/no_nested_borrows.rs', lines 74:0-74:37 *) Definition get_max (x : u32) (y : u32) : result u32 := if x s>= y then Ok x else Ok y . (** [no_nested_borrows::test3]: - Source: 'tests/src/no_nested_borrows.rs', lines 80:0-80:14 *) + Source: 'tests/src/no_nested_borrows.rs', lines 82:0-82:14 *) Definition test3 : result unit := x <- get_max 4%u32 3%u32; y <- get_max 10%u32 11%u32; @@ -96,7 +96,7 @@ Definition test3 : result unit := Check (test3 )%return. (** [no_nested_borrows::test_neg1]: - Source: 'tests/src/no_nested_borrows.rs', lines 87:0-87:18 *) + Source: 'tests/src/no_nested_borrows.rs', lines 89:0-89:18 *) Definition test_neg1 : result unit := y <- i32_neg 3%i32; if negb (y s= (-3)%i32) then Fail_ Failure else Ok tt . @@ -105,7 +105,7 @@ Definition test_neg1 : result unit := Check (test_neg1 )%return. (** [no_nested_borrows::refs_test1]: - Source: 'tests/src/no_nested_borrows.rs', lines 94:0-94:19 *) + Source: 'tests/src/no_nested_borrows.rs', lines 96:0-96:19 *) Definition refs_test1 : result unit := if negb (1%i32 s= 1%i32) then Fail_ Failure else Ok tt . @@ -114,7 +114,7 @@ Definition refs_test1 : result unit := Check (refs_test1 )%return. (** [no_nested_borrows::refs_test2]: - Source: 'tests/src/no_nested_borrows.rs', lines 105:0-105:19 *) + Source: 'tests/src/no_nested_borrows.rs', lines 107:0-107:19 *) Definition refs_test2 : result unit := if negb (2%i32 s= 2%i32) then Fail_ Failure @@ -131,7 +131,7 @@ Definition refs_test2 : result unit := Check (refs_test2 )%return. (** [no_nested_borrows::test_list1]: - Source: 'tests/src/no_nested_borrows.rs', lines 121:0-121:19 *) + Source: 'tests/src/no_nested_borrows.rs', lines 123:0-123:19 *) Definition test_list1 : result unit := Ok tt. @@ -139,7 +139,7 @@ Definition test_list1 : result unit := Check (test_list1 )%return. (** [no_nested_borrows::test_box1]: - Source: 'tests/src/no_nested_borrows.rs', lines 126:0-126:18 *) + Source: 'tests/src/no_nested_borrows.rs', lines 128:0-128:18 *) Definition test_box1 : result unit := p <- alloc_boxed_Box_deref_mut i32 0%i32; let (_, deref_mut_back) := p in @@ -152,24 +152,24 @@ Definition test_box1 : result unit := Check (test_box1 )%return. (** [no_nested_borrows::copy_int]: - Source: 'tests/src/no_nested_borrows.rs', lines 136:0-136:30 *) + Source: 'tests/src/no_nested_borrows.rs', lines 138:0-138:30 *) Definition copy_int (x : i32) : result i32 := Ok x. (** [no_nested_borrows::test_unreachable]: - Source: 'tests/src/no_nested_borrows.rs', lines 142:0-142:32 *) + Source: 'tests/src/no_nested_borrows.rs', lines 144:0-144:32 *) Definition test_unreachable (b : bool) : result unit := if b then Fail_ Failure else Ok tt . (** [no_nested_borrows::test_panic]: - Source: 'tests/src/no_nested_borrows.rs', lines 150:0-150:26 *) + Source: 'tests/src/no_nested_borrows.rs', lines 152:0-152:26 *) Definition test_panic (b : bool) : result unit := if b then Fail_ Failure else Ok tt . (** [no_nested_borrows::test_copy_int]: - Source: 'tests/src/no_nested_borrows.rs', lines 157:0-157:22 *) + Source: 'tests/src/no_nested_borrows.rs', lines 159:0-159:22 *) Definition test_copy_int : result unit := y <- copy_int 0%i32; if negb (0%i32 s= y) then Fail_ Failure else Ok tt . @@ -178,13 +178,13 @@ Definition test_copy_int : result unit := Check (test_copy_int )%return. (** [no_nested_borrows::is_cons]: - Source: 'tests/src/no_nested_borrows.rs', lines 164:0-164:38 *) + Source: 'tests/src/no_nested_borrows.rs', lines 166:0-166:38 *) Definition is_cons (T : Type) (l : List_t T) : result bool := match l with | List_Cons _ _ => Ok true | List_Nil => Ok false end . (** [no_nested_borrows::test_is_cons]: - Source: 'tests/src/no_nested_borrows.rs', lines 171:0-171:21 *) + Source: 'tests/src/no_nested_borrows.rs', lines 173:0-173:21 *) Definition test_is_cons : result unit := b <- is_cons i32 (List_Cons 0%i32 List_Nil); if negb b then Fail_ Failure else Ok tt @@ -194,13 +194,13 @@ Definition test_is_cons : result unit := Check (test_is_cons )%return. (** [no_nested_borrows::split_list]: - Source: 'tests/src/no_nested_borrows.rs', lines 177:0-177:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 179:0-179:48 *) Definition split_list (T : Type) (l : List_t T) : result (T * (List_t T)) := match l with | List_Cons hd tl => Ok (hd, tl) | List_Nil => Fail_ Failure end . (** [no_nested_borrows::test_split_list]: - Source: 'tests/src/no_nested_borrows.rs', lines 185:0-185:24 *) + Source: 'tests/src/no_nested_borrows.rs', lines 187:0-187:24 *) Definition test_split_list : result unit := p <- split_list i32 (List_Cons 0%i32 List_Nil); let (hd, _) := p in @@ -211,7 +211,7 @@ Definition test_split_list : result unit := Check (test_split_list )%return. (** [no_nested_borrows::choose]: - Source: 'tests/src/no_nested_borrows.rs', lines 192:0-192:70 *) + Source: 'tests/src/no_nested_borrows.rs', lines 194:0-194:70 *) Definition choose (T : Type) (b : bool) (x : T) (y : T) : result (T * (T -> result (T * T))) := if b @@ -220,7 +220,7 @@ Definition choose . (** [no_nested_borrows::choose_test]: - Source: 'tests/src/no_nested_borrows.rs', lines 200:0-200:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 202:0-202:20 *) Definition choose_test : result unit := p <- choose i32 true 0%i32 0%i32; let (z, choose_back) := p in @@ -239,18 +239,18 @@ Definition choose_test : result unit := Check (choose_test )%return. (** [no_nested_borrows::test_char]: - Source: 'tests/src/no_nested_borrows.rs', lines 212:0-212:26 *) + Source: 'tests/src/no_nested_borrows.rs', lines 214:0-214:26 *) Definition test_char : result char := Ok (char_of_byte Coq.Init.Byte.x61). (** [no_nested_borrows::Tree] - Source: 'tests/src/no_nested_borrows.rs', lines 217:0-217:16 *) + Source: 'tests/src/no_nested_borrows.rs', lines 219:0-219:16 *) Inductive Tree_t (T : Type) := | Tree_Leaf : T -> Tree_t T | Tree_Node : T -> NodeElem_t T -> Tree_t T -> Tree_t T (** [no_nested_borrows::NodeElem] - Source: 'tests/src/no_nested_borrows.rs', lines 222:0-222:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 224:0-224:20 *) with NodeElem_t (T : Type) := | NodeElem_Cons : Tree_t T -> NodeElem_t T -> NodeElem_t T | NodeElem_Nil : NodeElem_t T @@ -263,7 +263,7 @@ Arguments NodeElem_Cons { _ }. Arguments NodeElem_Nil { _ }. (** [no_nested_borrows::list_length]: - Source: 'tests/src/no_nested_borrows.rs', lines 257:0-257:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 259:0-259:48 *) Fixpoint list_length (T : Type) (l : List_t T) : result u32 := match l with | List_Cons _ l1 => i <- list_length T l1; u32_add 1%u32 i @@ -272,7 +272,7 @@ Fixpoint list_length (T : Type) (l : List_t T) : result u32 := . (** [no_nested_borrows::list_nth_shared]: - Source: 'tests/src/no_nested_borrows.rs', lines 265:0-265:62 *) + Source: 'tests/src/no_nested_borrows.rs', lines 267:0-267:62 *) Fixpoint list_nth_shared (T : Type) (l : List_t T) (i : u32) : result T := match l with | List_Cons x tl => @@ -284,7 +284,7 @@ Fixpoint list_nth_shared (T : Type) (l : List_t T) (i : u32) : result T := . (** [no_nested_borrows::list_nth_mut]: - Source: 'tests/src/no_nested_borrows.rs', lines 281:0-281:67 *) + Source: 'tests/src/no_nested_borrows.rs', lines 283:0-283:67 *) Fixpoint list_nth_mut (T : Type) (l : List_t T) (i : u32) : result (T * (T -> result (List_t T))) @@ -305,7 +305,7 @@ Fixpoint list_nth_mut . (** [no_nested_borrows::list_rev_aux]: - Source: 'tests/src/no_nested_borrows.rs', lines 297:0-297:63 *) + Source: 'tests/src/no_nested_borrows.rs', lines 299:0-299:63 *) Fixpoint list_rev_aux (T : Type) (li : List_t T) (lo : List_t T) : result (List_t T) := match li with @@ -315,14 +315,14 @@ Fixpoint list_rev_aux . (** [no_nested_borrows::list_rev]: - Source: 'tests/src/no_nested_borrows.rs', lines 311:0-311:42 *) + Source: 'tests/src/no_nested_borrows.rs', lines 313:0-313:42 *) Definition list_rev (T : Type) (l : List_t T) : result (List_t T) := let (li, _) := core_mem_replace (List_t T) l List_Nil in list_rev_aux T li List_Nil . (** [no_nested_borrows::test_list_functions]: - Source: 'tests/src/no_nested_borrows.rs', lines 316:0-316:28 *) + Source: 'tests/src/no_nested_borrows.rs', lines 318:0-318:28 *) Definition test_list_functions : result unit := let l := List_Cons 2%i32 List_Nil in let l1 := List_Cons 1%i32 l in @@ -361,7 +361,7 @@ Definition test_list_functions : result unit := Check (test_list_functions )%return. (** [no_nested_borrows::id_mut_pair1]: - Source: 'tests/src/no_nested_borrows.rs', lines 332:0-332:89 *) + Source: 'tests/src/no_nested_borrows.rs', lines 334:0-334:89 *) Definition id_mut_pair1 (T1 T2 : Type) (x : T1) (y : T2) : result ((T1 * T2) * ((T1 * T2) -> result (T1 * T2))) @@ -370,7 +370,7 @@ Definition id_mut_pair1 . (** [no_nested_borrows::id_mut_pair2]: - Source: 'tests/src/no_nested_borrows.rs', lines 336:0-336:88 *) + Source: 'tests/src/no_nested_borrows.rs', lines 338:0-338:88 *) Definition id_mut_pair2 (T1 T2 : Type) (p : (T1 * T2)) : result ((T1 * T2) * ((T1 * T2) -> result (T1 * T2))) @@ -379,7 +379,7 @@ Definition id_mut_pair2 . (** [no_nested_borrows::id_mut_pair3]: - Source: 'tests/src/no_nested_borrows.rs', lines 340:0-340:93 *) + Source: 'tests/src/no_nested_borrows.rs', lines 342:0-342:93 *) Definition id_mut_pair3 (T1 T2 : Type) (x : T1) (y : T2) : result ((T1 * T2) * (T1 -> result T1) * (T2 -> result T2)) @@ -388,7 +388,7 @@ Definition id_mut_pair3 . (** [no_nested_borrows::id_mut_pair4]: - Source: 'tests/src/no_nested_borrows.rs', lines 344:0-344:92 *) + Source: 'tests/src/no_nested_borrows.rs', lines 346:0-346:92 *) Definition id_mut_pair4 (T1 T2 : Type) (p : (T1 * T2)) : result ((T1 * T2) * (T1 -> result T1) * (T2 -> result T2)) @@ -397,7 +397,7 @@ Definition id_mut_pair4 . (** [no_nested_borrows::StructWithTuple] - Source: 'tests/src/no_nested_borrows.rs', lines 351:0-351:34 *) + Source: 'tests/src/no_nested_borrows.rs', lines 353:0-353:34 *) Record StructWithTuple_t (T1 T2 : Type) := mkStructWithTuple_t { structWithTuple_p : (T1 * T2); @@ -408,25 +408,25 @@ Arguments mkStructWithTuple_t { _ _ }. Arguments structWithTuple_p { _ _ }. (** [no_nested_borrows::new_tuple1]: - Source: 'tests/src/no_nested_borrows.rs', lines 355:0-355:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 357:0-357:48 *) Definition new_tuple1 : result (StructWithTuple_t u32 u32) := Ok {| structWithTuple_p := (1%u32, 2%u32) |} . (** [no_nested_borrows::new_tuple2]: - Source: 'tests/src/no_nested_borrows.rs', lines 359:0-359:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 361:0-361:48 *) Definition new_tuple2 : result (StructWithTuple_t i16 i16) := Ok {| structWithTuple_p := (1%i16, 2%i16) |} . (** [no_nested_borrows::new_tuple3]: - Source: 'tests/src/no_nested_borrows.rs', lines 363:0-363:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 365:0-365:48 *) Definition new_tuple3 : result (StructWithTuple_t u64 i64) := Ok {| structWithTuple_p := (1%u64, 2%i64) |} . (** [no_nested_borrows::StructWithPair] - Source: 'tests/src/no_nested_borrows.rs', lines 368:0-368:33 *) + Source: 'tests/src/no_nested_borrows.rs', lines 370:0-370:33 *) Record StructWithPair_t (T1 T2 : Type) := mkStructWithPair_t { structWithPair_p : Pair_t T1 T2; @@ -437,13 +437,13 @@ Arguments mkStructWithPair_t { _ _ }. Arguments structWithPair_p { _ _ }. (** [no_nested_borrows::new_pair1]: - Source: 'tests/src/no_nested_borrows.rs', lines 372:0-372:46 *) + Source: 'tests/src/no_nested_borrows.rs', lines 374:0-374:46 *) Definition new_pair1 : result (StructWithPair_t u32 u32) := Ok {| structWithPair_p := {| pair_x := 1%u32; pair_y := 2%u32 |} |} . (** [no_nested_borrows::test_constants]: - Source: 'tests/src/no_nested_borrows.rs', lines 380:0-380:23 *) + Source: 'tests/src/no_nested_borrows.rs', lines 382:0-382:23 *) Definition test_constants : result unit := swt <- new_tuple1; let (i, _) := swt.(structWithTuple_p) in @@ -470,7 +470,7 @@ Definition test_constants : result unit := Check (test_constants )%return. (** [no_nested_borrows::test_weird_borrows1]: - Source: 'tests/src/no_nested_borrows.rs', lines 389:0-389:28 *) + Source: 'tests/src/no_nested_borrows.rs', lines 391:0-391:28 *) Definition test_weird_borrows1 : result unit := Ok tt. @@ -478,78 +478,78 @@ Definition test_weird_borrows1 : result unit := Check (test_weird_borrows1 )%return. (** [no_nested_borrows::test_mem_replace]: - Source: 'tests/src/no_nested_borrows.rs', lines 399:0-399:37 *) + Source: 'tests/src/no_nested_borrows.rs', lines 401:0-401:37 *) Definition test_mem_replace (px : u32) : result u32 := let (y, _) := core_mem_replace u32 px 1%u32 in if negb (y s= 0%u32) then Fail_ Failure else Ok 2%u32 . (** [no_nested_borrows::test_shared_borrow_bool1]: - Source: 'tests/src/no_nested_borrows.rs', lines 406:0-406:47 *) + Source: 'tests/src/no_nested_borrows.rs', lines 408:0-408:47 *) Definition test_shared_borrow_bool1 (b : bool) : result u32 := if b then Ok 0%u32 else Ok 1%u32 . (** [no_nested_borrows::test_shared_borrow_bool2]: - Source: 'tests/src/no_nested_borrows.rs', lines 419:0-419:40 *) + Source: 'tests/src/no_nested_borrows.rs', lines 421:0-421:40 *) Definition test_shared_borrow_bool2 : result u32 := Ok 0%u32. (** [no_nested_borrows::test_shared_borrow_enum1]: - Source: 'tests/src/no_nested_borrows.rs', lines 434:0-434:52 *) + Source: 'tests/src/no_nested_borrows.rs', lines 436:0-436:52 *) Definition test_shared_borrow_enum1 (l : List_t u32) : result u32 := match l with | List_Cons _ _ => Ok 1%u32 | List_Nil => Ok 0%u32 end . (** [no_nested_borrows::test_shared_borrow_enum2]: - Source: 'tests/src/no_nested_borrows.rs', lines 446:0-446:40 *) + Source: 'tests/src/no_nested_borrows.rs', lines 448:0-448:40 *) Definition test_shared_borrow_enum2 : result u32 := Ok 0%u32. (** [no_nested_borrows::incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 457:0-457:24 *) + Source: 'tests/src/no_nested_borrows.rs', lines 459:0-459:24 *) Definition incr (x : u32) : result u32 := u32_add x 1%u32. (** [no_nested_borrows::call_incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 461:0-461:35 *) + Source: 'tests/src/no_nested_borrows.rs', lines 463:0-463:35 *) Definition call_incr (x : u32) : result u32 := incr x. (** [no_nested_borrows::read_then_incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 466:0-466:41 *) + Source: 'tests/src/no_nested_borrows.rs', lines 468:0-468:41 *) Definition read_then_incr (x : u32) : result (u32 * u32) := x1 <- u32_add x 1%u32; Ok (x, x1) . (** [no_nested_borrows::Tuple] - Source: 'tests/src/no_nested_borrows.rs', lines 472:0-472:24 *) + Source: 'tests/src/no_nested_borrows.rs', lines 474:0-474:24 *) Definition Tuple_t (T1 T2 : Type) : Type := T1 * T2. (** [no_nested_borrows::use_tuple_struct]: - Source: 'tests/src/no_nested_borrows.rs', lines 474:0-474:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 476:0-476:48 *) Definition use_tuple_struct (x : Tuple_t u32 u32) : result (Tuple_t u32 u32) := let (_, i) := x in Ok (1%u32, i) . (** [no_nested_borrows::create_tuple_struct]: - Source: 'tests/src/no_nested_borrows.rs', lines 478:0-478:61 *) + Source: 'tests/src/no_nested_borrows.rs', lines 480:0-480:61 *) Definition create_tuple_struct (x : u32) (y : u64) : result (Tuple_t u32 u64) := Ok (x, y) . (** [no_nested_borrows::IdType] - Source: 'tests/src/no_nested_borrows.rs', lines 483:0-483:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 485:0-485:20 *) Definition IdType_t (T : Type) : Type := T. (** [no_nested_borrows::use_id_type]: - Source: 'tests/src/no_nested_borrows.rs', lines 485:0-485:40 *) + Source: 'tests/src/no_nested_borrows.rs', lines 487:0-487:40 *) Definition use_id_type (T : Type) (x : IdType_t T) : result T := Ok x. (** [no_nested_borrows::create_id_type]: - Source: 'tests/src/no_nested_borrows.rs', lines 489:0-489:43 *) + Source: 'tests/src/no_nested_borrows.rs', lines 491:0-491:43 *) Definition create_id_type (T : Type) (x : T) : result (IdType_t T) := Ok x. diff --git a/tests/coq/misc/Paper.v b/tests/coq/misc/Paper.v index 661a8cc2..21e86542 100644 --- a/tests/coq/misc/Paper.v +++ b/tests/coq/misc/Paper.v @@ -9,12 +9,12 @@ Local Open Scope Primitives_scope. Module Paper. (** [paper::ref_incr]: - Source: 'tests/src/paper.rs', lines 4:0-4:28 *) + Source: 'tests/src/paper.rs', lines 6:0-6:28 *) Definition ref_incr (x : i32) : result i32 := i32_add x 1%i32. (** [paper::test_incr]: - Source: 'tests/src/paper.rs', lines 8:0-8:18 *) + Source: 'tests/src/paper.rs', lines 10:0-10:18 *) Definition test_incr : result unit := x <- ref_incr 0%i32; if negb (x s= 1%i32) then Fail_ Failure else Ok tt . @@ -23,7 +23,7 @@ Definition test_incr : result unit := Check (test_incr )%return. (** [paper::choose]: - Source: 'tests/src/paper.rs', lines 15:0-15:70 *) + Source: 'tests/src/paper.rs', lines 17:0-17:70 *) Definition choose (T : Type) (b : bool) (x : T) (y : T) : result (T * (T -> result (T * T))) := if b @@ -32,7 +32,7 @@ Definition choose . (** [paper::test_choose]: - Source: 'tests/src/paper.rs', lines 23:0-23:20 *) + Source: 'tests/src/paper.rs', lines 25:0-25:20 *) Definition test_choose : result unit := p <- choose i32 true 0%i32 0%i32; let (z, choose_back) := p in @@ -51,7 +51,7 @@ Definition test_choose : result unit := Check (test_choose )%return. (** [paper::List] - Source: 'tests/src/paper.rs', lines 35:0-35:16 *) + Source: 'tests/src/paper.rs', lines 37:0-37:16 *) Inductive List_t (T : Type) := | List_Cons : T -> List_t T -> List_t T | List_Nil : List_t T @@ -61,7 +61,7 @@ Arguments List_Cons { _ }. Arguments List_Nil { _ }. (** [paper::list_nth_mut]: - Source: 'tests/src/paper.rs', lines 42:0-42:67 *) + Source: 'tests/src/paper.rs', lines 44:0-44:67 *) Fixpoint list_nth_mut (T : Type) (l : List_t T) (i : u32) : result (T * (T -> result (List_t T))) @@ -82,7 +82,7 @@ Fixpoint list_nth_mut . (** [paper::sum]: - Source: 'tests/src/paper.rs', lines 57:0-57:32 *) + Source: 'tests/src/paper.rs', lines 59:0-59:32 *) Fixpoint sum (l : List_t i32) : result i32 := match l with | List_Cons x tl => i <- sum tl; i32_add x i @@ -91,7 +91,7 @@ Fixpoint sum (l : List_t i32) : result i32 := . (** [paper::test_nth]: - Source: 'tests/src/paper.rs', lines 68:0-68:17 *) + Source: 'tests/src/paper.rs', lines 70:0-70:17 *) Definition test_nth : result unit := let l := List_Cons 3%i32 List_Nil in let l1 := List_Cons 2%i32 l in @@ -107,7 +107,7 @@ Definition test_nth : result unit := Check (test_nth )%return. (** [paper::call_choose]: - Source: 'tests/src/paper.rs', lines 76:0-76:44 *) + Source: 'tests/src/paper.rs', lines 78:0-78:44 *) Definition call_choose (p : (u32 * u32)) : result u32 := let (px, py) := p in p1 <- choose u32 true px py; diff --git a/tests/coq/misc/PoloniusList.v b/tests/coq/misc/PoloniusList.v index c89f9655..91cfcdb7 100644 --- a/tests/coq/misc/PoloniusList.v +++ b/tests/coq/misc/PoloniusList.v @@ -9,7 +9,7 @@ Local Open Scope Primitives_scope. Module PoloniusList. (** [polonius_list::List] - Source: 'tests/src/polonius_list.rs', lines 3:0-3:16 *) + Source: 'tests/src/polonius_list.rs', lines 5:0-5:16 *) Inductive List_t (T : Type) := | List_Cons : T -> List_t T -> List_t T | List_Nil : List_t T @@ -19,7 +19,7 @@ Arguments List_Cons { _ }. Arguments List_Nil { _ }. (** [polonius_list::get_list_at_x]: - Source: 'tests/src/polonius_list.rs', lines 13:0-13:76 *) + Source: 'tests/src/polonius_list.rs', lines 15:0-15:76 *) Fixpoint get_list_at_x (ls : List_t u32) (x : u32) : result ((List_t u32) * (List_t u32 -> result (List_t u32))) diff --git a/tests/coq/traits/Traits.v b/tests/coq/traits/Traits.v index e583c2be..ad1be7ef 100644 --- a/tests/coq/traits/Traits.v +++ b/tests/coq/traits/Traits.v @@ -9,7 +9,7 @@ Local Open Scope Primitives_scope. Module Traits. (** Trait declaration: [traits::BoolTrait] - Source: 'tests/src/traits.rs', lines 1:0-1:19 *) + Source: 'tests/src/traits.rs', lines 2:0-2:19 *) Record BoolTrait_t (Self : Type) := mkBoolTrait_t { BoolTrait_t_get_bool : Self -> result bool; }. @@ -18,59 +18,59 @@ Arguments mkBoolTrait_t { _ }. Arguments BoolTrait_t_get_bool { _ }. (** [traits::{(traits::BoolTrait for bool)}::get_bool]: - Source: 'tests/src/traits.rs', lines 12:4-12:30 *) + Source: 'tests/src/traits.rs', lines 13:4-13:30 *) Definition boolTraitBool_get_bool (self : bool) : result bool := Ok self. (** Trait implementation: [traits::{(traits::BoolTrait for bool)}] - Source: 'tests/src/traits.rs', lines 11:0-11:23 *) + Source: 'tests/src/traits.rs', lines 12:0-12:23 *) Definition BoolTraitBool : BoolTrait_t bool := {| BoolTrait_t_get_bool := boolTraitBool_get_bool; |}. (** [traits::BoolTrait::ret_true]: - Source: 'tests/src/traits.rs', lines 6:4-6:30 *) + Source: 'tests/src/traits.rs', lines 7:4-7:30 *) Definition boolTrait_ret_true {Self : Type} (self_clause : BoolTrait_t Self) (self : Self) : result bool := Ok true . (** [traits::test_bool_trait_bool]: - Source: 'tests/src/traits.rs', lines 17:0-17:44 *) + Source: 'tests/src/traits.rs', lines 18:0-18:44 *) Definition test_bool_trait_bool (x : bool) : result bool := b <- boolTraitBool_get_bool x; if b then boolTrait_ret_true BoolTraitBool x else Ok false . (** [traits::{(traits::BoolTrait for core::option::Option)#1}::get_bool]: - Source: 'tests/src/traits.rs', lines 23:4-23:30 *) + Source: 'tests/src/traits.rs', lines 24:4-24:30 *) Definition boolTraitOption_get_bool (T : Type) (self : option T) : result bool := match self with | None => Ok false | Some _ => Ok true end . (** Trait implementation: [traits::{(traits::BoolTrait for core::option::Option)#1}] - Source: 'tests/src/traits.rs', lines 22:0-22:31 *) + Source: 'tests/src/traits.rs', lines 23:0-23:31 *) Definition BoolTraitOption (T : Type) : BoolTrait_t (option T) := {| BoolTrait_t_get_bool := boolTraitOption_get_bool T; |}. (** [traits::test_bool_trait_option]: - Source: 'tests/src/traits.rs', lines 31:0-31:54 *) + Source: 'tests/src/traits.rs', lines 32:0-32:54 *) Definition test_bool_trait_option (T : Type) (x : option T) : result bool := b <- boolTraitOption_get_bool T x; if b then boolTrait_ret_true (BoolTraitOption T) x else Ok false . (** [traits::test_bool_trait]: - Source: 'tests/src/traits.rs', lines 35:0-35:50 *) + Source: 'tests/src/traits.rs', lines 36:0-36:50 *) Definition test_bool_trait (T : Type) (boolTraitInst : BoolTrait_t T) (x : T) : result bool := boolTraitInst.(BoolTrait_t_get_bool) x . (** Trait declaration: [traits::ToU64] - Source: 'tests/src/traits.rs', lines 39:0-39:15 *) + Source: 'tests/src/traits.rs', lines 40:0-40:15 *) Record ToU64_t (Self : Type) := mkToU64_t { ToU64_t_to_u64 : Self -> result u64; }. @@ -79,16 +79,16 @@ Arguments mkToU64_t { _ }. Arguments ToU64_t_to_u64 { _ }. (** [traits::{(traits::ToU64 for u64)#2}::to_u64]: - Source: 'tests/src/traits.rs', lines 44:4-44:26 *) + Source: 'tests/src/traits.rs', lines 45:4-45:26 *) Definition toU64U64_to_u64 (self : u64) : result u64 := Ok self. (** Trait implementation: [traits::{(traits::ToU64 for u64)#2}] - Source: 'tests/src/traits.rs', lines 43:0-43:18 *) + Source: 'tests/src/traits.rs', lines 44:0-44:18 *) Definition ToU64U64 : ToU64_t u64 := {| ToU64_t_to_u64 := toU64U64_to_u64; |}. (** [traits::{(traits::ToU64 for (A, A))#3}::to_u64]: - Source: 'tests/src/traits.rs', lines 50:4-50:26 *) + Source: 'tests/src/traits.rs', lines 51:4-51:26 *) Definition toU64Pair_to_u64 (A : Type) (toU64Inst : ToU64_t A) (self : (A * A)) : result u64 := let (t, t1) := self in @@ -98,65 +98,65 @@ Definition toU64Pair_to_u64 . (** Trait implementation: [traits::{(traits::ToU64 for (A, A))#3}] - Source: 'tests/src/traits.rs', lines 49:0-49:31 *) + Source: 'tests/src/traits.rs', lines 50:0-50:31 *) Definition ToU64Pair (A : Type) (toU64Inst : ToU64_t A) : ToU64_t (A * A) := {| ToU64_t_to_u64 := toU64Pair_to_u64 A toU64Inst; |}. (** [traits::f]: - Source: 'tests/src/traits.rs', lines 55:0-55:36 *) + Source: 'tests/src/traits.rs', lines 56:0-56:36 *) Definition f (T : Type) (toU64Inst : ToU64_t T) (x : (T * T)) : result u64 := toU64Pair_to_u64 T toU64Inst x . (** [traits::g]: - Source: 'tests/src/traits.rs', lines 59:0-61:18 *) + Source: 'tests/src/traits.rs', lines 60:0-62:18 *) Definition g (T : Type) (toU64PairInst : ToU64_t (T * T)) (x : (T * T)) : result u64 := toU64PairInst.(ToU64_t_to_u64) x . (** [traits::h0]: - Source: 'tests/src/traits.rs', lines 66:0-66:24 *) + Source: 'tests/src/traits.rs', lines 67:0-67:24 *) Definition h0 (x : u64) : result u64 := toU64U64_to_u64 x. (** [traits::Wrapper] - Source: 'tests/src/traits.rs', lines 70:0-70:21 *) + Source: 'tests/src/traits.rs', lines 71:0-71:21 *) Record Wrapper_t (T : Type) := mkWrapper_t { wrapper_x : T; }. Arguments mkWrapper_t { _ }. Arguments wrapper_x { _ }. (** [traits::{(traits::ToU64 for traits::Wrapper)#4}::to_u64]: - Source: 'tests/src/traits.rs', lines 75:4-75:26 *) + Source: 'tests/src/traits.rs', lines 76:4-76:26 *) Definition toU64traitsWrapper_to_u64 (T : Type) (toU64Inst : ToU64_t T) (self : Wrapper_t T) : result u64 := toU64Inst.(ToU64_t_to_u64) self.(wrapper_x) . (** Trait implementation: [traits::{(traits::ToU64 for traits::Wrapper)#4}] - Source: 'tests/src/traits.rs', lines 74:0-74:35 *) + Source: 'tests/src/traits.rs', lines 75:0-75:35 *) Definition ToU64traitsWrapper (T : Type) (toU64Inst : ToU64_t T) : ToU64_t (Wrapper_t T) := {| ToU64_t_to_u64 := toU64traitsWrapper_to_u64 T toU64Inst; |}. (** [traits::h1]: - Source: 'tests/src/traits.rs', lines 80:0-80:33 *) + Source: 'tests/src/traits.rs', lines 81:0-81:33 *) Definition h1 (x : Wrapper_t u64) : result u64 := toU64traitsWrapper_to_u64 u64 ToU64U64 x . (** [traits::h2]: - Source: 'tests/src/traits.rs', lines 84:0-84:41 *) + Source: 'tests/src/traits.rs', lines 85:0-85:41 *) Definition h2 (T : Type) (toU64Inst : ToU64_t T) (x : Wrapper_t T) : result u64 := toU64traitsWrapper_to_u64 T toU64Inst x . (** Trait declaration: [traits::ToType] - Source: 'tests/src/traits.rs', lines 88:0-88:19 *) + Source: 'tests/src/traits.rs', lines 89:0-89:19 *) Record ToType_t (Self T : Type) := mkToType_t { ToType_t_to_type : Self -> result T; }. @@ -165,19 +165,19 @@ Arguments mkToType_t { _ _ }. Arguments ToType_t_to_type { _ _ }. (** [traits::{(traits::ToType for u64)#5}::to_type]: - Source: 'tests/src/traits.rs', lines 93:4-93:28 *) + Source: 'tests/src/traits.rs', lines 94:4-94:28 *) Definition toTypeU64Bool_to_type (self : u64) : result bool := Ok (self s> 0%u64) . (** Trait implementation: [traits::{(traits::ToType for u64)#5}] - Source: 'tests/src/traits.rs', lines 92:0-92:25 *) + Source: 'tests/src/traits.rs', lines 93:0-93:25 *) Definition ToTypeU64Bool : ToType_t u64 bool := {| ToType_t_to_type := toTypeU64Bool_to_type; |}. (** Trait declaration: [traits::OfType] - Source: 'tests/src/traits.rs', lines 98:0-98:16 *) + Source: 'tests/src/traits.rs', lines 99:0-99:16 *) Record OfType_t (Self : Type) := mkOfType_t { OfType_t_of_type : forall (T : Type) (toTypeInst : ToType_t T Self), T -> result Self; @@ -187,7 +187,7 @@ Arguments mkOfType_t { _ }. Arguments OfType_t_of_type { _ }. (** [traits::h3]: - Source: 'tests/src/traits.rs', lines 104:0-104:50 *) + Source: 'tests/src/traits.rs', lines 105:0-105:50 *) Definition h3 (T1 T2 : Type) (ofTypeInst : OfType_t T1) (toTypeInst : ToType_t T2 T1) (y : T2) : @@ -197,7 +197,7 @@ Definition h3 . (** Trait declaration: [traits::OfTypeBis] - Source: 'tests/src/traits.rs', lines 109:0-109:36 *) + Source: 'tests/src/traits.rs', lines 110:0-110:36 *) Record OfTypeBis_t (Self T : Type) := mkOfTypeBis_t { OfTypeBis_tOfTypeBis_t_ToTypeInst : ToType_t T Self; OfTypeBis_t_of_type : T -> result Self; @@ -208,7 +208,7 @@ Arguments OfTypeBis_tOfTypeBis_t_ToTypeInst { _ _ }. Arguments OfTypeBis_t_of_type { _ _ }. (** [traits::h4]: - Source: 'tests/src/traits.rs', lines 118:0-118:57 *) + Source: 'tests/src/traits.rs', lines 119:0-119:57 *) Definition h4 (T1 T2 : Type) (ofTypeBisInst : OfTypeBis_t T1 T2) (toTypeInst : ToType_t T2 T1) (y : T2) : @@ -218,15 +218,15 @@ Definition h4 . (** [traits::TestType] - Source: 'tests/src/traits.rs', lines 122:0-122:22 *) + Source: 'tests/src/traits.rs', lines 123:0-123:22 *) Definition TestType_t (T : Type) : Type := T. (** [traits::{traits::TestType#6}::test::TestType1] - Source: 'tests/src/traits.rs', lines 127:8-127:24 *) + Source: 'tests/src/traits.rs', lines 128:8-128:24 *) Definition TestType_test_TestType1_t : Type := u64. (** Trait declaration: [traits::{traits::TestType#6}::test::TestTrait] - Source: 'tests/src/traits.rs', lines 128:8-128:23 *) + Source: 'tests/src/traits.rs', lines 129:8-129:23 *) Record TestType_test_TestTrait_t (Self : Type) := mkTestType_test_TestTrait_t { TestType_test_TestTrait_t_test : Self -> result bool; }. @@ -235,14 +235,14 @@ Arguments mkTestType_test_TestTrait_t { _ }. Arguments TestType_test_TestTrait_t_test { _ }. (** [traits::{traits::TestType#6}::test::{(traits::{traits::TestType#6}::test::TestTrait for traits::{traits::TestType#6}::test::TestType1)}::test]: - Source: 'tests/src/traits.rs', lines 139:12-139:34 *) + Source: 'tests/src/traits.rs', lines 140:12-140:34 *) Definition testType_test_TestTraittraitsTestTypetestTestType1_test (self : TestType_test_TestType1_t) : result bool := Ok (self s> 1%u64) . (** Trait implementation: [traits::{traits::TestType#6}::test::{(traits::{traits::TestType#6}::test::TestTrait for traits::{traits::TestType#6}::test::TestType1)}] - Source: 'tests/src/traits.rs', lines 138:8-138:36 *) + Source: 'tests/src/traits.rs', lines 139:8-139:36 *) Definition TestType_test_TestTraittraitsTestTypetestTestType1 : TestType_test_TestTrait_t TestType_test_TestType1_t := {| TestType_test_TestTrait_t_test := @@ -250,7 +250,7 @@ Definition TestType_test_TestTraittraitsTestTypetestTestType1 : |}. (** [traits::{traits::TestType#6}::test]: - Source: 'tests/src/traits.rs', lines 126:4-126:36 *) + Source: 'tests/src/traits.rs', lines 127:4-127:36 *) Definition testType_test (T : Type) (toU64Inst : ToU64_t T) (self : TestType_t T) (x : T) : result bool @@ -262,11 +262,11 @@ Definition testType_test . (** [traits::BoolWrapper] - Source: 'tests/src/traits.rs', lines 150:0-150:22 *) + Source: 'tests/src/traits.rs', lines 151:0-151:22 *) Definition BoolWrapper_t : Type := bool. (** [traits::{(traits::ToType for traits::BoolWrapper)#7}::to_type]: - Source: 'tests/src/traits.rs', lines 156:4-156:25 *) + Source: 'tests/src/traits.rs', lines 157:4-157:25 *) Definition toTypetraitsBoolWrapperT_to_type (T : Type) (toTypeBoolTInst : ToType_t bool T) (self : BoolWrapper_t) : result T @@ -275,14 +275,14 @@ Definition toTypetraitsBoolWrapperT_to_type . (** Trait implementation: [traits::{(traits::ToType for traits::BoolWrapper)#7}] - Source: 'tests/src/traits.rs', lines 152:0-152:33 *) + Source: 'tests/src/traits.rs', lines 153:0-153:33 *) Definition ToTypetraitsBoolWrapperT (T : Type) (toTypeBoolTInst : ToType_t bool T) : ToType_t BoolWrapper_t T := {| ToType_t_to_type := toTypetraitsBoolWrapperT_to_type T toTypeBoolTInst; |}. (** [traits::WithConstTy::LEN2] - Source: 'tests/src/traits.rs', lines 164:4-164:21 *) + Source: 'tests/src/traits.rs', lines 165:4-165:21 *) Definition with_const_ty_len2_default_body (Self : Type) (LEN : usize) : result usize := Ok 32%usize @@ -292,7 +292,7 @@ Definition with_const_ty_len2_default (Self : Type) (LEN : usize) : usize := . (** Trait declaration: [traits::WithConstTy] - Source: 'tests/src/traits.rs', lines 161:0-161:39 *) + Source: 'tests/src/traits.rs', lines 162:0-162:39 *) Record WithConstTy_t (Self : Type) (LEN : usize) := mkWithConstTy_t { WithConstTy_tWithConstTy_t_LEN1 : usize; WithConstTy_tWithConstTy_t_LEN2 : usize; @@ -312,21 +312,21 @@ Arguments WithConstTy_tWithConstTy_t_W_clause_0 { _ _ }. Arguments WithConstTy_t_f { _ _ }. (** [traits::{(traits::WithConstTy<32: usize> for bool)#8}::LEN1] - Source: 'tests/src/traits.rs', lines 175:4-175:21 *) + Source: 'tests/src/traits.rs', lines 176:4-176:21 *) Definition with_const_ty_bool32_len1_body : result usize := Ok 12%usize. Definition with_const_ty_bool32_len1 : usize := with_const_ty_bool32_len1_body%global . (** [traits::{(traits::WithConstTy<32: usize> for bool)#8}::f]: - Source: 'tests/src/traits.rs', lines 180:4-180:39 *) + Source: 'tests/src/traits.rs', lines 181:4-181:39 *) Definition withConstTyBool32_f (i : u64) (a : array u8 32%usize) : result u64 := Ok i . (** Trait implementation: [traits::{(traits::WithConstTy<32: usize> for bool)#8}] - Source: 'tests/src/traits.rs', lines 174:0-174:29 *) + Source: 'tests/src/traits.rs', lines 175:0-175:29 *) Definition WithConstTyBool32 : WithConstTy_t bool 32%usize := {| WithConstTy_tWithConstTy_t_LEN1 := with_const_ty_bool32_len1; WithConstTy_tWithConstTy_t_LEN2 := with_const_ty_len2_default bool 32%usize; @@ -337,7 +337,7 @@ Definition WithConstTyBool32 : WithConstTy_t bool 32%usize := {| |}. (** [traits::use_with_const_ty1]: - Source: 'tests/src/traits.rs', lines 183:0-183:75 *) + Source: 'tests/src/traits.rs', lines 184:0-184:75 *) Definition use_with_const_ty1 (H : Type) (LEN : usize) (withConstTyInst : WithConstTy_t H LEN) : result usize @@ -346,7 +346,7 @@ Definition use_with_const_ty1 . (** [traits::use_with_const_ty2]: - Source: 'tests/src/traits.rs', lines 187:0-187:73 *) + Source: 'tests/src/traits.rs', lines 188:0-188:73 *) Definition use_with_const_ty2 (H : Type) (LEN : usize) (withConstTyInst : WithConstTy_t H LEN) (w : withConstTyInst.(WithConstTy_tWithConstTy_t_W)) : @@ -356,7 +356,7 @@ Definition use_with_const_ty2 . (** [traits::use_with_const_ty3]: - Source: 'tests/src/traits.rs', lines 189:0-189:80 *) + Source: 'tests/src/traits.rs', lines 190:0-190:80 *) Definition use_with_const_ty3 (H : Type) (LEN : usize) (withConstTyInst : WithConstTy_t H LEN) (x : withConstTyInst.(WithConstTy_tWithConstTy_t_W)) : @@ -366,12 +366,12 @@ Definition use_with_const_ty3 . (** [traits::test_where1]: - Source: 'tests/src/traits.rs', lines 193:0-193:40 *) + Source: 'tests/src/traits.rs', lines 194:0-194:40 *) Definition test_where1 (T : Type) (_x : T) : result unit := Ok tt. (** [traits::test_where2]: - Source: 'tests/src/traits.rs', lines 194:0-194:57 *) + Source: 'tests/src/traits.rs', lines 195:0-195:57 *) Definition test_where2 (T : Type) (withConstTyT32Inst : WithConstTy_t T 32%usize) (_x : u32) : result unit @@ -380,7 +380,7 @@ Definition test_where2 . (** Trait declaration: [traits::ParentTrait0] - Source: 'tests/src/traits.rs', lines 200:0-200:22 *) + Source: 'tests/src/traits.rs', lines 201:0-201:22 *) Record ParentTrait0_t (Self : Type) := mkParentTrait0_t { ParentTrait0_tParentTrait0_t_W : Type; ParentTrait0_t_get_name : Self -> result string; @@ -393,13 +393,13 @@ Arguments ParentTrait0_t_get_name { _ }. Arguments ParentTrait0_t_get_w { _ }. (** Trait declaration: [traits::ParentTrait1] - Source: 'tests/src/traits.rs', lines 205:0-205:22 *) + Source: 'tests/src/traits.rs', lines 206:0-206:22 *) Record ParentTrait1_t (Self : Type) := mkParentTrait1_t{}. Arguments mkParentTrait1_t { _ }. (** Trait declaration: [traits::ChildTrait] - Source: 'tests/src/traits.rs', lines 206:0-206:49 *) + Source: 'tests/src/traits.rs', lines 207:0-207:49 *) Record ChildTrait_t (Self : Type) := mkChildTrait_t { ChildTrait_tChildTrait_t_ParentTrait0Inst : ParentTrait0_t Self; ChildTrait_tChildTrait_t_ParentTrait1Inst : ParentTrait1_t Self; @@ -410,7 +410,7 @@ Arguments ChildTrait_tChildTrait_t_ParentTrait0Inst { _ }. Arguments ChildTrait_tChildTrait_t_ParentTrait1Inst { _ }. (** [traits::test_child_trait1]: - Source: 'tests/src/traits.rs', lines 209:0-209:56 *) + Source: 'tests/src/traits.rs', lines 210:0-210:56 *) Definition test_child_trait1 (T : Type) (childTraitInst : ChildTrait_t T) (x : T) : result string := childTraitInst.(ChildTrait_tChildTrait_t_ParentTrait0Inst).(ParentTrait0_t_get_name) @@ -418,7 +418,7 @@ Definition test_child_trait1 . (** [traits::test_child_trait2]: - Source: 'tests/src/traits.rs', lines 213:0-213:54 *) + Source: 'tests/src/traits.rs', lines 214:0-214:54 *) Definition test_child_trait2 (T : Type) (childTraitInst : ChildTrait_t T) (x : T) : result @@ -429,7 +429,7 @@ Definition test_child_trait2 . (** [traits::order1]: - Source: 'tests/src/traits.rs', lines 219:0-219:59 *) + Source: 'tests/src/traits.rs', lines 220:0-220:59 *) Definition order1 (T U : Type) (parentTrait0Inst : ParentTrait0_t T) (parentTrait0Inst1 : ParentTrait0_t U) : @@ -439,7 +439,7 @@ Definition order1 . (** Trait declaration: [traits::ChildTrait1] - Source: 'tests/src/traits.rs', lines 222:0-222:35 *) + Source: 'tests/src/traits.rs', lines 223:0-223:35 *) Record ChildTrait1_t (Self : Type) := mkChildTrait1_t { ChildTrait1_tChildTrait1_t_ParentTrait1Inst : ParentTrait1_t Self; }. @@ -448,17 +448,17 @@ Arguments mkChildTrait1_t { _ }. Arguments ChildTrait1_tChildTrait1_t_ParentTrait1Inst { _ }. (** Trait implementation: [traits::{(traits::ParentTrait1 for usize)#9}] - Source: 'tests/src/traits.rs', lines 224:0-224:27 *) + Source: 'tests/src/traits.rs', lines 225:0-225:27 *) Definition ParentTrait1Usize : ParentTrait1_t usize := mkParentTrait1_t. (** Trait implementation: [traits::{(traits::ChildTrait1 for usize)#10}] - Source: 'tests/src/traits.rs', lines 225:0-225:26 *) + Source: 'tests/src/traits.rs', lines 226:0-226:26 *) Definition ChildTrait1Usize : ChildTrait1_t usize := {| ChildTrait1_tChildTrait1_t_ParentTrait1Inst := ParentTrait1Usize; |}. (** Trait declaration: [traits::Iterator] - Source: 'tests/src/traits.rs', lines 229:0-229:18 *) + Source: 'tests/src/traits.rs', lines 230:0-230:18 *) Record Iterator_t (Self : Type) := mkIterator_t { Iterator_tIterator_t_Item : Type; }. @@ -467,7 +467,7 @@ Arguments mkIterator_t { _ }. Arguments Iterator_tIterator_t_Item { _ }. (** Trait declaration: [traits::IntoIterator] - Source: 'tests/src/traits.rs', lines 233:0-233:22 *) + Source: 'tests/src/traits.rs', lines 234:0-234:22 *) Record IntoIterator_t (Self : Type) := mkIntoIterator_t { IntoIterator_tIntoIterator_t_Item : Type; IntoIterator_tIntoIterator_t_IntoIter : Type; @@ -484,13 +484,13 @@ Arguments IntoIterator_tIntoIterator_t_IntoIter_clause_0 { _ }. Arguments IntoIterator_t_into_iter { _ }. (** Trait declaration: [traits::FromResidual] - Source: 'tests/src/traits.rs', lines 250:0-250:21 *) + Source: 'tests/src/traits.rs', lines 251:0-251:21 *) Record FromResidual_t (Self T : Type) := mkFromResidual_t{}. Arguments mkFromResidual_t { _ _ }. (** Trait declaration: [traits::Try] - Source: 'tests/src/traits.rs', lines 246:0-246:48 *) + Source: 'tests/src/traits.rs', lines 247:0-247:48 *) Record Try_t (Self : Type) := mkTry_t { Try_tTry_t_Residual : Type; Try_tTry_t_FromResidualSelftraitsTryResidualInst : FromResidual_t Self @@ -502,7 +502,7 @@ Arguments Try_tTry_t_Residual { _ }. Arguments Try_tTry_t_FromResidualSelftraitsTryResidualInst { _ }. (** Trait declaration: [traits::WithTarget] - Source: 'tests/src/traits.rs', lines 252:0-252:20 *) + Source: 'tests/src/traits.rs', lines 253:0-253:20 *) Record WithTarget_t (Self : Type) := mkWithTarget_t { WithTarget_tWithTarget_t_Target : Type; }. @@ -511,7 +511,7 @@ Arguments mkWithTarget_t { _ }. Arguments WithTarget_tWithTarget_t_Target { _ }. (** Trait declaration: [traits::ParentTrait2] - Source: 'tests/src/traits.rs', lines 256:0-256:22 *) + Source: 'tests/src/traits.rs', lines 257:0-257:22 *) Record ParentTrait2_t (Self : Type) := mkParentTrait2_t { ParentTrait2_tParentTrait2_t_U : Type; ParentTrait2_tParentTrait2_t_U_clause_0 : WithTarget_t @@ -523,7 +523,7 @@ Arguments ParentTrait2_tParentTrait2_t_U { _ }. Arguments ParentTrait2_tParentTrait2_t_U_clause_0 { _ }. (** Trait declaration: [traits::ChildTrait2] - Source: 'tests/src/traits.rs', lines 260:0-260:35 *) + Source: 'tests/src/traits.rs', lines 261:0-261:35 *) Record ChildTrait2_t (Self : Type) := mkChildTrait2_t { ChildTrait2_tChildTrait2_t_ParentTrait2Inst : ParentTrait2_t Self; ChildTrait2_t_convert : @@ -537,32 +537,32 @@ Arguments ChildTrait2_tChildTrait2_t_ParentTrait2Inst { _ }. Arguments ChildTrait2_t_convert { _ }. (** Trait implementation: [traits::{(traits::WithTarget for u32)#11}] - Source: 'tests/src/traits.rs', lines 264:0-264:23 *) + Source: 'tests/src/traits.rs', lines 265:0-265:23 *) Definition WithTargetU32 : WithTarget_t u32 := {| WithTarget_tWithTarget_t_Target := u32; |}. (** Trait implementation: [traits::{(traits::ParentTrait2 for u32)#12}] - Source: 'tests/src/traits.rs', lines 268:0-268:25 *) + Source: 'tests/src/traits.rs', lines 269:0-269:25 *) Definition ParentTrait2U32 : ParentTrait2_t u32 := {| ParentTrait2_tParentTrait2_t_U := u32; ParentTrait2_tParentTrait2_t_U_clause_0 := WithTargetU32; |}. (** [traits::{(traits::ChildTrait2 for u32)#13}::convert]: - Source: 'tests/src/traits.rs', lines 273:4-273:29 *) + Source: 'tests/src/traits.rs', lines 274:4-274:29 *) Definition childTrait2U32_convert (x : u32) : result u32 := Ok x. (** Trait implementation: [traits::{(traits::ChildTrait2 for u32)#13}] - Source: 'tests/src/traits.rs', lines 272:0-272:24 *) + Source: 'tests/src/traits.rs', lines 273:0-273:24 *) Definition ChildTrait2U32 : ChildTrait2_t u32 := {| ChildTrait2_tChildTrait2_t_ParentTrait2Inst := ParentTrait2U32; ChildTrait2_t_convert := childTrait2U32_convert; |}. (** Trait declaration: [traits::CFnOnce] - Source: 'tests/src/traits.rs', lines 286:0-286:23 *) + Source: 'tests/src/traits.rs', lines 287:0-287:23 *) Record CFnOnce_t (Self Args : Type) := mkCFnOnce_t { CFnOnce_tCFnOnce_t_Output : Type; CFnOnce_t_call_once : Self -> Args -> result CFnOnce_tCFnOnce_t_Output; @@ -573,7 +573,7 @@ Arguments CFnOnce_tCFnOnce_t_Output { _ _ }. Arguments CFnOnce_t_call_once { _ _ }. (** Trait declaration: [traits::CFnMut] - Source: 'tests/src/traits.rs', lines 292:0-292:37 *) + Source: 'tests/src/traits.rs', lines 293:0-293:37 *) Record CFnMut_t (Self Args : Type) := mkCFnMut_t { CFnMut_tCFnMut_t_CFnOnceInst : CFnOnce_t Self Args; CFnMut_t_call_mut : Self -> Args -> result @@ -585,7 +585,7 @@ Arguments CFnMut_tCFnMut_t_CFnOnceInst { _ _ }. Arguments CFnMut_t_call_mut { _ _ }. (** Trait declaration: [traits::CFn] - Source: 'tests/src/traits.rs', lines 296:0-296:33 *) + Source: 'tests/src/traits.rs', lines 297:0-297:33 *) Record CFn_t (Self Args : Type) := mkCFn_t { CFn_tCFn_t_CFnMutInst : CFnMut_t Self Args; CFn_t_call : Self -> Args -> result @@ -597,7 +597,7 @@ Arguments CFn_tCFn_t_CFnMutInst { _ _ }. Arguments CFn_t_call { _ _ }. (** Trait declaration: [traits::GetTrait] - Source: 'tests/src/traits.rs', lines 300:0-300:18 *) + Source: 'tests/src/traits.rs', lines 301:0-301:18 *) Record GetTrait_t (Self : Type) := mkGetTrait_t { GetTrait_tGetTrait_t_W : Type; GetTrait_t_get_w : Self -> result GetTrait_tGetTrait_t_W; @@ -608,7 +608,7 @@ Arguments GetTrait_tGetTrait_t_W { _ }. Arguments GetTrait_t_get_w { _ }. (** [traits::test_get_trait]: - Source: 'tests/src/traits.rs', lines 305:0-305:49 *) + Source: 'tests/src/traits.rs', lines 306:0-306:49 *) Definition test_get_trait (T : Type) (getTraitInst : GetTrait_t T) (x : T) : result getTraitInst.(GetTrait_tGetTrait_t_W) @@ -617,27 +617,27 @@ Definition test_get_trait . (** Trait declaration: [traits::Trait] - Source: 'tests/src/traits.rs', lines 310:0-310:15 *) + Source: 'tests/src/traits.rs', lines 311:0-311:15 *) Record Trait_t (Self : Type) := mkTrait_t { Trait_tTrait_t_LEN : usize; }. Arguments mkTrait_t { _ }. Arguments Trait_tTrait_t_LEN { _ }. (** [traits::{(traits::Trait for @Array)#14}::LEN] - Source: 'tests/src/traits.rs', lines 315:4-315:20 *) + Source: 'tests/src/traits.rs', lines 316:4-316:20 *) Definition trait_array_len_body (T : Type) (N : usize) : result usize := Ok N. Definition trait_array_len (T : Type) (N : usize) : usize := (trait_array_len_body T N)%global . (** Trait implementation: [traits::{(traits::Trait for @Array)#14}] - Source: 'tests/src/traits.rs', lines 314:0-314:40 *) + Source: 'tests/src/traits.rs', lines 315:0-315:40 *) Definition TraitArray (T : Type) (N : usize) : Trait_t (array T N) := {| Trait_tTrait_t_LEN := trait_array_len T N; |}. (** [traits::{(traits::Trait for traits::Wrapper)#15}::LEN] - Source: 'tests/src/traits.rs', lines 319:4-319:20 *) + Source: 'tests/src/traits.rs', lines 320:4-320:20 *) Definition traittraits_wrapper_len_body (T : Type) (traitInst : Trait_t T) : result usize := Ok 0%usize @@ -648,20 +648,20 @@ Definition traittraits_wrapper_len (T : Type) (traitInst : Trait_t T) . (** Trait implementation: [traits::{(traits::Trait for traits::Wrapper)#15}] - Source: 'tests/src/traits.rs', lines 318:0-318:35 *) + Source: 'tests/src/traits.rs', lines 319:0-319:35 *) Definition TraittraitsWrapper (T : Type) (traitInst : Trait_t T) : Trait_t (Wrapper_t T) := {| Trait_tTrait_t_LEN := traittraits_wrapper_len T traitInst; |}. (** [traits::use_wrapper_len]: - Source: 'tests/src/traits.rs', lines 322:0-322:43 *) + Source: 'tests/src/traits.rs', lines 323:0-323:43 *) Definition use_wrapper_len (T : Type) (traitInst : Trait_t T) : result usize := Ok (TraittraitsWrapper T traitInst).(Trait_tTrait_t_LEN) . (** [traits::Foo] - Source: 'tests/src/traits.rs', lines 326:0-326:20 *) + Source: 'tests/src/traits.rs', lines 327:0-327:20 *) Record Foo_t (T U : Type) := mkFoo_t { foo_x : T; foo_y : U; }. Arguments mkFoo_t { _ _ }. @@ -680,7 +680,7 @@ Arguments Core_result_Result_Ok { _ _ }. Arguments Core_result_Result_Err { _ _ }. (** [traits::{traits::Foo#16}::FOO] - Source: 'tests/src/traits.rs', lines 332:4-332:33 *) + Source: 'tests/src/traits.rs', lines 333:4-333:33 *) Definition foo_foo_body (T U : Type) (traitInst : Trait_t T) : result (core_result_Result_t T i32) := Ok (Core_result_Result_Err 0%i32) @@ -691,14 +691,14 @@ Definition foo_foo (T U : Type) (traitInst : Trait_t T) . (** [traits::use_foo1]: - Source: 'tests/src/traits.rs', lines 335:0-335:48 *) + Source: 'tests/src/traits.rs', lines 336:0-336:48 *) Definition use_foo1 (T U : Type) (traitInst : Trait_t T) : result (core_result_Result_t T i32) := Ok (foo_foo T U traitInst) . (** [traits::use_foo2]: - Source: 'tests/src/traits.rs', lines 339:0-339:48 *) + Source: 'tests/src/traits.rs', lines 340:0-340:48 *) Definition use_foo2 (T U : Type) (traitInst : Trait_t U) : result (core_result_Result_t U i32) := Ok (foo_foo U T traitInst) diff --git a/tests/fstar/arrays/Arrays.Clauses.Template.fst b/tests/fstar/arrays/Arrays.Clauses.Template.fst index e695b89b..914ef44e 100644 --- a/tests/fstar/arrays/Arrays.Clauses.Template.fst +++ b/tests/fstar/arrays/Arrays.Clauses.Template.fst @@ -7,31 +7,31 @@ open Arrays.Types #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [arrays::sum]: decreases clause - Source: 'tests/src/arrays.rs', lines 242:0-250:1 *) + Source: 'tests/src/arrays.rs', lines 245:0-253:1 *) unfold let sum_loop_decreases (s : slice u32) (sum1 : u32) (i : usize) : nat = admit () (** [arrays::sum2]: decreases clause - Source: 'tests/src/arrays.rs', lines 252:0-261:1 *) + Source: 'tests/src/arrays.rs', lines 255:0-264:1 *) unfold let sum2_loop_decreases (s : slice u32) (s2 : slice u32) (sum1 : u32) (i : usize) : nat = admit () (** [arrays::zero_slice]: decreases clause - Source: 'tests/src/arrays.rs', lines 303:0-310:1 *) + Source: 'tests/src/arrays.rs', lines 306:0-313:1 *) unfold let zero_slice_loop_decreases (a : slice u8) (i : usize) (len : usize) : nat = admit () (** [arrays::iter_mut_slice]: decreases clause - Source: 'tests/src/arrays.rs', lines 312:0-318:1 *) + Source: 'tests/src/arrays.rs', lines 315:0-321:1 *) unfold let iter_mut_slice_loop_decreases (len : usize) (i : usize) : nat = admit () (** [arrays::sum_mut_slice]: decreases clause - Source: 'tests/src/arrays.rs', lines 320:0-328:1 *) + Source: 'tests/src/arrays.rs', lines 323:0-331:1 *) unfold let sum_mut_slice_loop_decreases (a : slice u32) (i : usize) (s : u32) : nat = admit () diff --git a/tests/fstar/arrays/Arrays.Funs.fst b/tests/fstar/arrays/Arrays.Funs.fst index 6196e3b7..289e603d 100644 --- a/tests/fstar/arrays/Arrays.Funs.fst +++ b/tests/fstar/arrays/Arrays.Funs.fst @@ -8,17 +8,17 @@ include Arrays.Clauses #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [arrays::incr]: - Source: 'tests/src/arrays.rs', lines 8:0-8:24 *) + Source: 'tests/src/arrays.rs', lines 11:0-11:24 *) let incr (x : u32) : result u32 = u32_add x 1 (** [arrays::array_to_shared_slice_]: - Source: 'tests/src/arrays.rs', lines 16:0-16:53 *) + Source: 'tests/src/arrays.rs', lines 19:0-19:53 *) let array_to_shared_slice_ (t : Type0) (s : array t 32) : result (slice t) = array_to_slice t 32 s (** [arrays::array_to_mut_slice_]: - Source: 'tests/src/arrays.rs', lines 21:0-21:58 *) + Source: 'tests/src/arrays.rs', lines 24:0-24:58 *) let array_to_mut_slice_ (t : Type0) (s : array t 32) : result ((slice t) & (slice t -> result (array t 32))) @@ -26,37 +26,37 @@ let array_to_mut_slice_ array_to_slice_mut t 32 s (** [arrays::array_len]: - Source: 'tests/src/arrays.rs', lines 25:0-25:40 *) + Source: 'tests/src/arrays.rs', lines 28:0-28:40 *) let array_len (t : Type0) (s : array t 32) : result usize = let* s1 = array_to_slice t 32 s in Ok (slice_len t s1) (** [arrays::shared_array_len]: - Source: 'tests/src/arrays.rs', lines 29:0-29:48 *) + Source: 'tests/src/arrays.rs', lines 32:0-32:48 *) let shared_array_len (t : Type0) (s : array t 32) : result usize = let* s1 = array_to_slice t 32 s in Ok (slice_len t s1) (** [arrays::shared_slice_len]: - Source: 'tests/src/arrays.rs', lines 33:0-33:44 *) + Source: 'tests/src/arrays.rs', lines 36:0-36:44 *) let shared_slice_len (t : Type0) (s : slice t) : result usize = Ok (slice_len t s) (** [arrays::index_array_shared]: - Source: 'tests/src/arrays.rs', lines 37:0-37:57 *) + Source: 'tests/src/arrays.rs', lines 40:0-40:57 *) let index_array_shared (t : Type0) (s : array t 32) (i : usize) : result t = array_index_usize t 32 s i (** [arrays::index_array_u32]: - Source: 'tests/src/arrays.rs', lines 44:0-44:53 *) + Source: 'tests/src/arrays.rs', lines 47:0-47:53 *) let index_array_u32 (s : array u32 32) (i : usize) : result u32 = array_index_usize u32 32 s i (** [arrays::index_array_copy]: - Source: 'tests/src/arrays.rs', lines 48:0-48:45 *) + Source: 'tests/src/arrays.rs', lines 51:0-51:45 *) let index_array_copy (x : array u32 32) : result u32 = array_index_usize u32 32 x 0 (** [arrays::index_mut_array]: - Source: 'tests/src/arrays.rs', lines 52:0-52:62 *) + Source: 'tests/src/arrays.rs', lines 55:0-55:62 *) let index_mut_array (t : Type0) (s : array t 32) (i : usize) : result (t & (t -> result (array t 32))) @@ -64,12 +64,12 @@ let index_mut_array array_index_mut_usize t 32 s i (** [arrays::index_slice]: - Source: 'tests/src/arrays.rs', lines 56:0-56:46 *) + Source: 'tests/src/arrays.rs', lines 59:0-59:46 *) let index_slice (t : Type0) (s : slice t) (i : usize) : result t = slice_index_usize t s i (** [arrays::index_mut_slice]: - Source: 'tests/src/arrays.rs', lines 60:0-60:58 *) + Source: 'tests/src/arrays.rs', lines 63:0-63:58 *) let index_mut_slice (t : Type0) (s : slice t) (i : usize) : result (t & (t -> result (slice t))) @@ -77,7 +77,7 @@ let index_mut_slice slice_index_mut_usize t s i (** [arrays::slice_subslice_shared_]: - Source: 'tests/src/arrays.rs', lines 64:0-64:70 *) + Source: 'tests/src/arrays.rs', lines 67:0-67:70 *) let slice_subslice_shared_ (x : slice u32) (y : usize) (z : usize) : result (slice u32) = core_slice_index_Slice_index u32 (core_ops_range_Range usize) @@ -85,7 +85,7 @@ let slice_subslice_shared_ { start = y; end_ = z } (** [arrays::slice_subslice_mut_]: - Source: 'tests/src/arrays.rs', lines 68:0-68:75 *) + Source: 'tests/src/arrays.rs', lines 71:0-71:75 *) let slice_subslice_mut_ (x : slice u32) (y : usize) (z : usize) : result ((slice u32) & (slice u32 -> result (slice u32))) @@ -97,12 +97,12 @@ let slice_subslice_mut_ Ok (s, index_mut_back) (** [arrays::array_to_slice_shared_]: - Source: 'tests/src/arrays.rs', lines 72:0-72:54 *) + Source: 'tests/src/arrays.rs', lines 75:0-75:54 *) let array_to_slice_shared_ (x : array u32 32) : result (slice u32) = array_to_slice u32 32 x (** [arrays::array_to_slice_mut_]: - Source: 'tests/src/arrays.rs', lines 76:0-76:59 *) + Source: 'tests/src/arrays.rs', lines 79:0-79:59 *) let array_to_slice_mut_ (x : array u32 32) : result ((slice u32) & (slice u32 -> result (array u32 32))) @@ -110,7 +110,7 @@ let array_to_slice_mut_ array_to_slice_mut u32 32 x (** [arrays::array_subslice_shared_]: - Source: 'tests/src/arrays.rs', lines 80:0-80:74 *) + Source: 'tests/src/arrays.rs', lines 83:0-83:74 *) let array_subslice_shared_ (x : array u32 32) (y : usize) (z : usize) : result (slice u32) = core_array_Array_index u32 (core_ops_range_Range usize) 32 @@ -119,7 +119,7 @@ let array_subslice_shared_ { start = y; end_ = z } (** [arrays::array_subslice_mut_]: - Source: 'tests/src/arrays.rs', lines 84:0-84:79 *) + Source: 'tests/src/arrays.rs', lines 87:0-87:79 *) let array_subslice_mut_ (x : array u32 32) (y : usize) (z : usize) : result ((slice u32) & (slice u32 -> result (array u32 32))) @@ -132,24 +132,24 @@ let array_subslice_mut_ Ok (s, index_mut_back) (** [arrays::index_slice_0]: - Source: 'tests/src/arrays.rs', lines 88:0-88:38 *) + Source: 'tests/src/arrays.rs', lines 91:0-91:38 *) let index_slice_0 (t : Type0) (s : slice t) : result t = slice_index_usize t s 0 (** [arrays::index_array_0]: - Source: 'tests/src/arrays.rs', lines 92:0-92:42 *) + Source: 'tests/src/arrays.rs', lines 95:0-95:42 *) let index_array_0 (t : Type0) (s : array t 32) : result t = array_index_usize t 32 s 0 (** [arrays::index_index_array]: - Source: 'tests/src/arrays.rs', lines 103:0-103:71 *) + Source: 'tests/src/arrays.rs', lines 106:0-106:71 *) let index_index_array (s : array (array u32 32) 32) (i : usize) (j : usize) : result u32 = let* a = array_index_usize (array u32 32) 32 s i in array_index_usize u32 32 a j (** [arrays::update_update_array]: - Source: 'tests/src/arrays.rs', lines 114:0-114:70 *) + Source: 'tests/src/arrays.rs', lines 117:0-117:70 *) let update_update_array (s : array (array u32 32) 32) (i : usize) (j : usize) : result unit = let* (a, index_mut_back) = array_index_mut_usize (array u32 32) 32 s i in @@ -159,42 +159,42 @@ let update_update_array Ok () (** [arrays::array_local_deep_copy]: - Source: 'tests/src/arrays.rs', lines 118:0-118:43 *) + Source: 'tests/src/arrays.rs', lines 121:0-121:43 *) let array_local_deep_copy (x : array u32 32) : result unit = Ok () (** [arrays::take_array]: - Source: 'tests/src/arrays.rs', lines 122:0-122:30 *) + Source: 'tests/src/arrays.rs', lines 125:0-125:30 *) let take_array (a : array u32 2) : result unit = Ok () (** [arrays::take_array_borrow]: - Source: 'tests/src/arrays.rs', lines 123:0-123:38 *) + Source: 'tests/src/arrays.rs', lines 126:0-126:38 *) let take_array_borrow (a : array u32 2) : result unit = Ok () (** [arrays::take_slice]: - Source: 'tests/src/arrays.rs', lines 124:0-124:28 *) + Source: 'tests/src/arrays.rs', lines 127:0-127:28 *) let take_slice (s : slice u32) : result unit = Ok () (** [arrays::take_mut_slice]: - Source: 'tests/src/arrays.rs', lines 125:0-125:36 *) + Source: 'tests/src/arrays.rs', lines 128:0-128:36 *) let take_mut_slice (s : slice u32) : result (slice u32) = Ok s (** [arrays::const_array]: - Source: 'tests/src/arrays.rs', lines 127:0-127:32 *) + Source: 'tests/src/arrays.rs', lines 130:0-130:32 *) let const_array : result (array u32 2) = Ok (mk_array u32 2 [ 0; 0 ]) (** [arrays::const_slice]: - Source: 'tests/src/arrays.rs', lines 131:0-131:20 *) + Source: 'tests/src/arrays.rs', lines 134:0-134:20 *) let const_slice : result unit = let* _ = array_to_slice u32 2 (mk_array u32 2 [ 0; 0 ]) in Ok () (** [arrays::take_all]: - Source: 'tests/src/arrays.rs', lines 141:0-141:17 *) + Source: 'tests/src/arrays.rs', lines 144:0-144:17 *) let take_all : result unit = let* _ = take_array (mk_array u32 2 [ 0; 0 ]) in let* _ = take_array (mk_array u32 2 [ 0; 0 ]) in @@ -208,27 +208,27 @@ let take_all : result unit = Ok () (** [arrays::index_array]: - Source: 'tests/src/arrays.rs', lines 155:0-155:38 *) + Source: 'tests/src/arrays.rs', lines 158:0-158:38 *) let index_array (x : array u32 2) : result u32 = array_index_usize u32 2 x 0 (** [arrays::index_array_borrow]: - Source: 'tests/src/arrays.rs', lines 158:0-158:46 *) + Source: 'tests/src/arrays.rs', lines 161:0-161:46 *) let index_array_borrow (x : array u32 2) : result u32 = array_index_usize u32 2 x 0 (** [arrays::index_slice_u32_0]: - Source: 'tests/src/arrays.rs', lines 162:0-162:42 *) + Source: 'tests/src/arrays.rs', lines 165:0-165:42 *) let index_slice_u32_0 (x : slice u32) : result u32 = slice_index_usize u32 x 0 (** [arrays::index_mut_slice_u32_0]: - Source: 'tests/src/arrays.rs', lines 166:0-166:50 *) + Source: 'tests/src/arrays.rs', lines 169:0-169:50 *) let index_mut_slice_u32_0 (x : slice u32) : result (u32 & (slice u32)) = let* i = slice_index_usize u32 x 0 in Ok (i, x) (** [arrays::index_all]: - Source: 'tests/src/arrays.rs', lines 170:0-170:25 *) + Source: 'tests/src/arrays.rs', lines 173:0-173:25 *) let index_all : result u32 = let* i = index_array (mk_array u32 2 [ 0; 0 ]) in let* i1 = index_array (mk_array u32 2 [ 0; 0 ]) in @@ -246,25 +246,25 @@ let index_all : result u32 = Ok i8 (** [arrays::update_array]: - Source: 'tests/src/arrays.rs', lines 184:0-184:36 *) + Source: 'tests/src/arrays.rs', lines 187:0-187:36 *) let update_array (x : array u32 2) : result unit = let* (_, index_mut_back) = array_index_mut_usize u32 2 x 0 in let* _ = index_mut_back 1 in Ok () (** [arrays::update_array_mut_borrow]: - Source: 'tests/src/arrays.rs', lines 187:0-187:48 *) + Source: 'tests/src/arrays.rs', lines 190:0-190:48 *) let update_array_mut_borrow (x : array u32 2) : result (array u32 2) = let* (_, index_mut_back) = array_index_mut_usize u32 2 x 0 in index_mut_back 1 (** [arrays::update_mut_slice]: - Source: 'tests/src/arrays.rs', lines 190:0-190:38 *) + Source: 'tests/src/arrays.rs', lines 193:0-193:38 *) let update_mut_slice (x : slice u32) : result (slice u32) = let* (_, index_mut_back) = slice_index_mut_usize u32 x 0 in index_mut_back 1 (** [arrays::update_all]: - Source: 'tests/src/arrays.rs', lines 194:0-194:19 *) + Source: 'tests/src/arrays.rs', lines 197:0-197:19 *) let update_all : result unit = let* _ = update_array (mk_array u32 2 [ 0; 0 ]) in let* _ = update_array (mk_array u32 2 [ 0; 0 ]) in @@ -275,7 +275,7 @@ let update_all : result unit = Ok () (** [arrays::range_all]: - Source: 'tests/src/arrays.rs', lines 205:0-205:18 *) + Source: 'tests/src/arrays.rs', lines 208:0-208:18 *) let range_all : result unit = let* (s, index_mut_back) = core_array_Array_index_mut u32 (core_ops_range_Range usize) 4 @@ -287,27 +287,27 @@ let range_all : result unit = Ok () (** [arrays::deref_array_borrow]: - Source: 'tests/src/arrays.rs', lines 214:0-214:46 *) + Source: 'tests/src/arrays.rs', lines 217:0-217:46 *) let deref_array_borrow (x : array u32 2) : result u32 = array_index_usize u32 2 x 0 (** [arrays::deref_array_mut_borrow]: - Source: 'tests/src/arrays.rs', lines 219:0-219:54 *) + Source: 'tests/src/arrays.rs', lines 222:0-222:54 *) let deref_array_mut_borrow (x : array u32 2) : result (u32 & (array u32 2)) = let* i = array_index_usize u32 2 x 0 in Ok (i, x) (** [arrays::take_array_t]: - Source: 'tests/src/arrays.rs', lines 227:0-227:31 *) + Source: 'tests/src/arrays.rs', lines 230:0-230:31 *) let take_array_t (a : array aB_t 2) : result unit = Ok () (** [arrays::non_copyable_array]: - Source: 'tests/src/arrays.rs', lines 229:0-229:27 *) + Source: 'tests/src/arrays.rs', lines 232:0-232:27 *) let non_copyable_array : result unit = take_array_t (mk_array aB_t 2 [ AB_A; AB_B ]) (** [arrays::sum]: loop 0: - Source: 'tests/src/arrays.rs', lines 242:0-250:1 *) + Source: 'tests/src/arrays.rs', lines 245:0-253:1 *) let rec sum_loop (s : slice u32) (sum1 : u32) (i : usize) : Tot (result u32) (decreases (sum_loop_decreases s sum1 i)) @@ -322,12 +322,12 @@ let rec sum_loop else Ok sum1 (** [arrays::sum]: - Source: 'tests/src/arrays.rs', lines 242:0-242:28 *) + Source: 'tests/src/arrays.rs', lines 245:0-245:28 *) let sum (s : slice u32) : result u32 = sum_loop s 0 0 (** [arrays::sum2]: loop 0: - Source: 'tests/src/arrays.rs', lines 252:0-261:1 *) + Source: 'tests/src/arrays.rs', lines 255:0-264:1 *) let rec sum2_loop (s : slice u32) (s2 : slice u32) (sum1 : u32) (i : usize) : Tot (result u32) (decreases (sum2_loop_decreases s s2 sum1 i)) @@ -344,14 +344,14 @@ let rec sum2_loop else Ok sum1 (** [arrays::sum2]: - Source: 'tests/src/arrays.rs', lines 252:0-252:41 *) + Source: 'tests/src/arrays.rs', lines 255:0-255:41 *) let sum2 (s : slice u32) (s2 : slice u32) : result u32 = let i = slice_len u32 s in let i1 = slice_len u32 s2 in if not (i = i1) then Fail Failure else sum2_loop s s2 0 0 (** [arrays::f0]: - Source: 'tests/src/arrays.rs', lines 263:0-263:11 *) + Source: 'tests/src/arrays.rs', lines 266:0-266:11 *) let f0 : result unit = let* (s, to_slice_mut_back) = array_to_slice_mut u32 2 (mk_array u32 2 [ 1; 2 ]) in @@ -361,7 +361,7 @@ let f0 : result unit = Ok () (** [arrays::f1]: - Source: 'tests/src/arrays.rs', lines 268:0-268:11 *) + Source: 'tests/src/arrays.rs', lines 271:0-271:11 *) let f1 : result unit = let* (_, index_mut_back) = array_index_mut_usize u32 2 (mk_array u32 2 [ 1; 2 ]) 0 in @@ -369,12 +369,12 @@ let f1 : result unit = Ok () (** [arrays::f2]: - Source: 'tests/src/arrays.rs', lines 273:0-273:17 *) + Source: 'tests/src/arrays.rs', lines 276:0-276:17 *) let f2 (i : u32) : result unit = Ok () (** [arrays::f4]: - Source: 'tests/src/arrays.rs', lines 282:0-282:54 *) + Source: 'tests/src/arrays.rs', lines 285:0-285:54 *) let f4 (x : array u32 32) (y : usize) (z : usize) : result (slice u32) = core_array_Array_index u32 (core_ops_range_Range usize) 32 (core_ops_index_IndexSliceTIInst u32 (core_ops_range_Range usize) @@ -382,7 +382,7 @@ let f4 (x : array u32 32) (y : usize) (z : usize) : result (slice u32) = { start = y; end_ = z } (** [arrays::f3]: - Source: 'tests/src/arrays.rs', lines 275:0-275:18 *) + Source: 'tests/src/arrays.rs', lines 278:0-278:18 *) let f3 : result u32 = let* i = array_index_usize u32 2 (mk_array u32 2 [ 1; 2 ]) 0 in let* _ = f2 i in @@ -392,17 +392,17 @@ let f3 : result u32 = sum2 s s1 (** [arrays::SZ] - Source: 'tests/src/arrays.rs', lines 286:0-286:19 *) + Source: 'tests/src/arrays.rs', lines 289:0-289:19 *) let sz_body : result usize = Ok 32 let sz : usize = eval_global sz_body (** [arrays::f5]: - Source: 'tests/src/arrays.rs', lines 289:0-289:31 *) + Source: 'tests/src/arrays.rs', lines 292:0-292:31 *) let f5 (x : array u32 32) : result u32 = array_index_usize u32 32 x 0 (** [arrays::ite]: - Source: 'tests/src/arrays.rs', lines 294:0-294:12 *) + Source: 'tests/src/arrays.rs', lines 297:0-297:12 *) let ite : result unit = let* (s, to_slice_mut_back) = array_to_slice_mut u32 2 (mk_array u32 2 [ 0; 0 ]) in @@ -415,7 +415,7 @@ let ite : result unit = Ok () (** [arrays::zero_slice]: loop 0: - Source: 'tests/src/arrays.rs', lines 303:0-310:1 *) + Source: 'tests/src/arrays.rs', lines 306:0-313:1 *) let rec zero_slice_loop (a : slice u8) (i : usize) (len : usize) : Tot (result (slice u8)) (decreases (zero_slice_loop_decreases a i len)) @@ -429,12 +429,12 @@ let rec zero_slice_loop else Ok a (** [arrays::zero_slice]: - Source: 'tests/src/arrays.rs', lines 303:0-303:31 *) + Source: 'tests/src/arrays.rs', lines 306:0-306:31 *) let zero_slice (a : slice u8) : result (slice u8) = let len = slice_len u8 a in zero_slice_loop a 0 len (** [arrays::iter_mut_slice]: loop 0: - Source: 'tests/src/arrays.rs', lines 312:0-318:1 *) + Source: 'tests/src/arrays.rs', lines 315:0-321:1 *) let rec iter_mut_slice_loop (len : usize) (i : usize) : Tot (result unit) (decreases (iter_mut_slice_loop_decreases len i)) @@ -444,12 +444,12 @@ let rec iter_mut_slice_loop else Ok () (** [arrays::iter_mut_slice]: - Source: 'tests/src/arrays.rs', lines 312:0-312:35 *) + Source: 'tests/src/arrays.rs', lines 315:0-315:35 *) let iter_mut_slice (a : slice u8) : result (slice u8) = let len = slice_len u8 a in let* _ = iter_mut_slice_loop len 0 in Ok a (** [arrays::sum_mut_slice]: loop 0: - Source: 'tests/src/arrays.rs', lines 320:0-328:1 *) + Source: 'tests/src/arrays.rs', lines 323:0-331:1 *) let rec sum_mut_slice_loop (a : slice u32) (i : usize) (s : u32) : Tot (result u32) (decreases (sum_mut_slice_loop_decreases a i s)) @@ -464,7 +464,7 @@ let rec sum_mut_slice_loop else Ok s (** [arrays::sum_mut_slice]: - Source: 'tests/src/arrays.rs', lines 320:0-320:42 *) + Source: 'tests/src/arrays.rs', lines 323:0-323:42 *) let sum_mut_slice (a : slice u32) : result (u32 & (slice u32)) = let* i = sum_mut_slice_loop a 0 0 in Ok (i, a) diff --git a/tests/fstar/arrays/Arrays.Types.fst b/tests/fstar/arrays/Arrays.Types.fst index be9c224f..41f892ad 100644 --- a/tests/fstar/arrays/Arrays.Types.fst +++ b/tests/fstar/arrays/Arrays.Types.fst @@ -6,6 +6,6 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [arrays::AB] - Source: 'tests/src/arrays.rs', lines 3:0-3:11 *) + Source: 'tests/src/arrays.rs', lines 6:0-6:11 *) type aB_t = | AB_A : aB_t | AB_B : aB_t diff --git a/tests/fstar/demo/Demo.fst b/tests/fstar/demo/Demo.fst index 41fd9804..60722f46 100644 --- a/tests/fstar/demo/Demo.fst +++ b/tests/fstar/demo/Demo.fst @@ -6,7 +6,7 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [demo::choose]: - Source: 'tests/src/demo.rs', lines 5:0-5:70 *) + Source: 'tests/src/demo.rs', lines 6:0-6:70 *) let choose (t : Type0) (b : bool) (x : t) (y : t) : result (t & (t -> result (t & t))) = if b @@ -14,33 +14,33 @@ let choose else let back = fun ret -> Ok (x, ret) in Ok (y, back) (** [demo::mul2_add1]: - Source: 'tests/src/demo.rs', lines 13:0-13:31 *) + Source: 'tests/src/demo.rs', lines 14:0-14:31 *) let mul2_add1 (x : u32) : result u32 = let* i = u32_add x x in u32_add i 1 (** [demo::use_mul2_add1]: - Source: 'tests/src/demo.rs', lines 17:0-17:43 *) + Source: 'tests/src/demo.rs', lines 18:0-18:43 *) let use_mul2_add1 (x : u32) (y : u32) : result u32 = let* i = mul2_add1 x in u32_add i y (** [demo::incr]: - Source: 'tests/src/demo.rs', lines 21:0-21:31 *) + Source: 'tests/src/demo.rs', lines 22:0-22:31 *) let incr (x : u32) : result u32 = u32_add x 1 (** [demo::use_incr]: - Source: 'tests/src/demo.rs', lines 25:0-25:17 *) + Source: 'tests/src/demo.rs', lines 26:0-26:17 *) let use_incr : result unit = let* x = incr 0 in let* x1 = incr x in let* _ = incr x1 in Ok () (** [demo::CList] - Source: 'tests/src/demo.rs', lines 34:0-34:17 *) + Source: 'tests/src/demo.rs', lines 35:0-35:17 *) type cList_t (t : Type0) = | CList_CCons : t -> cList_t t -> cList_t t | CList_CNil : cList_t t (** [demo::list_nth]: - Source: 'tests/src/demo.rs', lines 39:0-39:56 *) + Source: 'tests/src/demo.rs', lines 40:0-40:56 *) let rec list_nth (t : Type0) (n : nat) (l : cList_t t) (i : u32) : result t = if is_zero n then Fail OutOfFuel @@ -53,7 +53,7 @@ let rec list_nth (t : Type0) (n : nat) (l : cList_t t) (i : u32) : result t = end (** [demo::list_nth_mut]: - Source: 'tests/src/demo.rs', lines 54:0-54:68 *) + Source: 'tests/src/demo.rs', lines 55:0-55:68 *) let rec list_nth_mut (t : Type0) (n : nat) (l : cList_t t) (i : u32) : result (t & (t -> result (cList_t t))) @@ -77,7 +77,7 @@ let rec list_nth_mut end (** [demo::list_nth_mut1]: loop 0: - Source: 'tests/src/demo.rs', lines 69:0-78:1 *) + Source: 'tests/src/demo.rs', lines 70:0-79:1 *) let rec list_nth_mut1_loop (t : Type0) (n : nat) (l : cList_t t) (i : u32) : result (t & (t -> result (cList_t t))) @@ -99,7 +99,7 @@ let rec list_nth_mut1_loop end (** [demo::list_nth_mut1]: - Source: 'tests/src/demo.rs', lines 69:0-69:77 *) + Source: 'tests/src/demo.rs', lines 70:0-70:77 *) let list_nth_mut1 (t : Type0) (n : nat) (l : cList_t t) (i : u32) : result (t & (t -> result (cList_t t))) @@ -107,7 +107,7 @@ let list_nth_mut1 list_nth_mut1_loop t n l i (** [demo::i32_id]: - Source: 'tests/src/demo.rs', lines 80:0-80:28 *) + Source: 'tests/src/demo.rs', lines 81:0-81:28 *) let rec i32_id (n : nat) (i : i32) : result i32 = if is_zero n then Fail OutOfFuel @@ -118,7 +118,7 @@ let rec i32_id (n : nat) (i : i32) : result i32 = else let* i1 = i32_sub i 1 in let* i2 = i32_id n1 i1 in i32_add i2 1 (** [demo::list_tail]: - Source: 'tests/src/demo.rs', lines 88:0-88:64 *) + Source: 'tests/src/demo.rs', lines 89:0-89:64 *) let rec list_tail (t : Type0) (n : nat) (l : cList_t t) : result ((cList_t t) & (cList_t t -> result (cList_t t))) @@ -137,20 +137,20 @@ let rec list_tail end (** Trait declaration: [demo::Counter] - Source: 'tests/src/demo.rs', lines 97:0-97:17 *) + Source: 'tests/src/demo.rs', lines 98:0-98:17 *) noeq type counter_t (self : Type0) = { incr : self -> result (usize & self); } (** [demo::{(demo::Counter for usize)}::incr]: - Source: 'tests/src/demo.rs', lines 102:4-102:31 *) + Source: 'tests/src/demo.rs', lines 103:4-103:31 *) let counterUsize_incr (self : usize) : result (usize & usize) = let* self1 = usize_add self 1 in Ok (self, self1) (** Trait implementation: [demo::{(demo::Counter for usize)}] - Source: 'tests/src/demo.rs', lines 101:0-101:22 *) + Source: 'tests/src/demo.rs', lines 102:0-102:22 *) let counterUsize : counter_t usize = { incr = counterUsize_incr; } (** [demo::use_counter]: - Source: 'tests/src/demo.rs', lines 109:0-109:59 *) + Source: 'tests/src/demo.rs', lines 110:0-110:59 *) let use_counter (t : Type0) (counterInst : counter_t t) (cnt : t) : result (usize & t) = counterInst.incr cnt diff --git a/tests/fstar/hashmap/Hashmap.Clauses.Template.fst b/tests/fstar/hashmap/Hashmap.Clauses.Template.fst index b96f6784..3119ded8 100644 --- a/tests/fstar/hashmap/Hashmap.Clauses.Template.fst +++ b/tests/fstar/hashmap/Hashmap.Clauses.Template.fst @@ -7,63 +7,63 @@ open Hashmap.Types #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [hashmap::{hashmap::HashMap}::allocate_slots]: decreases clause - Source: 'tests/src/hashmap.rs', lines 50:4-56:5 *) + Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) unfold let hashMap_allocate_slots_loop_decreases (t : Type0) (slots : alloc_vec_Vec (list_t t)) (n : usize) : nat = admit () (** [hashmap::{hashmap::HashMap}::clear]: decreases clause - Source: 'tests/src/hashmap.rs', lines 80:4-88:5 *) + Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) unfold let hashMap_clear_loop_decreases (t : Type0) (slots : alloc_vec_Vec (list_t t)) (i : usize) : nat = admit () (** [hashmap::{hashmap::HashMap}::insert_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 97:4-114:5 *) + Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) unfold let hashMap_insert_in_list_loop_decreases (t : Type0) (key : usize) (value : t) (ls : list_t t) : nat = admit () (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 183:4-196:5 *) + Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) unfold let hashMap_move_elements_from_list_loop_decreases (t : Type0) (ntable : hashMap_t t) (ls : list_t t) : nat = admit () (** [hashmap::{hashmap::HashMap}::move_elements]: decreases clause - Source: 'tests/src/hashmap.rs', lines 171:4-180:5 *) + Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) unfold let hashMap_move_elements_loop_decreases (t : Type0) (ntable : hashMap_t t) (slots : alloc_vec_Vec (list_t t)) (i : usize) : nat = admit () (** [hashmap::{hashmap::HashMap}::contains_key_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 206:4-219:5 *) + Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) unfold let hashMap_contains_key_in_list_loop_decreases (t : Type0) (key : usize) (ls : list_t t) : nat = admit () (** [hashmap::{hashmap::HashMap}::get_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 224:4-237:5 *) + Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) unfold let hashMap_get_in_list_loop_decreases (t : Type0) (key : usize) (ls : list_t t) : nat = admit () (** [hashmap::{hashmap::HashMap}::get_mut_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 245:4-254:5 *) + Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) unfold let hashMap_get_mut_in_list_loop_decreases (t : Type0) (ls : list_t t) (key : usize) : nat = admit () (** [hashmap::{hashmap::HashMap}::remove_from_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 265:4-291:5 *) + Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) unfold let hashMap_remove_from_list_loop_decreases (t : Type0) (key : usize) (ls : list_t t) : nat = diff --git a/tests/fstar/hashmap/Hashmap.Funs.fst b/tests/fstar/hashmap/Hashmap.Funs.fst index 2aca9fbe..06cdf7f3 100644 --- a/tests/fstar/hashmap/Hashmap.Funs.fst +++ b/tests/fstar/hashmap/Hashmap.Funs.fst @@ -8,12 +8,12 @@ include Hashmap.Clauses #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 27:0-27:32 *) + Source: 'tests/src/hashmap.rs', lines 35:0-35:32 *) let hash_key (k : usize) : result usize = Ok k (** [hashmap::{hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 50:4-56:5 *) + Source: 'tests/src/hashmap.rs', lines 58:4-64: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))) @@ -27,7 +27,7 @@ let rec hashMap_allocate_slots_loop else Ok slots (** [hashmap::{hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 50:4-50:76 *) + Source: 'tests/src/hashmap.rs', lines 58:4-58:76 *) let hashMap_allocate_slots (t : Type0) (slots : alloc_vec_Vec (list_t t)) (n : usize) : result (alloc_vec_Vec (list_t t)) @@ -35,7 +35,7 @@ let hashMap_allocate_slots hashMap_allocate_slots_loop t slots n (** [hashmap::{hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 59:4-63:13 *) + Source: 'tests/src/hashmap.rs', lines 67:4-71:13 *) let hashMap_new_with_capacity (t : Type0) (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : @@ -54,12 +54,12 @@ let hashMap_new_with_capacity } (** [hashmap::{hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 75:4-75:24 *) + Source: 'tests/src/hashmap.rs', lines 83:4-83:24 *) let hashMap_new (t : Type0) : result (hashMap_t t) = hashMap_new_with_capacity t 32 4 5 (** [hashmap::{hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 80:4-88:5 *) + Source: 'tests/src/hashmap.rs', lines 88:4-96: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))) @@ -77,18 +77,18 @@ let rec hashMap_clear_loop else Ok slots (** [hashmap::{hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 80:4-80:27 *) + Source: 'tests/src/hashmap.rs', lines 88:4-88: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}::len]: - Source: 'tests/src/hashmap.rs', lines 90:4-90:30 *) + Source: 'tests/src/hashmap.rs', lines 98:4-98:30 *) let hashMap_len (t : Type0) (self : hashMap_t t) : result usize = Ok self.num_entries (** [hashmap::{hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 97:4-114:5 *) + Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) let rec hashMap_insert_in_list_loop (t : Type0) (key : usize) (value : t) (ls : list_t t) : Tot (result (bool & (list_t t))) @@ -105,7 +105,7 @@ let rec hashMap_insert_in_list_loop end (** [hashmap::{hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 97:4-97:71 *) + Source: 'tests/src/hashmap.rs', lines 105:4-105:71 *) let hashMap_insert_in_list (t : Type0) (key : usize) (value : t) (ls : list_t t) : result (bool & (list_t t)) @@ -113,7 +113,7 @@ let hashMap_insert_in_list hashMap_insert_in_list_loop t key value ls (** [hashmap::{hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 117:4-117:54 *) + Source: 'tests/src/hashmap.rs', lines 125:4-125:54 *) let hashMap_insert_no_resize (t : Type0) (self : hashMap_t t) (key : usize) (value : t) : result (hashMap_t t) @@ -134,7 +134,7 @@ let hashMap_insert_no_resize else let* v = index_mut_back l1 in Ok { self with slots = v } (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 183:4-196:5 *) + Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) let rec hashMap_move_elements_from_list_loop (t : Type0) (ntable : hashMap_t t) (ls : list_t t) : Tot (result (hashMap_t t)) @@ -148,13 +148,13 @@ let rec hashMap_move_elements_from_list_loop end (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 183:4-183:72 *) + Source: 'tests/src/hashmap.rs', lines 191:4-191:72 *) let hashMap_move_elements_from_list (t : Type0) (ntable : hashMap_t t) (ls : list_t t) : result (hashMap_t t) = hashMap_move_elements_from_list_loop t ntable ls (** [hashmap::{hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 171:4-180:5 *) + Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) let rec hashMap_move_elements_loop (t : Type0) (ntable : hashMap_t t) (slots : alloc_vec_Vec (list_t t)) (i : usize) : @@ -175,7 +175,7 @@ let rec hashMap_move_elements_loop else Ok (ntable, slots) (** [hashmap::{hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 171:4-171:95 *) + Source: 'tests/src/hashmap.rs', lines 179:4-179:95 *) let hashMap_move_elements (t : Type0) (ntable : hashMap_t t) (slots : alloc_vec_Vec (list_t t)) (i : usize) : @@ -184,7 +184,7 @@ let hashMap_move_elements hashMap_move_elements_loop t ntable slots i (** [hashmap::{hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 140:4-140:28 *) + Source: 'tests/src/hashmap.rs', lines 148:4-148: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 @@ -204,7 +204,7 @@ let hashMap_try_resize else Ok { self with max_load_factor = (i, i1) } (** [hashmap::{hashmap::HashMap}::insert]: - Source: 'tests/src/hashmap.rs', lines 129:4-129:48 *) + Source: 'tests/src/hashmap.rs', lines 137:4-137:48 *) let hashMap_insert (t : Type0) (self : hashMap_t t) (key : usize) (value : t) : result (hashMap_t t) @@ -214,7 +214,7 @@ let hashMap_insert if i > self1.max_load then hashMap_try_resize t self1 else Ok self1 (** [hashmap::{hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 206:4-219:5 *) + Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) let rec hashMap_contains_key_in_list_loop (t : Type0) (key : usize) (ls : list_t t) : Tot (result bool) @@ -227,13 +227,13 @@ let rec hashMap_contains_key_in_list_loop end (** [hashmap::{hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 206:4-206:68 *) + Source: 'tests/src/hashmap.rs', lines 214:4-214:68 *) let hashMap_contains_key_in_list (t : Type0) (key : usize) (ls : list_t t) : result bool = hashMap_contains_key_in_list_loop t key ls (** [hashmap::{hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 199:4-199:49 *) + Source: 'tests/src/hashmap.rs', lines 207:4-207:49 *) let hashMap_contains_key (t : Type0) (self : hashMap_t t) (key : usize) : result bool = let* hash = hash_key key in @@ -246,7 +246,7 @@ let hashMap_contains_key hashMap_contains_key_in_list t key l (** [hashmap::{hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 224:4-237:5 *) + Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) let rec hashMap_get_in_list_loop (t : Type0) (key : usize) (ls : list_t t) : Tot (result t) (decreases (hashMap_get_in_list_loop_decreases t key ls)) @@ -258,12 +258,12 @@ let rec hashMap_get_in_list_loop end (** [hashmap::{hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 224:4-224:70 *) + Source: 'tests/src/hashmap.rs', lines 232:4-232:70 *) let hashMap_get_in_list (t : Type0) (key : usize) (ls : list_t t) : result t = hashMap_get_in_list_loop t key ls (** [hashmap::{hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 239:4-239:55 *) + Source: 'tests/src/hashmap.rs', lines 247:4-247: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 @@ -275,7 +275,7 @@ let hashMap_get (t : Type0) (self : hashMap_t t) (key : usize) : result t = hashMap_get_in_list t key l (** [hashmap::{hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 245:4-254:5 *) + Source: 'tests/src/hashmap.rs', lines 253:4-262: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)))) @@ -294,7 +294,7 @@ let rec hashMap_get_mut_in_list_loop end (** [hashmap::{hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 245:4-245:86 *) + Source: 'tests/src/hashmap.rs', lines 253:4-253:86 *) let hashMap_get_mut_in_list (t : Type0) (ls : list_t t) (key : usize) : result (t & (t -> result (list_t t))) @@ -302,7 +302,7 @@ let hashMap_get_mut_in_list hashMap_get_mut_in_list_loop t ls key (** [hashmap::{hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 257:4-257:67 *) + Source: 'tests/src/hashmap.rs', lines 265:4-265:67 *) let hashMap_get_mut (t : Type0) (self : hashMap_t t) (key : usize) : result (t & (t -> result (hashMap_t t))) @@ -323,7 +323,7 @@ let hashMap_get_mut Ok (x, back) (** [hashmap::{hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 265:4-291:5 *) + Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) let rec hashMap_remove_from_list_loop (t : Type0) (key : usize) (ls : list_t t) : Tot (result ((option t) & (list_t t))) @@ -346,7 +346,7 @@ let rec hashMap_remove_from_list_loop end (** [hashmap::{hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:69 *) + Source: 'tests/src/hashmap.rs', lines 273:4-273:69 *) let hashMap_remove_from_list (t : Type0) (key : usize) (ls : list_t t) : result ((option t) & (list_t t)) @@ -354,7 +354,7 @@ let hashMap_remove_from_list hashMap_remove_from_list_loop t key ls (** [hashmap::{hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 294:4-294:52 *) + Source: 'tests/src/hashmap.rs', lines 302:4-302:52 *) let hashMap_remove (t : Type0) (self : hashMap_t t) (key : usize) : result ((option t) & (hashMap_t t)) @@ -376,7 +376,7 @@ let hashMap_remove end (** [hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 315:0-315:10 *) + Source: 'tests/src/hashmap.rs', lines 323:0-323: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.Types.fst b/tests/fstar/hashmap/Hashmap.Types.fst index ee06442f..962fbee2 100644 --- a/tests/fstar/hashmap/Hashmap.Types.fst +++ b/tests/fstar/hashmap/Hashmap.Types.fst @@ -6,13 +6,13 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [hashmap::List] - Source: 'tests/src/hashmap.rs', lines 19:0-19:16 *) + Source: 'tests/src/hashmap.rs', lines 27:0-27:16 *) type list_t (t : Type0) = | List_Cons : usize -> t -> list_t t -> list_t t | List_Nil : list_t t (** [hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 35:0-35:21 *) + Source: 'tests/src/hashmap.rs', lines 43:0-43:21 *) type hashMap_t (t : Type0) = { num_entries : usize; diff --git a/tests/fstar/hashmap_on_disk/HashmapMain.Clauses.Template.fst b/tests/fstar/hashmap_on_disk/HashmapMain.Clauses.Template.fst index 0715bdcb..cdd73210 100644 --- a/tests/fstar/hashmap_on_disk/HashmapMain.Clauses.Template.fst +++ b/tests/fstar/hashmap_on_disk/HashmapMain.Clauses.Template.fst @@ -7,35 +7,35 @@ open HashmapMain.Types #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: decreases clause - Source: 'tests/src/hashmap.rs', lines 50:4-56:5 *) + Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) unfold let hashmap_HashMap_allocate_slots_loop_decreases (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (n : usize) : nat = admit () (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: decreases clause - Source: 'tests/src/hashmap.rs', lines 80:4-88:5 *) + Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) unfold let hashmap_HashMap_clear_loop_decreases (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : nat = admit () (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 97:4-114:5 *) + Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) unfold let hashmap_HashMap_insert_in_list_loop_decreases (t : Type0) (key : usize) (value : t) (ls : hashmap_List_t t) : nat = admit () (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 183:4-196:5 *) + Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) unfold let hashmap_HashMap_move_elements_from_list_loop_decreases (t : Type0) (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : nat = admit () (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: decreases clause - Source: 'tests/src/hashmap.rs', lines 171:4-180:5 *) + Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) unfold let hashmap_HashMap_move_elements_loop_decreases (t : Type0) (ntable : hashmap_HashMap_t t) (slots : alloc_vec_Vec (hashmap_List_t t)) @@ -43,28 +43,28 @@ let hashmap_HashMap_move_elements_loop_decreases (t : Type0) admit () (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 206:4-219:5 *) + Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) unfold let hashmap_HashMap_contains_key_in_list_loop_decreases (t : Type0) (key : usize) (ls : hashmap_List_t t) : nat = admit () (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 224:4-237:5 *) + Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) unfold let hashmap_HashMap_get_in_list_loop_decreases (t : Type0) (key : usize) (ls : hashmap_List_t t) : nat = admit () (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 245:4-254:5 *) + Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) unfold let hashmap_HashMap_get_mut_in_list_loop_decreases (t : Type0) (ls : hashmap_List_t t) (key : usize) : nat = admit () (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 265:4-291:5 *) + Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) unfold let hashmap_HashMap_remove_from_list_loop_decreases (t : Type0) (key : usize) (ls : hashmap_List_t t) : nat = diff --git a/tests/fstar/hashmap_on_disk/HashmapMain.Funs.fst b/tests/fstar/hashmap_on_disk/HashmapMain.Funs.fst index 4a032207..c88a746e 100644 --- a/tests/fstar/hashmap_on_disk/HashmapMain.Funs.fst +++ b/tests/fstar/hashmap_on_disk/HashmapMain.Funs.fst @@ -9,12 +9,12 @@ include HashmapMain.Clauses #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [hashmap_main::hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 27:0-27:32 *) + Source: 'tests/src/hashmap.rs', lines 35:0-35:32 *) let hashmap_hash_key (k : usize) : result usize = Ok k (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 50:4-56:5 *) + Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) let rec hashmap_HashMap_allocate_slots_loop (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (n : usize) : Tot (result (alloc_vec_Vec (hashmap_List_t t))) @@ -29,7 +29,7 @@ let rec hashmap_HashMap_allocate_slots_loop else Ok slots (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 50:4-50:76 *) + Source: 'tests/src/hashmap.rs', lines 58:4-58:76 *) let hashmap_HashMap_allocate_slots (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (n : usize) : result (alloc_vec_Vec (hashmap_List_t t)) @@ -37,7 +37,7 @@ let hashmap_HashMap_allocate_slots hashmap_HashMap_allocate_slots_loop t slots n (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 59:4-63:13 *) + Source: 'tests/src/hashmap.rs', lines 67:4-71:13 *) let hashmap_HashMap_new_with_capacity (t : Type0) (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : @@ -57,12 +57,12 @@ let hashmap_HashMap_new_with_capacity } (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 75:4-75:24 *) + Source: 'tests/src/hashmap.rs', lines 83:4-83:24 *) let hashmap_HashMap_new (t : Type0) : result (hashmap_HashMap_t t) = hashmap_HashMap_new_with_capacity t 32 4 5 (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 80:4-88:5 *) + Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) let rec hashmap_HashMap_clear_loop (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : Tot (result (alloc_vec_Vec (hashmap_List_t t))) @@ -81,20 +81,20 @@ let rec hashmap_HashMap_clear_loop else Ok slots (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 80:4-80:27 *) + Source: 'tests/src/hashmap.rs', lines 88:4-88:27 *) let hashmap_HashMap_clear (t : Type0) (self : hashmap_HashMap_t t) : result (hashmap_HashMap_t t) = let* hm = hashmap_HashMap_clear_loop t self.slots 0 in Ok { self with num_entries = 0; slots = hm } (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::len]: - Source: 'tests/src/hashmap.rs', lines 90:4-90:30 *) + Source: 'tests/src/hashmap.rs', lines 98:4-98:30 *) let hashmap_HashMap_len (t : Type0) (self : hashmap_HashMap_t t) : result usize = Ok self.num_entries (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 97:4-114:5 *) + Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) let rec hashmap_HashMap_insert_in_list_loop (t : Type0) (key : usize) (value : t) (ls : hashmap_List_t t) : Tot (result (bool & (hashmap_List_t t))) @@ -111,7 +111,7 @@ let rec hashmap_HashMap_insert_in_list_loop end (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 97:4-97:71 *) + Source: 'tests/src/hashmap.rs', lines 105:4-105:71 *) let hashmap_HashMap_insert_in_list (t : Type0) (key : usize) (value : t) (ls : hashmap_List_t t) : result (bool & (hashmap_List_t t)) @@ -119,7 +119,7 @@ let hashmap_HashMap_insert_in_list hashmap_HashMap_insert_in_list_loop t key value ls (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 117:4-117:54 *) + Source: 'tests/src/hashmap.rs', lines 125:4-125:54 *) let hashmap_HashMap_insert_no_resize (t : Type0) (self : hashmap_HashMap_t t) (key : usize) (value : t) : result (hashmap_HashMap_t t) @@ -140,7 +140,7 @@ let hashmap_HashMap_insert_no_resize else let* v = index_mut_back l1 in Ok { self with slots = v } (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 183:4-196:5 *) + Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) let rec hashmap_HashMap_move_elements_from_list_loop (t : Type0) (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : Tot (result (hashmap_HashMap_t t)) @@ -155,7 +155,7 @@ let rec hashmap_HashMap_move_elements_from_list_loop end (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 183:4-183:72 *) + Source: 'tests/src/hashmap.rs', lines 191:4-191:72 *) let hashmap_HashMap_move_elements_from_list (t : Type0) (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : result (hashmap_HashMap_t t) @@ -163,7 +163,7 @@ let hashmap_HashMap_move_elements_from_list hashmap_HashMap_move_elements_from_list_loop t ntable ls (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 171:4-180:5 *) + Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) let rec hashmap_HashMap_move_elements_loop (t : Type0) (ntable : hashmap_HashMap_t t) (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : @@ -185,7 +185,7 @@ let rec hashmap_HashMap_move_elements_loop else Ok (ntable, slots) (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 171:4-171:95 *) + Source: 'tests/src/hashmap.rs', lines 179:4-179:95 *) let hashmap_HashMap_move_elements (t : Type0) (ntable : hashmap_HashMap_t t) (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : @@ -194,7 +194,7 @@ let hashmap_HashMap_move_elements hashmap_HashMap_move_elements_loop t ntable slots i (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 140:4-140:28 *) + Source: 'tests/src/hashmap.rs', lines 148:4-148:28 *) let hashmap_HashMap_try_resize (t : Type0) (self : hashmap_HashMap_t t) : result (hashmap_HashMap_t t) = let* max_usize = scalar_cast U32 Usize core_u32_max in @@ -214,7 +214,7 @@ let hashmap_HashMap_try_resize else Ok { self with max_load_factor = (i, i1) } (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert]: - Source: 'tests/src/hashmap.rs', lines 129:4-129:48 *) + Source: 'tests/src/hashmap.rs', lines 137:4-137:48 *) let hashmap_HashMap_insert (t : Type0) (self : hashmap_HashMap_t t) (key : usize) (value : t) : result (hashmap_HashMap_t t) @@ -224,7 +224,7 @@ let hashmap_HashMap_insert if i > self1.max_load then hashmap_HashMap_try_resize t self1 else Ok self1 (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 206:4-219:5 *) + Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) let rec hashmap_HashMap_contains_key_in_list_loop (t : Type0) (key : usize) (ls : hashmap_List_t t) : Tot (result bool) @@ -239,13 +239,13 @@ let rec hashmap_HashMap_contains_key_in_list_loop end (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 206:4-206:68 *) + Source: 'tests/src/hashmap.rs', lines 214:4-214:68 *) let hashmap_HashMap_contains_key_in_list (t : Type0) (key : usize) (ls : hashmap_List_t t) : result bool = hashmap_HashMap_contains_key_in_list_loop t key ls (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 199:4-199:49 *) + Source: 'tests/src/hashmap.rs', lines 207:4-207:49 *) let hashmap_HashMap_contains_key (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : result bool = let* hash = hashmap_hash_key key in @@ -258,7 +258,7 @@ let hashmap_HashMap_contains_key hashmap_HashMap_contains_key_in_list t key l (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 224:4-237:5 *) + Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) let rec hashmap_HashMap_get_in_list_loop (t : Type0) (key : usize) (ls : hashmap_List_t t) : Tot (result t) @@ -271,13 +271,13 @@ let rec hashmap_HashMap_get_in_list_loop end (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 224:4-224:70 *) + Source: 'tests/src/hashmap.rs', lines 232:4-232:70 *) let hashmap_HashMap_get_in_list (t : Type0) (key : usize) (ls : hashmap_List_t t) : result t = hashmap_HashMap_get_in_list_loop t key ls (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 239:4-239:55 *) + Source: 'tests/src/hashmap.rs', lines 247:4-247:55 *) let hashmap_HashMap_get (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : result t = let* hash = hashmap_hash_key key in @@ -290,7 +290,7 @@ let hashmap_HashMap_get hashmap_HashMap_get_in_list t key l (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 245:4-254:5 *) + Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) let rec hashmap_HashMap_get_mut_in_list_loop (t : Type0) (ls : hashmap_List_t t) (key : usize) : Tot (result (t & (t -> result (hashmap_List_t t)))) @@ -312,7 +312,7 @@ let rec hashmap_HashMap_get_mut_in_list_loop end (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 245:4-245:86 *) + Source: 'tests/src/hashmap.rs', lines 253:4-253:86 *) let hashmap_HashMap_get_mut_in_list (t : Type0) (ls : hashmap_List_t t) (key : usize) : result (t & (t -> result (hashmap_List_t t))) @@ -320,7 +320,7 @@ let hashmap_HashMap_get_mut_in_list hashmap_HashMap_get_mut_in_list_loop t ls key (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 257:4-257:67 *) + Source: 'tests/src/hashmap.rs', lines 265:4-265:67 *) let hashmap_HashMap_get_mut (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : result (t & (t -> result (hashmap_HashMap_t t))) @@ -341,7 +341,7 @@ let hashmap_HashMap_get_mut Ok (x, back) (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 265:4-291:5 *) + Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) let rec hashmap_HashMap_remove_from_list_loop (t : Type0) (key : usize) (ls : hashmap_List_t t) : Tot (result ((option t) & (hashmap_List_t t))) @@ -365,7 +365,7 @@ let rec hashmap_HashMap_remove_from_list_loop end (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:69 *) + Source: 'tests/src/hashmap.rs', lines 273:4-273:69 *) let hashmap_HashMap_remove_from_list (t : Type0) (key : usize) (ls : hashmap_List_t t) : result ((option t) & (hashmap_List_t t)) @@ -373,7 +373,7 @@ let hashmap_HashMap_remove_from_list hashmap_HashMap_remove_from_list_loop t key ls (** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 294:4-294:52 *) + Source: 'tests/src/hashmap.rs', lines 302:4-302:52 *) let hashmap_HashMap_remove (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : result ((option t) & (hashmap_HashMap_t t)) @@ -395,7 +395,7 @@ let hashmap_HashMap_remove end (** [hashmap_main::hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 315:0-315:10 *) + Source: 'tests/src/hashmap.rs', lines 323:0-323:10 *) let hashmap_test1 : result unit = let* hm = hashmap_HashMap_new u64 in let* hm1 = hashmap_HashMap_insert u64 hm 0 42 in @@ -432,7 +432,7 @@ let hashmap_test1 : result unit = end (** [hashmap_main::insert_on_disk]: - Source: 'tests/src/hashmap_main.rs', lines 7:0-7:43 *) + Source: 'tests/src/hashmap_main.rs', lines 13:0-13:43 *) let insert_on_disk (key : usize) (value : u64) (st : state) : result (state & unit) = let* (st1, hm) = hashmap_utils_deserialize st in @@ -440,7 +440,7 @@ let insert_on_disk hashmap_utils_serialize hm1 st1 (** [hashmap_main::main]: - Source: 'tests/src/hashmap_main.rs', lines 16:0-16:13 *) + Source: 'tests/src/hashmap_main.rs', lines 22:0-22:13 *) let main : result unit = Ok () diff --git a/tests/fstar/hashmap_on_disk/HashmapMain.FunsExternal.fsti b/tests/fstar/hashmap_on_disk/HashmapMain.FunsExternal.fsti index 3afc4689..cc20d988 100644 --- a/tests/fstar/hashmap_on_disk/HashmapMain.FunsExternal.fsti +++ b/tests/fstar/hashmap_on_disk/HashmapMain.FunsExternal.fsti @@ -7,12 +7,12 @@ include HashmapMain.Types #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [hashmap_main::hashmap_utils::deserialize]: - Source: 'tests/src/hashmap_utils.rs', lines 10:0-10:43 *) + Source: 'tests/src/hashmap_utils.rs', lines 11:0-11:43 *) val hashmap_utils_deserialize : state -> result (state & (hashmap_HashMap_t u64)) (** [hashmap_main::hashmap_utils::serialize]: - Source: 'tests/src/hashmap_utils.rs', lines 5:0-5:42 *) + Source: 'tests/src/hashmap_utils.rs', lines 6:0-6:42 *) val hashmap_utils_serialize : hashmap_HashMap_t u64 -> state -> result (state & unit) diff --git a/tests/fstar/hashmap_on_disk/HashmapMain.Types.fst b/tests/fstar/hashmap_on_disk/HashmapMain.Types.fst index 21d32541..85bcaeea 100644 --- a/tests/fstar/hashmap_on_disk/HashmapMain.Types.fst +++ b/tests/fstar/hashmap_on_disk/HashmapMain.Types.fst @@ -7,13 +7,13 @@ include HashmapMain.TypesExternal #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [hashmap_main::hashmap::List] - Source: 'tests/src/hashmap.rs', lines 19:0-19:16 *) + Source: 'tests/src/hashmap.rs', lines 27:0-27:16 *) type hashmap_List_t (t : Type0) = | Hashmap_List_Cons : usize -> t -> hashmap_List_t t -> hashmap_List_t t | Hashmap_List_Nil : hashmap_List_t t (** [hashmap_main::hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 35:0-35:21 *) + Source: 'tests/src/hashmap.rs', lines 43:0-43:21 *) type hashmap_HashMap_t (t : Type0) = { num_entries : usize; diff --git a/tests/fstar/misc/Bitwise.fst b/tests/fstar/misc/Bitwise.fst index b17c9233..8d6bab58 100644 --- a/tests/fstar/misc/Bitwise.fst +++ b/tests/fstar/misc/Bitwise.fst @@ -6,27 +6,27 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [bitwise::shift_u32]: - Source: 'tests/src/bitwise.rs', lines 3:0-3:31 *) + Source: 'tests/src/bitwise.rs', lines 4:0-4:31 *) let shift_u32 (a : u32) : result u32 = let* t = u32_shr #Usize a 16 in u32_shl #Usize t 16 (** [bitwise::shift_i32]: - Source: 'tests/src/bitwise.rs', lines 10:0-10:31 *) + Source: 'tests/src/bitwise.rs', lines 11:0-11:31 *) let shift_i32 (a : i32) : result i32 = let* t = i32_shr #Isize a 16 in i32_shl #Isize t 16 (** [bitwise::xor_u32]: - Source: 'tests/src/bitwise.rs', lines 17:0-17:37 *) + Source: 'tests/src/bitwise.rs', lines 18:0-18:37 *) let xor_u32 (a : u32) (b : u32) : result u32 = Ok (u32_xor a b) (** [bitwise::or_u32]: - Source: 'tests/src/bitwise.rs', lines 21:0-21:36 *) + Source: 'tests/src/bitwise.rs', lines 22:0-22:36 *) let or_u32 (a : u32) (b : u32) : result u32 = Ok (u32_or a b) (** [bitwise::and_u32]: - Source: 'tests/src/bitwise.rs', lines 25:0-25:37 *) + Source: 'tests/src/bitwise.rs', lines 26:0-26:37 *) let and_u32 (a : u32) (b : u32) : result u32 = Ok (u32_and a b) diff --git a/tests/fstar/misc/Constants.fst b/tests/fstar/misc/Constants.fst index 78562ae1..4ff5e883 100644 --- a/tests/fstar/misc/Constants.fst +++ b/tests/fstar/misc/Constants.fst @@ -6,154 +6,154 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [constants::X0] - Source: 'tests/src/constants.rs', lines 5:0-5:17 *) + Source: 'tests/src/constants.rs', lines 7:0-7:17 *) let x0_body : result u32 = Ok 0 let x0 : u32 = eval_global x0_body (** [constants::X1] - Source: 'tests/src/constants.rs', lines 7:0-7:17 *) + Source: 'tests/src/constants.rs', lines 9:0-9:17 *) let x1_body : result u32 = Ok core_u32_max let x1 : u32 = eval_global x1_body (** [constants::X2] - Source: 'tests/src/constants.rs', lines 10:0-10:17 *) + Source: 'tests/src/constants.rs', lines 12:0-12:17 *) let x2_body : result u32 = Ok 3 let x2 : u32 = eval_global x2_body (** [constants::incr]: - Source: 'tests/src/constants.rs', lines 17:0-17:32 *) + Source: 'tests/src/constants.rs', lines 19:0-19:32 *) let incr (n : u32) : result u32 = u32_add n 1 (** [constants::X3] - Source: 'tests/src/constants.rs', lines 15:0-15:17 *) + Source: 'tests/src/constants.rs', lines 17:0-17:17 *) let x3_body : result u32 = incr 32 let x3 : u32 = eval_global x3_body (** [constants::mk_pair0]: - Source: 'tests/src/constants.rs', lines 23:0-23:51 *) + Source: 'tests/src/constants.rs', lines 25:0-25:51 *) let mk_pair0 (x : u32) (y1 : u32) : result (u32 & u32) = Ok (x, y1) (** [constants::Pair] - Source: 'tests/src/constants.rs', lines 36:0-36:23 *) + Source: 'tests/src/constants.rs', lines 38:0-38:23 *) type pair_t (t1 t2 : Type0) = { x : t1; y : t2; } (** [constants::mk_pair1]: - Source: 'tests/src/constants.rs', lines 27:0-27:55 *) + Source: 'tests/src/constants.rs', lines 29:0-29:55 *) let mk_pair1 (x : u32) (y1 : u32) : result (pair_t u32 u32) = Ok { x = x; y = y1 } (** [constants::P0] - Source: 'tests/src/constants.rs', lines 31:0-31:24 *) + Source: 'tests/src/constants.rs', lines 33:0-33:24 *) let p0_body : result (u32 & u32) = mk_pair0 0 1 let p0 : (u32 & u32) = eval_global p0_body (** [constants::P1] - Source: 'tests/src/constants.rs', lines 32:0-32:28 *) + Source: 'tests/src/constants.rs', lines 34:0-34:28 *) let p1_body : result (pair_t u32 u32) = mk_pair1 0 1 let p1 : pair_t u32 u32 = eval_global p1_body (** [constants::P2] - Source: 'tests/src/constants.rs', lines 33:0-33:24 *) + Source: 'tests/src/constants.rs', lines 35:0-35:24 *) let p2_body : result (u32 & u32) = Ok (0, 1) let p2 : (u32 & u32) = eval_global p2_body (** [constants::P3] - Source: 'tests/src/constants.rs', lines 34:0-34:28 *) + Source: 'tests/src/constants.rs', lines 36:0-36:28 *) let p3_body : result (pair_t u32 u32) = Ok { x = 0; y = 1 } let p3 : pair_t u32 u32 = eval_global p3_body (** [constants::Wrap] - Source: 'tests/src/constants.rs', lines 49:0-49:18 *) + Source: 'tests/src/constants.rs', lines 51:0-51:18 *) type wrap_t (t : Type0) = { value : t; } (** [constants::{constants::Wrap}::new]: - Source: 'tests/src/constants.rs', lines 54:4-54:41 *) + Source: 'tests/src/constants.rs', lines 56:4-56:41 *) let wrap_new (t : Type0) (value : t) : result (wrap_t t) = Ok { value = value } (** [constants::Y] - Source: 'tests/src/constants.rs', lines 41:0-41:22 *) + Source: 'tests/src/constants.rs', lines 43:0-43:22 *) let y_body : result (wrap_t i32) = wrap_new i32 2 let y : wrap_t i32 = eval_global y_body (** [constants::unwrap_y]: - Source: 'tests/src/constants.rs', lines 43:0-43:30 *) + Source: 'tests/src/constants.rs', lines 45:0-45:30 *) let unwrap_y : result i32 = Ok y.value (** [constants::YVAL] - Source: 'tests/src/constants.rs', lines 47:0-47:19 *) + Source: 'tests/src/constants.rs', lines 49:0-49:19 *) let yval_body : result i32 = unwrap_y let yval : i32 = eval_global yval_body (** [constants::get_z1::Z1] - Source: 'tests/src/constants.rs', lines 62:4-62:17 *) + Source: 'tests/src/constants.rs', lines 64:4-64:17 *) let get_z1_z1_body : result i32 = Ok 3 let get_z1_z1 : i32 = eval_global get_z1_z1_body (** [constants::get_z1]: - Source: 'tests/src/constants.rs', lines 61:0-61:28 *) + Source: 'tests/src/constants.rs', lines 63:0-63:28 *) let get_z1 : result i32 = Ok get_z1_z1 (** [constants::add]: - Source: 'tests/src/constants.rs', lines 66:0-66:39 *) + Source: 'tests/src/constants.rs', lines 68:0-68:39 *) let add (a : i32) (b : i32) : result i32 = i32_add a b (** [constants::Q1] - Source: 'tests/src/constants.rs', lines 74:0-74:17 *) + Source: 'tests/src/constants.rs', lines 76:0-76:17 *) let q1_body : result i32 = Ok 5 let q1 : i32 = eval_global q1_body (** [constants::Q2] - Source: 'tests/src/constants.rs', lines 75:0-75:17 *) + Source: 'tests/src/constants.rs', lines 77:0-77:17 *) let q2_body : result i32 = Ok q1 let q2 : i32 = eval_global q2_body (** [constants::Q3] - Source: 'tests/src/constants.rs', lines 76:0-76:17 *) + Source: 'tests/src/constants.rs', lines 78:0-78:17 *) let q3_body : result i32 = add q2 3 let q3 : i32 = eval_global q3_body (** [constants::get_z2]: - Source: 'tests/src/constants.rs', lines 70:0-70:28 *) + Source: 'tests/src/constants.rs', lines 72:0-72:28 *) let get_z2 : result i32 = let* i = get_z1 in let* i1 = add i q3 in add q1 i1 (** [constants::S1] - Source: 'tests/src/constants.rs', lines 80:0-80:18 *) + Source: 'tests/src/constants.rs', lines 82:0-82:18 *) let s1_body : result u32 = Ok 6 let s1 : u32 = eval_global s1_body (** [constants::S2] - Source: 'tests/src/constants.rs', lines 81:0-81:18 *) + Source: 'tests/src/constants.rs', lines 83:0-83:18 *) let s2_body : result u32 = incr s1 let s2 : u32 = eval_global s2_body (** [constants::S3] - Source: 'tests/src/constants.rs', lines 82:0-82:29 *) + Source: 'tests/src/constants.rs', lines 84:0-84:29 *) let s3_body : result (pair_t u32 u32) = Ok p3 let s3 : pair_t u32 u32 = eval_global s3_body (** [constants::S4] - Source: 'tests/src/constants.rs', lines 83:0-83:29 *) + Source: 'tests/src/constants.rs', lines 85:0-85:29 *) let s4_body : result (pair_t u32 u32) = mk_pair1 7 8 let s4 : pair_t u32 u32 = eval_global s4_body (** [constants::V] - Source: 'tests/src/constants.rs', lines 86:0-86:31 *) + Source: 'tests/src/constants.rs', lines 88:0-88:31 *) type v_t (t : Type0) (n : usize) = { x : array t n; } (** [constants::{constants::V#1}::LEN] - Source: 'tests/src/constants.rs', lines 91:4-91:24 *) + Source: 'tests/src/constants.rs', lines 93:4-93:24 *) let v_len_body (t : Type0) (n : usize) : result usize = Ok n let v_len (t : Type0) (n : usize) : usize = eval_global (v_len_body t n) (** [constants::use_v]: - Source: 'tests/src/constants.rs', lines 94:0-94:42 *) + Source: 'tests/src/constants.rs', lines 96:0-96:42 *) let use_v (t : Type0) (n : usize) : result usize = Ok (v_len t n) diff --git a/tests/fstar/misc/External.Funs.fst b/tests/fstar/misc/External.Funs.fst index 5a7a9f32..18fc901f 100644 --- a/tests/fstar/misc/External.Funs.fst +++ b/tests/fstar/misc/External.Funs.fst @@ -15,12 +15,12 @@ let core_marker_CopyU32 : core_marker_Copy_t u32 = { } (** [external::use_get]: - Source: 'tests/src/external.rs', lines 5:0-5:37 *) + Source: 'tests/src/external.rs', lines 8:0-8:37 *) let use_get (rc : core_cell_Cell_t u32) (st : state) : result (state & u32) = core_cell_Cell_get u32 core_marker_CopyU32 rc st (** [external::incr]: - Source: 'tests/src/external.rs', lines 9:0-9:31 *) + Source: 'tests/src/external.rs', lines 12:0-12:31 *) let incr (rc : core_cell_Cell_t u32) (st : state) : result (state & (core_cell_Cell_t u32)) diff --git a/tests/fstar/misc/Loops.Clauses.Template.fst b/tests/fstar/misc/Loops.Clauses.Template.fst index 77f9c3e4..7b042375 100644 --- a/tests/fstar/misc/Loops.Clauses.Template.fst +++ b/tests/fstar/misc/Loops.Clauses.Template.fst @@ -7,144 +7,144 @@ open Loops.Types #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [loops::sum]: decreases clause - Source: 'tests/src/loops.rs', lines 4:0-14:1 *) + Source: 'tests/src/loops.rs', lines 7:0-17:1 *) unfold let sum_loop_decreases (max : u32) (i : u32) (s : u32) : nat = admit () (** [loops::sum_with_mut_borrows]: decreases clause - Source: 'tests/src/loops.rs', lines 19:0-31:1 *) + Source: 'tests/src/loops.rs', lines 22:0-34:1 *) unfold let sum_with_mut_borrows_loop_decreases (max : u32) (i : u32) (s : u32) : nat = admit () (** [loops::sum_with_shared_borrows]: decreases clause - Source: 'tests/src/loops.rs', lines 34:0-48:1 *) + Source: 'tests/src/loops.rs', lines 37:0-51:1 *) unfold let sum_with_shared_borrows_loop_decreases (max : u32) (i : u32) (s : u32) : nat = admit () (** [loops::sum_array]: decreases clause - Source: 'tests/src/loops.rs', lines 50:0-58:1 *) + Source: 'tests/src/loops.rs', lines 53:0-61:1 *) unfold let sum_array_loop_decreases (n : usize) (a : array u32 n) (i : usize) (s : u32) : nat = admit () (** [loops::clear]: decreases clause - Source: 'tests/src/loops.rs', lines 62:0-68:1 *) + Source: 'tests/src/loops.rs', lines 65:0-71:1 *) unfold let clear_loop_decreases (v : alloc_vec_Vec u32) (i : usize) : nat = admit () (** [loops::list_mem]: decreases clause - Source: 'tests/src/loops.rs', lines 76:0-85:1 *) + Source: 'tests/src/loops.rs', lines 79:0-88:1 *) unfold let list_mem_loop_decreases (x : u32) (ls : list_t u32) : nat = admit () (** [loops::list_nth_mut_loop]: decreases clause - Source: 'tests/src/loops.rs', lines 88:0-98:1 *) + Source: 'tests/src/loops.rs', lines 91:0-101:1 *) unfold let list_nth_mut_loop_loop_decreases (t : Type0) (ls : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_shared_loop]: decreases clause - Source: 'tests/src/loops.rs', lines 101:0-111:1 *) + Source: 'tests/src/loops.rs', lines 104:0-114:1 *) unfold let list_nth_shared_loop_loop_decreases (t : Type0) (ls : list_t t) (i : u32) : nat = admit () (** [loops::get_elem_mut]: decreases clause - Source: 'tests/src/loops.rs', lines 113:0-127:1 *) + Source: 'tests/src/loops.rs', lines 116:0-130:1 *) unfold let get_elem_mut_loop_decreases (x : usize) (ls : list_t usize) : nat = admit () (** [loops::get_elem_shared]: decreases clause - Source: 'tests/src/loops.rs', lines 129:0-143:1 *) + Source: 'tests/src/loops.rs', lines 132:0-146:1 *) unfold let get_elem_shared_loop_decreases (x : usize) (ls : list_t usize) : nat = admit () (** [loops::list_nth_mut_loop_with_id]: decreases clause - Source: 'tests/src/loops.rs', lines 154:0-165:1 *) + Source: 'tests/src/loops.rs', lines 157:0-168:1 *) unfold let list_nth_mut_loop_with_id_loop_decreases (t : Type0) (i : u32) (ls : list_t t) : nat = admit () (** [loops::list_nth_shared_loop_with_id]: decreases clause - Source: 'tests/src/loops.rs', lines 168:0-179:1 *) + Source: 'tests/src/loops.rs', lines 171:0-182:1 *) unfold let list_nth_shared_loop_with_id_loop_decreases (t : Type0) (i : u32) (ls : list_t t) : nat = admit () (** [loops::list_nth_mut_loop_pair]: decreases clause - Source: 'tests/src/loops.rs', lines 184:0-205:1 *) + Source: 'tests/src/loops.rs', lines 187:0-208:1 *) unfold let list_nth_mut_loop_pair_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_shared_loop_pair]: decreases clause - Source: 'tests/src/loops.rs', lines 208:0-229:1 *) + Source: 'tests/src/loops.rs', lines 211:0-232:1 *) unfold let list_nth_shared_loop_pair_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_mut_loop_pair_merge]: decreases clause - Source: 'tests/src/loops.rs', lines 233:0-248:1 *) + Source: 'tests/src/loops.rs', lines 236:0-251:1 *) unfold let list_nth_mut_loop_pair_merge_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_shared_loop_pair_merge]: decreases clause - Source: 'tests/src/loops.rs', lines 251:0-266:1 *) + Source: 'tests/src/loops.rs', lines 254:0-269:1 *) unfold let list_nth_shared_loop_pair_merge_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_mut_shared_loop_pair]: decreases clause - Source: 'tests/src/loops.rs', lines 269:0-284:1 *) + Source: 'tests/src/loops.rs', lines 272:0-287:1 *) unfold let list_nth_mut_shared_loop_pair_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_mut_shared_loop_pair_merge]: decreases clause - Source: 'tests/src/loops.rs', lines 288:0-303:1 *) + Source: 'tests/src/loops.rs', lines 291:0-306:1 *) unfold let list_nth_mut_shared_loop_pair_merge_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_shared_mut_loop_pair]: decreases clause - Source: 'tests/src/loops.rs', lines 307:0-322:1 *) + Source: 'tests/src/loops.rs', lines 310:0-325:1 *) unfold let list_nth_shared_mut_loop_pair_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_shared_mut_loop_pair_merge]: decreases clause - Source: 'tests/src/loops.rs', lines 326:0-341:1 *) + Source: 'tests/src/loops.rs', lines 329:0-344:1 *) unfold let list_nth_shared_mut_loop_pair_merge_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::ignore_input_mut_borrow]: decreases clause - Source: 'tests/src/loops.rs', lines 345:0-349:1 *) + Source: 'tests/src/loops.rs', lines 348:0-352:1 *) unfold let ignore_input_mut_borrow_loop_decreases (i : u32) : nat = admit () (** [loops::incr_ignore_input_mut_borrow]: decreases clause - Source: 'tests/src/loops.rs', lines 353:0-358:1 *) + Source: 'tests/src/loops.rs', lines 356:0-361:1 *) unfold let incr_ignore_input_mut_borrow_loop_decreases (i : u32) : nat = admit () (** [loops::ignore_input_shared_borrow]: decreases clause - Source: 'tests/src/loops.rs', lines 362:0-366:1 *) + Source: 'tests/src/loops.rs', lines 365:0-369:1 *) unfold let ignore_input_shared_borrow_loop_decreases (i : u32) : nat = admit () diff --git a/tests/fstar/misc/Loops.Funs.fst b/tests/fstar/misc/Loops.Funs.fst index 0eafeebb..84e9634d 100644 --- a/tests/fstar/misc/Loops.Funs.fst +++ b/tests/fstar/misc/Loops.Funs.fst @@ -8,7 +8,7 @@ include Loops.Clauses #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [loops::sum]: loop 0: - Source: 'tests/src/loops.rs', lines 4:0-14:1 *) + Source: 'tests/src/loops.rs', lines 7:0-17:1 *) let rec sum_loop (max : u32) (i : u32) (s : u32) : Tot (result u32) (decreases (sum_loop_decreases max i s)) @@ -18,12 +18,12 @@ let rec sum_loop else u32_mul s 2 (** [loops::sum]: - Source: 'tests/src/loops.rs', lines 4:0-4:27 *) + Source: 'tests/src/loops.rs', lines 7:0-7:27 *) let sum (max : u32) : result u32 = sum_loop max 0 0 (** [loops::sum_with_mut_borrows]: loop 0: - Source: 'tests/src/loops.rs', lines 19:0-31:1 *) + Source: 'tests/src/loops.rs', lines 22:0-34:1 *) let rec sum_with_mut_borrows_loop (max : u32) (i : u32) (s : u32) : Tot (result u32) (decreases (sum_with_mut_borrows_loop_decreases max i s)) @@ -36,12 +36,12 @@ let rec sum_with_mut_borrows_loop else u32_mul s 2 (** [loops::sum_with_mut_borrows]: - Source: 'tests/src/loops.rs', lines 19:0-19:44 *) + Source: 'tests/src/loops.rs', lines 22:0-22:44 *) let sum_with_mut_borrows (max : u32) : result u32 = sum_with_mut_borrows_loop max 0 0 (** [loops::sum_with_shared_borrows]: loop 0: - Source: 'tests/src/loops.rs', lines 34:0-48:1 *) + Source: 'tests/src/loops.rs', lines 37:0-51:1 *) let rec sum_with_shared_borrows_loop (max : u32) (i : u32) (s : u32) : Tot (result u32) (decreases (sum_with_shared_borrows_loop_decreases max i s)) @@ -54,12 +54,12 @@ let rec sum_with_shared_borrows_loop else u32_mul s 2 (** [loops::sum_with_shared_borrows]: - Source: 'tests/src/loops.rs', lines 34:0-34:47 *) + Source: 'tests/src/loops.rs', lines 37:0-37:47 *) let sum_with_shared_borrows (max : u32) : result u32 = sum_with_shared_borrows_loop max 0 0 (** [loops::sum_array]: loop 0: - Source: 'tests/src/loops.rs', lines 50:0-58:1 *) + Source: 'tests/src/loops.rs', lines 53:0-61:1 *) let rec sum_array_loop (n : usize) (a : array u32 n) (i : usize) (s : u32) : Tot (result u32) (decreases (sum_array_loop_decreases n a i s)) @@ -73,12 +73,12 @@ let rec sum_array_loop else Ok s (** [loops::sum_array]: - Source: 'tests/src/loops.rs', lines 50:0-50:52 *) + Source: 'tests/src/loops.rs', lines 53:0-53:52 *) let sum_array (n : usize) (a : array u32 n) : result u32 = sum_array_loop n a 0 0 (** [loops::clear]: loop 0: - Source: 'tests/src/loops.rs', lines 62:0-68:1 *) + Source: 'tests/src/loops.rs', lines 65:0-71:1 *) let rec clear_loop (v : alloc_vec_Vec u32) (i : usize) : Tot (result (alloc_vec_Vec u32)) (decreases (clear_loop_decreases v i)) @@ -95,12 +95,12 @@ let rec clear_loop else Ok v (** [loops::clear]: - Source: 'tests/src/loops.rs', lines 62:0-62:30 *) + Source: 'tests/src/loops.rs', lines 65:0-65:30 *) let clear (v : alloc_vec_Vec u32) : result (alloc_vec_Vec u32) = clear_loop v 0 (** [loops::list_mem]: loop 0: - Source: 'tests/src/loops.rs', lines 76:0-85:1 *) + Source: 'tests/src/loops.rs', lines 79:0-88:1 *) let rec list_mem_loop (x : u32) (ls : list_t u32) : Tot (result bool) (decreases (list_mem_loop_decreases x ls)) @@ -111,12 +111,12 @@ let rec list_mem_loop end (** [loops::list_mem]: - Source: 'tests/src/loops.rs', lines 76:0-76:52 *) + Source: 'tests/src/loops.rs', lines 79:0-79:52 *) let list_mem (x : u32) (ls : list_t u32) : result bool = list_mem_loop x ls (** [loops::list_nth_mut_loop]: loop 0: - Source: 'tests/src/loops.rs', lines 88:0-98:1 *) + Source: 'tests/src/loops.rs', lines 91:0-101:1 *) let rec list_nth_mut_loop_loop (t : Type0) (ls : list_t t) (i : u32) : Tot (result (t & (t -> result (list_t t)))) @@ -135,7 +135,7 @@ let rec list_nth_mut_loop_loop end (** [loops::list_nth_mut_loop]: - Source: 'tests/src/loops.rs', lines 88:0-88:71 *) + Source: 'tests/src/loops.rs', lines 91:0-91:71 *) let list_nth_mut_loop (t : Type0) (ls : list_t t) (i : u32) : result (t & (t -> result (list_t t))) @@ -143,7 +143,7 @@ let list_nth_mut_loop list_nth_mut_loop_loop t ls i (** [loops::list_nth_shared_loop]: loop 0: - Source: 'tests/src/loops.rs', lines 101:0-111:1 *) + Source: 'tests/src/loops.rs', lines 104:0-114:1 *) let rec list_nth_shared_loop_loop (t : Type0) (ls : list_t t) (i : u32) : Tot (result t) (decreases (list_nth_shared_loop_loop_decreases t ls i)) @@ -157,12 +157,12 @@ let rec list_nth_shared_loop_loop end (** [loops::list_nth_shared_loop]: - Source: 'tests/src/loops.rs', lines 101:0-101:66 *) + Source: 'tests/src/loops.rs', lines 104:0-104:66 *) let list_nth_shared_loop (t : Type0) (ls : list_t t) (i : u32) : result t = list_nth_shared_loop_loop t ls i (** [loops::get_elem_mut]: loop 0: - Source: 'tests/src/loops.rs', lines 113:0-127:1 *) + Source: 'tests/src/loops.rs', lines 116:0-130:1 *) let rec get_elem_mut_loop (x : usize) (ls : list_t usize) : Tot (result (usize & (usize -> result (list_t usize)))) @@ -180,7 +180,7 @@ let rec get_elem_mut_loop end (** [loops::get_elem_mut]: - Source: 'tests/src/loops.rs', lines 113:0-113:73 *) + Source: 'tests/src/loops.rs', lines 116:0-116:73 *) let get_elem_mut (slots : alloc_vec_Vec (list_t usize)) (x : usize) : result (usize & (usize -> result (alloc_vec_Vec (list_t usize)))) @@ -193,7 +193,7 @@ let get_elem_mut Ok (i, back1) (** [loops::get_elem_shared]: loop 0: - Source: 'tests/src/loops.rs', lines 129:0-143:1 *) + Source: 'tests/src/loops.rs', lines 132:0-146:1 *) let rec get_elem_shared_loop (x : usize) (ls : list_t usize) : Tot (result usize) (decreases (get_elem_shared_loop_decreases x ls)) @@ -204,7 +204,7 @@ let rec get_elem_shared_loop end (** [loops::get_elem_shared]: - Source: 'tests/src/loops.rs', lines 129:0-129:68 *) + Source: 'tests/src/loops.rs', lines 132:0-132:68 *) let get_elem_shared (slots : alloc_vec_Vec (list_t usize)) (x : usize) : result usize = let* ls = @@ -213,7 +213,7 @@ let get_elem_shared get_elem_shared_loop x ls (** [loops::id_mut]: - Source: 'tests/src/loops.rs', lines 145:0-145:50 *) + Source: 'tests/src/loops.rs', lines 148:0-148:50 *) let id_mut (t : Type0) (ls : list_t t) : result ((list_t t) & (list_t t -> result (list_t t))) @@ -221,12 +221,12 @@ let id_mut Ok (ls, Ok) (** [loops::id_shared]: - Source: 'tests/src/loops.rs', lines 149:0-149:45 *) + Source: 'tests/src/loops.rs', lines 152:0-152:45 *) let id_shared (t : Type0) (ls : list_t t) : result (list_t t) = Ok ls (** [loops::list_nth_mut_loop_with_id]: loop 0: - Source: 'tests/src/loops.rs', lines 154:0-165:1 *) + Source: 'tests/src/loops.rs', lines 157:0-168:1 *) let rec list_nth_mut_loop_with_id_loop (t : Type0) (i : u32) (ls : list_t t) : Tot (result (t & (t -> result (list_t t)))) @@ -245,7 +245,7 @@ let rec list_nth_mut_loop_with_id_loop end (** [loops::list_nth_mut_loop_with_id]: - Source: 'tests/src/loops.rs', lines 154:0-154:75 *) + Source: 'tests/src/loops.rs', lines 157:0-157:75 *) let list_nth_mut_loop_with_id (t : Type0) (ls : list_t t) (i : u32) : result (t & (t -> result (list_t t))) @@ -256,7 +256,7 @@ let list_nth_mut_loop_with_id Ok (x, back1) (** [loops::list_nth_shared_loop_with_id]: loop 0: - Source: 'tests/src/loops.rs', lines 168:0-179:1 *) + Source: 'tests/src/loops.rs', lines 171:0-182:1 *) let rec list_nth_shared_loop_with_id_loop (t : Type0) (i : u32) (ls : list_t t) : Tot (result t) @@ -271,13 +271,13 @@ let rec list_nth_shared_loop_with_id_loop end (** [loops::list_nth_shared_loop_with_id]: - Source: 'tests/src/loops.rs', lines 168:0-168:70 *) + Source: 'tests/src/loops.rs', lines 171:0-171:70 *) let list_nth_shared_loop_with_id (t : Type0) (ls : list_t t) (i : u32) : result t = let* ls1 = id_shared t ls in list_nth_shared_loop_with_id_loop t i ls1 (** [loops::list_nth_mut_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 184:0-205:1 *) + Source: 'tests/src/loops.rs', lines 187:0-208:1 *) let rec list_nth_mut_loop_pair_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result ((t & t) & (t -> result (list_t t)) & (t -> result (list_t t)))) @@ -306,7 +306,7 @@ let rec list_nth_mut_loop_pair_loop end (** [loops::list_nth_mut_loop_pair]: - Source: 'tests/src/loops.rs', lines 184:0-188:27 *) + Source: 'tests/src/loops.rs', lines 187:0-191:27 *) let list_nth_mut_loop_pair (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result ((t & t) & (t -> result (list_t t)) & (t -> result (list_t t))) @@ -314,7 +314,7 @@ let list_nth_mut_loop_pair list_nth_mut_loop_pair_loop t ls0 ls1 i (** [loops::list_nth_shared_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 208:0-229:1 *) + Source: 'tests/src/loops.rs', lines 211:0-232:1 *) let rec list_nth_shared_loop_pair_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result (t & t)) @@ -333,13 +333,13 @@ let rec list_nth_shared_loop_pair_loop end (** [loops::list_nth_shared_loop_pair]: - Source: 'tests/src/loops.rs', lines 208:0-212:19 *) + Source: 'tests/src/loops.rs', lines 211:0-215:19 *) let list_nth_shared_loop_pair (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result (t & t) = list_nth_shared_loop_pair_loop t ls0 ls1 i (** [loops::list_nth_mut_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 233:0-248:1 *) + Source: 'tests/src/loops.rs', lines 236:0-251:1 *) let rec list_nth_mut_loop_pair_merge_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result ((t & t) & ((t & t) -> result ((list_t t) & (list_t t))))) @@ -369,7 +369,7 @@ let rec list_nth_mut_loop_pair_merge_loop end (** [loops::list_nth_mut_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 233:0-237:27 *) + Source: 'tests/src/loops.rs', lines 236:0-240:27 *) let list_nth_mut_loop_pair_merge (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result ((t & t) & ((t & t) -> result ((list_t t) & (list_t t)))) @@ -377,7 +377,7 @@ let list_nth_mut_loop_pair_merge list_nth_mut_loop_pair_merge_loop t ls0 ls1 i (** [loops::list_nth_shared_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 251:0-266:1 *) + Source: 'tests/src/loops.rs', lines 254:0-269:1 *) let rec list_nth_shared_loop_pair_merge_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result (t & t)) @@ -398,13 +398,13 @@ let rec list_nth_shared_loop_pair_merge_loop end (** [loops::list_nth_shared_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 251:0-255:19 *) + Source: 'tests/src/loops.rs', lines 254:0-258:19 *) let list_nth_shared_loop_pair_merge (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result (t & t) = list_nth_shared_loop_pair_merge_loop t ls0 ls1 i (** [loops::list_nth_mut_shared_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 269:0-284:1 *) + Source: 'tests/src/loops.rs', lines 272:0-287:1 *) let rec list_nth_mut_shared_loop_pair_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result ((t & t) & (t -> result (list_t t)))) @@ -428,7 +428,7 @@ let rec list_nth_mut_shared_loop_pair_loop end (** [loops::list_nth_mut_shared_loop_pair]: - Source: 'tests/src/loops.rs', lines 269:0-273:23 *) + Source: 'tests/src/loops.rs', lines 272:0-276:23 *) let list_nth_mut_shared_loop_pair (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result ((t & t) & (t -> result (list_t t))) @@ -436,7 +436,7 @@ let list_nth_mut_shared_loop_pair list_nth_mut_shared_loop_pair_loop t ls0 ls1 i (** [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 288:0-303:1 *) + Source: 'tests/src/loops.rs', lines 291:0-306:1 *) let rec list_nth_mut_shared_loop_pair_merge_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result ((t & t) & (t -> result (list_t t)))) @@ -461,7 +461,7 @@ let rec list_nth_mut_shared_loop_pair_merge_loop end (** [loops::list_nth_mut_shared_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 288:0-292:23 *) + Source: 'tests/src/loops.rs', lines 291:0-295:23 *) let list_nth_mut_shared_loop_pair_merge (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result ((t & t) & (t -> result (list_t t))) @@ -469,7 +469,7 @@ let list_nth_mut_shared_loop_pair_merge list_nth_mut_shared_loop_pair_merge_loop t ls0 ls1 i (** [loops::list_nth_shared_mut_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 307:0-322:1 *) + Source: 'tests/src/loops.rs', lines 310:0-325:1 *) let rec list_nth_shared_mut_loop_pair_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result ((t & t) & (t -> result (list_t t)))) @@ -493,7 +493,7 @@ let rec list_nth_shared_mut_loop_pair_loop end (** [loops::list_nth_shared_mut_loop_pair]: - Source: 'tests/src/loops.rs', lines 307:0-311:23 *) + Source: 'tests/src/loops.rs', lines 310:0-314:23 *) let list_nth_shared_mut_loop_pair (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result ((t & t) & (t -> result (list_t t))) @@ -501,7 +501,7 @@ let list_nth_shared_mut_loop_pair list_nth_shared_mut_loop_pair_loop t ls0 ls1 i (** [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 326:0-341:1 *) + Source: 'tests/src/loops.rs', lines 329:0-344:1 *) let rec list_nth_shared_mut_loop_pair_merge_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result ((t & t) & (t -> result (list_t t)))) @@ -526,7 +526,7 @@ let rec list_nth_shared_mut_loop_pair_merge_loop end (** [loops::list_nth_shared_mut_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 326:0-330:23 *) + Source: 'tests/src/loops.rs', lines 329:0-333:23 *) let list_nth_shared_mut_loop_pair_merge (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result ((t & t) & (t -> result (list_t t))) @@ -534,7 +534,7 @@ let list_nth_shared_mut_loop_pair_merge list_nth_shared_mut_loop_pair_merge_loop t ls0 ls1 i (** [loops::ignore_input_mut_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 345:0-349:1 *) + Source: 'tests/src/loops.rs', lines 348:0-352:1 *) let rec ignore_input_mut_borrow_loop (i : u32) : Tot (result unit) (decreases (ignore_input_mut_borrow_loop_decreases i)) @@ -544,12 +544,12 @@ let rec ignore_input_mut_borrow_loop else Ok () (** [loops::ignore_input_mut_borrow]: - Source: 'tests/src/loops.rs', lines 345:0-345:56 *) + Source: 'tests/src/loops.rs', lines 348:0-348:56 *) let ignore_input_mut_borrow (_a : u32) (i : u32) : result u32 = let* _ = ignore_input_mut_borrow_loop i in Ok _a (** [loops::incr_ignore_input_mut_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 353:0-358:1 *) + Source: 'tests/src/loops.rs', lines 356:0-361:1 *) let rec incr_ignore_input_mut_borrow_loop (i : u32) : Tot (result unit) (decreases (incr_ignore_input_mut_borrow_loop_decreases i)) @@ -559,14 +559,14 @@ let rec incr_ignore_input_mut_borrow_loop else Ok () (** [loops::incr_ignore_input_mut_borrow]: - Source: 'tests/src/loops.rs', lines 353:0-353:60 *) + Source: 'tests/src/loops.rs', lines 356:0-356:60 *) let incr_ignore_input_mut_borrow (a : u32) (i : u32) : result u32 = let* a1 = u32_add a 1 in let* _ = incr_ignore_input_mut_borrow_loop i in Ok a1 (** [loops::ignore_input_shared_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 362:0-366:1 *) + Source: 'tests/src/loops.rs', lines 365:0-369:1 *) let rec ignore_input_shared_borrow_loop (i : u32) : Tot (result unit) (decreases (ignore_input_shared_borrow_loop_decreases i)) @@ -576,7 +576,7 @@ let rec ignore_input_shared_borrow_loop else Ok () (** [loops::ignore_input_shared_borrow]: - Source: 'tests/src/loops.rs', lines 362:0-362:59 *) + Source: 'tests/src/loops.rs', lines 365:0-365:59 *) let ignore_input_shared_borrow (_a : u32) (i : u32) : result u32 = let* _ = ignore_input_shared_borrow_loop i in Ok _a diff --git a/tests/fstar/misc/Loops.Types.fst b/tests/fstar/misc/Loops.Types.fst index 534f5df7..c844d0e0 100644 --- a/tests/fstar/misc/Loops.Types.fst +++ b/tests/fstar/misc/Loops.Types.fst @@ -6,7 +6,7 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [loops::List] - Source: 'tests/src/loops.rs', lines 70:0-70:16 *) + Source: 'tests/src/loops.rs', lines 73:0-73:16 *) type list_t (t : Type0) = | List_Cons : t -> list_t t -> list_t t | List_Nil : list_t t diff --git a/tests/fstar/misc/NoNestedBorrows.fst b/tests/fstar/misc/NoNestedBorrows.fst index 340dd293..b9fbd669 100644 --- a/tests/fstar/misc/NoNestedBorrows.fst +++ b/tests/fstar/misc/NoNestedBorrows.fst @@ -6,54 +6,54 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [no_nested_borrows::Pair] - Source: 'tests/src/no_nested_borrows.rs', lines 4:0-4:23 *) + Source: 'tests/src/no_nested_borrows.rs', lines 6:0-6:23 *) type pair_t (t1 t2 : Type0) = { x : t1; y : t2; } (** [no_nested_borrows::List] - Source: 'tests/src/no_nested_borrows.rs', lines 9:0-9:16 *) + Source: 'tests/src/no_nested_borrows.rs', lines 11:0-11:16 *) type list_t (t : Type0) = | List_Cons : t -> list_t t -> list_t t | List_Nil : list_t t (** [no_nested_borrows::One] - Source: 'tests/src/no_nested_borrows.rs', lines 20:0-20:16 *) + Source: 'tests/src/no_nested_borrows.rs', lines 22:0-22:16 *) type one_t (t1 : Type0) = | One_One : t1 -> one_t t1 (** [no_nested_borrows::EmptyEnum] - Source: 'tests/src/no_nested_borrows.rs', lines 26:0-26:18 *) + Source: 'tests/src/no_nested_borrows.rs', lines 28:0-28:18 *) type emptyEnum_t = | EmptyEnum_Empty : emptyEnum_t (** [no_nested_borrows::Enum] - Source: 'tests/src/no_nested_borrows.rs', lines 32:0-32:13 *) + Source: 'tests/src/no_nested_borrows.rs', lines 34:0-34:13 *) type enum_t = | Enum_Variant1 : enum_t | Enum_Variant2 : enum_t (** [no_nested_borrows::EmptyStruct] - Source: 'tests/src/no_nested_borrows.rs', lines 39:0-39:22 *) + Source: 'tests/src/no_nested_borrows.rs', lines 41:0-41:22 *) type emptyStruct_t = unit (** [no_nested_borrows::Sum] - Source: 'tests/src/no_nested_borrows.rs', lines 41:0-41:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 43:0-43:20 *) type sum_t (t1 t2 : Type0) = | Sum_Left : t1 -> sum_t t1 t2 | Sum_Right : t2 -> sum_t t1 t2 (** [no_nested_borrows::cast_u32_to_i32]: - Source: 'tests/src/no_nested_borrows.rs', lines 46:0-46:37 *) + Source: 'tests/src/no_nested_borrows.rs', lines 48:0-48:37 *) let cast_u32_to_i32 (x : u32) : result i32 = scalar_cast U32 I32 x (** [no_nested_borrows::cast_bool_to_i32]: - Source: 'tests/src/no_nested_borrows.rs', lines 50:0-50:39 *) + Source: 'tests/src/no_nested_borrows.rs', lines 52:0-52:39 *) let cast_bool_to_i32 (x : bool) : result i32 = scalar_cast_bool I32 x (** [no_nested_borrows::cast_bool_to_bool]: - Source: 'tests/src/no_nested_borrows.rs', lines 55:0-55:41 *) + Source: 'tests/src/no_nested_borrows.rs', lines 57:0-57:41 *) let cast_bool_to_bool (x : bool) : result bool = Ok x (** [no_nested_borrows::test2]: - Source: 'tests/src/no_nested_borrows.rs', lines 60:0-60:14 *) + Source: 'tests/src/no_nested_borrows.rs', lines 62:0-62:14 *) let test2 : result unit = let* _ = u32_add 23 44 in Ok () @@ -61,12 +61,12 @@ let test2 : result unit = let _ = assert_norm (test2 = Ok ()) (** [no_nested_borrows::get_max]: - Source: 'tests/src/no_nested_borrows.rs', lines 72:0-72:37 *) + Source: 'tests/src/no_nested_borrows.rs', lines 74:0-74:37 *) let get_max (x : u32) (y : u32) : result u32 = if x >= y then Ok x else Ok y (** [no_nested_borrows::test3]: - Source: 'tests/src/no_nested_borrows.rs', lines 80:0-80:14 *) + Source: 'tests/src/no_nested_borrows.rs', lines 82:0-82:14 *) let test3 : result unit = let* x = get_max 4 3 in let* y = get_max 10 11 in @@ -77,7 +77,7 @@ let test3 : result unit = let _ = assert_norm (test3 = Ok ()) (** [no_nested_borrows::test_neg1]: - Source: 'tests/src/no_nested_borrows.rs', lines 87:0-87:18 *) + Source: 'tests/src/no_nested_borrows.rs', lines 89:0-89:18 *) let test_neg1 : result unit = let* y = i32_neg 3 in if not (y = -3) then Fail Failure else Ok () @@ -85,7 +85,7 @@ let test_neg1 : result unit = let _ = assert_norm (test_neg1 = Ok ()) (** [no_nested_borrows::refs_test1]: - Source: 'tests/src/no_nested_borrows.rs', lines 94:0-94:19 *) + Source: 'tests/src/no_nested_borrows.rs', lines 96:0-96:19 *) let refs_test1 : result unit = if not (1 = 1) then Fail Failure else Ok () @@ -93,7 +93,7 @@ let refs_test1 : result unit = let _ = assert_norm (refs_test1 = Ok ()) (** [no_nested_borrows::refs_test2]: - Source: 'tests/src/no_nested_borrows.rs', lines 105:0-105:19 *) + Source: 'tests/src/no_nested_borrows.rs', lines 107:0-107:19 *) let refs_test2 : result unit = if not (2 = 2) then Fail Failure @@ -109,7 +109,7 @@ let refs_test2 : result unit = let _ = assert_norm (refs_test2 = Ok ()) (** [no_nested_borrows::test_list1]: - Source: 'tests/src/no_nested_borrows.rs', lines 121:0-121:19 *) + Source: 'tests/src/no_nested_borrows.rs', lines 123:0-123:19 *) let test_list1 : result unit = Ok () @@ -117,7 +117,7 @@ let test_list1 : result unit = let _ = assert_norm (test_list1 = Ok ()) (** [no_nested_borrows::test_box1]: - Source: 'tests/src/no_nested_borrows.rs', lines 126:0-126:18 *) + Source: 'tests/src/no_nested_borrows.rs', lines 128:0-128:18 *) let test_box1 : result unit = let* (_, deref_mut_back) = alloc_boxed_Box_deref_mut i32 0 in let* b = deref_mut_back 1 in @@ -128,22 +128,22 @@ let test_box1 : result unit = let _ = assert_norm (test_box1 = Ok ()) (** [no_nested_borrows::copy_int]: - Source: 'tests/src/no_nested_borrows.rs', lines 136:0-136:30 *) + Source: 'tests/src/no_nested_borrows.rs', lines 138:0-138:30 *) let copy_int (x : i32) : result i32 = Ok x (** [no_nested_borrows::test_unreachable]: - Source: 'tests/src/no_nested_borrows.rs', lines 142:0-142:32 *) + Source: 'tests/src/no_nested_borrows.rs', lines 144:0-144:32 *) let test_unreachable (b : bool) : result unit = if b then Fail Failure else Ok () (** [no_nested_borrows::test_panic]: - Source: 'tests/src/no_nested_borrows.rs', lines 150:0-150:26 *) + Source: 'tests/src/no_nested_borrows.rs', lines 152:0-152:26 *) let test_panic (b : bool) : result unit = if b then Fail Failure else Ok () (** [no_nested_borrows::test_copy_int]: - Source: 'tests/src/no_nested_borrows.rs', lines 157:0-157:22 *) + Source: 'tests/src/no_nested_borrows.rs', lines 159:0-159:22 *) let test_copy_int : result unit = let* y = copy_int 0 in if not (0 = y) then Fail Failure else Ok () @@ -151,12 +151,12 @@ let test_copy_int : result unit = let _ = assert_norm (test_copy_int = Ok ()) (** [no_nested_borrows::is_cons]: - Source: 'tests/src/no_nested_borrows.rs', lines 164:0-164:38 *) + Source: 'tests/src/no_nested_borrows.rs', lines 166:0-166:38 *) let is_cons (t : Type0) (l : list_t t) : result bool = begin match l with | List_Cons _ _ -> Ok true | List_Nil -> Ok false end (** [no_nested_borrows::test_is_cons]: - Source: 'tests/src/no_nested_borrows.rs', lines 171:0-171:21 *) + Source: 'tests/src/no_nested_borrows.rs', lines 173:0-173:21 *) let test_is_cons : result unit = let* b = is_cons i32 (List_Cons 0 List_Nil) in if not b then Fail Failure else Ok () @@ -165,7 +165,7 @@ let test_is_cons : result unit = let _ = assert_norm (test_is_cons = Ok ()) (** [no_nested_borrows::split_list]: - Source: 'tests/src/no_nested_borrows.rs', lines 177:0-177:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 179:0-179:48 *) let split_list (t : Type0) (l : list_t t) : result (t & (list_t t)) = begin match l with | List_Cons hd tl -> Ok (hd, tl) @@ -173,7 +173,7 @@ let split_list (t : Type0) (l : list_t t) : result (t & (list_t t)) = end (** [no_nested_borrows::test_split_list]: - Source: 'tests/src/no_nested_borrows.rs', lines 185:0-185:24 *) + Source: 'tests/src/no_nested_borrows.rs', lines 187:0-187:24 *) let test_split_list : result unit = let* p = split_list i32 (List_Cons 0 List_Nil) in let (hd, _) = p in @@ -183,7 +183,7 @@ let test_split_list : result unit = let _ = assert_norm (test_split_list = Ok ()) (** [no_nested_borrows::choose]: - Source: 'tests/src/no_nested_borrows.rs', lines 192:0-192:70 *) + Source: 'tests/src/no_nested_borrows.rs', lines 194:0-194:70 *) let choose (t : Type0) (b : bool) (x : t) (y : t) : result (t & (t -> result (t & t))) = if b @@ -191,7 +191,7 @@ let choose else let back = fun ret -> Ok (x, ret) in Ok (y, back) (** [no_nested_borrows::choose_test]: - Source: 'tests/src/no_nested_borrows.rs', lines 200:0-200:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 202:0-202:20 *) let choose_test : result unit = let* (z, choose_back) = choose i32 true 0 0 in let* z1 = i32_add z 1 in @@ -207,24 +207,24 @@ let choose_test : result unit = let _ = assert_norm (choose_test = Ok ()) (** [no_nested_borrows::test_char]: - Source: 'tests/src/no_nested_borrows.rs', lines 212:0-212:26 *) + Source: 'tests/src/no_nested_borrows.rs', lines 214:0-214:26 *) let test_char : result char = Ok 'a' (** [no_nested_borrows::Tree] - Source: 'tests/src/no_nested_borrows.rs', lines 217:0-217:16 *) + Source: 'tests/src/no_nested_borrows.rs', lines 219:0-219:16 *) type tree_t (t : Type0) = | Tree_Leaf : t -> tree_t t | Tree_Node : t -> nodeElem_t t -> tree_t t -> tree_t t (** [no_nested_borrows::NodeElem] - Source: 'tests/src/no_nested_borrows.rs', lines 222:0-222:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 224:0-224:20 *) and nodeElem_t (t : Type0) = | NodeElem_Cons : tree_t t -> nodeElem_t t -> nodeElem_t t | NodeElem_Nil : nodeElem_t t (** [no_nested_borrows::list_length]: - Source: 'tests/src/no_nested_borrows.rs', lines 257:0-257:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 259:0-259:48 *) let rec list_length (t : Type0) (l : list_t t) : result u32 = begin match l with | List_Cons _ l1 -> let* i = list_length t l1 in u32_add 1 i @@ -232,7 +232,7 @@ let rec list_length (t : Type0) (l : list_t t) : result u32 = end (** [no_nested_borrows::list_nth_shared]: - Source: 'tests/src/no_nested_borrows.rs', lines 265:0-265:62 *) + Source: 'tests/src/no_nested_borrows.rs', lines 267:0-267:62 *) let rec list_nth_shared (t : Type0) (l : list_t t) (i : u32) : result t = begin match l with | List_Cons x tl -> @@ -241,7 +241,7 @@ let rec list_nth_shared (t : Type0) (l : list_t t) (i : u32) : result t = end (** [no_nested_borrows::list_nth_mut]: - Source: 'tests/src/no_nested_borrows.rs', lines 281:0-281:67 *) + Source: 'tests/src/no_nested_borrows.rs', lines 283:0-283:67 *) let rec list_nth_mut (t : Type0) (l : list_t t) (i : u32) : result (t & (t -> result (list_t t))) @@ -260,7 +260,7 @@ let rec list_nth_mut end (** [no_nested_borrows::list_rev_aux]: - Source: 'tests/src/no_nested_borrows.rs', lines 297:0-297:63 *) + Source: 'tests/src/no_nested_borrows.rs', lines 299:0-299:63 *) let rec list_rev_aux (t : Type0) (li : list_t t) (lo : list_t t) : result (list_t t) = begin match li with @@ -269,13 +269,13 @@ let rec list_rev_aux end (** [no_nested_borrows::list_rev]: - Source: 'tests/src/no_nested_borrows.rs', lines 311:0-311:42 *) + Source: 'tests/src/no_nested_borrows.rs', lines 313:0-313:42 *) let list_rev (t : Type0) (l : list_t t) : result (list_t t) = let (li, _) = core_mem_replace (list_t t) l List_Nil in list_rev_aux t li List_Nil (** [no_nested_borrows::test_list_functions]: - Source: 'tests/src/no_nested_borrows.rs', lines 316:0-316:28 *) + Source: 'tests/src/no_nested_borrows.rs', lines 318:0-318:28 *) let test_list_functions : result unit = let l = List_Cons 2 List_Nil in let l1 = List_Cons 1 l in @@ -312,7 +312,7 @@ let test_list_functions : result unit = let _ = assert_norm (test_list_functions = Ok ()) (** [no_nested_borrows::id_mut_pair1]: - Source: 'tests/src/no_nested_borrows.rs', lines 332:0-332:89 *) + Source: 'tests/src/no_nested_borrows.rs', lines 334:0-334:89 *) let id_mut_pair1 (t1 t2 : Type0) (x : t1) (y : t2) : result ((t1 & t2) & ((t1 & t2) -> result (t1 & t2))) @@ -320,7 +320,7 @@ let id_mut_pair1 Ok ((x, y), Ok) (** [no_nested_borrows::id_mut_pair2]: - Source: 'tests/src/no_nested_borrows.rs', lines 336:0-336:88 *) + Source: 'tests/src/no_nested_borrows.rs', lines 338:0-338:88 *) let id_mut_pair2 (t1 t2 : Type0) (p : (t1 & t2)) : result ((t1 & t2) & ((t1 & t2) -> result (t1 & t2))) @@ -328,7 +328,7 @@ let id_mut_pair2 let (x, x1) = p in Ok ((x, x1), Ok) (** [no_nested_borrows::id_mut_pair3]: - Source: 'tests/src/no_nested_borrows.rs', lines 340:0-340:93 *) + Source: 'tests/src/no_nested_borrows.rs', lines 342:0-342:93 *) let id_mut_pair3 (t1 t2 : Type0) (x : t1) (y : t2) : result ((t1 & t2) & (t1 -> result t1) & (t2 -> result t2)) @@ -336,7 +336,7 @@ let id_mut_pair3 Ok ((x, y), Ok, Ok) (** [no_nested_borrows::id_mut_pair4]: - Source: 'tests/src/no_nested_borrows.rs', lines 344:0-344:92 *) + Source: 'tests/src/no_nested_borrows.rs', lines 346:0-346:92 *) let id_mut_pair4 (t1 t2 : Type0) (p : (t1 & t2)) : result ((t1 & t2) & (t1 -> result t1) & (t2 -> result t2)) @@ -344,35 +344,35 @@ let id_mut_pair4 let (x, x1) = p in Ok ((x, x1), Ok, Ok) (** [no_nested_borrows::StructWithTuple] - Source: 'tests/src/no_nested_borrows.rs', lines 351:0-351:34 *) + Source: 'tests/src/no_nested_borrows.rs', lines 353:0-353:34 *) type structWithTuple_t (t1 t2 : Type0) = { p : (t1 & t2); } (** [no_nested_borrows::new_tuple1]: - Source: 'tests/src/no_nested_borrows.rs', lines 355:0-355:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 357:0-357:48 *) let new_tuple1 : result (structWithTuple_t u32 u32) = Ok { p = (1, 2) } (** [no_nested_borrows::new_tuple2]: - Source: 'tests/src/no_nested_borrows.rs', lines 359:0-359:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 361:0-361:48 *) let new_tuple2 : result (structWithTuple_t i16 i16) = Ok { p = (1, 2) } (** [no_nested_borrows::new_tuple3]: - Source: 'tests/src/no_nested_borrows.rs', lines 363:0-363:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 365:0-365:48 *) let new_tuple3 : result (structWithTuple_t u64 i64) = Ok { p = (1, 2) } (** [no_nested_borrows::StructWithPair] - Source: 'tests/src/no_nested_borrows.rs', lines 368:0-368:33 *) + Source: 'tests/src/no_nested_borrows.rs', lines 370:0-370:33 *) type structWithPair_t (t1 t2 : Type0) = { p : pair_t t1 t2; } (** [no_nested_borrows::new_pair1]: - Source: 'tests/src/no_nested_borrows.rs', lines 372:0-372:46 *) + Source: 'tests/src/no_nested_borrows.rs', lines 374:0-374:46 *) let new_pair1 : result (structWithPair_t u32 u32) = Ok { p = { x = 1; y = 2 } } (** [no_nested_borrows::test_constants]: - Source: 'tests/src/no_nested_borrows.rs', lines 380:0-380:23 *) + Source: 'tests/src/no_nested_borrows.rs', lines 382:0-382:23 *) let test_constants : result unit = let* swt = new_tuple1 in let (i, _) = swt.p in @@ -396,7 +396,7 @@ let test_constants : result unit = let _ = assert_norm (test_constants = Ok ()) (** [no_nested_borrows::test_weird_borrows1]: - Source: 'tests/src/no_nested_borrows.rs', lines 389:0-389:28 *) + Source: 'tests/src/no_nested_borrows.rs', lines 391:0-391:28 *) let test_weird_borrows1 : result unit = Ok () @@ -404,71 +404,71 @@ let test_weird_borrows1 : result unit = let _ = assert_norm (test_weird_borrows1 = Ok ()) (** [no_nested_borrows::test_mem_replace]: - Source: 'tests/src/no_nested_borrows.rs', lines 399:0-399:37 *) + Source: 'tests/src/no_nested_borrows.rs', lines 401:0-401:37 *) let test_mem_replace (px : u32) : result u32 = let (y, _) = core_mem_replace u32 px 1 in if not (y = 0) then Fail Failure else Ok 2 (** [no_nested_borrows::test_shared_borrow_bool1]: - Source: 'tests/src/no_nested_borrows.rs', lines 406:0-406:47 *) + Source: 'tests/src/no_nested_borrows.rs', lines 408:0-408:47 *) let test_shared_borrow_bool1 (b : bool) : result u32 = if b then Ok 0 else Ok 1 (** [no_nested_borrows::test_shared_borrow_bool2]: - Source: 'tests/src/no_nested_borrows.rs', lines 419:0-419:40 *) + Source: 'tests/src/no_nested_borrows.rs', lines 421:0-421:40 *) let test_shared_borrow_bool2 : result u32 = Ok 0 (** [no_nested_borrows::test_shared_borrow_enum1]: - Source: 'tests/src/no_nested_borrows.rs', lines 434:0-434:52 *) + Source: 'tests/src/no_nested_borrows.rs', lines 436:0-436:52 *) let test_shared_borrow_enum1 (l : list_t u32) : result u32 = begin match l with | List_Cons _ _ -> Ok 1 | List_Nil -> Ok 0 end (** [no_nested_borrows::test_shared_borrow_enum2]: - Source: 'tests/src/no_nested_borrows.rs', lines 446:0-446:40 *) + Source: 'tests/src/no_nested_borrows.rs', lines 448:0-448:40 *) let test_shared_borrow_enum2 : result u32 = Ok 0 (** [no_nested_borrows::incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 457:0-457:24 *) + Source: 'tests/src/no_nested_borrows.rs', lines 459:0-459:24 *) let incr (x : u32) : result u32 = u32_add x 1 (** [no_nested_borrows::call_incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 461:0-461:35 *) + Source: 'tests/src/no_nested_borrows.rs', lines 463:0-463:35 *) let call_incr (x : u32) : result u32 = incr x (** [no_nested_borrows::read_then_incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 466:0-466:41 *) + Source: 'tests/src/no_nested_borrows.rs', lines 468:0-468:41 *) let read_then_incr (x : u32) : result (u32 & u32) = let* x1 = u32_add x 1 in Ok (x, x1) (** [no_nested_borrows::Tuple] - Source: 'tests/src/no_nested_borrows.rs', lines 472:0-472:24 *) + Source: 'tests/src/no_nested_borrows.rs', lines 474:0-474:24 *) type tuple_t (t1 t2 : Type0) = t1 * t2 (** [no_nested_borrows::use_tuple_struct]: - Source: 'tests/src/no_nested_borrows.rs', lines 474:0-474:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 476:0-476:48 *) let use_tuple_struct (x : tuple_t u32 u32) : result (tuple_t u32 u32) = let (_, i) = x in Ok (1, i) (** [no_nested_borrows::create_tuple_struct]: - Source: 'tests/src/no_nested_borrows.rs', lines 478:0-478:61 *) + Source: 'tests/src/no_nested_borrows.rs', lines 480:0-480:61 *) let create_tuple_struct (x : u32) (y : u64) : result (tuple_t u32 u64) = Ok (x, y) (** [no_nested_borrows::IdType] - Source: 'tests/src/no_nested_borrows.rs', lines 483:0-483:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 485:0-485:20 *) type idType_t (t : Type0) = t (** [no_nested_borrows::use_id_type]: - Source: 'tests/src/no_nested_borrows.rs', lines 485:0-485:40 *) + Source: 'tests/src/no_nested_borrows.rs', lines 487:0-487:40 *) let use_id_type (t : Type0) (x : idType_t t) : result t = Ok x (** [no_nested_borrows::create_id_type]: - Source: 'tests/src/no_nested_borrows.rs', lines 489:0-489:43 *) + Source: 'tests/src/no_nested_borrows.rs', lines 491:0-491:43 *) let create_id_type (t : Type0) (x : t) : result (idType_t t) = Ok x diff --git a/tests/fstar/misc/Paper.fst b/tests/fstar/misc/Paper.fst index 2e74e21a..c78293f1 100644 --- a/tests/fstar/misc/Paper.fst +++ b/tests/fstar/misc/Paper.fst @@ -6,12 +6,12 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [paper::ref_incr]: - Source: 'tests/src/paper.rs', lines 4:0-4:28 *) + Source: 'tests/src/paper.rs', lines 6:0-6:28 *) let ref_incr (x : i32) : result i32 = i32_add x 1 (** [paper::test_incr]: - Source: 'tests/src/paper.rs', lines 8:0-8:18 *) + Source: 'tests/src/paper.rs', lines 10:0-10:18 *) let test_incr : result unit = let* x = ref_incr 0 in if not (x = 1) then Fail Failure else Ok () @@ -19,7 +19,7 @@ let test_incr : result unit = let _ = assert_norm (test_incr = Ok ()) (** [paper::choose]: - Source: 'tests/src/paper.rs', lines 15:0-15:70 *) + Source: 'tests/src/paper.rs', lines 17:0-17:70 *) let choose (t : Type0) (b : bool) (x : t) (y : t) : result (t & (t -> result (t & t))) = if b @@ -27,7 +27,7 @@ let choose else let back = fun ret -> Ok (x, ret) in Ok (y, back) (** [paper::test_choose]: - Source: 'tests/src/paper.rs', lines 23:0-23:20 *) + Source: 'tests/src/paper.rs', lines 25:0-25:20 *) let test_choose : result unit = let* (z, choose_back) = choose i32 true 0 0 in let* z1 = i32_add z 1 in @@ -43,13 +43,13 @@ let test_choose : result unit = let _ = assert_norm (test_choose = Ok ()) (** [paper::List] - Source: 'tests/src/paper.rs', lines 35:0-35:16 *) + Source: 'tests/src/paper.rs', lines 37:0-37:16 *) type list_t (t : Type0) = | List_Cons : t -> list_t t -> list_t t | List_Nil : list_t t (** [paper::list_nth_mut]: - Source: 'tests/src/paper.rs', lines 42:0-42:67 *) + Source: 'tests/src/paper.rs', lines 44:0-44:67 *) let rec list_nth_mut (t : Type0) (l : list_t t) (i : u32) : result (t & (t -> result (list_t t))) @@ -68,7 +68,7 @@ let rec list_nth_mut end (** [paper::sum]: - Source: 'tests/src/paper.rs', lines 57:0-57:32 *) + Source: 'tests/src/paper.rs', lines 59:0-59:32 *) let rec sum (l : list_t i32) : result i32 = begin match l with | List_Cons x tl -> let* i = sum tl in i32_add x i @@ -76,7 +76,7 @@ let rec sum (l : list_t i32) : result i32 = end (** [paper::test_nth]: - Source: 'tests/src/paper.rs', lines 68:0-68:17 *) + Source: 'tests/src/paper.rs', lines 70:0-70:17 *) let test_nth : result unit = let l = List_Cons 3 List_Nil in let l1 = List_Cons 2 l in @@ -90,7 +90,7 @@ let test_nth : result unit = let _ = assert_norm (test_nth = Ok ()) (** [paper::call_choose]: - Source: 'tests/src/paper.rs', lines 76:0-76:44 *) + Source: 'tests/src/paper.rs', lines 78:0-78:44 *) let call_choose (p : (u32 & u32)) : result u32 = let (px, py) = p in let* (pz, choose_back) = choose u32 true px py in diff --git a/tests/fstar/misc/PoloniusList.fst b/tests/fstar/misc/PoloniusList.fst index 071fd3ef..94c1b445 100644 --- a/tests/fstar/misc/PoloniusList.fst +++ b/tests/fstar/misc/PoloniusList.fst @@ -6,13 +6,13 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [polonius_list::List] - Source: 'tests/src/polonius_list.rs', lines 3:0-3:16 *) + Source: 'tests/src/polonius_list.rs', lines 5:0-5:16 *) type list_t (t : Type0) = | List_Cons : t -> list_t t -> list_t t | List_Nil : list_t t (** [polonius_list::get_list_at_x]: - Source: 'tests/src/polonius_list.rs', lines 13:0-13:76 *) + Source: 'tests/src/polonius_list.rs', lines 15:0-15:76 *) let rec get_list_at_x (ls : list_t u32) (x : u32) : result ((list_t u32) & (list_t u32 -> result (list_t u32))) diff --git a/tests/fstar/traits/Traits.fst b/tests/fstar/traits/Traits.fst index fa841e8a..70c345ba 100644 --- a/tests/fstar/traits/Traits.fst +++ b/tests/fstar/traits/Traits.fst @@ -6,20 +6,20 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** Trait declaration: [traits::BoolTrait] - Source: 'tests/src/traits.rs', lines 1:0-1:19 *) + Source: 'tests/src/traits.rs', lines 2:0-2:19 *) noeq type boolTrait_t (self : Type0) = { get_bool : self -> result bool; } (** [traits::{(traits::BoolTrait for bool)}::get_bool]: - Source: 'tests/src/traits.rs', lines 12:4-12:30 *) + Source: 'tests/src/traits.rs', lines 13:4-13:30 *) let boolTraitBool_get_bool (self : bool) : result bool = Ok self (** Trait implementation: [traits::{(traits::BoolTrait for bool)}] - Source: 'tests/src/traits.rs', lines 11:0-11:23 *) + Source: 'tests/src/traits.rs', lines 12:0-12:23 *) let boolTraitBool : boolTrait_t bool = { get_bool = boolTraitBool_get_bool; } (** [traits::BoolTrait::ret_true]: - Source: 'tests/src/traits.rs', lines 6:4-6:30 *) + Source: 'tests/src/traits.rs', lines 7:4-7:30 *) let boolTrait_ret_true (#self : Type0) (self_clause : boolTrait_t self) (self1 : self) : result bool @@ -27,49 +27,49 @@ let boolTrait_ret_true Ok true (** [traits::test_bool_trait_bool]: - Source: 'tests/src/traits.rs', lines 17:0-17:44 *) + Source: 'tests/src/traits.rs', lines 18:0-18:44 *) let test_bool_trait_bool (x : bool) : result bool = let* b = boolTraitBool_get_bool x in if b then boolTrait_ret_true boolTraitBool x else Ok false (** [traits::{(traits::BoolTrait for core::option::Option)#1}::get_bool]: - Source: 'tests/src/traits.rs', lines 23:4-23:30 *) + Source: 'tests/src/traits.rs', lines 24:4-24:30 *) let boolTraitOption_get_bool (t : Type0) (self : option t) : result bool = begin match self with | None -> Ok false | Some _ -> Ok true end (** Trait implementation: [traits::{(traits::BoolTrait for core::option::Option)#1}] - Source: 'tests/src/traits.rs', lines 22:0-22:31 *) + Source: 'tests/src/traits.rs', lines 23:0-23:31 *) let boolTraitOption (t : Type0) : boolTrait_t (option t) = { get_bool = boolTraitOption_get_bool t; } (** [traits::test_bool_trait_option]: - Source: 'tests/src/traits.rs', lines 31:0-31:54 *) + Source: 'tests/src/traits.rs', lines 32:0-32:54 *) let test_bool_trait_option (t : Type0) (x : option t) : result bool = let* b = boolTraitOption_get_bool t x in if b then boolTrait_ret_true (boolTraitOption t) x else Ok false (** [traits::test_bool_trait]: - Source: 'tests/src/traits.rs', lines 35:0-35:50 *) + Source: 'tests/src/traits.rs', lines 36:0-36:50 *) let test_bool_trait (t : Type0) (boolTraitInst : boolTrait_t t) (x : t) : result bool = boolTraitInst.get_bool x (** Trait declaration: [traits::ToU64] - Source: 'tests/src/traits.rs', lines 39:0-39:15 *) + Source: 'tests/src/traits.rs', lines 40:0-40:15 *) noeq type toU64_t (self : Type0) = { to_u64 : self -> result u64; } (** [traits::{(traits::ToU64 for u64)#2}::to_u64]: - Source: 'tests/src/traits.rs', lines 44:4-44:26 *) + Source: 'tests/src/traits.rs', lines 45:4-45:26 *) let toU64U64_to_u64 (self : u64) : result u64 = Ok self (** Trait implementation: [traits::{(traits::ToU64 for u64)#2}] - Source: 'tests/src/traits.rs', lines 43:0-43:18 *) + Source: 'tests/src/traits.rs', lines 44:0-44:18 *) let toU64U64 : toU64_t u64 = { to_u64 = toU64U64_to_u64; } (** [traits::{(traits::ToU64 for (A, A))#3}::to_u64]: - Source: 'tests/src/traits.rs', lines 50:4-50:26 *) + Source: 'tests/src/traits.rs', lines 51:4-51:26 *) let toU64Pair_to_u64 (a : Type0) (toU64Inst : toU64_t a) (self : (a & a)) : result u64 = let (x, x1) = self in @@ -78,75 +78,75 @@ let toU64Pair_to_u64 u64_add i i1 (** Trait implementation: [traits::{(traits::ToU64 for (A, A))#3}] - Source: 'tests/src/traits.rs', lines 49:0-49:31 *) + Source: 'tests/src/traits.rs', lines 50:0-50:31 *) let toU64Pair (a : Type0) (toU64Inst : toU64_t a) : toU64_t (a & a) = { to_u64 = toU64Pair_to_u64 a toU64Inst; } (** [traits::f]: - Source: 'tests/src/traits.rs', lines 55:0-55:36 *) + Source: 'tests/src/traits.rs', lines 56:0-56:36 *) let f (t : Type0) (toU64Inst : toU64_t t) (x : (t & t)) : result u64 = toU64Pair_to_u64 t toU64Inst x (** [traits::g]: - Source: 'tests/src/traits.rs', lines 59:0-61:18 *) + Source: 'tests/src/traits.rs', lines 60:0-62:18 *) let g (t : Type0) (toU64PairInst : toU64_t (t & t)) (x : (t & t)) : result u64 = toU64PairInst.to_u64 x (** [traits::h0]: - Source: 'tests/src/traits.rs', lines 66:0-66:24 *) + Source: 'tests/src/traits.rs', lines 67:0-67:24 *) let h0 (x : u64) : result u64 = toU64U64_to_u64 x (** [traits::Wrapper] - Source: 'tests/src/traits.rs', lines 70:0-70:21 *) + Source: 'tests/src/traits.rs', lines 71:0-71:21 *) type wrapper_t (t : Type0) = { x : t; } (** [traits::{(traits::ToU64 for traits::Wrapper)#4}::to_u64]: - Source: 'tests/src/traits.rs', lines 75:4-75:26 *) + Source: 'tests/src/traits.rs', lines 76:4-76:26 *) let toU64traitsWrapper_to_u64 (t : Type0) (toU64Inst : toU64_t t) (self : wrapper_t t) : result u64 = toU64Inst.to_u64 self.x (** Trait implementation: [traits::{(traits::ToU64 for traits::Wrapper)#4}] - Source: 'tests/src/traits.rs', lines 74:0-74:35 *) + Source: 'tests/src/traits.rs', lines 75:0-75:35 *) let toU64traitsWrapper (t : Type0) (toU64Inst : toU64_t t) : toU64_t (wrapper_t t) = { to_u64 = toU64traitsWrapper_to_u64 t toU64Inst; } (** [traits::h1]: - Source: 'tests/src/traits.rs', lines 80:0-80:33 *) + Source: 'tests/src/traits.rs', lines 81:0-81:33 *) let h1 (x : wrapper_t u64) : result u64 = toU64traitsWrapper_to_u64 u64 toU64U64 x (** [traits::h2]: - Source: 'tests/src/traits.rs', lines 84:0-84:41 *) + Source: 'tests/src/traits.rs', lines 85:0-85:41 *) let h2 (t : Type0) (toU64Inst : toU64_t t) (x : wrapper_t t) : result u64 = toU64traitsWrapper_to_u64 t toU64Inst x (** Trait declaration: [traits::ToType] - Source: 'tests/src/traits.rs', lines 88:0-88:19 *) + Source: 'tests/src/traits.rs', lines 89:0-89:19 *) noeq type toType_t (self t : Type0) = { to_type : self -> result t; } (** [traits::{(traits::ToType for u64)#5}::to_type]: - Source: 'tests/src/traits.rs', lines 93:4-93:28 *) + Source: 'tests/src/traits.rs', lines 94:4-94:28 *) let toTypeU64Bool_to_type (self : u64) : result bool = Ok (self > 0) (** Trait implementation: [traits::{(traits::ToType for u64)#5}] - Source: 'tests/src/traits.rs', lines 92:0-92:25 *) + Source: 'tests/src/traits.rs', lines 93:0-93:25 *) let toTypeU64Bool : toType_t u64 bool = { to_type = toTypeU64Bool_to_type; } (** Trait declaration: [traits::OfType] - Source: 'tests/src/traits.rs', lines 98:0-98:16 *) + Source: 'tests/src/traits.rs', lines 99:0-99:16 *) noeq type ofType_t (self : Type0) = { of_type : (t : Type0) -> (toTypeInst : toType_t t self) -> t -> result self; } (** [traits::h3]: - Source: 'tests/src/traits.rs', lines 104:0-104:50 *) + Source: 'tests/src/traits.rs', lines 105:0-105:50 *) let h3 (t1 t2 : Type0) (ofTypeInst : ofType_t t1) (toTypeInst : toType_t t2 t1) (y : t2) : @@ -155,14 +155,14 @@ let h3 ofTypeInst.of_type t2 toTypeInst y (** Trait declaration: [traits::OfTypeBis] - Source: 'tests/src/traits.rs', lines 109:0-109:36 *) + Source: 'tests/src/traits.rs', lines 110:0-110:36 *) noeq type ofTypeBis_t (self t : Type0) = { toTypeInst : toType_t t self; of_type : t -> result self; } (** [traits::h4]: - Source: 'tests/src/traits.rs', lines 118:0-118:57 *) + Source: 'tests/src/traits.rs', lines 119:0-119:57 *) let h4 (t1 t2 : Type0) (ofTypeBisInst : ofTypeBis_t t1 t2) (toTypeInst : toType_t t2 t1) (y : t2) : @@ -171,34 +171,34 @@ let h4 ofTypeBisInst.of_type y (** [traits::TestType] - Source: 'tests/src/traits.rs', lines 122:0-122:22 *) + Source: 'tests/src/traits.rs', lines 123:0-123:22 *) type testType_t (t : Type0) = t (** [traits::{traits::TestType#6}::test::TestType1] - Source: 'tests/src/traits.rs', lines 127:8-127:24 *) + Source: 'tests/src/traits.rs', lines 128:8-128:24 *) type testType_test_TestType1_t = u64 (** Trait declaration: [traits::{traits::TestType#6}::test::TestTrait] - Source: 'tests/src/traits.rs', lines 128:8-128:23 *) + Source: 'tests/src/traits.rs', lines 129:8-129:23 *) noeq type testType_test_TestTrait_t (self : Type0) = { test : self -> result bool; } (** [traits::{traits::TestType#6}::test::{(traits::{traits::TestType#6}::test::TestTrait for traits::{traits::TestType#6}::test::TestType1)}::test]: - Source: 'tests/src/traits.rs', lines 139:12-139:34 *) + Source: 'tests/src/traits.rs', lines 140:12-140:34 *) let testType_test_TestTraittraitsTestTypetestTestType1_test (self : testType_test_TestType1_t) : result bool = Ok (self > 1) (** Trait implementation: [traits::{traits::TestType#6}::test::{(traits::{traits::TestType#6}::test::TestTrait for traits::{traits::TestType#6}::test::TestType1)}] - Source: 'tests/src/traits.rs', lines 138:8-138:36 *) + Source: 'tests/src/traits.rs', lines 139:8-139:36 *) let testType_test_TestTraittraitsTestTypetestTestType1 : testType_test_TestTrait_t testType_test_TestType1_t = { test = testType_test_TestTraittraitsTestTypetestTestType1_test; } (** [traits::{traits::TestType#6}::test]: - Source: 'tests/src/traits.rs', lines 126:4-126:36 *) + Source: 'tests/src/traits.rs', lines 127:4-127:36 *) let testType_test (t : Type0) (toU64Inst : toU64_t t) (self : testType_t t) (x : t) : result bool @@ -209,11 +209,11 @@ let testType_test else Ok false (** [traits::BoolWrapper] - Source: 'tests/src/traits.rs', lines 150:0-150:22 *) + Source: 'tests/src/traits.rs', lines 151:0-151:22 *) type boolWrapper_t = bool (** [traits::{(traits::ToType for traits::BoolWrapper)#7}::to_type]: - Source: 'tests/src/traits.rs', lines 156:4-156:25 *) + Source: 'tests/src/traits.rs', lines 157:4-157:25 *) let toTypetraitsBoolWrapperT_to_type (t : Type0) (toTypeBoolTInst : toType_t bool t) (self : boolWrapper_t) : result t @@ -221,14 +221,14 @@ let toTypetraitsBoolWrapperT_to_type toTypeBoolTInst.to_type self (** Trait implementation: [traits::{(traits::ToType for traits::BoolWrapper)#7}] - Source: 'tests/src/traits.rs', lines 152:0-152:33 *) + Source: 'tests/src/traits.rs', lines 153:0-153:33 *) let toTypetraitsBoolWrapperT (t : Type0) (toTypeBoolTInst : toType_t bool t) : toType_t boolWrapper_t t = { to_type = toTypetraitsBoolWrapperT_to_type t toTypeBoolTInst; } (** [traits::WithConstTy::LEN2] - Source: 'tests/src/traits.rs', lines 164:4-164:21 *) + Source: 'tests/src/traits.rs', lines 165:4-165:21 *) let with_const_ty_len2_default_body (self : Type0) (len : usize) : result usize = Ok 32 @@ -236,7 +236,7 @@ let with_const_ty_len2_default (self : Type0) (len : usize) : usize = eval_global (with_const_ty_len2_default_body self len) (** Trait declaration: [traits::WithConstTy] - Source: 'tests/src/traits.rs', lines 161:0-161:39 *) + Source: 'tests/src/traits.rs', lines 162:0-162:39 *) noeq type withConstTy_t (self : Type0) (len : usize) = { cLEN1 : usize; cLEN2 : usize; @@ -247,18 +247,18 @@ noeq type withConstTy_t (self : Type0) (len : usize) = { } (** [traits::{(traits::WithConstTy<32: usize> for bool)#8}::LEN1] - Source: 'tests/src/traits.rs', lines 175:4-175:21 *) + Source: 'tests/src/traits.rs', lines 176:4-176:21 *) let with_const_ty_bool32_len1_body : result usize = Ok 12 let with_const_ty_bool32_len1 : usize = eval_global with_const_ty_bool32_len1_body (** [traits::{(traits::WithConstTy<32: usize> for bool)#8}::f]: - Source: 'tests/src/traits.rs', lines 180:4-180:39 *) + Source: 'tests/src/traits.rs', lines 181:4-181:39 *) let withConstTyBool32_f (i : u64) (a : array u8 32) : result u64 = Ok i (** Trait implementation: [traits::{(traits::WithConstTy<32: usize> for bool)#8}] - Source: 'tests/src/traits.rs', lines 174:0-174:29 *) + Source: 'tests/src/traits.rs', lines 175:0-175:29 *) let withConstTyBool32 : withConstTy_t bool 32 = { cLEN1 = with_const_ty_bool32_len1; cLEN2 = with_const_ty_len2_default bool 32; @@ -269,7 +269,7 @@ let withConstTyBool32 : withConstTy_t bool 32 = { } (** [traits::use_with_const_ty1]: - Source: 'tests/src/traits.rs', lines 183:0-183:75 *) + Source: 'tests/src/traits.rs', lines 184:0-184:75 *) let use_with_const_ty1 (h : Type0) (len : usize) (withConstTyInst : withConstTy_t h len) : result usize @@ -277,7 +277,7 @@ let use_with_const_ty1 Ok withConstTyInst.cLEN1 (** [traits::use_with_const_ty2]: - Source: 'tests/src/traits.rs', lines 187:0-187:73 *) + Source: 'tests/src/traits.rs', lines 188:0-188:73 *) let use_with_const_ty2 (h : Type0) (len : usize) (withConstTyInst : withConstTy_t h len) (w : withConstTyInst.tW) : @@ -286,7 +286,7 @@ let use_with_const_ty2 Ok () (** [traits::use_with_const_ty3]: - Source: 'tests/src/traits.rs', lines 189:0-189:80 *) + Source: 'tests/src/traits.rs', lines 190:0-190:80 *) let use_with_const_ty3 (h : Type0) (len : usize) (withConstTyInst : withConstTy_t h len) (x : withConstTyInst.tW) : @@ -295,12 +295,12 @@ let use_with_const_ty3 withConstTyInst.tW_clause_0.to_u64 x (** [traits::test_where1]: - Source: 'tests/src/traits.rs', lines 193:0-193:40 *) + Source: 'tests/src/traits.rs', lines 194:0-194:40 *) let test_where1 (t : Type0) (_x : t) : result unit = Ok () (** [traits::test_where2]: - Source: 'tests/src/traits.rs', lines 194:0-194:57 *) + Source: 'tests/src/traits.rs', lines 195:0-195:57 *) let test_where2 (t : Type0) (withConstTyT32Inst : withConstTy_t t 32) (_x : u32) : result unit @@ -308,7 +308,7 @@ let test_where2 Ok () (** Trait declaration: [traits::ParentTrait0] - Source: 'tests/src/traits.rs', lines 200:0-200:22 *) + Source: 'tests/src/traits.rs', lines 201:0-201:22 *) noeq type parentTrait0_t (self : Type0) = { tW : Type0; get_name : self -> result string; @@ -316,24 +316,24 @@ noeq type parentTrait0_t (self : Type0) = { } (** Trait declaration: [traits::ParentTrait1] - Source: 'tests/src/traits.rs', lines 205:0-205:22 *) + Source: 'tests/src/traits.rs', lines 206:0-206:22 *) type parentTrait1_t (self : Type0) = unit (** Trait declaration: [traits::ChildTrait] - Source: 'tests/src/traits.rs', lines 206:0-206:49 *) + Source: 'tests/src/traits.rs', lines 207:0-207:49 *) noeq type childTrait_t (self : Type0) = { parentTrait0Inst : parentTrait0_t self; parentTrait1Inst : parentTrait1_t self; } (** [traits::test_child_trait1]: - Source: 'tests/src/traits.rs', lines 209:0-209:56 *) + Source: 'tests/src/traits.rs', lines 210:0-210:56 *) let test_child_trait1 (t : Type0) (childTraitInst : childTrait_t t) (x : t) : result string = childTraitInst.parentTrait0Inst.get_name x (** [traits::test_child_trait2]: - Source: 'tests/src/traits.rs', lines 213:0-213:54 *) + Source: 'tests/src/traits.rs', lines 214:0-214:54 *) let test_child_trait2 (t : Type0) (childTraitInst : childTrait_t t) (x : t) : result childTraitInst.parentTrait0Inst.tW @@ -341,7 +341,7 @@ let test_child_trait2 childTraitInst.parentTrait0Inst.get_w x (** [traits::order1]: - Source: 'tests/src/traits.rs', lines 219:0-219:59 *) + Source: 'tests/src/traits.rs', lines 220:0-220:59 *) let order1 (t u : Type0) (parentTrait0Inst : parentTrait0_t t) (parentTrait0Inst1 : parentTrait0_t u) : @@ -350,27 +350,27 @@ let order1 Ok () (** Trait declaration: [traits::ChildTrait1] - Source: 'tests/src/traits.rs', lines 222:0-222:35 *) + Source: 'tests/src/traits.rs', lines 223:0-223:35 *) noeq type childTrait1_t (self : Type0) = { parentTrait1Inst : parentTrait1_t self; } (** Trait implementation: [traits::{(traits::ParentTrait1 for usize)#9}] - Source: 'tests/src/traits.rs', lines 224:0-224:27 *) + Source: 'tests/src/traits.rs', lines 225:0-225:27 *) let parentTrait1Usize : parentTrait1_t usize = () (** Trait implementation: [traits::{(traits::ChildTrait1 for usize)#10}] - Source: 'tests/src/traits.rs', lines 225:0-225:26 *) + Source: 'tests/src/traits.rs', lines 226:0-226:26 *) let childTrait1Usize : childTrait1_t usize = { parentTrait1Inst = parentTrait1Usize; } (** Trait declaration: [traits::Iterator] - Source: 'tests/src/traits.rs', lines 229:0-229:18 *) + Source: 'tests/src/traits.rs', lines 230:0-230:18 *) noeq type iterator_t (self : Type0) = { tItem : Type0; } (** Trait declaration: [traits::IntoIterator] - Source: 'tests/src/traits.rs', lines 233:0-233:22 *) + Source: 'tests/src/traits.rs', lines 234:0-234:22 *) noeq type intoIterator_t (self : Type0) = { tItem : Type0; tIntoIter : Type0; @@ -379,107 +379,107 @@ noeq type intoIterator_t (self : Type0) = { } (** Trait declaration: [traits::FromResidual] - Source: 'tests/src/traits.rs', lines 250:0-250:21 *) + Source: 'tests/src/traits.rs', lines 251:0-251:21 *) type fromResidual_t (self t : Type0) = unit (** Trait declaration: [traits::Try] - Source: 'tests/src/traits.rs', lines 246:0-246:48 *) + Source: 'tests/src/traits.rs', lines 247:0-247:48 *) noeq type try_t (self : Type0) = { tResidual : Type0; fromResidualSelftraitsTryResidualInst : fromResidual_t self tResidual; } (** Trait declaration: [traits::WithTarget] - Source: 'tests/src/traits.rs', lines 252:0-252:20 *) + Source: 'tests/src/traits.rs', lines 253:0-253:20 *) noeq type withTarget_t (self : Type0) = { tTarget : Type0; } (** Trait declaration: [traits::ParentTrait2] - Source: 'tests/src/traits.rs', lines 256:0-256:22 *) + Source: 'tests/src/traits.rs', lines 257:0-257:22 *) noeq type parentTrait2_t (self : Type0) = { tU : Type0; tU_clause_0 : withTarget_t tU; } (** Trait declaration: [traits::ChildTrait2] - Source: 'tests/src/traits.rs', lines 260:0-260:35 *) + Source: 'tests/src/traits.rs', lines 261:0-261:35 *) noeq type childTrait2_t (self : Type0) = { parentTrait2Inst : parentTrait2_t self; convert : parentTrait2Inst.tU -> result parentTrait2Inst.tU_clause_0.tTarget; } (** Trait implementation: [traits::{(traits::WithTarget for u32)#11}] - Source: 'tests/src/traits.rs', lines 264:0-264:23 *) + Source: 'tests/src/traits.rs', lines 265:0-265:23 *) let withTargetU32 : withTarget_t u32 = { tTarget = u32; } (** Trait implementation: [traits::{(traits::ParentTrait2 for u32)#12}] - Source: 'tests/src/traits.rs', lines 268:0-268:25 *) + Source: 'tests/src/traits.rs', lines 269:0-269:25 *) let parentTrait2U32 : parentTrait2_t u32 = { tU = u32; tU_clause_0 = withTargetU32; } (** [traits::{(traits::ChildTrait2 for u32)#13}::convert]: - Source: 'tests/src/traits.rs', lines 273:4-273:29 *) + Source: 'tests/src/traits.rs', lines 274:4-274:29 *) let childTrait2U32_convert (x : u32) : result u32 = Ok x (** Trait implementation: [traits::{(traits::ChildTrait2 for u32)#13}] - Source: 'tests/src/traits.rs', lines 272:0-272:24 *) + Source: 'tests/src/traits.rs', lines 273:0-273:24 *) let childTrait2U32 : childTrait2_t u32 = { parentTrait2Inst = parentTrait2U32; convert = childTrait2U32_convert; } (** Trait declaration: [traits::CFnOnce] - Source: 'tests/src/traits.rs', lines 286:0-286:23 *) + Source: 'tests/src/traits.rs', lines 287:0-287:23 *) noeq type cFnOnce_t (self args : Type0) = { tOutput : Type0; call_once : self -> args -> result tOutput; } (** Trait declaration: [traits::CFnMut] - Source: 'tests/src/traits.rs', lines 292:0-292:37 *) + Source: 'tests/src/traits.rs', lines 293:0-293:37 *) noeq type cFnMut_t (self args : Type0) = { cFnOnceInst : cFnOnce_t self args; call_mut : self -> args -> result (cFnOnceInst.tOutput & self); } (** Trait declaration: [traits::CFn] - Source: 'tests/src/traits.rs', lines 296:0-296:33 *) + Source: 'tests/src/traits.rs', lines 297:0-297:33 *) noeq type cFn_t (self args : Type0) = { cFnMutInst : cFnMut_t self args; call : self -> args -> result cFnMutInst.cFnOnceInst.tOutput; } (** Trait declaration: [traits::GetTrait] - Source: 'tests/src/traits.rs', lines 300:0-300:18 *) + Source: 'tests/src/traits.rs', lines 301:0-301:18 *) noeq type getTrait_t (self : Type0) = { tW : Type0; get_w : self -> result tW; } (** [traits::test_get_trait]: - Source: 'tests/src/traits.rs', lines 305:0-305:49 *) + Source: 'tests/src/traits.rs', lines 306:0-306:49 *) let test_get_trait (t : Type0) (getTraitInst : getTrait_t t) (x : t) : result getTraitInst.tW = getTraitInst.get_w x (** Trait declaration: [traits::Trait] - Source: 'tests/src/traits.rs', lines 310:0-310:15 *) + Source: 'tests/src/traits.rs', lines 311:0-311:15 *) noeq type trait_t (self : Type0) = { cLEN : usize; } (** [traits::{(traits::Trait for @Array)#14}::LEN] - Source: 'tests/src/traits.rs', lines 315:4-315:20 *) + Source: 'tests/src/traits.rs', lines 316:4-316:20 *) let trait_array_len_body (t : Type0) (n : usize) : result usize = Ok n let trait_array_len (t : Type0) (n : usize) : usize = eval_global (trait_array_len_body t n) (** Trait implementation: [traits::{(traits::Trait for @Array)#14}] - Source: 'tests/src/traits.rs', lines 314:0-314:40 *) + Source: 'tests/src/traits.rs', lines 315:0-315:40 *) let traitArray (t : Type0) (n : usize) : trait_t (array t n) = { cLEN = trait_array_len t n; } (** [traits::{(traits::Trait for traits::Wrapper)#15}::LEN] - Source: 'tests/src/traits.rs', lines 319:4-319:20 *) + Source: 'tests/src/traits.rs', lines 320:4-320:20 *) let traittraits_wrapper_len_body (t : Type0) (traitInst : trait_t t) : result usize = Ok 0 @@ -487,19 +487,19 @@ let traittraits_wrapper_len (t : Type0) (traitInst : trait_t t) : usize = eval_global (traittraits_wrapper_len_body t traitInst) (** Trait implementation: [traits::{(traits::Trait for traits::Wrapper)#15}] - Source: 'tests/src/traits.rs', lines 318:0-318:35 *) + Source: 'tests/src/traits.rs', lines 319:0-319:35 *) let traittraitsWrapper (t : Type0) (traitInst : trait_t t) : trait_t (wrapper_t t) = { cLEN = traittraits_wrapper_len t traitInst; } (** [traits::use_wrapper_len]: - Source: 'tests/src/traits.rs', lines 322:0-322:43 *) + Source: 'tests/src/traits.rs', lines 323:0-323:43 *) let use_wrapper_len (t : Type0) (traitInst : trait_t t) : result usize = Ok (traittraitsWrapper t traitInst).cLEN (** [traits::Foo] - Source: 'tests/src/traits.rs', lines 326:0-326:20 *) + Source: 'tests/src/traits.rs', lines 327:0-327:20 *) type foo_t (t u : Type0) = { x : t; y : u; } (** [core::result::Result] @@ -510,7 +510,7 @@ type core_result_Result_t (t e : Type0) = | Core_result_Result_Err : e -> core_result_Result_t t e (** [traits::{traits::Foo#16}::FOO] - Source: 'tests/src/traits.rs', lines 332:4-332:33 *) + Source: 'tests/src/traits.rs', lines 333:4-333:33 *) let foo_foo_body (t u : Type0) (traitInst : trait_t t) : result (core_result_Result_t t i32) = Ok (Core_result_Result_Err 0) @@ -519,13 +519,13 @@ let foo_foo (t u : Type0) (traitInst : trait_t t) eval_global (foo_foo_body t u traitInst) (** [traits::use_foo1]: - Source: 'tests/src/traits.rs', lines 335:0-335:48 *) + Source: 'tests/src/traits.rs', lines 336:0-336:48 *) let use_foo1 (t u : Type0) (traitInst : trait_t t) : result (core_result_Result_t t i32) = Ok (foo_foo t u traitInst) (** [traits::use_foo2]: - Source: 'tests/src/traits.rs', lines 339:0-339:48 *) + Source: 'tests/src/traits.rs', lines 340:0-340:48 *) let use_foo2 (t u : Type0) (traitInst : trait_t u) : result (core_result_Result_t u i32) = Ok (foo_foo u t traitInst) diff --git a/tests/lean/Arrays.lean b/tests/lean/Arrays.lean index 5ffcce51..c2c3ac90 100644 --- a/tests/lean/Arrays.lean +++ b/tests/lean/Arrays.lean @@ -6,24 +6,24 @@ open Primitives namespace arrays /- [arrays::AB] - Source: 'tests/src/arrays.rs', lines 3:0-3:11 -/ + Source: 'tests/src/arrays.rs', lines 6:0-6:11 -/ inductive AB := | A : AB | B : AB /- [arrays::incr]: - Source: 'tests/src/arrays.rs', lines 8:0-8:24 -/ + Source: 'tests/src/arrays.rs', lines 11:0-11:24 -/ def incr (x : U32) : Result U32 := x + 1#u32 /- [arrays::array_to_shared_slice_]: - Source: 'tests/src/arrays.rs', lines 16:0-16:53 -/ + Source: 'tests/src/arrays.rs', lines 19:0-19:53 -/ def array_to_shared_slice_ (T : Type) (s : Array T 32#usize) : Result (Slice T) := Array.to_slice T 32#usize s /- [arrays::array_to_mut_slice_]: - Source: 'tests/src/arrays.rs', lines 21:0-21:58 -/ + Source: 'tests/src/arrays.rs', lines 24:0-24:58 -/ def array_to_mut_slice_ (T : Type) (s : Array T 32#usize) : Result ((Slice T) × (Slice T → Result (Array T 32#usize))) @@ -31,42 +31,42 @@ def array_to_mut_slice_ Array.to_slice_mut T 32#usize s /- [arrays::array_len]: - Source: 'tests/src/arrays.rs', lines 25:0-25:40 -/ + Source: 'tests/src/arrays.rs', lines 28:0-28:40 -/ def array_len (T : Type) (s : Array T 32#usize) : Result Usize := do let s1 ← Array.to_slice T 32#usize s Result.ok (Slice.len T s1) /- [arrays::shared_array_len]: - Source: 'tests/src/arrays.rs', lines 29:0-29:48 -/ + Source: 'tests/src/arrays.rs', lines 32:0-32:48 -/ def shared_array_len (T : Type) (s : Array T 32#usize) : Result Usize := do let s1 ← Array.to_slice T 32#usize s Result.ok (Slice.len T s1) /- [arrays::shared_slice_len]: - Source: 'tests/src/arrays.rs', lines 33:0-33:44 -/ + Source: 'tests/src/arrays.rs', lines 36:0-36:44 -/ def shared_slice_len (T : Type) (s : Slice T) : Result Usize := Result.ok (Slice.len T s) /- [arrays::index_array_shared]: - Source: 'tests/src/arrays.rs', lines 37:0-37:57 -/ + Source: 'tests/src/arrays.rs', lines 40:0-40:57 -/ def index_array_shared (T : Type) (s : Array T 32#usize) (i : Usize) : Result T := Array.index_usize T 32#usize s i /- [arrays::index_array_u32]: - Source: 'tests/src/arrays.rs', lines 44:0-44:53 -/ + Source: 'tests/src/arrays.rs', lines 47:0-47:53 -/ def index_array_u32 (s : Array U32 32#usize) (i : Usize) : Result U32 := Array.index_usize U32 32#usize s i /- [arrays::index_array_copy]: - Source: 'tests/src/arrays.rs', lines 48:0-48:45 -/ + Source: 'tests/src/arrays.rs', lines 51:0-51:45 -/ def index_array_copy (x : Array U32 32#usize) : Result U32 := Array.index_usize U32 32#usize x 0#usize /- [arrays::index_mut_array]: - Source: 'tests/src/arrays.rs', lines 52:0-52:62 -/ + Source: 'tests/src/arrays.rs', lines 55:0-55:62 -/ def index_mut_array (T : Type) (s : Array T 32#usize) (i : Usize) : Result (T × (T → Result (Array T 32#usize))) @@ -74,12 +74,12 @@ def index_mut_array Array.index_mut_usize T 32#usize s i /- [arrays::index_slice]: - Source: 'tests/src/arrays.rs', lines 56:0-56:46 -/ + Source: 'tests/src/arrays.rs', lines 59:0-59:46 -/ def index_slice (T : Type) (s : Slice T) (i : Usize) : Result T := Slice.index_usize T s i /- [arrays::index_mut_slice]: - Source: 'tests/src/arrays.rs', lines 60:0-60:58 -/ + Source: 'tests/src/arrays.rs', lines 63:0-63:58 -/ def index_mut_slice (T : Type) (s : Slice T) (i : Usize) : Result (T × (T → Result (Slice T))) @@ -87,7 +87,7 @@ def index_mut_slice Slice.index_mut_usize T s i /- [arrays::slice_subslice_shared_]: - Source: 'tests/src/arrays.rs', lines 64:0-64:70 -/ + Source: 'tests/src/arrays.rs', lines 67:0-67:70 -/ def slice_subslice_shared_ (x : Slice U32) (y : Usize) (z : Usize) : Result (Slice U32) := core.slice.index.Slice.index U32 (core.ops.range.Range Usize) @@ -95,7 +95,7 @@ def slice_subslice_shared_ { start := y, end_ := z } /- [arrays::slice_subslice_mut_]: - Source: 'tests/src/arrays.rs', lines 68:0-68:75 -/ + Source: 'tests/src/arrays.rs', lines 71:0-71:75 -/ def slice_subslice_mut_ (x : Slice U32) (y : Usize) (z : Usize) : Result ((Slice U32) × (Slice U32 → Result (Slice U32))) @@ -108,12 +108,12 @@ def slice_subslice_mut_ Result.ok (s, index_mut_back) /- [arrays::array_to_slice_shared_]: - Source: 'tests/src/arrays.rs', lines 72:0-72:54 -/ + Source: 'tests/src/arrays.rs', lines 75:0-75:54 -/ def array_to_slice_shared_ (x : Array U32 32#usize) : Result (Slice U32) := Array.to_slice U32 32#usize x /- [arrays::array_to_slice_mut_]: - Source: 'tests/src/arrays.rs', lines 76:0-76:59 -/ + Source: 'tests/src/arrays.rs', lines 79:0-79:59 -/ def array_to_slice_mut_ (x : Array U32 32#usize) : Result ((Slice U32) × (Slice U32 → Result (Array U32 32#usize))) @@ -121,7 +121,7 @@ def array_to_slice_mut_ Array.to_slice_mut U32 32#usize x /- [arrays::array_subslice_shared_]: - Source: 'tests/src/arrays.rs', lines 80:0-80:74 -/ + Source: 'tests/src/arrays.rs', lines 83:0-83:74 -/ def array_subslice_shared_ (x : Array U32 32#usize) (y : Usize) (z : Usize) : Result (Slice U32) := core.array.Array.index U32 (core.ops.range.Range Usize) 32#usize @@ -130,7 +130,7 @@ def array_subslice_shared_ { start := y, end_ := z } /- [arrays::array_subslice_mut_]: - Source: 'tests/src/arrays.rs', lines 84:0-84:79 -/ + Source: 'tests/src/arrays.rs', lines 87:0-87:79 -/ def array_subslice_mut_ (x : Array U32 32#usize) (y : Usize) (z : Usize) : Result ((Slice U32) × (Slice U32 → Result (Array U32 32#usize))) @@ -144,17 +144,17 @@ def array_subslice_mut_ Result.ok (s, index_mut_back) /- [arrays::index_slice_0]: - Source: 'tests/src/arrays.rs', lines 88:0-88:38 -/ + Source: 'tests/src/arrays.rs', lines 91:0-91:38 -/ def index_slice_0 (T : Type) (s : Slice T) : Result T := Slice.index_usize T s 0#usize /- [arrays::index_array_0]: - Source: 'tests/src/arrays.rs', lines 92:0-92:42 -/ + Source: 'tests/src/arrays.rs', lines 95:0-95:42 -/ def index_array_0 (T : Type) (s : Array T 32#usize) : Result T := Array.index_usize T 32#usize s 0#usize /- [arrays::index_index_array]: - Source: 'tests/src/arrays.rs', lines 103:0-103:71 -/ + Source: 'tests/src/arrays.rs', lines 106:0-106:71 -/ def index_index_array (s : Array (Array U32 32#usize) 32#usize) (i : Usize) (j : Usize) : Result U32 @@ -164,7 +164,7 @@ def index_index_array Array.index_usize U32 32#usize a j /- [arrays::update_update_array]: - Source: 'tests/src/arrays.rs', lines 114:0-114:70 -/ + Source: 'tests/src/arrays.rs', lines 117:0-117:70 -/ def update_update_array (s : Array (Array U32 32#usize) 32#usize) (i : Usize) (j : Usize) : Result Unit @@ -178,37 +178,37 @@ def update_update_array Result.ok () /- [arrays::array_local_deep_copy]: - Source: 'tests/src/arrays.rs', lines 118:0-118:43 -/ + Source: 'tests/src/arrays.rs', lines 121:0-121:43 -/ def array_local_deep_copy (x : Array U32 32#usize) : Result Unit := Result.ok () /- [arrays::take_array]: - Source: 'tests/src/arrays.rs', lines 122:0-122:30 -/ + Source: 'tests/src/arrays.rs', lines 125:0-125:30 -/ def take_array (a : Array U32 2#usize) : Result Unit := Result.ok () /- [arrays::take_array_borrow]: - Source: 'tests/src/arrays.rs', lines 123:0-123:38 -/ + Source: 'tests/src/arrays.rs', lines 126:0-126:38 -/ def take_array_borrow (a : Array U32 2#usize) : Result Unit := Result.ok () /- [arrays::take_slice]: - Source: 'tests/src/arrays.rs', lines 124:0-124:28 -/ + Source: 'tests/src/arrays.rs', lines 127:0-127:28 -/ def take_slice (s : Slice U32) : Result Unit := Result.ok () /- [arrays::take_mut_slice]: - Source: 'tests/src/arrays.rs', lines 125:0-125:36 -/ + Source: 'tests/src/arrays.rs', lines 128:0-128:36 -/ def take_mut_slice (s : Slice U32) : Result (Slice U32) := Result.ok s /- [arrays::const_array]: - Source: 'tests/src/arrays.rs', lines 127:0-127:32 -/ + Source: 'tests/src/arrays.rs', lines 130:0-130:32 -/ def const_array : Result (Array U32 2#usize) := Result.ok (Array.make U32 2#usize [ 0#u32, 0#u32 ]) /- [arrays::const_slice]: - Source: 'tests/src/arrays.rs', lines 131:0-131:20 -/ + Source: 'tests/src/arrays.rs', lines 134:0-134:20 -/ def const_slice : Result Unit := do let _ ← @@ -216,7 +216,7 @@ def const_slice : Result Unit := Result.ok () /- [arrays::take_all]: - Source: 'tests/src/arrays.rs', lines 141:0-141:17 -/ + Source: 'tests/src/arrays.rs', lines 144:0-144:17 -/ def take_all : Result Unit := do let _ ← take_array (Array.make U32 2#usize [ 0#u32, 0#u32 ]) @@ -232,29 +232,29 @@ def take_all : Result Unit := Result.ok () /- [arrays::index_array]: - Source: 'tests/src/arrays.rs', lines 155:0-155:38 -/ + Source: 'tests/src/arrays.rs', lines 158:0-158:38 -/ def index_array (x : Array U32 2#usize) : Result U32 := Array.index_usize U32 2#usize x 0#usize /- [arrays::index_array_borrow]: - Source: 'tests/src/arrays.rs', lines 158:0-158:46 -/ + Source: 'tests/src/arrays.rs', lines 161:0-161:46 -/ def index_array_borrow (x : Array U32 2#usize) : Result U32 := Array.index_usize U32 2#usize x 0#usize /- [arrays::index_slice_u32_0]: - Source: 'tests/src/arrays.rs', lines 162:0-162:42 -/ + Source: 'tests/src/arrays.rs', lines 165:0-165:42 -/ def index_slice_u32_0 (x : Slice U32) : Result U32 := Slice.index_usize U32 x 0#usize /- [arrays::index_mut_slice_u32_0]: - Source: 'tests/src/arrays.rs', lines 166:0-166:50 -/ + Source: 'tests/src/arrays.rs', lines 169:0-169:50 -/ def index_mut_slice_u32_0 (x : Slice U32) : Result (U32 × (Slice U32)) := do let i ← Slice.index_usize U32 x 0#usize Result.ok (i, x) /- [arrays::index_all]: - Source: 'tests/src/arrays.rs', lines 170:0-170:25 -/ + Source: 'tests/src/arrays.rs', lines 173:0-173:25 -/ def index_all : Result U32 := do let i ← index_array (Array.make U32 2#usize [ 0#u32, 0#u32 ]) @@ -274,7 +274,7 @@ def index_all : Result U32 := Result.ok i8 /- [arrays::update_array]: - Source: 'tests/src/arrays.rs', lines 184:0-184:36 -/ + Source: 'tests/src/arrays.rs', lines 187:0-187:36 -/ def update_array (x : Array U32 2#usize) : Result Unit := do let (_, index_mut_back) ← Array.index_mut_usize U32 2#usize x 0#usize @@ -282,7 +282,7 @@ def update_array (x : Array U32 2#usize) : Result Unit := Result.ok () /- [arrays::update_array_mut_borrow]: - Source: 'tests/src/arrays.rs', lines 187:0-187:48 -/ + Source: 'tests/src/arrays.rs', lines 190:0-190:48 -/ def update_array_mut_borrow (x : Array U32 2#usize) : Result (Array U32 2#usize) := do @@ -290,14 +290,14 @@ def update_array_mut_borrow index_mut_back 1#u32 /- [arrays::update_mut_slice]: - Source: 'tests/src/arrays.rs', lines 190:0-190:38 -/ + Source: 'tests/src/arrays.rs', lines 193:0-193:38 -/ def update_mut_slice (x : Slice U32) : Result (Slice U32) := do let (_, index_mut_back) ← Slice.index_mut_usize U32 x 0#usize index_mut_back 1#u32 /- [arrays::update_all]: - Source: 'tests/src/arrays.rs', lines 194:0-194:19 -/ + Source: 'tests/src/arrays.rs', lines 197:0-197:19 -/ def update_all : Result Unit := do let _ ← update_array (Array.make U32 2#usize [ 0#u32, 0#u32 ]) @@ -309,7 +309,7 @@ def update_all : Result Unit := Result.ok () /- [arrays::range_all]: - Source: 'tests/src/arrays.rs', lines 205:0-205:18 -/ + Source: 'tests/src/arrays.rs', lines 208:0-208:18 -/ def range_all : Result Unit := do let (s, index_mut_back) ← @@ -323,12 +323,12 @@ def range_all : Result Unit := Result.ok () /- [arrays::deref_array_borrow]: - Source: 'tests/src/arrays.rs', lines 214:0-214:46 -/ + Source: 'tests/src/arrays.rs', lines 217:0-217:46 -/ def deref_array_borrow (x : Array U32 2#usize) : Result U32 := Array.index_usize U32 2#usize x 0#usize /- [arrays::deref_array_mut_borrow]: - Source: 'tests/src/arrays.rs', lines 219:0-219:54 -/ + Source: 'tests/src/arrays.rs', lines 222:0-222:54 -/ def deref_array_mut_borrow (x : Array U32 2#usize) : Result (U32 × (Array U32 2#usize)) := do @@ -336,17 +336,17 @@ def deref_array_mut_borrow Result.ok (i, x) /- [arrays::take_array_t]: - Source: 'tests/src/arrays.rs', lines 227:0-227:31 -/ + Source: 'tests/src/arrays.rs', lines 230:0-230:31 -/ def take_array_t (a : Array AB 2#usize) : Result Unit := Result.ok () /- [arrays::non_copyable_array]: - Source: 'tests/src/arrays.rs', lines 229:0-229:27 -/ + Source: 'tests/src/arrays.rs', lines 232:0-232:27 -/ def non_copyable_array : Result Unit := take_array_t (Array.make AB 2#usize [ AB.A, AB.B ]) /- [arrays::sum]: loop 0: - Source: 'tests/src/arrays.rs', lines 242:0-250:1 -/ + Source: 'tests/src/arrays.rs', lines 245:0-253:1 -/ divergent def sum_loop (s : Slice U32) (sum1 : U32) (i : Usize) : Result U32 := let i1 := Slice.len U32 s if i < i1 @@ -359,12 +359,12 @@ divergent def sum_loop (s : Slice U32) (sum1 : U32) (i : Usize) : Result U32 := else Result.ok sum1 /- [arrays::sum]: - Source: 'tests/src/arrays.rs', lines 242:0-242:28 -/ + Source: 'tests/src/arrays.rs', lines 245:0-245:28 -/ def sum (s : Slice U32) : Result U32 := sum_loop s 0#u32 0#usize /- [arrays::sum2]: loop 0: - Source: 'tests/src/arrays.rs', lines 252:0-261:1 -/ + Source: 'tests/src/arrays.rs', lines 255:0-264:1 -/ divergent def sum2_loop (s : Slice U32) (s2 : Slice U32) (sum1 : U32) (i : Usize) : Result U32 := let i1 := Slice.len U32 s @@ -380,7 +380,7 @@ divergent def sum2_loop else Result.ok sum1 /- [arrays::sum2]: - Source: 'tests/src/arrays.rs', lines 252:0-252:41 -/ + Source: 'tests/src/arrays.rs', lines 255:0-255:41 -/ def sum2 (s : Slice U32) (s2 : Slice U32) : Result U32 := let i := Slice.len U32 s let i1 := Slice.len U32 s2 @@ -389,7 +389,7 @@ def sum2 (s : Slice U32) (s2 : Slice U32) : Result U32 := else sum2_loop s s2 0#u32 0#usize /- [arrays::f0]: - Source: 'tests/src/arrays.rs', lines 263:0-263:11 -/ + Source: 'tests/src/arrays.rs', lines 266:0-266:11 -/ def f0 : Result Unit := do let (s, to_slice_mut_back) ← @@ -400,7 +400,7 @@ def f0 : Result Unit := Result.ok () /- [arrays::f1]: - Source: 'tests/src/arrays.rs', lines 268:0-268:11 -/ + Source: 'tests/src/arrays.rs', lines 271:0-271:11 -/ def f1 : Result Unit := do let (_, index_mut_back) ← @@ -410,12 +410,12 @@ def f1 : Result Unit := Result.ok () /- [arrays::f2]: - Source: 'tests/src/arrays.rs', lines 273:0-273:17 -/ + Source: 'tests/src/arrays.rs', lines 276:0-276:17 -/ def f2 (i : U32) : Result Unit := Result.ok () /- [arrays::f4]: - Source: 'tests/src/arrays.rs', lines 282:0-282:54 -/ + Source: 'tests/src/arrays.rs', lines 285:0-285:54 -/ def f4 (x : Array U32 32#usize) (y : Usize) (z : Usize) : Result (Slice U32) := core.array.Array.index U32 (core.ops.range.Range Usize) 32#usize (core.ops.index.IndexSliceTIInst U32 (core.ops.range.Range Usize) @@ -423,7 +423,7 @@ def f4 (x : Array U32 32#usize) (y : Usize) (z : Usize) : Result (Slice U32) := { start := y, end_ := z } /- [arrays::f3]: - Source: 'tests/src/arrays.rs', lines 275:0-275:18 -/ + Source: 'tests/src/arrays.rs', lines 278:0-278:18 -/ def f3 : Result U32 := do let i ← @@ -437,17 +437,17 @@ def f3 : Result U32 := sum2 s s1 /- [arrays::SZ] - Source: 'tests/src/arrays.rs', lines 286:0-286:19 -/ + Source: 'tests/src/arrays.rs', lines 289:0-289:19 -/ def SZ_body : Result Usize := Result.ok 32#usize def SZ : Usize := eval_global SZ_body /- [arrays::f5]: - Source: 'tests/src/arrays.rs', lines 289:0-289:31 -/ + Source: 'tests/src/arrays.rs', lines 292:0-292:31 -/ def f5 (x : Array U32 32#usize) : Result U32 := Array.index_usize U32 32#usize x 0#usize /- [arrays::ite]: - Source: 'tests/src/arrays.rs', lines 294:0-294:12 -/ + Source: 'tests/src/arrays.rs', lines 297:0-297:12 -/ def ite : Result Unit := do let (s, to_slice_mut_back) ← @@ -461,7 +461,7 @@ def ite : Result Unit := Result.ok () /- [arrays::zero_slice]: loop 0: - Source: 'tests/src/arrays.rs', lines 303:0-310:1 -/ + Source: 'tests/src/arrays.rs', lines 306:0-313:1 -/ divergent def zero_slice_loop (a : Slice U8) (i : Usize) (len : Usize) : Result (Slice U8) := if i < len @@ -474,13 +474,13 @@ divergent def zero_slice_loop else Result.ok a /- [arrays::zero_slice]: - Source: 'tests/src/arrays.rs', lines 303:0-303:31 -/ + Source: 'tests/src/arrays.rs', lines 306:0-306:31 -/ def zero_slice (a : Slice U8) : Result (Slice U8) := let len := Slice.len U8 a zero_slice_loop a 0#usize len /- [arrays::iter_mut_slice]: loop 0: - Source: 'tests/src/arrays.rs', lines 312:0-318:1 -/ + Source: 'tests/src/arrays.rs', lines 315:0-321:1 -/ divergent def iter_mut_slice_loop (len : Usize) (i : Usize) : Result Unit := if i < len then do @@ -489,7 +489,7 @@ divergent def iter_mut_slice_loop (len : Usize) (i : Usize) : Result Unit := else Result.ok () /- [arrays::iter_mut_slice]: - Source: 'tests/src/arrays.rs', lines 312:0-312:35 -/ + Source: 'tests/src/arrays.rs', lines 315:0-315:35 -/ def iter_mut_slice (a : Slice U8) : Result (Slice U8) := do let len := Slice.len U8 a @@ -497,7 +497,7 @@ def iter_mut_slice (a : Slice U8) : Result (Slice U8) := Result.ok a /- [arrays::sum_mut_slice]: loop 0: - Source: 'tests/src/arrays.rs', lines 320:0-328:1 -/ + Source: 'tests/src/arrays.rs', lines 323:0-331:1 -/ divergent def sum_mut_slice_loop (a : Slice U32) (i : Usize) (s : U32) : Result U32 := let i1 := Slice.len U32 a @@ -511,7 +511,7 @@ divergent def sum_mut_slice_loop else Result.ok s /- [arrays::sum_mut_slice]: - Source: 'tests/src/arrays.rs', lines 320:0-320:42 -/ + Source: 'tests/src/arrays.rs', lines 323:0-323:42 -/ def sum_mut_slice (a : Slice U32) : Result (U32 × (Slice U32)) := do let i ← sum_mut_slice_loop a 0#usize 0#u32 diff --git a/tests/lean/Bitwise.lean b/tests/lean/Bitwise.lean index 4ed56a0a..3c381233 100644 --- a/tests/lean/Bitwise.lean +++ b/tests/lean/Bitwise.lean @@ -6,31 +6,31 @@ open Primitives namespace bitwise /- [bitwise::shift_u32]: - Source: 'tests/src/bitwise.rs', lines 3:0-3:31 -/ + Source: 'tests/src/bitwise.rs', lines 4:0-4:31 -/ def shift_u32 (a : U32) : Result U32 := do let t ← a >>> 16#usize t <<< 16#usize /- [bitwise::shift_i32]: - Source: 'tests/src/bitwise.rs', lines 10:0-10:31 -/ + Source: 'tests/src/bitwise.rs', lines 11:0-11:31 -/ def shift_i32 (a : I32) : Result I32 := do let t ← a >>> 16#isize t <<< 16#isize /- [bitwise::xor_u32]: - Source: 'tests/src/bitwise.rs', lines 17:0-17:37 -/ + Source: 'tests/src/bitwise.rs', lines 18:0-18:37 -/ def xor_u32 (a : U32) (b : U32) : Result U32 := Result.ok (a ^^^ b) /- [bitwise::or_u32]: - Source: 'tests/src/bitwise.rs', lines 21:0-21:36 -/ + Source: 'tests/src/bitwise.rs', lines 22:0-22:36 -/ def or_u32 (a : U32) (b : U32) : Result U32 := Result.ok (a ||| b) /- [bitwise::and_u32]: - Source: 'tests/src/bitwise.rs', lines 25:0-25:37 -/ + Source: 'tests/src/bitwise.rs', lines 26:0-26:37 -/ def and_u32 (a : U32) (b : U32) : Result U32 := Result.ok (a &&& b) diff --git a/tests/lean/Constants.lean b/tests/lean/Constants.lean index 532f265a..334f4707 100644 --- a/tests/lean/Constants.lean +++ b/tests/lean/Constants.lean @@ -6,123 +6,123 @@ open Primitives namespace constants /- [constants::X0] - Source: 'tests/src/constants.rs', lines 5:0-5:17 -/ + Source: 'tests/src/constants.rs', lines 7:0-7:17 -/ def X0_body : Result U32 := Result.ok 0#u32 def X0 : U32 := eval_global X0_body /- [constants::X1] - Source: 'tests/src/constants.rs', lines 7:0-7:17 -/ + Source: 'tests/src/constants.rs', lines 9:0-9:17 -/ def X1_body : Result U32 := Result.ok core_u32_max def X1 : U32 := eval_global X1_body /- [constants::X2] - Source: 'tests/src/constants.rs', lines 10:0-10:17 -/ + Source: 'tests/src/constants.rs', lines 12:0-12:17 -/ def X2_body : Result U32 := Result.ok 3#u32 def X2 : U32 := eval_global X2_body /- [constants::incr]: - Source: 'tests/src/constants.rs', lines 17:0-17:32 -/ + Source: 'tests/src/constants.rs', lines 19:0-19:32 -/ def incr (n : U32) : Result U32 := n + 1#u32 /- [constants::X3] - Source: 'tests/src/constants.rs', lines 15:0-15:17 -/ + Source: 'tests/src/constants.rs', lines 17:0-17:17 -/ def X3_body : Result U32 := incr 32#u32 def X3 : U32 := eval_global X3_body /- [constants::mk_pair0]: - Source: 'tests/src/constants.rs', lines 23:0-23:51 -/ + Source: 'tests/src/constants.rs', lines 25:0-25:51 -/ def mk_pair0 (x : U32) (y : U32) : Result (U32 × U32) := Result.ok (x, y) /- [constants::Pair] - Source: 'tests/src/constants.rs', lines 36:0-36:23 -/ + Source: 'tests/src/constants.rs', lines 38:0-38:23 -/ structure Pair (T1 T2 : Type) where x : T1 y : T2 /- [constants::mk_pair1]: - Source: 'tests/src/constants.rs', lines 27:0-27:55 -/ + Source: 'tests/src/constants.rs', lines 29:0-29:55 -/ def mk_pair1 (x : U32) (y : U32) : Result (Pair U32 U32) := Result.ok { x := x, y := y } /- [constants::P0] - Source: 'tests/src/constants.rs', lines 31:0-31:24 -/ + Source: 'tests/src/constants.rs', lines 33:0-33:24 -/ def P0_body : Result (U32 × U32) := mk_pair0 0#u32 1#u32 def P0 : (U32 × U32) := eval_global P0_body /- [constants::P1] - Source: 'tests/src/constants.rs', lines 32:0-32:28 -/ + Source: 'tests/src/constants.rs', lines 34:0-34:28 -/ def P1_body : Result (Pair U32 U32) := mk_pair1 0#u32 1#u32 def P1 : Pair U32 U32 := eval_global P1_body /- [constants::P2] - Source: 'tests/src/constants.rs', lines 33:0-33:24 -/ + Source: 'tests/src/constants.rs', lines 35:0-35:24 -/ def P2_body : Result (U32 × U32) := Result.ok (0#u32, 1#u32) def P2 : (U32 × U32) := eval_global P2_body /- [constants::P3] - Source: 'tests/src/constants.rs', lines 34:0-34:28 -/ + Source: 'tests/src/constants.rs', lines 36:0-36:28 -/ def P3_body : Result (Pair U32 U32) := Result.ok { x := 0#u32, y := 1#u32 } def P3 : Pair U32 U32 := eval_global P3_body /- [constants::Wrap] - Source: 'tests/src/constants.rs', lines 49:0-49:18 -/ + Source: 'tests/src/constants.rs', lines 51:0-51:18 -/ structure Wrap (T : Type) where value : T /- [constants::{constants::Wrap}::new]: - Source: 'tests/src/constants.rs', lines 54:4-54:41 -/ + Source: 'tests/src/constants.rs', lines 56:4-56:41 -/ def Wrap.new (T : Type) (value : T) : Result (Wrap T) := Result.ok { value := value } /- [constants::Y] - Source: 'tests/src/constants.rs', lines 41:0-41:22 -/ + Source: 'tests/src/constants.rs', lines 43:0-43:22 -/ def Y_body : Result (Wrap I32) := Wrap.new I32 2#i32 def Y : Wrap I32 := eval_global Y_body /- [constants::unwrap_y]: - Source: 'tests/src/constants.rs', lines 43:0-43:30 -/ + Source: 'tests/src/constants.rs', lines 45:0-45:30 -/ def unwrap_y : Result I32 := Result.ok Y.value /- [constants::YVAL] - Source: 'tests/src/constants.rs', lines 47:0-47:19 -/ + Source: 'tests/src/constants.rs', lines 49:0-49:19 -/ def YVAL_body : Result I32 := unwrap_y def YVAL : I32 := eval_global YVAL_body /- [constants::get_z1::Z1] - Source: 'tests/src/constants.rs', lines 62:4-62:17 -/ + Source: 'tests/src/constants.rs', lines 64:4-64:17 -/ def get_z1.Z1_body : Result I32 := Result.ok 3#i32 def get_z1.Z1 : I32 := eval_global get_z1.Z1_body /- [constants::get_z1]: - Source: 'tests/src/constants.rs', lines 61:0-61:28 -/ + Source: 'tests/src/constants.rs', lines 63:0-63:28 -/ def get_z1 : Result I32 := Result.ok get_z1.Z1 /- [constants::add]: - Source: 'tests/src/constants.rs', lines 66:0-66:39 -/ + Source: 'tests/src/constants.rs', lines 68:0-68:39 -/ def add (a : I32) (b : I32) : Result I32 := a + b /- [constants::Q1] - Source: 'tests/src/constants.rs', lines 74:0-74:17 -/ + Source: 'tests/src/constants.rs', lines 76:0-76:17 -/ def Q1_body : Result I32 := Result.ok 5#i32 def Q1 : I32 := eval_global Q1_body /- [constants::Q2] - Source: 'tests/src/constants.rs', lines 75:0-75:17 -/ + Source: 'tests/src/constants.rs', lines 77:0-77:17 -/ def Q2_body : Result I32 := Result.ok Q1 def Q2 : I32 := eval_global Q2_body /- [constants::Q3] - Source: 'tests/src/constants.rs', lines 76:0-76:17 -/ + Source: 'tests/src/constants.rs', lines 78:0-78:17 -/ def Q3_body : Result I32 := add Q2 3#i32 def Q3 : I32 := eval_global Q3_body /- [constants::get_z2]: - Source: 'tests/src/constants.rs', lines 70:0-70:28 -/ + Source: 'tests/src/constants.rs', lines 72:0-72:28 -/ def get_z2 : Result I32 := do let i ← get_z1 @@ -130,37 +130,37 @@ def get_z2 : Result I32 := add Q1 i1 /- [constants::S1] - Source: 'tests/src/constants.rs', lines 80:0-80:18 -/ + Source: 'tests/src/constants.rs', lines 82:0-82:18 -/ def S1_body : Result U32 := Result.ok 6#u32 def S1 : U32 := eval_global S1_body /- [constants::S2] - Source: 'tests/src/constants.rs', lines 81:0-81:18 -/ + Source: 'tests/src/constants.rs', lines 83:0-83:18 -/ def S2_body : Result U32 := incr S1 def S2 : U32 := eval_global S2_body /- [constants::S3] - Source: 'tests/src/constants.rs', lines 82:0-82:29 -/ + Source: 'tests/src/constants.rs', lines 84:0-84:29 -/ def S3_body : Result (Pair U32 U32) := Result.ok P3 def S3 : Pair U32 U32 := eval_global S3_body /- [constants::S4] - Source: 'tests/src/constants.rs', lines 83:0-83:29 -/ + Source: 'tests/src/constants.rs', lines 85:0-85:29 -/ def S4_body : Result (Pair U32 U32) := mk_pair1 7#u32 8#u32 def S4 : Pair U32 U32 := eval_global S4_body /- [constants::V] - Source: 'tests/src/constants.rs', lines 86:0-86:31 -/ + Source: 'tests/src/constants.rs', lines 88:0-88:31 -/ structure V (T : Type) (N : Usize) where x : Array T N /- [constants::{constants::V#1}::LEN] - Source: 'tests/src/constants.rs', lines 91:4-91:24 -/ + Source: 'tests/src/constants.rs', lines 93:4-93:24 -/ def V.LEN_body (T : Type) (N : Usize) : Result Usize := Result.ok N def V.LEN (T : Type) (N : Usize) : Usize := eval_global (V.LEN_body T N) /- [constants::use_v]: - Source: 'tests/src/constants.rs', lines 94:0-94:42 -/ + Source: 'tests/src/constants.rs', lines 96:0-96:42 -/ def use_v (T : Type) (N : Usize) : Result Usize := Result.ok (V.LEN T N) diff --git a/tests/lean/Demo/Demo.lean b/tests/lean/Demo/Demo.lean index a9b349b3..a7683eb0 100644 --- a/tests/lean/Demo/Demo.lean +++ b/tests/lean/Demo/Demo.lean @@ -6,7 +6,7 @@ open Primitives namespace demo /- [demo::choose]: - Source: 'tests/src/demo.rs', lines 5:0-5:70 -/ + Source: 'tests/src/demo.rs', lines 6:0-6:70 -/ def choose (T : Type) (b : Bool) (x : T) (y : T) : Result (T × (T → Result (T × T))) @@ -18,26 +18,26 @@ def choose Result.ok (y, back) /- [demo::mul2_add1]: - Source: 'tests/src/demo.rs', lines 13:0-13:31 -/ + Source: 'tests/src/demo.rs', lines 14:0-14:31 -/ def mul2_add1 (x : U32) : Result U32 := do let i ← x + x i + 1#u32 /- [demo::use_mul2_add1]: - Source: 'tests/src/demo.rs', lines 17:0-17:43 -/ + Source: 'tests/src/demo.rs', lines 18:0-18:43 -/ def use_mul2_add1 (x : U32) (y : U32) : Result U32 := do let i ← mul2_add1 x i + y /- [demo::incr]: - Source: 'tests/src/demo.rs', lines 21:0-21:31 -/ + Source: 'tests/src/demo.rs', lines 22:0-22:31 -/ def incr (x : U32) : Result U32 := x + 1#u32 /- [demo::use_incr]: - Source: 'tests/src/demo.rs', lines 25:0-25:17 -/ + Source: 'tests/src/demo.rs', lines 26:0-26:17 -/ def use_incr : Result Unit := do let x ← incr 0#u32 @@ -46,13 +46,13 @@ def use_incr : Result Unit := Result.ok () /- [demo::CList] - Source: 'tests/src/demo.rs', lines 34:0-34:17 -/ + Source: 'tests/src/demo.rs', lines 35:0-35:17 -/ inductive CList (T : Type) := | CCons : T → CList T → CList T | CNil : CList T /- [demo::list_nth]: - Source: 'tests/src/demo.rs', lines 39:0-39:56 -/ + Source: 'tests/src/demo.rs', lines 40:0-40:56 -/ divergent def list_nth (T : Type) (l : CList T) (i : U32) : Result T := match l with | CList.CCons x tl => @@ -64,7 +64,7 @@ divergent def list_nth (T : Type) (l : CList T) (i : U32) : Result T := | CList.CNil => Result.fail .panic /- [demo::list_nth_mut]: - Source: 'tests/src/demo.rs', lines 54:0-54:68 -/ + Source: 'tests/src/demo.rs', lines 55:0-55:68 -/ divergent def list_nth_mut (T : Type) (l : CList T) (i : U32) : Result (T × (T → Result (CList T))) @@ -88,7 +88,7 @@ divergent def list_nth_mut | CList.CNil => Result.fail .panic /- [demo::list_nth_mut1]: loop 0: - Source: 'tests/src/demo.rs', lines 69:0-78:1 -/ + Source: 'tests/src/demo.rs', lines 70:0-79:1 -/ divergent def list_nth_mut1_loop (T : Type) (l : CList T) (i : U32) : Result (T × (T → Result (CList T))) @@ -111,7 +111,7 @@ divergent def list_nth_mut1_loop | CList.CNil => Result.fail .panic /- [demo::list_nth_mut1]: - Source: 'tests/src/demo.rs', lines 69:0-69:77 -/ + Source: 'tests/src/demo.rs', lines 70:0-70:77 -/ def list_nth_mut1 (T : Type) (l : CList T) (i : U32) : Result (T × (T → Result (CList T))) @@ -119,7 +119,7 @@ def list_nth_mut1 list_nth_mut1_loop T l i /- [demo::i32_id]: - Source: 'tests/src/demo.rs', lines 80:0-80:28 -/ + Source: 'tests/src/demo.rs', lines 81:0-81:28 -/ divergent def i32_id (i : I32) : Result I32 := if i = 0#i32 then Result.ok 0#i32 @@ -129,7 +129,7 @@ divergent def i32_id (i : I32) : Result I32 := i2 + 1#i32 /- [demo::list_tail]: - Source: 'tests/src/demo.rs', lines 88:0-88:64 -/ + Source: 'tests/src/demo.rs', lines 89:0-89:64 -/ divergent def list_tail (T : Type) (l : CList T) : Result ((CList T) × (CList T → Result (CList T))) @@ -147,25 +147,25 @@ divergent def list_tail | CList.CNil => Result.ok (CList.CNil, Result.ok) /- Trait declaration: [demo::Counter] - Source: 'tests/src/demo.rs', lines 97:0-97:17 -/ + Source: 'tests/src/demo.rs', lines 98:0-98:17 -/ structure Counter (Self : Type) where incr : Self → Result (Usize × Self) /- [demo::{(demo::Counter for usize)}::incr]: - Source: 'tests/src/demo.rs', lines 102:4-102:31 -/ + Source: 'tests/src/demo.rs', lines 103:4-103:31 -/ def CounterUsize.incr (self : Usize) : Result (Usize × Usize) := do let self1 ← self + 1#usize Result.ok (self, self1) /- Trait implementation: [demo::{(demo::Counter for usize)}] - Source: 'tests/src/demo.rs', lines 101:0-101:22 -/ + Source: 'tests/src/demo.rs', lines 102:0-102:22 -/ def CounterUsize : Counter Usize := { incr := CounterUsize.incr } /- [demo::use_counter]: - Source: 'tests/src/demo.rs', lines 109:0-109:59 -/ + Source: 'tests/src/demo.rs', lines 110:0-110:59 -/ def use_counter (T : Type) (CounterInst : Counter T) (cnt : T) : Result (Usize × T) := CounterInst.incr cnt diff --git a/tests/lean/External/Funs.lean b/tests/lean/External/Funs.lean index 5e3c9b1d..84fe1a28 100644 --- a/tests/lean/External/Funs.lean +++ b/tests/lean/External/Funs.lean @@ -15,12 +15,12 @@ def core.marker.CopyU32 : core.marker.Copy U32 := { } /- [external::use_get]: - Source: 'tests/src/external.rs', lines 5:0-5:37 -/ + Source: 'tests/src/external.rs', lines 8:0-8:37 -/ def use_get (rc : core.cell.Cell U32) (st : State) : Result (State × U32) := core.cell.Cell.get U32 core.marker.CopyU32 rc st /- [external::incr]: - Source: 'tests/src/external.rs', lines 9:0-9:31 -/ + Source: 'tests/src/external.rs', lines 12:0-12:31 -/ def incr (rc : core.cell.Cell U32) (st : State) : Result (State × (core.cell.Cell U32)) diff --git a/tests/lean/Hashmap/Funs.lean b/tests/lean/Hashmap/Funs.lean index cb11e5cf..4f05fbc8 100644 --- a/tests/lean/Hashmap/Funs.lean +++ b/tests/lean/Hashmap/Funs.lean @@ -7,12 +7,12 @@ open Primitives namespace hashmap /- [hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 27:0-27:32 -/ + Source: 'tests/src/hashmap.rs', lines 35:0-35:32 -/ def hash_key (k : Usize) : Result Usize := Result.ok k /- [hashmap::{hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 50:4-56:5 -/ + Source: 'tests/src/hashmap.rs', lines 58:4-64:5 -/ divergent def HashMap.allocate_slots_loop (T : Type) (slots : alloc.vec.Vec (List T)) (n : Usize) : Result (alloc.vec.Vec (List T)) @@ -26,7 +26,7 @@ divergent def HashMap.allocate_slots_loop else Result.ok slots /- [hashmap::{hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 50:4-50:76 -/ + Source: 'tests/src/hashmap.rs', lines 58:4-58:76 -/ def HashMap.allocate_slots (T : Type) (slots : alloc.vec.Vec (List T)) (n : Usize) : Result (alloc.vec.Vec (List T)) @@ -34,7 +34,7 @@ def HashMap.allocate_slots HashMap.allocate_slots_loop T slots n /- [hashmap::{hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 59:4-63:13 -/ + Source: 'tests/src/hashmap.rs', lines 67:4-71:13 -/ def HashMap.new_with_capacity (T : Type) (capacity : Usize) (max_load_dividend : Usize) (max_load_divisor : Usize) : @@ -53,12 +53,12 @@ def HashMap.new_with_capacity } /- [hashmap::{hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 75:4-75:24 -/ + Source: 'tests/src/hashmap.rs', lines 83:4-83:24 -/ def HashMap.new (T : Type) : Result (HashMap T) := HashMap.new_with_capacity T 32#usize 4#usize 5#usize /- [hashmap::{hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 80:4-88:5 -/ + Source: 'tests/src/hashmap.rs', lines 88:4-96:5 -/ divergent def HashMap.clear_loop (T : Type) (slots : alloc.vec.Vec (List T)) (i : Usize) : Result (alloc.vec.Vec (List T)) @@ -76,19 +76,19 @@ divergent def HashMap.clear_loop else Result.ok slots /- [hashmap::{hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 80:4-80:27 -/ + Source: 'tests/src/hashmap.rs', lines 88:4-88: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}::len]: - Source: 'tests/src/hashmap.rs', lines 90:4-90:30 -/ + Source: 'tests/src/hashmap.rs', lines 98:4-98:30 -/ def HashMap.len (T : Type) (self : HashMap T) : Result Usize := Result.ok self.num_entries /- [hashmap::{hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 97:4-114:5 -/ + Source: 'tests/src/hashmap.rs', lines 105:4-122:5 -/ divergent def HashMap.insert_in_list_loop (T : Type) (key : Usize) (value : T) (ls : List T) : Result (Bool × (List T)) @@ -104,7 +104,7 @@ divergent def HashMap.insert_in_list_loop | List.Nil => Result.ok (true, List.Cons key value List.Nil) /- [hashmap::{hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 97:4-97:71 -/ + Source: 'tests/src/hashmap.rs', lines 105:4-105:71 -/ def HashMap.insert_in_list (T : Type) (key : Usize) (value : T) (ls : List T) : Result (Bool × (List T)) @@ -112,7 +112,7 @@ def HashMap.insert_in_list HashMap.insert_in_list_loop T key value ls /- [hashmap::{hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 117:4-117:54 -/ + Source: 'tests/src/hashmap.rs', lines 125:4-125:54 -/ def HashMap.insert_no_resize (T : Type) (self : HashMap T) (key : Usize) (value : T) : Result (HashMap T) @@ -136,7 +136,7 @@ def HashMap.insert_no_resize Result.ok { self with slots := v } /- [hashmap::{hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 183:4-196:5 -/ + Source: 'tests/src/hashmap.rs', lines 191:4-204:5 -/ divergent def HashMap.move_elements_from_list_loop (T : Type) (ntable : HashMap T) (ls : List T) : Result (HashMap T) := match ls with @@ -147,13 +147,13 @@ divergent def HashMap.move_elements_from_list_loop | List.Nil => Result.ok ntable /- [hashmap::{hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 183:4-183:72 -/ + Source: 'tests/src/hashmap.rs', lines 191:4-191:72 -/ def HashMap.move_elements_from_list (T : Type) (ntable : HashMap T) (ls : List T) : Result (HashMap T) := HashMap.move_elements_from_list_loop T ntable ls /- [hashmap::{hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 171:4-180:5 -/ + Source: 'tests/src/hashmap.rs', lines 179:4-188:5 -/ divergent def HashMap.move_elements_loop (T : Type) (ntable : HashMap T) (slots : alloc.vec.Vec (List T)) (i : Usize) : @@ -174,7 +174,7 @@ divergent def HashMap.move_elements_loop else Result.ok (ntable, slots) /- [hashmap::{hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 171:4-171:95 -/ + Source: 'tests/src/hashmap.rs', lines 179:4-179:95 -/ def HashMap.move_elements (T : Type) (ntable : HashMap T) (slots : alloc.vec.Vec (List T)) (i : Usize) : @@ -183,7 +183,7 @@ def HashMap.move_elements HashMap.move_elements_loop T ntable slots i /- [hashmap::{hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 140:4-140:28 -/ + Source: 'tests/src/hashmap.rs', lines 148:4-148:28 -/ def HashMap.try_resize (T : Type) (self : HashMap T) : Result (HashMap T) := do let max_usize ← Scalar.cast .Usize core_u32_max @@ -207,7 +207,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}::insert]: - Source: 'tests/src/hashmap.rs', lines 129:4-129:48 -/ + Source: 'tests/src/hashmap.rs', lines 137:4-137:48 -/ def HashMap.insert (T : Type) (self : HashMap T) (key : Usize) (value : T) : Result (HashMap T) @@ -220,7 +220,7 @@ def HashMap.insert else Result.ok self1 /- [hashmap::{hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 206:4-219:5 -/ + Source: 'tests/src/hashmap.rs', lines 214:4-227:5 -/ divergent def HashMap.contains_key_in_list_loop (T : Type) (key : Usize) (ls : List T) : Result Bool := match ls with @@ -231,13 +231,13 @@ divergent def HashMap.contains_key_in_list_loop | List.Nil => Result.ok false /- [hashmap::{hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 206:4-206:68 -/ + Source: 'tests/src/hashmap.rs', lines 214:4-214:68 -/ def HashMap.contains_key_in_list (T : Type) (key : Usize) (ls : List T) : Result Bool := HashMap.contains_key_in_list_loop T key ls /- [hashmap::{hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 199:4-199:49 -/ + Source: 'tests/src/hashmap.rs', lines 207:4-207:49 -/ def HashMap.contains_key (T : Type) (self : HashMap T) (key : Usize) : Result Bool := do @@ -250,7 +250,7 @@ def HashMap.contains_key HashMap.contains_key_in_list T key l /- [hashmap::{hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 224:4-237:5 -/ + Source: 'tests/src/hashmap.rs', lines 232:4-245:5 -/ divergent def HashMap.get_in_list_loop (T : Type) (key : Usize) (ls : List T) : Result T := match ls with @@ -261,12 +261,12 @@ divergent def HashMap.get_in_list_loop | List.Nil => Result.fail .panic /- [hashmap::{hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 224:4-224:70 -/ + Source: 'tests/src/hashmap.rs', lines 232:4-232:70 -/ def HashMap.get_in_list (T : Type) (key : Usize) (ls : List T) : Result T := HashMap.get_in_list_loop T key ls /- [hashmap::{hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 239:4-239:55 -/ + Source: 'tests/src/hashmap.rs', lines 247:4-247:55 -/ def HashMap.get (T : Type) (self : HashMap T) (key : Usize) : Result T := do let hash ← hash_key key @@ -278,7 +278,7 @@ def HashMap.get (T : Type) (self : HashMap T) (key : Usize) : Result T := HashMap.get_in_list T key l /- [hashmap::{hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 245:4-254:5 -/ + Source: 'tests/src/hashmap.rs', lines 253:4-262:5 -/ divergent def HashMap.get_mut_in_list_loop (T : Type) (ls : List T) (key : Usize) : Result (T × (T → Result (List T))) @@ -301,7 +301,7 @@ divergent def HashMap.get_mut_in_list_loop | List.Nil => Result.fail .panic /- [hashmap::{hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 245:4-245:86 -/ + Source: 'tests/src/hashmap.rs', lines 253:4-253:86 -/ def HashMap.get_mut_in_list (T : Type) (ls : List T) (key : Usize) : Result (T × (T → Result (List T))) @@ -309,7 +309,7 @@ def HashMap.get_mut_in_list HashMap.get_mut_in_list_loop T ls key /- [hashmap::{hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 257:4-257:67 -/ + Source: 'tests/src/hashmap.rs', lines 265:4-265:67 -/ def HashMap.get_mut (T : Type) (self : HashMap T) (key : Usize) : Result (T × (T → Result (HashMap T))) @@ -331,7 +331,7 @@ def HashMap.get_mut Result.ok (t, back) /- [hashmap::{hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 265:4-291:5 -/ + Source: 'tests/src/hashmap.rs', lines 273:4-299:5 -/ divergent def HashMap.remove_from_list_loop (T : Type) (key : Usize) (ls : List T) : Result ((Option T) × (List T)) := match ls with @@ -350,13 +350,13 @@ divergent def HashMap.remove_from_list_loop | List.Nil => Result.ok (none, List.Nil) /- [hashmap::{hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:69 -/ + Source: 'tests/src/hashmap.rs', lines 273:4-273:69 -/ def HashMap.remove_from_list (T : Type) (key : Usize) (ls : List T) : Result ((Option T) × (List T)) := HashMap.remove_from_list_loop T key ls /- [hashmap::{hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 294:4-294:52 -/ + Source: 'tests/src/hashmap.rs', lines 302:4-302:52 -/ def HashMap.remove (T : Type) (self : HashMap T) (key : Usize) : Result ((Option T) × (HashMap T)) @@ -381,7 +381,7 @@ def HashMap.remove Result.ok (some x1, { self with num_entries := i1, slots := v }) /- [hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 315:0-315:10 -/ + Source: 'tests/src/hashmap.rs', lines 323:0-323:10 -/ def test1 : Result Unit := do let hm ← HashMap.new U64 diff --git a/tests/lean/Hashmap/Types.lean b/tests/lean/Hashmap/Types.lean index c99ba8d0..a98b972f 100644 --- a/tests/lean/Hashmap/Types.lean +++ b/tests/lean/Hashmap/Types.lean @@ -6,13 +6,13 @@ open Primitives namespace hashmap /- [hashmap::List] - Source: 'tests/src/hashmap.rs', lines 19:0-19:16 -/ + Source: 'tests/src/hashmap.rs', lines 27:0-27:16 -/ inductive List (T : Type) := | Cons : Usize → T → List T → List T | Nil : List T /- [hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 35:0-35:21 -/ + Source: 'tests/src/hashmap.rs', lines 43:0-43:21 -/ structure HashMap (T : Type) where num_entries : Usize max_load_factor : (Usize × Usize) diff --git a/tests/lean/HashmapMain/Funs.lean b/tests/lean/HashmapMain/Funs.lean index e27305b1..45d6b058 100644 --- a/tests/lean/HashmapMain/Funs.lean +++ b/tests/lean/HashmapMain/Funs.lean @@ -8,12 +8,12 @@ open Primitives namespace hashmap_main /- [hashmap_main::hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 27:0-27:32 -/ + Source: 'tests/src/hashmap.rs', lines 35:0-35:32 -/ def hashmap.hash_key (k : Usize) : Result Usize := Result.ok k /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 50:4-56:5 -/ + Source: 'tests/src/hashmap.rs', lines 58:4-64:5 -/ divergent def hashmap.HashMap.allocate_slots_loop (T : Type) (slots : alloc.vec.Vec (hashmap.List T)) (n : Usize) : Result (alloc.vec.Vec (hashmap.List T)) @@ -27,7 +27,7 @@ divergent def hashmap.HashMap.allocate_slots_loop else Result.ok slots /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 50:4-50:76 -/ + Source: 'tests/src/hashmap.rs', lines 58:4-58:76 -/ def hashmap.HashMap.allocate_slots (T : Type) (slots : alloc.vec.Vec (hashmap.List T)) (n : Usize) : Result (alloc.vec.Vec (hashmap.List T)) @@ -35,7 +35,7 @@ def hashmap.HashMap.allocate_slots hashmap.HashMap.allocate_slots_loop T slots n /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 59:4-63:13 -/ + Source: 'tests/src/hashmap.rs', lines 67:4-71:13 -/ def hashmap.HashMap.new_with_capacity (T : Type) (capacity : Usize) (max_load_dividend : Usize) (max_load_divisor : Usize) : @@ -56,12 +56,12 @@ def hashmap.HashMap.new_with_capacity } /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 75:4-75:24 -/ + Source: 'tests/src/hashmap.rs', lines 83:4-83:24 -/ def hashmap.HashMap.new (T : Type) : Result (hashmap.HashMap T) := hashmap.HashMap.new_with_capacity T 32#usize 4#usize 5#usize /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 80:4-88:5 -/ + Source: 'tests/src/hashmap.rs', lines 88:4-96:5 -/ divergent def hashmap.HashMap.clear_loop (T : Type) (slots : alloc.vec.Vec (hashmap.List T)) (i : Usize) : Result (alloc.vec.Vec (hashmap.List T)) @@ -79,7 +79,7 @@ divergent def hashmap.HashMap.clear_loop else Result.ok slots /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 80:4-80:27 -/ + Source: 'tests/src/hashmap.rs', lines 88:4-88:27 -/ def hashmap.HashMap.clear (T : Type) (self : hashmap.HashMap T) : Result (hashmap.HashMap T) := do @@ -87,12 +87,12 @@ def hashmap.HashMap.clear Result.ok { self with num_entries := 0#usize, slots := hm } /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::len]: - Source: 'tests/src/hashmap.rs', lines 90:4-90:30 -/ + Source: 'tests/src/hashmap.rs', lines 98:4-98:30 -/ def hashmap.HashMap.len (T : Type) (self : hashmap.HashMap T) : Result Usize := Result.ok self.num_entries /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 97:4-114:5 -/ + Source: 'tests/src/hashmap.rs', lines 105:4-122:5 -/ divergent def hashmap.HashMap.insert_in_list_loop (T : Type) (key : Usize) (value : T) (ls : hashmap.List T) : Result (Bool × (hashmap.List T)) @@ -109,7 +109,7 @@ divergent def hashmap.HashMap.insert_in_list_loop Result.ok (true, hashmap.List.Cons key value hashmap.List.Nil) /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 97:4-97:71 -/ + Source: 'tests/src/hashmap.rs', lines 105:4-105:71 -/ def hashmap.HashMap.insert_in_list (T : Type) (key : Usize) (value : T) (ls : hashmap.List T) : Result (Bool × (hashmap.List T)) @@ -117,7 +117,7 @@ def hashmap.HashMap.insert_in_list hashmap.HashMap.insert_in_list_loop T key value ls /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 117:4-117:54 -/ + Source: 'tests/src/hashmap.rs', lines 125:4-125:54 -/ def hashmap.HashMap.insert_no_resize (T : Type) (self : hashmap.HashMap T) (key : Usize) (value : T) : Result (hashmap.HashMap T) @@ -142,7 +142,7 @@ def hashmap.HashMap.insert_no_resize Result.ok { self with slots := v } /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 183:4-196:5 -/ + Source: 'tests/src/hashmap.rs', lines 191:4-204:5 -/ divergent def hashmap.HashMap.move_elements_from_list_loop (T : Type) (ntable : hashmap.HashMap T) (ls : hashmap.List T) : Result (hashmap.HashMap T) @@ -155,7 +155,7 @@ divergent def hashmap.HashMap.move_elements_from_list_loop | hashmap.List.Nil => Result.ok ntable /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 183:4-183:72 -/ + Source: 'tests/src/hashmap.rs', lines 191:4-191:72 -/ def hashmap.HashMap.move_elements_from_list (T : Type) (ntable : hashmap.HashMap T) (ls : hashmap.List T) : Result (hashmap.HashMap T) @@ -163,7 +163,7 @@ def hashmap.HashMap.move_elements_from_list hashmap.HashMap.move_elements_from_list_loop T ntable ls /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 171:4-180:5 -/ + Source: 'tests/src/hashmap.rs', lines 179:4-188:5 -/ divergent def hashmap.HashMap.move_elements_loop (T : Type) (ntable : hashmap.HashMap T) (slots : alloc.vec.Vec (hashmap.List T)) (i : Usize) : @@ -184,7 +184,7 @@ divergent def hashmap.HashMap.move_elements_loop else Result.ok (ntable, slots) /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 171:4-171:95 -/ + Source: 'tests/src/hashmap.rs', lines 179:4-179:95 -/ def hashmap.HashMap.move_elements (T : Type) (ntable : hashmap.HashMap T) (slots : alloc.vec.Vec (hashmap.List T)) (i : Usize) : @@ -193,7 +193,7 @@ def hashmap.HashMap.move_elements hashmap.HashMap.move_elements_loop T ntable slots i /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 140:4-140:28 -/ + Source: 'tests/src/hashmap.rs', lines 148:4-148:28 -/ def hashmap.HashMap.try_resize (T : Type) (self : hashmap.HashMap T) : Result (hashmap.HashMap T) := do @@ -218,7 +218,7 @@ def hashmap.HashMap.try_resize else Result.ok { self with max_load_factor := (i, i1) } /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert]: - Source: 'tests/src/hashmap.rs', lines 129:4-129:48 -/ + Source: 'tests/src/hashmap.rs', lines 137:4-137:48 -/ def hashmap.HashMap.insert (T : Type) (self : hashmap.HashMap T) (key : Usize) (value : T) : Result (hashmap.HashMap T) @@ -231,7 +231,7 @@ def hashmap.HashMap.insert else Result.ok self1 /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 206:4-219:5 -/ + Source: 'tests/src/hashmap.rs', lines 214:4-227:5 -/ divergent def hashmap.HashMap.contains_key_in_list_loop (T : Type) (key : Usize) (ls : hashmap.List T) : Result Bool := match ls with @@ -242,13 +242,13 @@ divergent def hashmap.HashMap.contains_key_in_list_loop | hashmap.List.Nil => Result.ok false /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 206:4-206:68 -/ + Source: 'tests/src/hashmap.rs', lines 214:4-214:68 -/ def hashmap.HashMap.contains_key_in_list (T : Type) (key : Usize) (ls : hashmap.List T) : Result Bool := hashmap.HashMap.contains_key_in_list_loop T key ls /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 199:4-199:49 -/ + Source: 'tests/src/hashmap.rs', lines 207:4-207:49 -/ def hashmap.HashMap.contains_key (T : Type) (self : hashmap.HashMap T) (key : Usize) : Result Bool := do @@ -262,7 +262,7 @@ def hashmap.HashMap.contains_key hashmap.HashMap.contains_key_in_list T key l /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 224:4-237:5 -/ + Source: 'tests/src/hashmap.rs', lines 232:4-245:5 -/ divergent def hashmap.HashMap.get_in_list_loop (T : Type) (key : Usize) (ls : hashmap.List T) : Result T := match ls with @@ -273,13 +273,13 @@ divergent def hashmap.HashMap.get_in_list_loop | hashmap.List.Nil => Result.fail .panic /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 224:4-224:70 -/ + Source: 'tests/src/hashmap.rs', lines 232:4-232:70 -/ def hashmap.HashMap.get_in_list (T : Type) (key : Usize) (ls : hashmap.List T) : Result T := hashmap.HashMap.get_in_list_loop T key ls /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 239:4-239:55 -/ + Source: 'tests/src/hashmap.rs', lines 247:4-247:55 -/ def hashmap.HashMap.get (T : Type) (self : hashmap.HashMap T) (key : Usize) : Result T := do @@ -293,7 +293,7 @@ def hashmap.HashMap.get hashmap.HashMap.get_in_list T key l /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 245:4-254:5 -/ + Source: 'tests/src/hashmap.rs', lines 253:4-262:5 -/ divergent def hashmap.HashMap.get_mut_in_list_loop (T : Type) (ls : hashmap.List T) (key : Usize) : Result (T × (T → Result (hashmap.List T))) @@ -316,7 +316,7 @@ divergent def hashmap.HashMap.get_mut_in_list_loop | hashmap.List.Nil => Result.fail .panic /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 245:4-245:86 -/ + Source: 'tests/src/hashmap.rs', lines 253:4-253:86 -/ def hashmap.HashMap.get_mut_in_list (T : Type) (ls : hashmap.List T) (key : Usize) : Result (T × (T → Result (hashmap.List T))) @@ -324,7 +324,7 @@ def hashmap.HashMap.get_mut_in_list hashmap.HashMap.get_mut_in_list_loop T ls key /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 257:4-257:67 -/ + Source: 'tests/src/hashmap.rs', lines 265:4-265:67 -/ def hashmap.HashMap.get_mut (T : Type) (self : hashmap.HashMap T) (key : Usize) : Result (T × (T → Result (hashmap.HashMap T))) @@ -347,7 +347,7 @@ def hashmap.HashMap.get_mut Result.ok (t, back) /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 265:4-291:5 -/ + Source: 'tests/src/hashmap.rs', lines 273:4-299:5 -/ divergent def hashmap.HashMap.remove_from_list_loop (T : Type) (key : Usize) (ls : hashmap.List T) : Result ((Option T) × (hashmap.List T)) @@ -369,7 +369,7 @@ divergent def hashmap.HashMap.remove_from_list_loop | hashmap.List.Nil => Result.ok (none, hashmap.List.Nil) /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:69 -/ + Source: 'tests/src/hashmap.rs', lines 273:4-273:69 -/ def hashmap.HashMap.remove_from_list (T : Type) (key : Usize) (ls : hashmap.List T) : Result ((Option T) × (hashmap.List T)) @@ -377,7 +377,7 @@ def hashmap.HashMap.remove_from_list hashmap.HashMap.remove_from_list_loop T key ls /- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 294:4-294:52 -/ + Source: 'tests/src/hashmap.rs', lines 302:4-302:52 -/ def hashmap.HashMap.remove (T : Type) (self : hashmap.HashMap T) (key : Usize) : Result ((Option T) × (hashmap.HashMap T)) @@ -403,7 +403,7 @@ def hashmap.HashMap.remove Result.ok (some x1, { self with num_entries := i1, slots := v }) /- [hashmap_main::hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 315:0-315:10 -/ + Source: 'tests/src/hashmap.rs', lines 323:0-323:10 -/ def hashmap.test1 : Result Unit := do let hm ← hashmap.HashMap.new U64 @@ -447,7 +447,7 @@ def hashmap.test1 : Result Unit := else Result.ok () /- [hashmap_main::insert_on_disk]: - Source: 'tests/src/hashmap_main.rs', lines 7:0-7:43 -/ + Source: 'tests/src/hashmap_main.rs', lines 13:0-13:43 -/ def insert_on_disk (key : Usize) (value : U64) (st : State) : Result (State × Unit) := do @@ -456,7 +456,7 @@ def insert_on_disk hashmap_utils.serialize hm1 st1 /- [hashmap_main::main]: - Source: 'tests/src/hashmap_main.rs', lines 16:0-16:13 -/ + Source: 'tests/src/hashmap_main.rs', lines 22:0-22:13 -/ def main : Result Unit := Result.ok () diff --git a/tests/lean/HashmapMain/FunsExternal_Template.lean b/tests/lean/HashmapMain/FunsExternal_Template.lean index c9c41608..1a6c40d5 100644 --- a/tests/lean/HashmapMain/FunsExternal_Template.lean +++ b/tests/lean/HashmapMain/FunsExternal_Template.lean @@ -7,12 +7,12 @@ open Primitives open hashmap_main /- [hashmap_main::hashmap_utils::deserialize]: - Source: 'tests/src/hashmap_utils.rs', lines 10:0-10:43 -/ + Source: 'tests/src/hashmap_utils.rs', lines 11:0-11:43 -/ axiom hashmap_utils.deserialize : State → Result (State × (hashmap.HashMap U64)) /- [hashmap_main::hashmap_utils::serialize]: - Source: 'tests/src/hashmap_utils.rs', lines 5:0-5:42 -/ + Source: 'tests/src/hashmap_utils.rs', lines 6:0-6:42 -/ axiom hashmap_utils.serialize : hashmap.HashMap U64 → State → Result (State × Unit) diff --git a/tests/lean/HashmapMain/Types.lean b/tests/lean/HashmapMain/Types.lean index 076fb9ea..54f11e1e 100644 --- a/tests/lean/HashmapMain/Types.lean +++ b/tests/lean/HashmapMain/Types.lean @@ -7,13 +7,13 @@ open Primitives namespace hashmap_main /- [hashmap_main::hashmap::List] - Source: 'tests/src/hashmap.rs', lines 19:0-19:16 -/ + Source: 'tests/src/hashmap.rs', lines 27:0-27:16 -/ inductive hashmap.List (T : Type) := | Cons : Usize → T → hashmap.List T → hashmap.List T | Nil : hashmap.List T /- [hashmap_main::hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 35:0-35:21 -/ + Source: 'tests/src/hashmap.rs', lines 43:0-43:21 -/ structure hashmap.HashMap (T : Type) where num_entries : Usize max_load_factor : (Usize × Usize) diff --git a/tests/lean/Loops.lean b/tests/lean/Loops.lean index d5bda6ab..54f1b24f 100644 --- a/tests/lean/Loops.lean +++ b/tests/lean/Loops.lean @@ -6,7 +6,7 @@ open Primitives namespace loops /- [loops::sum]: loop 0: - Source: 'tests/src/loops.rs', lines 4:0-14:1 -/ + Source: 'tests/src/loops.rs', lines 7:0-17:1 -/ divergent def sum_loop (max : U32) (i : U32) (s : U32) : Result U32 := if i < max then do @@ -16,12 +16,12 @@ divergent def sum_loop (max : U32) (i : U32) (s : U32) : Result U32 := else s * 2#u32 /- [loops::sum]: - Source: 'tests/src/loops.rs', lines 4:0-4:27 -/ + Source: 'tests/src/loops.rs', lines 7:0-7:27 -/ def sum (max : U32) : Result U32 := sum_loop max 0#u32 0#u32 /- [loops::sum_with_mut_borrows]: loop 0: - Source: 'tests/src/loops.rs', lines 19:0-31:1 -/ + Source: 'tests/src/loops.rs', lines 22:0-34:1 -/ divergent def sum_with_mut_borrows_loop (max : U32) (i : U32) (s : U32) : Result U32 := if i < max @@ -33,12 +33,12 @@ divergent def sum_with_mut_borrows_loop else s * 2#u32 /- [loops::sum_with_mut_borrows]: - Source: 'tests/src/loops.rs', lines 19:0-19:44 -/ + Source: 'tests/src/loops.rs', lines 22:0-22:44 -/ def sum_with_mut_borrows (max : U32) : Result U32 := sum_with_mut_borrows_loop max 0#u32 0#u32 /- [loops::sum_with_shared_borrows]: loop 0: - Source: 'tests/src/loops.rs', lines 34:0-48:1 -/ + Source: 'tests/src/loops.rs', lines 37:0-51:1 -/ divergent def sum_with_shared_borrows_loop (max : U32) (i : U32) (s : U32) : Result U32 := if i < max @@ -50,12 +50,12 @@ divergent def sum_with_shared_borrows_loop else s * 2#u32 /- [loops::sum_with_shared_borrows]: - Source: 'tests/src/loops.rs', lines 34:0-34:47 -/ + Source: 'tests/src/loops.rs', lines 37:0-37:47 -/ def sum_with_shared_borrows (max : U32) : Result U32 := sum_with_shared_borrows_loop max 0#u32 0#u32 /- [loops::sum_array]: loop 0: - Source: 'tests/src/loops.rs', lines 50:0-58:1 -/ + Source: 'tests/src/loops.rs', lines 53:0-61:1 -/ divergent def sum_array_loop (N : Usize) (a : Array U32 N) (i : Usize) (s : U32) : Result U32 := if i < N @@ -68,12 +68,12 @@ divergent def sum_array_loop else Result.ok s /- [loops::sum_array]: - Source: 'tests/src/loops.rs', lines 50:0-50:52 -/ + Source: 'tests/src/loops.rs', lines 53:0-53:52 -/ def sum_array (N : Usize) (a : Array U32 N) : Result U32 := sum_array_loop N a 0#usize 0#u32 /- [loops::clear]: loop 0: - Source: 'tests/src/loops.rs', lines 62:0-68:1 -/ + Source: 'tests/src/loops.rs', lines 65:0-71:1 -/ divergent def clear_loop (v : alloc.vec.Vec U32) (i : Usize) : Result (alloc.vec.Vec U32) := let i1 := alloc.vec.Vec.len U32 v @@ -89,18 +89,18 @@ divergent def clear_loop else Result.ok v /- [loops::clear]: - Source: 'tests/src/loops.rs', lines 62:0-62:30 -/ + Source: 'tests/src/loops.rs', lines 65:0-65:30 -/ def clear (v : alloc.vec.Vec U32) : Result (alloc.vec.Vec U32) := clear_loop v 0#usize /- [loops::List] - Source: 'tests/src/loops.rs', lines 70:0-70:16 -/ + Source: 'tests/src/loops.rs', lines 73:0-73:16 -/ inductive List (T : Type) := | Cons : T → List T → List T | Nil : List T /- [loops::list_mem]: loop 0: - Source: 'tests/src/loops.rs', lines 76:0-85:1 -/ + Source: 'tests/src/loops.rs', lines 79:0-88:1 -/ divergent def list_mem_loop (x : U32) (ls : List U32) : Result Bool := match ls with | List.Cons y tl => if y = x @@ -109,12 +109,12 @@ divergent def list_mem_loop (x : U32) (ls : List U32) : Result Bool := | List.Nil => Result.ok false /- [loops::list_mem]: - Source: 'tests/src/loops.rs', lines 76:0-76:52 -/ + Source: 'tests/src/loops.rs', lines 79:0-79:52 -/ def list_mem (x : U32) (ls : List U32) : Result Bool := list_mem_loop x ls /- [loops::list_nth_mut_loop]: loop 0: - Source: 'tests/src/loops.rs', lines 88:0-98:1 -/ + Source: 'tests/src/loops.rs', lines 91:0-101:1 -/ divergent def list_nth_mut_loop_loop (T : Type) (ls : List T) (i : U32) : Result (T × (T → Result (List T))) := match ls with @@ -135,13 +135,13 @@ divergent def list_nth_mut_loop_loop | List.Nil => Result.fail .panic /- [loops::list_nth_mut_loop]: - Source: 'tests/src/loops.rs', lines 88:0-88:71 -/ + Source: 'tests/src/loops.rs', lines 91:0-91:71 -/ 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 /- [loops::list_nth_shared_loop]: loop 0: - Source: 'tests/src/loops.rs', lines 101:0-111:1 -/ + Source: 'tests/src/loops.rs', lines 104:0-114:1 -/ divergent def list_nth_shared_loop_loop (T : Type) (ls : List T) (i : U32) : Result T := match ls with @@ -154,12 +154,12 @@ divergent def list_nth_shared_loop_loop | List.Nil => Result.fail .panic /- [loops::list_nth_shared_loop]: - Source: 'tests/src/loops.rs', lines 101:0-101:66 -/ + Source: 'tests/src/loops.rs', lines 104:0-104:66 -/ def list_nth_shared_loop (T : Type) (ls : List T) (i : U32) : Result T := list_nth_shared_loop_loop T ls i /- [loops::get_elem_mut]: loop 0: - Source: 'tests/src/loops.rs', lines 113:0-127:1 -/ + Source: 'tests/src/loops.rs', lines 116:0-130:1 -/ divergent def get_elem_mut_loop (x : Usize) (ls : List Usize) : Result (Usize × (Usize → Result (List Usize))) @@ -181,7 +181,7 @@ divergent def get_elem_mut_loop | List.Nil => Result.fail .panic /- [loops::get_elem_mut]: - Source: 'tests/src/loops.rs', lines 113:0-113:73 -/ + Source: 'tests/src/loops.rs', lines 116:0-116:73 -/ def get_elem_mut (slots : alloc.vec.Vec (List Usize)) (x : Usize) : Result (Usize × (Usize → Result (alloc.vec.Vec (List Usize)))) @@ -197,7 +197,7 @@ def get_elem_mut Result.ok (i, back1) /- [loops::get_elem_shared]: loop 0: - Source: 'tests/src/loops.rs', lines 129:0-143:1 -/ + Source: 'tests/src/loops.rs', lines 132:0-146:1 -/ divergent def get_elem_shared_loop (x : Usize) (ls : List Usize) : Result Usize := match ls with @@ -207,7 +207,7 @@ divergent def get_elem_shared_loop | List.Nil => Result.fail .panic /- [loops::get_elem_shared]: - Source: 'tests/src/loops.rs', lines 129:0-129:68 -/ + Source: 'tests/src/loops.rs', lines 132:0-132:68 -/ def get_elem_shared (slots : alloc.vec.Vec (List Usize)) (x : Usize) : Result Usize := do @@ -217,7 +217,7 @@ def get_elem_shared get_elem_shared_loop x ls /- [loops::id_mut]: - Source: 'tests/src/loops.rs', lines 145:0-145:50 -/ + Source: 'tests/src/loops.rs', lines 148:0-148:50 -/ def id_mut (T : Type) (ls : List T) : Result ((List T) × (List T → Result (List T))) @@ -225,12 +225,12 @@ def id_mut Result.ok (ls, Result.ok) /- [loops::id_shared]: - Source: 'tests/src/loops.rs', lines 149:0-149:45 -/ + Source: 'tests/src/loops.rs', lines 152:0-152:45 -/ def id_shared (T : Type) (ls : List T) : Result (List T) := Result.ok ls /- [loops::list_nth_mut_loop_with_id]: loop 0: - Source: 'tests/src/loops.rs', lines 154:0-165:1 -/ + Source: 'tests/src/loops.rs', lines 157:0-168:1 -/ divergent def list_nth_mut_loop_with_id_loop (T : Type) (i : U32) (ls : List T) : Result (T × (T → Result (List T))) := match ls with @@ -251,7 +251,7 @@ divergent def list_nth_mut_loop_with_id_loop | List.Nil => Result.fail .panic /- [loops::list_nth_mut_loop_with_id]: - Source: 'tests/src/loops.rs', lines 154:0-154:75 -/ + Source: 'tests/src/loops.rs', lines 157:0-157:75 -/ def list_nth_mut_loop_with_id (T : Type) (ls : List T) (i : U32) : Result (T × (T → Result (List T))) := do @@ -263,7 +263,7 @@ def list_nth_mut_loop_with_id Result.ok (t, back1) /- [loops::list_nth_shared_loop_with_id]: loop 0: - Source: 'tests/src/loops.rs', lines 168:0-179:1 -/ + Source: 'tests/src/loops.rs', lines 171:0-182:1 -/ divergent def list_nth_shared_loop_with_id_loop (T : Type) (i : U32) (ls : List T) : Result T := match ls with @@ -276,7 +276,7 @@ divergent def list_nth_shared_loop_with_id_loop | List.Nil => Result.fail .panic /- [loops::list_nth_shared_loop_with_id]: - Source: 'tests/src/loops.rs', lines 168:0-168:70 -/ + Source: 'tests/src/loops.rs', lines 171:0-171:70 -/ def list_nth_shared_loop_with_id (T : Type) (ls : List T) (i : U32) : Result T := do @@ -284,7 +284,7 @@ def list_nth_shared_loop_with_id list_nth_shared_loop_with_id_loop T i ls1 /- [loops::list_nth_mut_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 184:0-205:1 -/ + Source: 'tests/src/loops.rs', lines 187:0-208:1 -/ divergent def list_nth_mut_loop_pair_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T)) × (T → Result (List T))) @@ -315,7 +315,7 @@ divergent def list_nth_mut_loop_pair_loop | List.Nil => Result.fail .panic /- [loops::list_nth_mut_loop_pair]: - Source: 'tests/src/loops.rs', lines 184:0-188:27 -/ + Source: 'tests/src/loops.rs', lines 187:0-191:27 -/ 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))) @@ -323,7 +323,7 @@ def list_nth_mut_loop_pair list_nth_mut_loop_pair_loop T ls0 ls1 i /- [loops::list_nth_shared_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 208:0-229:1 -/ + Source: 'tests/src/loops.rs', lines 211:0-232:1 -/ divergent def list_nth_shared_loop_pair_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result (T × T) := match ls0 with @@ -339,13 +339,13 @@ divergent def list_nth_shared_loop_pair_loop | List.Nil => Result.fail .panic /- [loops::list_nth_shared_loop_pair]: - Source: 'tests/src/loops.rs', lines 208:0-212:19 -/ + Source: 'tests/src/loops.rs', lines 211:0-215:19 -/ 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 /- [loops::list_nth_mut_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 233:0-248:1 -/ + Source: 'tests/src/loops.rs', lines 236:0-251:1 -/ divergent def list_nth_mut_loop_pair_merge_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × ((T × T) → Result ((List T) × (List T)))) @@ -375,7 +375,7 @@ divergent def list_nth_mut_loop_pair_merge_loop | List.Nil => Result.fail .panic /- [loops::list_nth_mut_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 233:0-237:27 -/ + Source: 'tests/src/loops.rs', lines 236:0-240:27 -/ 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)))) @@ -383,7 +383,7 @@ def list_nth_mut_loop_pair_merge list_nth_mut_loop_pair_merge_loop T ls0 ls1 i /- [loops::list_nth_shared_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 251:0-266:1 -/ + Source: 'tests/src/loops.rs', lines 254:0-269:1 -/ divergent def list_nth_shared_loop_pair_merge_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result (T × T) := match ls0 with @@ -400,13 +400,13 @@ divergent def list_nth_shared_loop_pair_merge_loop | List.Nil => Result.fail .panic /- [loops::list_nth_shared_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 251:0-255:19 -/ + Source: 'tests/src/loops.rs', lines 254:0-258:19 -/ 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 /- [loops::list_nth_mut_shared_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 269:0-284:1 -/ + Source: 'tests/src/loops.rs', lines 272:0-287:1 -/ divergent def list_nth_mut_shared_loop_pair_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -432,7 +432,7 @@ divergent def list_nth_mut_shared_loop_pair_loop | List.Nil => Result.fail .panic /- [loops::list_nth_mut_shared_loop_pair]: - Source: 'tests/src/loops.rs', lines 269:0-273:23 -/ + Source: 'tests/src/loops.rs', lines 272:0-276:23 -/ def list_nth_mut_shared_loop_pair (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -440,7 +440,7 @@ def list_nth_mut_shared_loop_pair list_nth_mut_shared_loop_pair_loop T ls0 ls1 i /- [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 288:0-303:1 -/ + Source: 'tests/src/loops.rs', lines 291:0-306:1 -/ divergent def list_nth_mut_shared_loop_pair_merge_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -466,7 +466,7 @@ divergent def list_nth_mut_shared_loop_pair_merge_loop | List.Nil => Result.fail .panic /- [loops::list_nth_mut_shared_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 288:0-292:23 -/ + Source: 'tests/src/loops.rs', lines 291:0-295:23 -/ 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))) @@ -474,7 +474,7 @@ def list_nth_mut_shared_loop_pair_merge list_nth_mut_shared_loop_pair_merge_loop T ls0 ls1 i /- [loops::list_nth_shared_mut_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 307:0-322:1 -/ + Source: 'tests/src/loops.rs', lines 310:0-325:1 -/ divergent def list_nth_shared_mut_loop_pair_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -500,7 +500,7 @@ divergent def list_nth_shared_mut_loop_pair_loop | List.Nil => Result.fail .panic /- [loops::list_nth_shared_mut_loop_pair]: - Source: 'tests/src/loops.rs', lines 307:0-311:23 -/ + Source: 'tests/src/loops.rs', lines 310:0-314:23 -/ def list_nth_shared_mut_loop_pair (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -508,7 +508,7 @@ def list_nth_shared_mut_loop_pair list_nth_shared_mut_loop_pair_loop T ls0 ls1 i /- [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 326:0-341:1 -/ + Source: 'tests/src/loops.rs', lines 329:0-344:1 -/ divergent def list_nth_shared_mut_loop_pair_merge_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -534,7 +534,7 @@ divergent def list_nth_shared_mut_loop_pair_merge_loop | List.Nil => Result.fail .panic /- [loops::list_nth_shared_mut_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 326:0-330:23 -/ + Source: 'tests/src/loops.rs', lines 329:0-333:23 -/ 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))) @@ -542,7 +542,7 @@ def list_nth_shared_mut_loop_pair_merge list_nth_shared_mut_loop_pair_merge_loop T ls0 ls1 i /- [loops::ignore_input_mut_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 345:0-349:1 -/ + Source: 'tests/src/loops.rs', lines 348:0-352:1 -/ divergent def ignore_input_mut_borrow_loop (i : U32) : Result Unit := if i > 0#u32 then do @@ -551,14 +551,14 @@ divergent def ignore_input_mut_borrow_loop (i : U32) : Result Unit := else Result.ok () /- [loops::ignore_input_mut_borrow]: - Source: 'tests/src/loops.rs', lines 345:0-345:56 -/ + Source: 'tests/src/loops.rs', lines 348:0-348:56 -/ def ignore_input_mut_borrow (_a : U32) (i : U32) : Result U32 := do let _ ← ignore_input_mut_borrow_loop i Result.ok _a /- [loops::incr_ignore_input_mut_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 353:0-358:1 -/ + Source: 'tests/src/loops.rs', lines 356:0-361:1 -/ divergent def incr_ignore_input_mut_borrow_loop (i : U32) : Result Unit := if i > 0#u32 then do @@ -567,7 +567,7 @@ divergent def incr_ignore_input_mut_borrow_loop (i : U32) : Result Unit := else Result.ok () /- [loops::incr_ignore_input_mut_borrow]: - Source: 'tests/src/loops.rs', lines 353:0-353:60 -/ + Source: 'tests/src/loops.rs', lines 356:0-356:60 -/ def incr_ignore_input_mut_borrow (a : U32) (i : U32) : Result U32 := do let a1 ← a + 1#u32 @@ -575,7 +575,7 @@ def incr_ignore_input_mut_borrow (a : U32) (i : U32) : Result U32 := Result.ok a1 /- [loops::ignore_input_shared_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 362:0-366:1 -/ + Source: 'tests/src/loops.rs', lines 365:0-369:1 -/ divergent def ignore_input_shared_borrow_loop (i : U32) : Result Unit := if i > 0#u32 then do @@ -584,7 +584,7 @@ divergent def ignore_input_shared_borrow_loop (i : U32) : Result Unit := else Result.ok () /- [loops::ignore_input_shared_borrow]: - Source: 'tests/src/loops.rs', lines 362:0-362:59 -/ + Source: 'tests/src/loops.rs', lines 365:0-365:59 -/ def ignore_input_shared_borrow (_a : U32) (i : U32) : Result U32 := do let _ ← ignore_input_shared_borrow_loop i diff --git a/tests/lean/NoNestedBorrows.lean b/tests/lean/NoNestedBorrows.lean index 5ae22055..1781ac71 100644 --- a/tests/lean/NoNestedBorrows.lean +++ b/tests/lean/NoNestedBorrows.lean @@ -6,60 +6,60 @@ open Primitives namespace no_nested_borrows /- [no_nested_borrows::Pair] - Source: 'tests/src/no_nested_borrows.rs', lines 4:0-4:23 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 6:0-6:23 -/ structure Pair (T1 T2 : Type) where x : T1 y : T2 /- [no_nested_borrows::List] - Source: 'tests/src/no_nested_borrows.rs', lines 9:0-9:16 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 11:0-11:16 -/ inductive List (T : Type) := | Cons : T → List T → List T | Nil : List T /- [no_nested_borrows::One] - Source: 'tests/src/no_nested_borrows.rs', lines 20:0-20:16 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 22:0-22:16 -/ inductive One (T1 : Type) := | One : T1 → One T1 /- [no_nested_borrows::EmptyEnum] - Source: 'tests/src/no_nested_borrows.rs', lines 26:0-26:18 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 28:0-28:18 -/ inductive EmptyEnum := | Empty : EmptyEnum /- [no_nested_borrows::Enum] - Source: 'tests/src/no_nested_borrows.rs', lines 32:0-32:13 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 34:0-34:13 -/ inductive Enum := | Variant1 : Enum | Variant2 : Enum /- [no_nested_borrows::EmptyStruct] - Source: 'tests/src/no_nested_borrows.rs', lines 39:0-39:22 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 41:0-41:22 -/ @[reducible] def EmptyStruct := Unit /- [no_nested_borrows::Sum] - Source: 'tests/src/no_nested_borrows.rs', lines 41:0-41:20 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 43:0-43:20 -/ inductive Sum (T1 T2 : Type) := | Left : T1 → Sum T1 T2 | Right : T2 → Sum T1 T2 /- [no_nested_borrows::cast_u32_to_i32]: - Source: 'tests/src/no_nested_borrows.rs', lines 46:0-46:37 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 48:0-48:37 -/ def cast_u32_to_i32 (x : U32) : Result I32 := Scalar.cast .I32 x /- [no_nested_borrows::cast_bool_to_i32]: - Source: 'tests/src/no_nested_borrows.rs', lines 50:0-50:39 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 52:0-52:39 -/ def cast_bool_to_i32 (x : Bool) : Result I32 := Scalar.cast_bool .I32 x /- [no_nested_borrows::cast_bool_to_bool]: - Source: 'tests/src/no_nested_borrows.rs', lines 55:0-55:41 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 57:0-57:41 -/ def cast_bool_to_bool (x : Bool) : Result Bool := Result.ok x /- [no_nested_borrows::test2]: - Source: 'tests/src/no_nested_borrows.rs', lines 60:0-60:14 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 62:0-62:14 -/ def test2 : Result Unit := do let _ ← 23#u32 + 44#u32 @@ -69,14 +69,14 @@ def test2 : Result Unit := #assert (test2 == Result.ok ()) /- [no_nested_borrows::get_max]: - Source: 'tests/src/no_nested_borrows.rs', lines 72:0-72:37 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 74:0-74:37 -/ def get_max (x : U32) (y : U32) : Result U32 := if x >= y then Result.ok x else Result.ok y /- [no_nested_borrows::test3]: - Source: 'tests/src/no_nested_borrows.rs', lines 80:0-80:14 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 82:0-82:14 -/ def test3 : Result Unit := do let x ← get_max 4#u32 3#u32 @@ -90,7 +90,7 @@ def test3 : Result Unit := #assert (test3 == Result.ok ()) /- [no_nested_borrows::test_neg1]: - Source: 'tests/src/no_nested_borrows.rs', lines 87:0-87:18 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 89:0-89:18 -/ def test_neg1 : Result Unit := do let y ← -. 3#i32 @@ -102,7 +102,7 @@ def test_neg1 : Result Unit := #assert (test_neg1 == Result.ok ()) /- [no_nested_borrows::refs_test1]: - Source: 'tests/src/no_nested_borrows.rs', lines 94:0-94:19 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 96:0-96:19 -/ def refs_test1 : Result Unit := if ¬ (1#i32 = 1#i32) then Result.fail .panic @@ -112,7 +112,7 @@ def refs_test1 : Result Unit := #assert (refs_test1 == Result.ok ()) /- [no_nested_borrows::refs_test2]: - Source: 'tests/src/no_nested_borrows.rs', lines 105:0-105:19 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 107:0-107:19 -/ def refs_test2 : Result Unit := if ¬ (2#i32 = 2#i32) then Result.fail .panic @@ -130,7 +130,7 @@ def refs_test2 : Result Unit := #assert (refs_test2 == Result.ok ()) /- [no_nested_borrows::test_list1]: - Source: 'tests/src/no_nested_borrows.rs', lines 121:0-121:19 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 123:0-123:19 -/ def test_list1 : Result Unit := Result.ok () @@ -138,7 +138,7 @@ def test_list1 : Result Unit := #assert (test_list1 == Result.ok ()) /- [no_nested_borrows::test_box1]: - Source: 'tests/src/no_nested_borrows.rs', lines 126:0-126:18 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 128:0-128:18 -/ def test_box1 : Result Unit := do let (_, deref_mut_back) ← alloc.boxed.Box.deref_mut I32 0#i32 @@ -152,26 +152,26 @@ def test_box1 : Result Unit := #assert (test_box1 == Result.ok ()) /- [no_nested_borrows::copy_int]: - Source: 'tests/src/no_nested_borrows.rs', lines 136:0-136:30 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 138:0-138:30 -/ def copy_int (x : I32) : Result I32 := Result.ok x /- [no_nested_borrows::test_unreachable]: - Source: 'tests/src/no_nested_borrows.rs', lines 142:0-142:32 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 144:0-144:32 -/ def test_unreachable (b : Bool) : Result Unit := if b then Result.fail .panic else Result.ok () /- [no_nested_borrows::test_panic]: - Source: 'tests/src/no_nested_borrows.rs', lines 150:0-150:26 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 152:0-152:26 -/ def test_panic (b : Bool) : Result Unit := if b then Result.fail .panic else Result.ok () /- [no_nested_borrows::test_copy_int]: - Source: 'tests/src/no_nested_borrows.rs', lines 157:0-157:22 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 159:0-159:22 -/ def test_copy_int : Result Unit := do let y ← copy_int 0#i32 @@ -183,14 +183,14 @@ def test_copy_int : Result Unit := #assert (test_copy_int == Result.ok ()) /- [no_nested_borrows::is_cons]: - Source: 'tests/src/no_nested_borrows.rs', lines 164:0-164:38 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 166:0-166:38 -/ def is_cons (T : Type) (l : List T) : Result Bool := match l with | List.Cons _ _ => Result.ok true | List.Nil => Result.ok false /- [no_nested_borrows::test_is_cons]: - Source: 'tests/src/no_nested_borrows.rs', lines 171:0-171:21 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 173:0-173:21 -/ def test_is_cons : Result Unit := do let b ← is_cons I32 (List.Cons 0#i32 List.Nil) @@ -202,14 +202,14 @@ def test_is_cons : Result Unit := #assert (test_is_cons == Result.ok ()) /- [no_nested_borrows::split_list]: - Source: 'tests/src/no_nested_borrows.rs', lines 177:0-177:48 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 179:0-179:48 -/ def split_list (T : Type) (l : List T) : Result (T × (List T)) := match l with | List.Cons hd tl => Result.ok (hd, tl) | List.Nil => Result.fail .panic /- [no_nested_borrows::test_split_list]: - Source: 'tests/src/no_nested_borrows.rs', lines 185:0-185:24 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 187:0-187:24 -/ def test_split_list : Result Unit := do let p ← split_list I32 (List.Cons 0#i32 List.Nil) @@ -222,7 +222,7 @@ def test_split_list : Result Unit := #assert (test_split_list == Result.ok ()) /- [no_nested_borrows::choose]: - Source: 'tests/src/no_nested_borrows.rs', lines 192:0-192:70 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 194:0-194:70 -/ def choose (T : Type) (b : Bool) (x : T) (y : T) : Result (T × (T → Result (T × T))) @@ -234,7 +234,7 @@ def choose Result.ok (y, back) /- [no_nested_borrows::choose_test]: - Source: 'tests/src/no_nested_borrows.rs', lines 200:0-200:20 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 202:0-202:20 -/ def choose_test : Result Unit := do let (z, choose_back) ← choose I32 true 0#i32 0#i32 @@ -254,20 +254,20 @@ def choose_test : Result Unit := #assert (choose_test == Result.ok ()) /- [no_nested_borrows::test_char]: - Source: 'tests/src/no_nested_borrows.rs', lines 212:0-212:26 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 214:0-214:26 -/ def test_char : Result Char := Result.ok 'a' mutual /- [no_nested_borrows::Tree] - Source: 'tests/src/no_nested_borrows.rs', lines 217:0-217:16 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 219:0-219:16 -/ inductive Tree (T : Type) := | Leaf : T → Tree T | Node : T → NodeElem T → Tree T → Tree T /- [no_nested_borrows::NodeElem] - Source: 'tests/src/no_nested_borrows.rs', lines 222:0-222:20 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 224:0-224:20 -/ inductive NodeElem (T : Type) := | Cons : Tree T → NodeElem T → NodeElem T | Nil : NodeElem T @@ -275,7 +275,7 @@ inductive NodeElem (T : Type) := end /- [no_nested_borrows::list_length]: - Source: 'tests/src/no_nested_borrows.rs', lines 257:0-257:48 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 259:0-259:48 -/ divergent def list_length (T : Type) (l : List T) : Result U32 := match l with | List.Cons _ l1 => do @@ -284,7 +284,7 @@ divergent def list_length (T : Type) (l : List T) : Result U32 := | List.Nil => Result.ok 0#u32 /- [no_nested_borrows::list_nth_shared]: - Source: 'tests/src/no_nested_borrows.rs', lines 265:0-265:62 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 267:0-267:62 -/ divergent def list_nth_shared (T : Type) (l : List T) (i : U32) : Result T := match l with | List.Cons x tl => @@ -296,7 +296,7 @@ divergent def list_nth_shared (T : Type) (l : List T) (i : U32) : Result T := | List.Nil => Result.fail .panic /- [no_nested_borrows::list_nth_mut]: - Source: 'tests/src/no_nested_borrows.rs', lines 281:0-281:67 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 283:0-283:67 -/ divergent def list_nth_mut (T : Type) (l : List T) (i : U32) : Result (T × (T → Result (List T))) := match l with @@ -318,7 +318,7 @@ divergent def list_nth_mut | List.Nil => Result.fail .panic /- [no_nested_borrows::list_rev_aux]: - Source: 'tests/src/no_nested_borrows.rs', lines 297:0-297:63 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 299:0-299:63 -/ divergent def list_rev_aux (T : Type) (li : List T) (lo : List T) : Result (List T) := match li with @@ -326,13 +326,13 @@ divergent def list_rev_aux | List.Nil => Result.ok lo /- [no_nested_borrows::list_rev]: - Source: 'tests/src/no_nested_borrows.rs', lines 311:0-311:42 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 313:0-313:42 -/ def list_rev (T : Type) (l : List T) : Result (List T) := let (li, _) := core.mem.replace (List T) l List.Nil list_rev_aux T li List.Nil /- [no_nested_borrows::test_list_functions]: - Source: 'tests/src/no_nested_borrows.rs', lines 316:0-316:28 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 318:0-318:28 -/ def test_list_functions : Result Unit := do let l := List.Cons 2#i32 List.Nil @@ -379,7 +379,7 @@ def test_list_functions : Result Unit := #assert (test_list_functions == Result.ok ()) /- [no_nested_borrows::id_mut_pair1]: - Source: 'tests/src/no_nested_borrows.rs', lines 332:0-332:89 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 334:0-334:89 -/ def id_mut_pair1 (T1 T2 : Type) (x : T1) (y : T2) : Result ((T1 × T2) × ((T1 × T2) → Result (T1 × T2))) @@ -387,7 +387,7 @@ def id_mut_pair1 Result.ok ((x, y), Result.ok) /- [no_nested_borrows::id_mut_pair2]: - Source: 'tests/src/no_nested_borrows.rs', lines 336:0-336:88 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 338:0-338:88 -/ def id_mut_pair2 (T1 T2 : Type) (p : (T1 × T2)) : Result ((T1 × T2) × ((T1 × T2) → Result (T1 × T2))) @@ -396,7 +396,7 @@ def id_mut_pair2 Result.ok ((t, t1), Result.ok) /- [no_nested_borrows::id_mut_pair3]: - Source: 'tests/src/no_nested_borrows.rs', lines 340:0-340:93 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 342:0-342:93 -/ def id_mut_pair3 (T1 T2 : Type) (x : T1) (y : T2) : Result ((T1 × T2) × (T1 → Result T1) × (T2 → Result T2)) @@ -404,7 +404,7 @@ def id_mut_pair3 Result.ok ((x, y), Result.ok, Result.ok) /- [no_nested_borrows::id_mut_pair4]: - Source: 'tests/src/no_nested_borrows.rs', lines 344:0-344:92 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 346:0-346:92 -/ def id_mut_pair4 (T1 T2 : Type) (p : (T1 × T2)) : Result ((T1 × T2) × (T1 → Result T1) × (T2 → Result T2)) @@ -413,37 +413,37 @@ def id_mut_pair4 Result.ok ((t, t1), Result.ok, Result.ok) /- [no_nested_borrows::StructWithTuple] - Source: 'tests/src/no_nested_borrows.rs', lines 351:0-351:34 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 353:0-353:34 -/ structure StructWithTuple (T1 T2 : Type) where p : (T1 × T2) /- [no_nested_borrows::new_tuple1]: - Source: 'tests/src/no_nested_borrows.rs', lines 355:0-355:48 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 357:0-357:48 -/ def new_tuple1 : Result (StructWithTuple U32 U32) := Result.ok { p := (1#u32, 2#u32) } /- [no_nested_borrows::new_tuple2]: - Source: 'tests/src/no_nested_borrows.rs', lines 359:0-359:48 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 361:0-361:48 -/ def new_tuple2 : Result (StructWithTuple I16 I16) := Result.ok { p := (1#i16, 2#i16) } /- [no_nested_borrows::new_tuple3]: - Source: 'tests/src/no_nested_borrows.rs', lines 363:0-363:48 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 365:0-365:48 -/ def new_tuple3 : Result (StructWithTuple U64 I64) := Result.ok { p := (1#u64, 2#i64) } /- [no_nested_borrows::StructWithPair] - Source: 'tests/src/no_nested_borrows.rs', lines 368:0-368:33 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 370:0-370:33 -/ structure StructWithPair (T1 T2 : Type) where p : Pair T1 T2 /- [no_nested_borrows::new_pair1]: - Source: 'tests/src/no_nested_borrows.rs', lines 372:0-372:46 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 374:0-374:46 -/ def new_pair1 : Result (StructWithPair U32 U32) := Result.ok { p := { x := 1#u32, y := 2#u32 } } /- [no_nested_borrows::test_constants]: - Source: 'tests/src/no_nested_borrows.rs', lines 380:0-380:23 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 382:0-382:23 -/ def test_constants : Result Unit := do let swt ← new_tuple1 @@ -473,7 +473,7 @@ def test_constants : Result Unit := #assert (test_constants == Result.ok ()) /- [no_nested_borrows::test_weird_borrows1]: - Source: 'tests/src/no_nested_borrows.rs', lines 389:0-389:28 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 391:0-391:28 -/ def test_weird_borrows1 : Result Unit := Result.ok () @@ -481,7 +481,7 @@ def test_weird_borrows1 : Result Unit := #assert (test_weird_borrows1 == Result.ok ()) /- [no_nested_borrows::test_mem_replace]: - Source: 'tests/src/no_nested_borrows.rs', lines 399:0-399:37 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 401:0-401:37 -/ def test_mem_replace (px : U32) : Result U32 := let (y, _) := core.mem.replace U32 px 1#u32 if ¬ (y = 0#u32) @@ -489,71 +489,71 @@ def test_mem_replace (px : U32) : Result U32 := else Result.ok 2#u32 /- [no_nested_borrows::test_shared_borrow_bool1]: - Source: 'tests/src/no_nested_borrows.rs', lines 406:0-406:47 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 408:0-408:47 -/ def test_shared_borrow_bool1 (b : Bool) : Result U32 := if b then Result.ok 0#u32 else Result.ok 1#u32 /- [no_nested_borrows::test_shared_borrow_bool2]: - Source: 'tests/src/no_nested_borrows.rs', lines 419:0-419:40 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 421:0-421:40 -/ def test_shared_borrow_bool2 : Result U32 := Result.ok 0#u32 /- [no_nested_borrows::test_shared_borrow_enum1]: - Source: 'tests/src/no_nested_borrows.rs', lines 434:0-434:52 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 436:0-436:52 -/ def test_shared_borrow_enum1 (l : List U32) : Result U32 := match l with | List.Cons _ _ => Result.ok 1#u32 | List.Nil => Result.ok 0#u32 /- [no_nested_borrows::test_shared_borrow_enum2]: - Source: 'tests/src/no_nested_borrows.rs', lines 446:0-446:40 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 448:0-448:40 -/ def test_shared_borrow_enum2 : Result U32 := Result.ok 0#u32 /- [no_nested_borrows::incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 457:0-457:24 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 459:0-459:24 -/ def incr (x : U32) : Result U32 := x + 1#u32 /- [no_nested_borrows::call_incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 461:0-461:35 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 463:0-463:35 -/ def call_incr (x : U32) : Result U32 := incr x /- [no_nested_borrows::read_then_incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 466:0-466:41 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 468:0-468:41 -/ def read_then_incr (x : U32) : Result (U32 × U32) := do let x1 ← x + 1#u32 Result.ok (x, x1) /- [no_nested_borrows::Tuple] - Source: 'tests/src/no_nested_borrows.rs', lines 472:0-472:24 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 474:0-474:24 -/ def Tuple (T1 T2 : Type) := T1 × T2 /- [no_nested_borrows::use_tuple_struct]: - Source: 'tests/src/no_nested_borrows.rs', lines 474:0-474:48 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 476:0-476:48 -/ def use_tuple_struct (x : Tuple U32 U32) : Result (Tuple U32 U32) := Result.ok (1#u32, x.#1) /- [no_nested_borrows::create_tuple_struct]: - Source: 'tests/src/no_nested_borrows.rs', lines 478:0-478:61 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 480:0-480:61 -/ def create_tuple_struct (x : U32) (y : U64) : Result (Tuple U32 U64) := Result.ok (x, y) /- [no_nested_borrows::IdType] - Source: 'tests/src/no_nested_borrows.rs', lines 483:0-483:20 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 485:0-485:20 -/ @[reducible] def IdType (T : Type) := T /- [no_nested_borrows::use_id_type]: - Source: 'tests/src/no_nested_borrows.rs', lines 485:0-485:40 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 487:0-487:40 -/ def use_id_type (T : Type) (x : IdType T) : Result T := Result.ok x /- [no_nested_borrows::create_id_type]: - Source: 'tests/src/no_nested_borrows.rs', lines 489:0-489:43 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 491:0-491:43 -/ def create_id_type (T : Type) (x : T) : Result (IdType T) := Result.ok x diff --git a/tests/lean/Paper.lean b/tests/lean/Paper.lean index bede3dd9..e98ada42 100644 --- a/tests/lean/Paper.lean +++ b/tests/lean/Paper.lean @@ -6,12 +6,12 @@ open Primitives namespace paper /- [paper::ref_incr]: - Source: 'tests/src/paper.rs', lines 4:0-4:28 -/ + Source: 'tests/src/paper.rs', lines 6:0-6:28 -/ def ref_incr (x : I32) : Result I32 := x + 1#i32 /- [paper::test_incr]: - Source: 'tests/src/paper.rs', lines 8:0-8:18 -/ + Source: 'tests/src/paper.rs', lines 10:0-10:18 -/ def test_incr : Result Unit := do let x ← ref_incr 0#i32 @@ -23,7 +23,7 @@ def test_incr : Result Unit := #assert (test_incr == Result.ok ()) /- [paper::choose]: - Source: 'tests/src/paper.rs', lines 15:0-15:70 -/ + Source: 'tests/src/paper.rs', lines 17:0-17:70 -/ def choose (T : Type) (b : Bool) (x : T) (y : T) : Result (T × (T → Result (T × T))) @@ -35,7 +35,7 @@ def choose Result.ok (y, back) /- [paper::test_choose]: - Source: 'tests/src/paper.rs', lines 23:0-23:20 -/ + Source: 'tests/src/paper.rs', lines 25:0-25:20 -/ def test_choose : Result Unit := do let (z, choose_back) ← choose I32 true 0#i32 0#i32 @@ -55,13 +55,13 @@ def test_choose : Result Unit := #assert (test_choose == Result.ok ()) /- [paper::List] - Source: 'tests/src/paper.rs', lines 35:0-35:16 -/ + Source: 'tests/src/paper.rs', lines 37:0-37:16 -/ inductive List (T : Type) := | Cons : T → List T → List T | Nil : List T /- [paper::list_nth_mut]: - Source: 'tests/src/paper.rs', lines 42:0-42:67 -/ + Source: 'tests/src/paper.rs', lines 44:0-44:67 -/ divergent def list_nth_mut (T : Type) (l : List T) (i : U32) : Result (T × (T → Result (List T))) := match l with @@ -83,7 +83,7 @@ divergent def list_nth_mut | List.Nil => Result.fail .panic /- [paper::sum]: - Source: 'tests/src/paper.rs', lines 57:0-57:32 -/ + Source: 'tests/src/paper.rs', lines 59:0-59:32 -/ divergent def sum (l : List I32) : Result I32 := match l with | List.Cons x tl => do @@ -92,7 +92,7 @@ divergent def sum (l : List I32) : Result I32 := | List.Nil => Result.ok 0#i32 /- [paper::test_nth]: - Source: 'tests/src/paper.rs', lines 68:0-68:17 -/ + Source: 'tests/src/paper.rs', lines 70:0-70:17 -/ def test_nth : Result Unit := do let l := List.Cons 3#i32 List.Nil @@ -109,7 +109,7 @@ def test_nth : Result Unit := #assert (test_nth == Result.ok ()) /- [paper::call_choose]: - Source: 'tests/src/paper.rs', lines 76:0-76:44 -/ + Source: 'tests/src/paper.rs', lines 78:0-78:44 -/ def call_choose (p : (U32 × U32)) : Result U32 := do let (px, py) := p diff --git a/tests/lean/PoloniusList.lean b/tests/lean/PoloniusList.lean index d1ed10d2..defa48c7 100644 --- a/tests/lean/PoloniusList.lean +++ b/tests/lean/PoloniusList.lean @@ -6,13 +6,13 @@ open Primitives namespace polonius_list /- [polonius_list::List] - Source: 'tests/src/polonius_list.rs', lines 3:0-3:16 -/ + Source: 'tests/src/polonius_list.rs', lines 5:0-5:16 -/ inductive List (T : Type) := | Cons : T → List T → List T | Nil : List T /- [polonius_list::get_list_at_x]: - Source: 'tests/src/polonius_list.rs', lines 13:0-13:76 -/ + Source: 'tests/src/polonius_list.rs', lines 15:0-15:76 -/ divergent def get_list_at_x (ls : List U32) (x : U32) : Result ((List U32) × (List U32 → Result (List U32))) diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean index 32f84676..7cacb836 100644 --- a/tests/lean/Traits.lean +++ b/tests/lean/Traits.lean @@ -6,29 +6,29 @@ open Primitives namespace traits /- Trait declaration: [traits::BoolTrait] - Source: 'tests/src/traits.rs', lines 1:0-1:19 -/ + Source: 'tests/src/traits.rs', lines 2:0-2:19 -/ structure BoolTrait (Self : Type) where get_bool : Self → Result Bool /- [traits::{(traits::BoolTrait for bool)}::get_bool]: - Source: 'tests/src/traits.rs', lines 12:4-12:30 -/ + Source: 'tests/src/traits.rs', lines 13:4-13:30 -/ def BoolTraitBool.get_bool (self : Bool) : Result Bool := Result.ok self /- Trait implementation: [traits::{(traits::BoolTrait for bool)}] - Source: 'tests/src/traits.rs', lines 11:0-11:23 -/ + Source: 'tests/src/traits.rs', lines 12:0-12:23 -/ def BoolTraitBool : BoolTrait Bool := { get_bool := BoolTraitBool.get_bool } /- [traits::BoolTrait::ret_true]: - Source: 'tests/src/traits.rs', lines 6:4-6:30 -/ + Source: 'tests/src/traits.rs', lines 7:4-7:30 -/ def BoolTrait.ret_true {Self : Type} (self_clause : BoolTrait Self) (self : Self) : Result Bool := Result.ok true /- [traits::test_bool_trait_bool]: - Source: 'tests/src/traits.rs', lines 17:0-17:44 -/ + Source: 'tests/src/traits.rs', lines 18:0-18:44 -/ def test_bool_trait_bool (x : Bool) : Result Bool := do let b ← BoolTraitBool.get_bool x @@ -37,20 +37,20 @@ def test_bool_trait_bool (x : Bool) : Result Bool := else Result.ok false /- [traits::{(traits::BoolTrait for core::option::Option)#1}::get_bool]: - Source: 'tests/src/traits.rs', lines 23:4-23:30 -/ + Source: 'tests/src/traits.rs', lines 24:4-24:30 -/ def BoolTraitOption.get_bool (T : Type) (self : Option T) : Result Bool := match self with | none => Result.ok false | some _ => Result.ok true /- Trait implementation: [traits::{(traits::BoolTrait for core::option::Option)#1}] - Source: 'tests/src/traits.rs', lines 22:0-22:31 -/ + Source: 'tests/src/traits.rs', lines 23:0-23:31 -/ def BoolTraitOption (T : Type) : BoolTrait (Option T) := { get_bool := BoolTraitOption.get_bool T } /- [traits::test_bool_trait_option]: - Source: 'tests/src/traits.rs', lines 31:0-31:54 -/ + Source: 'tests/src/traits.rs', lines 32:0-32:54 -/ def test_bool_trait_option (T : Type) (x : Option T) : Result Bool := do let b ← BoolTraitOption.get_bool T x @@ -59,29 +59,29 @@ def test_bool_trait_option (T : Type) (x : Option T) : Result Bool := else Result.ok false /- [traits::test_bool_trait]: - Source: 'tests/src/traits.rs', lines 35:0-35:50 -/ + Source: 'tests/src/traits.rs', lines 36:0-36:50 -/ def test_bool_trait (T : Type) (BoolTraitInst : BoolTrait T) (x : T) : Result Bool := BoolTraitInst.get_bool x /- Trait declaration: [traits::ToU64] - Source: 'tests/src/traits.rs', lines 39:0-39:15 -/ + Source: 'tests/src/traits.rs', lines 40:0-40:15 -/ structure ToU64 (Self : Type) where to_u64 : Self → Result U64 /- [traits::{(traits::ToU64 for u64)#2}::to_u64]: - Source: 'tests/src/traits.rs', lines 44:4-44:26 -/ + Source: 'tests/src/traits.rs', lines 45:4-45:26 -/ def ToU64U64.to_u64 (self : U64) : Result U64 := Result.ok self /- Trait implementation: [traits::{(traits::ToU64 for u64)#2}] - Source: 'tests/src/traits.rs', lines 43:0-43:18 -/ + Source: 'tests/src/traits.rs', lines 44:0-44:18 -/ def ToU64U64 : ToU64 U64 := { to_u64 := ToU64U64.to_u64 } /- [traits::{(traits::ToU64 for (A, A))#3}::to_u64]: - Source: 'tests/src/traits.rs', lines 50:4-50:26 -/ + Source: 'tests/src/traits.rs', lines 51:4-51:26 -/ def ToU64Pair.to_u64 (A : Type) (ToU64Inst : ToU64 A) (self : (A × A)) : Result U64 := do @@ -91,78 +91,78 @@ def ToU64Pair.to_u64 i + i1 /- Trait implementation: [traits::{(traits::ToU64 for (A, A))#3}] - Source: 'tests/src/traits.rs', lines 49:0-49:31 -/ + Source: 'tests/src/traits.rs', lines 50:0-50:31 -/ def ToU64Pair (A : Type) (ToU64Inst : ToU64 A) : ToU64 (A × A) := { to_u64 := ToU64Pair.to_u64 A ToU64Inst } /- [traits::f]: - Source: 'tests/src/traits.rs', lines 55:0-55:36 -/ + Source: 'tests/src/traits.rs', lines 56:0-56:36 -/ def f (T : Type) (ToU64Inst : ToU64 T) (x : (T × T)) : Result U64 := ToU64Pair.to_u64 T ToU64Inst x /- [traits::g]: - Source: 'tests/src/traits.rs', lines 59:0-61:18 -/ + Source: 'tests/src/traits.rs', lines 60:0-62:18 -/ def g (T : Type) (ToU64PairInst : ToU64 (T × T)) (x : (T × T)) : Result U64 := ToU64PairInst.to_u64 x /- [traits::h0]: - Source: 'tests/src/traits.rs', lines 66:0-66:24 -/ + Source: 'tests/src/traits.rs', lines 67:0-67:24 -/ def h0 (x : U64) : Result U64 := ToU64U64.to_u64 x /- [traits::Wrapper] - Source: 'tests/src/traits.rs', lines 70:0-70:21 -/ + Source: 'tests/src/traits.rs', lines 71:0-71:21 -/ structure Wrapper (T : Type) where x : T /- [traits::{(traits::ToU64 for traits::Wrapper)#4}::to_u64]: - Source: 'tests/src/traits.rs', lines 75:4-75:26 -/ + Source: 'tests/src/traits.rs', lines 76:4-76:26 -/ def ToU64traitsWrapper.to_u64 (T : Type) (ToU64Inst : ToU64 T) (self : Wrapper T) : Result U64 := ToU64Inst.to_u64 self.x /- Trait implementation: [traits::{(traits::ToU64 for traits::Wrapper)#4}] - Source: 'tests/src/traits.rs', lines 74:0-74:35 -/ + Source: 'tests/src/traits.rs', lines 75:0-75:35 -/ def ToU64traitsWrapper (T : Type) (ToU64Inst : ToU64 T) : ToU64 (Wrapper T) := { to_u64 := ToU64traitsWrapper.to_u64 T ToU64Inst } /- [traits::h1]: - Source: 'tests/src/traits.rs', lines 80:0-80:33 -/ + Source: 'tests/src/traits.rs', lines 81:0-81:33 -/ def h1 (x : Wrapper U64) : Result U64 := ToU64traitsWrapper.to_u64 U64 ToU64U64 x /- [traits::h2]: - Source: 'tests/src/traits.rs', lines 84:0-84:41 -/ + Source: 'tests/src/traits.rs', lines 85:0-85:41 -/ def h2 (T : Type) (ToU64Inst : ToU64 T) (x : Wrapper T) : Result U64 := ToU64traitsWrapper.to_u64 T ToU64Inst x /- Trait declaration: [traits::ToType] - Source: 'tests/src/traits.rs', lines 88:0-88:19 -/ + Source: 'tests/src/traits.rs', lines 89:0-89:19 -/ structure ToType (Self T : Type) where to_type : Self → Result T /- [traits::{(traits::ToType for u64)#5}::to_type]: - Source: 'tests/src/traits.rs', lines 93:4-93:28 -/ + Source: 'tests/src/traits.rs', lines 94:4-94:28 -/ def ToTypeU64Bool.to_type (self : U64) : Result Bool := Result.ok (self > 0#u64) /- Trait implementation: [traits::{(traits::ToType for u64)#5}] - Source: 'tests/src/traits.rs', lines 92:0-92:25 -/ + Source: 'tests/src/traits.rs', lines 93:0-93:25 -/ def ToTypeU64Bool : ToType U64 Bool := { to_type := ToTypeU64Bool.to_type } /- Trait declaration: [traits::OfType] - Source: 'tests/src/traits.rs', lines 98:0-98:16 -/ + Source: 'tests/src/traits.rs', lines 99:0-99:16 -/ structure OfType (Self : Type) where of_type : forall (T : Type) (ToTypeInst : ToType T Self), T → Result Self /- [traits::h3]: - Source: 'tests/src/traits.rs', lines 104:0-104:50 -/ + Source: 'tests/src/traits.rs', lines 105:0-105:50 -/ def h3 (T1 T2 : Type) (OfTypeInst : OfType T1) (ToTypeInst : ToType T2 T1) (y : T2) : @@ -171,13 +171,13 @@ def h3 OfTypeInst.of_type T2 ToTypeInst y /- Trait declaration: [traits::OfTypeBis] - Source: 'tests/src/traits.rs', lines 109:0-109:36 -/ + Source: 'tests/src/traits.rs', lines 110:0-110:36 -/ structure OfTypeBis (Self T : Type) where ToTypeInst : ToType T Self of_type : T → Result Self /- [traits::h4]: - Source: 'tests/src/traits.rs', lines 118:0-118:57 -/ + Source: 'tests/src/traits.rs', lines 119:0-119:57 -/ def h4 (T1 T2 : Type) (OfTypeBisInst : OfTypeBis T1 T2) (ToTypeInst : ToType T2 T1) (y : T2) : @@ -186,33 +186,33 @@ def h4 OfTypeBisInst.of_type y /- [traits::TestType] - Source: 'tests/src/traits.rs', lines 122:0-122:22 -/ + Source: 'tests/src/traits.rs', lines 123:0-123:22 -/ @[reducible] def TestType (T : Type) := T /- [traits::{traits::TestType#6}::test::TestType1] - Source: 'tests/src/traits.rs', lines 127:8-127:24 -/ + Source: 'tests/src/traits.rs', lines 128:8-128:24 -/ @[reducible] def TestType.test.TestType1 := U64 /- Trait declaration: [traits::{traits::TestType#6}::test::TestTrait] - Source: 'tests/src/traits.rs', lines 128:8-128:23 -/ + Source: 'tests/src/traits.rs', lines 129:8-129:23 -/ structure TestType.test.TestTrait (Self : Type) where test : Self → Result Bool /- [traits::{traits::TestType#6}::test::{(traits::{traits::TestType#6}::test::TestTrait for traits::{traits::TestType#6}::test::TestType1)}::test]: - Source: 'tests/src/traits.rs', lines 139:12-139:34 -/ + Source: 'tests/src/traits.rs', lines 140:12-140:34 -/ def TestType.test.TestTraittraitsTestTypetestTestType1.test (self : TestType.test.TestType1) : Result Bool := Result.ok (self > 1#u64) /- Trait implementation: [traits::{traits::TestType#6}::test::{(traits::{traits::TestType#6}::test::TestTrait for traits::{traits::TestType#6}::test::TestType1)}] - Source: 'tests/src/traits.rs', lines 138:8-138:36 -/ + Source: 'tests/src/traits.rs', lines 139:8-139:36 -/ def TestType.test.TestTraittraitsTestTypetestTestType1 : TestType.test.TestTrait TestType.test.TestType1 := { test := TestType.test.TestTraittraitsTestTypetestTestType1.test } /- [traits::{traits::TestType#6}::test]: - Source: 'tests/src/traits.rs', lines 126:4-126:36 -/ + Source: 'tests/src/traits.rs', lines 127:4-127:36 -/ def TestType.test (T : Type) (ToU64Inst : ToU64 T) (self : TestType T) (x : T) : Result Bool := do @@ -222,11 +222,11 @@ def TestType.test else Result.ok false /- [traits::BoolWrapper] - Source: 'tests/src/traits.rs', lines 150:0-150:22 -/ + Source: 'tests/src/traits.rs', lines 151:0-151:22 -/ @[reducible] def BoolWrapper := Bool /- [traits::{(traits::ToType for traits::BoolWrapper)#7}::to_type]: - Source: 'tests/src/traits.rs', lines 156:4-156:25 -/ + Source: 'tests/src/traits.rs', lines 157:4-157:25 -/ def ToTypetraitsBoolWrapperT.to_type (T : Type) (ToTypeBoolTInst : ToType Bool T) (self : BoolWrapper) : Result T @@ -234,21 +234,21 @@ def ToTypetraitsBoolWrapperT.to_type ToTypeBoolTInst.to_type self /- Trait implementation: [traits::{(traits::ToType for traits::BoolWrapper)#7}] - Source: 'tests/src/traits.rs', lines 152:0-152:33 -/ + Source: 'tests/src/traits.rs', lines 153:0-153:33 -/ def ToTypetraitsBoolWrapperT (T : Type) (ToTypeBoolTInst : ToType Bool T) : ToType BoolWrapper T := { to_type := ToTypetraitsBoolWrapperT.to_type T ToTypeBoolTInst } /- [traits::WithConstTy::LEN2] - Source: 'tests/src/traits.rs', lines 164:4-164:21 -/ + Source: 'tests/src/traits.rs', lines 165:4-165:21 -/ def WithConstTy.LEN2_default_body (Self : Type) (LEN : Usize) : Result Usize := Result.ok 32#usize def WithConstTy.LEN2_default (Self : Type) (LEN : Usize) : Usize := eval_global (WithConstTy.LEN2_default_body Self LEN) /- Trait declaration: [traits::WithConstTy] - Source: 'tests/src/traits.rs', lines 161:0-161:39 -/ + Source: 'tests/src/traits.rs', lines 162:0-162:39 -/ structure WithConstTy (Self : Type) (LEN : Usize) where LEN1 : Usize LEN2 : Usize @@ -258,17 +258,17 @@ structure WithConstTy (Self : Type) (LEN : Usize) where f : W → Array U8 LEN → Result W /- [traits::{(traits::WithConstTy<32: usize> for bool)#8}::LEN1] - Source: 'tests/src/traits.rs', lines 175:4-175:21 -/ + Source: 'tests/src/traits.rs', lines 176:4-176:21 -/ def WithConstTyBool32.LEN1_body : Result Usize := Result.ok 12#usize def WithConstTyBool32.LEN1 : Usize := eval_global WithConstTyBool32.LEN1_body /- [traits::{(traits::WithConstTy<32: usize> for bool)#8}::f]: - Source: 'tests/src/traits.rs', lines 180:4-180:39 -/ + Source: 'tests/src/traits.rs', lines 181:4-181:39 -/ def WithConstTyBool32.f (i : U64) (a : Array U8 32#usize) : Result U64 := Result.ok i /- Trait implementation: [traits::{(traits::WithConstTy<32: usize> for bool)#8}] - Source: 'tests/src/traits.rs', lines 174:0-174:29 -/ + Source: 'tests/src/traits.rs', lines 175:0-175:29 -/ def WithConstTyBool32 : WithConstTy Bool 32#usize := { LEN1 := WithConstTyBool32.LEN1 LEN2 := WithConstTy.LEN2_default Bool 32#usize @@ -279,7 +279,7 @@ def WithConstTyBool32 : WithConstTy Bool 32#usize := { } /- [traits::use_with_const_ty1]: - Source: 'tests/src/traits.rs', lines 183:0-183:75 -/ + Source: 'tests/src/traits.rs', lines 184:0-184:75 -/ def use_with_const_ty1 (H : Type) (LEN : Usize) (WithConstTyInst : WithConstTy H LEN) : Result Usize @@ -287,7 +287,7 @@ def use_with_const_ty1 Result.ok WithConstTyInst.LEN1 /- [traits::use_with_const_ty2]: - Source: 'tests/src/traits.rs', lines 187:0-187:73 -/ + Source: 'tests/src/traits.rs', lines 188:0-188:73 -/ def use_with_const_ty2 (H : Type) (LEN : Usize) (WithConstTyInst : WithConstTy H LEN) (w : WithConstTyInst.W) : @@ -296,7 +296,7 @@ def use_with_const_ty2 Result.ok () /- [traits::use_with_const_ty3]: - Source: 'tests/src/traits.rs', lines 189:0-189:80 -/ + Source: 'tests/src/traits.rs', lines 190:0-190:80 -/ def use_with_const_ty3 (H : Type) (LEN : Usize) (WithConstTyInst : WithConstTy H LEN) (x : WithConstTyInst.W) : @@ -305,12 +305,12 @@ def use_with_const_ty3 WithConstTyInst.W_clause_0.to_u64 x /- [traits::test_where1]: - Source: 'tests/src/traits.rs', lines 193:0-193:40 -/ + Source: 'tests/src/traits.rs', lines 194:0-194:40 -/ def test_where1 (T : Type) (_x : T) : Result Unit := Result.ok () /- [traits::test_where2]: - Source: 'tests/src/traits.rs', lines 194:0-194:57 -/ + Source: 'tests/src/traits.rs', lines 195:0-195:57 -/ def test_where2 (T : Type) (WithConstTyT32Inst : WithConstTy T 32#usize) (_x : U32) : Result Unit @@ -318,30 +318,30 @@ def test_where2 Result.ok () /- Trait declaration: [traits::ParentTrait0] - Source: 'tests/src/traits.rs', lines 200:0-200:22 -/ + Source: 'tests/src/traits.rs', lines 201:0-201:22 -/ structure ParentTrait0 (Self : Type) where W : Type get_name : Self → Result String get_w : Self → Result W /- Trait declaration: [traits::ParentTrait1] - Source: 'tests/src/traits.rs', lines 205:0-205:22 -/ + Source: 'tests/src/traits.rs', lines 206:0-206:22 -/ structure ParentTrait1 (Self : Type) where /- Trait declaration: [traits::ChildTrait] - Source: 'tests/src/traits.rs', lines 206:0-206:49 -/ + Source: 'tests/src/traits.rs', lines 207:0-207:49 -/ structure ChildTrait (Self : Type) where ParentTrait0Inst : ParentTrait0 Self ParentTrait1Inst : ParentTrait1 Self /- [traits::test_child_trait1]: - Source: 'tests/src/traits.rs', lines 209:0-209:56 -/ + Source: 'tests/src/traits.rs', lines 210:0-210:56 -/ def test_child_trait1 (T : Type) (ChildTraitInst : ChildTrait T) (x : T) : Result String := ChildTraitInst.ParentTrait0Inst.get_name x /- [traits::test_child_trait2]: - Source: 'tests/src/traits.rs', lines 213:0-213:54 -/ + Source: 'tests/src/traits.rs', lines 214:0-214:54 -/ def test_child_trait2 (T : Type) (ChildTraitInst : ChildTrait T) (x : T) : Result ChildTraitInst.ParentTrait0Inst.W @@ -349,7 +349,7 @@ def test_child_trait2 ChildTraitInst.ParentTrait0Inst.get_w x /- [traits::order1]: - Source: 'tests/src/traits.rs', lines 219:0-219:59 -/ + Source: 'tests/src/traits.rs', lines 220:0-220:59 -/ def order1 (T U : Type) (ParentTrait0Inst : ParentTrait0 T) (ParentTrait0Inst1 : ParentTrait0 U) : @@ -358,28 +358,28 @@ def order1 Result.ok () /- Trait declaration: [traits::ChildTrait1] - Source: 'tests/src/traits.rs', lines 222:0-222:35 -/ + Source: 'tests/src/traits.rs', lines 223:0-223:35 -/ structure ChildTrait1 (Self : Type) where ParentTrait1Inst : ParentTrait1 Self /- Trait implementation: [traits::{(traits::ParentTrait1 for usize)#9}] - Source: 'tests/src/traits.rs', lines 224:0-224:27 -/ + Source: 'tests/src/traits.rs', lines 225:0-225:27 -/ def ParentTrait1Usize : ParentTrait1 Usize := { } /- Trait implementation: [traits::{(traits::ChildTrait1 for usize)#10}] - Source: 'tests/src/traits.rs', lines 225:0-225:26 -/ + Source: 'tests/src/traits.rs', lines 226:0-226:26 -/ def ChildTrait1Usize : ChildTrait1 Usize := { ParentTrait1Inst := ParentTrait1Usize } /- Trait declaration: [traits::Iterator] - Source: 'tests/src/traits.rs', lines 229:0-229:18 -/ + Source: 'tests/src/traits.rs', lines 230:0-230:18 -/ structure Iterator (Self : Type) where Item : Type /- Trait declaration: [traits::IntoIterator] - Source: 'tests/src/traits.rs', lines 233:0-233:22 -/ + Source: 'tests/src/traits.rs', lines 234:0-234:22 -/ structure IntoIterator (Self : Type) where Item : Type IntoIter : Type @@ -387,106 +387,106 @@ structure IntoIterator (Self : Type) where into_iter : Self → Result IntoIter /- Trait declaration: [traits::FromResidual] - Source: 'tests/src/traits.rs', lines 250:0-250:21 -/ + Source: 'tests/src/traits.rs', lines 251:0-251:21 -/ structure FromResidual (Self T : Type) where /- Trait declaration: [traits::Try] - Source: 'tests/src/traits.rs', lines 246:0-246:48 -/ + Source: 'tests/src/traits.rs', lines 247:0-247:48 -/ structure Try (Self : Type) where Residual : Type FromResidualSelftraitsTryResidualInst : FromResidual Self Residual /- Trait declaration: [traits::WithTarget] - Source: 'tests/src/traits.rs', lines 252:0-252:20 -/ + Source: 'tests/src/traits.rs', lines 253:0-253:20 -/ structure WithTarget (Self : Type) where Target : Type /- Trait declaration: [traits::ParentTrait2] - Source: 'tests/src/traits.rs', lines 256:0-256:22 -/ + Source: 'tests/src/traits.rs', lines 257:0-257:22 -/ structure ParentTrait2 (Self : Type) where U : Type U_clause_0 : WithTarget U /- Trait declaration: [traits::ChildTrait2] - Source: 'tests/src/traits.rs', lines 260:0-260:35 -/ + Source: 'tests/src/traits.rs', lines 261:0-261:35 -/ structure ChildTrait2 (Self : Type) where ParentTrait2Inst : ParentTrait2 Self convert : ParentTrait2Inst.U → Result ParentTrait2Inst.U_clause_0.Target /- Trait implementation: [traits::{(traits::WithTarget for u32)#11}] - Source: 'tests/src/traits.rs', lines 264:0-264:23 -/ + Source: 'tests/src/traits.rs', lines 265:0-265:23 -/ def WithTargetU32 : WithTarget U32 := { Target := U32 } /- Trait implementation: [traits::{(traits::ParentTrait2 for u32)#12}] - Source: 'tests/src/traits.rs', lines 268:0-268:25 -/ + Source: 'tests/src/traits.rs', lines 269:0-269:25 -/ def ParentTrait2U32 : ParentTrait2 U32 := { U := U32 U_clause_0 := WithTargetU32 } /- [traits::{(traits::ChildTrait2 for u32)#13}::convert]: - Source: 'tests/src/traits.rs', lines 273:4-273:29 -/ + Source: 'tests/src/traits.rs', lines 274:4-274:29 -/ def ChildTrait2U32.convert (x : U32) : Result U32 := Result.ok x /- Trait implementation: [traits::{(traits::ChildTrait2 for u32)#13}] - Source: 'tests/src/traits.rs', lines 272:0-272:24 -/ + Source: 'tests/src/traits.rs', lines 273:0-273:24 -/ def ChildTrait2U32 : ChildTrait2 U32 := { ParentTrait2Inst := ParentTrait2U32 convert := ChildTrait2U32.convert } /- Trait declaration: [traits::CFnOnce] - Source: 'tests/src/traits.rs', lines 286:0-286:23 -/ + Source: 'tests/src/traits.rs', lines 287:0-287:23 -/ structure CFnOnce (Self Args : Type) where Output : Type call_once : Self → Args → Result Output /- Trait declaration: [traits::CFnMut] - Source: 'tests/src/traits.rs', lines 292:0-292:37 -/ + Source: 'tests/src/traits.rs', lines 293:0-293:37 -/ structure CFnMut (Self Args : Type) where CFnOnceInst : CFnOnce Self Args call_mut : Self → Args → Result (CFnOnceInst.Output × Self) /- Trait declaration: [traits::CFn] - Source: 'tests/src/traits.rs', lines 296:0-296:33 -/ + Source: 'tests/src/traits.rs', lines 297:0-297:33 -/ structure CFn (Self Args : Type) where CFnMutInst : CFnMut Self Args call : Self → Args → Result CFnMutInst.CFnOnceInst.Output /- Trait declaration: [traits::GetTrait] - Source: 'tests/src/traits.rs', lines 300:0-300:18 -/ + Source: 'tests/src/traits.rs', lines 301:0-301:18 -/ structure GetTrait (Self : Type) where W : Type get_w : Self → Result W /- [traits::test_get_trait]: - Source: 'tests/src/traits.rs', lines 305:0-305:49 -/ + Source: 'tests/src/traits.rs', lines 306:0-306:49 -/ def test_get_trait (T : Type) (GetTraitInst : GetTrait T) (x : T) : Result GetTraitInst.W := GetTraitInst.get_w x /- Trait declaration: [traits::Trait] - Source: 'tests/src/traits.rs', lines 310:0-310:15 -/ + Source: 'tests/src/traits.rs', lines 311:0-311:15 -/ structure Trait (Self : Type) where LEN : Usize /- [traits::{(traits::Trait for @Array)#14}::LEN] - Source: 'tests/src/traits.rs', lines 315:4-315:20 -/ + Source: 'tests/src/traits.rs', lines 316:4-316:20 -/ def TraitArray.LEN_body (T : Type) (N : Usize) : Result Usize := Result.ok N def TraitArray.LEN (T : Type) (N : Usize) : Usize := eval_global (TraitArray.LEN_body T N) /- Trait implementation: [traits::{(traits::Trait for @Array)#14}] - Source: 'tests/src/traits.rs', lines 314:0-314:40 -/ + Source: 'tests/src/traits.rs', lines 315:0-315:40 -/ def TraitArray (T : Type) (N : Usize) : Trait (Array T N) := { LEN := TraitArray.LEN T N } /- [traits::{(traits::Trait for traits::Wrapper)#15}::LEN] - Source: 'tests/src/traits.rs', lines 319:4-319:20 -/ + Source: 'tests/src/traits.rs', lines 320:4-320:20 -/ def TraittraitsWrapper.LEN_body (T : Type) (TraitInst : Trait T) : Result Usize := Result.ok 0#usize @@ -494,19 +494,19 @@ def TraittraitsWrapper.LEN (T : Type) (TraitInst : Trait T) : Usize := eval_global (TraittraitsWrapper.LEN_body T TraitInst) /- Trait implementation: [traits::{(traits::Trait for traits::Wrapper)#15}] - Source: 'tests/src/traits.rs', lines 318:0-318:35 -/ + Source: 'tests/src/traits.rs', lines 319:0-319:35 -/ def TraittraitsWrapper (T : Type) (TraitInst : Trait T) : Trait (Wrapper T) := { LEN := TraittraitsWrapper.LEN T TraitInst } /- [traits::use_wrapper_len]: - Source: 'tests/src/traits.rs', lines 322:0-322:43 -/ + Source: 'tests/src/traits.rs', lines 323:0-323:43 -/ def use_wrapper_len (T : Type) (TraitInst : Trait T) : Result Usize := Result.ok (TraittraitsWrapper T TraitInst).LEN /- [traits::Foo] - Source: 'tests/src/traits.rs', lines 326:0-326:20 -/ + Source: 'tests/src/traits.rs', lines 327:0-327:20 -/ structure Foo (T U : Type) where x : T y : U @@ -519,7 +519,7 @@ inductive core.result.Result (T E : Type) := | Err : E → core.result.Result T E /- [traits::{traits::Foo#16}::FOO] - Source: 'tests/src/traits.rs', lines 332:4-332:33 -/ + Source: 'tests/src/traits.rs', lines 333:4-333:33 -/ def Foo.FOO_body (T U : Type) (TraitInst : Trait T) : Result (core.result.Result T I32) := Result.ok (core.result.Result.Err 0#i32) @@ -527,13 +527,13 @@ def Foo.FOO (T U : Type) (TraitInst : Trait T) : core.result.Result T I32 := eval_global (Foo.FOO_body T U TraitInst) /- [traits::use_foo1]: - Source: 'tests/src/traits.rs', lines 335:0-335:48 -/ + Source: 'tests/src/traits.rs', lines 336:0-336:48 -/ def use_foo1 (T U : Type) (TraitInst : Trait T) : Result (core.result.Result T I32) := Result.ok (Foo.FOO T U TraitInst) /- [traits::use_foo2]: - Source: 'tests/src/traits.rs', lines 339:0-339:48 -/ + Source: 'tests/src/traits.rs', lines 340:0-340:48 -/ def use_foo2 (T U : Type) (TraitInst : Trait U) : Result (core.result.Result U I32) := Result.ok (Foo.FOO U T TraitInst) -- cgit v1.2.3 From c4af12c1c34406720d8173f2972d4cf1f42f8f5b Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Fri, 24 May 2024 16:20:04 +0200 Subject: fix generated file --- tests/test_runner/aeneas_test_runner.opam | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/tests/test_runner/aeneas_test_runner.opam b/tests/test_runner/aeneas_test_runner.opam index b57cc9f6..1539c521 100644 --- a/tests/test_runner/aeneas_test_runner.opam +++ b/tests/test_runner/aeneas_test_runner.opam @@ -7,7 +7,7 @@ homepage: "https://github.com/AeneasVerif/aeneas" bug-reports: "https://github.com/AeneasVerif/aeneas/issues" depends: [ "ocaml" - "dune" {>= "3.12"} + "dune" {>= "3.7"} "odoc" {with-doc} ] build: [ -- cgit v1.2.3 From 0ebfa7a15f4d7218389488ff8a92206c0d6642ec Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Fri, 24 May 2024 16:20:26 +0200 Subject: runner: Allow filenames with dashes --- tests/test_runner/dune | 2 +- tests/test_runner/run_test.ml | 1 + 2 files changed, 2 insertions(+), 1 deletion(-) diff --git a/tests/test_runner/dune b/tests/test_runner/dune index e8b29d66..1c719532 100644 --- a/tests/test_runner/dune +++ b/tests/test_runner/dune @@ -1,6 +1,6 @@ (executable (public_name test_runner) - (libraries core_unix.sys_unix re unix) + (libraries core_unix.sys_unix re str unix) (preprocess (pps ppx_deriving.show ppx_deriving.ord ppx_sexp_conv)) (name run_test)) diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index 25efbcfd..3bda4e29 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -167,6 +167,7 @@ module Input = struct (* Given a path to a rust file or crate, gather the details and options about how to build the test. *) let build (path : string) : t = let name = Filename.remove_extension (Filename.basename path) in + let name = Str.global_replace (Str.regexp "-") "_" name in let kind = if Sys_unix.is_file_exn path then SingleFile else if Sys_unix.is_directory_exn path then Crate -- cgit v1.2.3 From ff5c20b1ef1309f3752eaf0ffd2789514a30589d Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Fri, 24 May 2024 16:09:05 +0200 Subject: runner: Allow per-backend skipping --- tests/test_runner/run_test.ml | 65 ++++++++++++++++++++++++++++++------------- 1 file changed, 46 insertions(+), 19 deletions(-) diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index 3bda4e29..71e2c32f 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -111,7 +111,7 @@ module Input = struct name : string; path : string; kind : kind; - action : action; + actions : action BackendMap.t; charon_options : string list; aeneas_options : string list BackendMap.t; subdir : string BackendMap.t; @@ -124,7 +124,7 @@ module Input = struct - `charon-args=...`: extra arguments to pass to charon; - `aeneas-args=...`: extra arguments to pass to aeneas; - `[backend,..]...`: where each `backend` is the name of a backend supported by - aeneas; this sets options for these backends only. Only supported for `aeneas-args`. + aeneas; this sets options for these backends only. *) let apply_special_comment comment input = let comment = String.trim comment in @@ -144,12 +144,28 @@ module Input = struct let charon_args = Core.String.chop_prefix comment ~prefix:"charon-args=" in let aeneas_args = Core.String.chop_prefix comment ~prefix:"aeneas-args=" in - if comment = "skip" then { input with action = Skip } - else if comment = "known-failure" then { input with action = KnownFailure } + if comment = "skip" then + { + input with + actions = + List.fold_left + (fun map backend -> BackendMap.add backend Skip map) + input.actions backends; + } + else if comment = "known-failure" then + { + input with + actions = + List.fold_left + (fun map backend -> BackendMap.add backend KnownFailure map) + input.actions backends; + } else if Option.is_some charon_args then let args = Option.get charon_args in let args = String.split_on_char ' ' args in - { input with charon_options = List.append input.charon_options args } + if backends != Backend.all then + failwith "Cannot set per-backend charon-args" + else { input with charon_options = List.append input.charon_options args } else if Option.is_some aeneas_args then let args = Option.get aeneas_args in let args = String.split_on_char ' ' args in @@ -173,7 +189,11 @@ module Input = struct else if Sys_unix.is_directory_exn path then Crate else failwith ("`" ^ path ^ "` is not a file or a directory.") in - let action = Normal in + let actions = + List.fold_left + (fun map backend -> BackendMap.add backend Normal map) + BackendMap.empty Backend.all + in let charon_options = charon_options_for_test name in let subdir = List.fold_left @@ -190,7 +210,7 @@ module Input = struct BackendMap.empty Backend.all in let input = - { path; name; kind; action; charon_options; subdir; aeneas_options } + { path; name; kind; actions; charon_options; subdir; aeneas_options } in match kind with | SingleFile -> @@ -267,7 +287,7 @@ let () = match Array.to_list Sys.argv with (* Ad-hoc argument passing for now. *) | _exe_path :: charon_path :: aeneas_path :: llbc_dir :: test_path - :: aeneas_options -> ( + :: aeneas_options -> let runner_env = { charon_path; aeneas_path; llbc_dir } in let test_case = Input.build test_path in let test_case = @@ -277,15 +297,22 @@ let () = BackendMap.map (List.append aeneas_options) test_case.aeneas_options; } in - - match test_case.action with - | Skip -> () - | Normal -> - (* Generate the llbc file *) - run_charon runner_env test_case; - (* Process the llbc file for the each backend *) - List.iter - (fun backend -> run_aeneas runner_env test_case backend) - Backend.all - | KnownFailure -> failwith "KnownFailure is unimplemented") + let skip_all = + List.for_all + (fun backend -> + BackendMap.find backend test_case.actions = Input.Skip) + Backend.all + in + if skip_all then () + else ( + (* Generate the llbc file *) + run_charon runner_env test_case; + (* Process the llbc file for the each backend *) + List.iter + (fun backend -> + match BackendMap.find backend test_case.actions with + | Skip -> () + | Normal -> run_aeneas runner_env test_case backend + | KnownFailure -> failwith "KnownFailure is unimplemented") + Backend.all) | _ -> failwith "Incorrect options passed to test runner" -- cgit v1.2.3 From 3b4139de5b9bcf5f43213330652ad01383e975d3 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Fri, 24 May 2024 16:50:18 +0200 Subject: Factor out the code for Lean and Coq --- compiler/ExtractTypes.ml | 184 +++++++++++------------------------------------ 1 file changed, 41 insertions(+), 143 deletions(-) diff --git a/compiler/ExtractTypes.ml b/compiler/ExtractTypes.ml index 2fc0c117..cc0c351d 100644 --- a/compiler/ExtractTypes.ml +++ b/compiler/ExtractTypes.ml @@ -1666,15 +1666,15 @@ let extract_type_decl_coq_arguments (ctx : extraction_ctx) (fmt : F.formatter) (** Auxiliary function. - Generate field projectors in Lean. + Generate field projectors for Lean and Coq. - Recursive structs are defined as inductives in Lean. + Recursive structs are defined as inductives in Lean and Coq. Field projectors allow to retrieve the facilities provided by Lean structures. *) -let extract_type_decl_lean_record_field_projectors (ctx : extraction_ctx) +let extract_type_decl_record_field_projectors (ctx : extraction_ctx) (fmt : F.formatter) (kind : decl_kind) (decl : type_decl) : unit = - sanity_check __FILE__ __LINE__ (!backend = Lean) decl.meta; + sanity_check __FILE__ __LINE__ (!backend = Coq || !backend = Lean) decl.span; match decl.kind with | Opaque | Enum _ -> () | Struct fields -> @@ -1683,18 +1683,18 @@ let extract_type_decl_lean_record_field_projectors (ctx : extraction_ctx) if is_rec then (* Add the type params *) let ctx, type_params, cg_params, trait_clauses = - ctx_add_generic_params decl.meta decl.llbc_name decl.llbc_generics + ctx_add_generic_params decl.span decl.llbc_name decl.llbc_generics decl.generics ctx in (* Record_var will be the ADT argument to the projector *) - let ctx, record_var = ctx_add_var decl.meta "x" (VarId.of_int 0) ctx in + let ctx, record_var = ctx_add_var decl.span "x" (VarId.of_int 0) ctx in (* Field_var will be the variable in the constructor that is returned by the projector *) - let ctx, field_var = ctx_add_var decl.meta "x" (VarId.of_int 1) ctx in + let ctx, field_var = ctx_add_var decl.span "x" (VarId.of_int 1) ctx in (* Name of the ADT *) - let def_name = ctx_get_local_type decl.meta decl.def_id ctx in + let def_name = ctx_get_local_type decl.span decl.def_id ctx in (* Name of the ADT constructor. As we are in the struct case, we only have one constructor *) - let cons_name = ctx_get_struct decl.meta (TAdtId decl.def_id) ctx in + let cons_name = ctx_get_struct decl.span (TAdtId decl.def_id) ctx in let extract_field_proj (field_id : FieldId.id) (_ : field) : unit = F.pp_print_space fmt (); @@ -1703,33 +1703,40 @@ let extract_type_decl_lean_record_field_projectors (ctx : extraction_ctx) (* Inner box for the projector definition *) F.pp_open_hvbox fmt ctx.indent_incr; - (* Box for the attributes *) - F.pp_open_vbox fmt 0; - (* Annotate the projectors with both simp and reducible. - The first one allows to automatically unfold when calling simp in proofs. - The second ensures that projectors will interact well with the unifier *) - F.pp_print_string fmt "@[simp, reducible]"; - F.pp_print_break fmt 0 0; - (* Close box for the attributes *) - F.pp_close_box fmt (); + (* For Lean: add some attributes *) + if !backend = Lean then ( + (* Box for the attributes *) + F.pp_open_vbox fmt 0; + (* Annotate the projectors with both simp and reducible. + The first one allows to automatically unfold when calling simp in proofs. + The second ensures that projectors will interact well with the unifier *) + F.pp_print_string fmt "@[simp, reducible]"; + F.pp_print_break fmt 0 0; + (* Close box for the attributes *) + F.pp_close_box fmt ()); (* Box for the [def ADT.proj ... :=] *) F.pp_open_hovbox fmt ctx.indent_incr; - F.pp_print_string fmt "def"; + (match !backend with + | Lean -> F.pp_print_string fmt "def" + | Coq -> F.pp_print_string fmt "Definition" + | _ -> internal_error __FILE__ __LINE__ decl.span); F.pp_print_space fmt (); (* Print the function name. In Lean, the syntax ADT.proj will - allow us to call x.proj for any x of type ADT *) + allow us to call x.proj for any x of type ADT. In Coq, + we will have to introduce a notation for the projector. *) let field_name = - ctx_get_field decl.meta (TAdtId decl.def_id) field_id ctx + ctx_get_field decl.span (TAdtId decl.def_id) field_id ctx in - F.pp_print_string fmt def_name; - F.pp_print_string fmt "."; + if !backend = Lean then ( + F.pp_print_string fmt def_name; + F.pp_print_string fmt "."); F.pp_print_string fmt field_name; (* Print the generics *) let as_implicits = true in - extract_generic_params decl.meta ctx fmt TypeDeclId.Set.empty + extract_generic_params decl.span ctx fmt TypeDeclId.Set.empty ~as_implicits decl.generics type_params cg_params trait_clauses; (* Print the record parameter as "(x : ADT)" *) @@ -1787,115 +1794,11 @@ let extract_type_decl_lean_record_field_projectors (ctx : extraction_ctx) (* Close the box for the branch *) F.pp_close_box fmt (); - (* Close the box for the whole match *) - F.pp_close_box fmt (); - - F.pp_close_box fmt (); - (* Close the outer box for projector definition *) - F.pp_close_box fmt (); - (* Add breaks to insert new lines between definitions *) - F.pp_print_break fmt 0 0 - in - - FieldId.iteri extract_field_proj fields - -(** Auxiliary function. - - Generate field projectors in Coq. - - Sometimes we extract records as inductives in Coq: when this happens we - have to define the field projectors afterwards. - *) -let extract_type_decl_coq_record_field_projectors (ctx : extraction_ctx) - (fmt : F.formatter) (kind : decl_kind) (decl : type_decl) : unit = - sanity_check __FILE__ __LINE__ (!backend = Coq) decl.span; - match decl.kind with - | Opaque | Enum _ -> () - | Struct fields -> - (* Records are extracted as inductives only if they are recursive *) - let is_rec = decl_is_from_rec_group kind in - if is_rec then - (* Add the type params *) - let ctx, type_params, cg_params, trait_clauses = - ctx_add_generic_params decl.span decl.llbc_name decl.llbc_generics - decl.generics ctx - in - let ctx, record_var = ctx_add_var decl.span "x" (VarId.of_int 0) ctx in - let ctx, field_var = ctx_add_var decl.span "x" (VarId.of_int 1) ctx in - let def_name = ctx_get_local_type decl.span decl.def_id ctx in - let cons_name = ctx_get_struct decl.span (TAdtId decl.def_id) ctx in - let extract_field_proj (field_id : FieldId.id) (_ : field) : unit = - F.pp_print_space fmt (); - (* Outer box for the projector definition *) - F.pp_open_hvbox fmt 0; - (* Inner box for the projector definition *) - F.pp_open_hvbox fmt ctx.indent_incr; - (* Open a box for the [Definition PROJ ... :=] *) - F.pp_open_hovbox fmt ctx.indent_incr; - F.pp_print_string fmt "Definition"; - F.pp_print_space fmt (); - let field_name = - ctx_get_field decl.span (TAdtId decl.def_id) field_id ctx - in - F.pp_print_string fmt field_name; - (* Print the generics *) - let as_implicits = true in - extract_generic_params decl.span ctx fmt TypeDeclId.Set.empty - ~as_implicits decl.generics type_params cg_params trait_clauses; - (* Print the record parameter *) - F.pp_print_space fmt (); - F.pp_print_string fmt "("; - F.pp_print_string fmt record_var; - F.pp_print_space fmt (); - F.pp_print_string fmt ":"; - F.pp_print_space fmt (); - F.pp_print_string fmt def_name; - List.iter - (fun p -> - F.pp_print_space fmt (); - F.pp_print_string fmt p) - type_params; - F.pp_print_string fmt ")"; - (* *) - F.pp_print_space fmt (); - F.pp_print_string fmt ":="; - (* Close the box for the [Definition PROJ ... :=] *) - F.pp_close_box fmt (); - F.pp_print_space fmt (); - (* Open a box for the whole match *) - F.pp_open_hvbox fmt 0; - (* Open a box for the [match ... with] *) - F.pp_open_hovbox fmt ctx.indent_incr; - F.pp_print_string fmt "match"; - F.pp_print_space fmt (); - F.pp_print_string fmt record_var; - F.pp_print_space fmt (); - F.pp_print_string fmt "with"; - (* Close the box for the [match ... with] *) - F.pp_close_box fmt (); - - (* Open a box for the branch *) - F.pp_open_hovbox fmt ctx.indent_incr; - (* Print the match branch *) - F.pp_print_space fmt (); - F.pp_print_string fmt "|"; - F.pp_print_space fmt (); - F.pp_print_string fmt cons_name; - FieldId.iteri - (fun id _ -> - F.pp_print_space fmt (); - if field_id = id then F.pp_print_string fmt field_var - else F.pp_print_string fmt "_") - fields; - F.pp_print_space fmt (); - F.pp_print_string fmt "=>"; - F.pp_print_space fmt (); - F.pp_print_string fmt field_var; - (* Close the box for the branch *) - F.pp_close_box fmt (); (* Print the [end] *) - F.pp_print_space fmt (); - F.pp_print_string fmt "end"; + if !backend = Coq then ( + F.pp_print_space fmt (); + F.pp_print_string fmt "end"); + (* Close the box for the whole match *) F.pp_close_box fmt (); (* Close the inner box projector *) @@ -1904,12 +1807,13 @@ let extract_type_decl_coq_record_field_projectors (ctx : extraction_ctx) if !backend = Coq then ( F.pp_print_cut fmt (); F.pp_print_string fmt "."); - (* Close the outer box projector *) + (* Close the outer box for projector definition *) F.pp_close_box fmt (); (* Add breaks to insert new lines between definitions *) F.pp_print_break fmt 0 0 in + (* Only for Coq: we need to define a notation for the projector *) let extract_proj_notation (field_id : FieldId.id) (_ : field) : unit = F.pp_print_space fmt (); (* Outer box for the projector definition *) @@ -1950,7 +1854,7 @@ let extract_type_decl_coq_record_field_projectors (ctx : extraction_ctx) let extract_field_proj_and_notation (field_id : FieldId.id) (field : field) : unit = extract_field_proj field_id field; - extract_proj_notation field_id field + if !backend = Coq then extract_proj_notation field_id field in FieldId.iteri extract_field_proj_and_notation fields @@ -1964,20 +1868,14 @@ let extract_type_decl_extra_info (ctx : extraction_ctx) (fmt : F.formatter) (kind : decl_kind) (decl : type_decl) : unit = match !backend with | FStar | HOL4 -> () - | Lean -> - if - not - (TypesUtils.type_decl_from_decl_id_is_tuple_struct - ctx.trans_ctx.type_ctx.type_infos decl.def_id) - then extract_type_decl_lean_record_field_projectors ctx fmt kind decl - | Coq -> + | Lean | Coq -> if not (TypesUtils.type_decl_from_decl_id_is_tuple_struct ctx.trans_ctx.type_ctx.type_infos decl.def_id) then ( - extract_type_decl_coq_arguments ctx fmt kind decl; - extract_type_decl_coq_record_field_projectors ctx fmt kind decl) + if !backend = Coq then extract_type_decl_coq_arguments ctx fmt kind decl; + extract_type_decl_record_field_projectors ctx fmt kind decl) (** Extract the state type declaration. *) let extract_state_type (fmt : F.formatter) (ctx : extraction_ctx) -- cgit v1.2.3 From 169af47945f013e61b14d67e7ebdc9c03636c5a2 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Fri, 24 May 2024 16:50:25 +0200 Subject: Update an .opam file --- tests/test_runner/aeneas_test_runner.opam | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/tests/test_runner/aeneas_test_runner.opam b/tests/test_runner/aeneas_test_runner.opam index b57cc9f6..1539c521 100644 --- a/tests/test_runner/aeneas_test_runner.opam +++ b/tests/test_runner/aeneas_test_runner.opam @@ -7,7 +7,7 @@ homepage: "https://github.com/AeneasVerif/aeneas" bug-reports: "https://github.com/AeneasVerif/aeneas/issues" depends: [ "ocaml" - "dune" {>= "3.12"} + "dune" {>= "3.7"} "odoc" {with-doc} ] build: [ -- cgit v1.2.3 From c1892d14fc9ec728422f05dbb9de13ee92984734 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Fri, 24 May 2024 16:45:45 +0200 Subject: runner: Factor out useful `BackendMap` operations --- tests/test_runner/run_test.ml | 59 +++++++++++++++++++------------------------ 1 file changed, 26 insertions(+), 33 deletions(-) diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index 71e2c32f..39330a77 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -37,7 +37,26 @@ module Backend = struct | HOL4 -> "hol4" end -module BackendMap = Map.Make (Backend) +module BackendMap = struct + include Map.Make (Backend) + + (* Make a new map with one entry per backend, given by `f` *) + let make (f : Backend.t -> 'a) : 'a t = + List.fold_left + (fun map backend -> add backend (f backend) map) + empty Backend.all + + (* Set this value for all the backends in `backends` *) + let add_each (backends : Backend.t list) (v : 'a) (map : 'a t) : 'a t = + List.fold_left (fun map backend -> add backend v map) map backends + + (* Updates all the backends in `backends` with `f` *) + let update_each (backends : Backend.t list) (f : 'a -> 'a) (map : 'a t) : 'a t + = + List.fold_left + (fun map backend -> update backend (Option.map f) map) + map backends +end let concat_path = List.fold_left Filename.concat "" @@ -145,20 +164,11 @@ module Input = struct let aeneas_args = Core.String.chop_prefix comment ~prefix:"aeneas-args=" in if comment = "skip" then - { - input with - actions = - List.fold_left - (fun map backend -> BackendMap.add backend Skip map) - input.actions backends; - } + { input with actions = BackendMap.add_each backends Skip input.actions } else if comment = "known-failure" then { input with - actions = - List.fold_left - (fun map backend -> BackendMap.add backend KnownFailure map) - input.actions backends; + actions = BackendMap.add_each backends KnownFailure input.actions; } else if Option.is_some charon_args then let args = Option.get charon_args in @@ -173,10 +183,7 @@ module Input = struct { input with aeneas_options = - List.fold_left - (fun map backend -> - BackendMap.update backend (Option.map add_args) map) - input.aeneas_options backends; + BackendMap.update_each backends add_args input.aeneas_options; } else failwith ("Unrecognized special comment: `" ^ comment ^ "`") @@ -189,25 +196,11 @@ module Input = struct else if Sys_unix.is_directory_exn path then Crate else failwith ("`" ^ path ^ "` is not a file or a directory.") in - let actions = - List.fold_left - (fun map backend -> BackendMap.add backend Normal map) - BackendMap.empty Backend.all - in + let actions = BackendMap.make (fun _ -> Normal) in let charon_options = charon_options_for_test name in - let subdir = - List.fold_left - (fun map backend -> - let subdir = test_subdir backend name in - BackendMap.add backend subdir map) - BackendMap.empty Backend.all - in + let subdir = BackendMap.make (fun backend -> test_subdir backend name) in let aeneas_options = - List.fold_left - (fun map backend -> - let opts = aeneas_options_for_test backend name in - BackendMap.add backend opts map) - BackendMap.empty Backend.all + BackendMap.make (fun backend -> aeneas_options_for_test backend name) in let input = { path; name; kind; actions; charon_options; subdir; aeneas_options } -- cgit v1.2.3 From 093ebced6d22f6b805147783c978af13a7a03caa Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Fri, 24 May 2024 16:51:54 +0200 Subject: runner: Specify subdirectory in magic comments --- tests/src/bitwise.rs | 1 + tests/src/constants.rs | 1 + tests/src/demo.rs | 1 + tests/src/external.rs | 1 + tests/src/loops.rs | 1 + tests/src/no_nested_borrows.rs | 1 + tests/src/paper.rs | 1 + tests/src/polonius_list.rs | 1 + tests/test_runner/run_test.ml | 31 ++++++++++++++++--------------- 9 files changed, 24 insertions(+), 15 deletions(-) diff --git a/tests/src/bitwise.rs b/tests/src/bitwise.rs index 15962047..fda8eff3 100644 --- a/tests/src/bitwise.rs +++ b/tests/src/bitwise.rs @@ -1,4 +1,5 @@ //@ aeneas-args=-test-trans-units +//@ [coq,fstar] subdir=misc //! Exercise the bitwise operations pub fn shift_u32(a: u32) -> u32 { diff --git a/tests/src/constants.rs b/tests/src/constants.rs index 925c62b1..ac24dcd4 100644 --- a/tests/src/constants.rs +++ b/tests/src/constants.rs @@ -1,5 +1,6 @@ //@ charon-args=--no-code-duplication //@ aeneas-args=-test-trans-units +//@ [coq,fstar] subdir=misc //! Tests with constants // Integers diff --git a/tests/src/demo.rs b/tests/src/demo.rs index b9bb7ca2..0a589338 100644 --- a/tests/src/demo.rs +++ b/tests/src/demo.rs @@ -1,4 +1,5 @@ //@ [coq,fstar] aeneas-args=-use-fuel +//@ [lean] subdir=Demo #![allow(clippy::needless_lifetimes)] /* Simple functions */ diff --git a/tests/src/external.rs b/tests/src/external.rs index ddd5539f..baea76e4 100644 --- a/tests/src/external.rs +++ b/tests/src/external.rs @@ -1,6 +1,7 @@ //@ charon-args=--no-code-duplication //@ aeneas-args=-state -split-files //@ aeneas-args=-test-trans-units +//@ [coq,fstar] subdir=misc //! This module uses external types and functions use std::cell::Cell; diff --git a/tests/src/loops.rs b/tests/src/loops.rs index 8692c60e..afc52ace 100644 --- a/tests/src/loops.rs +++ b/tests/src/loops.rs @@ -1,6 +1,7 @@ //@ [coq] aeneas-args=-use-fuel //@ [fstar] aeneas-args=-decreases-clauses -template-clauses //@ [fstar] aeneas-args=-split-files +//@ [coq,fstar] subdir=misc use std::vec::Vec; /// No borrows diff --git a/tests/src/no_nested_borrows.rs b/tests/src/no_nested_borrows.rs index 78163371..a250748c 100644 --- a/tests/src/no_nested_borrows.rs +++ b/tests/src/no_nested_borrows.rs @@ -1,5 +1,6 @@ //@ charon-args=--no-code-duplication //@ aeneas-args=-test-trans-units +//@ [coq,fstar] subdir=misc //! This module doesn't contain **functions which use nested borrows in their //! signatures**, and doesn't contain functions with loops. diff --git a/tests/src/paper.rs b/tests/src/paper.rs index 07453098..6a4a7c31 100644 --- a/tests/src/paper.rs +++ b/tests/src/paper.rs @@ -1,5 +1,6 @@ //@ charon-args=--no-code-duplication //@ aeneas-args=-test-trans-units +//@ [coq,fstar] subdir=misc //! The examples from the ICFP submission, all in one place. // 2.1 diff --git a/tests/src/polonius_list.rs b/tests/src/polonius_list.rs index a8d51e40..b029ad02 100644 --- a/tests/src/polonius_list.rs +++ b/tests/src/polonius_list.rs @@ -1,5 +1,6 @@ //@ charon-args=--polonius //@ aeneas-args=-test-trans-units +//@ [coq,fstar] subdir=misc #![allow(dead_code)] pub enum List { diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index 39330a77..8f356365 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -104,21 +104,16 @@ let charon_options_for_test test_name = [ "--polonius"; "--opaque=betree_utils"; "--crate"; "betree_main" ] | _ -> [] -(* The subdirectory in which to store the outputs. *) +(* The default subdirectory in which to store the outputs. *) (* This reproduces the file layout that was set by the old Makefile. FIXME: cleanup *) let test_subdir backend test_name = - let backend = Backend.to_string backend in match (backend, test_name) with - | "lean", "demo" -> "Demo" - | "lean", _ -> "." - | _, ("arrays" | "demo" | "hashmap" | "traits") -> test_name - | _, "betree" -> "betree" + | Backend.Lean, _ -> "." | _, "hashmap_main" -> "hashmap_on_disk" - | "hol4", _ -> "misc-" ^ test_name - | ( _, - ( "bitwise" | "constants" | "external" | "loops" | "no_nested_borrows" - | "paper" | "polonius_list" ) ) -> - "misc" + | ( Backend.HOL4, + ( "constants" | "external" | "loops" | "no_nested_borrows" | "paper" + | "polonius_list" ) ) -> + "misc-" ^ test_name | _ -> test_name (* The data for a specific test input *) @@ -133,13 +128,15 @@ module Input = struct actions : action BackendMap.t; charon_options : string list; aeneas_options : string list BackendMap.t; - subdir : string BackendMap.t; + subdirs : string BackendMap.t; } (* Parse lines that start `//@`. Each of them modifies the options we use for the test. Supported comments: - `skip`: don't process the file; - `known-failure`: TODO; + - `subdir=...: set the subdirectory in which to store the outputs. + Defaults to nothing for lean and to the test name for other backends; - `charon-args=...`: extra arguments to pass to charon; - `aeneas-args=...`: extra arguments to pass to aeneas; - `[backend,..]...`: where each `backend` is the name of a backend supported by @@ -162,6 +159,7 @@ module Input = struct (* Parse the other options *) let charon_args = Core.String.chop_prefix comment ~prefix:"charon-args=" in let aeneas_args = Core.String.chop_prefix comment ~prefix:"aeneas-args=" in + let subdir = Core.String.chop_prefix comment ~prefix:"subdir=" in if comment = "skip" then { input with actions = BackendMap.add_each backends Skip input.actions } @@ -185,6 +183,9 @@ module Input = struct aeneas_options = BackendMap.update_each backends add_args input.aeneas_options; } + else if Option.is_some subdir then + let subdir = Option.get subdir in + { input with subdirs = BackendMap.add_each backends subdir input.subdirs } else failwith ("Unrecognized special comment: `" ^ comment ^ "`") (* Given a path to a rust file or crate, gather the details and options about how to build the test. *) @@ -198,12 +199,12 @@ module Input = struct in let actions = BackendMap.make (fun _ -> Normal) in let charon_options = charon_options_for_test name in - let subdir = BackendMap.make (fun backend -> test_subdir backend name) in + let subdirs = BackendMap.make (fun backend -> test_subdir backend name) in let aeneas_options = BackendMap.make (fun backend -> aeneas_options_for_test backend name) in let input = - { path; name; kind; actions; charon_options; subdir; aeneas_options } + { path; name; kind; actions; charon_options; subdirs; aeneas_options } in match kind with | SingleFile -> @@ -226,7 +227,7 @@ let run_aeneas (env : runner_env) (case : Input.t) (backend : Backend.t) = (* FIXME: remove this special case *) let test_name = if case.name = "betree" then "betree_main" else case.name in let input_file = concat_path [ env.llbc_dir; test_name ] ^ ".llbc" in - let subdir = BackendMap.find backend case.subdir in + let subdir = BackendMap.find backend case.subdirs in let aeneas_options = BackendMap.find backend case.aeneas_options in let backend_str = Backend.to_string backend in let dest_dir = concat_path [ "tests"; backend_str; subdir ] in -- cgit v1.2.3 From e288482f437a5f259be5f81eb996b5b28158b300 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Fri, 24 May 2024 16:51:03 +0200 Subject: Update output files --- tests/coq/demo/Demo.v | 32 +++---- tests/coq/misc/Bitwise.v | 10 +-- tests/coq/misc/Constants.v | 62 +++++++------- tests/coq/misc/External_Funs.v | 4 +- tests/coq/misc/Loops.v | 98 +++++++++++----------- tests/coq/misc/NoNestedBorrows.v | 126 ++++++++++++++-------------- tests/coq/misc/Paper.v | 18 ++-- tests/coq/misc/PoloniusList.v | 4 +- tests/fstar/demo/Demo.fst | 32 +++---- tests/fstar/misc/Bitwise.fst | 10 +-- tests/fstar/misc/Constants.fst | 62 +++++++------- tests/fstar/misc/External.Funs.fst | 4 +- tests/fstar/misc/Loops.Clauses.Template.fst | 46 +++++----- tests/fstar/misc/Loops.Funs.fst | 96 ++++++++++----------- tests/fstar/misc/Loops.Types.fst | 2 +- tests/fstar/misc/NoNestedBorrows.fst | 126 ++++++++++++++-------------- tests/fstar/misc/Paper.fst | 18 ++-- tests/fstar/misc/PoloniusList.fst | 4 +- tests/lean/Bitwise.lean | 10 +-- tests/lean/Constants.lean | 62 +++++++------- tests/lean/Demo/Demo.lean | 32 +++---- tests/lean/External/Funs.lean | 4 +- tests/lean/Loops.lean | 98 +++++++++++----------- tests/lean/NoNestedBorrows.lean | 126 ++++++++++++++-------------- tests/lean/Paper.lean | 18 ++-- tests/lean/PoloniusList.lean | 4 +- 26 files changed, 554 insertions(+), 554 deletions(-) diff --git a/tests/coq/demo/Demo.v b/tests/coq/demo/Demo.v index 1cccbeda..e8c3a694 100644 --- a/tests/coq/demo/Demo.v +++ b/tests/coq/demo/Demo.v @@ -9,7 +9,7 @@ Local Open Scope Primitives_scope. Module Demo. (** [demo::choose]: - Source: 'tests/src/demo.rs', lines 6:0-6:70 *) + Source: 'tests/src/demo.rs', lines 7:0-7:70 *) Definition choose (T : Type) (b : bool) (x : T) (y : T) : result (T * (T -> result (T * T))) := if b @@ -18,30 +18,30 @@ Definition choose . (** [demo::mul2_add1]: - Source: 'tests/src/demo.rs', lines 14:0-14:31 *) + Source: 'tests/src/demo.rs', lines 15:0-15:31 *) Definition mul2_add1 (x : u32) : result u32 := i <- u32_add x x; u32_add i 1%u32 . (** [demo::use_mul2_add1]: - Source: 'tests/src/demo.rs', lines 18:0-18:43 *) + Source: 'tests/src/demo.rs', lines 19:0-19:43 *) Definition use_mul2_add1 (x : u32) (y : u32) : result u32 := i <- mul2_add1 x; u32_add i y . (** [demo::incr]: - Source: 'tests/src/demo.rs', lines 22:0-22:31 *) + Source: 'tests/src/demo.rs', lines 23:0-23:31 *) Definition incr (x : u32) : result u32 := u32_add x 1%u32. (** [demo::use_incr]: - Source: 'tests/src/demo.rs', lines 26:0-26:17 *) + Source: 'tests/src/demo.rs', lines 27:0-27:17 *) Definition use_incr : result unit := x <- incr 0%u32; x1 <- incr x; _ <- incr x1; Ok tt . (** [demo::CList] - Source: 'tests/src/demo.rs', lines 35:0-35:17 *) + Source: 'tests/src/demo.rs', lines 36:0-36:17 *) Inductive CList_t (T : Type) := | CList_CCons : T -> CList_t T -> CList_t T | CList_CNil : CList_t T @@ -51,7 +51,7 @@ Arguments CList_CCons { _ }. Arguments CList_CNil { _ }. (** [demo::list_nth]: - Source: 'tests/src/demo.rs', lines 40:0-40:56 *) + Source: 'tests/src/demo.rs', lines 41:0-41:56 *) Fixpoint list_nth (T : Type) (n : nat) (l : CList_t T) (i : u32) : result T := match n with | O => Fail_ OutOfFuel @@ -65,7 +65,7 @@ Fixpoint list_nth (T : Type) (n : nat) (l : CList_t T) (i : u32) : result T := . (** [demo::list_nth_mut]: - Source: 'tests/src/demo.rs', lines 55:0-55:68 *) + Source: 'tests/src/demo.rs', lines 56:0-56:68 *) Fixpoint list_nth_mut (T : Type) (n : nat) (l : CList_t T) (i : u32) : result (T * (T -> result (CList_t T))) @@ -91,7 +91,7 @@ Fixpoint list_nth_mut . (** [demo::list_nth_mut1]: loop 0: - Source: 'tests/src/demo.rs', lines 70:0-79:1 *) + Source: 'tests/src/demo.rs', lines 71:0-80:1 *) Fixpoint list_nth_mut1_loop (T : Type) (n : nat) (l : CList_t T) (i : u32) : result (T * (T -> result (CList_t T))) @@ -116,7 +116,7 @@ Fixpoint list_nth_mut1_loop . (** [demo::list_nth_mut1]: - Source: 'tests/src/demo.rs', lines 70:0-70:77 *) + Source: 'tests/src/demo.rs', lines 71:0-71:77 *) Definition list_nth_mut1 (T : Type) (n : nat) (l : CList_t T) (i : u32) : result (T * (T -> result (CList_t T))) @@ -125,7 +125,7 @@ Definition list_nth_mut1 . (** [demo::i32_id]: - Source: 'tests/src/demo.rs', lines 81:0-81:28 *) + Source: 'tests/src/demo.rs', lines 82:0-82:28 *) Fixpoint i32_id (n : nat) (i : i32) : result i32 := match n with | O => Fail_ OutOfFuel @@ -137,7 +137,7 @@ Fixpoint i32_id (n : nat) (i : i32) : result i32 := . (** [demo::list_tail]: - Source: 'tests/src/demo.rs', lines 89:0-89:64 *) + Source: 'tests/src/demo.rs', lines 90:0-90:64 *) Fixpoint list_tail (T : Type) (n : nat) (l : CList_t T) : result ((CList_t T) * (CList_t T -> result (CList_t T))) @@ -159,7 +159,7 @@ Fixpoint list_tail . (** Trait declaration: [demo::Counter] - Source: 'tests/src/demo.rs', lines 98:0-98:17 *) + Source: 'tests/src/demo.rs', lines 99:0-99:17 *) Record Counter_t (Self : Type) := mkCounter_t { Counter_t_incr : Self -> result (usize * Self); }. @@ -168,19 +168,19 @@ Arguments mkCounter_t { _ }. Arguments Counter_t_incr { _ }. (** [demo::{(demo::Counter for usize)}::incr]: - Source: 'tests/src/demo.rs', lines 103:4-103:31 *) + Source: 'tests/src/demo.rs', lines 104:4-104:31 *) Definition counterUsize_incr (self : usize) : result (usize * usize) := self1 <- usize_add self 1%usize; Ok (self, self1) . (** Trait implementation: [demo::{(demo::Counter for usize)}] - Source: 'tests/src/demo.rs', lines 102:0-102:22 *) + Source: 'tests/src/demo.rs', lines 103:0-103:22 *) Definition CounterUsize : Counter_t usize := {| Counter_t_incr := counterUsize_incr; |}. (** [demo::use_counter]: - Source: 'tests/src/demo.rs', lines 110:0-110:59 *) + Source: 'tests/src/demo.rs', lines 111:0-111:59 *) Definition use_counter (T : Type) (counterInst : Counter_t T) (cnt : T) : result (usize * T) := counterInst.(Counter_t_incr) cnt diff --git a/tests/coq/misc/Bitwise.v b/tests/coq/misc/Bitwise.v index d0dbfd51..610f4ea8 100644 --- a/tests/coq/misc/Bitwise.v +++ b/tests/coq/misc/Bitwise.v @@ -9,29 +9,29 @@ Local Open Scope Primitives_scope. Module Bitwise. (** [bitwise::shift_u32]: - Source: 'tests/src/bitwise.rs', lines 4:0-4:31 *) + Source: 'tests/src/bitwise.rs', lines 5:0-5:31 *) Definition shift_u32 (a : u32) : result u32 := t <- u32_shr a 16%usize; u32_shl t 16%usize . (** [bitwise::shift_i32]: - Source: 'tests/src/bitwise.rs', lines 11:0-11:31 *) + Source: 'tests/src/bitwise.rs', lines 12:0-12:31 *) Definition shift_i32 (a : i32) : result i32 := t <- i32_shr a 16%isize; i32_shl t 16%isize . (** [bitwise::xor_u32]: - Source: 'tests/src/bitwise.rs', lines 18:0-18:37 *) + Source: 'tests/src/bitwise.rs', lines 19:0-19:37 *) Definition xor_u32 (a : u32) (b : u32) : result u32 := Ok (u32_xor a b). (** [bitwise::or_u32]: - Source: 'tests/src/bitwise.rs', lines 22:0-22:36 *) + Source: 'tests/src/bitwise.rs', lines 23:0-23:36 *) Definition or_u32 (a : u32) (b : u32) : result u32 := Ok (u32_or a b). (** [bitwise::and_u32]: - Source: 'tests/src/bitwise.rs', lines 26:0-26:37 *) + Source: 'tests/src/bitwise.rs', lines 27:0-27:37 *) Definition and_u32 (a : u32) (b : u32) : result u32 := Ok (u32_and a b). diff --git a/tests/coq/misc/Constants.v b/tests/coq/misc/Constants.v index c3ecdb83..fb5f5a29 100644 --- a/tests/coq/misc/Constants.v +++ b/tests/coq/misc/Constants.v @@ -9,37 +9,37 @@ Local Open Scope Primitives_scope. Module Constants. (** [constants::X0] - Source: 'tests/src/constants.rs', lines 7:0-7:17 *) + Source: 'tests/src/constants.rs', lines 8:0-8:17 *) Definition x0_body : result u32 := Ok 0%u32. Definition x0 : u32 := x0_body%global. (** [constants::X1] - Source: 'tests/src/constants.rs', lines 9:0-9:17 *) + Source: 'tests/src/constants.rs', lines 10:0-10:17 *) Definition x1_body : result u32 := Ok core_u32_max. Definition x1 : u32 := x1_body%global. (** [constants::X2] - Source: 'tests/src/constants.rs', lines 12:0-12:17 *) + Source: 'tests/src/constants.rs', lines 13:0-13:17 *) Definition x2_body : result u32 := Ok 3%u32. Definition x2 : u32 := x2_body%global. (** [constants::incr]: - Source: 'tests/src/constants.rs', lines 19:0-19:32 *) + Source: 'tests/src/constants.rs', lines 20:0-20:32 *) Definition incr (n : u32) : result u32 := u32_add n 1%u32. (** [constants::X3] - Source: 'tests/src/constants.rs', lines 17:0-17:17 *) + Source: 'tests/src/constants.rs', lines 18:0-18:17 *) Definition x3_body : result u32 := incr 32%u32. Definition x3 : u32 := x3_body%global. (** [constants::mk_pair0]: - Source: 'tests/src/constants.rs', lines 25:0-25:51 *) + Source: 'tests/src/constants.rs', lines 26:0-26:51 *) Definition mk_pair0 (x : u32) (y1 : u32) : result (u32 * u32) := Ok (x, y1). (** [constants::Pair] - Source: 'tests/src/constants.rs', lines 38:0-38:23 *) + Source: 'tests/src/constants.rs', lines 39:0-39:23 *) Record Pair_t (T1 T2 : Type) := mkPair_t { pair_x : T1; pair_y : T2; }. Arguments mkPair_t { _ _ }. @@ -47,130 +47,130 @@ Arguments pair_x { _ _ }. Arguments pair_y { _ _ }. (** [constants::mk_pair1]: - Source: 'tests/src/constants.rs', lines 29:0-29:55 *) + Source: 'tests/src/constants.rs', lines 30:0-30:55 *) Definition mk_pair1 (x : u32) (y1 : u32) : result (Pair_t u32 u32) := Ok {| pair_x := x; pair_y := y1 |} . (** [constants::P0] - Source: 'tests/src/constants.rs', lines 33:0-33:24 *) + Source: 'tests/src/constants.rs', lines 34:0-34:24 *) Definition p0_body : result (u32 * u32) := mk_pair0 0%u32 1%u32. Definition p0 : (u32 * u32) := p0_body%global. (** [constants::P1] - Source: 'tests/src/constants.rs', lines 34:0-34:28 *) + Source: 'tests/src/constants.rs', lines 35:0-35:28 *) Definition p1_body : result (Pair_t u32 u32) := mk_pair1 0%u32 1%u32. Definition p1 : Pair_t u32 u32 := p1_body%global. (** [constants::P2] - Source: 'tests/src/constants.rs', lines 35:0-35:24 *) + Source: 'tests/src/constants.rs', lines 36:0-36:24 *) Definition p2_body : result (u32 * u32) := Ok (0%u32, 1%u32). Definition p2 : (u32 * u32) := p2_body%global. (** [constants::P3] - Source: 'tests/src/constants.rs', lines 36:0-36:28 *) + Source: 'tests/src/constants.rs', lines 37:0-37:28 *) Definition p3_body : result (Pair_t u32 u32) := Ok {| pair_x := 0%u32; pair_y := 1%u32 |} . Definition p3 : Pair_t u32 u32 := p3_body%global. (** [constants::Wrap] - Source: 'tests/src/constants.rs', lines 51:0-51:18 *) + Source: 'tests/src/constants.rs', lines 52:0-52:18 *) Record Wrap_t (T : Type) := mkWrap_t { wrap_value : T; }. Arguments mkWrap_t { _ }. Arguments wrap_value { _ }. (** [constants::{constants::Wrap}::new]: - Source: 'tests/src/constants.rs', lines 56:4-56:41 *) + Source: 'tests/src/constants.rs', lines 57:4-57:41 *) Definition wrap_new (T : Type) (value : T) : result (Wrap_t T) := Ok {| wrap_value := value |} . (** [constants::Y] - Source: 'tests/src/constants.rs', lines 43:0-43:22 *) + Source: 'tests/src/constants.rs', lines 44:0-44:22 *) Definition y_body : result (Wrap_t i32) := wrap_new i32 2%i32. Definition y : Wrap_t i32 := y_body%global. (** [constants::unwrap_y]: - Source: 'tests/src/constants.rs', lines 45:0-45:30 *) + Source: 'tests/src/constants.rs', lines 46:0-46:30 *) Definition unwrap_y : result i32 := Ok y.(wrap_value). (** [constants::YVAL] - Source: 'tests/src/constants.rs', lines 49:0-49:19 *) + Source: 'tests/src/constants.rs', lines 50:0-50:19 *) Definition yval_body : result i32 := unwrap_y. Definition yval : i32 := yval_body%global. (** [constants::get_z1::Z1] - Source: 'tests/src/constants.rs', lines 64:4-64:17 *) + Source: 'tests/src/constants.rs', lines 65:4-65:17 *) Definition get_z1_z1_body : result i32 := Ok 3%i32. Definition get_z1_z1 : i32 := get_z1_z1_body%global. (** [constants::get_z1]: - Source: 'tests/src/constants.rs', lines 63:0-63:28 *) + Source: 'tests/src/constants.rs', lines 64:0-64:28 *) Definition get_z1 : result i32 := Ok get_z1_z1. (** [constants::add]: - Source: 'tests/src/constants.rs', lines 68:0-68:39 *) + Source: 'tests/src/constants.rs', lines 69:0-69:39 *) Definition add (a : i32) (b : i32) : result i32 := i32_add a b. (** [constants::Q1] - Source: 'tests/src/constants.rs', lines 76:0-76:17 *) + Source: 'tests/src/constants.rs', lines 77:0-77:17 *) Definition q1_body : result i32 := Ok 5%i32. Definition q1 : i32 := q1_body%global. (** [constants::Q2] - Source: 'tests/src/constants.rs', lines 77:0-77:17 *) + Source: 'tests/src/constants.rs', lines 78:0-78:17 *) Definition q2_body : result i32 := Ok q1. Definition q2 : i32 := q2_body%global. (** [constants::Q3] - Source: 'tests/src/constants.rs', lines 78:0-78:17 *) + Source: 'tests/src/constants.rs', lines 79:0-79:17 *) Definition q3_body : result i32 := add q2 3%i32. Definition q3 : i32 := q3_body%global. (** [constants::get_z2]: - Source: 'tests/src/constants.rs', lines 72:0-72:28 *) + Source: 'tests/src/constants.rs', lines 73:0-73:28 *) Definition get_z2 : result i32 := i <- get_z1; i1 <- add i q3; add q1 i1. (** [constants::S1] - Source: 'tests/src/constants.rs', lines 82:0-82:18 *) + Source: 'tests/src/constants.rs', lines 83:0-83:18 *) Definition s1_body : result u32 := Ok 6%u32. Definition s1 : u32 := s1_body%global. (** [constants::S2] - Source: 'tests/src/constants.rs', lines 83:0-83:18 *) + Source: 'tests/src/constants.rs', lines 84:0-84:18 *) Definition s2_body : result u32 := incr s1. Definition s2 : u32 := s2_body%global. (** [constants::S3] - Source: 'tests/src/constants.rs', lines 84:0-84:29 *) + Source: 'tests/src/constants.rs', lines 85:0-85:29 *) Definition s3_body : result (Pair_t u32 u32) := Ok p3. Definition s3 : Pair_t u32 u32 := s3_body%global. (** [constants::S4] - Source: 'tests/src/constants.rs', lines 85:0-85:29 *) + Source: 'tests/src/constants.rs', lines 86:0-86:29 *) Definition s4_body : result (Pair_t u32 u32) := mk_pair1 7%u32 8%u32. Definition s4 : Pair_t u32 u32 := s4_body%global. (** [constants::V] - Source: 'tests/src/constants.rs', lines 88:0-88:31 *) + Source: 'tests/src/constants.rs', lines 89:0-89:31 *) Record V_t (T : Type) (N : usize) := mkV_t { v_x : array T N; }. Arguments mkV_t { _ _ }. Arguments v_x { _ _ }. (** [constants::{constants::V#1}::LEN] - Source: 'tests/src/constants.rs', lines 93:4-93:24 *) + Source: 'tests/src/constants.rs', lines 94:4-94:24 *) Definition v_len_body (T : Type) (N : usize) : result usize := Ok N. Definition v_len (T : Type) (N : usize) : usize := (v_len_body T N)%global. (** [constants::use_v]: - Source: 'tests/src/constants.rs', lines 96:0-96:42 *) + Source: 'tests/src/constants.rs', lines 97:0-97:42 *) Definition use_v (T : Type) (N : usize) : result usize := Ok (v_len T N). diff --git a/tests/coq/misc/External_Funs.v b/tests/coq/misc/External_Funs.v index 18586012..7b9a9842 100644 --- a/tests/coq/misc/External_Funs.v +++ b/tests/coq/misc/External_Funs.v @@ -20,14 +20,14 @@ Definition core_marker_CopyU32 : core_marker_Copy_t u32 := {| |}. (** [external::use_get]: - Source: 'tests/src/external.rs', lines 8:0-8:37 *) + Source: 'tests/src/external.rs', lines 9:0-9:37 *) Definition use_get (rc : core_cell_Cell_t u32) (st : state) : result (state * u32) := core_cell_Cell_get u32 core_marker_CopyU32 rc st . (** [external::incr]: - Source: 'tests/src/external.rs', lines 12:0-12:31 *) + Source: 'tests/src/external.rs', lines 13:0-13:31 *) Definition incr (rc : core_cell_Cell_t u32) (st : state) : result (state * (core_cell_Cell_t u32)) diff --git a/tests/coq/misc/Loops.v b/tests/coq/misc/Loops.v index bc8708f4..87b05193 100644 --- a/tests/coq/misc/Loops.v +++ b/tests/coq/misc/Loops.v @@ -9,7 +9,7 @@ Local Open Scope Primitives_scope. Module Loops. (** [loops::sum]: loop 0: - Source: 'tests/src/loops.rs', lines 7:0-17:1 *) + Source: 'tests/src/loops.rs', lines 8:0-18:1 *) Fixpoint sum_loop (n : nat) (max : u32) (i : u32) (s : u32) : result u32 := match n with | O => Fail_ OutOfFuel @@ -21,13 +21,13 @@ Fixpoint sum_loop (n : nat) (max : u32) (i : u32) (s : u32) : result u32 := . (** [loops::sum]: - Source: 'tests/src/loops.rs', lines 7:0-7:27 *) + Source: 'tests/src/loops.rs', lines 8:0-8:27 *) Definition sum (n : nat) (max : u32) : result u32 := sum_loop n max 0%u32 0%u32 . (** [loops::sum_with_mut_borrows]: loop 0: - Source: 'tests/src/loops.rs', lines 22:0-34:1 *) + Source: 'tests/src/loops.rs', lines 23:0-35:1 *) Fixpoint sum_with_mut_borrows_loop (n : nat) (max : u32) (i : u32) (s : u32) : result u32 := match n with @@ -43,13 +43,13 @@ Fixpoint sum_with_mut_borrows_loop . (** [loops::sum_with_mut_borrows]: - Source: 'tests/src/loops.rs', lines 22:0-22:44 *) + Source: 'tests/src/loops.rs', lines 23:0-23:44 *) Definition sum_with_mut_borrows (n : nat) (max : u32) : result u32 := sum_with_mut_borrows_loop n max 0%u32 0%u32 . (** [loops::sum_with_shared_borrows]: loop 0: - Source: 'tests/src/loops.rs', lines 37:0-51:1 *) + Source: 'tests/src/loops.rs', lines 38:0-52:1 *) Fixpoint sum_with_shared_borrows_loop (n : nat) (max : u32) (i : u32) (s : u32) : result u32 := match n with @@ -65,13 +65,13 @@ Fixpoint sum_with_shared_borrows_loop . (** [loops::sum_with_shared_borrows]: - Source: 'tests/src/loops.rs', lines 37:0-37:47 *) + Source: 'tests/src/loops.rs', lines 38:0-38:47 *) Definition sum_with_shared_borrows (n : nat) (max : u32) : result u32 := sum_with_shared_borrows_loop n max 0%u32 0%u32 . (** [loops::sum_array]: loop 0: - Source: 'tests/src/loops.rs', lines 53:0-61:1 *) + Source: 'tests/src/loops.rs', lines 54:0-62:1 *) Fixpoint sum_array_loop (N : usize) (n : nat) (a : array u32 N) (i : usize) (s : u32) : result u32 := match n with @@ -88,13 +88,13 @@ Fixpoint sum_array_loop . (** [loops::sum_array]: - Source: 'tests/src/loops.rs', lines 53:0-53:52 *) + Source: 'tests/src/loops.rs', lines 54:0-54:52 *) Definition sum_array (N : usize) (n : nat) (a : array u32 N) : result u32 := sum_array_loop N n a 0%usize 0%u32 . (** [loops::clear]: loop 0: - Source: 'tests/src/loops.rs', lines 65:0-71:1 *) + Source: 'tests/src/loops.rs', lines 66:0-72:1 *) Fixpoint clear_loop (n : nat) (v : alloc_vec_Vec u32) (i : usize) : result (alloc_vec_Vec u32) := match n with @@ -115,14 +115,14 @@ Fixpoint clear_loop . (** [loops::clear]: - Source: 'tests/src/loops.rs', lines 65:0-65:30 *) + Source: 'tests/src/loops.rs', lines 66:0-66:30 *) Definition clear (n : nat) (v : alloc_vec_Vec u32) : result (alloc_vec_Vec u32) := clear_loop n v 0%usize . (** [loops::List] - Source: 'tests/src/loops.rs', lines 73:0-73:16 *) + Source: 'tests/src/loops.rs', lines 74:0-74:16 *) Inductive List_t (T : Type) := | List_Cons : T -> List_t T -> List_t T | List_Nil : List_t T @@ -132,7 +132,7 @@ Arguments List_Cons { _ }. Arguments List_Nil { _ }. (** [loops::list_mem]: loop 0: - Source: 'tests/src/loops.rs', lines 79:0-88:1 *) + Source: 'tests/src/loops.rs', lines 80:0-89:1 *) Fixpoint list_mem_loop (n : nat) (x : u32) (ls : List_t u32) : result bool := match n with | O => Fail_ OutOfFuel @@ -145,13 +145,13 @@ Fixpoint list_mem_loop (n : nat) (x : u32) (ls : List_t u32) : result bool := . (** [loops::list_mem]: - Source: 'tests/src/loops.rs', lines 79:0-79:52 *) + Source: 'tests/src/loops.rs', lines 80:0-80:52 *) Definition list_mem (n : nat) (x : u32) (ls : List_t u32) : result bool := list_mem_loop n x ls . (** [loops::list_nth_mut_loop]: loop 0: - Source: 'tests/src/loops.rs', lines 91:0-101:1 *) + Source: 'tests/src/loops.rs', lines 92:0-102:1 *) Fixpoint list_nth_mut_loop_loop (T : Type) (n : nat) (ls : List_t T) (i : u32) : result (T * (T -> result (List_t T))) @@ -175,7 +175,7 @@ Fixpoint list_nth_mut_loop_loop . (** [loops::list_nth_mut_loop]: - Source: 'tests/src/loops.rs', lines 91:0-91:71 *) + Source: 'tests/src/loops.rs', lines 92:0-92:71 *) Definition list_nth_mut_loop (T : Type) (n : nat) (ls : List_t T) (i : u32) : result (T * (T -> result (List_t T))) @@ -184,7 +184,7 @@ Definition list_nth_mut_loop . (** [loops::list_nth_shared_loop]: loop 0: - Source: 'tests/src/loops.rs', lines 104:0-114:1 *) + Source: 'tests/src/loops.rs', lines 105:0-115:1 *) Fixpoint list_nth_shared_loop_loop (T : Type) (n : nat) (ls : List_t T) (i : u32) : result T := match n with @@ -201,14 +201,14 @@ Fixpoint list_nth_shared_loop_loop . (** [loops::list_nth_shared_loop]: - Source: 'tests/src/loops.rs', lines 104:0-104:66 *) + Source: 'tests/src/loops.rs', lines 105:0-105:66 *) Definition list_nth_shared_loop (T : Type) (n : nat) (ls : List_t T) (i : u32) : result T := list_nth_shared_loop_loop T n ls i . (** [loops::get_elem_mut]: loop 0: - Source: 'tests/src/loops.rs', lines 116:0-130:1 *) + Source: 'tests/src/loops.rs', lines 117:0-131:1 *) Fixpoint get_elem_mut_loop (n : nat) (x : usize) (ls : List_t usize) : result (usize * (usize -> result (List_t usize))) @@ -233,7 +233,7 @@ Fixpoint get_elem_mut_loop . (** [loops::get_elem_mut]: - Source: 'tests/src/loops.rs', lines 116:0-116:73 *) + Source: 'tests/src/loops.rs', lines 117:0-117:73 *) Definition get_elem_mut (n : nat) (slots : alloc_vec_Vec (List_t usize)) (x : usize) : result (usize * (usize -> result (alloc_vec_Vec (List_t usize)))) @@ -249,7 +249,7 @@ Definition get_elem_mut . (** [loops::get_elem_shared]: loop 0: - Source: 'tests/src/loops.rs', lines 132:0-146:1 *) + Source: 'tests/src/loops.rs', lines 133:0-147:1 *) Fixpoint get_elem_shared_loop (n : nat) (x : usize) (ls : List_t usize) : result usize := match n with @@ -263,7 +263,7 @@ Fixpoint get_elem_shared_loop . (** [loops::get_elem_shared]: - Source: 'tests/src/loops.rs', lines 132:0-132:68 *) + Source: 'tests/src/loops.rs', lines 133:0-133:68 *) Definition get_elem_shared (n : nat) (slots : alloc_vec_Vec (List_t usize)) (x : usize) : result usize @@ -275,7 +275,7 @@ Definition get_elem_shared . (** [loops::id_mut]: - Source: 'tests/src/loops.rs', lines 148:0-148:50 *) + Source: 'tests/src/loops.rs', lines 149:0-149:50 *) Definition id_mut (T : Type) (ls : List_t T) : result ((List_t T) * (List_t T -> result (List_t T))) @@ -284,12 +284,12 @@ Definition id_mut . (** [loops::id_shared]: - Source: 'tests/src/loops.rs', lines 152:0-152:45 *) + Source: 'tests/src/loops.rs', lines 153:0-153:45 *) Definition id_shared (T : Type) (ls : List_t T) : result (List_t T) := Ok ls. (** [loops::list_nth_mut_loop_with_id]: loop 0: - Source: 'tests/src/loops.rs', lines 157:0-168:1 *) + Source: 'tests/src/loops.rs', lines 158:0-169:1 *) Fixpoint list_nth_mut_loop_with_id_loop (T : Type) (n : nat) (i : u32) (ls : List_t T) : result (T * (T -> result (List_t T))) @@ -313,7 +313,7 @@ Fixpoint list_nth_mut_loop_with_id_loop . (** [loops::list_nth_mut_loop_with_id]: - Source: 'tests/src/loops.rs', lines 157:0-157:75 *) + Source: 'tests/src/loops.rs', lines 158:0-158:75 *) Definition list_nth_mut_loop_with_id (T : Type) (n : nat) (ls : List_t T) (i : u32) : result (T * (T -> result (List_t T))) @@ -327,7 +327,7 @@ Definition list_nth_mut_loop_with_id . (** [loops::list_nth_shared_loop_with_id]: loop 0: - Source: 'tests/src/loops.rs', lines 171:0-182:1 *) + Source: 'tests/src/loops.rs', lines 172:0-183:1 *) Fixpoint list_nth_shared_loop_with_id_loop (T : Type) (n : nat) (i : u32) (ls : List_t T) : result T := match n with @@ -345,14 +345,14 @@ Fixpoint list_nth_shared_loop_with_id_loop . (** [loops::list_nth_shared_loop_with_id]: - Source: 'tests/src/loops.rs', lines 171:0-171:70 *) + Source: 'tests/src/loops.rs', lines 172:0-172:70 *) Definition list_nth_shared_loop_with_id (T : Type) (n : nat) (ls : List_t T) (i : u32) : result T := ls1 <- id_shared T ls; list_nth_shared_loop_with_id_loop T n i ls1 . (** [loops::list_nth_mut_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 187:0-208:1 *) + Source: 'tests/src/loops.rs', lines 188:0-209:1 *) Fixpoint list_nth_mut_loop_pair_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T)) * (T -> result (List_t T))) @@ -386,7 +386,7 @@ Fixpoint list_nth_mut_loop_pair_loop . (** [loops::list_nth_mut_loop_pair]: - Source: 'tests/src/loops.rs', lines 187:0-191:27 *) + Source: 'tests/src/loops.rs', lines 188:0-192:27 *) Definition list_nth_mut_loop_pair (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T)) * (T -> result (List_t T))) @@ -395,7 +395,7 @@ Definition list_nth_mut_loop_pair . (** [loops::list_nth_shared_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 211:0-232:1 *) + Source: 'tests/src/loops.rs', lines 212:0-233:1 *) Fixpoint list_nth_shared_loop_pair_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result (T * T) @@ -419,7 +419,7 @@ Fixpoint list_nth_shared_loop_pair_loop . (** [loops::list_nth_shared_loop_pair]: - Source: 'tests/src/loops.rs', lines 211:0-215:19 *) + Source: 'tests/src/loops.rs', lines 212:0-216:19 *) Definition list_nth_shared_loop_pair (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result (T * T) @@ -428,7 +428,7 @@ Definition list_nth_shared_loop_pair . (** [loops::list_nth_mut_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 236:0-251:1 *) + Source: 'tests/src/loops.rs', lines 237:0-252:1 *) Fixpoint list_nth_mut_loop_pair_merge_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * ((T * T) -> result ((List_t T) * (List_t T)))) @@ -464,7 +464,7 @@ Fixpoint list_nth_mut_loop_pair_merge_loop . (** [loops::list_nth_mut_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 236:0-240:27 *) + Source: 'tests/src/loops.rs', lines 237:0-241:27 *) Definition list_nth_mut_loop_pair_merge (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * ((T * T) -> result ((List_t T) * (List_t T)))) @@ -473,7 +473,7 @@ Definition list_nth_mut_loop_pair_merge . (** [loops::list_nth_shared_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 254:0-269:1 *) + Source: 'tests/src/loops.rs', lines 255:0-270:1 *) Fixpoint list_nth_shared_loop_pair_merge_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result (T * T) @@ -498,7 +498,7 @@ Fixpoint list_nth_shared_loop_pair_merge_loop . (** [loops::list_nth_shared_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 254:0-258:19 *) + Source: 'tests/src/loops.rs', lines 255:0-259:19 *) Definition list_nth_shared_loop_pair_merge (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result (T * T) @@ -507,7 +507,7 @@ Definition list_nth_shared_loop_pair_merge . (** [loops::list_nth_mut_shared_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 272:0-287:1 *) + Source: 'tests/src/loops.rs', lines 273:0-288:1 *) Fixpoint list_nth_mut_shared_loop_pair_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -538,7 +538,7 @@ Fixpoint list_nth_mut_shared_loop_pair_loop . (** [loops::list_nth_mut_shared_loop_pair]: - Source: 'tests/src/loops.rs', lines 272:0-276:23 *) + Source: 'tests/src/loops.rs', lines 273:0-277:23 *) Definition list_nth_mut_shared_loop_pair (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -547,7 +547,7 @@ Definition list_nth_mut_shared_loop_pair . (** [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 291:0-306:1 *) + Source: 'tests/src/loops.rs', lines 292:0-307:1 *) Fixpoint list_nth_mut_shared_loop_pair_merge_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -578,7 +578,7 @@ Fixpoint list_nth_mut_shared_loop_pair_merge_loop . (** [loops::list_nth_mut_shared_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 291:0-295:23 *) + Source: 'tests/src/loops.rs', lines 292:0-296:23 *) Definition list_nth_mut_shared_loop_pair_merge (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -587,7 +587,7 @@ Definition list_nth_mut_shared_loop_pair_merge . (** [loops::list_nth_shared_mut_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 310:0-325:1 *) + Source: 'tests/src/loops.rs', lines 311:0-326:1 *) Fixpoint list_nth_shared_mut_loop_pair_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -618,7 +618,7 @@ Fixpoint list_nth_shared_mut_loop_pair_loop . (** [loops::list_nth_shared_mut_loop_pair]: - Source: 'tests/src/loops.rs', lines 310:0-314:23 *) + Source: 'tests/src/loops.rs', lines 311:0-315:23 *) Definition list_nth_shared_mut_loop_pair (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -627,7 +627,7 @@ Definition list_nth_shared_mut_loop_pair . (** [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 329:0-344:1 *) + Source: 'tests/src/loops.rs', lines 330:0-345:1 *) Fixpoint list_nth_shared_mut_loop_pair_merge_loop (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -658,7 +658,7 @@ Fixpoint list_nth_shared_mut_loop_pair_merge_loop . (** [loops::list_nth_shared_mut_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 329:0-333:23 *) + Source: 'tests/src/loops.rs', lines 330:0-334:23 *) Definition list_nth_shared_mut_loop_pair_merge (T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) : result ((T * T) * (T -> result (List_t T))) @@ -667,7 +667,7 @@ Definition list_nth_shared_mut_loop_pair_merge . (** [loops::ignore_input_mut_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 348:0-352:1 *) + Source: 'tests/src/loops.rs', lines 349:0-353:1 *) Fixpoint ignore_input_mut_borrow_loop (n : nat) (i : u32) : result unit := match n with | O => Fail_ OutOfFuel @@ -679,14 +679,14 @@ Fixpoint ignore_input_mut_borrow_loop (n : nat) (i : u32) : result unit := . (** [loops::ignore_input_mut_borrow]: - Source: 'tests/src/loops.rs', lines 348:0-348:56 *) + Source: 'tests/src/loops.rs', lines 349:0-349:56 *) Definition ignore_input_mut_borrow (n : nat) (_a : u32) (i : u32) : result u32 := _ <- ignore_input_mut_borrow_loop n i; Ok _a . (** [loops::incr_ignore_input_mut_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 356:0-361:1 *) + Source: 'tests/src/loops.rs', lines 357:0-362:1 *) Fixpoint incr_ignore_input_mut_borrow_loop (n : nat) (i : u32) : result unit := match n with | O => Fail_ OutOfFuel @@ -698,14 +698,14 @@ Fixpoint incr_ignore_input_mut_borrow_loop (n : nat) (i : u32) : result unit := . (** [loops::incr_ignore_input_mut_borrow]: - Source: 'tests/src/loops.rs', lines 356:0-356:60 *) + Source: 'tests/src/loops.rs', lines 357:0-357:60 *) Definition incr_ignore_input_mut_borrow (n : nat) (a : u32) (i : u32) : result u32 := a1 <- u32_add a 1%u32; _ <- incr_ignore_input_mut_borrow_loop n i; Ok a1 . (** [loops::ignore_input_shared_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 365:0-369:1 *) + Source: 'tests/src/loops.rs', lines 366:0-370:1 *) Fixpoint ignore_input_shared_borrow_loop (n : nat) (i : u32) : result unit := match n with | O => Fail_ OutOfFuel @@ -717,7 +717,7 @@ Fixpoint ignore_input_shared_borrow_loop (n : nat) (i : u32) : result unit := . (** [loops::ignore_input_shared_borrow]: - Source: 'tests/src/loops.rs', lines 365:0-365:59 *) + Source: 'tests/src/loops.rs', lines 366:0-366:59 *) Definition ignore_input_shared_borrow (n : nat) (_a : u32) (i : u32) : result u32 := _ <- ignore_input_shared_borrow_loop n i; Ok _a diff --git a/tests/coq/misc/NoNestedBorrows.v b/tests/coq/misc/NoNestedBorrows.v index 434b820c..2cc6af6c 100644 --- a/tests/coq/misc/NoNestedBorrows.v +++ b/tests/coq/misc/NoNestedBorrows.v @@ -9,7 +9,7 @@ Local Open Scope Primitives_scope. Module NoNestedBorrows. (** [no_nested_borrows::Pair] - Source: 'tests/src/no_nested_borrows.rs', lines 6:0-6:23 *) + Source: 'tests/src/no_nested_borrows.rs', lines 7:0-7:23 *) Record Pair_t (T1 T2 : Type) := mkPair_t { pair_x : T1; pair_y : T2; }. Arguments mkPair_t { _ _ }. @@ -17,7 +17,7 @@ Arguments pair_x { _ _ }. Arguments pair_y { _ _ }. (** [no_nested_borrows::List] - Source: 'tests/src/no_nested_borrows.rs', lines 11:0-11:16 *) + Source: 'tests/src/no_nested_borrows.rs', lines 12:0-12:16 *) Inductive List_t (T : Type) := | List_Cons : T -> List_t T -> List_t T | List_Nil : List_t T @@ -27,25 +27,25 @@ Arguments List_Cons { _ }. Arguments List_Nil { _ }. (** [no_nested_borrows::One] - Source: 'tests/src/no_nested_borrows.rs', lines 22:0-22:16 *) + Source: 'tests/src/no_nested_borrows.rs', lines 23:0-23:16 *) Inductive One_t (T1 : Type) := | One_One : T1 -> One_t T1. Arguments One_One { _ }. (** [no_nested_borrows::EmptyEnum] - Source: 'tests/src/no_nested_borrows.rs', lines 28:0-28:18 *) + Source: 'tests/src/no_nested_borrows.rs', lines 29:0-29:18 *) Inductive EmptyEnum_t := | EmptyEnum_Empty : EmptyEnum_t. (** [no_nested_borrows::Enum] - Source: 'tests/src/no_nested_borrows.rs', lines 34:0-34:13 *) + Source: 'tests/src/no_nested_borrows.rs', lines 35:0-35:13 *) Inductive Enum_t := | Enum_Variant1 : Enum_t | Enum_Variant2 : Enum_t. (** [no_nested_borrows::EmptyStruct] - Source: 'tests/src/no_nested_borrows.rs', lines 41:0-41:22 *) + Source: 'tests/src/no_nested_borrows.rs', lines 42:0-42:22 *) Definition EmptyStruct_t : Type := unit. (** [no_nested_borrows::Sum] - Source: 'tests/src/no_nested_borrows.rs', lines 43:0-43:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 44:0-44:20 *) Inductive Sum_t (T1 T2 : Type) := | Sum_Left : T1 -> Sum_t T1 T2 | Sum_Right : T2 -> Sum_t T1 T2 @@ -55,22 +55,22 @@ Arguments Sum_Left { _ _ }. Arguments Sum_Right { _ _ }. (** [no_nested_borrows::cast_u32_to_i32]: - Source: 'tests/src/no_nested_borrows.rs', lines 48:0-48:37 *) + Source: 'tests/src/no_nested_borrows.rs', lines 49:0-49:37 *) Definition cast_u32_to_i32 (x : u32) : result i32 := scalar_cast U32 I32 x. (** [no_nested_borrows::cast_bool_to_i32]: - Source: 'tests/src/no_nested_borrows.rs', lines 52:0-52:39 *) + Source: 'tests/src/no_nested_borrows.rs', lines 53:0-53:39 *) Definition cast_bool_to_i32 (x : bool) : result i32 := scalar_cast_bool I32 x. (** [no_nested_borrows::cast_bool_to_bool]: - Source: 'tests/src/no_nested_borrows.rs', lines 57:0-57:41 *) + Source: 'tests/src/no_nested_borrows.rs', lines 58:0-58:41 *) Definition cast_bool_to_bool (x : bool) : result bool := Ok x. (** [no_nested_borrows::test2]: - Source: 'tests/src/no_nested_borrows.rs', lines 62:0-62:14 *) + Source: 'tests/src/no_nested_borrows.rs', lines 63:0-63:14 *) Definition test2 : result unit := _ <- u32_add 23%u32 44%u32; Ok tt. @@ -78,13 +78,13 @@ Definition test2 : result unit := Check (test2 )%return. (** [no_nested_borrows::get_max]: - Source: 'tests/src/no_nested_borrows.rs', lines 74:0-74:37 *) + Source: 'tests/src/no_nested_borrows.rs', lines 75:0-75:37 *) Definition get_max (x : u32) (y : u32) : result u32 := if x s>= y then Ok x else Ok y . (** [no_nested_borrows::test3]: - Source: 'tests/src/no_nested_borrows.rs', lines 82:0-82:14 *) + Source: 'tests/src/no_nested_borrows.rs', lines 83:0-83:14 *) Definition test3 : result unit := x <- get_max 4%u32 3%u32; y <- get_max 10%u32 11%u32; @@ -96,7 +96,7 @@ Definition test3 : result unit := Check (test3 )%return. (** [no_nested_borrows::test_neg1]: - Source: 'tests/src/no_nested_borrows.rs', lines 89:0-89:18 *) + Source: 'tests/src/no_nested_borrows.rs', lines 90:0-90:18 *) Definition test_neg1 : result unit := y <- i32_neg 3%i32; if negb (y s= (-3)%i32) then Fail_ Failure else Ok tt . @@ -105,7 +105,7 @@ Definition test_neg1 : result unit := Check (test_neg1 )%return. (** [no_nested_borrows::refs_test1]: - Source: 'tests/src/no_nested_borrows.rs', lines 96:0-96:19 *) + Source: 'tests/src/no_nested_borrows.rs', lines 97:0-97:19 *) Definition refs_test1 : result unit := if negb (1%i32 s= 1%i32) then Fail_ Failure else Ok tt . @@ -114,7 +114,7 @@ Definition refs_test1 : result unit := Check (refs_test1 )%return. (** [no_nested_borrows::refs_test2]: - Source: 'tests/src/no_nested_borrows.rs', lines 107:0-107:19 *) + Source: 'tests/src/no_nested_borrows.rs', lines 108:0-108:19 *) Definition refs_test2 : result unit := if negb (2%i32 s= 2%i32) then Fail_ Failure @@ -131,7 +131,7 @@ Definition refs_test2 : result unit := Check (refs_test2 )%return. (** [no_nested_borrows::test_list1]: - Source: 'tests/src/no_nested_borrows.rs', lines 123:0-123:19 *) + Source: 'tests/src/no_nested_borrows.rs', lines 124:0-124:19 *) Definition test_list1 : result unit := Ok tt. @@ -139,7 +139,7 @@ Definition test_list1 : result unit := Check (test_list1 )%return. (** [no_nested_borrows::test_box1]: - Source: 'tests/src/no_nested_borrows.rs', lines 128:0-128:18 *) + Source: 'tests/src/no_nested_borrows.rs', lines 129:0-129:18 *) Definition test_box1 : result unit := p <- alloc_boxed_Box_deref_mut i32 0%i32; let (_, deref_mut_back) := p in @@ -152,24 +152,24 @@ Definition test_box1 : result unit := Check (test_box1 )%return. (** [no_nested_borrows::copy_int]: - Source: 'tests/src/no_nested_borrows.rs', lines 138:0-138:30 *) + Source: 'tests/src/no_nested_borrows.rs', lines 139:0-139:30 *) Definition copy_int (x : i32) : result i32 := Ok x. (** [no_nested_borrows::test_unreachable]: - Source: 'tests/src/no_nested_borrows.rs', lines 144:0-144:32 *) + Source: 'tests/src/no_nested_borrows.rs', lines 145:0-145:32 *) Definition test_unreachable (b : bool) : result unit := if b then Fail_ Failure else Ok tt . (** [no_nested_borrows::test_panic]: - Source: 'tests/src/no_nested_borrows.rs', lines 152:0-152:26 *) + Source: 'tests/src/no_nested_borrows.rs', lines 153:0-153:26 *) Definition test_panic (b : bool) : result unit := if b then Fail_ Failure else Ok tt . (** [no_nested_borrows::test_copy_int]: - Source: 'tests/src/no_nested_borrows.rs', lines 159:0-159:22 *) + Source: 'tests/src/no_nested_borrows.rs', lines 160:0-160:22 *) Definition test_copy_int : result unit := y <- copy_int 0%i32; if negb (0%i32 s= y) then Fail_ Failure else Ok tt . @@ -178,13 +178,13 @@ Definition test_copy_int : result unit := Check (test_copy_int )%return. (** [no_nested_borrows::is_cons]: - Source: 'tests/src/no_nested_borrows.rs', lines 166:0-166:38 *) + Source: 'tests/src/no_nested_borrows.rs', lines 167:0-167:38 *) Definition is_cons (T : Type) (l : List_t T) : result bool := match l with | List_Cons _ _ => Ok true | List_Nil => Ok false end . (** [no_nested_borrows::test_is_cons]: - Source: 'tests/src/no_nested_borrows.rs', lines 173:0-173:21 *) + Source: 'tests/src/no_nested_borrows.rs', lines 174:0-174:21 *) Definition test_is_cons : result unit := b <- is_cons i32 (List_Cons 0%i32 List_Nil); if negb b then Fail_ Failure else Ok tt @@ -194,13 +194,13 @@ Definition test_is_cons : result unit := Check (test_is_cons )%return. (** [no_nested_borrows::split_list]: - Source: 'tests/src/no_nested_borrows.rs', lines 179:0-179:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 180:0-180:48 *) Definition split_list (T : Type) (l : List_t T) : result (T * (List_t T)) := match l with | List_Cons hd tl => Ok (hd, tl) | List_Nil => Fail_ Failure end . (** [no_nested_borrows::test_split_list]: - Source: 'tests/src/no_nested_borrows.rs', lines 187:0-187:24 *) + Source: 'tests/src/no_nested_borrows.rs', lines 188:0-188:24 *) Definition test_split_list : result unit := p <- split_list i32 (List_Cons 0%i32 List_Nil); let (hd, _) := p in @@ -211,7 +211,7 @@ Definition test_split_list : result unit := Check (test_split_list )%return. (** [no_nested_borrows::choose]: - Source: 'tests/src/no_nested_borrows.rs', lines 194:0-194:70 *) + Source: 'tests/src/no_nested_borrows.rs', lines 195:0-195:70 *) Definition choose (T : Type) (b : bool) (x : T) (y : T) : result (T * (T -> result (T * T))) := if b @@ -220,7 +220,7 @@ Definition choose . (** [no_nested_borrows::choose_test]: - Source: 'tests/src/no_nested_borrows.rs', lines 202:0-202:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 203:0-203:20 *) Definition choose_test : result unit := p <- choose i32 true 0%i32 0%i32; let (z, choose_back) := p in @@ -239,18 +239,18 @@ Definition choose_test : result unit := Check (choose_test )%return. (** [no_nested_borrows::test_char]: - Source: 'tests/src/no_nested_borrows.rs', lines 214:0-214:26 *) + Source: 'tests/src/no_nested_borrows.rs', lines 215:0-215:26 *) Definition test_char : result char := Ok (char_of_byte Coq.Init.Byte.x61). (** [no_nested_borrows::Tree] - Source: 'tests/src/no_nested_borrows.rs', lines 219:0-219:16 *) + Source: 'tests/src/no_nested_borrows.rs', lines 220:0-220:16 *) Inductive Tree_t (T : Type) := | Tree_Leaf : T -> Tree_t T | Tree_Node : T -> NodeElem_t T -> Tree_t T -> Tree_t T (** [no_nested_borrows::NodeElem] - Source: 'tests/src/no_nested_borrows.rs', lines 224:0-224:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 225:0-225:20 *) with NodeElem_t (T : Type) := | NodeElem_Cons : Tree_t T -> NodeElem_t T -> NodeElem_t T | NodeElem_Nil : NodeElem_t T @@ -263,7 +263,7 @@ Arguments NodeElem_Cons { _ }. Arguments NodeElem_Nil { _ }. (** [no_nested_borrows::list_length]: - Source: 'tests/src/no_nested_borrows.rs', lines 259:0-259:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 260:0-260:48 *) Fixpoint list_length (T : Type) (l : List_t T) : result u32 := match l with | List_Cons _ l1 => i <- list_length T l1; u32_add 1%u32 i @@ -272,7 +272,7 @@ Fixpoint list_length (T : Type) (l : List_t T) : result u32 := . (** [no_nested_borrows::list_nth_shared]: - Source: 'tests/src/no_nested_borrows.rs', lines 267:0-267:62 *) + Source: 'tests/src/no_nested_borrows.rs', lines 268:0-268:62 *) Fixpoint list_nth_shared (T : Type) (l : List_t T) (i : u32) : result T := match l with | List_Cons x tl => @@ -284,7 +284,7 @@ Fixpoint list_nth_shared (T : Type) (l : List_t T) (i : u32) : result T := . (** [no_nested_borrows::list_nth_mut]: - Source: 'tests/src/no_nested_borrows.rs', lines 283:0-283:67 *) + Source: 'tests/src/no_nested_borrows.rs', lines 284:0-284:67 *) Fixpoint list_nth_mut (T : Type) (l : List_t T) (i : u32) : result (T * (T -> result (List_t T))) @@ -305,7 +305,7 @@ Fixpoint list_nth_mut . (** [no_nested_borrows::list_rev_aux]: - Source: 'tests/src/no_nested_borrows.rs', lines 299:0-299:63 *) + Source: 'tests/src/no_nested_borrows.rs', lines 300:0-300:63 *) Fixpoint list_rev_aux (T : Type) (li : List_t T) (lo : List_t T) : result (List_t T) := match li with @@ -315,14 +315,14 @@ Fixpoint list_rev_aux . (** [no_nested_borrows::list_rev]: - Source: 'tests/src/no_nested_borrows.rs', lines 313:0-313:42 *) + Source: 'tests/src/no_nested_borrows.rs', lines 314:0-314:42 *) Definition list_rev (T : Type) (l : List_t T) : result (List_t T) := let (li, _) := core_mem_replace (List_t T) l List_Nil in list_rev_aux T li List_Nil . (** [no_nested_borrows::test_list_functions]: - Source: 'tests/src/no_nested_borrows.rs', lines 318:0-318:28 *) + Source: 'tests/src/no_nested_borrows.rs', lines 319:0-319:28 *) Definition test_list_functions : result unit := let l := List_Cons 2%i32 List_Nil in let l1 := List_Cons 1%i32 l in @@ -361,7 +361,7 @@ Definition test_list_functions : result unit := Check (test_list_functions )%return. (** [no_nested_borrows::id_mut_pair1]: - Source: 'tests/src/no_nested_borrows.rs', lines 334:0-334:89 *) + Source: 'tests/src/no_nested_borrows.rs', lines 335:0-335:89 *) Definition id_mut_pair1 (T1 T2 : Type) (x : T1) (y : T2) : result ((T1 * T2) * ((T1 * T2) -> result (T1 * T2))) @@ -370,7 +370,7 @@ Definition id_mut_pair1 . (** [no_nested_borrows::id_mut_pair2]: - Source: 'tests/src/no_nested_borrows.rs', lines 338:0-338:88 *) + Source: 'tests/src/no_nested_borrows.rs', lines 339:0-339:88 *) Definition id_mut_pair2 (T1 T2 : Type) (p : (T1 * T2)) : result ((T1 * T2) * ((T1 * T2) -> result (T1 * T2))) @@ -379,7 +379,7 @@ Definition id_mut_pair2 . (** [no_nested_borrows::id_mut_pair3]: - Source: 'tests/src/no_nested_borrows.rs', lines 342:0-342:93 *) + Source: 'tests/src/no_nested_borrows.rs', lines 343:0-343:93 *) Definition id_mut_pair3 (T1 T2 : Type) (x : T1) (y : T2) : result ((T1 * T2) * (T1 -> result T1) * (T2 -> result T2)) @@ -388,7 +388,7 @@ Definition id_mut_pair3 . (** [no_nested_borrows::id_mut_pair4]: - Source: 'tests/src/no_nested_borrows.rs', lines 346:0-346:92 *) + Source: 'tests/src/no_nested_borrows.rs', lines 347:0-347:92 *) Definition id_mut_pair4 (T1 T2 : Type) (p : (T1 * T2)) : result ((T1 * T2) * (T1 -> result T1) * (T2 -> result T2)) @@ -397,7 +397,7 @@ Definition id_mut_pair4 . (** [no_nested_borrows::StructWithTuple] - Source: 'tests/src/no_nested_borrows.rs', lines 353:0-353:34 *) + Source: 'tests/src/no_nested_borrows.rs', lines 354:0-354:34 *) Record StructWithTuple_t (T1 T2 : Type) := mkStructWithTuple_t { structWithTuple_p : (T1 * T2); @@ -408,25 +408,25 @@ Arguments mkStructWithTuple_t { _ _ }. Arguments structWithTuple_p { _ _ }. (** [no_nested_borrows::new_tuple1]: - Source: 'tests/src/no_nested_borrows.rs', lines 357:0-357:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 358:0-358:48 *) Definition new_tuple1 : result (StructWithTuple_t u32 u32) := Ok {| structWithTuple_p := (1%u32, 2%u32) |} . (** [no_nested_borrows::new_tuple2]: - Source: 'tests/src/no_nested_borrows.rs', lines 361:0-361:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 362:0-362:48 *) Definition new_tuple2 : result (StructWithTuple_t i16 i16) := Ok {| structWithTuple_p := (1%i16, 2%i16) |} . (** [no_nested_borrows::new_tuple3]: - Source: 'tests/src/no_nested_borrows.rs', lines 365:0-365:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 366:0-366:48 *) Definition new_tuple3 : result (StructWithTuple_t u64 i64) := Ok {| structWithTuple_p := (1%u64, 2%i64) |} . (** [no_nested_borrows::StructWithPair] - Source: 'tests/src/no_nested_borrows.rs', lines 370:0-370:33 *) + Source: 'tests/src/no_nested_borrows.rs', lines 371:0-371:33 *) Record StructWithPair_t (T1 T2 : Type) := mkStructWithPair_t { structWithPair_p : Pair_t T1 T2; @@ -437,13 +437,13 @@ Arguments mkStructWithPair_t { _ _ }. Arguments structWithPair_p { _ _ }. (** [no_nested_borrows::new_pair1]: - Source: 'tests/src/no_nested_borrows.rs', lines 374:0-374:46 *) + Source: 'tests/src/no_nested_borrows.rs', lines 375:0-375:46 *) Definition new_pair1 : result (StructWithPair_t u32 u32) := Ok {| structWithPair_p := {| pair_x := 1%u32; pair_y := 2%u32 |} |} . (** [no_nested_borrows::test_constants]: - Source: 'tests/src/no_nested_borrows.rs', lines 382:0-382:23 *) + Source: 'tests/src/no_nested_borrows.rs', lines 383:0-383:23 *) Definition test_constants : result unit := swt <- new_tuple1; let (i, _) := swt.(structWithTuple_p) in @@ -470,7 +470,7 @@ Definition test_constants : result unit := Check (test_constants )%return. (** [no_nested_borrows::test_weird_borrows1]: - Source: 'tests/src/no_nested_borrows.rs', lines 391:0-391:28 *) + Source: 'tests/src/no_nested_borrows.rs', lines 392:0-392:28 *) Definition test_weird_borrows1 : result unit := Ok tt. @@ -478,78 +478,78 @@ Definition test_weird_borrows1 : result unit := Check (test_weird_borrows1 )%return. (** [no_nested_borrows::test_mem_replace]: - Source: 'tests/src/no_nested_borrows.rs', lines 401:0-401:37 *) + Source: 'tests/src/no_nested_borrows.rs', lines 402:0-402:37 *) Definition test_mem_replace (px : u32) : result u32 := let (y, _) := core_mem_replace u32 px 1%u32 in if negb (y s= 0%u32) then Fail_ Failure else Ok 2%u32 . (** [no_nested_borrows::test_shared_borrow_bool1]: - Source: 'tests/src/no_nested_borrows.rs', lines 408:0-408:47 *) + Source: 'tests/src/no_nested_borrows.rs', lines 409:0-409:47 *) Definition test_shared_borrow_bool1 (b : bool) : result u32 := if b then Ok 0%u32 else Ok 1%u32 . (** [no_nested_borrows::test_shared_borrow_bool2]: - Source: 'tests/src/no_nested_borrows.rs', lines 421:0-421:40 *) + Source: 'tests/src/no_nested_borrows.rs', lines 422:0-422:40 *) Definition test_shared_borrow_bool2 : result u32 := Ok 0%u32. (** [no_nested_borrows::test_shared_borrow_enum1]: - Source: 'tests/src/no_nested_borrows.rs', lines 436:0-436:52 *) + Source: 'tests/src/no_nested_borrows.rs', lines 437:0-437:52 *) Definition test_shared_borrow_enum1 (l : List_t u32) : result u32 := match l with | List_Cons _ _ => Ok 1%u32 | List_Nil => Ok 0%u32 end . (** [no_nested_borrows::test_shared_borrow_enum2]: - Source: 'tests/src/no_nested_borrows.rs', lines 448:0-448:40 *) + Source: 'tests/src/no_nested_borrows.rs', lines 449:0-449:40 *) Definition test_shared_borrow_enum2 : result u32 := Ok 0%u32. (** [no_nested_borrows::incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 459:0-459:24 *) + Source: 'tests/src/no_nested_borrows.rs', lines 460:0-460:24 *) Definition incr (x : u32) : result u32 := u32_add x 1%u32. (** [no_nested_borrows::call_incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 463:0-463:35 *) + Source: 'tests/src/no_nested_borrows.rs', lines 464:0-464:35 *) Definition call_incr (x : u32) : result u32 := incr x. (** [no_nested_borrows::read_then_incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 468:0-468:41 *) + Source: 'tests/src/no_nested_borrows.rs', lines 469:0-469:41 *) Definition read_then_incr (x : u32) : result (u32 * u32) := x1 <- u32_add x 1%u32; Ok (x, x1) . (** [no_nested_borrows::Tuple] - Source: 'tests/src/no_nested_borrows.rs', lines 474:0-474:24 *) + Source: 'tests/src/no_nested_borrows.rs', lines 475:0-475:24 *) Definition Tuple_t (T1 T2 : Type) : Type := T1 * T2. (** [no_nested_borrows::use_tuple_struct]: - Source: 'tests/src/no_nested_borrows.rs', lines 476:0-476:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 477:0-477:48 *) Definition use_tuple_struct (x : Tuple_t u32 u32) : result (Tuple_t u32 u32) := let (_, i) := x in Ok (1%u32, i) . (** [no_nested_borrows::create_tuple_struct]: - Source: 'tests/src/no_nested_borrows.rs', lines 480:0-480:61 *) + Source: 'tests/src/no_nested_borrows.rs', lines 481:0-481:61 *) Definition create_tuple_struct (x : u32) (y : u64) : result (Tuple_t u32 u64) := Ok (x, y) . (** [no_nested_borrows::IdType] - Source: 'tests/src/no_nested_borrows.rs', lines 485:0-485:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 486:0-486:20 *) Definition IdType_t (T : Type) : Type := T. (** [no_nested_borrows::use_id_type]: - Source: 'tests/src/no_nested_borrows.rs', lines 487:0-487:40 *) + Source: 'tests/src/no_nested_borrows.rs', lines 488:0-488:40 *) Definition use_id_type (T : Type) (x : IdType_t T) : result T := Ok x. (** [no_nested_borrows::create_id_type]: - Source: 'tests/src/no_nested_borrows.rs', lines 491:0-491:43 *) + Source: 'tests/src/no_nested_borrows.rs', lines 492:0-492:43 *) Definition create_id_type (T : Type) (x : T) : result (IdType_t T) := Ok x. diff --git a/tests/coq/misc/Paper.v b/tests/coq/misc/Paper.v index 21e86542..e5728364 100644 --- a/tests/coq/misc/Paper.v +++ b/tests/coq/misc/Paper.v @@ -9,12 +9,12 @@ Local Open Scope Primitives_scope. Module Paper. (** [paper::ref_incr]: - Source: 'tests/src/paper.rs', lines 6:0-6:28 *) + Source: 'tests/src/paper.rs', lines 7:0-7:28 *) Definition ref_incr (x : i32) : result i32 := i32_add x 1%i32. (** [paper::test_incr]: - Source: 'tests/src/paper.rs', lines 10:0-10:18 *) + Source: 'tests/src/paper.rs', lines 11:0-11:18 *) Definition test_incr : result unit := x <- ref_incr 0%i32; if negb (x s= 1%i32) then Fail_ Failure else Ok tt . @@ -23,7 +23,7 @@ Definition test_incr : result unit := Check (test_incr )%return. (** [paper::choose]: - Source: 'tests/src/paper.rs', lines 17:0-17:70 *) + Source: 'tests/src/paper.rs', lines 18:0-18:70 *) Definition choose (T : Type) (b : bool) (x : T) (y : T) : result (T * (T -> result (T * T))) := if b @@ -32,7 +32,7 @@ Definition choose . (** [paper::test_choose]: - Source: 'tests/src/paper.rs', lines 25:0-25:20 *) + Source: 'tests/src/paper.rs', lines 26:0-26:20 *) Definition test_choose : result unit := p <- choose i32 true 0%i32 0%i32; let (z, choose_back) := p in @@ -51,7 +51,7 @@ Definition test_choose : result unit := Check (test_choose )%return. (** [paper::List] - Source: 'tests/src/paper.rs', lines 37:0-37:16 *) + Source: 'tests/src/paper.rs', lines 38:0-38:16 *) Inductive List_t (T : Type) := | List_Cons : T -> List_t T -> List_t T | List_Nil : List_t T @@ -61,7 +61,7 @@ Arguments List_Cons { _ }. Arguments List_Nil { _ }. (** [paper::list_nth_mut]: - Source: 'tests/src/paper.rs', lines 44:0-44:67 *) + Source: 'tests/src/paper.rs', lines 45:0-45:67 *) Fixpoint list_nth_mut (T : Type) (l : List_t T) (i : u32) : result (T * (T -> result (List_t T))) @@ -82,7 +82,7 @@ Fixpoint list_nth_mut . (** [paper::sum]: - Source: 'tests/src/paper.rs', lines 59:0-59:32 *) + Source: 'tests/src/paper.rs', lines 60:0-60:32 *) Fixpoint sum (l : List_t i32) : result i32 := match l with | List_Cons x tl => i <- sum tl; i32_add x i @@ -91,7 +91,7 @@ Fixpoint sum (l : List_t i32) : result i32 := . (** [paper::test_nth]: - Source: 'tests/src/paper.rs', lines 70:0-70:17 *) + Source: 'tests/src/paper.rs', lines 71:0-71:17 *) Definition test_nth : result unit := let l := List_Cons 3%i32 List_Nil in let l1 := List_Cons 2%i32 l in @@ -107,7 +107,7 @@ Definition test_nth : result unit := Check (test_nth )%return. (** [paper::call_choose]: - Source: 'tests/src/paper.rs', lines 78:0-78:44 *) + Source: 'tests/src/paper.rs', lines 79:0-79:44 *) Definition call_choose (p : (u32 * u32)) : result u32 := let (px, py) := p in p1 <- choose u32 true px py; diff --git a/tests/coq/misc/PoloniusList.v b/tests/coq/misc/PoloniusList.v index 91cfcdb7..a600deaa 100644 --- a/tests/coq/misc/PoloniusList.v +++ b/tests/coq/misc/PoloniusList.v @@ -9,7 +9,7 @@ Local Open Scope Primitives_scope. Module PoloniusList. (** [polonius_list::List] - Source: 'tests/src/polonius_list.rs', lines 5:0-5:16 *) + Source: 'tests/src/polonius_list.rs', lines 6:0-6:16 *) Inductive List_t (T : Type) := | List_Cons : T -> List_t T -> List_t T | List_Nil : List_t T @@ -19,7 +19,7 @@ Arguments List_Cons { _ }. Arguments List_Nil { _ }. (** [polonius_list::get_list_at_x]: - Source: 'tests/src/polonius_list.rs', lines 15:0-15:76 *) + Source: 'tests/src/polonius_list.rs', lines 16:0-16:76 *) Fixpoint get_list_at_x (ls : List_t u32) (x : u32) : result ((List_t u32) * (List_t u32 -> result (List_t u32))) diff --git a/tests/fstar/demo/Demo.fst b/tests/fstar/demo/Demo.fst index 60722f46..d89f32e0 100644 --- a/tests/fstar/demo/Demo.fst +++ b/tests/fstar/demo/Demo.fst @@ -6,7 +6,7 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [demo::choose]: - Source: 'tests/src/demo.rs', lines 6:0-6:70 *) + Source: 'tests/src/demo.rs', lines 7:0-7:70 *) let choose (t : Type0) (b : bool) (x : t) (y : t) : result (t & (t -> result (t & t))) = if b @@ -14,33 +14,33 @@ let choose else let back = fun ret -> Ok (x, ret) in Ok (y, back) (** [demo::mul2_add1]: - Source: 'tests/src/demo.rs', lines 14:0-14:31 *) + Source: 'tests/src/demo.rs', lines 15:0-15:31 *) let mul2_add1 (x : u32) : result u32 = let* i = u32_add x x in u32_add i 1 (** [demo::use_mul2_add1]: - Source: 'tests/src/demo.rs', lines 18:0-18:43 *) + Source: 'tests/src/demo.rs', lines 19:0-19:43 *) let use_mul2_add1 (x : u32) (y : u32) : result u32 = let* i = mul2_add1 x in u32_add i y (** [demo::incr]: - Source: 'tests/src/demo.rs', lines 22:0-22:31 *) + Source: 'tests/src/demo.rs', lines 23:0-23:31 *) let incr (x : u32) : result u32 = u32_add x 1 (** [demo::use_incr]: - Source: 'tests/src/demo.rs', lines 26:0-26:17 *) + Source: 'tests/src/demo.rs', lines 27:0-27:17 *) let use_incr : result unit = let* x = incr 0 in let* x1 = incr x in let* _ = incr x1 in Ok () (** [demo::CList] - Source: 'tests/src/demo.rs', lines 35:0-35:17 *) + Source: 'tests/src/demo.rs', lines 36:0-36:17 *) type cList_t (t : Type0) = | CList_CCons : t -> cList_t t -> cList_t t | CList_CNil : cList_t t (** [demo::list_nth]: - Source: 'tests/src/demo.rs', lines 40:0-40:56 *) + Source: 'tests/src/demo.rs', lines 41:0-41:56 *) let rec list_nth (t : Type0) (n : nat) (l : cList_t t) (i : u32) : result t = if is_zero n then Fail OutOfFuel @@ -53,7 +53,7 @@ let rec list_nth (t : Type0) (n : nat) (l : cList_t t) (i : u32) : result t = end (** [demo::list_nth_mut]: - Source: 'tests/src/demo.rs', lines 55:0-55:68 *) + Source: 'tests/src/demo.rs', lines 56:0-56:68 *) let rec list_nth_mut (t : Type0) (n : nat) (l : cList_t t) (i : u32) : result (t & (t -> result (cList_t t))) @@ -77,7 +77,7 @@ let rec list_nth_mut end (** [demo::list_nth_mut1]: loop 0: - Source: 'tests/src/demo.rs', lines 70:0-79:1 *) + Source: 'tests/src/demo.rs', lines 71:0-80:1 *) let rec list_nth_mut1_loop (t : Type0) (n : nat) (l : cList_t t) (i : u32) : result (t & (t -> result (cList_t t))) @@ -99,7 +99,7 @@ let rec list_nth_mut1_loop end (** [demo::list_nth_mut1]: - Source: 'tests/src/demo.rs', lines 70:0-70:77 *) + Source: 'tests/src/demo.rs', lines 71:0-71:77 *) let list_nth_mut1 (t : Type0) (n : nat) (l : cList_t t) (i : u32) : result (t & (t -> result (cList_t t))) @@ -107,7 +107,7 @@ let list_nth_mut1 list_nth_mut1_loop t n l i (** [demo::i32_id]: - Source: 'tests/src/demo.rs', lines 81:0-81:28 *) + Source: 'tests/src/demo.rs', lines 82:0-82:28 *) let rec i32_id (n : nat) (i : i32) : result i32 = if is_zero n then Fail OutOfFuel @@ -118,7 +118,7 @@ let rec i32_id (n : nat) (i : i32) : result i32 = else let* i1 = i32_sub i 1 in let* i2 = i32_id n1 i1 in i32_add i2 1 (** [demo::list_tail]: - Source: 'tests/src/demo.rs', lines 89:0-89:64 *) + Source: 'tests/src/demo.rs', lines 90:0-90:64 *) let rec list_tail (t : Type0) (n : nat) (l : cList_t t) : result ((cList_t t) & (cList_t t -> result (cList_t t))) @@ -137,20 +137,20 @@ let rec list_tail end (** Trait declaration: [demo::Counter] - Source: 'tests/src/demo.rs', lines 98:0-98:17 *) + Source: 'tests/src/demo.rs', lines 99:0-99:17 *) noeq type counter_t (self : Type0) = { incr : self -> result (usize & self); } (** [demo::{(demo::Counter for usize)}::incr]: - Source: 'tests/src/demo.rs', lines 103:4-103:31 *) + Source: 'tests/src/demo.rs', lines 104:4-104:31 *) let counterUsize_incr (self : usize) : result (usize & usize) = let* self1 = usize_add self 1 in Ok (self, self1) (** Trait implementation: [demo::{(demo::Counter for usize)}] - Source: 'tests/src/demo.rs', lines 102:0-102:22 *) + Source: 'tests/src/demo.rs', lines 103:0-103:22 *) let counterUsize : counter_t usize = { incr = counterUsize_incr; } (** [demo::use_counter]: - Source: 'tests/src/demo.rs', lines 110:0-110:59 *) + Source: 'tests/src/demo.rs', lines 111:0-111:59 *) let use_counter (t : Type0) (counterInst : counter_t t) (cnt : t) : result (usize & t) = counterInst.incr cnt diff --git a/tests/fstar/misc/Bitwise.fst b/tests/fstar/misc/Bitwise.fst index 8d6bab58..7945e142 100644 --- a/tests/fstar/misc/Bitwise.fst +++ b/tests/fstar/misc/Bitwise.fst @@ -6,27 +6,27 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [bitwise::shift_u32]: - Source: 'tests/src/bitwise.rs', lines 4:0-4:31 *) + Source: 'tests/src/bitwise.rs', lines 5:0-5:31 *) let shift_u32 (a : u32) : result u32 = let* t = u32_shr #Usize a 16 in u32_shl #Usize t 16 (** [bitwise::shift_i32]: - Source: 'tests/src/bitwise.rs', lines 11:0-11:31 *) + Source: 'tests/src/bitwise.rs', lines 12:0-12:31 *) let shift_i32 (a : i32) : result i32 = let* t = i32_shr #Isize a 16 in i32_shl #Isize t 16 (** [bitwise::xor_u32]: - Source: 'tests/src/bitwise.rs', lines 18:0-18:37 *) + Source: 'tests/src/bitwise.rs', lines 19:0-19:37 *) let xor_u32 (a : u32) (b : u32) : result u32 = Ok (u32_xor a b) (** [bitwise::or_u32]: - Source: 'tests/src/bitwise.rs', lines 22:0-22:36 *) + Source: 'tests/src/bitwise.rs', lines 23:0-23:36 *) let or_u32 (a : u32) (b : u32) : result u32 = Ok (u32_or a b) (** [bitwise::and_u32]: - Source: 'tests/src/bitwise.rs', lines 26:0-26:37 *) + Source: 'tests/src/bitwise.rs', lines 27:0-27:37 *) let and_u32 (a : u32) (b : u32) : result u32 = Ok (u32_and a b) diff --git a/tests/fstar/misc/Constants.fst b/tests/fstar/misc/Constants.fst index 4ff5e883..1d72f5ff 100644 --- a/tests/fstar/misc/Constants.fst +++ b/tests/fstar/misc/Constants.fst @@ -6,154 +6,154 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [constants::X0] - Source: 'tests/src/constants.rs', lines 7:0-7:17 *) + Source: 'tests/src/constants.rs', lines 8:0-8:17 *) let x0_body : result u32 = Ok 0 let x0 : u32 = eval_global x0_body (** [constants::X1] - Source: 'tests/src/constants.rs', lines 9:0-9:17 *) + Source: 'tests/src/constants.rs', lines 10:0-10:17 *) let x1_body : result u32 = Ok core_u32_max let x1 : u32 = eval_global x1_body (** [constants::X2] - Source: 'tests/src/constants.rs', lines 12:0-12:17 *) + Source: 'tests/src/constants.rs', lines 13:0-13:17 *) let x2_body : result u32 = Ok 3 let x2 : u32 = eval_global x2_body (** [constants::incr]: - Source: 'tests/src/constants.rs', lines 19:0-19:32 *) + Source: 'tests/src/constants.rs', lines 20:0-20:32 *) let incr (n : u32) : result u32 = u32_add n 1 (** [constants::X3] - Source: 'tests/src/constants.rs', lines 17:0-17:17 *) + Source: 'tests/src/constants.rs', lines 18:0-18:17 *) let x3_body : result u32 = incr 32 let x3 : u32 = eval_global x3_body (** [constants::mk_pair0]: - Source: 'tests/src/constants.rs', lines 25:0-25:51 *) + Source: 'tests/src/constants.rs', lines 26:0-26:51 *) let mk_pair0 (x : u32) (y1 : u32) : result (u32 & u32) = Ok (x, y1) (** [constants::Pair] - Source: 'tests/src/constants.rs', lines 38:0-38:23 *) + Source: 'tests/src/constants.rs', lines 39:0-39:23 *) type pair_t (t1 t2 : Type0) = { x : t1; y : t2; } (** [constants::mk_pair1]: - Source: 'tests/src/constants.rs', lines 29:0-29:55 *) + Source: 'tests/src/constants.rs', lines 30:0-30:55 *) let mk_pair1 (x : u32) (y1 : u32) : result (pair_t u32 u32) = Ok { x = x; y = y1 } (** [constants::P0] - Source: 'tests/src/constants.rs', lines 33:0-33:24 *) + Source: 'tests/src/constants.rs', lines 34:0-34:24 *) let p0_body : result (u32 & u32) = mk_pair0 0 1 let p0 : (u32 & u32) = eval_global p0_body (** [constants::P1] - Source: 'tests/src/constants.rs', lines 34:0-34:28 *) + Source: 'tests/src/constants.rs', lines 35:0-35:28 *) let p1_body : result (pair_t u32 u32) = mk_pair1 0 1 let p1 : pair_t u32 u32 = eval_global p1_body (** [constants::P2] - Source: 'tests/src/constants.rs', lines 35:0-35:24 *) + Source: 'tests/src/constants.rs', lines 36:0-36:24 *) let p2_body : result (u32 & u32) = Ok (0, 1) let p2 : (u32 & u32) = eval_global p2_body (** [constants::P3] - Source: 'tests/src/constants.rs', lines 36:0-36:28 *) + Source: 'tests/src/constants.rs', lines 37:0-37:28 *) let p3_body : result (pair_t u32 u32) = Ok { x = 0; y = 1 } let p3 : pair_t u32 u32 = eval_global p3_body (** [constants::Wrap] - Source: 'tests/src/constants.rs', lines 51:0-51:18 *) + Source: 'tests/src/constants.rs', lines 52:0-52:18 *) type wrap_t (t : Type0) = { value : t; } (** [constants::{constants::Wrap}::new]: - Source: 'tests/src/constants.rs', lines 56:4-56:41 *) + Source: 'tests/src/constants.rs', lines 57:4-57:41 *) let wrap_new (t : Type0) (value : t) : result (wrap_t t) = Ok { value = value } (** [constants::Y] - Source: 'tests/src/constants.rs', lines 43:0-43:22 *) + Source: 'tests/src/constants.rs', lines 44:0-44:22 *) let y_body : result (wrap_t i32) = wrap_new i32 2 let y : wrap_t i32 = eval_global y_body (** [constants::unwrap_y]: - Source: 'tests/src/constants.rs', lines 45:0-45:30 *) + Source: 'tests/src/constants.rs', lines 46:0-46:30 *) let unwrap_y : result i32 = Ok y.value (** [constants::YVAL] - Source: 'tests/src/constants.rs', lines 49:0-49:19 *) + Source: 'tests/src/constants.rs', lines 50:0-50:19 *) let yval_body : result i32 = unwrap_y let yval : i32 = eval_global yval_body (** [constants::get_z1::Z1] - Source: 'tests/src/constants.rs', lines 64:4-64:17 *) + Source: 'tests/src/constants.rs', lines 65:4-65:17 *) let get_z1_z1_body : result i32 = Ok 3 let get_z1_z1 : i32 = eval_global get_z1_z1_body (** [constants::get_z1]: - Source: 'tests/src/constants.rs', lines 63:0-63:28 *) + Source: 'tests/src/constants.rs', lines 64:0-64:28 *) let get_z1 : result i32 = Ok get_z1_z1 (** [constants::add]: - Source: 'tests/src/constants.rs', lines 68:0-68:39 *) + Source: 'tests/src/constants.rs', lines 69:0-69:39 *) let add (a : i32) (b : i32) : result i32 = i32_add a b (** [constants::Q1] - Source: 'tests/src/constants.rs', lines 76:0-76:17 *) + Source: 'tests/src/constants.rs', lines 77:0-77:17 *) let q1_body : result i32 = Ok 5 let q1 : i32 = eval_global q1_body (** [constants::Q2] - Source: 'tests/src/constants.rs', lines 77:0-77:17 *) + Source: 'tests/src/constants.rs', lines 78:0-78:17 *) let q2_body : result i32 = Ok q1 let q2 : i32 = eval_global q2_body (** [constants::Q3] - Source: 'tests/src/constants.rs', lines 78:0-78:17 *) + Source: 'tests/src/constants.rs', lines 79:0-79:17 *) let q3_body : result i32 = add q2 3 let q3 : i32 = eval_global q3_body (** [constants::get_z2]: - Source: 'tests/src/constants.rs', lines 72:0-72:28 *) + Source: 'tests/src/constants.rs', lines 73:0-73:28 *) let get_z2 : result i32 = let* i = get_z1 in let* i1 = add i q3 in add q1 i1 (** [constants::S1] - Source: 'tests/src/constants.rs', lines 82:0-82:18 *) + Source: 'tests/src/constants.rs', lines 83:0-83:18 *) let s1_body : result u32 = Ok 6 let s1 : u32 = eval_global s1_body (** [constants::S2] - Source: 'tests/src/constants.rs', lines 83:0-83:18 *) + Source: 'tests/src/constants.rs', lines 84:0-84:18 *) let s2_body : result u32 = incr s1 let s2 : u32 = eval_global s2_body (** [constants::S3] - Source: 'tests/src/constants.rs', lines 84:0-84:29 *) + Source: 'tests/src/constants.rs', lines 85:0-85:29 *) let s3_body : result (pair_t u32 u32) = Ok p3 let s3 : pair_t u32 u32 = eval_global s3_body (** [constants::S4] - Source: 'tests/src/constants.rs', lines 85:0-85:29 *) + Source: 'tests/src/constants.rs', lines 86:0-86:29 *) let s4_body : result (pair_t u32 u32) = mk_pair1 7 8 let s4 : pair_t u32 u32 = eval_global s4_body (** [constants::V] - Source: 'tests/src/constants.rs', lines 88:0-88:31 *) + Source: 'tests/src/constants.rs', lines 89:0-89:31 *) type v_t (t : Type0) (n : usize) = { x : array t n; } (** [constants::{constants::V#1}::LEN] - Source: 'tests/src/constants.rs', lines 93:4-93:24 *) + Source: 'tests/src/constants.rs', lines 94:4-94:24 *) let v_len_body (t : Type0) (n : usize) : result usize = Ok n let v_len (t : Type0) (n : usize) : usize = eval_global (v_len_body t n) (** [constants::use_v]: - Source: 'tests/src/constants.rs', lines 96:0-96:42 *) + Source: 'tests/src/constants.rs', lines 97:0-97:42 *) let use_v (t : Type0) (n : usize) : result usize = Ok (v_len t n) diff --git a/tests/fstar/misc/External.Funs.fst b/tests/fstar/misc/External.Funs.fst index 18fc901f..f5e8d4b7 100644 --- a/tests/fstar/misc/External.Funs.fst +++ b/tests/fstar/misc/External.Funs.fst @@ -15,12 +15,12 @@ let core_marker_CopyU32 : core_marker_Copy_t u32 = { } (** [external::use_get]: - Source: 'tests/src/external.rs', lines 8:0-8:37 *) + Source: 'tests/src/external.rs', lines 9:0-9:37 *) let use_get (rc : core_cell_Cell_t u32) (st : state) : result (state & u32) = core_cell_Cell_get u32 core_marker_CopyU32 rc st (** [external::incr]: - Source: 'tests/src/external.rs', lines 12:0-12:31 *) + Source: 'tests/src/external.rs', lines 13:0-13:31 *) let incr (rc : core_cell_Cell_t u32) (st : state) : result (state & (core_cell_Cell_t u32)) diff --git a/tests/fstar/misc/Loops.Clauses.Template.fst b/tests/fstar/misc/Loops.Clauses.Template.fst index 7b042375..5d450275 100644 --- a/tests/fstar/misc/Loops.Clauses.Template.fst +++ b/tests/fstar/misc/Loops.Clauses.Template.fst @@ -7,144 +7,144 @@ open Loops.Types #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [loops::sum]: decreases clause - Source: 'tests/src/loops.rs', lines 7:0-17:1 *) + Source: 'tests/src/loops.rs', lines 8:0-18:1 *) unfold let sum_loop_decreases (max : u32) (i : u32) (s : u32) : nat = admit () (** [loops::sum_with_mut_borrows]: decreases clause - Source: 'tests/src/loops.rs', lines 22:0-34:1 *) + Source: 'tests/src/loops.rs', lines 23:0-35:1 *) unfold let sum_with_mut_borrows_loop_decreases (max : u32) (i : u32) (s : u32) : nat = admit () (** [loops::sum_with_shared_borrows]: decreases clause - Source: 'tests/src/loops.rs', lines 37:0-51:1 *) + Source: 'tests/src/loops.rs', lines 38:0-52:1 *) unfold let sum_with_shared_borrows_loop_decreases (max : u32) (i : u32) (s : u32) : nat = admit () (** [loops::sum_array]: decreases clause - Source: 'tests/src/loops.rs', lines 53:0-61:1 *) + Source: 'tests/src/loops.rs', lines 54:0-62:1 *) unfold let sum_array_loop_decreases (n : usize) (a : array u32 n) (i : usize) (s : u32) : nat = admit () (** [loops::clear]: decreases clause - Source: 'tests/src/loops.rs', lines 65:0-71:1 *) + Source: 'tests/src/loops.rs', lines 66:0-72:1 *) unfold let clear_loop_decreases (v : alloc_vec_Vec u32) (i : usize) : nat = admit () (** [loops::list_mem]: decreases clause - Source: 'tests/src/loops.rs', lines 79:0-88:1 *) + Source: 'tests/src/loops.rs', lines 80:0-89:1 *) unfold let list_mem_loop_decreases (x : u32) (ls : list_t u32) : nat = admit () (** [loops::list_nth_mut_loop]: decreases clause - Source: 'tests/src/loops.rs', lines 91:0-101:1 *) + Source: 'tests/src/loops.rs', lines 92:0-102:1 *) unfold let list_nth_mut_loop_loop_decreases (t : Type0) (ls : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_shared_loop]: decreases clause - Source: 'tests/src/loops.rs', lines 104:0-114:1 *) + Source: 'tests/src/loops.rs', lines 105:0-115:1 *) unfold let list_nth_shared_loop_loop_decreases (t : Type0) (ls : list_t t) (i : u32) : nat = admit () (** [loops::get_elem_mut]: decreases clause - Source: 'tests/src/loops.rs', lines 116:0-130:1 *) + Source: 'tests/src/loops.rs', lines 117:0-131:1 *) unfold let get_elem_mut_loop_decreases (x : usize) (ls : list_t usize) : nat = admit () (** [loops::get_elem_shared]: decreases clause - Source: 'tests/src/loops.rs', lines 132:0-146:1 *) + Source: 'tests/src/loops.rs', lines 133:0-147:1 *) unfold let get_elem_shared_loop_decreases (x : usize) (ls : list_t usize) : nat = admit () (** [loops::list_nth_mut_loop_with_id]: decreases clause - Source: 'tests/src/loops.rs', lines 157:0-168:1 *) + Source: 'tests/src/loops.rs', lines 158:0-169:1 *) unfold let list_nth_mut_loop_with_id_loop_decreases (t : Type0) (i : u32) (ls : list_t t) : nat = admit () (** [loops::list_nth_shared_loop_with_id]: decreases clause - Source: 'tests/src/loops.rs', lines 171:0-182:1 *) + Source: 'tests/src/loops.rs', lines 172:0-183:1 *) unfold let list_nth_shared_loop_with_id_loop_decreases (t : Type0) (i : u32) (ls : list_t t) : nat = admit () (** [loops::list_nth_mut_loop_pair]: decreases clause - Source: 'tests/src/loops.rs', lines 187:0-208:1 *) + Source: 'tests/src/loops.rs', lines 188:0-209:1 *) unfold let list_nth_mut_loop_pair_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_shared_loop_pair]: decreases clause - Source: 'tests/src/loops.rs', lines 211:0-232:1 *) + Source: 'tests/src/loops.rs', lines 212:0-233:1 *) unfold let list_nth_shared_loop_pair_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_mut_loop_pair_merge]: decreases clause - Source: 'tests/src/loops.rs', lines 236:0-251:1 *) + Source: 'tests/src/loops.rs', lines 237:0-252:1 *) unfold let list_nth_mut_loop_pair_merge_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_shared_loop_pair_merge]: decreases clause - Source: 'tests/src/loops.rs', lines 254:0-269:1 *) + Source: 'tests/src/loops.rs', lines 255:0-270:1 *) unfold let list_nth_shared_loop_pair_merge_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_mut_shared_loop_pair]: decreases clause - Source: 'tests/src/loops.rs', lines 272:0-287:1 *) + Source: 'tests/src/loops.rs', lines 273:0-288:1 *) unfold let list_nth_mut_shared_loop_pair_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_mut_shared_loop_pair_merge]: decreases clause - Source: 'tests/src/loops.rs', lines 291:0-306:1 *) + Source: 'tests/src/loops.rs', lines 292:0-307:1 *) unfold let list_nth_mut_shared_loop_pair_merge_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_shared_mut_loop_pair]: decreases clause - Source: 'tests/src/loops.rs', lines 310:0-325:1 *) + Source: 'tests/src/loops.rs', lines 311:0-326:1 *) unfold let list_nth_shared_mut_loop_pair_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::list_nth_shared_mut_loop_pair_merge]: decreases clause - Source: 'tests/src/loops.rs', lines 329:0-344:1 *) + Source: 'tests/src/loops.rs', lines 330:0-345:1 *) unfold let list_nth_shared_mut_loop_pair_merge_loop_decreases (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : nat = admit () (** [loops::ignore_input_mut_borrow]: decreases clause - Source: 'tests/src/loops.rs', lines 348:0-352:1 *) + Source: 'tests/src/loops.rs', lines 349:0-353:1 *) unfold let ignore_input_mut_borrow_loop_decreases (i : u32) : nat = admit () (** [loops::incr_ignore_input_mut_borrow]: decreases clause - Source: 'tests/src/loops.rs', lines 356:0-361:1 *) + Source: 'tests/src/loops.rs', lines 357:0-362:1 *) unfold let incr_ignore_input_mut_borrow_loop_decreases (i : u32) : nat = admit () (** [loops::ignore_input_shared_borrow]: decreases clause - Source: 'tests/src/loops.rs', lines 365:0-369:1 *) + Source: 'tests/src/loops.rs', lines 366:0-370:1 *) unfold let ignore_input_shared_borrow_loop_decreases (i : u32) : nat = admit () diff --git a/tests/fstar/misc/Loops.Funs.fst b/tests/fstar/misc/Loops.Funs.fst index 84e9634d..05e77046 100644 --- a/tests/fstar/misc/Loops.Funs.fst +++ b/tests/fstar/misc/Loops.Funs.fst @@ -8,7 +8,7 @@ include Loops.Clauses #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [loops::sum]: loop 0: - Source: 'tests/src/loops.rs', lines 7:0-17:1 *) + Source: 'tests/src/loops.rs', lines 8:0-18:1 *) let rec sum_loop (max : u32) (i : u32) (s : u32) : Tot (result u32) (decreases (sum_loop_decreases max i s)) @@ -18,12 +18,12 @@ let rec sum_loop else u32_mul s 2 (** [loops::sum]: - Source: 'tests/src/loops.rs', lines 7:0-7:27 *) + Source: 'tests/src/loops.rs', lines 8:0-8:27 *) let sum (max : u32) : result u32 = sum_loop max 0 0 (** [loops::sum_with_mut_borrows]: loop 0: - Source: 'tests/src/loops.rs', lines 22:0-34:1 *) + Source: 'tests/src/loops.rs', lines 23:0-35:1 *) let rec sum_with_mut_borrows_loop (max : u32) (i : u32) (s : u32) : Tot (result u32) (decreases (sum_with_mut_borrows_loop_decreases max i s)) @@ -36,12 +36,12 @@ let rec sum_with_mut_borrows_loop else u32_mul s 2 (** [loops::sum_with_mut_borrows]: - Source: 'tests/src/loops.rs', lines 22:0-22:44 *) + Source: 'tests/src/loops.rs', lines 23:0-23:44 *) let sum_with_mut_borrows (max : u32) : result u32 = sum_with_mut_borrows_loop max 0 0 (** [loops::sum_with_shared_borrows]: loop 0: - Source: 'tests/src/loops.rs', lines 37:0-51:1 *) + Source: 'tests/src/loops.rs', lines 38:0-52:1 *) let rec sum_with_shared_borrows_loop (max : u32) (i : u32) (s : u32) : Tot (result u32) (decreases (sum_with_shared_borrows_loop_decreases max i s)) @@ -54,12 +54,12 @@ let rec sum_with_shared_borrows_loop else u32_mul s 2 (** [loops::sum_with_shared_borrows]: - Source: 'tests/src/loops.rs', lines 37:0-37:47 *) + Source: 'tests/src/loops.rs', lines 38:0-38:47 *) let sum_with_shared_borrows (max : u32) : result u32 = sum_with_shared_borrows_loop max 0 0 (** [loops::sum_array]: loop 0: - Source: 'tests/src/loops.rs', lines 53:0-61:1 *) + Source: 'tests/src/loops.rs', lines 54:0-62:1 *) let rec sum_array_loop (n : usize) (a : array u32 n) (i : usize) (s : u32) : Tot (result u32) (decreases (sum_array_loop_decreases n a i s)) @@ -73,12 +73,12 @@ let rec sum_array_loop else Ok s (** [loops::sum_array]: - Source: 'tests/src/loops.rs', lines 53:0-53:52 *) + Source: 'tests/src/loops.rs', lines 54:0-54:52 *) let sum_array (n : usize) (a : array u32 n) : result u32 = sum_array_loop n a 0 0 (** [loops::clear]: loop 0: - Source: 'tests/src/loops.rs', lines 65:0-71:1 *) + Source: 'tests/src/loops.rs', lines 66:0-72:1 *) let rec clear_loop (v : alloc_vec_Vec u32) (i : usize) : Tot (result (alloc_vec_Vec u32)) (decreases (clear_loop_decreases v i)) @@ -95,12 +95,12 @@ let rec clear_loop else Ok v (** [loops::clear]: - Source: 'tests/src/loops.rs', lines 65:0-65:30 *) + Source: 'tests/src/loops.rs', lines 66:0-66:30 *) let clear (v : alloc_vec_Vec u32) : result (alloc_vec_Vec u32) = clear_loop v 0 (** [loops::list_mem]: loop 0: - Source: 'tests/src/loops.rs', lines 79:0-88:1 *) + Source: 'tests/src/loops.rs', lines 80:0-89:1 *) let rec list_mem_loop (x : u32) (ls : list_t u32) : Tot (result bool) (decreases (list_mem_loop_decreases x ls)) @@ -111,12 +111,12 @@ let rec list_mem_loop end (** [loops::list_mem]: - Source: 'tests/src/loops.rs', lines 79:0-79:52 *) + Source: 'tests/src/loops.rs', lines 80:0-80:52 *) let list_mem (x : u32) (ls : list_t u32) : result bool = list_mem_loop x ls (** [loops::list_nth_mut_loop]: loop 0: - Source: 'tests/src/loops.rs', lines 91:0-101:1 *) + Source: 'tests/src/loops.rs', lines 92:0-102:1 *) let rec list_nth_mut_loop_loop (t : Type0) (ls : list_t t) (i : u32) : Tot (result (t & (t -> result (list_t t)))) @@ -135,7 +135,7 @@ let rec list_nth_mut_loop_loop end (** [loops::list_nth_mut_loop]: - Source: 'tests/src/loops.rs', lines 91:0-91:71 *) + Source: 'tests/src/loops.rs', lines 92:0-92:71 *) let list_nth_mut_loop (t : Type0) (ls : list_t t) (i : u32) : result (t & (t -> result (list_t t))) @@ -143,7 +143,7 @@ let list_nth_mut_loop list_nth_mut_loop_loop t ls i (** [loops::list_nth_shared_loop]: loop 0: - Source: 'tests/src/loops.rs', lines 104:0-114:1 *) + Source: 'tests/src/loops.rs', lines 105:0-115:1 *) let rec list_nth_shared_loop_loop (t : Type0) (ls : list_t t) (i : u32) : Tot (result t) (decreases (list_nth_shared_loop_loop_decreases t ls i)) @@ -157,12 +157,12 @@ let rec list_nth_shared_loop_loop end (** [loops::list_nth_shared_loop]: - Source: 'tests/src/loops.rs', lines 104:0-104:66 *) + Source: 'tests/src/loops.rs', lines 105:0-105:66 *) let list_nth_shared_loop (t : Type0) (ls : list_t t) (i : u32) : result t = list_nth_shared_loop_loop t ls i (** [loops::get_elem_mut]: loop 0: - Source: 'tests/src/loops.rs', lines 116:0-130:1 *) + Source: 'tests/src/loops.rs', lines 117:0-131:1 *) let rec get_elem_mut_loop (x : usize) (ls : list_t usize) : Tot (result (usize & (usize -> result (list_t usize)))) @@ -180,7 +180,7 @@ let rec get_elem_mut_loop end (** [loops::get_elem_mut]: - Source: 'tests/src/loops.rs', lines 116:0-116:73 *) + Source: 'tests/src/loops.rs', lines 117:0-117:73 *) let get_elem_mut (slots : alloc_vec_Vec (list_t usize)) (x : usize) : result (usize & (usize -> result (alloc_vec_Vec (list_t usize)))) @@ -193,7 +193,7 @@ let get_elem_mut Ok (i, back1) (** [loops::get_elem_shared]: loop 0: - Source: 'tests/src/loops.rs', lines 132:0-146:1 *) + Source: 'tests/src/loops.rs', lines 133:0-147:1 *) let rec get_elem_shared_loop (x : usize) (ls : list_t usize) : Tot (result usize) (decreases (get_elem_shared_loop_decreases x ls)) @@ -204,7 +204,7 @@ let rec get_elem_shared_loop end (** [loops::get_elem_shared]: - Source: 'tests/src/loops.rs', lines 132:0-132:68 *) + Source: 'tests/src/loops.rs', lines 133:0-133:68 *) let get_elem_shared (slots : alloc_vec_Vec (list_t usize)) (x : usize) : result usize = let* ls = @@ -213,7 +213,7 @@ let get_elem_shared get_elem_shared_loop x ls (** [loops::id_mut]: - Source: 'tests/src/loops.rs', lines 148:0-148:50 *) + Source: 'tests/src/loops.rs', lines 149:0-149:50 *) let id_mut (t : Type0) (ls : list_t t) : result ((list_t t) & (list_t t -> result (list_t t))) @@ -221,12 +221,12 @@ let id_mut Ok (ls, Ok) (** [loops::id_shared]: - Source: 'tests/src/loops.rs', lines 152:0-152:45 *) + Source: 'tests/src/loops.rs', lines 153:0-153:45 *) let id_shared (t : Type0) (ls : list_t t) : result (list_t t) = Ok ls (** [loops::list_nth_mut_loop_with_id]: loop 0: - Source: 'tests/src/loops.rs', lines 157:0-168:1 *) + Source: 'tests/src/loops.rs', lines 158:0-169:1 *) let rec list_nth_mut_loop_with_id_loop (t : Type0) (i : u32) (ls : list_t t) : Tot (result (t & (t -> result (list_t t)))) @@ -245,7 +245,7 @@ let rec list_nth_mut_loop_with_id_loop end (** [loops::list_nth_mut_loop_with_id]: - Source: 'tests/src/loops.rs', lines 157:0-157:75 *) + Source: 'tests/src/loops.rs', lines 158:0-158:75 *) let list_nth_mut_loop_with_id (t : Type0) (ls : list_t t) (i : u32) : result (t & (t -> result (list_t t))) @@ -256,7 +256,7 @@ let list_nth_mut_loop_with_id Ok (x, back1) (** [loops::list_nth_shared_loop_with_id]: loop 0: - Source: 'tests/src/loops.rs', lines 171:0-182:1 *) + Source: 'tests/src/loops.rs', lines 172:0-183:1 *) let rec list_nth_shared_loop_with_id_loop (t : Type0) (i : u32) (ls : list_t t) : Tot (result t) @@ -271,13 +271,13 @@ let rec list_nth_shared_loop_with_id_loop end (** [loops::list_nth_shared_loop_with_id]: - Source: 'tests/src/loops.rs', lines 171:0-171:70 *) + Source: 'tests/src/loops.rs', lines 172:0-172:70 *) let list_nth_shared_loop_with_id (t : Type0) (ls : list_t t) (i : u32) : result t = let* ls1 = id_shared t ls in list_nth_shared_loop_with_id_loop t i ls1 (** [loops::list_nth_mut_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 187:0-208:1 *) + Source: 'tests/src/loops.rs', lines 188:0-209:1 *) let rec list_nth_mut_loop_pair_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result ((t & t) & (t -> result (list_t t)) & (t -> result (list_t t)))) @@ -306,7 +306,7 @@ let rec list_nth_mut_loop_pair_loop end (** [loops::list_nth_mut_loop_pair]: - Source: 'tests/src/loops.rs', lines 187:0-191:27 *) + Source: 'tests/src/loops.rs', lines 188:0-192:27 *) let list_nth_mut_loop_pair (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result ((t & t) & (t -> result (list_t t)) & (t -> result (list_t t))) @@ -314,7 +314,7 @@ let list_nth_mut_loop_pair list_nth_mut_loop_pair_loop t ls0 ls1 i (** [loops::list_nth_shared_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 211:0-232:1 *) + Source: 'tests/src/loops.rs', lines 212:0-233:1 *) let rec list_nth_shared_loop_pair_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result (t & t)) @@ -333,13 +333,13 @@ let rec list_nth_shared_loop_pair_loop end (** [loops::list_nth_shared_loop_pair]: - Source: 'tests/src/loops.rs', lines 211:0-215:19 *) + Source: 'tests/src/loops.rs', lines 212:0-216:19 *) let list_nth_shared_loop_pair (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result (t & t) = list_nth_shared_loop_pair_loop t ls0 ls1 i (** [loops::list_nth_mut_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 236:0-251:1 *) + Source: 'tests/src/loops.rs', lines 237:0-252:1 *) let rec list_nth_mut_loop_pair_merge_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result ((t & t) & ((t & t) -> result ((list_t t) & (list_t t))))) @@ -369,7 +369,7 @@ let rec list_nth_mut_loop_pair_merge_loop end (** [loops::list_nth_mut_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 236:0-240:27 *) + Source: 'tests/src/loops.rs', lines 237:0-241:27 *) let list_nth_mut_loop_pair_merge (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result ((t & t) & ((t & t) -> result ((list_t t) & (list_t t)))) @@ -377,7 +377,7 @@ let list_nth_mut_loop_pair_merge list_nth_mut_loop_pair_merge_loop t ls0 ls1 i (** [loops::list_nth_shared_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 254:0-269:1 *) + Source: 'tests/src/loops.rs', lines 255:0-270:1 *) let rec list_nth_shared_loop_pair_merge_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result (t & t)) @@ -398,13 +398,13 @@ let rec list_nth_shared_loop_pair_merge_loop end (** [loops::list_nth_shared_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 254:0-258:19 *) + Source: 'tests/src/loops.rs', lines 255:0-259:19 *) let list_nth_shared_loop_pair_merge (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result (t & t) = list_nth_shared_loop_pair_merge_loop t ls0 ls1 i (** [loops::list_nth_mut_shared_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 272:0-287:1 *) + Source: 'tests/src/loops.rs', lines 273:0-288:1 *) let rec list_nth_mut_shared_loop_pair_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result ((t & t) & (t -> result (list_t t)))) @@ -428,7 +428,7 @@ let rec list_nth_mut_shared_loop_pair_loop end (** [loops::list_nth_mut_shared_loop_pair]: - Source: 'tests/src/loops.rs', lines 272:0-276:23 *) + Source: 'tests/src/loops.rs', lines 273:0-277:23 *) let list_nth_mut_shared_loop_pair (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result ((t & t) & (t -> result (list_t t))) @@ -436,7 +436,7 @@ let list_nth_mut_shared_loop_pair list_nth_mut_shared_loop_pair_loop t ls0 ls1 i (** [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 291:0-306:1 *) + Source: 'tests/src/loops.rs', lines 292:0-307:1 *) let rec list_nth_mut_shared_loop_pair_merge_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result ((t & t) & (t -> result (list_t t)))) @@ -461,7 +461,7 @@ let rec list_nth_mut_shared_loop_pair_merge_loop end (** [loops::list_nth_mut_shared_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 291:0-295:23 *) + Source: 'tests/src/loops.rs', lines 292:0-296:23 *) let list_nth_mut_shared_loop_pair_merge (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result ((t & t) & (t -> result (list_t t))) @@ -469,7 +469,7 @@ let list_nth_mut_shared_loop_pair_merge list_nth_mut_shared_loop_pair_merge_loop t ls0 ls1 i (** [loops::list_nth_shared_mut_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 310:0-325:1 *) + Source: 'tests/src/loops.rs', lines 311:0-326:1 *) let rec list_nth_shared_mut_loop_pair_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result ((t & t) & (t -> result (list_t t)))) @@ -493,7 +493,7 @@ let rec list_nth_shared_mut_loop_pair_loop end (** [loops::list_nth_shared_mut_loop_pair]: - Source: 'tests/src/loops.rs', lines 310:0-314:23 *) + Source: 'tests/src/loops.rs', lines 311:0-315:23 *) let list_nth_shared_mut_loop_pair (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result ((t & t) & (t -> result (list_t t))) @@ -501,7 +501,7 @@ let list_nth_shared_mut_loop_pair list_nth_shared_mut_loop_pair_loop t ls0 ls1 i (** [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 329:0-344:1 *) + Source: 'tests/src/loops.rs', lines 330:0-345:1 *) let rec list_nth_shared_mut_loop_pair_merge_loop (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : Tot (result ((t & t) & (t -> result (list_t t)))) @@ -526,7 +526,7 @@ let rec list_nth_shared_mut_loop_pair_merge_loop end (** [loops::list_nth_shared_mut_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 329:0-333:23 *) + Source: 'tests/src/loops.rs', lines 330:0-334:23 *) let list_nth_shared_mut_loop_pair_merge (t : Type0) (ls0 : list_t t) (ls1 : list_t t) (i : u32) : result ((t & t) & (t -> result (list_t t))) @@ -534,7 +534,7 @@ let list_nth_shared_mut_loop_pair_merge list_nth_shared_mut_loop_pair_merge_loop t ls0 ls1 i (** [loops::ignore_input_mut_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 348:0-352:1 *) + Source: 'tests/src/loops.rs', lines 349:0-353:1 *) let rec ignore_input_mut_borrow_loop (i : u32) : Tot (result unit) (decreases (ignore_input_mut_borrow_loop_decreases i)) @@ -544,12 +544,12 @@ let rec ignore_input_mut_borrow_loop else Ok () (** [loops::ignore_input_mut_borrow]: - Source: 'tests/src/loops.rs', lines 348:0-348:56 *) + Source: 'tests/src/loops.rs', lines 349:0-349:56 *) let ignore_input_mut_borrow (_a : u32) (i : u32) : result u32 = let* _ = ignore_input_mut_borrow_loop i in Ok _a (** [loops::incr_ignore_input_mut_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 356:0-361:1 *) + Source: 'tests/src/loops.rs', lines 357:0-362:1 *) let rec incr_ignore_input_mut_borrow_loop (i : u32) : Tot (result unit) (decreases (incr_ignore_input_mut_borrow_loop_decreases i)) @@ -559,14 +559,14 @@ let rec incr_ignore_input_mut_borrow_loop else Ok () (** [loops::incr_ignore_input_mut_borrow]: - Source: 'tests/src/loops.rs', lines 356:0-356:60 *) + Source: 'tests/src/loops.rs', lines 357:0-357:60 *) let incr_ignore_input_mut_borrow (a : u32) (i : u32) : result u32 = let* a1 = u32_add a 1 in let* _ = incr_ignore_input_mut_borrow_loop i in Ok a1 (** [loops::ignore_input_shared_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 365:0-369:1 *) + Source: 'tests/src/loops.rs', lines 366:0-370:1 *) let rec ignore_input_shared_borrow_loop (i : u32) : Tot (result unit) (decreases (ignore_input_shared_borrow_loop_decreases i)) @@ -576,7 +576,7 @@ let rec ignore_input_shared_borrow_loop else Ok () (** [loops::ignore_input_shared_borrow]: - Source: 'tests/src/loops.rs', lines 365:0-365:59 *) + Source: 'tests/src/loops.rs', lines 366:0-366:59 *) let ignore_input_shared_borrow (_a : u32) (i : u32) : result u32 = let* _ = ignore_input_shared_borrow_loop i in Ok _a diff --git a/tests/fstar/misc/Loops.Types.fst b/tests/fstar/misc/Loops.Types.fst index c844d0e0..c9e2f1cf 100644 --- a/tests/fstar/misc/Loops.Types.fst +++ b/tests/fstar/misc/Loops.Types.fst @@ -6,7 +6,7 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [loops::List] - Source: 'tests/src/loops.rs', lines 73:0-73:16 *) + Source: 'tests/src/loops.rs', lines 74:0-74:16 *) type list_t (t : Type0) = | List_Cons : t -> list_t t -> list_t t | List_Nil : list_t t diff --git a/tests/fstar/misc/NoNestedBorrows.fst b/tests/fstar/misc/NoNestedBorrows.fst index b9fbd669..7e333b56 100644 --- a/tests/fstar/misc/NoNestedBorrows.fst +++ b/tests/fstar/misc/NoNestedBorrows.fst @@ -6,54 +6,54 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [no_nested_borrows::Pair] - Source: 'tests/src/no_nested_borrows.rs', lines 6:0-6:23 *) + Source: 'tests/src/no_nested_borrows.rs', lines 7:0-7:23 *) type pair_t (t1 t2 : Type0) = { x : t1; y : t2; } (** [no_nested_borrows::List] - Source: 'tests/src/no_nested_borrows.rs', lines 11:0-11:16 *) + Source: 'tests/src/no_nested_borrows.rs', lines 12:0-12:16 *) type list_t (t : Type0) = | List_Cons : t -> list_t t -> list_t t | List_Nil : list_t t (** [no_nested_borrows::One] - Source: 'tests/src/no_nested_borrows.rs', lines 22:0-22:16 *) + Source: 'tests/src/no_nested_borrows.rs', lines 23:0-23:16 *) type one_t (t1 : Type0) = | One_One : t1 -> one_t t1 (** [no_nested_borrows::EmptyEnum] - Source: 'tests/src/no_nested_borrows.rs', lines 28:0-28:18 *) + Source: 'tests/src/no_nested_borrows.rs', lines 29:0-29:18 *) type emptyEnum_t = | EmptyEnum_Empty : emptyEnum_t (** [no_nested_borrows::Enum] - Source: 'tests/src/no_nested_borrows.rs', lines 34:0-34:13 *) + Source: 'tests/src/no_nested_borrows.rs', lines 35:0-35:13 *) type enum_t = | Enum_Variant1 : enum_t | Enum_Variant2 : enum_t (** [no_nested_borrows::EmptyStruct] - Source: 'tests/src/no_nested_borrows.rs', lines 41:0-41:22 *) + Source: 'tests/src/no_nested_borrows.rs', lines 42:0-42:22 *) type emptyStruct_t = unit (** [no_nested_borrows::Sum] - Source: 'tests/src/no_nested_borrows.rs', lines 43:0-43:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 44:0-44:20 *) type sum_t (t1 t2 : Type0) = | Sum_Left : t1 -> sum_t t1 t2 | Sum_Right : t2 -> sum_t t1 t2 (** [no_nested_borrows::cast_u32_to_i32]: - Source: 'tests/src/no_nested_borrows.rs', lines 48:0-48:37 *) + Source: 'tests/src/no_nested_borrows.rs', lines 49:0-49:37 *) let cast_u32_to_i32 (x : u32) : result i32 = scalar_cast U32 I32 x (** [no_nested_borrows::cast_bool_to_i32]: - Source: 'tests/src/no_nested_borrows.rs', lines 52:0-52:39 *) + Source: 'tests/src/no_nested_borrows.rs', lines 53:0-53:39 *) let cast_bool_to_i32 (x : bool) : result i32 = scalar_cast_bool I32 x (** [no_nested_borrows::cast_bool_to_bool]: - Source: 'tests/src/no_nested_borrows.rs', lines 57:0-57:41 *) + Source: 'tests/src/no_nested_borrows.rs', lines 58:0-58:41 *) let cast_bool_to_bool (x : bool) : result bool = Ok x (** [no_nested_borrows::test2]: - Source: 'tests/src/no_nested_borrows.rs', lines 62:0-62:14 *) + Source: 'tests/src/no_nested_borrows.rs', lines 63:0-63:14 *) let test2 : result unit = let* _ = u32_add 23 44 in Ok () @@ -61,12 +61,12 @@ let test2 : result unit = let _ = assert_norm (test2 = Ok ()) (** [no_nested_borrows::get_max]: - Source: 'tests/src/no_nested_borrows.rs', lines 74:0-74:37 *) + Source: 'tests/src/no_nested_borrows.rs', lines 75:0-75:37 *) let get_max (x : u32) (y : u32) : result u32 = if x >= y then Ok x else Ok y (** [no_nested_borrows::test3]: - Source: 'tests/src/no_nested_borrows.rs', lines 82:0-82:14 *) + Source: 'tests/src/no_nested_borrows.rs', lines 83:0-83:14 *) let test3 : result unit = let* x = get_max 4 3 in let* y = get_max 10 11 in @@ -77,7 +77,7 @@ let test3 : result unit = let _ = assert_norm (test3 = Ok ()) (** [no_nested_borrows::test_neg1]: - Source: 'tests/src/no_nested_borrows.rs', lines 89:0-89:18 *) + Source: 'tests/src/no_nested_borrows.rs', lines 90:0-90:18 *) let test_neg1 : result unit = let* y = i32_neg 3 in if not (y = -3) then Fail Failure else Ok () @@ -85,7 +85,7 @@ let test_neg1 : result unit = let _ = assert_norm (test_neg1 = Ok ()) (** [no_nested_borrows::refs_test1]: - Source: 'tests/src/no_nested_borrows.rs', lines 96:0-96:19 *) + Source: 'tests/src/no_nested_borrows.rs', lines 97:0-97:19 *) let refs_test1 : result unit = if not (1 = 1) then Fail Failure else Ok () @@ -93,7 +93,7 @@ let refs_test1 : result unit = let _ = assert_norm (refs_test1 = Ok ()) (** [no_nested_borrows::refs_test2]: - Source: 'tests/src/no_nested_borrows.rs', lines 107:0-107:19 *) + Source: 'tests/src/no_nested_borrows.rs', lines 108:0-108:19 *) let refs_test2 : result unit = if not (2 = 2) then Fail Failure @@ -109,7 +109,7 @@ let refs_test2 : result unit = let _ = assert_norm (refs_test2 = Ok ()) (** [no_nested_borrows::test_list1]: - Source: 'tests/src/no_nested_borrows.rs', lines 123:0-123:19 *) + Source: 'tests/src/no_nested_borrows.rs', lines 124:0-124:19 *) let test_list1 : result unit = Ok () @@ -117,7 +117,7 @@ let test_list1 : result unit = let _ = assert_norm (test_list1 = Ok ()) (** [no_nested_borrows::test_box1]: - Source: 'tests/src/no_nested_borrows.rs', lines 128:0-128:18 *) + Source: 'tests/src/no_nested_borrows.rs', lines 129:0-129:18 *) let test_box1 : result unit = let* (_, deref_mut_back) = alloc_boxed_Box_deref_mut i32 0 in let* b = deref_mut_back 1 in @@ -128,22 +128,22 @@ let test_box1 : result unit = let _ = assert_norm (test_box1 = Ok ()) (** [no_nested_borrows::copy_int]: - Source: 'tests/src/no_nested_borrows.rs', lines 138:0-138:30 *) + Source: 'tests/src/no_nested_borrows.rs', lines 139:0-139:30 *) let copy_int (x : i32) : result i32 = Ok x (** [no_nested_borrows::test_unreachable]: - Source: 'tests/src/no_nested_borrows.rs', lines 144:0-144:32 *) + Source: 'tests/src/no_nested_borrows.rs', lines 145:0-145:32 *) let test_unreachable (b : bool) : result unit = if b then Fail Failure else Ok () (** [no_nested_borrows::test_panic]: - Source: 'tests/src/no_nested_borrows.rs', lines 152:0-152:26 *) + Source: 'tests/src/no_nested_borrows.rs', lines 153:0-153:26 *) let test_panic (b : bool) : result unit = if b then Fail Failure else Ok () (** [no_nested_borrows::test_copy_int]: - Source: 'tests/src/no_nested_borrows.rs', lines 159:0-159:22 *) + Source: 'tests/src/no_nested_borrows.rs', lines 160:0-160:22 *) let test_copy_int : result unit = let* y = copy_int 0 in if not (0 = y) then Fail Failure else Ok () @@ -151,12 +151,12 @@ let test_copy_int : result unit = let _ = assert_norm (test_copy_int = Ok ()) (** [no_nested_borrows::is_cons]: - Source: 'tests/src/no_nested_borrows.rs', lines 166:0-166:38 *) + Source: 'tests/src/no_nested_borrows.rs', lines 167:0-167:38 *) let is_cons (t : Type0) (l : list_t t) : result bool = begin match l with | List_Cons _ _ -> Ok true | List_Nil -> Ok false end (** [no_nested_borrows::test_is_cons]: - Source: 'tests/src/no_nested_borrows.rs', lines 173:0-173:21 *) + Source: 'tests/src/no_nested_borrows.rs', lines 174:0-174:21 *) let test_is_cons : result unit = let* b = is_cons i32 (List_Cons 0 List_Nil) in if not b then Fail Failure else Ok () @@ -165,7 +165,7 @@ let test_is_cons : result unit = let _ = assert_norm (test_is_cons = Ok ()) (** [no_nested_borrows::split_list]: - Source: 'tests/src/no_nested_borrows.rs', lines 179:0-179:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 180:0-180:48 *) let split_list (t : Type0) (l : list_t t) : result (t & (list_t t)) = begin match l with | List_Cons hd tl -> Ok (hd, tl) @@ -173,7 +173,7 @@ let split_list (t : Type0) (l : list_t t) : result (t & (list_t t)) = end (** [no_nested_borrows::test_split_list]: - Source: 'tests/src/no_nested_borrows.rs', lines 187:0-187:24 *) + Source: 'tests/src/no_nested_borrows.rs', lines 188:0-188:24 *) let test_split_list : result unit = let* p = split_list i32 (List_Cons 0 List_Nil) in let (hd, _) = p in @@ -183,7 +183,7 @@ let test_split_list : result unit = let _ = assert_norm (test_split_list = Ok ()) (** [no_nested_borrows::choose]: - Source: 'tests/src/no_nested_borrows.rs', lines 194:0-194:70 *) + Source: 'tests/src/no_nested_borrows.rs', lines 195:0-195:70 *) let choose (t : Type0) (b : bool) (x : t) (y : t) : result (t & (t -> result (t & t))) = if b @@ -191,7 +191,7 @@ let choose else let back = fun ret -> Ok (x, ret) in Ok (y, back) (** [no_nested_borrows::choose_test]: - Source: 'tests/src/no_nested_borrows.rs', lines 202:0-202:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 203:0-203:20 *) let choose_test : result unit = let* (z, choose_back) = choose i32 true 0 0 in let* z1 = i32_add z 1 in @@ -207,24 +207,24 @@ let choose_test : result unit = let _ = assert_norm (choose_test = Ok ()) (** [no_nested_borrows::test_char]: - Source: 'tests/src/no_nested_borrows.rs', lines 214:0-214:26 *) + Source: 'tests/src/no_nested_borrows.rs', lines 215:0-215:26 *) let test_char : result char = Ok 'a' (** [no_nested_borrows::Tree] - Source: 'tests/src/no_nested_borrows.rs', lines 219:0-219:16 *) + Source: 'tests/src/no_nested_borrows.rs', lines 220:0-220:16 *) type tree_t (t : Type0) = | Tree_Leaf : t -> tree_t t | Tree_Node : t -> nodeElem_t t -> tree_t t -> tree_t t (** [no_nested_borrows::NodeElem] - Source: 'tests/src/no_nested_borrows.rs', lines 224:0-224:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 225:0-225:20 *) and nodeElem_t (t : Type0) = | NodeElem_Cons : tree_t t -> nodeElem_t t -> nodeElem_t t | NodeElem_Nil : nodeElem_t t (** [no_nested_borrows::list_length]: - Source: 'tests/src/no_nested_borrows.rs', lines 259:0-259:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 260:0-260:48 *) let rec list_length (t : Type0) (l : list_t t) : result u32 = begin match l with | List_Cons _ l1 -> let* i = list_length t l1 in u32_add 1 i @@ -232,7 +232,7 @@ let rec list_length (t : Type0) (l : list_t t) : result u32 = end (** [no_nested_borrows::list_nth_shared]: - Source: 'tests/src/no_nested_borrows.rs', lines 267:0-267:62 *) + Source: 'tests/src/no_nested_borrows.rs', lines 268:0-268:62 *) let rec list_nth_shared (t : Type0) (l : list_t t) (i : u32) : result t = begin match l with | List_Cons x tl -> @@ -241,7 +241,7 @@ let rec list_nth_shared (t : Type0) (l : list_t t) (i : u32) : result t = end (** [no_nested_borrows::list_nth_mut]: - Source: 'tests/src/no_nested_borrows.rs', lines 283:0-283:67 *) + Source: 'tests/src/no_nested_borrows.rs', lines 284:0-284:67 *) let rec list_nth_mut (t : Type0) (l : list_t t) (i : u32) : result (t & (t -> result (list_t t))) @@ -260,7 +260,7 @@ let rec list_nth_mut end (** [no_nested_borrows::list_rev_aux]: - Source: 'tests/src/no_nested_borrows.rs', lines 299:0-299:63 *) + Source: 'tests/src/no_nested_borrows.rs', lines 300:0-300:63 *) let rec list_rev_aux (t : Type0) (li : list_t t) (lo : list_t t) : result (list_t t) = begin match li with @@ -269,13 +269,13 @@ let rec list_rev_aux end (** [no_nested_borrows::list_rev]: - Source: 'tests/src/no_nested_borrows.rs', lines 313:0-313:42 *) + Source: 'tests/src/no_nested_borrows.rs', lines 314:0-314:42 *) let list_rev (t : Type0) (l : list_t t) : result (list_t t) = let (li, _) = core_mem_replace (list_t t) l List_Nil in list_rev_aux t li List_Nil (** [no_nested_borrows::test_list_functions]: - Source: 'tests/src/no_nested_borrows.rs', lines 318:0-318:28 *) + Source: 'tests/src/no_nested_borrows.rs', lines 319:0-319:28 *) let test_list_functions : result unit = let l = List_Cons 2 List_Nil in let l1 = List_Cons 1 l in @@ -312,7 +312,7 @@ let test_list_functions : result unit = let _ = assert_norm (test_list_functions = Ok ()) (** [no_nested_borrows::id_mut_pair1]: - Source: 'tests/src/no_nested_borrows.rs', lines 334:0-334:89 *) + Source: 'tests/src/no_nested_borrows.rs', lines 335:0-335:89 *) let id_mut_pair1 (t1 t2 : Type0) (x : t1) (y : t2) : result ((t1 & t2) & ((t1 & t2) -> result (t1 & t2))) @@ -320,7 +320,7 @@ let id_mut_pair1 Ok ((x, y), Ok) (** [no_nested_borrows::id_mut_pair2]: - Source: 'tests/src/no_nested_borrows.rs', lines 338:0-338:88 *) + Source: 'tests/src/no_nested_borrows.rs', lines 339:0-339:88 *) let id_mut_pair2 (t1 t2 : Type0) (p : (t1 & t2)) : result ((t1 & t2) & ((t1 & t2) -> result (t1 & t2))) @@ -328,7 +328,7 @@ let id_mut_pair2 let (x, x1) = p in Ok ((x, x1), Ok) (** [no_nested_borrows::id_mut_pair3]: - Source: 'tests/src/no_nested_borrows.rs', lines 342:0-342:93 *) + Source: 'tests/src/no_nested_borrows.rs', lines 343:0-343:93 *) let id_mut_pair3 (t1 t2 : Type0) (x : t1) (y : t2) : result ((t1 & t2) & (t1 -> result t1) & (t2 -> result t2)) @@ -336,7 +336,7 @@ let id_mut_pair3 Ok ((x, y), Ok, Ok) (** [no_nested_borrows::id_mut_pair4]: - Source: 'tests/src/no_nested_borrows.rs', lines 346:0-346:92 *) + Source: 'tests/src/no_nested_borrows.rs', lines 347:0-347:92 *) let id_mut_pair4 (t1 t2 : Type0) (p : (t1 & t2)) : result ((t1 & t2) & (t1 -> result t1) & (t2 -> result t2)) @@ -344,35 +344,35 @@ let id_mut_pair4 let (x, x1) = p in Ok ((x, x1), Ok, Ok) (** [no_nested_borrows::StructWithTuple] - Source: 'tests/src/no_nested_borrows.rs', lines 353:0-353:34 *) + Source: 'tests/src/no_nested_borrows.rs', lines 354:0-354:34 *) type structWithTuple_t (t1 t2 : Type0) = { p : (t1 & t2); } (** [no_nested_borrows::new_tuple1]: - Source: 'tests/src/no_nested_borrows.rs', lines 357:0-357:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 358:0-358:48 *) let new_tuple1 : result (structWithTuple_t u32 u32) = Ok { p = (1, 2) } (** [no_nested_borrows::new_tuple2]: - Source: 'tests/src/no_nested_borrows.rs', lines 361:0-361:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 362:0-362:48 *) let new_tuple2 : result (structWithTuple_t i16 i16) = Ok { p = (1, 2) } (** [no_nested_borrows::new_tuple3]: - Source: 'tests/src/no_nested_borrows.rs', lines 365:0-365:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 366:0-366:48 *) let new_tuple3 : result (structWithTuple_t u64 i64) = Ok { p = (1, 2) } (** [no_nested_borrows::StructWithPair] - Source: 'tests/src/no_nested_borrows.rs', lines 370:0-370:33 *) + Source: 'tests/src/no_nested_borrows.rs', lines 371:0-371:33 *) type structWithPair_t (t1 t2 : Type0) = { p : pair_t t1 t2; } (** [no_nested_borrows::new_pair1]: - Source: 'tests/src/no_nested_borrows.rs', lines 374:0-374:46 *) + Source: 'tests/src/no_nested_borrows.rs', lines 375:0-375:46 *) let new_pair1 : result (structWithPair_t u32 u32) = Ok { p = { x = 1; y = 2 } } (** [no_nested_borrows::test_constants]: - Source: 'tests/src/no_nested_borrows.rs', lines 382:0-382:23 *) + Source: 'tests/src/no_nested_borrows.rs', lines 383:0-383:23 *) let test_constants : result unit = let* swt = new_tuple1 in let (i, _) = swt.p in @@ -396,7 +396,7 @@ let test_constants : result unit = let _ = assert_norm (test_constants = Ok ()) (** [no_nested_borrows::test_weird_borrows1]: - Source: 'tests/src/no_nested_borrows.rs', lines 391:0-391:28 *) + Source: 'tests/src/no_nested_borrows.rs', lines 392:0-392:28 *) let test_weird_borrows1 : result unit = Ok () @@ -404,71 +404,71 @@ let test_weird_borrows1 : result unit = let _ = assert_norm (test_weird_borrows1 = Ok ()) (** [no_nested_borrows::test_mem_replace]: - Source: 'tests/src/no_nested_borrows.rs', lines 401:0-401:37 *) + Source: 'tests/src/no_nested_borrows.rs', lines 402:0-402:37 *) let test_mem_replace (px : u32) : result u32 = let (y, _) = core_mem_replace u32 px 1 in if not (y = 0) then Fail Failure else Ok 2 (** [no_nested_borrows::test_shared_borrow_bool1]: - Source: 'tests/src/no_nested_borrows.rs', lines 408:0-408:47 *) + Source: 'tests/src/no_nested_borrows.rs', lines 409:0-409:47 *) let test_shared_borrow_bool1 (b : bool) : result u32 = if b then Ok 0 else Ok 1 (** [no_nested_borrows::test_shared_borrow_bool2]: - Source: 'tests/src/no_nested_borrows.rs', lines 421:0-421:40 *) + Source: 'tests/src/no_nested_borrows.rs', lines 422:0-422:40 *) let test_shared_borrow_bool2 : result u32 = Ok 0 (** [no_nested_borrows::test_shared_borrow_enum1]: - Source: 'tests/src/no_nested_borrows.rs', lines 436:0-436:52 *) + Source: 'tests/src/no_nested_borrows.rs', lines 437:0-437:52 *) let test_shared_borrow_enum1 (l : list_t u32) : result u32 = begin match l with | List_Cons _ _ -> Ok 1 | List_Nil -> Ok 0 end (** [no_nested_borrows::test_shared_borrow_enum2]: - Source: 'tests/src/no_nested_borrows.rs', lines 448:0-448:40 *) + Source: 'tests/src/no_nested_borrows.rs', lines 449:0-449:40 *) let test_shared_borrow_enum2 : result u32 = Ok 0 (** [no_nested_borrows::incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 459:0-459:24 *) + Source: 'tests/src/no_nested_borrows.rs', lines 460:0-460:24 *) let incr (x : u32) : result u32 = u32_add x 1 (** [no_nested_borrows::call_incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 463:0-463:35 *) + Source: 'tests/src/no_nested_borrows.rs', lines 464:0-464:35 *) let call_incr (x : u32) : result u32 = incr x (** [no_nested_borrows::read_then_incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 468:0-468:41 *) + Source: 'tests/src/no_nested_borrows.rs', lines 469:0-469:41 *) let read_then_incr (x : u32) : result (u32 & u32) = let* x1 = u32_add x 1 in Ok (x, x1) (** [no_nested_borrows::Tuple] - Source: 'tests/src/no_nested_borrows.rs', lines 474:0-474:24 *) + Source: 'tests/src/no_nested_borrows.rs', lines 475:0-475:24 *) type tuple_t (t1 t2 : Type0) = t1 * t2 (** [no_nested_borrows::use_tuple_struct]: - Source: 'tests/src/no_nested_borrows.rs', lines 476:0-476:48 *) + Source: 'tests/src/no_nested_borrows.rs', lines 477:0-477:48 *) let use_tuple_struct (x : tuple_t u32 u32) : result (tuple_t u32 u32) = let (_, i) = x in Ok (1, i) (** [no_nested_borrows::create_tuple_struct]: - Source: 'tests/src/no_nested_borrows.rs', lines 480:0-480:61 *) + Source: 'tests/src/no_nested_borrows.rs', lines 481:0-481:61 *) let create_tuple_struct (x : u32) (y : u64) : result (tuple_t u32 u64) = Ok (x, y) (** [no_nested_borrows::IdType] - Source: 'tests/src/no_nested_borrows.rs', lines 485:0-485:20 *) + Source: 'tests/src/no_nested_borrows.rs', lines 486:0-486:20 *) type idType_t (t : Type0) = t (** [no_nested_borrows::use_id_type]: - Source: 'tests/src/no_nested_borrows.rs', lines 487:0-487:40 *) + Source: 'tests/src/no_nested_borrows.rs', lines 488:0-488:40 *) let use_id_type (t : Type0) (x : idType_t t) : result t = Ok x (** [no_nested_borrows::create_id_type]: - Source: 'tests/src/no_nested_borrows.rs', lines 491:0-491:43 *) + Source: 'tests/src/no_nested_borrows.rs', lines 492:0-492:43 *) let create_id_type (t : Type0) (x : t) : result (idType_t t) = Ok x diff --git a/tests/fstar/misc/Paper.fst b/tests/fstar/misc/Paper.fst index c78293f1..e2412076 100644 --- a/tests/fstar/misc/Paper.fst +++ b/tests/fstar/misc/Paper.fst @@ -6,12 +6,12 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [paper::ref_incr]: - Source: 'tests/src/paper.rs', lines 6:0-6:28 *) + Source: 'tests/src/paper.rs', lines 7:0-7:28 *) let ref_incr (x : i32) : result i32 = i32_add x 1 (** [paper::test_incr]: - Source: 'tests/src/paper.rs', lines 10:0-10:18 *) + Source: 'tests/src/paper.rs', lines 11:0-11:18 *) let test_incr : result unit = let* x = ref_incr 0 in if not (x = 1) then Fail Failure else Ok () @@ -19,7 +19,7 @@ let test_incr : result unit = let _ = assert_norm (test_incr = Ok ()) (** [paper::choose]: - Source: 'tests/src/paper.rs', lines 17:0-17:70 *) + Source: 'tests/src/paper.rs', lines 18:0-18:70 *) let choose (t : Type0) (b : bool) (x : t) (y : t) : result (t & (t -> result (t & t))) = if b @@ -27,7 +27,7 @@ let choose else let back = fun ret -> Ok (x, ret) in Ok (y, back) (** [paper::test_choose]: - Source: 'tests/src/paper.rs', lines 25:0-25:20 *) + Source: 'tests/src/paper.rs', lines 26:0-26:20 *) let test_choose : result unit = let* (z, choose_back) = choose i32 true 0 0 in let* z1 = i32_add z 1 in @@ -43,13 +43,13 @@ let test_choose : result unit = let _ = assert_norm (test_choose = Ok ()) (** [paper::List] - Source: 'tests/src/paper.rs', lines 37:0-37:16 *) + Source: 'tests/src/paper.rs', lines 38:0-38:16 *) type list_t (t : Type0) = | List_Cons : t -> list_t t -> list_t t | List_Nil : list_t t (** [paper::list_nth_mut]: - Source: 'tests/src/paper.rs', lines 44:0-44:67 *) + Source: 'tests/src/paper.rs', lines 45:0-45:67 *) let rec list_nth_mut (t : Type0) (l : list_t t) (i : u32) : result (t & (t -> result (list_t t))) @@ -68,7 +68,7 @@ let rec list_nth_mut end (** [paper::sum]: - Source: 'tests/src/paper.rs', lines 59:0-59:32 *) + Source: 'tests/src/paper.rs', lines 60:0-60:32 *) let rec sum (l : list_t i32) : result i32 = begin match l with | List_Cons x tl -> let* i = sum tl in i32_add x i @@ -76,7 +76,7 @@ let rec sum (l : list_t i32) : result i32 = end (** [paper::test_nth]: - Source: 'tests/src/paper.rs', lines 70:0-70:17 *) + Source: 'tests/src/paper.rs', lines 71:0-71:17 *) let test_nth : result unit = let l = List_Cons 3 List_Nil in let l1 = List_Cons 2 l in @@ -90,7 +90,7 @@ let test_nth : result unit = let _ = assert_norm (test_nth = Ok ()) (** [paper::call_choose]: - Source: 'tests/src/paper.rs', lines 78:0-78:44 *) + Source: 'tests/src/paper.rs', lines 79:0-79:44 *) let call_choose (p : (u32 & u32)) : result u32 = let (px, py) = p in let* (pz, choose_back) = choose u32 true px py in diff --git a/tests/fstar/misc/PoloniusList.fst b/tests/fstar/misc/PoloniusList.fst index 94c1b445..810388db 100644 --- a/tests/fstar/misc/PoloniusList.fst +++ b/tests/fstar/misc/PoloniusList.fst @@ -6,13 +6,13 @@ open Primitives #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [polonius_list::List] - Source: 'tests/src/polonius_list.rs', lines 5:0-5:16 *) + Source: 'tests/src/polonius_list.rs', lines 6:0-6:16 *) type list_t (t : Type0) = | List_Cons : t -> list_t t -> list_t t | List_Nil : list_t t (** [polonius_list::get_list_at_x]: - Source: 'tests/src/polonius_list.rs', lines 15:0-15:76 *) + Source: 'tests/src/polonius_list.rs', lines 16:0-16:76 *) let rec get_list_at_x (ls : list_t u32) (x : u32) : result ((list_t u32) & (list_t u32 -> result (list_t u32))) diff --git a/tests/lean/Bitwise.lean b/tests/lean/Bitwise.lean index 3c381233..23ec66b4 100644 --- a/tests/lean/Bitwise.lean +++ b/tests/lean/Bitwise.lean @@ -6,31 +6,31 @@ open Primitives namespace bitwise /- [bitwise::shift_u32]: - Source: 'tests/src/bitwise.rs', lines 4:0-4:31 -/ + Source: 'tests/src/bitwise.rs', lines 5:0-5:31 -/ def shift_u32 (a : U32) : Result U32 := do let t ← a >>> 16#usize t <<< 16#usize /- [bitwise::shift_i32]: - Source: 'tests/src/bitwise.rs', lines 11:0-11:31 -/ + Source: 'tests/src/bitwise.rs', lines 12:0-12:31 -/ def shift_i32 (a : I32) : Result I32 := do let t ← a >>> 16#isize t <<< 16#isize /- [bitwise::xor_u32]: - Source: 'tests/src/bitwise.rs', lines 18:0-18:37 -/ + Source: 'tests/src/bitwise.rs', lines 19:0-19:37 -/ def xor_u32 (a : U32) (b : U32) : Result U32 := Result.ok (a ^^^ b) /- [bitwise::or_u32]: - Source: 'tests/src/bitwise.rs', lines 22:0-22:36 -/ + Source: 'tests/src/bitwise.rs', lines 23:0-23:36 -/ def or_u32 (a : U32) (b : U32) : Result U32 := Result.ok (a ||| b) /- [bitwise::and_u32]: - Source: 'tests/src/bitwise.rs', lines 26:0-26:37 -/ + Source: 'tests/src/bitwise.rs', lines 27:0-27:37 -/ def and_u32 (a : U32) (b : U32) : Result U32 := Result.ok (a &&& b) diff --git a/tests/lean/Constants.lean b/tests/lean/Constants.lean index 334f4707..ecb91c16 100644 --- a/tests/lean/Constants.lean +++ b/tests/lean/Constants.lean @@ -6,123 +6,123 @@ open Primitives namespace constants /- [constants::X0] - Source: 'tests/src/constants.rs', lines 7:0-7:17 -/ + Source: 'tests/src/constants.rs', lines 8:0-8:17 -/ def X0_body : Result U32 := Result.ok 0#u32 def X0 : U32 := eval_global X0_body /- [constants::X1] - Source: 'tests/src/constants.rs', lines 9:0-9:17 -/ + Source: 'tests/src/constants.rs', lines 10:0-10:17 -/ def X1_body : Result U32 := Result.ok core_u32_max def X1 : U32 := eval_global X1_body /- [constants::X2] - Source: 'tests/src/constants.rs', lines 12:0-12:17 -/ + Source: 'tests/src/constants.rs', lines 13:0-13:17 -/ def X2_body : Result U32 := Result.ok 3#u32 def X2 : U32 := eval_global X2_body /- [constants::incr]: - Source: 'tests/src/constants.rs', lines 19:0-19:32 -/ + Source: 'tests/src/constants.rs', lines 20:0-20:32 -/ def incr (n : U32) : Result U32 := n + 1#u32 /- [constants::X3] - Source: 'tests/src/constants.rs', lines 17:0-17:17 -/ + Source: 'tests/src/constants.rs', lines 18:0-18:17 -/ def X3_body : Result U32 := incr 32#u32 def X3 : U32 := eval_global X3_body /- [constants::mk_pair0]: - Source: 'tests/src/constants.rs', lines 25:0-25:51 -/ + Source: 'tests/src/constants.rs', lines 26:0-26:51 -/ def mk_pair0 (x : U32) (y : U32) : Result (U32 × U32) := Result.ok (x, y) /- [constants::Pair] - Source: 'tests/src/constants.rs', lines 38:0-38:23 -/ + Source: 'tests/src/constants.rs', lines 39:0-39:23 -/ structure Pair (T1 T2 : Type) where x : T1 y : T2 /- [constants::mk_pair1]: - Source: 'tests/src/constants.rs', lines 29:0-29:55 -/ + Source: 'tests/src/constants.rs', lines 30:0-30:55 -/ def mk_pair1 (x : U32) (y : U32) : Result (Pair U32 U32) := Result.ok { x := x, y := y } /- [constants::P0] - Source: 'tests/src/constants.rs', lines 33:0-33:24 -/ + Source: 'tests/src/constants.rs', lines 34:0-34:24 -/ def P0_body : Result (U32 × U32) := mk_pair0 0#u32 1#u32 def P0 : (U32 × U32) := eval_global P0_body /- [constants::P1] - Source: 'tests/src/constants.rs', lines 34:0-34:28 -/ + Source: 'tests/src/constants.rs', lines 35:0-35:28 -/ def P1_body : Result (Pair U32 U32) := mk_pair1 0#u32 1#u32 def P1 : Pair U32 U32 := eval_global P1_body /- [constants::P2] - Source: 'tests/src/constants.rs', lines 35:0-35:24 -/ + Source: 'tests/src/constants.rs', lines 36:0-36:24 -/ def P2_body : Result (U32 × U32) := Result.ok (0#u32, 1#u32) def P2 : (U32 × U32) := eval_global P2_body /- [constants::P3] - Source: 'tests/src/constants.rs', lines 36:0-36:28 -/ + Source: 'tests/src/constants.rs', lines 37:0-37:28 -/ def P3_body : Result (Pair U32 U32) := Result.ok { x := 0#u32, y := 1#u32 } def P3 : Pair U32 U32 := eval_global P3_body /- [constants::Wrap] - Source: 'tests/src/constants.rs', lines 51:0-51:18 -/ + Source: 'tests/src/constants.rs', lines 52:0-52:18 -/ structure Wrap (T : Type) where value : T /- [constants::{constants::Wrap}::new]: - Source: 'tests/src/constants.rs', lines 56:4-56:41 -/ + Source: 'tests/src/constants.rs', lines 57:4-57:41 -/ def Wrap.new (T : Type) (value : T) : Result (Wrap T) := Result.ok { value := value } /- [constants::Y] - Source: 'tests/src/constants.rs', lines 43:0-43:22 -/ + Source: 'tests/src/constants.rs', lines 44:0-44:22 -/ def Y_body : Result (Wrap I32) := Wrap.new I32 2#i32 def Y : Wrap I32 := eval_global Y_body /- [constants::unwrap_y]: - Source: 'tests/src/constants.rs', lines 45:0-45:30 -/ + Source: 'tests/src/constants.rs', lines 46:0-46:30 -/ def unwrap_y : Result I32 := Result.ok Y.value /- [constants::YVAL] - Source: 'tests/src/constants.rs', lines 49:0-49:19 -/ + Source: 'tests/src/constants.rs', lines 50:0-50:19 -/ def YVAL_body : Result I32 := unwrap_y def YVAL : I32 := eval_global YVAL_body /- [constants::get_z1::Z1] - Source: 'tests/src/constants.rs', lines 64:4-64:17 -/ + Source: 'tests/src/constants.rs', lines 65:4-65:17 -/ def get_z1.Z1_body : Result I32 := Result.ok 3#i32 def get_z1.Z1 : I32 := eval_global get_z1.Z1_body /- [constants::get_z1]: - Source: 'tests/src/constants.rs', lines 63:0-63:28 -/ + Source: 'tests/src/constants.rs', lines 64:0-64:28 -/ def get_z1 : Result I32 := Result.ok get_z1.Z1 /- [constants::add]: - Source: 'tests/src/constants.rs', lines 68:0-68:39 -/ + Source: 'tests/src/constants.rs', lines 69:0-69:39 -/ def add (a : I32) (b : I32) : Result I32 := a + b /- [constants::Q1] - Source: 'tests/src/constants.rs', lines 76:0-76:17 -/ + Source: 'tests/src/constants.rs', lines 77:0-77:17 -/ def Q1_body : Result I32 := Result.ok 5#i32 def Q1 : I32 := eval_global Q1_body /- [constants::Q2] - Source: 'tests/src/constants.rs', lines 77:0-77:17 -/ + Source: 'tests/src/constants.rs', lines 78:0-78:17 -/ def Q2_body : Result I32 := Result.ok Q1 def Q2 : I32 := eval_global Q2_body /- [constants::Q3] - Source: 'tests/src/constants.rs', lines 78:0-78:17 -/ + Source: 'tests/src/constants.rs', lines 79:0-79:17 -/ def Q3_body : Result I32 := add Q2 3#i32 def Q3 : I32 := eval_global Q3_body /- [constants::get_z2]: - Source: 'tests/src/constants.rs', lines 72:0-72:28 -/ + Source: 'tests/src/constants.rs', lines 73:0-73:28 -/ def get_z2 : Result I32 := do let i ← get_z1 @@ -130,37 +130,37 @@ def get_z2 : Result I32 := add Q1 i1 /- [constants::S1] - Source: 'tests/src/constants.rs', lines 82:0-82:18 -/ + Source: 'tests/src/constants.rs', lines 83:0-83:18 -/ def S1_body : Result U32 := Result.ok 6#u32 def S1 : U32 := eval_global S1_body /- [constants::S2] - Source: 'tests/src/constants.rs', lines 83:0-83:18 -/ + Source: 'tests/src/constants.rs', lines 84:0-84:18 -/ def S2_body : Result U32 := incr S1 def S2 : U32 := eval_global S2_body /- [constants::S3] - Source: 'tests/src/constants.rs', lines 84:0-84:29 -/ + Source: 'tests/src/constants.rs', lines 85:0-85:29 -/ def S3_body : Result (Pair U32 U32) := Result.ok P3 def S3 : Pair U32 U32 := eval_global S3_body /- [constants::S4] - Source: 'tests/src/constants.rs', lines 85:0-85:29 -/ + Source: 'tests/src/constants.rs', lines 86:0-86:29 -/ def S4_body : Result (Pair U32 U32) := mk_pair1 7#u32 8#u32 def S4 : Pair U32 U32 := eval_global S4_body /- [constants::V] - Source: 'tests/src/constants.rs', lines 88:0-88:31 -/ + Source: 'tests/src/constants.rs', lines 89:0-89:31 -/ structure V (T : Type) (N : Usize) where x : Array T N /- [constants::{constants::V#1}::LEN] - Source: 'tests/src/constants.rs', lines 93:4-93:24 -/ + Source: 'tests/src/constants.rs', lines 94:4-94:24 -/ def V.LEN_body (T : Type) (N : Usize) : Result Usize := Result.ok N def V.LEN (T : Type) (N : Usize) : Usize := eval_global (V.LEN_body T N) /- [constants::use_v]: - Source: 'tests/src/constants.rs', lines 96:0-96:42 -/ + Source: 'tests/src/constants.rs', lines 97:0-97:42 -/ def use_v (T : Type) (N : Usize) : Result Usize := Result.ok (V.LEN T N) diff --git a/tests/lean/Demo/Demo.lean b/tests/lean/Demo/Demo.lean index a7683eb0..48ac2062 100644 --- a/tests/lean/Demo/Demo.lean +++ b/tests/lean/Demo/Demo.lean @@ -6,7 +6,7 @@ open Primitives namespace demo /- [demo::choose]: - Source: 'tests/src/demo.rs', lines 6:0-6:70 -/ + Source: 'tests/src/demo.rs', lines 7:0-7:70 -/ def choose (T : Type) (b : Bool) (x : T) (y : T) : Result (T × (T → Result (T × T))) @@ -18,26 +18,26 @@ def choose Result.ok (y, back) /- [demo::mul2_add1]: - Source: 'tests/src/demo.rs', lines 14:0-14:31 -/ + Source: 'tests/src/demo.rs', lines 15:0-15:31 -/ def mul2_add1 (x : U32) : Result U32 := do let i ← x + x i + 1#u32 /- [demo::use_mul2_add1]: - Source: 'tests/src/demo.rs', lines 18:0-18:43 -/ + Source: 'tests/src/demo.rs', lines 19:0-19:43 -/ def use_mul2_add1 (x : U32) (y : U32) : Result U32 := do let i ← mul2_add1 x i + y /- [demo::incr]: - Source: 'tests/src/demo.rs', lines 22:0-22:31 -/ + Source: 'tests/src/demo.rs', lines 23:0-23:31 -/ def incr (x : U32) : Result U32 := x + 1#u32 /- [demo::use_incr]: - Source: 'tests/src/demo.rs', lines 26:0-26:17 -/ + Source: 'tests/src/demo.rs', lines 27:0-27:17 -/ def use_incr : Result Unit := do let x ← incr 0#u32 @@ -46,13 +46,13 @@ def use_incr : Result Unit := Result.ok () /- [demo::CList] - Source: 'tests/src/demo.rs', lines 35:0-35:17 -/ + Source: 'tests/src/demo.rs', lines 36:0-36:17 -/ inductive CList (T : Type) := | CCons : T → CList T → CList T | CNil : CList T /- [demo::list_nth]: - Source: 'tests/src/demo.rs', lines 40:0-40:56 -/ + Source: 'tests/src/demo.rs', lines 41:0-41:56 -/ divergent def list_nth (T : Type) (l : CList T) (i : U32) : Result T := match l with | CList.CCons x tl => @@ -64,7 +64,7 @@ divergent def list_nth (T : Type) (l : CList T) (i : U32) : Result T := | CList.CNil => Result.fail .panic /- [demo::list_nth_mut]: - Source: 'tests/src/demo.rs', lines 55:0-55:68 -/ + Source: 'tests/src/demo.rs', lines 56:0-56:68 -/ divergent def list_nth_mut (T : Type) (l : CList T) (i : U32) : Result (T × (T → Result (CList T))) @@ -88,7 +88,7 @@ divergent def list_nth_mut | CList.CNil => Result.fail .panic /- [demo::list_nth_mut1]: loop 0: - Source: 'tests/src/demo.rs', lines 70:0-79:1 -/ + Source: 'tests/src/demo.rs', lines 71:0-80:1 -/ divergent def list_nth_mut1_loop (T : Type) (l : CList T) (i : U32) : Result (T × (T → Result (CList T))) @@ -111,7 +111,7 @@ divergent def list_nth_mut1_loop | CList.CNil => Result.fail .panic /- [demo::list_nth_mut1]: - Source: 'tests/src/demo.rs', lines 70:0-70:77 -/ + Source: 'tests/src/demo.rs', lines 71:0-71:77 -/ def list_nth_mut1 (T : Type) (l : CList T) (i : U32) : Result (T × (T → Result (CList T))) @@ -119,7 +119,7 @@ def list_nth_mut1 list_nth_mut1_loop T l i /- [demo::i32_id]: - Source: 'tests/src/demo.rs', lines 81:0-81:28 -/ + Source: 'tests/src/demo.rs', lines 82:0-82:28 -/ divergent def i32_id (i : I32) : Result I32 := if i = 0#i32 then Result.ok 0#i32 @@ -129,7 +129,7 @@ divergent def i32_id (i : I32) : Result I32 := i2 + 1#i32 /- [demo::list_tail]: - Source: 'tests/src/demo.rs', lines 89:0-89:64 -/ + Source: 'tests/src/demo.rs', lines 90:0-90:64 -/ divergent def list_tail (T : Type) (l : CList T) : Result ((CList T) × (CList T → Result (CList T))) @@ -147,25 +147,25 @@ divergent def list_tail | CList.CNil => Result.ok (CList.CNil, Result.ok) /- Trait declaration: [demo::Counter] - Source: 'tests/src/demo.rs', lines 98:0-98:17 -/ + Source: 'tests/src/demo.rs', lines 99:0-99:17 -/ structure Counter (Self : Type) where incr : Self → Result (Usize × Self) /- [demo::{(demo::Counter for usize)}::incr]: - Source: 'tests/src/demo.rs', lines 103:4-103:31 -/ + Source: 'tests/src/demo.rs', lines 104:4-104:31 -/ def CounterUsize.incr (self : Usize) : Result (Usize × Usize) := do let self1 ← self + 1#usize Result.ok (self, self1) /- Trait implementation: [demo::{(demo::Counter for usize)}] - Source: 'tests/src/demo.rs', lines 102:0-102:22 -/ + Source: 'tests/src/demo.rs', lines 103:0-103:22 -/ def CounterUsize : Counter Usize := { incr := CounterUsize.incr } /- [demo::use_counter]: - Source: 'tests/src/demo.rs', lines 110:0-110:59 -/ + Source: 'tests/src/demo.rs', lines 111:0-111:59 -/ def use_counter (T : Type) (CounterInst : Counter T) (cnt : T) : Result (Usize × T) := CounterInst.incr cnt diff --git a/tests/lean/External/Funs.lean b/tests/lean/External/Funs.lean index 84fe1a28..1acb9707 100644 --- a/tests/lean/External/Funs.lean +++ b/tests/lean/External/Funs.lean @@ -15,12 +15,12 @@ def core.marker.CopyU32 : core.marker.Copy U32 := { } /- [external::use_get]: - Source: 'tests/src/external.rs', lines 8:0-8:37 -/ + Source: 'tests/src/external.rs', lines 9:0-9:37 -/ def use_get (rc : core.cell.Cell U32) (st : State) : Result (State × U32) := core.cell.Cell.get U32 core.marker.CopyU32 rc st /- [external::incr]: - Source: 'tests/src/external.rs', lines 12:0-12:31 -/ + Source: 'tests/src/external.rs', lines 13:0-13:31 -/ def incr (rc : core.cell.Cell U32) (st : State) : Result (State × (core.cell.Cell U32)) diff --git a/tests/lean/Loops.lean b/tests/lean/Loops.lean index 54f1b24f..199605d5 100644 --- a/tests/lean/Loops.lean +++ b/tests/lean/Loops.lean @@ -6,7 +6,7 @@ open Primitives namespace loops /- [loops::sum]: loop 0: - Source: 'tests/src/loops.rs', lines 7:0-17:1 -/ + Source: 'tests/src/loops.rs', lines 8:0-18:1 -/ divergent def sum_loop (max : U32) (i : U32) (s : U32) : Result U32 := if i < max then do @@ -16,12 +16,12 @@ divergent def sum_loop (max : U32) (i : U32) (s : U32) : Result U32 := else s * 2#u32 /- [loops::sum]: - Source: 'tests/src/loops.rs', lines 7:0-7:27 -/ + Source: 'tests/src/loops.rs', lines 8:0-8:27 -/ def sum (max : U32) : Result U32 := sum_loop max 0#u32 0#u32 /- [loops::sum_with_mut_borrows]: loop 0: - Source: 'tests/src/loops.rs', lines 22:0-34:1 -/ + Source: 'tests/src/loops.rs', lines 23:0-35:1 -/ divergent def sum_with_mut_borrows_loop (max : U32) (i : U32) (s : U32) : Result U32 := if i < max @@ -33,12 +33,12 @@ divergent def sum_with_mut_borrows_loop else s * 2#u32 /- [loops::sum_with_mut_borrows]: - Source: 'tests/src/loops.rs', lines 22:0-22:44 -/ + Source: 'tests/src/loops.rs', lines 23:0-23:44 -/ def sum_with_mut_borrows (max : U32) : Result U32 := sum_with_mut_borrows_loop max 0#u32 0#u32 /- [loops::sum_with_shared_borrows]: loop 0: - Source: 'tests/src/loops.rs', lines 37:0-51:1 -/ + Source: 'tests/src/loops.rs', lines 38:0-52:1 -/ divergent def sum_with_shared_borrows_loop (max : U32) (i : U32) (s : U32) : Result U32 := if i < max @@ -50,12 +50,12 @@ divergent def sum_with_shared_borrows_loop else s * 2#u32 /- [loops::sum_with_shared_borrows]: - Source: 'tests/src/loops.rs', lines 37:0-37:47 -/ + Source: 'tests/src/loops.rs', lines 38:0-38:47 -/ def sum_with_shared_borrows (max : U32) : Result U32 := sum_with_shared_borrows_loop max 0#u32 0#u32 /- [loops::sum_array]: loop 0: - Source: 'tests/src/loops.rs', lines 53:0-61:1 -/ + Source: 'tests/src/loops.rs', lines 54:0-62:1 -/ divergent def sum_array_loop (N : Usize) (a : Array U32 N) (i : Usize) (s : U32) : Result U32 := if i < N @@ -68,12 +68,12 @@ divergent def sum_array_loop else Result.ok s /- [loops::sum_array]: - Source: 'tests/src/loops.rs', lines 53:0-53:52 -/ + Source: 'tests/src/loops.rs', lines 54:0-54:52 -/ def sum_array (N : Usize) (a : Array U32 N) : Result U32 := sum_array_loop N a 0#usize 0#u32 /- [loops::clear]: loop 0: - Source: 'tests/src/loops.rs', lines 65:0-71:1 -/ + Source: 'tests/src/loops.rs', lines 66:0-72:1 -/ divergent def clear_loop (v : alloc.vec.Vec U32) (i : Usize) : Result (alloc.vec.Vec U32) := let i1 := alloc.vec.Vec.len U32 v @@ -89,18 +89,18 @@ divergent def clear_loop else Result.ok v /- [loops::clear]: - Source: 'tests/src/loops.rs', lines 65:0-65:30 -/ + Source: 'tests/src/loops.rs', lines 66:0-66:30 -/ def clear (v : alloc.vec.Vec U32) : Result (alloc.vec.Vec U32) := clear_loop v 0#usize /- [loops::List] - Source: 'tests/src/loops.rs', lines 73:0-73:16 -/ + Source: 'tests/src/loops.rs', lines 74:0-74:16 -/ inductive List (T : Type) := | Cons : T → List T → List T | Nil : List T /- [loops::list_mem]: loop 0: - Source: 'tests/src/loops.rs', lines 79:0-88:1 -/ + Source: 'tests/src/loops.rs', lines 80:0-89:1 -/ divergent def list_mem_loop (x : U32) (ls : List U32) : Result Bool := match ls with | List.Cons y tl => if y = x @@ -109,12 +109,12 @@ divergent def list_mem_loop (x : U32) (ls : List U32) : Result Bool := | List.Nil => Result.ok false /- [loops::list_mem]: - Source: 'tests/src/loops.rs', lines 79:0-79:52 -/ + Source: 'tests/src/loops.rs', lines 80:0-80:52 -/ def list_mem (x : U32) (ls : List U32) : Result Bool := list_mem_loop x ls /- [loops::list_nth_mut_loop]: loop 0: - Source: 'tests/src/loops.rs', lines 91:0-101:1 -/ + Source: 'tests/src/loops.rs', lines 92:0-102:1 -/ divergent def list_nth_mut_loop_loop (T : Type) (ls : List T) (i : U32) : Result (T × (T → Result (List T))) := match ls with @@ -135,13 +135,13 @@ divergent def list_nth_mut_loop_loop | List.Nil => Result.fail .panic /- [loops::list_nth_mut_loop]: - Source: 'tests/src/loops.rs', lines 91:0-91:71 -/ + Source: 'tests/src/loops.rs', lines 92:0-92:71 -/ 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 /- [loops::list_nth_shared_loop]: loop 0: - Source: 'tests/src/loops.rs', lines 104:0-114:1 -/ + Source: 'tests/src/loops.rs', lines 105:0-115:1 -/ divergent def list_nth_shared_loop_loop (T : Type) (ls : List T) (i : U32) : Result T := match ls with @@ -154,12 +154,12 @@ divergent def list_nth_shared_loop_loop | List.Nil => Result.fail .panic /- [loops::list_nth_shared_loop]: - Source: 'tests/src/loops.rs', lines 104:0-104:66 -/ + Source: 'tests/src/loops.rs', lines 105:0-105:66 -/ def list_nth_shared_loop (T : Type) (ls : List T) (i : U32) : Result T := list_nth_shared_loop_loop T ls i /- [loops::get_elem_mut]: loop 0: - Source: 'tests/src/loops.rs', lines 116:0-130:1 -/ + Source: 'tests/src/loops.rs', lines 117:0-131:1 -/ divergent def get_elem_mut_loop (x : Usize) (ls : List Usize) : Result (Usize × (Usize → Result (List Usize))) @@ -181,7 +181,7 @@ divergent def get_elem_mut_loop | List.Nil => Result.fail .panic /- [loops::get_elem_mut]: - Source: 'tests/src/loops.rs', lines 116:0-116:73 -/ + Source: 'tests/src/loops.rs', lines 117:0-117:73 -/ def get_elem_mut (slots : alloc.vec.Vec (List Usize)) (x : Usize) : Result (Usize × (Usize → Result (alloc.vec.Vec (List Usize)))) @@ -197,7 +197,7 @@ def get_elem_mut Result.ok (i, back1) /- [loops::get_elem_shared]: loop 0: - Source: 'tests/src/loops.rs', lines 132:0-146:1 -/ + Source: 'tests/src/loops.rs', lines 133:0-147:1 -/ divergent def get_elem_shared_loop (x : Usize) (ls : List Usize) : Result Usize := match ls with @@ -207,7 +207,7 @@ divergent def get_elem_shared_loop | List.Nil => Result.fail .panic /- [loops::get_elem_shared]: - Source: 'tests/src/loops.rs', lines 132:0-132:68 -/ + Source: 'tests/src/loops.rs', lines 133:0-133:68 -/ def get_elem_shared (slots : alloc.vec.Vec (List Usize)) (x : Usize) : Result Usize := do @@ -217,7 +217,7 @@ def get_elem_shared get_elem_shared_loop x ls /- [loops::id_mut]: - Source: 'tests/src/loops.rs', lines 148:0-148:50 -/ + Source: 'tests/src/loops.rs', lines 149:0-149:50 -/ def id_mut (T : Type) (ls : List T) : Result ((List T) × (List T → Result (List T))) @@ -225,12 +225,12 @@ def id_mut Result.ok (ls, Result.ok) /- [loops::id_shared]: - Source: 'tests/src/loops.rs', lines 152:0-152:45 -/ + Source: 'tests/src/loops.rs', lines 153:0-153:45 -/ def id_shared (T : Type) (ls : List T) : Result (List T) := Result.ok ls /- [loops::list_nth_mut_loop_with_id]: loop 0: - Source: 'tests/src/loops.rs', lines 157:0-168:1 -/ + Source: 'tests/src/loops.rs', lines 158:0-169:1 -/ divergent def list_nth_mut_loop_with_id_loop (T : Type) (i : U32) (ls : List T) : Result (T × (T → Result (List T))) := match ls with @@ -251,7 +251,7 @@ divergent def list_nth_mut_loop_with_id_loop | List.Nil => Result.fail .panic /- [loops::list_nth_mut_loop_with_id]: - Source: 'tests/src/loops.rs', lines 157:0-157:75 -/ + Source: 'tests/src/loops.rs', lines 158:0-158:75 -/ def list_nth_mut_loop_with_id (T : Type) (ls : List T) (i : U32) : Result (T × (T → Result (List T))) := do @@ -263,7 +263,7 @@ def list_nth_mut_loop_with_id Result.ok (t, back1) /- [loops::list_nth_shared_loop_with_id]: loop 0: - Source: 'tests/src/loops.rs', lines 171:0-182:1 -/ + Source: 'tests/src/loops.rs', lines 172:0-183:1 -/ divergent def list_nth_shared_loop_with_id_loop (T : Type) (i : U32) (ls : List T) : Result T := match ls with @@ -276,7 +276,7 @@ divergent def list_nth_shared_loop_with_id_loop | List.Nil => Result.fail .panic /- [loops::list_nth_shared_loop_with_id]: - Source: 'tests/src/loops.rs', lines 171:0-171:70 -/ + Source: 'tests/src/loops.rs', lines 172:0-172:70 -/ def list_nth_shared_loop_with_id (T : Type) (ls : List T) (i : U32) : Result T := do @@ -284,7 +284,7 @@ def list_nth_shared_loop_with_id list_nth_shared_loop_with_id_loop T i ls1 /- [loops::list_nth_mut_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 187:0-208:1 -/ + Source: 'tests/src/loops.rs', lines 188:0-209:1 -/ divergent def list_nth_mut_loop_pair_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T)) × (T → Result (List T))) @@ -315,7 +315,7 @@ divergent def list_nth_mut_loop_pair_loop | List.Nil => Result.fail .panic /- [loops::list_nth_mut_loop_pair]: - Source: 'tests/src/loops.rs', lines 187:0-191:27 -/ + Source: 'tests/src/loops.rs', lines 188:0-192:27 -/ 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))) @@ -323,7 +323,7 @@ def list_nth_mut_loop_pair list_nth_mut_loop_pair_loop T ls0 ls1 i /- [loops::list_nth_shared_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 211:0-232:1 -/ + Source: 'tests/src/loops.rs', lines 212:0-233:1 -/ divergent def list_nth_shared_loop_pair_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result (T × T) := match ls0 with @@ -339,13 +339,13 @@ divergent def list_nth_shared_loop_pair_loop | List.Nil => Result.fail .panic /- [loops::list_nth_shared_loop_pair]: - Source: 'tests/src/loops.rs', lines 211:0-215:19 -/ + Source: 'tests/src/loops.rs', lines 212:0-216:19 -/ 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 /- [loops::list_nth_mut_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 236:0-251:1 -/ + Source: 'tests/src/loops.rs', lines 237:0-252:1 -/ divergent def list_nth_mut_loop_pair_merge_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × ((T × T) → Result ((List T) × (List T)))) @@ -375,7 +375,7 @@ divergent def list_nth_mut_loop_pair_merge_loop | List.Nil => Result.fail .panic /- [loops::list_nth_mut_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 236:0-240:27 -/ + Source: 'tests/src/loops.rs', lines 237:0-241:27 -/ 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)))) @@ -383,7 +383,7 @@ def list_nth_mut_loop_pair_merge list_nth_mut_loop_pair_merge_loop T ls0 ls1 i /- [loops::list_nth_shared_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 254:0-269:1 -/ + Source: 'tests/src/loops.rs', lines 255:0-270:1 -/ divergent def list_nth_shared_loop_pair_merge_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result (T × T) := match ls0 with @@ -400,13 +400,13 @@ divergent def list_nth_shared_loop_pair_merge_loop | List.Nil => Result.fail .panic /- [loops::list_nth_shared_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 254:0-258:19 -/ + Source: 'tests/src/loops.rs', lines 255:0-259:19 -/ 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 /- [loops::list_nth_mut_shared_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 272:0-287:1 -/ + Source: 'tests/src/loops.rs', lines 273:0-288:1 -/ divergent def list_nth_mut_shared_loop_pair_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -432,7 +432,7 @@ divergent def list_nth_mut_shared_loop_pair_loop | List.Nil => Result.fail .panic /- [loops::list_nth_mut_shared_loop_pair]: - Source: 'tests/src/loops.rs', lines 272:0-276:23 -/ + Source: 'tests/src/loops.rs', lines 273:0-277:23 -/ def list_nth_mut_shared_loop_pair (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -440,7 +440,7 @@ def list_nth_mut_shared_loop_pair list_nth_mut_shared_loop_pair_loop T ls0 ls1 i /- [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 291:0-306:1 -/ + Source: 'tests/src/loops.rs', lines 292:0-307:1 -/ divergent def list_nth_mut_shared_loop_pair_merge_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -466,7 +466,7 @@ divergent def list_nth_mut_shared_loop_pair_merge_loop | List.Nil => Result.fail .panic /- [loops::list_nth_mut_shared_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 291:0-295:23 -/ + Source: 'tests/src/loops.rs', lines 292:0-296:23 -/ 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))) @@ -474,7 +474,7 @@ def list_nth_mut_shared_loop_pair_merge list_nth_mut_shared_loop_pair_merge_loop T ls0 ls1 i /- [loops::list_nth_shared_mut_loop_pair]: loop 0: - Source: 'tests/src/loops.rs', lines 310:0-325:1 -/ + Source: 'tests/src/loops.rs', lines 311:0-326:1 -/ divergent def list_nth_shared_mut_loop_pair_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -500,7 +500,7 @@ divergent def list_nth_shared_mut_loop_pair_loop | List.Nil => Result.fail .panic /- [loops::list_nth_shared_mut_loop_pair]: - Source: 'tests/src/loops.rs', lines 310:0-314:23 -/ + Source: 'tests/src/loops.rs', lines 311:0-315:23 -/ def list_nth_shared_mut_loop_pair (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -508,7 +508,7 @@ def list_nth_shared_mut_loop_pair list_nth_shared_mut_loop_pair_loop T ls0 ls1 i /- [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: - Source: 'tests/src/loops.rs', lines 329:0-344:1 -/ + Source: 'tests/src/loops.rs', lines 330:0-345:1 -/ divergent def list_nth_shared_mut_loop_pair_merge_loop (T : Type) (ls0 : List T) (ls1 : List T) (i : U32) : Result ((T × T) × (T → Result (List T))) @@ -534,7 +534,7 @@ divergent def list_nth_shared_mut_loop_pair_merge_loop | List.Nil => Result.fail .panic /- [loops::list_nth_shared_mut_loop_pair_merge]: - Source: 'tests/src/loops.rs', lines 329:0-333:23 -/ + Source: 'tests/src/loops.rs', lines 330:0-334:23 -/ 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))) @@ -542,7 +542,7 @@ def list_nth_shared_mut_loop_pair_merge list_nth_shared_mut_loop_pair_merge_loop T ls0 ls1 i /- [loops::ignore_input_mut_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 348:0-352:1 -/ + Source: 'tests/src/loops.rs', lines 349:0-353:1 -/ divergent def ignore_input_mut_borrow_loop (i : U32) : Result Unit := if i > 0#u32 then do @@ -551,14 +551,14 @@ divergent def ignore_input_mut_borrow_loop (i : U32) : Result Unit := else Result.ok () /- [loops::ignore_input_mut_borrow]: - Source: 'tests/src/loops.rs', lines 348:0-348:56 -/ + Source: 'tests/src/loops.rs', lines 349:0-349:56 -/ def ignore_input_mut_borrow (_a : U32) (i : U32) : Result U32 := do let _ ← ignore_input_mut_borrow_loop i Result.ok _a /- [loops::incr_ignore_input_mut_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 356:0-361:1 -/ + Source: 'tests/src/loops.rs', lines 357:0-362:1 -/ divergent def incr_ignore_input_mut_borrow_loop (i : U32) : Result Unit := if i > 0#u32 then do @@ -567,7 +567,7 @@ divergent def incr_ignore_input_mut_borrow_loop (i : U32) : Result Unit := else Result.ok () /- [loops::incr_ignore_input_mut_borrow]: - Source: 'tests/src/loops.rs', lines 356:0-356:60 -/ + Source: 'tests/src/loops.rs', lines 357:0-357:60 -/ def incr_ignore_input_mut_borrow (a : U32) (i : U32) : Result U32 := do let a1 ← a + 1#u32 @@ -575,7 +575,7 @@ def incr_ignore_input_mut_borrow (a : U32) (i : U32) : Result U32 := Result.ok a1 /- [loops::ignore_input_shared_borrow]: loop 0: - Source: 'tests/src/loops.rs', lines 365:0-369:1 -/ + Source: 'tests/src/loops.rs', lines 366:0-370:1 -/ divergent def ignore_input_shared_borrow_loop (i : U32) : Result Unit := if i > 0#u32 then do @@ -584,7 +584,7 @@ divergent def ignore_input_shared_borrow_loop (i : U32) : Result Unit := else Result.ok () /- [loops::ignore_input_shared_borrow]: - Source: 'tests/src/loops.rs', lines 365:0-365:59 -/ + Source: 'tests/src/loops.rs', lines 366:0-366:59 -/ def ignore_input_shared_borrow (_a : U32) (i : U32) : Result U32 := do let _ ← ignore_input_shared_borrow_loop i diff --git a/tests/lean/NoNestedBorrows.lean b/tests/lean/NoNestedBorrows.lean index 1781ac71..022b32fb 100644 --- a/tests/lean/NoNestedBorrows.lean +++ b/tests/lean/NoNestedBorrows.lean @@ -6,60 +6,60 @@ open Primitives namespace no_nested_borrows /- [no_nested_borrows::Pair] - Source: 'tests/src/no_nested_borrows.rs', lines 6:0-6:23 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 7:0-7:23 -/ structure Pair (T1 T2 : Type) where x : T1 y : T2 /- [no_nested_borrows::List] - Source: 'tests/src/no_nested_borrows.rs', lines 11:0-11:16 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 12:0-12:16 -/ inductive List (T : Type) := | Cons : T → List T → List T | Nil : List T /- [no_nested_borrows::One] - Source: 'tests/src/no_nested_borrows.rs', lines 22:0-22:16 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 23:0-23:16 -/ inductive One (T1 : Type) := | One : T1 → One T1 /- [no_nested_borrows::EmptyEnum] - Source: 'tests/src/no_nested_borrows.rs', lines 28:0-28:18 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 29:0-29:18 -/ inductive EmptyEnum := | Empty : EmptyEnum /- [no_nested_borrows::Enum] - Source: 'tests/src/no_nested_borrows.rs', lines 34:0-34:13 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 35:0-35:13 -/ inductive Enum := | Variant1 : Enum | Variant2 : Enum /- [no_nested_borrows::EmptyStruct] - Source: 'tests/src/no_nested_borrows.rs', lines 41:0-41:22 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 42:0-42:22 -/ @[reducible] def EmptyStruct := Unit /- [no_nested_borrows::Sum] - Source: 'tests/src/no_nested_borrows.rs', lines 43:0-43:20 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 44:0-44:20 -/ inductive Sum (T1 T2 : Type) := | Left : T1 → Sum T1 T2 | Right : T2 → Sum T1 T2 /- [no_nested_borrows::cast_u32_to_i32]: - Source: 'tests/src/no_nested_borrows.rs', lines 48:0-48:37 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 49:0-49:37 -/ def cast_u32_to_i32 (x : U32) : Result I32 := Scalar.cast .I32 x /- [no_nested_borrows::cast_bool_to_i32]: - Source: 'tests/src/no_nested_borrows.rs', lines 52:0-52:39 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 53:0-53:39 -/ def cast_bool_to_i32 (x : Bool) : Result I32 := Scalar.cast_bool .I32 x /- [no_nested_borrows::cast_bool_to_bool]: - Source: 'tests/src/no_nested_borrows.rs', lines 57:0-57:41 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 58:0-58:41 -/ def cast_bool_to_bool (x : Bool) : Result Bool := Result.ok x /- [no_nested_borrows::test2]: - Source: 'tests/src/no_nested_borrows.rs', lines 62:0-62:14 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 63:0-63:14 -/ def test2 : Result Unit := do let _ ← 23#u32 + 44#u32 @@ -69,14 +69,14 @@ def test2 : Result Unit := #assert (test2 == Result.ok ()) /- [no_nested_borrows::get_max]: - Source: 'tests/src/no_nested_borrows.rs', lines 74:0-74:37 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 75:0-75:37 -/ def get_max (x : U32) (y : U32) : Result U32 := if x >= y then Result.ok x else Result.ok y /- [no_nested_borrows::test3]: - Source: 'tests/src/no_nested_borrows.rs', lines 82:0-82:14 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 83:0-83:14 -/ def test3 : Result Unit := do let x ← get_max 4#u32 3#u32 @@ -90,7 +90,7 @@ def test3 : Result Unit := #assert (test3 == Result.ok ()) /- [no_nested_borrows::test_neg1]: - Source: 'tests/src/no_nested_borrows.rs', lines 89:0-89:18 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 90:0-90:18 -/ def test_neg1 : Result Unit := do let y ← -. 3#i32 @@ -102,7 +102,7 @@ def test_neg1 : Result Unit := #assert (test_neg1 == Result.ok ()) /- [no_nested_borrows::refs_test1]: - Source: 'tests/src/no_nested_borrows.rs', lines 96:0-96:19 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 97:0-97:19 -/ def refs_test1 : Result Unit := if ¬ (1#i32 = 1#i32) then Result.fail .panic @@ -112,7 +112,7 @@ def refs_test1 : Result Unit := #assert (refs_test1 == Result.ok ()) /- [no_nested_borrows::refs_test2]: - Source: 'tests/src/no_nested_borrows.rs', lines 107:0-107:19 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 108:0-108:19 -/ def refs_test2 : Result Unit := if ¬ (2#i32 = 2#i32) then Result.fail .panic @@ -130,7 +130,7 @@ def refs_test2 : Result Unit := #assert (refs_test2 == Result.ok ()) /- [no_nested_borrows::test_list1]: - Source: 'tests/src/no_nested_borrows.rs', lines 123:0-123:19 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 124:0-124:19 -/ def test_list1 : Result Unit := Result.ok () @@ -138,7 +138,7 @@ def test_list1 : Result Unit := #assert (test_list1 == Result.ok ()) /- [no_nested_borrows::test_box1]: - Source: 'tests/src/no_nested_borrows.rs', lines 128:0-128:18 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 129:0-129:18 -/ def test_box1 : Result Unit := do let (_, deref_mut_back) ← alloc.boxed.Box.deref_mut I32 0#i32 @@ -152,26 +152,26 @@ def test_box1 : Result Unit := #assert (test_box1 == Result.ok ()) /- [no_nested_borrows::copy_int]: - Source: 'tests/src/no_nested_borrows.rs', lines 138:0-138:30 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 139:0-139:30 -/ def copy_int (x : I32) : Result I32 := Result.ok x /- [no_nested_borrows::test_unreachable]: - Source: 'tests/src/no_nested_borrows.rs', lines 144:0-144:32 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 145:0-145:32 -/ def test_unreachable (b : Bool) : Result Unit := if b then Result.fail .panic else Result.ok () /- [no_nested_borrows::test_panic]: - Source: 'tests/src/no_nested_borrows.rs', lines 152:0-152:26 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 153:0-153:26 -/ def test_panic (b : Bool) : Result Unit := if b then Result.fail .panic else Result.ok () /- [no_nested_borrows::test_copy_int]: - Source: 'tests/src/no_nested_borrows.rs', lines 159:0-159:22 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 160:0-160:22 -/ def test_copy_int : Result Unit := do let y ← copy_int 0#i32 @@ -183,14 +183,14 @@ def test_copy_int : Result Unit := #assert (test_copy_int == Result.ok ()) /- [no_nested_borrows::is_cons]: - Source: 'tests/src/no_nested_borrows.rs', lines 166:0-166:38 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 167:0-167:38 -/ def is_cons (T : Type) (l : List T) : Result Bool := match l with | List.Cons _ _ => Result.ok true | List.Nil => Result.ok false /- [no_nested_borrows::test_is_cons]: - Source: 'tests/src/no_nested_borrows.rs', lines 173:0-173:21 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 174:0-174:21 -/ def test_is_cons : Result Unit := do let b ← is_cons I32 (List.Cons 0#i32 List.Nil) @@ -202,14 +202,14 @@ def test_is_cons : Result Unit := #assert (test_is_cons == Result.ok ()) /- [no_nested_borrows::split_list]: - Source: 'tests/src/no_nested_borrows.rs', lines 179:0-179:48 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 180:0-180:48 -/ def split_list (T : Type) (l : List T) : Result (T × (List T)) := match l with | List.Cons hd tl => Result.ok (hd, tl) | List.Nil => Result.fail .panic /- [no_nested_borrows::test_split_list]: - Source: 'tests/src/no_nested_borrows.rs', lines 187:0-187:24 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 188:0-188:24 -/ def test_split_list : Result Unit := do let p ← split_list I32 (List.Cons 0#i32 List.Nil) @@ -222,7 +222,7 @@ def test_split_list : Result Unit := #assert (test_split_list == Result.ok ()) /- [no_nested_borrows::choose]: - Source: 'tests/src/no_nested_borrows.rs', lines 194:0-194:70 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 195:0-195:70 -/ def choose (T : Type) (b : Bool) (x : T) (y : T) : Result (T × (T → Result (T × T))) @@ -234,7 +234,7 @@ def choose Result.ok (y, back) /- [no_nested_borrows::choose_test]: - Source: 'tests/src/no_nested_borrows.rs', lines 202:0-202:20 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 203:0-203:20 -/ def choose_test : Result Unit := do let (z, choose_back) ← choose I32 true 0#i32 0#i32 @@ -254,20 +254,20 @@ def choose_test : Result Unit := #assert (choose_test == Result.ok ()) /- [no_nested_borrows::test_char]: - Source: 'tests/src/no_nested_borrows.rs', lines 214:0-214:26 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 215:0-215:26 -/ def test_char : Result Char := Result.ok 'a' mutual /- [no_nested_borrows::Tree] - Source: 'tests/src/no_nested_borrows.rs', lines 219:0-219:16 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 220:0-220:16 -/ inductive Tree (T : Type) := | Leaf : T → Tree T | Node : T → NodeElem T → Tree T → Tree T /- [no_nested_borrows::NodeElem] - Source: 'tests/src/no_nested_borrows.rs', lines 224:0-224:20 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 225:0-225:20 -/ inductive NodeElem (T : Type) := | Cons : Tree T → NodeElem T → NodeElem T | Nil : NodeElem T @@ -275,7 +275,7 @@ inductive NodeElem (T : Type) := end /- [no_nested_borrows::list_length]: - Source: 'tests/src/no_nested_borrows.rs', lines 259:0-259:48 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 260:0-260:48 -/ divergent def list_length (T : Type) (l : List T) : Result U32 := match l with | List.Cons _ l1 => do @@ -284,7 +284,7 @@ divergent def list_length (T : Type) (l : List T) : Result U32 := | List.Nil => Result.ok 0#u32 /- [no_nested_borrows::list_nth_shared]: - Source: 'tests/src/no_nested_borrows.rs', lines 267:0-267:62 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 268:0-268:62 -/ divergent def list_nth_shared (T : Type) (l : List T) (i : U32) : Result T := match l with | List.Cons x tl => @@ -296,7 +296,7 @@ divergent def list_nth_shared (T : Type) (l : List T) (i : U32) : Result T := | List.Nil => Result.fail .panic /- [no_nested_borrows::list_nth_mut]: - Source: 'tests/src/no_nested_borrows.rs', lines 283:0-283:67 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 284:0-284:67 -/ divergent def list_nth_mut (T : Type) (l : List T) (i : U32) : Result (T × (T → Result (List T))) := match l with @@ -318,7 +318,7 @@ divergent def list_nth_mut | List.Nil => Result.fail .panic /- [no_nested_borrows::list_rev_aux]: - Source: 'tests/src/no_nested_borrows.rs', lines 299:0-299:63 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 300:0-300:63 -/ divergent def list_rev_aux (T : Type) (li : List T) (lo : List T) : Result (List T) := match li with @@ -326,13 +326,13 @@ divergent def list_rev_aux | List.Nil => Result.ok lo /- [no_nested_borrows::list_rev]: - Source: 'tests/src/no_nested_borrows.rs', lines 313:0-313:42 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 314:0-314:42 -/ def list_rev (T : Type) (l : List T) : Result (List T) := let (li, _) := core.mem.replace (List T) l List.Nil list_rev_aux T li List.Nil /- [no_nested_borrows::test_list_functions]: - Source: 'tests/src/no_nested_borrows.rs', lines 318:0-318:28 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 319:0-319:28 -/ def test_list_functions : Result Unit := do let l := List.Cons 2#i32 List.Nil @@ -379,7 +379,7 @@ def test_list_functions : Result Unit := #assert (test_list_functions == Result.ok ()) /- [no_nested_borrows::id_mut_pair1]: - Source: 'tests/src/no_nested_borrows.rs', lines 334:0-334:89 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 335:0-335:89 -/ def id_mut_pair1 (T1 T2 : Type) (x : T1) (y : T2) : Result ((T1 × T2) × ((T1 × T2) → Result (T1 × T2))) @@ -387,7 +387,7 @@ def id_mut_pair1 Result.ok ((x, y), Result.ok) /- [no_nested_borrows::id_mut_pair2]: - Source: 'tests/src/no_nested_borrows.rs', lines 338:0-338:88 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 339:0-339:88 -/ def id_mut_pair2 (T1 T2 : Type) (p : (T1 × T2)) : Result ((T1 × T2) × ((T1 × T2) → Result (T1 × T2))) @@ -396,7 +396,7 @@ def id_mut_pair2 Result.ok ((t, t1), Result.ok) /- [no_nested_borrows::id_mut_pair3]: - Source: 'tests/src/no_nested_borrows.rs', lines 342:0-342:93 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 343:0-343:93 -/ def id_mut_pair3 (T1 T2 : Type) (x : T1) (y : T2) : Result ((T1 × T2) × (T1 → Result T1) × (T2 → Result T2)) @@ -404,7 +404,7 @@ def id_mut_pair3 Result.ok ((x, y), Result.ok, Result.ok) /- [no_nested_borrows::id_mut_pair4]: - Source: 'tests/src/no_nested_borrows.rs', lines 346:0-346:92 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 347:0-347:92 -/ def id_mut_pair4 (T1 T2 : Type) (p : (T1 × T2)) : Result ((T1 × T2) × (T1 → Result T1) × (T2 → Result T2)) @@ -413,37 +413,37 @@ def id_mut_pair4 Result.ok ((t, t1), Result.ok, Result.ok) /- [no_nested_borrows::StructWithTuple] - Source: 'tests/src/no_nested_borrows.rs', lines 353:0-353:34 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 354:0-354:34 -/ structure StructWithTuple (T1 T2 : Type) where p : (T1 × T2) /- [no_nested_borrows::new_tuple1]: - Source: 'tests/src/no_nested_borrows.rs', lines 357:0-357:48 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 358:0-358:48 -/ def new_tuple1 : Result (StructWithTuple U32 U32) := Result.ok { p := (1#u32, 2#u32) } /- [no_nested_borrows::new_tuple2]: - Source: 'tests/src/no_nested_borrows.rs', lines 361:0-361:48 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 362:0-362:48 -/ def new_tuple2 : Result (StructWithTuple I16 I16) := Result.ok { p := (1#i16, 2#i16) } /- [no_nested_borrows::new_tuple3]: - Source: 'tests/src/no_nested_borrows.rs', lines 365:0-365:48 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 366:0-366:48 -/ def new_tuple3 : Result (StructWithTuple U64 I64) := Result.ok { p := (1#u64, 2#i64) } /- [no_nested_borrows::StructWithPair] - Source: 'tests/src/no_nested_borrows.rs', lines 370:0-370:33 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 371:0-371:33 -/ structure StructWithPair (T1 T2 : Type) where p : Pair T1 T2 /- [no_nested_borrows::new_pair1]: - Source: 'tests/src/no_nested_borrows.rs', lines 374:0-374:46 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 375:0-375:46 -/ def new_pair1 : Result (StructWithPair U32 U32) := Result.ok { p := { x := 1#u32, y := 2#u32 } } /- [no_nested_borrows::test_constants]: - Source: 'tests/src/no_nested_borrows.rs', lines 382:0-382:23 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 383:0-383:23 -/ def test_constants : Result Unit := do let swt ← new_tuple1 @@ -473,7 +473,7 @@ def test_constants : Result Unit := #assert (test_constants == Result.ok ()) /- [no_nested_borrows::test_weird_borrows1]: - Source: 'tests/src/no_nested_borrows.rs', lines 391:0-391:28 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 392:0-392:28 -/ def test_weird_borrows1 : Result Unit := Result.ok () @@ -481,7 +481,7 @@ def test_weird_borrows1 : Result Unit := #assert (test_weird_borrows1 == Result.ok ()) /- [no_nested_borrows::test_mem_replace]: - Source: 'tests/src/no_nested_borrows.rs', lines 401:0-401:37 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 402:0-402:37 -/ def test_mem_replace (px : U32) : Result U32 := let (y, _) := core.mem.replace U32 px 1#u32 if ¬ (y = 0#u32) @@ -489,71 +489,71 @@ def test_mem_replace (px : U32) : Result U32 := else Result.ok 2#u32 /- [no_nested_borrows::test_shared_borrow_bool1]: - Source: 'tests/src/no_nested_borrows.rs', lines 408:0-408:47 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 409:0-409:47 -/ def test_shared_borrow_bool1 (b : Bool) : Result U32 := if b then Result.ok 0#u32 else Result.ok 1#u32 /- [no_nested_borrows::test_shared_borrow_bool2]: - Source: 'tests/src/no_nested_borrows.rs', lines 421:0-421:40 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 422:0-422:40 -/ def test_shared_borrow_bool2 : Result U32 := Result.ok 0#u32 /- [no_nested_borrows::test_shared_borrow_enum1]: - Source: 'tests/src/no_nested_borrows.rs', lines 436:0-436:52 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 437:0-437:52 -/ def test_shared_borrow_enum1 (l : List U32) : Result U32 := match l with | List.Cons _ _ => Result.ok 1#u32 | List.Nil => Result.ok 0#u32 /- [no_nested_borrows::test_shared_borrow_enum2]: - Source: 'tests/src/no_nested_borrows.rs', lines 448:0-448:40 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 449:0-449:40 -/ def test_shared_borrow_enum2 : Result U32 := Result.ok 0#u32 /- [no_nested_borrows::incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 459:0-459:24 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 460:0-460:24 -/ def incr (x : U32) : Result U32 := x + 1#u32 /- [no_nested_borrows::call_incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 463:0-463:35 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 464:0-464:35 -/ def call_incr (x : U32) : Result U32 := incr x /- [no_nested_borrows::read_then_incr]: - Source: 'tests/src/no_nested_borrows.rs', lines 468:0-468:41 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 469:0-469:41 -/ def read_then_incr (x : U32) : Result (U32 × U32) := do let x1 ← x + 1#u32 Result.ok (x, x1) /- [no_nested_borrows::Tuple] - Source: 'tests/src/no_nested_borrows.rs', lines 474:0-474:24 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 475:0-475:24 -/ def Tuple (T1 T2 : Type) := T1 × T2 /- [no_nested_borrows::use_tuple_struct]: - Source: 'tests/src/no_nested_borrows.rs', lines 476:0-476:48 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 477:0-477:48 -/ def use_tuple_struct (x : Tuple U32 U32) : Result (Tuple U32 U32) := Result.ok (1#u32, x.#1) /- [no_nested_borrows::create_tuple_struct]: - Source: 'tests/src/no_nested_borrows.rs', lines 480:0-480:61 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 481:0-481:61 -/ def create_tuple_struct (x : U32) (y : U64) : Result (Tuple U32 U64) := Result.ok (x, y) /- [no_nested_borrows::IdType] - Source: 'tests/src/no_nested_borrows.rs', lines 485:0-485:20 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 486:0-486:20 -/ @[reducible] def IdType (T : Type) := T /- [no_nested_borrows::use_id_type]: - Source: 'tests/src/no_nested_borrows.rs', lines 487:0-487:40 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 488:0-488:40 -/ def use_id_type (T : Type) (x : IdType T) : Result T := Result.ok x /- [no_nested_borrows::create_id_type]: - Source: 'tests/src/no_nested_borrows.rs', lines 491:0-491:43 -/ + Source: 'tests/src/no_nested_borrows.rs', lines 492:0-492:43 -/ def create_id_type (T : Type) (x : T) : Result (IdType T) := Result.ok x diff --git a/tests/lean/Paper.lean b/tests/lean/Paper.lean index e98ada42..dbd56f3e 100644 --- a/tests/lean/Paper.lean +++ b/tests/lean/Paper.lean @@ -6,12 +6,12 @@ open Primitives namespace paper /- [paper::ref_incr]: - Source: 'tests/src/paper.rs', lines 6:0-6:28 -/ + Source: 'tests/src/paper.rs', lines 7:0-7:28 -/ def ref_incr (x : I32) : Result I32 := x + 1#i32 /- [paper::test_incr]: - Source: 'tests/src/paper.rs', lines 10:0-10:18 -/ + Source: 'tests/src/paper.rs', lines 11:0-11:18 -/ def test_incr : Result Unit := do let x ← ref_incr 0#i32 @@ -23,7 +23,7 @@ def test_incr : Result Unit := #assert (test_incr == Result.ok ()) /- [paper::choose]: - Source: 'tests/src/paper.rs', lines 17:0-17:70 -/ + Source: 'tests/src/paper.rs', lines 18:0-18:70 -/ def choose (T : Type) (b : Bool) (x : T) (y : T) : Result (T × (T → Result (T × T))) @@ -35,7 +35,7 @@ def choose Result.ok (y, back) /- [paper::test_choose]: - Source: 'tests/src/paper.rs', lines 25:0-25:20 -/ + Source: 'tests/src/paper.rs', lines 26:0-26:20 -/ def test_choose : Result Unit := do let (z, choose_back) ← choose I32 true 0#i32 0#i32 @@ -55,13 +55,13 @@ def test_choose : Result Unit := #assert (test_choose == Result.ok ()) /- [paper::List] - Source: 'tests/src/paper.rs', lines 37:0-37:16 -/ + Source: 'tests/src/paper.rs', lines 38:0-38:16 -/ inductive List (T : Type) := | Cons : T → List T → List T | Nil : List T /- [paper::list_nth_mut]: - Source: 'tests/src/paper.rs', lines 44:0-44:67 -/ + Source: 'tests/src/paper.rs', lines 45:0-45:67 -/ divergent def list_nth_mut (T : Type) (l : List T) (i : U32) : Result (T × (T → Result (List T))) := match l with @@ -83,7 +83,7 @@ divergent def list_nth_mut | List.Nil => Result.fail .panic /- [paper::sum]: - Source: 'tests/src/paper.rs', lines 59:0-59:32 -/ + Source: 'tests/src/paper.rs', lines 60:0-60:32 -/ divergent def sum (l : List I32) : Result I32 := match l with | List.Cons x tl => do @@ -92,7 +92,7 @@ divergent def sum (l : List I32) : Result I32 := | List.Nil => Result.ok 0#i32 /- [paper::test_nth]: - Source: 'tests/src/paper.rs', lines 70:0-70:17 -/ + Source: 'tests/src/paper.rs', lines 71:0-71:17 -/ def test_nth : Result Unit := do let l := List.Cons 3#i32 List.Nil @@ -109,7 +109,7 @@ def test_nth : Result Unit := #assert (test_nth == Result.ok ()) /- [paper::call_choose]: - Source: 'tests/src/paper.rs', lines 78:0-78:44 -/ + Source: 'tests/src/paper.rs', lines 79:0-79:44 -/ def call_choose (p : (U32 × U32)) : Result U32 := do let (px, py) := p diff --git a/tests/lean/PoloniusList.lean b/tests/lean/PoloniusList.lean index defa48c7..ed279d58 100644 --- a/tests/lean/PoloniusList.lean +++ b/tests/lean/PoloniusList.lean @@ -6,13 +6,13 @@ open Primitives namespace polonius_list /- [polonius_list::List] - Source: 'tests/src/polonius_list.rs', lines 5:0-5:16 -/ + Source: 'tests/src/polonius_list.rs', lines 6:0-6:16 -/ inductive List (T : Type) := | Cons : T → List T → List T | Nil : List T /- [polonius_list::get_list_at_x]: - Source: 'tests/src/polonius_list.rs', lines 15:0-15:76 -/ + Source: 'tests/src/polonius_list.rs', lines 16:0-16:76 -/ divergent def get_list_at_x (ls : List U32) (x : U32) : Result ((List U32) × (List U32 → Result (List U32))) -- cgit v1.2.3 From 3adbe18d36df3767e98f30b760ccd9c6ace640ad Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Fri, 24 May 2024 17:01:16 +0200 Subject: Rename some subdirectories for consistency --- tests/coq/hashmap_main/HashmapMain_Funs.v | 589 ++++++++ tests/coq/hashmap_main/HashmapMain_FunsExternal.v | 24 + .../HashmapMain_FunsExternal_Template.v | 26 + tests/coq/hashmap_main/HashmapMain_Types.v | 40 + tests/coq/hashmap_main/HashmapMain_TypesExternal.v | 14 + .../HashmapMain_TypesExternal_Template.v | 15 + tests/coq/hashmap_main/Makefile | 23 + tests/coq/hashmap_main/Primitives.v | 981 ++++++++++++ tests/coq/hashmap_main/_CoqProject | 12 + tests/coq/hashmap_on_disk/HashmapMain_Funs.v | 589 -------- .../coq/hashmap_on_disk/HashmapMain_FunsExternal.v | 24 - .../HashmapMain_FunsExternal_Template.v | 26 - tests/coq/hashmap_on_disk/HashmapMain_Types.v | 40 - .../hashmap_on_disk/HashmapMain_TypesExternal.v | 14 - .../HashmapMain_TypesExternal_Template.v | 15 - tests/coq/hashmap_on_disk/Makefile | 23 - tests/coq/hashmap_on_disk/Primitives.v | 981 ------------ tests/coq/hashmap_on_disk/_CoqProject | 12 - .../hashmap_main/HashmapMain.Clauses.Template.fst | 72 + tests/fstar/hashmap_main/HashmapMain.Clauses.fst | 61 + tests/fstar/hashmap_main/HashmapMain.Funs.fst | 446 ++++++ .../hashmap_main/HashmapMain.FunsExternal.fsti | 18 + .../fstar/hashmap_main/HashmapMain.Properties.fst | 48 + tests/fstar/hashmap_main/HashmapMain.Types.fst | 24 + .../hashmap_main/HashmapMain.TypesExternal.fsti | 10 + tests/fstar/hashmap_main/Makefile | 49 + tests/fstar/hashmap_main/Primitives.fst | 929 ++++++++++++ .../HashmapMain.Clauses.Template.fst | 72 - .../fstar/hashmap_on_disk/HashmapMain.Clauses.fst | 61 - tests/fstar/hashmap_on_disk/HashmapMain.Funs.fst | 446 ------ .../hashmap_on_disk/HashmapMain.FunsExternal.fsti | 18 - .../hashmap_on_disk/HashmapMain.Properties.fst | 48 - tests/fstar/hashmap_on_disk/HashmapMain.Types.fst | 24 - .../hashmap_on_disk/HashmapMain.TypesExternal.fsti | 10 - tests/fstar/hashmap_on_disk/Makefile | 49 - tests/fstar/hashmap_on_disk/Primitives.fst | 929 ------------ tests/hol4/constants/Holmakefile | 5 + tests/hol4/constants/constantsScript.sml | 217 +++ tests/hol4/constants/constantsTheory.sig | 538 +++++++ tests/hol4/external/Holmakefile | 5 + tests/hol4/external/external_FunsScript.sml | 108 ++ tests/hol4/external/external_FunsTheory.sig | 105 ++ tests/hol4/external/external_OpaqueScript.sml | 25 + tests/hol4/external/external_OpaqueTheory.sig | 11 + tests/hol4/external/external_TypesScript.sml | 13 + tests/hol4/external/external_TypesTheory.sig | 11 + tests/hol4/hashmap_main/Holmakefile | 5 + tests/hol4/hashmap_main/hashmapMain_FunsScript.sml | 647 ++++++++ tests/hol4/hashmap_main/hashmapMain_FunsTheory.sig | 598 ++++++++ .../hol4/hashmap_main/hashmapMain_OpaqueScript.sml | 15 + .../hol4/hashmap_main/hashmapMain_OpaqueTheory.sig | 11 + .../hol4/hashmap_main/hashmapMain_TypesScript.sml | 27 + .../hol4/hashmap_main/hashmapMain_TypesTheory.sig | 568 +++++++ tests/hol4/hashmap_on_disk/Holmakefile | 5 - .../hashmap_on_disk/hashmapMain_FunsScript.sml | 647 -------- .../hashmap_on_disk/hashmapMain_FunsTheory.sig | 598 -------- .../hashmap_on_disk/hashmapMain_OpaqueScript.sml | 15 - .../hashmap_on_disk/hashmapMain_OpaqueTheory.sig | 11 - .../hashmap_on_disk/hashmapMain_TypesScript.sml | 27 - .../hashmap_on_disk/hashmapMain_TypesTheory.sig | 568 ------- tests/hol4/loops/Holmakefile | 5 + tests/hol4/loops/loops_FunsScript.sml | 755 +++++++++ tests/hol4/loops/loops_FunsTheory.sig | 731 +++++++++ tests/hol4/loops/loops_TypesScript.sml | 13 + tests/hol4/loops/loops_TypesTheory.sig | 94 ++ tests/hol4/misc-constants/Holmakefile | 5 - tests/hol4/misc-constants/constantsScript.sml | 217 --- tests/hol4/misc-constants/constantsTheory.sig | 538 ------- tests/hol4/misc-external/Holmakefile | 5 - tests/hol4/misc-external/external_FunsScript.sml | 108 -- tests/hol4/misc-external/external_FunsTheory.sig | 105 -- tests/hol4/misc-external/external_OpaqueScript.sml | 25 - tests/hol4/misc-external/external_OpaqueTheory.sig | 11 - tests/hol4/misc-external/external_TypesScript.sml | 13 - tests/hol4/misc-external/external_TypesTheory.sig | 11 - tests/hol4/misc-loops/Holmakefile | 5 - tests/hol4/misc-loops/loops_FunsScript.sml | 755 --------- tests/hol4/misc-loops/loops_FunsTheory.sig | 731 --------- tests/hol4/misc-loops/loops_TypesScript.sml | 13 - tests/hol4/misc-loops/loops_TypesTheory.sig | 94 -- tests/hol4/misc-no_nested_borrows/Holmakefile | 5 - .../noNestedBorrowsScript.sml | 592 -------- .../noNestedBorrowsTheory.sig | 1598 -------------------- tests/hol4/misc-paper/Holmakefile | 5 - tests/hol4/misc-paper/paperScript.sml | 136 -- tests/hol4/misc-paper/paperTheory.sig | 210 --- tests/hol4/misc-polonius_list/Holmakefile | 5 - .../hol4/misc-polonius_list/poloniusListScript.sml | 37 - .../hol4/misc-polonius_list/poloniusListTheory.sig | 120 -- tests/hol4/no_nested_borrows/Holmakefile | 5 + .../no_nested_borrows/noNestedBorrowsScript.sml | 592 ++++++++ .../no_nested_borrows/noNestedBorrowsTheory.sig | 1598 ++++++++++++++++++++ tests/hol4/paper/Holmakefile | 5 + tests/hol4/paper/paperScript.sml | 136 ++ tests/hol4/paper/paperTheory.sig | 210 +++ tests/hol4/polonius_list/Holmakefile | 5 + tests/hol4/polonius_list/poloniusListScript.sml | 37 + tests/hol4/polonius_list/poloniusListTheory.sig | 120 ++ tests/test_runner/run_test.ml | 20 +- 99 files changed, 10603 insertions(+), 10609 deletions(-) create mode 100644 tests/coq/hashmap_main/HashmapMain_Funs.v create mode 100644 tests/coq/hashmap_main/HashmapMain_FunsExternal.v create mode 100644 tests/coq/hashmap_main/HashmapMain_FunsExternal_Template.v create mode 100644 tests/coq/hashmap_main/HashmapMain_Types.v create mode 100644 tests/coq/hashmap_main/HashmapMain_TypesExternal.v create mode 100644 tests/coq/hashmap_main/HashmapMain_TypesExternal_Template.v create mode 100644 tests/coq/hashmap_main/Makefile create mode 100644 tests/coq/hashmap_main/Primitives.v create mode 100644 tests/coq/hashmap_main/_CoqProject delete mode 100644 tests/coq/hashmap_on_disk/HashmapMain_Funs.v delete mode 100644 tests/coq/hashmap_on_disk/HashmapMain_FunsExternal.v delete mode 100644 tests/coq/hashmap_on_disk/HashmapMain_FunsExternal_Template.v delete mode 100644 tests/coq/hashmap_on_disk/HashmapMain_Types.v delete mode 100644 tests/coq/hashmap_on_disk/HashmapMain_TypesExternal.v delete mode 100644 tests/coq/hashmap_on_disk/HashmapMain_TypesExternal_Template.v delete mode 100644 tests/coq/hashmap_on_disk/Makefile delete mode 100644 tests/coq/hashmap_on_disk/Primitives.v delete mode 100644 tests/coq/hashmap_on_disk/_CoqProject create mode 100644 tests/fstar/hashmap_main/HashmapMain.Clauses.Template.fst create mode 100644 tests/fstar/hashmap_main/HashmapMain.Clauses.fst create mode 100644 tests/fstar/hashmap_main/HashmapMain.Funs.fst create mode 100644 tests/fstar/hashmap_main/HashmapMain.FunsExternal.fsti create mode 100644 tests/fstar/hashmap_main/HashmapMain.Properties.fst create mode 100644 tests/fstar/hashmap_main/HashmapMain.Types.fst create mode 100644 tests/fstar/hashmap_main/HashmapMain.TypesExternal.fsti create mode 100644 tests/fstar/hashmap_main/Makefile create mode 100644 tests/fstar/hashmap_main/Primitives.fst delete mode 100644 tests/fstar/hashmap_on_disk/HashmapMain.Clauses.Template.fst delete mode 100644 tests/fstar/hashmap_on_disk/HashmapMain.Clauses.fst delete mode 100644 tests/fstar/hashmap_on_disk/HashmapMain.Funs.fst delete mode 100644 tests/fstar/hashmap_on_disk/HashmapMain.FunsExternal.fsti delete mode 100644 tests/fstar/hashmap_on_disk/HashmapMain.Properties.fst delete mode 100644 tests/fstar/hashmap_on_disk/HashmapMain.Types.fst delete mode 100644 tests/fstar/hashmap_on_disk/HashmapMain.TypesExternal.fsti delete mode 100644 tests/fstar/hashmap_on_disk/Makefile delete mode 100644 tests/fstar/hashmap_on_disk/Primitives.fst create mode 100644 tests/hol4/constants/Holmakefile create mode 100644 tests/hol4/constants/constantsScript.sml create mode 100644 tests/hol4/constants/constantsTheory.sig create mode 100644 tests/hol4/external/Holmakefile create mode 100644 tests/hol4/external/external_FunsScript.sml create mode 100644 tests/hol4/external/external_FunsTheory.sig create mode 100644 tests/hol4/external/external_OpaqueScript.sml create mode 100644 tests/hol4/external/external_OpaqueTheory.sig create mode 100644 tests/hol4/external/external_TypesScript.sml create mode 100644 tests/hol4/external/external_TypesTheory.sig create mode 100644 tests/hol4/hashmap_main/Holmakefile create mode 100644 tests/hol4/hashmap_main/hashmapMain_FunsScript.sml create mode 100644 tests/hol4/hashmap_main/hashmapMain_FunsTheory.sig create mode 100644 tests/hol4/hashmap_main/hashmapMain_OpaqueScript.sml create mode 100644 tests/hol4/hashmap_main/hashmapMain_OpaqueTheory.sig create mode 100644 tests/hol4/hashmap_main/hashmapMain_TypesScript.sml create mode 100644 tests/hol4/hashmap_main/hashmapMain_TypesTheory.sig delete mode 100644 tests/hol4/hashmap_on_disk/Holmakefile delete mode 100644 tests/hol4/hashmap_on_disk/hashmapMain_FunsScript.sml delete mode 100644 tests/hol4/hashmap_on_disk/hashmapMain_FunsTheory.sig delete mode 100644 tests/hol4/hashmap_on_disk/hashmapMain_OpaqueScript.sml delete mode 100644 tests/hol4/hashmap_on_disk/hashmapMain_OpaqueTheory.sig delete mode 100644 tests/hol4/hashmap_on_disk/hashmapMain_TypesScript.sml delete mode 100644 tests/hol4/hashmap_on_disk/hashmapMain_TypesTheory.sig create mode 100644 tests/hol4/loops/Holmakefile create mode 100644 tests/hol4/loops/loops_FunsScript.sml create mode 100644 tests/hol4/loops/loops_FunsTheory.sig create mode 100644 tests/hol4/loops/loops_TypesScript.sml create mode 100644 tests/hol4/loops/loops_TypesTheory.sig delete mode 100644 tests/hol4/misc-constants/Holmakefile delete mode 100644 tests/hol4/misc-constants/constantsScript.sml delete mode 100644 tests/hol4/misc-constants/constantsTheory.sig delete mode 100644 tests/hol4/misc-external/Holmakefile delete mode 100644 tests/hol4/misc-external/external_FunsScript.sml delete mode 100644 tests/hol4/misc-external/external_FunsTheory.sig delete mode 100644 tests/hol4/misc-external/external_OpaqueScript.sml delete mode 100644 tests/hol4/misc-external/external_OpaqueTheory.sig delete mode 100644 tests/hol4/misc-external/external_TypesScript.sml delete mode 100644 tests/hol4/misc-external/external_TypesTheory.sig delete mode 100644 tests/hol4/misc-loops/Holmakefile delete mode 100644 tests/hol4/misc-loops/loops_FunsScript.sml delete mode 100644 tests/hol4/misc-loops/loops_FunsTheory.sig delete mode 100644 tests/hol4/misc-loops/loops_TypesScript.sml delete mode 100644 tests/hol4/misc-loops/loops_TypesTheory.sig delete mode 100644 tests/hol4/misc-no_nested_borrows/Holmakefile delete mode 100644 tests/hol4/misc-no_nested_borrows/noNestedBorrowsScript.sml delete mode 100644 tests/hol4/misc-no_nested_borrows/noNestedBorrowsTheory.sig delete mode 100644 tests/hol4/misc-paper/Holmakefile delete mode 100644 tests/hol4/misc-paper/paperScript.sml delete mode 100644 tests/hol4/misc-paper/paperTheory.sig delete mode 100644 tests/hol4/misc-polonius_list/Holmakefile delete mode 100644 tests/hol4/misc-polonius_list/poloniusListScript.sml delete mode 100644 tests/hol4/misc-polonius_list/poloniusListTheory.sig create mode 100644 tests/hol4/no_nested_borrows/Holmakefile create mode 100644 tests/hol4/no_nested_borrows/noNestedBorrowsScript.sml create mode 100644 tests/hol4/no_nested_borrows/noNestedBorrowsTheory.sig create mode 100644 tests/hol4/paper/Holmakefile create mode 100644 tests/hol4/paper/paperScript.sml create mode 100644 tests/hol4/paper/paperTheory.sig create mode 100644 tests/hol4/polonius_list/Holmakefile create mode 100644 tests/hol4/polonius_list/poloniusListScript.sml create mode 100644 tests/hol4/polonius_list/poloniusListTheory.sig diff --git a/tests/coq/hashmap_main/HashmapMain_Funs.v b/tests/coq/hashmap_main/HashmapMain_Funs.v new file mode 100644 index 00000000..f6467d5a --- /dev/null +++ b/tests/coq/hashmap_main/HashmapMain_Funs.v @@ -0,0 +1,589 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: function definitions *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Require Import HashmapMain_Types. +Include HashmapMain_Types. +Require Import HashmapMain_FunsExternal. +Include HashmapMain_FunsExternal. +Module HashmapMain_Funs. + +(** [hashmap_main::hashmap::hash_key]: + Source: 'tests/src/hashmap.rs', lines 35:0-35:32 *) +Definition hashmap_hash_key (k : usize) : result usize := + Ok k. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: loop 0: + Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) +Fixpoint hashmap_HashMap_allocate_slots_loop + (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (n1 : usize) + : + result (alloc_vec_Vec (hashmap_List_t T)) + := + match n with + | O => Fail_ OutOfFuel + | S n2 => + if n1 s> 0%usize + then ( + slots1 <- alloc_vec_Vec_push (hashmap_List_t T) slots Hashmap_List_Nil; + n3 <- usize_sub n1 1%usize; + hashmap_HashMap_allocate_slots_loop T n2 slots1 n3) + else Ok slots + end +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: + Source: 'tests/src/hashmap.rs', lines 58:4-58:76 *) +Definition hashmap_HashMap_allocate_slots + (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (n1 : usize) + : + result (alloc_vec_Vec (hashmap_List_t T)) + := + hashmap_HashMap_allocate_slots_loop T n slots n1 +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new_with_capacity]: + Source: 'tests/src/hashmap.rs', lines 67:4-71:13 *) +Definition hashmap_HashMap_new_with_capacity + (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) + (max_load_divisor : usize) : + result (hashmap_HashMap_t T) + := + slots <- + hashmap_HashMap_allocate_slots T n (alloc_vec_Vec_new (hashmap_List_t T)) + capacity; + i <- usize_mul capacity max_load_dividend; + i1 <- usize_div i max_load_divisor; + Ok + {| + hashmap_HashMap_num_entries := 0%usize; + hashmap_HashMap_max_load_factor := (max_load_dividend, max_load_divisor); + hashmap_HashMap_max_load := i1; + hashmap_HashMap_slots := slots + |} +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new]: + Source: 'tests/src/hashmap.rs', lines 83:4-83:24 *) +Definition hashmap_HashMap_new + (T : Type) (n : nat) : result (hashmap_HashMap_t T) := + hashmap_HashMap_new_with_capacity T n 32%usize 4%usize 5%usize +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: loop 0: + Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) +Fixpoint hashmap_HashMap_clear_loop + (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) : + result (alloc_vec_Vec (hashmap_List_t T)) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + let i1 := alloc_vec_Vec_len (hashmap_List_t T) slots in + if i s< i1 + then ( + p <- + alloc_vec_Vec_index_mut (hashmap_List_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) slots + i; + let (_, index_mut_back) := p in + i2 <- usize_add i 1%usize; + slots1 <- index_mut_back Hashmap_List_Nil; + hashmap_HashMap_clear_loop T n1 slots1 i2) + else Ok slots + end +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: + Source: 'tests/src/hashmap.rs', lines 88:4-88:27 *) +Definition hashmap_HashMap_clear + (T : Type) (n : nat) (self : hashmap_HashMap_t T) : + result (hashmap_HashMap_t T) + := + hm <- hashmap_HashMap_clear_loop T n self.(hashmap_HashMap_slots) 0%usize; + Ok + {| + hashmap_HashMap_num_entries := 0%usize; + hashmap_HashMap_max_load_factor := self.(hashmap_HashMap_max_load_factor); + hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); + hashmap_HashMap_slots := hm + |} +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::len]: + Source: 'tests/src/hashmap.rs', lines 98:4-98:30 *) +Definition hashmap_HashMap_len + (T : Type) (self : hashmap_HashMap_t T) : result usize := + Ok self.(hashmap_HashMap_num_entries) +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: loop 0: + Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) +Fixpoint hashmap_HashMap_insert_in_list_loop + (T : Type) (n : nat) (key : usize) (value : T) (ls : hashmap_List_t T) : + result (bool * (hashmap_List_t T)) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match ls with + | Hashmap_List_Cons ckey cvalue tl => + if ckey s= key + then Ok (false, Hashmap_List_Cons ckey value tl) + else ( + p <- hashmap_HashMap_insert_in_list_loop T n1 key value tl; + let (b, tl1) := p in + Ok (b, Hashmap_List_Cons ckey cvalue tl1)) + | Hashmap_List_Nil => + Ok (true, Hashmap_List_Cons key value Hashmap_List_Nil) + end + end +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: + Source: 'tests/src/hashmap.rs', lines 105:4-105:71 *) +Definition hashmap_HashMap_insert_in_list + (T : Type) (n : nat) (key : usize) (value : T) (ls : hashmap_List_t T) : + result (bool * (hashmap_List_t T)) + := + hashmap_HashMap_insert_in_list_loop T n key value ls +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_no_resize]: + Source: 'tests/src/hashmap.rs', lines 125:4-125:54 *) +Definition hashmap_HashMap_insert_no_resize + (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) (value : T) : + result (hashmap_HashMap_t T) + := + hash <- hashmap_hash_key key; + let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in + hash_mod <- usize_rem hash i; + p <- + alloc_vec_Vec_index_mut (hashmap_List_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) + self.(hashmap_HashMap_slots) hash_mod; + let (l, index_mut_back) := p in + p1 <- hashmap_HashMap_insert_in_list T n key value l; + let (inserted, l1) := p1 in + if inserted + then ( + i1 <- usize_add self.(hashmap_HashMap_num_entries) 1%usize; + v <- index_mut_back l1; + Ok + {| + hashmap_HashMap_num_entries := i1; + hashmap_HashMap_max_load_factor := + self.(hashmap_HashMap_max_load_factor); + hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); + hashmap_HashMap_slots := v + |}) + else ( + v <- index_mut_back l1; + Ok + {| + hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); + hashmap_HashMap_max_load_factor := + self.(hashmap_HashMap_max_load_factor); + hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); + hashmap_HashMap_slots := v + |}) +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: loop 0: + Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) +Fixpoint hashmap_HashMap_move_elements_from_list_loop + (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) (ls : hashmap_List_t T) : + result (hashmap_HashMap_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match ls with + | Hashmap_List_Cons k v tl => + ntable1 <- hashmap_HashMap_insert_no_resize T n1 ntable k v; + hashmap_HashMap_move_elements_from_list_loop T n1 ntable1 tl + | Hashmap_List_Nil => Ok ntable + end + end +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: + Source: 'tests/src/hashmap.rs', lines 191:4-191:72 *) +Definition hashmap_HashMap_move_elements_from_list + (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) (ls : hashmap_List_t T) : + result (hashmap_HashMap_t T) + := + hashmap_HashMap_move_elements_from_list_loop T n ntable ls +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: loop 0: + Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) +Fixpoint hashmap_HashMap_move_elements_loop + (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) + (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) : + result ((hashmap_HashMap_t T) * (alloc_vec_Vec (hashmap_List_t T))) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + let i1 := alloc_vec_Vec_len (hashmap_List_t T) slots in + if i s< i1 + then ( + p <- + alloc_vec_Vec_index_mut (hashmap_List_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) slots + i; + let (l, index_mut_back) := p in + let (ls, l1) := core_mem_replace (hashmap_List_t T) l Hashmap_List_Nil in + ntable1 <- hashmap_HashMap_move_elements_from_list T n1 ntable ls; + i2 <- usize_add i 1%usize; + slots1 <- index_mut_back l1; + hashmap_HashMap_move_elements_loop T n1 ntable1 slots1 i2) + else Ok (ntable, slots) + end +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: + Source: 'tests/src/hashmap.rs', lines 179:4-179:95 *) +Definition hashmap_HashMap_move_elements + (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) + (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) : + result ((hashmap_HashMap_t T) * (alloc_vec_Vec (hashmap_List_t T))) + := + hashmap_HashMap_move_elements_loop T n ntable slots i +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::try_resize]: + Source: 'tests/src/hashmap.rs', lines 148:4-148:28 *) +Definition hashmap_HashMap_try_resize + (T : Type) (n : nat) (self : hashmap_HashMap_t T) : + result (hashmap_HashMap_t T) + := + max_usize <- scalar_cast U32 Usize core_u32_max; + let capacity := + alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in + n1 <- usize_div max_usize 2%usize; + let (i, i1) := self.(hashmap_HashMap_max_load_factor) in + i2 <- usize_div n1 i; + if capacity s<= i2 + then ( + i3 <- usize_mul capacity 2%usize; + ntable <- hashmap_HashMap_new_with_capacity T n i3 i i1; + p <- + hashmap_HashMap_move_elements T n ntable self.(hashmap_HashMap_slots) + 0%usize; + let (ntable1, _) := p in + Ok + {| + hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); + hashmap_HashMap_max_load_factor := (i, i1); + hashmap_HashMap_max_load := ntable1.(hashmap_HashMap_max_load); + hashmap_HashMap_slots := ntable1.(hashmap_HashMap_slots) + |}) + else + Ok + {| + hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); + hashmap_HashMap_max_load_factor := (i, i1); + hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); + hashmap_HashMap_slots := self.(hashmap_HashMap_slots) + |} +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert]: + Source: 'tests/src/hashmap.rs', lines 137:4-137:48 *) +Definition hashmap_HashMap_insert + (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) (value : T) : + result (hashmap_HashMap_t T) + := + self1 <- hashmap_HashMap_insert_no_resize T n self key value; + i <- hashmap_HashMap_len T self1; + if i s> self1.(hashmap_HashMap_max_load) + then hashmap_HashMap_try_resize T n self1 + else Ok self1 +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: loop 0: + Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) +Fixpoint hashmap_HashMap_contains_key_in_list_loop + (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result bool := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match ls with + | Hashmap_List_Cons ckey _ tl => + if ckey s= key + then Ok true + else hashmap_HashMap_contains_key_in_list_loop T n1 key tl + | Hashmap_List_Nil => Ok false + end + end +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: + Source: 'tests/src/hashmap.rs', lines 214:4-214:68 *) +Definition hashmap_HashMap_contains_key_in_list + (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result bool := + hashmap_HashMap_contains_key_in_list_loop T n key ls +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key]: + Source: 'tests/src/hashmap.rs', lines 207:4-207:49 *) +Definition hashmap_HashMap_contains_key + (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : + result bool + := + hash <- hashmap_hash_key key; + let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in + hash_mod <- usize_rem hash i; + l <- + alloc_vec_Vec_index (hashmap_List_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) + self.(hashmap_HashMap_slots) hash_mod; + hashmap_HashMap_contains_key_in_list T n key l +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: loop 0: + Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) +Fixpoint hashmap_HashMap_get_in_list_loop + (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result T := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match ls with + | Hashmap_List_Cons ckey cvalue tl => + if ckey s= key + then Ok cvalue + else hashmap_HashMap_get_in_list_loop T n1 key tl + | Hashmap_List_Nil => Fail_ Failure + end + end +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: + Source: 'tests/src/hashmap.rs', lines 232:4-232:70 *) +Definition hashmap_HashMap_get_in_list + (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result T := + hashmap_HashMap_get_in_list_loop T n key ls +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get]: + Source: 'tests/src/hashmap.rs', lines 247:4-247:55 *) +Definition hashmap_HashMap_get + (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : result T := + hash <- hashmap_hash_key key; + let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in + hash_mod <- usize_rem hash i; + l <- + alloc_vec_Vec_index (hashmap_List_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) + self.(hashmap_HashMap_slots) hash_mod; + hashmap_HashMap_get_in_list T n key l +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: loop 0: + Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) +Fixpoint hashmap_HashMap_get_mut_in_list_loop + (T : Type) (n : nat) (ls : hashmap_List_t T) (key : usize) : + result (T * (T -> result (hashmap_List_t T))) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match ls with + | Hashmap_List_Cons ckey cvalue tl => + if ckey s= key + then + let back := fun (ret : T) => Ok (Hashmap_List_Cons ckey ret tl) in + Ok (cvalue, back) + else ( + p <- hashmap_HashMap_get_mut_in_list_loop T n1 tl key; + let (t, back) := p in + let back1 := + fun (ret : T) => + tl1 <- back ret; Ok (Hashmap_List_Cons ckey cvalue tl1) in + Ok (t, back1)) + | Hashmap_List_Nil => Fail_ Failure + end + end +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: + Source: 'tests/src/hashmap.rs', lines 253:4-253:86 *) +Definition hashmap_HashMap_get_mut_in_list + (T : Type) (n : nat) (ls : hashmap_List_t T) (key : usize) : + result (T * (T -> result (hashmap_List_t T))) + := + hashmap_HashMap_get_mut_in_list_loop T n ls key +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut]: + Source: 'tests/src/hashmap.rs', lines 265:4-265:67 *) +Definition hashmap_HashMap_get_mut + (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : + result (T * (T -> result (hashmap_HashMap_t T))) + := + hash <- hashmap_hash_key key; + let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in + hash_mod <- usize_rem hash i; + p <- + alloc_vec_Vec_index_mut (hashmap_List_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) + self.(hashmap_HashMap_slots) hash_mod; + let (l, index_mut_back) := p in + p1 <- hashmap_HashMap_get_mut_in_list T n l 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; + Ok + {| + hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); + hashmap_HashMap_max_load_factor := + self.(hashmap_HashMap_max_load_factor); + hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); + hashmap_HashMap_slots := v + |} in + Ok (t, back) +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: loop 0: + Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) +Fixpoint hashmap_HashMap_remove_from_list_loop + (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : + result ((option T) * (hashmap_List_t T)) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match ls with + | Hashmap_List_Cons ckey t tl => + if ckey s= key + then + let (mv_ls, _) := + core_mem_replace (hashmap_List_t T) (Hashmap_List_Cons ckey t tl) + Hashmap_List_Nil in + match mv_ls with + | Hashmap_List_Cons _ cvalue tl1 => Ok (Some cvalue, tl1) + | Hashmap_List_Nil => Fail_ Failure + end + else ( + p <- hashmap_HashMap_remove_from_list_loop T n1 key tl; + let (o, tl1) := p in + Ok (o, Hashmap_List_Cons ckey t tl1)) + | Hashmap_List_Nil => Ok (None, Hashmap_List_Nil) + end + end +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: + Source: 'tests/src/hashmap.rs', lines 273:4-273:69 *) +Definition hashmap_HashMap_remove_from_list + (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : + result ((option T) * (hashmap_List_t T)) + := + hashmap_HashMap_remove_from_list_loop T n key ls +. + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove]: + Source: 'tests/src/hashmap.rs', lines 302:4-302:52 *) +Definition hashmap_HashMap_remove + (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : + result ((option T) * (hashmap_HashMap_t T)) + := + hash <- hashmap_hash_key key; + let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in + hash_mod <- usize_rem hash i; + p <- + alloc_vec_Vec_index_mut (hashmap_List_t T) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) + self.(hashmap_HashMap_slots) hash_mod; + let (l, index_mut_back) := p in + p1 <- hashmap_HashMap_remove_from_list T n key l; + let (x, l1) := p1 in + match x with + | None => + v <- index_mut_back l1; + Ok (None, + {| + hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); + hashmap_HashMap_max_load_factor := + self.(hashmap_HashMap_max_load_factor); + hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); + hashmap_HashMap_slots := v + |}) + | Some x1 => + i1 <- usize_sub self.(hashmap_HashMap_num_entries) 1%usize; + v <- index_mut_back l1; + Ok (Some x1, + {| + hashmap_HashMap_num_entries := i1; + hashmap_HashMap_max_load_factor := + self.(hashmap_HashMap_max_load_factor); + hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); + hashmap_HashMap_slots := v + |}) + end +. + +(** [hashmap_main::hashmap::test1]: + Source: 'tests/src/hashmap.rs', lines 323:0-323:10 *) +Definition hashmap_test1 (n : nat) : result unit := + hm <- hashmap_HashMap_new u64 n; + hm1 <- hashmap_HashMap_insert u64 n hm 0%usize 42%u64; + hm2 <- hashmap_HashMap_insert u64 n hm1 128%usize 18%u64; + hm3 <- hashmap_HashMap_insert u64 n hm2 1024%usize 138%u64; + hm4 <- hashmap_HashMap_insert u64 n hm3 1056%usize 256%u64; + i <- hashmap_HashMap_get u64 n hm4 128%usize; + if negb (i s= 18%u64) + then Fail_ Failure + else ( + p <- hashmap_HashMap_get_mut u64 n hm4 1024%usize; + let (_, get_mut_back) := p in + hm5 <- get_mut_back 56%u64; + i1 <- hashmap_HashMap_get u64 n hm5 1024%usize; + if negb (i1 s= 56%u64) + then Fail_ Failure + else ( + p1 <- hashmap_HashMap_remove u64 n hm5 1024%usize; + let (x, hm6) := p1 in + match x with + | None => Fail_ Failure + | Some x1 => + if negb (x1 s= 56%u64) + then Fail_ Failure + else ( + i2 <- hashmap_HashMap_get u64 n hm6 0%usize; + if negb (i2 s= 42%u64) + then Fail_ Failure + else ( + i3 <- hashmap_HashMap_get u64 n hm6 128%usize; + if negb (i3 s= 18%u64) + then Fail_ Failure + else ( + i4 <- hashmap_HashMap_get u64 n hm6 1056%usize; + if negb (i4 s= 256%u64) then Fail_ Failure else Ok tt))) + end)) +. + +(** [hashmap_main::insert_on_disk]: + Source: 'tests/src/hashmap_main.rs', lines 13:0-13:43 *) +Definition insert_on_disk + (n : nat) (key : usize) (value : u64) (st : state) : result (state * unit) := + p <- hashmap_utils_deserialize st; + let (st1, hm) := p in + hm1 <- hashmap_HashMap_insert u64 n hm key value; + hashmap_utils_serialize hm1 st1 +. + +(** [hashmap_main::main]: + Source: 'tests/src/hashmap_main.rs', lines 22:0-22:13 *) +Definition main : result unit := + Ok tt. + +End HashmapMain_Funs. diff --git a/tests/coq/hashmap_main/HashmapMain_FunsExternal.v b/tests/coq/hashmap_main/HashmapMain_FunsExternal.v new file mode 100644 index 00000000..fb5f23cd --- /dev/null +++ b/tests/coq/hashmap_main/HashmapMain_FunsExternal.v @@ -0,0 +1,24 @@ +(** [hashmap_main]: external function declarations *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Require Export HashmapMain_Types. +Import HashmapMain_Types. +Module HashmapMain_FunsExternal. + +(** [hashmap_main::hashmap_utils::deserialize]: forward function + Source: 'src/hashmap_utils.rs', lines 10:0-10:43 *) +Axiom hashmap_utils_deserialize + : state -> result (state * (hashmap_HashMap_t u64)) +. + +(** [hashmap_main::hashmap_utils::serialize]: forward function + Source: 'src/hashmap_utils.rs', lines 5:0-5:42 *) +Axiom hashmap_utils_serialize + : hashmap_HashMap_t u64 -> state -> result (state * unit) +. + +End HashmapMain_FunsExternal. diff --git a/tests/coq/hashmap_main/HashmapMain_FunsExternal_Template.v b/tests/coq/hashmap_main/HashmapMain_FunsExternal_Template.v new file mode 100644 index 00000000..66835e8c --- /dev/null +++ b/tests/coq/hashmap_main/HashmapMain_FunsExternal_Template.v @@ -0,0 +1,26 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: external functions. +-- This is a template file: rename it to "FunsExternal.lean" and fill the holes. *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Require Import HashmapMain_Types. +Include HashmapMain_Types. +Module HashmapMain_FunsExternal_Template. + +(** [hashmap_main::hashmap_utils::deserialize]: + Source: 'tests/src/hashmap_utils.rs', lines 11:0-11:43 *) +Axiom hashmap_utils_deserialize + : state -> result (state * (hashmap_HashMap_t u64)) +. + +(** [hashmap_main::hashmap_utils::serialize]: + Source: 'tests/src/hashmap_utils.rs', lines 6:0-6:42 *) +Axiom hashmap_utils_serialize + : hashmap_HashMap_t u64 -> state -> result (state * unit) +. + +End HashmapMain_FunsExternal_Template. diff --git a/tests/coq/hashmap_main/HashmapMain_Types.v b/tests/coq/hashmap_main/HashmapMain_Types.v new file mode 100644 index 00000000..5656bd9c --- /dev/null +++ b/tests/coq/hashmap_main/HashmapMain_Types.v @@ -0,0 +1,40 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: type definitions *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Require Import HashmapMain_TypesExternal. +Include HashmapMain_TypesExternal. +Module HashmapMain_Types. + +(** [hashmap_main::hashmap::List] + Source: 'tests/src/hashmap.rs', lines 27:0-27:16 *) +Inductive hashmap_List_t (T : Type) := +| Hashmap_List_Cons : usize -> T -> hashmap_List_t T -> hashmap_List_t T +| Hashmap_List_Nil : hashmap_List_t T +. + +Arguments Hashmap_List_Cons { _ }. +Arguments Hashmap_List_Nil { _ }. + +(** [hashmap_main::hashmap::HashMap] + Source: 'tests/src/hashmap.rs', lines 43:0-43:21 *) +Record hashmap_HashMap_t (T : Type) := +mkhashmap_HashMap_t { + hashmap_HashMap_num_entries : usize; + hashmap_HashMap_max_load_factor : (usize * usize); + hashmap_HashMap_max_load : usize; + hashmap_HashMap_slots : alloc_vec_Vec (hashmap_List_t T); +} +. + +Arguments mkhashmap_HashMap_t { _ }. +Arguments hashmap_HashMap_num_entries { _ }. +Arguments hashmap_HashMap_max_load_factor { _ }. +Arguments hashmap_HashMap_max_load { _ }. +Arguments hashmap_HashMap_slots { _ }. + +End HashmapMain_Types. diff --git a/tests/coq/hashmap_main/HashmapMain_TypesExternal.v b/tests/coq/hashmap_main/HashmapMain_TypesExternal.v new file mode 100644 index 00000000..28651c14 --- /dev/null +++ b/tests/coq/hashmap_main/HashmapMain_TypesExternal.v @@ -0,0 +1,14 @@ +(** [hashmap_main]: external types. +-- This is a template file: rename it to "TypesExternal.lean" and fill the holes. *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Module HashmapMain_TypesExternal. + +(** The state type used in the state-error monad *) +Axiom state : Type. + +End HashmapMain_TypesExternal. diff --git a/tests/coq/hashmap_main/HashmapMain_TypesExternal_Template.v b/tests/coq/hashmap_main/HashmapMain_TypesExternal_Template.v new file mode 100644 index 00000000..391b2775 --- /dev/null +++ b/tests/coq/hashmap_main/HashmapMain_TypesExternal_Template.v @@ -0,0 +1,15 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: external types. +-- This is a template file: rename it to "TypesExternal.lean" and fill the holes. *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Module HashmapMain_TypesExternal_Template. + +(** The state type used in the state-error monad *) +Axiom state : Type. + +End HashmapMain_TypesExternal_Template. diff --git a/tests/coq/hashmap_main/Makefile b/tests/coq/hashmap_main/Makefile new file mode 100644 index 00000000..1a5aee4a --- /dev/null +++ b/tests/coq/hashmap_main/Makefile @@ -0,0 +1,23 @@ +# This file was automatically generated - modify ../Makefile.template instead +# Makefile originally taken from coq-club + +%: Makefile.coq phony + +make -f Makefile.coq $@ + +all: Makefile.coq + +make -f Makefile.coq all + +clean: Makefile.coq + +make -f Makefile.coq clean + rm -f Makefile.coq + +Makefile.coq: _CoqProject Makefile + coq_makefile -f _CoqProject | sed 's/$$(COQCHK) $$(COQCHKFLAGS) $$(COQLIBS)/$$(COQCHK) $$(COQCHKFLAGS) $$(subst -Q,-R,$$(COQLIBS))/' > Makefile.coq + +_CoqProject: ; + +Makefile: ; + +phony: ; + +.PHONY: all clean phony diff --git a/tests/coq/hashmap_main/Primitives.v b/tests/coq/hashmap_main/Primitives.v new file mode 100644 index 00000000..b29fce43 --- /dev/null +++ b/tests/coq/hashmap_main/Primitives.v @@ -0,0 +1,981 @@ +Require Import Lia. +Require Coq.Strings.Ascii. +Require Coq.Strings.String. +Require Import Coq.Program.Equality. +Require Import Coq.ZArith.ZArith. +Require Import Coq.ZArith.Znat. +Require Import List. +Import ListNotations. + +Module Primitives. + + (* TODO: use more *) +Declare Scope Primitives_scope. + +(*** Result *) + +Inductive error := + | Failure + | OutOfFuel. + +Inductive result A := + | Ok : A -> result A + | Fail_ : error -> result A. + +Arguments Ok {_} a. +Arguments Fail_ {_}. + +Definition bind {A B} (m: result A) (f: A -> result B) : result B := + match m with + | Fail_ e => Fail_ e + | Ok x => f x + end. + +Definition return_ {A: Type} (x: A) : result A := Ok x. +Definition fail_ {A: Type} (e: error) : result A := Fail_ e. + +Notation "x <- c1 ; c2" := (bind c1 (fun x => c2)) + (at level 61, c1 at next level, right associativity). + +(** Monadic assert *) +Definition massert (b: bool) : result unit := + if b then Ok tt else Fail_ Failure. + +(** Normalize and unwrap a successful result (used for globals) *) +Definition eval_result_refl {A} {x} (a: result A) (p: a = Ok x) : A := + match a as r return (r = Ok x -> A) with + | Ok a' => fun _ => a' + | Fail_ e => fun p' => + False_rect _ (eq_ind (Fail_ e) + (fun e : result A => + match e with + | Ok _ => False + | Fail_ e => True + end) + I (Ok x) p') + end p. + +Notation "x %global" := (eval_result_refl x eq_refl) (at level 40). +Notation "x %return" := (eval_result_refl x eq_refl) (at level 40). + +(* Sanity check *) +Check (if true then Ok (1 + 2) else Fail_ Failure)%global = 3. + +(*** Misc *) + +Definition string := Coq.Strings.String.string. +Definition char := Coq.Strings.Ascii.ascii. +Definition char_of_byte := Coq.Strings.Ascii.ascii_of_byte. + +Definition core_mem_replace (a : Type) (x : a) (y : a) : a * a := (x, x) . + +Record mut_raw_ptr (T : Type) := { mut_raw_ptr_v : T }. +Record const_raw_ptr (T : Type) := { const_raw_ptr_v : T }. + +(*** Scalars *) + +Definition i8_min : Z := -128%Z. +Definition i8_max : Z := 127%Z. +Definition i16_min : Z := -32768%Z. +Definition i16_max : Z := 32767%Z. +Definition i32_min : Z := -2147483648%Z. +Definition i32_max : Z := 2147483647%Z. +Definition i64_min : Z := -9223372036854775808%Z. +Definition i64_max : Z := 9223372036854775807%Z. +Definition i128_min : Z := -170141183460469231731687303715884105728%Z. +Definition i128_max : Z := 170141183460469231731687303715884105727%Z. +Definition u8_min : Z := 0%Z. +Definition u8_max : Z := 255%Z. +Definition u16_min : Z := 0%Z. +Definition u16_max : Z := 65535%Z. +Definition u32_min : Z := 0%Z. +Definition u32_max : Z := 4294967295%Z. +Definition u64_min : Z := 0%Z. +Definition u64_max : Z := 18446744073709551615%Z. +Definition u128_min : Z := 0%Z. +Definition u128_max : Z := 340282366920938463463374607431768211455%Z. + +(** The bounds of [isize] and [usize] vary with the architecture. *) +Axiom isize_min : Z. +Axiom isize_max : Z. +Definition usize_min : Z := 0%Z. +Axiom usize_max : Z. + +Open Scope Z_scope. + +(** We provide those lemmas to reason about the bounds of [isize] and [usize] *) +Axiom isize_min_bound : isize_min <= i32_min. +Axiom isize_max_bound : i32_max <= isize_max. +Axiom usize_max_bound : u32_max <= usize_max. + +Inductive scalar_ty := + | Isize + | I8 + | I16 + | I32 + | I64 + | I128 + | Usize + | U8 + | U16 + | U32 + | U64 + | U128 +. + +Definition scalar_min (ty: scalar_ty) : Z := + match ty with + | Isize => isize_min + | I8 => i8_min + | I16 => i16_min + | I32 => i32_min + | I64 => i64_min + | I128 => i128_min + | Usize => usize_min + | U8 => u8_min + | U16 => u16_min + | U32 => u32_min + | U64 => u64_min + | U128 => u128_min +end. + +Definition scalar_max (ty: scalar_ty) : Z := + match ty with + | Isize => isize_max + | I8 => i8_max + | I16 => i16_max + | I32 => i32_max + | I64 => i64_max + | I128 => i128_max + | Usize => usize_max + | U8 => u8_max + | U16 => u16_max + | U32 => u32_max + | U64 => u64_max + | U128 => u128_max +end. + +(** We use the following conservative bounds to make sure we can compute bound + checks in most situations *) +Definition scalar_min_cons (ty: scalar_ty) : Z := + match ty with + | Isize => i32_min + | Usize => u32_min + | _ => scalar_min ty +end. + +Definition scalar_max_cons (ty: scalar_ty) : Z := + match ty with + | Isize => i32_max + | Usize => u32_max + | _ => scalar_max ty +end. + +Lemma scalar_min_cons_valid : forall ty, scalar_min ty <= scalar_min_cons ty . +Proof. + destruct ty; unfold scalar_min_cons, scalar_min; try lia. + - pose isize_min_bound; lia. + - apply Z.le_refl. +Qed. + +Lemma scalar_max_cons_valid : forall ty, scalar_max ty >= scalar_max_cons ty . +Proof. + destruct ty; unfold scalar_max_cons, scalar_max; try lia. + - pose isize_max_bound; lia. + - pose usize_max_bound. lia. +Qed. + +Definition scalar (ty: scalar_ty) : Type := + { x: Z | scalar_min ty <= x <= scalar_max ty }. + +Definition to_Z {ty} (x: scalar ty) : Z := proj1_sig x. + +(** Bounds checks: we start by using the conservative bounds, to make sure we + can compute in most situations, then we use the real bounds (for [isize] + and [usize]). *) +Definition scalar_ge_min (ty: scalar_ty) (x: Z) : bool := + Z.leb (scalar_min_cons ty) x || Z.leb (scalar_min ty) x. + +Definition scalar_le_max (ty: scalar_ty) (x: Z) : bool := + Z.leb x (scalar_max_cons ty) || Z.leb x (scalar_max ty). + +Lemma scalar_ge_min_valid (ty: scalar_ty) (x: Z) : + scalar_ge_min ty x = true -> scalar_min ty <= x . +Proof. + unfold scalar_ge_min. + pose (scalar_min_cons_valid ty). + lia. +Qed. + +Lemma scalar_le_max_valid (ty: scalar_ty) (x: Z) : + scalar_le_max ty x = true -> x <= scalar_max ty . +Proof. + unfold scalar_le_max. + pose (scalar_max_cons_valid ty). + lia. +Qed. + +Definition scalar_in_bounds (ty: scalar_ty) (x: Z) : bool := + scalar_ge_min ty x && scalar_le_max ty x . + +Lemma scalar_in_bounds_valid (ty: scalar_ty) (x: Z) : + scalar_in_bounds ty x = true -> scalar_min ty <= x <= scalar_max ty . +Proof. + unfold scalar_in_bounds. + intros H. + destruct (scalar_ge_min ty x) eqn:Hmin. + - destruct (scalar_le_max ty x) eqn:Hmax. + + pose (scalar_ge_min_valid ty x Hmin). + pose (scalar_le_max_valid ty x Hmax). + lia. + + inversion H. + - inversion H. +Qed. + +Import Sumbool. + +Definition mk_scalar (ty: scalar_ty) (x: Z) : result (scalar ty) := + match sumbool_of_bool (scalar_in_bounds ty x) with + | left H => Ok (exist _ x (scalar_in_bounds_valid _ _ H)) + | right _ => Fail_ Failure + end. + +Definition scalar_add {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x + to_Z y). + +Definition scalar_sub {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x - to_Z y). + +Definition scalar_mul {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x * to_Z y). + +Definition scalar_div {ty} (x y: scalar ty) : result (scalar ty) := + if to_Z y =? 0 then Fail_ Failure else + mk_scalar ty (to_Z x / to_Z y). + +Definition scalar_rem {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (Z.rem (to_Z x) (to_Z y)). + +Definition scalar_neg {ty} (x: scalar ty) : result (scalar ty) := mk_scalar ty (-(to_Z x)). + +Axiom scalar_xor : forall ty, scalar ty -> scalar ty -> scalar ty. (* TODO *) +Axiom scalar_or : forall ty, scalar ty -> scalar ty -> scalar ty. (* TODO *) +Axiom scalar_and : forall ty, scalar ty -> scalar ty -> scalar ty. (* TODO *) +Axiom scalar_shl : forall ty0 ty1, scalar ty0 -> scalar ty1 -> result (scalar ty0). (* TODO *) +Axiom scalar_shr : forall ty0 ty1, scalar ty0 -> scalar ty1 -> result (scalar ty0). (* TODO *) + +(** Cast an integer from a [src_ty] to a [tgt_ty] *) +(* TODO: check the semantics of casts in Rust *) +Definition scalar_cast (src_ty tgt_ty : scalar_ty) (x : scalar src_ty) : result (scalar tgt_ty) := + mk_scalar tgt_ty (to_Z x). + +(* This can't fail, but for now we make all casts faillible (easier for the translation) *) +Definition scalar_cast_bool (tgt_ty : scalar_ty) (x : bool) : result (scalar tgt_ty) := + mk_scalar tgt_ty (if x then 1 else 0). + +(** Comparisons *) +Definition scalar_leb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.leb (to_Z x) (to_Z y) . + +Definition scalar_ltb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.ltb (to_Z x) (to_Z y) . + +Definition scalar_geb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.geb (to_Z x) (to_Z y) . + +Definition scalar_gtb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.gtb (to_Z x) (to_Z y) . + +Definition scalar_eqb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.eqb (to_Z x) (to_Z y) . + +Definition scalar_neqb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + negb (Z.eqb (to_Z x) (to_Z y)) . + + +(** The scalar types *) +Definition isize := scalar Isize. +Definition i8 := scalar I8. +Definition i16 := scalar I16. +Definition i32 := scalar I32. +Definition i64 := scalar I64. +Definition i128 := scalar I128. +Definition usize := scalar Usize. +Definition u8 := scalar U8. +Definition u16 := scalar U16. +Definition u32 := scalar U32. +Definition u64 := scalar U64. +Definition u128 := scalar U128. + +(** Negaion *) +Definition isize_neg := @scalar_neg Isize. +Definition i8_neg := @scalar_neg I8. +Definition i16_neg := @scalar_neg I16. +Definition i32_neg := @scalar_neg I32. +Definition i64_neg := @scalar_neg I64. +Definition i128_neg := @scalar_neg I128. + +(** Division *) +Definition isize_div := @scalar_div Isize. +Definition i8_div := @scalar_div I8. +Definition i16_div := @scalar_div I16. +Definition i32_div := @scalar_div I32. +Definition i64_div := @scalar_div I64. +Definition i128_div := @scalar_div I128. +Definition usize_div := @scalar_div Usize. +Definition u8_div := @scalar_div U8. +Definition u16_div := @scalar_div U16. +Definition u32_div := @scalar_div U32. +Definition u64_div := @scalar_div U64. +Definition u128_div := @scalar_div U128. + +(** Remainder *) +Definition isize_rem := @scalar_rem Isize. +Definition i8_rem := @scalar_rem I8. +Definition i16_rem := @scalar_rem I16. +Definition i32_rem := @scalar_rem I32. +Definition i64_rem := @scalar_rem I64. +Definition i128_rem := @scalar_rem I128. +Definition usize_rem := @scalar_rem Usize. +Definition u8_rem := @scalar_rem U8. +Definition u16_rem := @scalar_rem U16. +Definition u32_rem := @scalar_rem U32. +Definition u64_rem := @scalar_rem U64. +Definition u128_rem := @scalar_rem U128. + +(** Addition *) +Definition isize_add := @scalar_add Isize. +Definition i8_add := @scalar_add I8. +Definition i16_add := @scalar_add I16. +Definition i32_add := @scalar_add I32. +Definition i64_add := @scalar_add I64. +Definition i128_add := @scalar_add I128. +Definition usize_add := @scalar_add Usize. +Definition u8_add := @scalar_add U8. +Definition u16_add := @scalar_add U16. +Definition u32_add := @scalar_add U32. +Definition u64_add := @scalar_add U64. +Definition u128_add := @scalar_add U128. + +(** Substraction *) +Definition isize_sub := @scalar_sub Isize. +Definition i8_sub := @scalar_sub I8. +Definition i16_sub := @scalar_sub I16. +Definition i32_sub := @scalar_sub I32. +Definition i64_sub := @scalar_sub I64. +Definition i128_sub := @scalar_sub I128. +Definition usize_sub := @scalar_sub Usize. +Definition u8_sub := @scalar_sub U8. +Definition u16_sub := @scalar_sub U16. +Definition u32_sub := @scalar_sub U32. +Definition u64_sub := @scalar_sub U64. +Definition u128_sub := @scalar_sub U128. + +(** Multiplication *) +Definition isize_mul := @scalar_mul Isize. +Definition i8_mul := @scalar_mul I8. +Definition i16_mul := @scalar_mul I16. +Definition i32_mul := @scalar_mul I32. +Definition i64_mul := @scalar_mul I64. +Definition i128_mul := @scalar_mul I128. +Definition usize_mul := @scalar_mul Usize. +Definition u8_mul := @scalar_mul U8. +Definition u16_mul := @scalar_mul U16. +Definition u32_mul := @scalar_mul U32. +Definition u64_mul := @scalar_mul U64. +Definition u128_mul := @scalar_mul U128. + +(** Xor *) +Definition u8_xor := @scalar_xor U8. +Definition u16_xor := @scalar_xor U16. +Definition u32_xor := @scalar_xor U32. +Definition u64_xor := @scalar_xor U64. +Definition u128_xor := @scalar_xor U128. +Definition usize_xor := @scalar_xor Usize. +Definition i8_xor := @scalar_xor I8. +Definition i16_xor := @scalar_xor I16. +Definition i32_xor := @scalar_xor I32. +Definition i64_xor := @scalar_xor I64. +Definition i128_xor := @scalar_xor I128. +Definition isize_xor := @scalar_xor Isize. + +(** Or *) +Definition u8_or := @scalar_or U8. +Definition u16_or := @scalar_or U16. +Definition u32_or := @scalar_or U32. +Definition u64_or := @scalar_or U64. +Definition u128_or := @scalar_or U128. +Definition usize_or := @scalar_or Usize. +Definition i8_or := @scalar_or I8. +Definition i16_or := @scalar_or I16. +Definition i32_or := @scalar_or I32. +Definition i64_or := @scalar_or I64. +Definition i128_or := @scalar_or I128. +Definition isize_or := @scalar_or Isize. + +(** And *) +Definition u8_and := @scalar_and U8. +Definition u16_and := @scalar_and U16. +Definition u32_and := @scalar_and U32. +Definition u64_and := @scalar_and U64. +Definition u128_and := @scalar_and U128. +Definition usize_and := @scalar_and Usize. +Definition i8_and := @scalar_and I8. +Definition i16_and := @scalar_and I16. +Definition i32_and := @scalar_and I32. +Definition i64_and := @scalar_and I64. +Definition i128_and := @scalar_and I128. +Definition isize_and := @scalar_and Isize. + +(** Shift left *) +Definition u8_shl {ty} := @scalar_shl U8 ty. +Definition u16_shl {ty} := @scalar_shl U16 ty. +Definition u32_shl {ty} := @scalar_shl U32 ty. +Definition u64_shl {ty} := @scalar_shl U64 ty. +Definition u128_shl {ty} := @scalar_shl U128 ty. +Definition usize_shl {ty} := @scalar_shl Usize ty. +Definition i8_shl {ty} := @scalar_shl I8 ty. +Definition i16_shl {ty} := @scalar_shl I16 ty. +Definition i32_shl {ty} := @scalar_shl I32 ty. +Definition i64_shl {ty} := @scalar_shl I64 ty. +Definition i128_shl {ty} := @scalar_shl I128 ty. +Definition isize_shl {ty} := @scalar_shl Isize ty. + +(** Shift right *) +Definition u8_shr {ty} := @scalar_shr U8 ty. +Definition u16_shr {ty} := @scalar_shr U16 ty. +Definition u32_shr {ty} := @scalar_shr U32 ty. +Definition u64_shr {ty} := @scalar_shr U64 ty. +Definition u128_shr {ty} := @scalar_shr U128 ty. +Definition usize_shr {ty} := @scalar_shr Usize ty. +Definition i8_shr {ty} := @scalar_shr I8 ty. +Definition i16_shr {ty} := @scalar_shr I16 ty. +Definition i32_shr {ty} := @scalar_shr I32 ty. +Definition i64_shr {ty} := @scalar_shr I64 ty. +Definition i128_shr {ty} := @scalar_shr I128 ty. +Definition isize_shr {ty} := @scalar_shr Isize ty. + +(** Small utility *) +Definition usize_to_nat (x: usize) : nat := Z.to_nat (to_Z x). + +(** Notations *) +Notation "x %isize" := ((mk_scalar Isize x)%return) (at level 9). +Notation "x %i8" := ((mk_scalar I8 x)%return) (at level 9). +Notation "x %i16" := ((mk_scalar I16 x)%return) (at level 9). +Notation "x %i32" := ((mk_scalar I32 x)%return) (at level 9). +Notation "x %i64" := ((mk_scalar I64 x)%return) (at level 9). +Notation "x %i128" := ((mk_scalar I128 x)%return) (at level 9). +Notation "x %usize" := ((mk_scalar Usize x)%return) (at level 9). +Notation "x %u8" := ((mk_scalar U8 x)%return) (at level 9). +Notation "x %u16" := ((mk_scalar U16 x)%return) (at level 9). +Notation "x %u32" := ((mk_scalar U32 x)%return) (at level 9). +Notation "x %u64" := ((mk_scalar U64 x)%return) (at level 9). +Notation "x %u128" := ((mk_scalar U128 x)%return) (at level 9). + +Notation "x s= y" := (scalar_eqb x y) (at level 80) : Primitives_scope. +Notation "x s<> y" := (scalar_neqb x y) (at level 80) : Primitives_scope. +Notation "x s<= y" := (scalar_leb x y) (at level 80) : Primitives_scope. +Notation "x s< y" := (scalar_ltb x y) (at level 80) : Primitives_scope. +Notation "x s>= y" := (scalar_geb x y) (at level 80) : Primitives_scope. +Notation "x s> y" := (scalar_gtb x y) (at level 80) : Primitives_scope. + +(** Constants *) +Definition core_u8_max := u8_max %u32. +Definition core_u16_max := u16_max %u32. +Definition core_u32_max := u32_max %u32. +Definition core_u64_max := u64_max %u64. +Definition core_u128_max := u64_max %u128. +Axiom core_usize_max : usize. (** TODO *) +Definition core_i8_max := i8_max %i32. +Definition core_i16_max := i16_max %i32. +Definition core_i32_max := i32_max %i32. +Definition core_i64_max := i64_max %i64. +Definition core_i128_max := i64_max %i128. +Axiom core_isize_max : isize. (** TODO *) + +(*** core *) + +(** Trait declaration: [core::clone::Clone] *) +Record core_clone_Clone (self : Type) := { + clone : self -> result self +}. + +Definition core_clone_impls_CloneBool_clone (b : bool) : bool := b. + +Definition core_clone_CloneBool : core_clone_Clone bool := {| + clone := fun b => Ok (core_clone_impls_CloneBool_clone b) +|}. + +Definition core_clone_impls_CloneUsize_clone (x : usize) : usize := x. +Definition core_clone_impls_CloneU8_clone (x : u8) : u8 := x. +Definition core_clone_impls_CloneU16_clone (x : u16) : u16 := x. +Definition core_clone_impls_CloneU32_clone (x : u32) : u32 := x. +Definition core_clone_impls_CloneU64_clone (x : u64) : u64 := x. +Definition core_clone_impls_CloneU128_clone (x : u128) : u128 := x. + +Definition core_clone_impls_CloneIsize_clone (x : isize) : isize := x. +Definition core_clone_impls_CloneI8_clone (x : i8) : i8 := x. +Definition core_clone_impls_CloneI16_clone (x : i16) : i16 := x. +Definition core_clone_impls_CloneI32_clone (x : i32) : i32 := x. +Definition core_clone_impls_CloneI64_clone (x : i64) : i64 := x. +Definition core_clone_impls_CloneI128_clone (x : i128) : i128 := x. + +Definition core_clone_CloneUsize : core_clone_Clone usize := {| + clone := fun x => Ok (core_clone_impls_CloneUsize_clone x) +|}. + +Definition core_clone_CloneU8 : core_clone_Clone u8 := {| + clone := fun x => Ok (core_clone_impls_CloneU8_clone x) +|}. + +Definition core_clone_CloneU16 : core_clone_Clone u16 := {| + clone := fun x => Ok (core_clone_impls_CloneU16_clone x) +|}. + +Definition core_clone_CloneU32 : core_clone_Clone u32 := {| + clone := fun x => Ok (core_clone_impls_CloneU32_clone x) +|}. + +Definition core_clone_CloneU64 : core_clone_Clone u64 := {| + clone := fun x => Ok (core_clone_impls_CloneU64_clone x) +|}. + +Definition core_clone_CloneU128 : core_clone_Clone u128 := {| + clone := fun x => Ok (core_clone_impls_CloneU128_clone x) +|}. + +Definition core_clone_CloneIsize : core_clone_Clone isize := {| + clone := fun x => Ok (core_clone_impls_CloneIsize_clone x) +|}. + +Definition core_clone_CloneI8 : core_clone_Clone i8 := {| + clone := fun x => Ok (core_clone_impls_CloneI8_clone x) +|}. + +Definition core_clone_CloneI16 : core_clone_Clone i16 := {| + clone := fun x => Ok (core_clone_impls_CloneI16_clone x) +|}. + +Definition core_clone_CloneI32 : core_clone_Clone i32 := {| + clone := fun x => Ok (core_clone_impls_CloneI32_clone x) +|}. + +Definition core_clone_CloneI64 : core_clone_Clone i64 := {| + clone := fun x => Ok (core_clone_impls_CloneI64_clone x) +|}. + +Definition core_clone_CloneI128 : core_clone_Clone i128 := {| + clone := fun x => Ok (core_clone_impls_CloneI128_clone x) +|}. + +(** [core::option::{core::option::Option}::unwrap] *) +Definition core_option_Option_unwrap (T : Type) (x : option T) : result T := + match x with + | None => Fail_ Failure + | Some x => Ok x + end. + +(*** core::ops *) + +(* Trait declaration: [core::ops::index::Index] *) +Record core_ops_index_Index (Self Idx : Type) := mk_core_ops_index_Index { + core_ops_index_Index_Output : Type; + core_ops_index_Index_index : Self -> Idx -> result core_ops_index_Index_Output; +}. +Arguments mk_core_ops_index_Index {_ _}. +Arguments core_ops_index_Index_Output {_ _}. +Arguments core_ops_index_Index_index {_ _}. + +(* Trait declaration: [core::ops::index::IndexMut] *) +Record core_ops_index_IndexMut (Self Idx : Type) := mk_core_ops_index_IndexMut { + core_ops_index_IndexMut_indexInst : core_ops_index_Index Self Idx; + core_ops_index_IndexMut_index_mut : + Self -> + Idx -> + result (core_ops_index_IndexMut_indexInst.(core_ops_index_Index_Output) * + (core_ops_index_IndexMut_indexInst.(core_ops_index_Index_Output) -> result Self)); +}. +Arguments mk_core_ops_index_IndexMut {_ _}. +Arguments core_ops_index_IndexMut_indexInst {_ _}. +Arguments core_ops_index_IndexMut_index_mut {_ _}. + +(* Trait declaration [core::ops::deref::Deref] *) +Record core_ops_deref_Deref (Self : Type) := mk_core_ops_deref_Deref { + core_ops_deref_Deref_target : Type; + core_ops_deref_Deref_deref : Self -> result core_ops_deref_Deref_target; +}. +Arguments mk_core_ops_deref_Deref {_}. +Arguments core_ops_deref_Deref_target {_}. +Arguments core_ops_deref_Deref_deref {_}. + +(* Trait declaration [core::ops::deref::DerefMut] *) +Record core_ops_deref_DerefMut (Self : Type) := mk_core_ops_deref_DerefMut { + core_ops_deref_DerefMut_derefInst : core_ops_deref_Deref Self; + core_ops_deref_DerefMut_deref_mut : + Self -> + result (core_ops_deref_DerefMut_derefInst.(core_ops_deref_Deref_target) * + (core_ops_deref_DerefMut_derefInst.(core_ops_deref_Deref_target) -> result Self)); +}. +Arguments mk_core_ops_deref_DerefMut {_}. +Arguments core_ops_deref_DerefMut_derefInst {_}. +Arguments core_ops_deref_DerefMut_deref_mut {_}. + +Record core_ops_range_Range (T : Type) := mk_core_ops_range_Range { + core_ops_range_Range_start : T; + core_ops_range_Range_end_ : T; +}. +Arguments mk_core_ops_range_Range {_}. +Arguments core_ops_range_Range_start {_}. +Arguments core_ops_range_Range_end_ {_}. + +(*** [alloc] *) + +Definition alloc_boxed_Box_deref (T : Type) (x : T) : result T := Ok x. +Definition alloc_boxed_Box_deref_mut (T : Type) (x : T) : result (T * (T -> result T)) := + Ok (x, fun x => Ok x). + +(* Trait instance *) +Definition alloc_boxed_Box_coreopsDerefInst (Self : Type) : core_ops_deref_Deref Self := {| + core_ops_deref_Deref_target := Self; + core_ops_deref_Deref_deref := alloc_boxed_Box_deref Self; +|}. + +(* Trait instance *) +Definition alloc_boxed_Box_coreopsDerefMutInst (Self : Type) : core_ops_deref_DerefMut Self := {| + core_ops_deref_DerefMut_derefInst := alloc_boxed_Box_coreopsDerefInst Self; + core_ops_deref_DerefMut_deref_mut := alloc_boxed_Box_deref_mut Self; +|}. + + +(*** Arrays *) +Definition array T (n : usize) := { l: list T | Z.of_nat (length l) = to_Z n}. + +Lemma le_0_usize_max : 0 <= usize_max. +Proof. + pose (H := usize_max_bound). + unfold u32_max in H. + lia. +Qed. + +Lemma eqb_imp_eq (x y : Z) : Z.eqb x y = true -> x = y. +Proof. + lia. +Qed. + +(* TODO: finish the definitions *) +Axiom mk_array : forall (T : Type) (n : usize) (l : list T), array T n. + +(* For initialization *) +Axiom array_repeat : forall (T : Type) (n : usize) (x : T), array T n. + +Axiom array_index_usize : forall (T : Type) (n : usize) (x : array T n) (i : usize), result T. +Axiom array_update_usize : forall (T : Type) (n : usize) (x : array T n) (i : usize) (nx : T), result (array T n). + +Definition array_index_mut_usize (T : Type) (n : usize) (a : array T n) (i : usize) : + result (T * (T -> result (array T n))) := + match array_index_usize T n a i with + | Fail_ e => Fail_ e + | Ok x => Ok (x, array_update_usize T n a i) + end. + +(*** Slice *) +Definition slice T := { l: list T | Z.of_nat (length l) <= usize_max}. + +Axiom slice_len : forall (T : Type) (s : slice T), usize. +Axiom slice_index_usize : forall (T : Type) (x : slice T) (i : usize), result T. +Axiom slice_update_usize : forall (T : Type) (x : slice T) (i : usize) (nx : T), result (slice T). + +Definition slice_index_mut_usize (T : Type) (s : slice T) (i : usize) : + result (T * (T -> result (slice T))) := + match slice_index_usize T s i with + | Fail_ e => Fail_ e + | Ok x => Ok (x, slice_update_usize T s i) + end. + +(*** Subslices *) + +Axiom array_to_slice : forall (T : Type) (n : usize) (x : array T n), result (slice T). +Axiom array_from_slice : forall (T : Type) (n : usize) (x : array T n) (s : slice T), result (array T n). + +Definition array_to_slice_mut (T : Type) (n : usize) (a : array T n) : + result (slice T * (slice T -> result (array T n))) := + match array_to_slice T n a with + | Fail_ e => Fail_ e + | Ok x => Ok (x, array_from_slice T n a) + end. + +Axiom array_subslice: forall (T : Type) (n : usize) (x : array T n) (r : core_ops_range_Range usize), result (slice T). +Axiom array_update_subslice: forall (T : Type) (n : usize) (x : array T n) (r : core_ops_range_Range usize) (ns : slice T), result (array T n). + +Axiom slice_subslice: forall (T : Type) (x : slice T) (r : core_ops_range_Range usize), result (slice T). +Axiom slice_update_subslice: forall (T : Type) (x : slice T) (r : core_ops_range_Range usize) (ns : slice T), result (slice T). + +(*** Vectors *) + +Definition alloc_vec_Vec T := { l: list T | Z.of_nat (length l) <= usize_max }. + +Definition alloc_vec_Vec_to_list {T: Type} (v: alloc_vec_Vec T) : list T := proj1_sig v. + +Definition alloc_vec_Vec_length {T: Type} (v: alloc_vec_Vec T) : Z := Z.of_nat (length (alloc_vec_Vec_to_list v)). + +Definition alloc_vec_Vec_new (T: Type) : alloc_vec_Vec T := (exist _ [] le_0_usize_max). + +Lemma alloc_vec_Vec_len_in_usize {T} (v: alloc_vec_Vec T) : usize_min <= alloc_vec_Vec_length v <= usize_max. +Proof. + unfold alloc_vec_Vec_length, usize_min. + split. + - lia. + - apply (proj2_sig v). +Qed. + +Definition alloc_vec_Vec_len (T: Type) (v: alloc_vec_Vec T) : usize := + exist _ (alloc_vec_Vec_length v) (alloc_vec_Vec_len_in_usize v). + +Fixpoint list_update {A} (l: list A) (n: nat) (a: A) + : list A := + match l with + | [] => [] + | x :: t => match n with + | 0%nat => a :: t + | S m => x :: (list_update t m a) +end end. + +Definition alloc_vec_Vec_bind {A B} (v: alloc_vec_Vec A) (f: list A -> result (list B)) : result (alloc_vec_Vec B) := + l <- f (alloc_vec_Vec_to_list v) ; + match sumbool_of_bool (scalar_le_max Usize (Z.of_nat (length l))) with + | left H => Ok (exist _ l (scalar_le_max_valid _ _ H)) + | right _ => Fail_ Failure + end. + +Definition alloc_vec_Vec_push (T: Type) (v: alloc_vec_Vec T) (x: T) : result (alloc_vec_Vec T) := + alloc_vec_Vec_bind v (fun l => Ok (l ++ [x])). + +Definition alloc_vec_Vec_insert (T: Type) (v: alloc_vec_Vec T) (i: usize) (x: T) : result (alloc_vec_Vec T) := + alloc_vec_Vec_bind v (fun l => + if to_Z i result (alloc_vec_Vec T))) := + match alloc_vec_Vec_index_usize v i with + | Ok x => + Ok (x, alloc_vec_Vec_update_usize v i) + | Fail_ e => Fail_ e + end. + +(* Trait declaration: [core::slice::index::private_slice_index::Sealed] *) +Definition core_slice_index_private_slice_index_Sealed (self : Type) := unit. + +(* Trait declaration: [core::slice::index::SliceIndex] *) +Record core_slice_index_SliceIndex (Self T : Type) := mk_core_slice_index_SliceIndex { + core_slice_index_SliceIndex_sealedInst : core_slice_index_private_slice_index_Sealed Self; + core_slice_index_SliceIndex_Output : Type; + core_slice_index_SliceIndex_get : Self -> T -> result (option core_slice_index_SliceIndex_Output); + core_slice_index_SliceIndex_get_mut : + Self -> T -> result (option core_slice_index_SliceIndex_Output * (option core_slice_index_SliceIndex_Output -> result T)); + core_slice_index_SliceIndex_get_unchecked : Self -> const_raw_ptr T -> result (const_raw_ptr core_slice_index_SliceIndex_Output); + core_slice_index_SliceIndex_get_unchecked_mut : Self -> mut_raw_ptr T -> result (mut_raw_ptr core_slice_index_SliceIndex_Output); + core_slice_index_SliceIndex_index : Self -> T -> result core_slice_index_SliceIndex_Output; + core_slice_index_SliceIndex_index_mut : + Self -> T -> result (core_slice_index_SliceIndex_Output * (core_slice_index_SliceIndex_Output -> result T)); +}. +Arguments mk_core_slice_index_SliceIndex {_ _}. +Arguments core_slice_index_SliceIndex_sealedInst {_ _}. +Arguments core_slice_index_SliceIndex_Output {_ _}. +Arguments core_slice_index_SliceIndex_get {_ _}. +Arguments core_slice_index_SliceIndex_get_mut {_ _}. +Arguments core_slice_index_SliceIndex_get_unchecked {_ _}. +Arguments core_slice_index_SliceIndex_get_unchecked_mut {_ _}. +Arguments core_slice_index_SliceIndex_index {_ _}. +Arguments core_slice_index_SliceIndex_index_mut {_ _}. + +(* [core::slice::index::[T]::index]: forward function *) +Definition core_slice_index_Slice_index + (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)) + (s : slice T) (i : Idx) : result inst.(core_slice_index_SliceIndex_Output) := + x <- inst.(core_slice_index_SliceIndex_get) i s; + match x with + | None => Fail_ Failure + | Some x => Ok x + end. + +(* [core::slice::index::Range:::get]: forward function *) +Axiom core_slice_index_RangeUsize_get : forall (T : Type) (i : core_ops_range_Range usize) (s : slice T), result (option (slice T)). + +(* [core::slice::index::Range::get_mut]: forward function *) +Axiom core_slice_index_RangeUsize_get_mut : + forall (T : Type), + core_ops_range_Range usize -> slice T -> + result (option (slice T) * (option (slice T) -> result (slice T))). + +(* [core::slice::index::Range::get_unchecked]: forward function *) +Definition core_slice_index_RangeUsize_get_unchecked + (T : Type) : + core_ops_range_Range usize -> const_raw_ptr (slice T) -> result (const_raw_ptr (slice T)) := + (* Don't know what the model should be - for now we always fail to make + sure code which uses it fails *) + fun _ _ => Fail_ Failure. + +(* [core::slice::index::Range::get_unchecked_mut]: forward function *) +Definition core_slice_index_RangeUsize_get_unchecked_mut + (T : Type) : + core_ops_range_Range usize -> mut_raw_ptr (slice T) -> result (mut_raw_ptr (slice T)) := + (* Don't know what the model should be - for now we always fail to make + sure code which uses it fails *) + fun _ _ => Fail_ Failure. + +(* [core::slice::index::Range::index]: forward function *) +Axiom core_slice_index_RangeUsize_index : + forall (T : Type), core_ops_range_Range usize -> slice T -> result (slice T). + +(* [core::slice::index::Range::index_mut]: forward function *) +Axiom core_slice_index_RangeUsize_index_mut : + forall (T : Type), core_ops_range_Range usize -> slice T -> result (slice T * (slice T -> result (slice T))). + +(* [core::slice::index::[T]::index_mut]: forward function *) +Axiom core_slice_index_Slice_index_mut : + forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)), + slice T -> Idx -> + result (inst.(core_slice_index_SliceIndex_Output) * + (inst.(core_slice_index_SliceIndex_Output) -> result (slice T))). + +(* [core::array::[T; N]::index]: forward function *) +Axiom core_array_Array_index : + forall (T Idx : Type) (N : usize) (inst : core_ops_index_Index (slice T) Idx) + (a : array T N) (i : Idx), result inst.(core_ops_index_Index_Output). + +(* [core::array::[T; N]::index_mut]: forward function *) +Axiom core_array_Array_index_mut : + forall (T Idx : Type) (N : usize) (inst : core_ops_index_IndexMut (slice T) Idx) + (a : array T N) (i : Idx), + result (inst.(core_ops_index_IndexMut_indexInst).(core_ops_index_Index_Output) * + (inst.(core_ops_index_IndexMut_indexInst).(core_ops_index_Index_Output) -> result (array T N))). + +(* Trait implementation: [core::slice::index::private_slice_index::Range] *) +Definition core_slice_index_private_slice_index_SealedRangeUsizeInst + : core_slice_index_private_slice_index_Sealed (core_ops_range_Range usize) := tt. + +(* Trait implementation: [core::slice::index::Range] *) +Definition core_slice_index_SliceIndexRangeUsizeSliceTInst (T : Type) : + core_slice_index_SliceIndex (core_ops_range_Range usize) (slice T) := {| + core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_SealedRangeUsizeInst; + core_slice_index_SliceIndex_Output := slice T; + core_slice_index_SliceIndex_get := core_slice_index_RangeUsize_get T; + core_slice_index_SliceIndex_get_mut := core_slice_index_RangeUsize_get_mut T; + core_slice_index_SliceIndex_get_unchecked := core_slice_index_RangeUsize_get_unchecked T; + core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_RangeUsize_get_unchecked_mut T; + core_slice_index_SliceIndex_index := core_slice_index_RangeUsize_index T; + core_slice_index_SliceIndex_index_mut := core_slice_index_RangeUsize_index_mut T; +|}. + +(* Trait implementation: [core::slice::index::[T]] *) +Definition core_ops_index_IndexSliceTIInst (T Idx : Type) + (inst : core_slice_index_SliceIndex Idx (slice T)) : + core_ops_index_Index (slice T) Idx := {| + core_ops_index_Index_Output := inst.(core_slice_index_SliceIndex_Output); + core_ops_index_Index_index := core_slice_index_Slice_index T Idx inst; +|}. + +(* Trait implementation: [core::slice::index::[T]] *) +Definition core_ops_index_IndexMutSliceTIInst (T Idx : Type) + (inst : core_slice_index_SliceIndex Idx (slice T)) : + core_ops_index_IndexMut (slice T) Idx := {| + core_ops_index_IndexMut_indexInst := core_ops_index_IndexSliceTIInst T Idx inst; + core_ops_index_IndexMut_index_mut := core_slice_index_Slice_index_mut T Idx inst; +|}. + +(* Trait implementation: [core::array::[T; N]] *) +Definition core_ops_index_IndexArrayInst (T Idx : Type) (N : usize) + (inst : core_ops_index_Index (slice T) Idx) : + core_ops_index_Index (array T N) Idx := {| + core_ops_index_Index_Output := inst.(core_ops_index_Index_Output); + core_ops_index_Index_index := core_array_Array_index T Idx N inst; +|}. + +(* Trait implementation: [core::array::[T; N]] *) +Definition core_ops_index_IndexMutArrayInst (T Idx : Type) (N : usize) + (inst : core_ops_index_IndexMut (slice T) Idx) : + core_ops_index_IndexMut (array T N) Idx := {| + core_ops_index_IndexMut_indexInst := core_ops_index_IndexArrayInst T Idx N inst.(core_ops_index_IndexMut_indexInst); + core_ops_index_IndexMut_index_mut := core_array_Array_index_mut T Idx N inst; +|}. + +(* [core::slice::index::usize::get]: forward function *) +Axiom core_slice_index_usize_get : forall (T : Type), usize -> slice T -> result (option T). + +(* [core::slice::index::usize::get_mut]: forward function *) +Axiom core_slice_index_usize_get_mut : + forall (T : Type), usize -> slice T -> result (option T * (option T -> result (slice T))). + +(* [core::slice::index::usize::get_unchecked]: forward function *) +Axiom core_slice_index_usize_get_unchecked : + forall (T : Type), usize -> const_raw_ptr (slice T) -> result (const_raw_ptr T). + +(* [core::slice::index::usize::get_unchecked_mut]: forward function *) +Axiom core_slice_index_usize_get_unchecked_mut : + forall (T : Type), usize -> mut_raw_ptr (slice T) -> result (mut_raw_ptr T). + +(* [core::slice::index::usize::index]: forward function *) +Axiom core_slice_index_usize_index : forall (T : Type), usize -> slice T -> result T. + +(* [core::slice::index::usize::index_mut]: forward function *) +Axiom core_slice_index_usize_index_mut : + forall (T : Type), usize -> slice T -> result (T * (T -> result (slice T))). + +(* Trait implementation: [core::slice::index::private_slice_index::usize] *) +Definition core_slice_index_private_slice_index_SealedUsizeInst + : core_slice_index_private_slice_index_Sealed usize := tt. + +(* Trait implementation: [core::slice::index::usize] *) +Definition core_slice_index_SliceIndexUsizeSliceTInst (T : Type) : + core_slice_index_SliceIndex usize (slice T) := {| + core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_SealedUsizeInst; + core_slice_index_SliceIndex_Output := T; + core_slice_index_SliceIndex_get := core_slice_index_usize_get T; + core_slice_index_SliceIndex_get_mut := core_slice_index_usize_get_mut T; + core_slice_index_SliceIndex_get_unchecked := core_slice_index_usize_get_unchecked T; + core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_usize_get_unchecked_mut T; + core_slice_index_SliceIndex_index := core_slice_index_usize_index T; + core_slice_index_SliceIndex_index_mut := core_slice_index_usize_index_mut T; +|}. + +(* [alloc::vec::Vec::index]: forward function *) +Axiom alloc_vec_Vec_index : forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)) + (Self : alloc_vec_Vec T) (i : Idx), result inst.(core_slice_index_SliceIndex_Output). + +(* [alloc::vec::Vec::index_mut]: forward function *) +Axiom alloc_vec_Vec_index_mut : forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)) + (Self : alloc_vec_Vec T) (i : Idx), + result (inst.(core_slice_index_SliceIndex_Output) * + (inst.(core_slice_index_SliceIndex_Output) -> result (alloc_vec_Vec T))). + +(* Trait implementation: [alloc::vec::Vec] *) +Definition alloc_vec_Vec_coreopsindexIndexInst (T Idx : Type) + (inst : core_slice_index_SliceIndex Idx (slice T)) : + core_ops_index_Index (alloc_vec_Vec T) Idx := {| + core_ops_index_Index_Output := inst.(core_slice_index_SliceIndex_Output); + core_ops_index_Index_index := alloc_vec_Vec_index T Idx inst; +|}. + +(* Trait implementation: [alloc::vec::Vec] *) +Definition alloc_vec_Vec_coreopsindexIndexMutInst (T Idx : Type) + (inst : core_slice_index_SliceIndex Idx (slice T)) : + core_ops_index_IndexMut (alloc_vec_Vec T) Idx := {| + core_ops_index_IndexMut_indexInst := alloc_vec_Vec_coreopsindexIndexInst T Idx inst; + core_ops_index_IndexMut_index_mut := alloc_vec_Vec_index_mut T Idx inst; +|}. + +(*** Theorems *) + +Axiom alloc_vec_Vec_index_eq : forall {a : Type} (v : alloc_vec_Vec a) (i : usize) (x : a), + alloc_vec_Vec_index a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i = + alloc_vec_Vec_index_usize v i. + +Axiom alloc_vec_Vec_index_mut_eq : forall {a : Type} (v : alloc_vec_Vec a) (i : usize) (x : a), + alloc_vec_Vec_index_mut a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i = + alloc_vec_Vec_index_mut_usize v i. + +End Primitives. diff --git a/tests/coq/hashmap_main/_CoqProject b/tests/coq/hashmap_main/_CoqProject new file mode 100644 index 00000000..d73541d9 --- /dev/null +++ b/tests/coq/hashmap_main/_CoqProject @@ -0,0 +1,12 @@ +# This file was automatically generated - see ../Makefile +-R . Lib +-arg -w +-arg all + +HashmapMain_Types.v +HashmapMain_FunsExternal_Template.v +Primitives.v +HashmapMain_Funs.v +HashmapMain_TypesExternal.v +HashmapMain_FunsExternal.v +HashmapMain_TypesExternal_Template.v diff --git a/tests/coq/hashmap_on_disk/HashmapMain_Funs.v b/tests/coq/hashmap_on_disk/HashmapMain_Funs.v deleted file mode 100644 index f6467d5a..00000000 --- a/tests/coq/hashmap_on_disk/HashmapMain_Funs.v +++ /dev/null @@ -1,589 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: function definitions *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Require Import HashmapMain_Types. -Include HashmapMain_Types. -Require Import HashmapMain_FunsExternal. -Include HashmapMain_FunsExternal. -Module HashmapMain_Funs. - -(** [hashmap_main::hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 35:0-35:32 *) -Definition hashmap_hash_key (k : usize) : result usize := - Ok k. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) -Fixpoint hashmap_HashMap_allocate_slots_loop - (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (n1 : usize) - : - result (alloc_vec_Vec (hashmap_List_t T)) - := - match n with - | O => Fail_ OutOfFuel - | S n2 => - if n1 s> 0%usize - then ( - slots1 <- alloc_vec_Vec_push (hashmap_List_t T) slots Hashmap_List_Nil; - n3 <- usize_sub n1 1%usize; - hashmap_HashMap_allocate_slots_loop T n2 slots1 n3) - else Ok slots - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 58:4-58:76 *) -Definition hashmap_HashMap_allocate_slots - (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (n1 : usize) - : - result (alloc_vec_Vec (hashmap_List_t T)) - := - hashmap_HashMap_allocate_slots_loop T n slots n1 -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 67:4-71:13 *) -Definition hashmap_HashMap_new_with_capacity - (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) - (max_load_divisor : usize) : - result (hashmap_HashMap_t T) - := - slots <- - hashmap_HashMap_allocate_slots T n (alloc_vec_Vec_new (hashmap_List_t T)) - capacity; - i <- usize_mul capacity max_load_dividend; - i1 <- usize_div i max_load_divisor; - Ok - {| - hashmap_HashMap_num_entries := 0%usize; - hashmap_HashMap_max_load_factor := (max_load_dividend, max_load_divisor); - hashmap_HashMap_max_load := i1; - hashmap_HashMap_slots := slots - |} -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 83:4-83:24 *) -Definition hashmap_HashMap_new - (T : Type) (n : nat) : result (hashmap_HashMap_t T) := - hashmap_HashMap_new_with_capacity T n 32%usize 4%usize 5%usize -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) -Fixpoint hashmap_HashMap_clear_loop - (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) : - result (alloc_vec_Vec (hashmap_List_t T)) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - let i1 := alloc_vec_Vec_len (hashmap_List_t T) slots in - if i s< i1 - then ( - p <- - alloc_vec_Vec_index_mut (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) slots - i; - let (_, index_mut_back) := p in - i2 <- usize_add i 1%usize; - slots1 <- index_mut_back Hashmap_List_Nil; - hashmap_HashMap_clear_loop T n1 slots1 i2) - else Ok slots - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 88:4-88:27 *) -Definition hashmap_HashMap_clear - (T : Type) (n : nat) (self : hashmap_HashMap_t T) : - result (hashmap_HashMap_t T) - := - hm <- hashmap_HashMap_clear_loop T n self.(hashmap_HashMap_slots) 0%usize; - Ok - {| - hashmap_HashMap_num_entries := 0%usize; - hashmap_HashMap_max_load_factor := self.(hashmap_HashMap_max_load_factor); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := hm - |} -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::len]: - Source: 'tests/src/hashmap.rs', lines 98:4-98:30 *) -Definition hashmap_HashMap_len - (T : Type) (self : hashmap_HashMap_t T) : result usize := - Ok self.(hashmap_HashMap_num_entries) -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) -Fixpoint hashmap_HashMap_insert_in_list_loop - (T : Type) (n : nat) (key : usize) (value : T) (ls : hashmap_List_t T) : - result (bool * (hashmap_List_t T)) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match ls with - | Hashmap_List_Cons ckey cvalue tl => - if ckey s= key - then Ok (false, Hashmap_List_Cons ckey value tl) - else ( - p <- hashmap_HashMap_insert_in_list_loop T n1 key value tl; - let (b, tl1) := p in - Ok (b, Hashmap_List_Cons ckey cvalue tl1)) - | Hashmap_List_Nil => - Ok (true, Hashmap_List_Cons key value Hashmap_List_Nil) - end - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 105:4-105:71 *) -Definition hashmap_HashMap_insert_in_list - (T : Type) (n : nat) (key : usize) (value : T) (ls : hashmap_List_t T) : - result (bool * (hashmap_List_t T)) - := - hashmap_HashMap_insert_in_list_loop T n key value ls -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 125:4-125:54 *) -Definition hashmap_HashMap_insert_no_resize - (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) (value : T) : - result (hashmap_HashMap_t T) - := - hash <- hashmap_hash_key key; - let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in - hash_mod <- usize_rem hash i; - p <- - alloc_vec_Vec_index_mut (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) - self.(hashmap_HashMap_slots) hash_mod; - let (l, index_mut_back) := p in - p1 <- hashmap_HashMap_insert_in_list T n key value l; - let (inserted, l1) := p1 in - if inserted - then ( - i1 <- usize_add self.(hashmap_HashMap_num_entries) 1%usize; - v <- index_mut_back l1; - Ok - {| - hashmap_HashMap_num_entries := i1; - hashmap_HashMap_max_load_factor := - self.(hashmap_HashMap_max_load_factor); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := v - |}) - else ( - v <- index_mut_back l1; - Ok - {| - hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); - hashmap_HashMap_max_load_factor := - self.(hashmap_HashMap_max_load_factor); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := v - |}) -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) -Fixpoint hashmap_HashMap_move_elements_from_list_loop - (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) (ls : hashmap_List_t T) : - result (hashmap_HashMap_t T) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match ls with - | Hashmap_List_Cons k v tl => - ntable1 <- hashmap_HashMap_insert_no_resize T n1 ntable k v; - hashmap_HashMap_move_elements_from_list_loop T n1 ntable1 tl - | Hashmap_List_Nil => Ok ntable - end - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 191:4-191:72 *) -Definition hashmap_HashMap_move_elements_from_list - (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) (ls : hashmap_List_t T) : - result (hashmap_HashMap_t T) - := - hashmap_HashMap_move_elements_from_list_loop T n ntable ls -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) -Fixpoint hashmap_HashMap_move_elements_loop - (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) - (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) : - result ((hashmap_HashMap_t T) * (alloc_vec_Vec (hashmap_List_t T))) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - let i1 := alloc_vec_Vec_len (hashmap_List_t T) slots in - if i s< i1 - then ( - p <- - alloc_vec_Vec_index_mut (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) slots - i; - let (l, index_mut_back) := p in - let (ls, l1) := core_mem_replace (hashmap_List_t T) l Hashmap_List_Nil in - ntable1 <- hashmap_HashMap_move_elements_from_list T n1 ntable ls; - i2 <- usize_add i 1%usize; - slots1 <- index_mut_back l1; - hashmap_HashMap_move_elements_loop T n1 ntable1 slots1 i2) - else Ok (ntable, slots) - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 179:4-179:95 *) -Definition hashmap_HashMap_move_elements - (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) - (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) : - result ((hashmap_HashMap_t T) * (alloc_vec_Vec (hashmap_List_t T))) - := - hashmap_HashMap_move_elements_loop T n ntable slots i -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 148:4-148:28 *) -Definition hashmap_HashMap_try_resize - (T : Type) (n : nat) (self : hashmap_HashMap_t T) : - result (hashmap_HashMap_t T) - := - max_usize <- scalar_cast U32 Usize core_u32_max; - let capacity := - alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in - n1 <- usize_div max_usize 2%usize; - let (i, i1) := self.(hashmap_HashMap_max_load_factor) in - i2 <- usize_div n1 i; - if capacity s<= i2 - then ( - i3 <- usize_mul capacity 2%usize; - ntable <- hashmap_HashMap_new_with_capacity T n i3 i i1; - p <- - hashmap_HashMap_move_elements T n ntable self.(hashmap_HashMap_slots) - 0%usize; - let (ntable1, _) := p in - Ok - {| - hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); - hashmap_HashMap_max_load_factor := (i, i1); - hashmap_HashMap_max_load := ntable1.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := ntable1.(hashmap_HashMap_slots) - |}) - else - Ok - {| - hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); - hashmap_HashMap_max_load_factor := (i, i1); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := self.(hashmap_HashMap_slots) - |} -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert]: - Source: 'tests/src/hashmap.rs', lines 137:4-137:48 *) -Definition hashmap_HashMap_insert - (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) (value : T) : - result (hashmap_HashMap_t T) - := - self1 <- hashmap_HashMap_insert_no_resize T n self key value; - i <- hashmap_HashMap_len T self1; - if i s> self1.(hashmap_HashMap_max_load) - then hashmap_HashMap_try_resize T n self1 - else Ok self1 -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) -Fixpoint hashmap_HashMap_contains_key_in_list_loop - (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result bool := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match ls with - | Hashmap_List_Cons ckey _ tl => - if ckey s= key - then Ok true - else hashmap_HashMap_contains_key_in_list_loop T n1 key tl - | Hashmap_List_Nil => Ok false - end - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 214:4-214:68 *) -Definition hashmap_HashMap_contains_key_in_list - (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result bool := - hashmap_HashMap_contains_key_in_list_loop T n key ls -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 207:4-207:49 *) -Definition hashmap_HashMap_contains_key - (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : - result bool - := - hash <- hashmap_hash_key key; - let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in - hash_mod <- usize_rem hash i; - l <- - alloc_vec_Vec_index (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) - self.(hashmap_HashMap_slots) hash_mod; - hashmap_HashMap_contains_key_in_list T n key l -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) -Fixpoint hashmap_HashMap_get_in_list_loop - (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result T := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match ls with - | Hashmap_List_Cons ckey cvalue tl => - if ckey s= key - then Ok cvalue - else hashmap_HashMap_get_in_list_loop T n1 key tl - | Hashmap_List_Nil => Fail_ Failure - end - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 232:4-232:70 *) -Definition hashmap_HashMap_get_in_list - (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result T := - hashmap_HashMap_get_in_list_loop T n key ls -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 247:4-247:55 *) -Definition hashmap_HashMap_get - (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : result T := - hash <- hashmap_hash_key key; - let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in - hash_mod <- usize_rem hash i; - l <- - alloc_vec_Vec_index (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) - self.(hashmap_HashMap_slots) hash_mod; - hashmap_HashMap_get_in_list T n key l -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) -Fixpoint hashmap_HashMap_get_mut_in_list_loop - (T : Type) (n : nat) (ls : hashmap_List_t T) (key : usize) : - result (T * (T -> result (hashmap_List_t T))) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match ls with - | Hashmap_List_Cons ckey cvalue tl => - if ckey s= key - then - let back := fun (ret : T) => Ok (Hashmap_List_Cons ckey ret tl) in - Ok (cvalue, back) - else ( - p <- hashmap_HashMap_get_mut_in_list_loop T n1 tl key; - let (t, back) := p in - let back1 := - fun (ret : T) => - tl1 <- back ret; Ok (Hashmap_List_Cons ckey cvalue tl1) in - Ok (t, back1)) - | Hashmap_List_Nil => Fail_ Failure - end - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 253:4-253:86 *) -Definition hashmap_HashMap_get_mut_in_list - (T : Type) (n : nat) (ls : hashmap_List_t T) (key : usize) : - result (T * (T -> result (hashmap_List_t T))) - := - hashmap_HashMap_get_mut_in_list_loop T n ls key -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:67 *) -Definition hashmap_HashMap_get_mut - (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : - result (T * (T -> result (hashmap_HashMap_t T))) - := - hash <- hashmap_hash_key key; - let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in - hash_mod <- usize_rem hash i; - p <- - alloc_vec_Vec_index_mut (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) - self.(hashmap_HashMap_slots) hash_mod; - let (l, index_mut_back) := p in - p1 <- hashmap_HashMap_get_mut_in_list T n l 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; - Ok - {| - hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); - hashmap_HashMap_max_load_factor := - self.(hashmap_HashMap_max_load_factor); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := v - |} in - Ok (t, back) -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) -Fixpoint hashmap_HashMap_remove_from_list_loop - (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : - result ((option T) * (hashmap_List_t T)) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match ls with - | Hashmap_List_Cons ckey t tl => - if ckey s= key - then - let (mv_ls, _) := - core_mem_replace (hashmap_List_t T) (Hashmap_List_Cons ckey t tl) - Hashmap_List_Nil in - match mv_ls with - | Hashmap_List_Cons _ cvalue tl1 => Ok (Some cvalue, tl1) - | Hashmap_List_Nil => Fail_ Failure - end - else ( - p <- hashmap_HashMap_remove_from_list_loop T n1 key tl; - let (o, tl1) := p in - Ok (o, Hashmap_List_Cons ckey t tl1)) - | Hashmap_List_Nil => Ok (None, Hashmap_List_Nil) - end - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 273:4-273:69 *) -Definition hashmap_HashMap_remove_from_list - (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : - result ((option T) * (hashmap_List_t T)) - := - hashmap_HashMap_remove_from_list_loop T n key ls -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 302:4-302:52 *) -Definition hashmap_HashMap_remove - (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : - result ((option T) * (hashmap_HashMap_t T)) - := - hash <- hashmap_hash_key key; - let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in - hash_mod <- usize_rem hash i; - p <- - alloc_vec_Vec_index_mut (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) - self.(hashmap_HashMap_slots) hash_mod; - let (l, index_mut_back) := p in - p1 <- hashmap_HashMap_remove_from_list T n key l; - let (x, l1) := p1 in - match x with - | None => - v <- index_mut_back l1; - Ok (None, - {| - hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); - hashmap_HashMap_max_load_factor := - self.(hashmap_HashMap_max_load_factor); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := v - |}) - | Some x1 => - i1 <- usize_sub self.(hashmap_HashMap_num_entries) 1%usize; - v <- index_mut_back l1; - Ok (Some x1, - {| - hashmap_HashMap_num_entries := i1; - hashmap_HashMap_max_load_factor := - self.(hashmap_HashMap_max_load_factor); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := v - |}) - end -. - -(** [hashmap_main::hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 323:0-323:10 *) -Definition hashmap_test1 (n : nat) : result unit := - hm <- hashmap_HashMap_new u64 n; - hm1 <- hashmap_HashMap_insert u64 n hm 0%usize 42%u64; - hm2 <- hashmap_HashMap_insert u64 n hm1 128%usize 18%u64; - hm3 <- hashmap_HashMap_insert u64 n hm2 1024%usize 138%u64; - hm4 <- hashmap_HashMap_insert u64 n hm3 1056%usize 256%u64; - i <- hashmap_HashMap_get u64 n hm4 128%usize; - if negb (i s= 18%u64) - then Fail_ Failure - else ( - p <- hashmap_HashMap_get_mut u64 n hm4 1024%usize; - let (_, get_mut_back) := p in - hm5 <- get_mut_back 56%u64; - i1 <- hashmap_HashMap_get u64 n hm5 1024%usize; - if negb (i1 s= 56%u64) - then Fail_ Failure - else ( - p1 <- hashmap_HashMap_remove u64 n hm5 1024%usize; - let (x, hm6) := p1 in - match x with - | None => Fail_ Failure - | Some x1 => - if negb (x1 s= 56%u64) - then Fail_ Failure - else ( - i2 <- hashmap_HashMap_get u64 n hm6 0%usize; - if negb (i2 s= 42%u64) - then Fail_ Failure - else ( - i3 <- hashmap_HashMap_get u64 n hm6 128%usize; - if negb (i3 s= 18%u64) - then Fail_ Failure - else ( - i4 <- hashmap_HashMap_get u64 n hm6 1056%usize; - if negb (i4 s= 256%u64) then Fail_ Failure else Ok tt))) - end)) -. - -(** [hashmap_main::insert_on_disk]: - Source: 'tests/src/hashmap_main.rs', lines 13:0-13:43 *) -Definition insert_on_disk - (n : nat) (key : usize) (value : u64) (st : state) : result (state * unit) := - p <- hashmap_utils_deserialize st; - let (st1, hm) := p in - hm1 <- hashmap_HashMap_insert u64 n hm key value; - hashmap_utils_serialize hm1 st1 -. - -(** [hashmap_main::main]: - Source: 'tests/src/hashmap_main.rs', lines 22:0-22:13 *) -Definition main : result unit := - Ok tt. - -End HashmapMain_Funs. diff --git a/tests/coq/hashmap_on_disk/HashmapMain_FunsExternal.v b/tests/coq/hashmap_on_disk/HashmapMain_FunsExternal.v deleted file mode 100644 index fb5f23cd..00000000 --- a/tests/coq/hashmap_on_disk/HashmapMain_FunsExternal.v +++ /dev/null @@ -1,24 +0,0 @@ -(** [hashmap_main]: external function declarations *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Require Export HashmapMain_Types. -Import HashmapMain_Types. -Module HashmapMain_FunsExternal. - -(** [hashmap_main::hashmap_utils::deserialize]: forward function - Source: 'src/hashmap_utils.rs', lines 10:0-10:43 *) -Axiom hashmap_utils_deserialize - : state -> result (state * (hashmap_HashMap_t u64)) -. - -(** [hashmap_main::hashmap_utils::serialize]: forward function - Source: 'src/hashmap_utils.rs', lines 5:0-5:42 *) -Axiom hashmap_utils_serialize - : hashmap_HashMap_t u64 -> state -> result (state * unit) -. - -End HashmapMain_FunsExternal. diff --git a/tests/coq/hashmap_on_disk/HashmapMain_FunsExternal_Template.v b/tests/coq/hashmap_on_disk/HashmapMain_FunsExternal_Template.v deleted file mode 100644 index 66835e8c..00000000 --- a/tests/coq/hashmap_on_disk/HashmapMain_FunsExternal_Template.v +++ /dev/null @@ -1,26 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: external functions. --- This is a template file: rename it to "FunsExternal.lean" and fill the holes. *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Require Import HashmapMain_Types. -Include HashmapMain_Types. -Module HashmapMain_FunsExternal_Template. - -(** [hashmap_main::hashmap_utils::deserialize]: - Source: 'tests/src/hashmap_utils.rs', lines 11:0-11:43 *) -Axiom hashmap_utils_deserialize - : state -> result (state * (hashmap_HashMap_t u64)) -. - -(** [hashmap_main::hashmap_utils::serialize]: - Source: 'tests/src/hashmap_utils.rs', lines 6:0-6:42 *) -Axiom hashmap_utils_serialize - : hashmap_HashMap_t u64 -> state -> result (state * unit) -. - -End HashmapMain_FunsExternal_Template. diff --git a/tests/coq/hashmap_on_disk/HashmapMain_Types.v b/tests/coq/hashmap_on_disk/HashmapMain_Types.v deleted file mode 100644 index 5656bd9c..00000000 --- a/tests/coq/hashmap_on_disk/HashmapMain_Types.v +++ /dev/null @@ -1,40 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: type definitions *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Require Import HashmapMain_TypesExternal. -Include HashmapMain_TypesExternal. -Module HashmapMain_Types. - -(** [hashmap_main::hashmap::List] - Source: 'tests/src/hashmap.rs', lines 27:0-27:16 *) -Inductive hashmap_List_t (T : Type) := -| Hashmap_List_Cons : usize -> T -> hashmap_List_t T -> hashmap_List_t T -| Hashmap_List_Nil : hashmap_List_t T -. - -Arguments Hashmap_List_Cons { _ }. -Arguments Hashmap_List_Nil { _ }. - -(** [hashmap_main::hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 43:0-43:21 *) -Record hashmap_HashMap_t (T : Type) := -mkhashmap_HashMap_t { - hashmap_HashMap_num_entries : usize; - hashmap_HashMap_max_load_factor : (usize * usize); - hashmap_HashMap_max_load : usize; - hashmap_HashMap_slots : alloc_vec_Vec (hashmap_List_t T); -} -. - -Arguments mkhashmap_HashMap_t { _ }. -Arguments hashmap_HashMap_num_entries { _ }. -Arguments hashmap_HashMap_max_load_factor { _ }. -Arguments hashmap_HashMap_max_load { _ }. -Arguments hashmap_HashMap_slots { _ }. - -End HashmapMain_Types. diff --git a/tests/coq/hashmap_on_disk/HashmapMain_TypesExternal.v b/tests/coq/hashmap_on_disk/HashmapMain_TypesExternal.v deleted file mode 100644 index 28651c14..00000000 --- a/tests/coq/hashmap_on_disk/HashmapMain_TypesExternal.v +++ /dev/null @@ -1,14 +0,0 @@ -(** [hashmap_main]: external types. --- This is a template file: rename it to "TypesExternal.lean" and fill the holes. *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Module HashmapMain_TypesExternal. - -(** The state type used in the state-error monad *) -Axiom state : Type. - -End HashmapMain_TypesExternal. diff --git a/tests/coq/hashmap_on_disk/HashmapMain_TypesExternal_Template.v b/tests/coq/hashmap_on_disk/HashmapMain_TypesExternal_Template.v deleted file mode 100644 index 391b2775..00000000 --- a/tests/coq/hashmap_on_disk/HashmapMain_TypesExternal_Template.v +++ /dev/null @@ -1,15 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: external types. --- This is a template file: rename it to "TypesExternal.lean" and fill the holes. *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Module HashmapMain_TypesExternal_Template. - -(** The state type used in the state-error monad *) -Axiom state : Type. - -End HashmapMain_TypesExternal_Template. diff --git a/tests/coq/hashmap_on_disk/Makefile b/tests/coq/hashmap_on_disk/Makefile deleted file mode 100644 index 1a5aee4a..00000000 --- a/tests/coq/hashmap_on_disk/Makefile +++ /dev/null @@ -1,23 +0,0 @@ -# This file was automatically generated - modify ../Makefile.template instead -# Makefile originally taken from coq-club - -%: Makefile.coq phony - +make -f Makefile.coq $@ - -all: Makefile.coq - +make -f Makefile.coq all - -clean: Makefile.coq - +make -f Makefile.coq clean - rm -f Makefile.coq - -Makefile.coq: _CoqProject Makefile - coq_makefile -f _CoqProject | sed 's/$$(COQCHK) $$(COQCHKFLAGS) $$(COQLIBS)/$$(COQCHK) $$(COQCHKFLAGS) $$(subst -Q,-R,$$(COQLIBS))/' > Makefile.coq - -_CoqProject: ; - -Makefile: ; - -phony: ; - -.PHONY: all clean phony diff --git a/tests/coq/hashmap_on_disk/Primitives.v b/tests/coq/hashmap_on_disk/Primitives.v deleted file mode 100644 index b29fce43..00000000 --- a/tests/coq/hashmap_on_disk/Primitives.v +++ /dev/null @@ -1,981 +0,0 @@ -Require Import Lia. -Require Coq.Strings.Ascii. -Require Coq.Strings.String. -Require Import Coq.Program.Equality. -Require Import Coq.ZArith.ZArith. -Require Import Coq.ZArith.Znat. -Require Import List. -Import ListNotations. - -Module Primitives. - - (* TODO: use more *) -Declare Scope Primitives_scope. - -(*** Result *) - -Inductive error := - | Failure - | OutOfFuel. - -Inductive result A := - | Ok : A -> result A - | Fail_ : error -> result A. - -Arguments Ok {_} a. -Arguments Fail_ {_}. - -Definition bind {A B} (m: result A) (f: A -> result B) : result B := - match m with - | Fail_ e => Fail_ e - | Ok x => f x - end. - -Definition return_ {A: Type} (x: A) : result A := Ok x. -Definition fail_ {A: Type} (e: error) : result A := Fail_ e. - -Notation "x <- c1 ; c2" := (bind c1 (fun x => c2)) - (at level 61, c1 at next level, right associativity). - -(** Monadic assert *) -Definition massert (b: bool) : result unit := - if b then Ok tt else Fail_ Failure. - -(** Normalize and unwrap a successful result (used for globals) *) -Definition eval_result_refl {A} {x} (a: result A) (p: a = Ok x) : A := - match a as r return (r = Ok x -> A) with - | Ok a' => fun _ => a' - | Fail_ e => fun p' => - False_rect _ (eq_ind (Fail_ e) - (fun e : result A => - match e with - | Ok _ => False - | Fail_ e => True - end) - I (Ok x) p') - end p. - -Notation "x %global" := (eval_result_refl x eq_refl) (at level 40). -Notation "x %return" := (eval_result_refl x eq_refl) (at level 40). - -(* Sanity check *) -Check (if true then Ok (1 + 2) else Fail_ Failure)%global = 3. - -(*** Misc *) - -Definition string := Coq.Strings.String.string. -Definition char := Coq.Strings.Ascii.ascii. -Definition char_of_byte := Coq.Strings.Ascii.ascii_of_byte. - -Definition core_mem_replace (a : Type) (x : a) (y : a) : a * a := (x, x) . - -Record mut_raw_ptr (T : Type) := { mut_raw_ptr_v : T }. -Record const_raw_ptr (T : Type) := { const_raw_ptr_v : T }. - -(*** Scalars *) - -Definition i8_min : Z := -128%Z. -Definition i8_max : Z := 127%Z. -Definition i16_min : Z := -32768%Z. -Definition i16_max : Z := 32767%Z. -Definition i32_min : Z := -2147483648%Z. -Definition i32_max : Z := 2147483647%Z. -Definition i64_min : Z := -9223372036854775808%Z. -Definition i64_max : Z := 9223372036854775807%Z. -Definition i128_min : Z := -170141183460469231731687303715884105728%Z. -Definition i128_max : Z := 170141183460469231731687303715884105727%Z. -Definition u8_min : Z := 0%Z. -Definition u8_max : Z := 255%Z. -Definition u16_min : Z := 0%Z. -Definition u16_max : Z := 65535%Z. -Definition u32_min : Z := 0%Z. -Definition u32_max : Z := 4294967295%Z. -Definition u64_min : Z := 0%Z. -Definition u64_max : Z := 18446744073709551615%Z. -Definition u128_min : Z := 0%Z. -Definition u128_max : Z := 340282366920938463463374607431768211455%Z. - -(** The bounds of [isize] and [usize] vary with the architecture. *) -Axiom isize_min : Z. -Axiom isize_max : Z. -Definition usize_min : Z := 0%Z. -Axiom usize_max : Z. - -Open Scope Z_scope. - -(** We provide those lemmas to reason about the bounds of [isize] and [usize] *) -Axiom isize_min_bound : isize_min <= i32_min. -Axiom isize_max_bound : i32_max <= isize_max. -Axiom usize_max_bound : u32_max <= usize_max. - -Inductive scalar_ty := - | Isize - | I8 - | I16 - | I32 - | I64 - | I128 - | Usize - | U8 - | U16 - | U32 - | U64 - | U128 -. - -Definition scalar_min (ty: scalar_ty) : Z := - match ty with - | Isize => isize_min - | I8 => i8_min - | I16 => i16_min - | I32 => i32_min - | I64 => i64_min - | I128 => i128_min - | Usize => usize_min - | U8 => u8_min - | U16 => u16_min - | U32 => u32_min - | U64 => u64_min - | U128 => u128_min -end. - -Definition scalar_max (ty: scalar_ty) : Z := - match ty with - | Isize => isize_max - | I8 => i8_max - | I16 => i16_max - | I32 => i32_max - | I64 => i64_max - | I128 => i128_max - | Usize => usize_max - | U8 => u8_max - | U16 => u16_max - | U32 => u32_max - | U64 => u64_max - | U128 => u128_max -end. - -(** We use the following conservative bounds to make sure we can compute bound - checks in most situations *) -Definition scalar_min_cons (ty: scalar_ty) : Z := - match ty with - | Isize => i32_min - | Usize => u32_min - | _ => scalar_min ty -end. - -Definition scalar_max_cons (ty: scalar_ty) : Z := - match ty with - | Isize => i32_max - | Usize => u32_max - | _ => scalar_max ty -end. - -Lemma scalar_min_cons_valid : forall ty, scalar_min ty <= scalar_min_cons ty . -Proof. - destruct ty; unfold scalar_min_cons, scalar_min; try lia. - - pose isize_min_bound; lia. - - apply Z.le_refl. -Qed. - -Lemma scalar_max_cons_valid : forall ty, scalar_max ty >= scalar_max_cons ty . -Proof. - destruct ty; unfold scalar_max_cons, scalar_max; try lia. - - pose isize_max_bound; lia. - - pose usize_max_bound. lia. -Qed. - -Definition scalar (ty: scalar_ty) : Type := - { x: Z | scalar_min ty <= x <= scalar_max ty }. - -Definition to_Z {ty} (x: scalar ty) : Z := proj1_sig x. - -(** Bounds checks: we start by using the conservative bounds, to make sure we - can compute in most situations, then we use the real bounds (for [isize] - and [usize]). *) -Definition scalar_ge_min (ty: scalar_ty) (x: Z) : bool := - Z.leb (scalar_min_cons ty) x || Z.leb (scalar_min ty) x. - -Definition scalar_le_max (ty: scalar_ty) (x: Z) : bool := - Z.leb x (scalar_max_cons ty) || Z.leb x (scalar_max ty). - -Lemma scalar_ge_min_valid (ty: scalar_ty) (x: Z) : - scalar_ge_min ty x = true -> scalar_min ty <= x . -Proof. - unfold scalar_ge_min. - pose (scalar_min_cons_valid ty). - lia. -Qed. - -Lemma scalar_le_max_valid (ty: scalar_ty) (x: Z) : - scalar_le_max ty x = true -> x <= scalar_max ty . -Proof. - unfold scalar_le_max. - pose (scalar_max_cons_valid ty). - lia. -Qed. - -Definition scalar_in_bounds (ty: scalar_ty) (x: Z) : bool := - scalar_ge_min ty x && scalar_le_max ty x . - -Lemma scalar_in_bounds_valid (ty: scalar_ty) (x: Z) : - scalar_in_bounds ty x = true -> scalar_min ty <= x <= scalar_max ty . -Proof. - unfold scalar_in_bounds. - intros H. - destruct (scalar_ge_min ty x) eqn:Hmin. - - destruct (scalar_le_max ty x) eqn:Hmax. - + pose (scalar_ge_min_valid ty x Hmin). - pose (scalar_le_max_valid ty x Hmax). - lia. - + inversion H. - - inversion H. -Qed. - -Import Sumbool. - -Definition mk_scalar (ty: scalar_ty) (x: Z) : result (scalar ty) := - match sumbool_of_bool (scalar_in_bounds ty x) with - | left H => Ok (exist _ x (scalar_in_bounds_valid _ _ H)) - | right _ => Fail_ Failure - end. - -Definition scalar_add {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x + to_Z y). - -Definition scalar_sub {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x - to_Z y). - -Definition scalar_mul {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x * to_Z y). - -Definition scalar_div {ty} (x y: scalar ty) : result (scalar ty) := - if to_Z y =? 0 then Fail_ Failure else - mk_scalar ty (to_Z x / to_Z y). - -Definition scalar_rem {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (Z.rem (to_Z x) (to_Z y)). - -Definition scalar_neg {ty} (x: scalar ty) : result (scalar ty) := mk_scalar ty (-(to_Z x)). - -Axiom scalar_xor : forall ty, scalar ty -> scalar ty -> scalar ty. (* TODO *) -Axiom scalar_or : forall ty, scalar ty -> scalar ty -> scalar ty. (* TODO *) -Axiom scalar_and : forall ty, scalar ty -> scalar ty -> scalar ty. (* TODO *) -Axiom scalar_shl : forall ty0 ty1, scalar ty0 -> scalar ty1 -> result (scalar ty0). (* TODO *) -Axiom scalar_shr : forall ty0 ty1, scalar ty0 -> scalar ty1 -> result (scalar ty0). (* TODO *) - -(** Cast an integer from a [src_ty] to a [tgt_ty] *) -(* TODO: check the semantics of casts in Rust *) -Definition scalar_cast (src_ty tgt_ty : scalar_ty) (x : scalar src_ty) : result (scalar tgt_ty) := - mk_scalar tgt_ty (to_Z x). - -(* This can't fail, but for now we make all casts faillible (easier for the translation) *) -Definition scalar_cast_bool (tgt_ty : scalar_ty) (x : bool) : result (scalar tgt_ty) := - mk_scalar tgt_ty (if x then 1 else 0). - -(** Comparisons *) -Definition scalar_leb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := - Z.leb (to_Z x) (to_Z y) . - -Definition scalar_ltb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := - Z.ltb (to_Z x) (to_Z y) . - -Definition scalar_geb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := - Z.geb (to_Z x) (to_Z y) . - -Definition scalar_gtb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := - Z.gtb (to_Z x) (to_Z y) . - -Definition scalar_eqb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := - Z.eqb (to_Z x) (to_Z y) . - -Definition scalar_neqb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := - negb (Z.eqb (to_Z x) (to_Z y)) . - - -(** The scalar types *) -Definition isize := scalar Isize. -Definition i8 := scalar I8. -Definition i16 := scalar I16. -Definition i32 := scalar I32. -Definition i64 := scalar I64. -Definition i128 := scalar I128. -Definition usize := scalar Usize. -Definition u8 := scalar U8. -Definition u16 := scalar U16. -Definition u32 := scalar U32. -Definition u64 := scalar U64. -Definition u128 := scalar U128. - -(** Negaion *) -Definition isize_neg := @scalar_neg Isize. -Definition i8_neg := @scalar_neg I8. -Definition i16_neg := @scalar_neg I16. -Definition i32_neg := @scalar_neg I32. -Definition i64_neg := @scalar_neg I64. -Definition i128_neg := @scalar_neg I128. - -(** Division *) -Definition isize_div := @scalar_div Isize. -Definition i8_div := @scalar_div I8. -Definition i16_div := @scalar_div I16. -Definition i32_div := @scalar_div I32. -Definition i64_div := @scalar_div I64. -Definition i128_div := @scalar_div I128. -Definition usize_div := @scalar_div Usize. -Definition u8_div := @scalar_div U8. -Definition u16_div := @scalar_div U16. -Definition u32_div := @scalar_div U32. -Definition u64_div := @scalar_div U64. -Definition u128_div := @scalar_div U128. - -(** Remainder *) -Definition isize_rem := @scalar_rem Isize. -Definition i8_rem := @scalar_rem I8. -Definition i16_rem := @scalar_rem I16. -Definition i32_rem := @scalar_rem I32. -Definition i64_rem := @scalar_rem I64. -Definition i128_rem := @scalar_rem I128. -Definition usize_rem := @scalar_rem Usize. -Definition u8_rem := @scalar_rem U8. -Definition u16_rem := @scalar_rem U16. -Definition u32_rem := @scalar_rem U32. -Definition u64_rem := @scalar_rem U64. -Definition u128_rem := @scalar_rem U128. - -(** Addition *) -Definition isize_add := @scalar_add Isize. -Definition i8_add := @scalar_add I8. -Definition i16_add := @scalar_add I16. -Definition i32_add := @scalar_add I32. -Definition i64_add := @scalar_add I64. -Definition i128_add := @scalar_add I128. -Definition usize_add := @scalar_add Usize. -Definition u8_add := @scalar_add U8. -Definition u16_add := @scalar_add U16. -Definition u32_add := @scalar_add U32. -Definition u64_add := @scalar_add U64. -Definition u128_add := @scalar_add U128. - -(** Substraction *) -Definition isize_sub := @scalar_sub Isize. -Definition i8_sub := @scalar_sub I8. -Definition i16_sub := @scalar_sub I16. -Definition i32_sub := @scalar_sub I32. -Definition i64_sub := @scalar_sub I64. -Definition i128_sub := @scalar_sub I128. -Definition usize_sub := @scalar_sub Usize. -Definition u8_sub := @scalar_sub U8. -Definition u16_sub := @scalar_sub U16. -Definition u32_sub := @scalar_sub U32. -Definition u64_sub := @scalar_sub U64. -Definition u128_sub := @scalar_sub U128. - -(** Multiplication *) -Definition isize_mul := @scalar_mul Isize. -Definition i8_mul := @scalar_mul I8. -Definition i16_mul := @scalar_mul I16. -Definition i32_mul := @scalar_mul I32. -Definition i64_mul := @scalar_mul I64. -Definition i128_mul := @scalar_mul I128. -Definition usize_mul := @scalar_mul Usize. -Definition u8_mul := @scalar_mul U8. -Definition u16_mul := @scalar_mul U16. -Definition u32_mul := @scalar_mul U32. -Definition u64_mul := @scalar_mul U64. -Definition u128_mul := @scalar_mul U128. - -(** Xor *) -Definition u8_xor := @scalar_xor U8. -Definition u16_xor := @scalar_xor U16. -Definition u32_xor := @scalar_xor U32. -Definition u64_xor := @scalar_xor U64. -Definition u128_xor := @scalar_xor U128. -Definition usize_xor := @scalar_xor Usize. -Definition i8_xor := @scalar_xor I8. -Definition i16_xor := @scalar_xor I16. -Definition i32_xor := @scalar_xor I32. -Definition i64_xor := @scalar_xor I64. -Definition i128_xor := @scalar_xor I128. -Definition isize_xor := @scalar_xor Isize. - -(** Or *) -Definition u8_or := @scalar_or U8. -Definition u16_or := @scalar_or U16. -Definition u32_or := @scalar_or U32. -Definition u64_or := @scalar_or U64. -Definition u128_or := @scalar_or U128. -Definition usize_or := @scalar_or Usize. -Definition i8_or := @scalar_or I8. -Definition i16_or := @scalar_or I16. -Definition i32_or := @scalar_or I32. -Definition i64_or := @scalar_or I64. -Definition i128_or := @scalar_or I128. -Definition isize_or := @scalar_or Isize. - -(** And *) -Definition u8_and := @scalar_and U8. -Definition u16_and := @scalar_and U16. -Definition u32_and := @scalar_and U32. -Definition u64_and := @scalar_and U64. -Definition u128_and := @scalar_and U128. -Definition usize_and := @scalar_and Usize. -Definition i8_and := @scalar_and I8. -Definition i16_and := @scalar_and I16. -Definition i32_and := @scalar_and I32. -Definition i64_and := @scalar_and I64. -Definition i128_and := @scalar_and I128. -Definition isize_and := @scalar_and Isize. - -(** Shift left *) -Definition u8_shl {ty} := @scalar_shl U8 ty. -Definition u16_shl {ty} := @scalar_shl U16 ty. -Definition u32_shl {ty} := @scalar_shl U32 ty. -Definition u64_shl {ty} := @scalar_shl U64 ty. -Definition u128_shl {ty} := @scalar_shl U128 ty. -Definition usize_shl {ty} := @scalar_shl Usize ty. -Definition i8_shl {ty} := @scalar_shl I8 ty. -Definition i16_shl {ty} := @scalar_shl I16 ty. -Definition i32_shl {ty} := @scalar_shl I32 ty. -Definition i64_shl {ty} := @scalar_shl I64 ty. -Definition i128_shl {ty} := @scalar_shl I128 ty. -Definition isize_shl {ty} := @scalar_shl Isize ty. - -(** Shift right *) -Definition u8_shr {ty} := @scalar_shr U8 ty. -Definition u16_shr {ty} := @scalar_shr U16 ty. -Definition u32_shr {ty} := @scalar_shr U32 ty. -Definition u64_shr {ty} := @scalar_shr U64 ty. -Definition u128_shr {ty} := @scalar_shr U128 ty. -Definition usize_shr {ty} := @scalar_shr Usize ty. -Definition i8_shr {ty} := @scalar_shr I8 ty. -Definition i16_shr {ty} := @scalar_shr I16 ty. -Definition i32_shr {ty} := @scalar_shr I32 ty. -Definition i64_shr {ty} := @scalar_shr I64 ty. -Definition i128_shr {ty} := @scalar_shr I128 ty. -Definition isize_shr {ty} := @scalar_shr Isize ty. - -(** Small utility *) -Definition usize_to_nat (x: usize) : nat := Z.to_nat (to_Z x). - -(** Notations *) -Notation "x %isize" := ((mk_scalar Isize x)%return) (at level 9). -Notation "x %i8" := ((mk_scalar I8 x)%return) (at level 9). -Notation "x %i16" := ((mk_scalar I16 x)%return) (at level 9). -Notation "x %i32" := ((mk_scalar I32 x)%return) (at level 9). -Notation "x %i64" := ((mk_scalar I64 x)%return) (at level 9). -Notation "x %i128" := ((mk_scalar I128 x)%return) (at level 9). -Notation "x %usize" := ((mk_scalar Usize x)%return) (at level 9). -Notation "x %u8" := ((mk_scalar U8 x)%return) (at level 9). -Notation "x %u16" := ((mk_scalar U16 x)%return) (at level 9). -Notation "x %u32" := ((mk_scalar U32 x)%return) (at level 9). -Notation "x %u64" := ((mk_scalar U64 x)%return) (at level 9). -Notation "x %u128" := ((mk_scalar U128 x)%return) (at level 9). - -Notation "x s= y" := (scalar_eqb x y) (at level 80) : Primitives_scope. -Notation "x s<> y" := (scalar_neqb x y) (at level 80) : Primitives_scope. -Notation "x s<= y" := (scalar_leb x y) (at level 80) : Primitives_scope. -Notation "x s< y" := (scalar_ltb x y) (at level 80) : Primitives_scope. -Notation "x s>= y" := (scalar_geb x y) (at level 80) : Primitives_scope. -Notation "x s> y" := (scalar_gtb x y) (at level 80) : Primitives_scope. - -(** Constants *) -Definition core_u8_max := u8_max %u32. -Definition core_u16_max := u16_max %u32. -Definition core_u32_max := u32_max %u32. -Definition core_u64_max := u64_max %u64. -Definition core_u128_max := u64_max %u128. -Axiom core_usize_max : usize. (** TODO *) -Definition core_i8_max := i8_max %i32. -Definition core_i16_max := i16_max %i32. -Definition core_i32_max := i32_max %i32. -Definition core_i64_max := i64_max %i64. -Definition core_i128_max := i64_max %i128. -Axiom core_isize_max : isize. (** TODO *) - -(*** core *) - -(** Trait declaration: [core::clone::Clone] *) -Record core_clone_Clone (self : Type) := { - clone : self -> result self -}. - -Definition core_clone_impls_CloneBool_clone (b : bool) : bool := b. - -Definition core_clone_CloneBool : core_clone_Clone bool := {| - clone := fun b => Ok (core_clone_impls_CloneBool_clone b) -|}. - -Definition core_clone_impls_CloneUsize_clone (x : usize) : usize := x. -Definition core_clone_impls_CloneU8_clone (x : u8) : u8 := x. -Definition core_clone_impls_CloneU16_clone (x : u16) : u16 := x. -Definition core_clone_impls_CloneU32_clone (x : u32) : u32 := x. -Definition core_clone_impls_CloneU64_clone (x : u64) : u64 := x. -Definition core_clone_impls_CloneU128_clone (x : u128) : u128 := x. - -Definition core_clone_impls_CloneIsize_clone (x : isize) : isize := x. -Definition core_clone_impls_CloneI8_clone (x : i8) : i8 := x. -Definition core_clone_impls_CloneI16_clone (x : i16) : i16 := x. -Definition core_clone_impls_CloneI32_clone (x : i32) : i32 := x. -Definition core_clone_impls_CloneI64_clone (x : i64) : i64 := x. -Definition core_clone_impls_CloneI128_clone (x : i128) : i128 := x. - -Definition core_clone_CloneUsize : core_clone_Clone usize := {| - clone := fun x => Ok (core_clone_impls_CloneUsize_clone x) -|}. - -Definition core_clone_CloneU8 : core_clone_Clone u8 := {| - clone := fun x => Ok (core_clone_impls_CloneU8_clone x) -|}. - -Definition core_clone_CloneU16 : core_clone_Clone u16 := {| - clone := fun x => Ok (core_clone_impls_CloneU16_clone x) -|}. - -Definition core_clone_CloneU32 : core_clone_Clone u32 := {| - clone := fun x => Ok (core_clone_impls_CloneU32_clone x) -|}. - -Definition core_clone_CloneU64 : core_clone_Clone u64 := {| - clone := fun x => Ok (core_clone_impls_CloneU64_clone x) -|}. - -Definition core_clone_CloneU128 : core_clone_Clone u128 := {| - clone := fun x => Ok (core_clone_impls_CloneU128_clone x) -|}. - -Definition core_clone_CloneIsize : core_clone_Clone isize := {| - clone := fun x => Ok (core_clone_impls_CloneIsize_clone x) -|}. - -Definition core_clone_CloneI8 : core_clone_Clone i8 := {| - clone := fun x => Ok (core_clone_impls_CloneI8_clone x) -|}. - -Definition core_clone_CloneI16 : core_clone_Clone i16 := {| - clone := fun x => Ok (core_clone_impls_CloneI16_clone x) -|}. - -Definition core_clone_CloneI32 : core_clone_Clone i32 := {| - clone := fun x => Ok (core_clone_impls_CloneI32_clone x) -|}. - -Definition core_clone_CloneI64 : core_clone_Clone i64 := {| - clone := fun x => Ok (core_clone_impls_CloneI64_clone x) -|}. - -Definition core_clone_CloneI128 : core_clone_Clone i128 := {| - clone := fun x => Ok (core_clone_impls_CloneI128_clone x) -|}. - -(** [core::option::{core::option::Option}::unwrap] *) -Definition core_option_Option_unwrap (T : Type) (x : option T) : result T := - match x with - | None => Fail_ Failure - | Some x => Ok x - end. - -(*** core::ops *) - -(* Trait declaration: [core::ops::index::Index] *) -Record core_ops_index_Index (Self Idx : Type) := mk_core_ops_index_Index { - core_ops_index_Index_Output : Type; - core_ops_index_Index_index : Self -> Idx -> result core_ops_index_Index_Output; -}. -Arguments mk_core_ops_index_Index {_ _}. -Arguments core_ops_index_Index_Output {_ _}. -Arguments core_ops_index_Index_index {_ _}. - -(* Trait declaration: [core::ops::index::IndexMut] *) -Record core_ops_index_IndexMut (Self Idx : Type) := mk_core_ops_index_IndexMut { - core_ops_index_IndexMut_indexInst : core_ops_index_Index Self Idx; - core_ops_index_IndexMut_index_mut : - Self -> - Idx -> - result (core_ops_index_IndexMut_indexInst.(core_ops_index_Index_Output) * - (core_ops_index_IndexMut_indexInst.(core_ops_index_Index_Output) -> result Self)); -}. -Arguments mk_core_ops_index_IndexMut {_ _}. -Arguments core_ops_index_IndexMut_indexInst {_ _}. -Arguments core_ops_index_IndexMut_index_mut {_ _}. - -(* Trait declaration [core::ops::deref::Deref] *) -Record core_ops_deref_Deref (Self : Type) := mk_core_ops_deref_Deref { - core_ops_deref_Deref_target : Type; - core_ops_deref_Deref_deref : Self -> result core_ops_deref_Deref_target; -}. -Arguments mk_core_ops_deref_Deref {_}. -Arguments core_ops_deref_Deref_target {_}. -Arguments core_ops_deref_Deref_deref {_}. - -(* Trait declaration [core::ops::deref::DerefMut] *) -Record core_ops_deref_DerefMut (Self : Type) := mk_core_ops_deref_DerefMut { - core_ops_deref_DerefMut_derefInst : core_ops_deref_Deref Self; - core_ops_deref_DerefMut_deref_mut : - Self -> - result (core_ops_deref_DerefMut_derefInst.(core_ops_deref_Deref_target) * - (core_ops_deref_DerefMut_derefInst.(core_ops_deref_Deref_target) -> result Self)); -}. -Arguments mk_core_ops_deref_DerefMut {_}. -Arguments core_ops_deref_DerefMut_derefInst {_}. -Arguments core_ops_deref_DerefMut_deref_mut {_}. - -Record core_ops_range_Range (T : Type) := mk_core_ops_range_Range { - core_ops_range_Range_start : T; - core_ops_range_Range_end_ : T; -}. -Arguments mk_core_ops_range_Range {_}. -Arguments core_ops_range_Range_start {_}. -Arguments core_ops_range_Range_end_ {_}. - -(*** [alloc] *) - -Definition alloc_boxed_Box_deref (T : Type) (x : T) : result T := Ok x. -Definition alloc_boxed_Box_deref_mut (T : Type) (x : T) : result (T * (T -> result T)) := - Ok (x, fun x => Ok x). - -(* Trait instance *) -Definition alloc_boxed_Box_coreopsDerefInst (Self : Type) : core_ops_deref_Deref Self := {| - core_ops_deref_Deref_target := Self; - core_ops_deref_Deref_deref := alloc_boxed_Box_deref Self; -|}. - -(* Trait instance *) -Definition alloc_boxed_Box_coreopsDerefMutInst (Self : Type) : core_ops_deref_DerefMut Self := {| - core_ops_deref_DerefMut_derefInst := alloc_boxed_Box_coreopsDerefInst Self; - core_ops_deref_DerefMut_deref_mut := alloc_boxed_Box_deref_mut Self; -|}. - - -(*** Arrays *) -Definition array T (n : usize) := { l: list T | Z.of_nat (length l) = to_Z n}. - -Lemma le_0_usize_max : 0 <= usize_max. -Proof. - pose (H := usize_max_bound). - unfold u32_max in H. - lia. -Qed. - -Lemma eqb_imp_eq (x y : Z) : Z.eqb x y = true -> x = y. -Proof. - lia. -Qed. - -(* TODO: finish the definitions *) -Axiom mk_array : forall (T : Type) (n : usize) (l : list T), array T n. - -(* For initialization *) -Axiom array_repeat : forall (T : Type) (n : usize) (x : T), array T n. - -Axiom array_index_usize : forall (T : Type) (n : usize) (x : array T n) (i : usize), result T. -Axiom array_update_usize : forall (T : Type) (n : usize) (x : array T n) (i : usize) (nx : T), result (array T n). - -Definition array_index_mut_usize (T : Type) (n : usize) (a : array T n) (i : usize) : - result (T * (T -> result (array T n))) := - match array_index_usize T n a i with - | Fail_ e => Fail_ e - | Ok x => Ok (x, array_update_usize T n a i) - end. - -(*** Slice *) -Definition slice T := { l: list T | Z.of_nat (length l) <= usize_max}. - -Axiom slice_len : forall (T : Type) (s : slice T), usize. -Axiom slice_index_usize : forall (T : Type) (x : slice T) (i : usize), result T. -Axiom slice_update_usize : forall (T : Type) (x : slice T) (i : usize) (nx : T), result (slice T). - -Definition slice_index_mut_usize (T : Type) (s : slice T) (i : usize) : - result (T * (T -> result (slice T))) := - match slice_index_usize T s i with - | Fail_ e => Fail_ e - | Ok x => Ok (x, slice_update_usize T s i) - end. - -(*** Subslices *) - -Axiom array_to_slice : forall (T : Type) (n : usize) (x : array T n), result (slice T). -Axiom array_from_slice : forall (T : Type) (n : usize) (x : array T n) (s : slice T), result (array T n). - -Definition array_to_slice_mut (T : Type) (n : usize) (a : array T n) : - result (slice T * (slice T -> result (array T n))) := - match array_to_slice T n a with - | Fail_ e => Fail_ e - | Ok x => Ok (x, array_from_slice T n a) - end. - -Axiom array_subslice: forall (T : Type) (n : usize) (x : array T n) (r : core_ops_range_Range usize), result (slice T). -Axiom array_update_subslice: forall (T : Type) (n : usize) (x : array T n) (r : core_ops_range_Range usize) (ns : slice T), result (array T n). - -Axiom slice_subslice: forall (T : Type) (x : slice T) (r : core_ops_range_Range usize), result (slice T). -Axiom slice_update_subslice: forall (T : Type) (x : slice T) (r : core_ops_range_Range usize) (ns : slice T), result (slice T). - -(*** Vectors *) - -Definition alloc_vec_Vec T := { l: list T | Z.of_nat (length l) <= usize_max }. - -Definition alloc_vec_Vec_to_list {T: Type} (v: alloc_vec_Vec T) : list T := proj1_sig v. - -Definition alloc_vec_Vec_length {T: Type} (v: alloc_vec_Vec T) : Z := Z.of_nat (length (alloc_vec_Vec_to_list v)). - -Definition alloc_vec_Vec_new (T: Type) : alloc_vec_Vec T := (exist _ [] le_0_usize_max). - -Lemma alloc_vec_Vec_len_in_usize {T} (v: alloc_vec_Vec T) : usize_min <= alloc_vec_Vec_length v <= usize_max. -Proof. - unfold alloc_vec_Vec_length, usize_min. - split. - - lia. - - apply (proj2_sig v). -Qed. - -Definition alloc_vec_Vec_len (T: Type) (v: alloc_vec_Vec T) : usize := - exist _ (alloc_vec_Vec_length v) (alloc_vec_Vec_len_in_usize v). - -Fixpoint list_update {A} (l: list A) (n: nat) (a: A) - : list A := - match l with - | [] => [] - | x :: t => match n with - | 0%nat => a :: t - | S m => x :: (list_update t m a) -end end. - -Definition alloc_vec_Vec_bind {A B} (v: alloc_vec_Vec A) (f: list A -> result (list B)) : result (alloc_vec_Vec B) := - l <- f (alloc_vec_Vec_to_list v) ; - match sumbool_of_bool (scalar_le_max Usize (Z.of_nat (length l))) with - | left H => Ok (exist _ l (scalar_le_max_valid _ _ H)) - | right _ => Fail_ Failure - end. - -Definition alloc_vec_Vec_push (T: Type) (v: alloc_vec_Vec T) (x: T) : result (alloc_vec_Vec T) := - alloc_vec_Vec_bind v (fun l => Ok (l ++ [x])). - -Definition alloc_vec_Vec_insert (T: Type) (v: alloc_vec_Vec T) (i: usize) (x: T) : result (alloc_vec_Vec T) := - alloc_vec_Vec_bind v (fun l => - if to_Z i result (alloc_vec_Vec T))) := - match alloc_vec_Vec_index_usize v i with - | Ok x => - Ok (x, alloc_vec_Vec_update_usize v i) - | Fail_ e => Fail_ e - end. - -(* Trait declaration: [core::slice::index::private_slice_index::Sealed] *) -Definition core_slice_index_private_slice_index_Sealed (self : Type) := unit. - -(* Trait declaration: [core::slice::index::SliceIndex] *) -Record core_slice_index_SliceIndex (Self T : Type) := mk_core_slice_index_SliceIndex { - core_slice_index_SliceIndex_sealedInst : core_slice_index_private_slice_index_Sealed Self; - core_slice_index_SliceIndex_Output : Type; - core_slice_index_SliceIndex_get : Self -> T -> result (option core_slice_index_SliceIndex_Output); - core_slice_index_SliceIndex_get_mut : - Self -> T -> result (option core_slice_index_SliceIndex_Output * (option core_slice_index_SliceIndex_Output -> result T)); - core_slice_index_SliceIndex_get_unchecked : Self -> const_raw_ptr T -> result (const_raw_ptr core_slice_index_SliceIndex_Output); - core_slice_index_SliceIndex_get_unchecked_mut : Self -> mut_raw_ptr T -> result (mut_raw_ptr core_slice_index_SliceIndex_Output); - core_slice_index_SliceIndex_index : Self -> T -> result core_slice_index_SliceIndex_Output; - core_slice_index_SliceIndex_index_mut : - Self -> T -> result (core_slice_index_SliceIndex_Output * (core_slice_index_SliceIndex_Output -> result T)); -}. -Arguments mk_core_slice_index_SliceIndex {_ _}. -Arguments core_slice_index_SliceIndex_sealedInst {_ _}. -Arguments core_slice_index_SliceIndex_Output {_ _}. -Arguments core_slice_index_SliceIndex_get {_ _}. -Arguments core_slice_index_SliceIndex_get_mut {_ _}. -Arguments core_slice_index_SliceIndex_get_unchecked {_ _}. -Arguments core_slice_index_SliceIndex_get_unchecked_mut {_ _}. -Arguments core_slice_index_SliceIndex_index {_ _}. -Arguments core_slice_index_SliceIndex_index_mut {_ _}. - -(* [core::slice::index::[T]::index]: forward function *) -Definition core_slice_index_Slice_index - (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)) - (s : slice T) (i : Idx) : result inst.(core_slice_index_SliceIndex_Output) := - x <- inst.(core_slice_index_SliceIndex_get) i s; - match x with - | None => Fail_ Failure - | Some x => Ok x - end. - -(* [core::slice::index::Range:::get]: forward function *) -Axiom core_slice_index_RangeUsize_get : forall (T : Type) (i : core_ops_range_Range usize) (s : slice T), result (option (slice T)). - -(* [core::slice::index::Range::get_mut]: forward function *) -Axiom core_slice_index_RangeUsize_get_mut : - forall (T : Type), - core_ops_range_Range usize -> slice T -> - result (option (slice T) * (option (slice T) -> result (slice T))). - -(* [core::slice::index::Range::get_unchecked]: forward function *) -Definition core_slice_index_RangeUsize_get_unchecked - (T : Type) : - core_ops_range_Range usize -> const_raw_ptr (slice T) -> result (const_raw_ptr (slice T)) := - (* Don't know what the model should be - for now we always fail to make - sure code which uses it fails *) - fun _ _ => Fail_ Failure. - -(* [core::slice::index::Range::get_unchecked_mut]: forward function *) -Definition core_slice_index_RangeUsize_get_unchecked_mut - (T : Type) : - core_ops_range_Range usize -> mut_raw_ptr (slice T) -> result (mut_raw_ptr (slice T)) := - (* Don't know what the model should be - for now we always fail to make - sure code which uses it fails *) - fun _ _ => Fail_ Failure. - -(* [core::slice::index::Range::index]: forward function *) -Axiom core_slice_index_RangeUsize_index : - forall (T : Type), core_ops_range_Range usize -> slice T -> result (slice T). - -(* [core::slice::index::Range::index_mut]: forward function *) -Axiom core_slice_index_RangeUsize_index_mut : - forall (T : Type), core_ops_range_Range usize -> slice T -> result (slice T * (slice T -> result (slice T))). - -(* [core::slice::index::[T]::index_mut]: forward function *) -Axiom core_slice_index_Slice_index_mut : - forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)), - slice T -> Idx -> - result (inst.(core_slice_index_SliceIndex_Output) * - (inst.(core_slice_index_SliceIndex_Output) -> result (slice T))). - -(* [core::array::[T; N]::index]: forward function *) -Axiom core_array_Array_index : - forall (T Idx : Type) (N : usize) (inst : core_ops_index_Index (slice T) Idx) - (a : array T N) (i : Idx), result inst.(core_ops_index_Index_Output). - -(* [core::array::[T; N]::index_mut]: forward function *) -Axiom core_array_Array_index_mut : - forall (T Idx : Type) (N : usize) (inst : core_ops_index_IndexMut (slice T) Idx) - (a : array T N) (i : Idx), - result (inst.(core_ops_index_IndexMut_indexInst).(core_ops_index_Index_Output) * - (inst.(core_ops_index_IndexMut_indexInst).(core_ops_index_Index_Output) -> result (array T N))). - -(* Trait implementation: [core::slice::index::private_slice_index::Range] *) -Definition core_slice_index_private_slice_index_SealedRangeUsizeInst - : core_slice_index_private_slice_index_Sealed (core_ops_range_Range usize) := tt. - -(* Trait implementation: [core::slice::index::Range] *) -Definition core_slice_index_SliceIndexRangeUsizeSliceTInst (T : Type) : - core_slice_index_SliceIndex (core_ops_range_Range usize) (slice T) := {| - core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_SealedRangeUsizeInst; - core_slice_index_SliceIndex_Output := slice T; - core_slice_index_SliceIndex_get := core_slice_index_RangeUsize_get T; - core_slice_index_SliceIndex_get_mut := core_slice_index_RangeUsize_get_mut T; - core_slice_index_SliceIndex_get_unchecked := core_slice_index_RangeUsize_get_unchecked T; - core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_RangeUsize_get_unchecked_mut T; - core_slice_index_SliceIndex_index := core_slice_index_RangeUsize_index T; - core_slice_index_SliceIndex_index_mut := core_slice_index_RangeUsize_index_mut T; -|}. - -(* Trait implementation: [core::slice::index::[T]] *) -Definition core_ops_index_IndexSliceTIInst (T Idx : Type) - (inst : core_slice_index_SliceIndex Idx (slice T)) : - core_ops_index_Index (slice T) Idx := {| - core_ops_index_Index_Output := inst.(core_slice_index_SliceIndex_Output); - core_ops_index_Index_index := core_slice_index_Slice_index T Idx inst; -|}. - -(* Trait implementation: [core::slice::index::[T]] *) -Definition core_ops_index_IndexMutSliceTIInst (T Idx : Type) - (inst : core_slice_index_SliceIndex Idx (slice T)) : - core_ops_index_IndexMut (slice T) Idx := {| - core_ops_index_IndexMut_indexInst := core_ops_index_IndexSliceTIInst T Idx inst; - core_ops_index_IndexMut_index_mut := core_slice_index_Slice_index_mut T Idx inst; -|}. - -(* Trait implementation: [core::array::[T; N]] *) -Definition core_ops_index_IndexArrayInst (T Idx : Type) (N : usize) - (inst : core_ops_index_Index (slice T) Idx) : - core_ops_index_Index (array T N) Idx := {| - core_ops_index_Index_Output := inst.(core_ops_index_Index_Output); - core_ops_index_Index_index := core_array_Array_index T Idx N inst; -|}. - -(* Trait implementation: [core::array::[T; N]] *) -Definition core_ops_index_IndexMutArrayInst (T Idx : Type) (N : usize) - (inst : core_ops_index_IndexMut (slice T) Idx) : - core_ops_index_IndexMut (array T N) Idx := {| - core_ops_index_IndexMut_indexInst := core_ops_index_IndexArrayInst T Idx N inst.(core_ops_index_IndexMut_indexInst); - core_ops_index_IndexMut_index_mut := core_array_Array_index_mut T Idx N inst; -|}. - -(* [core::slice::index::usize::get]: forward function *) -Axiom core_slice_index_usize_get : forall (T : Type), usize -> slice T -> result (option T). - -(* [core::slice::index::usize::get_mut]: forward function *) -Axiom core_slice_index_usize_get_mut : - forall (T : Type), usize -> slice T -> result (option T * (option T -> result (slice T))). - -(* [core::slice::index::usize::get_unchecked]: forward function *) -Axiom core_slice_index_usize_get_unchecked : - forall (T : Type), usize -> const_raw_ptr (slice T) -> result (const_raw_ptr T). - -(* [core::slice::index::usize::get_unchecked_mut]: forward function *) -Axiom core_slice_index_usize_get_unchecked_mut : - forall (T : Type), usize -> mut_raw_ptr (slice T) -> result (mut_raw_ptr T). - -(* [core::slice::index::usize::index]: forward function *) -Axiom core_slice_index_usize_index : forall (T : Type), usize -> slice T -> result T. - -(* [core::slice::index::usize::index_mut]: forward function *) -Axiom core_slice_index_usize_index_mut : - forall (T : Type), usize -> slice T -> result (T * (T -> result (slice T))). - -(* Trait implementation: [core::slice::index::private_slice_index::usize] *) -Definition core_slice_index_private_slice_index_SealedUsizeInst - : core_slice_index_private_slice_index_Sealed usize := tt. - -(* Trait implementation: [core::slice::index::usize] *) -Definition core_slice_index_SliceIndexUsizeSliceTInst (T : Type) : - core_slice_index_SliceIndex usize (slice T) := {| - core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_SealedUsizeInst; - core_slice_index_SliceIndex_Output := T; - core_slice_index_SliceIndex_get := core_slice_index_usize_get T; - core_slice_index_SliceIndex_get_mut := core_slice_index_usize_get_mut T; - core_slice_index_SliceIndex_get_unchecked := core_slice_index_usize_get_unchecked T; - core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_usize_get_unchecked_mut T; - core_slice_index_SliceIndex_index := core_slice_index_usize_index T; - core_slice_index_SliceIndex_index_mut := core_slice_index_usize_index_mut T; -|}. - -(* [alloc::vec::Vec::index]: forward function *) -Axiom alloc_vec_Vec_index : forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)) - (Self : alloc_vec_Vec T) (i : Idx), result inst.(core_slice_index_SliceIndex_Output). - -(* [alloc::vec::Vec::index_mut]: forward function *) -Axiom alloc_vec_Vec_index_mut : forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)) - (Self : alloc_vec_Vec T) (i : Idx), - result (inst.(core_slice_index_SliceIndex_Output) * - (inst.(core_slice_index_SliceIndex_Output) -> result (alloc_vec_Vec T))). - -(* Trait implementation: [alloc::vec::Vec] *) -Definition alloc_vec_Vec_coreopsindexIndexInst (T Idx : Type) - (inst : core_slice_index_SliceIndex Idx (slice T)) : - core_ops_index_Index (alloc_vec_Vec T) Idx := {| - core_ops_index_Index_Output := inst.(core_slice_index_SliceIndex_Output); - core_ops_index_Index_index := alloc_vec_Vec_index T Idx inst; -|}. - -(* Trait implementation: [alloc::vec::Vec] *) -Definition alloc_vec_Vec_coreopsindexIndexMutInst (T Idx : Type) - (inst : core_slice_index_SliceIndex Idx (slice T)) : - core_ops_index_IndexMut (alloc_vec_Vec T) Idx := {| - core_ops_index_IndexMut_indexInst := alloc_vec_Vec_coreopsindexIndexInst T Idx inst; - core_ops_index_IndexMut_index_mut := alloc_vec_Vec_index_mut T Idx inst; -|}. - -(*** Theorems *) - -Axiom alloc_vec_Vec_index_eq : forall {a : Type} (v : alloc_vec_Vec a) (i : usize) (x : a), - alloc_vec_Vec_index a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i = - alloc_vec_Vec_index_usize v i. - -Axiom alloc_vec_Vec_index_mut_eq : forall {a : Type} (v : alloc_vec_Vec a) (i : usize) (x : a), - alloc_vec_Vec_index_mut a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i = - alloc_vec_Vec_index_mut_usize v i. - -End Primitives. diff --git a/tests/coq/hashmap_on_disk/_CoqProject b/tests/coq/hashmap_on_disk/_CoqProject deleted file mode 100644 index d73541d9..00000000 --- a/tests/coq/hashmap_on_disk/_CoqProject +++ /dev/null @@ -1,12 +0,0 @@ -# This file was automatically generated - see ../Makefile --R . Lib --arg -w --arg all - -HashmapMain_Types.v -HashmapMain_FunsExternal_Template.v -Primitives.v -HashmapMain_Funs.v -HashmapMain_TypesExternal.v -HashmapMain_FunsExternal.v -HashmapMain_TypesExternal_Template.v diff --git a/tests/fstar/hashmap_main/HashmapMain.Clauses.Template.fst b/tests/fstar/hashmap_main/HashmapMain.Clauses.Template.fst new file mode 100644 index 00000000..cdd73210 --- /dev/null +++ b/tests/fstar/hashmap_main/HashmapMain.Clauses.Template.fst @@ -0,0 +1,72 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: templates for the decreases clauses *) +module HashmapMain.Clauses.Template +open Primitives +open HashmapMain.Types + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: decreases clause + Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) +unfold +let hashmap_HashMap_allocate_slots_loop_decreases (t : Type0) + (slots : alloc_vec_Vec (hashmap_List_t t)) (n : usize) : nat = + admit () + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: decreases clause + Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) +unfold +let hashmap_HashMap_clear_loop_decreases (t : Type0) + (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : nat = + admit () + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: decreases clause + Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) +unfold +let hashmap_HashMap_insert_in_list_loop_decreases (t : Type0) (key : usize) + (value : t) (ls : hashmap_List_t t) : nat = + admit () + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: decreases clause + Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) +unfold +let hashmap_HashMap_move_elements_from_list_loop_decreases (t : Type0) + (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : nat = + admit () + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: decreases clause + Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) +unfold +let hashmap_HashMap_move_elements_loop_decreases (t : Type0) + (ntable : hashmap_HashMap_t t) (slots : alloc_vec_Vec (hashmap_List_t t)) + (i : usize) : nat = + admit () + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: decreases clause + Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) +unfold +let hashmap_HashMap_contains_key_in_list_loop_decreases (t : Type0) + (key : usize) (ls : hashmap_List_t t) : nat = + admit () + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: decreases clause + Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) +unfold +let hashmap_HashMap_get_in_list_loop_decreases (t : Type0) (key : usize) + (ls : hashmap_List_t t) : nat = + admit () + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: decreases clause + Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) +unfold +let hashmap_HashMap_get_mut_in_list_loop_decreases (t : Type0) + (ls : hashmap_List_t t) (key : usize) : nat = + admit () + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: decreases clause + Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) +unfold +let hashmap_HashMap_remove_from_list_loop_decreases (t : Type0) (key : usize) + (ls : hashmap_List_t t) : nat = + admit () + diff --git a/tests/fstar/hashmap_main/HashmapMain.Clauses.fst b/tests/fstar/hashmap_main/HashmapMain.Clauses.fst new file mode 100644 index 00000000..be5a4ab1 --- /dev/null +++ b/tests/fstar/hashmap_main/HashmapMain.Clauses.fst @@ -0,0 +1,61 @@ +(** [hashmap]: the decreases clauses *) +module HashmapMain.Clauses +open Primitives +open FStar.List.Tot +open HashmapMain.Types + +#set-options "--z3rlimit 50 --fuel 0 --ifuel 1" + +(** [hashmap::HashMap::allocate_slots]: decreases clause *) +unfold +let hashmap_HashMap_allocate_slots_loop_decreases (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) + (n : usize) : nat = n + +(** [hashmap::HashMap::clear]: decreases clause *) +unfold +let hashmap_HashMap_clear_loop_decreases (t : Type0) (slots : alloc_vec_Vec (hashmap_List_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_HashMap_insert_in_list_loop_decreases (t : Type0) (key : usize) (value : t) + (ls : hashmap_List_t t) : hashmap_List_t t = + ls + +(** [hashmap::HashMap::move_elements_from_list]: decreases clause *) +unfold +let hashmap_HashMap_move_elements_from_list_loop_decreases (t : Type0) + (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : hashmap_List_t t = + ls + +(** [hashmap::HashMap::move_elements]: decreases clause *) +unfold +let hashmap_HashMap_move_elements_loop_decreases (t : Type0) (ntable : hashmap_HashMap_t t) + (slots : alloc_vec_Vec (hashmap_List_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_HashMap_contains_key_in_list_loop_decreases (t : Type0) (key : usize) + (ls : hashmap_List_t t) : hashmap_List_t t = + ls + +(** [hashmap::HashMap::get_in_list]: decreases clause *) +unfold +let hashmap_HashMap_get_in_list_loop_decreases (t : Type0) (key : usize) (ls : hashmap_List_t t) : + hashmap_List_t t = + ls + +(** [hashmap::HashMap::get_mut_in_list]: decreases clause *) +unfold +let hashmap_HashMap_get_mut_in_list_loop_decreases (t : Type0) + (ls : hashmap_List_t t) (key : usize) : hashmap_List_t t = + ls + +(** [hashmap::HashMap::remove_from_list]: decreases clause *) +unfold +let hashmap_HashMap_remove_from_list_loop_decreases (t : Type0) (key : usize) + (ls : hashmap_List_t t) : hashmap_List_t t = + ls + diff --git a/tests/fstar/hashmap_main/HashmapMain.Funs.fst b/tests/fstar/hashmap_main/HashmapMain.Funs.fst new file mode 100644 index 00000000..c88a746e --- /dev/null +++ b/tests/fstar/hashmap_main/HashmapMain.Funs.fst @@ -0,0 +1,446 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: function definitions *) +module HashmapMain.Funs +open Primitives +include HashmapMain.Types +include HashmapMain.FunsExternal +include HashmapMain.Clauses + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** [hashmap_main::hashmap::hash_key]: + Source: 'tests/src/hashmap.rs', lines 35:0-35:32 *) +let hashmap_hash_key (k : usize) : result usize = + Ok k + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: loop 0: + Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) +let rec hashmap_HashMap_allocate_slots_loop + (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (n : usize) : + Tot (result (alloc_vec_Vec (hashmap_List_t t))) + (decreases (hashmap_HashMap_allocate_slots_loop_decreases t slots n)) + = + if n > 0 + then + let* slots1 = alloc_vec_Vec_push (hashmap_List_t t) slots Hashmap_List_Nil + in + let* n1 = usize_sub n 1 in + hashmap_HashMap_allocate_slots_loop t slots1 n1 + else Ok slots + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: + Source: 'tests/src/hashmap.rs', lines 58:4-58:76 *) +let hashmap_HashMap_allocate_slots + (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (n : usize) : + result (alloc_vec_Vec (hashmap_List_t t)) + = + hashmap_HashMap_allocate_slots_loop t slots n + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new_with_capacity]: + Source: 'tests/src/hashmap.rs', lines 67:4-71:13 *) +let hashmap_HashMap_new_with_capacity + (t : Type0) (capacity : usize) (max_load_dividend : usize) + (max_load_divisor : usize) : + result (hashmap_HashMap_t t) + = + let* slots = + hashmap_HashMap_allocate_slots t (alloc_vec_Vec_new (hashmap_List_t t)) + capacity in + let* i = usize_mul capacity max_load_dividend in + let* i1 = usize_div i max_load_divisor in + Ok + { + num_entries = 0; + max_load_factor = (max_load_dividend, max_load_divisor); + max_load = i1; + slots = slots + } + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new]: + Source: 'tests/src/hashmap.rs', lines 83:4-83:24 *) +let hashmap_HashMap_new (t : Type0) : result (hashmap_HashMap_t t) = + hashmap_HashMap_new_with_capacity t 32 4 5 + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: loop 0: + Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) +let rec hashmap_HashMap_clear_loop + (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : + Tot (result (alloc_vec_Vec (hashmap_List_t t))) + (decreases (hashmap_HashMap_clear_loop_decreases t slots i)) + = + let i1 = alloc_vec_Vec_len (hashmap_List_t t) slots in + if i < i1 + then + let* (_, index_mut_back) = + alloc_vec_Vec_index_mut (hashmap_List_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) slots i + in + let* i2 = usize_add i 1 in + let* slots1 = index_mut_back Hashmap_List_Nil in + hashmap_HashMap_clear_loop t slots1 i2 + else Ok slots + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: + Source: 'tests/src/hashmap.rs', lines 88:4-88:27 *) +let hashmap_HashMap_clear + (t : Type0) (self : hashmap_HashMap_t t) : result (hashmap_HashMap_t t) = + let* hm = hashmap_HashMap_clear_loop t self.slots 0 in + Ok { self with num_entries = 0; slots = hm } + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::len]: + Source: 'tests/src/hashmap.rs', lines 98:4-98:30 *) +let hashmap_HashMap_len + (t : Type0) (self : hashmap_HashMap_t t) : result usize = + Ok self.num_entries + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: loop 0: + Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) +let rec hashmap_HashMap_insert_in_list_loop + (t : Type0) (key : usize) (value : t) (ls : hashmap_List_t t) : + Tot (result (bool & (hashmap_List_t t))) + (decreases (hashmap_HashMap_insert_in_list_loop_decreases t key value ls)) + = + begin match ls with + | Hashmap_List_Cons ckey cvalue tl -> + if ckey = key + then Ok (false, Hashmap_List_Cons ckey value tl) + else + let* (b, tl1) = hashmap_HashMap_insert_in_list_loop t key value tl in + Ok (b, Hashmap_List_Cons ckey cvalue tl1) + | Hashmap_List_Nil -> Ok (true, Hashmap_List_Cons key value Hashmap_List_Nil) + end + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: + Source: 'tests/src/hashmap.rs', lines 105:4-105:71 *) +let hashmap_HashMap_insert_in_list + (t : Type0) (key : usize) (value : t) (ls : hashmap_List_t t) : + result (bool & (hashmap_List_t t)) + = + hashmap_HashMap_insert_in_list_loop t key value ls + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_no_resize]: + Source: 'tests/src/hashmap.rs', lines 125:4-125:54 *) +let hashmap_HashMap_insert_no_resize + (t : Type0) (self : hashmap_HashMap_t t) (key : usize) (value : t) : + result (hashmap_HashMap_t t) + = + let* hash = hashmap_hash_key key in + let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in + let* hash_mod = usize_rem hash i in + let* (l, index_mut_back) = + alloc_vec_Vec_index_mut (hashmap_List_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) + self.slots hash_mod in + let* (inserted, l1) = hashmap_HashMap_insert_in_list t key value l in + if inserted + then + let* i1 = usize_add self.num_entries 1 in + let* v = index_mut_back l1 in + Ok { self with num_entries = i1; slots = v } + else let* v = index_mut_back l1 in Ok { self with slots = v } + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: loop 0: + Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) +let rec hashmap_HashMap_move_elements_from_list_loop + (t : Type0) (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : + Tot (result (hashmap_HashMap_t t)) + (decreases ( + hashmap_HashMap_move_elements_from_list_loop_decreases t ntable ls)) + = + begin match ls with + | Hashmap_List_Cons k v tl -> + let* ntable1 = hashmap_HashMap_insert_no_resize t ntable k v in + hashmap_HashMap_move_elements_from_list_loop t ntable1 tl + | Hashmap_List_Nil -> Ok ntable + end + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: + Source: 'tests/src/hashmap.rs', lines 191:4-191:72 *) +let hashmap_HashMap_move_elements_from_list + (t : Type0) (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : + result (hashmap_HashMap_t t) + = + hashmap_HashMap_move_elements_from_list_loop t ntable ls + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: loop 0: + Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) +let rec hashmap_HashMap_move_elements_loop + (t : Type0) (ntable : hashmap_HashMap_t t) + (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : + Tot (result ((hashmap_HashMap_t t) & (alloc_vec_Vec (hashmap_List_t t)))) + (decreases (hashmap_HashMap_move_elements_loop_decreases t ntable slots i)) + = + let i1 = alloc_vec_Vec_len (hashmap_List_t t) slots in + if i < i1 + then + let* (l, index_mut_back) = + alloc_vec_Vec_index_mut (hashmap_List_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) slots i + in + let (ls, l1) = core_mem_replace (hashmap_List_t t) l Hashmap_List_Nil in + let* ntable1 = hashmap_HashMap_move_elements_from_list t ntable ls in + let* i2 = usize_add i 1 in + let* slots1 = index_mut_back l1 in + hashmap_HashMap_move_elements_loop t ntable1 slots1 i2 + else Ok (ntable, slots) + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: + Source: 'tests/src/hashmap.rs', lines 179:4-179:95 *) +let hashmap_HashMap_move_elements + (t : Type0) (ntable : hashmap_HashMap_t t) + (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : + result ((hashmap_HashMap_t t) & (alloc_vec_Vec (hashmap_List_t t))) + = + hashmap_HashMap_move_elements_loop t ntable slots i + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::try_resize]: + Source: 'tests/src/hashmap.rs', lines 148:4-148:28 *) +let hashmap_HashMap_try_resize + (t : Type0) (self : hashmap_HashMap_t t) : result (hashmap_HashMap_t t) = + let* max_usize = scalar_cast U32 Usize core_u32_max in + let capacity = alloc_vec_Vec_len (hashmap_List_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 + if capacity <= i2 + then + let* i3 = usize_mul capacity 2 in + let* ntable = hashmap_HashMap_new_with_capacity t i3 i i1 in + let* p = hashmap_HashMap_move_elements t ntable self.slots 0 in + let (ntable1, _) = p in + Ok + { ntable1 with num_entries = self.num_entries; max_load_factor = (i, i1) + } + else Ok { self with max_load_factor = (i, i1) } + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert]: + Source: 'tests/src/hashmap.rs', lines 137:4-137:48 *) +let hashmap_HashMap_insert + (t : Type0) (self : hashmap_HashMap_t t) (key : usize) (value : t) : + result (hashmap_HashMap_t t) + = + let* self1 = hashmap_HashMap_insert_no_resize t self key value in + let* i = hashmap_HashMap_len t self1 in + if i > self1.max_load then hashmap_HashMap_try_resize t self1 else Ok self1 + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: loop 0: + Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) +let rec hashmap_HashMap_contains_key_in_list_loop + (t : Type0) (key : usize) (ls : hashmap_List_t t) : + Tot (result bool) + (decreases (hashmap_HashMap_contains_key_in_list_loop_decreases t key ls)) + = + begin match ls with + | Hashmap_List_Cons ckey _ tl -> + if ckey = key + then Ok true + else hashmap_HashMap_contains_key_in_list_loop t key tl + | Hashmap_List_Nil -> Ok false + end + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: + Source: 'tests/src/hashmap.rs', lines 214:4-214:68 *) +let hashmap_HashMap_contains_key_in_list + (t : Type0) (key : usize) (ls : hashmap_List_t t) : result bool = + hashmap_HashMap_contains_key_in_list_loop t key ls + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key]: + Source: 'tests/src/hashmap.rs', lines 207:4-207:49 *) +let hashmap_HashMap_contains_key + (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : result bool = + let* hash = hashmap_hash_key key in + let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in + let* hash_mod = usize_rem hash i in + let* l = + alloc_vec_Vec_index (hashmap_List_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) + self.slots hash_mod in + hashmap_HashMap_contains_key_in_list t key l + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: loop 0: + Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) +let rec hashmap_HashMap_get_in_list_loop + (t : Type0) (key : usize) (ls : hashmap_List_t t) : + Tot (result t) + (decreases (hashmap_HashMap_get_in_list_loop_decreases t key ls)) + = + begin match ls with + | Hashmap_List_Cons ckey cvalue tl -> + if ckey = key then Ok cvalue else hashmap_HashMap_get_in_list_loop t key tl + | Hashmap_List_Nil -> Fail Failure + end + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: + Source: 'tests/src/hashmap.rs', lines 232:4-232:70 *) +let hashmap_HashMap_get_in_list + (t : Type0) (key : usize) (ls : hashmap_List_t t) : result t = + hashmap_HashMap_get_in_list_loop t key ls + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get]: + Source: 'tests/src/hashmap.rs', lines 247:4-247:55 *) +let hashmap_HashMap_get + (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : result t = + let* hash = hashmap_hash_key key in + let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in + let* hash_mod = usize_rem hash i in + let* l = + alloc_vec_Vec_index (hashmap_List_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) + self.slots hash_mod in + hashmap_HashMap_get_in_list t key l + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: loop 0: + Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) +let rec hashmap_HashMap_get_mut_in_list_loop + (t : Type0) (ls : hashmap_List_t t) (key : usize) : + Tot (result (t & (t -> result (hashmap_List_t t)))) + (decreases (hashmap_HashMap_get_mut_in_list_loop_decreases t ls key)) + = + begin match ls with + | Hashmap_List_Cons ckey cvalue tl -> + if ckey = key + then + let back = fun ret -> Ok (Hashmap_List_Cons ckey ret tl) in + Ok (cvalue, back) + else + let* (x, back) = hashmap_HashMap_get_mut_in_list_loop t tl key in + let back1 = + fun ret -> + let* tl1 = back ret in Ok (Hashmap_List_Cons ckey cvalue tl1) in + Ok (x, back1) + | Hashmap_List_Nil -> Fail Failure + end + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: + Source: 'tests/src/hashmap.rs', lines 253:4-253:86 *) +let hashmap_HashMap_get_mut_in_list + (t : Type0) (ls : hashmap_List_t t) (key : usize) : + result (t & (t -> result (hashmap_List_t t))) + = + hashmap_HashMap_get_mut_in_list_loop t ls key + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut]: + Source: 'tests/src/hashmap.rs', lines 265:4-265:67 *) +let hashmap_HashMap_get_mut + (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : + result (t & (t -> result (hashmap_HashMap_t t))) + = + let* hash = hashmap_hash_key key in + let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in + let* hash_mod = usize_rem hash i in + let* (l, index_mut_back) = + alloc_vec_Vec_index_mut (hashmap_List_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) + self.slots hash_mod in + let* (x, get_mut_in_list_back) = hashmap_HashMap_get_mut_in_list t l key in + let back = + fun ret -> + let* l1 = get_mut_in_list_back ret in + let* v = index_mut_back l1 in + Ok { self with slots = v } in + Ok (x, back) + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: loop 0: + Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) +let rec hashmap_HashMap_remove_from_list_loop + (t : Type0) (key : usize) (ls : hashmap_List_t t) : + Tot (result ((option t) & (hashmap_List_t t))) + (decreases (hashmap_HashMap_remove_from_list_loop_decreases t key ls)) + = + begin match ls with + | Hashmap_List_Cons ckey x tl -> + if ckey = key + then + let (mv_ls, _) = + core_mem_replace (hashmap_List_t t) (Hashmap_List_Cons ckey x tl) + Hashmap_List_Nil in + begin match mv_ls with + | Hashmap_List_Cons _ cvalue tl1 -> Ok (Some cvalue, tl1) + | Hashmap_List_Nil -> Fail Failure + end + else + let* (o, tl1) = hashmap_HashMap_remove_from_list_loop t key tl in + Ok (o, Hashmap_List_Cons ckey x tl1) + | Hashmap_List_Nil -> Ok (None, Hashmap_List_Nil) + end + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: + Source: 'tests/src/hashmap.rs', lines 273:4-273:69 *) +let hashmap_HashMap_remove_from_list + (t : Type0) (key : usize) (ls : hashmap_List_t t) : + result ((option t) & (hashmap_List_t t)) + = + hashmap_HashMap_remove_from_list_loop t key ls + +(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove]: + Source: 'tests/src/hashmap.rs', lines 302:4-302:52 *) +let hashmap_HashMap_remove + (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : + result ((option t) & (hashmap_HashMap_t t)) + = + let* hash = hashmap_hash_key key in + let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in + let* hash_mod = usize_rem hash i in + let* (l, index_mut_back) = + alloc_vec_Vec_index_mut (hashmap_List_t t) usize + (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) + self.slots hash_mod in + let* (x, l1) = hashmap_HashMap_remove_from_list t key l in + begin match x with + | None -> let* v = index_mut_back l1 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 + Ok (Some x1, { self with num_entries = i1; slots = v }) + end + +(** [hashmap_main::hashmap::test1]: + Source: 'tests/src/hashmap.rs', lines 323:0-323:10 *) +let hashmap_test1 : result unit = + let* hm = hashmap_HashMap_new u64 in + let* hm1 = hashmap_HashMap_insert u64 hm 0 42 in + let* hm2 = hashmap_HashMap_insert u64 hm1 128 18 in + let* hm3 = hashmap_HashMap_insert u64 hm2 1024 138 in + let* hm4 = hashmap_HashMap_insert u64 hm3 1056 256 in + let* i = hashmap_HashMap_get u64 hm4 128 in + if not (i = 18) + then Fail Failure + else + let* (_, get_mut_back) = hashmap_HashMap_get_mut u64 hm4 1024 in + let* hm5 = get_mut_back 56 in + let* i1 = hashmap_HashMap_get u64 hm5 1024 in + if not (i1 = 56) + then Fail Failure + else + let* (x, hm6) = hashmap_HashMap_remove u64 hm5 1024 in + begin match x with + | None -> Fail Failure + | Some x1 -> + if not (x1 = 56) + then Fail Failure + else + let* i2 = hashmap_HashMap_get u64 hm6 0 in + if not (i2 = 42) + then Fail Failure + else + let* i3 = hashmap_HashMap_get u64 hm6 128 in + if not (i3 = 18) + then Fail Failure + else + let* i4 = hashmap_HashMap_get u64 hm6 1056 in + if not (i4 = 256) then Fail Failure else Ok () + end + +(** [hashmap_main::insert_on_disk]: + Source: 'tests/src/hashmap_main.rs', lines 13:0-13:43 *) +let insert_on_disk + (key : usize) (value : u64) (st : state) : result (state & unit) = + let* (st1, hm) = hashmap_utils_deserialize st in + let* hm1 = hashmap_HashMap_insert u64 hm key value in + hashmap_utils_serialize hm1 st1 + +(** [hashmap_main::main]: + Source: 'tests/src/hashmap_main.rs', lines 22:0-22:13 *) +let main : result unit = + Ok () + diff --git a/tests/fstar/hashmap_main/HashmapMain.FunsExternal.fsti b/tests/fstar/hashmap_main/HashmapMain.FunsExternal.fsti new file mode 100644 index 00000000..cc20d988 --- /dev/null +++ b/tests/fstar/hashmap_main/HashmapMain.FunsExternal.fsti @@ -0,0 +1,18 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: external function declarations *) +module HashmapMain.FunsExternal +open Primitives +include HashmapMain.Types + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** [hashmap_main::hashmap_utils::deserialize]: + Source: 'tests/src/hashmap_utils.rs', lines 11:0-11:43 *) +val hashmap_utils_deserialize + : state -> result (state & (hashmap_HashMap_t u64)) + +(** [hashmap_main::hashmap_utils::serialize]: + Source: 'tests/src/hashmap_utils.rs', lines 6:0-6:42 *) +val hashmap_utils_serialize + : hashmap_HashMap_t u64 -> state -> result (state & unit) + diff --git a/tests/fstar/hashmap_main/HashmapMain.Properties.fst b/tests/fstar/hashmap_main/HashmapMain.Properties.fst new file mode 100644 index 00000000..beb3dc2c --- /dev/null +++ b/tests/fstar/hashmap_main/HashmapMain.Properties.fst @@ -0,0 +1,48 @@ +(** Properties about the hashmap written on disk *) +module HashmapMain.Properties +open Primitives +open HashmapMain.Funs + +#set-options "--z3rlimit 50 --fuel 0 --ifuel 1" + +/// Below, we focus on the functions to read from disk/write to disk to showcase +/// how such reasoning which mixes opaque functions together with a state-error +/// monad can be performed. + +(*** Hypotheses *) + +/// [state_v] gives us the hash map currently stored on disk +assume +val state_v : state -> hashmap_HashMap_t u64 + +/// [serialize] updates the hash map stored on disk +assume +val serialize_lem (hm : hashmap_HashMap_t u64) (st : state) : Lemma ( + match hashmap_utils_serialize hm st with + | Fail _ -> True + | Ok (st', ()) -> state_v st' == hm) + [SMTPat (hashmap_utils_serialize hm st)] + +/// [deserialize] gives us the hash map stored on disk, without updating it +assume +val deserialize_lem (st : state) : Lemma ( + match hashmap_utils_deserialize st with + | Fail _ -> True + | Ok (st', hm) -> hm == state_v st /\ st' == st) + [SMTPat (hashmap_utils_deserialize st)] + +(*** Lemmas *) + +/// The obvious lemma about [insert_on_disk]: the updated hash map stored on disk +/// is exactly the hash map produced from inserting the binding ([key], [value]) +/// in the hash map previously stored on disk. +val insert_on_disk_lem (key : usize) (value : u64) (st : state) : Lemma ( + match insert_on_disk key value st with + | Fail _ -> True + | Ok (st', ()) -> + let hm = state_v st in + match hashmap_HashMap_insert u64 hm key value with + | Fail _ -> False + | Ok hm' -> hm' == state_v st') + +let insert_on_disk_lem key value st = () diff --git a/tests/fstar/hashmap_main/HashmapMain.Types.fst b/tests/fstar/hashmap_main/HashmapMain.Types.fst new file mode 100644 index 00000000..85bcaeea --- /dev/null +++ b/tests/fstar/hashmap_main/HashmapMain.Types.fst @@ -0,0 +1,24 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: type definitions *) +module HashmapMain.Types +open Primitives +include HashmapMain.TypesExternal + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** [hashmap_main::hashmap::List] + Source: 'tests/src/hashmap.rs', lines 27:0-27:16 *) +type hashmap_List_t (t : Type0) = +| Hashmap_List_Cons : usize -> t -> hashmap_List_t t -> hashmap_List_t t +| Hashmap_List_Nil : hashmap_List_t t + +(** [hashmap_main::hashmap::HashMap] + Source: 'tests/src/hashmap.rs', lines 43:0-43:21 *) +type hashmap_HashMap_t (t : Type0) = +{ + num_entries : usize; + max_load_factor : (usize & usize); + max_load : usize; + slots : alloc_vec_Vec (hashmap_List_t t); +} + diff --git a/tests/fstar/hashmap_main/HashmapMain.TypesExternal.fsti b/tests/fstar/hashmap_main/HashmapMain.TypesExternal.fsti new file mode 100644 index 00000000..75747408 --- /dev/null +++ b/tests/fstar/hashmap_main/HashmapMain.TypesExternal.fsti @@ -0,0 +1,10 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: external type declarations *) +module HashmapMain.TypesExternal +open Primitives + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** The state type used in the state-error monad *) +val state : Type0 + diff --git a/tests/fstar/hashmap_main/Makefile b/tests/fstar/hashmap_main/Makefile new file mode 100644 index 00000000..fa7d1f36 --- /dev/null +++ b/tests/fstar/hashmap_main/Makefile @@ -0,0 +1,49 @@ +# This file was automatically generated - modify ../Makefile.template instead +INCLUDE_DIRS = . + +FSTAR_INCLUDES = $(addprefix --include ,$(INCLUDE_DIRS)) + +FSTAR_HINTS ?= --use_hints --use_hint_hashes --record_hints + +FSTAR_OPTIONS = $(FSTAR_HINTS) \ + --cache_checked_modules $(FSTAR_INCLUDES) --cmi \ + --warn_error '+241@247+285-274' \ + +FSTAR_EXE ?= fstar.exe +FSTAR_NO_FLAGS = $(FSTAR_EXE) --already_cached 'Prims FStar LowStar Steel' --odir obj --cache_dir obj + +FSTAR = $(FSTAR_NO_FLAGS) $(FSTAR_OPTIONS) + +# The F* roots are used to compute the dependency graph, and generate the .depend file +FSTAR_ROOTS ?= $(wildcard *.fst *.fsti) + +# Build all the files +all: $(addprefix obj/,$(addsuffix .checked,$(FSTAR_ROOTS))) + +# This is the right way to ensure the .depend file always gets re-built. +ifeq (,$(filter %-in,$(MAKECMDGOALS))) +ifndef NODEPEND +ifndef MAKE_RESTARTS +.depend: .FORCE + $(FSTAR_NO_FLAGS) --dep full $(notdir $(FSTAR_ROOTS)) > $@ + +.PHONY: .FORCE +.FORCE: +endif +endif + +include .depend +endif + +# For the interactive mode +%.fst-in %.fsti-in: + @echo $(FSTAR_OPTIONS) + +# Generete the .checked files in batch mode +%.checked: + $(FSTAR) $(FSTAR_OPTIONS) $< && \ + touch -c $@ + +.PHONY: clean +clean: + rm -f obj/* diff --git a/tests/fstar/hashmap_main/Primitives.fst b/tests/fstar/hashmap_main/Primitives.fst new file mode 100644 index 00000000..9951ccc3 --- /dev/null +++ b/tests/fstar/hashmap_main/Primitives.fst @@ -0,0 +1,929 @@ +/// This file lists primitive and assumed functions and types +module Primitives +open FStar.Mul +open FStar.List.Tot + +#set-options "--z3rlimit 15 --fuel 0 --ifuel 1" + +(*** Utilities *) +val list_update (#a : Type0) (ls : list a) (i : nat{i < length ls}) (x : a) : + ls':list a{ + length ls' = length ls /\ + index ls' i == x + } +#push-options "--fuel 1" +let rec list_update #a ls i x = + match ls with + | x' :: ls -> if i = 0 then x :: ls else x' :: list_update ls (i-1) x +#pop-options + +(*** Result *) +type error : Type0 = +| Failure +| OutOfFuel + +type result (a : Type0) : Type0 = +| Ok : v:a -> result a +| Fail : e:error -> result a + +// Monadic return operator +unfold let return (#a : Type0) (x : a) : result a = Ok x + +// Monadic bind operator. +// Allows to use the notation: +// ``` +// let* x = y in +// ... +// ``` +unfold let (let*) (#a #b : Type0) (m: result a) + (f: (x:a) -> Pure (result b) (requires (m == Ok x)) (ensures fun _ -> True)) : + result b = + match m with + | Ok x -> f x + | Fail e -> Fail e + +// Monadic assert(...) +let massert (b:bool) : result unit = if b then Ok () else Fail Failure + +// Normalize and unwrap a successful result (used for globals). +let eval_global (#a : Type0) (x : result a{Ok? (normalize_term x)}) : a = Ok?.v x + +(*** Misc *) +type char = FStar.Char.char +type string = string + +let is_zero (n: nat) : bool = n = 0 +let decrease (n: nat{n > 0}) : nat = n - 1 + +let core_mem_replace (a : Type0) (x : a) (y : a) : a & a = (x, x) + +// We don't really use raw pointers for now +type mut_raw_ptr (t : Type0) = { v : t } +type const_raw_ptr (t : Type0) = { v : t } + +(*** Scalars *) +/// Rem.: most of the following code was partially generated + +assume val size_numbits : pos + +// TODO: we could use FStar.Int.int_t and FStar.UInt.int_t + +let isize_min : int = -9223372036854775808 // TODO: should be opaque +let isize_max : int = 9223372036854775807 // TODO: should be opaque +let i8_min : int = -128 +let i8_max : int = 127 +let i16_min : int = -32768 +let i16_max : int = 32767 +let i32_min : int = -2147483648 +let i32_max : int = 2147483647 +let i64_min : int = -9223372036854775808 +let i64_max : int = 9223372036854775807 +let i128_min : int = -170141183460469231731687303715884105728 +let i128_max : int = 170141183460469231731687303715884105727 +let usize_min : int = 0 +let usize_max : int = 4294967295 // TODO: should be opaque +let u8_min : int = 0 +let u8_max : int = 255 +let u16_min : int = 0 +let u16_max : int = 65535 +let u32_min : int = 0 +let u32_max : int = 4294967295 +let u64_min : int = 0 +let u64_max : int = 18446744073709551615 +let u128_min : int = 0 +let u128_max : int = 340282366920938463463374607431768211455 + +type scalar_ty = +| Isize +| I8 +| I16 +| I32 +| I64 +| I128 +| Usize +| U8 +| U16 +| U32 +| U64 +| U128 + +let is_unsigned = function + | Isize | I8 | I16 | I32 | I64 | I128 -> false + | Usize | U8 | U16 | U32 | U64 | U128 -> true + +let scalar_min (ty : scalar_ty) : int = + match ty with + | Isize -> isize_min + | I8 -> i8_min + | I16 -> i16_min + | I32 -> i32_min + | I64 -> i64_min + | I128 -> i128_min + | Usize -> usize_min + | U8 -> u8_min + | U16 -> u16_min + | U32 -> u32_min + | U64 -> u64_min + | U128 -> u128_min + +let scalar_max (ty : scalar_ty) : int = + match ty with + | Isize -> isize_max + | I8 -> i8_max + | I16 -> i16_max + | I32 -> i32_max + | I64 -> i64_max + | I128 -> i128_max + | Usize -> usize_max + | U8 -> u8_max + | U16 -> u16_max + | U32 -> u32_max + | U64 -> u64_max + | U128 -> u128_max + +type scalar (ty : scalar_ty) : eqtype = x:int{scalar_min ty <= x && x <= scalar_max ty} + +let mk_scalar (ty : scalar_ty) (x : int) : result (scalar ty) = + if scalar_min ty <= x && scalar_max ty >= x then Ok x else Fail Failure + +let scalar_neg (#ty : scalar_ty) (x : scalar ty) : result (scalar ty) = mk_scalar ty (-x) + +let scalar_div (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = + if y <> 0 then mk_scalar ty (x / y) else Fail Failure + +/// The remainder operation +let int_rem (x : int) (y : int{y <> 0}) : int = + if x >= 0 then (x % y) else -(x % y) + +(* Checking consistency with Rust *) +let _ = assert_norm(int_rem 1 2 = 1) +let _ = assert_norm(int_rem (-1) 2 = -1) +let _ = assert_norm(int_rem 1 (-2) = 1) +let _ = assert_norm(int_rem (-1) (-2) = -1) + +let scalar_rem (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = + if y <> 0 then mk_scalar ty (int_rem x y) else Fail Failure + +let scalar_add (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = + mk_scalar ty (x + y) + +let scalar_sub (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = + mk_scalar ty (x - y) + +let scalar_mul (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = + mk_scalar ty (x * y) + +let scalar_xor (#ty : scalar_ty) + (x : scalar ty) (y : scalar ty) : scalar ty = + match ty with + | U8 -> FStar.UInt.logxor #8 x y + | U16 -> FStar.UInt.logxor #16 x y + | U32 -> FStar.UInt.logxor #32 x y + | U64 -> FStar.UInt.logxor #64 x y + | U128 -> FStar.UInt.logxor #128 x y + | Usize -> admit() // TODO + | I8 -> + // Encoding issues... + normalize_spec (FStar.Int.int_t 8); + normalize_spec (scalar I8); + FStar.Int.logxor #8 x y + | I16 -> + // Encoding issues... + normalize_spec (FStar.Int.int_t 16); + normalize_spec (scalar I16); + FStar.Int.logxor #16 x y + | I32 -> FStar.Int.logxor #32 x y + | I64 -> FStar.Int.logxor #64 x y + | I128 -> + // Encoding issues... + normalize_spec (FStar.Int.int_t 128); + normalize_spec (scalar I128); + FStar.Int.logxor #128 x y + | Isize -> admit() // TODO + +let scalar_or (#ty : scalar_ty) + (x : scalar ty) (y : scalar ty) : scalar ty = + match ty with + | U8 -> FStar.UInt.logor #8 x y + | U16 -> FStar.UInt.logor #16 x y + | U32 -> FStar.UInt.logor #32 x y + | U64 -> FStar.UInt.logor #64 x y + | U128 -> FStar.UInt.logor #128 x y + | Usize -> admit() // TODO + | I8 -> + // Encoding issues... + normalize_spec (FStar.Int.int_t 8); + normalize_spec (scalar I8); + FStar.Int.logor #8 x y + | I16 -> + // Encoding issues... + normalize_spec (FStar.Int.int_t 16); + normalize_spec (scalar I16); + FStar.Int.logor #16 x y + | I32 -> FStar.Int.logor #32 x y + | I64 -> FStar.Int.logor #64 x y + | I128 -> + // Encoding issues... + normalize_spec (FStar.Int.int_t 128); + normalize_spec (scalar I128); + FStar.Int.logor #128 x y + | Isize -> admit() // TODO + +let scalar_and (#ty : scalar_ty) + (x : scalar ty) (y : scalar ty) : scalar ty = + match ty with + | U8 -> FStar.UInt.logand #8 x y + | U16 -> FStar.UInt.logand #16 x y + | U32 -> FStar.UInt.logand #32 x y + | U64 -> FStar.UInt.logand #64 x y + | U128 -> FStar.UInt.logand #128 x y + | Usize -> admit() // TODO + | I8 -> + // Encoding issues... + normalize_spec (FStar.Int.int_t 8); + normalize_spec (scalar I8); + FStar.Int.logand #8 x y + | I16 -> + // Encoding issues... + normalize_spec (FStar.Int.int_t 16); + normalize_spec (scalar I16); + FStar.Int.logand #16 x y + | I32 -> FStar.Int.logand #32 x y + | I64 -> FStar.Int.logand #64 x y + | I128 -> + // Encoding issues... + normalize_spec (FStar.Int.int_t 128); + normalize_spec (scalar I128); + FStar.Int.logand #128 x y + | Isize -> admit() // TODO + +// Shift left +let scalar_shl (#ty0 #ty1 : scalar_ty) + (x : scalar ty0) (y : scalar ty1) : result (scalar ty0) = + admit() + +// Shift right +let scalar_shr (#ty0 #ty1 : scalar_ty) + (x : scalar ty0) (y : scalar ty1) : result (scalar ty0) = + admit() + +(** Cast an integer from a [src_ty] to a [tgt_ty] *) +// TODO: check the semantics of casts in Rust +let scalar_cast (src_ty : scalar_ty) (tgt_ty : scalar_ty) (x : scalar src_ty) : result (scalar tgt_ty) = + mk_scalar tgt_ty x + +// This can't fail, but for now we make all casts faillible (easier for the translation) +let scalar_cast_bool (tgt_ty : scalar_ty) (x : bool) : result (scalar tgt_ty) = + mk_scalar tgt_ty (if x then 1 else 0) + +/// The scalar types +type isize : eqtype = scalar Isize +type i8 : eqtype = scalar I8 +type i16 : eqtype = scalar I16 +type i32 : eqtype = scalar I32 +type i64 : eqtype = scalar I64 +type i128 : eqtype = scalar I128 +type usize : eqtype = scalar Usize +type u8 : eqtype = scalar U8 +type u16 : eqtype = scalar U16 +type u32 : eqtype = scalar U32 +type u64 : eqtype = scalar U64 +type u128 : eqtype = scalar U128 + + +let core_isize_min : isize = isize_min +let core_isize_max : isize = isize_max +let core_i8_min : i8 = i8_min +let core_i8_max : i8 = i8_max +let core_i16_min : i16 = i16_min +let core_i16_max : i16 = i16_max +let core_i32_min : i32 = i32_min +let core_i32_max : i32 = i32_max +let core_i64_min : i64 = i64_min +let core_i64_max : i64 = i64_max +let core_i128_min : i128 = i128_min +let core_i128_max : i128 = i128_max + +let core_usize_min : usize = usize_min +let core_usize_max : usize = usize_max +let core_u8_min : u8 = u8_min +let core_u8_max : u8 = u8_max +let core_u16_min : u16 = u16_min +let core_u16_max : u16 = u16_max +let core_u32_min : u32 = u32_min +let core_u32_max : u32 = u32_max +let core_u64_min : u64 = u64_min +let core_u64_max : u64 = u64_max +let core_u128_min : u128 = u128_min +let core_u128_max : u128 = u128_max + +/// Negation +let isize_neg = scalar_neg #Isize +let i8_neg = scalar_neg #I8 +let i16_neg = scalar_neg #I16 +let i32_neg = scalar_neg #I32 +let i64_neg = scalar_neg #I64 +let i128_neg = scalar_neg #I128 + +/// Division +let isize_div = scalar_div #Isize +let i8_div = scalar_div #I8 +let i16_div = scalar_div #I16 +let i32_div = scalar_div #I32 +let i64_div = scalar_div #I64 +let i128_div = scalar_div #I128 +let usize_div = scalar_div #Usize +let u8_div = scalar_div #U8 +let u16_div = scalar_div #U16 +let u32_div = scalar_div #U32 +let u64_div = scalar_div #U64 +let u128_div = scalar_div #U128 + +/// Remainder +let isize_rem = scalar_rem #Isize +let i8_rem = scalar_rem #I8 +let i16_rem = scalar_rem #I16 +let i32_rem = scalar_rem #I32 +let i64_rem = scalar_rem #I64 +let i128_rem = scalar_rem #I128 +let usize_rem = scalar_rem #Usize +let u8_rem = scalar_rem #U8 +let u16_rem = scalar_rem #U16 +let u32_rem = scalar_rem #U32 +let u64_rem = scalar_rem #U64 +let u128_rem = scalar_rem #U128 + +/// Addition +let isize_add = scalar_add #Isize +let i8_add = scalar_add #I8 +let i16_add = scalar_add #I16 +let i32_add = scalar_add #I32 +let i64_add = scalar_add #I64 +let i128_add = scalar_add #I128 +let usize_add = scalar_add #Usize +let u8_add = scalar_add #U8 +let u16_add = scalar_add #U16 +let u32_add = scalar_add #U32 +let u64_add = scalar_add #U64 +let u128_add = scalar_add #U128 + +/// Subtraction +let isize_sub = scalar_sub #Isize +let i8_sub = scalar_sub #I8 +let i16_sub = scalar_sub #I16 +let i32_sub = scalar_sub #I32 +let i64_sub = scalar_sub #I64 +let i128_sub = scalar_sub #I128 +let usize_sub = scalar_sub #Usize +let u8_sub = scalar_sub #U8 +let u16_sub = scalar_sub #U16 +let u32_sub = scalar_sub #U32 +let u64_sub = scalar_sub #U64 +let u128_sub = scalar_sub #U128 + +/// Multiplication +let isize_mul = scalar_mul #Isize +let i8_mul = scalar_mul #I8 +let i16_mul = scalar_mul #I16 +let i32_mul = scalar_mul #I32 +let i64_mul = scalar_mul #I64 +let i128_mul = scalar_mul #I128 +let usize_mul = scalar_mul #Usize +let u8_mul = scalar_mul #U8 +let u16_mul = scalar_mul #U16 +let u32_mul = scalar_mul #U32 +let u64_mul = scalar_mul #U64 +let u128_mul = scalar_mul #U128 + +/// Xor +let u8_xor = scalar_xor #U8 +let u16_xor = scalar_xor #U16 +let u32_xor = scalar_xor #U32 +let u64_xor = scalar_xor #U64 +let u128_xor = scalar_xor #U128 +let usize_xor = scalar_xor #Usize +let i8_xor = scalar_xor #I8 +let i16_xor = scalar_xor #I16 +let i32_xor = scalar_xor #I32 +let i64_xor = scalar_xor #I64 +let i128_xor = scalar_xor #I128 +let isize_xor = scalar_xor #Isize + +/// Or +let u8_or = scalar_or #U8 +let u16_or = scalar_or #U16 +let u32_or = scalar_or #U32 +let u64_or = scalar_or #U64 +let u128_or = scalar_or #U128 +let usize_or = scalar_or #Usize +let i8_or = scalar_or #I8 +let i16_or = scalar_or #I16 +let i32_or = scalar_or #I32 +let i64_or = scalar_or #I64 +let i128_or = scalar_or #I128 +let isize_or = scalar_or #Isize + +/// And +let u8_and = scalar_and #U8 +let u16_and = scalar_and #U16 +let u32_and = scalar_and #U32 +let u64_and = scalar_and #U64 +let u128_and = scalar_and #U128 +let usize_and = scalar_and #Usize +let i8_and = scalar_and #I8 +let i16_and = scalar_and #I16 +let i32_and = scalar_and #I32 +let i64_and = scalar_and #I64 +let i128_and = scalar_and #I128 +let isize_and = scalar_and #Isize + +/// Shift left +let u8_shl #ty = scalar_shl #U8 #ty +let u16_shl #ty = scalar_shl #U16 #ty +let u32_shl #ty = scalar_shl #U32 #ty +let u64_shl #ty = scalar_shl #U64 #ty +let u128_shl #ty = scalar_shl #U128 #ty +let usize_shl #ty = scalar_shl #Usize #ty +let i8_shl #ty = scalar_shl #I8 #ty +let i16_shl #ty = scalar_shl #I16 #ty +let i32_shl #ty = scalar_shl #I32 #ty +let i64_shl #ty = scalar_shl #I64 #ty +let i128_shl #ty = scalar_shl #I128 #ty +let isize_shl #ty = scalar_shl #Isize #ty + +/// Shift right +let u8_shr #ty = scalar_shr #U8 #ty +let u16_shr #ty = scalar_shr #U16 #ty +let u32_shr #ty = scalar_shr #U32 #ty +let u64_shr #ty = scalar_shr #U64 #ty +let u128_shr #ty = scalar_shr #U128 #ty +let usize_shr #ty = scalar_shr #Usize #ty +let i8_shr #ty = scalar_shr #I8 #ty +let i16_shr #ty = scalar_shr #I16 #ty +let i32_shr #ty = scalar_shr #I32 #ty +let i64_shr #ty = scalar_shr #I64 #ty +let i128_shr #ty = scalar_shr #I128 #ty +let isize_shr #ty = scalar_shr #Isize #ty + +(*** core *) + +/// Trait declaration: [core::clone::Clone] +noeq type core_clone_Clone (self : Type0) = { + clone : self → result self +} + +let core_clone_impls_CloneBool_clone (b : bool) : bool = b + +let core_clone_CloneBool : core_clone_Clone bool = { + clone = fun b -> Ok (core_clone_impls_CloneBool_clone b) +} + +let core_clone_impls_CloneUsize_clone (x : usize) : usize = x +let core_clone_impls_CloneU8_clone (x : u8) : u8 = x +let core_clone_impls_CloneU16_clone (x : u16) : u16 = x +let core_clone_impls_CloneU32_clone (x : u32) : u32 = x +let core_clone_impls_CloneU64_clone (x : u64) : u64 = x +let core_clone_impls_CloneU128_clone (x : u128) : u128 = x + +let core_clone_impls_CloneIsize_clone (x : isize) : isize = x +let core_clone_impls_CloneI8_clone (x : i8) : i8 = x +let core_clone_impls_CloneI16_clone (x : i16) : i16 = x +let core_clone_impls_CloneI32_clone (x : i32) : i32 = x +let core_clone_impls_CloneI64_clone (x : i64) : i64 = x +let core_clone_impls_CloneI128_clone (x : i128) : i128 = x + +let core_clone_CloneUsize : core_clone_Clone usize = { + clone = fun x -> Ok (core_clone_impls_CloneUsize_clone x) +} + +let core_clone_CloneU8 : core_clone_Clone u8 = { + clone = fun x -> Ok (core_clone_impls_CloneU8_clone x) +} + +let core_clone_CloneU16 : core_clone_Clone u16 = { + clone = fun x -> Ok (core_clone_impls_CloneU16_clone x) +} + +let core_clone_CloneU32 : core_clone_Clone u32 = { + clone = fun x -> Ok (core_clone_impls_CloneU32_clone x) +} + +let core_clone_CloneU64 : core_clone_Clone u64 = { + clone = fun x -> Ok (core_clone_impls_CloneU64_clone x) +} + +let core_clone_CloneU128 : core_clone_Clone u128 = { + clone = fun x -> Ok (core_clone_impls_CloneU128_clone x) +} + +let core_clone_CloneIsize : core_clone_Clone isize = { + clone = fun x -> Ok (core_clone_impls_CloneIsize_clone x) +} + +let core_clone_CloneI8 : core_clone_Clone i8 = { + clone = fun x -> Ok (core_clone_impls_CloneI8_clone x) +} + +let core_clone_CloneI16 : core_clone_Clone i16 = { + clone = fun x -> Ok (core_clone_impls_CloneI16_clone x) +} + +let core_clone_CloneI32 : core_clone_Clone i32 = { + clone = fun x -> Ok (core_clone_impls_CloneI32_clone x) +} + +let core_clone_CloneI64 : core_clone_Clone i64 = { + clone = fun x -> Ok (core_clone_impls_CloneI64_clone x) +} + +let core_clone_CloneI128 : core_clone_Clone i128 = { + clone = fun x -> Ok (core_clone_impls_CloneI128_clone x) +} + +(** [core::option::{core::option::Option}::unwrap] *) +let core_option_Option_unwrap (t : Type0) (x : option t) : result t = + match x with + | None -> Fail Failure + | Some x -> Ok x + +(*** core::ops *) + +// Trait declaration: [core::ops::index::Index] +noeq type core_ops_index_Index (self idx : Type0) = { + output : Type0; + index : self → idx → result output +} + +// Trait declaration: [core::ops::index::IndexMut] +noeq type core_ops_index_IndexMut (self idx : Type0) = { + indexInst : core_ops_index_Index self idx; + index_mut : self → idx → result (indexInst.output & (indexInst.output → result self)); +} + +// Trait declaration [core::ops::deref::Deref] +noeq type core_ops_deref_Deref (self : Type0) = { + target : Type0; + deref : self → result target; +} + +// Trait declaration [core::ops::deref::DerefMut] +noeq type core_ops_deref_DerefMut (self : Type0) = { + derefInst : core_ops_deref_Deref self; + deref_mut : self → result (derefInst.target & (derefInst.target → result self)); +} + +type core_ops_range_Range (a : Type0) = { + start : a; + end_ : a; +} + +(*** [alloc] *) + +let alloc_boxed_Box_deref (t : Type0) (x : t) : result t = Ok x +let alloc_boxed_Box_deref_mut (t : Type0) (x : t) : result (t & (t -> result t)) = + Ok (x, (fun x -> Ok x)) + +// Trait instance +let alloc_boxed_Box_coreopsDerefInst (self : Type0) : core_ops_deref_Deref self = { + target = self; + deref = alloc_boxed_Box_deref self; +} + +// Trait instance +let alloc_boxed_Box_coreopsDerefMutInst (self : Type0) : core_ops_deref_DerefMut self = { + derefInst = alloc_boxed_Box_coreopsDerefInst self; + deref_mut = alloc_boxed_Box_deref_mut self; +} + +(*** Array *) +type array (a : Type0) (n : usize) = s:list a{length s = n} + +// We tried putting the normalize_term condition as a refinement on the list +// but it didn't work. It works with the requires clause. +let mk_array (a : Type0) (n : usize) + (l : list a) : + Pure (array a n) + (requires (normalize_term(FStar.List.Tot.length l) = n)) + (ensures (fun _ -> True)) = + normalize_term_spec (FStar.List.Tot.length l); + l + +let array_index_usize (a : Type0) (n : usize) (x : array a n) (i : usize) : result a = + if i < length x then Ok (index x i) + else Fail Failure + +let array_update_usize (a : Type0) (n : usize) (x : array a n) (i : usize) (nx : a) : + result (array a n) = + if i < length x then Ok (list_update x i nx) + else Fail Failure + +let array_index_mut_usize (a : Type0) (n : usize) (x : array a n) (i : usize) : + result (a & (a -> result (array a n))) = + match array_index_usize a n x i with + | Fail e -> Fail e + | Ok v -> + Ok (v, array_update_usize a n x i) + +(*** Slice *) +type slice (a : Type0) = s:list a{length s <= usize_max} + +let slice_len (a : Type0) (s : slice a) : usize = length s + +let slice_index_usize (a : Type0) (x : slice a) (i : usize) : result a = + if i < length x then Ok (index x i) + else Fail Failure + +let slice_update_usize (a : Type0) (x : slice a) (i : usize) (nx : a) : result (slice a) = + if i < length x then Ok (list_update x i nx) + else Fail Failure + +let slice_index_mut_usize (a : Type0) (s : slice a) (i : usize) : + result (a & (a -> result (slice a))) = + match slice_index_usize a s i with + | Fail e -> Fail e + | Ok x -> + Ok (x, slice_update_usize a s i) + +(*** Subslices *) + +let array_to_slice (a : Type0) (n : usize) (x : array a n) : result (slice a) = Ok x +let array_from_slice (a : Type0) (n : usize) (x : array a n) (s : slice a) : result (array a n) = + if length s = n then Ok s + else Fail Failure + +let array_to_slice_mut (a : Type0) (n : usize) (x : array a n) : + result (slice a & (slice a -> result (array a n))) = + Ok (x, array_from_slice a n x) + +// TODO: finish the definitions below (there lacks [List.drop] and [List.take] in the standard library *) +let array_subslice (a : Type0) (n : usize) (x : array a n) (r : core_ops_range_Range usize) : result (slice a) = + admit() + +let array_update_subslice (a : Type0) (n : usize) (x : array a n) (r : core_ops_range_Range usize) (ns : slice a) : result (array a n) = + admit() + +let array_repeat (a : Type0) (n : usize) (x : a) : array a n = + admit() + +let slice_subslice (a : Type0) (x : slice a) (r : core_ops_range_Range usize) : result (slice a) = + admit() + +let slice_update_subslice (a : Type0) (x : slice a) (r : core_ops_range_Range usize) (ns : slice a) : result (slice a) = + admit() + +(*** Vector *) +type alloc_vec_Vec (a : Type0) = v:list a{length v <= usize_max} + +let alloc_vec_Vec_new (a : Type0) : alloc_vec_Vec a = assert_norm(length #a [] == 0); [] +let alloc_vec_Vec_len (a : Type0) (v : alloc_vec_Vec a) : usize = length v + +// Helper +let alloc_vec_Vec_index_usize (#a : Type0) (v : alloc_vec_Vec a) (i : usize) : result a = + if i < length v then Ok (index v i) else Fail Failure +// Helper +let alloc_vec_Vec_update_usize (#a : Type0) (v : alloc_vec_Vec a) (i : usize) (x : a) : result (alloc_vec_Vec a) = + if i < length v then Ok (list_update v i x) else Fail Failure + +let alloc_vec_Vec_index_mut_usize (#a : Type0) (v: alloc_vec_Vec a) (i: usize) : + result (a & (a → result (alloc_vec_Vec a))) = + match alloc_vec_Vec_index_usize v i with + | Ok x -> + Ok (x, alloc_vec_Vec_update_usize v i) + | Fail e -> Fail e + +let alloc_vec_Vec_push (a : Type0) (v : alloc_vec_Vec a) (x : a) : + Pure (result (alloc_vec_Vec a)) + (requires True) + (ensures (fun res -> + match res with + | Fail e -> e == Failure + | Ok v' -> length v' = length v + 1)) = + if length v < usize_max then begin + (**) assert_norm(length [x] == 1); + (**) append_length v [x]; + (**) assert(length (append v [x]) = length v + 1); + Ok (append v [x]) + end + else Fail Failure + +let alloc_vec_Vec_insert (a : Type0) (v : alloc_vec_Vec a) (i : usize) (x : a) : result (alloc_vec_Vec a) = + if i < length v then Ok (list_update v i x) else Fail Failure + +// Trait declaration: [core::slice::index::private_slice_index::Sealed] +type core_slice_index_private_slice_index_Sealed (self : Type0) = unit + +// Trait declaration: [core::slice::index::SliceIndex] +noeq type core_slice_index_SliceIndex (self t : Type0) = { + sealedInst : core_slice_index_private_slice_index_Sealed self; + output : Type0; + get : self → t → result (option output); + get_mut : self → t → result (option output & (option output -> result t)); + get_unchecked : self → const_raw_ptr t → result (const_raw_ptr output); + get_unchecked_mut : self → mut_raw_ptr t → result (mut_raw_ptr output); + index : self → t → result output; + index_mut : self → t → result (output & (output -> result t)); +} + +// [core::slice::index::[T]::index]: forward function +let core_slice_index_Slice_index + (t idx : Type0) (inst : core_slice_index_SliceIndex idx (slice t)) + (s : slice t) (i : idx) : result inst.output = + let* x = inst.get i s in + match x with + | None -> Fail Failure + | Some x -> Ok x + +// [core::slice::index::Range:::get]: forward function +let core_slice_index_RangeUsize_get (t : Type0) (i : core_ops_range_Range usize) (s : slice t) : + result (option (slice t)) = + admit () // TODO + +// [core::slice::index::Range::get_mut]: forward function +let core_slice_index_RangeUsize_get_mut (t : Type0) : + core_ops_range_Range usize → slice t → result (option (slice t) & (option (slice t) -> result (slice t))) = + admit () // TODO + +// [core::slice::index::Range::get_unchecked]: forward function +let core_slice_index_RangeUsize_get_unchecked + (t : Type0) : + core_ops_range_Range usize → const_raw_ptr (slice t) → result (const_raw_ptr (slice t)) = + // Don't know what the model should be - for now we always fail to make + // sure code which uses it fails + fun _ _ -> Fail Failure + +// [core::slice::index::Range::get_unchecked_mut]: forward function +let core_slice_index_RangeUsize_get_unchecked_mut + (t : Type0) : + core_ops_range_Range usize → mut_raw_ptr (slice t) → result (mut_raw_ptr (slice t)) = + // Don't know what the model should be - for now we always fail to make + // sure code which uses it fails + fun _ _ -> Fail Failure + +// [core::slice::index::Range::index]: forward function +let core_slice_index_RangeUsize_index + (t : Type0) : core_ops_range_Range usize → slice t → result (slice t) = + admit () // TODO + +// [core::slice::index::Range::index_mut]: forward function +let core_slice_index_RangeUsize_index_mut (t : Type0) : + core_ops_range_Range usize → slice t → result (slice t & (slice t -> result (slice t))) = + admit () // TODO + +// [core::slice::index::[T]::index_mut]: forward function +let core_slice_index_Slice_index_mut + (t idx : Type0) (inst : core_slice_index_SliceIndex idx (slice t)) : + slice t → idx → result (inst.output & (inst.output -> result (slice t))) = + admit () // + +// [core::array::[T; N]::index]: forward function +let core_array_Array_index + (t idx : Type0) (n : usize) (inst : core_ops_index_Index (slice t) idx) + (a : array t n) (i : idx) : result inst.output = + admit () // TODO + +// [core::array::[T; N]::index_mut]: forward function +let core_array_Array_index_mut + (t idx : Type0) (n : usize) (inst : core_ops_index_IndexMut (slice t) idx) + (a : array t n) (i : idx) : + result (inst.indexInst.output & (inst.indexInst.output -> result (array t n))) = + admit () // TODO + +// Trait implementation: [core::slice::index::private_slice_index::Range] +let core_slice_index_private_slice_index_SealedRangeUsizeInst + : core_slice_index_private_slice_index_Sealed (core_ops_range_Range usize) = () + +// Trait implementation: [core::slice::index::Range] +let core_slice_index_SliceIndexRangeUsizeSliceTInst (t : Type0) : + core_slice_index_SliceIndex (core_ops_range_Range usize) (slice t) = { + sealedInst = core_slice_index_private_slice_index_SealedRangeUsizeInst; + output = slice t; + get = core_slice_index_RangeUsize_get t; + get_mut = core_slice_index_RangeUsize_get_mut t; + get_unchecked = core_slice_index_RangeUsize_get_unchecked t; + get_unchecked_mut = core_slice_index_RangeUsize_get_unchecked_mut t; + index = core_slice_index_RangeUsize_index t; + index_mut = core_slice_index_RangeUsize_index_mut t; +} + +// Trait implementation: [core::slice::index::[T]] +let core_ops_index_IndexSliceTIInst (t idx : Type0) + (inst : core_slice_index_SliceIndex idx (slice t)) : + core_ops_index_Index (slice t) idx = { + output = inst.output; + index = core_slice_index_Slice_index t idx inst; +} + +// Trait implementation: [core::slice::index::[T]] +let core_ops_index_IndexMutSliceTIInst (t idx : Type0) + (inst : core_slice_index_SliceIndex idx (slice t)) : + core_ops_index_IndexMut (slice t) idx = { + indexInst = core_ops_index_IndexSliceTIInst t idx inst; + index_mut = core_slice_index_Slice_index_mut t idx inst; +} + +// Trait implementation: [core::array::[T; N]] +let core_ops_index_IndexArrayInst (t idx : Type0) (n : usize) + (inst : core_ops_index_Index (slice t) idx) : + core_ops_index_Index (array t n) idx = { + output = inst.output; + index = core_array_Array_index t idx n inst; +} + +// Trait implementation: [core::array::[T; N]] +let core_ops_index_IndexMutArrayIInst (t idx : Type0) (n : usize) + (inst : core_ops_index_IndexMut (slice t) idx) : + core_ops_index_IndexMut (array t n) idx = { + indexInst = core_ops_index_IndexArrayInst t idx n inst.indexInst; + index_mut = core_array_Array_index_mut t idx n inst; +} + +// [core::slice::index::usize::get]: forward function +let core_slice_index_usize_get + (t : Type0) : usize → slice t → result (option t) = + admit () // TODO + +// [core::slice::index::usize::get_mut]: forward function +let core_slice_index_usize_get_mut (t : Type0) : + usize → slice t → result (option t & (option t -> result (slice t))) = + admit () // TODO + +// [core::slice::index::usize::get_unchecked]: forward function +let core_slice_index_usize_get_unchecked + (t : Type0) : usize → const_raw_ptr (slice t) → result (const_raw_ptr t) = + admit () // TODO + +// [core::slice::index::usize::get_unchecked_mut]: forward function +let core_slice_index_usize_get_unchecked_mut + (t : Type0) : usize → mut_raw_ptr (slice t) → result (mut_raw_ptr t) = + admit () // TODO + +// [core::slice::index::usize::index]: forward function +let core_slice_index_usize_index (t : Type0) : usize → slice t → result t = + admit () // TODO + +// [core::slice::index::usize::index_mut]: forward function +let core_slice_index_usize_index_mut (t : Type0) : + usize → slice t → result (t & (t -> result (slice t))) = + admit () // TODO + +// Trait implementation: [core::slice::index::private_slice_index::usize] +let core_slice_index_private_slice_index_SealedUsizeInst + : core_slice_index_private_slice_index_Sealed usize = () + +// Trait implementation: [core::slice::index::usize] +let core_slice_index_SliceIndexUsizeSliceTInst (t : Type0) : + core_slice_index_SliceIndex usize (slice t) = { + sealedInst = core_slice_index_private_slice_index_SealedUsizeInst; + output = t; + get = core_slice_index_usize_get t; + get_mut = core_slice_index_usize_get_mut t; + get_unchecked = core_slice_index_usize_get_unchecked t; + get_unchecked_mut = core_slice_index_usize_get_unchecked_mut t; + index = core_slice_index_usize_index t; + index_mut = core_slice_index_usize_index_mut t; +} + +// [alloc::vec::Vec::index]: forward function +let alloc_vec_Vec_index (t idx : Type0) (inst : core_slice_index_SliceIndex idx (slice t)) + (self : alloc_vec_Vec t) (i : idx) : result inst.output = + admit () // TODO + +// [alloc::vec::Vec::index_mut]: forward function +let alloc_vec_Vec_index_mut (t idx : Type0) (inst : core_slice_index_SliceIndex idx (slice t)) + (self : alloc_vec_Vec t) (i : idx) : + result (inst.output & (inst.output -> result (alloc_vec_Vec t))) = + admit () // TODO + +// Trait implementation: [alloc::vec::Vec] +let alloc_vec_Vec_coreopsindexIndexInst (t idx : Type0) + (inst : core_slice_index_SliceIndex idx (slice t)) : + core_ops_index_Index (alloc_vec_Vec t) idx = { + output = inst.output; + index = alloc_vec_Vec_index t idx inst; +} + +// Trait implementation: [alloc::vec::Vec] +let alloc_vec_Vec_coreopsindexIndexMutInst (t idx : Type0) + (inst : core_slice_index_SliceIndex idx (slice t)) : + core_ops_index_IndexMut (alloc_vec_Vec t) idx = { + indexInst = alloc_vec_Vec_coreopsindexIndexInst t idx inst; + index_mut = alloc_vec_Vec_index_mut t idx inst; +} + +(*** Theorems *) + +let alloc_vec_Vec_index_eq (#a : Type0) (v : alloc_vec_Vec a) (i : usize) : + Lemma ( + alloc_vec_Vec_index a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i == + alloc_vec_Vec_index_usize v i) + [SMTPat (alloc_vec_Vec_index a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i)] + = + admit() + +let alloc_vec_Vec_index_mut_eq (#a : Type0) (v : alloc_vec_Vec a) (i : usize) : + Lemma ( + alloc_vec_Vec_index_mut a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i == + alloc_vec_Vec_index_mut_usize v i) + [SMTPat (alloc_vec_Vec_index_mut a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i)] + = + admit() diff --git a/tests/fstar/hashmap_on_disk/HashmapMain.Clauses.Template.fst b/tests/fstar/hashmap_on_disk/HashmapMain.Clauses.Template.fst deleted file mode 100644 index cdd73210..00000000 --- a/tests/fstar/hashmap_on_disk/HashmapMain.Clauses.Template.fst +++ /dev/null @@ -1,72 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: templates for the decreases clauses *) -module HashmapMain.Clauses.Template -open Primitives -open HashmapMain.Types - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: decreases clause - Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) -unfold -let hashmap_HashMap_allocate_slots_loop_decreases (t : Type0) - (slots : alloc_vec_Vec (hashmap_List_t t)) (n : usize) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: decreases clause - Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) -unfold -let hashmap_HashMap_clear_loop_decreases (t : Type0) - (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) -unfold -let hashmap_HashMap_insert_in_list_loop_decreases (t : Type0) (key : usize) - (value : t) (ls : hashmap_List_t t) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) -unfold -let hashmap_HashMap_move_elements_from_list_loop_decreases (t : Type0) - (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: decreases clause - Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) -unfold -let hashmap_HashMap_move_elements_loop_decreases (t : Type0) - (ntable : hashmap_HashMap_t t) (slots : alloc_vec_Vec (hashmap_List_t t)) - (i : usize) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) -unfold -let hashmap_HashMap_contains_key_in_list_loop_decreases (t : Type0) - (key : usize) (ls : hashmap_List_t t) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) -unfold -let hashmap_HashMap_get_in_list_loop_decreases (t : Type0) (key : usize) - (ls : hashmap_List_t t) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) -unfold -let hashmap_HashMap_get_mut_in_list_loop_decreases (t : Type0) - (ls : hashmap_List_t t) (key : usize) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) -unfold -let hashmap_HashMap_remove_from_list_loop_decreases (t : Type0) (key : usize) - (ls : hashmap_List_t t) : nat = - admit () - diff --git a/tests/fstar/hashmap_on_disk/HashmapMain.Clauses.fst b/tests/fstar/hashmap_on_disk/HashmapMain.Clauses.fst deleted file mode 100644 index be5a4ab1..00000000 --- a/tests/fstar/hashmap_on_disk/HashmapMain.Clauses.fst +++ /dev/null @@ -1,61 +0,0 @@ -(** [hashmap]: the decreases clauses *) -module HashmapMain.Clauses -open Primitives -open FStar.List.Tot -open HashmapMain.Types - -#set-options "--z3rlimit 50 --fuel 0 --ifuel 1" - -(** [hashmap::HashMap::allocate_slots]: decreases clause *) -unfold -let hashmap_HashMap_allocate_slots_loop_decreases (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) - (n : usize) : nat = n - -(** [hashmap::HashMap::clear]: decreases clause *) -unfold -let hashmap_HashMap_clear_loop_decreases (t : Type0) (slots : alloc_vec_Vec (hashmap_List_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_HashMap_insert_in_list_loop_decreases (t : Type0) (key : usize) (value : t) - (ls : hashmap_List_t t) : hashmap_List_t t = - ls - -(** [hashmap::HashMap::move_elements_from_list]: decreases clause *) -unfold -let hashmap_HashMap_move_elements_from_list_loop_decreases (t : Type0) - (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : hashmap_List_t t = - ls - -(** [hashmap::HashMap::move_elements]: decreases clause *) -unfold -let hashmap_HashMap_move_elements_loop_decreases (t : Type0) (ntable : hashmap_HashMap_t t) - (slots : alloc_vec_Vec (hashmap_List_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_HashMap_contains_key_in_list_loop_decreases (t : Type0) (key : usize) - (ls : hashmap_List_t t) : hashmap_List_t t = - ls - -(** [hashmap::HashMap::get_in_list]: decreases clause *) -unfold -let hashmap_HashMap_get_in_list_loop_decreases (t : Type0) (key : usize) (ls : hashmap_List_t t) : - hashmap_List_t t = - ls - -(** [hashmap::HashMap::get_mut_in_list]: decreases clause *) -unfold -let hashmap_HashMap_get_mut_in_list_loop_decreases (t : Type0) - (ls : hashmap_List_t t) (key : usize) : hashmap_List_t t = - ls - -(** [hashmap::HashMap::remove_from_list]: decreases clause *) -unfold -let hashmap_HashMap_remove_from_list_loop_decreases (t : Type0) (key : usize) - (ls : hashmap_List_t t) : hashmap_List_t t = - ls - diff --git a/tests/fstar/hashmap_on_disk/HashmapMain.Funs.fst b/tests/fstar/hashmap_on_disk/HashmapMain.Funs.fst deleted file mode 100644 index c88a746e..00000000 --- a/tests/fstar/hashmap_on_disk/HashmapMain.Funs.fst +++ /dev/null @@ -1,446 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: function definitions *) -module HashmapMain.Funs -open Primitives -include HashmapMain.Types -include HashmapMain.FunsExternal -include HashmapMain.Clauses - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** [hashmap_main::hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 35:0-35:32 *) -let hashmap_hash_key (k : usize) : result usize = - Ok k - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) -let rec hashmap_HashMap_allocate_slots_loop - (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (n : usize) : - Tot (result (alloc_vec_Vec (hashmap_List_t t))) - (decreases (hashmap_HashMap_allocate_slots_loop_decreases t slots n)) - = - if n > 0 - then - let* slots1 = alloc_vec_Vec_push (hashmap_List_t t) slots Hashmap_List_Nil - in - let* n1 = usize_sub n 1 in - hashmap_HashMap_allocate_slots_loop t slots1 n1 - else Ok slots - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 58:4-58:76 *) -let hashmap_HashMap_allocate_slots - (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (n : usize) : - result (alloc_vec_Vec (hashmap_List_t t)) - = - hashmap_HashMap_allocate_slots_loop t slots n - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 67:4-71:13 *) -let hashmap_HashMap_new_with_capacity - (t : Type0) (capacity : usize) (max_load_dividend : usize) - (max_load_divisor : usize) : - result (hashmap_HashMap_t t) - = - let* slots = - hashmap_HashMap_allocate_slots t (alloc_vec_Vec_new (hashmap_List_t t)) - capacity in - let* i = usize_mul capacity max_load_dividend in - let* i1 = usize_div i max_load_divisor in - Ok - { - num_entries = 0; - max_load_factor = (max_load_dividend, max_load_divisor); - max_load = i1; - slots = slots - } - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 83:4-83:24 *) -let hashmap_HashMap_new (t : Type0) : result (hashmap_HashMap_t t) = - hashmap_HashMap_new_with_capacity t 32 4 5 - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) -let rec hashmap_HashMap_clear_loop - (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : - Tot (result (alloc_vec_Vec (hashmap_List_t t))) - (decreases (hashmap_HashMap_clear_loop_decreases t slots i)) - = - let i1 = alloc_vec_Vec_len (hashmap_List_t t) slots in - if i < i1 - then - let* (_, index_mut_back) = - alloc_vec_Vec_index_mut (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) slots i - in - let* i2 = usize_add i 1 in - let* slots1 = index_mut_back Hashmap_List_Nil in - hashmap_HashMap_clear_loop t slots1 i2 - else Ok slots - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 88:4-88:27 *) -let hashmap_HashMap_clear - (t : Type0) (self : hashmap_HashMap_t t) : result (hashmap_HashMap_t t) = - let* hm = hashmap_HashMap_clear_loop t self.slots 0 in - Ok { self with num_entries = 0; slots = hm } - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::len]: - Source: 'tests/src/hashmap.rs', lines 98:4-98:30 *) -let hashmap_HashMap_len - (t : Type0) (self : hashmap_HashMap_t t) : result usize = - Ok self.num_entries - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) -let rec hashmap_HashMap_insert_in_list_loop - (t : Type0) (key : usize) (value : t) (ls : hashmap_List_t t) : - Tot (result (bool & (hashmap_List_t t))) - (decreases (hashmap_HashMap_insert_in_list_loop_decreases t key value ls)) - = - begin match ls with - | Hashmap_List_Cons ckey cvalue tl -> - if ckey = key - then Ok (false, Hashmap_List_Cons ckey value tl) - else - let* (b, tl1) = hashmap_HashMap_insert_in_list_loop t key value tl in - Ok (b, Hashmap_List_Cons ckey cvalue tl1) - | Hashmap_List_Nil -> Ok (true, Hashmap_List_Cons key value Hashmap_List_Nil) - end - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 105:4-105:71 *) -let hashmap_HashMap_insert_in_list - (t : Type0) (key : usize) (value : t) (ls : hashmap_List_t t) : - result (bool & (hashmap_List_t t)) - = - hashmap_HashMap_insert_in_list_loop t key value ls - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 125:4-125:54 *) -let hashmap_HashMap_insert_no_resize - (t : Type0) (self : hashmap_HashMap_t t) (key : usize) (value : t) : - result (hashmap_HashMap_t t) - = - let* hash = hashmap_hash_key key in - let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in - let* hash_mod = usize_rem hash i in - let* (l, index_mut_back) = - alloc_vec_Vec_index_mut (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) - self.slots hash_mod in - let* (inserted, l1) = hashmap_HashMap_insert_in_list t key value l in - if inserted - then - let* i1 = usize_add self.num_entries 1 in - let* v = index_mut_back l1 in - Ok { self with num_entries = i1; slots = v } - else let* v = index_mut_back l1 in Ok { self with slots = v } - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) -let rec hashmap_HashMap_move_elements_from_list_loop - (t : Type0) (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : - Tot (result (hashmap_HashMap_t t)) - (decreases ( - hashmap_HashMap_move_elements_from_list_loop_decreases t ntable ls)) - = - begin match ls with - | Hashmap_List_Cons k v tl -> - let* ntable1 = hashmap_HashMap_insert_no_resize t ntable k v in - hashmap_HashMap_move_elements_from_list_loop t ntable1 tl - | Hashmap_List_Nil -> Ok ntable - end - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 191:4-191:72 *) -let hashmap_HashMap_move_elements_from_list - (t : Type0) (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : - result (hashmap_HashMap_t t) - = - hashmap_HashMap_move_elements_from_list_loop t ntable ls - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) -let rec hashmap_HashMap_move_elements_loop - (t : Type0) (ntable : hashmap_HashMap_t t) - (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : - Tot (result ((hashmap_HashMap_t t) & (alloc_vec_Vec (hashmap_List_t t)))) - (decreases (hashmap_HashMap_move_elements_loop_decreases t ntable slots i)) - = - let i1 = alloc_vec_Vec_len (hashmap_List_t t) slots in - if i < i1 - then - let* (l, index_mut_back) = - alloc_vec_Vec_index_mut (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) slots i - in - let (ls, l1) = core_mem_replace (hashmap_List_t t) l Hashmap_List_Nil in - let* ntable1 = hashmap_HashMap_move_elements_from_list t ntable ls in - let* i2 = usize_add i 1 in - let* slots1 = index_mut_back l1 in - hashmap_HashMap_move_elements_loop t ntable1 slots1 i2 - else Ok (ntable, slots) - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 179:4-179:95 *) -let hashmap_HashMap_move_elements - (t : Type0) (ntable : hashmap_HashMap_t t) - (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : - result ((hashmap_HashMap_t t) & (alloc_vec_Vec (hashmap_List_t t))) - = - hashmap_HashMap_move_elements_loop t ntable slots i - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 148:4-148:28 *) -let hashmap_HashMap_try_resize - (t : Type0) (self : hashmap_HashMap_t t) : result (hashmap_HashMap_t t) = - let* max_usize = scalar_cast U32 Usize core_u32_max in - let capacity = alloc_vec_Vec_len (hashmap_List_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 - if capacity <= i2 - then - let* i3 = usize_mul capacity 2 in - let* ntable = hashmap_HashMap_new_with_capacity t i3 i i1 in - let* p = hashmap_HashMap_move_elements t ntable self.slots 0 in - let (ntable1, _) = p in - Ok - { ntable1 with num_entries = self.num_entries; max_load_factor = (i, i1) - } - else Ok { self with max_load_factor = (i, i1) } - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert]: - Source: 'tests/src/hashmap.rs', lines 137:4-137:48 *) -let hashmap_HashMap_insert - (t : Type0) (self : hashmap_HashMap_t t) (key : usize) (value : t) : - result (hashmap_HashMap_t t) - = - let* self1 = hashmap_HashMap_insert_no_resize t self key value in - let* i = hashmap_HashMap_len t self1 in - if i > self1.max_load then hashmap_HashMap_try_resize t self1 else Ok self1 - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) -let rec hashmap_HashMap_contains_key_in_list_loop - (t : Type0) (key : usize) (ls : hashmap_List_t t) : - Tot (result bool) - (decreases (hashmap_HashMap_contains_key_in_list_loop_decreases t key ls)) - = - begin match ls with - | Hashmap_List_Cons ckey _ tl -> - if ckey = key - then Ok true - else hashmap_HashMap_contains_key_in_list_loop t key tl - | Hashmap_List_Nil -> Ok false - end - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 214:4-214:68 *) -let hashmap_HashMap_contains_key_in_list - (t : Type0) (key : usize) (ls : hashmap_List_t t) : result bool = - hashmap_HashMap_contains_key_in_list_loop t key ls - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 207:4-207:49 *) -let hashmap_HashMap_contains_key - (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : result bool = - let* hash = hashmap_hash_key key in - let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in - let* hash_mod = usize_rem hash i in - let* l = - alloc_vec_Vec_index (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) - self.slots hash_mod in - hashmap_HashMap_contains_key_in_list t key l - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) -let rec hashmap_HashMap_get_in_list_loop - (t : Type0) (key : usize) (ls : hashmap_List_t t) : - Tot (result t) - (decreases (hashmap_HashMap_get_in_list_loop_decreases t key ls)) - = - begin match ls with - | Hashmap_List_Cons ckey cvalue tl -> - if ckey = key then Ok cvalue else hashmap_HashMap_get_in_list_loop t key tl - | Hashmap_List_Nil -> Fail Failure - end - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 232:4-232:70 *) -let hashmap_HashMap_get_in_list - (t : Type0) (key : usize) (ls : hashmap_List_t t) : result t = - hashmap_HashMap_get_in_list_loop t key ls - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 247:4-247:55 *) -let hashmap_HashMap_get - (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : result t = - let* hash = hashmap_hash_key key in - let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in - let* hash_mod = usize_rem hash i in - let* l = - alloc_vec_Vec_index (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) - self.slots hash_mod in - hashmap_HashMap_get_in_list t key l - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) -let rec hashmap_HashMap_get_mut_in_list_loop - (t : Type0) (ls : hashmap_List_t t) (key : usize) : - Tot (result (t & (t -> result (hashmap_List_t t)))) - (decreases (hashmap_HashMap_get_mut_in_list_loop_decreases t ls key)) - = - begin match ls with - | Hashmap_List_Cons ckey cvalue tl -> - if ckey = key - then - let back = fun ret -> Ok (Hashmap_List_Cons ckey ret tl) in - Ok (cvalue, back) - else - let* (x, back) = hashmap_HashMap_get_mut_in_list_loop t tl key in - let back1 = - fun ret -> - let* tl1 = back ret in Ok (Hashmap_List_Cons ckey cvalue tl1) in - Ok (x, back1) - | Hashmap_List_Nil -> Fail Failure - end - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 253:4-253:86 *) -let hashmap_HashMap_get_mut_in_list - (t : Type0) (ls : hashmap_List_t t) (key : usize) : - result (t & (t -> result (hashmap_List_t t))) - = - hashmap_HashMap_get_mut_in_list_loop t ls key - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:67 *) -let hashmap_HashMap_get_mut - (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : - result (t & (t -> result (hashmap_HashMap_t t))) - = - let* hash = hashmap_hash_key key in - let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in - let* hash_mod = usize_rem hash i in - let* (l, index_mut_back) = - alloc_vec_Vec_index_mut (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) - self.slots hash_mod in - let* (x, get_mut_in_list_back) = hashmap_HashMap_get_mut_in_list t l key in - let back = - fun ret -> - let* l1 = get_mut_in_list_back ret in - let* v = index_mut_back l1 in - Ok { self with slots = v } in - Ok (x, back) - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) -let rec hashmap_HashMap_remove_from_list_loop - (t : Type0) (key : usize) (ls : hashmap_List_t t) : - Tot (result ((option t) & (hashmap_List_t t))) - (decreases (hashmap_HashMap_remove_from_list_loop_decreases t key ls)) - = - begin match ls with - | Hashmap_List_Cons ckey x tl -> - if ckey = key - then - let (mv_ls, _) = - core_mem_replace (hashmap_List_t t) (Hashmap_List_Cons ckey x tl) - Hashmap_List_Nil in - begin match mv_ls with - | Hashmap_List_Cons _ cvalue tl1 -> Ok (Some cvalue, tl1) - | Hashmap_List_Nil -> Fail Failure - end - else - let* (o, tl1) = hashmap_HashMap_remove_from_list_loop t key tl in - Ok (o, Hashmap_List_Cons ckey x tl1) - | Hashmap_List_Nil -> Ok (None, Hashmap_List_Nil) - end - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 273:4-273:69 *) -let hashmap_HashMap_remove_from_list - (t : Type0) (key : usize) (ls : hashmap_List_t t) : - result ((option t) & (hashmap_List_t t)) - = - hashmap_HashMap_remove_from_list_loop t key ls - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 302:4-302:52 *) -let hashmap_HashMap_remove - (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : - result ((option t) & (hashmap_HashMap_t t)) - = - let* hash = hashmap_hash_key key in - let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in - let* hash_mod = usize_rem hash i in - let* (l, index_mut_back) = - alloc_vec_Vec_index_mut (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) - self.slots hash_mod in - let* (x, l1) = hashmap_HashMap_remove_from_list t key l in - begin match x with - | None -> let* v = index_mut_back l1 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 - Ok (Some x1, { self with num_entries = i1; slots = v }) - end - -(** [hashmap_main::hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 323:0-323:10 *) -let hashmap_test1 : result unit = - let* hm = hashmap_HashMap_new u64 in - let* hm1 = hashmap_HashMap_insert u64 hm 0 42 in - let* hm2 = hashmap_HashMap_insert u64 hm1 128 18 in - let* hm3 = hashmap_HashMap_insert u64 hm2 1024 138 in - let* hm4 = hashmap_HashMap_insert u64 hm3 1056 256 in - let* i = hashmap_HashMap_get u64 hm4 128 in - if not (i = 18) - then Fail Failure - else - let* (_, get_mut_back) = hashmap_HashMap_get_mut u64 hm4 1024 in - let* hm5 = get_mut_back 56 in - let* i1 = hashmap_HashMap_get u64 hm5 1024 in - if not (i1 = 56) - then Fail Failure - else - let* (x, hm6) = hashmap_HashMap_remove u64 hm5 1024 in - begin match x with - | None -> Fail Failure - | Some x1 -> - if not (x1 = 56) - then Fail Failure - else - let* i2 = hashmap_HashMap_get u64 hm6 0 in - if not (i2 = 42) - then Fail Failure - else - let* i3 = hashmap_HashMap_get u64 hm6 128 in - if not (i3 = 18) - then Fail Failure - else - let* i4 = hashmap_HashMap_get u64 hm6 1056 in - if not (i4 = 256) then Fail Failure else Ok () - end - -(** [hashmap_main::insert_on_disk]: - Source: 'tests/src/hashmap_main.rs', lines 13:0-13:43 *) -let insert_on_disk - (key : usize) (value : u64) (st : state) : result (state & unit) = - let* (st1, hm) = hashmap_utils_deserialize st in - let* hm1 = hashmap_HashMap_insert u64 hm key value in - hashmap_utils_serialize hm1 st1 - -(** [hashmap_main::main]: - Source: 'tests/src/hashmap_main.rs', lines 22:0-22:13 *) -let main : result unit = - Ok () - diff --git a/tests/fstar/hashmap_on_disk/HashmapMain.FunsExternal.fsti b/tests/fstar/hashmap_on_disk/HashmapMain.FunsExternal.fsti deleted file mode 100644 index cc20d988..00000000 --- a/tests/fstar/hashmap_on_disk/HashmapMain.FunsExternal.fsti +++ /dev/null @@ -1,18 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: external function declarations *) -module HashmapMain.FunsExternal -open Primitives -include HashmapMain.Types - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** [hashmap_main::hashmap_utils::deserialize]: - Source: 'tests/src/hashmap_utils.rs', lines 11:0-11:43 *) -val hashmap_utils_deserialize - : state -> result (state & (hashmap_HashMap_t u64)) - -(** [hashmap_main::hashmap_utils::serialize]: - Source: 'tests/src/hashmap_utils.rs', lines 6:0-6:42 *) -val hashmap_utils_serialize - : hashmap_HashMap_t u64 -> state -> result (state & unit) - diff --git a/tests/fstar/hashmap_on_disk/HashmapMain.Properties.fst b/tests/fstar/hashmap_on_disk/HashmapMain.Properties.fst deleted file mode 100644 index beb3dc2c..00000000 --- a/tests/fstar/hashmap_on_disk/HashmapMain.Properties.fst +++ /dev/null @@ -1,48 +0,0 @@ -(** Properties about the hashmap written on disk *) -module HashmapMain.Properties -open Primitives -open HashmapMain.Funs - -#set-options "--z3rlimit 50 --fuel 0 --ifuel 1" - -/// Below, we focus on the functions to read from disk/write to disk to showcase -/// how such reasoning which mixes opaque functions together with a state-error -/// monad can be performed. - -(*** Hypotheses *) - -/// [state_v] gives us the hash map currently stored on disk -assume -val state_v : state -> hashmap_HashMap_t u64 - -/// [serialize] updates the hash map stored on disk -assume -val serialize_lem (hm : hashmap_HashMap_t u64) (st : state) : Lemma ( - match hashmap_utils_serialize hm st with - | Fail _ -> True - | Ok (st', ()) -> state_v st' == hm) - [SMTPat (hashmap_utils_serialize hm st)] - -/// [deserialize] gives us the hash map stored on disk, without updating it -assume -val deserialize_lem (st : state) : Lemma ( - match hashmap_utils_deserialize st with - | Fail _ -> True - | Ok (st', hm) -> hm == state_v st /\ st' == st) - [SMTPat (hashmap_utils_deserialize st)] - -(*** Lemmas *) - -/// The obvious lemma about [insert_on_disk]: the updated hash map stored on disk -/// is exactly the hash map produced from inserting the binding ([key], [value]) -/// in the hash map previously stored on disk. -val insert_on_disk_lem (key : usize) (value : u64) (st : state) : Lemma ( - match insert_on_disk key value st with - | Fail _ -> True - | Ok (st', ()) -> - let hm = state_v st in - match hashmap_HashMap_insert u64 hm key value with - | Fail _ -> False - | Ok hm' -> hm' == state_v st') - -let insert_on_disk_lem key value st = () diff --git a/tests/fstar/hashmap_on_disk/HashmapMain.Types.fst b/tests/fstar/hashmap_on_disk/HashmapMain.Types.fst deleted file mode 100644 index 85bcaeea..00000000 --- a/tests/fstar/hashmap_on_disk/HashmapMain.Types.fst +++ /dev/null @@ -1,24 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: type definitions *) -module HashmapMain.Types -open Primitives -include HashmapMain.TypesExternal - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** [hashmap_main::hashmap::List] - Source: 'tests/src/hashmap.rs', lines 27:0-27:16 *) -type hashmap_List_t (t : Type0) = -| Hashmap_List_Cons : usize -> t -> hashmap_List_t t -> hashmap_List_t t -| Hashmap_List_Nil : hashmap_List_t t - -(** [hashmap_main::hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 43:0-43:21 *) -type hashmap_HashMap_t (t : Type0) = -{ - num_entries : usize; - max_load_factor : (usize & usize); - max_load : usize; - slots : alloc_vec_Vec (hashmap_List_t t); -} - diff --git a/tests/fstar/hashmap_on_disk/HashmapMain.TypesExternal.fsti b/tests/fstar/hashmap_on_disk/HashmapMain.TypesExternal.fsti deleted file mode 100644 index 75747408..00000000 --- a/tests/fstar/hashmap_on_disk/HashmapMain.TypesExternal.fsti +++ /dev/null @@ -1,10 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: external type declarations *) -module HashmapMain.TypesExternal -open Primitives - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** The state type used in the state-error monad *) -val state : Type0 - diff --git a/tests/fstar/hashmap_on_disk/Makefile b/tests/fstar/hashmap_on_disk/Makefile deleted file mode 100644 index fa7d1f36..00000000 --- a/tests/fstar/hashmap_on_disk/Makefile +++ /dev/null @@ -1,49 +0,0 @@ -# This file was automatically generated - modify ../Makefile.template instead -INCLUDE_DIRS = . - -FSTAR_INCLUDES = $(addprefix --include ,$(INCLUDE_DIRS)) - -FSTAR_HINTS ?= --use_hints --use_hint_hashes --record_hints - -FSTAR_OPTIONS = $(FSTAR_HINTS) \ - --cache_checked_modules $(FSTAR_INCLUDES) --cmi \ - --warn_error '+241@247+285-274' \ - -FSTAR_EXE ?= fstar.exe -FSTAR_NO_FLAGS = $(FSTAR_EXE) --already_cached 'Prims FStar LowStar Steel' --odir obj --cache_dir obj - -FSTAR = $(FSTAR_NO_FLAGS) $(FSTAR_OPTIONS) - -# The F* roots are used to compute the dependency graph, and generate the .depend file -FSTAR_ROOTS ?= $(wildcard *.fst *.fsti) - -# Build all the files -all: $(addprefix obj/,$(addsuffix .checked,$(FSTAR_ROOTS))) - -# This is the right way to ensure the .depend file always gets re-built. -ifeq (,$(filter %-in,$(MAKECMDGOALS))) -ifndef NODEPEND -ifndef MAKE_RESTARTS -.depend: .FORCE - $(FSTAR_NO_FLAGS) --dep full $(notdir $(FSTAR_ROOTS)) > $@ - -.PHONY: .FORCE -.FORCE: -endif -endif - -include .depend -endif - -# For the interactive mode -%.fst-in %.fsti-in: - @echo $(FSTAR_OPTIONS) - -# Generete the .checked files in batch mode -%.checked: - $(FSTAR) $(FSTAR_OPTIONS) $< && \ - touch -c $@ - -.PHONY: clean -clean: - rm -f obj/* diff --git a/tests/fstar/hashmap_on_disk/Primitives.fst b/tests/fstar/hashmap_on_disk/Primitives.fst deleted file mode 100644 index 9951ccc3..00000000 --- a/tests/fstar/hashmap_on_disk/Primitives.fst +++ /dev/null @@ -1,929 +0,0 @@ -/// This file lists primitive and assumed functions and types -module Primitives -open FStar.Mul -open FStar.List.Tot - -#set-options "--z3rlimit 15 --fuel 0 --ifuel 1" - -(*** Utilities *) -val list_update (#a : Type0) (ls : list a) (i : nat{i < length ls}) (x : a) : - ls':list a{ - length ls' = length ls /\ - index ls' i == x - } -#push-options "--fuel 1" -let rec list_update #a ls i x = - match ls with - | x' :: ls -> if i = 0 then x :: ls else x' :: list_update ls (i-1) x -#pop-options - -(*** Result *) -type error : Type0 = -| Failure -| OutOfFuel - -type result (a : Type0) : Type0 = -| Ok : v:a -> result a -| Fail : e:error -> result a - -// Monadic return operator -unfold let return (#a : Type0) (x : a) : result a = Ok x - -// Monadic bind operator. -// Allows to use the notation: -// ``` -// let* x = y in -// ... -// ``` -unfold let (let*) (#a #b : Type0) (m: result a) - (f: (x:a) -> Pure (result b) (requires (m == Ok x)) (ensures fun _ -> True)) : - result b = - match m with - | Ok x -> f x - | Fail e -> Fail e - -// Monadic assert(...) -let massert (b:bool) : result unit = if b then Ok () else Fail Failure - -// Normalize and unwrap a successful result (used for globals). -let eval_global (#a : Type0) (x : result a{Ok? (normalize_term x)}) : a = Ok?.v x - -(*** Misc *) -type char = FStar.Char.char -type string = string - -let is_zero (n: nat) : bool = n = 0 -let decrease (n: nat{n > 0}) : nat = n - 1 - -let core_mem_replace (a : Type0) (x : a) (y : a) : a & a = (x, x) - -// We don't really use raw pointers for now -type mut_raw_ptr (t : Type0) = { v : t } -type const_raw_ptr (t : Type0) = { v : t } - -(*** Scalars *) -/// Rem.: most of the following code was partially generated - -assume val size_numbits : pos - -// TODO: we could use FStar.Int.int_t and FStar.UInt.int_t - -let isize_min : int = -9223372036854775808 // TODO: should be opaque -let isize_max : int = 9223372036854775807 // TODO: should be opaque -let i8_min : int = -128 -let i8_max : int = 127 -let i16_min : int = -32768 -let i16_max : int = 32767 -let i32_min : int = -2147483648 -let i32_max : int = 2147483647 -let i64_min : int = -9223372036854775808 -let i64_max : int = 9223372036854775807 -let i128_min : int = -170141183460469231731687303715884105728 -let i128_max : int = 170141183460469231731687303715884105727 -let usize_min : int = 0 -let usize_max : int = 4294967295 // TODO: should be opaque -let u8_min : int = 0 -let u8_max : int = 255 -let u16_min : int = 0 -let u16_max : int = 65535 -let u32_min : int = 0 -let u32_max : int = 4294967295 -let u64_min : int = 0 -let u64_max : int = 18446744073709551615 -let u128_min : int = 0 -let u128_max : int = 340282366920938463463374607431768211455 - -type scalar_ty = -| Isize -| I8 -| I16 -| I32 -| I64 -| I128 -| Usize -| U8 -| U16 -| U32 -| U64 -| U128 - -let is_unsigned = function - | Isize | I8 | I16 | I32 | I64 | I128 -> false - | Usize | U8 | U16 | U32 | U64 | U128 -> true - -let scalar_min (ty : scalar_ty) : int = - match ty with - | Isize -> isize_min - | I8 -> i8_min - | I16 -> i16_min - | I32 -> i32_min - | I64 -> i64_min - | I128 -> i128_min - | Usize -> usize_min - | U8 -> u8_min - | U16 -> u16_min - | U32 -> u32_min - | U64 -> u64_min - | U128 -> u128_min - -let scalar_max (ty : scalar_ty) : int = - match ty with - | Isize -> isize_max - | I8 -> i8_max - | I16 -> i16_max - | I32 -> i32_max - | I64 -> i64_max - | I128 -> i128_max - | Usize -> usize_max - | U8 -> u8_max - | U16 -> u16_max - | U32 -> u32_max - | U64 -> u64_max - | U128 -> u128_max - -type scalar (ty : scalar_ty) : eqtype = x:int{scalar_min ty <= x && x <= scalar_max ty} - -let mk_scalar (ty : scalar_ty) (x : int) : result (scalar ty) = - if scalar_min ty <= x && scalar_max ty >= x then Ok x else Fail Failure - -let scalar_neg (#ty : scalar_ty) (x : scalar ty) : result (scalar ty) = mk_scalar ty (-x) - -let scalar_div (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = - if y <> 0 then mk_scalar ty (x / y) else Fail Failure - -/// The remainder operation -let int_rem (x : int) (y : int{y <> 0}) : int = - if x >= 0 then (x % y) else -(x % y) - -(* Checking consistency with Rust *) -let _ = assert_norm(int_rem 1 2 = 1) -let _ = assert_norm(int_rem (-1) 2 = -1) -let _ = assert_norm(int_rem 1 (-2) = 1) -let _ = assert_norm(int_rem (-1) (-2) = -1) - -let scalar_rem (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = - if y <> 0 then mk_scalar ty (int_rem x y) else Fail Failure - -let scalar_add (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = - mk_scalar ty (x + y) - -let scalar_sub (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = - mk_scalar ty (x - y) - -let scalar_mul (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = - mk_scalar ty (x * y) - -let scalar_xor (#ty : scalar_ty) - (x : scalar ty) (y : scalar ty) : scalar ty = - match ty with - | U8 -> FStar.UInt.logxor #8 x y - | U16 -> FStar.UInt.logxor #16 x y - | U32 -> FStar.UInt.logxor #32 x y - | U64 -> FStar.UInt.logxor #64 x y - | U128 -> FStar.UInt.logxor #128 x y - | Usize -> admit() // TODO - | I8 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 8); - normalize_spec (scalar I8); - FStar.Int.logxor #8 x y - | I16 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 16); - normalize_spec (scalar I16); - FStar.Int.logxor #16 x y - | I32 -> FStar.Int.logxor #32 x y - | I64 -> FStar.Int.logxor #64 x y - | I128 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 128); - normalize_spec (scalar I128); - FStar.Int.logxor #128 x y - | Isize -> admit() // TODO - -let scalar_or (#ty : scalar_ty) - (x : scalar ty) (y : scalar ty) : scalar ty = - match ty with - | U8 -> FStar.UInt.logor #8 x y - | U16 -> FStar.UInt.logor #16 x y - | U32 -> FStar.UInt.logor #32 x y - | U64 -> FStar.UInt.logor #64 x y - | U128 -> FStar.UInt.logor #128 x y - | Usize -> admit() // TODO - | I8 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 8); - normalize_spec (scalar I8); - FStar.Int.logor #8 x y - | I16 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 16); - normalize_spec (scalar I16); - FStar.Int.logor #16 x y - | I32 -> FStar.Int.logor #32 x y - | I64 -> FStar.Int.logor #64 x y - | I128 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 128); - normalize_spec (scalar I128); - FStar.Int.logor #128 x y - | Isize -> admit() // TODO - -let scalar_and (#ty : scalar_ty) - (x : scalar ty) (y : scalar ty) : scalar ty = - match ty with - | U8 -> FStar.UInt.logand #8 x y - | U16 -> FStar.UInt.logand #16 x y - | U32 -> FStar.UInt.logand #32 x y - | U64 -> FStar.UInt.logand #64 x y - | U128 -> FStar.UInt.logand #128 x y - | Usize -> admit() // TODO - | I8 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 8); - normalize_spec (scalar I8); - FStar.Int.logand #8 x y - | I16 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 16); - normalize_spec (scalar I16); - FStar.Int.logand #16 x y - | I32 -> FStar.Int.logand #32 x y - | I64 -> FStar.Int.logand #64 x y - | I128 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 128); - normalize_spec (scalar I128); - FStar.Int.logand #128 x y - | Isize -> admit() // TODO - -// Shift left -let scalar_shl (#ty0 #ty1 : scalar_ty) - (x : scalar ty0) (y : scalar ty1) : result (scalar ty0) = - admit() - -// Shift right -let scalar_shr (#ty0 #ty1 : scalar_ty) - (x : scalar ty0) (y : scalar ty1) : result (scalar ty0) = - admit() - -(** Cast an integer from a [src_ty] to a [tgt_ty] *) -// TODO: check the semantics of casts in Rust -let scalar_cast (src_ty : scalar_ty) (tgt_ty : scalar_ty) (x : scalar src_ty) : result (scalar tgt_ty) = - mk_scalar tgt_ty x - -// This can't fail, but for now we make all casts faillible (easier for the translation) -let scalar_cast_bool (tgt_ty : scalar_ty) (x : bool) : result (scalar tgt_ty) = - mk_scalar tgt_ty (if x then 1 else 0) - -/// The scalar types -type isize : eqtype = scalar Isize -type i8 : eqtype = scalar I8 -type i16 : eqtype = scalar I16 -type i32 : eqtype = scalar I32 -type i64 : eqtype = scalar I64 -type i128 : eqtype = scalar I128 -type usize : eqtype = scalar Usize -type u8 : eqtype = scalar U8 -type u16 : eqtype = scalar U16 -type u32 : eqtype = scalar U32 -type u64 : eqtype = scalar U64 -type u128 : eqtype = scalar U128 - - -let core_isize_min : isize = isize_min -let core_isize_max : isize = isize_max -let core_i8_min : i8 = i8_min -let core_i8_max : i8 = i8_max -let core_i16_min : i16 = i16_min -let core_i16_max : i16 = i16_max -let core_i32_min : i32 = i32_min -let core_i32_max : i32 = i32_max -let core_i64_min : i64 = i64_min -let core_i64_max : i64 = i64_max -let core_i128_min : i128 = i128_min -let core_i128_max : i128 = i128_max - -let core_usize_min : usize = usize_min -let core_usize_max : usize = usize_max -let core_u8_min : u8 = u8_min -let core_u8_max : u8 = u8_max -let core_u16_min : u16 = u16_min -let core_u16_max : u16 = u16_max -let core_u32_min : u32 = u32_min -let core_u32_max : u32 = u32_max -let core_u64_min : u64 = u64_min -let core_u64_max : u64 = u64_max -let core_u128_min : u128 = u128_min -let core_u128_max : u128 = u128_max - -/// Negation -let isize_neg = scalar_neg #Isize -let i8_neg = scalar_neg #I8 -let i16_neg = scalar_neg #I16 -let i32_neg = scalar_neg #I32 -let i64_neg = scalar_neg #I64 -let i128_neg = scalar_neg #I128 - -/// Division -let isize_div = scalar_div #Isize -let i8_div = scalar_div #I8 -let i16_div = scalar_div #I16 -let i32_div = scalar_div #I32 -let i64_div = scalar_div #I64 -let i128_div = scalar_div #I128 -let usize_div = scalar_div #Usize -let u8_div = scalar_div #U8 -let u16_div = scalar_div #U16 -let u32_div = scalar_div #U32 -let u64_div = scalar_div #U64 -let u128_div = scalar_div #U128 - -/// Remainder -let isize_rem = scalar_rem #Isize -let i8_rem = scalar_rem #I8 -let i16_rem = scalar_rem #I16 -let i32_rem = scalar_rem #I32 -let i64_rem = scalar_rem #I64 -let i128_rem = scalar_rem #I128 -let usize_rem = scalar_rem #Usize -let u8_rem = scalar_rem #U8 -let u16_rem = scalar_rem #U16 -let u32_rem = scalar_rem #U32 -let u64_rem = scalar_rem #U64 -let u128_rem = scalar_rem #U128 - -/// Addition -let isize_add = scalar_add #Isize -let i8_add = scalar_add #I8 -let i16_add = scalar_add #I16 -let i32_add = scalar_add #I32 -let i64_add = scalar_add #I64 -let i128_add = scalar_add #I128 -let usize_add = scalar_add #Usize -let u8_add = scalar_add #U8 -let u16_add = scalar_add #U16 -let u32_add = scalar_add #U32 -let u64_add = scalar_add #U64 -let u128_add = scalar_add #U128 - -/// Subtraction -let isize_sub = scalar_sub #Isize -let i8_sub = scalar_sub #I8 -let i16_sub = scalar_sub #I16 -let i32_sub = scalar_sub #I32 -let i64_sub = scalar_sub #I64 -let i128_sub = scalar_sub #I128 -let usize_sub = scalar_sub #Usize -let u8_sub = scalar_sub #U8 -let u16_sub = scalar_sub #U16 -let u32_sub = scalar_sub #U32 -let u64_sub = scalar_sub #U64 -let u128_sub = scalar_sub #U128 - -/// Multiplication -let isize_mul = scalar_mul #Isize -let i8_mul = scalar_mul #I8 -let i16_mul = scalar_mul #I16 -let i32_mul = scalar_mul #I32 -let i64_mul = scalar_mul #I64 -let i128_mul = scalar_mul #I128 -let usize_mul = scalar_mul #Usize -let u8_mul = scalar_mul #U8 -let u16_mul = scalar_mul #U16 -let u32_mul = scalar_mul #U32 -let u64_mul = scalar_mul #U64 -let u128_mul = scalar_mul #U128 - -/// Xor -let u8_xor = scalar_xor #U8 -let u16_xor = scalar_xor #U16 -let u32_xor = scalar_xor #U32 -let u64_xor = scalar_xor #U64 -let u128_xor = scalar_xor #U128 -let usize_xor = scalar_xor #Usize -let i8_xor = scalar_xor #I8 -let i16_xor = scalar_xor #I16 -let i32_xor = scalar_xor #I32 -let i64_xor = scalar_xor #I64 -let i128_xor = scalar_xor #I128 -let isize_xor = scalar_xor #Isize - -/// Or -let u8_or = scalar_or #U8 -let u16_or = scalar_or #U16 -let u32_or = scalar_or #U32 -let u64_or = scalar_or #U64 -let u128_or = scalar_or #U128 -let usize_or = scalar_or #Usize -let i8_or = scalar_or #I8 -let i16_or = scalar_or #I16 -let i32_or = scalar_or #I32 -let i64_or = scalar_or #I64 -let i128_or = scalar_or #I128 -let isize_or = scalar_or #Isize - -/// And -let u8_and = scalar_and #U8 -let u16_and = scalar_and #U16 -let u32_and = scalar_and #U32 -let u64_and = scalar_and #U64 -let u128_and = scalar_and #U128 -let usize_and = scalar_and #Usize -let i8_and = scalar_and #I8 -let i16_and = scalar_and #I16 -let i32_and = scalar_and #I32 -let i64_and = scalar_and #I64 -let i128_and = scalar_and #I128 -let isize_and = scalar_and #Isize - -/// Shift left -let u8_shl #ty = scalar_shl #U8 #ty -let u16_shl #ty = scalar_shl #U16 #ty -let u32_shl #ty = scalar_shl #U32 #ty -let u64_shl #ty = scalar_shl #U64 #ty -let u128_shl #ty = scalar_shl #U128 #ty -let usize_shl #ty = scalar_shl #Usize #ty -let i8_shl #ty = scalar_shl #I8 #ty -let i16_shl #ty = scalar_shl #I16 #ty -let i32_shl #ty = scalar_shl #I32 #ty -let i64_shl #ty = scalar_shl #I64 #ty -let i128_shl #ty = scalar_shl #I128 #ty -let isize_shl #ty = scalar_shl #Isize #ty - -/// Shift right -let u8_shr #ty = scalar_shr #U8 #ty -let u16_shr #ty = scalar_shr #U16 #ty -let u32_shr #ty = scalar_shr #U32 #ty -let u64_shr #ty = scalar_shr #U64 #ty -let u128_shr #ty = scalar_shr #U128 #ty -let usize_shr #ty = scalar_shr #Usize #ty -let i8_shr #ty = scalar_shr #I8 #ty -let i16_shr #ty = scalar_shr #I16 #ty -let i32_shr #ty = scalar_shr #I32 #ty -let i64_shr #ty = scalar_shr #I64 #ty -let i128_shr #ty = scalar_shr #I128 #ty -let isize_shr #ty = scalar_shr #Isize #ty - -(*** core *) - -/// Trait declaration: [core::clone::Clone] -noeq type core_clone_Clone (self : Type0) = { - clone : self → result self -} - -let core_clone_impls_CloneBool_clone (b : bool) : bool = b - -let core_clone_CloneBool : core_clone_Clone bool = { - clone = fun b -> Ok (core_clone_impls_CloneBool_clone b) -} - -let core_clone_impls_CloneUsize_clone (x : usize) : usize = x -let core_clone_impls_CloneU8_clone (x : u8) : u8 = x -let core_clone_impls_CloneU16_clone (x : u16) : u16 = x -let core_clone_impls_CloneU32_clone (x : u32) : u32 = x -let core_clone_impls_CloneU64_clone (x : u64) : u64 = x -let core_clone_impls_CloneU128_clone (x : u128) : u128 = x - -let core_clone_impls_CloneIsize_clone (x : isize) : isize = x -let core_clone_impls_CloneI8_clone (x : i8) : i8 = x -let core_clone_impls_CloneI16_clone (x : i16) : i16 = x -let core_clone_impls_CloneI32_clone (x : i32) : i32 = x -let core_clone_impls_CloneI64_clone (x : i64) : i64 = x -let core_clone_impls_CloneI128_clone (x : i128) : i128 = x - -let core_clone_CloneUsize : core_clone_Clone usize = { - clone = fun x -> Ok (core_clone_impls_CloneUsize_clone x) -} - -let core_clone_CloneU8 : core_clone_Clone u8 = { - clone = fun x -> Ok (core_clone_impls_CloneU8_clone x) -} - -let core_clone_CloneU16 : core_clone_Clone u16 = { - clone = fun x -> Ok (core_clone_impls_CloneU16_clone x) -} - -let core_clone_CloneU32 : core_clone_Clone u32 = { - clone = fun x -> Ok (core_clone_impls_CloneU32_clone x) -} - -let core_clone_CloneU64 : core_clone_Clone u64 = { - clone = fun x -> Ok (core_clone_impls_CloneU64_clone x) -} - -let core_clone_CloneU128 : core_clone_Clone u128 = { - clone = fun x -> Ok (core_clone_impls_CloneU128_clone x) -} - -let core_clone_CloneIsize : core_clone_Clone isize = { - clone = fun x -> Ok (core_clone_impls_CloneIsize_clone x) -} - -let core_clone_CloneI8 : core_clone_Clone i8 = { - clone = fun x -> Ok (core_clone_impls_CloneI8_clone x) -} - -let core_clone_CloneI16 : core_clone_Clone i16 = { - clone = fun x -> Ok (core_clone_impls_CloneI16_clone x) -} - -let core_clone_CloneI32 : core_clone_Clone i32 = { - clone = fun x -> Ok (core_clone_impls_CloneI32_clone x) -} - -let core_clone_CloneI64 : core_clone_Clone i64 = { - clone = fun x -> Ok (core_clone_impls_CloneI64_clone x) -} - -let core_clone_CloneI128 : core_clone_Clone i128 = { - clone = fun x -> Ok (core_clone_impls_CloneI128_clone x) -} - -(** [core::option::{core::option::Option}::unwrap] *) -let core_option_Option_unwrap (t : Type0) (x : option t) : result t = - match x with - | None -> Fail Failure - | Some x -> Ok x - -(*** core::ops *) - -// Trait declaration: [core::ops::index::Index] -noeq type core_ops_index_Index (self idx : Type0) = { - output : Type0; - index : self → idx → result output -} - -// Trait declaration: [core::ops::index::IndexMut] -noeq type core_ops_index_IndexMut (self idx : Type0) = { - indexInst : core_ops_index_Index self idx; - index_mut : self → idx → result (indexInst.output & (indexInst.output → result self)); -} - -// Trait declaration [core::ops::deref::Deref] -noeq type core_ops_deref_Deref (self : Type0) = { - target : Type0; - deref : self → result target; -} - -// Trait declaration [core::ops::deref::DerefMut] -noeq type core_ops_deref_DerefMut (self : Type0) = { - derefInst : core_ops_deref_Deref self; - deref_mut : self → result (derefInst.target & (derefInst.target → result self)); -} - -type core_ops_range_Range (a : Type0) = { - start : a; - end_ : a; -} - -(*** [alloc] *) - -let alloc_boxed_Box_deref (t : Type0) (x : t) : result t = Ok x -let alloc_boxed_Box_deref_mut (t : Type0) (x : t) : result (t & (t -> result t)) = - Ok (x, (fun x -> Ok x)) - -// Trait instance -let alloc_boxed_Box_coreopsDerefInst (self : Type0) : core_ops_deref_Deref self = { - target = self; - deref = alloc_boxed_Box_deref self; -} - -// Trait instance -let alloc_boxed_Box_coreopsDerefMutInst (self : Type0) : core_ops_deref_DerefMut self = { - derefInst = alloc_boxed_Box_coreopsDerefInst self; - deref_mut = alloc_boxed_Box_deref_mut self; -} - -(*** Array *) -type array (a : Type0) (n : usize) = s:list a{length s = n} - -// We tried putting the normalize_term condition as a refinement on the list -// but it didn't work. It works with the requires clause. -let mk_array (a : Type0) (n : usize) - (l : list a) : - Pure (array a n) - (requires (normalize_term(FStar.List.Tot.length l) = n)) - (ensures (fun _ -> True)) = - normalize_term_spec (FStar.List.Tot.length l); - l - -let array_index_usize (a : Type0) (n : usize) (x : array a n) (i : usize) : result a = - if i < length x then Ok (index x i) - else Fail Failure - -let array_update_usize (a : Type0) (n : usize) (x : array a n) (i : usize) (nx : a) : - result (array a n) = - if i < length x then Ok (list_update x i nx) - else Fail Failure - -let array_index_mut_usize (a : Type0) (n : usize) (x : array a n) (i : usize) : - result (a & (a -> result (array a n))) = - match array_index_usize a n x i with - | Fail e -> Fail e - | Ok v -> - Ok (v, array_update_usize a n x i) - -(*** Slice *) -type slice (a : Type0) = s:list a{length s <= usize_max} - -let slice_len (a : Type0) (s : slice a) : usize = length s - -let slice_index_usize (a : Type0) (x : slice a) (i : usize) : result a = - if i < length x then Ok (index x i) - else Fail Failure - -let slice_update_usize (a : Type0) (x : slice a) (i : usize) (nx : a) : result (slice a) = - if i < length x then Ok (list_update x i nx) - else Fail Failure - -let slice_index_mut_usize (a : Type0) (s : slice a) (i : usize) : - result (a & (a -> result (slice a))) = - match slice_index_usize a s i with - | Fail e -> Fail e - | Ok x -> - Ok (x, slice_update_usize a s i) - -(*** Subslices *) - -let array_to_slice (a : Type0) (n : usize) (x : array a n) : result (slice a) = Ok x -let array_from_slice (a : Type0) (n : usize) (x : array a n) (s : slice a) : result (array a n) = - if length s = n then Ok s - else Fail Failure - -let array_to_slice_mut (a : Type0) (n : usize) (x : array a n) : - result (slice a & (slice a -> result (array a n))) = - Ok (x, array_from_slice a n x) - -// TODO: finish the definitions below (there lacks [List.drop] and [List.take] in the standard library *) -let array_subslice (a : Type0) (n : usize) (x : array a n) (r : core_ops_range_Range usize) : result (slice a) = - admit() - -let array_update_subslice (a : Type0) (n : usize) (x : array a n) (r : core_ops_range_Range usize) (ns : slice a) : result (array a n) = - admit() - -let array_repeat (a : Type0) (n : usize) (x : a) : array a n = - admit() - -let slice_subslice (a : Type0) (x : slice a) (r : core_ops_range_Range usize) : result (slice a) = - admit() - -let slice_update_subslice (a : Type0) (x : slice a) (r : core_ops_range_Range usize) (ns : slice a) : result (slice a) = - admit() - -(*** Vector *) -type alloc_vec_Vec (a : Type0) = v:list a{length v <= usize_max} - -let alloc_vec_Vec_new (a : Type0) : alloc_vec_Vec a = assert_norm(length #a [] == 0); [] -let alloc_vec_Vec_len (a : Type0) (v : alloc_vec_Vec a) : usize = length v - -// Helper -let alloc_vec_Vec_index_usize (#a : Type0) (v : alloc_vec_Vec a) (i : usize) : result a = - if i < length v then Ok (index v i) else Fail Failure -// Helper -let alloc_vec_Vec_update_usize (#a : Type0) (v : alloc_vec_Vec a) (i : usize) (x : a) : result (alloc_vec_Vec a) = - if i < length v then Ok (list_update v i x) else Fail Failure - -let alloc_vec_Vec_index_mut_usize (#a : Type0) (v: alloc_vec_Vec a) (i: usize) : - result (a & (a → result (alloc_vec_Vec a))) = - match alloc_vec_Vec_index_usize v i with - | Ok x -> - Ok (x, alloc_vec_Vec_update_usize v i) - | Fail e -> Fail e - -let alloc_vec_Vec_push (a : Type0) (v : alloc_vec_Vec a) (x : a) : - Pure (result (alloc_vec_Vec a)) - (requires True) - (ensures (fun res -> - match res with - | Fail e -> e == Failure - | Ok v' -> length v' = length v + 1)) = - if length v < usize_max then begin - (**) assert_norm(length [x] == 1); - (**) append_length v [x]; - (**) assert(length (append v [x]) = length v + 1); - Ok (append v [x]) - end - else Fail Failure - -let alloc_vec_Vec_insert (a : Type0) (v : alloc_vec_Vec a) (i : usize) (x : a) : result (alloc_vec_Vec a) = - if i < length v then Ok (list_update v i x) else Fail Failure - -// Trait declaration: [core::slice::index::private_slice_index::Sealed] -type core_slice_index_private_slice_index_Sealed (self : Type0) = unit - -// Trait declaration: [core::slice::index::SliceIndex] -noeq type core_slice_index_SliceIndex (self t : Type0) = { - sealedInst : core_slice_index_private_slice_index_Sealed self; - output : Type0; - get : self → t → result (option output); - get_mut : self → t → result (option output & (option output -> result t)); - get_unchecked : self → const_raw_ptr t → result (const_raw_ptr output); - get_unchecked_mut : self → mut_raw_ptr t → result (mut_raw_ptr output); - index : self → t → result output; - index_mut : self → t → result (output & (output -> result t)); -} - -// [core::slice::index::[T]::index]: forward function -let core_slice_index_Slice_index - (t idx : Type0) (inst : core_slice_index_SliceIndex idx (slice t)) - (s : slice t) (i : idx) : result inst.output = - let* x = inst.get i s in - match x with - | None -> Fail Failure - | Some x -> Ok x - -// [core::slice::index::Range:::get]: forward function -let core_slice_index_RangeUsize_get (t : Type0) (i : core_ops_range_Range usize) (s : slice t) : - result (option (slice t)) = - admit () // TODO - -// [core::slice::index::Range::get_mut]: forward function -let core_slice_index_RangeUsize_get_mut (t : Type0) : - core_ops_range_Range usize → slice t → result (option (slice t) & (option (slice t) -> result (slice t))) = - admit () // TODO - -// [core::slice::index::Range::get_unchecked]: forward function -let core_slice_index_RangeUsize_get_unchecked - (t : Type0) : - core_ops_range_Range usize → const_raw_ptr (slice t) → result (const_raw_ptr (slice t)) = - // Don't know what the model should be - for now we always fail to make - // sure code which uses it fails - fun _ _ -> Fail Failure - -// [core::slice::index::Range::get_unchecked_mut]: forward function -let core_slice_index_RangeUsize_get_unchecked_mut - (t : Type0) : - core_ops_range_Range usize → mut_raw_ptr (slice t) → result (mut_raw_ptr (slice t)) = - // Don't know what the model should be - for now we always fail to make - // sure code which uses it fails - fun _ _ -> Fail Failure - -// [core::slice::index::Range::index]: forward function -let core_slice_index_RangeUsize_index - (t : Type0) : core_ops_range_Range usize → slice t → result (slice t) = - admit () // TODO - -// [core::slice::index::Range::index_mut]: forward function -let core_slice_index_RangeUsize_index_mut (t : Type0) : - core_ops_range_Range usize → slice t → result (slice t & (slice t -> result (slice t))) = - admit () // TODO - -// [core::slice::index::[T]::index_mut]: forward function -let core_slice_index_Slice_index_mut - (t idx : Type0) (inst : core_slice_index_SliceIndex idx (slice t)) : - slice t → idx → result (inst.output & (inst.output -> result (slice t))) = - admit () // - -// [core::array::[T; N]::index]: forward function -let core_array_Array_index - (t idx : Type0) (n : usize) (inst : core_ops_index_Index (slice t) idx) - (a : array t n) (i : idx) : result inst.output = - admit () // TODO - -// [core::array::[T; N]::index_mut]: forward function -let core_array_Array_index_mut - (t idx : Type0) (n : usize) (inst : core_ops_index_IndexMut (slice t) idx) - (a : array t n) (i : idx) : - result (inst.indexInst.output & (inst.indexInst.output -> result (array t n))) = - admit () // TODO - -// Trait implementation: [core::slice::index::private_slice_index::Range] -let core_slice_index_private_slice_index_SealedRangeUsizeInst - : core_slice_index_private_slice_index_Sealed (core_ops_range_Range usize) = () - -// Trait implementation: [core::slice::index::Range] -let core_slice_index_SliceIndexRangeUsizeSliceTInst (t : Type0) : - core_slice_index_SliceIndex (core_ops_range_Range usize) (slice t) = { - sealedInst = core_slice_index_private_slice_index_SealedRangeUsizeInst; - output = slice t; - get = core_slice_index_RangeUsize_get t; - get_mut = core_slice_index_RangeUsize_get_mut t; - get_unchecked = core_slice_index_RangeUsize_get_unchecked t; - get_unchecked_mut = core_slice_index_RangeUsize_get_unchecked_mut t; - index = core_slice_index_RangeUsize_index t; - index_mut = core_slice_index_RangeUsize_index_mut t; -} - -// Trait implementation: [core::slice::index::[T]] -let core_ops_index_IndexSliceTIInst (t idx : Type0) - (inst : core_slice_index_SliceIndex idx (slice t)) : - core_ops_index_Index (slice t) idx = { - output = inst.output; - index = core_slice_index_Slice_index t idx inst; -} - -// Trait implementation: [core::slice::index::[T]] -let core_ops_index_IndexMutSliceTIInst (t idx : Type0) - (inst : core_slice_index_SliceIndex idx (slice t)) : - core_ops_index_IndexMut (slice t) idx = { - indexInst = core_ops_index_IndexSliceTIInst t idx inst; - index_mut = core_slice_index_Slice_index_mut t idx inst; -} - -// Trait implementation: [core::array::[T; N]] -let core_ops_index_IndexArrayInst (t idx : Type0) (n : usize) - (inst : core_ops_index_Index (slice t) idx) : - core_ops_index_Index (array t n) idx = { - output = inst.output; - index = core_array_Array_index t idx n inst; -} - -// Trait implementation: [core::array::[T; N]] -let core_ops_index_IndexMutArrayIInst (t idx : Type0) (n : usize) - (inst : core_ops_index_IndexMut (slice t) idx) : - core_ops_index_IndexMut (array t n) idx = { - indexInst = core_ops_index_IndexArrayInst t idx n inst.indexInst; - index_mut = core_array_Array_index_mut t idx n inst; -} - -// [core::slice::index::usize::get]: forward function -let core_slice_index_usize_get - (t : Type0) : usize → slice t → result (option t) = - admit () // TODO - -// [core::slice::index::usize::get_mut]: forward function -let core_slice_index_usize_get_mut (t : Type0) : - usize → slice t → result (option t & (option t -> result (slice t))) = - admit () // TODO - -// [core::slice::index::usize::get_unchecked]: forward function -let core_slice_index_usize_get_unchecked - (t : Type0) : usize → const_raw_ptr (slice t) → result (const_raw_ptr t) = - admit () // TODO - -// [core::slice::index::usize::get_unchecked_mut]: forward function -let core_slice_index_usize_get_unchecked_mut - (t : Type0) : usize → mut_raw_ptr (slice t) → result (mut_raw_ptr t) = - admit () // TODO - -// [core::slice::index::usize::index]: forward function -let core_slice_index_usize_index (t : Type0) : usize → slice t → result t = - admit () // TODO - -// [core::slice::index::usize::index_mut]: forward function -let core_slice_index_usize_index_mut (t : Type0) : - usize → slice t → result (t & (t -> result (slice t))) = - admit () // TODO - -// Trait implementation: [core::slice::index::private_slice_index::usize] -let core_slice_index_private_slice_index_SealedUsizeInst - : core_slice_index_private_slice_index_Sealed usize = () - -// Trait implementation: [core::slice::index::usize] -let core_slice_index_SliceIndexUsizeSliceTInst (t : Type0) : - core_slice_index_SliceIndex usize (slice t) = { - sealedInst = core_slice_index_private_slice_index_SealedUsizeInst; - output = t; - get = core_slice_index_usize_get t; - get_mut = core_slice_index_usize_get_mut t; - get_unchecked = core_slice_index_usize_get_unchecked t; - get_unchecked_mut = core_slice_index_usize_get_unchecked_mut t; - index = core_slice_index_usize_index t; - index_mut = core_slice_index_usize_index_mut t; -} - -// [alloc::vec::Vec::index]: forward function -let alloc_vec_Vec_index (t idx : Type0) (inst : core_slice_index_SliceIndex idx (slice t)) - (self : alloc_vec_Vec t) (i : idx) : result inst.output = - admit () // TODO - -// [alloc::vec::Vec::index_mut]: forward function -let alloc_vec_Vec_index_mut (t idx : Type0) (inst : core_slice_index_SliceIndex idx (slice t)) - (self : alloc_vec_Vec t) (i : idx) : - result (inst.output & (inst.output -> result (alloc_vec_Vec t))) = - admit () // TODO - -// Trait implementation: [alloc::vec::Vec] -let alloc_vec_Vec_coreopsindexIndexInst (t idx : Type0) - (inst : core_slice_index_SliceIndex idx (slice t)) : - core_ops_index_Index (alloc_vec_Vec t) idx = { - output = inst.output; - index = alloc_vec_Vec_index t idx inst; -} - -// Trait implementation: [alloc::vec::Vec] -let alloc_vec_Vec_coreopsindexIndexMutInst (t idx : Type0) - (inst : core_slice_index_SliceIndex idx (slice t)) : - core_ops_index_IndexMut (alloc_vec_Vec t) idx = { - indexInst = alloc_vec_Vec_coreopsindexIndexInst t idx inst; - index_mut = alloc_vec_Vec_index_mut t idx inst; -} - -(*** Theorems *) - -let alloc_vec_Vec_index_eq (#a : Type0) (v : alloc_vec_Vec a) (i : usize) : - Lemma ( - alloc_vec_Vec_index a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i == - alloc_vec_Vec_index_usize v i) - [SMTPat (alloc_vec_Vec_index a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i)] - = - admit() - -let alloc_vec_Vec_index_mut_eq (#a : Type0) (v : alloc_vec_Vec a) (i : usize) : - Lemma ( - alloc_vec_Vec_index_mut a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i == - alloc_vec_Vec_index_mut_usize v i) - [SMTPat (alloc_vec_Vec_index_mut a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i)] - = - admit() diff --git a/tests/hol4/constants/Holmakefile b/tests/hol4/constants/Holmakefile new file mode 100644 index 00000000..3c4b8973 --- /dev/null +++ b/tests/hol4/constants/Holmakefile @@ -0,0 +1,5 @@ +# This file was automatically generated - modify ../Holmakefile.template instead +INCLUDES = ../../../backends/hol4 + +all: $(DEFAULT_TARGETS) +.PHONY: all diff --git a/tests/hol4/constants/constantsScript.sml b/tests/hol4/constants/constantsScript.sml new file mode 100644 index 00000000..40a319c6 --- /dev/null +++ b/tests/hol4/constants/constantsScript.sml @@ -0,0 +1,217 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [constants] *) +open primitivesLib divDefLib + +val _ = new_theory "constants" + + +(** [constants::X0] *) +Definition x0_body_def: + x0_body : u32 result = Return (int_to_u32 0) +End +Definition x0_c_def: + x0_c : u32 = get_return_value x0_body +End + +(** [constants::X1] *) +Definition x1_body_def: + x1_body : u32 result = Return core_u32_max +End +Definition x1_c_def: + x1_c : u32 = get_return_value x1_body +End + +(** [constants::X2] *) +Definition x2_body_def: + x2_body : u32 result = Return (int_to_u32 3) +End +Definition x2_c_def: + x2_c : u32 = get_return_value x2_body +End + +val incr_fwd_def = Define ‘ + (** [constants::incr]: forward function *) + incr_fwd (n : u32) : u32 result = + u32_add n (int_to_u32 1) +’ + +(** [constants::X3] *) +Definition x3_body_def: + x3_body : u32 result = incr_fwd (int_to_u32 32) +End +Definition x3_c_def: + x3_c : u32 = get_return_value x3_body +End + +val mk_pair0_fwd_def = Define ‘ + (** [constants::mk_pair0]: forward function *) + mk_pair0_fwd (x : u32) (y : u32) : (u32 # u32) result = + Return (x, y) +’ + +Datatype: + (** [constants::Pair] *) + pair_t = <| pair_x : 't1; pair_y : 't2; |> +End + +val mk_pair1_fwd_def = Define ‘ + (** [constants::mk_pair1]: forward function *) + mk_pair1_fwd (x : u32) (y : u32) : (u32, u32) pair_t result = + Return (<| pair_x := x; pair_y := y |>) +’ + +(** [constants::P0] *) +Definition p0_body_def: + p0_body : (u32 # u32) result = mk_pair0_fwd (int_to_u32 0) (int_to_u32 1) +End +Definition p0_c_def: + p0_c : (u32 # u32) = get_return_value p0_body +End + +(** [constants::P1] *) +Definition p1_body_def: + p1_body : (u32, u32) pair_t result = + mk_pair1_fwd (int_to_u32 0) (int_to_u32 1) +End +Definition p1_c_def: + p1_c : (u32, u32) pair_t = get_return_value p1_body +End + +(** [constants::P2] *) +Definition p2_body_def: + p2_body : (u32 # u32) result = Return (int_to_u32 0, int_to_u32 1) +End +Definition p2_c_def: + p2_c : (u32 # u32) = get_return_value p2_body +End + +(** [constants::P3] *) +Definition p3_body_def: + p3_body : (u32, u32) pair_t result = + Return (<| pair_x := (int_to_u32 0); pair_y := (int_to_u32 1) |>) +End +Definition p3_c_def: + p3_c : (u32, u32) pair_t = get_return_value p3_body +End + +Datatype: + (** [constants::Wrap] *) + wrap_t = <| wrap_val : 't; |> +End + +val wrap_new_fwd_def = Define ‘ + (** [constants::Wrap::{0}::new]: forward function *) + wrap_new_fwd (val : 't) : 't wrap_t result = + Return (<| wrap_val := val |>) +’ + +(** [constants::Y] *) +Definition y_body_def: + y_body : i32 wrap_t result = wrap_new_fwd (int_to_i32 2) +End +Definition y_c_def: + y_c : i32 wrap_t = get_return_value y_body +End + +val unwrap_y_fwd_def = Define ‘ + (** [constants::unwrap_y]: forward function *) + unwrap_y_fwd : i32 result = + Return y_c.wrap_val +’ + +(** [constants::YVAL] *) +Definition yval_body_def: + yval_body : i32 result = unwrap_y_fwd +End +Definition yval_c_def: + yval_c : i32 = get_return_value yval_body +End + +(** [constants::get_z1::Z1] *) +Definition get_z1_z1_body_def: + get_z1_z1_body : i32 result = Return (int_to_i32 3) +End +Definition get_z1_z1_c_def: + get_z1_z1_c : i32 = get_return_value get_z1_z1_body +End + +val get_z1_fwd_def = Define ‘ + (** [constants::get_z1]: forward function *) + get_z1_fwd : i32 result = + Return get_z1_z1_c +’ + +val add_fwd_def = Define ‘ + (** [constants::add]: forward function *) + add_fwd (a : i32) (b : i32) : i32 result = + i32_add a b +’ + +(** [constants::Q1] *) +Definition q1_body_def: + q1_body : i32 result = Return (int_to_i32 5) +End +Definition q1_c_def: + q1_c : i32 = get_return_value q1_body +End + +(** [constants::Q2] *) +Definition q2_body_def: + q2_body : i32 result = Return q1_c +End +Definition q2_c_def: + q2_c : i32 = get_return_value q2_body +End + +(** [constants::Q3] *) +Definition q3_body_def: + q3_body : i32 result = add_fwd q2_c (int_to_i32 3) +End +Definition q3_c_def: + q3_c : i32 = get_return_value q3_body +End + +val get_z2_fwd_def = Define ‘ + (** [constants::get_z2]: forward function *) + get_z2_fwd : i32 result = + do + i <- get_z1_fwd; + i0 <- add_fwd i q3_c; + add_fwd q1_c i0 + od +’ + +(** [constants::S1] *) +Definition s1_body_def: + s1_body : u32 result = Return (int_to_u32 6) +End +Definition s1_c_def: + s1_c : u32 = get_return_value s1_body +End + +(** [constants::S2] *) +Definition s2_body_def: + s2_body : u32 result = incr_fwd s1_c +End +Definition s2_c_def: + s2_c : u32 = get_return_value s2_body +End + +(** [constants::S3] *) +Definition s3_body_def: + s3_body : (u32, u32) pair_t result = Return p3_c +End +Definition s3_c_def: + s3_c : (u32, u32) pair_t = get_return_value s3_body +End + +(** [constants::S4] *) +Definition s4_body_def: + s4_body : (u32, u32) pair_t result = + mk_pair1_fwd (int_to_u32 7) (int_to_u32 8) +End +Definition s4_c_def: + s4_c : (u32, u32) pair_t = get_return_value s4_body +End + +val _ = export_theory () diff --git a/tests/hol4/constants/constantsTheory.sig b/tests/hol4/constants/constantsTheory.sig new file mode 100644 index 00000000..287ad5f5 --- /dev/null +++ b/tests/hol4/constants/constantsTheory.sig @@ -0,0 +1,538 @@ +signature constantsTheory = +sig + type thm = Thm.thm + + (* Definitions *) + val add_fwd_def : thm + val get_z1_fwd_def : thm + val get_z1_z1_body_def : thm + val get_z1_z1_c_def : thm + val get_z2_fwd_def : thm + val incr_fwd_def : thm + val mk_pair0_fwd_def : thm + val mk_pair1_fwd_def : thm + val p0_body_def : thm + val p0_c_def : thm + val p1_body_def : thm + val p1_c_def : thm + val p2_body_def : thm + val p2_c_def : thm + val p3_body_def : thm + val p3_c_def : thm + val pair_t_TY_DEF : thm + val pair_t_case_def : thm + val pair_t_pair_x : thm + val pair_t_pair_x_fupd : thm + val pair_t_pair_y : thm + val pair_t_pair_y_fupd : thm + val pair_t_size_def : thm + val q1_body_def : thm + val q1_c_def : thm + val q2_body_def : thm + val q2_c_def : thm + val q3_body_def : thm + val q3_c_def : thm + val s1_body_def : thm + val s1_c_def : thm + val s2_body_def : thm + val s2_c_def : thm + val s3_body_def : thm + val s3_c_def : thm + val s4_body_def : thm + val s4_c_def : thm + val unwrap_y_fwd_def : thm + val wrap_new_fwd_def : thm + val wrap_t_TY_DEF : thm + val wrap_t_case_def : thm + val wrap_t_size_def : thm + val wrap_t_wrap_val : thm + val wrap_t_wrap_val_fupd : thm + val x0_body_def : thm + val x0_c_def : thm + val x1_body_def : thm + val x1_c_def : thm + val x2_body_def : thm + val x2_c_def : thm + val x3_body_def : thm + val x3_c_def : thm + val y_body_def : thm + val y_c_def : thm + val yval_body_def : thm + val yval_c_def : thm + + (* Theorems *) + val EXISTS_pair_t : thm + val EXISTS_wrap_t : thm + val FORALL_pair_t : thm + val FORALL_wrap_t : thm + val datatype_pair_t : thm + val datatype_wrap_t : thm + val pair_t_11 : thm + val pair_t_Axiom : thm + val pair_t_accessors : thm + val pair_t_accfupds : thm + val pair_t_case_cong : thm + val pair_t_case_eq : thm + val pair_t_component_equality : thm + val pair_t_fn_updates : thm + val pair_t_fupdcanon : thm + val pair_t_fupdcanon_comp : thm + val pair_t_fupdfupds : thm + val pair_t_fupdfupds_comp : thm + val pair_t_induction : thm + val pair_t_literal_11 : thm + val pair_t_literal_nchotomy : thm + val pair_t_nchotomy : thm + val pair_t_updates_eq_literal : thm + val wrap_t_11 : thm + val wrap_t_Axiom : thm + val wrap_t_accessors : thm + val wrap_t_accfupds : thm + val wrap_t_case_cong : thm + val wrap_t_case_eq : thm + val wrap_t_component_equality : thm + val wrap_t_fn_updates : thm + val wrap_t_fupdfupds : thm + val wrap_t_fupdfupds_comp : thm + val wrap_t_induction : thm + val wrap_t_literal_11 : thm + val wrap_t_literal_nchotomy : thm + val wrap_t_nchotomy : thm + val wrap_t_updates_eq_literal : thm + + val constants_grammars : type_grammar.grammar * term_grammar.grammar +(* + [divDef] Parent theory of "constants" + + [add_fwd_def] Definition + + ⊢ ∀a b. add_fwd a b = i32_add a b + + [get_z1_fwd_def] Definition + + ⊢ get_z1_fwd = Return get_z1_z1_c + + [get_z1_z1_body_def] Definition + + ⊢ get_z1_z1_body = Return (int_to_i32 3) + + [get_z1_z1_c_def] Definition + + ⊢ get_z1_z1_c = get_return_value get_z1_z1_body + + [get_z2_fwd_def] Definition + + ⊢ get_z2_fwd = + do i <- get_z1_fwd; i0 <- add_fwd i q3_c; add_fwd q1_c i0 od + + [incr_fwd_def] Definition + + ⊢ ∀n. incr_fwd n = u32_add n (int_to_u32 1) + + [mk_pair0_fwd_def] Definition + + ⊢ ∀x y. mk_pair0_fwd x y = Return (x,y) + + [mk_pair1_fwd_def] Definition + + ⊢ ∀x y. mk_pair1_fwd x y = Return <|pair_x := x; pair_y := y|> + + [p0_body_def] Definition + + ⊢ p0_body = mk_pair0_fwd (int_to_u32 0) (int_to_u32 1) + + [p0_c_def] Definition + + ⊢ p0_c = get_return_value p0_body + + [p1_body_def] Definition + + ⊢ p1_body = mk_pair1_fwd (int_to_u32 0) (int_to_u32 1) + + [p1_c_def] Definition + + ⊢ p1_c = get_return_value p1_body + + [p2_body_def] Definition + + ⊢ p2_body = Return (int_to_u32 0,int_to_u32 1) + + [p2_c_def] Definition + + ⊢ p2_c = get_return_value p2_body + + [p3_body_def] Definition + + ⊢ p3_body = Return <|pair_x := int_to_u32 0; pair_y := int_to_u32 1|> + + [p3_c_def] Definition + + ⊢ p3_c = get_return_value p3_body + + [pair_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('pair_t'). + (∀a0'. + (∃a0 a1. + a0' = + (λa0 a1. + ind_type$CONSTR 0 (a0,a1) + (λn. ind_type$BOTTOM)) a0 a1) ⇒ + $var$('pair_t') a0') ⇒ + $var$('pair_t') a0') rep + + [pair_t_case_def] Definition + + ⊢ ∀a0 a1 f. pair_t_CASE (pair_t a0 a1) f = f a0 a1 + + [pair_t_pair_x] Definition + + ⊢ ∀t t0. (pair_t t t0).pair_x = t + + [pair_t_pair_x_fupd] Definition + + ⊢ ∀f t t0. pair_t t t0 with pair_x updated_by f = pair_t (f t) t0 + + [pair_t_pair_y] Definition + + ⊢ ∀t t0. (pair_t t t0).pair_y = t0 + + [pair_t_pair_y_fupd] Definition + + ⊢ ∀f t t0. pair_t t t0 with pair_y updated_by f = pair_t t (f t0) + + [pair_t_size_def] Definition + + ⊢ ∀f f1 a0 a1. pair_t_size f f1 (pair_t a0 a1) = 1 + (f a0 + f1 a1) + + [q1_body_def] Definition + + ⊢ q1_body = Return (int_to_i32 5) + + [q1_c_def] Definition + + ⊢ q1_c = get_return_value q1_body + + [q2_body_def] Definition + + ⊢ q2_body = Return q1_c + + [q2_c_def] Definition + + ⊢ q2_c = get_return_value q2_body + + [q3_body_def] Definition + + ⊢ q3_body = add_fwd q2_c (int_to_i32 3) + + [q3_c_def] Definition + + ⊢ q3_c = get_return_value q3_body + + [s1_body_def] Definition + + ⊢ s1_body = Return (int_to_u32 6) + + [s1_c_def] Definition + + ⊢ s1_c = get_return_value s1_body + + [s2_body_def] Definition + + ⊢ s2_body = incr_fwd s1_c + + [s2_c_def] Definition + + ⊢ s2_c = get_return_value s2_body + + [s3_body_def] Definition + + ⊢ s3_body = Return p3_c + + [s3_c_def] Definition + + ⊢ s3_c = get_return_value s3_body + + [s4_body_def] Definition + + ⊢ s4_body = mk_pair1_fwd (int_to_u32 7) (int_to_u32 8) + + [s4_c_def] Definition + + ⊢ s4_c = get_return_value s4_body + + [unwrap_y_fwd_def] Definition + + ⊢ unwrap_y_fwd = Return y_c.wrap_val + + [wrap_new_fwd_def] Definition + + ⊢ ∀val. wrap_new_fwd val = Return <|wrap_val := val|> + + [wrap_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0. + ∀ $var$('wrap_t'). + (∀a0. + (∃a. a0 = + (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) + a) ⇒ + $var$('wrap_t') a0) ⇒ + $var$('wrap_t') a0) rep + + [wrap_t_case_def] Definition + + ⊢ ∀a f. wrap_t_CASE (wrap_t a) f = f a + + [wrap_t_size_def] Definition + + ⊢ ∀f a. wrap_t_size f (wrap_t a) = 1 + f a + + [wrap_t_wrap_val] Definition + + ⊢ ∀t. (wrap_t t).wrap_val = t + + [wrap_t_wrap_val_fupd] Definition + + ⊢ ∀f t. wrap_t t with wrap_val updated_by f = wrap_t (f t) + + [x0_body_def] Definition + + ⊢ x0_body = Return (int_to_u32 0) + + [x0_c_def] Definition + + ⊢ x0_c = get_return_value x0_body + + [x1_body_def] Definition + + ⊢ x1_body = Return core_u32_max + + [x1_c_def] Definition + + ⊢ x1_c = get_return_value x1_body + + [x2_body_def] Definition + + ⊢ x2_body = Return (int_to_u32 3) + + [x2_c_def] Definition + + ⊢ x2_c = get_return_value x2_body + + [x3_body_def] Definition + + ⊢ x3_body = incr_fwd (int_to_u32 32) + + [x3_c_def] Definition + + ⊢ x3_c = get_return_value x3_body + + [y_body_def] Definition + + ⊢ y_body = wrap_new_fwd (int_to_i32 2) + + [y_c_def] Definition + + ⊢ y_c = get_return_value y_body + + [yval_body_def] Definition + + ⊢ yval_body = unwrap_y_fwd + + [yval_c_def] Definition + + ⊢ yval_c = get_return_value yval_body + + [EXISTS_pair_t] Theorem + + ⊢ ∀P. (∃p. P p) ⇔ ∃t0 t. P <|pair_x := t0; pair_y := t|> + + [EXISTS_wrap_t] Theorem + + ⊢ ∀P. (∃w. P w) ⇔ ∃u. P <|wrap_val := u|> + + [FORALL_pair_t] Theorem + + ⊢ ∀P. (∀p. P p) ⇔ ∀t0 t. P <|pair_x := t0; pair_y := t|> + + [FORALL_wrap_t] Theorem + + ⊢ ∀P. (∀w. P w) ⇔ ∀u. P <|wrap_val := u|> + + [datatype_pair_t] Theorem + + ⊢ DATATYPE (record pair_t pair_x pair_y) + + [datatype_wrap_t] Theorem + + ⊢ DATATYPE (record wrap_t wrap_val) + + [pair_t_11] Theorem + + ⊢ ∀a0 a1 a0' a1'. pair_t a0 a1 = pair_t a0' a1' ⇔ a0 = a0' ∧ a1 = a1' + + [pair_t_Axiom] Theorem + + ⊢ ∀f. ∃fn. ∀a0 a1. fn (pair_t a0 a1) = f a0 a1 + + [pair_t_accessors] Theorem + + ⊢ (∀t t0. (pair_t t t0).pair_x = t) ∧ + ∀t t0. (pair_t t t0).pair_y = t0 + + [pair_t_accfupds] Theorem + + ⊢ (∀p f. (p with pair_y updated_by f).pair_x = p.pair_x) ∧ + (∀p f. (p with pair_x updated_by f).pair_y = p.pair_y) ∧ + (∀p f. (p with pair_x updated_by f).pair_x = f p.pair_x) ∧ + ∀p f. (p with pair_y updated_by f).pair_y = f p.pair_y + + [pair_t_case_cong] Theorem + + ⊢ ∀M M' f. + M = M' ∧ (∀a0 a1. M' = pair_t a0 a1 ⇒ f a0 a1 = f' a0 a1) ⇒ + pair_t_CASE M f = pair_t_CASE M' f' + + [pair_t_case_eq] Theorem + + ⊢ pair_t_CASE x f = v ⇔ ∃t t0. x = pair_t t t0 ∧ f t t0 = v + + [pair_t_component_equality] Theorem + + ⊢ ∀p1 p2. p1 = p2 ⇔ p1.pair_x = p2.pair_x ∧ p1.pair_y = p2.pair_y + + [pair_t_fn_updates] Theorem + + ⊢ (∀f t t0. pair_t t t0 with pair_x updated_by f = pair_t (f t) t0) ∧ + ∀f t t0. pair_t t t0 with pair_y updated_by f = pair_t t (f t0) + + [pair_t_fupdcanon] Theorem + + ⊢ ∀p g f. + p with <|pair_y updated_by f; pair_x updated_by g|> = + p with <|pair_x updated_by g; pair_y updated_by f|> + + [pair_t_fupdcanon_comp] Theorem + + ⊢ (∀g f. + pair_y_fupd f ∘ pair_x_fupd g = pair_x_fupd g ∘ pair_y_fupd f) ∧ + ∀h g f. + pair_y_fupd f ∘ pair_x_fupd g ∘ h = + pair_x_fupd g ∘ pair_y_fupd f ∘ h + + [pair_t_fupdfupds] Theorem + + ⊢ (∀p g f. + p with <|pair_x updated_by f; pair_x updated_by g|> = + p with pair_x updated_by f ∘ g) ∧ + ∀p g f. + p with <|pair_y updated_by f; pair_y updated_by g|> = + p with pair_y updated_by f ∘ g + + [pair_t_fupdfupds_comp] Theorem + + ⊢ ((∀g f. pair_x_fupd f ∘ pair_x_fupd g = pair_x_fupd (f ∘ g)) ∧ + ∀h g f. + pair_x_fupd f ∘ pair_x_fupd g ∘ h = pair_x_fupd (f ∘ g) ∘ h) ∧ + (∀g f. pair_y_fupd f ∘ pair_y_fupd g = pair_y_fupd (f ∘ g)) ∧ + ∀h g f. pair_y_fupd f ∘ pair_y_fupd g ∘ h = pair_y_fupd (f ∘ g) ∘ h + + [pair_t_induction] Theorem + + ⊢ ∀P. (∀t t0. P (pair_t t t0)) ⇒ ∀p. P p + + [pair_t_literal_11] Theorem + + ⊢ ∀t01 t1 t02 t2. + <|pair_x := t01; pair_y := t1|> = <|pair_x := t02; pair_y := t2|> ⇔ + t01 = t02 ∧ t1 = t2 + + [pair_t_literal_nchotomy] Theorem + + ⊢ ∀p. ∃t0 t. p = <|pair_x := t0; pair_y := t|> + + [pair_t_nchotomy] Theorem + + ⊢ ∀pp. ∃t t0. pp = pair_t t t0 + + [pair_t_updates_eq_literal] Theorem + + ⊢ ∀p t0 t. + p with <|pair_x := t0; pair_y := t|> = + <|pair_x := t0; pair_y := t|> + + [wrap_t_11] Theorem + + ⊢ ∀a a'. wrap_t a = wrap_t a' ⇔ a = a' + + [wrap_t_Axiom] Theorem + + ⊢ ∀f. ∃fn. ∀a. fn (wrap_t a) = f a + + [wrap_t_accessors] Theorem + + ⊢ ∀t. (wrap_t t).wrap_val = t + + [wrap_t_accfupds] Theorem + + ⊢ ∀w f. (w with wrap_val updated_by f).wrap_val = f w.wrap_val + + [wrap_t_case_cong] Theorem + + ⊢ ∀M M' f. + M = M' ∧ (∀a. M' = wrap_t a ⇒ f a = f' a) ⇒ + wrap_t_CASE M f = wrap_t_CASE M' f' + + [wrap_t_case_eq] Theorem + + ⊢ wrap_t_CASE x f = v ⇔ ∃t. x = wrap_t t ∧ f t = v + + [wrap_t_component_equality] Theorem + + ⊢ ∀w1 w2. w1 = w2 ⇔ w1.wrap_val = w2.wrap_val + + [wrap_t_fn_updates] Theorem + + ⊢ ∀f t. wrap_t t with wrap_val updated_by f = wrap_t (f t) + + [wrap_t_fupdfupds] Theorem + + ⊢ ∀w g f. + w with <|wrap_val updated_by f; wrap_val updated_by g|> = + w with wrap_val updated_by f ∘ g + + [wrap_t_fupdfupds_comp] Theorem + + ⊢ (∀g f. wrap_val_fupd f ∘ wrap_val_fupd g = wrap_val_fupd (f ∘ g)) ∧ + ∀h g f. + wrap_val_fupd f ∘ wrap_val_fupd g ∘ h = wrap_val_fupd (f ∘ g) ∘ h + + [wrap_t_induction] Theorem + + ⊢ ∀P. (∀t. P (wrap_t t)) ⇒ ∀w. P w + + [wrap_t_literal_11] Theorem + + ⊢ ∀u1 u2. <|wrap_val := u1|> = <|wrap_val := u2|> ⇔ u1 = u2 + + [wrap_t_literal_nchotomy] Theorem + + ⊢ ∀w. ∃u. w = <|wrap_val := u|> + + [wrap_t_nchotomy] Theorem + + ⊢ ∀ww. ∃t. ww = wrap_t t + + [wrap_t_updates_eq_literal] Theorem + + ⊢ ∀w u. w with wrap_val := u = <|wrap_val := u|> + + +*) +end diff --git a/tests/hol4/external/Holmakefile b/tests/hol4/external/Holmakefile new file mode 100644 index 00000000..3c4b8973 --- /dev/null +++ b/tests/hol4/external/Holmakefile @@ -0,0 +1,5 @@ +# This file was automatically generated - modify ../Holmakefile.template instead +INCLUDES = ../../../backends/hol4 + +all: $(DEFAULT_TARGETS) +.PHONY: all diff --git a/tests/hol4/external/external_FunsScript.sml b/tests/hol4/external/external_FunsScript.sml new file mode 100644 index 00000000..f3692ee2 --- /dev/null +++ b/tests/hol4/external/external_FunsScript.sml @@ -0,0 +1,108 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [external]: function definitions *) +open primitivesLib divDefLib +open external_TypesTheory external_OpaqueTheory + +val _ = new_theory "external_Funs" + + +val swap_fwd_def = Define ‘ + (** [external::swap]: forward function *) + swap_fwd (x : 't) (y : 't) (st : state) : (state # unit) result = + do + (st0, _) <- core_mem_swap_fwd x y st; + (st1, _) <- core_mem_swap_back0 x y st st0; + (st2, _) <- core_mem_swap_back1 x y st st1; + Return (st2, ()) + od +’ + +val swap_back_def = Define ‘ + (** [external::swap]: backward function 0 *) + swap_back + (x : 't) (y : 't) (st : state) (st0 : state) : (state # ('t # 't)) result = + do + (st1, _) <- core_mem_swap_fwd x y st; + (st2, x0) <- core_mem_swap_back0 x y st st1; + (_, y0) <- core_mem_swap_back1 x y st st2; + Return (st0, (x0, y0)) + od +’ + +val test_new_non_zero_u32_fwd_def = Define ‘ + (** [external::test_new_non_zero_u32]: forward function *) + test_new_non_zero_u32_fwd + (x : u32) (st : state) : (state # core_num_nonzero_non_zero_u32_t) result = + do + (st0, opt) <- core_num_nonzero_non_zero_u32_new_fwd x st; + core_option_option_unwrap_fwd opt st0 + od +’ + +val test_vec_fwd_def = Define ‘ + (** [external::test_vec]: forward function *) + test_vec_fwd : unit result = + let v = vec_new in do + _ <- vec_push_back v (int_to_u32 0); + Return () + od +’ + +(** Unit test for [external::test_vec] *) +val _ = assert_return (“test_vec_fwd”) + +val custom_swap_fwd_def = Define ‘ + (** [external::custom_swap]: forward function *) + custom_swap_fwd (x : 't) (y : 't) (st : state) : (state # 't) result = + do + (st0, _) <- core_mem_swap_fwd x y st; + (st1, x0) <- core_mem_swap_back0 x y st st0; + (st2, _) <- core_mem_swap_back1 x y st st1; + Return (st2, x0) + od +’ + +val custom_swap_back_def = Define ‘ + (** [external::custom_swap]: backward function 0 *) + custom_swap_back + (x : 't) (y : 't) (st : state) (ret : 't) (st0 : state) : + (state # ('t # 't)) result + = + do + (st1, _) <- core_mem_swap_fwd x y st; + (st2, _) <- core_mem_swap_back0 x y st st1; + (_, y0) <- core_mem_swap_back1 x y st st2; + Return (st0, (ret, y0)) + od +’ + +val test_custom_swap_fwd_def = Define ‘ + (** [external::test_custom_swap]: forward function *) + test_custom_swap_fwd + (x : u32) (y : u32) (st : state) : (state # unit) result = + do + (st0, _) <- custom_swap_fwd x y st; + Return (st0, ()) + od +’ + +val test_custom_swap_back_def = Define ‘ + (** [external::test_custom_swap]: backward function 0 *) + test_custom_swap_back + (x : u32) (y : u32) (st : state) (st0 : state) : + (state # (u32 # u32)) result + = + custom_swap_back x y st (int_to_u32 1) st0 +’ + +val test_swap_non_zero_fwd_def = Define ‘ + (** [external::test_swap_non_zero]: forward function *) + test_swap_non_zero_fwd (x : u32) (st : state) : (state # u32) result = + do + (st0, _) <- swap_fwd x (int_to_u32 0) st; + (st1, (x0, _)) <- swap_back x (int_to_u32 0) st st0; + if x0 = int_to_u32 0 then Fail Failure else Return (st1, x0) + od +’ + +val _ = export_theory () diff --git a/tests/hol4/external/external_FunsTheory.sig b/tests/hol4/external/external_FunsTheory.sig new file mode 100644 index 00000000..490f9d06 --- /dev/null +++ b/tests/hol4/external/external_FunsTheory.sig @@ -0,0 +1,105 @@ +signature external_FunsTheory = +sig + type thm = Thm.thm + + (* Definitions *) + val custom_swap_back_def : thm + val custom_swap_fwd_def : thm + val swap_back_def : thm + val swap_fwd_def : thm + val test_custom_swap_back_def : thm + val test_custom_swap_fwd_def : thm + val test_new_non_zero_u32_fwd_def : thm + val test_swap_non_zero_fwd_def : thm + val test_vec_fwd_def : thm + + val external_Funs_grammars : type_grammar.grammar * term_grammar.grammar +(* + [external_Opaque] Parent theory of "external_Funs" + + [custom_swap_back_def] Definition + + ⊢ ∀x y st ret st0. + custom_swap_back x y st ret st0 = + do + (st1,_) <- core_mem_swap_fwd x y st; + (st2,_) <- core_mem_swap_back0 x y st st1; + (_,y0) <- core_mem_swap_back1 x y st st2; + Return (st0,ret,y0) + od + + [custom_swap_fwd_def] Definition + + ⊢ ∀x y st. + custom_swap_fwd x y st = + do + (st0,_) <- core_mem_swap_fwd x y st; + (st1,x0) <- core_mem_swap_back0 x y st st0; + (st2,_) <- core_mem_swap_back1 x y st st1; + Return (st2,x0) + od + + [swap_back_def] Definition + + ⊢ ∀x y st st0. + swap_back x y st st0 = + do + (st1,_) <- core_mem_swap_fwd x y st; + (st2,x0) <- core_mem_swap_back0 x y st st1; + (_,y0) <- core_mem_swap_back1 x y st st2; + Return (st0,x0,y0) + od + + [swap_fwd_def] Definition + + ⊢ ∀x y st. + swap_fwd x y st = + do + (st0,_) <- core_mem_swap_fwd x y st; + (st1,_) <- core_mem_swap_back0 x y st st0; + (st2,_) <- core_mem_swap_back1 x y st st1; + Return (st2,()) + od + + [test_custom_swap_back_def] Definition + + ⊢ ∀x y st st0. + test_custom_swap_back x y st st0 = + custom_swap_back x y st (int_to_u32 1) st0 + + [test_custom_swap_fwd_def] Definition + + ⊢ ∀x y st. + test_custom_swap_fwd x y st = + do (st0,_) <- custom_swap_fwd x y st; Return (st0,()) od + + [test_new_non_zero_u32_fwd_def] Definition + + ⊢ ∀x st. + test_new_non_zero_u32_fwd x st = + do + (st0,opt) <- core_num_nonzero_non_zero_u32_new_fwd x st; + core_option_option_unwrap_fwd opt st0 + od + + [test_swap_non_zero_fwd_def] Definition + + ⊢ ∀x st. + test_swap_non_zero_fwd x st = + do + (st0,_) <- swap_fwd x (int_to_u32 0) st; + (st1,x0,_) <- swap_back x (int_to_u32 0) st st0; + if x0 = int_to_u32 0 then Fail Failure else Return (st1,x0) + od + + [test_vec_fwd_def] Definition + + ⊢ test_vec_fwd = + (let + v = vec_new + in + monad_ignore_bind (vec_push_back v (int_to_u32 0)) (Return ())) + + +*) +end diff --git a/tests/hol4/external/external_OpaqueScript.sml b/tests/hol4/external/external_OpaqueScript.sml new file mode 100644 index 00000000..b5a6d91d --- /dev/null +++ b/tests/hol4/external/external_OpaqueScript.sml @@ -0,0 +1,25 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [external]: external function declarations *) +open primitivesLib divDefLib +open external_TypesTheory + +val _ = new_theory "external_Opaque" + + +(** [core::mem::swap]: forward function *)val _ = new_constant ("core_mem_swap_fwd", + “:'t -> 't -> state -> (state # unit) result”) + +(** [core::mem::swap]: backward function 0 *)val _ = new_constant ("core_mem_swap_back0", + “:'t -> 't -> state -> state -> (state # 't) result”) + +(** [core::mem::swap]: backward function 1 *)val _ = new_constant ("core_mem_swap_back1", + “:'t -> 't -> state -> state -> (state # 't) result”) + +(** [core::num::nonzero::NonZeroU32::{14}::new]: forward function *)val _ = new_constant ("core_num_nonzero_non_zero_u32_new_fwd", + “:u32 -> state -> (state # core_num_nonzero_non_zero_u32_t option) + result”) + +(** [core::option::Option::{0}::unwrap]: forward function *)val _ = new_constant ("core_option_option_unwrap_fwd", + “:'t option -> state -> (state # 't) result”) + +val _ = export_theory () diff --git a/tests/hol4/external/external_OpaqueTheory.sig b/tests/hol4/external/external_OpaqueTheory.sig new file mode 100644 index 00000000..7cd7a08c --- /dev/null +++ b/tests/hol4/external/external_OpaqueTheory.sig @@ -0,0 +1,11 @@ +signature external_OpaqueTheory = +sig + type thm = Thm.thm + + val external_Opaque_grammars : type_grammar.grammar * term_grammar.grammar +(* + [external_Types] Parent theory of "external_Opaque" + + +*) +end diff --git a/tests/hol4/external/external_TypesScript.sml b/tests/hol4/external/external_TypesScript.sml new file mode 100644 index 00000000..d290c3f6 --- /dev/null +++ b/tests/hol4/external/external_TypesScript.sml @@ -0,0 +1,13 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [external]: type definitions *) +open primitivesLib divDefLib + +val _ = new_theory "external_Types" + + +val _ = new_type ("core_num_nonzero_non_zero_u32_t", 0) + +(** The state type used in the state-error monad *) +val _ = new_type ("state", 0) + +val _ = export_theory () diff --git a/tests/hol4/external/external_TypesTheory.sig b/tests/hol4/external/external_TypesTheory.sig new file mode 100644 index 00000000..17e2e8e3 --- /dev/null +++ b/tests/hol4/external/external_TypesTheory.sig @@ -0,0 +1,11 @@ +signature external_TypesTheory = +sig + type thm = Thm.thm + + val external_Types_grammars : type_grammar.grammar * term_grammar.grammar +(* + [divDef] Parent theory of "external_Types" + + +*) +end diff --git a/tests/hol4/hashmap_main/Holmakefile b/tests/hol4/hashmap_main/Holmakefile new file mode 100644 index 00000000..3c4b8973 --- /dev/null +++ b/tests/hol4/hashmap_main/Holmakefile @@ -0,0 +1,5 @@ +# This file was automatically generated - modify ../Holmakefile.template instead +INCLUDES = ../../../backends/hol4 + +all: $(DEFAULT_TARGETS) +.PHONY: all diff --git a/tests/hol4/hashmap_main/hashmapMain_FunsScript.sml b/tests/hol4/hashmap_main/hashmapMain_FunsScript.sml new file mode 100644 index 00000000..c1e30aa6 --- /dev/null +++ b/tests/hol4/hashmap_main/hashmapMain_FunsScript.sml @@ -0,0 +1,647 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: function definitions *) +open primitivesLib divDefLib +open hashmapMain_TypesTheory hashmapMain_OpaqueTheory + +val _ = new_theory "hashmapMain_Funs" + + +val hashmap_hash_key_fwd_def = Define ‘ + (** [hashmap_main::hashmap::hash_key]: forward function *) + hashmap_hash_key_fwd (k : usize) : usize result = + Return k +’ + +val [hashmap_hash_map_allocate_slots_loop_fwd_def] = DefineDiv ‘ + (** [hashmap_main::hashmap::HashMap::{0}::allocate_slots]: loop 0: forward function *) + hashmap_hash_map_allocate_slots_loop_fwd + (slots : 't hashmap_list_t vec) (n : usize) : + 't hashmap_list_t vec result + = + if usize_gt n (int_to_usize 0) + then ( + do + slots0 <- vec_push_back slots HashmapListNil; + n0 <- usize_sub n (int_to_usize 1); + hashmap_hash_map_allocate_slots_loop_fwd slots0 n0 + od) + else Return slots +’ + +val hashmap_hash_map_allocate_slots_fwd_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::allocate_slots]: forward function *) + hashmap_hash_map_allocate_slots_fwd + (slots : 't hashmap_list_t vec) (n : usize) : + 't hashmap_list_t vec result + = + hashmap_hash_map_allocate_slots_loop_fwd slots n +’ + +val hashmap_hash_map_new_with_capacity_fwd_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::new_with_capacity]: forward function *) + hashmap_hash_map_new_with_capacity_fwd + (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : + 't hashmap_hash_map_t result + = + let v = vec_new in + do + slots <- hashmap_hash_map_allocate_slots_fwd v capacity; + i <- usize_mul capacity max_load_dividend; + i0 <- usize_div i max_load_divisor; + Return + (<| + hashmap_hash_map_num_entries := (int_to_usize 0); + hashmap_hash_map_max_load_factor := + (max_load_dividend, max_load_divisor); + hashmap_hash_map_max_load := i0; + hashmap_hash_map_slots := slots + |>) + od +’ + +val hashmap_hash_map_new_fwd_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::new]: forward function *) + hashmap_hash_map_new_fwd : 't hashmap_hash_map_t result = + hashmap_hash_map_new_with_capacity_fwd (int_to_usize 32) (int_to_usize 4) + (int_to_usize 5) +’ + +val [hashmap_hash_map_clear_loop_fwd_back_def] = DefineDiv ‘ + (** [hashmap_main::hashmap::HashMap::{0}::clear]: loop 0: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + hashmap_hash_map_clear_loop_fwd_back + (slots : 't hashmap_list_t vec) (i : usize) : + 't hashmap_list_t vec result + = + let i0 = vec_len slots in + if usize_lt i i0 + then ( + do + i1 <- usize_add i (int_to_usize 1); + slots0 <- vec_index_mut_back slots i HashmapListNil; + hashmap_hash_map_clear_loop_fwd_back slots0 i1 + od) + else Return slots +’ + +val hashmap_hash_map_clear_fwd_back_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::clear]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + hashmap_hash_map_clear_fwd_back + (self : 't hashmap_hash_map_t) : 't hashmap_hash_map_t result = + do + v <- + hashmap_hash_map_clear_loop_fwd_back self.hashmap_hash_map_slots + (int_to_usize 0); + Return + (self + with + <| + hashmap_hash_map_num_entries := (int_to_usize 0); + hashmap_hash_map_slots := v + |>) + od +’ + +val hashmap_hash_map_len_fwd_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::len]: forward function *) + hashmap_hash_map_len_fwd (self : 't hashmap_hash_map_t) : usize result = + Return self.hashmap_hash_map_num_entries +’ + +val [hashmap_hash_map_insert_in_list_loop_fwd_def] = DefineDiv ‘ + (** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: loop 0: forward function *) + hashmap_hash_map_insert_in_list_loop_fwd + (key : usize) (value : 't) (ls : 't hashmap_list_t) : bool result = + (case ls of + | HashmapListCons ckey cvalue tl => + if ckey = key + then Return F + else hashmap_hash_map_insert_in_list_loop_fwd key value tl + | HashmapListNil => Return T) +’ + +val hashmap_hash_map_insert_in_list_fwd_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: forward function *) + hashmap_hash_map_insert_in_list_fwd + (key : usize) (value : 't) (ls : 't hashmap_list_t) : bool result = + hashmap_hash_map_insert_in_list_loop_fwd key value ls +’ + +val [hashmap_hash_map_insert_in_list_loop_back_def] = DefineDiv ‘ + (** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: loop 0: backward function 0 *) + hashmap_hash_map_insert_in_list_loop_back + (key : usize) (value : 't) (ls : 't hashmap_list_t) : + 't hashmap_list_t result + = + (case ls of + | HashmapListCons ckey cvalue tl => + if ckey = key + then Return (HashmapListCons ckey value tl) + else ( + do + tl0 <- hashmap_hash_map_insert_in_list_loop_back key value tl; + Return (HashmapListCons ckey cvalue tl0) + od) + | HashmapListNil => + let l = HashmapListNil in Return (HashmapListCons key value l)) +’ + +val hashmap_hash_map_insert_in_list_back_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: backward function 0 *) + hashmap_hash_map_insert_in_list_back + (key : usize) (value : 't) (ls : 't hashmap_list_t) : + 't hashmap_list_t result + = + hashmap_hash_map_insert_in_list_loop_back key value ls +’ + +val hashmap_hash_map_insert_no_resize_fwd_back_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::insert_no_resize]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + hashmap_hash_map_insert_no_resize_fwd_back + (self : 't hashmap_hash_map_t) (key : usize) (value : 't) : + 't hashmap_hash_map_t result + = + do + hash <- hashmap_hash_key_fwd key; + let i = vec_len self.hashmap_hash_map_slots in + do + hash_mod <- usize_rem hash i; + l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; + inserted <- hashmap_hash_map_insert_in_list_fwd key value l; + if inserted + then ( + do + i0 <- usize_add self.hashmap_hash_map_num_entries (int_to_usize 1); + l0 <- hashmap_hash_map_insert_in_list_back key value l; + v <- vec_index_mut_back self.hashmap_hash_map_slots hash_mod l0; + Return + (self + with + <| + hashmap_hash_map_num_entries := i0; hashmap_hash_map_slots := v + |>) + od) + else ( + do + l0 <- hashmap_hash_map_insert_in_list_back key value l; + v <- vec_index_mut_back self.hashmap_hash_map_slots hash_mod l0; + Return (self with <| hashmap_hash_map_slots := v |>) + od) + od + od +’ + +val [hashmap_hash_map_move_elements_from_list_loop_fwd_back_def] = DefineDiv ‘ + (** [hashmap_main::hashmap::HashMap::{0}::move_elements_from_list]: loop 0: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + hashmap_hash_map_move_elements_from_list_loop_fwd_back + (ntable : 't hashmap_hash_map_t) (ls : 't hashmap_list_t) : + 't hashmap_hash_map_t result + = + (case ls of + | HashmapListCons k v tl => + do + ntable0 <- hashmap_hash_map_insert_no_resize_fwd_back ntable k v; + hashmap_hash_map_move_elements_from_list_loop_fwd_back ntable0 tl + od + | HashmapListNil => Return ntable) +’ + +val hashmap_hash_map_move_elements_from_list_fwd_back_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::move_elements_from_list]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + hashmap_hash_map_move_elements_from_list_fwd_back + (ntable : 't hashmap_hash_map_t) (ls : 't hashmap_list_t) : + 't hashmap_hash_map_t result + = + hashmap_hash_map_move_elements_from_list_loop_fwd_back ntable ls +’ + +val [hashmap_hash_map_move_elements_loop_fwd_back_def] = DefineDiv ‘ + (** [hashmap_main::hashmap::HashMap::{0}::move_elements]: loop 0: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + hashmap_hash_map_move_elements_loop_fwd_back + (ntable : 't hashmap_hash_map_t) (slots : 't hashmap_list_t vec) + (i : usize) : + ('t hashmap_hash_map_t # 't hashmap_list_t vec) result + = + let i0 = vec_len slots in + if usize_lt i i0 + then ( + do + l <- vec_index_mut_fwd slots i; + let ls = mem_replace_fwd l HashmapListNil in + do + ntable0 <- hashmap_hash_map_move_elements_from_list_fwd_back ntable ls; + i1 <- usize_add i (int_to_usize 1); + let l0 = mem_replace_back l HashmapListNil in + do + slots0 <- vec_index_mut_back slots i l0; + hashmap_hash_map_move_elements_loop_fwd_back ntable0 slots0 i1 + od + od + od) + else Return (ntable, slots) +’ + +val hashmap_hash_map_move_elements_fwd_back_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::move_elements]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + hashmap_hash_map_move_elements_fwd_back + (ntable : 't hashmap_hash_map_t) (slots : 't hashmap_list_t vec) + (i : usize) : + ('t hashmap_hash_map_t # 't hashmap_list_t vec) result + = + hashmap_hash_map_move_elements_loop_fwd_back ntable slots i +’ + +val hashmap_hash_map_try_resize_fwd_back_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::try_resize]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + hashmap_hash_map_try_resize_fwd_back + (self : 't hashmap_hash_map_t) : 't hashmap_hash_map_t result = + do + max_usize <- mk_usize (u32_to_int core_u32_max); + let capacity = vec_len self.hashmap_hash_map_slots in + do + n1 <- usize_div max_usize (int_to_usize 2); + let (i, i0) = self.hashmap_hash_map_max_load_factor in + do + i1 <- usize_div n1 i; + if usize_le capacity i1 + then ( + do + i2 <- usize_mul capacity (int_to_usize 2); + ntable <- hashmap_hash_map_new_with_capacity_fwd i2 i i0; + (ntable0, _) <- + hashmap_hash_map_move_elements_fwd_back ntable + self.hashmap_hash_map_slots (int_to_usize 0); + Return + (ntable0 + with + <| + hashmap_hash_map_num_entries := self.hashmap_hash_map_num_entries; + hashmap_hash_map_max_load_factor := (i, i0) + |>) + od) + else Return (self with <| hashmap_hash_map_max_load_factor := (i, i0) |>) + od + od + od +’ + +val hashmap_hash_map_insert_fwd_back_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::insert]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + hashmap_hash_map_insert_fwd_back + (self : 't hashmap_hash_map_t) (key : usize) (value : 't) : + 't hashmap_hash_map_t result + = + do + self0 <- hashmap_hash_map_insert_no_resize_fwd_back self key value; + i <- hashmap_hash_map_len_fwd self0; + if usize_gt i self0.hashmap_hash_map_max_load + then hashmap_hash_map_try_resize_fwd_back self0 + else Return self0 + od +’ + +val [hashmap_hash_map_contains_key_in_list_loop_fwd_def] = DefineDiv ‘ + (** [hashmap_main::hashmap::HashMap::{0}::contains_key_in_list]: loop 0: forward function *) + hashmap_hash_map_contains_key_in_list_loop_fwd + (key : usize) (ls : 't hashmap_list_t) : bool result = + (case ls of + | HashmapListCons ckey t tl => + if ckey = key + then Return T + else hashmap_hash_map_contains_key_in_list_loop_fwd key tl + | HashmapListNil => Return F) +’ + +val hashmap_hash_map_contains_key_in_list_fwd_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::contains_key_in_list]: forward function *) + hashmap_hash_map_contains_key_in_list_fwd + (key : usize) (ls : 't hashmap_list_t) : bool result = + hashmap_hash_map_contains_key_in_list_loop_fwd key ls +’ + +val hashmap_hash_map_contains_key_fwd_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::contains_key]: forward function *) + hashmap_hash_map_contains_key_fwd + (self : 't hashmap_hash_map_t) (key : usize) : bool result = + do + hash <- hashmap_hash_key_fwd key; + let i = vec_len self.hashmap_hash_map_slots in + do + hash_mod <- usize_rem hash i; + l <- vec_index_fwd self.hashmap_hash_map_slots hash_mod; + hashmap_hash_map_contains_key_in_list_fwd key l + od + od +’ + +val [hashmap_hash_map_get_in_list_loop_fwd_def] = DefineDiv ‘ + (** [hashmap_main::hashmap::HashMap::{0}::get_in_list]: loop 0: forward function *) + hashmap_hash_map_get_in_list_loop_fwd + (key : usize) (ls : 't hashmap_list_t) : 't result = + (case ls of + | HashmapListCons ckey cvalue tl => + if ckey = key + then Return cvalue + else hashmap_hash_map_get_in_list_loop_fwd key tl + | HashmapListNil => Fail Failure) +’ + +val hashmap_hash_map_get_in_list_fwd_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::get_in_list]: forward function *) + hashmap_hash_map_get_in_list_fwd + (key : usize) (ls : 't hashmap_list_t) : 't result = + hashmap_hash_map_get_in_list_loop_fwd key ls +’ + +val hashmap_hash_map_get_fwd_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::get]: forward function *) + hashmap_hash_map_get_fwd + (self : 't hashmap_hash_map_t) (key : usize) : 't result = + do + hash <- hashmap_hash_key_fwd key; + let i = vec_len self.hashmap_hash_map_slots in + do + hash_mod <- usize_rem hash i; + l <- vec_index_fwd self.hashmap_hash_map_slots hash_mod; + hashmap_hash_map_get_in_list_fwd key l + od + od +’ + +val [hashmap_hash_map_get_mut_in_list_loop_fwd_def] = DefineDiv ‘ + (** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: loop 0: forward function *) + hashmap_hash_map_get_mut_in_list_loop_fwd + (ls : 't hashmap_list_t) (key : usize) : 't result = + (case ls of + | HashmapListCons ckey cvalue tl => + if ckey = key + then Return cvalue + else hashmap_hash_map_get_mut_in_list_loop_fwd tl key + | HashmapListNil => Fail Failure) +’ + +val hashmap_hash_map_get_mut_in_list_fwd_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: forward function *) + hashmap_hash_map_get_mut_in_list_fwd + (ls : 't hashmap_list_t) (key : usize) : 't result = + hashmap_hash_map_get_mut_in_list_loop_fwd ls key +’ + +val [hashmap_hash_map_get_mut_in_list_loop_back_def] = DefineDiv ‘ + (** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: loop 0: backward function 0 *) + hashmap_hash_map_get_mut_in_list_loop_back + (ls : 't hashmap_list_t) (key : usize) (ret : 't) : + 't hashmap_list_t result + = + (case ls of + | HashmapListCons ckey cvalue tl => + if ckey = key + then Return (HashmapListCons ckey ret tl) + else ( + do + tl0 <- hashmap_hash_map_get_mut_in_list_loop_back tl key ret; + Return (HashmapListCons ckey cvalue tl0) + od) + | HashmapListNil => Fail Failure) +’ + +val hashmap_hash_map_get_mut_in_list_back_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: backward function 0 *) + hashmap_hash_map_get_mut_in_list_back + (ls : 't hashmap_list_t) (key : usize) (ret : 't) : + 't hashmap_list_t result + = + hashmap_hash_map_get_mut_in_list_loop_back ls key ret +’ + +val hashmap_hash_map_get_mut_fwd_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::get_mut]: forward function *) + hashmap_hash_map_get_mut_fwd + (self : 't hashmap_hash_map_t) (key : usize) : 't result = + do + hash <- hashmap_hash_key_fwd key; + let i = vec_len self.hashmap_hash_map_slots in + do + hash_mod <- usize_rem hash i; + l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; + hashmap_hash_map_get_mut_in_list_fwd l key + od + od +’ + +val hashmap_hash_map_get_mut_back_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::get_mut]: backward function 0 *) + hashmap_hash_map_get_mut_back + (self : 't hashmap_hash_map_t) (key : usize) (ret : 't) : + 't hashmap_hash_map_t result + = + do + hash <- hashmap_hash_key_fwd key; + let i = vec_len self.hashmap_hash_map_slots in + do + hash_mod <- usize_rem hash i; + l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; + l0 <- hashmap_hash_map_get_mut_in_list_back l key ret; + v <- vec_index_mut_back self.hashmap_hash_map_slots hash_mod l0; + Return (self with <| hashmap_hash_map_slots := v |>) + od + od +’ + +val [hashmap_hash_map_remove_from_list_loop_fwd_def] = DefineDiv ‘ + (** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: loop 0: forward function *) + hashmap_hash_map_remove_from_list_loop_fwd + (key : usize) (ls : 't hashmap_list_t) : 't option result = + (case ls of + | HashmapListCons ckey t tl => + if ckey = key + then + let mv_ls = mem_replace_fwd (HashmapListCons ckey t tl) HashmapListNil + in + (case mv_ls of + | HashmapListCons i cvalue tl0 => Return (SOME cvalue) + | HashmapListNil => Fail Failure) + else hashmap_hash_map_remove_from_list_loop_fwd key tl + | HashmapListNil => Return NONE) +’ + +val hashmap_hash_map_remove_from_list_fwd_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: forward function *) + hashmap_hash_map_remove_from_list_fwd + (key : usize) (ls : 't hashmap_list_t) : 't option result = + hashmap_hash_map_remove_from_list_loop_fwd key ls +’ + +val [hashmap_hash_map_remove_from_list_loop_back_def] = DefineDiv ‘ + (** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: loop 0: backward function 1 *) + hashmap_hash_map_remove_from_list_loop_back + (key : usize) (ls : 't hashmap_list_t) : 't hashmap_list_t result = + (case ls of + | HashmapListCons ckey t tl => + if ckey = key + then + let mv_ls = mem_replace_fwd (HashmapListCons ckey t tl) HashmapListNil + in + (case mv_ls of + | HashmapListCons i cvalue tl0 => Return tl0 + | HashmapListNil => Fail Failure) + else ( + do + tl0 <- hashmap_hash_map_remove_from_list_loop_back key tl; + Return (HashmapListCons ckey t tl0) + od) + | HashmapListNil => Return HashmapListNil) +’ + +val hashmap_hash_map_remove_from_list_back_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: backward function 1 *) + hashmap_hash_map_remove_from_list_back + (key : usize) (ls : 't hashmap_list_t) : 't hashmap_list_t result = + hashmap_hash_map_remove_from_list_loop_back key ls +’ + +val hashmap_hash_map_remove_fwd_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::remove]: forward function *) + hashmap_hash_map_remove_fwd + (self : 't hashmap_hash_map_t) (key : usize) : 't option result = + do + hash <- hashmap_hash_key_fwd key; + let i = vec_len self.hashmap_hash_map_slots in + do + hash_mod <- usize_rem hash i; + l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; + x <- hashmap_hash_map_remove_from_list_fwd key l; + (case x of + | NONE => Return NONE + | SOME x0 => + do + _ <- usize_sub self.hashmap_hash_map_num_entries (int_to_usize 1); + Return (SOME x0) + od) + od + od +’ + +val hashmap_hash_map_remove_back_def = Define ‘ + (** [hashmap_main::hashmap::HashMap::{0}::remove]: backward function 0 *) + hashmap_hash_map_remove_back + (self : 't hashmap_hash_map_t) (key : usize) : + 't hashmap_hash_map_t result + = + do + hash <- hashmap_hash_key_fwd key; + let i = vec_len self.hashmap_hash_map_slots in + do + hash_mod <- usize_rem hash i; + l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; + x <- hashmap_hash_map_remove_from_list_fwd key l; + (case x of + | NONE => + do + l0 <- hashmap_hash_map_remove_from_list_back key l; + v <- vec_index_mut_back self.hashmap_hash_map_slots hash_mod l0; + Return (self with <| hashmap_hash_map_slots := v |>) + od + | SOME x0 => + do + i0 <- usize_sub self.hashmap_hash_map_num_entries (int_to_usize 1); + l0 <- hashmap_hash_map_remove_from_list_back key l; + v <- vec_index_mut_back self.hashmap_hash_map_slots hash_mod l0; + Return + (self + with + <| + hashmap_hash_map_num_entries := i0; hashmap_hash_map_slots := v + |>) + od) + od + od +’ + +val hashmap_test1_fwd_def = Define ‘ + (** [hashmap_main::hashmap::test1]: forward function *) + hashmap_test1_fwd : unit result = + do + hm <- hashmap_hash_map_new_fwd; + hm0 <- + hashmap_hash_map_insert_fwd_back hm (int_to_usize 0) (int_to_u64 42); + hm1 <- + hashmap_hash_map_insert_fwd_back hm0 (int_to_usize 128) (int_to_u64 18); + hm2 <- + hashmap_hash_map_insert_fwd_back hm1 (int_to_usize 1024) (int_to_u64 138); + hm3 <- + hashmap_hash_map_insert_fwd_back hm2 (int_to_usize 1056) (int_to_u64 256); + i <- hashmap_hash_map_get_fwd hm3 (int_to_usize 128); + if ~ (i = int_to_u64 18) + then Fail Failure + else ( + do + hm4 <- + hashmap_hash_map_get_mut_back hm3 (int_to_usize 1024) (int_to_u64 56); + i0 <- hashmap_hash_map_get_fwd hm4 (int_to_usize 1024); + if ~ (i0 = int_to_u64 56) + then Fail Failure + else ( + do + x <- hashmap_hash_map_remove_fwd hm4 (int_to_usize 1024); + (case x of + | NONE => Fail Failure + | SOME x0 => + if ~ (x0 = int_to_u64 56) + then Fail Failure + else ( + do + hm5 <- hashmap_hash_map_remove_back hm4 (int_to_usize 1024); + i1 <- hashmap_hash_map_get_fwd hm5 (int_to_usize 0); + if ~ (i1 = int_to_u64 42) + then Fail Failure + else ( + do + i2 <- hashmap_hash_map_get_fwd hm5 (int_to_usize 128); + if ~ (i2 = int_to_u64 18) + then Fail Failure + else ( + do + i3 <- hashmap_hash_map_get_fwd hm5 (int_to_usize 1056); + if ~ (i3 = int_to_u64 256) then Fail Failure else Return () + od) + od) + od)) + od) + od) + od +’ + +(** Unit test for [hashmap_main::hashmap::test1] *) +val _ = assert_return (“hashmap_test1_fwd”) + +val insert_on_disk_fwd_def = Define ‘ + (** [hashmap_main::insert_on_disk]: forward function *) + insert_on_disk_fwd + (key : usize) (value : u64) (st : state) : (state # unit) result = + do + (st0, hm) <- hashmap_utils_deserialize_fwd st; + hm0 <- hashmap_hash_map_insert_fwd_back hm key value; + (st1, _) <- hashmap_utils_serialize_fwd hm0 st0; + Return (st1, ()) + od +’ + +val main_fwd_def = Define ‘ + (** [hashmap_main::main]: forward function *) + main_fwd : unit result = + Return () +’ + +(** Unit test for [hashmap_main::main] *) +val _ = assert_return (“main_fwd”) + +val _ = export_theory () diff --git a/tests/hol4/hashmap_main/hashmapMain_FunsTheory.sig b/tests/hol4/hashmap_main/hashmapMain_FunsTheory.sig new file mode 100644 index 00000000..d4e43d9a --- /dev/null +++ b/tests/hol4/hashmap_main/hashmapMain_FunsTheory.sig @@ -0,0 +1,598 @@ +signature hashmapMain_FunsTheory = +sig + type thm = Thm.thm + + (* Definitions *) + val hashmap_hash_key_fwd_def : thm + val hashmap_hash_map_allocate_slots_fwd_def : thm + val hashmap_hash_map_allocate_slots_loop_fwd_def : thm + val hashmap_hash_map_clear_fwd_back_def : thm + val hashmap_hash_map_clear_loop_fwd_back_def : thm + val hashmap_hash_map_contains_key_fwd_def : thm + val hashmap_hash_map_contains_key_in_list_fwd_def : thm + val hashmap_hash_map_contains_key_in_list_loop_fwd_def : thm + val hashmap_hash_map_get_fwd_def : thm + val hashmap_hash_map_get_in_list_fwd_def : thm + val hashmap_hash_map_get_in_list_loop_fwd_def : thm + val hashmap_hash_map_get_mut_back_def : thm + val hashmap_hash_map_get_mut_fwd_def : thm + val hashmap_hash_map_get_mut_in_list_back_def : thm + val hashmap_hash_map_get_mut_in_list_fwd_def : thm + val hashmap_hash_map_get_mut_in_list_loop_back_def : thm + val hashmap_hash_map_get_mut_in_list_loop_fwd_def : thm + val hashmap_hash_map_insert_fwd_back_def : thm + val hashmap_hash_map_insert_in_list_back_def : thm + val hashmap_hash_map_insert_in_list_fwd_def : thm + val hashmap_hash_map_insert_in_list_loop_back_def : thm + val hashmap_hash_map_insert_in_list_loop_fwd_def : thm + val hashmap_hash_map_insert_no_resize_fwd_back_def : thm + val hashmap_hash_map_len_fwd_def : thm + val hashmap_hash_map_move_elements_from_list_fwd_back_def : thm + val hashmap_hash_map_move_elements_from_list_loop_fwd_back_def : thm + val hashmap_hash_map_move_elements_fwd_back_def : thm + val hashmap_hash_map_move_elements_loop_fwd_back_def : thm + val hashmap_hash_map_new_fwd_def : thm + val hashmap_hash_map_new_with_capacity_fwd_def : thm + val hashmap_hash_map_remove_back_def : thm + val hashmap_hash_map_remove_from_list_back_def : thm + val hashmap_hash_map_remove_from_list_fwd_def : thm + val hashmap_hash_map_remove_from_list_loop_back_def : thm + val hashmap_hash_map_remove_from_list_loop_fwd_def : thm + val hashmap_hash_map_remove_fwd_def : thm + val hashmap_hash_map_try_resize_fwd_back_def : thm + val hashmap_test1_fwd_def : thm + val insert_on_disk_fwd_def : thm + val main_fwd_def : thm + + val hashmapMain_Funs_grammars : type_grammar.grammar * term_grammar.grammar +(* + [hashmapMain_Opaque] Parent theory of "hashmapMain_Funs" + + [hashmap_hash_key_fwd_def] Definition + + ⊢ ∀k. hashmap_hash_key_fwd k = Return k + + [hashmap_hash_map_allocate_slots_fwd_def] Definition + + ⊢ ∀slots n. + hashmap_hash_map_allocate_slots_fwd slots n = + hashmap_hash_map_allocate_slots_loop_fwd slots n + + [hashmap_hash_map_allocate_slots_loop_fwd_def] Definition + + ⊢ ∀slots n. + hashmap_hash_map_allocate_slots_loop_fwd slots n = + if usize_gt n (int_to_usize 0) then + do + slots0 <- vec_push_back slots HashmapListNil; + n0 <- usize_sub n (int_to_usize 1); + hashmap_hash_map_allocate_slots_loop_fwd slots0 n0 + od + else Return slots + + [hashmap_hash_map_clear_fwd_back_def] Definition + + ⊢ ∀self. + hashmap_hash_map_clear_fwd_back self = + do + v <- + hashmap_hash_map_clear_loop_fwd_back + self.hashmap_hash_map_slots (int_to_usize 0); + Return + (self with + <|hashmap_hash_map_num_entries := int_to_usize 0; + hashmap_hash_map_slots := v|>) + od + + [hashmap_hash_map_clear_loop_fwd_back_def] Definition + + ⊢ ∀slots i. + hashmap_hash_map_clear_loop_fwd_back slots i = + (let + i0 = vec_len slots + in + if usize_lt i i0 then + do + i1 <- usize_add i (int_to_usize 1); + slots0 <- vec_index_mut_back slots i HashmapListNil; + hashmap_hash_map_clear_loop_fwd_back slots0 i1 + od + else Return slots) + + [hashmap_hash_map_contains_key_fwd_def] Definition + + ⊢ ∀self key. + hashmap_hash_map_contains_key_fwd self key = + do + hash <- hashmap_hash_key_fwd key; + i <<- vec_len self.hashmap_hash_map_slots; + hash_mod <- usize_rem hash i; + l <- vec_index_fwd self.hashmap_hash_map_slots hash_mod; + hashmap_hash_map_contains_key_in_list_fwd key l + od + + [hashmap_hash_map_contains_key_in_list_fwd_def] Definition + + ⊢ ∀key ls. + hashmap_hash_map_contains_key_in_list_fwd key ls = + hashmap_hash_map_contains_key_in_list_loop_fwd key ls + + [hashmap_hash_map_contains_key_in_list_loop_fwd_def] Definition + + ⊢ ∀key ls. + hashmap_hash_map_contains_key_in_list_loop_fwd key ls = + case ls of + HashmapListCons ckey t tl => + if ckey = key then Return T + else hashmap_hash_map_contains_key_in_list_loop_fwd key tl + | HashmapListNil => Return F + + [hashmap_hash_map_get_fwd_def] Definition + + ⊢ ∀self key. + hashmap_hash_map_get_fwd self key = + do + hash <- hashmap_hash_key_fwd key; + i <<- vec_len self.hashmap_hash_map_slots; + hash_mod <- usize_rem hash i; + l <- vec_index_fwd self.hashmap_hash_map_slots hash_mod; + hashmap_hash_map_get_in_list_fwd key l + od + + [hashmap_hash_map_get_in_list_fwd_def] Definition + + ⊢ ∀key ls. + hashmap_hash_map_get_in_list_fwd key ls = + hashmap_hash_map_get_in_list_loop_fwd key ls + + [hashmap_hash_map_get_in_list_loop_fwd_def] Definition + + ⊢ ∀key ls. + hashmap_hash_map_get_in_list_loop_fwd key ls = + case ls of + HashmapListCons ckey cvalue tl => + if ckey = key then Return cvalue + else hashmap_hash_map_get_in_list_loop_fwd key tl + | HashmapListNil => Fail Failure + + [hashmap_hash_map_get_mut_back_def] Definition + + ⊢ ∀self key ret. + hashmap_hash_map_get_mut_back self key ret = + do + hash <- hashmap_hash_key_fwd key; + i <<- vec_len self.hashmap_hash_map_slots; + hash_mod <- usize_rem hash i; + l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; + l0 <- hashmap_hash_map_get_mut_in_list_back l key ret; + v <- vec_index_mut_back self.hashmap_hash_map_slots hash_mod l0; + Return (self with hashmap_hash_map_slots := v) + od + + [hashmap_hash_map_get_mut_fwd_def] Definition + + ⊢ ∀self key. + hashmap_hash_map_get_mut_fwd self key = + do + hash <- hashmap_hash_key_fwd key; + i <<- vec_len self.hashmap_hash_map_slots; + hash_mod <- usize_rem hash i; + l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; + hashmap_hash_map_get_mut_in_list_fwd l key + od + + [hashmap_hash_map_get_mut_in_list_back_def] Definition + + ⊢ ∀ls key ret. + hashmap_hash_map_get_mut_in_list_back ls key ret = + hashmap_hash_map_get_mut_in_list_loop_back ls key ret + + [hashmap_hash_map_get_mut_in_list_fwd_def] Definition + + ⊢ ∀ls key. + hashmap_hash_map_get_mut_in_list_fwd ls key = + hashmap_hash_map_get_mut_in_list_loop_fwd ls key + + [hashmap_hash_map_get_mut_in_list_loop_back_def] Definition + + ⊢ ∀ls key ret. + hashmap_hash_map_get_mut_in_list_loop_back ls key ret = + case ls of + HashmapListCons ckey cvalue tl => + if ckey = key then Return (HashmapListCons ckey ret tl) + else + do + tl0 <- + hashmap_hash_map_get_mut_in_list_loop_back tl key ret; + Return (HashmapListCons ckey cvalue tl0) + od + | HashmapListNil => Fail Failure + + [hashmap_hash_map_get_mut_in_list_loop_fwd_def] Definition + + ⊢ ∀ls key. + hashmap_hash_map_get_mut_in_list_loop_fwd ls key = + case ls of + HashmapListCons ckey cvalue tl => + if ckey = key then Return cvalue + else hashmap_hash_map_get_mut_in_list_loop_fwd tl key + | HashmapListNil => Fail Failure + + [hashmap_hash_map_insert_fwd_back_def] Definition + + ⊢ ∀self key value. + hashmap_hash_map_insert_fwd_back self key value = + do + self0 <- + hashmap_hash_map_insert_no_resize_fwd_back self key value; + i <- hashmap_hash_map_len_fwd self0; + if usize_gt i self0.hashmap_hash_map_max_load then + hashmap_hash_map_try_resize_fwd_back self0 + else Return self0 + od + + [hashmap_hash_map_insert_in_list_back_def] Definition + + ⊢ ∀key value ls. + hashmap_hash_map_insert_in_list_back key value ls = + hashmap_hash_map_insert_in_list_loop_back key value ls + + [hashmap_hash_map_insert_in_list_fwd_def] Definition + + ⊢ ∀key value ls. + hashmap_hash_map_insert_in_list_fwd key value ls = + hashmap_hash_map_insert_in_list_loop_fwd key value ls + + [hashmap_hash_map_insert_in_list_loop_back_def] Definition + + ⊢ ∀key value ls. + hashmap_hash_map_insert_in_list_loop_back key value ls = + case ls of + HashmapListCons ckey cvalue tl => + if ckey = key then Return (HashmapListCons ckey value tl) + else + do + tl0 <- + hashmap_hash_map_insert_in_list_loop_back key value tl; + Return (HashmapListCons ckey cvalue tl0) + od + | HashmapListNil => + (let + l = HashmapListNil + in + Return (HashmapListCons key value l)) + + [hashmap_hash_map_insert_in_list_loop_fwd_def] Definition + + ⊢ ∀key value ls. + hashmap_hash_map_insert_in_list_loop_fwd key value ls = + case ls of + HashmapListCons ckey cvalue tl => + if ckey = key then Return F + else hashmap_hash_map_insert_in_list_loop_fwd key value tl + | HashmapListNil => Return T + + [hashmap_hash_map_insert_no_resize_fwd_back_def] Definition + + ⊢ ∀self key value. + hashmap_hash_map_insert_no_resize_fwd_back self key value = + do + hash <- hashmap_hash_key_fwd key; + i <<- vec_len self.hashmap_hash_map_slots; + hash_mod <- usize_rem hash i; + l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; + inserted <- hashmap_hash_map_insert_in_list_fwd key value l; + if inserted then + do + i0 <- + usize_add self.hashmap_hash_map_num_entries + (int_to_usize 1); + l0 <- hashmap_hash_map_insert_in_list_back key value l; + v <- + vec_index_mut_back self.hashmap_hash_map_slots hash_mod + l0; + Return + (self with + <|hashmap_hash_map_num_entries := i0; + hashmap_hash_map_slots := v|>) + od + else + do + l0 <- hashmap_hash_map_insert_in_list_back key value l; + v <- + vec_index_mut_back self.hashmap_hash_map_slots hash_mod + l0; + Return (self with hashmap_hash_map_slots := v) + od + od + + [hashmap_hash_map_len_fwd_def] Definition + + ⊢ ∀self. + hashmap_hash_map_len_fwd self = + Return self.hashmap_hash_map_num_entries + + [hashmap_hash_map_move_elements_from_list_fwd_back_def] Definition + + ⊢ ∀ntable ls. + hashmap_hash_map_move_elements_from_list_fwd_back ntable ls = + hashmap_hash_map_move_elements_from_list_loop_fwd_back ntable ls + + [hashmap_hash_map_move_elements_from_list_loop_fwd_back_def] Definition + + ⊢ ∀ntable ls. + hashmap_hash_map_move_elements_from_list_loop_fwd_back ntable ls = + case ls of + HashmapListCons k v tl => + do + ntable0 <- + hashmap_hash_map_insert_no_resize_fwd_back ntable k v; + hashmap_hash_map_move_elements_from_list_loop_fwd_back + ntable0 tl + od + | HashmapListNil => Return ntable + + [hashmap_hash_map_move_elements_fwd_back_def] Definition + + ⊢ ∀ntable slots i. + hashmap_hash_map_move_elements_fwd_back ntable slots i = + hashmap_hash_map_move_elements_loop_fwd_back ntable slots i + + [hashmap_hash_map_move_elements_loop_fwd_back_def] Definition + + ⊢ ∀ntable slots i. + hashmap_hash_map_move_elements_loop_fwd_back ntable slots i = + (let + i0 = vec_len slots + in + if usize_lt i i0 then + do + l <- vec_index_mut_fwd slots i; + ls <<- mem_replace_fwd l HashmapListNil; + ntable0 <- + hashmap_hash_map_move_elements_from_list_fwd_back ntable + ls; + i1 <- usize_add i (int_to_usize 1); + l0 <<- mem_replace_back l HashmapListNil; + slots0 <- vec_index_mut_back slots i l0; + hashmap_hash_map_move_elements_loop_fwd_back ntable0 + slots0 i1 + od + else Return (ntable,slots)) + + [hashmap_hash_map_new_fwd_def] Definition + + ⊢ hashmap_hash_map_new_fwd = + hashmap_hash_map_new_with_capacity_fwd (int_to_usize 32) + (int_to_usize 4) (int_to_usize 5) + + [hashmap_hash_map_new_with_capacity_fwd_def] Definition + + ⊢ ∀capacity max_load_dividend max_load_divisor. + hashmap_hash_map_new_with_capacity_fwd capacity max_load_dividend + max_load_divisor = + (let + v = vec_new + in + do + slots <- hashmap_hash_map_allocate_slots_fwd v capacity; + i <- usize_mul capacity max_load_dividend; + i0 <- usize_div i max_load_divisor; + Return + <|hashmap_hash_map_num_entries := int_to_usize 0; + hashmap_hash_map_max_load_factor := + (max_load_dividend,max_load_divisor); + hashmap_hash_map_max_load := i0; + hashmap_hash_map_slots := slots|> + od) + + [hashmap_hash_map_remove_back_def] Definition + + ⊢ ∀self key. + hashmap_hash_map_remove_back self key = + do + hash <- hashmap_hash_key_fwd key; + i <<- vec_len self.hashmap_hash_map_slots; + hash_mod <- usize_rem hash i; + l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; + x <- hashmap_hash_map_remove_from_list_fwd key l; + case x of + NONE => + do + l0 <- hashmap_hash_map_remove_from_list_back key l; + v <- + vec_index_mut_back self.hashmap_hash_map_slots hash_mod + l0; + Return (self with hashmap_hash_map_slots := v) + od + | SOME x0 => + do + i0 <- + usize_sub self.hashmap_hash_map_num_entries + (int_to_usize 1); + l0 <- hashmap_hash_map_remove_from_list_back key l; + v <- + vec_index_mut_back self.hashmap_hash_map_slots hash_mod + l0; + Return + (self with + <|hashmap_hash_map_num_entries := i0; + hashmap_hash_map_slots := v|>) + od + od + + [hashmap_hash_map_remove_from_list_back_def] Definition + + ⊢ ∀key ls. + hashmap_hash_map_remove_from_list_back key ls = + hashmap_hash_map_remove_from_list_loop_back key ls + + [hashmap_hash_map_remove_from_list_fwd_def] Definition + + ⊢ ∀key ls. + hashmap_hash_map_remove_from_list_fwd key ls = + hashmap_hash_map_remove_from_list_loop_fwd key ls + + [hashmap_hash_map_remove_from_list_loop_back_def] Definition + + ⊢ ∀key ls. + hashmap_hash_map_remove_from_list_loop_back key ls = + case ls of + HashmapListCons ckey t tl => + if ckey = key then + (let + mv_ls = + mem_replace_fwd (HashmapListCons ckey t tl) + HashmapListNil + in + case mv_ls of + HashmapListCons i cvalue tl0 => Return tl0 + | HashmapListNil => Fail Failure) + else + do + tl0 <- hashmap_hash_map_remove_from_list_loop_back key tl; + Return (HashmapListCons ckey t tl0) + od + | HashmapListNil => Return HashmapListNil + + [hashmap_hash_map_remove_from_list_loop_fwd_def] Definition + + ⊢ ∀key ls. + hashmap_hash_map_remove_from_list_loop_fwd key ls = + case ls of + HashmapListCons ckey t tl => + if ckey = key then + (let + mv_ls = + mem_replace_fwd (HashmapListCons ckey t tl) + HashmapListNil + in + case mv_ls of + HashmapListCons i cvalue tl0 => Return (SOME cvalue) + | HashmapListNil => Fail Failure) + else hashmap_hash_map_remove_from_list_loop_fwd key tl + | HashmapListNil => Return NONE + + [hashmap_hash_map_remove_fwd_def] Definition + + ⊢ ∀self key. + hashmap_hash_map_remove_fwd self key = + do + hash <- hashmap_hash_key_fwd key; + i <<- vec_len self.hashmap_hash_map_slots; + hash_mod <- usize_rem hash i; + l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; + x <- hashmap_hash_map_remove_from_list_fwd key l; + case x of + NONE => Return NONE + | SOME x0 => + monad_ignore_bind + (usize_sub self.hashmap_hash_map_num_entries + (int_to_usize 1)) (Return (SOME x0)) + od + + [hashmap_hash_map_try_resize_fwd_back_def] Definition + + ⊢ ∀self. + hashmap_hash_map_try_resize_fwd_back self = + do + max_usize <- mk_usize (u32_to_int core_u32_max); + capacity <<- vec_len self.hashmap_hash_map_slots; + n1 <- usize_div max_usize (int_to_usize 2); + (i,i0) <<- self.hashmap_hash_map_max_load_factor; + i1 <- usize_div n1 i; + if usize_le capacity i1 then + do + i2 <- usize_mul capacity (int_to_usize 2); + ntable <- hashmap_hash_map_new_with_capacity_fwd i2 i i0; + (ntable0,_) <- + hashmap_hash_map_move_elements_fwd_back ntable + self.hashmap_hash_map_slots (int_to_usize 0); + Return + (ntable0 with + <|hashmap_hash_map_num_entries := + self.hashmap_hash_map_num_entries; + hashmap_hash_map_max_load_factor := (i,i0)|>) + od + else + Return (self with hashmap_hash_map_max_load_factor := (i,i0)) + od + + [hashmap_test1_fwd_def] Definition + + ⊢ hashmap_test1_fwd = + do + hm <- hashmap_hash_map_new_fwd; + hm0 <- + hashmap_hash_map_insert_fwd_back hm (int_to_usize 0) + (int_to_u64 42); + hm1 <- + hashmap_hash_map_insert_fwd_back hm0 (int_to_usize 128) + (int_to_u64 18); + hm2 <- + hashmap_hash_map_insert_fwd_back hm1 (int_to_usize 1024) + (int_to_u64 138); + hm3 <- + hashmap_hash_map_insert_fwd_back hm2 (int_to_usize 1056) + (int_to_u64 256); + i <- hashmap_hash_map_get_fwd hm3 (int_to_usize 128); + if i ≠ int_to_u64 18 then Fail Failure + else + do + hm4 <- + hashmap_hash_map_get_mut_back hm3 (int_to_usize 1024) + (int_to_u64 56); + i0 <- hashmap_hash_map_get_fwd hm4 (int_to_usize 1024); + if i0 ≠ int_to_u64 56 then Fail Failure + else + do + x <- hashmap_hash_map_remove_fwd hm4 (int_to_usize 1024); + case x of + NONE => Fail Failure + | SOME x0 => + if x0 ≠ int_to_u64 56 then Fail Failure + else + do + hm5 <- + hashmap_hash_map_remove_back hm4 + (int_to_usize 1024); + i1 <- hashmap_hash_map_get_fwd hm5 (int_to_usize 0); + if i1 ≠ int_to_u64 42 then Fail Failure + else + do + i2 <- + hashmap_hash_map_get_fwd hm5 + (int_to_usize 128); + if i2 ≠ int_to_u64 18 then Fail Failure + else + do + i3 <- + hashmap_hash_map_get_fwd hm5 + (int_to_usize 1056); + if i3 ≠ int_to_u64 256 then Fail Failure + else Return () + od + od + od + od + od + od + + [insert_on_disk_fwd_def] Definition + + ⊢ ∀key value st. + insert_on_disk_fwd key value st = + do + (st0,hm) <- hashmap_utils_deserialize_fwd st; + hm0 <- hashmap_hash_map_insert_fwd_back hm key value; + (st1,_) <- hashmap_utils_serialize_fwd hm0 st0; + Return (st1,()) + od + + [main_fwd_def] Definition + + ⊢ main_fwd = Return () + + +*) +end diff --git a/tests/hol4/hashmap_main/hashmapMain_OpaqueScript.sml b/tests/hol4/hashmap_main/hashmapMain_OpaqueScript.sml new file mode 100644 index 00000000..f7221d92 --- /dev/null +++ b/tests/hol4/hashmap_main/hashmapMain_OpaqueScript.sml @@ -0,0 +1,15 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: external function declarations *) +open primitivesLib divDefLib +open hashmapMain_TypesTheory + +val _ = new_theory "hashmapMain_Opaque" + + +(** [hashmap_main::hashmap_utils::deserialize]: forward function *)val _ = new_constant ("hashmap_utils_deserialize_fwd", + “:state -> (state # u64 hashmap_hash_map_t) result”) + +(** [hashmap_main::hashmap_utils::serialize]: forward function *)val _ = new_constant ("hashmap_utils_serialize_fwd", + “:u64 hashmap_hash_map_t -> state -> (state # unit) result”) + +val _ = export_theory () diff --git a/tests/hol4/hashmap_main/hashmapMain_OpaqueTheory.sig b/tests/hol4/hashmap_main/hashmapMain_OpaqueTheory.sig new file mode 100644 index 00000000..f7373609 --- /dev/null +++ b/tests/hol4/hashmap_main/hashmapMain_OpaqueTheory.sig @@ -0,0 +1,11 @@ +signature hashmapMain_OpaqueTheory = +sig + type thm = Thm.thm + + val hashmapMain_Opaque_grammars : type_grammar.grammar * term_grammar.grammar +(* + [hashmapMain_Types] Parent theory of "hashmapMain_Opaque" + + +*) +end diff --git a/tests/hol4/hashmap_main/hashmapMain_TypesScript.sml b/tests/hol4/hashmap_main/hashmapMain_TypesScript.sml new file mode 100644 index 00000000..3f8ca9b9 --- /dev/null +++ b/tests/hol4/hashmap_main/hashmapMain_TypesScript.sml @@ -0,0 +1,27 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: type definitions *) +open primitivesLib divDefLib + +val _ = new_theory "hashmapMain_Types" + + +Datatype: + (** [hashmap_main::hashmap::List] *) + hashmap_list_t = | HashmapListCons usize 't hashmap_list_t | HashmapListNil +End + +Datatype: + (** [hashmap_main::hashmap::HashMap] *) + hashmap_hash_map_t = + <| + hashmap_hash_map_num_entries : usize; + hashmap_hash_map_max_load_factor : (usize # usize); + hashmap_hash_map_max_load : usize; + hashmap_hash_map_slots : 't hashmap_list_t vec; + |> +End + +(** The state type used in the state-error monad *) +val _ = new_type ("state", 0) + +val _ = export_theory () diff --git a/tests/hol4/hashmap_main/hashmapMain_TypesTheory.sig b/tests/hol4/hashmap_main/hashmapMain_TypesTheory.sig new file mode 100644 index 00000000..a3e770ea --- /dev/null +++ b/tests/hol4/hashmap_main/hashmapMain_TypesTheory.sig @@ -0,0 +1,568 @@ +signature hashmapMain_TypesTheory = +sig + type thm = Thm.thm + + (* Definitions *) + val hashmap_hash_map_t_TY_DEF : thm + val hashmap_hash_map_t_case_def : thm + val hashmap_hash_map_t_hashmap_hash_map_max_load : thm + val hashmap_hash_map_t_hashmap_hash_map_max_load_factor : thm + val hashmap_hash_map_t_hashmap_hash_map_max_load_factor_fupd : thm + val hashmap_hash_map_t_hashmap_hash_map_max_load_fupd : thm + val hashmap_hash_map_t_hashmap_hash_map_num_entries : thm + val hashmap_hash_map_t_hashmap_hash_map_num_entries_fupd : thm + val hashmap_hash_map_t_hashmap_hash_map_slots : thm + val hashmap_hash_map_t_hashmap_hash_map_slots_fupd : thm + val hashmap_hash_map_t_size_def : thm + val hashmap_list_t_TY_DEF : thm + val hashmap_list_t_case_def : thm + val hashmap_list_t_size_def : thm + + (* Theorems *) + val EXISTS_hashmap_hash_map_t : thm + val FORALL_hashmap_hash_map_t : thm + val datatype_hashmap_hash_map_t : thm + val datatype_hashmap_list_t : thm + val hashmap_hash_map_t_11 : thm + val hashmap_hash_map_t_Axiom : thm + val hashmap_hash_map_t_accessors : thm + val hashmap_hash_map_t_accfupds : thm + val hashmap_hash_map_t_case_cong : thm + val hashmap_hash_map_t_case_eq : thm + val hashmap_hash_map_t_component_equality : thm + val hashmap_hash_map_t_fn_updates : thm + val hashmap_hash_map_t_fupdcanon : thm + val hashmap_hash_map_t_fupdcanon_comp : thm + val hashmap_hash_map_t_fupdfupds : thm + val hashmap_hash_map_t_fupdfupds_comp : thm + val hashmap_hash_map_t_induction : thm + val hashmap_hash_map_t_literal_11 : thm + val hashmap_hash_map_t_literal_nchotomy : thm + val hashmap_hash_map_t_nchotomy : thm + val hashmap_hash_map_t_updates_eq_literal : thm + val hashmap_list_t_11 : thm + val hashmap_list_t_Axiom : thm + val hashmap_list_t_case_cong : thm + val hashmap_list_t_case_eq : thm + val hashmap_list_t_distinct : thm + val hashmap_list_t_induction : thm + val hashmap_list_t_nchotomy : thm + + val hashmapMain_Types_grammars : type_grammar.grammar * term_grammar.grammar +(* + [divDef] Parent theory of "hashmapMain_Types" + + [hashmap_hash_map_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('hashmap_hash_map_t'). + (∀a0'. + (∃a0 a1 a2 a3. + a0' = + (λa0 a1 a2 a3. + ind_type$CONSTR 0 (a0,a1,a2,a3) + (λn. ind_type$BOTTOM)) a0 a1 a2 a3) ⇒ + $var$('hashmap_hash_map_t') a0') ⇒ + $var$('hashmap_hash_map_t') a0') rep + + [hashmap_hash_map_t_case_def] Definition + + ⊢ ∀a0 a1 a2 a3 f. + hashmap_hash_map_t_CASE (hashmap_hash_map_t a0 a1 a2 a3) f = + f a0 a1 a2 a3 + + [hashmap_hash_map_t_hashmap_hash_map_max_load] Definition + + ⊢ ∀u p u0 v. + (hashmap_hash_map_t u p u0 v).hashmap_hash_map_max_load = u0 + + [hashmap_hash_map_t_hashmap_hash_map_max_load_factor] Definition + + ⊢ ∀u p u0 v. + (hashmap_hash_map_t u p u0 v).hashmap_hash_map_max_load_factor = + p + + [hashmap_hash_map_t_hashmap_hash_map_max_load_factor_fupd] Definition + + ⊢ ∀f u p u0 v. + hashmap_hash_map_t u p u0 v with + hashmap_hash_map_max_load_factor updated_by f = + hashmap_hash_map_t u (f p) u0 v + + [hashmap_hash_map_t_hashmap_hash_map_max_load_fupd] Definition + + ⊢ ∀f u p u0 v. + hashmap_hash_map_t u p u0 v with + hashmap_hash_map_max_load updated_by f = + hashmap_hash_map_t u p (f u0) v + + [hashmap_hash_map_t_hashmap_hash_map_num_entries] Definition + + ⊢ ∀u p u0 v. + (hashmap_hash_map_t u p u0 v).hashmap_hash_map_num_entries = u + + [hashmap_hash_map_t_hashmap_hash_map_num_entries_fupd] Definition + + ⊢ ∀f u p u0 v. + hashmap_hash_map_t u p u0 v with + hashmap_hash_map_num_entries updated_by f = + hashmap_hash_map_t (f u) p u0 v + + [hashmap_hash_map_t_hashmap_hash_map_slots] Definition + + ⊢ ∀u p u0 v. (hashmap_hash_map_t u p u0 v).hashmap_hash_map_slots = v + + [hashmap_hash_map_t_hashmap_hash_map_slots_fupd] Definition + + ⊢ ∀f u p u0 v. + hashmap_hash_map_t u p u0 v with + hashmap_hash_map_slots updated_by f = + hashmap_hash_map_t u p u0 (f v) + + [hashmap_hash_map_t_size_def] Definition + + ⊢ ∀f a0 a1 a2 a3. + hashmap_hash_map_t_size f (hashmap_hash_map_t a0 a1 a2 a3) = + 1 + pair_size (λv. 0) (λv. 0) a1 + + [hashmap_list_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('hashmap_list_t'). + (∀a0'. + (∃a0 a1 a2. + a0' = + (λa0 a1 a2. + ind_type$CONSTR 0 (a0,a1) + (ind_type$FCONS a2 (λn. ind_type$BOTTOM))) + a0 a1 a2 ∧ $var$('hashmap_list_t') a2) ∨ + a0' = + ind_type$CONSTR (SUC 0) (ARB,ARB) + (λn. ind_type$BOTTOM) ⇒ + $var$('hashmap_list_t') a0') ⇒ + $var$('hashmap_list_t') a0') rep + + [hashmap_list_t_case_def] Definition + + ⊢ (∀a0 a1 a2 f v. + hashmap_list_t_CASE (HashmapListCons a0 a1 a2) f v = f a0 a1 a2) ∧ + ∀f v. hashmap_list_t_CASE HashmapListNil f v = v + + [hashmap_list_t_size_def] Definition + + ⊢ (∀f a0 a1 a2. + hashmap_list_t_size f (HashmapListCons a0 a1 a2) = + 1 + (f a1 + hashmap_list_t_size f a2)) ∧ + ∀f. hashmap_list_t_size f HashmapListNil = 0 + + [EXISTS_hashmap_hash_map_t] Theorem + + ⊢ ∀P. (∃h. P h) ⇔ + ∃u0 p u v. + P + <|hashmap_hash_map_num_entries := u0; + hashmap_hash_map_max_load_factor := p; + hashmap_hash_map_max_load := u; + hashmap_hash_map_slots := v|> + + [FORALL_hashmap_hash_map_t] Theorem + + ⊢ ∀P. (∀h. P h) ⇔ + ∀u0 p u v. + P + <|hashmap_hash_map_num_entries := u0; + hashmap_hash_map_max_load_factor := p; + hashmap_hash_map_max_load := u; + hashmap_hash_map_slots := v|> + + [datatype_hashmap_hash_map_t] Theorem + + ⊢ DATATYPE + (record hashmap_hash_map_t hashmap_hash_map_num_entries + hashmap_hash_map_max_load_factor hashmap_hash_map_max_load + hashmap_hash_map_slots) + + [datatype_hashmap_list_t] Theorem + + ⊢ DATATYPE (hashmap_list_t HashmapListCons HashmapListNil) + + [hashmap_hash_map_t_11] Theorem + + ⊢ ∀a0 a1 a2 a3 a0' a1' a2' a3'. + hashmap_hash_map_t a0 a1 a2 a3 = + hashmap_hash_map_t a0' a1' a2' a3' ⇔ + a0 = a0' ∧ a1 = a1' ∧ a2 = a2' ∧ a3 = a3' + + [hashmap_hash_map_t_Axiom] Theorem + + ⊢ ∀f. ∃fn. ∀a0 a1 a2 a3. + fn (hashmap_hash_map_t a0 a1 a2 a3) = f a0 a1 a2 a3 + + [hashmap_hash_map_t_accessors] Theorem + + ⊢ (∀u p u0 v. + (hashmap_hash_map_t u p u0 v).hashmap_hash_map_num_entries = u) ∧ + (∀u p u0 v. + (hashmap_hash_map_t u p u0 v).hashmap_hash_map_max_load_factor = + p) ∧ + (∀u p u0 v. + (hashmap_hash_map_t u p u0 v).hashmap_hash_map_max_load = u0) ∧ + ∀u p u0 v. (hashmap_hash_map_t u p u0 v).hashmap_hash_map_slots = v + + [hashmap_hash_map_t_accfupds] Theorem + + ⊢ (∀h f. + (h with hashmap_hash_map_max_load_factor updated_by f). + hashmap_hash_map_num_entries = + h.hashmap_hash_map_num_entries) ∧ + (∀h f. + (h with hashmap_hash_map_max_load updated_by f). + hashmap_hash_map_num_entries = + h.hashmap_hash_map_num_entries) ∧ + (∀h f. + (h with hashmap_hash_map_slots updated_by f). + hashmap_hash_map_num_entries = + h.hashmap_hash_map_num_entries) ∧ + (∀h f. + (h with hashmap_hash_map_num_entries updated_by f). + hashmap_hash_map_max_load_factor = + h.hashmap_hash_map_max_load_factor) ∧ + (∀h f. + (h with hashmap_hash_map_max_load updated_by f). + hashmap_hash_map_max_load_factor = + h.hashmap_hash_map_max_load_factor) ∧ + (∀h f. + (h with hashmap_hash_map_slots updated_by f). + hashmap_hash_map_max_load_factor = + h.hashmap_hash_map_max_load_factor) ∧ + (∀h f. + (h with hashmap_hash_map_num_entries updated_by f). + hashmap_hash_map_max_load = + h.hashmap_hash_map_max_load) ∧ + (∀h f. + (h with hashmap_hash_map_max_load_factor updated_by f). + hashmap_hash_map_max_load = + h.hashmap_hash_map_max_load) ∧ + (∀h f. + (h with hashmap_hash_map_slots updated_by f). + hashmap_hash_map_max_load = + h.hashmap_hash_map_max_load) ∧ + (∀h f. + (h with hashmap_hash_map_num_entries updated_by f). + hashmap_hash_map_slots = + h.hashmap_hash_map_slots) ∧ + (∀h f. + (h with hashmap_hash_map_max_load_factor updated_by f). + hashmap_hash_map_slots = + h.hashmap_hash_map_slots) ∧ + (∀h f. + (h with hashmap_hash_map_max_load updated_by f). + hashmap_hash_map_slots = + h.hashmap_hash_map_slots) ∧ + (∀h f. + (h with hashmap_hash_map_num_entries updated_by f). + hashmap_hash_map_num_entries = + f h.hashmap_hash_map_num_entries) ∧ + (∀h f. + (h with hashmap_hash_map_max_load_factor updated_by f). + hashmap_hash_map_max_load_factor = + f h.hashmap_hash_map_max_load_factor) ∧ + (∀h f. + (h with hashmap_hash_map_max_load updated_by f). + hashmap_hash_map_max_load = + f h.hashmap_hash_map_max_load) ∧ + ∀h f. + (h with hashmap_hash_map_slots updated_by f). + hashmap_hash_map_slots = + f h.hashmap_hash_map_slots + + [hashmap_hash_map_t_case_cong] Theorem + + ⊢ ∀M M' f. + M = M' ∧ + (∀a0 a1 a2 a3. + M' = hashmap_hash_map_t a0 a1 a2 a3 ⇒ + f a0 a1 a2 a3 = f' a0 a1 a2 a3) ⇒ + hashmap_hash_map_t_CASE M f = hashmap_hash_map_t_CASE M' f' + + [hashmap_hash_map_t_case_eq] Theorem + + ⊢ hashmap_hash_map_t_CASE x f = v ⇔ + ∃u p u0 v'. x = hashmap_hash_map_t u p u0 v' ∧ f u p u0 v' = v + + [hashmap_hash_map_t_component_equality] Theorem + + ⊢ ∀h1 h2. + h1 = h2 ⇔ + h1.hashmap_hash_map_num_entries = h2.hashmap_hash_map_num_entries ∧ + h1.hashmap_hash_map_max_load_factor = + h2.hashmap_hash_map_max_load_factor ∧ + h1.hashmap_hash_map_max_load = h2.hashmap_hash_map_max_load ∧ + h1.hashmap_hash_map_slots = h2.hashmap_hash_map_slots + + [hashmap_hash_map_t_fn_updates] Theorem + + ⊢ (∀f u p u0 v. + hashmap_hash_map_t u p u0 v with + hashmap_hash_map_num_entries updated_by f = + hashmap_hash_map_t (f u) p u0 v) ∧ + (∀f u p u0 v. + hashmap_hash_map_t u p u0 v with + hashmap_hash_map_max_load_factor updated_by f = + hashmap_hash_map_t u (f p) u0 v) ∧ + (∀f u p u0 v. + hashmap_hash_map_t u p u0 v with + hashmap_hash_map_max_load updated_by f = + hashmap_hash_map_t u p (f u0) v) ∧ + ∀f u p u0 v. + hashmap_hash_map_t u p u0 v with + hashmap_hash_map_slots updated_by f = + hashmap_hash_map_t u p u0 (f v) + + [hashmap_hash_map_t_fupdcanon] Theorem + + ⊢ (∀h g f. + h with + <|hashmap_hash_map_max_load_factor updated_by f; + hashmap_hash_map_num_entries updated_by g|> = + h with + <|hashmap_hash_map_num_entries updated_by g; + hashmap_hash_map_max_load_factor updated_by f|>) ∧ + (∀h g f. + h with + <|hashmap_hash_map_max_load updated_by f; + hashmap_hash_map_num_entries updated_by g|> = + h with + <|hashmap_hash_map_num_entries updated_by g; + hashmap_hash_map_max_load updated_by f|>) ∧ + (∀h g f. + h with + <|hashmap_hash_map_max_load updated_by f; + hashmap_hash_map_max_load_factor updated_by g|> = + h with + <|hashmap_hash_map_max_load_factor updated_by g; + hashmap_hash_map_max_load updated_by f|>) ∧ + (∀h g f. + h with + <|hashmap_hash_map_slots updated_by f; + hashmap_hash_map_num_entries updated_by g|> = + h with + <|hashmap_hash_map_num_entries updated_by g; + hashmap_hash_map_slots updated_by f|>) ∧ + (∀h g f. + h with + <|hashmap_hash_map_slots updated_by f; + hashmap_hash_map_max_load_factor updated_by g|> = + h with + <|hashmap_hash_map_max_load_factor updated_by g; + hashmap_hash_map_slots updated_by f|>) ∧ + ∀h g f. + h with + <|hashmap_hash_map_slots updated_by f; + hashmap_hash_map_max_load updated_by g|> = + h with + <|hashmap_hash_map_max_load updated_by g; + hashmap_hash_map_slots updated_by f|> + + [hashmap_hash_map_t_fupdcanon_comp] Theorem + + ⊢ ((∀g f. + hashmap_hash_map_max_load_factor_fupd f ∘ + hashmap_hash_map_num_entries_fupd g = + hashmap_hash_map_num_entries_fupd g ∘ + hashmap_hash_map_max_load_factor_fupd f) ∧ + ∀h g f. + hashmap_hash_map_max_load_factor_fupd f ∘ + hashmap_hash_map_num_entries_fupd g ∘ h = + hashmap_hash_map_num_entries_fupd g ∘ + hashmap_hash_map_max_load_factor_fupd f ∘ h) ∧ + ((∀g f. + hashmap_hash_map_max_load_fupd f ∘ + hashmap_hash_map_num_entries_fupd g = + hashmap_hash_map_num_entries_fupd g ∘ + hashmap_hash_map_max_load_fupd f) ∧ + ∀h g f. + hashmap_hash_map_max_load_fupd f ∘ + hashmap_hash_map_num_entries_fupd g ∘ h = + hashmap_hash_map_num_entries_fupd g ∘ + hashmap_hash_map_max_load_fupd f ∘ h) ∧ + ((∀g f. + hashmap_hash_map_max_load_fupd f ∘ + hashmap_hash_map_max_load_factor_fupd g = + hashmap_hash_map_max_load_factor_fupd g ∘ + hashmap_hash_map_max_load_fupd f) ∧ + ∀h g f. + hashmap_hash_map_max_load_fupd f ∘ + hashmap_hash_map_max_load_factor_fupd g ∘ h = + hashmap_hash_map_max_load_factor_fupd g ∘ + hashmap_hash_map_max_load_fupd f ∘ h) ∧ + ((∀g f. + hashmap_hash_map_slots_fupd f ∘ + hashmap_hash_map_num_entries_fupd g = + hashmap_hash_map_num_entries_fupd g ∘ + hashmap_hash_map_slots_fupd f) ∧ + ∀h g f. + hashmap_hash_map_slots_fupd f ∘ + hashmap_hash_map_num_entries_fupd g ∘ h = + hashmap_hash_map_num_entries_fupd g ∘ + hashmap_hash_map_slots_fupd f ∘ h) ∧ + ((∀g f. + hashmap_hash_map_slots_fupd f ∘ + hashmap_hash_map_max_load_factor_fupd g = + hashmap_hash_map_max_load_factor_fupd g ∘ + hashmap_hash_map_slots_fupd f) ∧ + ∀h g f. + hashmap_hash_map_slots_fupd f ∘ + hashmap_hash_map_max_load_factor_fupd g ∘ h = + hashmap_hash_map_max_load_factor_fupd g ∘ + hashmap_hash_map_slots_fupd f ∘ h) ∧ + (∀g f. + hashmap_hash_map_slots_fupd f ∘ hashmap_hash_map_max_load_fupd g = + hashmap_hash_map_max_load_fupd g ∘ hashmap_hash_map_slots_fupd f) ∧ + ∀h g f. + hashmap_hash_map_slots_fupd f ∘ + hashmap_hash_map_max_load_fupd g ∘ h = + hashmap_hash_map_max_load_fupd g ∘ + hashmap_hash_map_slots_fupd f ∘ h + + [hashmap_hash_map_t_fupdfupds] Theorem + + ⊢ (∀h g f. + h with + <|hashmap_hash_map_num_entries updated_by f; + hashmap_hash_map_num_entries updated_by g|> = + h with hashmap_hash_map_num_entries updated_by f ∘ g) ∧ + (∀h g f. + h with + <|hashmap_hash_map_max_load_factor updated_by f; + hashmap_hash_map_max_load_factor updated_by g|> = + h with hashmap_hash_map_max_load_factor updated_by f ∘ g) ∧ + (∀h g f. + h with + <|hashmap_hash_map_max_load updated_by f; + hashmap_hash_map_max_load updated_by g|> = + h with hashmap_hash_map_max_load updated_by f ∘ g) ∧ + ∀h g f. + h with + <|hashmap_hash_map_slots updated_by f; + hashmap_hash_map_slots updated_by g|> = + h with hashmap_hash_map_slots updated_by f ∘ g + + [hashmap_hash_map_t_fupdfupds_comp] Theorem + + ⊢ ((∀g f. + hashmap_hash_map_num_entries_fupd f ∘ + hashmap_hash_map_num_entries_fupd g = + hashmap_hash_map_num_entries_fupd (f ∘ g)) ∧ + ∀h g f. + hashmap_hash_map_num_entries_fupd f ∘ + hashmap_hash_map_num_entries_fupd g ∘ h = + hashmap_hash_map_num_entries_fupd (f ∘ g) ∘ h) ∧ + ((∀g f. + hashmap_hash_map_max_load_factor_fupd f ∘ + hashmap_hash_map_max_load_factor_fupd g = + hashmap_hash_map_max_load_factor_fupd (f ∘ g)) ∧ + ∀h g f. + hashmap_hash_map_max_load_factor_fupd f ∘ + hashmap_hash_map_max_load_factor_fupd g ∘ h = + hashmap_hash_map_max_load_factor_fupd (f ∘ g) ∘ h) ∧ + ((∀g f. + hashmap_hash_map_max_load_fupd f ∘ + hashmap_hash_map_max_load_fupd g = + hashmap_hash_map_max_load_fupd (f ∘ g)) ∧ + ∀h g f. + hashmap_hash_map_max_load_fupd f ∘ + hashmap_hash_map_max_load_fupd g ∘ h = + hashmap_hash_map_max_load_fupd (f ∘ g) ∘ h) ∧ + (∀g f. + hashmap_hash_map_slots_fupd f ∘ hashmap_hash_map_slots_fupd g = + hashmap_hash_map_slots_fupd (f ∘ g)) ∧ + ∀h g f. + hashmap_hash_map_slots_fupd f ∘ hashmap_hash_map_slots_fupd g ∘ h = + hashmap_hash_map_slots_fupd (f ∘ g) ∘ h + + [hashmap_hash_map_t_induction] Theorem + + ⊢ ∀P. (∀u p u0 v. P (hashmap_hash_map_t u p u0 v)) ⇒ ∀h. P h + + [hashmap_hash_map_t_literal_11] Theorem + + ⊢ ∀u01 p1 u1 v1 u02 p2 u2 v2. + <|hashmap_hash_map_num_entries := u01; + hashmap_hash_map_max_load_factor := p1; + hashmap_hash_map_max_load := u1; hashmap_hash_map_slots := v1|> = + <|hashmap_hash_map_num_entries := u02; + hashmap_hash_map_max_load_factor := p2; + hashmap_hash_map_max_load := u2; hashmap_hash_map_slots := v2|> ⇔ + u01 = u02 ∧ p1 = p2 ∧ u1 = u2 ∧ v1 = v2 + + [hashmap_hash_map_t_literal_nchotomy] Theorem + + ⊢ ∀h. ∃u0 p u v. + h = + <|hashmap_hash_map_num_entries := u0; + hashmap_hash_map_max_load_factor := p; + hashmap_hash_map_max_load := u; hashmap_hash_map_slots := v|> + + [hashmap_hash_map_t_nchotomy] Theorem + + ⊢ ∀hh. ∃u p u0 v. hh = hashmap_hash_map_t u p u0 v + + [hashmap_hash_map_t_updates_eq_literal] Theorem + + ⊢ ∀h u0 p u v. + h with + <|hashmap_hash_map_num_entries := u0; + hashmap_hash_map_max_load_factor := p; + hashmap_hash_map_max_load := u; hashmap_hash_map_slots := v|> = + <|hashmap_hash_map_num_entries := u0; + hashmap_hash_map_max_load_factor := p; + hashmap_hash_map_max_load := u; hashmap_hash_map_slots := v|> + + [hashmap_list_t_11] Theorem + + ⊢ ∀a0 a1 a2 a0' a1' a2'. + HashmapListCons a0 a1 a2 = HashmapListCons a0' a1' a2' ⇔ + a0 = a0' ∧ a1 = a1' ∧ a2 = a2' + + [hashmap_list_t_Axiom] Theorem + + ⊢ ∀f0 f1. ∃fn. + (∀a0 a1 a2. fn (HashmapListCons a0 a1 a2) = f0 a0 a1 a2 (fn a2)) ∧ + fn HashmapListNil = f1 + + [hashmap_list_t_case_cong] Theorem + + ⊢ ∀M M' f v. + M = M' ∧ + (∀a0 a1 a2. + M' = HashmapListCons a0 a1 a2 ⇒ f a0 a1 a2 = f' a0 a1 a2) ∧ + (M' = HashmapListNil ⇒ v = v') ⇒ + hashmap_list_t_CASE M f v = hashmap_list_t_CASE M' f' v' + + [hashmap_list_t_case_eq] Theorem + + ⊢ hashmap_list_t_CASE x f v = v' ⇔ + (∃u t h. x = HashmapListCons u t h ∧ f u t h = v') ∨ + x = HashmapListNil ∧ v = v' + + [hashmap_list_t_distinct] Theorem + + ⊢ ∀a2 a1 a0. HashmapListCons a0 a1 a2 ≠ HashmapListNil + + [hashmap_list_t_induction] Theorem + + ⊢ ∀P. (∀h. P h ⇒ ∀t u. P (HashmapListCons u t h)) ∧ P HashmapListNil ⇒ + ∀h. P h + + [hashmap_list_t_nchotomy] Theorem + + ⊢ ∀hh. (∃u t h. hh = HashmapListCons u t h) ∨ hh = HashmapListNil + + +*) +end diff --git a/tests/hol4/hashmap_on_disk/Holmakefile b/tests/hol4/hashmap_on_disk/Holmakefile deleted file mode 100644 index 3c4b8973..00000000 --- a/tests/hol4/hashmap_on_disk/Holmakefile +++ /dev/null @@ -1,5 +0,0 @@ -# This file was automatically generated - modify ../Holmakefile.template instead -INCLUDES = ../../../backends/hol4 - -all: $(DEFAULT_TARGETS) -.PHONY: all diff --git a/tests/hol4/hashmap_on_disk/hashmapMain_FunsScript.sml b/tests/hol4/hashmap_on_disk/hashmapMain_FunsScript.sml deleted file mode 100644 index c1e30aa6..00000000 --- a/tests/hol4/hashmap_on_disk/hashmapMain_FunsScript.sml +++ /dev/null @@ -1,647 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: function definitions *) -open primitivesLib divDefLib -open hashmapMain_TypesTheory hashmapMain_OpaqueTheory - -val _ = new_theory "hashmapMain_Funs" - - -val hashmap_hash_key_fwd_def = Define ‘ - (** [hashmap_main::hashmap::hash_key]: forward function *) - hashmap_hash_key_fwd (k : usize) : usize result = - Return k -’ - -val [hashmap_hash_map_allocate_slots_loop_fwd_def] = DefineDiv ‘ - (** [hashmap_main::hashmap::HashMap::{0}::allocate_slots]: loop 0: forward function *) - hashmap_hash_map_allocate_slots_loop_fwd - (slots : 't hashmap_list_t vec) (n : usize) : - 't hashmap_list_t vec result - = - if usize_gt n (int_to_usize 0) - then ( - do - slots0 <- vec_push_back slots HashmapListNil; - n0 <- usize_sub n (int_to_usize 1); - hashmap_hash_map_allocate_slots_loop_fwd slots0 n0 - od) - else Return slots -’ - -val hashmap_hash_map_allocate_slots_fwd_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::allocate_slots]: forward function *) - hashmap_hash_map_allocate_slots_fwd - (slots : 't hashmap_list_t vec) (n : usize) : - 't hashmap_list_t vec result - = - hashmap_hash_map_allocate_slots_loop_fwd slots n -’ - -val hashmap_hash_map_new_with_capacity_fwd_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::new_with_capacity]: forward function *) - hashmap_hash_map_new_with_capacity_fwd - (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : - 't hashmap_hash_map_t result - = - let v = vec_new in - do - slots <- hashmap_hash_map_allocate_slots_fwd v capacity; - i <- usize_mul capacity max_load_dividend; - i0 <- usize_div i max_load_divisor; - Return - (<| - hashmap_hash_map_num_entries := (int_to_usize 0); - hashmap_hash_map_max_load_factor := - (max_load_dividend, max_load_divisor); - hashmap_hash_map_max_load := i0; - hashmap_hash_map_slots := slots - |>) - od -’ - -val hashmap_hash_map_new_fwd_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::new]: forward function *) - hashmap_hash_map_new_fwd : 't hashmap_hash_map_t result = - hashmap_hash_map_new_with_capacity_fwd (int_to_usize 32) (int_to_usize 4) - (int_to_usize 5) -’ - -val [hashmap_hash_map_clear_loop_fwd_back_def] = DefineDiv ‘ - (** [hashmap_main::hashmap::HashMap::{0}::clear]: loop 0: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - hashmap_hash_map_clear_loop_fwd_back - (slots : 't hashmap_list_t vec) (i : usize) : - 't hashmap_list_t vec result - = - let i0 = vec_len slots in - if usize_lt i i0 - then ( - do - i1 <- usize_add i (int_to_usize 1); - slots0 <- vec_index_mut_back slots i HashmapListNil; - hashmap_hash_map_clear_loop_fwd_back slots0 i1 - od) - else Return slots -’ - -val hashmap_hash_map_clear_fwd_back_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::clear]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - hashmap_hash_map_clear_fwd_back - (self : 't hashmap_hash_map_t) : 't hashmap_hash_map_t result = - do - v <- - hashmap_hash_map_clear_loop_fwd_back self.hashmap_hash_map_slots - (int_to_usize 0); - Return - (self - with - <| - hashmap_hash_map_num_entries := (int_to_usize 0); - hashmap_hash_map_slots := v - |>) - od -’ - -val hashmap_hash_map_len_fwd_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::len]: forward function *) - hashmap_hash_map_len_fwd (self : 't hashmap_hash_map_t) : usize result = - Return self.hashmap_hash_map_num_entries -’ - -val [hashmap_hash_map_insert_in_list_loop_fwd_def] = DefineDiv ‘ - (** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: loop 0: forward function *) - hashmap_hash_map_insert_in_list_loop_fwd - (key : usize) (value : 't) (ls : 't hashmap_list_t) : bool result = - (case ls of - | HashmapListCons ckey cvalue tl => - if ckey = key - then Return F - else hashmap_hash_map_insert_in_list_loop_fwd key value tl - | HashmapListNil => Return T) -’ - -val hashmap_hash_map_insert_in_list_fwd_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: forward function *) - hashmap_hash_map_insert_in_list_fwd - (key : usize) (value : 't) (ls : 't hashmap_list_t) : bool result = - hashmap_hash_map_insert_in_list_loop_fwd key value ls -’ - -val [hashmap_hash_map_insert_in_list_loop_back_def] = DefineDiv ‘ - (** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: loop 0: backward function 0 *) - hashmap_hash_map_insert_in_list_loop_back - (key : usize) (value : 't) (ls : 't hashmap_list_t) : - 't hashmap_list_t result - = - (case ls of - | HashmapListCons ckey cvalue tl => - if ckey = key - then Return (HashmapListCons ckey value tl) - else ( - do - tl0 <- hashmap_hash_map_insert_in_list_loop_back key value tl; - Return (HashmapListCons ckey cvalue tl0) - od) - | HashmapListNil => - let l = HashmapListNil in Return (HashmapListCons key value l)) -’ - -val hashmap_hash_map_insert_in_list_back_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::insert_in_list]: backward function 0 *) - hashmap_hash_map_insert_in_list_back - (key : usize) (value : 't) (ls : 't hashmap_list_t) : - 't hashmap_list_t result - = - hashmap_hash_map_insert_in_list_loop_back key value ls -’ - -val hashmap_hash_map_insert_no_resize_fwd_back_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::insert_no_resize]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - hashmap_hash_map_insert_no_resize_fwd_back - (self : 't hashmap_hash_map_t) (key : usize) (value : 't) : - 't hashmap_hash_map_t result - = - do - hash <- hashmap_hash_key_fwd key; - let i = vec_len self.hashmap_hash_map_slots in - do - hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; - inserted <- hashmap_hash_map_insert_in_list_fwd key value l; - if inserted - then ( - do - i0 <- usize_add self.hashmap_hash_map_num_entries (int_to_usize 1); - l0 <- hashmap_hash_map_insert_in_list_back key value l; - v <- vec_index_mut_back self.hashmap_hash_map_slots hash_mod l0; - Return - (self - with - <| - hashmap_hash_map_num_entries := i0; hashmap_hash_map_slots := v - |>) - od) - else ( - do - l0 <- hashmap_hash_map_insert_in_list_back key value l; - v <- vec_index_mut_back self.hashmap_hash_map_slots hash_mod l0; - Return (self with <| hashmap_hash_map_slots := v |>) - od) - od - od -’ - -val [hashmap_hash_map_move_elements_from_list_loop_fwd_back_def] = DefineDiv ‘ - (** [hashmap_main::hashmap::HashMap::{0}::move_elements_from_list]: loop 0: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - hashmap_hash_map_move_elements_from_list_loop_fwd_back - (ntable : 't hashmap_hash_map_t) (ls : 't hashmap_list_t) : - 't hashmap_hash_map_t result - = - (case ls of - | HashmapListCons k v tl => - do - ntable0 <- hashmap_hash_map_insert_no_resize_fwd_back ntable k v; - hashmap_hash_map_move_elements_from_list_loop_fwd_back ntable0 tl - od - | HashmapListNil => Return ntable) -’ - -val hashmap_hash_map_move_elements_from_list_fwd_back_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::move_elements_from_list]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - hashmap_hash_map_move_elements_from_list_fwd_back - (ntable : 't hashmap_hash_map_t) (ls : 't hashmap_list_t) : - 't hashmap_hash_map_t result - = - hashmap_hash_map_move_elements_from_list_loop_fwd_back ntable ls -’ - -val [hashmap_hash_map_move_elements_loop_fwd_back_def] = DefineDiv ‘ - (** [hashmap_main::hashmap::HashMap::{0}::move_elements]: loop 0: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - hashmap_hash_map_move_elements_loop_fwd_back - (ntable : 't hashmap_hash_map_t) (slots : 't hashmap_list_t vec) - (i : usize) : - ('t hashmap_hash_map_t # 't hashmap_list_t vec) result - = - let i0 = vec_len slots in - if usize_lt i i0 - then ( - do - l <- vec_index_mut_fwd slots i; - let ls = mem_replace_fwd l HashmapListNil in - do - ntable0 <- hashmap_hash_map_move_elements_from_list_fwd_back ntable ls; - i1 <- usize_add i (int_to_usize 1); - let l0 = mem_replace_back l HashmapListNil in - do - slots0 <- vec_index_mut_back slots i l0; - hashmap_hash_map_move_elements_loop_fwd_back ntable0 slots0 i1 - od - od - od) - else Return (ntable, slots) -’ - -val hashmap_hash_map_move_elements_fwd_back_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::move_elements]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - hashmap_hash_map_move_elements_fwd_back - (ntable : 't hashmap_hash_map_t) (slots : 't hashmap_list_t vec) - (i : usize) : - ('t hashmap_hash_map_t # 't hashmap_list_t vec) result - = - hashmap_hash_map_move_elements_loop_fwd_back ntable slots i -’ - -val hashmap_hash_map_try_resize_fwd_back_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::try_resize]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - hashmap_hash_map_try_resize_fwd_back - (self : 't hashmap_hash_map_t) : 't hashmap_hash_map_t result = - do - max_usize <- mk_usize (u32_to_int core_u32_max); - let capacity = vec_len self.hashmap_hash_map_slots in - do - n1 <- usize_div max_usize (int_to_usize 2); - let (i, i0) = self.hashmap_hash_map_max_load_factor in - do - i1 <- usize_div n1 i; - if usize_le capacity i1 - then ( - do - i2 <- usize_mul capacity (int_to_usize 2); - ntable <- hashmap_hash_map_new_with_capacity_fwd i2 i i0; - (ntable0, _) <- - hashmap_hash_map_move_elements_fwd_back ntable - self.hashmap_hash_map_slots (int_to_usize 0); - Return - (ntable0 - with - <| - hashmap_hash_map_num_entries := self.hashmap_hash_map_num_entries; - hashmap_hash_map_max_load_factor := (i, i0) - |>) - od) - else Return (self with <| hashmap_hash_map_max_load_factor := (i, i0) |>) - od - od - od -’ - -val hashmap_hash_map_insert_fwd_back_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::insert]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - hashmap_hash_map_insert_fwd_back - (self : 't hashmap_hash_map_t) (key : usize) (value : 't) : - 't hashmap_hash_map_t result - = - do - self0 <- hashmap_hash_map_insert_no_resize_fwd_back self key value; - i <- hashmap_hash_map_len_fwd self0; - if usize_gt i self0.hashmap_hash_map_max_load - then hashmap_hash_map_try_resize_fwd_back self0 - else Return self0 - od -’ - -val [hashmap_hash_map_contains_key_in_list_loop_fwd_def] = DefineDiv ‘ - (** [hashmap_main::hashmap::HashMap::{0}::contains_key_in_list]: loop 0: forward function *) - hashmap_hash_map_contains_key_in_list_loop_fwd - (key : usize) (ls : 't hashmap_list_t) : bool result = - (case ls of - | HashmapListCons ckey t tl => - if ckey = key - then Return T - else hashmap_hash_map_contains_key_in_list_loop_fwd key tl - | HashmapListNil => Return F) -’ - -val hashmap_hash_map_contains_key_in_list_fwd_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::contains_key_in_list]: forward function *) - hashmap_hash_map_contains_key_in_list_fwd - (key : usize) (ls : 't hashmap_list_t) : bool result = - hashmap_hash_map_contains_key_in_list_loop_fwd key ls -’ - -val hashmap_hash_map_contains_key_fwd_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::contains_key]: forward function *) - hashmap_hash_map_contains_key_fwd - (self : 't hashmap_hash_map_t) (key : usize) : bool result = - do - hash <- hashmap_hash_key_fwd key; - let i = vec_len self.hashmap_hash_map_slots in - do - hash_mod <- usize_rem hash i; - l <- vec_index_fwd self.hashmap_hash_map_slots hash_mod; - hashmap_hash_map_contains_key_in_list_fwd key l - od - od -’ - -val [hashmap_hash_map_get_in_list_loop_fwd_def] = DefineDiv ‘ - (** [hashmap_main::hashmap::HashMap::{0}::get_in_list]: loop 0: forward function *) - hashmap_hash_map_get_in_list_loop_fwd - (key : usize) (ls : 't hashmap_list_t) : 't result = - (case ls of - | HashmapListCons ckey cvalue tl => - if ckey = key - then Return cvalue - else hashmap_hash_map_get_in_list_loop_fwd key tl - | HashmapListNil => Fail Failure) -’ - -val hashmap_hash_map_get_in_list_fwd_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::get_in_list]: forward function *) - hashmap_hash_map_get_in_list_fwd - (key : usize) (ls : 't hashmap_list_t) : 't result = - hashmap_hash_map_get_in_list_loop_fwd key ls -’ - -val hashmap_hash_map_get_fwd_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::get]: forward function *) - hashmap_hash_map_get_fwd - (self : 't hashmap_hash_map_t) (key : usize) : 't result = - do - hash <- hashmap_hash_key_fwd key; - let i = vec_len self.hashmap_hash_map_slots in - do - hash_mod <- usize_rem hash i; - l <- vec_index_fwd self.hashmap_hash_map_slots hash_mod; - hashmap_hash_map_get_in_list_fwd key l - od - od -’ - -val [hashmap_hash_map_get_mut_in_list_loop_fwd_def] = DefineDiv ‘ - (** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: loop 0: forward function *) - hashmap_hash_map_get_mut_in_list_loop_fwd - (ls : 't hashmap_list_t) (key : usize) : 't result = - (case ls of - | HashmapListCons ckey cvalue tl => - if ckey = key - then Return cvalue - else hashmap_hash_map_get_mut_in_list_loop_fwd tl key - | HashmapListNil => Fail Failure) -’ - -val hashmap_hash_map_get_mut_in_list_fwd_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: forward function *) - hashmap_hash_map_get_mut_in_list_fwd - (ls : 't hashmap_list_t) (key : usize) : 't result = - hashmap_hash_map_get_mut_in_list_loop_fwd ls key -’ - -val [hashmap_hash_map_get_mut_in_list_loop_back_def] = DefineDiv ‘ - (** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: loop 0: backward function 0 *) - hashmap_hash_map_get_mut_in_list_loop_back - (ls : 't hashmap_list_t) (key : usize) (ret : 't) : - 't hashmap_list_t result - = - (case ls of - | HashmapListCons ckey cvalue tl => - if ckey = key - then Return (HashmapListCons ckey ret tl) - else ( - do - tl0 <- hashmap_hash_map_get_mut_in_list_loop_back tl key ret; - Return (HashmapListCons ckey cvalue tl0) - od) - | HashmapListNil => Fail Failure) -’ - -val hashmap_hash_map_get_mut_in_list_back_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list]: backward function 0 *) - hashmap_hash_map_get_mut_in_list_back - (ls : 't hashmap_list_t) (key : usize) (ret : 't) : - 't hashmap_list_t result - = - hashmap_hash_map_get_mut_in_list_loop_back ls key ret -’ - -val hashmap_hash_map_get_mut_fwd_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::get_mut]: forward function *) - hashmap_hash_map_get_mut_fwd - (self : 't hashmap_hash_map_t) (key : usize) : 't result = - do - hash <- hashmap_hash_key_fwd key; - let i = vec_len self.hashmap_hash_map_slots in - do - hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; - hashmap_hash_map_get_mut_in_list_fwd l key - od - od -’ - -val hashmap_hash_map_get_mut_back_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::get_mut]: backward function 0 *) - hashmap_hash_map_get_mut_back - (self : 't hashmap_hash_map_t) (key : usize) (ret : 't) : - 't hashmap_hash_map_t result - = - do - hash <- hashmap_hash_key_fwd key; - let i = vec_len self.hashmap_hash_map_slots in - do - hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; - l0 <- hashmap_hash_map_get_mut_in_list_back l key ret; - v <- vec_index_mut_back self.hashmap_hash_map_slots hash_mod l0; - Return (self with <| hashmap_hash_map_slots := v |>) - od - od -’ - -val [hashmap_hash_map_remove_from_list_loop_fwd_def] = DefineDiv ‘ - (** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: loop 0: forward function *) - hashmap_hash_map_remove_from_list_loop_fwd - (key : usize) (ls : 't hashmap_list_t) : 't option result = - (case ls of - | HashmapListCons ckey t tl => - if ckey = key - then - let mv_ls = mem_replace_fwd (HashmapListCons ckey t tl) HashmapListNil - in - (case mv_ls of - | HashmapListCons i cvalue tl0 => Return (SOME cvalue) - | HashmapListNil => Fail Failure) - else hashmap_hash_map_remove_from_list_loop_fwd key tl - | HashmapListNil => Return NONE) -’ - -val hashmap_hash_map_remove_from_list_fwd_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: forward function *) - hashmap_hash_map_remove_from_list_fwd - (key : usize) (ls : 't hashmap_list_t) : 't option result = - hashmap_hash_map_remove_from_list_loop_fwd key ls -’ - -val [hashmap_hash_map_remove_from_list_loop_back_def] = DefineDiv ‘ - (** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: loop 0: backward function 1 *) - hashmap_hash_map_remove_from_list_loop_back - (key : usize) (ls : 't hashmap_list_t) : 't hashmap_list_t result = - (case ls of - | HashmapListCons ckey t tl => - if ckey = key - then - let mv_ls = mem_replace_fwd (HashmapListCons ckey t tl) HashmapListNil - in - (case mv_ls of - | HashmapListCons i cvalue tl0 => Return tl0 - | HashmapListNil => Fail Failure) - else ( - do - tl0 <- hashmap_hash_map_remove_from_list_loop_back key tl; - Return (HashmapListCons ckey t tl0) - od) - | HashmapListNil => Return HashmapListNil) -’ - -val hashmap_hash_map_remove_from_list_back_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::remove_from_list]: backward function 1 *) - hashmap_hash_map_remove_from_list_back - (key : usize) (ls : 't hashmap_list_t) : 't hashmap_list_t result = - hashmap_hash_map_remove_from_list_loop_back key ls -’ - -val hashmap_hash_map_remove_fwd_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::remove]: forward function *) - hashmap_hash_map_remove_fwd - (self : 't hashmap_hash_map_t) (key : usize) : 't option result = - do - hash <- hashmap_hash_key_fwd key; - let i = vec_len self.hashmap_hash_map_slots in - do - hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; - x <- hashmap_hash_map_remove_from_list_fwd key l; - (case x of - | NONE => Return NONE - | SOME x0 => - do - _ <- usize_sub self.hashmap_hash_map_num_entries (int_to_usize 1); - Return (SOME x0) - od) - od - od -’ - -val hashmap_hash_map_remove_back_def = Define ‘ - (** [hashmap_main::hashmap::HashMap::{0}::remove]: backward function 0 *) - hashmap_hash_map_remove_back - (self : 't hashmap_hash_map_t) (key : usize) : - 't hashmap_hash_map_t result - = - do - hash <- hashmap_hash_key_fwd key; - let i = vec_len self.hashmap_hash_map_slots in - do - hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; - x <- hashmap_hash_map_remove_from_list_fwd key l; - (case x of - | NONE => - do - l0 <- hashmap_hash_map_remove_from_list_back key l; - v <- vec_index_mut_back self.hashmap_hash_map_slots hash_mod l0; - Return (self with <| hashmap_hash_map_slots := v |>) - od - | SOME x0 => - do - i0 <- usize_sub self.hashmap_hash_map_num_entries (int_to_usize 1); - l0 <- hashmap_hash_map_remove_from_list_back key l; - v <- vec_index_mut_back self.hashmap_hash_map_slots hash_mod l0; - Return - (self - with - <| - hashmap_hash_map_num_entries := i0; hashmap_hash_map_slots := v - |>) - od) - od - od -’ - -val hashmap_test1_fwd_def = Define ‘ - (** [hashmap_main::hashmap::test1]: forward function *) - hashmap_test1_fwd : unit result = - do - hm <- hashmap_hash_map_new_fwd; - hm0 <- - hashmap_hash_map_insert_fwd_back hm (int_to_usize 0) (int_to_u64 42); - hm1 <- - hashmap_hash_map_insert_fwd_back hm0 (int_to_usize 128) (int_to_u64 18); - hm2 <- - hashmap_hash_map_insert_fwd_back hm1 (int_to_usize 1024) (int_to_u64 138); - hm3 <- - hashmap_hash_map_insert_fwd_back hm2 (int_to_usize 1056) (int_to_u64 256); - i <- hashmap_hash_map_get_fwd hm3 (int_to_usize 128); - if ~ (i = int_to_u64 18) - then Fail Failure - else ( - do - hm4 <- - hashmap_hash_map_get_mut_back hm3 (int_to_usize 1024) (int_to_u64 56); - i0 <- hashmap_hash_map_get_fwd hm4 (int_to_usize 1024); - if ~ (i0 = int_to_u64 56) - then Fail Failure - else ( - do - x <- hashmap_hash_map_remove_fwd hm4 (int_to_usize 1024); - (case x of - | NONE => Fail Failure - | SOME x0 => - if ~ (x0 = int_to_u64 56) - then Fail Failure - else ( - do - hm5 <- hashmap_hash_map_remove_back hm4 (int_to_usize 1024); - i1 <- hashmap_hash_map_get_fwd hm5 (int_to_usize 0); - if ~ (i1 = int_to_u64 42) - then Fail Failure - else ( - do - i2 <- hashmap_hash_map_get_fwd hm5 (int_to_usize 128); - if ~ (i2 = int_to_u64 18) - then Fail Failure - else ( - do - i3 <- hashmap_hash_map_get_fwd hm5 (int_to_usize 1056); - if ~ (i3 = int_to_u64 256) then Fail Failure else Return () - od) - od) - od)) - od) - od) - od -’ - -(** Unit test for [hashmap_main::hashmap::test1] *) -val _ = assert_return (“hashmap_test1_fwd”) - -val insert_on_disk_fwd_def = Define ‘ - (** [hashmap_main::insert_on_disk]: forward function *) - insert_on_disk_fwd - (key : usize) (value : u64) (st : state) : (state # unit) result = - do - (st0, hm) <- hashmap_utils_deserialize_fwd st; - hm0 <- hashmap_hash_map_insert_fwd_back hm key value; - (st1, _) <- hashmap_utils_serialize_fwd hm0 st0; - Return (st1, ()) - od -’ - -val main_fwd_def = Define ‘ - (** [hashmap_main::main]: forward function *) - main_fwd : unit result = - Return () -’ - -(** Unit test for [hashmap_main::main] *) -val _ = assert_return (“main_fwd”) - -val _ = export_theory () diff --git a/tests/hol4/hashmap_on_disk/hashmapMain_FunsTheory.sig b/tests/hol4/hashmap_on_disk/hashmapMain_FunsTheory.sig deleted file mode 100644 index d4e43d9a..00000000 --- a/tests/hol4/hashmap_on_disk/hashmapMain_FunsTheory.sig +++ /dev/null @@ -1,598 +0,0 @@ -signature hashmapMain_FunsTheory = -sig - type thm = Thm.thm - - (* Definitions *) - val hashmap_hash_key_fwd_def : thm - val hashmap_hash_map_allocate_slots_fwd_def : thm - val hashmap_hash_map_allocate_slots_loop_fwd_def : thm - val hashmap_hash_map_clear_fwd_back_def : thm - val hashmap_hash_map_clear_loop_fwd_back_def : thm - val hashmap_hash_map_contains_key_fwd_def : thm - val hashmap_hash_map_contains_key_in_list_fwd_def : thm - val hashmap_hash_map_contains_key_in_list_loop_fwd_def : thm - val hashmap_hash_map_get_fwd_def : thm - val hashmap_hash_map_get_in_list_fwd_def : thm - val hashmap_hash_map_get_in_list_loop_fwd_def : thm - val hashmap_hash_map_get_mut_back_def : thm - val hashmap_hash_map_get_mut_fwd_def : thm - val hashmap_hash_map_get_mut_in_list_back_def : thm - val hashmap_hash_map_get_mut_in_list_fwd_def : thm - val hashmap_hash_map_get_mut_in_list_loop_back_def : thm - val hashmap_hash_map_get_mut_in_list_loop_fwd_def : thm - val hashmap_hash_map_insert_fwd_back_def : thm - val hashmap_hash_map_insert_in_list_back_def : thm - val hashmap_hash_map_insert_in_list_fwd_def : thm - val hashmap_hash_map_insert_in_list_loop_back_def : thm - val hashmap_hash_map_insert_in_list_loop_fwd_def : thm - val hashmap_hash_map_insert_no_resize_fwd_back_def : thm - val hashmap_hash_map_len_fwd_def : thm - val hashmap_hash_map_move_elements_from_list_fwd_back_def : thm - val hashmap_hash_map_move_elements_from_list_loop_fwd_back_def : thm - val hashmap_hash_map_move_elements_fwd_back_def : thm - val hashmap_hash_map_move_elements_loop_fwd_back_def : thm - val hashmap_hash_map_new_fwd_def : thm - val hashmap_hash_map_new_with_capacity_fwd_def : thm - val hashmap_hash_map_remove_back_def : thm - val hashmap_hash_map_remove_from_list_back_def : thm - val hashmap_hash_map_remove_from_list_fwd_def : thm - val hashmap_hash_map_remove_from_list_loop_back_def : thm - val hashmap_hash_map_remove_from_list_loop_fwd_def : thm - val hashmap_hash_map_remove_fwd_def : thm - val hashmap_hash_map_try_resize_fwd_back_def : thm - val hashmap_test1_fwd_def : thm - val insert_on_disk_fwd_def : thm - val main_fwd_def : thm - - val hashmapMain_Funs_grammars : type_grammar.grammar * term_grammar.grammar -(* - [hashmapMain_Opaque] Parent theory of "hashmapMain_Funs" - - [hashmap_hash_key_fwd_def] Definition - - ⊢ ∀k. hashmap_hash_key_fwd k = Return k - - [hashmap_hash_map_allocate_slots_fwd_def] Definition - - ⊢ ∀slots n. - hashmap_hash_map_allocate_slots_fwd slots n = - hashmap_hash_map_allocate_slots_loop_fwd slots n - - [hashmap_hash_map_allocate_slots_loop_fwd_def] Definition - - ⊢ ∀slots n. - hashmap_hash_map_allocate_slots_loop_fwd slots n = - if usize_gt n (int_to_usize 0) then - do - slots0 <- vec_push_back slots HashmapListNil; - n0 <- usize_sub n (int_to_usize 1); - hashmap_hash_map_allocate_slots_loop_fwd slots0 n0 - od - else Return slots - - [hashmap_hash_map_clear_fwd_back_def] Definition - - ⊢ ∀self. - hashmap_hash_map_clear_fwd_back self = - do - v <- - hashmap_hash_map_clear_loop_fwd_back - self.hashmap_hash_map_slots (int_to_usize 0); - Return - (self with - <|hashmap_hash_map_num_entries := int_to_usize 0; - hashmap_hash_map_slots := v|>) - od - - [hashmap_hash_map_clear_loop_fwd_back_def] Definition - - ⊢ ∀slots i. - hashmap_hash_map_clear_loop_fwd_back slots i = - (let - i0 = vec_len slots - in - if usize_lt i i0 then - do - i1 <- usize_add i (int_to_usize 1); - slots0 <- vec_index_mut_back slots i HashmapListNil; - hashmap_hash_map_clear_loop_fwd_back slots0 i1 - od - else Return slots) - - [hashmap_hash_map_contains_key_fwd_def] Definition - - ⊢ ∀self key. - hashmap_hash_map_contains_key_fwd self key = - do - hash <- hashmap_hash_key_fwd key; - i <<- vec_len self.hashmap_hash_map_slots; - hash_mod <- usize_rem hash i; - l <- vec_index_fwd self.hashmap_hash_map_slots hash_mod; - hashmap_hash_map_contains_key_in_list_fwd key l - od - - [hashmap_hash_map_contains_key_in_list_fwd_def] Definition - - ⊢ ∀key ls. - hashmap_hash_map_contains_key_in_list_fwd key ls = - hashmap_hash_map_contains_key_in_list_loop_fwd key ls - - [hashmap_hash_map_contains_key_in_list_loop_fwd_def] Definition - - ⊢ ∀key ls. - hashmap_hash_map_contains_key_in_list_loop_fwd key ls = - case ls of - HashmapListCons ckey t tl => - if ckey = key then Return T - else hashmap_hash_map_contains_key_in_list_loop_fwd key tl - | HashmapListNil => Return F - - [hashmap_hash_map_get_fwd_def] Definition - - ⊢ ∀self key. - hashmap_hash_map_get_fwd self key = - do - hash <- hashmap_hash_key_fwd key; - i <<- vec_len self.hashmap_hash_map_slots; - hash_mod <- usize_rem hash i; - l <- vec_index_fwd self.hashmap_hash_map_slots hash_mod; - hashmap_hash_map_get_in_list_fwd key l - od - - [hashmap_hash_map_get_in_list_fwd_def] Definition - - ⊢ ∀key ls. - hashmap_hash_map_get_in_list_fwd key ls = - hashmap_hash_map_get_in_list_loop_fwd key ls - - [hashmap_hash_map_get_in_list_loop_fwd_def] Definition - - ⊢ ∀key ls. - hashmap_hash_map_get_in_list_loop_fwd key ls = - case ls of - HashmapListCons ckey cvalue tl => - if ckey = key then Return cvalue - else hashmap_hash_map_get_in_list_loop_fwd key tl - | HashmapListNil => Fail Failure - - [hashmap_hash_map_get_mut_back_def] Definition - - ⊢ ∀self key ret. - hashmap_hash_map_get_mut_back self key ret = - do - hash <- hashmap_hash_key_fwd key; - i <<- vec_len self.hashmap_hash_map_slots; - hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; - l0 <- hashmap_hash_map_get_mut_in_list_back l key ret; - v <- vec_index_mut_back self.hashmap_hash_map_slots hash_mod l0; - Return (self with hashmap_hash_map_slots := v) - od - - [hashmap_hash_map_get_mut_fwd_def] Definition - - ⊢ ∀self key. - hashmap_hash_map_get_mut_fwd self key = - do - hash <- hashmap_hash_key_fwd key; - i <<- vec_len self.hashmap_hash_map_slots; - hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; - hashmap_hash_map_get_mut_in_list_fwd l key - od - - [hashmap_hash_map_get_mut_in_list_back_def] Definition - - ⊢ ∀ls key ret. - hashmap_hash_map_get_mut_in_list_back ls key ret = - hashmap_hash_map_get_mut_in_list_loop_back ls key ret - - [hashmap_hash_map_get_mut_in_list_fwd_def] Definition - - ⊢ ∀ls key. - hashmap_hash_map_get_mut_in_list_fwd ls key = - hashmap_hash_map_get_mut_in_list_loop_fwd ls key - - [hashmap_hash_map_get_mut_in_list_loop_back_def] Definition - - ⊢ ∀ls key ret. - hashmap_hash_map_get_mut_in_list_loop_back ls key ret = - case ls of - HashmapListCons ckey cvalue tl => - if ckey = key then Return (HashmapListCons ckey ret tl) - else - do - tl0 <- - hashmap_hash_map_get_mut_in_list_loop_back tl key ret; - Return (HashmapListCons ckey cvalue tl0) - od - | HashmapListNil => Fail Failure - - [hashmap_hash_map_get_mut_in_list_loop_fwd_def] Definition - - ⊢ ∀ls key. - hashmap_hash_map_get_mut_in_list_loop_fwd ls key = - case ls of - HashmapListCons ckey cvalue tl => - if ckey = key then Return cvalue - else hashmap_hash_map_get_mut_in_list_loop_fwd tl key - | HashmapListNil => Fail Failure - - [hashmap_hash_map_insert_fwd_back_def] Definition - - ⊢ ∀self key value. - hashmap_hash_map_insert_fwd_back self key value = - do - self0 <- - hashmap_hash_map_insert_no_resize_fwd_back self key value; - i <- hashmap_hash_map_len_fwd self0; - if usize_gt i self0.hashmap_hash_map_max_load then - hashmap_hash_map_try_resize_fwd_back self0 - else Return self0 - od - - [hashmap_hash_map_insert_in_list_back_def] Definition - - ⊢ ∀key value ls. - hashmap_hash_map_insert_in_list_back key value ls = - hashmap_hash_map_insert_in_list_loop_back key value ls - - [hashmap_hash_map_insert_in_list_fwd_def] Definition - - ⊢ ∀key value ls. - hashmap_hash_map_insert_in_list_fwd key value ls = - hashmap_hash_map_insert_in_list_loop_fwd key value ls - - [hashmap_hash_map_insert_in_list_loop_back_def] Definition - - ⊢ ∀key value ls. - hashmap_hash_map_insert_in_list_loop_back key value ls = - case ls of - HashmapListCons ckey cvalue tl => - if ckey = key then Return (HashmapListCons ckey value tl) - else - do - tl0 <- - hashmap_hash_map_insert_in_list_loop_back key value tl; - Return (HashmapListCons ckey cvalue tl0) - od - | HashmapListNil => - (let - l = HashmapListNil - in - Return (HashmapListCons key value l)) - - [hashmap_hash_map_insert_in_list_loop_fwd_def] Definition - - ⊢ ∀key value ls. - hashmap_hash_map_insert_in_list_loop_fwd key value ls = - case ls of - HashmapListCons ckey cvalue tl => - if ckey = key then Return F - else hashmap_hash_map_insert_in_list_loop_fwd key value tl - | HashmapListNil => Return T - - [hashmap_hash_map_insert_no_resize_fwd_back_def] Definition - - ⊢ ∀self key value. - hashmap_hash_map_insert_no_resize_fwd_back self key value = - do - hash <- hashmap_hash_key_fwd key; - i <<- vec_len self.hashmap_hash_map_slots; - hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; - inserted <- hashmap_hash_map_insert_in_list_fwd key value l; - if inserted then - do - i0 <- - usize_add self.hashmap_hash_map_num_entries - (int_to_usize 1); - l0 <- hashmap_hash_map_insert_in_list_back key value l; - v <- - vec_index_mut_back self.hashmap_hash_map_slots hash_mod - l0; - Return - (self with - <|hashmap_hash_map_num_entries := i0; - hashmap_hash_map_slots := v|>) - od - else - do - l0 <- hashmap_hash_map_insert_in_list_back key value l; - v <- - vec_index_mut_back self.hashmap_hash_map_slots hash_mod - l0; - Return (self with hashmap_hash_map_slots := v) - od - od - - [hashmap_hash_map_len_fwd_def] Definition - - ⊢ ∀self. - hashmap_hash_map_len_fwd self = - Return self.hashmap_hash_map_num_entries - - [hashmap_hash_map_move_elements_from_list_fwd_back_def] Definition - - ⊢ ∀ntable ls. - hashmap_hash_map_move_elements_from_list_fwd_back ntable ls = - hashmap_hash_map_move_elements_from_list_loop_fwd_back ntable ls - - [hashmap_hash_map_move_elements_from_list_loop_fwd_back_def] Definition - - ⊢ ∀ntable ls. - hashmap_hash_map_move_elements_from_list_loop_fwd_back ntable ls = - case ls of - HashmapListCons k v tl => - do - ntable0 <- - hashmap_hash_map_insert_no_resize_fwd_back ntable k v; - hashmap_hash_map_move_elements_from_list_loop_fwd_back - ntable0 tl - od - | HashmapListNil => Return ntable - - [hashmap_hash_map_move_elements_fwd_back_def] Definition - - ⊢ ∀ntable slots i. - hashmap_hash_map_move_elements_fwd_back ntable slots i = - hashmap_hash_map_move_elements_loop_fwd_back ntable slots i - - [hashmap_hash_map_move_elements_loop_fwd_back_def] Definition - - ⊢ ∀ntable slots i. - hashmap_hash_map_move_elements_loop_fwd_back ntable slots i = - (let - i0 = vec_len slots - in - if usize_lt i i0 then - do - l <- vec_index_mut_fwd slots i; - ls <<- mem_replace_fwd l HashmapListNil; - ntable0 <- - hashmap_hash_map_move_elements_from_list_fwd_back ntable - ls; - i1 <- usize_add i (int_to_usize 1); - l0 <<- mem_replace_back l HashmapListNil; - slots0 <- vec_index_mut_back slots i l0; - hashmap_hash_map_move_elements_loop_fwd_back ntable0 - slots0 i1 - od - else Return (ntable,slots)) - - [hashmap_hash_map_new_fwd_def] Definition - - ⊢ hashmap_hash_map_new_fwd = - hashmap_hash_map_new_with_capacity_fwd (int_to_usize 32) - (int_to_usize 4) (int_to_usize 5) - - [hashmap_hash_map_new_with_capacity_fwd_def] Definition - - ⊢ ∀capacity max_load_dividend max_load_divisor. - hashmap_hash_map_new_with_capacity_fwd capacity max_load_dividend - max_load_divisor = - (let - v = vec_new - in - do - slots <- hashmap_hash_map_allocate_slots_fwd v capacity; - i <- usize_mul capacity max_load_dividend; - i0 <- usize_div i max_load_divisor; - Return - <|hashmap_hash_map_num_entries := int_to_usize 0; - hashmap_hash_map_max_load_factor := - (max_load_dividend,max_load_divisor); - hashmap_hash_map_max_load := i0; - hashmap_hash_map_slots := slots|> - od) - - [hashmap_hash_map_remove_back_def] Definition - - ⊢ ∀self key. - hashmap_hash_map_remove_back self key = - do - hash <- hashmap_hash_key_fwd key; - i <<- vec_len self.hashmap_hash_map_slots; - hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; - x <- hashmap_hash_map_remove_from_list_fwd key l; - case x of - NONE => - do - l0 <- hashmap_hash_map_remove_from_list_back key l; - v <- - vec_index_mut_back self.hashmap_hash_map_slots hash_mod - l0; - Return (self with hashmap_hash_map_slots := v) - od - | SOME x0 => - do - i0 <- - usize_sub self.hashmap_hash_map_num_entries - (int_to_usize 1); - l0 <- hashmap_hash_map_remove_from_list_back key l; - v <- - vec_index_mut_back self.hashmap_hash_map_slots hash_mod - l0; - Return - (self with - <|hashmap_hash_map_num_entries := i0; - hashmap_hash_map_slots := v|>) - od - od - - [hashmap_hash_map_remove_from_list_back_def] Definition - - ⊢ ∀key ls. - hashmap_hash_map_remove_from_list_back key ls = - hashmap_hash_map_remove_from_list_loop_back key ls - - [hashmap_hash_map_remove_from_list_fwd_def] Definition - - ⊢ ∀key ls. - hashmap_hash_map_remove_from_list_fwd key ls = - hashmap_hash_map_remove_from_list_loop_fwd key ls - - [hashmap_hash_map_remove_from_list_loop_back_def] Definition - - ⊢ ∀key ls. - hashmap_hash_map_remove_from_list_loop_back key ls = - case ls of - HashmapListCons ckey t tl => - if ckey = key then - (let - mv_ls = - mem_replace_fwd (HashmapListCons ckey t tl) - HashmapListNil - in - case mv_ls of - HashmapListCons i cvalue tl0 => Return tl0 - | HashmapListNil => Fail Failure) - else - do - tl0 <- hashmap_hash_map_remove_from_list_loop_back key tl; - Return (HashmapListCons ckey t tl0) - od - | HashmapListNil => Return HashmapListNil - - [hashmap_hash_map_remove_from_list_loop_fwd_def] Definition - - ⊢ ∀key ls. - hashmap_hash_map_remove_from_list_loop_fwd key ls = - case ls of - HashmapListCons ckey t tl => - if ckey = key then - (let - mv_ls = - mem_replace_fwd (HashmapListCons ckey t tl) - HashmapListNil - in - case mv_ls of - HashmapListCons i cvalue tl0 => Return (SOME cvalue) - | HashmapListNil => Fail Failure) - else hashmap_hash_map_remove_from_list_loop_fwd key tl - | HashmapListNil => Return NONE - - [hashmap_hash_map_remove_fwd_def] Definition - - ⊢ ∀self key. - hashmap_hash_map_remove_fwd self key = - do - hash <- hashmap_hash_key_fwd key; - i <<- vec_len self.hashmap_hash_map_slots; - hash_mod <- usize_rem hash i; - l <- vec_index_mut_fwd self.hashmap_hash_map_slots hash_mod; - x <- hashmap_hash_map_remove_from_list_fwd key l; - case x of - NONE => Return NONE - | SOME x0 => - monad_ignore_bind - (usize_sub self.hashmap_hash_map_num_entries - (int_to_usize 1)) (Return (SOME x0)) - od - - [hashmap_hash_map_try_resize_fwd_back_def] Definition - - ⊢ ∀self. - hashmap_hash_map_try_resize_fwd_back self = - do - max_usize <- mk_usize (u32_to_int core_u32_max); - capacity <<- vec_len self.hashmap_hash_map_slots; - n1 <- usize_div max_usize (int_to_usize 2); - (i,i0) <<- self.hashmap_hash_map_max_load_factor; - i1 <- usize_div n1 i; - if usize_le capacity i1 then - do - i2 <- usize_mul capacity (int_to_usize 2); - ntable <- hashmap_hash_map_new_with_capacity_fwd i2 i i0; - (ntable0,_) <- - hashmap_hash_map_move_elements_fwd_back ntable - self.hashmap_hash_map_slots (int_to_usize 0); - Return - (ntable0 with - <|hashmap_hash_map_num_entries := - self.hashmap_hash_map_num_entries; - hashmap_hash_map_max_load_factor := (i,i0)|>) - od - else - Return (self with hashmap_hash_map_max_load_factor := (i,i0)) - od - - [hashmap_test1_fwd_def] Definition - - ⊢ hashmap_test1_fwd = - do - hm <- hashmap_hash_map_new_fwd; - hm0 <- - hashmap_hash_map_insert_fwd_back hm (int_to_usize 0) - (int_to_u64 42); - hm1 <- - hashmap_hash_map_insert_fwd_back hm0 (int_to_usize 128) - (int_to_u64 18); - hm2 <- - hashmap_hash_map_insert_fwd_back hm1 (int_to_usize 1024) - (int_to_u64 138); - hm3 <- - hashmap_hash_map_insert_fwd_back hm2 (int_to_usize 1056) - (int_to_u64 256); - i <- hashmap_hash_map_get_fwd hm3 (int_to_usize 128); - if i ≠ int_to_u64 18 then Fail Failure - else - do - hm4 <- - hashmap_hash_map_get_mut_back hm3 (int_to_usize 1024) - (int_to_u64 56); - i0 <- hashmap_hash_map_get_fwd hm4 (int_to_usize 1024); - if i0 ≠ int_to_u64 56 then Fail Failure - else - do - x <- hashmap_hash_map_remove_fwd hm4 (int_to_usize 1024); - case x of - NONE => Fail Failure - | SOME x0 => - if x0 ≠ int_to_u64 56 then Fail Failure - else - do - hm5 <- - hashmap_hash_map_remove_back hm4 - (int_to_usize 1024); - i1 <- hashmap_hash_map_get_fwd hm5 (int_to_usize 0); - if i1 ≠ int_to_u64 42 then Fail Failure - else - do - i2 <- - hashmap_hash_map_get_fwd hm5 - (int_to_usize 128); - if i2 ≠ int_to_u64 18 then Fail Failure - else - do - i3 <- - hashmap_hash_map_get_fwd hm5 - (int_to_usize 1056); - if i3 ≠ int_to_u64 256 then Fail Failure - else Return () - od - od - od - od - od - od - - [insert_on_disk_fwd_def] Definition - - ⊢ ∀key value st. - insert_on_disk_fwd key value st = - do - (st0,hm) <- hashmap_utils_deserialize_fwd st; - hm0 <- hashmap_hash_map_insert_fwd_back hm key value; - (st1,_) <- hashmap_utils_serialize_fwd hm0 st0; - Return (st1,()) - od - - [main_fwd_def] Definition - - ⊢ main_fwd = Return () - - -*) -end diff --git a/tests/hol4/hashmap_on_disk/hashmapMain_OpaqueScript.sml b/tests/hol4/hashmap_on_disk/hashmapMain_OpaqueScript.sml deleted file mode 100644 index f7221d92..00000000 --- a/tests/hol4/hashmap_on_disk/hashmapMain_OpaqueScript.sml +++ /dev/null @@ -1,15 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: external function declarations *) -open primitivesLib divDefLib -open hashmapMain_TypesTheory - -val _ = new_theory "hashmapMain_Opaque" - - -(** [hashmap_main::hashmap_utils::deserialize]: forward function *)val _ = new_constant ("hashmap_utils_deserialize_fwd", - “:state -> (state # u64 hashmap_hash_map_t) result”) - -(** [hashmap_main::hashmap_utils::serialize]: forward function *)val _ = new_constant ("hashmap_utils_serialize_fwd", - “:u64 hashmap_hash_map_t -> state -> (state # unit) result”) - -val _ = export_theory () diff --git a/tests/hol4/hashmap_on_disk/hashmapMain_OpaqueTheory.sig b/tests/hol4/hashmap_on_disk/hashmapMain_OpaqueTheory.sig deleted file mode 100644 index f7373609..00000000 --- a/tests/hol4/hashmap_on_disk/hashmapMain_OpaqueTheory.sig +++ /dev/null @@ -1,11 +0,0 @@ -signature hashmapMain_OpaqueTheory = -sig - type thm = Thm.thm - - val hashmapMain_Opaque_grammars : type_grammar.grammar * term_grammar.grammar -(* - [hashmapMain_Types] Parent theory of "hashmapMain_Opaque" - - -*) -end diff --git a/tests/hol4/hashmap_on_disk/hashmapMain_TypesScript.sml b/tests/hol4/hashmap_on_disk/hashmapMain_TypesScript.sml deleted file mode 100644 index 3f8ca9b9..00000000 --- a/tests/hol4/hashmap_on_disk/hashmapMain_TypesScript.sml +++ /dev/null @@ -1,27 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: type definitions *) -open primitivesLib divDefLib - -val _ = new_theory "hashmapMain_Types" - - -Datatype: - (** [hashmap_main::hashmap::List] *) - hashmap_list_t = | HashmapListCons usize 't hashmap_list_t | HashmapListNil -End - -Datatype: - (** [hashmap_main::hashmap::HashMap] *) - hashmap_hash_map_t = - <| - hashmap_hash_map_num_entries : usize; - hashmap_hash_map_max_load_factor : (usize # usize); - hashmap_hash_map_max_load : usize; - hashmap_hash_map_slots : 't hashmap_list_t vec; - |> -End - -(** The state type used in the state-error monad *) -val _ = new_type ("state", 0) - -val _ = export_theory () diff --git a/tests/hol4/hashmap_on_disk/hashmapMain_TypesTheory.sig b/tests/hol4/hashmap_on_disk/hashmapMain_TypesTheory.sig deleted file mode 100644 index a3e770ea..00000000 --- a/tests/hol4/hashmap_on_disk/hashmapMain_TypesTheory.sig +++ /dev/null @@ -1,568 +0,0 @@ -signature hashmapMain_TypesTheory = -sig - type thm = Thm.thm - - (* Definitions *) - val hashmap_hash_map_t_TY_DEF : thm - val hashmap_hash_map_t_case_def : thm - val hashmap_hash_map_t_hashmap_hash_map_max_load : thm - val hashmap_hash_map_t_hashmap_hash_map_max_load_factor : thm - val hashmap_hash_map_t_hashmap_hash_map_max_load_factor_fupd : thm - val hashmap_hash_map_t_hashmap_hash_map_max_load_fupd : thm - val hashmap_hash_map_t_hashmap_hash_map_num_entries : thm - val hashmap_hash_map_t_hashmap_hash_map_num_entries_fupd : thm - val hashmap_hash_map_t_hashmap_hash_map_slots : thm - val hashmap_hash_map_t_hashmap_hash_map_slots_fupd : thm - val hashmap_hash_map_t_size_def : thm - val hashmap_list_t_TY_DEF : thm - val hashmap_list_t_case_def : thm - val hashmap_list_t_size_def : thm - - (* Theorems *) - val EXISTS_hashmap_hash_map_t : thm - val FORALL_hashmap_hash_map_t : thm - val datatype_hashmap_hash_map_t : thm - val datatype_hashmap_list_t : thm - val hashmap_hash_map_t_11 : thm - val hashmap_hash_map_t_Axiom : thm - val hashmap_hash_map_t_accessors : thm - val hashmap_hash_map_t_accfupds : thm - val hashmap_hash_map_t_case_cong : thm - val hashmap_hash_map_t_case_eq : thm - val hashmap_hash_map_t_component_equality : thm - val hashmap_hash_map_t_fn_updates : thm - val hashmap_hash_map_t_fupdcanon : thm - val hashmap_hash_map_t_fupdcanon_comp : thm - val hashmap_hash_map_t_fupdfupds : thm - val hashmap_hash_map_t_fupdfupds_comp : thm - val hashmap_hash_map_t_induction : thm - val hashmap_hash_map_t_literal_11 : thm - val hashmap_hash_map_t_literal_nchotomy : thm - val hashmap_hash_map_t_nchotomy : thm - val hashmap_hash_map_t_updates_eq_literal : thm - val hashmap_list_t_11 : thm - val hashmap_list_t_Axiom : thm - val hashmap_list_t_case_cong : thm - val hashmap_list_t_case_eq : thm - val hashmap_list_t_distinct : thm - val hashmap_list_t_induction : thm - val hashmap_list_t_nchotomy : thm - - val hashmapMain_Types_grammars : type_grammar.grammar * term_grammar.grammar -(* - [divDef] Parent theory of "hashmapMain_Types" - - [hashmap_hash_map_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('hashmap_hash_map_t'). - (∀a0'. - (∃a0 a1 a2 a3. - a0' = - (λa0 a1 a2 a3. - ind_type$CONSTR 0 (a0,a1,a2,a3) - (λn. ind_type$BOTTOM)) a0 a1 a2 a3) ⇒ - $var$('hashmap_hash_map_t') a0') ⇒ - $var$('hashmap_hash_map_t') a0') rep - - [hashmap_hash_map_t_case_def] Definition - - ⊢ ∀a0 a1 a2 a3 f. - hashmap_hash_map_t_CASE (hashmap_hash_map_t a0 a1 a2 a3) f = - f a0 a1 a2 a3 - - [hashmap_hash_map_t_hashmap_hash_map_max_load] Definition - - ⊢ ∀u p u0 v. - (hashmap_hash_map_t u p u0 v).hashmap_hash_map_max_load = u0 - - [hashmap_hash_map_t_hashmap_hash_map_max_load_factor] Definition - - ⊢ ∀u p u0 v. - (hashmap_hash_map_t u p u0 v).hashmap_hash_map_max_load_factor = - p - - [hashmap_hash_map_t_hashmap_hash_map_max_load_factor_fupd] Definition - - ⊢ ∀f u p u0 v. - hashmap_hash_map_t u p u0 v with - hashmap_hash_map_max_load_factor updated_by f = - hashmap_hash_map_t u (f p) u0 v - - [hashmap_hash_map_t_hashmap_hash_map_max_load_fupd] Definition - - ⊢ ∀f u p u0 v. - hashmap_hash_map_t u p u0 v with - hashmap_hash_map_max_load updated_by f = - hashmap_hash_map_t u p (f u0) v - - [hashmap_hash_map_t_hashmap_hash_map_num_entries] Definition - - ⊢ ∀u p u0 v. - (hashmap_hash_map_t u p u0 v).hashmap_hash_map_num_entries = u - - [hashmap_hash_map_t_hashmap_hash_map_num_entries_fupd] Definition - - ⊢ ∀f u p u0 v. - hashmap_hash_map_t u p u0 v with - hashmap_hash_map_num_entries updated_by f = - hashmap_hash_map_t (f u) p u0 v - - [hashmap_hash_map_t_hashmap_hash_map_slots] Definition - - ⊢ ∀u p u0 v. (hashmap_hash_map_t u p u0 v).hashmap_hash_map_slots = v - - [hashmap_hash_map_t_hashmap_hash_map_slots_fupd] Definition - - ⊢ ∀f u p u0 v. - hashmap_hash_map_t u p u0 v with - hashmap_hash_map_slots updated_by f = - hashmap_hash_map_t u p u0 (f v) - - [hashmap_hash_map_t_size_def] Definition - - ⊢ ∀f a0 a1 a2 a3. - hashmap_hash_map_t_size f (hashmap_hash_map_t a0 a1 a2 a3) = - 1 + pair_size (λv. 0) (λv. 0) a1 - - [hashmap_list_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('hashmap_list_t'). - (∀a0'. - (∃a0 a1 a2. - a0' = - (λa0 a1 a2. - ind_type$CONSTR 0 (a0,a1) - (ind_type$FCONS a2 (λn. ind_type$BOTTOM))) - a0 a1 a2 ∧ $var$('hashmap_list_t') a2) ∨ - a0' = - ind_type$CONSTR (SUC 0) (ARB,ARB) - (λn. ind_type$BOTTOM) ⇒ - $var$('hashmap_list_t') a0') ⇒ - $var$('hashmap_list_t') a0') rep - - [hashmap_list_t_case_def] Definition - - ⊢ (∀a0 a1 a2 f v. - hashmap_list_t_CASE (HashmapListCons a0 a1 a2) f v = f a0 a1 a2) ∧ - ∀f v. hashmap_list_t_CASE HashmapListNil f v = v - - [hashmap_list_t_size_def] Definition - - ⊢ (∀f a0 a1 a2. - hashmap_list_t_size f (HashmapListCons a0 a1 a2) = - 1 + (f a1 + hashmap_list_t_size f a2)) ∧ - ∀f. hashmap_list_t_size f HashmapListNil = 0 - - [EXISTS_hashmap_hash_map_t] Theorem - - ⊢ ∀P. (∃h. P h) ⇔ - ∃u0 p u v. - P - <|hashmap_hash_map_num_entries := u0; - hashmap_hash_map_max_load_factor := p; - hashmap_hash_map_max_load := u; - hashmap_hash_map_slots := v|> - - [FORALL_hashmap_hash_map_t] Theorem - - ⊢ ∀P. (∀h. P h) ⇔ - ∀u0 p u v. - P - <|hashmap_hash_map_num_entries := u0; - hashmap_hash_map_max_load_factor := p; - hashmap_hash_map_max_load := u; - hashmap_hash_map_slots := v|> - - [datatype_hashmap_hash_map_t] Theorem - - ⊢ DATATYPE - (record hashmap_hash_map_t hashmap_hash_map_num_entries - hashmap_hash_map_max_load_factor hashmap_hash_map_max_load - hashmap_hash_map_slots) - - [datatype_hashmap_list_t] Theorem - - ⊢ DATATYPE (hashmap_list_t HashmapListCons HashmapListNil) - - [hashmap_hash_map_t_11] Theorem - - ⊢ ∀a0 a1 a2 a3 a0' a1' a2' a3'. - hashmap_hash_map_t a0 a1 a2 a3 = - hashmap_hash_map_t a0' a1' a2' a3' ⇔ - a0 = a0' ∧ a1 = a1' ∧ a2 = a2' ∧ a3 = a3' - - [hashmap_hash_map_t_Axiom] Theorem - - ⊢ ∀f. ∃fn. ∀a0 a1 a2 a3. - fn (hashmap_hash_map_t a0 a1 a2 a3) = f a0 a1 a2 a3 - - [hashmap_hash_map_t_accessors] Theorem - - ⊢ (∀u p u0 v. - (hashmap_hash_map_t u p u0 v).hashmap_hash_map_num_entries = u) ∧ - (∀u p u0 v. - (hashmap_hash_map_t u p u0 v).hashmap_hash_map_max_load_factor = - p) ∧ - (∀u p u0 v. - (hashmap_hash_map_t u p u0 v).hashmap_hash_map_max_load = u0) ∧ - ∀u p u0 v. (hashmap_hash_map_t u p u0 v).hashmap_hash_map_slots = v - - [hashmap_hash_map_t_accfupds] Theorem - - ⊢ (∀h f. - (h with hashmap_hash_map_max_load_factor updated_by f). - hashmap_hash_map_num_entries = - h.hashmap_hash_map_num_entries) ∧ - (∀h f. - (h with hashmap_hash_map_max_load updated_by f). - hashmap_hash_map_num_entries = - h.hashmap_hash_map_num_entries) ∧ - (∀h f. - (h with hashmap_hash_map_slots updated_by f). - hashmap_hash_map_num_entries = - h.hashmap_hash_map_num_entries) ∧ - (∀h f. - (h with hashmap_hash_map_num_entries updated_by f). - hashmap_hash_map_max_load_factor = - h.hashmap_hash_map_max_load_factor) ∧ - (∀h f. - (h with hashmap_hash_map_max_load updated_by f). - hashmap_hash_map_max_load_factor = - h.hashmap_hash_map_max_load_factor) ∧ - (∀h f. - (h with hashmap_hash_map_slots updated_by f). - hashmap_hash_map_max_load_factor = - h.hashmap_hash_map_max_load_factor) ∧ - (∀h f. - (h with hashmap_hash_map_num_entries updated_by f). - hashmap_hash_map_max_load = - h.hashmap_hash_map_max_load) ∧ - (∀h f. - (h with hashmap_hash_map_max_load_factor updated_by f). - hashmap_hash_map_max_load = - h.hashmap_hash_map_max_load) ∧ - (∀h f. - (h with hashmap_hash_map_slots updated_by f). - hashmap_hash_map_max_load = - h.hashmap_hash_map_max_load) ∧ - (∀h f. - (h with hashmap_hash_map_num_entries updated_by f). - hashmap_hash_map_slots = - h.hashmap_hash_map_slots) ∧ - (∀h f. - (h with hashmap_hash_map_max_load_factor updated_by f). - hashmap_hash_map_slots = - h.hashmap_hash_map_slots) ∧ - (∀h f. - (h with hashmap_hash_map_max_load updated_by f). - hashmap_hash_map_slots = - h.hashmap_hash_map_slots) ∧ - (∀h f. - (h with hashmap_hash_map_num_entries updated_by f). - hashmap_hash_map_num_entries = - f h.hashmap_hash_map_num_entries) ∧ - (∀h f. - (h with hashmap_hash_map_max_load_factor updated_by f). - hashmap_hash_map_max_load_factor = - f h.hashmap_hash_map_max_load_factor) ∧ - (∀h f. - (h with hashmap_hash_map_max_load updated_by f). - hashmap_hash_map_max_load = - f h.hashmap_hash_map_max_load) ∧ - ∀h f. - (h with hashmap_hash_map_slots updated_by f). - hashmap_hash_map_slots = - f h.hashmap_hash_map_slots - - [hashmap_hash_map_t_case_cong] Theorem - - ⊢ ∀M M' f. - M = M' ∧ - (∀a0 a1 a2 a3. - M' = hashmap_hash_map_t a0 a1 a2 a3 ⇒ - f a0 a1 a2 a3 = f' a0 a1 a2 a3) ⇒ - hashmap_hash_map_t_CASE M f = hashmap_hash_map_t_CASE M' f' - - [hashmap_hash_map_t_case_eq] Theorem - - ⊢ hashmap_hash_map_t_CASE x f = v ⇔ - ∃u p u0 v'. x = hashmap_hash_map_t u p u0 v' ∧ f u p u0 v' = v - - [hashmap_hash_map_t_component_equality] Theorem - - ⊢ ∀h1 h2. - h1 = h2 ⇔ - h1.hashmap_hash_map_num_entries = h2.hashmap_hash_map_num_entries ∧ - h1.hashmap_hash_map_max_load_factor = - h2.hashmap_hash_map_max_load_factor ∧ - h1.hashmap_hash_map_max_load = h2.hashmap_hash_map_max_load ∧ - h1.hashmap_hash_map_slots = h2.hashmap_hash_map_slots - - [hashmap_hash_map_t_fn_updates] Theorem - - ⊢ (∀f u p u0 v. - hashmap_hash_map_t u p u0 v with - hashmap_hash_map_num_entries updated_by f = - hashmap_hash_map_t (f u) p u0 v) ∧ - (∀f u p u0 v. - hashmap_hash_map_t u p u0 v with - hashmap_hash_map_max_load_factor updated_by f = - hashmap_hash_map_t u (f p) u0 v) ∧ - (∀f u p u0 v. - hashmap_hash_map_t u p u0 v with - hashmap_hash_map_max_load updated_by f = - hashmap_hash_map_t u p (f u0) v) ∧ - ∀f u p u0 v. - hashmap_hash_map_t u p u0 v with - hashmap_hash_map_slots updated_by f = - hashmap_hash_map_t u p u0 (f v) - - [hashmap_hash_map_t_fupdcanon] Theorem - - ⊢ (∀h g f. - h with - <|hashmap_hash_map_max_load_factor updated_by f; - hashmap_hash_map_num_entries updated_by g|> = - h with - <|hashmap_hash_map_num_entries updated_by g; - hashmap_hash_map_max_load_factor updated_by f|>) ∧ - (∀h g f. - h with - <|hashmap_hash_map_max_load updated_by f; - hashmap_hash_map_num_entries updated_by g|> = - h with - <|hashmap_hash_map_num_entries updated_by g; - hashmap_hash_map_max_load updated_by f|>) ∧ - (∀h g f. - h with - <|hashmap_hash_map_max_load updated_by f; - hashmap_hash_map_max_load_factor updated_by g|> = - h with - <|hashmap_hash_map_max_load_factor updated_by g; - hashmap_hash_map_max_load updated_by f|>) ∧ - (∀h g f. - h with - <|hashmap_hash_map_slots updated_by f; - hashmap_hash_map_num_entries updated_by g|> = - h with - <|hashmap_hash_map_num_entries updated_by g; - hashmap_hash_map_slots updated_by f|>) ∧ - (∀h g f. - h with - <|hashmap_hash_map_slots updated_by f; - hashmap_hash_map_max_load_factor updated_by g|> = - h with - <|hashmap_hash_map_max_load_factor updated_by g; - hashmap_hash_map_slots updated_by f|>) ∧ - ∀h g f. - h with - <|hashmap_hash_map_slots updated_by f; - hashmap_hash_map_max_load updated_by g|> = - h with - <|hashmap_hash_map_max_load updated_by g; - hashmap_hash_map_slots updated_by f|> - - [hashmap_hash_map_t_fupdcanon_comp] Theorem - - ⊢ ((∀g f. - hashmap_hash_map_max_load_factor_fupd f ∘ - hashmap_hash_map_num_entries_fupd g = - hashmap_hash_map_num_entries_fupd g ∘ - hashmap_hash_map_max_load_factor_fupd f) ∧ - ∀h g f. - hashmap_hash_map_max_load_factor_fupd f ∘ - hashmap_hash_map_num_entries_fupd g ∘ h = - hashmap_hash_map_num_entries_fupd g ∘ - hashmap_hash_map_max_load_factor_fupd f ∘ h) ∧ - ((∀g f. - hashmap_hash_map_max_load_fupd f ∘ - hashmap_hash_map_num_entries_fupd g = - hashmap_hash_map_num_entries_fupd g ∘ - hashmap_hash_map_max_load_fupd f) ∧ - ∀h g f. - hashmap_hash_map_max_load_fupd f ∘ - hashmap_hash_map_num_entries_fupd g ∘ h = - hashmap_hash_map_num_entries_fupd g ∘ - hashmap_hash_map_max_load_fupd f ∘ h) ∧ - ((∀g f. - hashmap_hash_map_max_load_fupd f ∘ - hashmap_hash_map_max_load_factor_fupd g = - hashmap_hash_map_max_load_factor_fupd g ∘ - hashmap_hash_map_max_load_fupd f) ∧ - ∀h g f. - hashmap_hash_map_max_load_fupd f ∘ - hashmap_hash_map_max_load_factor_fupd g ∘ h = - hashmap_hash_map_max_load_factor_fupd g ∘ - hashmap_hash_map_max_load_fupd f ∘ h) ∧ - ((∀g f. - hashmap_hash_map_slots_fupd f ∘ - hashmap_hash_map_num_entries_fupd g = - hashmap_hash_map_num_entries_fupd g ∘ - hashmap_hash_map_slots_fupd f) ∧ - ∀h g f. - hashmap_hash_map_slots_fupd f ∘ - hashmap_hash_map_num_entries_fupd g ∘ h = - hashmap_hash_map_num_entries_fupd g ∘ - hashmap_hash_map_slots_fupd f ∘ h) ∧ - ((∀g f. - hashmap_hash_map_slots_fupd f ∘ - hashmap_hash_map_max_load_factor_fupd g = - hashmap_hash_map_max_load_factor_fupd g ∘ - hashmap_hash_map_slots_fupd f) ∧ - ∀h g f. - hashmap_hash_map_slots_fupd f ∘ - hashmap_hash_map_max_load_factor_fupd g ∘ h = - hashmap_hash_map_max_load_factor_fupd g ∘ - hashmap_hash_map_slots_fupd f ∘ h) ∧ - (∀g f. - hashmap_hash_map_slots_fupd f ∘ hashmap_hash_map_max_load_fupd g = - hashmap_hash_map_max_load_fupd g ∘ hashmap_hash_map_slots_fupd f) ∧ - ∀h g f. - hashmap_hash_map_slots_fupd f ∘ - hashmap_hash_map_max_load_fupd g ∘ h = - hashmap_hash_map_max_load_fupd g ∘ - hashmap_hash_map_slots_fupd f ∘ h - - [hashmap_hash_map_t_fupdfupds] Theorem - - ⊢ (∀h g f. - h with - <|hashmap_hash_map_num_entries updated_by f; - hashmap_hash_map_num_entries updated_by g|> = - h with hashmap_hash_map_num_entries updated_by f ∘ g) ∧ - (∀h g f. - h with - <|hashmap_hash_map_max_load_factor updated_by f; - hashmap_hash_map_max_load_factor updated_by g|> = - h with hashmap_hash_map_max_load_factor updated_by f ∘ g) ∧ - (∀h g f. - h with - <|hashmap_hash_map_max_load updated_by f; - hashmap_hash_map_max_load updated_by g|> = - h with hashmap_hash_map_max_load updated_by f ∘ g) ∧ - ∀h g f. - h with - <|hashmap_hash_map_slots updated_by f; - hashmap_hash_map_slots updated_by g|> = - h with hashmap_hash_map_slots updated_by f ∘ g - - [hashmap_hash_map_t_fupdfupds_comp] Theorem - - ⊢ ((∀g f. - hashmap_hash_map_num_entries_fupd f ∘ - hashmap_hash_map_num_entries_fupd g = - hashmap_hash_map_num_entries_fupd (f ∘ g)) ∧ - ∀h g f. - hashmap_hash_map_num_entries_fupd f ∘ - hashmap_hash_map_num_entries_fupd g ∘ h = - hashmap_hash_map_num_entries_fupd (f ∘ g) ∘ h) ∧ - ((∀g f. - hashmap_hash_map_max_load_factor_fupd f ∘ - hashmap_hash_map_max_load_factor_fupd g = - hashmap_hash_map_max_load_factor_fupd (f ∘ g)) ∧ - ∀h g f. - hashmap_hash_map_max_load_factor_fupd f ∘ - hashmap_hash_map_max_load_factor_fupd g ∘ h = - hashmap_hash_map_max_load_factor_fupd (f ∘ g) ∘ h) ∧ - ((∀g f. - hashmap_hash_map_max_load_fupd f ∘ - hashmap_hash_map_max_load_fupd g = - hashmap_hash_map_max_load_fupd (f ∘ g)) ∧ - ∀h g f. - hashmap_hash_map_max_load_fupd f ∘ - hashmap_hash_map_max_load_fupd g ∘ h = - hashmap_hash_map_max_load_fupd (f ∘ g) ∘ h) ∧ - (∀g f. - hashmap_hash_map_slots_fupd f ∘ hashmap_hash_map_slots_fupd g = - hashmap_hash_map_slots_fupd (f ∘ g)) ∧ - ∀h g f. - hashmap_hash_map_slots_fupd f ∘ hashmap_hash_map_slots_fupd g ∘ h = - hashmap_hash_map_slots_fupd (f ∘ g) ∘ h - - [hashmap_hash_map_t_induction] Theorem - - ⊢ ∀P. (∀u p u0 v. P (hashmap_hash_map_t u p u0 v)) ⇒ ∀h. P h - - [hashmap_hash_map_t_literal_11] Theorem - - ⊢ ∀u01 p1 u1 v1 u02 p2 u2 v2. - <|hashmap_hash_map_num_entries := u01; - hashmap_hash_map_max_load_factor := p1; - hashmap_hash_map_max_load := u1; hashmap_hash_map_slots := v1|> = - <|hashmap_hash_map_num_entries := u02; - hashmap_hash_map_max_load_factor := p2; - hashmap_hash_map_max_load := u2; hashmap_hash_map_slots := v2|> ⇔ - u01 = u02 ∧ p1 = p2 ∧ u1 = u2 ∧ v1 = v2 - - [hashmap_hash_map_t_literal_nchotomy] Theorem - - ⊢ ∀h. ∃u0 p u v. - h = - <|hashmap_hash_map_num_entries := u0; - hashmap_hash_map_max_load_factor := p; - hashmap_hash_map_max_load := u; hashmap_hash_map_slots := v|> - - [hashmap_hash_map_t_nchotomy] Theorem - - ⊢ ∀hh. ∃u p u0 v. hh = hashmap_hash_map_t u p u0 v - - [hashmap_hash_map_t_updates_eq_literal] Theorem - - ⊢ ∀h u0 p u v. - h with - <|hashmap_hash_map_num_entries := u0; - hashmap_hash_map_max_load_factor := p; - hashmap_hash_map_max_load := u; hashmap_hash_map_slots := v|> = - <|hashmap_hash_map_num_entries := u0; - hashmap_hash_map_max_load_factor := p; - hashmap_hash_map_max_load := u; hashmap_hash_map_slots := v|> - - [hashmap_list_t_11] Theorem - - ⊢ ∀a0 a1 a2 a0' a1' a2'. - HashmapListCons a0 a1 a2 = HashmapListCons a0' a1' a2' ⇔ - a0 = a0' ∧ a1 = a1' ∧ a2 = a2' - - [hashmap_list_t_Axiom] Theorem - - ⊢ ∀f0 f1. ∃fn. - (∀a0 a1 a2. fn (HashmapListCons a0 a1 a2) = f0 a0 a1 a2 (fn a2)) ∧ - fn HashmapListNil = f1 - - [hashmap_list_t_case_cong] Theorem - - ⊢ ∀M M' f v. - M = M' ∧ - (∀a0 a1 a2. - M' = HashmapListCons a0 a1 a2 ⇒ f a0 a1 a2 = f' a0 a1 a2) ∧ - (M' = HashmapListNil ⇒ v = v') ⇒ - hashmap_list_t_CASE M f v = hashmap_list_t_CASE M' f' v' - - [hashmap_list_t_case_eq] Theorem - - ⊢ hashmap_list_t_CASE x f v = v' ⇔ - (∃u t h. x = HashmapListCons u t h ∧ f u t h = v') ∨ - x = HashmapListNil ∧ v = v' - - [hashmap_list_t_distinct] Theorem - - ⊢ ∀a2 a1 a0. HashmapListCons a0 a1 a2 ≠ HashmapListNil - - [hashmap_list_t_induction] Theorem - - ⊢ ∀P. (∀h. P h ⇒ ∀t u. P (HashmapListCons u t h)) ∧ P HashmapListNil ⇒ - ∀h. P h - - [hashmap_list_t_nchotomy] Theorem - - ⊢ ∀hh. (∃u t h. hh = HashmapListCons u t h) ∨ hh = HashmapListNil - - -*) -end diff --git a/tests/hol4/loops/Holmakefile b/tests/hol4/loops/Holmakefile new file mode 100644 index 00000000..3c4b8973 --- /dev/null +++ b/tests/hol4/loops/Holmakefile @@ -0,0 +1,5 @@ +# This file was automatically generated - modify ../Holmakefile.template instead +INCLUDES = ../../../backends/hol4 + +all: $(DEFAULT_TARGETS) +.PHONY: all diff --git a/tests/hol4/loops/loops_FunsScript.sml b/tests/hol4/loops/loops_FunsScript.sml new file mode 100644 index 00000000..65cf77d4 --- /dev/null +++ b/tests/hol4/loops/loops_FunsScript.sml @@ -0,0 +1,755 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [loops]: function definitions *) +open primitivesLib divDefLib +open loops_TypesTheory + +val _ = new_theory "loops_Funs" + + +val [sum_loop_fwd_def] = DefineDiv ‘ + (** [loops::sum]: loop 0: forward function *) + sum_loop_fwd (max : u32) (i : u32) (s : u32) : u32 result = + if u32_lt i max + then ( + do + s0 <- u32_add s i; + i0 <- u32_add i (int_to_u32 1); + sum_loop_fwd max i0 s0 + od) + else u32_mul s (int_to_u32 2) +’ + +val sum_fwd_def = Define ‘ + (** [loops::sum]: forward function *) + sum_fwd (max : u32) : u32 result = + sum_loop_fwd max (int_to_u32 0) (int_to_u32 0) +’ + +val [sum_with_mut_borrows_loop_fwd_def] = DefineDiv ‘ + (** [loops::sum_with_mut_borrows]: loop 0: forward function *) + sum_with_mut_borrows_loop_fwd + (max : u32) (mi : u32) (ms : u32) : u32 result = + if u32_lt mi max + then ( + do + ms0 <- u32_add ms mi; + mi0 <- u32_add mi (int_to_u32 1); + sum_with_mut_borrows_loop_fwd max mi0 ms0 + od) + else u32_mul ms (int_to_u32 2) +’ + +val sum_with_mut_borrows_fwd_def = Define ‘ + (** [loops::sum_with_mut_borrows]: forward function *) + sum_with_mut_borrows_fwd (max : u32) : u32 result = + sum_with_mut_borrows_loop_fwd max (int_to_u32 0) (int_to_u32 0) +’ + +val [sum_with_shared_borrows_loop_fwd_def] = DefineDiv ‘ + (** [loops::sum_with_shared_borrows]: loop 0: forward function *) + sum_with_shared_borrows_loop_fwd + (max : u32) (i : u32) (s : u32) : u32 result = + if u32_lt i max + then ( + do + i0 <- u32_add i (int_to_u32 1); + s0 <- u32_add s i0; + sum_with_shared_borrows_loop_fwd max i0 s0 + od) + else u32_mul s (int_to_u32 2) +’ + +val sum_with_shared_borrows_fwd_def = Define ‘ + (** [loops::sum_with_shared_borrows]: forward function *) + sum_with_shared_borrows_fwd (max : u32) : u32 result = + sum_with_shared_borrows_loop_fwd max (int_to_u32 0) (int_to_u32 0) +’ + +val [clear_loop_fwd_back_def] = DefineDiv ‘ + (** [loops::clear]: loop 0: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + clear_loop_fwd_back (v : u32 vec) (i : usize) : u32 vec result = + let i0 = vec_len v in + if usize_lt i i0 + then ( + do + i1 <- usize_add i (int_to_usize 1); + v0 <- vec_index_mut_back v i (int_to_u32 0); + clear_loop_fwd_back v0 i1 + od) + else Return v +’ + +val clear_fwd_back_def = Define ‘ + (** [loops::clear]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + clear_fwd_back (v : u32 vec) : u32 vec result = + clear_loop_fwd_back v (int_to_usize 0) +’ + +val [list_mem_loop_fwd_def] = DefineDiv ‘ + (** [loops::list_mem]: loop 0: forward function *) + list_mem_loop_fwd (x : u32) (ls : u32 list_t) : bool result = + (case ls of + | ListCons y tl => if y = x then Return T else list_mem_loop_fwd x tl + | ListNil => Return F) +’ + +val list_mem_fwd_def = Define ‘ + (** [loops::list_mem]: forward function *) + list_mem_fwd (x : u32) (ls : u32 list_t) : bool result = + list_mem_loop_fwd x ls +’ + +val [list_nth_mut_loop_loop_fwd_def] = DefineDiv ‘ + (** [loops::list_nth_mut_loop]: loop 0: forward function *) + list_nth_mut_loop_loop_fwd (ls : 't list_t) (i : u32) : 't result = + (case ls of + | ListCons x tl => + if i = int_to_u32 0 + then Return x + else ( + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_loop_loop_fwd tl i0 + od) + | ListNil => Fail Failure) +’ + +val list_nth_mut_loop_fwd_def = Define ‘ + (** [loops::list_nth_mut_loop]: forward function *) + list_nth_mut_loop_fwd (ls : 't list_t) (i : u32) : 't result = + list_nth_mut_loop_loop_fwd ls i +’ + +val [list_nth_mut_loop_loop_back_def] = DefineDiv ‘ + (** [loops::list_nth_mut_loop]: loop 0: backward function 0 *) + list_nth_mut_loop_loop_back + (ls : 't list_t) (i : u32) (ret : 't) : 't list_t result = + (case ls of + | ListCons x tl => + if i = int_to_u32 0 + then Return (ListCons ret tl) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + tl0 <- list_nth_mut_loop_loop_back tl i0 ret; + Return (ListCons x tl0) + od) + | ListNil => Fail Failure) +’ + +val list_nth_mut_loop_back_def = Define ‘ + (** [loops::list_nth_mut_loop]: backward function 0 *) + list_nth_mut_loop_back + (ls : 't list_t) (i : u32) (ret : 't) : 't list_t result = + list_nth_mut_loop_loop_back ls i ret +’ + +val [list_nth_shared_loop_loop_fwd_def] = DefineDiv ‘ + (** [loops::list_nth_shared_loop]: loop 0: forward function *) + list_nth_shared_loop_loop_fwd (ls : 't list_t) (i : u32) : 't result = + (case ls of + | ListCons x tl => + if i = int_to_u32 0 + then Return x + else ( + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_loop_loop_fwd tl i0 + od) + | ListNil => Fail Failure) +’ + +val list_nth_shared_loop_fwd_def = Define ‘ + (** [loops::list_nth_shared_loop]: forward function *) + list_nth_shared_loop_fwd (ls : 't list_t) (i : u32) : 't result = + list_nth_shared_loop_loop_fwd ls i +’ + +val [get_elem_mut_loop_fwd_def] = DefineDiv ‘ + (** [loops::get_elem_mut]: loop 0: forward function *) + get_elem_mut_loop_fwd (x : usize) (ls : usize list_t) : usize result = + (case ls of + | ListCons y tl => if y = x then Return y else get_elem_mut_loop_fwd x tl + | ListNil => Fail Failure) +’ + +val get_elem_mut_fwd_def = Define ‘ + (** [loops::get_elem_mut]: forward function *) + get_elem_mut_fwd (slots : usize list_t vec) (x : usize) : usize result = + do + l <- vec_index_mut_fwd slots (int_to_usize 0); + get_elem_mut_loop_fwd x l + od +’ + +val [get_elem_mut_loop_back_def] = DefineDiv ‘ + (** [loops::get_elem_mut]: loop 0: backward function 0 *) + get_elem_mut_loop_back + (x : usize) (ls : usize list_t) (ret : usize) : usize list_t result = + (case ls of + | ListCons y tl => + if y = x + then Return (ListCons ret tl) + else ( + do + tl0 <- get_elem_mut_loop_back x tl ret; + Return (ListCons y tl0) + od) + | ListNil => Fail Failure) +’ + +val get_elem_mut_back_def = Define ‘ + (** [loops::get_elem_mut]: backward function 0 *) + get_elem_mut_back + (slots : usize list_t vec) (x : usize) (ret : usize) : + usize list_t vec result + = + do + l <- vec_index_mut_fwd slots (int_to_usize 0); + l0 <- get_elem_mut_loop_back x l ret; + vec_index_mut_back slots (int_to_usize 0) l0 + od +’ + +val [get_elem_shared_loop_fwd_def] = DefineDiv ‘ + (** [loops::get_elem_shared]: loop 0: forward function *) + get_elem_shared_loop_fwd (x : usize) (ls : usize list_t) : usize result = + (case ls of + | ListCons y tl => + if y = x then Return y else get_elem_shared_loop_fwd x tl + | ListNil => Fail Failure) +’ + +val get_elem_shared_fwd_def = Define ‘ + (** [loops::get_elem_shared]: forward function *) + get_elem_shared_fwd (slots : usize list_t vec) (x : usize) : usize result = + do + l <- vec_index_fwd slots (int_to_usize 0); + get_elem_shared_loop_fwd x l + od +’ + +val id_mut_fwd_def = Define ‘ + (** [loops::id_mut]: forward function *) + id_mut_fwd (ls : 't list_t) : 't list_t result = + Return ls +’ + +val id_mut_back_def = Define ‘ + (** [loops::id_mut]: backward function 0 *) + id_mut_back (ls : 't list_t) (ret : 't list_t) : 't list_t result = + Return ret +’ + +val id_shared_fwd_def = Define ‘ + (** [loops::id_shared]: forward function *) + id_shared_fwd (ls : 't list_t) : 't list_t result = + Return ls +’ + +val [list_nth_mut_loop_with_id_loop_fwd_def] = DefineDiv ‘ + (** [loops::list_nth_mut_loop_with_id]: loop 0: forward function *) + list_nth_mut_loop_with_id_loop_fwd (i : u32) (ls : 't list_t) : 't result = + (case ls of + | ListCons x tl => + if i = int_to_u32 0 + then Return x + else ( + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_loop_with_id_loop_fwd i0 tl + od) + | ListNil => Fail Failure) +’ + +val list_nth_mut_loop_with_id_fwd_def = Define ‘ + (** [loops::list_nth_mut_loop_with_id]: forward function *) + list_nth_mut_loop_with_id_fwd (ls : 't list_t) (i : u32) : 't result = + do + ls0 <- id_mut_fwd ls; + list_nth_mut_loop_with_id_loop_fwd i ls0 + od +’ + +val [list_nth_mut_loop_with_id_loop_back_def] = DefineDiv ‘ + (** [loops::list_nth_mut_loop_with_id]: loop 0: backward function 0 *) + list_nth_mut_loop_with_id_loop_back + (i : u32) (ls : 't list_t) (ret : 't) : 't list_t result = + (case ls of + | ListCons x tl => + if i = int_to_u32 0 + then Return (ListCons ret tl) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + tl0 <- list_nth_mut_loop_with_id_loop_back i0 tl ret; + Return (ListCons x tl0) + od) + | ListNil => Fail Failure) +’ + +val list_nth_mut_loop_with_id_back_def = Define ‘ + (** [loops::list_nth_mut_loop_with_id]: backward function 0 *) + list_nth_mut_loop_with_id_back + (ls : 't list_t) (i : u32) (ret : 't) : 't list_t result = + do + ls0 <- id_mut_fwd ls; + l <- list_nth_mut_loop_with_id_loop_back i ls0 ret; + id_mut_back ls l + od +’ + +val [list_nth_shared_loop_with_id_loop_fwd_def] = DefineDiv ‘ + (** [loops::list_nth_shared_loop_with_id]: loop 0: forward function *) + list_nth_shared_loop_with_id_loop_fwd + (i : u32) (ls : 't list_t) : 't result = + (case ls of + | ListCons x tl => + if i = int_to_u32 0 + then Return x + else ( + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_loop_with_id_loop_fwd i0 tl + od) + | ListNil => Fail Failure) +’ + +val list_nth_shared_loop_with_id_fwd_def = Define ‘ + (** [loops::list_nth_shared_loop_with_id]: forward function *) + list_nth_shared_loop_with_id_fwd (ls : 't list_t) (i : u32) : 't result = + do + ls0 <- id_shared_fwd ls; + list_nth_shared_loop_with_id_loop_fwd i ls0 + od +’ + +val [list_nth_mut_loop_pair_loop_fwd_def] = DefineDiv ‘ + (** [loops::list_nth_mut_loop_pair]: loop 0: forward function *) + list_nth_mut_loop_pair_loop_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (x0, x1) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_loop_pair_loop_fwd tl0 tl1 i0 + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_mut_loop_pair_fwd_def = Define ‘ + (** [loops::list_nth_mut_loop_pair]: forward function *) + list_nth_mut_loop_pair_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + list_nth_mut_loop_pair_loop_fwd ls0 ls1 i +’ + +val [list_nth_mut_loop_pair_loop_back'a_def] = DefineDiv ‘ + (** [loops::list_nth_mut_loop_pair]: loop 0: backward function 0 *) + list_nth_mut_loop_pair_loop_back'a + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : + 't list_t result + = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (ListCons ret tl0) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + tl00 <- list_nth_mut_loop_pair_loop_back'a tl0 tl1 i0 ret; + Return (ListCons x0 tl00) + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_mut_loop_pair_back'a_def = Define ‘ + (** [loops::list_nth_mut_loop_pair]: backward function 0 *) + list_nth_mut_loop_pair_back'a + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : + 't list_t result + = + list_nth_mut_loop_pair_loop_back'a ls0 ls1 i ret +’ + +val [list_nth_mut_loop_pair_loop_back'b_def] = DefineDiv ‘ + (** [loops::list_nth_mut_loop_pair]: loop 0: backward function 1 *) + list_nth_mut_loop_pair_loop_back'b + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : + 't list_t result + = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (ListCons ret tl1) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + tl10 <- list_nth_mut_loop_pair_loop_back'b tl0 tl1 i0 ret; + Return (ListCons x1 tl10) + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_mut_loop_pair_back'b_def = Define ‘ + (** [loops::list_nth_mut_loop_pair]: backward function 1 *) + list_nth_mut_loop_pair_back'b + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : + 't list_t result + = + list_nth_mut_loop_pair_loop_back'b ls0 ls1 i ret +’ + +val [list_nth_shared_loop_pair_loop_fwd_def] = DefineDiv ‘ + (** [loops::list_nth_shared_loop_pair]: loop 0: forward function *) + list_nth_shared_loop_pair_loop_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (x0, x1) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_loop_pair_loop_fwd tl0 tl1 i0 + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_shared_loop_pair_fwd_def = Define ‘ + (** [loops::list_nth_shared_loop_pair]: forward function *) + list_nth_shared_loop_pair_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + list_nth_shared_loop_pair_loop_fwd ls0 ls1 i +’ + +val [list_nth_mut_loop_pair_merge_loop_fwd_def] = DefineDiv ‘ + (** [loops::list_nth_mut_loop_pair_merge]: loop 0: forward function *) + list_nth_mut_loop_pair_merge_loop_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (x0, x1) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_loop_pair_merge_loop_fwd tl0 tl1 i0 + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_mut_loop_pair_merge_fwd_def = Define ‘ + (** [loops::list_nth_mut_loop_pair_merge]: forward function *) + list_nth_mut_loop_pair_merge_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + list_nth_mut_loop_pair_merge_loop_fwd ls0 ls1 i +’ + +val [list_nth_mut_loop_pair_merge_loop_back_def] = DefineDiv ‘ + (** [loops::list_nth_mut_loop_pair_merge]: loop 0: backward function 0 *) + list_nth_mut_loop_pair_merge_loop_back + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : ('t # 't)) : + ('t list_t # 't list_t) result + = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then let (t, t0) = ret in Return (ListCons t tl0, ListCons t0 tl1) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + (tl00, tl10) <- + list_nth_mut_loop_pair_merge_loop_back tl0 tl1 i0 ret; + Return (ListCons x0 tl00, ListCons x1 tl10) + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_mut_loop_pair_merge_back_def = Define ‘ + (** [loops::list_nth_mut_loop_pair_merge]: backward function 0 *) + list_nth_mut_loop_pair_merge_back + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : ('t # 't)) : + ('t list_t # 't list_t) result + = + list_nth_mut_loop_pair_merge_loop_back ls0 ls1 i ret +’ + +val [list_nth_shared_loop_pair_merge_loop_fwd_def] = DefineDiv ‘ + (** [loops::list_nth_shared_loop_pair_merge]: loop 0: forward function *) + list_nth_shared_loop_pair_merge_loop_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (x0, x1) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_loop_pair_merge_loop_fwd tl0 tl1 i0 + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_shared_loop_pair_merge_fwd_def = Define ‘ + (** [loops::list_nth_shared_loop_pair_merge]: forward function *) + list_nth_shared_loop_pair_merge_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + list_nth_shared_loop_pair_merge_loop_fwd ls0 ls1 i +’ + +val [list_nth_mut_shared_loop_pair_loop_fwd_def] = DefineDiv ‘ + (** [loops::list_nth_mut_shared_loop_pair]: loop 0: forward function *) + list_nth_mut_shared_loop_pair_loop_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (x0, x1) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_shared_loop_pair_loop_fwd tl0 tl1 i0 + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_mut_shared_loop_pair_fwd_def = Define ‘ + (** [loops::list_nth_mut_shared_loop_pair]: forward function *) + list_nth_mut_shared_loop_pair_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + list_nth_mut_shared_loop_pair_loop_fwd ls0 ls1 i +’ + +val [list_nth_mut_shared_loop_pair_loop_back_def] = DefineDiv ‘ + (** [loops::list_nth_mut_shared_loop_pair]: loop 0: backward function 0 *) + list_nth_mut_shared_loop_pair_loop_back + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : + 't list_t result + = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (ListCons ret tl0) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + tl00 <- list_nth_mut_shared_loop_pair_loop_back tl0 tl1 i0 ret; + Return (ListCons x0 tl00) + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_mut_shared_loop_pair_back_def = Define ‘ + (** [loops::list_nth_mut_shared_loop_pair]: backward function 0 *) + list_nth_mut_shared_loop_pair_back + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : + 't list_t result + = + list_nth_mut_shared_loop_pair_loop_back ls0 ls1 i ret +’ + +val [list_nth_mut_shared_loop_pair_merge_loop_fwd_def] = DefineDiv ‘ + (** [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: forward function *) + list_nth_mut_shared_loop_pair_merge_loop_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (x0, x1) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_shared_loop_pair_merge_loop_fwd tl0 tl1 i0 + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_mut_shared_loop_pair_merge_fwd_def = Define ‘ + (** [loops::list_nth_mut_shared_loop_pair_merge]: forward function *) + list_nth_mut_shared_loop_pair_merge_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + list_nth_mut_shared_loop_pair_merge_loop_fwd ls0 ls1 i +’ + +val [list_nth_mut_shared_loop_pair_merge_loop_back_def] = DefineDiv ‘ + (** [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: backward function 0 *) + list_nth_mut_shared_loop_pair_merge_loop_back + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : + 't list_t result + = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (ListCons ret tl0) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + tl00 <- list_nth_mut_shared_loop_pair_merge_loop_back tl0 tl1 i0 ret; + Return (ListCons x0 tl00) + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_mut_shared_loop_pair_merge_back_def = Define ‘ + (** [loops::list_nth_mut_shared_loop_pair_merge]: backward function 0 *) + list_nth_mut_shared_loop_pair_merge_back + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : + 't list_t result + = + list_nth_mut_shared_loop_pair_merge_loop_back ls0 ls1 i ret +’ + +val [list_nth_shared_mut_loop_pair_loop_fwd_def] = DefineDiv ‘ + (** [loops::list_nth_shared_mut_loop_pair]: loop 0: forward function *) + list_nth_shared_mut_loop_pair_loop_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (x0, x1) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_mut_loop_pair_loop_fwd tl0 tl1 i0 + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_shared_mut_loop_pair_fwd_def = Define ‘ + (** [loops::list_nth_shared_mut_loop_pair]: forward function *) + list_nth_shared_mut_loop_pair_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + list_nth_shared_mut_loop_pair_loop_fwd ls0 ls1 i +’ + +val [list_nth_shared_mut_loop_pair_loop_back_def] = DefineDiv ‘ + (** [loops::list_nth_shared_mut_loop_pair]: loop 0: backward function 1 *) + list_nth_shared_mut_loop_pair_loop_back + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : + 't list_t result + = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (ListCons ret tl1) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + tl10 <- list_nth_shared_mut_loop_pair_loop_back tl0 tl1 i0 ret; + Return (ListCons x1 tl10) + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_shared_mut_loop_pair_back_def = Define ‘ + (** [loops::list_nth_shared_mut_loop_pair]: backward function 1 *) + list_nth_shared_mut_loop_pair_back + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : + 't list_t result + = + list_nth_shared_mut_loop_pair_loop_back ls0 ls1 i ret +’ + +val [list_nth_shared_mut_loop_pair_merge_loop_fwd_def] = DefineDiv ‘ + (** [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: forward function *) + list_nth_shared_mut_loop_pair_merge_loop_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (x0, x1) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_mut_loop_pair_merge_loop_fwd tl0 tl1 i0 + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_shared_mut_loop_pair_merge_fwd_def = Define ‘ + (** [loops::list_nth_shared_mut_loop_pair_merge]: forward function *) + list_nth_shared_mut_loop_pair_merge_fwd + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = + list_nth_shared_mut_loop_pair_merge_loop_fwd ls0 ls1 i +’ + +val [list_nth_shared_mut_loop_pair_merge_loop_back_def] = DefineDiv ‘ + (** [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: backward function 0 *) + list_nth_shared_mut_loop_pair_merge_loop_back + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : + 't list_t result + = + (case ls0 of + | ListCons x0 tl0 => + (case ls1 of + | ListCons x1 tl1 => + if i = int_to_u32 0 + then Return (ListCons ret tl1) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + tl10 <- list_nth_shared_mut_loop_pair_merge_loop_back tl0 tl1 i0 ret; + Return (ListCons x1 tl10) + od) + | ListNil => Fail Failure) + | ListNil => Fail Failure) +’ + +val list_nth_shared_mut_loop_pair_merge_back_def = Define ‘ + (** [loops::list_nth_shared_mut_loop_pair_merge]: backward function 0 *) + list_nth_shared_mut_loop_pair_merge_back + (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : + 't list_t result + = + list_nth_shared_mut_loop_pair_merge_loop_back ls0 ls1 i ret +’ + +val _ = export_theory () diff --git a/tests/hol4/loops/loops_FunsTheory.sig b/tests/hol4/loops/loops_FunsTheory.sig new file mode 100644 index 00000000..63fe56c9 --- /dev/null +++ b/tests/hol4/loops/loops_FunsTheory.sig @@ -0,0 +1,731 @@ +signature loops_FunsTheory = +sig + type thm = Thm.thm + + (* Definitions *) + val clear_fwd_back_def : thm + val clear_loop_fwd_back_def : thm + val get_elem_mut_back_def : thm + val get_elem_mut_fwd_def : thm + val get_elem_mut_loop_back_def : thm + val get_elem_mut_loop_fwd_def : thm + val get_elem_shared_fwd_def : thm + val get_elem_shared_loop_fwd_def : thm + val id_mut_back_def : thm + val id_mut_fwd_def : thm + val id_shared_fwd_def : thm + val list_mem_fwd_def : thm + val list_mem_loop_fwd_def : thm + val list_nth_mut_loop_back_def : thm + val list_nth_mut_loop_fwd_def : thm + val list_nth_mut_loop_loop_back_def : thm + val list_nth_mut_loop_loop_fwd_def : thm + val list_nth_mut_loop_pair_back'a_def : thm + val list_nth_mut_loop_pair_back'b_def : thm + val list_nth_mut_loop_pair_fwd_def : thm + val list_nth_mut_loop_pair_loop_back'a_def : thm + val list_nth_mut_loop_pair_loop_back'b_def : thm + val list_nth_mut_loop_pair_loop_fwd_def : thm + val list_nth_mut_loop_pair_merge_back_def : thm + val list_nth_mut_loop_pair_merge_fwd_def : thm + val list_nth_mut_loop_pair_merge_loop_back_def : thm + val list_nth_mut_loop_pair_merge_loop_fwd_def : thm + val list_nth_mut_loop_with_id_back_def : thm + val list_nth_mut_loop_with_id_fwd_def : thm + val list_nth_mut_loop_with_id_loop_back_def : thm + val list_nth_mut_loop_with_id_loop_fwd_def : thm + val list_nth_mut_shared_loop_pair_back_def : thm + val list_nth_mut_shared_loop_pair_fwd_def : thm + val list_nth_mut_shared_loop_pair_loop_back_def : thm + val list_nth_mut_shared_loop_pair_loop_fwd_def : thm + val list_nth_mut_shared_loop_pair_merge_back_def : thm + val list_nth_mut_shared_loop_pair_merge_fwd_def : thm + val list_nth_mut_shared_loop_pair_merge_loop_back_def : thm + val list_nth_mut_shared_loop_pair_merge_loop_fwd_def : thm + val list_nth_shared_loop_fwd_def : thm + val list_nth_shared_loop_loop_fwd_def : thm + val list_nth_shared_loop_pair_fwd_def : thm + val list_nth_shared_loop_pair_loop_fwd_def : thm + val list_nth_shared_loop_pair_merge_fwd_def : thm + val list_nth_shared_loop_pair_merge_loop_fwd_def : thm + val list_nth_shared_loop_with_id_fwd_def : thm + val list_nth_shared_loop_with_id_loop_fwd_def : thm + val list_nth_shared_mut_loop_pair_back_def : thm + val list_nth_shared_mut_loop_pair_fwd_def : thm + val list_nth_shared_mut_loop_pair_loop_back_def : thm + val list_nth_shared_mut_loop_pair_loop_fwd_def : thm + val list_nth_shared_mut_loop_pair_merge_back_def : thm + val list_nth_shared_mut_loop_pair_merge_fwd_def : thm + val list_nth_shared_mut_loop_pair_merge_loop_back_def : thm + val list_nth_shared_mut_loop_pair_merge_loop_fwd_def : thm + val sum_fwd_def : thm + val sum_loop_fwd_def : thm + val sum_with_mut_borrows_fwd_def : thm + val sum_with_mut_borrows_loop_fwd_def : thm + val sum_with_shared_borrows_fwd_def : thm + val sum_with_shared_borrows_loop_fwd_def : thm + + val loops_Funs_grammars : type_grammar.grammar * term_grammar.grammar +(* + [loops_Types] Parent theory of "loops_Funs" + + [clear_fwd_back_def] Definition + + ⊢ ∀v. clear_fwd_back v = clear_loop_fwd_back v (int_to_usize 0) + + [clear_loop_fwd_back_def] Definition + + ⊢ ∀v i. + clear_loop_fwd_back v i = + (let + i0 = vec_len v + in + if usize_lt i i0 then + do + i1 <- usize_add i (int_to_usize 1); + v0 <- vec_index_mut_back v i (int_to_u32 0); + clear_loop_fwd_back v0 i1 + od + else Return v) + + [get_elem_mut_back_def] Definition + + ⊢ ∀slots x ret. + get_elem_mut_back slots x ret = + do + l <- vec_index_mut_fwd slots (int_to_usize 0); + l0 <- get_elem_mut_loop_back x l ret; + vec_index_mut_back slots (int_to_usize 0) l0 + od + + [get_elem_mut_fwd_def] Definition + + ⊢ ∀slots x. + get_elem_mut_fwd slots x = + do + l <- vec_index_mut_fwd slots (int_to_usize 0); + get_elem_mut_loop_fwd x l + od + + [get_elem_mut_loop_back_def] Definition + + ⊢ ∀x ls ret. + get_elem_mut_loop_back x ls ret = + case ls of + ListCons y tl => + if y = x then Return (ListCons ret tl) + else + do + tl0 <- get_elem_mut_loop_back x tl ret; + Return (ListCons y tl0) + od + | ListNil => Fail Failure + + [get_elem_mut_loop_fwd_def] Definition + + ⊢ ∀x ls. + get_elem_mut_loop_fwd x ls = + case ls of + ListCons y tl => + if y = x then Return y else get_elem_mut_loop_fwd x tl + | ListNil => Fail Failure + + [get_elem_shared_fwd_def] Definition + + ⊢ ∀slots x. + get_elem_shared_fwd slots x = + do + l <- vec_index_fwd slots (int_to_usize 0); + get_elem_shared_loop_fwd x l + od + + [get_elem_shared_loop_fwd_def] Definition + + ⊢ ∀x ls. + get_elem_shared_loop_fwd x ls = + case ls of + ListCons y tl => + if y = x then Return y else get_elem_shared_loop_fwd x tl + | ListNil => Fail Failure + + [id_mut_back_def] Definition + + ⊢ ∀ls ret. id_mut_back ls ret = Return ret + + [id_mut_fwd_def] Definition + + ⊢ ∀ls. id_mut_fwd ls = Return ls + + [id_shared_fwd_def] Definition + + ⊢ ∀ls. id_shared_fwd ls = Return ls + + [list_mem_fwd_def] Definition + + ⊢ ∀x ls. list_mem_fwd x ls = list_mem_loop_fwd x ls + + [list_mem_loop_fwd_def] Definition + + ⊢ ∀x ls. + list_mem_loop_fwd x ls = + case ls of + ListCons y tl => + if y = x then Return T else list_mem_loop_fwd x tl + | ListNil => Return F + + [list_nth_mut_loop_back_def] Definition + + ⊢ ∀ls i ret. + list_nth_mut_loop_back ls i ret = + list_nth_mut_loop_loop_back ls i ret + + [list_nth_mut_loop_fwd_def] Definition + + ⊢ ∀ls i. list_nth_mut_loop_fwd ls i = list_nth_mut_loop_loop_fwd ls i + + [list_nth_mut_loop_loop_back_def] Definition + + ⊢ ∀ls i ret. + list_nth_mut_loop_loop_back ls i ret = + case ls of + ListCons x tl => + if i = int_to_u32 0 then Return (ListCons ret tl) + else + do + i0 <- u32_sub i (int_to_u32 1); + tl0 <- list_nth_mut_loop_loop_back tl i0 ret; + Return (ListCons x tl0) + od + | ListNil => Fail Failure + + [list_nth_mut_loop_loop_fwd_def] Definition + + ⊢ ∀ls i. + list_nth_mut_loop_loop_fwd ls i = + case ls of + ListCons x tl => + if i = int_to_u32 0 then Return x + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_loop_loop_fwd tl i0 + od + | ListNil => Fail Failure + + [list_nth_mut_loop_pair_back'a_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_mut_loop_pair_back'a ls0 ls1 i ret = + list_nth_mut_loop_pair_loop_back'a ls0 ls1 i ret + + [list_nth_mut_loop_pair_back'b_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_mut_loop_pair_back'b ls0 ls1 i ret = + list_nth_mut_loop_pair_loop_back'b ls0 ls1 i ret + + [list_nth_mut_loop_pair_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_mut_loop_pair_fwd ls0 ls1 i = + list_nth_mut_loop_pair_loop_fwd ls0 ls1 i + + [list_nth_mut_loop_pair_loop_back'a_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_mut_loop_pair_loop_back'a ls0 ls1 i ret = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (ListCons ret tl0) + else + do + i0 <- u32_sub i (int_to_u32 1); + tl00 <- + list_nth_mut_loop_pair_loop_back'a tl0 tl1 i0 ret; + Return (ListCons x0 tl00) + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_mut_loop_pair_loop_back'b_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_mut_loop_pair_loop_back'b ls0 ls1 i ret = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (ListCons ret tl1) + else + do + i0 <- u32_sub i (int_to_u32 1); + tl10 <- + list_nth_mut_loop_pair_loop_back'b tl0 tl1 i0 ret; + Return (ListCons x1 tl10) + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_mut_loop_pair_loop_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_mut_loop_pair_loop_fwd ls0 ls1 i = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (x0,x1) + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_loop_pair_loop_fwd tl0 tl1 i0 + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_mut_loop_pair_merge_back_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_mut_loop_pair_merge_back ls0 ls1 i ret = + list_nth_mut_loop_pair_merge_loop_back ls0 ls1 i ret + + [list_nth_mut_loop_pair_merge_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_mut_loop_pair_merge_fwd ls0 ls1 i = + list_nth_mut_loop_pair_merge_loop_fwd ls0 ls1 i + + [list_nth_mut_loop_pair_merge_loop_back_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_mut_loop_pair_merge_loop_back ls0 ls1 i ret = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then + (let + (t,t0) = ret + in + Return (ListCons t tl0,ListCons t0 tl1)) + else + do + i0 <- u32_sub i (int_to_u32 1); + (tl00,tl10) <- + list_nth_mut_loop_pair_merge_loop_back tl0 tl1 i0 + ret; + Return (ListCons x0 tl00,ListCons x1 tl10) + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_mut_loop_pair_merge_loop_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_mut_loop_pair_merge_loop_fwd ls0 ls1 i = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (x0,x1) + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_loop_pair_merge_loop_fwd tl0 tl1 i0 + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_mut_loop_with_id_back_def] Definition + + ⊢ ∀ls i ret. + list_nth_mut_loop_with_id_back ls i ret = + do + ls0 <- id_mut_fwd ls; + l <- list_nth_mut_loop_with_id_loop_back i ls0 ret; + id_mut_back ls l + od + + [list_nth_mut_loop_with_id_fwd_def] Definition + + ⊢ ∀ls i. + list_nth_mut_loop_with_id_fwd ls i = + do + ls0 <- id_mut_fwd ls; + list_nth_mut_loop_with_id_loop_fwd i ls0 + od + + [list_nth_mut_loop_with_id_loop_back_def] Definition + + ⊢ ∀i ls ret. + list_nth_mut_loop_with_id_loop_back i ls ret = + case ls of + ListCons x tl => + if i = int_to_u32 0 then Return (ListCons ret tl) + else + do + i0 <- u32_sub i (int_to_u32 1); + tl0 <- list_nth_mut_loop_with_id_loop_back i0 tl ret; + Return (ListCons x tl0) + od + | ListNil => Fail Failure + + [list_nth_mut_loop_with_id_loop_fwd_def] Definition + + ⊢ ∀i ls. + list_nth_mut_loop_with_id_loop_fwd i ls = + case ls of + ListCons x tl => + if i = int_to_u32 0 then Return x + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_loop_with_id_loop_fwd i0 tl + od + | ListNil => Fail Failure + + [list_nth_mut_shared_loop_pair_back_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_mut_shared_loop_pair_back ls0 ls1 i ret = + list_nth_mut_shared_loop_pair_loop_back ls0 ls1 i ret + + [list_nth_mut_shared_loop_pair_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_mut_shared_loop_pair_fwd ls0 ls1 i = + list_nth_mut_shared_loop_pair_loop_fwd ls0 ls1 i + + [list_nth_mut_shared_loop_pair_loop_back_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_mut_shared_loop_pair_loop_back ls0 ls1 i ret = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (ListCons ret tl0) + else + do + i0 <- u32_sub i (int_to_u32 1); + tl00 <- + list_nth_mut_shared_loop_pair_loop_back tl0 tl1 i0 + ret; + Return (ListCons x0 tl00) + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_mut_shared_loop_pair_loop_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_mut_shared_loop_pair_loop_fwd ls0 ls1 i = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (x0,x1) + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_shared_loop_pair_loop_fwd tl0 tl1 i0 + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_mut_shared_loop_pair_merge_back_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_mut_shared_loop_pair_merge_back ls0 ls1 i ret = + list_nth_mut_shared_loop_pair_merge_loop_back ls0 ls1 i ret + + [list_nth_mut_shared_loop_pair_merge_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_mut_shared_loop_pair_merge_fwd ls0 ls1 i = + list_nth_mut_shared_loop_pair_merge_loop_fwd ls0 ls1 i + + [list_nth_mut_shared_loop_pair_merge_loop_back_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_mut_shared_loop_pair_merge_loop_back ls0 ls1 i ret = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (ListCons ret tl0) + else + do + i0 <- u32_sub i (int_to_u32 1); + tl00 <- + list_nth_mut_shared_loop_pair_merge_loop_back tl0 + tl1 i0 ret; + Return (ListCons x0 tl00) + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_mut_shared_loop_pair_merge_loop_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_mut_shared_loop_pair_merge_loop_fwd ls0 ls1 i = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (x0,x1) + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_shared_loop_pair_merge_loop_fwd tl0 tl1 + i0 + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_shared_loop_fwd_def] Definition + + ⊢ ∀ls i. + list_nth_shared_loop_fwd ls i = + list_nth_shared_loop_loop_fwd ls i + + [list_nth_shared_loop_loop_fwd_def] Definition + + ⊢ ∀ls i. + list_nth_shared_loop_loop_fwd ls i = + case ls of + ListCons x tl => + if i = int_to_u32 0 then Return x + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_loop_loop_fwd tl i0 + od + | ListNil => Fail Failure + + [list_nth_shared_loop_pair_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_shared_loop_pair_fwd ls0 ls1 i = + list_nth_shared_loop_pair_loop_fwd ls0 ls1 i + + [list_nth_shared_loop_pair_loop_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_shared_loop_pair_loop_fwd ls0 ls1 i = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (x0,x1) + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_loop_pair_loop_fwd tl0 tl1 i0 + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_shared_loop_pair_merge_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_shared_loop_pair_merge_fwd ls0 ls1 i = + list_nth_shared_loop_pair_merge_loop_fwd ls0 ls1 i + + [list_nth_shared_loop_pair_merge_loop_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_shared_loop_pair_merge_loop_fwd ls0 ls1 i = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (x0,x1) + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_loop_pair_merge_loop_fwd tl0 tl1 i0 + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_shared_loop_with_id_fwd_def] Definition + + ⊢ ∀ls i. + list_nth_shared_loop_with_id_fwd ls i = + do + ls0 <- id_shared_fwd ls; + list_nth_shared_loop_with_id_loop_fwd i ls0 + od + + [list_nth_shared_loop_with_id_loop_fwd_def] Definition + + ⊢ ∀i ls. + list_nth_shared_loop_with_id_loop_fwd i ls = + case ls of + ListCons x tl => + if i = int_to_u32 0 then Return x + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_loop_with_id_loop_fwd i0 tl + od + | ListNil => Fail Failure + + [list_nth_shared_mut_loop_pair_back_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_shared_mut_loop_pair_back ls0 ls1 i ret = + list_nth_shared_mut_loop_pair_loop_back ls0 ls1 i ret + + [list_nth_shared_mut_loop_pair_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_shared_mut_loop_pair_fwd ls0 ls1 i = + list_nth_shared_mut_loop_pair_loop_fwd ls0 ls1 i + + [list_nth_shared_mut_loop_pair_loop_back_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_shared_mut_loop_pair_loop_back ls0 ls1 i ret = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (ListCons ret tl1) + else + do + i0 <- u32_sub i (int_to_u32 1); + tl10 <- + list_nth_shared_mut_loop_pair_loop_back tl0 tl1 i0 + ret; + Return (ListCons x1 tl10) + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_shared_mut_loop_pair_loop_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_shared_mut_loop_pair_loop_fwd ls0 ls1 i = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (x0,x1) + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_mut_loop_pair_loop_fwd tl0 tl1 i0 + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_shared_mut_loop_pair_merge_back_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_shared_mut_loop_pair_merge_back ls0 ls1 i ret = + list_nth_shared_mut_loop_pair_merge_loop_back ls0 ls1 i ret + + [list_nth_shared_mut_loop_pair_merge_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_shared_mut_loop_pair_merge_fwd ls0 ls1 i = + list_nth_shared_mut_loop_pair_merge_loop_fwd ls0 ls1 i + + [list_nth_shared_mut_loop_pair_merge_loop_back_def] Definition + + ⊢ ∀ls0 ls1 i ret. + list_nth_shared_mut_loop_pair_merge_loop_back ls0 ls1 i ret = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (ListCons ret tl1) + else + do + i0 <- u32_sub i (int_to_u32 1); + tl10 <- + list_nth_shared_mut_loop_pair_merge_loop_back tl0 + tl1 i0 ret; + Return (ListCons x1 tl10) + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [list_nth_shared_mut_loop_pair_merge_loop_fwd_def] Definition + + ⊢ ∀ls0 ls1 i. + list_nth_shared_mut_loop_pair_merge_loop_fwd ls0 ls1 i = + case ls0 of + ListCons x0 tl0 => + (case ls1 of + ListCons x1 tl1 => + if i = int_to_u32 0 then Return (x0,x1) + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_mut_loop_pair_merge_loop_fwd tl0 tl1 + i0 + od + | ListNil => Fail Failure) + | ListNil => Fail Failure + + [sum_fwd_def] Definition + + ⊢ ∀max. sum_fwd max = sum_loop_fwd max (int_to_u32 0) (int_to_u32 0) + + [sum_loop_fwd_def] Definition + + ⊢ ∀max i s. + sum_loop_fwd max i s = + if u32_lt i max then + do + s0 <- u32_add s i; + i0 <- u32_add i (int_to_u32 1); + sum_loop_fwd max i0 s0 + od + else u32_mul s (int_to_u32 2) + + [sum_with_mut_borrows_fwd_def] Definition + + ⊢ ∀max. + sum_with_mut_borrows_fwd max = + sum_with_mut_borrows_loop_fwd max (int_to_u32 0) (int_to_u32 0) + + [sum_with_mut_borrows_loop_fwd_def] Definition + + ⊢ ∀max mi ms. + sum_with_mut_borrows_loop_fwd max mi ms = + if u32_lt mi max then + do + ms0 <- u32_add ms mi; + mi0 <- u32_add mi (int_to_u32 1); + sum_with_mut_borrows_loop_fwd max mi0 ms0 + od + else u32_mul ms (int_to_u32 2) + + [sum_with_shared_borrows_fwd_def] Definition + + ⊢ ∀max. + sum_with_shared_borrows_fwd max = + sum_with_shared_borrows_loop_fwd max (int_to_u32 0) + (int_to_u32 0) + + [sum_with_shared_borrows_loop_fwd_def] Definition + + ⊢ ∀max i s. + sum_with_shared_borrows_loop_fwd max i s = + if u32_lt i max then + do + i0 <- u32_add i (int_to_u32 1); + s0 <- u32_add s i0; + sum_with_shared_borrows_loop_fwd max i0 s0 + od + else u32_mul s (int_to_u32 2) + + +*) +end diff --git a/tests/hol4/loops/loops_TypesScript.sml b/tests/hol4/loops/loops_TypesScript.sml new file mode 100644 index 00000000..e3e5b8d1 --- /dev/null +++ b/tests/hol4/loops/loops_TypesScript.sml @@ -0,0 +1,13 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [loops]: type definitions *) +open primitivesLib divDefLib + +val _ = new_theory "loops_Types" + + +Datatype: + (** [loops::List] *) + list_t = | ListCons 't list_t | ListNil +End + +val _ = export_theory () diff --git a/tests/hol4/loops/loops_TypesTheory.sig b/tests/hol4/loops/loops_TypesTheory.sig new file mode 100644 index 00000000..c3e638d8 --- /dev/null +++ b/tests/hol4/loops/loops_TypesTheory.sig @@ -0,0 +1,94 @@ +signature loops_TypesTheory = +sig + type thm = Thm.thm + + (* Definitions *) + val list_t_TY_DEF : thm + val list_t_case_def : thm + val list_t_size_def : thm + + (* Theorems *) + val datatype_list_t : thm + val list_t_11 : thm + val list_t_Axiom : thm + val list_t_case_cong : thm + val list_t_case_eq : thm + val list_t_distinct : thm + val list_t_induction : thm + val list_t_nchotomy : thm + + val loops_Types_grammars : type_grammar.grammar * term_grammar.grammar +(* + [divDef] Parent theory of "loops_Types" + + [list_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('list_t'). + (∀a0'. + (∃a0 a1. + a0' = + (λa0 a1. + ind_type$CONSTR 0 a0 + (ind_type$FCONS a1 (λn. ind_type$BOTTOM))) + a0 a1 ∧ $var$('list_t') a1) ∨ + a0' = + ind_type$CONSTR (SUC 0) ARB (λn. ind_type$BOTTOM) ⇒ + $var$('list_t') a0') ⇒ + $var$('list_t') a0') rep + + [list_t_case_def] Definition + + ⊢ (∀a0 a1 f v. list_t_CASE (ListCons a0 a1) f v = f a0 a1) ∧ + ∀f v. list_t_CASE ListNil f v = v + + [list_t_size_def] Definition + + ⊢ (∀f a0 a1. + list_t_size f (ListCons a0 a1) = 1 + (f a0 + list_t_size f a1)) ∧ + ∀f. list_t_size f ListNil = 0 + + [datatype_list_t] Theorem + + ⊢ DATATYPE (list_t ListCons ListNil) + + [list_t_11] Theorem + + ⊢ ∀a0 a1 a0' a1'. + ListCons a0 a1 = ListCons a0' a1' ⇔ a0 = a0' ∧ a1 = a1' + + [list_t_Axiom] Theorem + + ⊢ ∀f0 f1. ∃fn. + (∀a0 a1. fn (ListCons a0 a1) = f0 a0 a1 (fn a1)) ∧ + fn ListNil = f1 + + [list_t_case_cong] Theorem + + ⊢ ∀M M' f v. + M = M' ∧ (∀a0 a1. M' = ListCons a0 a1 ⇒ f a0 a1 = f' a0 a1) ∧ + (M' = ListNil ⇒ v = v') ⇒ + list_t_CASE M f v = list_t_CASE M' f' v' + + [list_t_case_eq] Theorem + + ⊢ list_t_CASE x f v = v' ⇔ + (∃t l. x = ListCons t l ∧ f t l = v') ∨ x = ListNil ∧ v = v' + + [list_t_distinct] Theorem + + ⊢ ∀a1 a0. ListCons a0 a1 ≠ ListNil + + [list_t_induction] Theorem + + ⊢ ∀P. (∀l. P l ⇒ ∀t. P (ListCons t l)) ∧ P ListNil ⇒ ∀l. P l + + [list_t_nchotomy] Theorem + + ⊢ ∀ll. (∃t l. ll = ListCons t l) ∨ ll = ListNil + + +*) +end diff --git a/tests/hol4/misc-constants/Holmakefile b/tests/hol4/misc-constants/Holmakefile deleted file mode 100644 index 3c4b8973..00000000 --- a/tests/hol4/misc-constants/Holmakefile +++ /dev/null @@ -1,5 +0,0 @@ -# This file was automatically generated - modify ../Holmakefile.template instead -INCLUDES = ../../../backends/hol4 - -all: $(DEFAULT_TARGETS) -.PHONY: all diff --git a/tests/hol4/misc-constants/constantsScript.sml b/tests/hol4/misc-constants/constantsScript.sml deleted file mode 100644 index 40a319c6..00000000 --- a/tests/hol4/misc-constants/constantsScript.sml +++ /dev/null @@ -1,217 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [constants] *) -open primitivesLib divDefLib - -val _ = new_theory "constants" - - -(** [constants::X0] *) -Definition x0_body_def: - x0_body : u32 result = Return (int_to_u32 0) -End -Definition x0_c_def: - x0_c : u32 = get_return_value x0_body -End - -(** [constants::X1] *) -Definition x1_body_def: - x1_body : u32 result = Return core_u32_max -End -Definition x1_c_def: - x1_c : u32 = get_return_value x1_body -End - -(** [constants::X2] *) -Definition x2_body_def: - x2_body : u32 result = Return (int_to_u32 3) -End -Definition x2_c_def: - x2_c : u32 = get_return_value x2_body -End - -val incr_fwd_def = Define ‘ - (** [constants::incr]: forward function *) - incr_fwd (n : u32) : u32 result = - u32_add n (int_to_u32 1) -’ - -(** [constants::X3] *) -Definition x3_body_def: - x3_body : u32 result = incr_fwd (int_to_u32 32) -End -Definition x3_c_def: - x3_c : u32 = get_return_value x3_body -End - -val mk_pair0_fwd_def = Define ‘ - (** [constants::mk_pair0]: forward function *) - mk_pair0_fwd (x : u32) (y : u32) : (u32 # u32) result = - Return (x, y) -’ - -Datatype: - (** [constants::Pair] *) - pair_t = <| pair_x : 't1; pair_y : 't2; |> -End - -val mk_pair1_fwd_def = Define ‘ - (** [constants::mk_pair1]: forward function *) - mk_pair1_fwd (x : u32) (y : u32) : (u32, u32) pair_t result = - Return (<| pair_x := x; pair_y := y |>) -’ - -(** [constants::P0] *) -Definition p0_body_def: - p0_body : (u32 # u32) result = mk_pair0_fwd (int_to_u32 0) (int_to_u32 1) -End -Definition p0_c_def: - p0_c : (u32 # u32) = get_return_value p0_body -End - -(** [constants::P1] *) -Definition p1_body_def: - p1_body : (u32, u32) pair_t result = - mk_pair1_fwd (int_to_u32 0) (int_to_u32 1) -End -Definition p1_c_def: - p1_c : (u32, u32) pair_t = get_return_value p1_body -End - -(** [constants::P2] *) -Definition p2_body_def: - p2_body : (u32 # u32) result = Return (int_to_u32 0, int_to_u32 1) -End -Definition p2_c_def: - p2_c : (u32 # u32) = get_return_value p2_body -End - -(** [constants::P3] *) -Definition p3_body_def: - p3_body : (u32, u32) pair_t result = - Return (<| pair_x := (int_to_u32 0); pair_y := (int_to_u32 1) |>) -End -Definition p3_c_def: - p3_c : (u32, u32) pair_t = get_return_value p3_body -End - -Datatype: - (** [constants::Wrap] *) - wrap_t = <| wrap_val : 't; |> -End - -val wrap_new_fwd_def = Define ‘ - (** [constants::Wrap::{0}::new]: forward function *) - wrap_new_fwd (val : 't) : 't wrap_t result = - Return (<| wrap_val := val |>) -’ - -(** [constants::Y] *) -Definition y_body_def: - y_body : i32 wrap_t result = wrap_new_fwd (int_to_i32 2) -End -Definition y_c_def: - y_c : i32 wrap_t = get_return_value y_body -End - -val unwrap_y_fwd_def = Define ‘ - (** [constants::unwrap_y]: forward function *) - unwrap_y_fwd : i32 result = - Return y_c.wrap_val -’ - -(** [constants::YVAL] *) -Definition yval_body_def: - yval_body : i32 result = unwrap_y_fwd -End -Definition yval_c_def: - yval_c : i32 = get_return_value yval_body -End - -(** [constants::get_z1::Z1] *) -Definition get_z1_z1_body_def: - get_z1_z1_body : i32 result = Return (int_to_i32 3) -End -Definition get_z1_z1_c_def: - get_z1_z1_c : i32 = get_return_value get_z1_z1_body -End - -val get_z1_fwd_def = Define ‘ - (** [constants::get_z1]: forward function *) - get_z1_fwd : i32 result = - Return get_z1_z1_c -’ - -val add_fwd_def = Define ‘ - (** [constants::add]: forward function *) - add_fwd (a : i32) (b : i32) : i32 result = - i32_add a b -’ - -(** [constants::Q1] *) -Definition q1_body_def: - q1_body : i32 result = Return (int_to_i32 5) -End -Definition q1_c_def: - q1_c : i32 = get_return_value q1_body -End - -(** [constants::Q2] *) -Definition q2_body_def: - q2_body : i32 result = Return q1_c -End -Definition q2_c_def: - q2_c : i32 = get_return_value q2_body -End - -(** [constants::Q3] *) -Definition q3_body_def: - q3_body : i32 result = add_fwd q2_c (int_to_i32 3) -End -Definition q3_c_def: - q3_c : i32 = get_return_value q3_body -End - -val get_z2_fwd_def = Define ‘ - (** [constants::get_z2]: forward function *) - get_z2_fwd : i32 result = - do - i <- get_z1_fwd; - i0 <- add_fwd i q3_c; - add_fwd q1_c i0 - od -’ - -(** [constants::S1] *) -Definition s1_body_def: - s1_body : u32 result = Return (int_to_u32 6) -End -Definition s1_c_def: - s1_c : u32 = get_return_value s1_body -End - -(** [constants::S2] *) -Definition s2_body_def: - s2_body : u32 result = incr_fwd s1_c -End -Definition s2_c_def: - s2_c : u32 = get_return_value s2_body -End - -(** [constants::S3] *) -Definition s3_body_def: - s3_body : (u32, u32) pair_t result = Return p3_c -End -Definition s3_c_def: - s3_c : (u32, u32) pair_t = get_return_value s3_body -End - -(** [constants::S4] *) -Definition s4_body_def: - s4_body : (u32, u32) pair_t result = - mk_pair1_fwd (int_to_u32 7) (int_to_u32 8) -End -Definition s4_c_def: - s4_c : (u32, u32) pair_t = get_return_value s4_body -End - -val _ = export_theory () diff --git a/tests/hol4/misc-constants/constantsTheory.sig b/tests/hol4/misc-constants/constantsTheory.sig deleted file mode 100644 index 287ad5f5..00000000 --- a/tests/hol4/misc-constants/constantsTheory.sig +++ /dev/null @@ -1,538 +0,0 @@ -signature constantsTheory = -sig - type thm = Thm.thm - - (* Definitions *) - val add_fwd_def : thm - val get_z1_fwd_def : thm - val get_z1_z1_body_def : thm - val get_z1_z1_c_def : thm - val get_z2_fwd_def : thm - val incr_fwd_def : thm - val mk_pair0_fwd_def : thm - val mk_pair1_fwd_def : thm - val p0_body_def : thm - val p0_c_def : thm - val p1_body_def : thm - val p1_c_def : thm - val p2_body_def : thm - val p2_c_def : thm - val p3_body_def : thm - val p3_c_def : thm - val pair_t_TY_DEF : thm - val pair_t_case_def : thm - val pair_t_pair_x : thm - val pair_t_pair_x_fupd : thm - val pair_t_pair_y : thm - val pair_t_pair_y_fupd : thm - val pair_t_size_def : thm - val q1_body_def : thm - val q1_c_def : thm - val q2_body_def : thm - val q2_c_def : thm - val q3_body_def : thm - val q3_c_def : thm - val s1_body_def : thm - val s1_c_def : thm - val s2_body_def : thm - val s2_c_def : thm - val s3_body_def : thm - val s3_c_def : thm - val s4_body_def : thm - val s4_c_def : thm - val unwrap_y_fwd_def : thm - val wrap_new_fwd_def : thm - val wrap_t_TY_DEF : thm - val wrap_t_case_def : thm - val wrap_t_size_def : thm - val wrap_t_wrap_val : thm - val wrap_t_wrap_val_fupd : thm - val x0_body_def : thm - val x0_c_def : thm - val x1_body_def : thm - val x1_c_def : thm - val x2_body_def : thm - val x2_c_def : thm - val x3_body_def : thm - val x3_c_def : thm - val y_body_def : thm - val y_c_def : thm - val yval_body_def : thm - val yval_c_def : thm - - (* Theorems *) - val EXISTS_pair_t : thm - val EXISTS_wrap_t : thm - val FORALL_pair_t : thm - val FORALL_wrap_t : thm - val datatype_pair_t : thm - val datatype_wrap_t : thm - val pair_t_11 : thm - val pair_t_Axiom : thm - val pair_t_accessors : thm - val pair_t_accfupds : thm - val pair_t_case_cong : thm - val pair_t_case_eq : thm - val pair_t_component_equality : thm - val pair_t_fn_updates : thm - val pair_t_fupdcanon : thm - val pair_t_fupdcanon_comp : thm - val pair_t_fupdfupds : thm - val pair_t_fupdfupds_comp : thm - val pair_t_induction : thm - val pair_t_literal_11 : thm - val pair_t_literal_nchotomy : thm - val pair_t_nchotomy : thm - val pair_t_updates_eq_literal : thm - val wrap_t_11 : thm - val wrap_t_Axiom : thm - val wrap_t_accessors : thm - val wrap_t_accfupds : thm - val wrap_t_case_cong : thm - val wrap_t_case_eq : thm - val wrap_t_component_equality : thm - val wrap_t_fn_updates : thm - val wrap_t_fupdfupds : thm - val wrap_t_fupdfupds_comp : thm - val wrap_t_induction : thm - val wrap_t_literal_11 : thm - val wrap_t_literal_nchotomy : thm - val wrap_t_nchotomy : thm - val wrap_t_updates_eq_literal : thm - - val constants_grammars : type_grammar.grammar * term_grammar.grammar -(* - [divDef] Parent theory of "constants" - - [add_fwd_def] Definition - - ⊢ ∀a b. add_fwd a b = i32_add a b - - [get_z1_fwd_def] Definition - - ⊢ get_z1_fwd = Return get_z1_z1_c - - [get_z1_z1_body_def] Definition - - ⊢ get_z1_z1_body = Return (int_to_i32 3) - - [get_z1_z1_c_def] Definition - - ⊢ get_z1_z1_c = get_return_value get_z1_z1_body - - [get_z2_fwd_def] Definition - - ⊢ get_z2_fwd = - do i <- get_z1_fwd; i0 <- add_fwd i q3_c; add_fwd q1_c i0 od - - [incr_fwd_def] Definition - - ⊢ ∀n. incr_fwd n = u32_add n (int_to_u32 1) - - [mk_pair0_fwd_def] Definition - - ⊢ ∀x y. mk_pair0_fwd x y = Return (x,y) - - [mk_pair1_fwd_def] Definition - - ⊢ ∀x y. mk_pair1_fwd x y = Return <|pair_x := x; pair_y := y|> - - [p0_body_def] Definition - - ⊢ p0_body = mk_pair0_fwd (int_to_u32 0) (int_to_u32 1) - - [p0_c_def] Definition - - ⊢ p0_c = get_return_value p0_body - - [p1_body_def] Definition - - ⊢ p1_body = mk_pair1_fwd (int_to_u32 0) (int_to_u32 1) - - [p1_c_def] Definition - - ⊢ p1_c = get_return_value p1_body - - [p2_body_def] Definition - - ⊢ p2_body = Return (int_to_u32 0,int_to_u32 1) - - [p2_c_def] Definition - - ⊢ p2_c = get_return_value p2_body - - [p3_body_def] Definition - - ⊢ p3_body = Return <|pair_x := int_to_u32 0; pair_y := int_to_u32 1|> - - [p3_c_def] Definition - - ⊢ p3_c = get_return_value p3_body - - [pair_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('pair_t'). - (∀a0'. - (∃a0 a1. - a0' = - (λa0 a1. - ind_type$CONSTR 0 (a0,a1) - (λn. ind_type$BOTTOM)) a0 a1) ⇒ - $var$('pair_t') a0') ⇒ - $var$('pair_t') a0') rep - - [pair_t_case_def] Definition - - ⊢ ∀a0 a1 f. pair_t_CASE (pair_t a0 a1) f = f a0 a1 - - [pair_t_pair_x] Definition - - ⊢ ∀t t0. (pair_t t t0).pair_x = t - - [pair_t_pair_x_fupd] Definition - - ⊢ ∀f t t0. pair_t t t0 with pair_x updated_by f = pair_t (f t) t0 - - [pair_t_pair_y] Definition - - ⊢ ∀t t0. (pair_t t t0).pair_y = t0 - - [pair_t_pair_y_fupd] Definition - - ⊢ ∀f t t0. pair_t t t0 with pair_y updated_by f = pair_t t (f t0) - - [pair_t_size_def] Definition - - ⊢ ∀f f1 a0 a1. pair_t_size f f1 (pair_t a0 a1) = 1 + (f a0 + f1 a1) - - [q1_body_def] Definition - - ⊢ q1_body = Return (int_to_i32 5) - - [q1_c_def] Definition - - ⊢ q1_c = get_return_value q1_body - - [q2_body_def] Definition - - ⊢ q2_body = Return q1_c - - [q2_c_def] Definition - - ⊢ q2_c = get_return_value q2_body - - [q3_body_def] Definition - - ⊢ q3_body = add_fwd q2_c (int_to_i32 3) - - [q3_c_def] Definition - - ⊢ q3_c = get_return_value q3_body - - [s1_body_def] Definition - - ⊢ s1_body = Return (int_to_u32 6) - - [s1_c_def] Definition - - ⊢ s1_c = get_return_value s1_body - - [s2_body_def] Definition - - ⊢ s2_body = incr_fwd s1_c - - [s2_c_def] Definition - - ⊢ s2_c = get_return_value s2_body - - [s3_body_def] Definition - - ⊢ s3_body = Return p3_c - - [s3_c_def] Definition - - ⊢ s3_c = get_return_value s3_body - - [s4_body_def] Definition - - ⊢ s4_body = mk_pair1_fwd (int_to_u32 7) (int_to_u32 8) - - [s4_c_def] Definition - - ⊢ s4_c = get_return_value s4_body - - [unwrap_y_fwd_def] Definition - - ⊢ unwrap_y_fwd = Return y_c.wrap_val - - [wrap_new_fwd_def] Definition - - ⊢ ∀val. wrap_new_fwd val = Return <|wrap_val := val|> - - [wrap_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0. - ∀ $var$('wrap_t'). - (∀a0. - (∃a. a0 = - (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) - a) ⇒ - $var$('wrap_t') a0) ⇒ - $var$('wrap_t') a0) rep - - [wrap_t_case_def] Definition - - ⊢ ∀a f. wrap_t_CASE (wrap_t a) f = f a - - [wrap_t_size_def] Definition - - ⊢ ∀f a. wrap_t_size f (wrap_t a) = 1 + f a - - [wrap_t_wrap_val] Definition - - ⊢ ∀t. (wrap_t t).wrap_val = t - - [wrap_t_wrap_val_fupd] Definition - - ⊢ ∀f t. wrap_t t with wrap_val updated_by f = wrap_t (f t) - - [x0_body_def] Definition - - ⊢ x0_body = Return (int_to_u32 0) - - [x0_c_def] Definition - - ⊢ x0_c = get_return_value x0_body - - [x1_body_def] Definition - - ⊢ x1_body = Return core_u32_max - - [x1_c_def] Definition - - ⊢ x1_c = get_return_value x1_body - - [x2_body_def] Definition - - ⊢ x2_body = Return (int_to_u32 3) - - [x2_c_def] Definition - - ⊢ x2_c = get_return_value x2_body - - [x3_body_def] Definition - - ⊢ x3_body = incr_fwd (int_to_u32 32) - - [x3_c_def] Definition - - ⊢ x3_c = get_return_value x3_body - - [y_body_def] Definition - - ⊢ y_body = wrap_new_fwd (int_to_i32 2) - - [y_c_def] Definition - - ⊢ y_c = get_return_value y_body - - [yval_body_def] Definition - - ⊢ yval_body = unwrap_y_fwd - - [yval_c_def] Definition - - ⊢ yval_c = get_return_value yval_body - - [EXISTS_pair_t] Theorem - - ⊢ ∀P. (∃p. P p) ⇔ ∃t0 t. P <|pair_x := t0; pair_y := t|> - - [EXISTS_wrap_t] Theorem - - ⊢ ∀P. (∃w. P w) ⇔ ∃u. P <|wrap_val := u|> - - [FORALL_pair_t] Theorem - - ⊢ ∀P. (∀p. P p) ⇔ ∀t0 t. P <|pair_x := t0; pair_y := t|> - - [FORALL_wrap_t] Theorem - - ⊢ ∀P. (∀w. P w) ⇔ ∀u. P <|wrap_val := u|> - - [datatype_pair_t] Theorem - - ⊢ DATATYPE (record pair_t pair_x pair_y) - - [datatype_wrap_t] Theorem - - ⊢ DATATYPE (record wrap_t wrap_val) - - [pair_t_11] Theorem - - ⊢ ∀a0 a1 a0' a1'. pair_t a0 a1 = pair_t a0' a1' ⇔ a0 = a0' ∧ a1 = a1' - - [pair_t_Axiom] Theorem - - ⊢ ∀f. ∃fn. ∀a0 a1. fn (pair_t a0 a1) = f a0 a1 - - [pair_t_accessors] Theorem - - ⊢ (∀t t0. (pair_t t t0).pair_x = t) ∧ - ∀t t0. (pair_t t t0).pair_y = t0 - - [pair_t_accfupds] Theorem - - ⊢ (∀p f. (p with pair_y updated_by f).pair_x = p.pair_x) ∧ - (∀p f. (p with pair_x updated_by f).pair_y = p.pair_y) ∧ - (∀p f. (p with pair_x updated_by f).pair_x = f p.pair_x) ∧ - ∀p f. (p with pair_y updated_by f).pair_y = f p.pair_y - - [pair_t_case_cong] Theorem - - ⊢ ∀M M' f. - M = M' ∧ (∀a0 a1. M' = pair_t a0 a1 ⇒ f a0 a1 = f' a0 a1) ⇒ - pair_t_CASE M f = pair_t_CASE M' f' - - [pair_t_case_eq] Theorem - - ⊢ pair_t_CASE x f = v ⇔ ∃t t0. x = pair_t t t0 ∧ f t t0 = v - - [pair_t_component_equality] Theorem - - ⊢ ∀p1 p2. p1 = p2 ⇔ p1.pair_x = p2.pair_x ∧ p1.pair_y = p2.pair_y - - [pair_t_fn_updates] Theorem - - ⊢ (∀f t t0. pair_t t t0 with pair_x updated_by f = pair_t (f t) t0) ∧ - ∀f t t0. pair_t t t0 with pair_y updated_by f = pair_t t (f t0) - - [pair_t_fupdcanon] Theorem - - ⊢ ∀p g f. - p with <|pair_y updated_by f; pair_x updated_by g|> = - p with <|pair_x updated_by g; pair_y updated_by f|> - - [pair_t_fupdcanon_comp] Theorem - - ⊢ (∀g f. - pair_y_fupd f ∘ pair_x_fupd g = pair_x_fupd g ∘ pair_y_fupd f) ∧ - ∀h g f. - pair_y_fupd f ∘ pair_x_fupd g ∘ h = - pair_x_fupd g ∘ pair_y_fupd f ∘ h - - [pair_t_fupdfupds] Theorem - - ⊢ (∀p g f. - p with <|pair_x updated_by f; pair_x updated_by g|> = - p with pair_x updated_by f ∘ g) ∧ - ∀p g f. - p with <|pair_y updated_by f; pair_y updated_by g|> = - p with pair_y updated_by f ∘ g - - [pair_t_fupdfupds_comp] Theorem - - ⊢ ((∀g f. pair_x_fupd f ∘ pair_x_fupd g = pair_x_fupd (f ∘ g)) ∧ - ∀h g f. - pair_x_fupd f ∘ pair_x_fupd g ∘ h = pair_x_fupd (f ∘ g) ∘ h) ∧ - (∀g f. pair_y_fupd f ∘ pair_y_fupd g = pair_y_fupd (f ∘ g)) ∧ - ∀h g f. pair_y_fupd f ∘ pair_y_fupd g ∘ h = pair_y_fupd (f ∘ g) ∘ h - - [pair_t_induction] Theorem - - ⊢ ∀P. (∀t t0. P (pair_t t t0)) ⇒ ∀p. P p - - [pair_t_literal_11] Theorem - - ⊢ ∀t01 t1 t02 t2. - <|pair_x := t01; pair_y := t1|> = <|pair_x := t02; pair_y := t2|> ⇔ - t01 = t02 ∧ t1 = t2 - - [pair_t_literal_nchotomy] Theorem - - ⊢ ∀p. ∃t0 t. p = <|pair_x := t0; pair_y := t|> - - [pair_t_nchotomy] Theorem - - ⊢ ∀pp. ∃t t0. pp = pair_t t t0 - - [pair_t_updates_eq_literal] Theorem - - ⊢ ∀p t0 t. - p with <|pair_x := t0; pair_y := t|> = - <|pair_x := t0; pair_y := t|> - - [wrap_t_11] Theorem - - ⊢ ∀a a'. wrap_t a = wrap_t a' ⇔ a = a' - - [wrap_t_Axiom] Theorem - - ⊢ ∀f. ∃fn. ∀a. fn (wrap_t a) = f a - - [wrap_t_accessors] Theorem - - ⊢ ∀t. (wrap_t t).wrap_val = t - - [wrap_t_accfupds] Theorem - - ⊢ ∀w f. (w with wrap_val updated_by f).wrap_val = f w.wrap_val - - [wrap_t_case_cong] Theorem - - ⊢ ∀M M' f. - M = M' ∧ (∀a. M' = wrap_t a ⇒ f a = f' a) ⇒ - wrap_t_CASE M f = wrap_t_CASE M' f' - - [wrap_t_case_eq] Theorem - - ⊢ wrap_t_CASE x f = v ⇔ ∃t. x = wrap_t t ∧ f t = v - - [wrap_t_component_equality] Theorem - - ⊢ ∀w1 w2. w1 = w2 ⇔ w1.wrap_val = w2.wrap_val - - [wrap_t_fn_updates] Theorem - - ⊢ ∀f t. wrap_t t with wrap_val updated_by f = wrap_t (f t) - - [wrap_t_fupdfupds] Theorem - - ⊢ ∀w g f. - w with <|wrap_val updated_by f; wrap_val updated_by g|> = - w with wrap_val updated_by f ∘ g - - [wrap_t_fupdfupds_comp] Theorem - - ⊢ (∀g f. wrap_val_fupd f ∘ wrap_val_fupd g = wrap_val_fupd (f ∘ g)) ∧ - ∀h g f. - wrap_val_fupd f ∘ wrap_val_fupd g ∘ h = wrap_val_fupd (f ∘ g) ∘ h - - [wrap_t_induction] Theorem - - ⊢ ∀P. (∀t. P (wrap_t t)) ⇒ ∀w. P w - - [wrap_t_literal_11] Theorem - - ⊢ ∀u1 u2. <|wrap_val := u1|> = <|wrap_val := u2|> ⇔ u1 = u2 - - [wrap_t_literal_nchotomy] Theorem - - ⊢ ∀w. ∃u. w = <|wrap_val := u|> - - [wrap_t_nchotomy] Theorem - - ⊢ ∀ww. ∃t. ww = wrap_t t - - [wrap_t_updates_eq_literal] Theorem - - ⊢ ∀w u. w with wrap_val := u = <|wrap_val := u|> - - -*) -end diff --git a/tests/hol4/misc-external/Holmakefile b/tests/hol4/misc-external/Holmakefile deleted file mode 100644 index 3c4b8973..00000000 --- a/tests/hol4/misc-external/Holmakefile +++ /dev/null @@ -1,5 +0,0 @@ -# This file was automatically generated - modify ../Holmakefile.template instead -INCLUDES = ../../../backends/hol4 - -all: $(DEFAULT_TARGETS) -.PHONY: all diff --git a/tests/hol4/misc-external/external_FunsScript.sml b/tests/hol4/misc-external/external_FunsScript.sml deleted file mode 100644 index f3692ee2..00000000 --- a/tests/hol4/misc-external/external_FunsScript.sml +++ /dev/null @@ -1,108 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [external]: function definitions *) -open primitivesLib divDefLib -open external_TypesTheory external_OpaqueTheory - -val _ = new_theory "external_Funs" - - -val swap_fwd_def = Define ‘ - (** [external::swap]: forward function *) - swap_fwd (x : 't) (y : 't) (st : state) : (state # unit) result = - do - (st0, _) <- core_mem_swap_fwd x y st; - (st1, _) <- core_mem_swap_back0 x y st st0; - (st2, _) <- core_mem_swap_back1 x y st st1; - Return (st2, ()) - od -’ - -val swap_back_def = Define ‘ - (** [external::swap]: backward function 0 *) - swap_back - (x : 't) (y : 't) (st : state) (st0 : state) : (state # ('t # 't)) result = - do - (st1, _) <- core_mem_swap_fwd x y st; - (st2, x0) <- core_mem_swap_back0 x y st st1; - (_, y0) <- core_mem_swap_back1 x y st st2; - Return (st0, (x0, y0)) - od -’ - -val test_new_non_zero_u32_fwd_def = Define ‘ - (** [external::test_new_non_zero_u32]: forward function *) - test_new_non_zero_u32_fwd - (x : u32) (st : state) : (state # core_num_nonzero_non_zero_u32_t) result = - do - (st0, opt) <- core_num_nonzero_non_zero_u32_new_fwd x st; - core_option_option_unwrap_fwd opt st0 - od -’ - -val test_vec_fwd_def = Define ‘ - (** [external::test_vec]: forward function *) - test_vec_fwd : unit result = - let v = vec_new in do - _ <- vec_push_back v (int_to_u32 0); - Return () - od -’ - -(** Unit test for [external::test_vec] *) -val _ = assert_return (“test_vec_fwd”) - -val custom_swap_fwd_def = Define ‘ - (** [external::custom_swap]: forward function *) - custom_swap_fwd (x : 't) (y : 't) (st : state) : (state # 't) result = - do - (st0, _) <- core_mem_swap_fwd x y st; - (st1, x0) <- core_mem_swap_back0 x y st st0; - (st2, _) <- core_mem_swap_back1 x y st st1; - Return (st2, x0) - od -’ - -val custom_swap_back_def = Define ‘ - (** [external::custom_swap]: backward function 0 *) - custom_swap_back - (x : 't) (y : 't) (st : state) (ret : 't) (st0 : state) : - (state # ('t # 't)) result - = - do - (st1, _) <- core_mem_swap_fwd x y st; - (st2, _) <- core_mem_swap_back0 x y st st1; - (_, y0) <- core_mem_swap_back1 x y st st2; - Return (st0, (ret, y0)) - od -’ - -val test_custom_swap_fwd_def = Define ‘ - (** [external::test_custom_swap]: forward function *) - test_custom_swap_fwd - (x : u32) (y : u32) (st : state) : (state # unit) result = - do - (st0, _) <- custom_swap_fwd x y st; - Return (st0, ()) - od -’ - -val test_custom_swap_back_def = Define ‘ - (** [external::test_custom_swap]: backward function 0 *) - test_custom_swap_back - (x : u32) (y : u32) (st : state) (st0 : state) : - (state # (u32 # u32)) result - = - custom_swap_back x y st (int_to_u32 1) st0 -’ - -val test_swap_non_zero_fwd_def = Define ‘ - (** [external::test_swap_non_zero]: forward function *) - test_swap_non_zero_fwd (x : u32) (st : state) : (state # u32) result = - do - (st0, _) <- swap_fwd x (int_to_u32 0) st; - (st1, (x0, _)) <- swap_back x (int_to_u32 0) st st0; - if x0 = int_to_u32 0 then Fail Failure else Return (st1, x0) - od -’ - -val _ = export_theory () diff --git a/tests/hol4/misc-external/external_FunsTheory.sig b/tests/hol4/misc-external/external_FunsTheory.sig deleted file mode 100644 index 490f9d06..00000000 --- a/tests/hol4/misc-external/external_FunsTheory.sig +++ /dev/null @@ -1,105 +0,0 @@ -signature external_FunsTheory = -sig - type thm = Thm.thm - - (* Definitions *) - val custom_swap_back_def : thm - val custom_swap_fwd_def : thm - val swap_back_def : thm - val swap_fwd_def : thm - val test_custom_swap_back_def : thm - val test_custom_swap_fwd_def : thm - val test_new_non_zero_u32_fwd_def : thm - val test_swap_non_zero_fwd_def : thm - val test_vec_fwd_def : thm - - val external_Funs_grammars : type_grammar.grammar * term_grammar.grammar -(* - [external_Opaque] Parent theory of "external_Funs" - - [custom_swap_back_def] Definition - - ⊢ ∀x y st ret st0. - custom_swap_back x y st ret st0 = - do - (st1,_) <- core_mem_swap_fwd x y st; - (st2,_) <- core_mem_swap_back0 x y st st1; - (_,y0) <- core_mem_swap_back1 x y st st2; - Return (st0,ret,y0) - od - - [custom_swap_fwd_def] Definition - - ⊢ ∀x y st. - custom_swap_fwd x y st = - do - (st0,_) <- core_mem_swap_fwd x y st; - (st1,x0) <- core_mem_swap_back0 x y st st0; - (st2,_) <- core_mem_swap_back1 x y st st1; - Return (st2,x0) - od - - [swap_back_def] Definition - - ⊢ ∀x y st st0. - swap_back x y st st0 = - do - (st1,_) <- core_mem_swap_fwd x y st; - (st2,x0) <- core_mem_swap_back0 x y st st1; - (_,y0) <- core_mem_swap_back1 x y st st2; - Return (st0,x0,y0) - od - - [swap_fwd_def] Definition - - ⊢ ∀x y st. - swap_fwd x y st = - do - (st0,_) <- core_mem_swap_fwd x y st; - (st1,_) <- core_mem_swap_back0 x y st st0; - (st2,_) <- core_mem_swap_back1 x y st st1; - Return (st2,()) - od - - [test_custom_swap_back_def] Definition - - ⊢ ∀x y st st0. - test_custom_swap_back x y st st0 = - custom_swap_back x y st (int_to_u32 1) st0 - - [test_custom_swap_fwd_def] Definition - - ⊢ ∀x y st. - test_custom_swap_fwd x y st = - do (st0,_) <- custom_swap_fwd x y st; Return (st0,()) od - - [test_new_non_zero_u32_fwd_def] Definition - - ⊢ ∀x st. - test_new_non_zero_u32_fwd x st = - do - (st0,opt) <- core_num_nonzero_non_zero_u32_new_fwd x st; - core_option_option_unwrap_fwd opt st0 - od - - [test_swap_non_zero_fwd_def] Definition - - ⊢ ∀x st. - test_swap_non_zero_fwd x st = - do - (st0,_) <- swap_fwd x (int_to_u32 0) st; - (st1,x0,_) <- swap_back x (int_to_u32 0) st st0; - if x0 = int_to_u32 0 then Fail Failure else Return (st1,x0) - od - - [test_vec_fwd_def] Definition - - ⊢ test_vec_fwd = - (let - v = vec_new - in - monad_ignore_bind (vec_push_back v (int_to_u32 0)) (Return ())) - - -*) -end diff --git a/tests/hol4/misc-external/external_OpaqueScript.sml b/tests/hol4/misc-external/external_OpaqueScript.sml deleted file mode 100644 index b5a6d91d..00000000 --- a/tests/hol4/misc-external/external_OpaqueScript.sml +++ /dev/null @@ -1,25 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [external]: external function declarations *) -open primitivesLib divDefLib -open external_TypesTheory - -val _ = new_theory "external_Opaque" - - -(** [core::mem::swap]: forward function *)val _ = new_constant ("core_mem_swap_fwd", - “:'t -> 't -> state -> (state # unit) result”) - -(** [core::mem::swap]: backward function 0 *)val _ = new_constant ("core_mem_swap_back0", - “:'t -> 't -> state -> state -> (state # 't) result”) - -(** [core::mem::swap]: backward function 1 *)val _ = new_constant ("core_mem_swap_back1", - “:'t -> 't -> state -> state -> (state # 't) result”) - -(** [core::num::nonzero::NonZeroU32::{14}::new]: forward function *)val _ = new_constant ("core_num_nonzero_non_zero_u32_new_fwd", - “:u32 -> state -> (state # core_num_nonzero_non_zero_u32_t option) - result”) - -(** [core::option::Option::{0}::unwrap]: forward function *)val _ = new_constant ("core_option_option_unwrap_fwd", - “:'t option -> state -> (state # 't) result”) - -val _ = export_theory () diff --git a/tests/hol4/misc-external/external_OpaqueTheory.sig b/tests/hol4/misc-external/external_OpaqueTheory.sig deleted file mode 100644 index 7cd7a08c..00000000 --- a/tests/hol4/misc-external/external_OpaqueTheory.sig +++ /dev/null @@ -1,11 +0,0 @@ -signature external_OpaqueTheory = -sig - type thm = Thm.thm - - val external_Opaque_grammars : type_grammar.grammar * term_grammar.grammar -(* - [external_Types] Parent theory of "external_Opaque" - - -*) -end diff --git a/tests/hol4/misc-external/external_TypesScript.sml b/tests/hol4/misc-external/external_TypesScript.sml deleted file mode 100644 index d290c3f6..00000000 --- a/tests/hol4/misc-external/external_TypesScript.sml +++ /dev/null @@ -1,13 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [external]: type definitions *) -open primitivesLib divDefLib - -val _ = new_theory "external_Types" - - -val _ = new_type ("core_num_nonzero_non_zero_u32_t", 0) - -(** The state type used in the state-error monad *) -val _ = new_type ("state", 0) - -val _ = export_theory () diff --git a/tests/hol4/misc-external/external_TypesTheory.sig b/tests/hol4/misc-external/external_TypesTheory.sig deleted file mode 100644 index 17e2e8e3..00000000 --- a/tests/hol4/misc-external/external_TypesTheory.sig +++ /dev/null @@ -1,11 +0,0 @@ -signature external_TypesTheory = -sig - type thm = Thm.thm - - val external_Types_grammars : type_grammar.grammar * term_grammar.grammar -(* - [divDef] Parent theory of "external_Types" - - -*) -end diff --git a/tests/hol4/misc-loops/Holmakefile b/tests/hol4/misc-loops/Holmakefile deleted file mode 100644 index 3c4b8973..00000000 --- a/tests/hol4/misc-loops/Holmakefile +++ /dev/null @@ -1,5 +0,0 @@ -# This file was automatically generated - modify ../Holmakefile.template instead -INCLUDES = ../../../backends/hol4 - -all: $(DEFAULT_TARGETS) -.PHONY: all diff --git a/tests/hol4/misc-loops/loops_FunsScript.sml b/tests/hol4/misc-loops/loops_FunsScript.sml deleted file mode 100644 index 65cf77d4..00000000 --- a/tests/hol4/misc-loops/loops_FunsScript.sml +++ /dev/null @@ -1,755 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [loops]: function definitions *) -open primitivesLib divDefLib -open loops_TypesTheory - -val _ = new_theory "loops_Funs" - - -val [sum_loop_fwd_def] = DefineDiv ‘ - (** [loops::sum]: loop 0: forward function *) - sum_loop_fwd (max : u32) (i : u32) (s : u32) : u32 result = - if u32_lt i max - then ( - do - s0 <- u32_add s i; - i0 <- u32_add i (int_to_u32 1); - sum_loop_fwd max i0 s0 - od) - else u32_mul s (int_to_u32 2) -’ - -val sum_fwd_def = Define ‘ - (** [loops::sum]: forward function *) - sum_fwd (max : u32) : u32 result = - sum_loop_fwd max (int_to_u32 0) (int_to_u32 0) -’ - -val [sum_with_mut_borrows_loop_fwd_def] = DefineDiv ‘ - (** [loops::sum_with_mut_borrows]: loop 0: forward function *) - sum_with_mut_borrows_loop_fwd - (max : u32) (mi : u32) (ms : u32) : u32 result = - if u32_lt mi max - then ( - do - ms0 <- u32_add ms mi; - mi0 <- u32_add mi (int_to_u32 1); - sum_with_mut_borrows_loop_fwd max mi0 ms0 - od) - else u32_mul ms (int_to_u32 2) -’ - -val sum_with_mut_borrows_fwd_def = Define ‘ - (** [loops::sum_with_mut_borrows]: forward function *) - sum_with_mut_borrows_fwd (max : u32) : u32 result = - sum_with_mut_borrows_loop_fwd max (int_to_u32 0) (int_to_u32 0) -’ - -val [sum_with_shared_borrows_loop_fwd_def] = DefineDiv ‘ - (** [loops::sum_with_shared_borrows]: loop 0: forward function *) - sum_with_shared_borrows_loop_fwd - (max : u32) (i : u32) (s : u32) : u32 result = - if u32_lt i max - then ( - do - i0 <- u32_add i (int_to_u32 1); - s0 <- u32_add s i0; - sum_with_shared_borrows_loop_fwd max i0 s0 - od) - else u32_mul s (int_to_u32 2) -’ - -val sum_with_shared_borrows_fwd_def = Define ‘ - (** [loops::sum_with_shared_borrows]: forward function *) - sum_with_shared_borrows_fwd (max : u32) : u32 result = - sum_with_shared_borrows_loop_fwd max (int_to_u32 0) (int_to_u32 0) -’ - -val [clear_loop_fwd_back_def] = DefineDiv ‘ - (** [loops::clear]: loop 0: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - clear_loop_fwd_back (v : u32 vec) (i : usize) : u32 vec result = - let i0 = vec_len v in - if usize_lt i i0 - then ( - do - i1 <- usize_add i (int_to_usize 1); - v0 <- vec_index_mut_back v i (int_to_u32 0); - clear_loop_fwd_back v0 i1 - od) - else Return v -’ - -val clear_fwd_back_def = Define ‘ - (** [loops::clear]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - clear_fwd_back (v : u32 vec) : u32 vec result = - clear_loop_fwd_back v (int_to_usize 0) -’ - -val [list_mem_loop_fwd_def] = DefineDiv ‘ - (** [loops::list_mem]: loop 0: forward function *) - list_mem_loop_fwd (x : u32) (ls : u32 list_t) : bool result = - (case ls of - | ListCons y tl => if y = x then Return T else list_mem_loop_fwd x tl - | ListNil => Return F) -’ - -val list_mem_fwd_def = Define ‘ - (** [loops::list_mem]: forward function *) - list_mem_fwd (x : u32) (ls : u32 list_t) : bool result = - list_mem_loop_fwd x ls -’ - -val [list_nth_mut_loop_loop_fwd_def] = DefineDiv ‘ - (** [loops::list_nth_mut_loop]: loop 0: forward function *) - list_nth_mut_loop_loop_fwd (ls : 't list_t) (i : u32) : 't result = - (case ls of - | ListCons x tl => - if i = int_to_u32 0 - then Return x - else ( - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_loop_loop_fwd tl i0 - od) - | ListNil => Fail Failure) -’ - -val list_nth_mut_loop_fwd_def = Define ‘ - (** [loops::list_nth_mut_loop]: forward function *) - list_nth_mut_loop_fwd (ls : 't list_t) (i : u32) : 't result = - list_nth_mut_loop_loop_fwd ls i -’ - -val [list_nth_mut_loop_loop_back_def] = DefineDiv ‘ - (** [loops::list_nth_mut_loop]: loop 0: backward function 0 *) - list_nth_mut_loop_loop_back - (ls : 't list_t) (i : u32) (ret : 't) : 't list_t result = - (case ls of - | ListCons x tl => - if i = int_to_u32 0 - then Return (ListCons ret tl) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - tl0 <- list_nth_mut_loop_loop_back tl i0 ret; - Return (ListCons x tl0) - od) - | ListNil => Fail Failure) -’ - -val list_nth_mut_loop_back_def = Define ‘ - (** [loops::list_nth_mut_loop]: backward function 0 *) - list_nth_mut_loop_back - (ls : 't list_t) (i : u32) (ret : 't) : 't list_t result = - list_nth_mut_loop_loop_back ls i ret -’ - -val [list_nth_shared_loop_loop_fwd_def] = DefineDiv ‘ - (** [loops::list_nth_shared_loop]: loop 0: forward function *) - list_nth_shared_loop_loop_fwd (ls : 't list_t) (i : u32) : 't result = - (case ls of - | ListCons x tl => - if i = int_to_u32 0 - then Return x - else ( - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_loop_loop_fwd tl i0 - od) - | ListNil => Fail Failure) -’ - -val list_nth_shared_loop_fwd_def = Define ‘ - (** [loops::list_nth_shared_loop]: forward function *) - list_nth_shared_loop_fwd (ls : 't list_t) (i : u32) : 't result = - list_nth_shared_loop_loop_fwd ls i -’ - -val [get_elem_mut_loop_fwd_def] = DefineDiv ‘ - (** [loops::get_elem_mut]: loop 0: forward function *) - get_elem_mut_loop_fwd (x : usize) (ls : usize list_t) : usize result = - (case ls of - | ListCons y tl => if y = x then Return y else get_elem_mut_loop_fwd x tl - | ListNil => Fail Failure) -’ - -val get_elem_mut_fwd_def = Define ‘ - (** [loops::get_elem_mut]: forward function *) - get_elem_mut_fwd (slots : usize list_t vec) (x : usize) : usize result = - do - l <- vec_index_mut_fwd slots (int_to_usize 0); - get_elem_mut_loop_fwd x l - od -’ - -val [get_elem_mut_loop_back_def] = DefineDiv ‘ - (** [loops::get_elem_mut]: loop 0: backward function 0 *) - get_elem_mut_loop_back - (x : usize) (ls : usize list_t) (ret : usize) : usize list_t result = - (case ls of - | ListCons y tl => - if y = x - then Return (ListCons ret tl) - else ( - do - tl0 <- get_elem_mut_loop_back x tl ret; - Return (ListCons y tl0) - od) - | ListNil => Fail Failure) -’ - -val get_elem_mut_back_def = Define ‘ - (** [loops::get_elem_mut]: backward function 0 *) - get_elem_mut_back - (slots : usize list_t vec) (x : usize) (ret : usize) : - usize list_t vec result - = - do - l <- vec_index_mut_fwd slots (int_to_usize 0); - l0 <- get_elem_mut_loop_back x l ret; - vec_index_mut_back slots (int_to_usize 0) l0 - od -’ - -val [get_elem_shared_loop_fwd_def] = DefineDiv ‘ - (** [loops::get_elem_shared]: loop 0: forward function *) - get_elem_shared_loop_fwd (x : usize) (ls : usize list_t) : usize result = - (case ls of - | ListCons y tl => - if y = x then Return y else get_elem_shared_loop_fwd x tl - | ListNil => Fail Failure) -’ - -val get_elem_shared_fwd_def = Define ‘ - (** [loops::get_elem_shared]: forward function *) - get_elem_shared_fwd (slots : usize list_t vec) (x : usize) : usize result = - do - l <- vec_index_fwd slots (int_to_usize 0); - get_elem_shared_loop_fwd x l - od -’ - -val id_mut_fwd_def = Define ‘ - (** [loops::id_mut]: forward function *) - id_mut_fwd (ls : 't list_t) : 't list_t result = - Return ls -’ - -val id_mut_back_def = Define ‘ - (** [loops::id_mut]: backward function 0 *) - id_mut_back (ls : 't list_t) (ret : 't list_t) : 't list_t result = - Return ret -’ - -val id_shared_fwd_def = Define ‘ - (** [loops::id_shared]: forward function *) - id_shared_fwd (ls : 't list_t) : 't list_t result = - Return ls -’ - -val [list_nth_mut_loop_with_id_loop_fwd_def] = DefineDiv ‘ - (** [loops::list_nth_mut_loop_with_id]: loop 0: forward function *) - list_nth_mut_loop_with_id_loop_fwd (i : u32) (ls : 't list_t) : 't result = - (case ls of - | ListCons x tl => - if i = int_to_u32 0 - then Return x - else ( - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_loop_with_id_loop_fwd i0 tl - od) - | ListNil => Fail Failure) -’ - -val list_nth_mut_loop_with_id_fwd_def = Define ‘ - (** [loops::list_nth_mut_loop_with_id]: forward function *) - list_nth_mut_loop_with_id_fwd (ls : 't list_t) (i : u32) : 't result = - do - ls0 <- id_mut_fwd ls; - list_nth_mut_loop_with_id_loop_fwd i ls0 - od -’ - -val [list_nth_mut_loop_with_id_loop_back_def] = DefineDiv ‘ - (** [loops::list_nth_mut_loop_with_id]: loop 0: backward function 0 *) - list_nth_mut_loop_with_id_loop_back - (i : u32) (ls : 't list_t) (ret : 't) : 't list_t result = - (case ls of - | ListCons x tl => - if i = int_to_u32 0 - then Return (ListCons ret tl) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - tl0 <- list_nth_mut_loop_with_id_loop_back i0 tl ret; - Return (ListCons x tl0) - od) - | ListNil => Fail Failure) -’ - -val list_nth_mut_loop_with_id_back_def = Define ‘ - (** [loops::list_nth_mut_loop_with_id]: backward function 0 *) - list_nth_mut_loop_with_id_back - (ls : 't list_t) (i : u32) (ret : 't) : 't list_t result = - do - ls0 <- id_mut_fwd ls; - l <- list_nth_mut_loop_with_id_loop_back i ls0 ret; - id_mut_back ls l - od -’ - -val [list_nth_shared_loop_with_id_loop_fwd_def] = DefineDiv ‘ - (** [loops::list_nth_shared_loop_with_id]: loop 0: forward function *) - list_nth_shared_loop_with_id_loop_fwd - (i : u32) (ls : 't list_t) : 't result = - (case ls of - | ListCons x tl => - if i = int_to_u32 0 - then Return x - else ( - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_loop_with_id_loop_fwd i0 tl - od) - | ListNil => Fail Failure) -’ - -val list_nth_shared_loop_with_id_fwd_def = Define ‘ - (** [loops::list_nth_shared_loop_with_id]: forward function *) - list_nth_shared_loop_with_id_fwd (ls : 't list_t) (i : u32) : 't result = - do - ls0 <- id_shared_fwd ls; - list_nth_shared_loop_with_id_loop_fwd i ls0 - od -’ - -val [list_nth_mut_loop_pair_loop_fwd_def] = DefineDiv ‘ - (** [loops::list_nth_mut_loop_pair]: loop 0: forward function *) - list_nth_mut_loop_pair_loop_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (x0, x1) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_loop_pair_loop_fwd tl0 tl1 i0 - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_mut_loop_pair_fwd_def = Define ‘ - (** [loops::list_nth_mut_loop_pair]: forward function *) - list_nth_mut_loop_pair_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - list_nth_mut_loop_pair_loop_fwd ls0 ls1 i -’ - -val [list_nth_mut_loop_pair_loop_back'a_def] = DefineDiv ‘ - (** [loops::list_nth_mut_loop_pair]: loop 0: backward function 0 *) - list_nth_mut_loop_pair_loop_back'a - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : - 't list_t result - = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (ListCons ret tl0) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - tl00 <- list_nth_mut_loop_pair_loop_back'a tl0 tl1 i0 ret; - Return (ListCons x0 tl00) - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_mut_loop_pair_back'a_def = Define ‘ - (** [loops::list_nth_mut_loop_pair]: backward function 0 *) - list_nth_mut_loop_pair_back'a - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : - 't list_t result - = - list_nth_mut_loop_pair_loop_back'a ls0 ls1 i ret -’ - -val [list_nth_mut_loop_pair_loop_back'b_def] = DefineDiv ‘ - (** [loops::list_nth_mut_loop_pair]: loop 0: backward function 1 *) - list_nth_mut_loop_pair_loop_back'b - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : - 't list_t result - = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (ListCons ret tl1) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - tl10 <- list_nth_mut_loop_pair_loop_back'b tl0 tl1 i0 ret; - Return (ListCons x1 tl10) - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_mut_loop_pair_back'b_def = Define ‘ - (** [loops::list_nth_mut_loop_pair]: backward function 1 *) - list_nth_mut_loop_pair_back'b - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : - 't list_t result - = - list_nth_mut_loop_pair_loop_back'b ls0 ls1 i ret -’ - -val [list_nth_shared_loop_pair_loop_fwd_def] = DefineDiv ‘ - (** [loops::list_nth_shared_loop_pair]: loop 0: forward function *) - list_nth_shared_loop_pair_loop_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (x0, x1) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_loop_pair_loop_fwd tl0 tl1 i0 - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_shared_loop_pair_fwd_def = Define ‘ - (** [loops::list_nth_shared_loop_pair]: forward function *) - list_nth_shared_loop_pair_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - list_nth_shared_loop_pair_loop_fwd ls0 ls1 i -’ - -val [list_nth_mut_loop_pair_merge_loop_fwd_def] = DefineDiv ‘ - (** [loops::list_nth_mut_loop_pair_merge]: loop 0: forward function *) - list_nth_mut_loop_pair_merge_loop_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (x0, x1) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_loop_pair_merge_loop_fwd tl0 tl1 i0 - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_mut_loop_pair_merge_fwd_def = Define ‘ - (** [loops::list_nth_mut_loop_pair_merge]: forward function *) - list_nth_mut_loop_pair_merge_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - list_nth_mut_loop_pair_merge_loop_fwd ls0 ls1 i -’ - -val [list_nth_mut_loop_pair_merge_loop_back_def] = DefineDiv ‘ - (** [loops::list_nth_mut_loop_pair_merge]: loop 0: backward function 0 *) - list_nth_mut_loop_pair_merge_loop_back - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : ('t # 't)) : - ('t list_t # 't list_t) result - = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then let (t, t0) = ret in Return (ListCons t tl0, ListCons t0 tl1) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - (tl00, tl10) <- - list_nth_mut_loop_pair_merge_loop_back tl0 tl1 i0 ret; - Return (ListCons x0 tl00, ListCons x1 tl10) - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_mut_loop_pair_merge_back_def = Define ‘ - (** [loops::list_nth_mut_loop_pair_merge]: backward function 0 *) - list_nth_mut_loop_pair_merge_back - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : ('t # 't)) : - ('t list_t # 't list_t) result - = - list_nth_mut_loop_pair_merge_loop_back ls0 ls1 i ret -’ - -val [list_nth_shared_loop_pair_merge_loop_fwd_def] = DefineDiv ‘ - (** [loops::list_nth_shared_loop_pair_merge]: loop 0: forward function *) - list_nth_shared_loop_pair_merge_loop_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (x0, x1) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_loop_pair_merge_loop_fwd tl0 tl1 i0 - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_shared_loop_pair_merge_fwd_def = Define ‘ - (** [loops::list_nth_shared_loop_pair_merge]: forward function *) - list_nth_shared_loop_pair_merge_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - list_nth_shared_loop_pair_merge_loop_fwd ls0 ls1 i -’ - -val [list_nth_mut_shared_loop_pair_loop_fwd_def] = DefineDiv ‘ - (** [loops::list_nth_mut_shared_loop_pair]: loop 0: forward function *) - list_nth_mut_shared_loop_pair_loop_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (x0, x1) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_shared_loop_pair_loop_fwd tl0 tl1 i0 - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_mut_shared_loop_pair_fwd_def = Define ‘ - (** [loops::list_nth_mut_shared_loop_pair]: forward function *) - list_nth_mut_shared_loop_pair_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - list_nth_mut_shared_loop_pair_loop_fwd ls0 ls1 i -’ - -val [list_nth_mut_shared_loop_pair_loop_back_def] = DefineDiv ‘ - (** [loops::list_nth_mut_shared_loop_pair]: loop 0: backward function 0 *) - list_nth_mut_shared_loop_pair_loop_back - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : - 't list_t result - = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (ListCons ret tl0) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - tl00 <- list_nth_mut_shared_loop_pair_loop_back tl0 tl1 i0 ret; - Return (ListCons x0 tl00) - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_mut_shared_loop_pair_back_def = Define ‘ - (** [loops::list_nth_mut_shared_loop_pair]: backward function 0 *) - list_nth_mut_shared_loop_pair_back - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : - 't list_t result - = - list_nth_mut_shared_loop_pair_loop_back ls0 ls1 i ret -’ - -val [list_nth_mut_shared_loop_pair_merge_loop_fwd_def] = DefineDiv ‘ - (** [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: forward function *) - list_nth_mut_shared_loop_pair_merge_loop_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (x0, x1) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_shared_loop_pair_merge_loop_fwd tl0 tl1 i0 - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_mut_shared_loop_pair_merge_fwd_def = Define ‘ - (** [loops::list_nth_mut_shared_loop_pair_merge]: forward function *) - list_nth_mut_shared_loop_pair_merge_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - list_nth_mut_shared_loop_pair_merge_loop_fwd ls0 ls1 i -’ - -val [list_nth_mut_shared_loop_pair_merge_loop_back_def] = DefineDiv ‘ - (** [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: backward function 0 *) - list_nth_mut_shared_loop_pair_merge_loop_back - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : - 't list_t result - = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (ListCons ret tl0) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - tl00 <- list_nth_mut_shared_loop_pair_merge_loop_back tl0 tl1 i0 ret; - Return (ListCons x0 tl00) - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_mut_shared_loop_pair_merge_back_def = Define ‘ - (** [loops::list_nth_mut_shared_loop_pair_merge]: backward function 0 *) - list_nth_mut_shared_loop_pair_merge_back - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : - 't list_t result - = - list_nth_mut_shared_loop_pair_merge_loop_back ls0 ls1 i ret -’ - -val [list_nth_shared_mut_loop_pair_loop_fwd_def] = DefineDiv ‘ - (** [loops::list_nth_shared_mut_loop_pair]: loop 0: forward function *) - list_nth_shared_mut_loop_pair_loop_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (x0, x1) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_mut_loop_pair_loop_fwd tl0 tl1 i0 - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_shared_mut_loop_pair_fwd_def = Define ‘ - (** [loops::list_nth_shared_mut_loop_pair]: forward function *) - list_nth_shared_mut_loop_pair_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - list_nth_shared_mut_loop_pair_loop_fwd ls0 ls1 i -’ - -val [list_nth_shared_mut_loop_pair_loop_back_def] = DefineDiv ‘ - (** [loops::list_nth_shared_mut_loop_pair]: loop 0: backward function 1 *) - list_nth_shared_mut_loop_pair_loop_back - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : - 't list_t result - = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (ListCons ret tl1) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - tl10 <- list_nth_shared_mut_loop_pair_loop_back tl0 tl1 i0 ret; - Return (ListCons x1 tl10) - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_shared_mut_loop_pair_back_def = Define ‘ - (** [loops::list_nth_shared_mut_loop_pair]: backward function 1 *) - list_nth_shared_mut_loop_pair_back - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : - 't list_t result - = - list_nth_shared_mut_loop_pair_loop_back ls0 ls1 i ret -’ - -val [list_nth_shared_mut_loop_pair_merge_loop_fwd_def] = DefineDiv ‘ - (** [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: forward function *) - list_nth_shared_mut_loop_pair_merge_loop_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (x0, x1) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_mut_loop_pair_merge_loop_fwd tl0 tl1 i0 - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_shared_mut_loop_pair_merge_fwd_def = Define ‘ - (** [loops::list_nth_shared_mut_loop_pair_merge]: forward function *) - list_nth_shared_mut_loop_pair_merge_fwd - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) : ('t # 't) result = - list_nth_shared_mut_loop_pair_merge_loop_fwd ls0 ls1 i -’ - -val [list_nth_shared_mut_loop_pair_merge_loop_back_def] = DefineDiv ‘ - (** [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: backward function 0 *) - list_nth_shared_mut_loop_pair_merge_loop_back - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : - 't list_t result - = - (case ls0 of - | ListCons x0 tl0 => - (case ls1 of - | ListCons x1 tl1 => - if i = int_to_u32 0 - then Return (ListCons ret tl1) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - tl10 <- list_nth_shared_mut_loop_pair_merge_loop_back tl0 tl1 i0 ret; - Return (ListCons x1 tl10) - od) - | ListNil => Fail Failure) - | ListNil => Fail Failure) -’ - -val list_nth_shared_mut_loop_pair_merge_back_def = Define ‘ - (** [loops::list_nth_shared_mut_loop_pair_merge]: backward function 0 *) - list_nth_shared_mut_loop_pair_merge_back - (ls0 : 't list_t) (ls1 : 't list_t) (i : u32) (ret : 't) : - 't list_t result - = - list_nth_shared_mut_loop_pair_merge_loop_back ls0 ls1 i ret -’ - -val _ = export_theory () diff --git a/tests/hol4/misc-loops/loops_FunsTheory.sig b/tests/hol4/misc-loops/loops_FunsTheory.sig deleted file mode 100644 index 63fe56c9..00000000 --- a/tests/hol4/misc-loops/loops_FunsTheory.sig +++ /dev/null @@ -1,731 +0,0 @@ -signature loops_FunsTheory = -sig - type thm = Thm.thm - - (* Definitions *) - val clear_fwd_back_def : thm - val clear_loop_fwd_back_def : thm - val get_elem_mut_back_def : thm - val get_elem_mut_fwd_def : thm - val get_elem_mut_loop_back_def : thm - val get_elem_mut_loop_fwd_def : thm - val get_elem_shared_fwd_def : thm - val get_elem_shared_loop_fwd_def : thm - val id_mut_back_def : thm - val id_mut_fwd_def : thm - val id_shared_fwd_def : thm - val list_mem_fwd_def : thm - val list_mem_loop_fwd_def : thm - val list_nth_mut_loop_back_def : thm - val list_nth_mut_loop_fwd_def : thm - val list_nth_mut_loop_loop_back_def : thm - val list_nth_mut_loop_loop_fwd_def : thm - val list_nth_mut_loop_pair_back'a_def : thm - val list_nth_mut_loop_pair_back'b_def : thm - val list_nth_mut_loop_pair_fwd_def : thm - val list_nth_mut_loop_pair_loop_back'a_def : thm - val list_nth_mut_loop_pair_loop_back'b_def : thm - val list_nth_mut_loop_pair_loop_fwd_def : thm - val list_nth_mut_loop_pair_merge_back_def : thm - val list_nth_mut_loop_pair_merge_fwd_def : thm - val list_nth_mut_loop_pair_merge_loop_back_def : thm - val list_nth_mut_loop_pair_merge_loop_fwd_def : thm - val list_nth_mut_loop_with_id_back_def : thm - val list_nth_mut_loop_with_id_fwd_def : thm - val list_nth_mut_loop_with_id_loop_back_def : thm - val list_nth_mut_loop_with_id_loop_fwd_def : thm - val list_nth_mut_shared_loop_pair_back_def : thm - val list_nth_mut_shared_loop_pair_fwd_def : thm - val list_nth_mut_shared_loop_pair_loop_back_def : thm - val list_nth_mut_shared_loop_pair_loop_fwd_def : thm - val list_nth_mut_shared_loop_pair_merge_back_def : thm - val list_nth_mut_shared_loop_pair_merge_fwd_def : thm - val list_nth_mut_shared_loop_pair_merge_loop_back_def : thm - val list_nth_mut_shared_loop_pair_merge_loop_fwd_def : thm - val list_nth_shared_loop_fwd_def : thm - val list_nth_shared_loop_loop_fwd_def : thm - val list_nth_shared_loop_pair_fwd_def : thm - val list_nth_shared_loop_pair_loop_fwd_def : thm - val list_nth_shared_loop_pair_merge_fwd_def : thm - val list_nth_shared_loop_pair_merge_loop_fwd_def : thm - val list_nth_shared_loop_with_id_fwd_def : thm - val list_nth_shared_loop_with_id_loop_fwd_def : thm - val list_nth_shared_mut_loop_pair_back_def : thm - val list_nth_shared_mut_loop_pair_fwd_def : thm - val list_nth_shared_mut_loop_pair_loop_back_def : thm - val list_nth_shared_mut_loop_pair_loop_fwd_def : thm - val list_nth_shared_mut_loop_pair_merge_back_def : thm - val list_nth_shared_mut_loop_pair_merge_fwd_def : thm - val list_nth_shared_mut_loop_pair_merge_loop_back_def : thm - val list_nth_shared_mut_loop_pair_merge_loop_fwd_def : thm - val sum_fwd_def : thm - val sum_loop_fwd_def : thm - val sum_with_mut_borrows_fwd_def : thm - val sum_with_mut_borrows_loop_fwd_def : thm - val sum_with_shared_borrows_fwd_def : thm - val sum_with_shared_borrows_loop_fwd_def : thm - - val loops_Funs_grammars : type_grammar.grammar * term_grammar.grammar -(* - [loops_Types] Parent theory of "loops_Funs" - - [clear_fwd_back_def] Definition - - ⊢ ∀v. clear_fwd_back v = clear_loop_fwd_back v (int_to_usize 0) - - [clear_loop_fwd_back_def] Definition - - ⊢ ∀v i. - clear_loop_fwd_back v i = - (let - i0 = vec_len v - in - if usize_lt i i0 then - do - i1 <- usize_add i (int_to_usize 1); - v0 <- vec_index_mut_back v i (int_to_u32 0); - clear_loop_fwd_back v0 i1 - od - else Return v) - - [get_elem_mut_back_def] Definition - - ⊢ ∀slots x ret. - get_elem_mut_back slots x ret = - do - l <- vec_index_mut_fwd slots (int_to_usize 0); - l0 <- get_elem_mut_loop_back x l ret; - vec_index_mut_back slots (int_to_usize 0) l0 - od - - [get_elem_mut_fwd_def] Definition - - ⊢ ∀slots x. - get_elem_mut_fwd slots x = - do - l <- vec_index_mut_fwd slots (int_to_usize 0); - get_elem_mut_loop_fwd x l - od - - [get_elem_mut_loop_back_def] Definition - - ⊢ ∀x ls ret. - get_elem_mut_loop_back x ls ret = - case ls of - ListCons y tl => - if y = x then Return (ListCons ret tl) - else - do - tl0 <- get_elem_mut_loop_back x tl ret; - Return (ListCons y tl0) - od - | ListNil => Fail Failure - - [get_elem_mut_loop_fwd_def] Definition - - ⊢ ∀x ls. - get_elem_mut_loop_fwd x ls = - case ls of - ListCons y tl => - if y = x then Return y else get_elem_mut_loop_fwd x tl - | ListNil => Fail Failure - - [get_elem_shared_fwd_def] Definition - - ⊢ ∀slots x. - get_elem_shared_fwd slots x = - do - l <- vec_index_fwd slots (int_to_usize 0); - get_elem_shared_loop_fwd x l - od - - [get_elem_shared_loop_fwd_def] Definition - - ⊢ ∀x ls. - get_elem_shared_loop_fwd x ls = - case ls of - ListCons y tl => - if y = x then Return y else get_elem_shared_loop_fwd x tl - | ListNil => Fail Failure - - [id_mut_back_def] Definition - - ⊢ ∀ls ret. id_mut_back ls ret = Return ret - - [id_mut_fwd_def] Definition - - ⊢ ∀ls. id_mut_fwd ls = Return ls - - [id_shared_fwd_def] Definition - - ⊢ ∀ls. id_shared_fwd ls = Return ls - - [list_mem_fwd_def] Definition - - ⊢ ∀x ls. list_mem_fwd x ls = list_mem_loop_fwd x ls - - [list_mem_loop_fwd_def] Definition - - ⊢ ∀x ls. - list_mem_loop_fwd x ls = - case ls of - ListCons y tl => - if y = x then Return T else list_mem_loop_fwd x tl - | ListNil => Return F - - [list_nth_mut_loop_back_def] Definition - - ⊢ ∀ls i ret. - list_nth_mut_loop_back ls i ret = - list_nth_mut_loop_loop_back ls i ret - - [list_nth_mut_loop_fwd_def] Definition - - ⊢ ∀ls i. list_nth_mut_loop_fwd ls i = list_nth_mut_loop_loop_fwd ls i - - [list_nth_mut_loop_loop_back_def] Definition - - ⊢ ∀ls i ret. - list_nth_mut_loop_loop_back ls i ret = - case ls of - ListCons x tl => - if i = int_to_u32 0 then Return (ListCons ret tl) - else - do - i0 <- u32_sub i (int_to_u32 1); - tl0 <- list_nth_mut_loop_loop_back tl i0 ret; - Return (ListCons x tl0) - od - | ListNil => Fail Failure - - [list_nth_mut_loop_loop_fwd_def] Definition - - ⊢ ∀ls i. - list_nth_mut_loop_loop_fwd ls i = - case ls of - ListCons x tl => - if i = int_to_u32 0 then Return x - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_loop_loop_fwd tl i0 - od - | ListNil => Fail Failure - - [list_nth_mut_loop_pair_back'a_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_mut_loop_pair_back'a ls0 ls1 i ret = - list_nth_mut_loop_pair_loop_back'a ls0 ls1 i ret - - [list_nth_mut_loop_pair_back'b_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_mut_loop_pair_back'b ls0 ls1 i ret = - list_nth_mut_loop_pair_loop_back'b ls0 ls1 i ret - - [list_nth_mut_loop_pair_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_mut_loop_pair_fwd ls0 ls1 i = - list_nth_mut_loop_pair_loop_fwd ls0 ls1 i - - [list_nth_mut_loop_pair_loop_back'a_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_mut_loop_pair_loop_back'a ls0 ls1 i ret = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (ListCons ret tl0) - else - do - i0 <- u32_sub i (int_to_u32 1); - tl00 <- - list_nth_mut_loop_pair_loop_back'a tl0 tl1 i0 ret; - Return (ListCons x0 tl00) - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_mut_loop_pair_loop_back'b_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_mut_loop_pair_loop_back'b ls0 ls1 i ret = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (ListCons ret tl1) - else - do - i0 <- u32_sub i (int_to_u32 1); - tl10 <- - list_nth_mut_loop_pair_loop_back'b tl0 tl1 i0 ret; - Return (ListCons x1 tl10) - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_mut_loop_pair_loop_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_mut_loop_pair_loop_fwd ls0 ls1 i = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (x0,x1) - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_loop_pair_loop_fwd tl0 tl1 i0 - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_mut_loop_pair_merge_back_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_mut_loop_pair_merge_back ls0 ls1 i ret = - list_nth_mut_loop_pair_merge_loop_back ls0 ls1 i ret - - [list_nth_mut_loop_pair_merge_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_mut_loop_pair_merge_fwd ls0 ls1 i = - list_nth_mut_loop_pair_merge_loop_fwd ls0 ls1 i - - [list_nth_mut_loop_pair_merge_loop_back_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_mut_loop_pair_merge_loop_back ls0 ls1 i ret = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then - (let - (t,t0) = ret - in - Return (ListCons t tl0,ListCons t0 tl1)) - else - do - i0 <- u32_sub i (int_to_u32 1); - (tl00,tl10) <- - list_nth_mut_loop_pair_merge_loop_back tl0 tl1 i0 - ret; - Return (ListCons x0 tl00,ListCons x1 tl10) - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_mut_loop_pair_merge_loop_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_mut_loop_pair_merge_loop_fwd ls0 ls1 i = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (x0,x1) - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_loop_pair_merge_loop_fwd tl0 tl1 i0 - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_mut_loop_with_id_back_def] Definition - - ⊢ ∀ls i ret. - list_nth_mut_loop_with_id_back ls i ret = - do - ls0 <- id_mut_fwd ls; - l <- list_nth_mut_loop_with_id_loop_back i ls0 ret; - id_mut_back ls l - od - - [list_nth_mut_loop_with_id_fwd_def] Definition - - ⊢ ∀ls i. - list_nth_mut_loop_with_id_fwd ls i = - do - ls0 <- id_mut_fwd ls; - list_nth_mut_loop_with_id_loop_fwd i ls0 - od - - [list_nth_mut_loop_with_id_loop_back_def] Definition - - ⊢ ∀i ls ret. - list_nth_mut_loop_with_id_loop_back i ls ret = - case ls of - ListCons x tl => - if i = int_to_u32 0 then Return (ListCons ret tl) - else - do - i0 <- u32_sub i (int_to_u32 1); - tl0 <- list_nth_mut_loop_with_id_loop_back i0 tl ret; - Return (ListCons x tl0) - od - | ListNil => Fail Failure - - [list_nth_mut_loop_with_id_loop_fwd_def] Definition - - ⊢ ∀i ls. - list_nth_mut_loop_with_id_loop_fwd i ls = - case ls of - ListCons x tl => - if i = int_to_u32 0 then Return x - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_loop_with_id_loop_fwd i0 tl - od - | ListNil => Fail Failure - - [list_nth_mut_shared_loop_pair_back_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_mut_shared_loop_pair_back ls0 ls1 i ret = - list_nth_mut_shared_loop_pair_loop_back ls0 ls1 i ret - - [list_nth_mut_shared_loop_pair_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_mut_shared_loop_pair_fwd ls0 ls1 i = - list_nth_mut_shared_loop_pair_loop_fwd ls0 ls1 i - - [list_nth_mut_shared_loop_pair_loop_back_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_mut_shared_loop_pair_loop_back ls0 ls1 i ret = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (ListCons ret tl0) - else - do - i0 <- u32_sub i (int_to_u32 1); - tl00 <- - list_nth_mut_shared_loop_pair_loop_back tl0 tl1 i0 - ret; - Return (ListCons x0 tl00) - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_mut_shared_loop_pair_loop_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_mut_shared_loop_pair_loop_fwd ls0 ls1 i = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (x0,x1) - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_shared_loop_pair_loop_fwd tl0 tl1 i0 - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_mut_shared_loop_pair_merge_back_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_mut_shared_loop_pair_merge_back ls0 ls1 i ret = - list_nth_mut_shared_loop_pair_merge_loop_back ls0 ls1 i ret - - [list_nth_mut_shared_loop_pair_merge_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_mut_shared_loop_pair_merge_fwd ls0 ls1 i = - list_nth_mut_shared_loop_pair_merge_loop_fwd ls0 ls1 i - - [list_nth_mut_shared_loop_pair_merge_loop_back_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_mut_shared_loop_pair_merge_loop_back ls0 ls1 i ret = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (ListCons ret tl0) - else - do - i0 <- u32_sub i (int_to_u32 1); - tl00 <- - list_nth_mut_shared_loop_pair_merge_loop_back tl0 - tl1 i0 ret; - Return (ListCons x0 tl00) - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_mut_shared_loop_pair_merge_loop_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_mut_shared_loop_pair_merge_loop_fwd ls0 ls1 i = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (x0,x1) - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_shared_loop_pair_merge_loop_fwd tl0 tl1 - i0 - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_shared_loop_fwd_def] Definition - - ⊢ ∀ls i. - list_nth_shared_loop_fwd ls i = - list_nth_shared_loop_loop_fwd ls i - - [list_nth_shared_loop_loop_fwd_def] Definition - - ⊢ ∀ls i. - list_nth_shared_loop_loop_fwd ls i = - case ls of - ListCons x tl => - if i = int_to_u32 0 then Return x - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_loop_loop_fwd tl i0 - od - | ListNil => Fail Failure - - [list_nth_shared_loop_pair_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_shared_loop_pair_fwd ls0 ls1 i = - list_nth_shared_loop_pair_loop_fwd ls0 ls1 i - - [list_nth_shared_loop_pair_loop_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_shared_loop_pair_loop_fwd ls0 ls1 i = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (x0,x1) - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_loop_pair_loop_fwd tl0 tl1 i0 - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_shared_loop_pair_merge_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_shared_loop_pair_merge_fwd ls0 ls1 i = - list_nth_shared_loop_pair_merge_loop_fwd ls0 ls1 i - - [list_nth_shared_loop_pair_merge_loop_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_shared_loop_pair_merge_loop_fwd ls0 ls1 i = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (x0,x1) - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_loop_pair_merge_loop_fwd tl0 tl1 i0 - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_shared_loop_with_id_fwd_def] Definition - - ⊢ ∀ls i. - list_nth_shared_loop_with_id_fwd ls i = - do - ls0 <- id_shared_fwd ls; - list_nth_shared_loop_with_id_loop_fwd i ls0 - od - - [list_nth_shared_loop_with_id_loop_fwd_def] Definition - - ⊢ ∀i ls. - list_nth_shared_loop_with_id_loop_fwd i ls = - case ls of - ListCons x tl => - if i = int_to_u32 0 then Return x - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_loop_with_id_loop_fwd i0 tl - od - | ListNil => Fail Failure - - [list_nth_shared_mut_loop_pair_back_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_shared_mut_loop_pair_back ls0 ls1 i ret = - list_nth_shared_mut_loop_pair_loop_back ls0 ls1 i ret - - [list_nth_shared_mut_loop_pair_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_shared_mut_loop_pair_fwd ls0 ls1 i = - list_nth_shared_mut_loop_pair_loop_fwd ls0 ls1 i - - [list_nth_shared_mut_loop_pair_loop_back_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_shared_mut_loop_pair_loop_back ls0 ls1 i ret = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (ListCons ret tl1) - else - do - i0 <- u32_sub i (int_to_u32 1); - tl10 <- - list_nth_shared_mut_loop_pair_loop_back tl0 tl1 i0 - ret; - Return (ListCons x1 tl10) - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_shared_mut_loop_pair_loop_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_shared_mut_loop_pair_loop_fwd ls0 ls1 i = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (x0,x1) - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_mut_loop_pair_loop_fwd tl0 tl1 i0 - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_shared_mut_loop_pair_merge_back_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_shared_mut_loop_pair_merge_back ls0 ls1 i ret = - list_nth_shared_mut_loop_pair_merge_loop_back ls0 ls1 i ret - - [list_nth_shared_mut_loop_pair_merge_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_shared_mut_loop_pair_merge_fwd ls0 ls1 i = - list_nth_shared_mut_loop_pair_merge_loop_fwd ls0 ls1 i - - [list_nth_shared_mut_loop_pair_merge_loop_back_def] Definition - - ⊢ ∀ls0 ls1 i ret. - list_nth_shared_mut_loop_pair_merge_loop_back ls0 ls1 i ret = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (ListCons ret tl1) - else - do - i0 <- u32_sub i (int_to_u32 1); - tl10 <- - list_nth_shared_mut_loop_pair_merge_loop_back tl0 - tl1 i0 ret; - Return (ListCons x1 tl10) - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [list_nth_shared_mut_loop_pair_merge_loop_fwd_def] Definition - - ⊢ ∀ls0 ls1 i. - list_nth_shared_mut_loop_pair_merge_loop_fwd ls0 ls1 i = - case ls0 of - ListCons x0 tl0 => - (case ls1 of - ListCons x1 tl1 => - if i = int_to_u32 0 then Return (x0,x1) - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_mut_loop_pair_merge_loop_fwd tl0 tl1 - i0 - od - | ListNil => Fail Failure) - | ListNil => Fail Failure - - [sum_fwd_def] Definition - - ⊢ ∀max. sum_fwd max = sum_loop_fwd max (int_to_u32 0) (int_to_u32 0) - - [sum_loop_fwd_def] Definition - - ⊢ ∀max i s. - sum_loop_fwd max i s = - if u32_lt i max then - do - s0 <- u32_add s i; - i0 <- u32_add i (int_to_u32 1); - sum_loop_fwd max i0 s0 - od - else u32_mul s (int_to_u32 2) - - [sum_with_mut_borrows_fwd_def] Definition - - ⊢ ∀max. - sum_with_mut_borrows_fwd max = - sum_with_mut_borrows_loop_fwd max (int_to_u32 0) (int_to_u32 0) - - [sum_with_mut_borrows_loop_fwd_def] Definition - - ⊢ ∀max mi ms. - sum_with_mut_borrows_loop_fwd max mi ms = - if u32_lt mi max then - do - ms0 <- u32_add ms mi; - mi0 <- u32_add mi (int_to_u32 1); - sum_with_mut_borrows_loop_fwd max mi0 ms0 - od - else u32_mul ms (int_to_u32 2) - - [sum_with_shared_borrows_fwd_def] Definition - - ⊢ ∀max. - sum_with_shared_borrows_fwd max = - sum_with_shared_borrows_loop_fwd max (int_to_u32 0) - (int_to_u32 0) - - [sum_with_shared_borrows_loop_fwd_def] Definition - - ⊢ ∀max i s. - sum_with_shared_borrows_loop_fwd max i s = - if u32_lt i max then - do - i0 <- u32_add i (int_to_u32 1); - s0 <- u32_add s i0; - sum_with_shared_borrows_loop_fwd max i0 s0 - od - else u32_mul s (int_to_u32 2) - - -*) -end diff --git a/tests/hol4/misc-loops/loops_TypesScript.sml b/tests/hol4/misc-loops/loops_TypesScript.sml deleted file mode 100644 index e3e5b8d1..00000000 --- a/tests/hol4/misc-loops/loops_TypesScript.sml +++ /dev/null @@ -1,13 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [loops]: type definitions *) -open primitivesLib divDefLib - -val _ = new_theory "loops_Types" - - -Datatype: - (** [loops::List] *) - list_t = | ListCons 't list_t | ListNil -End - -val _ = export_theory () diff --git a/tests/hol4/misc-loops/loops_TypesTheory.sig b/tests/hol4/misc-loops/loops_TypesTheory.sig deleted file mode 100644 index c3e638d8..00000000 --- a/tests/hol4/misc-loops/loops_TypesTheory.sig +++ /dev/null @@ -1,94 +0,0 @@ -signature loops_TypesTheory = -sig - type thm = Thm.thm - - (* Definitions *) - val list_t_TY_DEF : thm - val list_t_case_def : thm - val list_t_size_def : thm - - (* Theorems *) - val datatype_list_t : thm - val list_t_11 : thm - val list_t_Axiom : thm - val list_t_case_cong : thm - val list_t_case_eq : thm - val list_t_distinct : thm - val list_t_induction : thm - val list_t_nchotomy : thm - - val loops_Types_grammars : type_grammar.grammar * term_grammar.grammar -(* - [divDef] Parent theory of "loops_Types" - - [list_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('list_t'). - (∀a0'. - (∃a0 a1. - a0' = - (λa0 a1. - ind_type$CONSTR 0 a0 - (ind_type$FCONS a1 (λn. ind_type$BOTTOM))) - a0 a1 ∧ $var$('list_t') a1) ∨ - a0' = - ind_type$CONSTR (SUC 0) ARB (λn. ind_type$BOTTOM) ⇒ - $var$('list_t') a0') ⇒ - $var$('list_t') a0') rep - - [list_t_case_def] Definition - - ⊢ (∀a0 a1 f v. list_t_CASE (ListCons a0 a1) f v = f a0 a1) ∧ - ∀f v. list_t_CASE ListNil f v = v - - [list_t_size_def] Definition - - ⊢ (∀f a0 a1. - list_t_size f (ListCons a0 a1) = 1 + (f a0 + list_t_size f a1)) ∧ - ∀f. list_t_size f ListNil = 0 - - [datatype_list_t] Theorem - - ⊢ DATATYPE (list_t ListCons ListNil) - - [list_t_11] Theorem - - ⊢ ∀a0 a1 a0' a1'. - ListCons a0 a1 = ListCons a0' a1' ⇔ a0 = a0' ∧ a1 = a1' - - [list_t_Axiom] Theorem - - ⊢ ∀f0 f1. ∃fn. - (∀a0 a1. fn (ListCons a0 a1) = f0 a0 a1 (fn a1)) ∧ - fn ListNil = f1 - - [list_t_case_cong] Theorem - - ⊢ ∀M M' f v. - M = M' ∧ (∀a0 a1. M' = ListCons a0 a1 ⇒ f a0 a1 = f' a0 a1) ∧ - (M' = ListNil ⇒ v = v') ⇒ - list_t_CASE M f v = list_t_CASE M' f' v' - - [list_t_case_eq] Theorem - - ⊢ list_t_CASE x f v = v' ⇔ - (∃t l. x = ListCons t l ∧ f t l = v') ∨ x = ListNil ∧ v = v' - - [list_t_distinct] Theorem - - ⊢ ∀a1 a0. ListCons a0 a1 ≠ ListNil - - [list_t_induction] Theorem - - ⊢ ∀P. (∀l. P l ⇒ ∀t. P (ListCons t l)) ∧ P ListNil ⇒ ∀l. P l - - [list_t_nchotomy] Theorem - - ⊢ ∀ll. (∃t l. ll = ListCons t l) ∨ ll = ListNil - - -*) -end diff --git a/tests/hol4/misc-no_nested_borrows/Holmakefile b/tests/hol4/misc-no_nested_borrows/Holmakefile deleted file mode 100644 index 3c4b8973..00000000 --- a/tests/hol4/misc-no_nested_borrows/Holmakefile +++ /dev/null @@ -1,5 +0,0 @@ -# This file was automatically generated - modify ../Holmakefile.template instead -INCLUDES = ../../../backends/hol4 - -all: $(DEFAULT_TARGETS) -.PHONY: all diff --git a/tests/hol4/misc-no_nested_borrows/noNestedBorrowsScript.sml b/tests/hol4/misc-no_nested_borrows/noNestedBorrowsScript.sml deleted file mode 100644 index 1b2d6121..00000000 --- a/tests/hol4/misc-no_nested_borrows/noNestedBorrowsScript.sml +++ /dev/null @@ -1,592 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [no_nested_borrows] *) -open primitivesLib divDefLib - -val _ = new_theory "noNestedBorrows" - - -Datatype: - (** [no_nested_borrows::Pair] *) - pair_t = <| pair_x : 't1; pair_y : 't2; |> -End - -Datatype: - (** [no_nested_borrows::List] *) - list_t = | ListCons 't list_t | ListNil -End - -Datatype: - (** [no_nested_borrows::One] *) - one_t = | OneOne 't1 -End - -Datatype: - (** [no_nested_borrows::EmptyEnum] *) - empty_enum_t = | EmptyEnumEmpty -End - -Datatype: - (** [no_nested_borrows::Enum] *) - enum_t = | EnumVariant1 | EnumVariant2 -End - -Type empty_struct_t = “: unit” - -Datatype: - (** [no_nested_borrows::Sum] *) - sum_t = | SumLeft 't1 | SumRight 't2 -End - -val neg_test_fwd_def = Define ‘ - (** [no_nested_borrows::neg_test]: forward function *) - neg_test_fwd (x : i32) : i32 result = - i32_neg x -’ - -val add_test_fwd_def = Define ‘ - (** [no_nested_borrows::add_test]: forward function *) - add_test_fwd (x : u32) (y : u32) : u32 result = - u32_add x y -’ - -val subs_test_fwd_def = Define ‘ - (** [no_nested_borrows::subs_test]: forward function *) - subs_test_fwd (x : u32) (y : u32) : u32 result = - u32_sub x y -’ - -val div_test_fwd_def = Define ‘ - (** [no_nested_borrows::div_test]: forward function *) - div_test_fwd (x : u32) (y : u32) : u32 result = - u32_div x y -’ - -val div_test1_fwd_def = Define ‘ - (** [no_nested_borrows::div_test1]: forward function *) - div_test1_fwd (x : u32) : u32 result = - u32_div x (int_to_u32 2) -’ - -val rem_test_fwd_def = Define ‘ - (** [no_nested_borrows::rem_test]: forward function *) - rem_test_fwd (x : u32) (y : u32) : u32 result = - u32_rem x y -’ - -val cast_test_fwd_def = Define ‘ - (** [no_nested_borrows::cast_test]: forward function *) - cast_test_fwd (x : u32) : i32 result = - mk_i32 (u32_to_int x) -’ - -val test2_fwd_def = Define ‘ - (** [no_nested_borrows::test2]: forward function *) - test2_fwd : unit result = - do - _ <- u32_add (int_to_u32 23) (int_to_u32 44); - Return () - od -’ - -(** Unit test for [no_nested_borrows::test2] *) -val _ = assert_return (“test2_fwd”) - -val get_max_fwd_def = Define ‘ - (** [no_nested_borrows::get_max]: forward function *) - get_max_fwd (x : u32) (y : u32) : u32 result = - if u32_ge x y then Return x else Return y -’ - -val test3_fwd_def = Define ‘ - (** [no_nested_borrows::test3]: forward function *) - test3_fwd : unit result = - do - x <- get_max_fwd (int_to_u32 4) (int_to_u32 3); - y <- get_max_fwd (int_to_u32 10) (int_to_u32 11); - z <- u32_add x y; - if ~ (z = int_to_u32 15) then Fail Failure else Return () - od -’ - -(** Unit test for [no_nested_borrows::test3] *) -val _ = assert_return (“test3_fwd”) - -val test_neg1_fwd_def = Define ‘ - (** [no_nested_borrows::test_neg1]: forward function *) - test_neg1_fwd : unit result = - do - y <- i32_neg (int_to_i32 3); - if ~ (y = int_to_i32 (-3)) then Fail Failure else Return () - od -’ - -(** Unit test for [no_nested_borrows::test_neg1] *) -val _ = assert_return (“test_neg1_fwd”) - -val refs_test1_fwd_def = Define ‘ - (** [no_nested_borrows::refs_test1]: forward function *) - refs_test1_fwd : unit result = - if ~ (int_to_i32 1 = int_to_i32 1) then Fail Failure else Return () -’ - -(** Unit test for [no_nested_borrows::refs_test1] *) -val _ = assert_return (“refs_test1_fwd”) - -val refs_test2_fwd_def = Define ‘ - (** [no_nested_borrows::refs_test2]: forward function *) - refs_test2_fwd : unit result = - if ~ (int_to_i32 2 = int_to_i32 2) - then Fail Failure - else - if ~ (int_to_i32 0 = int_to_i32 0) - then Fail Failure - else - if ~ (int_to_i32 2 = int_to_i32 2) - then Fail Failure - else - if ~ (int_to_i32 2 = int_to_i32 2) then Fail Failure else Return () -’ - -(** Unit test for [no_nested_borrows::refs_test2] *) -val _ = assert_return (“refs_test2_fwd”) - -val test_list1_fwd_def = Define ‘ - (** [no_nested_borrows::test_list1]: forward function *) - test_list1_fwd : unit result = - Return () -’ - -(** Unit test for [no_nested_borrows::test_list1] *) -val _ = assert_return (“test_list1_fwd”) - -val test_box1_fwd_def = Define ‘ - (** [no_nested_borrows::test_box1]: forward function *) - test_box1_fwd : unit result = - let b = int_to_i32 1 in - let x = b in - if ~ (x = int_to_i32 1) then Fail Failure else Return () -’ - -(** Unit test for [no_nested_borrows::test_box1] *) -val _ = assert_return (“test_box1_fwd”) - -val copy_int_fwd_def = Define ‘ - (** [no_nested_borrows::copy_int]: forward function *) - copy_int_fwd (x : i32) : i32 result = - Return x -’ - -val test_unreachable_fwd_def = Define ‘ - (** [no_nested_borrows::test_unreachable]: forward function *) - test_unreachable_fwd (b : bool) : unit result = - if b then Fail Failure else Return () -’ - -val test_panic_fwd_def = Define ‘ - (** [no_nested_borrows::test_panic]: forward function *) - test_panic_fwd (b : bool) : unit result = - if b then Fail Failure else Return () -’ - -val test_copy_int_fwd_def = Define ‘ - (** [no_nested_borrows::test_copy_int]: forward function *) - test_copy_int_fwd : unit result = - do - y <- copy_int_fwd (int_to_i32 0); - if ~ (int_to_i32 0 = y) then Fail Failure else Return () - od -’ - -(** Unit test for [no_nested_borrows::test_copy_int] *) -val _ = assert_return (“test_copy_int_fwd”) - -val is_cons_fwd_def = Define ‘ - (** [no_nested_borrows::is_cons]: forward function *) - is_cons_fwd (l : 't list_t) : bool result = - (case l of | ListCons t l0 => Return T | ListNil => Return F) -’ - -val test_is_cons_fwd_def = Define ‘ - (** [no_nested_borrows::test_is_cons]: forward function *) - test_is_cons_fwd : unit result = - let l = ListNil in - do - b <- is_cons_fwd (ListCons (int_to_i32 0) l); - if ~ b then Fail Failure else Return () - od -’ - -(** Unit test for [no_nested_borrows::test_is_cons] *) -val _ = assert_return (“test_is_cons_fwd”) - -val split_list_fwd_def = Define ‘ - (** [no_nested_borrows::split_list]: forward function *) - split_list_fwd (l : 't list_t) : ('t # 't list_t) result = - (case l of | ListCons hd tl => Return (hd, tl) | ListNil => Fail Failure) -’ - -val test_split_list_fwd_def = Define ‘ - (** [no_nested_borrows::test_split_list]: forward function *) - test_split_list_fwd : unit result = - let l = ListNil in - do - p <- split_list_fwd (ListCons (int_to_i32 0) l); - let (hd, _) = p in - if ~ (hd = int_to_i32 0) then Fail Failure else Return () - od -’ - -(** Unit test for [no_nested_borrows::test_split_list] *) -val _ = assert_return (“test_split_list_fwd”) - -val choose_fwd_def = Define ‘ - (** [no_nested_borrows::choose]: forward function *) - choose_fwd (b : bool) (x : 't) (y : 't) : 't result = - if b then Return x else Return y -’ - -val choose_back_def = Define ‘ - (** [no_nested_borrows::choose]: backward function 0 *) - choose_back (b : bool) (x : 't) (y : 't) (ret : 't) : ('t # 't) result = - if b then Return (ret, y) else Return (x, ret) -’ - -val choose_test_fwd_def = Define ‘ - (** [no_nested_borrows::choose_test]: forward function *) - choose_test_fwd : unit result = - do - z <- choose_fwd T (int_to_i32 0) (int_to_i32 0); - z0 <- i32_add z (int_to_i32 1); - if ~ (z0 = int_to_i32 1) - then Fail Failure - else ( - do - (x, y) <- choose_back T (int_to_i32 0) (int_to_i32 0) z0; - if ~ (x = int_to_i32 1) - then Fail Failure - else if ~ (y = int_to_i32 0) then Fail Failure else Return () - od) - od -’ - -(** Unit test for [no_nested_borrows::choose_test] *) -val _ = assert_return (“choose_test_fwd”) - -val test_char_fwd_def = Define ‘ - (** [no_nested_borrows::test_char]: forward function *) - test_char_fwd : char result = - Return #"a" -’ - -Datatype: - (** [no_nested_borrows::Tree] *) - tree_t = | TreeLeaf 't | TreeNode 't node_elem_t tree_t ; - - (** [no_nested_borrows::NodeElem] *) - node_elem_t = | NodeElemCons tree_t node_elem_t | NodeElemNil -End - -val [list_length_fwd_def] = DefineDiv ‘ - (** [no_nested_borrows::list_length]: forward function *) - list_length_fwd (l : 't list_t) : u32 result = - (case l of - | ListCons t l1 => do - i <- list_length_fwd l1; - u32_add (int_to_u32 1) i - od - | ListNil => Return (int_to_u32 0)) -’ - -val [list_nth_shared_fwd_def] = DefineDiv ‘ - (** [no_nested_borrows::list_nth_shared]: forward function *) - list_nth_shared_fwd (l : 't list_t) (i : u32) : 't result = - (case l of - | ListCons x tl => - if i = int_to_u32 0 - then Return x - else (do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_fwd tl i0 - od) - | ListNil => Fail Failure) -’ - -val [list_nth_mut_fwd_def] = DefineDiv ‘ - (** [no_nested_borrows::list_nth_mut]: forward function *) - list_nth_mut_fwd (l : 't list_t) (i : u32) : 't result = - (case l of - | ListCons x tl => - if i = int_to_u32 0 - then Return x - else (do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_fwd tl i0 - od) - | ListNil => Fail Failure) -’ - -val [list_nth_mut_back_def] = DefineDiv ‘ - (** [no_nested_borrows::list_nth_mut]: backward function 0 *) - list_nth_mut_back (l : 't list_t) (i : u32) (ret : 't) : 't list_t result = - (case l of - | ListCons x tl => - if i = int_to_u32 0 - then Return (ListCons ret tl) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - tl0 <- list_nth_mut_back tl i0 ret; - Return (ListCons x tl0) - od) - | ListNil => Fail Failure) -’ - -val [list_rev_aux_fwd_def] = DefineDiv ‘ - (** [no_nested_borrows::list_rev_aux]: forward function *) - list_rev_aux_fwd (li : 't list_t) (lo : 't list_t) : 't list_t result = - (case li of - | ListCons hd tl => list_rev_aux_fwd tl (ListCons hd lo) - | ListNil => Return lo) -’ - -val list_rev_fwd_back_def = Define ‘ - (** [no_nested_borrows::list_rev]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - list_rev_fwd_back (l : 't list_t) : 't list_t result = - let li = mem_replace_fwd l ListNil in list_rev_aux_fwd li ListNil -’ - -val test_list_functions_fwd_def = Define ‘ - (** [no_nested_borrows::test_list_functions]: forward function *) - test_list_functions_fwd : unit result = - let l = ListNil in - let l0 = ListCons (int_to_i32 2) l in - let l1 = ListCons (int_to_i32 1) l0 in - do - i <- list_length_fwd (ListCons (int_to_i32 0) l1); - if ~ (i = int_to_u32 3) - then Fail Failure - else ( - do - i0 <- list_nth_shared_fwd (ListCons (int_to_i32 0) l1) (int_to_u32 0); - if ~ (i0 = int_to_i32 0) - then Fail Failure - else ( - do - i1 <- list_nth_shared_fwd (ListCons (int_to_i32 0) l1) (int_to_u32 1); - if ~ (i1 = int_to_i32 1) - then Fail Failure - else ( - do - i2 <- - list_nth_shared_fwd (ListCons (int_to_i32 0) l1) (int_to_u32 2); - if ~ (i2 = int_to_i32 2) - then Fail Failure - else ( - do - ls <- - list_nth_mut_back (ListCons (int_to_i32 0) l1) (int_to_u32 1) - (int_to_i32 3); - i3 <- list_nth_shared_fwd ls (int_to_u32 0); - if ~ (i3 = int_to_i32 0) - then Fail Failure - else ( - do - i4 <- list_nth_shared_fwd ls (int_to_u32 1); - if ~ (i4 = int_to_i32 3) - then Fail Failure - else ( - do - i5 <- list_nth_shared_fwd ls (int_to_u32 2); - if ~ (i5 = int_to_i32 2) then Fail Failure else Return () - od) - od) - od) - od) - od) - od) - od -’ - -(** Unit test for [no_nested_borrows::test_list_functions] *) -val _ = assert_return (“test_list_functions_fwd”) - -val id_mut_pair1_fwd_def = Define ‘ - (** [no_nested_borrows::id_mut_pair1]: forward function *) - id_mut_pair1_fwd (x : 't1) (y : 't2) : ('t1 # 't2) result = - Return (x, y) -’ - -val id_mut_pair1_back_def = Define ‘ - (** [no_nested_borrows::id_mut_pair1]: backward function 0 *) - id_mut_pair1_back - (x : 't1) (y : 't2) (ret : ('t1 # 't2)) : ('t1 # 't2) result = - let (t, t0) = ret in Return (t, t0) -’ - -val id_mut_pair2_fwd_def = Define ‘ - (** [no_nested_borrows::id_mut_pair2]: forward function *) - id_mut_pair2_fwd (p : ('t1 # 't2)) : ('t1 # 't2) result = - let (t, t0) = p in Return (t, t0) -’ - -val id_mut_pair2_back_def = Define ‘ - (** [no_nested_borrows::id_mut_pair2]: backward function 0 *) - id_mut_pair2_back - (p : ('t1 # 't2)) (ret : ('t1 # 't2)) : ('t1 # 't2) result = - let (t, t0) = ret in Return (t, t0) -’ - -val id_mut_pair3_fwd_def = Define ‘ - (** [no_nested_borrows::id_mut_pair3]: forward function *) - id_mut_pair3_fwd (x : 't1) (y : 't2) : ('t1 # 't2) result = - Return (x, y) -’ - -val id_mut_pair3_back'a_def = Define ‘ - (** [no_nested_borrows::id_mut_pair3]: backward function 0 *) - id_mut_pair3_back'a (x : 't1) (y : 't2) (ret : 't1) : 't1 result = - Return ret -’ - -val id_mut_pair3_back'b_def = Define ‘ - (** [no_nested_borrows::id_mut_pair3]: backward function 1 *) - id_mut_pair3_back'b (x : 't1) (y : 't2) (ret : 't2) : 't2 result = - Return ret -’ - -val id_mut_pair4_fwd_def = Define ‘ - (** [no_nested_borrows::id_mut_pair4]: forward function *) - id_mut_pair4_fwd (p : ('t1 # 't2)) : ('t1 # 't2) result = - let (t, t0) = p in Return (t, t0) -’ - -val id_mut_pair4_back'a_def = Define ‘ - (** [no_nested_borrows::id_mut_pair4]: backward function 0 *) - id_mut_pair4_back'a (p : ('t1 # 't2)) (ret : 't1) : 't1 result = - Return ret -’ - -val id_mut_pair4_back'b_def = Define ‘ - (** [no_nested_borrows::id_mut_pair4]: backward function 1 *) - id_mut_pair4_back'b (p : ('t1 # 't2)) (ret : 't2) : 't2 result = - Return ret -’ - -Datatype: - (** [no_nested_borrows::StructWithTuple] *) - struct_with_tuple_t = <| struct_with_tuple_p : ('t1 # 't2); |> -End - -val new_tuple1_fwd_def = Define ‘ - (** [no_nested_borrows::new_tuple1]: forward function *) - new_tuple1_fwd : (u32, u32) struct_with_tuple_t result = - Return (<| struct_with_tuple_p := (int_to_u32 1, int_to_u32 2) |>) -’ - -val new_tuple2_fwd_def = Define ‘ - (** [no_nested_borrows::new_tuple2]: forward function *) - new_tuple2_fwd : (i16, i16) struct_with_tuple_t result = - Return (<| struct_with_tuple_p := (int_to_i16 1, int_to_i16 2) |>) -’ - -val new_tuple3_fwd_def = Define ‘ - (** [no_nested_borrows::new_tuple3]: forward function *) - new_tuple3_fwd : (u64, i64) struct_with_tuple_t result = - Return (<| struct_with_tuple_p := (int_to_u64 1, int_to_i64 2) |>) -’ - -Datatype: - (** [no_nested_borrows::StructWithPair] *) - struct_with_pair_t = <| struct_with_pair_p : ('t1, 't2) pair_t; |> -End - -val new_pair1_fwd_def = Define ‘ - (** [no_nested_borrows::new_pair1]: forward function *) - new_pair1_fwd : (u32, u32) struct_with_pair_t result = - Return - (<| - struct_with_pair_p := - (<| pair_x := (int_to_u32 1); pair_y := (int_to_u32 2) |>) - |>) -’ - -val test_constants_fwd_def = Define ‘ - (** [no_nested_borrows::test_constants]: forward function *) - test_constants_fwd : unit result = - do - swt <- new_tuple1_fwd; - let (i, _) = swt.struct_with_tuple_p in - if ~ (i = int_to_u32 1) - then Fail Failure - else ( - do - swt0 <- new_tuple2_fwd; - let (i0, _) = swt0.struct_with_tuple_p in - if ~ (i0 = int_to_i16 1) - then Fail Failure - else ( - do - swt1 <- new_tuple3_fwd; - let (i1, _) = swt1.struct_with_tuple_p in - if ~ (i1 = int_to_u64 1) - then Fail Failure - else ( - do - swp <- new_pair1_fwd; - if ~ (swp.struct_with_pair_p.pair_x = int_to_u32 1) - then Fail Failure - else Return () - od) - od) - od) - od -’ - -(** Unit test for [no_nested_borrows::test_constants] *) -val _ = assert_return (“test_constants_fwd”) - -val test_weird_borrows1_fwd_def = Define ‘ - (** [no_nested_borrows::test_weird_borrows1]: forward function *) - test_weird_borrows1_fwd : unit result = - Return () -’ - -(** Unit test for [no_nested_borrows::test_weird_borrows1] *) -val _ = assert_return (“test_weird_borrows1_fwd”) - -val test_mem_replace_fwd_back_def = Define ‘ - (** [no_nested_borrows::test_mem_replace]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - test_mem_replace_fwd_back (px : u32) : u32 result = - let y = mem_replace_fwd px (int_to_u32 1) in - if ~ (y = int_to_u32 0) then Fail Failure else Return (int_to_u32 2) -’ - -val test_shared_borrow_bool1_fwd_def = Define ‘ - (** [no_nested_borrows::test_shared_borrow_bool1]: forward function *) - test_shared_borrow_bool1_fwd (b : bool) : u32 result = - if b then Return (int_to_u32 0) else Return (int_to_u32 1) -’ - -val test_shared_borrow_bool2_fwd_def = Define ‘ - (** [no_nested_borrows::test_shared_borrow_bool2]: forward function *) - test_shared_borrow_bool2_fwd : u32 result = - Return (int_to_u32 0) -’ - -val test_shared_borrow_enum1_fwd_def = Define ‘ - (** [no_nested_borrows::test_shared_borrow_enum1]: forward function *) - test_shared_borrow_enum1_fwd (l : u32 list_t) : u32 result = - (case l of - | ListCons i l0 => Return (int_to_u32 1) - | ListNil => Return (int_to_u32 0)) -’ - -val test_shared_borrow_enum2_fwd_def = Define ‘ - (** [no_nested_borrows::test_shared_borrow_enum2]: forward function *) - test_shared_borrow_enum2_fwd : u32 result = - Return (int_to_u32 0) -’ - -val _ = export_theory () diff --git a/tests/hol4/misc-no_nested_borrows/noNestedBorrowsTheory.sig b/tests/hol4/misc-no_nested_borrows/noNestedBorrowsTheory.sig deleted file mode 100644 index 67368e38..00000000 --- a/tests/hol4/misc-no_nested_borrows/noNestedBorrowsTheory.sig +++ /dev/null @@ -1,1598 +0,0 @@ -signature noNestedBorrowsTheory = -sig - type thm = Thm.thm - - (* Definitions *) - val add_test_fwd_def : thm - val cast_test_fwd_def : thm - val choose_back_def : thm - val choose_fwd_def : thm - val choose_test_fwd_def : thm - val copy_int_fwd_def : thm - val div_test1_fwd_def : thm - val div_test_fwd_def : thm - val empty_enum_t_BIJ : thm - val empty_enum_t_CASE : thm - val empty_enum_t_TY_DEF : thm - val empty_enum_t_size_def : thm - val enum_t_BIJ : thm - val enum_t_CASE : thm - val enum_t_TY_DEF : thm - val enum_t_size_def : thm - val get_max_fwd_def : thm - val id_mut_pair1_back_def : thm - val id_mut_pair1_fwd_def : thm - val id_mut_pair2_back_def : thm - val id_mut_pair2_fwd_def : thm - val id_mut_pair3_back'a_def : thm - val id_mut_pair3_back'b_def : thm - val id_mut_pair3_fwd_def : thm - val id_mut_pair4_back'a_def : thm - val id_mut_pair4_back'b_def : thm - val id_mut_pair4_fwd_def : thm - val is_cons_fwd_def : thm - val list_length_fwd_def : thm - val list_nth_mut_back_def : thm - val list_nth_mut_fwd_def : thm - val list_nth_shared_fwd_def : thm - val list_rev_aux_fwd_def : thm - val list_rev_fwd_back_def : thm - val list_t_TY_DEF : thm - val list_t_case_def : thm - val list_t_size_def : thm - val neg_test_fwd_def : thm - val new_pair1_fwd_def : thm - val new_tuple1_fwd_def : thm - val new_tuple2_fwd_def : thm - val new_tuple3_fwd_def : thm - val node_elem_t_TY_DEF : thm - val node_elem_t_case_def : thm - val one_t_TY_DEF : thm - val one_t_case_def : thm - val one_t_size_def : thm - val pair_t_TY_DEF : thm - val pair_t_case_def : thm - val pair_t_pair_x : thm - val pair_t_pair_x_fupd : thm - val pair_t_pair_y : thm - val pair_t_pair_y_fupd : thm - val pair_t_size_def : thm - val refs_test1_fwd_def : thm - val refs_test2_fwd_def : thm - val rem_test_fwd_def : thm - val split_list_fwd_def : thm - val struct_with_pair_t_TY_DEF : thm - val struct_with_pair_t_case_def : thm - val struct_with_pair_t_size_def : thm - val struct_with_pair_t_struct_with_pair_p : thm - val struct_with_pair_t_struct_with_pair_p_fupd : thm - val struct_with_tuple_t_TY_DEF : thm - val struct_with_tuple_t_case_def : thm - val struct_with_tuple_t_size_def : thm - val struct_with_tuple_t_struct_with_tuple_p : thm - val struct_with_tuple_t_struct_with_tuple_p_fupd : thm - val subs_test_fwd_def : thm - val sum_t_TY_DEF : thm - val sum_t_case_def : thm - val sum_t_size_def : thm - val test2_fwd_def : thm - val test3_fwd_def : thm - val test_box1_fwd_def : thm - val test_char_fwd_def : thm - val test_constants_fwd_def : thm - val test_copy_int_fwd_def : thm - val test_is_cons_fwd_def : thm - val test_list1_fwd_def : thm - val test_list_functions_fwd_def : thm - val test_mem_replace_fwd_back_def : thm - val test_neg1_fwd_def : thm - val test_panic_fwd_def : thm - val test_shared_borrow_bool1_fwd_def : thm - val test_shared_borrow_bool2_fwd_def : thm - val test_shared_borrow_enum1_fwd_def : thm - val test_shared_borrow_enum2_fwd_def : thm - val test_split_list_fwd_def : thm - val test_unreachable_fwd_def : thm - val test_weird_borrows1_fwd_def : thm - val tree_t_TY_DEF : thm - val tree_t_case_def : thm - val tree_t_size_def : thm - - (* Theorems *) - val EXISTS_pair_t : thm - val EXISTS_struct_with_pair_t : thm - val EXISTS_struct_with_tuple_t : thm - val FORALL_pair_t : thm - val FORALL_struct_with_pair_t : thm - val FORALL_struct_with_tuple_t : thm - val datatype_empty_enum_t : thm - val datatype_enum_t : thm - val datatype_list_t : thm - val datatype_one_t : thm - val datatype_pair_t : thm - val datatype_struct_with_pair_t : thm - val datatype_struct_with_tuple_t : thm - val datatype_sum_t : thm - val datatype_tree_t : thm - val empty_enum_t2num_11 : thm - val empty_enum_t2num_ONTO : thm - val empty_enum_t2num_num2empty_enum_t : thm - val empty_enum_t2num_thm : thm - val empty_enum_t_Axiom : thm - val empty_enum_t_EQ_empty_enum_t : thm - val empty_enum_t_case_cong : thm - val empty_enum_t_case_def : thm - val empty_enum_t_case_eq : thm - val empty_enum_t_induction : thm - val empty_enum_t_nchotomy : thm - val enum_t2num_11 : thm - val enum_t2num_ONTO : thm - val enum_t2num_num2enum_t : thm - val enum_t2num_thm : thm - val enum_t_Axiom : thm - val enum_t_EQ_enum_t : thm - val enum_t_case_cong : thm - val enum_t_case_def : thm - val enum_t_case_eq : thm - val enum_t_distinct : thm - val enum_t_induction : thm - val enum_t_nchotomy : thm - val list_t_11 : thm - val list_t_Axiom : thm - val list_t_case_cong : thm - val list_t_case_eq : thm - val list_t_distinct : thm - val list_t_induction : thm - val list_t_nchotomy : thm - val node_elem_t_11 : thm - val node_elem_t_Axiom : thm - val node_elem_t_case_cong : thm - val node_elem_t_case_eq : thm - val node_elem_t_distinct : thm - val node_elem_t_induction : thm - val node_elem_t_nchotomy : thm - val num2empty_enum_t_11 : thm - val num2empty_enum_t_ONTO : thm - val num2empty_enum_t_empty_enum_t2num : thm - val num2empty_enum_t_thm : thm - val num2enum_t_11 : thm - val num2enum_t_ONTO : thm - val num2enum_t_enum_t2num : thm - val num2enum_t_thm : thm - val one_t_11 : thm - val one_t_Axiom : thm - val one_t_case_cong : thm - val one_t_case_eq : thm - val one_t_induction : thm - val one_t_nchotomy : thm - val pair_t_11 : thm - val pair_t_Axiom : thm - val pair_t_accessors : thm - val pair_t_accfupds : thm - val pair_t_case_cong : thm - val pair_t_case_eq : thm - val pair_t_component_equality : thm - val pair_t_fn_updates : thm - val pair_t_fupdcanon : thm - val pair_t_fupdcanon_comp : thm - val pair_t_fupdfupds : thm - val pair_t_fupdfupds_comp : thm - val pair_t_induction : thm - val pair_t_literal_11 : thm - val pair_t_literal_nchotomy : thm - val pair_t_nchotomy : thm - val pair_t_updates_eq_literal : thm - val struct_with_pair_t_11 : thm - val struct_with_pair_t_Axiom : thm - val struct_with_pair_t_accessors : thm - val struct_with_pair_t_accfupds : thm - val struct_with_pair_t_case_cong : thm - val struct_with_pair_t_case_eq : thm - val struct_with_pair_t_component_equality : thm - val struct_with_pair_t_fn_updates : thm - val struct_with_pair_t_fupdfupds : thm - val struct_with_pair_t_fupdfupds_comp : thm - val struct_with_pair_t_induction : thm - val struct_with_pair_t_literal_11 : thm - val struct_with_pair_t_literal_nchotomy : thm - val struct_with_pair_t_nchotomy : thm - val struct_with_pair_t_updates_eq_literal : thm - val struct_with_tuple_t_11 : thm - val struct_with_tuple_t_Axiom : thm - val struct_with_tuple_t_accessors : thm - val struct_with_tuple_t_accfupds : thm - val struct_with_tuple_t_case_cong : thm - val struct_with_tuple_t_case_eq : thm - val struct_with_tuple_t_component_equality : thm - val struct_with_tuple_t_fn_updates : thm - val struct_with_tuple_t_fupdfupds : thm - val struct_with_tuple_t_fupdfupds_comp : thm - val struct_with_tuple_t_induction : thm - val struct_with_tuple_t_literal_11 : thm - val struct_with_tuple_t_literal_nchotomy : thm - val struct_with_tuple_t_nchotomy : thm - val struct_with_tuple_t_updates_eq_literal : thm - val sum_t_11 : thm - val sum_t_Axiom : thm - val sum_t_case_cong : thm - val sum_t_case_eq : thm - val sum_t_distinct : thm - val sum_t_induction : thm - val sum_t_nchotomy : thm - val tree_t_11 : thm - val tree_t_Axiom : thm - val tree_t_case_cong : thm - val tree_t_case_eq : thm - val tree_t_distinct : thm - val tree_t_induction : thm - val tree_t_nchotomy : thm - - val noNestedBorrows_grammars : type_grammar.grammar * term_grammar.grammar -(* - [divDef] Parent theory of "noNestedBorrows" - - [add_test_fwd_def] Definition - - ⊢ ∀x y. add_test_fwd x y = u32_add x y - - [cast_test_fwd_def] Definition - - ⊢ ∀x. cast_test_fwd x = mk_i32 (u32_to_int x) - - [choose_back_def] Definition - - ⊢ ∀b x y ret. - choose_back b x y ret = - if b then Return (ret,y) else Return (x,ret) - - [choose_fwd_def] Definition - - ⊢ ∀b x y. choose_fwd b x y = if b then Return x else Return y - - [choose_test_fwd_def] Definition - - ⊢ choose_test_fwd = - do - z <- choose_fwd T (int_to_i32 0) (int_to_i32 0); - z0 <- i32_add z (int_to_i32 1); - if z0 ≠ int_to_i32 1 then Fail Failure - else - do - (x,y) <- choose_back T (int_to_i32 0) (int_to_i32 0) z0; - if x ≠ int_to_i32 1 then Fail Failure - else if y ≠ int_to_i32 0 then Fail Failure - else Return () - od - od - - [copy_int_fwd_def] Definition - - ⊢ ∀x. copy_int_fwd x = Return x - - [div_test1_fwd_def] Definition - - ⊢ ∀x. div_test1_fwd x = u32_div x (int_to_u32 2) - - [div_test_fwd_def] Definition - - ⊢ ∀x y. div_test_fwd x y = u32_div x y - - [empty_enum_t_BIJ] Definition - - ⊢ (∀a. num2empty_enum_t (empty_enum_t2num a) = a) ∧ - ∀r. (λn. n < 1) r ⇔ empty_enum_t2num (num2empty_enum_t r) = r - - [empty_enum_t_CASE] Definition - - ⊢ ∀x v0. - (case x of EmptyEnumEmpty => v0) = (λm. v0) (empty_enum_t2num x) - - [empty_enum_t_TY_DEF] Definition - - ⊢ ∃rep. TYPE_DEFINITION (λn. n < 1) rep - - [empty_enum_t_size_def] Definition - - ⊢ ∀x. empty_enum_t_size x = 0 - - [enum_t_BIJ] Definition - - ⊢ (∀a. num2enum_t (enum_t2num a) = a) ∧ - ∀r. (λn. n < 2) r ⇔ enum_t2num (num2enum_t r) = r - - [enum_t_CASE] Definition - - ⊢ ∀x v0 v1. - (case x of EnumVariant1 => v0 | EnumVariant2 => v1) = - (λm. if m = 0 then v0 else v1) (enum_t2num x) - - [enum_t_TY_DEF] Definition - - ⊢ ∃rep. TYPE_DEFINITION (λn. n < 2) rep - - [enum_t_size_def] Definition - - ⊢ ∀x. enum_t_size x = 0 - - [get_max_fwd_def] Definition - - ⊢ ∀x y. get_max_fwd x y = if u32_ge x y then Return x else Return y - - [id_mut_pair1_back_def] Definition - - ⊢ ∀x y ret. - id_mut_pair1_back x y ret = (let (t,t0) = ret in Return (t,t0)) - - [id_mut_pair1_fwd_def] Definition - - ⊢ ∀x y. id_mut_pair1_fwd x y = Return (x,y) - - [id_mut_pair2_back_def] Definition - - ⊢ ∀p ret. - id_mut_pair2_back p ret = (let (t,t0) = ret in Return (t,t0)) - - [id_mut_pair2_fwd_def] Definition - - ⊢ ∀p. id_mut_pair2_fwd p = (let (t,t0) = p in Return (t,t0)) - - [id_mut_pair3_back'a_def] Definition - - ⊢ ∀x y ret. id_mut_pair3_back'a x y ret = Return ret - - [id_mut_pair3_back'b_def] Definition - - ⊢ ∀x y ret. id_mut_pair3_back'b x y ret = Return ret - - [id_mut_pair3_fwd_def] Definition - - ⊢ ∀x y. id_mut_pair3_fwd x y = Return (x,y) - - [id_mut_pair4_back'a_def] Definition - - ⊢ ∀p ret. id_mut_pair4_back'a p ret = Return ret - - [id_mut_pair4_back'b_def] Definition - - ⊢ ∀p ret. id_mut_pair4_back'b p ret = Return ret - - [id_mut_pair4_fwd_def] Definition - - ⊢ ∀p. id_mut_pair4_fwd p = (let (t,t0) = p in Return (t,t0)) - - [is_cons_fwd_def] Definition - - ⊢ ∀l. is_cons_fwd l = - case l of ListCons t l0 => Return T | ListNil => Return F - - [list_length_fwd_def] Definition - - ⊢ ∀l. list_length_fwd l = - case l of - ListCons t l1 => - do i <- list_length_fwd l1; u32_add (int_to_u32 1) i od - | ListNil => Return (int_to_u32 0) - - [list_nth_mut_back_def] Definition - - ⊢ ∀l i ret. - list_nth_mut_back l i ret = - case l of - ListCons x tl => - if i = int_to_u32 0 then Return (ListCons ret tl) - else - do - i0 <- u32_sub i (int_to_u32 1); - tl0 <- list_nth_mut_back tl i0 ret; - Return (ListCons x tl0) - od - | ListNil => Fail Failure - - [list_nth_mut_fwd_def] Definition - - ⊢ ∀l i. - list_nth_mut_fwd l i = - case l of - ListCons x tl => - if i = int_to_u32 0 then Return x - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_fwd tl i0 - od - | ListNil => Fail Failure - - [list_nth_shared_fwd_def] Definition - - ⊢ ∀l i. - list_nth_shared_fwd l i = - case l of - ListCons x tl => - if i = int_to_u32 0 then Return x - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_shared_fwd tl i0 - od - | ListNil => Fail Failure - - [list_rev_aux_fwd_def] Definition - - ⊢ ∀li lo. - list_rev_aux_fwd li lo = - case li of - ListCons hd tl => list_rev_aux_fwd tl (ListCons hd lo) - | ListNil => Return lo - - [list_rev_fwd_back_def] Definition - - ⊢ ∀l. list_rev_fwd_back l = - (let - li = mem_replace_fwd l ListNil - in - list_rev_aux_fwd li ListNil) - - [list_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('list_t'). - (∀a0'. - (∃a0 a1. - a0' = - (λa0 a1. - ind_type$CONSTR 0 a0 - (ind_type$FCONS a1 (λn. ind_type$BOTTOM))) - a0 a1 ∧ $var$('list_t') a1) ∨ - a0' = - ind_type$CONSTR (SUC 0) ARB (λn. ind_type$BOTTOM) ⇒ - $var$('list_t') a0') ⇒ - $var$('list_t') a0') rep - - [list_t_case_def] Definition - - ⊢ (∀a0 a1 f v. list_t_CASE (ListCons a0 a1) f v = f a0 a1) ∧ - ∀f v. list_t_CASE ListNil f v = v - - [list_t_size_def] Definition - - ⊢ (∀f a0 a1. - list_t_size f (ListCons a0 a1) = 1 + (f a0 + list_t_size f a1)) ∧ - ∀f. list_t_size f ListNil = 0 - - [neg_test_fwd_def] Definition - - ⊢ ∀x. neg_test_fwd x = i32_neg x - - [new_pair1_fwd_def] Definition - - ⊢ new_pair1_fwd = - Return - <|struct_with_pair_p := - <|pair_x := int_to_u32 1; pair_y := int_to_u32 2|> |> - - [new_tuple1_fwd_def] Definition - - ⊢ new_tuple1_fwd = - Return <|struct_with_tuple_p := (int_to_u32 1,int_to_u32 2)|> - - [new_tuple2_fwd_def] Definition - - ⊢ new_tuple2_fwd = - Return <|struct_with_tuple_p := (int_to_i16 1,int_to_i16 2)|> - - [new_tuple3_fwd_def] Definition - - ⊢ new_tuple3_fwd = - Return <|struct_with_tuple_p := (int_to_u64 1,int_to_i64 2)|> - - [node_elem_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa1'. - ∀ $var$('tree_t') $var$('node_elem_t'). - (∀a0'. - (∃a. a0' = - (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) - a) ∨ - (∃a0 a1 a2. - a0' = - (λa0 a1 a2. - ind_type$CONSTR (SUC 0) a0 - (ind_type$FCONS a1 - (ind_type$FCONS a2 (λn. ind_type$BOTTOM)))) - a0 a1 a2 ∧ $var$('node_elem_t') a1 ∧ - $var$('tree_t') a2) ⇒ - $var$('tree_t') a0') ∧ - (∀a1'. - (∃a0 a1. - a1' = - (λa0 a1. - ind_type$CONSTR (SUC (SUC 0)) ARB - (ind_type$FCONS a0 - (ind_type$FCONS a1 (λn. ind_type$BOTTOM)))) - a0 a1 ∧ $var$('tree_t') a0 ∧ - $var$('node_elem_t') a1) ∨ - a1' = - ind_type$CONSTR (SUC (SUC (SUC 0))) ARB - (λn. ind_type$BOTTOM) ⇒ - $var$('node_elem_t') a1') ⇒ - $var$('node_elem_t') a1') rep - - [node_elem_t_case_def] Definition - - ⊢ (∀a0 a1 f v. node_elem_t_CASE (NodeElemCons a0 a1) f v = f a0 a1) ∧ - ∀f v. node_elem_t_CASE NodeElemNil f v = v - - [one_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0. - ∀ $var$('one_t'). - (∀a0. - (∃a. a0 = - (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) - a) ⇒ - $var$('one_t') a0) ⇒ - $var$('one_t') a0) rep - - [one_t_case_def] Definition - - ⊢ ∀a f. one_t_CASE (OneOne a) f = f a - - [one_t_size_def] Definition - - ⊢ ∀f a. one_t_size f (OneOne a) = 1 + f a - - [pair_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('pair_t'). - (∀a0'. - (∃a0 a1. - a0' = - (λa0 a1. - ind_type$CONSTR 0 (a0,a1) - (λn. ind_type$BOTTOM)) a0 a1) ⇒ - $var$('pair_t') a0') ⇒ - $var$('pair_t') a0') rep - - [pair_t_case_def] Definition - - ⊢ ∀a0 a1 f. pair_t_CASE (pair_t a0 a1) f = f a0 a1 - - [pair_t_pair_x] Definition - - ⊢ ∀t t0. (pair_t t t0).pair_x = t - - [pair_t_pair_x_fupd] Definition - - ⊢ ∀f t t0. pair_t t t0 with pair_x updated_by f = pair_t (f t) t0 - - [pair_t_pair_y] Definition - - ⊢ ∀t t0. (pair_t t t0).pair_y = t0 - - [pair_t_pair_y_fupd] Definition - - ⊢ ∀f t t0. pair_t t t0 with pair_y updated_by f = pair_t t (f t0) - - [pair_t_size_def] Definition - - ⊢ ∀f f1 a0 a1. pair_t_size f f1 (pair_t a0 a1) = 1 + (f a0 + f1 a1) - - [refs_test1_fwd_def] Definition - - ⊢ refs_test1_fwd = - if int_to_i32 1 ≠ int_to_i32 1 then Fail Failure else Return () - - [refs_test2_fwd_def] Definition - - ⊢ refs_test2_fwd = - if int_to_i32 2 ≠ int_to_i32 2 then Fail Failure - else if int_to_i32 0 ≠ int_to_i32 0 then Fail Failure - else if int_to_i32 2 ≠ int_to_i32 2 then Fail Failure - else if int_to_i32 2 ≠ int_to_i32 2 then Fail Failure - else Return () - - [rem_test_fwd_def] Definition - - ⊢ ∀x y. rem_test_fwd x y = u32_rem x y - - [split_list_fwd_def] Definition - - ⊢ ∀l. split_list_fwd l = - case l of - ListCons hd tl => Return (hd,tl) - | ListNil => Fail Failure - - [struct_with_pair_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0. - ∀ $var$('struct_with_pair_t'). - (∀a0. - (∃a. a0 = - (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) - a) ⇒ - $var$('struct_with_pair_t') a0) ⇒ - $var$('struct_with_pair_t') a0) rep - - [struct_with_pair_t_case_def] Definition - - ⊢ ∀a f. struct_with_pair_t_CASE (struct_with_pair_t a) f = f a - - [struct_with_pair_t_size_def] Definition - - ⊢ ∀f f1 a. - struct_with_pair_t_size f f1 (struct_with_pair_t a) = - 1 + pair_t_size f f1 a - - [struct_with_pair_t_struct_with_pair_p] Definition - - ⊢ ∀p. (struct_with_pair_t p).struct_with_pair_p = p - - [struct_with_pair_t_struct_with_pair_p_fupd] Definition - - ⊢ ∀f p. - struct_with_pair_t p with struct_with_pair_p updated_by f = - struct_with_pair_t (f p) - - [struct_with_tuple_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0. - ∀ $var$('struct_with_tuple_t'). - (∀a0. - (∃a. a0 = - (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) - a) ⇒ - $var$('struct_with_tuple_t') a0) ⇒ - $var$('struct_with_tuple_t') a0) rep - - [struct_with_tuple_t_case_def] Definition - - ⊢ ∀a f. struct_with_tuple_t_CASE (struct_with_tuple_t a) f = f a - - [struct_with_tuple_t_size_def] Definition - - ⊢ ∀f f1 a. - struct_with_tuple_t_size f f1 (struct_with_tuple_t a) = - 1 + pair_size f f1 a - - [struct_with_tuple_t_struct_with_tuple_p] Definition - - ⊢ ∀p. (struct_with_tuple_t p).struct_with_tuple_p = p - - [struct_with_tuple_t_struct_with_tuple_p_fupd] Definition - - ⊢ ∀f p. - struct_with_tuple_t p with struct_with_tuple_p updated_by f = - struct_with_tuple_t (f p) - - [subs_test_fwd_def] Definition - - ⊢ ∀x y. subs_test_fwd x y = u32_sub x y - - [sum_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0. - ∀ $var$('sum_t'). - (∀a0. - (∃a. a0 = - (λa. - ind_type$CONSTR 0 (a,ARB) - (λn. ind_type$BOTTOM)) a) ∨ - (∃a. a0 = - (λa. - ind_type$CONSTR (SUC 0) (ARB,a) - (λn. ind_type$BOTTOM)) a) ⇒ - $var$('sum_t') a0) ⇒ - $var$('sum_t') a0) rep - - [sum_t_case_def] Definition - - ⊢ (∀a f f1. sum_t_CASE (SumLeft a) f f1 = f a) ∧ - ∀a f f1. sum_t_CASE (SumRight a) f f1 = f1 a - - [sum_t_size_def] Definition - - ⊢ (∀f f1 a. sum_t_size f f1 (SumLeft a) = 1 + f a) ∧ - ∀f f1 a. sum_t_size f f1 (SumRight a) = 1 + f1 a - - [test2_fwd_def] Definition - - ⊢ test2_fwd = - monad_ignore_bind (u32_add (int_to_u32 23) (int_to_u32 44)) - (Return ()) - - [test3_fwd_def] Definition - - ⊢ test3_fwd = - do - x <- get_max_fwd (int_to_u32 4) (int_to_u32 3); - y <- get_max_fwd (int_to_u32 10) (int_to_u32 11); - z <- u32_add x y; - if z ≠ int_to_u32 15 then Fail Failure else Return () - od - - [test_box1_fwd_def] Definition - - ⊢ test_box1_fwd = - (let - b = int_to_i32 1; - x = b - in - if x ≠ int_to_i32 1 then Fail Failure else Return ()) - - [test_char_fwd_def] Definition - - ⊢ test_char_fwd = Return #"a" - - [test_constants_fwd_def] Definition - - ⊢ test_constants_fwd = - do - swt <- new_tuple1_fwd; - (i,_) <<- swt.struct_with_tuple_p; - if i ≠ int_to_u32 1 then Fail Failure - else - do - swt0 <- new_tuple2_fwd; - (i0,_) <<- swt0.struct_with_tuple_p; - if i0 ≠ int_to_i16 1 then Fail Failure - else - do - swt1 <- new_tuple3_fwd; - (i1,_) <<- swt1.struct_with_tuple_p; - if i1 ≠ int_to_u64 1 then Fail Failure - else - do - swp <- new_pair1_fwd; - if swp.struct_with_pair_p.pair_x ≠ int_to_u32 1 then - Fail Failure - else Return () - od - od - od - od - - [test_copy_int_fwd_def] Definition - - ⊢ test_copy_int_fwd = - do - y <- copy_int_fwd (int_to_i32 0); - if int_to_i32 0 ≠ y then Fail Failure else Return () - od - - [test_is_cons_fwd_def] Definition - - ⊢ test_is_cons_fwd = - (let - l = ListNil - in - do - b <- is_cons_fwd (ListCons (int_to_i32 0) l); - if ¬b then Fail Failure else Return () - od) - - [test_list1_fwd_def] Definition - - ⊢ test_list1_fwd = Return () - - [test_list_functions_fwd_def] Definition - - ⊢ test_list_functions_fwd = - (let - l = ListNil; - l0 = ListCons (int_to_i32 2) l; - l1 = ListCons (int_to_i32 1) l0 - in - do - i <- list_length_fwd (ListCons (int_to_i32 0) l1); - if i ≠ int_to_u32 3 then Fail Failure - else - do - i0 <- - list_nth_shared_fwd (ListCons (int_to_i32 0) l1) - (int_to_u32 0); - if i0 ≠ int_to_i32 0 then Fail Failure - else - do - i1 <- - list_nth_shared_fwd (ListCons (int_to_i32 0) l1) - (int_to_u32 1); - if i1 ≠ int_to_i32 1 then Fail Failure - else - do - i2 <- - list_nth_shared_fwd (ListCons (int_to_i32 0) l1) - (int_to_u32 2); - if i2 ≠ int_to_i32 2 then Fail Failure - else - do - ls <- - list_nth_mut_back - (ListCons (int_to_i32 0) l1) - (int_to_u32 1) (int_to_i32 3); - i3 <- list_nth_shared_fwd ls (int_to_u32 0); - if i3 ≠ int_to_i32 0 then Fail Failure - else - do - i4 <- - list_nth_shared_fwd ls (int_to_u32 1); - if i4 ≠ int_to_i32 3 then Fail Failure - else - do - i5 <- - list_nth_shared_fwd ls - (int_to_u32 2); - if i5 ≠ int_to_i32 2 then Fail Failure - else Return () - od - od - od - od - od - od - od) - - [test_mem_replace_fwd_back_def] Definition - - ⊢ ∀px. - test_mem_replace_fwd_back px = - (let - y = mem_replace_fwd px (int_to_u32 1) - in - if y ≠ int_to_u32 0 then Fail Failure - else Return (int_to_u32 2)) - - [test_neg1_fwd_def] Definition - - ⊢ test_neg1_fwd = - do - y <- i32_neg (int_to_i32 3); - if y ≠ int_to_i32 (-3) then Fail Failure else Return () - od - - [test_panic_fwd_def] Definition - - ⊢ ∀b. test_panic_fwd b = if b then Fail Failure else Return () - - [test_shared_borrow_bool1_fwd_def] Definition - - ⊢ ∀b. test_shared_borrow_bool1_fwd b = - if b then Return (int_to_u32 0) else Return (int_to_u32 1) - - [test_shared_borrow_bool2_fwd_def] Definition - - ⊢ test_shared_borrow_bool2_fwd = Return (int_to_u32 0) - - [test_shared_borrow_enum1_fwd_def] Definition - - ⊢ ∀l. test_shared_borrow_enum1_fwd l = - case l of - ListCons i l0 => Return (int_to_u32 1) - | ListNil => Return (int_to_u32 0) - - [test_shared_borrow_enum2_fwd_def] Definition - - ⊢ test_shared_borrow_enum2_fwd = Return (int_to_u32 0) - - [test_split_list_fwd_def] Definition - - ⊢ test_split_list_fwd = - (let - l = ListNil - in - do - p <- split_list_fwd (ListCons (int_to_i32 0) l); - (hd,_) <<- p; - if hd ≠ int_to_i32 0 then Fail Failure else Return () - od) - - [test_unreachable_fwd_def] Definition - - ⊢ ∀b. test_unreachable_fwd b = if b then Fail Failure else Return () - - [test_weird_borrows1_fwd_def] Definition - - ⊢ test_weird_borrows1_fwd = Return () - - [tree_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('tree_t') $var$('node_elem_t'). - (∀a0'. - (∃a. a0' = - (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) - a) ∨ - (∃a0 a1 a2. - a0' = - (λa0 a1 a2. - ind_type$CONSTR (SUC 0) a0 - (ind_type$FCONS a1 - (ind_type$FCONS a2 (λn. ind_type$BOTTOM)))) - a0 a1 a2 ∧ $var$('node_elem_t') a1 ∧ - $var$('tree_t') a2) ⇒ - $var$('tree_t') a0') ∧ - (∀a1'. - (∃a0 a1. - a1' = - (λa0 a1. - ind_type$CONSTR (SUC (SUC 0)) ARB - (ind_type$FCONS a0 - (ind_type$FCONS a1 (λn. ind_type$BOTTOM)))) - a0 a1 ∧ $var$('tree_t') a0 ∧ - $var$('node_elem_t') a1) ∨ - a1' = - ind_type$CONSTR (SUC (SUC (SUC 0))) ARB - (λn. ind_type$BOTTOM) ⇒ - $var$('node_elem_t') a1') ⇒ - $var$('tree_t') a0') rep - - [tree_t_case_def] Definition - - ⊢ (∀a f f1. tree_t_CASE (TreeLeaf a) f f1 = f a) ∧ - ∀a0 a1 a2 f f1. tree_t_CASE (TreeNode a0 a1 a2) f f1 = f1 a0 a1 a2 - - [tree_t_size_def] Definition - - ⊢ (∀f a. tree_t_size f (TreeLeaf a) = 1 + f a) ∧ - (∀f a0 a1 a2. - tree_t_size f (TreeNode a0 a1 a2) = - 1 + (f a0 + (node_elem_t_size f a1 + tree_t_size f a2))) ∧ - (∀f a0 a1. - node_elem_t_size f (NodeElemCons a0 a1) = - 1 + (tree_t_size f a0 + node_elem_t_size f a1)) ∧ - ∀f. node_elem_t_size f NodeElemNil = 0 - - [EXISTS_pair_t] Theorem - - ⊢ ∀P. (∃p. P p) ⇔ ∃t0 t. P <|pair_x := t0; pair_y := t|> - - [EXISTS_struct_with_pair_t] Theorem - - ⊢ ∀P. (∃s. P s) ⇔ ∃p. P <|struct_with_pair_p := p|> - - [EXISTS_struct_with_tuple_t] Theorem - - ⊢ ∀P. (∃s. P s) ⇔ ∃p. P <|struct_with_tuple_p := p|> - - [FORALL_pair_t] Theorem - - ⊢ ∀P. (∀p. P p) ⇔ ∀t0 t. P <|pair_x := t0; pair_y := t|> - - [FORALL_struct_with_pair_t] Theorem - - ⊢ ∀P. (∀s. P s) ⇔ ∀p. P <|struct_with_pair_p := p|> - - [FORALL_struct_with_tuple_t] Theorem - - ⊢ ∀P. (∀s. P s) ⇔ ∀p. P <|struct_with_tuple_p := p|> - - [datatype_empty_enum_t] Theorem - - ⊢ DATATYPE (empty_enum_t EmptyEnumEmpty) - - [datatype_enum_t] Theorem - - ⊢ DATATYPE (enum_t EnumVariant1 EnumVariant2) - - [datatype_list_t] Theorem - - ⊢ DATATYPE (list_t ListCons ListNil) - - [datatype_one_t] Theorem - - ⊢ DATATYPE (one_t OneOne) - - [datatype_pair_t] Theorem - - ⊢ DATATYPE (record pair_t pair_x pair_y) - - [datatype_struct_with_pair_t] Theorem - - ⊢ DATATYPE (record struct_with_pair_t struct_with_pair_p) - - [datatype_struct_with_tuple_t] Theorem - - ⊢ DATATYPE (record struct_with_tuple_t struct_with_tuple_p) - - [datatype_sum_t] Theorem - - ⊢ DATATYPE (sum_t SumLeft SumRight) - - [datatype_tree_t] Theorem - - ⊢ DATATYPE - (tree_t TreeLeaf TreeNode ∧ node_elem_t NodeElemCons NodeElemNil) - - [empty_enum_t2num_11] Theorem - - ⊢ ∀a a'. empty_enum_t2num a = empty_enum_t2num a' ⇔ a = a' - - [empty_enum_t2num_ONTO] Theorem - - ⊢ ∀r. r < 1 ⇔ ∃a. r = empty_enum_t2num a - - [empty_enum_t2num_num2empty_enum_t] Theorem - - ⊢ ∀r. r < 1 ⇔ empty_enum_t2num (num2empty_enum_t r) = r - - [empty_enum_t2num_thm] Theorem - - ⊢ empty_enum_t2num EmptyEnumEmpty = 0 - - [empty_enum_t_Axiom] Theorem - - ⊢ ∀x0. ∃f. f EmptyEnumEmpty = x0 - - [empty_enum_t_EQ_empty_enum_t] Theorem - - ⊢ ∀a a'. a = a' ⇔ empty_enum_t2num a = empty_enum_t2num a' - - [empty_enum_t_case_cong] Theorem - - ⊢ ∀M M' v0. - M = M' ∧ (M' = EmptyEnumEmpty ⇒ v0 = v0') ⇒ - (case M of EmptyEnumEmpty => v0) = - case M' of EmptyEnumEmpty => v0' - - [empty_enum_t_case_def] Theorem - - ⊢ ∀v0. (case EmptyEnumEmpty of EmptyEnumEmpty => v0) = v0 - - [empty_enum_t_case_eq] Theorem - - ⊢ (case x of EmptyEnumEmpty => v0) = v ⇔ x = EmptyEnumEmpty ∧ v0 = v - - [empty_enum_t_induction] Theorem - - ⊢ ∀P. P EmptyEnumEmpty ⇒ ∀a. P a - - [empty_enum_t_nchotomy] Theorem - - ⊢ ∀a. a = EmptyEnumEmpty - - [enum_t2num_11] Theorem - - ⊢ ∀a a'. enum_t2num a = enum_t2num a' ⇔ a = a' - - [enum_t2num_ONTO] Theorem - - ⊢ ∀r. r < 2 ⇔ ∃a. r = enum_t2num a - - [enum_t2num_num2enum_t] Theorem - - ⊢ ∀r. r < 2 ⇔ enum_t2num (num2enum_t r) = r - - [enum_t2num_thm] Theorem - - ⊢ enum_t2num EnumVariant1 = 0 ∧ enum_t2num EnumVariant2 = 1 - - [enum_t_Axiom] Theorem - - ⊢ ∀x0 x1. ∃f. f EnumVariant1 = x0 ∧ f EnumVariant2 = x1 - - [enum_t_EQ_enum_t] Theorem - - ⊢ ∀a a'. a = a' ⇔ enum_t2num a = enum_t2num a' - - [enum_t_case_cong] Theorem - - ⊢ ∀M M' v0 v1. - M = M' ∧ (M' = EnumVariant1 ⇒ v0 = v0') ∧ - (M' = EnumVariant2 ⇒ v1 = v1') ⇒ - (case M of EnumVariant1 => v0 | EnumVariant2 => v1) = - case M' of EnumVariant1 => v0' | EnumVariant2 => v1' - - [enum_t_case_def] Theorem - - ⊢ (∀v0 v1. - (case EnumVariant1 of EnumVariant1 => v0 | EnumVariant2 => v1) = - v0) ∧ - ∀v0 v1. - (case EnumVariant2 of EnumVariant1 => v0 | EnumVariant2 => v1) = - v1 - - [enum_t_case_eq] Theorem - - ⊢ (case x of EnumVariant1 => v0 | EnumVariant2 => v1) = v ⇔ - x = EnumVariant1 ∧ v0 = v ∨ x = EnumVariant2 ∧ v1 = v - - [enum_t_distinct] Theorem - - ⊢ EnumVariant1 ≠ EnumVariant2 - - [enum_t_induction] Theorem - - ⊢ ∀P. P EnumVariant1 ∧ P EnumVariant2 ⇒ ∀a. P a - - [enum_t_nchotomy] Theorem - - ⊢ ∀a. a = EnumVariant1 ∨ a = EnumVariant2 - - [list_t_11] Theorem - - ⊢ ∀a0 a1 a0' a1'. - ListCons a0 a1 = ListCons a0' a1' ⇔ a0 = a0' ∧ a1 = a1' - - [list_t_Axiom] Theorem - - ⊢ ∀f0 f1. ∃fn. - (∀a0 a1. fn (ListCons a0 a1) = f0 a0 a1 (fn a1)) ∧ - fn ListNil = f1 - - [list_t_case_cong] Theorem - - ⊢ ∀M M' f v. - M = M' ∧ (∀a0 a1. M' = ListCons a0 a1 ⇒ f a0 a1 = f' a0 a1) ∧ - (M' = ListNil ⇒ v = v') ⇒ - list_t_CASE M f v = list_t_CASE M' f' v' - - [list_t_case_eq] Theorem - - ⊢ list_t_CASE x f v = v' ⇔ - (∃t l. x = ListCons t l ∧ f t l = v') ∨ x = ListNil ∧ v = v' - - [list_t_distinct] Theorem - - ⊢ ∀a1 a0. ListCons a0 a1 ≠ ListNil - - [list_t_induction] Theorem - - ⊢ ∀P. (∀l. P l ⇒ ∀t. P (ListCons t l)) ∧ P ListNil ⇒ ∀l. P l - - [list_t_nchotomy] Theorem - - ⊢ ∀ll. (∃t l. ll = ListCons t l) ∨ ll = ListNil - - [node_elem_t_11] Theorem - - ⊢ ∀a0 a1 a0' a1'. - NodeElemCons a0 a1 = NodeElemCons a0' a1' ⇔ a0 = a0' ∧ a1 = a1' - - [node_elem_t_Axiom] Theorem - - ⊢ ∀f0 f1 f2 f3. ∃fn0 fn1. - (∀a. fn0 (TreeLeaf a) = f0 a) ∧ - (∀a0 a1 a2. - fn0 (TreeNode a0 a1 a2) = f1 a0 a1 a2 (fn1 a1) (fn0 a2)) ∧ - (∀a0 a1. fn1 (NodeElemCons a0 a1) = f2 a0 a1 (fn0 a0) (fn1 a1)) ∧ - fn1 NodeElemNil = f3 - - [node_elem_t_case_cong] Theorem - - ⊢ ∀M M' f v. - M = M' ∧ (∀a0 a1. M' = NodeElemCons a0 a1 ⇒ f a0 a1 = f' a0 a1) ∧ - (M' = NodeElemNil ⇒ v = v') ⇒ - node_elem_t_CASE M f v = node_elem_t_CASE M' f' v' - - [node_elem_t_case_eq] Theorem - - ⊢ node_elem_t_CASE x f v = v' ⇔ - (∃t n. x = NodeElemCons t n ∧ f t n = v') ∨ - x = NodeElemNil ∧ v = v' - - [node_elem_t_distinct] Theorem - - ⊢ ∀a1 a0. NodeElemCons a0 a1 ≠ NodeElemNil - - [node_elem_t_induction] Theorem - - ⊢ ∀P0 P1. - (∀t. P0 (TreeLeaf t)) ∧ - (∀n t. P1 n ∧ P0 t ⇒ ∀t0. P0 (TreeNode t0 n t)) ∧ - (∀t n. P0 t ∧ P1 n ⇒ P1 (NodeElemCons t n)) ∧ P1 NodeElemNil ⇒ - (∀t. P0 t) ∧ ∀n. P1 n - - [node_elem_t_nchotomy] Theorem - - ⊢ ∀nn. (∃t n. nn = NodeElemCons t n) ∨ nn = NodeElemNil - - [num2empty_enum_t_11] Theorem - - ⊢ ∀r r'. - r < 1 ⇒ - r' < 1 ⇒ - (num2empty_enum_t r = num2empty_enum_t r' ⇔ r = r') - - [num2empty_enum_t_ONTO] Theorem - - ⊢ ∀a. ∃r. a = num2empty_enum_t r ∧ r < 1 - - [num2empty_enum_t_empty_enum_t2num] Theorem - - ⊢ ∀a. num2empty_enum_t (empty_enum_t2num a) = a - - [num2empty_enum_t_thm] Theorem - - ⊢ num2empty_enum_t 0 = EmptyEnumEmpty - - [num2enum_t_11] Theorem - - ⊢ ∀r r'. r < 2 ⇒ r' < 2 ⇒ (num2enum_t r = num2enum_t r' ⇔ r = r') - - [num2enum_t_ONTO] Theorem - - ⊢ ∀a. ∃r. a = num2enum_t r ∧ r < 2 - - [num2enum_t_enum_t2num] Theorem - - ⊢ ∀a. num2enum_t (enum_t2num a) = a - - [num2enum_t_thm] Theorem - - ⊢ num2enum_t 0 = EnumVariant1 ∧ num2enum_t 1 = EnumVariant2 - - [one_t_11] Theorem - - ⊢ ∀a a'. OneOne a = OneOne a' ⇔ a = a' - - [one_t_Axiom] Theorem - - ⊢ ∀f. ∃fn. ∀a. fn (OneOne a) = f a - - [one_t_case_cong] Theorem - - ⊢ ∀M M' f. - M = M' ∧ (∀a. M' = OneOne a ⇒ f a = f' a) ⇒ - one_t_CASE M f = one_t_CASE M' f' - - [one_t_case_eq] Theorem - - ⊢ one_t_CASE x f = v ⇔ ∃t. x = OneOne t ∧ f t = v - - [one_t_induction] Theorem - - ⊢ ∀P. (∀t. P (OneOne t)) ⇒ ∀ $o. P $o - - [one_t_nchotomy] Theorem - - ⊢ ∀oo. ∃t. oo = OneOne t - - [pair_t_11] Theorem - - ⊢ ∀a0 a1 a0' a1'. pair_t a0 a1 = pair_t a0' a1' ⇔ a0 = a0' ∧ a1 = a1' - - [pair_t_Axiom] Theorem - - ⊢ ∀f. ∃fn. ∀a0 a1. fn (pair_t a0 a1) = f a0 a1 - - [pair_t_accessors] Theorem - - ⊢ (∀t t0. (pair_t t t0).pair_x = t) ∧ - ∀t t0. (pair_t t t0).pair_y = t0 - - [pair_t_accfupds] Theorem - - ⊢ (∀p f. (p with pair_y updated_by f).pair_x = p.pair_x) ∧ - (∀p f. (p with pair_x updated_by f).pair_y = p.pair_y) ∧ - (∀p f. (p with pair_x updated_by f).pair_x = f p.pair_x) ∧ - ∀p f. (p with pair_y updated_by f).pair_y = f p.pair_y - - [pair_t_case_cong] Theorem - - ⊢ ∀M M' f. - M = M' ∧ (∀a0 a1. M' = pair_t a0 a1 ⇒ f a0 a1 = f' a0 a1) ⇒ - pair_t_CASE M f = pair_t_CASE M' f' - - [pair_t_case_eq] Theorem - - ⊢ pair_t_CASE x f = v ⇔ ∃t t0. x = pair_t t t0 ∧ f t t0 = v - - [pair_t_component_equality] Theorem - - ⊢ ∀p1 p2. p1 = p2 ⇔ p1.pair_x = p2.pair_x ∧ p1.pair_y = p2.pair_y - - [pair_t_fn_updates] Theorem - - ⊢ (∀f t t0. pair_t t t0 with pair_x updated_by f = pair_t (f t) t0) ∧ - ∀f t t0. pair_t t t0 with pair_y updated_by f = pair_t t (f t0) - - [pair_t_fupdcanon] Theorem - - ⊢ ∀p g f. - p with <|pair_y updated_by f; pair_x updated_by g|> = - p with <|pair_x updated_by g; pair_y updated_by f|> - - [pair_t_fupdcanon_comp] Theorem - - ⊢ (∀g f. - pair_y_fupd f ∘ pair_x_fupd g = pair_x_fupd g ∘ pair_y_fupd f) ∧ - ∀h g f. - pair_y_fupd f ∘ pair_x_fupd g ∘ h = - pair_x_fupd g ∘ pair_y_fupd f ∘ h - - [pair_t_fupdfupds] Theorem - - ⊢ (∀p g f. - p with <|pair_x updated_by f; pair_x updated_by g|> = - p with pair_x updated_by f ∘ g) ∧ - ∀p g f. - p with <|pair_y updated_by f; pair_y updated_by g|> = - p with pair_y updated_by f ∘ g - - [pair_t_fupdfupds_comp] Theorem - - ⊢ ((∀g f. pair_x_fupd f ∘ pair_x_fupd g = pair_x_fupd (f ∘ g)) ∧ - ∀h g f. - pair_x_fupd f ∘ pair_x_fupd g ∘ h = pair_x_fupd (f ∘ g) ∘ h) ∧ - (∀g f. pair_y_fupd f ∘ pair_y_fupd g = pair_y_fupd (f ∘ g)) ∧ - ∀h g f. pair_y_fupd f ∘ pair_y_fupd g ∘ h = pair_y_fupd (f ∘ g) ∘ h - - [pair_t_induction] Theorem - - ⊢ ∀P. (∀t t0. P (pair_t t t0)) ⇒ ∀p. P p - - [pair_t_literal_11] Theorem - - ⊢ ∀t01 t1 t02 t2. - <|pair_x := t01; pair_y := t1|> = <|pair_x := t02; pair_y := t2|> ⇔ - t01 = t02 ∧ t1 = t2 - - [pair_t_literal_nchotomy] Theorem - - ⊢ ∀p. ∃t0 t. p = <|pair_x := t0; pair_y := t|> - - [pair_t_nchotomy] Theorem - - ⊢ ∀pp. ∃t t0. pp = pair_t t t0 - - [pair_t_updates_eq_literal] Theorem - - ⊢ ∀p t0 t. - p with <|pair_x := t0; pair_y := t|> = - <|pair_x := t0; pair_y := t|> - - [struct_with_pair_t_11] Theorem - - ⊢ ∀a a'. struct_with_pair_t a = struct_with_pair_t a' ⇔ a = a' - - [struct_with_pair_t_Axiom] Theorem - - ⊢ ∀f. ∃fn. ∀a. fn (struct_with_pair_t a) = f a - - [struct_with_pair_t_accessors] Theorem - - ⊢ ∀p. (struct_with_pair_t p).struct_with_pair_p = p - - [struct_with_pair_t_accfupds] Theorem - - ⊢ ∀s f. - (s with struct_with_pair_p updated_by f).struct_with_pair_p = - f s.struct_with_pair_p - - [struct_with_pair_t_case_cong] Theorem - - ⊢ ∀M M' f. - M = M' ∧ (∀a. M' = struct_with_pair_t a ⇒ f a = f' a) ⇒ - struct_with_pair_t_CASE M f = struct_with_pair_t_CASE M' f' - - [struct_with_pair_t_case_eq] Theorem - - ⊢ struct_with_pair_t_CASE x f = v ⇔ - ∃p. x = struct_with_pair_t p ∧ f p = v - - [struct_with_pair_t_component_equality] Theorem - - ⊢ ∀s1 s2. s1 = s2 ⇔ s1.struct_with_pair_p = s2.struct_with_pair_p - - [struct_with_pair_t_fn_updates] Theorem - - ⊢ ∀f p. - struct_with_pair_t p with struct_with_pair_p updated_by f = - struct_with_pair_t (f p) - - [struct_with_pair_t_fupdfupds] Theorem - - ⊢ ∀s g f. - s with - <|struct_with_pair_p updated_by f; - struct_with_pair_p updated_by g|> = - s with struct_with_pair_p updated_by f ∘ g - - [struct_with_pair_t_fupdfupds_comp] Theorem - - ⊢ (∀g f. - struct_with_pair_p_fupd f ∘ struct_with_pair_p_fupd g = - struct_with_pair_p_fupd (f ∘ g)) ∧ - ∀h g f. - struct_with_pair_p_fupd f ∘ struct_with_pair_p_fupd g ∘ h = - struct_with_pair_p_fupd (f ∘ g) ∘ h - - [struct_with_pair_t_induction] Theorem - - ⊢ ∀P. (∀p. P (struct_with_pair_t p)) ⇒ ∀s. P s - - [struct_with_pair_t_literal_11] Theorem - - ⊢ ∀p1 p2. - <|struct_with_pair_p := p1|> = <|struct_with_pair_p := p2|> ⇔ - p1 = p2 - - [struct_with_pair_t_literal_nchotomy] Theorem - - ⊢ ∀s. ∃p. s = <|struct_with_pair_p := p|> - - [struct_with_pair_t_nchotomy] Theorem - - ⊢ ∀ss. ∃p. ss = struct_with_pair_t p - - [struct_with_pair_t_updates_eq_literal] Theorem - - ⊢ ∀s p. s with struct_with_pair_p := p = <|struct_with_pair_p := p|> - - [struct_with_tuple_t_11] Theorem - - ⊢ ∀a a'. struct_with_tuple_t a = struct_with_tuple_t a' ⇔ a = a' - - [struct_with_tuple_t_Axiom] Theorem - - ⊢ ∀f. ∃fn. ∀a. fn (struct_with_tuple_t a) = f a - - [struct_with_tuple_t_accessors] Theorem - - ⊢ ∀p. (struct_with_tuple_t p).struct_with_tuple_p = p - - [struct_with_tuple_t_accfupds] Theorem - - ⊢ ∀s f. - (s with struct_with_tuple_p updated_by f).struct_with_tuple_p = - f s.struct_with_tuple_p - - [struct_with_tuple_t_case_cong] Theorem - - ⊢ ∀M M' f. - M = M' ∧ (∀a. M' = struct_with_tuple_t a ⇒ f a = f' a) ⇒ - struct_with_tuple_t_CASE M f = struct_with_tuple_t_CASE M' f' - - [struct_with_tuple_t_case_eq] Theorem - - ⊢ struct_with_tuple_t_CASE x f = v ⇔ - ∃p. x = struct_with_tuple_t p ∧ f p = v - - [struct_with_tuple_t_component_equality] Theorem - - ⊢ ∀s1 s2. s1 = s2 ⇔ s1.struct_with_tuple_p = s2.struct_with_tuple_p - - [struct_with_tuple_t_fn_updates] Theorem - - ⊢ ∀f p. - struct_with_tuple_t p with struct_with_tuple_p updated_by f = - struct_with_tuple_t (f p) - - [struct_with_tuple_t_fupdfupds] Theorem - - ⊢ ∀s g f. - s with - <|struct_with_tuple_p updated_by f; - struct_with_tuple_p updated_by g|> = - s with struct_with_tuple_p updated_by f ∘ g - - [struct_with_tuple_t_fupdfupds_comp] Theorem - - ⊢ (∀g f. - struct_with_tuple_p_fupd f ∘ struct_with_tuple_p_fupd g = - struct_with_tuple_p_fupd (f ∘ g)) ∧ - ∀h g f. - struct_with_tuple_p_fupd f ∘ struct_with_tuple_p_fupd g ∘ h = - struct_with_tuple_p_fupd (f ∘ g) ∘ h - - [struct_with_tuple_t_induction] Theorem - - ⊢ ∀P. (∀p. P (struct_with_tuple_t p)) ⇒ ∀s. P s - - [struct_with_tuple_t_literal_11] Theorem - - ⊢ ∀p1 p2. - <|struct_with_tuple_p := p1|> = <|struct_with_tuple_p := p2|> ⇔ - p1 = p2 - - [struct_with_tuple_t_literal_nchotomy] Theorem - - ⊢ ∀s. ∃p. s = <|struct_with_tuple_p := p|> - - [struct_with_tuple_t_nchotomy] Theorem - - ⊢ ∀ss. ∃p. ss = struct_with_tuple_t p - - [struct_with_tuple_t_updates_eq_literal] Theorem - - ⊢ ∀s p. - s with struct_with_tuple_p := p = <|struct_with_tuple_p := p|> - - [sum_t_11] Theorem - - ⊢ (∀a a'. SumLeft a = SumLeft a' ⇔ a = a') ∧ - ∀a a'. SumRight a = SumRight a' ⇔ a = a' - - [sum_t_Axiom] Theorem - - ⊢ ∀f0 f1. ∃fn. - (∀a. fn (SumLeft a) = f0 a) ∧ ∀a. fn (SumRight a) = f1 a - - [sum_t_case_cong] Theorem - - ⊢ ∀M M' f f1. - M = M' ∧ (∀a. M' = SumLeft a ⇒ f a = f' a) ∧ - (∀a. M' = SumRight a ⇒ f1 a = f1' a) ⇒ - sum_t_CASE M f f1 = sum_t_CASE M' f' f1' - - [sum_t_case_eq] Theorem - - ⊢ sum_t_CASE x f f1 = v ⇔ - (∃t. x = SumLeft t ∧ f t = v) ∨ ∃t. x = SumRight t ∧ f1 t = v - - [sum_t_distinct] Theorem - - ⊢ ∀a' a. SumLeft a ≠ SumRight a' - - [sum_t_induction] Theorem - - ⊢ ∀P. (∀t. P (SumLeft t)) ∧ (∀t. P (SumRight t)) ⇒ ∀s. P s - - [sum_t_nchotomy] Theorem - - ⊢ ∀ss. (∃t. ss = SumLeft t) ∨ ∃t. ss = SumRight t - - [tree_t_11] Theorem - - ⊢ (∀a a'. TreeLeaf a = TreeLeaf a' ⇔ a = a') ∧ - ∀a0 a1 a2 a0' a1' a2'. - TreeNode a0 a1 a2 = TreeNode a0' a1' a2' ⇔ - a0 = a0' ∧ a1 = a1' ∧ a2 = a2' - - [tree_t_Axiom] Theorem - - ⊢ ∀f0 f1 f2 f3. ∃fn0 fn1. - (∀a. fn0 (TreeLeaf a) = f0 a) ∧ - (∀a0 a1 a2. - fn0 (TreeNode a0 a1 a2) = f1 a0 a1 a2 (fn1 a1) (fn0 a2)) ∧ - (∀a0 a1. fn1 (NodeElemCons a0 a1) = f2 a0 a1 (fn0 a0) (fn1 a1)) ∧ - fn1 NodeElemNil = f3 - - [tree_t_case_cong] Theorem - - ⊢ ∀M M' f f1. - M = M' ∧ (∀a. M' = TreeLeaf a ⇒ f a = f' a) ∧ - (∀a0 a1 a2. M' = TreeNode a0 a1 a2 ⇒ f1 a0 a1 a2 = f1' a0 a1 a2) ⇒ - tree_t_CASE M f f1 = tree_t_CASE M' f' f1' - - [tree_t_case_eq] Theorem - - ⊢ tree_t_CASE x f f1 = v ⇔ - (∃t. x = TreeLeaf t ∧ f t = v) ∨ - ∃t0 n t. x = TreeNode t0 n t ∧ f1 t0 n t = v - - [tree_t_distinct] Theorem - - ⊢ ∀a2 a1 a0 a. TreeLeaf a ≠ TreeNode a0 a1 a2 - - [tree_t_induction] Theorem - - ⊢ ∀P0 P1. - (∀t. P0 (TreeLeaf t)) ∧ - (∀n t. P1 n ∧ P0 t ⇒ ∀t0. P0 (TreeNode t0 n t)) ∧ - (∀t n. P0 t ∧ P1 n ⇒ P1 (NodeElemCons t n)) ∧ P1 NodeElemNil ⇒ - (∀t. P0 t) ∧ ∀n. P1 n - - [tree_t_nchotomy] Theorem - - ⊢ ∀tt. (∃t. tt = TreeLeaf t) ∨ ∃t0 n t. tt = TreeNode t0 n t - - -*) -end diff --git a/tests/hol4/misc-paper/Holmakefile b/tests/hol4/misc-paper/Holmakefile deleted file mode 100644 index 3c4b8973..00000000 --- a/tests/hol4/misc-paper/Holmakefile +++ /dev/null @@ -1,5 +0,0 @@ -# This file was automatically generated - modify ../Holmakefile.template instead -INCLUDES = ../../../backends/hol4 - -all: $(DEFAULT_TARGETS) -.PHONY: all diff --git a/tests/hol4/misc-paper/paperScript.sml b/tests/hol4/misc-paper/paperScript.sml deleted file mode 100644 index 3ac5b6ca..00000000 --- a/tests/hol4/misc-paper/paperScript.sml +++ /dev/null @@ -1,136 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [paper] *) -open primitivesLib divDefLib - -val _ = new_theory "paper" - - -val ref_incr_fwd_back_def = Define ‘ - (** [paper::ref_incr]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - ref_incr_fwd_back (x : i32) : i32 result = - i32_add x (int_to_i32 1) -’ - -val test_incr_fwd_def = Define ‘ - (** [paper::test_incr]: forward function *) - test_incr_fwd : unit result = - do - x <- ref_incr_fwd_back (int_to_i32 0); - if ~ (x = int_to_i32 1) then Fail Failure else Return () - od -’ - -(** Unit test for [paper::test_incr] *) -val _ = assert_return (“test_incr_fwd”) - -val choose_fwd_def = Define ‘ - (** [paper::choose]: forward function *) - choose_fwd (b : bool) (x : 't) (y : 't) : 't result = - if b then Return x else Return y -’ - -val choose_back_def = Define ‘ - (** [paper::choose]: backward function 0 *) - choose_back (b : bool) (x : 't) (y : 't) (ret : 't) : ('t # 't) result = - if b then Return (ret, y) else Return (x, ret) -’ - -val test_choose_fwd_def = Define ‘ - (** [paper::test_choose]: forward function *) - test_choose_fwd : unit result = - do - z <- choose_fwd T (int_to_i32 0) (int_to_i32 0); - z0 <- i32_add z (int_to_i32 1); - if ~ (z0 = int_to_i32 1) - then Fail Failure - else ( - do - (x, y) <- choose_back T (int_to_i32 0) (int_to_i32 0) z0; - if ~ (x = int_to_i32 1) - then Fail Failure - else if ~ (y = int_to_i32 0) then Fail Failure else Return () - od) - od -’ - -(** Unit test for [paper::test_choose] *) -val _ = assert_return (“test_choose_fwd”) - -Datatype: - (** [paper::List] *) - list_t = | ListCons 't list_t | ListNil -End - -val [list_nth_mut_fwd_def] = DefineDiv ‘ - (** [paper::list_nth_mut]: forward function *) - list_nth_mut_fwd (l : 't list_t) (i : u32) : 't result = - (case l of - | ListCons x tl => - if i = int_to_u32 0 - then Return x - else (do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_fwd tl i0 - od) - | ListNil => Fail Failure) -’ - -val [list_nth_mut_back_def] = DefineDiv ‘ - (** [paper::list_nth_mut]: backward function 0 *) - list_nth_mut_back (l : 't list_t) (i : u32) (ret : 't) : 't list_t result = - (case l of - | ListCons x tl => - if i = int_to_u32 0 - then Return (ListCons ret tl) - else ( - do - i0 <- u32_sub i (int_to_u32 1); - tl0 <- list_nth_mut_back tl i0 ret; - Return (ListCons x tl0) - od) - | ListNil => Fail Failure) -’ - -val [sum_fwd_def] = DefineDiv ‘ - (** [paper::sum]: forward function *) - sum_fwd (l : i32 list_t) : i32 result = - (case l of - | ListCons x tl => do - i <- sum_fwd tl; - i32_add x i - od - | ListNil => Return (int_to_i32 0)) -’ - -val test_nth_fwd_def = Define ‘ - (** [paper::test_nth]: forward function *) - test_nth_fwd : unit result = - let l = ListNil in - let l0 = ListCons (int_to_i32 3) l in - let l1 = ListCons (int_to_i32 2) l0 in - do - x <- list_nth_mut_fwd (ListCons (int_to_i32 1) l1) (int_to_u32 2); - x0 <- i32_add x (int_to_i32 1); - l2 <- list_nth_mut_back (ListCons (int_to_i32 1) l1) (int_to_u32 2) x0; - i <- sum_fwd l2; - if ~ (i = int_to_i32 7) then Fail Failure else Return () - od -’ - -(** Unit test for [paper::test_nth] *) -val _ = assert_return (“test_nth_fwd”) - -val call_choose_fwd_def = Define ‘ - (** [paper::call_choose]: forward function *) - call_choose_fwd (p : (u32 # u32)) : u32 result = - let (px, py) = p in - do - pz <- choose_fwd T px py; - pz0 <- u32_add pz (int_to_u32 1); - (px0, _) <- choose_back T px py pz0; - Return px0 - od -’ - -val _ = export_theory () diff --git a/tests/hol4/misc-paper/paperTheory.sig b/tests/hol4/misc-paper/paperTheory.sig deleted file mode 100644 index 2da80da1..00000000 --- a/tests/hol4/misc-paper/paperTheory.sig +++ /dev/null @@ -1,210 +0,0 @@ -signature paperTheory = -sig - type thm = Thm.thm - - (* Definitions *) - val call_choose_fwd_def : thm - val choose_back_def : thm - val choose_fwd_def : thm - val list_nth_mut_back_def : thm - val list_nth_mut_fwd_def : thm - val list_t_TY_DEF : thm - val list_t_case_def : thm - val list_t_size_def : thm - val ref_incr_fwd_back_def : thm - val sum_fwd_def : thm - val test_choose_fwd_def : thm - val test_incr_fwd_def : thm - val test_nth_fwd_def : thm - - (* Theorems *) - val datatype_list_t : thm - val list_t_11 : thm - val list_t_Axiom : thm - val list_t_case_cong : thm - val list_t_case_eq : thm - val list_t_distinct : thm - val list_t_induction : thm - val list_t_nchotomy : thm - - val paper_grammars : type_grammar.grammar * term_grammar.grammar -(* - [divDef] Parent theory of "paper" - - [call_choose_fwd_def] Definition - - ⊢ ∀p. call_choose_fwd p = - (let - (px,py) = p - in - do - pz <- choose_fwd T px py; - pz0 <- u32_add pz (int_to_u32 1); - (px0,_) <- choose_back T px py pz0; - Return px0 - od) - - [choose_back_def] Definition - - ⊢ ∀b x y ret. - choose_back b x y ret = - if b then Return (ret,y) else Return (x,ret) - - [choose_fwd_def] Definition - - ⊢ ∀b x y. choose_fwd b x y = if b then Return x else Return y - - [list_nth_mut_back_def] Definition - - ⊢ ∀l i ret. - list_nth_mut_back l i ret = - case l of - ListCons x tl => - if i = int_to_u32 0 then Return (ListCons ret tl) - else - do - i0 <- u32_sub i (int_to_u32 1); - tl0 <- list_nth_mut_back tl i0 ret; - Return (ListCons x tl0) - od - | ListNil => Fail Failure - - [list_nth_mut_fwd_def] Definition - - ⊢ ∀l i. - list_nth_mut_fwd l i = - case l of - ListCons x tl => - if i = int_to_u32 0 then Return x - else - do - i0 <- u32_sub i (int_to_u32 1); - list_nth_mut_fwd tl i0 - od - | ListNil => Fail Failure - - [list_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('list_t'). - (∀a0'. - (∃a0 a1. - a0' = - (λa0 a1. - ind_type$CONSTR 0 a0 - (ind_type$FCONS a1 (λn. ind_type$BOTTOM))) - a0 a1 ∧ $var$('list_t') a1) ∨ - a0' = - ind_type$CONSTR (SUC 0) ARB (λn. ind_type$BOTTOM) ⇒ - $var$('list_t') a0') ⇒ - $var$('list_t') a0') rep - - [list_t_case_def] Definition - - ⊢ (∀a0 a1 f v. list_t_CASE (ListCons a0 a1) f v = f a0 a1) ∧ - ∀f v. list_t_CASE ListNil f v = v - - [list_t_size_def] Definition - - ⊢ (∀f a0 a1. - list_t_size f (ListCons a0 a1) = 1 + (f a0 + list_t_size f a1)) ∧ - ∀f. list_t_size f ListNil = 0 - - [ref_incr_fwd_back_def] Definition - - ⊢ ∀x. ref_incr_fwd_back x = i32_add x (int_to_i32 1) - - [sum_fwd_def] Definition - - ⊢ ∀l. sum_fwd l = - case l of - ListCons x tl => do i <- sum_fwd tl; i32_add x i od - | ListNil => Return (int_to_i32 0) - - [test_choose_fwd_def] Definition - - ⊢ test_choose_fwd = - do - z <- choose_fwd T (int_to_i32 0) (int_to_i32 0); - z0 <- i32_add z (int_to_i32 1); - if z0 ≠ int_to_i32 1 then Fail Failure - else - do - (x,y) <- choose_back T (int_to_i32 0) (int_to_i32 0) z0; - if x ≠ int_to_i32 1 then Fail Failure - else if y ≠ int_to_i32 0 then Fail Failure - else Return () - od - od - - [test_incr_fwd_def] Definition - - ⊢ test_incr_fwd = - do - x <- ref_incr_fwd_back (int_to_i32 0); - if x ≠ int_to_i32 1 then Fail Failure else Return () - od - - [test_nth_fwd_def] Definition - - ⊢ test_nth_fwd = - (let - l = ListNil; - l0 = ListCons (int_to_i32 3) l; - l1 = ListCons (int_to_i32 2) l0 - in - do - x <- - list_nth_mut_fwd (ListCons (int_to_i32 1) l1) (int_to_u32 2); - x0 <- i32_add x (int_to_i32 1); - l2 <- - list_nth_mut_back (ListCons (int_to_i32 1) l1) - (int_to_u32 2) x0; - i <- sum_fwd l2; - if i ≠ int_to_i32 7 then Fail Failure else Return () - od) - - [datatype_list_t] Theorem - - ⊢ DATATYPE (list_t ListCons ListNil) - - [list_t_11] Theorem - - ⊢ ∀a0 a1 a0' a1'. - ListCons a0 a1 = ListCons a0' a1' ⇔ a0 = a0' ∧ a1 = a1' - - [list_t_Axiom] Theorem - - ⊢ ∀f0 f1. ∃fn. - (∀a0 a1. fn (ListCons a0 a1) = f0 a0 a1 (fn a1)) ∧ - fn ListNil = f1 - - [list_t_case_cong] Theorem - - ⊢ ∀M M' f v. - M = M' ∧ (∀a0 a1. M' = ListCons a0 a1 ⇒ f a0 a1 = f' a0 a1) ∧ - (M' = ListNil ⇒ v = v') ⇒ - list_t_CASE M f v = list_t_CASE M' f' v' - - [list_t_case_eq] Theorem - - ⊢ list_t_CASE x f v = v' ⇔ - (∃t l. x = ListCons t l ∧ f t l = v') ∨ x = ListNil ∧ v = v' - - [list_t_distinct] Theorem - - ⊢ ∀a1 a0. ListCons a0 a1 ≠ ListNil - - [list_t_induction] Theorem - - ⊢ ∀P. (∀l. P l ⇒ ∀t. P (ListCons t l)) ∧ P ListNil ⇒ ∀l. P l - - [list_t_nchotomy] Theorem - - ⊢ ∀ll. (∃t l. ll = ListCons t l) ∨ ll = ListNil - - -*) -end diff --git a/tests/hol4/misc-polonius_list/Holmakefile b/tests/hol4/misc-polonius_list/Holmakefile deleted file mode 100644 index 3c4b8973..00000000 --- a/tests/hol4/misc-polonius_list/Holmakefile +++ /dev/null @@ -1,5 +0,0 @@ -# This file was automatically generated - modify ../Holmakefile.template instead -INCLUDES = ../../../backends/hol4 - -all: $(DEFAULT_TARGETS) -.PHONY: all diff --git a/tests/hol4/misc-polonius_list/poloniusListScript.sml b/tests/hol4/misc-polonius_list/poloniusListScript.sml deleted file mode 100644 index 06876ed4..00000000 --- a/tests/hol4/misc-polonius_list/poloniusListScript.sml +++ /dev/null @@ -1,37 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [polonius_list] *) -open primitivesLib divDefLib - -val _ = new_theory "poloniusList" - - -Datatype: - (** [polonius_list::List] *) - list_t = | ListCons 't list_t | ListNil -End - -val [get_list_at_x_fwd_def] = DefineDiv ‘ - (** [polonius_list::get_list_at_x]: forward function *) - get_list_at_x_fwd (ls : u32 list_t) (x : u32) : u32 list_t result = - (case ls of - | ListCons hd tl => - if hd = x then Return (ListCons hd tl) else get_list_at_x_fwd tl x - | ListNil => Return ListNil) -’ - -val [get_list_at_x_back_def] = DefineDiv ‘ - (** [polonius_list::get_list_at_x]: backward function 0 *) - get_list_at_x_back - (ls : u32 list_t) (x : u32) (ret : u32 list_t) : u32 list_t result = - (case ls of - | ListCons hd tl => - if hd = x - then Return ret - else (do - tl0 <- get_list_at_x_back tl x ret; - Return (ListCons hd tl0) - od) - | ListNil => Return ret) -’ - -val _ = export_theory () diff --git a/tests/hol4/misc-polonius_list/poloniusListTheory.sig b/tests/hol4/misc-polonius_list/poloniusListTheory.sig deleted file mode 100644 index 41f21df7..00000000 --- a/tests/hol4/misc-polonius_list/poloniusListTheory.sig +++ /dev/null @@ -1,120 +0,0 @@ -signature poloniusListTheory = -sig - type thm = Thm.thm - - (* Definitions *) - val get_list_at_x_back_def : thm - val get_list_at_x_fwd_def : thm - val list_t_TY_DEF : thm - val list_t_case_def : thm - val list_t_size_def : thm - - (* Theorems *) - val datatype_list_t : thm - val list_t_11 : thm - val list_t_Axiom : thm - val list_t_case_cong : thm - val list_t_case_eq : thm - val list_t_distinct : thm - val list_t_induction : thm - val list_t_nchotomy : thm - - val poloniusList_grammars : type_grammar.grammar * term_grammar.grammar -(* - [divDef] Parent theory of "poloniusList" - - [get_list_at_x_back_def] Definition - - ⊢ ∀ls x ret. - get_list_at_x_back ls x ret = - case ls of - ListCons hd tl => - if hd = x then Return ret - else - do - tl0 <- get_list_at_x_back tl x ret; - Return (ListCons hd tl0) - od - | ListNil => Return ret - - [get_list_at_x_fwd_def] Definition - - ⊢ ∀ls x. - get_list_at_x_fwd ls x = - case ls of - ListCons hd tl => - if hd = x then Return (ListCons hd tl) - else get_list_at_x_fwd tl x - | ListNil => Return ListNil - - [list_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('list_t'). - (∀a0'. - (∃a0 a1. - a0' = - (λa0 a1. - ind_type$CONSTR 0 a0 - (ind_type$FCONS a1 (λn. ind_type$BOTTOM))) - a0 a1 ∧ $var$('list_t') a1) ∨ - a0' = - ind_type$CONSTR (SUC 0) ARB (λn. ind_type$BOTTOM) ⇒ - $var$('list_t') a0') ⇒ - $var$('list_t') a0') rep - - [list_t_case_def] Definition - - ⊢ (∀a0 a1 f v. list_t_CASE (ListCons a0 a1) f v = f a0 a1) ∧ - ∀f v. list_t_CASE ListNil f v = v - - [list_t_size_def] Definition - - ⊢ (∀f a0 a1. - list_t_size f (ListCons a0 a1) = 1 + (f a0 + list_t_size f a1)) ∧ - ∀f. list_t_size f ListNil = 0 - - [datatype_list_t] Theorem - - ⊢ DATATYPE (list_t ListCons ListNil) - - [list_t_11] Theorem - - ⊢ ∀a0 a1 a0' a1'. - ListCons a0 a1 = ListCons a0' a1' ⇔ a0 = a0' ∧ a1 = a1' - - [list_t_Axiom] Theorem - - ⊢ ∀f0 f1. ∃fn. - (∀a0 a1. fn (ListCons a0 a1) = f0 a0 a1 (fn a1)) ∧ - fn ListNil = f1 - - [list_t_case_cong] Theorem - - ⊢ ∀M M' f v. - M = M' ∧ (∀a0 a1. M' = ListCons a0 a1 ⇒ f a0 a1 = f' a0 a1) ∧ - (M' = ListNil ⇒ v = v') ⇒ - list_t_CASE M f v = list_t_CASE M' f' v' - - [list_t_case_eq] Theorem - - ⊢ list_t_CASE x f v = v' ⇔ - (∃t l. x = ListCons t l ∧ f t l = v') ∨ x = ListNil ∧ v = v' - - [list_t_distinct] Theorem - - ⊢ ∀a1 a0. ListCons a0 a1 ≠ ListNil - - [list_t_induction] Theorem - - ⊢ ∀P. (∀l. P l ⇒ ∀t. P (ListCons t l)) ∧ P ListNil ⇒ ∀l. P l - - [list_t_nchotomy] Theorem - - ⊢ ∀ll. (∃t l. ll = ListCons t l) ∨ ll = ListNil - - -*) -end diff --git a/tests/hol4/no_nested_borrows/Holmakefile b/tests/hol4/no_nested_borrows/Holmakefile new file mode 100644 index 00000000..3c4b8973 --- /dev/null +++ b/tests/hol4/no_nested_borrows/Holmakefile @@ -0,0 +1,5 @@ +# This file was automatically generated - modify ../Holmakefile.template instead +INCLUDES = ../../../backends/hol4 + +all: $(DEFAULT_TARGETS) +.PHONY: all diff --git a/tests/hol4/no_nested_borrows/noNestedBorrowsScript.sml b/tests/hol4/no_nested_borrows/noNestedBorrowsScript.sml new file mode 100644 index 00000000..1b2d6121 --- /dev/null +++ b/tests/hol4/no_nested_borrows/noNestedBorrowsScript.sml @@ -0,0 +1,592 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [no_nested_borrows] *) +open primitivesLib divDefLib + +val _ = new_theory "noNestedBorrows" + + +Datatype: + (** [no_nested_borrows::Pair] *) + pair_t = <| pair_x : 't1; pair_y : 't2; |> +End + +Datatype: + (** [no_nested_borrows::List] *) + list_t = | ListCons 't list_t | ListNil +End + +Datatype: + (** [no_nested_borrows::One] *) + one_t = | OneOne 't1 +End + +Datatype: + (** [no_nested_borrows::EmptyEnum] *) + empty_enum_t = | EmptyEnumEmpty +End + +Datatype: + (** [no_nested_borrows::Enum] *) + enum_t = | EnumVariant1 | EnumVariant2 +End + +Type empty_struct_t = “: unit” + +Datatype: + (** [no_nested_borrows::Sum] *) + sum_t = | SumLeft 't1 | SumRight 't2 +End + +val neg_test_fwd_def = Define ‘ + (** [no_nested_borrows::neg_test]: forward function *) + neg_test_fwd (x : i32) : i32 result = + i32_neg x +’ + +val add_test_fwd_def = Define ‘ + (** [no_nested_borrows::add_test]: forward function *) + add_test_fwd (x : u32) (y : u32) : u32 result = + u32_add x y +’ + +val subs_test_fwd_def = Define ‘ + (** [no_nested_borrows::subs_test]: forward function *) + subs_test_fwd (x : u32) (y : u32) : u32 result = + u32_sub x y +’ + +val div_test_fwd_def = Define ‘ + (** [no_nested_borrows::div_test]: forward function *) + div_test_fwd (x : u32) (y : u32) : u32 result = + u32_div x y +’ + +val div_test1_fwd_def = Define ‘ + (** [no_nested_borrows::div_test1]: forward function *) + div_test1_fwd (x : u32) : u32 result = + u32_div x (int_to_u32 2) +’ + +val rem_test_fwd_def = Define ‘ + (** [no_nested_borrows::rem_test]: forward function *) + rem_test_fwd (x : u32) (y : u32) : u32 result = + u32_rem x y +’ + +val cast_test_fwd_def = Define ‘ + (** [no_nested_borrows::cast_test]: forward function *) + cast_test_fwd (x : u32) : i32 result = + mk_i32 (u32_to_int x) +’ + +val test2_fwd_def = Define ‘ + (** [no_nested_borrows::test2]: forward function *) + test2_fwd : unit result = + do + _ <- u32_add (int_to_u32 23) (int_to_u32 44); + Return () + od +’ + +(** Unit test for [no_nested_borrows::test2] *) +val _ = assert_return (“test2_fwd”) + +val get_max_fwd_def = Define ‘ + (** [no_nested_borrows::get_max]: forward function *) + get_max_fwd (x : u32) (y : u32) : u32 result = + if u32_ge x y then Return x else Return y +’ + +val test3_fwd_def = Define ‘ + (** [no_nested_borrows::test3]: forward function *) + test3_fwd : unit result = + do + x <- get_max_fwd (int_to_u32 4) (int_to_u32 3); + y <- get_max_fwd (int_to_u32 10) (int_to_u32 11); + z <- u32_add x y; + if ~ (z = int_to_u32 15) then Fail Failure else Return () + od +’ + +(** Unit test for [no_nested_borrows::test3] *) +val _ = assert_return (“test3_fwd”) + +val test_neg1_fwd_def = Define ‘ + (** [no_nested_borrows::test_neg1]: forward function *) + test_neg1_fwd : unit result = + do + y <- i32_neg (int_to_i32 3); + if ~ (y = int_to_i32 (-3)) then Fail Failure else Return () + od +’ + +(** Unit test for [no_nested_borrows::test_neg1] *) +val _ = assert_return (“test_neg1_fwd”) + +val refs_test1_fwd_def = Define ‘ + (** [no_nested_borrows::refs_test1]: forward function *) + refs_test1_fwd : unit result = + if ~ (int_to_i32 1 = int_to_i32 1) then Fail Failure else Return () +’ + +(** Unit test for [no_nested_borrows::refs_test1] *) +val _ = assert_return (“refs_test1_fwd”) + +val refs_test2_fwd_def = Define ‘ + (** [no_nested_borrows::refs_test2]: forward function *) + refs_test2_fwd : unit result = + if ~ (int_to_i32 2 = int_to_i32 2) + then Fail Failure + else + if ~ (int_to_i32 0 = int_to_i32 0) + then Fail Failure + else + if ~ (int_to_i32 2 = int_to_i32 2) + then Fail Failure + else + if ~ (int_to_i32 2 = int_to_i32 2) then Fail Failure else Return () +’ + +(** Unit test for [no_nested_borrows::refs_test2] *) +val _ = assert_return (“refs_test2_fwd”) + +val test_list1_fwd_def = Define ‘ + (** [no_nested_borrows::test_list1]: forward function *) + test_list1_fwd : unit result = + Return () +’ + +(** Unit test for [no_nested_borrows::test_list1] *) +val _ = assert_return (“test_list1_fwd”) + +val test_box1_fwd_def = Define ‘ + (** [no_nested_borrows::test_box1]: forward function *) + test_box1_fwd : unit result = + let b = int_to_i32 1 in + let x = b in + if ~ (x = int_to_i32 1) then Fail Failure else Return () +’ + +(** Unit test for [no_nested_borrows::test_box1] *) +val _ = assert_return (“test_box1_fwd”) + +val copy_int_fwd_def = Define ‘ + (** [no_nested_borrows::copy_int]: forward function *) + copy_int_fwd (x : i32) : i32 result = + Return x +’ + +val test_unreachable_fwd_def = Define ‘ + (** [no_nested_borrows::test_unreachable]: forward function *) + test_unreachable_fwd (b : bool) : unit result = + if b then Fail Failure else Return () +’ + +val test_panic_fwd_def = Define ‘ + (** [no_nested_borrows::test_panic]: forward function *) + test_panic_fwd (b : bool) : unit result = + if b then Fail Failure else Return () +’ + +val test_copy_int_fwd_def = Define ‘ + (** [no_nested_borrows::test_copy_int]: forward function *) + test_copy_int_fwd : unit result = + do + y <- copy_int_fwd (int_to_i32 0); + if ~ (int_to_i32 0 = y) then Fail Failure else Return () + od +’ + +(** Unit test for [no_nested_borrows::test_copy_int] *) +val _ = assert_return (“test_copy_int_fwd”) + +val is_cons_fwd_def = Define ‘ + (** [no_nested_borrows::is_cons]: forward function *) + is_cons_fwd (l : 't list_t) : bool result = + (case l of | ListCons t l0 => Return T | ListNil => Return F) +’ + +val test_is_cons_fwd_def = Define ‘ + (** [no_nested_borrows::test_is_cons]: forward function *) + test_is_cons_fwd : unit result = + let l = ListNil in + do + b <- is_cons_fwd (ListCons (int_to_i32 0) l); + if ~ b then Fail Failure else Return () + od +’ + +(** Unit test for [no_nested_borrows::test_is_cons] *) +val _ = assert_return (“test_is_cons_fwd”) + +val split_list_fwd_def = Define ‘ + (** [no_nested_borrows::split_list]: forward function *) + split_list_fwd (l : 't list_t) : ('t # 't list_t) result = + (case l of | ListCons hd tl => Return (hd, tl) | ListNil => Fail Failure) +’ + +val test_split_list_fwd_def = Define ‘ + (** [no_nested_borrows::test_split_list]: forward function *) + test_split_list_fwd : unit result = + let l = ListNil in + do + p <- split_list_fwd (ListCons (int_to_i32 0) l); + let (hd, _) = p in + if ~ (hd = int_to_i32 0) then Fail Failure else Return () + od +’ + +(** Unit test for [no_nested_borrows::test_split_list] *) +val _ = assert_return (“test_split_list_fwd”) + +val choose_fwd_def = Define ‘ + (** [no_nested_borrows::choose]: forward function *) + choose_fwd (b : bool) (x : 't) (y : 't) : 't result = + if b then Return x else Return y +’ + +val choose_back_def = Define ‘ + (** [no_nested_borrows::choose]: backward function 0 *) + choose_back (b : bool) (x : 't) (y : 't) (ret : 't) : ('t # 't) result = + if b then Return (ret, y) else Return (x, ret) +’ + +val choose_test_fwd_def = Define ‘ + (** [no_nested_borrows::choose_test]: forward function *) + choose_test_fwd : unit result = + do + z <- choose_fwd T (int_to_i32 0) (int_to_i32 0); + z0 <- i32_add z (int_to_i32 1); + if ~ (z0 = int_to_i32 1) + then Fail Failure + else ( + do + (x, y) <- choose_back T (int_to_i32 0) (int_to_i32 0) z0; + if ~ (x = int_to_i32 1) + then Fail Failure + else if ~ (y = int_to_i32 0) then Fail Failure else Return () + od) + od +’ + +(** Unit test for [no_nested_borrows::choose_test] *) +val _ = assert_return (“choose_test_fwd”) + +val test_char_fwd_def = Define ‘ + (** [no_nested_borrows::test_char]: forward function *) + test_char_fwd : char result = + Return #"a" +’ + +Datatype: + (** [no_nested_borrows::Tree] *) + tree_t = | TreeLeaf 't | TreeNode 't node_elem_t tree_t ; + + (** [no_nested_borrows::NodeElem] *) + node_elem_t = | NodeElemCons tree_t node_elem_t | NodeElemNil +End + +val [list_length_fwd_def] = DefineDiv ‘ + (** [no_nested_borrows::list_length]: forward function *) + list_length_fwd (l : 't list_t) : u32 result = + (case l of + | ListCons t l1 => do + i <- list_length_fwd l1; + u32_add (int_to_u32 1) i + od + | ListNil => Return (int_to_u32 0)) +’ + +val [list_nth_shared_fwd_def] = DefineDiv ‘ + (** [no_nested_borrows::list_nth_shared]: forward function *) + list_nth_shared_fwd (l : 't list_t) (i : u32) : 't result = + (case l of + | ListCons x tl => + if i = int_to_u32 0 + then Return x + else (do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_fwd tl i0 + od) + | ListNil => Fail Failure) +’ + +val [list_nth_mut_fwd_def] = DefineDiv ‘ + (** [no_nested_borrows::list_nth_mut]: forward function *) + list_nth_mut_fwd (l : 't list_t) (i : u32) : 't result = + (case l of + | ListCons x tl => + if i = int_to_u32 0 + then Return x + else (do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_fwd tl i0 + od) + | ListNil => Fail Failure) +’ + +val [list_nth_mut_back_def] = DefineDiv ‘ + (** [no_nested_borrows::list_nth_mut]: backward function 0 *) + list_nth_mut_back (l : 't list_t) (i : u32) (ret : 't) : 't list_t result = + (case l of + | ListCons x tl => + if i = int_to_u32 0 + then Return (ListCons ret tl) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + tl0 <- list_nth_mut_back tl i0 ret; + Return (ListCons x tl0) + od) + | ListNil => Fail Failure) +’ + +val [list_rev_aux_fwd_def] = DefineDiv ‘ + (** [no_nested_borrows::list_rev_aux]: forward function *) + list_rev_aux_fwd (li : 't list_t) (lo : 't list_t) : 't list_t result = + (case li of + | ListCons hd tl => list_rev_aux_fwd tl (ListCons hd lo) + | ListNil => Return lo) +’ + +val list_rev_fwd_back_def = Define ‘ + (** [no_nested_borrows::list_rev]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + list_rev_fwd_back (l : 't list_t) : 't list_t result = + let li = mem_replace_fwd l ListNil in list_rev_aux_fwd li ListNil +’ + +val test_list_functions_fwd_def = Define ‘ + (** [no_nested_borrows::test_list_functions]: forward function *) + test_list_functions_fwd : unit result = + let l = ListNil in + let l0 = ListCons (int_to_i32 2) l in + let l1 = ListCons (int_to_i32 1) l0 in + do + i <- list_length_fwd (ListCons (int_to_i32 0) l1); + if ~ (i = int_to_u32 3) + then Fail Failure + else ( + do + i0 <- list_nth_shared_fwd (ListCons (int_to_i32 0) l1) (int_to_u32 0); + if ~ (i0 = int_to_i32 0) + then Fail Failure + else ( + do + i1 <- list_nth_shared_fwd (ListCons (int_to_i32 0) l1) (int_to_u32 1); + if ~ (i1 = int_to_i32 1) + then Fail Failure + else ( + do + i2 <- + list_nth_shared_fwd (ListCons (int_to_i32 0) l1) (int_to_u32 2); + if ~ (i2 = int_to_i32 2) + then Fail Failure + else ( + do + ls <- + list_nth_mut_back (ListCons (int_to_i32 0) l1) (int_to_u32 1) + (int_to_i32 3); + i3 <- list_nth_shared_fwd ls (int_to_u32 0); + if ~ (i3 = int_to_i32 0) + then Fail Failure + else ( + do + i4 <- list_nth_shared_fwd ls (int_to_u32 1); + if ~ (i4 = int_to_i32 3) + then Fail Failure + else ( + do + i5 <- list_nth_shared_fwd ls (int_to_u32 2); + if ~ (i5 = int_to_i32 2) then Fail Failure else Return () + od) + od) + od) + od) + od) + od) + od +’ + +(** Unit test for [no_nested_borrows::test_list_functions] *) +val _ = assert_return (“test_list_functions_fwd”) + +val id_mut_pair1_fwd_def = Define ‘ + (** [no_nested_borrows::id_mut_pair1]: forward function *) + id_mut_pair1_fwd (x : 't1) (y : 't2) : ('t1 # 't2) result = + Return (x, y) +’ + +val id_mut_pair1_back_def = Define ‘ + (** [no_nested_borrows::id_mut_pair1]: backward function 0 *) + id_mut_pair1_back + (x : 't1) (y : 't2) (ret : ('t1 # 't2)) : ('t1 # 't2) result = + let (t, t0) = ret in Return (t, t0) +’ + +val id_mut_pair2_fwd_def = Define ‘ + (** [no_nested_borrows::id_mut_pair2]: forward function *) + id_mut_pair2_fwd (p : ('t1 # 't2)) : ('t1 # 't2) result = + let (t, t0) = p in Return (t, t0) +’ + +val id_mut_pair2_back_def = Define ‘ + (** [no_nested_borrows::id_mut_pair2]: backward function 0 *) + id_mut_pair2_back + (p : ('t1 # 't2)) (ret : ('t1 # 't2)) : ('t1 # 't2) result = + let (t, t0) = ret in Return (t, t0) +’ + +val id_mut_pair3_fwd_def = Define ‘ + (** [no_nested_borrows::id_mut_pair3]: forward function *) + id_mut_pair3_fwd (x : 't1) (y : 't2) : ('t1 # 't2) result = + Return (x, y) +’ + +val id_mut_pair3_back'a_def = Define ‘ + (** [no_nested_borrows::id_mut_pair3]: backward function 0 *) + id_mut_pair3_back'a (x : 't1) (y : 't2) (ret : 't1) : 't1 result = + Return ret +’ + +val id_mut_pair3_back'b_def = Define ‘ + (** [no_nested_borrows::id_mut_pair3]: backward function 1 *) + id_mut_pair3_back'b (x : 't1) (y : 't2) (ret : 't2) : 't2 result = + Return ret +’ + +val id_mut_pair4_fwd_def = Define ‘ + (** [no_nested_borrows::id_mut_pair4]: forward function *) + id_mut_pair4_fwd (p : ('t1 # 't2)) : ('t1 # 't2) result = + let (t, t0) = p in Return (t, t0) +’ + +val id_mut_pair4_back'a_def = Define ‘ + (** [no_nested_borrows::id_mut_pair4]: backward function 0 *) + id_mut_pair4_back'a (p : ('t1 # 't2)) (ret : 't1) : 't1 result = + Return ret +’ + +val id_mut_pair4_back'b_def = Define ‘ + (** [no_nested_borrows::id_mut_pair4]: backward function 1 *) + id_mut_pair4_back'b (p : ('t1 # 't2)) (ret : 't2) : 't2 result = + Return ret +’ + +Datatype: + (** [no_nested_borrows::StructWithTuple] *) + struct_with_tuple_t = <| struct_with_tuple_p : ('t1 # 't2); |> +End + +val new_tuple1_fwd_def = Define ‘ + (** [no_nested_borrows::new_tuple1]: forward function *) + new_tuple1_fwd : (u32, u32) struct_with_tuple_t result = + Return (<| struct_with_tuple_p := (int_to_u32 1, int_to_u32 2) |>) +’ + +val new_tuple2_fwd_def = Define ‘ + (** [no_nested_borrows::new_tuple2]: forward function *) + new_tuple2_fwd : (i16, i16) struct_with_tuple_t result = + Return (<| struct_with_tuple_p := (int_to_i16 1, int_to_i16 2) |>) +’ + +val new_tuple3_fwd_def = Define ‘ + (** [no_nested_borrows::new_tuple3]: forward function *) + new_tuple3_fwd : (u64, i64) struct_with_tuple_t result = + Return (<| struct_with_tuple_p := (int_to_u64 1, int_to_i64 2) |>) +’ + +Datatype: + (** [no_nested_borrows::StructWithPair] *) + struct_with_pair_t = <| struct_with_pair_p : ('t1, 't2) pair_t; |> +End + +val new_pair1_fwd_def = Define ‘ + (** [no_nested_borrows::new_pair1]: forward function *) + new_pair1_fwd : (u32, u32) struct_with_pair_t result = + Return + (<| + struct_with_pair_p := + (<| pair_x := (int_to_u32 1); pair_y := (int_to_u32 2) |>) + |>) +’ + +val test_constants_fwd_def = Define ‘ + (** [no_nested_borrows::test_constants]: forward function *) + test_constants_fwd : unit result = + do + swt <- new_tuple1_fwd; + let (i, _) = swt.struct_with_tuple_p in + if ~ (i = int_to_u32 1) + then Fail Failure + else ( + do + swt0 <- new_tuple2_fwd; + let (i0, _) = swt0.struct_with_tuple_p in + if ~ (i0 = int_to_i16 1) + then Fail Failure + else ( + do + swt1 <- new_tuple3_fwd; + let (i1, _) = swt1.struct_with_tuple_p in + if ~ (i1 = int_to_u64 1) + then Fail Failure + else ( + do + swp <- new_pair1_fwd; + if ~ (swp.struct_with_pair_p.pair_x = int_to_u32 1) + then Fail Failure + else Return () + od) + od) + od) + od +’ + +(** Unit test for [no_nested_borrows::test_constants] *) +val _ = assert_return (“test_constants_fwd”) + +val test_weird_borrows1_fwd_def = Define ‘ + (** [no_nested_borrows::test_weird_borrows1]: forward function *) + test_weird_borrows1_fwd : unit result = + Return () +’ + +(** Unit test for [no_nested_borrows::test_weird_borrows1] *) +val _ = assert_return (“test_weird_borrows1_fwd”) + +val test_mem_replace_fwd_back_def = Define ‘ + (** [no_nested_borrows::test_mem_replace]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + test_mem_replace_fwd_back (px : u32) : u32 result = + let y = mem_replace_fwd px (int_to_u32 1) in + if ~ (y = int_to_u32 0) then Fail Failure else Return (int_to_u32 2) +’ + +val test_shared_borrow_bool1_fwd_def = Define ‘ + (** [no_nested_borrows::test_shared_borrow_bool1]: forward function *) + test_shared_borrow_bool1_fwd (b : bool) : u32 result = + if b then Return (int_to_u32 0) else Return (int_to_u32 1) +’ + +val test_shared_borrow_bool2_fwd_def = Define ‘ + (** [no_nested_borrows::test_shared_borrow_bool2]: forward function *) + test_shared_borrow_bool2_fwd : u32 result = + Return (int_to_u32 0) +’ + +val test_shared_borrow_enum1_fwd_def = Define ‘ + (** [no_nested_borrows::test_shared_borrow_enum1]: forward function *) + test_shared_borrow_enum1_fwd (l : u32 list_t) : u32 result = + (case l of + | ListCons i l0 => Return (int_to_u32 1) + | ListNil => Return (int_to_u32 0)) +’ + +val test_shared_borrow_enum2_fwd_def = Define ‘ + (** [no_nested_borrows::test_shared_borrow_enum2]: forward function *) + test_shared_borrow_enum2_fwd : u32 result = + Return (int_to_u32 0) +’ + +val _ = export_theory () diff --git a/tests/hol4/no_nested_borrows/noNestedBorrowsTheory.sig b/tests/hol4/no_nested_borrows/noNestedBorrowsTheory.sig new file mode 100644 index 00000000..67368e38 --- /dev/null +++ b/tests/hol4/no_nested_borrows/noNestedBorrowsTheory.sig @@ -0,0 +1,1598 @@ +signature noNestedBorrowsTheory = +sig + type thm = Thm.thm + + (* Definitions *) + val add_test_fwd_def : thm + val cast_test_fwd_def : thm + val choose_back_def : thm + val choose_fwd_def : thm + val choose_test_fwd_def : thm + val copy_int_fwd_def : thm + val div_test1_fwd_def : thm + val div_test_fwd_def : thm + val empty_enum_t_BIJ : thm + val empty_enum_t_CASE : thm + val empty_enum_t_TY_DEF : thm + val empty_enum_t_size_def : thm + val enum_t_BIJ : thm + val enum_t_CASE : thm + val enum_t_TY_DEF : thm + val enum_t_size_def : thm + val get_max_fwd_def : thm + val id_mut_pair1_back_def : thm + val id_mut_pair1_fwd_def : thm + val id_mut_pair2_back_def : thm + val id_mut_pair2_fwd_def : thm + val id_mut_pair3_back'a_def : thm + val id_mut_pair3_back'b_def : thm + val id_mut_pair3_fwd_def : thm + val id_mut_pair4_back'a_def : thm + val id_mut_pair4_back'b_def : thm + val id_mut_pair4_fwd_def : thm + val is_cons_fwd_def : thm + val list_length_fwd_def : thm + val list_nth_mut_back_def : thm + val list_nth_mut_fwd_def : thm + val list_nth_shared_fwd_def : thm + val list_rev_aux_fwd_def : thm + val list_rev_fwd_back_def : thm + val list_t_TY_DEF : thm + val list_t_case_def : thm + val list_t_size_def : thm + val neg_test_fwd_def : thm + val new_pair1_fwd_def : thm + val new_tuple1_fwd_def : thm + val new_tuple2_fwd_def : thm + val new_tuple3_fwd_def : thm + val node_elem_t_TY_DEF : thm + val node_elem_t_case_def : thm + val one_t_TY_DEF : thm + val one_t_case_def : thm + val one_t_size_def : thm + val pair_t_TY_DEF : thm + val pair_t_case_def : thm + val pair_t_pair_x : thm + val pair_t_pair_x_fupd : thm + val pair_t_pair_y : thm + val pair_t_pair_y_fupd : thm + val pair_t_size_def : thm + val refs_test1_fwd_def : thm + val refs_test2_fwd_def : thm + val rem_test_fwd_def : thm + val split_list_fwd_def : thm + val struct_with_pair_t_TY_DEF : thm + val struct_with_pair_t_case_def : thm + val struct_with_pair_t_size_def : thm + val struct_with_pair_t_struct_with_pair_p : thm + val struct_with_pair_t_struct_with_pair_p_fupd : thm + val struct_with_tuple_t_TY_DEF : thm + val struct_with_tuple_t_case_def : thm + val struct_with_tuple_t_size_def : thm + val struct_with_tuple_t_struct_with_tuple_p : thm + val struct_with_tuple_t_struct_with_tuple_p_fupd : thm + val subs_test_fwd_def : thm + val sum_t_TY_DEF : thm + val sum_t_case_def : thm + val sum_t_size_def : thm + val test2_fwd_def : thm + val test3_fwd_def : thm + val test_box1_fwd_def : thm + val test_char_fwd_def : thm + val test_constants_fwd_def : thm + val test_copy_int_fwd_def : thm + val test_is_cons_fwd_def : thm + val test_list1_fwd_def : thm + val test_list_functions_fwd_def : thm + val test_mem_replace_fwd_back_def : thm + val test_neg1_fwd_def : thm + val test_panic_fwd_def : thm + val test_shared_borrow_bool1_fwd_def : thm + val test_shared_borrow_bool2_fwd_def : thm + val test_shared_borrow_enum1_fwd_def : thm + val test_shared_borrow_enum2_fwd_def : thm + val test_split_list_fwd_def : thm + val test_unreachable_fwd_def : thm + val test_weird_borrows1_fwd_def : thm + val tree_t_TY_DEF : thm + val tree_t_case_def : thm + val tree_t_size_def : thm + + (* Theorems *) + val EXISTS_pair_t : thm + val EXISTS_struct_with_pair_t : thm + val EXISTS_struct_with_tuple_t : thm + val FORALL_pair_t : thm + val FORALL_struct_with_pair_t : thm + val FORALL_struct_with_tuple_t : thm + val datatype_empty_enum_t : thm + val datatype_enum_t : thm + val datatype_list_t : thm + val datatype_one_t : thm + val datatype_pair_t : thm + val datatype_struct_with_pair_t : thm + val datatype_struct_with_tuple_t : thm + val datatype_sum_t : thm + val datatype_tree_t : thm + val empty_enum_t2num_11 : thm + val empty_enum_t2num_ONTO : thm + val empty_enum_t2num_num2empty_enum_t : thm + val empty_enum_t2num_thm : thm + val empty_enum_t_Axiom : thm + val empty_enum_t_EQ_empty_enum_t : thm + val empty_enum_t_case_cong : thm + val empty_enum_t_case_def : thm + val empty_enum_t_case_eq : thm + val empty_enum_t_induction : thm + val empty_enum_t_nchotomy : thm + val enum_t2num_11 : thm + val enum_t2num_ONTO : thm + val enum_t2num_num2enum_t : thm + val enum_t2num_thm : thm + val enum_t_Axiom : thm + val enum_t_EQ_enum_t : thm + val enum_t_case_cong : thm + val enum_t_case_def : thm + val enum_t_case_eq : thm + val enum_t_distinct : thm + val enum_t_induction : thm + val enum_t_nchotomy : thm + val list_t_11 : thm + val list_t_Axiom : thm + val list_t_case_cong : thm + val list_t_case_eq : thm + val list_t_distinct : thm + val list_t_induction : thm + val list_t_nchotomy : thm + val node_elem_t_11 : thm + val node_elem_t_Axiom : thm + val node_elem_t_case_cong : thm + val node_elem_t_case_eq : thm + val node_elem_t_distinct : thm + val node_elem_t_induction : thm + val node_elem_t_nchotomy : thm + val num2empty_enum_t_11 : thm + val num2empty_enum_t_ONTO : thm + val num2empty_enum_t_empty_enum_t2num : thm + val num2empty_enum_t_thm : thm + val num2enum_t_11 : thm + val num2enum_t_ONTO : thm + val num2enum_t_enum_t2num : thm + val num2enum_t_thm : thm + val one_t_11 : thm + val one_t_Axiom : thm + val one_t_case_cong : thm + val one_t_case_eq : thm + val one_t_induction : thm + val one_t_nchotomy : thm + val pair_t_11 : thm + val pair_t_Axiom : thm + val pair_t_accessors : thm + val pair_t_accfupds : thm + val pair_t_case_cong : thm + val pair_t_case_eq : thm + val pair_t_component_equality : thm + val pair_t_fn_updates : thm + val pair_t_fupdcanon : thm + val pair_t_fupdcanon_comp : thm + val pair_t_fupdfupds : thm + val pair_t_fupdfupds_comp : thm + val pair_t_induction : thm + val pair_t_literal_11 : thm + val pair_t_literal_nchotomy : thm + val pair_t_nchotomy : thm + val pair_t_updates_eq_literal : thm + val struct_with_pair_t_11 : thm + val struct_with_pair_t_Axiom : thm + val struct_with_pair_t_accessors : thm + val struct_with_pair_t_accfupds : thm + val struct_with_pair_t_case_cong : thm + val struct_with_pair_t_case_eq : thm + val struct_with_pair_t_component_equality : thm + val struct_with_pair_t_fn_updates : thm + val struct_with_pair_t_fupdfupds : thm + val struct_with_pair_t_fupdfupds_comp : thm + val struct_with_pair_t_induction : thm + val struct_with_pair_t_literal_11 : thm + val struct_with_pair_t_literal_nchotomy : thm + val struct_with_pair_t_nchotomy : thm + val struct_with_pair_t_updates_eq_literal : thm + val struct_with_tuple_t_11 : thm + val struct_with_tuple_t_Axiom : thm + val struct_with_tuple_t_accessors : thm + val struct_with_tuple_t_accfupds : thm + val struct_with_tuple_t_case_cong : thm + val struct_with_tuple_t_case_eq : thm + val struct_with_tuple_t_component_equality : thm + val struct_with_tuple_t_fn_updates : thm + val struct_with_tuple_t_fupdfupds : thm + val struct_with_tuple_t_fupdfupds_comp : thm + val struct_with_tuple_t_induction : thm + val struct_with_tuple_t_literal_11 : thm + val struct_with_tuple_t_literal_nchotomy : thm + val struct_with_tuple_t_nchotomy : thm + val struct_with_tuple_t_updates_eq_literal : thm + val sum_t_11 : thm + val sum_t_Axiom : thm + val sum_t_case_cong : thm + val sum_t_case_eq : thm + val sum_t_distinct : thm + val sum_t_induction : thm + val sum_t_nchotomy : thm + val tree_t_11 : thm + val tree_t_Axiom : thm + val tree_t_case_cong : thm + val tree_t_case_eq : thm + val tree_t_distinct : thm + val tree_t_induction : thm + val tree_t_nchotomy : thm + + val noNestedBorrows_grammars : type_grammar.grammar * term_grammar.grammar +(* + [divDef] Parent theory of "noNestedBorrows" + + [add_test_fwd_def] Definition + + ⊢ ∀x y. add_test_fwd x y = u32_add x y + + [cast_test_fwd_def] Definition + + ⊢ ∀x. cast_test_fwd x = mk_i32 (u32_to_int x) + + [choose_back_def] Definition + + ⊢ ∀b x y ret. + choose_back b x y ret = + if b then Return (ret,y) else Return (x,ret) + + [choose_fwd_def] Definition + + ⊢ ∀b x y. choose_fwd b x y = if b then Return x else Return y + + [choose_test_fwd_def] Definition + + ⊢ choose_test_fwd = + do + z <- choose_fwd T (int_to_i32 0) (int_to_i32 0); + z0 <- i32_add z (int_to_i32 1); + if z0 ≠ int_to_i32 1 then Fail Failure + else + do + (x,y) <- choose_back T (int_to_i32 0) (int_to_i32 0) z0; + if x ≠ int_to_i32 1 then Fail Failure + else if y ≠ int_to_i32 0 then Fail Failure + else Return () + od + od + + [copy_int_fwd_def] Definition + + ⊢ ∀x. copy_int_fwd x = Return x + + [div_test1_fwd_def] Definition + + ⊢ ∀x. div_test1_fwd x = u32_div x (int_to_u32 2) + + [div_test_fwd_def] Definition + + ⊢ ∀x y. div_test_fwd x y = u32_div x y + + [empty_enum_t_BIJ] Definition + + ⊢ (∀a. num2empty_enum_t (empty_enum_t2num a) = a) ∧ + ∀r. (λn. n < 1) r ⇔ empty_enum_t2num (num2empty_enum_t r) = r + + [empty_enum_t_CASE] Definition + + ⊢ ∀x v0. + (case x of EmptyEnumEmpty => v0) = (λm. v0) (empty_enum_t2num x) + + [empty_enum_t_TY_DEF] Definition + + ⊢ ∃rep. TYPE_DEFINITION (λn. n < 1) rep + + [empty_enum_t_size_def] Definition + + ⊢ ∀x. empty_enum_t_size x = 0 + + [enum_t_BIJ] Definition + + ⊢ (∀a. num2enum_t (enum_t2num a) = a) ∧ + ∀r. (λn. n < 2) r ⇔ enum_t2num (num2enum_t r) = r + + [enum_t_CASE] Definition + + ⊢ ∀x v0 v1. + (case x of EnumVariant1 => v0 | EnumVariant2 => v1) = + (λm. if m = 0 then v0 else v1) (enum_t2num x) + + [enum_t_TY_DEF] Definition + + ⊢ ∃rep. TYPE_DEFINITION (λn. n < 2) rep + + [enum_t_size_def] Definition + + ⊢ ∀x. enum_t_size x = 0 + + [get_max_fwd_def] Definition + + ⊢ ∀x y. get_max_fwd x y = if u32_ge x y then Return x else Return y + + [id_mut_pair1_back_def] Definition + + ⊢ ∀x y ret. + id_mut_pair1_back x y ret = (let (t,t0) = ret in Return (t,t0)) + + [id_mut_pair1_fwd_def] Definition + + ⊢ ∀x y. id_mut_pair1_fwd x y = Return (x,y) + + [id_mut_pair2_back_def] Definition + + ⊢ ∀p ret. + id_mut_pair2_back p ret = (let (t,t0) = ret in Return (t,t0)) + + [id_mut_pair2_fwd_def] Definition + + ⊢ ∀p. id_mut_pair2_fwd p = (let (t,t0) = p in Return (t,t0)) + + [id_mut_pair3_back'a_def] Definition + + ⊢ ∀x y ret. id_mut_pair3_back'a x y ret = Return ret + + [id_mut_pair3_back'b_def] Definition + + ⊢ ∀x y ret. id_mut_pair3_back'b x y ret = Return ret + + [id_mut_pair3_fwd_def] Definition + + ⊢ ∀x y. id_mut_pair3_fwd x y = Return (x,y) + + [id_mut_pair4_back'a_def] Definition + + ⊢ ∀p ret. id_mut_pair4_back'a p ret = Return ret + + [id_mut_pair4_back'b_def] Definition + + ⊢ ∀p ret. id_mut_pair4_back'b p ret = Return ret + + [id_mut_pair4_fwd_def] Definition + + ⊢ ∀p. id_mut_pair4_fwd p = (let (t,t0) = p in Return (t,t0)) + + [is_cons_fwd_def] Definition + + ⊢ ∀l. is_cons_fwd l = + case l of ListCons t l0 => Return T | ListNil => Return F + + [list_length_fwd_def] Definition + + ⊢ ∀l. list_length_fwd l = + case l of + ListCons t l1 => + do i <- list_length_fwd l1; u32_add (int_to_u32 1) i od + | ListNil => Return (int_to_u32 0) + + [list_nth_mut_back_def] Definition + + ⊢ ∀l i ret. + list_nth_mut_back l i ret = + case l of + ListCons x tl => + if i = int_to_u32 0 then Return (ListCons ret tl) + else + do + i0 <- u32_sub i (int_to_u32 1); + tl0 <- list_nth_mut_back tl i0 ret; + Return (ListCons x tl0) + od + | ListNil => Fail Failure + + [list_nth_mut_fwd_def] Definition + + ⊢ ∀l i. + list_nth_mut_fwd l i = + case l of + ListCons x tl => + if i = int_to_u32 0 then Return x + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_fwd tl i0 + od + | ListNil => Fail Failure + + [list_nth_shared_fwd_def] Definition + + ⊢ ∀l i. + list_nth_shared_fwd l i = + case l of + ListCons x tl => + if i = int_to_u32 0 then Return x + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_shared_fwd tl i0 + od + | ListNil => Fail Failure + + [list_rev_aux_fwd_def] Definition + + ⊢ ∀li lo. + list_rev_aux_fwd li lo = + case li of + ListCons hd tl => list_rev_aux_fwd tl (ListCons hd lo) + | ListNil => Return lo + + [list_rev_fwd_back_def] Definition + + ⊢ ∀l. list_rev_fwd_back l = + (let + li = mem_replace_fwd l ListNil + in + list_rev_aux_fwd li ListNil) + + [list_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('list_t'). + (∀a0'. + (∃a0 a1. + a0' = + (λa0 a1. + ind_type$CONSTR 0 a0 + (ind_type$FCONS a1 (λn. ind_type$BOTTOM))) + a0 a1 ∧ $var$('list_t') a1) ∨ + a0' = + ind_type$CONSTR (SUC 0) ARB (λn. ind_type$BOTTOM) ⇒ + $var$('list_t') a0') ⇒ + $var$('list_t') a0') rep + + [list_t_case_def] Definition + + ⊢ (∀a0 a1 f v. list_t_CASE (ListCons a0 a1) f v = f a0 a1) ∧ + ∀f v. list_t_CASE ListNil f v = v + + [list_t_size_def] Definition + + ⊢ (∀f a0 a1. + list_t_size f (ListCons a0 a1) = 1 + (f a0 + list_t_size f a1)) ∧ + ∀f. list_t_size f ListNil = 0 + + [neg_test_fwd_def] Definition + + ⊢ ∀x. neg_test_fwd x = i32_neg x + + [new_pair1_fwd_def] Definition + + ⊢ new_pair1_fwd = + Return + <|struct_with_pair_p := + <|pair_x := int_to_u32 1; pair_y := int_to_u32 2|> |> + + [new_tuple1_fwd_def] Definition + + ⊢ new_tuple1_fwd = + Return <|struct_with_tuple_p := (int_to_u32 1,int_to_u32 2)|> + + [new_tuple2_fwd_def] Definition + + ⊢ new_tuple2_fwd = + Return <|struct_with_tuple_p := (int_to_i16 1,int_to_i16 2)|> + + [new_tuple3_fwd_def] Definition + + ⊢ new_tuple3_fwd = + Return <|struct_with_tuple_p := (int_to_u64 1,int_to_i64 2)|> + + [node_elem_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa1'. + ∀ $var$('tree_t') $var$('node_elem_t'). + (∀a0'. + (∃a. a0' = + (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) + a) ∨ + (∃a0 a1 a2. + a0' = + (λa0 a1 a2. + ind_type$CONSTR (SUC 0) a0 + (ind_type$FCONS a1 + (ind_type$FCONS a2 (λn. ind_type$BOTTOM)))) + a0 a1 a2 ∧ $var$('node_elem_t') a1 ∧ + $var$('tree_t') a2) ⇒ + $var$('tree_t') a0') ∧ + (∀a1'. + (∃a0 a1. + a1' = + (λa0 a1. + ind_type$CONSTR (SUC (SUC 0)) ARB + (ind_type$FCONS a0 + (ind_type$FCONS a1 (λn. ind_type$BOTTOM)))) + a0 a1 ∧ $var$('tree_t') a0 ∧ + $var$('node_elem_t') a1) ∨ + a1' = + ind_type$CONSTR (SUC (SUC (SUC 0))) ARB + (λn. ind_type$BOTTOM) ⇒ + $var$('node_elem_t') a1') ⇒ + $var$('node_elem_t') a1') rep + + [node_elem_t_case_def] Definition + + ⊢ (∀a0 a1 f v. node_elem_t_CASE (NodeElemCons a0 a1) f v = f a0 a1) ∧ + ∀f v. node_elem_t_CASE NodeElemNil f v = v + + [one_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0. + ∀ $var$('one_t'). + (∀a0. + (∃a. a0 = + (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) + a) ⇒ + $var$('one_t') a0) ⇒ + $var$('one_t') a0) rep + + [one_t_case_def] Definition + + ⊢ ∀a f. one_t_CASE (OneOne a) f = f a + + [one_t_size_def] Definition + + ⊢ ∀f a. one_t_size f (OneOne a) = 1 + f a + + [pair_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('pair_t'). + (∀a0'. + (∃a0 a1. + a0' = + (λa0 a1. + ind_type$CONSTR 0 (a0,a1) + (λn. ind_type$BOTTOM)) a0 a1) ⇒ + $var$('pair_t') a0') ⇒ + $var$('pair_t') a0') rep + + [pair_t_case_def] Definition + + ⊢ ∀a0 a1 f. pair_t_CASE (pair_t a0 a1) f = f a0 a1 + + [pair_t_pair_x] Definition + + ⊢ ∀t t0. (pair_t t t0).pair_x = t + + [pair_t_pair_x_fupd] Definition + + ⊢ ∀f t t0. pair_t t t0 with pair_x updated_by f = pair_t (f t) t0 + + [pair_t_pair_y] Definition + + ⊢ ∀t t0. (pair_t t t0).pair_y = t0 + + [pair_t_pair_y_fupd] Definition + + ⊢ ∀f t t0. pair_t t t0 with pair_y updated_by f = pair_t t (f t0) + + [pair_t_size_def] Definition + + ⊢ ∀f f1 a0 a1. pair_t_size f f1 (pair_t a0 a1) = 1 + (f a0 + f1 a1) + + [refs_test1_fwd_def] Definition + + ⊢ refs_test1_fwd = + if int_to_i32 1 ≠ int_to_i32 1 then Fail Failure else Return () + + [refs_test2_fwd_def] Definition + + ⊢ refs_test2_fwd = + if int_to_i32 2 ≠ int_to_i32 2 then Fail Failure + else if int_to_i32 0 ≠ int_to_i32 0 then Fail Failure + else if int_to_i32 2 ≠ int_to_i32 2 then Fail Failure + else if int_to_i32 2 ≠ int_to_i32 2 then Fail Failure + else Return () + + [rem_test_fwd_def] Definition + + ⊢ ∀x y. rem_test_fwd x y = u32_rem x y + + [split_list_fwd_def] Definition + + ⊢ ∀l. split_list_fwd l = + case l of + ListCons hd tl => Return (hd,tl) + | ListNil => Fail Failure + + [struct_with_pair_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0. + ∀ $var$('struct_with_pair_t'). + (∀a0. + (∃a. a0 = + (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) + a) ⇒ + $var$('struct_with_pair_t') a0) ⇒ + $var$('struct_with_pair_t') a0) rep + + [struct_with_pair_t_case_def] Definition + + ⊢ ∀a f. struct_with_pair_t_CASE (struct_with_pair_t a) f = f a + + [struct_with_pair_t_size_def] Definition + + ⊢ ∀f f1 a. + struct_with_pair_t_size f f1 (struct_with_pair_t a) = + 1 + pair_t_size f f1 a + + [struct_with_pair_t_struct_with_pair_p] Definition + + ⊢ ∀p. (struct_with_pair_t p).struct_with_pair_p = p + + [struct_with_pair_t_struct_with_pair_p_fupd] Definition + + ⊢ ∀f p. + struct_with_pair_t p with struct_with_pair_p updated_by f = + struct_with_pair_t (f p) + + [struct_with_tuple_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0. + ∀ $var$('struct_with_tuple_t'). + (∀a0. + (∃a. a0 = + (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) + a) ⇒ + $var$('struct_with_tuple_t') a0) ⇒ + $var$('struct_with_tuple_t') a0) rep + + [struct_with_tuple_t_case_def] Definition + + ⊢ ∀a f. struct_with_tuple_t_CASE (struct_with_tuple_t a) f = f a + + [struct_with_tuple_t_size_def] Definition + + ⊢ ∀f f1 a. + struct_with_tuple_t_size f f1 (struct_with_tuple_t a) = + 1 + pair_size f f1 a + + [struct_with_tuple_t_struct_with_tuple_p] Definition + + ⊢ ∀p. (struct_with_tuple_t p).struct_with_tuple_p = p + + [struct_with_tuple_t_struct_with_tuple_p_fupd] Definition + + ⊢ ∀f p. + struct_with_tuple_t p with struct_with_tuple_p updated_by f = + struct_with_tuple_t (f p) + + [subs_test_fwd_def] Definition + + ⊢ ∀x y. subs_test_fwd x y = u32_sub x y + + [sum_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0. + ∀ $var$('sum_t'). + (∀a0. + (∃a. a0 = + (λa. + ind_type$CONSTR 0 (a,ARB) + (λn. ind_type$BOTTOM)) a) ∨ + (∃a. a0 = + (λa. + ind_type$CONSTR (SUC 0) (ARB,a) + (λn. ind_type$BOTTOM)) a) ⇒ + $var$('sum_t') a0) ⇒ + $var$('sum_t') a0) rep + + [sum_t_case_def] Definition + + ⊢ (∀a f f1. sum_t_CASE (SumLeft a) f f1 = f a) ∧ + ∀a f f1. sum_t_CASE (SumRight a) f f1 = f1 a + + [sum_t_size_def] Definition + + ⊢ (∀f f1 a. sum_t_size f f1 (SumLeft a) = 1 + f a) ∧ + ∀f f1 a. sum_t_size f f1 (SumRight a) = 1 + f1 a + + [test2_fwd_def] Definition + + ⊢ test2_fwd = + monad_ignore_bind (u32_add (int_to_u32 23) (int_to_u32 44)) + (Return ()) + + [test3_fwd_def] Definition + + ⊢ test3_fwd = + do + x <- get_max_fwd (int_to_u32 4) (int_to_u32 3); + y <- get_max_fwd (int_to_u32 10) (int_to_u32 11); + z <- u32_add x y; + if z ≠ int_to_u32 15 then Fail Failure else Return () + od + + [test_box1_fwd_def] Definition + + ⊢ test_box1_fwd = + (let + b = int_to_i32 1; + x = b + in + if x ≠ int_to_i32 1 then Fail Failure else Return ()) + + [test_char_fwd_def] Definition + + ⊢ test_char_fwd = Return #"a" + + [test_constants_fwd_def] Definition + + ⊢ test_constants_fwd = + do + swt <- new_tuple1_fwd; + (i,_) <<- swt.struct_with_tuple_p; + if i ≠ int_to_u32 1 then Fail Failure + else + do + swt0 <- new_tuple2_fwd; + (i0,_) <<- swt0.struct_with_tuple_p; + if i0 ≠ int_to_i16 1 then Fail Failure + else + do + swt1 <- new_tuple3_fwd; + (i1,_) <<- swt1.struct_with_tuple_p; + if i1 ≠ int_to_u64 1 then Fail Failure + else + do + swp <- new_pair1_fwd; + if swp.struct_with_pair_p.pair_x ≠ int_to_u32 1 then + Fail Failure + else Return () + od + od + od + od + + [test_copy_int_fwd_def] Definition + + ⊢ test_copy_int_fwd = + do + y <- copy_int_fwd (int_to_i32 0); + if int_to_i32 0 ≠ y then Fail Failure else Return () + od + + [test_is_cons_fwd_def] Definition + + ⊢ test_is_cons_fwd = + (let + l = ListNil + in + do + b <- is_cons_fwd (ListCons (int_to_i32 0) l); + if ¬b then Fail Failure else Return () + od) + + [test_list1_fwd_def] Definition + + ⊢ test_list1_fwd = Return () + + [test_list_functions_fwd_def] Definition + + ⊢ test_list_functions_fwd = + (let + l = ListNil; + l0 = ListCons (int_to_i32 2) l; + l1 = ListCons (int_to_i32 1) l0 + in + do + i <- list_length_fwd (ListCons (int_to_i32 0) l1); + if i ≠ int_to_u32 3 then Fail Failure + else + do + i0 <- + list_nth_shared_fwd (ListCons (int_to_i32 0) l1) + (int_to_u32 0); + if i0 ≠ int_to_i32 0 then Fail Failure + else + do + i1 <- + list_nth_shared_fwd (ListCons (int_to_i32 0) l1) + (int_to_u32 1); + if i1 ≠ int_to_i32 1 then Fail Failure + else + do + i2 <- + list_nth_shared_fwd (ListCons (int_to_i32 0) l1) + (int_to_u32 2); + if i2 ≠ int_to_i32 2 then Fail Failure + else + do + ls <- + list_nth_mut_back + (ListCons (int_to_i32 0) l1) + (int_to_u32 1) (int_to_i32 3); + i3 <- list_nth_shared_fwd ls (int_to_u32 0); + if i3 ≠ int_to_i32 0 then Fail Failure + else + do + i4 <- + list_nth_shared_fwd ls (int_to_u32 1); + if i4 ≠ int_to_i32 3 then Fail Failure + else + do + i5 <- + list_nth_shared_fwd ls + (int_to_u32 2); + if i5 ≠ int_to_i32 2 then Fail Failure + else Return () + od + od + od + od + od + od + od) + + [test_mem_replace_fwd_back_def] Definition + + ⊢ ∀px. + test_mem_replace_fwd_back px = + (let + y = mem_replace_fwd px (int_to_u32 1) + in + if y ≠ int_to_u32 0 then Fail Failure + else Return (int_to_u32 2)) + + [test_neg1_fwd_def] Definition + + ⊢ test_neg1_fwd = + do + y <- i32_neg (int_to_i32 3); + if y ≠ int_to_i32 (-3) then Fail Failure else Return () + od + + [test_panic_fwd_def] Definition + + ⊢ ∀b. test_panic_fwd b = if b then Fail Failure else Return () + + [test_shared_borrow_bool1_fwd_def] Definition + + ⊢ ∀b. test_shared_borrow_bool1_fwd b = + if b then Return (int_to_u32 0) else Return (int_to_u32 1) + + [test_shared_borrow_bool2_fwd_def] Definition + + ⊢ test_shared_borrow_bool2_fwd = Return (int_to_u32 0) + + [test_shared_borrow_enum1_fwd_def] Definition + + ⊢ ∀l. test_shared_borrow_enum1_fwd l = + case l of + ListCons i l0 => Return (int_to_u32 1) + | ListNil => Return (int_to_u32 0) + + [test_shared_borrow_enum2_fwd_def] Definition + + ⊢ test_shared_borrow_enum2_fwd = Return (int_to_u32 0) + + [test_split_list_fwd_def] Definition + + ⊢ test_split_list_fwd = + (let + l = ListNil + in + do + p <- split_list_fwd (ListCons (int_to_i32 0) l); + (hd,_) <<- p; + if hd ≠ int_to_i32 0 then Fail Failure else Return () + od) + + [test_unreachable_fwd_def] Definition + + ⊢ ∀b. test_unreachable_fwd b = if b then Fail Failure else Return () + + [test_weird_borrows1_fwd_def] Definition + + ⊢ test_weird_borrows1_fwd = Return () + + [tree_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('tree_t') $var$('node_elem_t'). + (∀a0'. + (∃a. a0' = + (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) + a) ∨ + (∃a0 a1 a2. + a0' = + (λa0 a1 a2. + ind_type$CONSTR (SUC 0) a0 + (ind_type$FCONS a1 + (ind_type$FCONS a2 (λn. ind_type$BOTTOM)))) + a0 a1 a2 ∧ $var$('node_elem_t') a1 ∧ + $var$('tree_t') a2) ⇒ + $var$('tree_t') a0') ∧ + (∀a1'. + (∃a0 a1. + a1' = + (λa0 a1. + ind_type$CONSTR (SUC (SUC 0)) ARB + (ind_type$FCONS a0 + (ind_type$FCONS a1 (λn. ind_type$BOTTOM)))) + a0 a1 ∧ $var$('tree_t') a0 ∧ + $var$('node_elem_t') a1) ∨ + a1' = + ind_type$CONSTR (SUC (SUC (SUC 0))) ARB + (λn. ind_type$BOTTOM) ⇒ + $var$('node_elem_t') a1') ⇒ + $var$('tree_t') a0') rep + + [tree_t_case_def] Definition + + ⊢ (∀a f f1. tree_t_CASE (TreeLeaf a) f f1 = f a) ∧ + ∀a0 a1 a2 f f1. tree_t_CASE (TreeNode a0 a1 a2) f f1 = f1 a0 a1 a2 + + [tree_t_size_def] Definition + + ⊢ (∀f a. tree_t_size f (TreeLeaf a) = 1 + f a) ∧ + (∀f a0 a1 a2. + tree_t_size f (TreeNode a0 a1 a2) = + 1 + (f a0 + (node_elem_t_size f a1 + tree_t_size f a2))) ∧ + (∀f a0 a1. + node_elem_t_size f (NodeElemCons a0 a1) = + 1 + (tree_t_size f a0 + node_elem_t_size f a1)) ∧ + ∀f. node_elem_t_size f NodeElemNil = 0 + + [EXISTS_pair_t] Theorem + + ⊢ ∀P. (∃p. P p) ⇔ ∃t0 t. P <|pair_x := t0; pair_y := t|> + + [EXISTS_struct_with_pair_t] Theorem + + ⊢ ∀P. (∃s. P s) ⇔ ∃p. P <|struct_with_pair_p := p|> + + [EXISTS_struct_with_tuple_t] Theorem + + ⊢ ∀P. (∃s. P s) ⇔ ∃p. P <|struct_with_tuple_p := p|> + + [FORALL_pair_t] Theorem + + ⊢ ∀P. (∀p. P p) ⇔ ∀t0 t. P <|pair_x := t0; pair_y := t|> + + [FORALL_struct_with_pair_t] Theorem + + ⊢ ∀P. (∀s. P s) ⇔ ∀p. P <|struct_with_pair_p := p|> + + [FORALL_struct_with_tuple_t] Theorem + + ⊢ ∀P. (∀s. P s) ⇔ ∀p. P <|struct_with_tuple_p := p|> + + [datatype_empty_enum_t] Theorem + + ⊢ DATATYPE (empty_enum_t EmptyEnumEmpty) + + [datatype_enum_t] Theorem + + ⊢ DATATYPE (enum_t EnumVariant1 EnumVariant2) + + [datatype_list_t] Theorem + + ⊢ DATATYPE (list_t ListCons ListNil) + + [datatype_one_t] Theorem + + ⊢ DATATYPE (one_t OneOne) + + [datatype_pair_t] Theorem + + ⊢ DATATYPE (record pair_t pair_x pair_y) + + [datatype_struct_with_pair_t] Theorem + + ⊢ DATATYPE (record struct_with_pair_t struct_with_pair_p) + + [datatype_struct_with_tuple_t] Theorem + + ⊢ DATATYPE (record struct_with_tuple_t struct_with_tuple_p) + + [datatype_sum_t] Theorem + + ⊢ DATATYPE (sum_t SumLeft SumRight) + + [datatype_tree_t] Theorem + + ⊢ DATATYPE + (tree_t TreeLeaf TreeNode ∧ node_elem_t NodeElemCons NodeElemNil) + + [empty_enum_t2num_11] Theorem + + ⊢ ∀a a'. empty_enum_t2num a = empty_enum_t2num a' ⇔ a = a' + + [empty_enum_t2num_ONTO] Theorem + + ⊢ ∀r. r < 1 ⇔ ∃a. r = empty_enum_t2num a + + [empty_enum_t2num_num2empty_enum_t] Theorem + + ⊢ ∀r. r < 1 ⇔ empty_enum_t2num (num2empty_enum_t r) = r + + [empty_enum_t2num_thm] Theorem + + ⊢ empty_enum_t2num EmptyEnumEmpty = 0 + + [empty_enum_t_Axiom] Theorem + + ⊢ ∀x0. ∃f. f EmptyEnumEmpty = x0 + + [empty_enum_t_EQ_empty_enum_t] Theorem + + ⊢ ∀a a'. a = a' ⇔ empty_enum_t2num a = empty_enum_t2num a' + + [empty_enum_t_case_cong] Theorem + + ⊢ ∀M M' v0. + M = M' ∧ (M' = EmptyEnumEmpty ⇒ v0 = v0') ⇒ + (case M of EmptyEnumEmpty => v0) = + case M' of EmptyEnumEmpty => v0' + + [empty_enum_t_case_def] Theorem + + ⊢ ∀v0. (case EmptyEnumEmpty of EmptyEnumEmpty => v0) = v0 + + [empty_enum_t_case_eq] Theorem + + ⊢ (case x of EmptyEnumEmpty => v0) = v ⇔ x = EmptyEnumEmpty ∧ v0 = v + + [empty_enum_t_induction] Theorem + + ⊢ ∀P. P EmptyEnumEmpty ⇒ ∀a. P a + + [empty_enum_t_nchotomy] Theorem + + ⊢ ∀a. a = EmptyEnumEmpty + + [enum_t2num_11] Theorem + + ⊢ ∀a a'. enum_t2num a = enum_t2num a' ⇔ a = a' + + [enum_t2num_ONTO] Theorem + + ⊢ ∀r. r < 2 ⇔ ∃a. r = enum_t2num a + + [enum_t2num_num2enum_t] Theorem + + ⊢ ∀r. r < 2 ⇔ enum_t2num (num2enum_t r) = r + + [enum_t2num_thm] Theorem + + ⊢ enum_t2num EnumVariant1 = 0 ∧ enum_t2num EnumVariant2 = 1 + + [enum_t_Axiom] Theorem + + ⊢ ∀x0 x1. ∃f. f EnumVariant1 = x0 ∧ f EnumVariant2 = x1 + + [enum_t_EQ_enum_t] Theorem + + ⊢ ∀a a'. a = a' ⇔ enum_t2num a = enum_t2num a' + + [enum_t_case_cong] Theorem + + ⊢ ∀M M' v0 v1. + M = M' ∧ (M' = EnumVariant1 ⇒ v0 = v0') ∧ + (M' = EnumVariant2 ⇒ v1 = v1') ⇒ + (case M of EnumVariant1 => v0 | EnumVariant2 => v1) = + case M' of EnumVariant1 => v0' | EnumVariant2 => v1' + + [enum_t_case_def] Theorem + + ⊢ (∀v0 v1. + (case EnumVariant1 of EnumVariant1 => v0 | EnumVariant2 => v1) = + v0) ∧ + ∀v0 v1. + (case EnumVariant2 of EnumVariant1 => v0 | EnumVariant2 => v1) = + v1 + + [enum_t_case_eq] Theorem + + ⊢ (case x of EnumVariant1 => v0 | EnumVariant2 => v1) = v ⇔ + x = EnumVariant1 ∧ v0 = v ∨ x = EnumVariant2 ∧ v1 = v + + [enum_t_distinct] Theorem + + ⊢ EnumVariant1 ≠ EnumVariant2 + + [enum_t_induction] Theorem + + ⊢ ∀P. P EnumVariant1 ∧ P EnumVariant2 ⇒ ∀a. P a + + [enum_t_nchotomy] Theorem + + ⊢ ∀a. a = EnumVariant1 ∨ a = EnumVariant2 + + [list_t_11] Theorem + + ⊢ ∀a0 a1 a0' a1'. + ListCons a0 a1 = ListCons a0' a1' ⇔ a0 = a0' ∧ a1 = a1' + + [list_t_Axiom] Theorem + + ⊢ ∀f0 f1. ∃fn. + (∀a0 a1. fn (ListCons a0 a1) = f0 a0 a1 (fn a1)) ∧ + fn ListNil = f1 + + [list_t_case_cong] Theorem + + ⊢ ∀M M' f v. + M = M' ∧ (∀a0 a1. M' = ListCons a0 a1 ⇒ f a0 a1 = f' a0 a1) ∧ + (M' = ListNil ⇒ v = v') ⇒ + list_t_CASE M f v = list_t_CASE M' f' v' + + [list_t_case_eq] Theorem + + ⊢ list_t_CASE x f v = v' ⇔ + (∃t l. x = ListCons t l ∧ f t l = v') ∨ x = ListNil ∧ v = v' + + [list_t_distinct] Theorem + + ⊢ ∀a1 a0. ListCons a0 a1 ≠ ListNil + + [list_t_induction] Theorem + + ⊢ ∀P. (∀l. P l ⇒ ∀t. P (ListCons t l)) ∧ P ListNil ⇒ ∀l. P l + + [list_t_nchotomy] Theorem + + ⊢ ∀ll. (∃t l. ll = ListCons t l) ∨ ll = ListNil + + [node_elem_t_11] Theorem + + ⊢ ∀a0 a1 a0' a1'. + NodeElemCons a0 a1 = NodeElemCons a0' a1' ⇔ a0 = a0' ∧ a1 = a1' + + [node_elem_t_Axiom] Theorem + + ⊢ ∀f0 f1 f2 f3. ∃fn0 fn1. + (∀a. fn0 (TreeLeaf a) = f0 a) ∧ + (∀a0 a1 a2. + fn0 (TreeNode a0 a1 a2) = f1 a0 a1 a2 (fn1 a1) (fn0 a2)) ∧ + (∀a0 a1. fn1 (NodeElemCons a0 a1) = f2 a0 a1 (fn0 a0) (fn1 a1)) ∧ + fn1 NodeElemNil = f3 + + [node_elem_t_case_cong] Theorem + + ⊢ ∀M M' f v. + M = M' ∧ (∀a0 a1. M' = NodeElemCons a0 a1 ⇒ f a0 a1 = f' a0 a1) ∧ + (M' = NodeElemNil ⇒ v = v') ⇒ + node_elem_t_CASE M f v = node_elem_t_CASE M' f' v' + + [node_elem_t_case_eq] Theorem + + ⊢ node_elem_t_CASE x f v = v' ⇔ + (∃t n. x = NodeElemCons t n ∧ f t n = v') ∨ + x = NodeElemNil ∧ v = v' + + [node_elem_t_distinct] Theorem + + ⊢ ∀a1 a0. NodeElemCons a0 a1 ≠ NodeElemNil + + [node_elem_t_induction] Theorem + + ⊢ ∀P0 P1. + (∀t. P0 (TreeLeaf t)) ∧ + (∀n t. P1 n ∧ P0 t ⇒ ∀t0. P0 (TreeNode t0 n t)) ∧ + (∀t n. P0 t ∧ P1 n ⇒ P1 (NodeElemCons t n)) ∧ P1 NodeElemNil ⇒ + (∀t. P0 t) ∧ ∀n. P1 n + + [node_elem_t_nchotomy] Theorem + + ⊢ ∀nn. (∃t n. nn = NodeElemCons t n) ∨ nn = NodeElemNil + + [num2empty_enum_t_11] Theorem + + ⊢ ∀r r'. + r < 1 ⇒ + r' < 1 ⇒ + (num2empty_enum_t r = num2empty_enum_t r' ⇔ r = r') + + [num2empty_enum_t_ONTO] Theorem + + ⊢ ∀a. ∃r. a = num2empty_enum_t r ∧ r < 1 + + [num2empty_enum_t_empty_enum_t2num] Theorem + + ⊢ ∀a. num2empty_enum_t (empty_enum_t2num a) = a + + [num2empty_enum_t_thm] Theorem + + ⊢ num2empty_enum_t 0 = EmptyEnumEmpty + + [num2enum_t_11] Theorem + + ⊢ ∀r r'. r < 2 ⇒ r' < 2 ⇒ (num2enum_t r = num2enum_t r' ⇔ r = r') + + [num2enum_t_ONTO] Theorem + + ⊢ ∀a. ∃r. a = num2enum_t r ∧ r < 2 + + [num2enum_t_enum_t2num] Theorem + + ⊢ ∀a. num2enum_t (enum_t2num a) = a + + [num2enum_t_thm] Theorem + + ⊢ num2enum_t 0 = EnumVariant1 ∧ num2enum_t 1 = EnumVariant2 + + [one_t_11] Theorem + + ⊢ ∀a a'. OneOne a = OneOne a' ⇔ a = a' + + [one_t_Axiom] Theorem + + ⊢ ∀f. ∃fn. ∀a. fn (OneOne a) = f a + + [one_t_case_cong] Theorem + + ⊢ ∀M M' f. + M = M' ∧ (∀a. M' = OneOne a ⇒ f a = f' a) ⇒ + one_t_CASE M f = one_t_CASE M' f' + + [one_t_case_eq] Theorem + + ⊢ one_t_CASE x f = v ⇔ ∃t. x = OneOne t ∧ f t = v + + [one_t_induction] Theorem + + ⊢ ∀P. (∀t. P (OneOne t)) ⇒ ∀ $o. P $o + + [one_t_nchotomy] Theorem + + ⊢ ∀oo. ∃t. oo = OneOne t + + [pair_t_11] Theorem + + ⊢ ∀a0 a1 a0' a1'. pair_t a0 a1 = pair_t a0' a1' ⇔ a0 = a0' ∧ a1 = a1' + + [pair_t_Axiom] Theorem + + ⊢ ∀f. ∃fn. ∀a0 a1. fn (pair_t a0 a1) = f a0 a1 + + [pair_t_accessors] Theorem + + ⊢ (∀t t0. (pair_t t t0).pair_x = t) ∧ + ∀t t0. (pair_t t t0).pair_y = t0 + + [pair_t_accfupds] Theorem + + ⊢ (∀p f. (p with pair_y updated_by f).pair_x = p.pair_x) ∧ + (∀p f. (p with pair_x updated_by f).pair_y = p.pair_y) ∧ + (∀p f. (p with pair_x updated_by f).pair_x = f p.pair_x) ∧ + ∀p f. (p with pair_y updated_by f).pair_y = f p.pair_y + + [pair_t_case_cong] Theorem + + ⊢ ∀M M' f. + M = M' ∧ (∀a0 a1. M' = pair_t a0 a1 ⇒ f a0 a1 = f' a0 a1) ⇒ + pair_t_CASE M f = pair_t_CASE M' f' + + [pair_t_case_eq] Theorem + + ⊢ pair_t_CASE x f = v ⇔ ∃t t0. x = pair_t t t0 ∧ f t t0 = v + + [pair_t_component_equality] Theorem + + ⊢ ∀p1 p2. p1 = p2 ⇔ p1.pair_x = p2.pair_x ∧ p1.pair_y = p2.pair_y + + [pair_t_fn_updates] Theorem + + ⊢ (∀f t t0. pair_t t t0 with pair_x updated_by f = pair_t (f t) t0) ∧ + ∀f t t0. pair_t t t0 with pair_y updated_by f = pair_t t (f t0) + + [pair_t_fupdcanon] Theorem + + ⊢ ∀p g f. + p with <|pair_y updated_by f; pair_x updated_by g|> = + p with <|pair_x updated_by g; pair_y updated_by f|> + + [pair_t_fupdcanon_comp] Theorem + + ⊢ (∀g f. + pair_y_fupd f ∘ pair_x_fupd g = pair_x_fupd g ∘ pair_y_fupd f) ∧ + ∀h g f. + pair_y_fupd f ∘ pair_x_fupd g ∘ h = + pair_x_fupd g ∘ pair_y_fupd f ∘ h + + [pair_t_fupdfupds] Theorem + + ⊢ (∀p g f. + p with <|pair_x updated_by f; pair_x updated_by g|> = + p with pair_x updated_by f ∘ g) ∧ + ∀p g f. + p with <|pair_y updated_by f; pair_y updated_by g|> = + p with pair_y updated_by f ∘ g + + [pair_t_fupdfupds_comp] Theorem + + ⊢ ((∀g f. pair_x_fupd f ∘ pair_x_fupd g = pair_x_fupd (f ∘ g)) ∧ + ∀h g f. + pair_x_fupd f ∘ pair_x_fupd g ∘ h = pair_x_fupd (f ∘ g) ∘ h) ∧ + (∀g f. pair_y_fupd f ∘ pair_y_fupd g = pair_y_fupd (f ∘ g)) ∧ + ∀h g f. pair_y_fupd f ∘ pair_y_fupd g ∘ h = pair_y_fupd (f ∘ g) ∘ h + + [pair_t_induction] Theorem + + ⊢ ∀P. (∀t t0. P (pair_t t t0)) ⇒ ∀p. P p + + [pair_t_literal_11] Theorem + + ⊢ ∀t01 t1 t02 t2. + <|pair_x := t01; pair_y := t1|> = <|pair_x := t02; pair_y := t2|> ⇔ + t01 = t02 ∧ t1 = t2 + + [pair_t_literal_nchotomy] Theorem + + ⊢ ∀p. ∃t0 t. p = <|pair_x := t0; pair_y := t|> + + [pair_t_nchotomy] Theorem + + ⊢ ∀pp. ∃t t0. pp = pair_t t t0 + + [pair_t_updates_eq_literal] Theorem + + ⊢ ∀p t0 t. + p with <|pair_x := t0; pair_y := t|> = + <|pair_x := t0; pair_y := t|> + + [struct_with_pair_t_11] Theorem + + ⊢ ∀a a'. struct_with_pair_t a = struct_with_pair_t a' ⇔ a = a' + + [struct_with_pair_t_Axiom] Theorem + + ⊢ ∀f. ∃fn. ∀a. fn (struct_with_pair_t a) = f a + + [struct_with_pair_t_accessors] Theorem + + ⊢ ∀p. (struct_with_pair_t p).struct_with_pair_p = p + + [struct_with_pair_t_accfupds] Theorem + + ⊢ ∀s f. + (s with struct_with_pair_p updated_by f).struct_with_pair_p = + f s.struct_with_pair_p + + [struct_with_pair_t_case_cong] Theorem + + ⊢ ∀M M' f. + M = M' ∧ (∀a. M' = struct_with_pair_t a ⇒ f a = f' a) ⇒ + struct_with_pair_t_CASE M f = struct_with_pair_t_CASE M' f' + + [struct_with_pair_t_case_eq] Theorem + + ⊢ struct_with_pair_t_CASE x f = v ⇔ + ∃p. x = struct_with_pair_t p ∧ f p = v + + [struct_with_pair_t_component_equality] Theorem + + ⊢ ∀s1 s2. s1 = s2 ⇔ s1.struct_with_pair_p = s2.struct_with_pair_p + + [struct_with_pair_t_fn_updates] Theorem + + ⊢ ∀f p. + struct_with_pair_t p with struct_with_pair_p updated_by f = + struct_with_pair_t (f p) + + [struct_with_pair_t_fupdfupds] Theorem + + ⊢ ∀s g f. + s with + <|struct_with_pair_p updated_by f; + struct_with_pair_p updated_by g|> = + s with struct_with_pair_p updated_by f ∘ g + + [struct_with_pair_t_fupdfupds_comp] Theorem + + ⊢ (∀g f. + struct_with_pair_p_fupd f ∘ struct_with_pair_p_fupd g = + struct_with_pair_p_fupd (f ∘ g)) ∧ + ∀h g f. + struct_with_pair_p_fupd f ∘ struct_with_pair_p_fupd g ∘ h = + struct_with_pair_p_fupd (f ∘ g) ∘ h + + [struct_with_pair_t_induction] Theorem + + ⊢ ∀P. (∀p. P (struct_with_pair_t p)) ⇒ ∀s. P s + + [struct_with_pair_t_literal_11] Theorem + + ⊢ ∀p1 p2. + <|struct_with_pair_p := p1|> = <|struct_with_pair_p := p2|> ⇔ + p1 = p2 + + [struct_with_pair_t_literal_nchotomy] Theorem + + ⊢ ∀s. ∃p. s = <|struct_with_pair_p := p|> + + [struct_with_pair_t_nchotomy] Theorem + + ⊢ ∀ss. ∃p. ss = struct_with_pair_t p + + [struct_with_pair_t_updates_eq_literal] Theorem + + ⊢ ∀s p. s with struct_with_pair_p := p = <|struct_with_pair_p := p|> + + [struct_with_tuple_t_11] Theorem + + ⊢ ∀a a'. struct_with_tuple_t a = struct_with_tuple_t a' ⇔ a = a' + + [struct_with_tuple_t_Axiom] Theorem + + ⊢ ∀f. ∃fn. ∀a. fn (struct_with_tuple_t a) = f a + + [struct_with_tuple_t_accessors] Theorem + + ⊢ ∀p. (struct_with_tuple_t p).struct_with_tuple_p = p + + [struct_with_tuple_t_accfupds] Theorem + + ⊢ ∀s f. + (s with struct_with_tuple_p updated_by f).struct_with_tuple_p = + f s.struct_with_tuple_p + + [struct_with_tuple_t_case_cong] Theorem + + ⊢ ∀M M' f. + M = M' ∧ (∀a. M' = struct_with_tuple_t a ⇒ f a = f' a) ⇒ + struct_with_tuple_t_CASE M f = struct_with_tuple_t_CASE M' f' + + [struct_with_tuple_t_case_eq] Theorem + + ⊢ struct_with_tuple_t_CASE x f = v ⇔ + ∃p. x = struct_with_tuple_t p ∧ f p = v + + [struct_with_tuple_t_component_equality] Theorem + + ⊢ ∀s1 s2. s1 = s2 ⇔ s1.struct_with_tuple_p = s2.struct_with_tuple_p + + [struct_with_tuple_t_fn_updates] Theorem + + ⊢ ∀f p. + struct_with_tuple_t p with struct_with_tuple_p updated_by f = + struct_with_tuple_t (f p) + + [struct_with_tuple_t_fupdfupds] Theorem + + ⊢ ∀s g f. + s with + <|struct_with_tuple_p updated_by f; + struct_with_tuple_p updated_by g|> = + s with struct_with_tuple_p updated_by f ∘ g + + [struct_with_tuple_t_fupdfupds_comp] Theorem + + ⊢ (∀g f. + struct_with_tuple_p_fupd f ∘ struct_with_tuple_p_fupd g = + struct_with_tuple_p_fupd (f ∘ g)) ∧ + ∀h g f. + struct_with_tuple_p_fupd f ∘ struct_with_tuple_p_fupd g ∘ h = + struct_with_tuple_p_fupd (f ∘ g) ∘ h + + [struct_with_tuple_t_induction] Theorem + + ⊢ ∀P. (∀p. P (struct_with_tuple_t p)) ⇒ ∀s. P s + + [struct_with_tuple_t_literal_11] Theorem + + ⊢ ∀p1 p2. + <|struct_with_tuple_p := p1|> = <|struct_with_tuple_p := p2|> ⇔ + p1 = p2 + + [struct_with_tuple_t_literal_nchotomy] Theorem + + ⊢ ∀s. ∃p. s = <|struct_with_tuple_p := p|> + + [struct_with_tuple_t_nchotomy] Theorem + + ⊢ ∀ss. ∃p. ss = struct_with_tuple_t p + + [struct_with_tuple_t_updates_eq_literal] Theorem + + ⊢ ∀s p. + s with struct_with_tuple_p := p = <|struct_with_tuple_p := p|> + + [sum_t_11] Theorem + + ⊢ (∀a a'. SumLeft a = SumLeft a' ⇔ a = a') ∧ + ∀a a'. SumRight a = SumRight a' ⇔ a = a' + + [sum_t_Axiom] Theorem + + ⊢ ∀f0 f1. ∃fn. + (∀a. fn (SumLeft a) = f0 a) ∧ ∀a. fn (SumRight a) = f1 a + + [sum_t_case_cong] Theorem + + ⊢ ∀M M' f f1. + M = M' ∧ (∀a. M' = SumLeft a ⇒ f a = f' a) ∧ + (∀a. M' = SumRight a ⇒ f1 a = f1' a) ⇒ + sum_t_CASE M f f1 = sum_t_CASE M' f' f1' + + [sum_t_case_eq] Theorem + + ⊢ sum_t_CASE x f f1 = v ⇔ + (∃t. x = SumLeft t ∧ f t = v) ∨ ∃t. x = SumRight t ∧ f1 t = v + + [sum_t_distinct] Theorem + + ⊢ ∀a' a. SumLeft a ≠ SumRight a' + + [sum_t_induction] Theorem + + ⊢ ∀P. (∀t. P (SumLeft t)) ∧ (∀t. P (SumRight t)) ⇒ ∀s. P s + + [sum_t_nchotomy] Theorem + + ⊢ ∀ss. (∃t. ss = SumLeft t) ∨ ∃t. ss = SumRight t + + [tree_t_11] Theorem + + ⊢ (∀a a'. TreeLeaf a = TreeLeaf a' ⇔ a = a') ∧ + ∀a0 a1 a2 a0' a1' a2'. + TreeNode a0 a1 a2 = TreeNode a0' a1' a2' ⇔ + a0 = a0' ∧ a1 = a1' ∧ a2 = a2' + + [tree_t_Axiom] Theorem + + ⊢ ∀f0 f1 f2 f3. ∃fn0 fn1. + (∀a. fn0 (TreeLeaf a) = f0 a) ∧ + (∀a0 a1 a2. + fn0 (TreeNode a0 a1 a2) = f1 a0 a1 a2 (fn1 a1) (fn0 a2)) ∧ + (∀a0 a1. fn1 (NodeElemCons a0 a1) = f2 a0 a1 (fn0 a0) (fn1 a1)) ∧ + fn1 NodeElemNil = f3 + + [tree_t_case_cong] Theorem + + ⊢ ∀M M' f f1. + M = M' ∧ (∀a. M' = TreeLeaf a ⇒ f a = f' a) ∧ + (∀a0 a1 a2. M' = TreeNode a0 a1 a2 ⇒ f1 a0 a1 a2 = f1' a0 a1 a2) ⇒ + tree_t_CASE M f f1 = tree_t_CASE M' f' f1' + + [tree_t_case_eq] Theorem + + ⊢ tree_t_CASE x f f1 = v ⇔ + (∃t. x = TreeLeaf t ∧ f t = v) ∨ + ∃t0 n t. x = TreeNode t0 n t ∧ f1 t0 n t = v + + [tree_t_distinct] Theorem + + ⊢ ∀a2 a1 a0 a. TreeLeaf a ≠ TreeNode a0 a1 a2 + + [tree_t_induction] Theorem + + ⊢ ∀P0 P1. + (∀t. P0 (TreeLeaf t)) ∧ + (∀n t. P1 n ∧ P0 t ⇒ ∀t0. P0 (TreeNode t0 n t)) ∧ + (∀t n. P0 t ∧ P1 n ⇒ P1 (NodeElemCons t n)) ∧ P1 NodeElemNil ⇒ + (∀t. P0 t) ∧ ∀n. P1 n + + [tree_t_nchotomy] Theorem + + ⊢ ∀tt. (∃t. tt = TreeLeaf t) ∨ ∃t0 n t. tt = TreeNode t0 n t + + +*) +end diff --git a/tests/hol4/paper/Holmakefile b/tests/hol4/paper/Holmakefile new file mode 100644 index 00000000..3c4b8973 --- /dev/null +++ b/tests/hol4/paper/Holmakefile @@ -0,0 +1,5 @@ +# This file was automatically generated - modify ../Holmakefile.template instead +INCLUDES = ../../../backends/hol4 + +all: $(DEFAULT_TARGETS) +.PHONY: all diff --git a/tests/hol4/paper/paperScript.sml b/tests/hol4/paper/paperScript.sml new file mode 100644 index 00000000..3ac5b6ca --- /dev/null +++ b/tests/hol4/paper/paperScript.sml @@ -0,0 +1,136 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [paper] *) +open primitivesLib divDefLib + +val _ = new_theory "paper" + + +val ref_incr_fwd_back_def = Define ‘ + (** [paper::ref_incr]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + ref_incr_fwd_back (x : i32) : i32 result = + i32_add x (int_to_i32 1) +’ + +val test_incr_fwd_def = Define ‘ + (** [paper::test_incr]: forward function *) + test_incr_fwd : unit result = + do + x <- ref_incr_fwd_back (int_to_i32 0); + if ~ (x = int_to_i32 1) then Fail Failure else Return () + od +’ + +(** Unit test for [paper::test_incr] *) +val _ = assert_return (“test_incr_fwd”) + +val choose_fwd_def = Define ‘ + (** [paper::choose]: forward function *) + choose_fwd (b : bool) (x : 't) (y : 't) : 't result = + if b then Return x else Return y +’ + +val choose_back_def = Define ‘ + (** [paper::choose]: backward function 0 *) + choose_back (b : bool) (x : 't) (y : 't) (ret : 't) : ('t # 't) result = + if b then Return (ret, y) else Return (x, ret) +’ + +val test_choose_fwd_def = Define ‘ + (** [paper::test_choose]: forward function *) + test_choose_fwd : unit result = + do + z <- choose_fwd T (int_to_i32 0) (int_to_i32 0); + z0 <- i32_add z (int_to_i32 1); + if ~ (z0 = int_to_i32 1) + then Fail Failure + else ( + do + (x, y) <- choose_back T (int_to_i32 0) (int_to_i32 0) z0; + if ~ (x = int_to_i32 1) + then Fail Failure + else if ~ (y = int_to_i32 0) then Fail Failure else Return () + od) + od +’ + +(** Unit test for [paper::test_choose] *) +val _ = assert_return (“test_choose_fwd”) + +Datatype: + (** [paper::List] *) + list_t = | ListCons 't list_t | ListNil +End + +val [list_nth_mut_fwd_def] = DefineDiv ‘ + (** [paper::list_nth_mut]: forward function *) + list_nth_mut_fwd (l : 't list_t) (i : u32) : 't result = + (case l of + | ListCons x tl => + if i = int_to_u32 0 + then Return x + else (do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_fwd tl i0 + od) + | ListNil => Fail Failure) +’ + +val [list_nth_mut_back_def] = DefineDiv ‘ + (** [paper::list_nth_mut]: backward function 0 *) + list_nth_mut_back (l : 't list_t) (i : u32) (ret : 't) : 't list_t result = + (case l of + | ListCons x tl => + if i = int_to_u32 0 + then Return (ListCons ret tl) + else ( + do + i0 <- u32_sub i (int_to_u32 1); + tl0 <- list_nth_mut_back tl i0 ret; + Return (ListCons x tl0) + od) + | ListNil => Fail Failure) +’ + +val [sum_fwd_def] = DefineDiv ‘ + (** [paper::sum]: forward function *) + sum_fwd (l : i32 list_t) : i32 result = + (case l of + | ListCons x tl => do + i <- sum_fwd tl; + i32_add x i + od + | ListNil => Return (int_to_i32 0)) +’ + +val test_nth_fwd_def = Define ‘ + (** [paper::test_nth]: forward function *) + test_nth_fwd : unit result = + let l = ListNil in + let l0 = ListCons (int_to_i32 3) l in + let l1 = ListCons (int_to_i32 2) l0 in + do + x <- list_nth_mut_fwd (ListCons (int_to_i32 1) l1) (int_to_u32 2); + x0 <- i32_add x (int_to_i32 1); + l2 <- list_nth_mut_back (ListCons (int_to_i32 1) l1) (int_to_u32 2) x0; + i <- sum_fwd l2; + if ~ (i = int_to_i32 7) then Fail Failure else Return () + od +’ + +(** Unit test for [paper::test_nth] *) +val _ = assert_return (“test_nth_fwd”) + +val call_choose_fwd_def = Define ‘ + (** [paper::call_choose]: forward function *) + call_choose_fwd (p : (u32 # u32)) : u32 result = + let (px, py) = p in + do + pz <- choose_fwd T px py; + pz0 <- u32_add pz (int_to_u32 1); + (px0, _) <- choose_back T px py pz0; + Return px0 + od +’ + +val _ = export_theory () diff --git a/tests/hol4/paper/paperTheory.sig b/tests/hol4/paper/paperTheory.sig new file mode 100644 index 00000000..2da80da1 --- /dev/null +++ b/tests/hol4/paper/paperTheory.sig @@ -0,0 +1,210 @@ +signature paperTheory = +sig + type thm = Thm.thm + + (* Definitions *) + val call_choose_fwd_def : thm + val choose_back_def : thm + val choose_fwd_def : thm + val list_nth_mut_back_def : thm + val list_nth_mut_fwd_def : thm + val list_t_TY_DEF : thm + val list_t_case_def : thm + val list_t_size_def : thm + val ref_incr_fwd_back_def : thm + val sum_fwd_def : thm + val test_choose_fwd_def : thm + val test_incr_fwd_def : thm + val test_nth_fwd_def : thm + + (* Theorems *) + val datatype_list_t : thm + val list_t_11 : thm + val list_t_Axiom : thm + val list_t_case_cong : thm + val list_t_case_eq : thm + val list_t_distinct : thm + val list_t_induction : thm + val list_t_nchotomy : thm + + val paper_grammars : type_grammar.grammar * term_grammar.grammar +(* + [divDef] Parent theory of "paper" + + [call_choose_fwd_def] Definition + + ⊢ ∀p. call_choose_fwd p = + (let + (px,py) = p + in + do + pz <- choose_fwd T px py; + pz0 <- u32_add pz (int_to_u32 1); + (px0,_) <- choose_back T px py pz0; + Return px0 + od) + + [choose_back_def] Definition + + ⊢ ∀b x y ret. + choose_back b x y ret = + if b then Return (ret,y) else Return (x,ret) + + [choose_fwd_def] Definition + + ⊢ ∀b x y. choose_fwd b x y = if b then Return x else Return y + + [list_nth_mut_back_def] Definition + + ⊢ ∀l i ret. + list_nth_mut_back l i ret = + case l of + ListCons x tl => + if i = int_to_u32 0 then Return (ListCons ret tl) + else + do + i0 <- u32_sub i (int_to_u32 1); + tl0 <- list_nth_mut_back tl i0 ret; + Return (ListCons x tl0) + od + | ListNil => Fail Failure + + [list_nth_mut_fwd_def] Definition + + ⊢ ∀l i. + list_nth_mut_fwd l i = + case l of + ListCons x tl => + if i = int_to_u32 0 then Return x + else + do + i0 <- u32_sub i (int_to_u32 1); + list_nth_mut_fwd tl i0 + od + | ListNil => Fail Failure + + [list_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('list_t'). + (∀a0'. + (∃a0 a1. + a0' = + (λa0 a1. + ind_type$CONSTR 0 a0 + (ind_type$FCONS a1 (λn. ind_type$BOTTOM))) + a0 a1 ∧ $var$('list_t') a1) ∨ + a0' = + ind_type$CONSTR (SUC 0) ARB (λn. ind_type$BOTTOM) ⇒ + $var$('list_t') a0') ⇒ + $var$('list_t') a0') rep + + [list_t_case_def] Definition + + ⊢ (∀a0 a1 f v. list_t_CASE (ListCons a0 a1) f v = f a0 a1) ∧ + ∀f v. list_t_CASE ListNil f v = v + + [list_t_size_def] Definition + + ⊢ (∀f a0 a1. + list_t_size f (ListCons a0 a1) = 1 + (f a0 + list_t_size f a1)) ∧ + ∀f. list_t_size f ListNil = 0 + + [ref_incr_fwd_back_def] Definition + + ⊢ ∀x. ref_incr_fwd_back x = i32_add x (int_to_i32 1) + + [sum_fwd_def] Definition + + ⊢ ∀l. sum_fwd l = + case l of + ListCons x tl => do i <- sum_fwd tl; i32_add x i od + | ListNil => Return (int_to_i32 0) + + [test_choose_fwd_def] Definition + + ⊢ test_choose_fwd = + do + z <- choose_fwd T (int_to_i32 0) (int_to_i32 0); + z0 <- i32_add z (int_to_i32 1); + if z0 ≠ int_to_i32 1 then Fail Failure + else + do + (x,y) <- choose_back T (int_to_i32 0) (int_to_i32 0) z0; + if x ≠ int_to_i32 1 then Fail Failure + else if y ≠ int_to_i32 0 then Fail Failure + else Return () + od + od + + [test_incr_fwd_def] Definition + + ⊢ test_incr_fwd = + do + x <- ref_incr_fwd_back (int_to_i32 0); + if x ≠ int_to_i32 1 then Fail Failure else Return () + od + + [test_nth_fwd_def] Definition + + ⊢ test_nth_fwd = + (let + l = ListNil; + l0 = ListCons (int_to_i32 3) l; + l1 = ListCons (int_to_i32 2) l0 + in + do + x <- + list_nth_mut_fwd (ListCons (int_to_i32 1) l1) (int_to_u32 2); + x0 <- i32_add x (int_to_i32 1); + l2 <- + list_nth_mut_back (ListCons (int_to_i32 1) l1) + (int_to_u32 2) x0; + i <- sum_fwd l2; + if i ≠ int_to_i32 7 then Fail Failure else Return () + od) + + [datatype_list_t] Theorem + + ⊢ DATATYPE (list_t ListCons ListNil) + + [list_t_11] Theorem + + ⊢ ∀a0 a1 a0' a1'. + ListCons a0 a1 = ListCons a0' a1' ⇔ a0 = a0' ∧ a1 = a1' + + [list_t_Axiom] Theorem + + ⊢ ∀f0 f1. ∃fn. + (∀a0 a1. fn (ListCons a0 a1) = f0 a0 a1 (fn a1)) ∧ + fn ListNil = f1 + + [list_t_case_cong] Theorem + + ⊢ ∀M M' f v. + M = M' ∧ (∀a0 a1. M' = ListCons a0 a1 ⇒ f a0 a1 = f' a0 a1) ∧ + (M' = ListNil ⇒ v = v') ⇒ + list_t_CASE M f v = list_t_CASE M' f' v' + + [list_t_case_eq] Theorem + + ⊢ list_t_CASE x f v = v' ⇔ + (∃t l. x = ListCons t l ∧ f t l = v') ∨ x = ListNil ∧ v = v' + + [list_t_distinct] Theorem + + ⊢ ∀a1 a0. ListCons a0 a1 ≠ ListNil + + [list_t_induction] Theorem + + ⊢ ∀P. (∀l. P l ⇒ ∀t. P (ListCons t l)) ∧ P ListNil ⇒ ∀l. P l + + [list_t_nchotomy] Theorem + + ⊢ ∀ll. (∃t l. ll = ListCons t l) ∨ ll = ListNil + + +*) +end diff --git a/tests/hol4/polonius_list/Holmakefile b/tests/hol4/polonius_list/Holmakefile new file mode 100644 index 00000000..3c4b8973 --- /dev/null +++ b/tests/hol4/polonius_list/Holmakefile @@ -0,0 +1,5 @@ +# This file was automatically generated - modify ../Holmakefile.template instead +INCLUDES = ../../../backends/hol4 + +all: $(DEFAULT_TARGETS) +.PHONY: all diff --git a/tests/hol4/polonius_list/poloniusListScript.sml b/tests/hol4/polonius_list/poloniusListScript.sml new file mode 100644 index 00000000..06876ed4 --- /dev/null +++ b/tests/hol4/polonius_list/poloniusListScript.sml @@ -0,0 +1,37 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [polonius_list] *) +open primitivesLib divDefLib + +val _ = new_theory "poloniusList" + + +Datatype: + (** [polonius_list::List] *) + list_t = | ListCons 't list_t | ListNil +End + +val [get_list_at_x_fwd_def] = DefineDiv ‘ + (** [polonius_list::get_list_at_x]: forward function *) + get_list_at_x_fwd (ls : u32 list_t) (x : u32) : u32 list_t result = + (case ls of + | ListCons hd tl => + if hd = x then Return (ListCons hd tl) else get_list_at_x_fwd tl x + | ListNil => Return ListNil) +’ + +val [get_list_at_x_back_def] = DefineDiv ‘ + (** [polonius_list::get_list_at_x]: backward function 0 *) + get_list_at_x_back + (ls : u32 list_t) (x : u32) (ret : u32 list_t) : u32 list_t result = + (case ls of + | ListCons hd tl => + if hd = x + then Return ret + else (do + tl0 <- get_list_at_x_back tl x ret; + Return (ListCons hd tl0) + od) + | ListNil => Return ret) +’ + +val _ = export_theory () diff --git a/tests/hol4/polonius_list/poloniusListTheory.sig b/tests/hol4/polonius_list/poloniusListTheory.sig new file mode 100644 index 00000000..41f21df7 --- /dev/null +++ b/tests/hol4/polonius_list/poloniusListTheory.sig @@ -0,0 +1,120 @@ +signature poloniusListTheory = +sig + type thm = Thm.thm + + (* Definitions *) + val get_list_at_x_back_def : thm + val get_list_at_x_fwd_def : thm + val list_t_TY_DEF : thm + val list_t_case_def : thm + val list_t_size_def : thm + + (* Theorems *) + val datatype_list_t : thm + val list_t_11 : thm + val list_t_Axiom : thm + val list_t_case_cong : thm + val list_t_case_eq : thm + val list_t_distinct : thm + val list_t_induction : thm + val list_t_nchotomy : thm + + val poloniusList_grammars : type_grammar.grammar * term_grammar.grammar +(* + [divDef] Parent theory of "poloniusList" + + [get_list_at_x_back_def] Definition + + ⊢ ∀ls x ret. + get_list_at_x_back ls x ret = + case ls of + ListCons hd tl => + if hd = x then Return ret + else + do + tl0 <- get_list_at_x_back tl x ret; + Return (ListCons hd tl0) + od + | ListNil => Return ret + + [get_list_at_x_fwd_def] Definition + + ⊢ ∀ls x. + get_list_at_x_fwd ls x = + case ls of + ListCons hd tl => + if hd = x then Return (ListCons hd tl) + else get_list_at_x_fwd tl x + | ListNil => Return ListNil + + [list_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('list_t'). + (∀a0'. + (∃a0 a1. + a0' = + (λa0 a1. + ind_type$CONSTR 0 a0 + (ind_type$FCONS a1 (λn. ind_type$BOTTOM))) + a0 a1 ∧ $var$('list_t') a1) ∨ + a0' = + ind_type$CONSTR (SUC 0) ARB (λn. ind_type$BOTTOM) ⇒ + $var$('list_t') a0') ⇒ + $var$('list_t') a0') rep + + [list_t_case_def] Definition + + ⊢ (∀a0 a1 f v. list_t_CASE (ListCons a0 a1) f v = f a0 a1) ∧ + ∀f v. list_t_CASE ListNil f v = v + + [list_t_size_def] Definition + + ⊢ (∀f a0 a1. + list_t_size f (ListCons a0 a1) = 1 + (f a0 + list_t_size f a1)) ∧ + ∀f. list_t_size f ListNil = 0 + + [datatype_list_t] Theorem + + ⊢ DATATYPE (list_t ListCons ListNil) + + [list_t_11] Theorem + + ⊢ ∀a0 a1 a0' a1'. + ListCons a0 a1 = ListCons a0' a1' ⇔ a0 = a0' ∧ a1 = a1' + + [list_t_Axiom] Theorem + + ⊢ ∀f0 f1. ∃fn. + (∀a0 a1. fn (ListCons a0 a1) = f0 a0 a1 (fn a1)) ∧ + fn ListNil = f1 + + [list_t_case_cong] Theorem + + ⊢ ∀M M' f v. + M = M' ∧ (∀a0 a1. M' = ListCons a0 a1 ⇒ f a0 a1 = f' a0 a1) ∧ + (M' = ListNil ⇒ v = v') ⇒ + list_t_CASE M f v = list_t_CASE M' f' v' + + [list_t_case_eq] Theorem + + ⊢ list_t_CASE x f v = v' ⇔ + (∃t l. x = ListCons t l ∧ f t l = v') ∨ x = ListNil ∧ v = v' + + [list_t_distinct] Theorem + + ⊢ ∀a1 a0. ListCons a0 a1 ≠ ListNil + + [list_t_induction] Theorem + + ⊢ ∀P. (∀l. P l ⇒ ∀t. P (ListCons t l)) ∧ P ListNil ⇒ ∀l. P l + + [list_t_nchotomy] Theorem + + ⊢ ∀ll. (∃t l. ll = ListCons t l) ∨ ll = ListNil + + +*) +end diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index 8f356365..5d77bf9e 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -104,18 +104,6 @@ let charon_options_for_test test_name = [ "--polonius"; "--opaque=betree_utils"; "--crate"; "betree_main" ] | _ -> [] -(* The default subdirectory in which to store the outputs. *) -(* This reproduces the file layout that was set by the old Makefile. FIXME: cleanup *) -let test_subdir backend test_name = - match (backend, test_name) with - | Backend.Lean, _ -> "." - | _, "hashmap_main" -> "hashmap_on_disk" - | ( Backend.HOL4, - ( "constants" | "external" | "loops" | "no_nested_borrows" | "paper" - | "polonius_list" ) ) -> - "misc-" ^ test_name - | _ -> test_name - (* The data for a specific test input *) module Input = struct type kind = SingleFile | Crate @@ -131,6 +119,10 @@ module Input = struct subdirs : string BackendMap.t; } + (* The default subdirectory in which to store the outputs. *) + let default_subdir backend test_name = + match backend with Backend.Lean -> "." | _ -> test_name + (* Parse lines that start `//@`. Each of them modifies the options we use for the test. Supported comments: - `skip`: don't process the file; @@ -199,7 +191,9 @@ module Input = struct in let actions = BackendMap.make (fun _ -> Normal) in let charon_options = charon_options_for_test name in - let subdirs = BackendMap.make (fun backend -> test_subdir backend name) in + let subdirs = + BackendMap.make (fun backend -> default_subdir backend name) + in let aeneas_options = BackendMap.make (fun backend -> aeneas_options_for_test backend name) in -- cgit v1.2.3 From aee6dc227c4ed041bbbae7cf38729a4b1a3a6869 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Mon, 27 May 2024 10:48:07 +0200 Subject: runner: Correctly catch command exit status --- tests/test_runner/dune | 2 +- tests/test_runner/run_test.ml | 67 +++++++++++++++++++++++++------------------ 2 files changed, 40 insertions(+), 29 deletions(-) diff --git a/tests/test_runner/dune b/tests/test_runner/dune index 1c719532..c38e009c 100644 --- a/tests/test_runner/dune +++ b/tests/test_runner/dune @@ -1,6 +1,6 @@ (executable (public_name test_runner) - (libraries core_unix.sys_unix re str unix) + (libraries core_unix core_unix.sys_unix re str unix) (preprocess (pps ppx_deriving.show ppx_deriving.ord ppx_sexp_conv)) (name run_test)) diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index 5d77bf9e..a5a89317 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -60,24 +60,37 @@ end let concat_path = List.fold_left Filename.concat "" -let run_command args = - (* Debug arguments *) - print_string "[test_runner] Running: "; - Array.iter - (fun x -> - print_string x; - print_string " ") - args; - print_endline ""; +module Command = struct + type t = { args : string array } + type status = Success | Failure - (* Run the command *) - let pid = - Unix.create_process args.(0) args Unix.stdin Unix.stdout Unix.stderr - in - let _ = Unix.waitpid [] pid in - () + let make (args : string list) : t = { args = Array.of_list args } + let to_string (cmd : t) = Core.String.concat_array ~sep:" " cmd.args + + (* Run the command and returns its exit status. *) + let run (cmd : t) : status = + let command_str = to_string cmd in + print_endline ("[test_runner] Running: " ^ command_str); + + (* Run the command *) + let pid = + Unix.create_process cmd.args.(0) cmd.args Unix.stdin Unix.stdout + Unix.stderr + in + let status = Core_unix.waitpid (Core.Pid.of_int pid) in + match status with + | Ok () -> Success + | Error (`Exit_non_zero _) -> Failure + | Error (`Signal _) -> + failwith ("Command `" ^ command_str ^ "` exited incorrectly.") + + (* Run the command and aborts the program if the command failed. *) + let run_command_expecting_success cmd = + match run cmd with + | Success -> () + | Failure -> failwith ("Command `" ^ to_string cmd ^ "` failed.") +end -(* File-specific options *) let aeneas_options_for_test backend test_name = if test_name = "betree" then let options = @@ -226,20 +239,18 @@ let run_aeneas (env : runner_env) (case : Input.t) (backend : Backend.t) = let backend_str = Backend.to_string backend in let dest_dir = concat_path [ "tests"; backend_str; subdir ] in let args = - [| - env.aeneas_path; input_file; "-dest"; dest_dir; "-backend"; backend_str; - |] + [ env.aeneas_path; input_file; "-dest"; dest_dir; "-backend"; backend_str ] in - let args = Array.append args (Array.of_list aeneas_options) in + let args = List.append args aeneas_options in (* Run Aeneas *) - run_command args + Command.run_command_expecting_success (Command.make args) (* Run Charon on a specific input with the given options *) let run_charon (env : runner_env) (case : Input.t) = match case.kind with | SingleFile -> let args = - [| + [ env.charon_path; "--no-cargo"; "--input"; @@ -248,20 +259,20 @@ let run_charon (env : runner_env) (case : Input.t) = case.name; "--dest"; env.llbc_dir; - |] + ] in - let args = Array.append args (Array.of_list case.charon_options) in + let args = List.append args case.charon_options in (* Run Charon on the rust file *) - run_command args + Command.run_command_expecting_success (Command.make args) | Crate -> ( match Sys.getenv_opt "IN_CI" with | None -> - let args = [| env.charon_path; "--dest"; env.llbc_dir |] in - let args = Array.append args (Array.of_list case.charon_options) in + let args = [ env.charon_path; "--dest"; env.llbc_dir ] in + let args = List.append args case.charon_options in (* Run Charon inside the crate *) let old_pwd = Unix.getcwd () in Unix.chdir case.path; - run_command args; + Command.run_command_expecting_success (Command.make args); Unix.chdir old_pwd | Some _ -> (* Crates with dependencies must be generated separately in CI. We skip -- cgit v1.2.3 From 9c09789c26dd8142b8a29b42e250a685aa983e58 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Mon, 27 May 2024 11:22:37 +0200 Subject: runner: Support negative tests --- Makefile | 4 +- flake.nix | 12 +++-- tests/lean/misc/MutuallyRecursiveTraits.lean | 6 +++ tests/src/mutually-recursive-traits.lean.out | 17 +++++++ tests/src/mutually-recursive-traits.rs | 11 +++++ tests/src/string-chars.lean.out | 21 +++++++++ tests/src/string-chars.rs | 6 +++ tests/test_runner/dune | 2 +- tests/test_runner/run_test.ml | 70 ++++++++++++++++++++-------- 9 files changed, 123 insertions(+), 26 deletions(-) create mode 100644 tests/lean/misc/MutuallyRecursiveTraits.lean create mode 100644 tests/src/mutually-recursive-traits.lean.out create mode 100644 tests/src/mutually-recursive-traits.rs create mode 100644 tests/src/string-chars.lean.out create mode 100644 tests/src/string-chars.rs diff --git a/Makefile b/Makefile index bda88c74..38da9f7e 100644 --- a/Makefile +++ b/Makefile @@ -25,7 +25,7 @@ CHARON_OPTIONS ?= # The directory thta contains the rust source files for tests. INPUTS_DIR ?= tests/src # The directory where to look for the .llbc files. -LLBC_DIR ?= $(PWD)/tests/llbc +LLBC_DIR ?= tests/llbc # In CI, we enforce formatting and activate the (expensive) sanity checks. IN_CI ?= @@ -157,6 +157,8 @@ verify: # List the files and directories in `INPUTS_DIR` INPUTS_LIST = $(wildcard $(INPUTS_DIR)/*) +# Remove the committed output files +INPUTS_LIST := $(filter-out %.out,$(INPUTS_LIST)) # Remove the directory prefix, replace with `test-` INPUTS_LIST := $(subst $(INPUTS_DIR)/,test-,$(INPUTS_LIST)) diff --git a/flake.nix b/flake.nix index 68378954..f1a2258a 100644 --- a/flake.nix +++ b/flake.nix @@ -118,17 +118,19 @@ export CHARON_EXE=${charon.packages.${system}.charon}/bin/charon export TEST_RUNNER_EXE=${test_runner}/bin/test_runner - mkdir llbc - export LLBC_DIR=llbc - # Copy over the llbc file we can't generate ourselves. - cp ${betree-llbc}/llbc/betree_main.llbc $LLBC_DIR - # Copy the tests cp -r tests tests-copy make clean-generated-aeneas + mkdir tests/llbc + export LLBC_DIR=tests/llbc + # Copy over the llbc file we can't generate ourselves. + cp ${betree-llbc}/llbc/betree_main.llbc $LLBC_DIR + # Run the tests with extra sanity checks enabled IN_CI=1 make test-all -j $NIX_BUILD_CORES + # Clean generated llbc files so we don't compare them. + rm -r tests/llbc # Check that there are no differences between the generated tests # and the original tests diff --git a/tests/lean/misc/MutuallyRecursiveTraits.lean b/tests/lean/misc/MutuallyRecursiveTraits.lean new file mode 100644 index 00000000..b0fcb9e9 --- /dev/null +++ b/tests/lean/misc/MutuallyRecursiveTraits.lean @@ -0,0 +1,6 @@ +-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS +-- [mutually_recursive_traits] +import Base +open Primitives + +namespace mutually_recursive_traits diff --git a/tests/src/mutually-recursive-traits.lean.out b/tests/src/mutually-recursive-traits.lean.out new file mode 100644 index 00000000..33206fe5 --- /dev/null +++ b/tests/src/mutually-recursive-traits.lean.out @@ -0,0 +1,17 @@ +[Info ] Imported: tests/llbc/mutually_recursive_traits.llbc +[Error] In file Translate.ml, line 813: +Mutually recursive trait declarations are not supported + +Uncaught exception: + + (Failure + "In file Translate.ml, line 813:\ + \nIn file Translate.ml, line 813:\ + \nMutually recursive trait declarations are not supported") + +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 1502, characters 5-42 +Called from Dune__exe__Main in file "Main.ml", line 276, characters 9-61 diff --git a/tests/src/mutually-recursive-traits.rs b/tests/src/mutually-recursive-traits.rs new file mode 100644 index 00000000..351763b2 --- /dev/null +++ b/tests/src/mutually-recursive-traits.rs @@ -0,0 +1,11 @@ +//@ [lean] known-failure +//@ [coq,fstar] skip +//@ subdir=misc +pub trait Trait1 { + type T: Trait2; +} + +pub trait Trait2: Trait1 {} + +pub trait T1>: Sized {} +pub trait T2>: Sized {} diff --git a/tests/src/string-chars.lean.out b/tests/src/string-chars.lean.out new file mode 100644 index 00000000..0d28744f --- /dev/null +++ b/tests/src/string-chars.lean.out @@ -0,0 +1,21 @@ +[Info ] Imported: tests/llbc/string_chars.llbc +[Error] In file SymbolicToPure.ml, line 588: +ADTs containing borrows are not supported yet +Source: '/rustc/ad963232d9b987d66a6f8e6ec4141f672b8b9900/library/core/src/str/iter.rs', lines 32:0-32:20 + +Uncaught exception: + + (Failure + "In file SymbolicToPure.ml, line 588:\ + \nIn file SymbolicToPure.ml, line 588:\ + \nADTs containing borrows are not supported yet\ + \nSource: '/rustc/ad963232d9b987d66a6f8e6ec4141f672b8b9900/library/core/src/str/iter.rs', lines 32:0-32:20\ + \nSource: '/rustc/ad963232d9b987d66a6f8e6ec4141f672b8b9900/library/core/src/str/iter.rs', lines 32:0-32:20") + +Raised at Aeneas__Errors.craise_opt_span in file "Errors.ml", line 46, characters 4-76 +Called from Aeneas__SymbolicToPure.translate_type_decl in file "SymbolicToPure.ml", line 588, characters 2-113 +Called from Aeneas__SymbolicToPure.translate_type_decls.(fun) in file "SymbolicToPure.ml", line 3923, characters 15-42 +Called from Stdlib__List.filter_map.aux in file "list.ml", line 258, characters 14-17 +Called from Aeneas__Translate.translate_crate_to_pure in file "Translate.ml", line 229, characters 19-64 +Called from Aeneas__Translate.translate_crate in file "Translate.ml", line 981, characters 4-33 +Called from Dune__exe__Main in file "Main.ml", line 276, characters 9-61 diff --git a/tests/src/string-chars.rs b/tests/src/string-chars.rs new file mode 100644 index 00000000..5ba9a2d6 --- /dev/null +++ b/tests/src/string-chars.rs @@ -0,0 +1,6 @@ +//@ [lean] known-failure +//@ [coq,fstar] skip +fn main() { + let s = "hello"; + let _chs: Vec = s.chars().collect(); +} diff --git a/tests/test_runner/dune b/tests/test_runner/dune index c38e009c..6bb3f7b2 100644 --- a/tests/test_runner/dune +++ b/tests/test_runner/dune @@ -1,6 +1,6 @@ (executable (public_name test_runner) - (libraries core_unix core_unix.sys_unix re str unix) + (libraries core_unix core_unix.filename_unix core_unix.sys_unix re str unix) (preprocess (pps ppx_deriving.show ppx_deriving.ord ppx_sexp_conv)) (name run_test)) diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index a5a89317..7a52b6f3 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -15,6 +15,7 @@ type runner_env = { charon_path : string; aeneas_path : string; llbc_dir : string; + dest_dir : string; } module Backend = struct @@ -61,10 +62,12 @@ end let concat_path = List.fold_left Filename.concat "" module Command = struct - type t = { args : string array } + type t = { args : string array; redirect_out : Unix.file_descr option } type status = Success | Failure - let make (args : string list) : t = { args = Array.of_list args } + let make (args : string list) : t = + { args = Array.of_list args; redirect_out = None } + let to_string (cmd : t) = Core.String.concat_array ~sep:" " cmd.args (* Run the command and returns its exit status. *) @@ -73,10 +76,8 @@ module Command = struct print_endline ("[test_runner] Running: " ^ command_str); (* Run the command *) - let pid = - Unix.create_process cmd.args.(0) cmd.args Unix.stdin Unix.stdout - Unix.stderr - in + let out = Option.value cmd.redirect_out ~default:Unix.stdout in + let pid = Unix.create_process cmd.args.(0) cmd.args Unix.stdin out out in let status = Core_unix.waitpid (Core.Pid.of_int pid) in match status with | Ok () -> Success @@ -89,6 +90,14 @@ module Command = struct match run cmd with | Success -> () | Failure -> failwith ("Command `" ^ to_string cmd ^ "` failed.") + + (* Run the command and aborts the program if the command succeeded. *) + let run_command_expecting_failure cmd = + match run cmd with + | Success -> + failwith + ("Command `" ^ to_string cmd ^ "` succeeded but was expected to fail.") + | Failure -> () end let aeneas_options_for_test backend test_name = @@ -232,18 +241,43 @@ end (* Run Aeneas on a specific input with the given options *) let run_aeneas (env : runner_env) (case : Input.t) (backend : Backend.t) = (* FIXME: remove this special case *) + let backend_str = Backend.to_string backend in let test_name = if case.name = "betree" then "betree_main" else case.name in let input_file = concat_path [ env.llbc_dir; test_name ] ^ ".llbc" in + let output_file = + Filename.remove_extension case.path ^ "." ^ backend_str ^ ".out" + in let subdir = BackendMap.find backend case.subdirs in + let dest_dir = concat_path [ env.dest_dir; backend_str; subdir ] in let aeneas_options = BackendMap.find backend case.aeneas_options in - let backend_str = Backend.to_string backend in - let dest_dir = concat_path [ "tests"; backend_str; subdir ] in + let action = BackendMap.find backend case.actions in + + (* Build the command *) let args = [ env.aeneas_path; input_file; "-dest"; dest_dir; "-backend"; backend_str ] in - let args = List.append args aeneas_options in + let abort_on_error = + match action with + | Skip | Normal -> [] + | KnownFailure -> [ "-abort-on-error" ] + in + let args = List.concat [ args; aeneas_options; abort_on_error ] in + let cmd = Command.make args in + (* Remove leftover files if they're not needed anymore *) + (match action with + | (Skip | Normal) when Sys.file_exists output_file -> Sys.remove output_file + | _ -> ()); (* Run Aeneas *) - Command.run_command_expecting_success (Command.make args) + match action with + | Skip -> () + | Normal -> Command.run_command_expecting_success cmd + | KnownFailure -> + let out = + Core_unix.openfile ~mode:[ O_CREAT; O_TRUNC; O_WRONLY ] output_file + in + let cmd = { cmd with redirect_out = Some out } in + Command.run_command_expecting_failure cmd; + Unix.close out (* Run Charon on a specific input with the given options *) let run_charon (env : runner_env) (case : Input.t) = @@ -267,7 +301,9 @@ let run_charon (env : runner_env) (case : Input.t) = | Crate -> ( match Sys.getenv_opt "IN_CI" with | None -> - let args = [ env.charon_path; "--dest"; env.llbc_dir ] in + let args = + [ env.charon_path; "--dest"; Filename_unix.realpath env.llbc_dir ] + in let args = List.append args case.charon_options in (* Run Charon inside the crate *) let old_pwd = Unix.getcwd () in @@ -287,7 +323,9 @@ let () = (* Ad-hoc argument passing for now. *) | _exe_path :: charon_path :: aeneas_path :: llbc_dir :: test_path :: aeneas_options -> - let runner_env = { charon_path; aeneas_path; llbc_dir } in + let runner_env = + { charon_path; aeneas_path; llbc_dir; dest_dir = "tests" } + in let test_case = Input.build test_path in let test_case = { @@ -307,11 +345,5 @@ let () = (* Generate the llbc file *) run_charon runner_env test_case; (* Process the llbc file for the each backend *) - List.iter - (fun backend -> - match BackendMap.find backend test_case.actions with - | Skip -> () - | Normal -> run_aeneas runner_env test_case backend - | KnownFailure -> failwith "KnownFailure is unimplemented") - Backend.all) + List.iter (run_aeneas runner_env test_case) Backend.all) | _ -> failwith "Incorrect options passed to test runner" -- cgit v1.2.3 From fb16d0af4caacd2c9a3463f6d2b455209b755697 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Mon, 27 May 2024 14:34:12 +0200 Subject: runner: Add `no-check-output` option for unstable outputs --- tests/src/string-chars.lean.out | 21 --------------------- tests/src/string-chars.rs | 1 + tests/test_runner/run_test.ml | 36 ++++++++++++++++++++++++++++++------ 3 files changed, 31 insertions(+), 27 deletions(-) delete mode 100644 tests/src/string-chars.lean.out diff --git a/tests/src/string-chars.lean.out b/tests/src/string-chars.lean.out deleted file mode 100644 index 0d28744f..00000000 --- a/tests/src/string-chars.lean.out +++ /dev/null @@ -1,21 +0,0 @@ -[Info ] Imported: tests/llbc/string_chars.llbc -[Error] In file SymbolicToPure.ml, line 588: -ADTs containing borrows are not supported yet -Source: '/rustc/ad963232d9b987d66a6f8e6ec4141f672b8b9900/library/core/src/str/iter.rs', lines 32:0-32:20 - -Uncaught exception: - - (Failure - "In file SymbolicToPure.ml, line 588:\ - \nIn file SymbolicToPure.ml, line 588:\ - \nADTs containing borrows are not supported yet\ - \nSource: '/rustc/ad963232d9b987d66a6f8e6ec4141f672b8b9900/library/core/src/str/iter.rs', lines 32:0-32:20\ - \nSource: '/rustc/ad963232d9b987d66a6f8e6ec4141f672b8b9900/library/core/src/str/iter.rs', lines 32:0-32:20") - -Raised at Aeneas__Errors.craise_opt_span in file "Errors.ml", line 46, characters 4-76 -Called from Aeneas__SymbolicToPure.translate_type_decl in file "SymbolicToPure.ml", line 588, characters 2-113 -Called from Aeneas__SymbolicToPure.translate_type_decls.(fun) in file "SymbolicToPure.ml", line 3923, characters 15-42 -Called from Stdlib__List.filter_map.aux in file "list.ml", line 258, characters 14-17 -Called from Aeneas__Translate.translate_crate_to_pure in file "Translate.ml", line 229, characters 19-64 -Called from Aeneas__Translate.translate_crate in file "Translate.ml", line 981, characters 4-33 -Called from Dune__exe__Main in file "Main.ml", line 276, characters 9-61 diff --git a/tests/src/string-chars.rs b/tests/src/string-chars.rs index 5ba9a2d6..f8490e76 100644 --- a/tests/src/string-chars.rs +++ b/tests/src/string-chars.rs @@ -1,5 +1,6 @@ //@ [lean] known-failure //@ [coq,fstar] skip +//@ no-check-output fn main() { let s = "hello"; let _chs: Vec = s.chars().collect(); diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index 7a52b6f3..1c885d4b 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -139,6 +139,8 @@ module Input = struct charon_options : string list; aeneas_options : string list BackendMap.t; subdirs : string BackendMap.t; + (* Whether to store the command output to a file. Only available for known-failure tests. *) + check_output : bool BackendMap.t; } (* The default subdirectory in which to store the outputs. *) @@ -182,6 +184,11 @@ module Input = struct input with actions = BackendMap.add_each backends KnownFailure input.actions; } + else if comment = "no-check-output" then + { + input with + check_output = BackendMap.add_each backends false input.check_output; + } else if Option.is_some charon_args then let args = Option.get charon_args in let args = String.split_on_char ' ' args in @@ -219,8 +226,18 @@ module Input = struct let aeneas_options = BackendMap.make (fun backend -> aeneas_options_for_test backend name) in + let check_output = BackendMap.make (fun _ -> true) in let input = - { path; name; kind; actions; charon_options; subdirs; aeneas_options } + { + path; + name; + kind; + actions; + charon_options; + subdirs; + aeneas_options; + check_output; + } in match kind with | SingleFile -> @@ -251,6 +268,7 @@ let run_aeneas (env : runner_env) (case : Input.t) (backend : Backend.t) = let dest_dir = concat_path [ env.dest_dir; backend_str; subdir ] in let aeneas_options = BackendMap.find backend case.aeneas_options in let action = BackendMap.find backend case.actions in + let check_output = BackendMap.find backend case.check_output in (* Build the command *) let args = @@ -264,21 +282,27 @@ let run_aeneas (env : runner_env) (case : Input.t) (backend : Backend.t) = let args = List.concat [ args; aeneas_options; abort_on_error ] in let cmd = Command.make args in (* Remove leftover files if they're not needed anymore *) - (match action with - | (Skip | Normal) when Sys.file_exists output_file -> Sys.remove output_file - | _ -> ()); + if + Sys.file_exists output_file + && + match action with + | Skip | Normal -> true + | KnownFailure when not check_output -> true + | _ -> false + then Sys.remove output_file; (* Run Aeneas *) match action with | Skip -> () | Normal -> Command.run_command_expecting_success cmd | KnownFailure -> let out = - Core_unix.openfile ~mode:[ O_CREAT; O_TRUNC; O_WRONLY ] output_file + if check_output then + Core_unix.openfile ~mode:[ O_CREAT; O_TRUNC; O_WRONLY ] output_file + else Core_unix.openfile ~mode:[ O_WRONLY ] "/dev/null" in let cmd = { cmd with redirect_out = Some out } in Command.run_command_expecting_failure cmd; Unix.close out - (* Run Charon on a specific input with the given options *) let run_charon (env : runner_env) (case : Input.t) = match case.kind with -- cgit v1.2.3 From 2b40c5c3de1ee2caca2c0072f812fea04b5a0238 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Mon, 27 May 2024 14:59:10 +0200 Subject: tests: Merge the hashmap test files --- tests/coq/hashmap/Hashmap_Funs.v | 72 +- tests/coq/hashmap/Hashmap_FunsExternal.v | 21 + tests/coq/hashmap/Hashmap_FunsExternal_Template.v | 23 + tests/coq/hashmap/Hashmap_Types.v | 6 +- tests/coq/hashmap/Hashmap_TypesExternal.v | 13 + tests/coq/hashmap/Hashmap_TypesExternal_Template.v | 15 + tests/coq/hashmap_main/HashmapMain_Funs.v | 589 ------------- tests/coq/hashmap_main/HashmapMain_FunsExternal.v | 24 - .../HashmapMain_FunsExternal_Template.v | 26 - tests/coq/hashmap_main/HashmapMain_Types.v | 40 - tests/coq/hashmap_main/HashmapMain_TypesExternal.v | 14 - .../HashmapMain_TypesExternal_Template.v | 15 - tests/coq/hashmap_main/Makefile | 23 - tests/coq/hashmap_main/Primitives.v | 981 --------------------- tests/coq/hashmap_main/_CoqProject | 12 - tests/fstar/hashmap/Hashmap.Clauses.Template.fst | 18 +- tests/fstar/hashmap/Hashmap.Funs.fst | 69 +- tests/fstar/hashmap/Hashmap.FunsExternal.fsti | 16 + tests/fstar/hashmap/Hashmap.Properties.fst | 48 + tests/fstar/hashmap/Hashmap.Types.fst | 5 +- tests/fstar/hashmap/Hashmap.TypesExternal.fsti | 10 + .../hashmap_main/HashmapMain.Clauses.Template.fst | 72 -- tests/fstar/hashmap_main/HashmapMain.Clauses.fst | 61 -- tests/fstar/hashmap_main/HashmapMain.Funs.fst | 446 ---------- .../hashmap_main/HashmapMain.FunsExternal.fsti | 18 - .../fstar/hashmap_main/HashmapMain.Properties.fst | 48 - tests/fstar/hashmap_main/HashmapMain.Types.fst | 24 - .../hashmap_main/HashmapMain.TypesExternal.fsti | 10 - tests/fstar/hashmap_main/Makefile | 49 - tests/fstar/hashmap_main/Primitives.fst | 929 ------------------- tests/lean/Hashmap/Funs.lean | 70 +- tests/lean/Hashmap/FunsExternal.lean | 19 + tests/lean/Hashmap/FunsExternal_Template.lean | 17 + tests/lean/Hashmap/Types.lean | 5 +- tests/lean/Hashmap/TypesExternal.lean | 7 + tests/lean/Hashmap/TypesExternal_Template.lean | 9 + tests/lean/HashmapMain.lean | 1 - tests/lean/HashmapMain/Funs.lean | 463 ---------- tests/lean/HashmapMain/FunsExternal.lean | 17 - tests/lean/HashmapMain/FunsExternal_Template.lean | 18 - tests/lean/HashmapMain/Opaque.lean | 18 - tests/lean/HashmapMain/Types.lean | 23 - tests/lean/HashmapMain/TypesExternal.lean | 7 - tests/lean/HashmapMain/TypesExternal_Template.lean | 9 - tests/lean/lakefile.lean | 1 - tests/src/hashmap.rs | 29 +- tests/src/hashmap_main.rs | 22 - tests/src/hashmap_utils.rs | 13 - 48 files changed, 366 insertions(+), 4079 deletions(-) create mode 100644 tests/coq/hashmap/Hashmap_FunsExternal.v create mode 100644 tests/coq/hashmap/Hashmap_FunsExternal_Template.v create mode 100644 tests/coq/hashmap/Hashmap_TypesExternal.v create mode 100644 tests/coq/hashmap/Hashmap_TypesExternal_Template.v delete mode 100644 tests/coq/hashmap_main/HashmapMain_Funs.v delete mode 100644 tests/coq/hashmap_main/HashmapMain_FunsExternal.v delete mode 100644 tests/coq/hashmap_main/HashmapMain_FunsExternal_Template.v delete mode 100644 tests/coq/hashmap_main/HashmapMain_Types.v delete mode 100644 tests/coq/hashmap_main/HashmapMain_TypesExternal.v delete mode 100644 tests/coq/hashmap_main/HashmapMain_TypesExternal_Template.v delete mode 100644 tests/coq/hashmap_main/Makefile delete mode 100644 tests/coq/hashmap_main/Primitives.v delete mode 100644 tests/coq/hashmap_main/_CoqProject create mode 100644 tests/fstar/hashmap/Hashmap.FunsExternal.fsti create mode 100644 tests/fstar/hashmap/Hashmap.Properties.fst create mode 100644 tests/fstar/hashmap/Hashmap.TypesExternal.fsti delete mode 100644 tests/fstar/hashmap_main/HashmapMain.Clauses.Template.fst delete mode 100644 tests/fstar/hashmap_main/HashmapMain.Clauses.fst delete mode 100644 tests/fstar/hashmap_main/HashmapMain.Funs.fst delete mode 100644 tests/fstar/hashmap_main/HashmapMain.FunsExternal.fsti delete mode 100644 tests/fstar/hashmap_main/HashmapMain.Properties.fst delete mode 100644 tests/fstar/hashmap_main/HashmapMain.Types.fst delete mode 100644 tests/fstar/hashmap_main/HashmapMain.TypesExternal.fsti delete mode 100644 tests/fstar/hashmap_main/Makefile delete mode 100644 tests/fstar/hashmap_main/Primitives.fst create mode 100644 tests/lean/Hashmap/FunsExternal.lean create mode 100644 tests/lean/Hashmap/FunsExternal_Template.lean create mode 100644 tests/lean/Hashmap/TypesExternal.lean create mode 100644 tests/lean/Hashmap/TypesExternal_Template.lean delete mode 100644 tests/lean/HashmapMain.lean delete mode 100644 tests/lean/HashmapMain/Funs.lean delete mode 100644 tests/lean/HashmapMain/FunsExternal.lean delete mode 100644 tests/lean/HashmapMain/FunsExternal_Template.lean delete mode 100644 tests/lean/HashmapMain/Opaque.lean delete mode 100644 tests/lean/HashmapMain/Types.lean delete mode 100644 tests/lean/HashmapMain/TypesExternal.lean delete mode 100644 tests/lean/HashmapMain/TypesExternal_Template.lean delete mode 100644 tests/src/hashmap_main.rs delete mode 100644 tests/src/hashmap_utils.rs diff --git a/tests/coq/hashmap/Hashmap_Funs.v b/tests/coq/hashmap/Hashmap_Funs.v index 778b9d56..89f90210 100644 --- a/tests/coq/hashmap/Hashmap_Funs.v +++ b/tests/coq/hashmap/Hashmap_Funs.v @@ -8,15 +8,17 @@ Import ListNotations. Local Open Scope Primitives_scope. Require Import Hashmap_Types. Include Hashmap_Types. +Require Import Hashmap_FunsExternal. +Include Hashmap_FunsExternal. Module Hashmap_Funs. (** [hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 35:0-35:32 *) + Source: 'tests/src/hashmap.rs', lines 37:0-37:32 *) Definition hash_key (k : usize) : result usize := Ok k. (** [hashmap::{hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) + Source: 'tests/src/hashmap.rs', lines 60:4-66: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)) @@ -34,7 +36,7 @@ Fixpoint hashMap_allocate_slots_loop . (** [hashmap::{hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 58:4-58:76 *) + Source: 'tests/src/hashmap.rs', lines 60:4-60:76 *) Definition hashMap_allocate_slots (T : Type) (n : nat) (slots : alloc_vec_Vec (List_t T)) (n1 : usize) : result (alloc_vec_Vec (List_t T)) @@ -43,7 +45,7 @@ Definition hashMap_allocate_slots . (** [hashmap::{hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 67:4-71:13 *) + Source: 'tests/src/hashmap.rs', lines 69:4-73:13 *) Definition hashMap_new_with_capacity (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : @@ -62,13 +64,13 @@ Definition hashMap_new_with_capacity . (** [hashmap::{hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 83:4-83:24 *) + Source: 'tests/src/hashmap.rs', lines 85:4-85:24 *) Definition hashMap_new (T : Type) (n : nat) : result (HashMap_t T) := hashMap_new_with_capacity T n 32%usize 4%usize 5%usize . (** [hashmap::{hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) + Source: 'tests/src/hashmap.rs', lines 90:4-98: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)) @@ -91,7 +93,7 @@ Fixpoint hashMap_clear_loop . (** [hashmap::{hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 88:4-88:27 *) + Source: 'tests/src/hashmap.rs', lines 90:4-90: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; @@ -105,13 +107,13 @@ Definition hashMap_clear . (** [hashmap::{hashmap::HashMap}::len]: - Source: 'tests/src/hashmap.rs', lines 98:4-98:30 *) + Source: 'tests/src/hashmap.rs', lines 100:4-100:30 *) Definition hashMap_len (T : Type) (self : HashMap_t T) : result usize := Ok self.(hashMap_num_entries) . (** [hashmap::{hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) + Source: 'tests/src/hashmap.rs', lines 107:4-124:5 *) Fixpoint hashMap_insert_in_list_loop (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : result (bool * (List_t T)) @@ -133,7 +135,7 @@ Fixpoint hashMap_insert_in_list_loop . (** [hashmap::{hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 105:4-105:71 *) + Source: 'tests/src/hashmap.rs', lines 107:4-107:71 *) Definition hashMap_insert_in_list (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : result (bool * (List_t T)) @@ -142,7 +144,7 @@ Definition hashMap_insert_in_list . (** [hashmap::{hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 125:4-125:54 *) + Source: 'tests/src/hashmap.rs', lines 127:4-127:54 *) Definition hashMap_insert_no_resize (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) : result (HashMap_t T) @@ -180,7 +182,7 @@ Definition hashMap_insert_no_resize . (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) + Source: 'tests/src/hashmap.rs', lines 193:4-206:5 *) Fixpoint hashMap_move_elements_from_list_loop (T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) : result (HashMap_t T) @@ -198,7 +200,7 @@ Fixpoint hashMap_move_elements_from_list_loop . (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 191:4-191:72 *) + Source: 'tests/src/hashmap.rs', lines 193:4-193:72 *) Definition hashMap_move_elements_from_list (T : Type) (n : nat) (ntable : HashMap_t T) (ls : List_t T) : result (HashMap_t T) @@ -207,7 +209,7 @@ Definition hashMap_move_elements_from_list . (** [hashmap::{hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) + Source: 'tests/src/hashmap.rs', lines 181:4-190:5 *) Fixpoint hashMap_move_elements_loop (T : Type) (n : nat) (ntable : HashMap_t T) (slots : alloc_vec_Vec (List_t T)) (i : usize) : @@ -233,7 +235,7 @@ Fixpoint hashMap_move_elements_loop . (** [hashmap::{hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 179:4-179:95 *) + Source: 'tests/src/hashmap.rs', lines 181:4-181:95 *) Definition hashMap_move_elements (T : Type) (n : nat) (ntable : HashMap_t T) (slots : alloc_vec_Vec (List_t T)) (i : usize) : @@ -243,7 +245,7 @@ Definition hashMap_move_elements . (** [hashmap::{hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 148:4-148:28 *) + Source: 'tests/src/hashmap.rs', lines 150:4-150: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; @@ -275,7 +277,7 @@ Definition hashMap_try_resize . (** [hashmap::{hashmap::HashMap}::insert]: - Source: 'tests/src/hashmap.rs', lines 137:4-137:48 *) + Source: 'tests/src/hashmap.rs', lines 139:4-139:48 *) Definition hashMap_insert (T : Type) (n : nat) (self : HashMap_t T) (key : usize) (value : T) : result (HashMap_t T) @@ -288,7 +290,7 @@ Definition hashMap_insert . (** [hashmap::{hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) + Source: 'tests/src/hashmap.rs', lines 216:4-229:5 *) Fixpoint hashMap_contains_key_in_list_loop (T : Type) (n : nat) (key : usize) (ls : List_t T) : result bool := match n with @@ -305,14 +307,14 @@ Fixpoint hashMap_contains_key_in_list_loop . (** [hashmap::{hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 214:4-214:68 *) + Source: 'tests/src/hashmap.rs', lines 216:4-216:68 *) Definition hashMap_contains_key_in_list (T : Type) (n : nat) (key : usize) (ls : List_t T) : result bool := hashMap_contains_key_in_list_loop T n key ls . (** [hashmap::{hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 207:4-207:49 *) + Source: 'tests/src/hashmap.rs', lines 209:4-209:49 *) Definition hashMap_contains_key (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result bool := hash <- hash_key key; @@ -326,7 +328,7 @@ Definition hashMap_contains_key . (** [hashmap::{hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) + Source: 'tests/src/hashmap.rs', lines 234:4-247:5 *) Fixpoint hashMap_get_in_list_loop (T : Type) (n : nat) (key : usize) (ls : List_t T) : result T := match n with @@ -341,14 +343,14 @@ Fixpoint hashMap_get_in_list_loop . (** [hashmap::{hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 232:4-232:70 *) + Source: 'tests/src/hashmap.rs', lines 234:4-234:70 *) Definition hashMap_get_in_list (T : Type) (n : nat) (key : usize) (ls : List_t T) : result T := hashMap_get_in_list_loop T n key ls . (** [hashmap::{hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 247:4-247:55 *) + Source: 'tests/src/hashmap.rs', lines 249:4-249:55 *) Definition hashMap_get (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result T := hash <- hash_key key; @@ -362,7 +364,7 @@ Definition hashMap_get . (** [hashmap::{hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) + Source: 'tests/src/hashmap.rs', lines 255:4-264: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))) @@ -388,7 +390,7 @@ Fixpoint hashMap_get_mut_in_list_loop . (** [hashmap::{hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 253:4-253:86 *) + Source: 'tests/src/hashmap.rs', lines 255:4-255:86 *) Definition hashMap_get_mut_in_list (T : Type) (n : nat) (ls : List_t T) (key : usize) : result (T * (T -> result (List_t T))) @@ -397,7 +399,7 @@ Definition hashMap_get_mut_in_list . (** [hashmap::{hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:67 *) + Source: 'tests/src/hashmap.rs', lines 267:4-267:67 *) Definition hashMap_get_mut (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result (T * (T -> result (HashMap_t T))) @@ -427,7 +429,7 @@ Definition hashMap_get_mut . (** [hashmap::{hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) + Source: 'tests/src/hashmap.rs', lines 275:4-301:5 *) Fixpoint hashMap_remove_from_list_loop (T : Type) (n : nat) (key : usize) (ls : List_t T) : result ((option T) * (List_t T)) @@ -455,7 +457,7 @@ Fixpoint hashMap_remove_from_list_loop . (** [hashmap::{hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 273:4-273:69 *) + Source: 'tests/src/hashmap.rs', lines 275:4-275:69 *) Definition hashMap_remove_from_list (T : Type) (n : nat) (key : usize) (ls : List_t T) : result ((option T) * (List_t T)) @@ -464,7 +466,7 @@ Definition hashMap_remove_from_list . (** [hashmap::{hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 302:4-302:52 *) + Source: 'tests/src/hashmap.rs', lines 304:4-304:52 *) Definition hashMap_remove (T : Type) (n : nat) (self : HashMap_t T) (key : usize) : result ((option T) * (HashMap_t T)) @@ -502,8 +504,18 @@ Definition hashMap_remove end . +(** [hashmap::insert_on_disk]: + Source: 'tests/src/hashmap.rs', lines 335:0-335:43 *) +Definition insert_on_disk + (n : nat) (key : usize) (value : u64) (st : state) : result (state * unit) := + p <- hashmap_utils_deserialize st; + let (st1, hm) := p in + hm1 <- hashMap_insert u64 n hm key value; + hashmap_utils_serialize hm1 st1 +. + (** [hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 323:0-323:10 *) + Source: 'tests/src/hashmap.rs', lines 350:0-350: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.v b/tests/coq/hashmap/Hashmap_FunsExternal.v new file mode 100644 index 00000000..bcf12c17 --- /dev/null +++ b/tests/coq/hashmap/Hashmap_FunsExternal.v @@ -0,0 +1,21 @@ +(** [hashmap]: external functions. *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Require Export Hashmap_Types. +Import Hashmap_Types. +Module Hashmap_FunsExternal. + +(** [hashmap::hashmap_utils::deserialize]: + Source: 'tests/src/hashmap.rs', lines 330:4-330:47 *) +Axiom hashmap_utils_deserialize : state -> result (state * (HashMap_t u64)). + +(** [hashmap::hashmap_utils::serialize]: + Source: 'tests/src/hashmap.rs', lines 325:4-325:46 *) +Axiom hashmap_utils_serialize : HashMap_t u64 -> state -> result (state * unit) +. + +End Hashmap_FunsExternal. diff --git a/tests/coq/hashmap/Hashmap_FunsExternal_Template.v b/tests/coq/hashmap/Hashmap_FunsExternal_Template.v new file mode 100644 index 00000000..c35b8e14 --- /dev/null +++ b/tests/coq/hashmap/Hashmap_FunsExternal_Template.v @@ -0,0 +1,23 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap]: external functions. +-- This is a template file: rename it to "FunsExternal.lean" and fill the holes. *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Require Import Hashmap_Types. +Include Hashmap_Types. +Module Hashmap_FunsExternal_Template. + +(** [hashmap::hashmap_utils::deserialize]: + Source: 'tests/src/hashmap.rs', lines 330:4-330:47 *) +Axiom hashmap_utils_deserialize : state -> result (state * (HashMap_t u64)). + +(** [hashmap::hashmap_utils::serialize]: + Source: 'tests/src/hashmap.rs', lines 325:4-325:46 *) +Axiom hashmap_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 8a8137d5..452d56f8 100644 --- a/tests/coq/hashmap/Hashmap_Types.v +++ b/tests/coq/hashmap/Hashmap_Types.v @@ -6,10 +6,12 @@ Require Import Coq.ZArith.ZArith. Require Import List. Import ListNotations. Local Open Scope Primitives_scope. +Require Import Hashmap_TypesExternal. +Include Hashmap_TypesExternal. Module Hashmap_Types. (** [hashmap::List] - Source: 'tests/src/hashmap.rs', lines 27:0-27:16 *) + 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 @@ -19,7 +21,7 @@ Arguments List_Cons { _ }. Arguments List_Nil { _ }. (** [hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 43:0-43:21 *) + Source: 'tests/src/hashmap.rs', lines 45:0-45:21 *) Record HashMap_t (T : Type) := mkHashMap_t { hashMap_num_entries : usize; diff --git a/tests/coq/hashmap/Hashmap_TypesExternal.v b/tests/coq/hashmap/Hashmap_TypesExternal.v new file mode 100644 index 00000000..0aa52a5d --- /dev/null +++ b/tests/coq/hashmap/Hashmap_TypesExternal.v @@ -0,0 +1,13 @@ +(** [hashmap]: external types. *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Module Hashmap_TypesExternal. + +(** The state type used in the state-error monad *) +Axiom state : Type. + +End Hashmap_TypesExternal. diff --git a/tests/coq/hashmap/Hashmap_TypesExternal_Template.v b/tests/coq/hashmap/Hashmap_TypesExternal_Template.v new file mode 100644 index 00000000..a8051adb --- /dev/null +++ b/tests/coq/hashmap/Hashmap_TypesExternal_Template.v @@ -0,0 +1,15 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap]: external types. +-- This is a template file: rename it to "TypesExternal.lean" and fill the holes. *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Module Hashmap_TypesExternal_Template. + +(** The state type used in the state-error monad *) +Axiom state : Type. + +End Hashmap_TypesExternal_Template. diff --git a/tests/coq/hashmap_main/HashmapMain_Funs.v b/tests/coq/hashmap_main/HashmapMain_Funs.v deleted file mode 100644 index f6467d5a..00000000 --- a/tests/coq/hashmap_main/HashmapMain_Funs.v +++ /dev/null @@ -1,589 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: function definitions *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Require Import HashmapMain_Types. -Include HashmapMain_Types. -Require Import HashmapMain_FunsExternal. -Include HashmapMain_FunsExternal. -Module HashmapMain_Funs. - -(** [hashmap_main::hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 35:0-35:32 *) -Definition hashmap_hash_key (k : usize) : result usize := - Ok k. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) -Fixpoint hashmap_HashMap_allocate_slots_loop - (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (n1 : usize) - : - result (alloc_vec_Vec (hashmap_List_t T)) - := - match n with - | O => Fail_ OutOfFuel - | S n2 => - if n1 s> 0%usize - then ( - slots1 <- alloc_vec_Vec_push (hashmap_List_t T) slots Hashmap_List_Nil; - n3 <- usize_sub n1 1%usize; - hashmap_HashMap_allocate_slots_loop T n2 slots1 n3) - else Ok slots - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 58:4-58:76 *) -Definition hashmap_HashMap_allocate_slots - (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (n1 : usize) - : - result (alloc_vec_Vec (hashmap_List_t T)) - := - hashmap_HashMap_allocate_slots_loop T n slots n1 -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 67:4-71:13 *) -Definition hashmap_HashMap_new_with_capacity - (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) - (max_load_divisor : usize) : - result (hashmap_HashMap_t T) - := - slots <- - hashmap_HashMap_allocate_slots T n (alloc_vec_Vec_new (hashmap_List_t T)) - capacity; - i <- usize_mul capacity max_load_dividend; - i1 <- usize_div i max_load_divisor; - Ok - {| - hashmap_HashMap_num_entries := 0%usize; - hashmap_HashMap_max_load_factor := (max_load_dividend, max_load_divisor); - hashmap_HashMap_max_load := i1; - hashmap_HashMap_slots := slots - |} -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 83:4-83:24 *) -Definition hashmap_HashMap_new - (T : Type) (n : nat) : result (hashmap_HashMap_t T) := - hashmap_HashMap_new_with_capacity T n 32%usize 4%usize 5%usize -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) -Fixpoint hashmap_HashMap_clear_loop - (T : Type) (n : nat) (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) : - result (alloc_vec_Vec (hashmap_List_t T)) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - let i1 := alloc_vec_Vec_len (hashmap_List_t T) slots in - if i s< i1 - then ( - p <- - alloc_vec_Vec_index_mut (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) slots - i; - let (_, index_mut_back) := p in - i2 <- usize_add i 1%usize; - slots1 <- index_mut_back Hashmap_List_Nil; - hashmap_HashMap_clear_loop T n1 slots1 i2) - else Ok slots - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 88:4-88:27 *) -Definition hashmap_HashMap_clear - (T : Type) (n : nat) (self : hashmap_HashMap_t T) : - result (hashmap_HashMap_t T) - := - hm <- hashmap_HashMap_clear_loop T n self.(hashmap_HashMap_slots) 0%usize; - Ok - {| - hashmap_HashMap_num_entries := 0%usize; - hashmap_HashMap_max_load_factor := self.(hashmap_HashMap_max_load_factor); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := hm - |} -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::len]: - Source: 'tests/src/hashmap.rs', lines 98:4-98:30 *) -Definition hashmap_HashMap_len - (T : Type) (self : hashmap_HashMap_t T) : result usize := - Ok self.(hashmap_HashMap_num_entries) -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) -Fixpoint hashmap_HashMap_insert_in_list_loop - (T : Type) (n : nat) (key : usize) (value : T) (ls : hashmap_List_t T) : - result (bool * (hashmap_List_t T)) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match ls with - | Hashmap_List_Cons ckey cvalue tl => - if ckey s= key - then Ok (false, Hashmap_List_Cons ckey value tl) - else ( - p <- hashmap_HashMap_insert_in_list_loop T n1 key value tl; - let (b, tl1) := p in - Ok (b, Hashmap_List_Cons ckey cvalue tl1)) - | Hashmap_List_Nil => - Ok (true, Hashmap_List_Cons key value Hashmap_List_Nil) - end - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 105:4-105:71 *) -Definition hashmap_HashMap_insert_in_list - (T : Type) (n : nat) (key : usize) (value : T) (ls : hashmap_List_t T) : - result (bool * (hashmap_List_t T)) - := - hashmap_HashMap_insert_in_list_loop T n key value ls -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 125:4-125:54 *) -Definition hashmap_HashMap_insert_no_resize - (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) (value : T) : - result (hashmap_HashMap_t T) - := - hash <- hashmap_hash_key key; - let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in - hash_mod <- usize_rem hash i; - p <- - alloc_vec_Vec_index_mut (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) - self.(hashmap_HashMap_slots) hash_mod; - let (l, index_mut_back) := p in - p1 <- hashmap_HashMap_insert_in_list T n key value l; - let (inserted, l1) := p1 in - if inserted - then ( - i1 <- usize_add self.(hashmap_HashMap_num_entries) 1%usize; - v <- index_mut_back l1; - Ok - {| - hashmap_HashMap_num_entries := i1; - hashmap_HashMap_max_load_factor := - self.(hashmap_HashMap_max_load_factor); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := v - |}) - else ( - v <- index_mut_back l1; - Ok - {| - hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); - hashmap_HashMap_max_load_factor := - self.(hashmap_HashMap_max_load_factor); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := v - |}) -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) -Fixpoint hashmap_HashMap_move_elements_from_list_loop - (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) (ls : hashmap_List_t T) : - result (hashmap_HashMap_t T) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match ls with - | Hashmap_List_Cons k v tl => - ntable1 <- hashmap_HashMap_insert_no_resize T n1 ntable k v; - hashmap_HashMap_move_elements_from_list_loop T n1 ntable1 tl - | Hashmap_List_Nil => Ok ntable - end - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 191:4-191:72 *) -Definition hashmap_HashMap_move_elements_from_list - (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) (ls : hashmap_List_t T) : - result (hashmap_HashMap_t T) - := - hashmap_HashMap_move_elements_from_list_loop T n ntable ls -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) -Fixpoint hashmap_HashMap_move_elements_loop - (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) - (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) : - result ((hashmap_HashMap_t T) * (alloc_vec_Vec (hashmap_List_t T))) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - let i1 := alloc_vec_Vec_len (hashmap_List_t T) slots in - if i s< i1 - then ( - p <- - alloc_vec_Vec_index_mut (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) slots - i; - let (l, index_mut_back) := p in - let (ls, l1) := core_mem_replace (hashmap_List_t T) l Hashmap_List_Nil in - ntable1 <- hashmap_HashMap_move_elements_from_list T n1 ntable ls; - i2 <- usize_add i 1%usize; - slots1 <- index_mut_back l1; - hashmap_HashMap_move_elements_loop T n1 ntable1 slots1 i2) - else Ok (ntable, slots) - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 179:4-179:95 *) -Definition hashmap_HashMap_move_elements - (T : Type) (n : nat) (ntable : hashmap_HashMap_t T) - (slots : alloc_vec_Vec (hashmap_List_t T)) (i : usize) : - result ((hashmap_HashMap_t T) * (alloc_vec_Vec (hashmap_List_t T))) - := - hashmap_HashMap_move_elements_loop T n ntable slots i -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 148:4-148:28 *) -Definition hashmap_HashMap_try_resize - (T : Type) (n : nat) (self : hashmap_HashMap_t T) : - result (hashmap_HashMap_t T) - := - max_usize <- scalar_cast U32 Usize core_u32_max; - let capacity := - alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in - n1 <- usize_div max_usize 2%usize; - let (i, i1) := self.(hashmap_HashMap_max_load_factor) in - i2 <- usize_div n1 i; - if capacity s<= i2 - then ( - i3 <- usize_mul capacity 2%usize; - ntable <- hashmap_HashMap_new_with_capacity T n i3 i i1; - p <- - hashmap_HashMap_move_elements T n ntable self.(hashmap_HashMap_slots) - 0%usize; - let (ntable1, _) := p in - Ok - {| - hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); - hashmap_HashMap_max_load_factor := (i, i1); - hashmap_HashMap_max_load := ntable1.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := ntable1.(hashmap_HashMap_slots) - |}) - else - Ok - {| - hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); - hashmap_HashMap_max_load_factor := (i, i1); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := self.(hashmap_HashMap_slots) - |} -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert]: - Source: 'tests/src/hashmap.rs', lines 137:4-137:48 *) -Definition hashmap_HashMap_insert - (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) (value : T) : - result (hashmap_HashMap_t T) - := - self1 <- hashmap_HashMap_insert_no_resize T n self key value; - i <- hashmap_HashMap_len T self1; - if i s> self1.(hashmap_HashMap_max_load) - then hashmap_HashMap_try_resize T n self1 - else Ok self1 -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) -Fixpoint hashmap_HashMap_contains_key_in_list_loop - (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result bool := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match ls with - | Hashmap_List_Cons ckey _ tl => - if ckey s= key - then Ok true - else hashmap_HashMap_contains_key_in_list_loop T n1 key tl - | Hashmap_List_Nil => Ok false - end - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 214:4-214:68 *) -Definition hashmap_HashMap_contains_key_in_list - (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result bool := - hashmap_HashMap_contains_key_in_list_loop T n key ls -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 207:4-207:49 *) -Definition hashmap_HashMap_contains_key - (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : - result bool - := - hash <- hashmap_hash_key key; - let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in - hash_mod <- usize_rem hash i; - l <- - alloc_vec_Vec_index (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) - self.(hashmap_HashMap_slots) hash_mod; - hashmap_HashMap_contains_key_in_list T n key l -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) -Fixpoint hashmap_HashMap_get_in_list_loop - (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result T := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match ls with - | Hashmap_List_Cons ckey cvalue tl => - if ckey s= key - then Ok cvalue - else hashmap_HashMap_get_in_list_loop T n1 key tl - | Hashmap_List_Nil => Fail_ Failure - end - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 232:4-232:70 *) -Definition hashmap_HashMap_get_in_list - (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : result T := - hashmap_HashMap_get_in_list_loop T n key ls -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 247:4-247:55 *) -Definition hashmap_HashMap_get - (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : result T := - hash <- hashmap_hash_key key; - let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in - hash_mod <- usize_rem hash i; - l <- - alloc_vec_Vec_index (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) - self.(hashmap_HashMap_slots) hash_mod; - hashmap_HashMap_get_in_list T n key l -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) -Fixpoint hashmap_HashMap_get_mut_in_list_loop - (T : Type) (n : nat) (ls : hashmap_List_t T) (key : usize) : - result (T * (T -> result (hashmap_List_t T))) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match ls with - | Hashmap_List_Cons ckey cvalue tl => - if ckey s= key - then - let back := fun (ret : T) => Ok (Hashmap_List_Cons ckey ret tl) in - Ok (cvalue, back) - else ( - p <- hashmap_HashMap_get_mut_in_list_loop T n1 tl key; - let (t, back) := p in - let back1 := - fun (ret : T) => - tl1 <- back ret; Ok (Hashmap_List_Cons ckey cvalue tl1) in - Ok (t, back1)) - | Hashmap_List_Nil => Fail_ Failure - end - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 253:4-253:86 *) -Definition hashmap_HashMap_get_mut_in_list - (T : Type) (n : nat) (ls : hashmap_List_t T) (key : usize) : - result (T * (T -> result (hashmap_List_t T))) - := - hashmap_HashMap_get_mut_in_list_loop T n ls key -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:67 *) -Definition hashmap_HashMap_get_mut - (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : - result (T * (T -> result (hashmap_HashMap_t T))) - := - hash <- hashmap_hash_key key; - let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in - hash_mod <- usize_rem hash i; - p <- - alloc_vec_Vec_index_mut (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) - self.(hashmap_HashMap_slots) hash_mod; - let (l, index_mut_back) := p in - p1 <- hashmap_HashMap_get_mut_in_list T n l 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; - Ok - {| - hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); - hashmap_HashMap_max_load_factor := - self.(hashmap_HashMap_max_load_factor); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := v - |} in - Ok (t, back) -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) -Fixpoint hashmap_HashMap_remove_from_list_loop - (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : - result ((option T) * (hashmap_List_t T)) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match ls with - | Hashmap_List_Cons ckey t tl => - if ckey s= key - then - let (mv_ls, _) := - core_mem_replace (hashmap_List_t T) (Hashmap_List_Cons ckey t tl) - Hashmap_List_Nil in - match mv_ls with - | Hashmap_List_Cons _ cvalue tl1 => Ok (Some cvalue, tl1) - | Hashmap_List_Nil => Fail_ Failure - end - else ( - p <- hashmap_HashMap_remove_from_list_loop T n1 key tl; - let (o, tl1) := p in - Ok (o, Hashmap_List_Cons ckey t tl1)) - | Hashmap_List_Nil => Ok (None, Hashmap_List_Nil) - end - end -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 273:4-273:69 *) -Definition hashmap_HashMap_remove_from_list - (T : Type) (n : nat) (key : usize) (ls : hashmap_List_t T) : - result ((option T) * (hashmap_List_t T)) - := - hashmap_HashMap_remove_from_list_loop T n key ls -. - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 302:4-302:52 *) -Definition hashmap_HashMap_remove - (T : Type) (n : nat) (self : hashmap_HashMap_t T) (key : usize) : - result ((option T) * (hashmap_HashMap_t T)) - := - hash <- hashmap_hash_key key; - let i := alloc_vec_Vec_len (hashmap_List_t T) self.(hashmap_HashMap_slots) in - hash_mod <- usize_rem hash i; - p <- - alloc_vec_Vec_index_mut (hashmap_List_t T) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t T)) - self.(hashmap_HashMap_slots) hash_mod; - let (l, index_mut_back) := p in - p1 <- hashmap_HashMap_remove_from_list T n key l; - let (x, l1) := p1 in - match x with - | None => - v <- index_mut_back l1; - Ok (None, - {| - hashmap_HashMap_num_entries := self.(hashmap_HashMap_num_entries); - hashmap_HashMap_max_load_factor := - self.(hashmap_HashMap_max_load_factor); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := v - |}) - | Some x1 => - i1 <- usize_sub self.(hashmap_HashMap_num_entries) 1%usize; - v <- index_mut_back l1; - Ok (Some x1, - {| - hashmap_HashMap_num_entries := i1; - hashmap_HashMap_max_load_factor := - self.(hashmap_HashMap_max_load_factor); - hashmap_HashMap_max_load := self.(hashmap_HashMap_max_load); - hashmap_HashMap_slots := v - |}) - end -. - -(** [hashmap_main::hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 323:0-323:10 *) -Definition hashmap_test1 (n : nat) : result unit := - hm <- hashmap_HashMap_new u64 n; - hm1 <- hashmap_HashMap_insert u64 n hm 0%usize 42%u64; - hm2 <- hashmap_HashMap_insert u64 n hm1 128%usize 18%u64; - hm3 <- hashmap_HashMap_insert u64 n hm2 1024%usize 138%u64; - hm4 <- hashmap_HashMap_insert u64 n hm3 1056%usize 256%u64; - i <- hashmap_HashMap_get u64 n hm4 128%usize; - if negb (i s= 18%u64) - then Fail_ Failure - else ( - p <- hashmap_HashMap_get_mut u64 n hm4 1024%usize; - let (_, get_mut_back) := p in - hm5 <- get_mut_back 56%u64; - i1 <- hashmap_HashMap_get u64 n hm5 1024%usize; - if negb (i1 s= 56%u64) - then Fail_ Failure - else ( - p1 <- hashmap_HashMap_remove u64 n hm5 1024%usize; - let (x, hm6) := p1 in - match x with - | None => Fail_ Failure - | Some x1 => - if negb (x1 s= 56%u64) - then Fail_ Failure - else ( - i2 <- hashmap_HashMap_get u64 n hm6 0%usize; - if negb (i2 s= 42%u64) - then Fail_ Failure - else ( - i3 <- hashmap_HashMap_get u64 n hm6 128%usize; - if negb (i3 s= 18%u64) - then Fail_ Failure - else ( - i4 <- hashmap_HashMap_get u64 n hm6 1056%usize; - if negb (i4 s= 256%u64) then Fail_ Failure else Ok tt))) - end)) -. - -(** [hashmap_main::insert_on_disk]: - Source: 'tests/src/hashmap_main.rs', lines 13:0-13:43 *) -Definition insert_on_disk - (n : nat) (key : usize) (value : u64) (st : state) : result (state * unit) := - p <- hashmap_utils_deserialize st; - let (st1, hm) := p in - hm1 <- hashmap_HashMap_insert u64 n hm key value; - hashmap_utils_serialize hm1 st1 -. - -(** [hashmap_main::main]: - Source: 'tests/src/hashmap_main.rs', lines 22:0-22:13 *) -Definition main : result unit := - Ok tt. - -End HashmapMain_Funs. diff --git a/tests/coq/hashmap_main/HashmapMain_FunsExternal.v b/tests/coq/hashmap_main/HashmapMain_FunsExternal.v deleted file mode 100644 index fb5f23cd..00000000 --- a/tests/coq/hashmap_main/HashmapMain_FunsExternal.v +++ /dev/null @@ -1,24 +0,0 @@ -(** [hashmap_main]: external function declarations *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Require Export HashmapMain_Types. -Import HashmapMain_Types. -Module HashmapMain_FunsExternal. - -(** [hashmap_main::hashmap_utils::deserialize]: forward function - Source: 'src/hashmap_utils.rs', lines 10:0-10:43 *) -Axiom hashmap_utils_deserialize - : state -> result (state * (hashmap_HashMap_t u64)) -. - -(** [hashmap_main::hashmap_utils::serialize]: forward function - Source: 'src/hashmap_utils.rs', lines 5:0-5:42 *) -Axiom hashmap_utils_serialize - : hashmap_HashMap_t u64 -> state -> result (state * unit) -. - -End HashmapMain_FunsExternal. diff --git a/tests/coq/hashmap_main/HashmapMain_FunsExternal_Template.v b/tests/coq/hashmap_main/HashmapMain_FunsExternal_Template.v deleted file mode 100644 index 66835e8c..00000000 --- a/tests/coq/hashmap_main/HashmapMain_FunsExternal_Template.v +++ /dev/null @@ -1,26 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: external functions. --- This is a template file: rename it to "FunsExternal.lean" and fill the holes. *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Require Import HashmapMain_Types. -Include HashmapMain_Types. -Module HashmapMain_FunsExternal_Template. - -(** [hashmap_main::hashmap_utils::deserialize]: - Source: 'tests/src/hashmap_utils.rs', lines 11:0-11:43 *) -Axiom hashmap_utils_deserialize - : state -> result (state * (hashmap_HashMap_t u64)) -. - -(** [hashmap_main::hashmap_utils::serialize]: - Source: 'tests/src/hashmap_utils.rs', lines 6:0-6:42 *) -Axiom hashmap_utils_serialize - : hashmap_HashMap_t u64 -> state -> result (state * unit) -. - -End HashmapMain_FunsExternal_Template. diff --git a/tests/coq/hashmap_main/HashmapMain_Types.v b/tests/coq/hashmap_main/HashmapMain_Types.v deleted file mode 100644 index 5656bd9c..00000000 --- a/tests/coq/hashmap_main/HashmapMain_Types.v +++ /dev/null @@ -1,40 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: type definitions *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Require Import HashmapMain_TypesExternal. -Include HashmapMain_TypesExternal. -Module HashmapMain_Types. - -(** [hashmap_main::hashmap::List] - Source: 'tests/src/hashmap.rs', lines 27:0-27:16 *) -Inductive hashmap_List_t (T : Type) := -| Hashmap_List_Cons : usize -> T -> hashmap_List_t T -> hashmap_List_t T -| Hashmap_List_Nil : hashmap_List_t T -. - -Arguments Hashmap_List_Cons { _ }. -Arguments Hashmap_List_Nil { _ }. - -(** [hashmap_main::hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 43:0-43:21 *) -Record hashmap_HashMap_t (T : Type) := -mkhashmap_HashMap_t { - hashmap_HashMap_num_entries : usize; - hashmap_HashMap_max_load_factor : (usize * usize); - hashmap_HashMap_max_load : usize; - hashmap_HashMap_slots : alloc_vec_Vec (hashmap_List_t T); -} -. - -Arguments mkhashmap_HashMap_t { _ }. -Arguments hashmap_HashMap_num_entries { _ }. -Arguments hashmap_HashMap_max_load_factor { _ }. -Arguments hashmap_HashMap_max_load { _ }. -Arguments hashmap_HashMap_slots { _ }. - -End HashmapMain_Types. diff --git a/tests/coq/hashmap_main/HashmapMain_TypesExternal.v b/tests/coq/hashmap_main/HashmapMain_TypesExternal.v deleted file mode 100644 index 28651c14..00000000 --- a/tests/coq/hashmap_main/HashmapMain_TypesExternal.v +++ /dev/null @@ -1,14 +0,0 @@ -(** [hashmap_main]: external types. --- This is a template file: rename it to "TypesExternal.lean" and fill the holes. *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Module HashmapMain_TypesExternal. - -(** The state type used in the state-error monad *) -Axiom state : Type. - -End HashmapMain_TypesExternal. diff --git a/tests/coq/hashmap_main/HashmapMain_TypesExternal_Template.v b/tests/coq/hashmap_main/HashmapMain_TypesExternal_Template.v deleted file mode 100644 index 391b2775..00000000 --- a/tests/coq/hashmap_main/HashmapMain_TypesExternal_Template.v +++ /dev/null @@ -1,15 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: external types. --- This is a template file: rename it to "TypesExternal.lean" and fill the holes. *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Module HashmapMain_TypesExternal_Template. - -(** The state type used in the state-error monad *) -Axiom state : Type. - -End HashmapMain_TypesExternal_Template. diff --git a/tests/coq/hashmap_main/Makefile b/tests/coq/hashmap_main/Makefile deleted file mode 100644 index 1a5aee4a..00000000 --- a/tests/coq/hashmap_main/Makefile +++ /dev/null @@ -1,23 +0,0 @@ -# This file was automatically generated - modify ../Makefile.template instead -# Makefile originally taken from coq-club - -%: Makefile.coq phony - +make -f Makefile.coq $@ - -all: Makefile.coq - +make -f Makefile.coq all - -clean: Makefile.coq - +make -f Makefile.coq clean - rm -f Makefile.coq - -Makefile.coq: _CoqProject Makefile - coq_makefile -f _CoqProject | sed 's/$$(COQCHK) $$(COQCHKFLAGS) $$(COQLIBS)/$$(COQCHK) $$(COQCHKFLAGS) $$(subst -Q,-R,$$(COQLIBS))/' > Makefile.coq - -_CoqProject: ; - -Makefile: ; - -phony: ; - -.PHONY: all clean phony diff --git a/tests/coq/hashmap_main/Primitives.v b/tests/coq/hashmap_main/Primitives.v deleted file mode 100644 index b29fce43..00000000 --- a/tests/coq/hashmap_main/Primitives.v +++ /dev/null @@ -1,981 +0,0 @@ -Require Import Lia. -Require Coq.Strings.Ascii. -Require Coq.Strings.String. -Require Import Coq.Program.Equality. -Require Import Coq.ZArith.ZArith. -Require Import Coq.ZArith.Znat. -Require Import List. -Import ListNotations. - -Module Primitives. - - (* TODO: use more *) -Declare Scope Primitives_scope. - -(*** Result *) - -Inductive error := - | Failure - | OutOfFuel. - -Inductive result A := - | Ok : A -> result A - | Fail_ : error -> result A. - -Arguments Ok {_} a. -Arguments Fail_ {_}. - -Definition bind {A B} (m: result A) (f: A -> result B) : result B := - match m with - | Fail_ e => Fail_ e - | Ok x => f x - end. - -Definition return_ {A: Type} (x: A) : result A := Ok x. -Definition fail_ {A: Type} (e: error) : result A := Fail_ e. - -Notation "x <- c1 ; c2" := (bind c1 (fun x => c2)) - (at level 61, c1 at next level, right associativity). - -(** Monadic assert *) -Definition massert (b: bool) : result unit := - if b then Ok tt else Fail_ Failure. - -(** Normalize and unwrap a successful result (used for globals) *) -Definition eval_result_refl {A} {x} (a: result A) (p: a = Ok x) : A := - match a as r return (r = Ok x -> A) with - | Ok a' => fun _ => a' - | Fail_ e => fun p' => - False_rect _ (eq_ind (Fail_ e) - (fun e : result A => - match e with - | Ok _ => False - | Fail_ e => True - end) - I (Ok x) p') - end p. - -Notation "x %global" := (eval_result_refl x eq_refl) (at level 40). -Notation "x %return" := (eval_result_refl x eq_refl) (at level 40). - -(* Sanity check *) -Check (if true then Ok (1 + 2) else Fail_ Failure)%global = 3. - -(*** Misc *) - -Definition string := Coq.Strings.String.string. -Definition char := Coq.Strings.Ascii.ascii. -Definition char_of_byte := Coq.Strings.Ascii.ascii_of_byte. - -Definition core_mem_replace (a : Type) (x : a) (y : a) : a * a := (x, x) . - -Record mut_raw_ptr (T : Type) := { mut_raw_ptr_v : T }. -Record const_raw_ptr (T : Type) := { const_raw_ptr_v : T }. - -(*** Scalars *) - -Definition i8_min : Z := -128%Z. -Definition i8_max : Z := 127%Z. -Definition i16_min : Z := -32768%Z. -Definition i16_max : Z := 32767%Z. -Definition i32_min : Z := -2147483648%Z. -Definition i32_max : Z := 2147483647%Z. -Definition i64_min : Z := -9223372036854775808%Z. -Definition i64_max : Z := 9223372036854775807%Z. -Definition i128_min : Z := -170141183460469231731687303715884105728%Z. -Definition i128_max : Z := 170141183460469231731687303715884105727%Z. -Definition u8_min : Z := 0%Z. -Definition u8_max : Z := 255%Z. -Definition u16_min : Z := 0%Z. -Definition u16_max : Z := 65535%Z. -Definition u32_min : Z := 0%Z. -Definition u32_max : Z := 4294967295%Z. -Definition u64_min : Z := 0%Z. -Definition u64_max : Z := 18446744073709551615%Z. -Definition u128_min : Z := 0%Z. -Definition u128_max : Z := 340282366920938463463374607431768211455%Z. - -(** The bounds of [isize] and [usize] vary with the architecture. *) -Axiom isize_min : Z. -Axiom isize_max : Z. -Definition usize_min : Z := 0%Z. -Axiom usize_max : Z. - -Open Scope Z_scope. - -(** We provide those lemmas to reason about the bounds of [isize] and [usize] *) -Axiom isize_min_bound : isize_min <= i32_min. -Axiom isize_max_bound : i32_max <= isize_max. -Axiom usize_max_bound : u32_max <= usize_max. - -Inductive scalar_ty := - | Isize - | I8 - | I16 - | I32 - | I64 - | I128 - | Usize - | U8 - | U16 - | U32 - | U64 - | U128 -. - -Definition scalar_min (ty: scalar_ty) : Z := - match ty with - | Isize => isize_min - | I8 => i8_min - | I16 => i16_min - | I32 => i32_min - | I64 => i64_min - | I128 => i128_min - | Usize => usize_min - | U8 => u8_min - | U16 => u16_min - | U32 => u32_min - | U64 => u64_min - | U128 => u128_min -end. - -Definition scalar_max (ty: scalar_ty) : Z := - match ty with - | Isize => isize_max - | I8 => i8_max - | I16 => i16_max - | I32 => i32_max - | I64 => i64_max - | I128 => i128_max - | Usize => usize_max - | U8 => u8_max - | U16 => u16_max - | U32 => u32_max - | U64 => u64_max - | U128 => u128_max -end. - -(** We use the following conservative bounds to make sure we can compute bound - checks in most situations *) -Definition scalar_min_cons (ty: scalar_ty) : Z := - match ty with - | Isize => i32_min - | Usize => u32_min - | _ => scalar_min ty -end. - -Definition scalar_max_cons (ty: scalar_ty) : Z := - match ty with - | Isize => i32_max - | Usize => u32_max - | _ => scalar_max ty -end. - -Lemma scalar_min_cons_valid : forall ty, scalar_min ty <= scalar_min_cons ty . -Proof. - destruct ty; unfold scalar_min_cons, scalar_min; try lia. - - pose isize_min_bound; lia. - - apply Z.le_refl. -Qed. - -Lemma scalar_max_cons_valid : forall ty, scalar_max ty >= scalar_max_cons ty . -Proof. - destruct ty; unfold scalar_max_cons, scalar_max; try lia. - - pose isize_max_bound; lia. - - pose usize_max_bound. lia. -Qed. - -Definition scalar (ty: scalar_ty) : Type := - { x: Z | scalar_min ty <= x <= scalar_max ty }. - -Definition to_Z {ty} (x: scalar ty) : Z := proj1_sig x. - -(** Bounds checks: we start by using the conservative bounds, to make sure we - can compute in most situations, then we use the real bounds (for [isize] - and [usize]). *) -Definition scalar_ge_min (ty: scalar_ty) (x: Z) : bool := - Z.leb (scalar_min_cons ty) x || Z.leb (scalar_min ty) x. - -Definition scalar_le_max (ty: scalar_ty) (x: Z) : bool := - Z.leb x (scalar_max_cons ty) || Z.leb x (scalar_max ty). - -Lemma scalar_ge_min_valid (ty: scalar_ty) (x: Z) : - scalar_ge_min ty x = true -> scalar_min ty <= x . -Proof. - unfold scalar_ge_min. - pose (scalar_min_cons_valid ty). - lia. -Qed. - -Lemma scalar_le_max_valid (ty: scalar_ty) (x: Z) : - scalar_le_max ty x = true -> x <= scalar_max ty . -Proof. - unfold scalar_le_max. - pose (scalar_max_cons_valid ty). - lia. -Qed. - -Definition scalar_in_bounds (ty: scalar_ty) (x: Z) : bool := - scalar_ge_min ty x && scalar_le_max ty x . - -Lemma scalar_in_bounds_valid (ty: scalar_ty) (x: Z) : - scalar_in_bounds ty x = true -> scalar_min ty <= x <= scalar_max ty . -Proof. - unfold scalar_in_bounds. - intros H. - destruct (scalar_ge_min ty x) eqn:Hmin. - - destruct (scalar_le_max ty x) eqn:Hmax. - + pose (scalar_ge_min_valid ty x Hmin). - pose (scalar_le_max_valid ty x Hmax). - lia. - + inversion H. - - inversion H. -Qed. - -Import Sumbool. - -Definition mk_scalar (ty: scalar_ty) (x: Z) : result (scalar ty) := - match sumbool_of_bool (scalar_in_bounds ty x) with - | left H => Ok (exist _ x (scalar_in_bounds_valid _ _ H)) - | right _ => Fail_ Failure - end. - -Definition scalar_add {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x + to_Z y). - -Definition scalar_sub {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x - to_Z y). - -Definition scalar_mul {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x * to_Z y). - -Definition scalar_div {ty} (x y: scalar ty) : result (scalar ty) := - if to_Z y =? 0 then Fail_ Failure else - mk_scalar ty (to_Z x / to_Z y). - -Definition scalar_rem {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (Z.rem (to_Z x) (to_Z y)). - -Definition scalar_neg {ty} (x: scalar ty) : result (scalar ty) := mk_scalar ty (-(to_Z x)). - -Axiom scalar_xor : forall ty, scalar ty -> scalar ty -> scalar ty. (* TODO *) -Axiom scalar_or : forall ty, scalar ty -> scalar ty -> scalar ty. (* TODO *) -Axiom scalar_and : forall ty, scalar ty -> scalar ty -> scalar ty. (* TODO *) -Axiom scalar_shl : forall ty0 ty1, scalar ty0 -> scalar ty1 -> result (scalar ty0). (* TODO *) -Axiom scalar_shr : forall ty0 ty1, scalar ty0 -> scalar ty1 -> result (scalar ty0). (* TODO *) - -(** Cast an integer from a [src_ty] to a [tgt_ty] *) -(* TODO: check the semantics of casts in Rust *) -Definition scalar_cast (src_ty tgt_ty : scalar_ty) (x : scalar src_ty) : result (scalar tgt_ty) := - mk_scalar tgt_ty (to_Z x). - -(* This can't fail, but for now we make all casts faillible (easier for the translation) *) -Definition scalar_cast_bool (tgt_ty : scalar_ty) (x : bool) : result (scalar tgt_ty) := - mk_scalar tgt_ty (if x then 1 else 0). - -(** Comparisons *) -Definition scalar_leb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := - Z.leb (to_Z x) (to_Z y) . - -Definition scalar_ltb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := - Z.ltb (to_Z x) (to_Z y) . - -Definition scalar_geb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := - Z.geb (to_Z x) (to_Z y) . - -Definition scalar_gtb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := - Z.gtb (to_Z x) (to_Z y) . - -Definition scalar_eqb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := - Z.eqb (to_Z x) (to_Z y) . - -Definition scalar_neqb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := - negb (Z.eqb (to_Z x) (to_Z y)) . - - -(** The scalar types *) -Definition isize := scalar Isize. -Definition i8 := scalar I8. -Definition i16 := scalar I16. -Definition i32 := scalar I32. -Definition i64 := scalar I64. -Definition i128 := scalar I128. -Definition usize := scalar Usize. -Definition u8 := scalar U8. -Definition u16 := scalar U16. -Definition u32 := scalar U32. -Definition u64 := scalar U64. -Definition u128 := scalar U128. - -(** Negaion *) -Definition isize_neg := @scalar_neg Isize. -Definition i8_neg := @scalar_neg I8. -Definition i16_neg := @scalar_neg I16. -Definition i32_neg := @scalar_neg I32. -Definition i64_neg := @scalar_neg I64. -Definition i128_neg := @scalar_neg I128. - -(** Division *) -Definition isize_div := @scalar_div Isize. -Definition i8_div := @scalar_div I8. -Definition i16_div := @scalar_div I16. -Definition i32_div := @scalar_div I32. -Definition i64_div := @scalar_div I64. -Definition i128_div := @scalar_div I128. -Definition usize_div := @scalar_div Usize. -Definition u8_div := @scalar_div U8. -Definition u16_div := @scalar_div U16. -Definition u32_div := @scalar_div U32. -Definition u64_div := @scalar_div U64. -Definition u128_div := @scalar_div U128. - -(** Remainder *) -Definition isize_rem := @scalar_rem Isize. -Definition i8_rem := @scalar_rem I8. -Definition i16_rem := @scalar_rem I16. -Definition i32_rem := @scalar_rem I32. -Definition i64_rem := @scalar_rem I64. -Definition i128_rem := @scalar_rem I128. -Definition usize_rem := @scalar_rem Usize. -Definition u8_rem := @scalar_rem U8. -Definition u16_rem := @scalar_rem U16. -Definition u32_rem := @scalar_rem U32. -Definition u64_rem := @scalar_rem U64. -Definition u128_rem := @scalar_rem U128. - -(** Addition *) -Definition isize_add := @scalar_add Isize. -Definition i8_add := @scalar_add I8. -Definition i16_add := @scalar_add I16. -Definition i32_add := @scalar_add I32. -Definition i64_add := @scalar_add I64. -Definition i128_add := @scalar_add I128. -Definition usize_add := @scalar_add Usize. -Definition u8_add := @scalar_add U8. -Definition u16_add := @scalar_add U16. -Definition u32_add := @scalar_add U32. -Definition u64_add := @scalar_add U64. -Definition u128_add := @scalar_add U128. - -(** Substraction *) -Definition isize_sub := @scalar_sub Isize. -Definition i8_sub := @scalar_sub I8. -Definition i16_sub := @scalar_sub I16. -Definition i32_sub := @scalar_sub I32. -Definition i64_sub := @scalar_sub I64. -Definition i128_sub := @scalar_sub I128. -Definition usize_sub := @scalar_sub Usize. -Definition u8_sub := @scalar_sub U8. -Definition u16_sub := @scalar_sub U16. -Definition u32_sub := @scalar_sub U32. -Definition u64_sub := @scalar_sub U64. -Definition u128_sub := @scalar_sub U128. - -(** Multiplication *) -Definition isize_mul := @scalar_mul Isize. -Definition i8_mul := @scalar_mul I8. -Definition i16_mul := @scalar_mul I16. -Definition i32_mul := @scalar_mul I32. -Definition i64_mul := @scalar_mul I64. -Definition i128_mul := @scalar_mul I128. -Definition usize_mul := @scalar_mul Usize. -Definition u8_mul := @scalar_mul U8. -Definition u16_mul := @scalar_mul U16. -Definition u32_mul := @scalar_mul U32. -Definition u64_mul := @scalar_mul U64. -Definition u128_mul := @scalar_mul U128. - -(** Xor *) -Definition u8_xor := @scalar_xor U8. -Definition u16_xor := @scalar_xor U16. -Definition u32_xor := @scalar_xor U32. -Definition u64_xor := @scalar_xor U64. -Definition u128_xor := @scalar_xor U128. -Definition usize_xor := @scalar_xor Usize. -Definition i8_xor := @scalar_xor I8. -Definition i16_xor := @scalar_xor I16. -Definition i32_xor := @scalar_xor I32. -Definition i64_xor := @scalar_xor I64. -Definition i128_xor := @scalar_xor I128. -Definition isize_xor := @scalar_xor Isize. - -(** Or *) -Definition u8_or := @scalar_or U8. -Definition u16_or := @scalar_or U16. -Definition u32_or := @scalar_or U32. -Definition u64_or := @scalar_or U64. -Definition u128_or := @scalar_or U128. -Definition usize_or := @scalar_or Usize. -Definition i8_or := @scalar_or I8. -Definition i16_or := @scalar_or I16. -Definition i32_or := @scalar_or I32. -Definition i64_or := @scalar_or I64. -Definition i128_or := @scalar_or I128. -Definition isize_or := @scalar_or Isize. - -(** And *) -Definition u8_and := @scalar_and U8. -Definition u16_and := @scalar_and U16. -Definition u32_and := @scalar_and U32. -Definition u64_and := @scalar_and U64. -Definition u128_and := @scalar_and U128. -Definition usize_and := @scalar_and Usize. -Definition i8_and := @scalar_and I8. -Definition i16_and := @scalar_and I16. -Definition i32_and := @scalar_and I32. -Definition i64_and := @scalar_and I64. -Definition i128_and := @scalar_and I128. -Definition isize_and := @scalar_and Isize. - -(** Shift left *) -Definition u8_shl {ty} := @scalar_shl U8 ty. -Definition u16_shl {ty} := @scalar_shl U16 ty. -Definition u32_shl {ty} := @scalar_shl U32 ty. -Definition u64_shl {ty} := @scalar_shl U64 ty. -Definition u128_shl {ty} := @scalar_shl U128 ty. -Definition usize_shl {ty} := @scalar_shl Usize ty. -Definition i8_shl {ty} := @scalar_shl I8 ty. -Definition i16_shl {ty} := @scalar_shl I16 ty. -Definition i32_shl {ty} := @scalar_shl I32 ty. -Definition i64_shl {ty} := @scalar_shl I64 ty. -Definition i128_shl {ty} := @scalar_shl I128 ty. -Definition isize_shl {ty} := @scalar_shl Isize ty. - -(** Shift right *) -Definition u8_shr {ty} := @scalar_shr U8 ty. -Definition u16_shr {ty} := @scalar_shr U16 ty. -Definition u32_shr {ty} := @scalar_shr U32 ty. -Definition u64_shr {ty} := @scalar_shr U64 ty. -Definition u128_shr {ty} := @scalar_shr U128 ty. -Definition usize_shr {ty} := @scalar_shr Usize ty. -Definition i8_shr {ty} := @scalar_shr I8 ty. -Definition i16_shr {ty} := @scalar_shr I16 ty. -Definition i32_shr {ty} := @scalar_shr I32 ty. -Definition i64_shr {ty} := @scalar_shr I64 ty. -Definition i128_shr {ty} := @scalar_shr I128 ty. -Definition isize_shr {ty} := @scalar_shr Isize ty. - -(** Small utility *) -Definition usize_to_nat (x: usize) : nat := Z.to_nat (to_Z x). - -(** Notations *) -Notation "x %isize" := ((mk_scalar Isize x)%return) (at level 9). -Notation "x %i8" := ((mk_scalar I8 x)%return) (at level 9). -Notation "x %i16" := ((mk_scalar I16 x)%return) (at level 9). -Notation "x %i32" := ((mk_scalar I32 x)%return) (at level 9). -Notation "x %i64" := ((mk_scalar I64 x)%return) (at level 9). -Notation "x %i128" := ((mk_scalar I128 x)%return) (at level 9). -Notation "x %usize" := ((mk_scalar Usize x)%return) (at level 9). -Notation "x %u8" := ((mk_scalar U8 x)%return) (at level 9). -Notation "x %u16" := ((mk_scalar U16 x)%return) (at level 9). -Notation "x %u32" := ((mk_scalar U32 x)%return) (at level 9). -Notation "x %u64" := ((mk_scalar U64 x)%return) (at level 9). -Notation "x %u128" := ((mk_scalar U128 x)%return) (at level 9). - -Notation "x s= y" := (scalar_eqb x y) (at level 80) : Primitives_scope. -Notation "x s<> y" := (scalar_neqb x y) (at level 80) : Primitives_scope. -Notation "x s<= y" := (scalar_leb x y) (at level 80) : Primitives_scope. -Notation "x s< y" := (scalar_ltb x y) (at level 80) : Primitives_scope. -Notation "x s>= y" := (scalar_geb x y) (at level 80) : Primitives_scope. -Notation "x s> y" := (scalar_gtb x y) (at level 80) : Primitives_scope. - -(** Constants *) -Definition core_u8_max := u8_max %u32. -Definition core_u16_max := u16_max %u32. -Definition core_u32_max := u32_max %u32. -Definition core_u64_max := u64_max %u64. -Definition core_u128_max := u64_max %u128. -Axiom core_usize_max : usize. (** TODO *) -Definition core_i8_max := i8_max %i32. -Definition core_i16_max := i16_max %i32. -Definition core_i32_max := i32_max %i32. -Definition core_i64_max := i64_max %i64. -Definition core_i128_max := i64_max %i128. -Axiom core_isize_max : isize. (** TODO *) - -(*** core *) - -(** Trait declaration: [core::clone::Clone] *) -Record core_clone_Clone (self : Type) := { - clone : self -> result self -}. - -Definition core_clone_impls_CloneBool_clone (b : bool) : bool := b. - -Definition core_clone_CloneBool : core_clone_Clone bool := {| - clone := fun b => Ok (core_clone_impls_CloneBool_clone b) -|}. - -Definition core_clone_impls_CloneUsize_clone (x : usize) : usize := x. -Definition core_clone_impls_CloneU8_clone (x : u8) : u8 := x. -Definition core_clone_impls_CloneU16_clone (x : u16) : u16 := x. -Definition core_clone_impls_CloneU32_clone (x : u32) : u32 := x. -Definition core_clone_impls_CloneU64_clone (x : u64) : u64 := x. -Definition core_clone_impls_CloneU128_clone (x : u128) : u128 := x. - -Definition core_clone_impls_CloneIsize_clone (x : isize) : isize := x. -Definition core_clone_impls_CloneI8_clone (x : i8) : i8 := x. -Definition core_clone_impls_CloneI16_clone (x : i16) : i16 := x. -Definition core_clone_impls_CloneI32_clone (x : i32) : i32 := x. -Definition core_clone_impls_CloneI64_clone (x : i64) : i64 := x. -Definition core_clone_impls_CloneI128_clone (x : i128) : i128 := x. - -Definition core_clone_CloneUsize : core_clone_Clone usize := {| - clone := fun x => Ok (core_clone_impls_CloneUsize_clone x) -|}. - -Definition core_clone_CloneU8 : core_clone_Clone u8 := {| - clone := fun x => Ok (core_clone_impls_CloneU8_clone x) -|}. - -Definition core_clone_CloneU16 : core_clone_Clone u16 := {| - clone := fun x => Ok (core_clone_impls_CloneU16_clone x) -|}. - -Definition core_clone_CloneU32 : core_clone_Clone u32 := {| - clone := fun x => Ok (core_clone_impls_CloneU32_clone x) -|}. - -Definition core_clone_CloneU64 : core_clone_Clone u64 := {| - clone := fun x => Ok (core_clone_impls_CloneU64_clone x) -|}. - -Definition core_clone_CloneU128 : core_clone_Clone u128 := {| - clone := fun x => Ok (core_clone_impls_CloneU128_clone x) -|}. - -Definition core_clone_CloneIsize : core_clone_Clone isize := {| - clone := fun x => Ok (core_clone_impls_CloneIsize_clone x) -|}. - -Definition core_clone_CloneI8 : core_clone_Clone i8 := {| - clone := fun x => Ok (core_clone_impls_CloneI8_clone x) -|}. - -Definition core_clone_CloneI16 : core_clone_Clone i16 := {| - clone := fun x => Ok (core_clone_impls_CloneI16_clone x) -|}. - -Definition core_clone_CloneI32 : core_clone_Clone i32 := {| - clone := fun x => Ok (core_clone_impls_CloneI32_clone x) -|}. - -Definition core_clone_CloneI64 : core_clone_Clone i64 := {| - clone := fun x => Ok (core_clone_impls_CloneI64_clone x) -|}. - -Definition core_clone_CloneI128 : core_clone_Clone i128 := {| - clone := fun x => Ok (core_clone_impls_CloneI128_clone x) -|}. - -(** [core::option::{core::option::Option}::unwrap] *) -Definition core_option_Option_unwrap (T : Type) (x : option T) : result T := - match x with - | None => Fail_ Failure - | Some x => Ok x - end. - -(*** core::ops *) - -(* Trait declaration: [core::ops::index::Index] *) -Record core_ops_index_Index (Self Idx : Type) := mk_core_ops_index_Index { - core_ops_index_Index_Output : Type; - core_ops_index_Index_index : Self -> Idx -> result core_ops_index_Index_Output; -}. -Arguments mk_core_ops_index_Index {_ _}. -Arguments core_ops_index_Index_Output {_ _}. -Arguments core_ops_index_Index_index {_ _}. - -(* Trait declaration: [core::ops::index::IndexMut] *) -Record core_ops_index_IndexMut (Self Idx : Type) := mk_core_ops_index_IndexMut { - core_ops_index_IndexMut_indexInst : core_ops_index_Index Self Idx; - core_ops_index_IndexMut_index_mut : - Self -> - Idx -> - result (core_ops_index_IndexMut_indexInst.(core_ops_index_Index_Output) * - (core_ops_index_IndexMut_indexInst.(core_ops_index_Index_Output) -> result Self)); -}. -Arguments mk_core_ops_index_IndexMut {_ _}. -Arguments core_ops_index_IndexMut_indexInst {_ _}. -Arguments core_ops_index_IndexMut_index_mut {_ _}. - -(* Trait declaration [core::ops::deref::Deref] *) -Record core_ops_deref_Deref (Self : Type) := mk_core_ops_deref_Deref { - core_ops_deref_Deref_target : Type; - core_ops_deref_Deref_deref : Self -> result core_ops_deref_Deref_target; -}. -Arguments mk_core_ops_deref_Deref {_}. -Arguments core_ops_deref_Deref_target {_}. -Arguments core_ops_deref_Deref_deref {_}. - -(* Trait declaration [core::ops::deref::DerefMut] *) -Record core_ops_deref_DerefMut (Self : Type) := mk_core_ops_deref_DerefMut { - core_ops_deref_DerefMut_derefInst : core_ops_deref_Deref Self; - core_ops_deref_DerefMut_deref_mut : - Self -> - result (core_ops_deref_DerefMut_derefInst.(core_ops_deref_Deref_target) * - (core_ops_deref_DerefMut_derefInst.(core_ops_deref_Deref_target) -> result Self)); -}. -Arguments mk_core_ops_deref_DerefMut {_}. -Arguments core_ops_deref_DerefMut_derefInst {_}. -Arguments core_ops_deref_DerefMut_deref_mut {_}. - -Record core_ops_range_Range (T : Type) := mk_core_ops_range_Range { - core_ops_range_Range_start : T; - core_ops_range_Range_end_ : T; -}. -Arguments mk_core_ops_range_Range {_}. -Arguments core_ops_range_Range_start {_}. -Arguments core_ops_range_Range_end_ {_}. - -(*** [alloc] *) - -Definition alloc_boxed_Box_deref (T : Type) (x : T) : result T := Ok x. -Definition alloc_boxed_Box_deref_mut (T : Type) (x : T) : result (T * (T -> result T)) := - Ok (x, fun x => Ok x). - -(* Trait instance *) -Definition alloc_boxed_Box_coreopsDerefInst (Self : Type) : core_ops_deref_Deref Self := {| - core_ops_deref_Deref_target := Self; - core_ops_deref_Deref_deref := alloc_boxed_Box_deref Self; -|}. - -(* Trait instance *) -Definition alloc_boxed_Box_coreopsDerefMutInst (Self : Type) : core_ops_deref_DerefMut Self := {| - core_ops_deref_DerefMut_derefInst := alloc_boxed_Box_coreopsDerefInst Self; - core_ops_deref_DerefMut_deref_mut := alloc_boxed_Box_deref_mut Self; -|}. - - -(*** Arrays *) -Definition array T (n : usize) := { l: list T | Z.of_nat (length l) = to_Z n}. - -Lemma le_0_usize_max : 0 <= usize_max. -Proof. - pose (H := usize_max_bound). - unfold u32_max in H. - lia. -Qed. - -Lemma eqb_imp_eq (x y : Z) : Z.eqb x y = true -> x = y. -Proof. - lia. -Qed. - -(* TODO: finish the definitions *) -Axiom mk_array : forall (T : Type) (n : usize) (l : list T), array T n. - -(* For initialization *) -Axiom array_repeat : forall (T : Type) (n : usize) (x : T), array T n. - -Axiom array_index_usize : forall (T : Type) (n : usize) (x : array T n) (i : usize), result T. -Axiom array_update_usize : forall (T : Type) (n : usize) (x : array T n) (i : usize) (nx : T), result (array T n). - -Definition array_index_mut_usize (T : Type) (n : usize) (a : array T n) (i : usize) : - result (T * (T -> result (array T n))) := - match array_index_usize T n a i with - | Fail_ e => Fail_ e - | Ok x => Ok (x, array_update_usize T n a i) - end. - -(*** Slice *) -Definition slice T := { l: list T | Z.of_nat (length l) <= usize_max}. - -Axiom slice_len : forall (T : Type) (s : slice T), usize. -Axiom slice_index_usize : forall (T : Type) (x : slice T) (i : usize), result T. -Axiom slice_update_usize : forall (T : Type) (x : slice T) (i : usize) (nx : T), result (slice T). - -Definition slice_index_mut_usize (T : Type) (s : slice T) (i : usize) : - result (T * (T -> result (slice T))) := - match slice_index_usize T s i with - | Fail_ e => Fail_ e - | Ok x => Ok (x, slice_update_usize T s i) - end. - -(*** Subslices *) - -Axiom array_to_slice : forall (T : Type) (n : usize) (x : array T n), result (slice T). -Axiom array_from_slice : forall (T : Type) (n : usize) (x : array T n) (s : slice T), result (array T n). - -Definition array_to_slice_mut (T : Type) (n : usize) (a : array T n) : - result (slice T * (slice T -> result (array T n))) := - match array_to_slice T n a with - | Fail_ e => Fail_ e - | Ok x => Ok (x, array_from_slice T n a) - end. - -Axiom array_subslice: forall (T : Type) (n : usize) (x : array T n) (r : core_ops_range_Range usize), result (slice T). -Axiom array_update_subslice: forall (T : Type) (n : usize) (x : array T n) (r : core_ops_range_Range usize) (ns : slice T), result (array T n). - -Axiom slice_subslice: forall (T : Type) (x : slice T) (r : core_ops_range_Range usize), result (slice T). -Axiom slice_update_subslice: forall (T : Type) (x : slice T) (r : core_ops_range_Range usize) (ns : slice T), result (slice T). - -(*** Vectors *) - -Definition alloc_vec_Vec T := { l: list T | Z.of_nat (length l) <= usize_max }. - -Definition alloc_vec_Vec_to_list {T: Type} (v: alloc_vec_Vec T) : list T := proj1_sig v. - -Definition alloc_vec_Vec_length {T: Type} (v: alloc_vec_Vec T) : Z := Z.of_nat (length (alloc_vec_Vec_to_list v)). - -Definition alloc_vec_Vec_new (T: Type) : alloc_vec_Vec T := (exist _ [] le_0_usize_max). - -Lemma alloc_vec_Vec_len_in_usize {T} (v: alloc_vec_Vec T) : usize_min <= alloc_vec_Vec_length v <= usize_max. -Proof. - unfold alloc_vec_Vec_length, usize_min. - split. - - lia. - - apply (proj2_sig v). -Qed. - -Definition alloc_vec_Vec_len (T: Type) (v: alloc_vec_Vec T) : usize := - exist _ (alloc_vec_Vec_length v) (alloc_vec_Vec_len_in_usize v). - -Fixpoint list_update {A} (l: list A) (n: nat) (a: A) - : list A := - match l with - | [] => [] - | x :: t => match n with - | 0%nat => a :: t - | S m => x :: (list_update t m a) -end end. - -Definition alloc_vec_Vec_bind {A B} (v: alloc_vec_Vec A) (f: list A -> result (list B)) : result (alloc_vec_Vec B) := - l <- f (alloc_vec_Vec_to_list v) ; - match sumbool_of_bool (scalar_le_max Usize (Z.of_nat (length l))) with - | left H => Ok (exist _ l (scalar_le_max_valid _ _ H)) - | right _ => Fail_ Failure - end. - -Definition alloc_vec_Vec_push (T: Type) (v: alloc_vec_Vec T) (x: T) : result (alloc_vec_Vec T) := - alloc_vec_Vec_bind v (fun l => Ok (l ++ [x])). - -Definition alloc_vec_Vec_insert (T: Type) (v: alloc_vec_Vec T) (i: usize) (x: T) : result (alloc_vec_Vec T) := - alloc_vec_Vec_bind v (fun l => - if to_Z i result (alloc_vec_Vec T))) := - match alloc_vec_Vec_index_usize v i with - | Ok x => - Ok (x, alloc_vec_Vec_update_usize v i) - | Fail_ e => Fail_ e - end. - -(* Trait declaration: [core::slice::index::private_slice_index::Sealed] *) -Definition core_slice_index_private_slice_index_Sealed (self : Type) := unit. - -(* Trait declaration: [core::slice::index::SliceIndex] *) -Record core_slice_index_SliceIndex (Self T : Type) := mk_core_slice_index_SliceIndex { - core_slice_index_SliceIndex_sealedInst : core_slice_index_private_slice_index_Sealed Self; - core_slice_index_SliceIndex_Output : Type; - core_slice_index_SliceIndex_get : Self -> T -> result (option core_slice_index_SliceIndex_Output); - core_slice_index_SliceIndex_get_mut : - Self -> T -> result (option core_slice_index_SliceIndex_Output * (option core_slice_index_SliceIndex_Output -> result T)); - core_slice_index_SliceIndex_get_unchecked : Self -> const_raw_ptr T -> result (const_raw_ptr core_slice_index_SliceIndex_Output); - core_slice_index_SliceIndex_get_unchecked_mut : Self -> mut_raw_ptr T -> result (mut_raw_ptr core_slice_index_SliceIndex_Output); - core_slice_index_SliceIndex_index : Self -> T -> result core_slice_index_SliceIndex_Output; - core_slice_index_SliceIndex_index_mut : - Self -> T -> result (core_slice_index_SliceIndex_Output * (core_slice_index_SliceIndex_Output -> result T)); -}. -Arguments mk_core_slice_index_SliceIndex {_ _}. -Arguments core_slice_index_SliceIndex_sealedInst {_ _}. -Arguments core_slice_index_SliceIndex_Output {_ _}. -Arguments core_slice_index_SliceIndex_get {_ _}. -Arguments core_slice_index_SliceIndex_get_mut {_ _}. -Arguments core_slice_index_SliceIndex_get_unchecked {_ _}. -Arguments core_slice_index_SliceIndex_get_unchecked_mut {_ _}. -Arguments core_slice_index_SliceIndex_index {_ _}. -Arguments core_slice_index_SliceIndex_index_mut {_ _}. - -(* [core::slice::index::[T]::index]: forward function *) -Definition core_slice_index_Slice_index - (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)) - (s : slice T) (i : Idx) : result inst.(core_slice_index_SliceIndex_Output) := - x <- inst.(core_slice_index_SliceIndex_get) i s; - match x with - | None => Fail_ Failure - | Some x => Ok x - end. - -(* [core::slice::index::Range:::get]: forward function *) -Axiom core_slice_index_RangeUsize_get : forall (T : Type) (i : core_ops_range_Range usize) (s : slice T), result (option (slice T)). - -(* [core::slice::index::Range::get_mut]: forward function *) -Axiom core_slice_index_RangeUsize_get_mut : - forall (T : Type), - core_ops_range_Range usize -> slice T -> - result (option (slice T) * (option (slice T) -> result (slice T))). - -(* [core::slice::index::Range::get_unchecked]: forward function *) -Definition core_slice_index_RangeUsize_get_unchecked - (T : Type) : - core_ops_range_Range usize -> const_raw_ptr (slice T) -> result (const_raw_ptr (slice T)) := - (* Don't know what the model should be - for now we always fail to make - sure code which uses it fails *) - fun _ _ => Fail_ Failure. - -(* [core::slice::index::Range::get_unchecked_mut]: forward function *) -Definition core_slice_index_RangeUsize_get_unchecked_mut - (T : Type) : - core_ops_range_Range usize -> mut_raw_ptr (slice T) -> result (mut_raw_ptr (slice T)) := - (* Don't know what the model should be - for now we always fail to make - sure code which uses it fails *) - fun _ _ => Fail_ Failure. - -(* [core::slice::index::Range::index]: forward function *) -Axiom core_slice_index_RangeUsize_index : - forall (T : Type), core_ops_range_Range usize -> slice T -> result (slice T). - -(* [core::slice::index::Range::index_mut]: forward function *) -Axiom core_slice_index_RangeUsize_index_mut : - forall (T : Type), core_ops_range_Range usize -> slice T -> result (slice T * (slice T -> result (slice T))). - -(* [core::slice::index::[T]::index_mut]: forward function *) -Axiom core_slice_index_Slice_index_mut : - forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)), - slice T -> Idx -> - result (inst.(core_slice_index_SliceIndex_Output) * - (inst.(core_slice_index_SliceIndex_Output) -> result (slice T))). - -(* [core::array::[T; N]::index]: forward function *) -Axiom core_array_Array_index : - forall (T Idx : Type) (N : usize) (inst : core_ops_index_Index (slice T) Idx) - (a : array T N) (i : Idx), result inst.(core_ops_index_Index_Output). - -(* [core::array::[T; N]::index_mut]: forward function *) -Axiom core_array_Array_index_mut : - forall (T Idx : Type) (N : usize) (inst : core_ops_index_IndexMut (slice T) Idx) - (a : array T N) (i : Idx), - result (inst.(core_ops_index_IndexMut_indexInst).(core_ops_index_Index_Output) * - (inst.(core_ops_index_IndexMut_indexInst).(core_ops_index_Index_Output) -> result (array T N))). - -(* Trait implementation: [core::slice::index::private_slice_index::Range] *) -Definition core_slice_index_private_slice_index_SealedRangeUsizeInst - : core_slice_index_private_slice_index_Sealed (core_ops_range_Range usize) := tt. - -(* Trait implementation: [core::slice::index::Range] *) -Definition core_slice_index_SliceIndexRangeUsizeSliceTInst (T : Type) : - core_slice_index_SliceIndex (core_ops_range_Range usize) (slice T) := {| - core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_SealedRangeUsizeInst; - core_slice_index_SliceIndex_Output := slice T; - core_slice_index_SliceIndex_get := core_slice_index_RangeUsize_get T; - core_slice_index_SliceIndex_get_mut := core_slice_index_RangeUsize_get_mut T; - core_slice_index_SliceIndex_get_unchecked := core_slice_index_RangeUsize_get_unchecked T; - core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_RangeUsize_get_unchecked_mut T; - core_slice_index_SliceIndex_index := core_slice_index_RangeUsize_index T; - core_slice_index_SliceIndex_index_mut := core_slice_index_RangeUsize_index_mut T; -|}. - -(* Trait implementation: [core::slice::index::[T]] *) -Definition core_ops_index_IndexSliceTIInst (T Idx : Type) - (inst : core_slice_index_SliceIndex Idx (slice T)) : - core_ops_index_Index (slice T) Idx := {| - core_ops_index_Index_Output := inst.(core_slice_index_SliceIndex_Output); - core_ops_index_Index_index := core_slice_index_Slice_index T Idx inst; -|}. - -(* Trait implementation: [core::slice::index::[T]] *) -Definition core_ops_index_IndexMutSliceTIInst (T Idx : Type) - (inst : core_slice_index_SliceIndex Idx (slice T)) : - core_ops_index_IndexMut (slice T) Idx := {| - core_ops_index_IndexMut_indexInst := core_ops_index_IndexSliceTIInst T Idx inst; - core_ops_index_IndexMut_index_mut := core_slice_index_Slice_index_mut T Idx inst; -|}. - -(* Trait implementation: [core::array::[T; N]] *) -Definition core_ops_index_IndexArrayInst (T Idx : Type) (N : usize) - (inst : core_ops_index_Index (slice T) Idx) : - core_ops_index_Index (array T N) Idx := {| - core_ops_index_Index_Output := inst.(core_ops_index_Index_Output); - core_ops_index_Index_index := core_array_Array_index T Idx N inst; -|}. - -(* Trait implementation: [core::array::[T; N]] *) -Definition core_ops_index_IndexMutArrayInst (T Idx : Type) (N : usize) - (inst : core_ops_index_IndexMut (slice T) Idx) : - core_ops_index_IndexMut (array T N) Idx := {| - core_ops_index_IndexMut_indexInst := core_ops_index_IndexArrayInst T Idx N inst.(core_ops_index_IndexMut_indexInst); - core_ops_index_IndexMut_index_mut := core_array_Array_index_mut T Idx N inst; -|}. - -(* [core::slice::index::usize::get]: forward function *) -Axiom core_slice_index_usize_get : forall (T : Type), usize -> slice T -> result (option T). - -(* [core::slice::index::usize::get_mut]: forward function *) -Axiom core_slice_index_usize_get_mut : - forall (T : Type), usize -> slice T -> result (option T * (option T -> result (slice T))). - -(* [core::slice::index::usize::get_unchecked]: forward function *) -Axiom core_slice_index_usize_get_unchecked : - forall (T : Type), usize -> const_raw_ptr (slice T) -> result (const_raw_ptr T). - -(* [core::slice::index::usize::get_unchecked_mut]: forward function *) -Axiom core_slice_index_usize_get_unchecked_mut : - forall (T : Type), usize -> mut_raw_ptr (slice T) -> result (mut_raw_ptr T). - -(* [core::slice::index::usize::index]: forward function *) -Axiom core_slice_index_usize_index : forall (T : Type), usize -> slice T -> result T. - -(* [core::slice::index::usize::index_mut]: forward function *) -Axiom core_slice_index_usize_index_mut : - forall (T : Type), usize -> slice T -> result (T * (T -> result (slice T))). - -(* Trait implementation: [core::slice::index::private_slice_index::usize] *) -Definition core_slice_index_private_slice_index_SealedUsizeInst - : core_slice_index_private_slice_index_Sealed usize := tt. - -(* Trait implementation: [core::slice::index::usize] *) -Definition core_slice_index_SliceIndexUsizeSliceTInst (T : Type) : - core_slice_index_SliceIndex usize (slice T) := {| - core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_SealedUsizeInst; - core_slice_index_SliceIndex_Output := T; - core_slice_index_SliceIndex_get := core_slice_index_usize_get T; - core_slice_index_SliceIndex_get_mut := core_slice_index_usize_get_mut T; - core_slice_index_SliceIndex_get_unchecked := core_slice_index_usize_get_unchecked T; - core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_usize_get_unchecked_mut T; - core_slice_index_SliceIndex_index := core_slice_index_usize_index T; - core_slice_index_SliceIndex_index_mut := core_slice_index_usize_index_mut T; -|}. - -(* [alloc::vec::Vec::index]: forward function *) -Axiom alloc_vec_Vec_index : forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)) - (Self : alloc_vec_Vec T) (i : Idx), result inst.(core_slice_index_SliceIndex_Output). - -(* [alloc::vec::Vec::index_mut]: forward function *) -Axiom alloc_vec_Vec_index_mut : forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)) - (Self : alloc_vec_Vec T) (i : Idx), - result (inst.(core_slice_index_SliceIndex_Output) * - (inst.(core_slice_index_SliceIndex_Output) -> result (alloc_vec_Vec T))). - -(* Trait implementation: [alloc::vec::Vec] *) -Definition alloc_vec_Vec_coreopsindexIndexInst (T Idx : Type) - (inst : core_slice_index_SliceIndex Idx (slice T)) : - core_ops_index_Index (alloc_vec_Vec T) Idx := {| - core_ops_index_Index_Output := inst.(core_slice_index_SliceIndex_Output); - core_ops_index_Index_index := alloc_vec_Vec_index T Idx inst; -|}. - -(* Trait implementation: [alloc::vec::Vec] *) -Definition alloc_vec_Vec_coreopsindexIndexMutInst (T Idx : Type) - (inst : core_slice_index_SliceIndex Idx (slice T)) : - core_ops_index_IndexMut (alloc_vec_Vec T) Idx := {| - core_ops_index_IndexMut_indexInst := alloc_vec_Vec_coreopsindexIndexInst T Idx inst; - core_ops_index_IndexMut_index_mut := alloc_vec_Vec_index_mut T Idx inst; -|}. - -(*** Theorems *) - -Axiom alloc_vec_Vec_index_eq : forall {a : Type} (v : alloc_vec_Vec a) (i : usize) (x : a), - alloc_vec_Vec_index a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i = - alloc_vec_Vec_index_usize v i. - -Axiom alloc_vec_Vec_index_mut_eq : forall {a : Type} (v : alloc_vec_Vec a) (i : usize) (x : a), - alloc_vec_Vec_index_mut a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i = - alloc_vec_Vec_index_mut_usize v i. - -End Primitives. diff --git a/tests/coq/hashmap_main/_CoqProject b/tests/coq/hashmap_main/_CoqProject deleted file mode 100644 index d73541d9..00000000 --- a/tests/coq/hashmap_main/_CoqProject +++ /dev/null @@ -1,12 +0,0 @@ -# This file was automatically generated - see ../Makefile --R . Lib --arg -w --arg all - -HashmapMain_Types.v -HashmapMain_FunsExternal_Template.v -Primitives.v -HashmapMain_Funs.v -HashmapMain_TypesExternal.v -HashmapMain_FunsExternal.v -HashmapMain_TypesExternal_Template.v diff --git a/tests/fstar/hashmap/Hashmap.Clauses.Template.fst b/tests/fstar/hashmap/Hashmap.Clauses.Template.fst index 3119ded8..57003613 100644 --- a/tests/fstar/hashmap/Hashmap.Clauses.Template.fst +++ b/tests/fstar/hashmap/Hashmap.Clauses.Template.fst @@ -7,63 +7,63 @@ open Hashmap.Types #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [hashmap::{hashmap::HashMap}::allocate_slots]: decreases clause - Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) + Source: 'tests/src/hashmap.rs', lines 60:4-66:5 *) unfold let hashMap_allocate_slots_loop_decreases (t : Type0) (slots : alloc_vec_Vec (list_t t)) (n : usize) : nat = admit () (** [hashmap::{hashmap::HashMap}::clear]: decreases clause - Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) + Source: 'tests/src/hashmap.rs', lines 90:4-98:5 *) unfold let hashMap_clear_loop_decreases (t : Type0) (slots : alloc_vec_Vec (list_t t)) (i : usize) : nat = admit () (** [hashmap::{hashmap::HashMap}::insert_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) + Source: 'tests/src/hashmap.rs', lines 107:4-124:5 *) unfold let hashMap_insert_in_list_loop_decreases (t : Type0) (key : usize) (value : t) (ls : list_t t) : nat = admit () (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) + Source: 'tests/src/hashmap.rs', lines 193:4-206:5 *) unfold let hashMap_move_elements_from_list_loop_decreases (t : Type0) (ntable : hashMap_t t) (ls : list_t t) : nat = admit () (** [hashmap::{hashmap::HashMap}::move_elements]: decreases clause - Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) + Source: 'tests/src/hashmap.rs', lines 181:4-190:5 *) unfold let hashMap_move_elements_loop_decreases (t : Type0) (ntable : hashMap_t t) (slots : alloc_vec_Vec (list_t t)) (i : usize) : nat = admit () (** [hashmap::{hashmap::HashMap}::contains_key_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) + Source: 'tests/src/hashmap.rs', lines 216:4-229:5 *) unfold let hashMap_contains_key_in_list_loop_decreases (t : Type0) (key : usize) (ls : list_t t) : nat = admit () (** [hashmap::{hashmap::HashMap}::get_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) + Source: 'tests/src/hashmap.rs', lines 234:4-247:5 *) unfold let hashMap_get_in_list_loop_decreases (t : Type0) (key : usize) (ls : list_t t) : nat = admit () (** [hashmap::{hashmap::HashMap}::get_mut_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) + Source: 'tests/src/hashmap.rs', lines 255:4-264:5 *) unfold let hashMap_get_mut_in_list_loop_decreases (t : Type0) (ls : list_t t) (key : usize) : nat = admit () (** [hashmap::{hashmap::HashMap}::remove_from_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) + Source: 'tests/src/hashmap.rs', lines 275:4-301:5 *) unfold let hashMap_remove_from_list_loop_decreases (t : Type0) (key : usize) (ls : list_t t) : nat = diff --git a/tests/fstar/hashmap/Hashmap.Funs.fst b/tests/fstar/hashmap/Hashmap.Funs.fst index 06cdf7f3..a85be4ed 100644 --- a/tests/fstar/hashmap/Hashmap.Funs.fst +++ b/tests/fstar/hashmap/Hashmap.Funs.fst @@ -3,17 +3,18 @@ module Hashmap.Funs open Primitives include Hashmap.Types +include Hashmap.FunsExternal include Hashmap.Clauses #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 35:0-35:32 *) + Source: 'tests/src/hashmap.rs', lines 37:0-37:32 *) let hash_key (k : usize) : result usize = Ok k (** [hashmap::{hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) + Source: 'tests/src/hashmap.rs', lines 60:4-66: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))) @@ -27,7 +28,7 @@ let rec hashMap_allocate_slots_loop else Ok slots (** [hashmap::{hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 58:4-58:76 *) + Source: 'tests/src/hashmap.rs', lines 60:4-60:76 *) let hashMap_allocate_slots (t : Type0) (slots : alloc_vec_Vec (list_t t)) (n : usize) : result (alloc_vec_Vec (list_t t)) @@ -35,7 +36,7 @@ let hashMap_allocate_slots hashMap_allocate_slots_loop t slots n (** [hashmap::{hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 67:4-71:13 *) + Source: 'tests/src/hashmap.rs', lines 69:4-73:13 *) let hashMap_new_with_capacity (t : Type0) (capacity : usize) (max_load_dividend : usize) (max_load_divisor : usize) : @@ -54,12 +55,12 @@ let hashMap_new_with_capacity } (** [hashmap::{hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 83:4-83:24 *) + Source: 'tests/src/hashmap.rs', lines 85:4-85:24 *) let hashMap_new (t : Type0) : result (hashMap_t t) = hashMap_new_with_capacity t 32 4 5 (** [hashmap::{hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) + Source: 'tests/src/hashmap.rs', lines 90:4-98: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))) @@ -77,18 +78,18 @@ let rec hashMap_clear_loop else Ok slots (** [hashmap::{hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 88:4-88:27 *) + Source: 'tests/src/hashmap.rs', lines 90:4-90: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}::len]: - Source: 'tests/src/hashmap.rs', lines 98:4-98:30 *) + Source: 'tests/src/hashmap.rs', lines 100:4-100:30 *) let hashMap_len (t : Type0) (self : hashMap_t t) : result usize = Ok self.num_entries (** [hashmap::{hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) + Source: 'tests/src/hashmap.rs', lines 107:4-124:5 *) let rec hashMap_insert_in_list_loop (t : Type0) (key : usize) (value : t) (ls : list_t t) : Tot (result (bool & (list_t t))) @@ -105,7 +106,7 @@ let rec hashMap_insert_in_list_loop end (** [hashmap::{hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 105:4-105:71 *) + Source: 'tests/src/hashmap.rs', lines 107:4-107:71 *) let hashMap_insert_in_list (t : Type0) (key : usize) (value : t) (ls : list_t t) : result (bool & (list_t t)) @@ -113,7 +114,7 @@ let hashMap_insert_in_list hashMap_insert_in_list_loop t key value ls (** [hashmap::{hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 125:4-125:54 *) + Source: 'tests/src/hashmap.rs', lines 127:4-127:54 *) let hashMap_insert_no_resize (t : Type0) (self : hashMap_t t) (key : usize) (value : t) : result (hashMap_t t) @@ -134,7 +135,7 @@ let hashMap_insert_no_resize else let* v = index_mut_back l1 in Ok { self with slots = v } (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) + Source: 'tests/src/hashmap.rs', lines 193:4-206:5 *) let rec hashMap_move_elements_from_list_loop (t : Type0) (ntable : hashMap_t t) (ls : list_t t) : Tot (result (hashMap_t t)) @@ -148,13 +149,13 @@ let rec hashMap_move_elements_from_list_loop end (** [hashmap::{hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 191:4-191:72 *) + Source: 'tests/src/hashmap.rs', lines 193:4-193:72 *) let hashMap_move_elements_from_list (t : Type0) (ntable : hashMap_t t) (ls : list_t t) : result (hashMap_t t) = hashMap_move_elements_from_list_loop t ntable ls (** [hashmap::{hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) + Source: 'tests/src/hashmap.rs', lines 181:4-190:5 *) let rec hashMap_move_elements_loop (t : Type0) (ntable : hashMap_t t) (slots : alloc_vec_Vec (list_t t)) (i : usize) : @@ -175,7 +176,7 @@ let rec hashMap_move_elements_loop else Ok (ntable, slots) (** [hashmap::{hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 179:4-179:95 *) + Source: 'tests/src/hashmap.rs', lines 181:4-181:95 *) let hashMap_move_elements (t : Type0) (ntable : hashMap_t t) (slots : alloc_vec_Vec (list_t t)) (i : usize) : @@ -184,7 +185,7 @@ let hashMap_move_elements hashMap_move_elements_loop t ntable slots i (** [hashmap::{hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 148:4-148:28 *) + Source: 'tests/src/hashmap.rs', lines 150:4-150: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 @@ -204,7 +205,7 @@ let hashMap_try_resize else Ok { self with max_load_factor = (i, i1) } (** [hashmap::{hashmap::HashMap}::insert]: - Source: 'tests/src/hashmap.rs', lines 137:4-137:48 *) + Source: 'tests/src/hashmap.rs', lines 139:4-139:48 *) let hashMap_insert (t : Type0) (self : hashMap_t t) (key : usize) (value : t) : result (hashMap_t t) @@ -214,7 +215,7 @@ let hashMap_insert if i > self1.max_load then hashMap_try_resize t self1 else Ok self1 (** [hashmap::{hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) + Source: 'tests/src/hashmap.rs', lines 216:4-229:5 *) let rec hashMap_contains_key_in_list_loop (t : Type0) (key : usize) (ls : list_t t) : Tot (result bool) @@ -227,13 +228,13 @@ let rec hashMap_contains_key_in_list_loop end (** [hashmap::{hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 214:4-214:68 *) + Source: 'tests/src/hashmap.rs', lines 216:4-216:68 *) let hashMap_contains_key_in_list (t : Type0) (key : usize) (ls : list_t t) : result bool = hashMap_contains_key_in_list_loop t key ls (** [hashmap::{hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 207:4-207:49 *) + Source: 'tests/src/hashmap.rs', lines 209:4-209:49 *) let hashMap_contains_key (t : Type0) (self : hashMap_t t) (key : usize) : result bool = let* hash = hash_key key in @@ -246,7 +247,7 @@ let hashMap_contains_key hashMap_contains_key_in_list t key l (** [hashmap::{hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) + Source: 'tests/src/hashmap.rs', lines 234:4-247:5 *) let rec hashMap_get_in_list_loop (t : Type0) (key : usize) (ls : list_t t) : Tot (result t) (decreases (hashMap_get_in_list_loop_decreases t key ls)) @@ -258,12 +259,12 @@ let rec hashMap_get_in_list_loop end (** [hashmap::{hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 232:4-232:70 *) + 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 = hashMap_get_in_list_loop t key ls (** [hashmap::{hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 247:4-247:55 *) + Source: 'tests/src/hashmap.rs', lines 249:4-249: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 @@ -275,7 +276,7 @@ let hashMap_get (t : Type0) (self : hashMap_t t) (key : usize) : result t = hashMap_get_in_list t key l (** [hashmap::{hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) + Source: 'tests/src/hashmap.rs', lines 255:4-264: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)))) @@ -294,7 +295,7 @@ let rec hashMap_get_mut_in_list_loop end (** [hashmap::{hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 253:4-253:86 *) + Source: 'tests/src/hashmap.rs', lines 255:4-255:86 *) let hashMap_get_mut_in_list (t : Type0) (ls : list_t t) (key : usize) : result (t & (t -> result (list_t t))) @@ -302,7 +303,7 @@ let hashMap_get_mut_in_list hashMap_get_mut_in_list_loop t ls key (** [hashmap::{hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:67 *) + Source: 'tests/src/hashmap.rs', lines 267:4-267:67 *) let hashMap_get_mut (t : Type0) (self : hashMap_t t) (key : usize) : result (t & (t -> result (hashMap_t t))) @@ -323,7 +324,7 @@ let hashMap_get_mut Ok (x, back) (** [hashmap::{hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) + Source: 'tests/src/hashmap.rs', lines 275:4-301:5 *) let rec hashMap_remove_from_list_loop (t : Type0) (key : usize) (ls : list_t t) : Tot (result ((option t) & (list_t t))) @@ -346,7 +347,7 @@ let rec hashMap_remove_from_list_loop end (** [hashmap::{hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 273:4-273:69 *) + Source: 'tests/src/hashmap.rs', lines 275:4-275:69 *) let hashMap_remove_from_list (t : Type0) (key : usize) (ls : list_t t) : result ((option t) & (list_t t)) @@ -354,7 +355,7 @@ let hashMap_remove_from_list hashMap_remove_from_list_loop t key ls (** [hashmap::{hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 302:4-302:52 *) + Source: 'tests/src/hashmap.rs', lines 304:4-304:52 *) let hashMap_remove (t : Type0) (self : hashMap_t t) (key : usize) : result ((option t) & (hashMap_t t)) @@ -375,8 +376,16 @@ let hashMap_remove 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 *) +let insert_on_disk + (key : usize) (value : u64) (st : state) : result (state & unit) = + let* (st1, hm) = hashmap_utils_deserialize st in + let* hm1 = hashMap_insert u64 hm key value in + hashmap_utils_serialize hm1 st1 + (** [hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 323:0-323:10 *) + Source: 'tests/src/hashmap.rs', lines 350:0-350: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 new file mode 100644 index 00000000..6e3964c7 --- /dev/null +++ b/tests/fstar/hashmap/Hashmap.FunsExternal.fsti @@ -0,0 +1,16 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap]: external function declarations *) +module Hashmap.FunsExternal +open Primitives +include Hashmap.Types + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** [hashmap::hashmap_utils::deserialize]: + Source: 'tests/src/hashmap.rs', lines 330:4-330:47 *) +val hashmap_utils_deserialize : state -> result (state & (hashMap_t u64)) + +(** [hashmap::hashmap_utils::serialize]: + Source: 'tests/src/hashmap.rs', lines 325:4-325:46 *) +val hashmap_utils_serialize : hashMap_t u64 -> state -> result (state & unit) + diff --git a/tests/fstar/hashmap/Hashmap.Properties.fst b/tests/fstar/hashmap/Hashmap.Properties.fst new file mode 100644 index 00000000..0d6372c1 --- /dev/null +++ b/tests/fstar/hashmap/Hashmap.Properties.fst @@ -0,0 +1,48 @@ +(** Properties about the hashmap written on disk *) +module Hashmap.Properties +open Primitives +open Hashmap.Funs + +#set-options "--z3rlimit 50 --fuel 0 --ifuel 1" + +/// Below, we focus on the functions to read from disk/write to disk to showcase +/// how such reasoning which mixes opaque functions together with a state-error +/// monad can be performed. + +(*** Hypotheses *) + +/// [state_v] gives us the hash map currently stored on disk +assume +val state_v : state -> hashMap_t u64 + +/// [serialize] updates the hash map stored on disk +assume +val serialize_lem (hm : hashMap_t u64) (st : state) : Lemma ( + match hashmap_utils_serialize hm st with + | Fail _ -> True + | Ok (st', ()) -> state_v st' == hm) + [SMTPat (hashmap_utils_serialize hm st)] + +/// [deserialize] gives us the hash map stored on disk, without updating it +assume +val deserialize_lem (st : state) : Lemma ( + match hashmap_utils_deserialize st with + | Fail _ -> True + | Ok (st', hm) -> hm == state_v st /\ st' == st) + [SMTPat (hashmap_utils_deserialize st)] + +(*** Lemmas *) + +/// The obvious lemma about [insert_on_disk]: the updated hash map stored on disk +/// is exactly the hash map produced from inserting the binding ([key], [value]) +/// in the hash map previously stored on disk. +val insert_on_disk_lem (key : usize) (value : u64) (st : state) : Lemma ( + match insert_on_disk key value st with + | Fail _ -> True + | Ok (st', ()) -> + let hm = state_v st in + match hashMap_insert u64 hm key value with + | Fail _ -> False + | Ok hm' -> hm' == state_v st') + +let insert_on_disk_lem key value st = () diff --git a/tests/fstar/hashmap/Hashmap.Types.fst b/tests/fstar/hashmap/Hashmap.Types.fst index 962fbee2..818cb9d1 100644 --- a/tests/fstar/hashmap/Hashmap.Types.fst +++ b/tests/fstar/hashmap/Hashmap.Types.fst @@ -2,17 +2,18 @@ (** [hashmap]: type definitions *) module Hashmap.Types open Primitives +include Hashmap.TypesExternal #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" (** [hashmap::List] - Source: 'tests/src/hashmap.rs', lines 27:0-27:16 *) + 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::HashMap] - Source: 'tests/src/hashmap.rs', lines 43:0-43:21 *) + Source: 'tests/src/hashmap.rs', lines 45:0-45:21 *) type hashMap_t (t : Type0) = { num_entries : usize; diff --git a/tests/fstar/hashmap/Hashmap.TypesExternal.fsti b/tests/fstar/hashmap/Hashmap.TypesExternal.fsti new file mode 100644 index 00000000..66cb659b --- /dev/null +++ b/tests/fstar/hashmap/Hashmap.TypesExternal.fsti @@ -0,0 +1,10 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap]: external type declarations *) +module Hashmap.TypesExternal +open Primitives + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** The state type used in the state-error monad *) +val state : Type0 + diff --git a/tests/fstar/hashmap_main/HashmapMain.Clauses.Template.fst b/tests/fstar/hashmap_main/HashmapMain.Clauses.Template.fst deleted file mode 100644 index cdd73210..00000000 --- a/tests/fstar/hashmap_main/HashmapMain.Clauses.Template.fst +++ /dev/null @@ -1,72 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: templates for the decreases clauses *) -module HashmapMain.Clauses.Template -open Primitives -open HashmapMain.Types - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: decreases clause - Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) -unfold -let hashmap_HashMap_allocate_slots_loop_decreases (t : Type0) - (slots : alloc_vec_Vec (hashmap_List_t t)) (n : usize) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: decreases clause - Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) -unfold -let hashmap_HashMap_clear_loop_decreases (t : Type0) - (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) -unfold -let hashmap_HashMap_insert_in_list_loop_decreases (t : Type0) (key : usize) - (value : t) (ls : hashmap_List_t t) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) -unfold -let hashmap_HashMap_move_elements_from_list_loop_decreases (t : Type0) - (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: decreases clause - Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) -unfold -let hashmap_HashMap_move_elements_loop_decreases (t : Type0) - (ntable : hashmap_HashMap_t t) (slots : alloc_vec_Vec (hashmap_List_t t)) - (i : usize) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) -unfold -let hashmap_HashMap_contains_key_in_list_loop_decreases (t : Type0) - (key : usize) (ls : hashmap_List_t t) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) -unfold -let hashmap_HashMap_get_in_list_loop_decreases (t : Type0) (key : usize) - (ls : hashmap_List_t t) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) -unfold -let hashmap_HashMap_get_mut_in_list_loop_decreases (t : Type0) - (ls : hashmap_List_t t) (key : usize) : nat = - admit () - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: decreases clause - Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) -unfold -let hashmap_HashMap_remove_from_list_loop_decreases (t : Type0) (key : usize) - (ls : hashmap_List_t t) : nat = - admit () - diff --git a/tests/fstar/hashmap_main/HashmapMain.Clauses.fst b/tests/fstar/hashmap_main/HashmapMain.Clauses.fst deleted file mode 100644 index be5a4ab1..00000000 --- a/tests/fstar/hashmap_main/HashmapMain.Clauses.fst +++ /dev/null @@ -1,61 +0,0 @@ -(** [hashmap]: the decreases clauses *) -module HashmapMain.Clauses -open Primitives -open FStar.List.Tot -open HashmapMain.Types - -#set-options "--z3rlimit 50 --fuel 0 --ifuel 1" - -(** [hashmap::HashMap::allocate_slots]: decreases clause *) -unfold -let hashmap_HashMap_allocate_slots_loop_decreases (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) - (n : usize) : nat = n - -(** [hashmap::HashMap::clear]: decreases clause *) -unfold -let hashmap_HashMap_clear_loop_decreases (t : Type0) (slots : alloc_vec_Vec (hashmap_List_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_HashMap_insert_in_list_loop_decreases (t : Type0) (key : usize) (value : t) - (ls : hashmap_List_t t) : hashmap_List_t t = - ls - -(** [hashmap::HashMap::move_elements_from_list]: decreases clause *) -unfold -let hashmap_HashMap_move_elements_from_list_loop_decreases (t : Type0) - (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : hashmap_List_t t = - ls - -(** [hashmap::HashMap::move_elements]: decreases clause *) -unfold -let hashmap_HashMap_move_elements_loop_decreases (t : Type0) (ntable : hashmap_HashMap_t t) - (slots : alloc_vec_Vec (hashmap_List_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_HashMap_contains_key_in_list_loop_decreases (t : Type0) (key : usize) - (ls : hashmap_List_t t) : hashmap_List_t t = - ls - -(** [hashmap::HashMap::get_in_list]: decreases clause *) -unfold -let hashmap_HashMap_get_in_list_loop_decreases (t : Type0) (key : usize) (ls : hashmap_List_t t) : - hashmap_List_t t = - ls - -(** [hashmap::HashMap::get_mut_in_list]: decreases clause *) -unfold -let hashmap_HashMap_get_mut_in_list_loop_decreases (t : Type0) - (ls : hashmap_List_t t) (key : usize) : hashmap_List_t t = - ls - -(** [hashmap::HashMap::remove_from_list]: decreases clause *) -unfold -let hashmap_HashMap_remove_from_list_loop_decreases (t : Type0) (key : usize) - (ls : hashmap_List_t t) : hashmap_List_t t = - ls - diff --git a/tests/fstar/hashmap_main/HashmapMain.Funs.fst b/tests/fstar/hashmap_main/HashmapMain.Funs.fst deleted file mode 100644 index c88a746e..00000000 --- a/tests/fstar/hashmap_main/HashmapMain.Funs.fst +++ /dev/null @@ -1,446 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: function definitions *) -module HashmapMain.Funs -open Primitives -include HashmapMain.Types -include HashmapMain.FunsExternal -include HashmapMain.Clauses - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** [hashmap_main::hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 35:0-35:32 *) -let hashmap_hash_key (k : usize) : result usize = - Ok k - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 58:4-64:5 *) -let rec hashmap_HashMap_allocate_slots_loop - (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (n : usize) : - Tot (result (alloc_vec_Vec (hashmap_List_t t))) - (decreases (hashmap_HashMap_allocate_slots_loop_decreases t slots n)) - = - if n > 0 - then - let* slots1 = alloc_vec_Vec_push (hashmap_List_t t) slots Hashmap_List_Nil - in - let* n1 = usize_sub n 1 in - hashmap_HashMap_allocate_slots_loop t slots1 n1 - else Ok slots - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 58:4-58:76 *) -let hashmap_HashMap_allocate_slots - (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (n : usize) : - result (alloc_vec_Vec (hashmap_List_t t)) - = - hashmap_HashMap_allocate_slots_loop t slots n - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 67:4-71:13 *) -let hashmap_HashMap_new_with_capacity - (t : Type0) (capacity : usize) (max_load_dividend : usize) - (max_load_divisor : usize) : - result (hashmap_HashMap_t t) - = - let* slots = - hashmap_HashMap_allocate_slots t (alloc_vec_Vec_new (hashmap_List_t t)) - capacity in - let* i = usize_mul capacity max_load_dividend in - let* i1 = usize_div i max_load_divisor in - Ok - { - num_entries = 0; - max_load_factor = (max_load_dividend, max_load_divisor); - max_load = i1; - slots = slots - } - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 83:4-83:24 *) -let hashmap_HashMap_new (t : Type0) : result (hashmap_HashMap_t t) = - hashmap_HashMap_new_with_capacity t 32 4 5 - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 88:4-96:5 *) -let rec hashmap_HashMap_clear_loop - (t : Type0) (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : - Tot (result (alloc_vec_Vec (hashmap_List_t t))) - (decreases (hashmap_HashMap_clear_loop_decreases t slots i)) - = - let i1 = alloc_vec_Vec_len (hashmap_List_t t) slots in - if i < i1 - then - let* (_, index_mut_back) = - alloc_vec_Vec_index_mut (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) slots i - in - let* i2 = usize_add i 1 in - let* slots1 = index_mut_back Hashmap_List_Nil in - hashmap_HashMap_clear_loop t slots1 i2 - else Ok slots - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 88:4-88:27 *) -let hashmap_HashMap_clear - (t : Type0) (self : hashmap_HashMap_t t) : result (hashmap_HashMap_t t) = - let* hm = hashmap_HashMap_clear_loop t self.slots 0 in - Ok { self with num_entries = 0; slots = hm } - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::len]: - Source: 'tests/src/hashmap.rs', lines 98:4-98:30 *) -let hashmap_HashMap_len - (t : Type0) (self : hashmap_HashMap_t t) : result usize = - Ok self.num_entries - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 105:4-122:5 *) -let rec hashmap_HashMap_insert_in_list_loop - (t : Type0) (key : usize) (value : t) (ls : hashmap_List_t t) : - Tot (result (bool & (hashmap_List_t t))) - (decreases (hashmap_HashMap_insert_in_list_loop_decreases t key value ls)) - = - begin match ls with - | Hashmap_List_Cons ckey cvalue tl -> - if ckey = key - then Ok (false, Hashmap_List_Cons ckey value tl) - else - let* (b, tl1) = hashmap_HashMap_insert_in_list_loop t key value tl in - Ok (b, Hashmap_List_Cons ckey cvalue tl1) - | Hashmap_List_Nil -> Ok (true, Hashmap_List_Cons key value Hashmap_List_Nil) - end - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 105:4-105:71 *) -let hashmap_HashMap_insert_in_list - (t : Type0) (key : usize) (value : t) (ls : hashmap_List_t t) : - result (bool & (hashmap_List_t t)) - = - hashmap_HashMap_insert_in_list_loop t key value ls - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 125:4-125:54 *) -let hashmap_HashMap_insert_no_resize - (t : Type0) (self : hashmap_HashMap_t t) (key : usize) (value : t) : - result (hashmap_HashMap_t t) - = - let* hash = hashmap_hash_key key in - let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in - let* hash_mod = usize_rem hash i in - let* (l, index_mut_back) = - alloc_vec_Vec_index_mut (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) - self.slots hash_mod in - let* (inserted, l1) = hashmap_HashMap_insert_in_list t key value l in - if inserted - then - let* i1 = usize_add self.num_entries 1 in - let* v = index_mut_back l1 in - Ok { self with num_entries = i1; slots = v } - else let* v = index_mut_back l1 in Ok { self with slots = v } - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 191:4-204:5 *) -let rec hashmap_HashMap_move_elements_from_list_loop - (t : Type0) (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : - Tot (result (hashmap_HashMap_t t)) - (decreases ( - hashmap_HashMap_move_elements_from_list_loop_decreases t ntable ls)) - = - begin match ls with - | Hashmap_List_Cons k v tl -> - let* ntable1 = hashmap_HashMap_insert_no_resize t ntable k v in - hashmap_HashMap_move_elements_from_list_loop t ntable1 tl - | Hashmap_List_Nil -> Ok ntable - end - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 191:4-191:72 *) -let hashmap_HashMap_move_elements_from_list - (t : Type0) (ntable : hashmap_HashMap_t t) (ls : hashmap_List_t t) : - result (hashmap_HashMap_t t) - = - hashmap_HashMap_move_elements_from_list_loop t ntable ls - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 179:4-188:5 *) -let rec hashmap_HashMap_move_elements_loop - (t : Type0) (ntable : hashmap_HashMap_t t) - (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : - Tot (result ((hashmap_HashMap_t t) & (alloc_vec_Vec (hashmap_List_t t)))) - (decreases (hashmap_HashMap_move_elements_loop_decreases t ntable slots i)) - = - let i1 = alloc_vec_Vec_len (hashmap_List_t t) slots in - if i < i1 - then - let* (l, index_mut_back) = - alloc_vec_Vec_index_mut (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) slots i - in - let (ls, l1) = core_mem_replace (hashmap_List_t t) l Hashmap_List_Nil in - let* ntable1 = hashmap_HashMap_move_elements_from_list t ntable ls in - let* i2 = usize_add i 1 in - let* slots1 = index_mut_back l1 in - hashmap_HashMap_move_elements_loop t ntable1 slots1 i2 - else Ok (ntable, slots) - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 179:4-179:95 *) -let hashmap_HashMap_move_elements - (t : Type0) (ntable : hashmap_HashMap_t t) - (slots : alloc_vec_Vec (hashmap_List_t t)) (i : usize) : - result ((hashmap_HashMap_t t) & (alloc_vec_Vec (hashmap_List_t t))) - = - hashmap_HashMap_move_elements_loop t ntable slots i - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 148:4-148:28 *) -let hashmap_HashMap_try_resize - (t : Type0) (self : hashmap_HashMap_t t) : result (hashmap_HashMap_t t) = - let* max_usize = scalar_cast U32 Usize core_u32_max in - let capacity = alloc_vec_Vec_len (hashmap_List_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 - if capacity <= i2 - then - let* i3 = usize_mul capacity 2 in - let* ntable = hashmap_HashMap_new_with_capacity t i3 i i1 in - let* p = hashmap_HashMap_move_elements t ntable self.slots 0 in - let (ntable1, _) = p in - Ok - { ntable1 with num_entries = self.num_entries; max_load_factor = (i, i1) - } - else Ok { self with max_load_factor = (i, i1) } - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert]: - Source: 'tests/src/hashmap.rs', lines 137:4-137:48 *) -let hashmap_HashMap_insert - (t : Type0) (self : hashmap_HashMap_t t) (key : usize) (value : t) : - result (hashmap_HashMap_t t) - = - let* self1 = hashmap_HashMap_insert_no_resize t self key value in - let* i = hashmap_HashMap_len t self1 in - if i > self1.max_load then hashmap_HashMap_try_resize t self1 else Ok self1 - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 214:4-227:5 *) -let rec hashmap_HashMap_contains_key_in_list_loop - (t : Type0) (key : usize) (ls : hashmap_List_t t) : - Tot (result bool) - (decreases (hashmap_HashMap_contains_key_in_list_loop_decreases t key ls)) - = - begin match ls with - | Hashmap_List_Cons ckey _ tl -> - if ckey = key - then Ok true - else hashmap_HashMap_contains_key_in_list_loop t key tl - | Hashmap_List_Nil -> Ok false - end - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 214:4-214:68 *) -let hashmap_HashMap_contains_key_in_list - (t : Type0) (key : usize) (ls : hashmap_List_t t) : result bool = - hashmap_HashMap_contains_key_in_list_loop t key ls - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 207:4-207:49 *) -let hashmap_HashMap_contains_key - (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : result bool = - let* hash = hashmap_hash_key key in - let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in - let* hash_mod = usize_rem hash i in - let* l = - alloc_vec_Vec_index (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) - self.slots hash_mod in - hashmap_HashMap_contains_key_in_list t key l - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 232:4-245:5 *) -let rec hashmap_HashMap_get_in_list_loop - (t : Type0) (key : usize) (ls : hashmap_List_t t) : - Tot (result t) - (decreases (hashmap_HashMap_get_in_list_loop_decreases t key ls)) - = - begin match ls with - | Hashmap_List_Cons ckey cvalue tl -> - if ckey = key then Ok cvalue else hashmap_HashMap_get_in_list_loop t key tl - | Hashmap_List_Nil -> Fail Failure - end - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 232:4-232:70 *) -let hashmap_HashMap_get_in_list - (t : Type0) (key : usize) (ls : hashmap_List_t t) : result t = - hashmap_HashMap_get_in_list_loop t key ls - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 247:4-247:55 *) -let hashmap_HashMap_get - (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : result t = - let* hash = hashmap_hash_key key in - let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in - let* hash_mod = usize_rem hash i in - let* l = - alloc_vec_Vec_index (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) - self.slots hash_mod in - hashmap_HashMap_get_in_list t key l - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 253:4-262:5 *) -let rec hashmap_HashMap_get_mut_in_list_loop - (t : Type0) (ls : hashmap_List_t t) (key : usize) : - Tot (result (t & (t -> result (hashmap_List_t t)))) - (decreases (hashmap_HashMap_get_mut_in_list_loop_decreases t ls key)) - = - begin match ls with - | Hashmap_List_Cons ckey cvalue tl -> - if ckey = key - then - let back = fun ret -> Ok (Hashmap_List_Cons ckey ret tl) in - Ok (cvalue, back) - else - let* (x, back) = hashmap_HashMap_get_mut_in_list_loop t tl key in - let back1 = - fun ret -> - let* tl1 = back ret in Ok (Hashmap_List_Cons ckey cvalue tl1) in - Ok (x, back1) - | Hashmap_List_Nil -> Fail Failure - end - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 253:4-253:86 *) -let hashmap_HashMap_get_mut_in_list - (t : Type0) (ls : hashmap_List_t t) (key : usize) : - result (t & (t -> result (hashmap_List_t t))) - = - hashmap_HashMap_get_mut_in_list_loop t ls key - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:67 *) -let hashmap_HashMap_get_mut - (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : - result (t & (t -> result (hashmap_HashMap_t t))) - = - let* hash = hashmap_hash_key key in - let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in - let* hash_mod = usize_rem hash i in - let* (l, index_mut_back) = - alloc_vec_Vec_index_mut (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) - self.slots hash_mod in - let* (x, get_mut_in_list_back) = hashmap_HashMap_get_mut_in_list t l key in - let back = - fun ret -> - let* l1 = get_mut_in_list_back ret in - let* v = index_mut_back l1 in - Ok { self with slots = v } in - Ok (x, back) - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 273:4-299:5 *) -let rec hashmap_HashMap_remove_from_list_loop - (t : Type0) (key : usize) (ls : hashmap_List_t t) : - Tot (result ((option t) & (hashmap_List_t t))) - (decreases (hashmap_HashMap_remove_from_list_loop_decreases t key ls)) - = - begin match ls with - | Hashmap_List_Cons ckey x tl -> - if ckey = key - then - let (mv_ls, _) = - core_mem_replace (hashmap_List_t t) (Hashmap_List_Cons ckey x tl) - Hashmap_List_Nil in - begin match mv_ls with - | Hashmap_List_Cons _ cvalue tl1 -> Ok (Some cvalue, tl1) - | Hashmap_List_Nil -> Fail Failure - end - else - let* (o, tl1) = hashmap_HashMap_remove_from_list_loop t key tl in - Ok (o, Hashmap_List_Cons ckey x tl1) - | Hashmap_List_Nil -> Ok (None, Hashmap_List_Nil) - end - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 273:4-273:69 *) -let hashmap_HashMap_remove_from_list - (t : Type0) (key : usize) (ls : hashmap_List_t t) : - result ((option t) & (hashmap_List_t t)) - = - hashmap_HashMap_remove_from_list_loop t key ls - -(** [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 302:4-302:52 *) -let hashmap_HashMap_remove - (t : Type0) (self : hashmap_HashMap_t t) (key : usize) : - result ((option t) & (hashmap_HashMap_t t)) - = - let* hash = hashmap_hash_key key in - let i = alloc_vec_Vec_len (hashmap_List_t t) self.slots in - let* hash_mod = usize_rem hash i in - let* (l, index_mut_back) = - alloc_vec_Vec_index_mut (hashmap_List_t t) usize - (core_slice_index_SliceIndexUsizeSliceTInst (hashmap_List_t t)) - self.slots hash_mod in - let* (x, l1) = hashmap_HashMap_remove_from_list t key l in - begin match x with - | None -> let* v = index_mut_back l1 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 - Ok (Some x1, { self with num_entries = i1; slots = v }) - end - -(** [hashmap_main::hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 323:0-323:10 *) -let hashmap_test1 : result unit = - let* hm = hashmap_HashMap_new u64 in - let* hm1 = hashmap_HashMap_insert u64 hm 0 42 in - let* hm2 = hashmap_HashMap_insert u64 hm1 128 18 in - let* hm3 = hashmap_HashMap_insert u64 hm2 1024 138 in - let* hm4 = hashmap_HashMap_insert u64 hm3 1056 256 in - let* i = hashmap_HashMap_get u64 hm4 128 in - if not (i = 18) - then Fail Failure - else - let* (_, get_mut_back) = hashmap_HashMap_get_mut u64 hm4 1024 in - let* hm5 = get_mut_back 56 in - let* i1 = hashmap_HashMap_get u64 hm5 1024 in - if not (i1 = 56) - then Fail Failure - else - let* (x, hm6) = hashmap_HashMap_remove u64 hm5 1024 in - begin match x with - | None -> Fail Failure - | Some x1 -> - if not (x1 = 56) - then Fail Failure - else - let* i2 = hashmap_HashMap_get u64 hm6 0 in - if not (i2 = 42) - then Fail Failure - else - let* i3 = hashmap_HashMap_get u64 hm6 128 in - if not (i3 = 18) - then Fail Failure - else - let* i4 = hashmap_HashMap_get u64 hm6 1056 in - if not (i4 = 256) then Fail Failure else Ok () - end - -(** [hashmap_main::insert_on_disk]: - Source: 'tests/src/hashmap_main.rs', lines 13:0-13:43 *) -let insert_on_disk - (key : usize) (value : u64) (st : state) : result (state & unit) = - let* (st1, hm) = hashmap_utils_deserialize st in - let* hm1 = hashmap_HashMap_insert u64 hm key value in - hashmap_utils_serialize hm1 st1 - -(** [hashmap_main::main]: - Source: 'tests/src/hashmap_main.rs', lines 22:0-22:13 *) -let main : result unit = - Ok () - diff --git a/tests/fstar/hashmap_main/HashmapMain.FunsExternal.fsti b/tests/fstar/hashmap_main/HashmapMain.FunsExternal.fsti deleted file mode 100644 index cc20d988..00000000 --- a/tests/fstar/hashmap_main/HashmapMain.FunsExternal.fsti +++ /dev/null @@ -1,18 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: external function declarations *) -module HashmapMain.FunsExternal -open Primitives -include HashmapMain.Types - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** [hashmap_main::hashmap_utils::deserialize]: - Source: 'tests/src/hashmap_utils.rs', lines 11:0-11:43 *) -val hashmap_utils_deserialize - : state -> result (state & (hashmap_HashMap_t u64)) - -(** [hashmap_main::hashmap_utils::serialize]: - Source: 'tests/src/hashmap_utils.rs', lines 6:0-6:42 *) -val hashmap_utils_serialize - : hashmap_HashMap_t u64 -> state -> result (state & unit) - diff --git a/tests/fstar/hashmap_main/HashmapMain.Properties.fst b/tests/fstar/hashmap_main/HashmapMain.Properties.fst deleted file mode 100644 index beb3dc2c..00000000 --- a/tests/fstar/hashmap_main/HashmapMain.Properties.fst +++ /dev/null @@ -1,48 +0,0 @@ -(** Properties about the hashmap written on disk *) -module HashmapMain.Properties -open Primitives -open HashmapMain.Funs - -#set-options "--z3rlimit 50 --fuel 0 --ifuel 1" - -/// Below, we focus on the functions to read from disk/write to disk to showcase -/// how such reasoning which mixes opaque functions together with a state-error -/// monad can be performed. - -(*** Hypotheses *) - -/// [state_v] gives us the hash map currently stored on disk -assume -val state_v : state -> hashmap_HashMap_t u64 - -/// [serialize] updates the hash map stored on disk -assume -val serialize_lem (hm : hashmap_HashMap_t u64) (st : state) : Lemma ( - match hashmap_utils_serialize hm st with - | Fail _ -> True - | Ok (st', ()) -> state_v st' == hm) - [SMTPat (hashmap_utils_serialize hm st)] - -/// [deserialize] gives us the hash map stored on disk, without updating it -assume -val deserialize_lem (st : state) : Lemma ( - match hashmap_utils_deserialize st with - | Fail _ -> True - | Ok (st', hm) -> hm == state_v st /\ st' == st) - [SMTPat (hashmap_utils_deserialize st)] - -(*** Lemmas *) - -/// The obvious lemma about [insert_on_disk]: the updated hash map stored on disk -/// is exactly the hash map produced from inserting the binding ([key], [value]) -/// in the hash map previously stored on disk. -val insert_on_disk_lem (key : usize) (value : u64) (st : state) : Lemma ( - match insert_on_disk key value st with - | Fail _ -> True - | Ok (st', ()) -> - let hm = state_v st in - match hashmap_HashMap_insert u64 hm key value with - | Fail _ -> False - | Ok hm' -> hm' == state_v st') - -let insert_on_disk_lem key value st = () diff --git a/tests/fstar/hashmap_main/HashmapMain.Types.fst b/tests/fstar/hashmap_main/HashmapMain.Types.fst deleted file mode 100644 index 85bcaeea..00000000 --- a/tests/fstar/hashmap_main/HashmapMain.Types.fst +++ /dev/null @@ -1,24 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: type definitions *) -module HashmapMain.Types -open Primitives -include HashmapMain.TypesExternal - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** [hashmap_main::hashmap::List] - Source: 'tests/src/hashmap.rs', lines 27:0-27:16 *) -type hashmap_List_t (t : Type0) = -| Hashmap_List_Cons : usize -> t -> hashmap_List_t t -> hashmap_List_t t -| Hashmap_List_Nil : hashmap_List_t t - -(** [hashmap_main::hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 43:0-43:21 *) -type hashmap_HashMap_t (t : Type0) = -{ - num_entries : usize; - max_load_factor : (usize & usize); - max_load : usize; - slots : alloc_vec_Vec (hashmap_List_t t); -} - diff --git a/tests/fstar/hashmap_main/HashmapMain.TypesExternal.fsti b/tests/fstar/hashmap_main/HashmapMain.TypesExternal.fsti deleted file mode 100644 index 75747408..00000000 --- a/tests/fstar/hashmap_main/HashmapMain.TypesExternal.fsti +++ /dev/null @@ -1,10 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [hashmap_main]: external type declarations *) -module HashmapMain.TypesExternal -open Primitives - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** The state type used in the state-error monad *) -val state : Type0 - diff --git a/tests/fstar/hashmap_main/Makefile b/tests/fstar/hashmap_main/Makefile deleted file mode 100644 index fa7d1f36..00000000 --- a/tests/fstar/hashmap_main/Makefile +++ /dev/null @@ -1,49 +0,0 @@ -# This file was automatically generated - modify ../Makefile.template instead -INCLUDE_DIRS = . - -FSTAR_INCLUDES = $(addprefix --include ,$(INCLUDE_DIRS)) - -FSTAR_HINTS ?= --use_hints --use_hint_hashes --record_hints - -FSTAR_OPTIONS = $(FSTAR_HINTS) \ - --cache_checked_modules $(FSTAR_INCLUDES) --cmi \ - --warn_error '+241@247+285-274' \ - -FSTAR_EXE ?= fstar.exe -FSTAR_NO_FLAGS = $(FSTAR_EXE) --already_cached 'Prims FStar LowStar Steel' --odir obj --cache_dir obj - -FSTAR = $(FSTAR_NO_FLAGS) $(FSTAR_OPTIONS) - -# The F* roots are used to compute the dependency graph, and generate the .depend file -FSTAR_ROOTS ?= $(wildcard *.fst *.fsti) - -# Build all the files -all: $(addprefix obj/,$(addsuffix .checked,$(FSTAR_ROOTS))) - -# This is the right way to ensure the .depend file always gets re-built. -ifeq (,$(filter %-in,$(MAKECMDGOALS))) -ifndef NODEPEND -ifndef MAKE_RESTARTS -.depend: .FORCE - $(FSTAR_NO_FLAGS) --dep full $(notdir $(FSTAR_ROOTS)) > $@ - -.PHONY: .FORCE -.FORCE: -endif -endif - -include .depend -endif - -# For the interactive mode -%.fst-in %.fsti-in: - @echo $(FSTAR_OPTIONS) - -# Generete the .checked files in batch mode -%.checked: - $(FSTAR) $(FSTAR_OPTIONS) $< && \ - touch -c $@ - -.PHONY: clean -clean: - rm -f obj/* diff --git a/tests/fstar/hashmap_main/Primitives.fst b/tests/fstar/hashmap_main/Primitives.fst deleted file mode 100644 index 9951ccc3..00000000 --- a/tests/fstar/hashmap_main/Primitives.fst +++ /dev/null @@ -1,929 +0,0 @@ -/// This file lists primitive and assumed functions and types -module Primitives -open FStar.Mul -open FStar.List.Tot - -#set-options "--z3rlimit 15 --fuel 0 --ifuel 1" - -(*** Utilities *) -val list_update (#a : Type0) (ls : list a) (i : nat{i < length ls}) (x : a) : - ls':list a{ - length ls' = length ls /\ - index ls' i == x - } -#push-options "--fuel 1" -let rec list_update #a ls i x = - match ls with - | x' :: ls -> if i = 0 then x :: ls else x' :: list_update ls (i-1) x -#pop-options - -(*** Result *) -type error : Type0 = -| Failure -| OutOfFuel - -type result (a : Type0) : Type0 = -| Ok : v:a -> result a -| Fail : e:error -> result a - -// Monadic return operator -unfold let return (#a : Type0) (x : a) : result a = Ok x - -// Monadic bind operator. -// Allows to use the notation: -// ``` -// let* x = y in -// ... -// ``` -unfold let (let*) (#a #b : Type0) (m: result a) - (f: (x:a) -> Pure (result b) (requires (m == Ok x)) (ensures fun _ -> True)) : - result b = - match m with - | Ok x -> f x - | Fail e -> Fail e - -// Monadic assert(...) -let massert (b:bool) : result unit = if b then Ok () else Fail Failure - -// Normalize and unwrap a successful result (used for globals). -let eval_global (#a : Type0) (x : result a{Ok? (normalize_term x)}) : a = Ok?.v x - -(*** Misc *) -type char = FStar.Char.char -type string = string - -let is_zero (n: nat) : bool = n = 0 -let decrease (n: nat{n > 0}) : nat = n - 1 - -let core_mem_replace (a : Type0) (x : a) (y : a) : a & a = (x, x) - -// We don't really use raw pointers for now -type mut_raw_ptr (t : Type0) = { v : t } -type const_raw_ptr (t : Type0) = { v : t } - -(*** Scalars *) -/// Rem.: most of the following code was partially generated - -assume val size_numbits : pos - -// TODO: we could use FStar.Int.int_t and FStar.UInt.int_t - -let isize_min : int = -9223372036854775808 // TODO: should be opaque -let isize_max : int = 9223372036854775807 // TODO: should be opaque -let i8_min : int = -128 -let i8_max : int = 127 -let i16_min : int = -32768 -let i16_max : int = 32767 -let i32_min : int = -2147483648 -let i32_max : int = 2147483647 -let i64_min : int = -9223372036854775808 -let i64_max : int = 9223372036854775807 -let i128_min : int = -170141183460469231731687303715884105728 -let i128_max : int = 170141183460469231731687303715884105727 -let usize_min : int = 0 -let usize_max : int = 4294967295 // TODO: should be opaque -let u8_min : int = 0 -let u8_max : int = 255 -let u16_min : int = 0 -let u16_max : int = 65535 -let u32_min : int = 0 -let u32_max : int = 4294967295 -let u64_min : int = 0 -let u64_max : int = 18446744073709551615 -let u128_min : int = 0 -let u128_max : int = 340282366920938463463374607431768211455 - -type scalar_ty = -| Isize -| I8 -| I16 -| I32 -| I64 -| I128 -| Usize -| U8 -| U16 -| U32 -| U64 -| U128 - -let is_unsigned = function - | Isize | I8 | I16 | I32 | I64 | I128 -> false - | Usize | U8 | U16 | U32 | U64 | U128 -> true - -let scalar_min (ty : scalar_ty) : int = - match ty with - | Isize -> isize_min - | I8 -> i8_min - | I16 -> i16_min - | I32 -> i32_min - | I64 -> i64_min - | I128 -> i128_min - | Usize -> usize_min - | U8 -> u8_min - | U16 -> u16_min - | U32 -> u32_min - | U64 -> u64_min - | U128 -> u128_min - -let scalar_max (ty : scalar_ty) : int = - match ty with - | Isize -> isize_max - | I8 -> i8_max - | I16 -> i16_max - | I32 -> i32_max - | I64 -> i64_max - | I128 -> i128_max - | Usize -> usize_max - | U8 -> u8_max - | U16 -> u16_max - | U32 -> u32_max - | U64 -> u64_max - | U128 -> u128_max - -type scalar (ty : scalar_ty) : eqtype = x:int{scalar_min ty <= x && x <= scalar_max ty} - -let mk_scalar (ty : scalar_ty) (x : int) : result (scalar ty) = - if scalar_min ty <= x && scalar_max ty >= x then Ok x else Fail Failure - -let scalar_neg (#ty : scalar_ty) (x : scalar ty) : result (scalar ty) = mk_scalar ty (-x) - -let scalar_div (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = - if y <> 0 then mk_scalar ty (x / y) else Fail Failure - -/// The remainder operation -let int_rem (x : int) (y : int{y <> 0}) : int = - if x >= 0 then (x % y) else -(x % y) - -(* Checking consistency with Rust *) -let _ = assert_norm(int_rem 1 2 = 1) -let _ = assert_norm(int_rem (-1) 2 = -1) -let _ = assert_norm(int_rem 1 (-2) = 1) -let _ = assert_norm(int_rem (-1) (-2) = -1) - -let scalar_rem (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = - if y <> 0 then mk_scalar ty (int_rem x y) else Fail Failure - -let scalar_add (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = - mk_scalar ty (x + y) - -let scalar_sub (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = - mk_scalar ty (x - y) - -let scalar_mul (#ty : scalar_ty) (x : scalar ty) (y : scalar ty) : result (scalar ty) = - mk_scalar ty (x * y) - -let scalar_xor (#ty : scalar_ty) - (x : scalar ty) (y : scalar ty) : scalar ty = - match ty with - | U8 -> FStar.UInt.logxor #8 x y - | U16 -> FStar.UInt.logxor #16 x y - | U32 -> FStar.UInt.logxor #32 x y - | U64 -> FStar.UInt.logxor #64 x y - | U128 -> FStar.UInt.logxor #128 x y - | Usize -> admit() // TODO - | I8 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 8); - normalize_spec (scalar I8); - FStar.Int.logxor #8 x y - | I16 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 16); - normalize_spec (scalar I16); - FStar.Int.logxor #16 x y - | I32 -> FStar.Int.logxor #32 x y - | I64 -> FStar.Int.logxor #64 x y - | I128 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 128); - normalize_spec (scalar I128); - FStar.Int.logxor #128 x y - | Isize -> admit() // TODO - -let scalar_or (#ty : scalar_ty) - (x : scalar ty) (y : scalar ty) : scalar ty = - match ty with - | U8 -> FStar.UInt.logor #8 x y - | U16 -> FStar.UInt.logor #16 x y - | U32 -> FStar.UInt.logor #32 x y - | U64 -> FStar.UInt.logor #64 x y - | U128 -> FStar.UInt.logor #128 x y - | Usize -> admit() // TODO - | I8 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 8); - normalize_spec (scalar I8); - FStar.Int.logor #8 x y - | I16 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 16); - normalize_spec (scalar I16); - FStar.Int.logor #16 x y - | I32 -> FStar.Int.logor #32 x y - | I64 -> FStar.Int.logor #64 x y - | I128 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 128); - normalize_spec (scalar I128); - FStar.Int.logor #128 x y - | Isize -> admit() // TODO - -let scalar_and (#ty : scalar_ty) - (x : scalar ty) (y : scalar ty) : scalar ty = - match ty with - | U8 -> FStar.UInt.logand #8 x y - | U16 -> FStar.UInt.logand #16 x y - | U32 -> FStar.UInt.logand #32 x y - | U64 -> FStar.UInt.logand #64 x y - | U128 -> FStar.UInt.logand #128 x y - | Usize -> admit() // TODO - | I8 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 8); - normalize_spec (scalar I8); - FStar.Int.logand #8 x y - | I16 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 16); - normalize_spec (scalar I16); - FStar.Int.logand #16 x y - | I32 -> FStar.Int.logand #32 x y - | I64 -> FStar.Int.logand #64 x y - | I128 -> - // Encoding issues... - normalize_spec (FStar.Int.int_t 128); - normalize_spec (scalar I128); - FStar.Int.logand #128 x y - | Isize -> admit() // TODO - -// Shift left -let scalar_shl (#ty0 #ty1 : scalar_ty) - (x : scalar ty0) (y : scalar ty1) : result (scalar ty0) = - admit() - -// Shift right -let scalar_shr (#ty0 #ty1 : scalar_ty) - (x : scalar ty0) (y : scalar ty1) : result (scalar ty0) = - admit() - -(** Cast an integer from a [src_ty] to a [tgt_ty] *) -// TODO: check the semantics of casts in Rust -let scalar_cast (src_ty : scalar_ty) (tgt_ty : scalar_ty) (x : scalar src_ty) : result (scalar tgt_ty) = - mk_scalar tgt_ty x - -// This can't fail, but for now we make all casts faillible (easier for the translation) -let scalar_cast_bool (tgt_ty : scalar_ty) (x : bool) : result (scalar tgt_ty) = - mk_scalar tgt_ty (if x then 1 else 0) - -/// The scalar types -type isize : eqtype = scalar Isize -type i8 : eqtype = scalar I8 -type i16 : eqtype = scalar I16 -type i32 : eqtype = scalar I32 -type i64 : eqtype = scalar I64 -type i128 : eqtype = scalar I128 -type usize : eqtype = scalar Usize -type u8 : eqtype = scalar U8 -type u16 : eqtype = scalar U16 -type u32 : eqtype = scalar U32 -type u64 : eqtype = scalar U64 -type u128 : eqtype = scalar U128 - - -let core_isize_min : isize = isize_min -let core_isize_max : isize = isize_max -let core_i8_min : i8 = i8_min -let core_i8_max : i8 = i8_max -let core_i16_min : i16 = i16_min -let core_i16_max : i16 = i16_max -let core_i32_min : i32 = i32_min -let core_i32_max : i32 = i32_max -let core_i64_min : i64 = i64_min -let core_i64_max : i64 = i64_max -let core_i128_min : i128 = i128_min -let core_i128_max : i128 = i128_max - -let core_usize_min : usize = usize_min -let core_usize_max : usize = usize_max -let core_u8_min : u8 = u8_min -let core_u8_max : u8 = u8_max -let core_u16_min : u16 = u16_min -let core_u16_max : u16 = u16_max -let core_u32_min : u32 = u32_min -let core_u32_max : u32 = u32_max -let core_u64_min : u64 = u64_min -let core_u64_max : u64 = u64_max -let core_u128_min : u128 = u128_min -let core_u128_max : u128 = u128_max - -/// Negation -let isize_neg = scalar_neg #Isize -let i8_neg = scalar_neg #I8 -let i16_neg = scalar_neg #I16 -let i32_neg = scalar_neg #I32 -let i64_neg = scalar_neg #I64 -let i128_neg = scalar_neg #I128 - -/// Division -let isize_div = scalar_div #Isize -let i8_div = scalar_div #I8 -let i16_div = scalar_div #I16 -let i32_div = scalar_div #I32 -let i64_div = scalar_div #I64 -let i128_div = scalar_div #I128 -let usize_div = scalar_div #Usize -let u8_div = scalar_div #U8 -let u16_div = scalar_div #U16 -let u32_div = scalar_div #U32 -let u64_div = scalar_div #U64 -let u128_div = scalar_div #U128 - -/// Remainder -let isize_rem = scalar_rem #Isize -let i8_rem = scalar_rem #I8 -let i16_rem = scalar_rem #I16 -let i32_rem = scalar_rem #I32 -let i64_rem = scalar_rem #I64 -let i128_rem = scalar_rem #I128 -let usize_rem = scalar_rem #Usize -let u8_rem = scalar_rem #U8 -let u16_rem = scalar_rem #U16 -let u32_rem = scalar_rem #U32 -let u64_rem = scalar_rem #U64 -let u128_rem = scalar_rem #U128 - -/// Addition -let isize_add = scalar_add #Isize -let i8_add = scalar_add #I8 -let i16_add = scalar_add #I16 -let i32_add = scalar_add #I32 -let i64_add = scalar_add #I64 -let i128_add = scalar_add #I128 -let usize_add = scalar_add #Usize -let u8_add = scalar_add #U8 -let u16_add = scalar_add #U16 -let u32_add = scalar_add #U32 -let u64_add = scalar_add #U64 -let u128_add = scalar_add #U128 - -/// Subtraction -let isize_sub = scalar_sub #Isize -let i8_sub = scalar_sub #I8 -let i16_sub = scalar_sub #I16 -let i32_sub = scalar_sub #I32 -let i64_sub = scalar_sub #I64 -let i128_sub = scalar_sub #I128 -let usize_sub = scalar_sub #Usize -let u8_sub = scalar_sub #U8 -let u16_sub = scalar_sub #U16 -let u32_sub = scalar_sub #U32 -let u64_sub = scalar_sub #U64 -let u128_sub = scalar_sub #U128 - -/// Multiplication -let isize_mul = scalar_mul #Isize -let i8_mul = scalar_mul #I8 -let i16_mul = scalar_mul #I16 -let i32_mul = scalar_mul #I32 -let i64_mul = scalar_mul #I64 -let i128_mul = scalar_mul #I128 -let usize_mul = scalar_mul #Usize -let u8_mul = scalar_mul #U8 -let u16_mul = scalar_mul #U16 -let u32_mul = scalar_mul #U32 -let u64_mul = scalar_mul #U64 -let u128_mul = scalar_mul #U128 - -/// Xor -let u8_xor = scalar_xor #U8 -let u16_xor = scalar_xor #U16 -let u32_xor = scalar_xor #U32 -let u64_xor = scalar_xor #U64 -let u128_xor = scalar_xor #U128 -let usize_xor = scalar_xor #Usize -let i8_xor = scalar_xor #I8 -let i16_xor = scalar_xor #I16 -let i32_xor = scalar_xor #I32 -let i64_xor = scalar_xor #I64 -let i128_xor = scalar_xor #I128 -let isize_xor = scalar_xor #Isize - -/// Or -let u8_or = scalar_or #U8 -let u16_or = scalar_or #U16 -let u32_or = scalar_or #U32 -let u64_or = scalar_or #U64 -let u128_or = scalar_or #U128 -let usize_or = scalar_or #Usize -let i8_or = scalar_or #I8 -let i16_or = scalar_or #I16 -let i32_or = scalar_or #I32 -let i64_or = scalar_or #I64 -let i128_or = scalar_or #I128 -let isize_or = scalar_or #Isize - -/// And -let u8_and = scalar_and #U8 -let u16_and = scalar_and #U16 -let u32_and = scalar_and #U32 -let u64_and = scalar_and #U64 -let u128_and = scalar_and #U128 -let usize_and = scalar_and #Usize -let i8_and = scalar_and #I8 -let i16_and = scalar_and #I16 -let i32_and = scalar_and #I32 -let i64_and = scalar_and #I64 -let i128_and = scalar_and #I128 -let isize_and = scalar_and #Isize - -/// Shift left -let u8_shl #ty = scalar_shl #U8 #ty -let u16_shl #ty = scalar_shl #U16 #ty -let u32_shl #ty = scalar_shl #U32 #ty -let u64_shl #ty = scalar_shl #U64 #ty -let u128_shl #ty = scalar_shl #U128 #ty -let usize_shl #ty = scalar_shl #Usize #ty -let i8_shl #ty = scalar_shl #I8 #ty -let i16_shl #ty = scalar_shl #I16 #ty -let i32_shl #ty = scalar_shl #I32 #ty -let i64_shl #ty = scalar_shl #I64 #ty -let i128_shl #ty = scalar_shl #I128 #ty -let isize_shl #ty = scalar_shl #Isize #ty - -/// Shift right -let u8_shr #ty = scalar_shr #U8 #ty -let u16_shr #ty = scalar_shr #U16 #ty -let u32_shr #ty = scalar_shr #U32 #ty -let u64_shr #ty = scalar_shr #U64 #ty -let u128_shr #ty = scalar_shr #U128 #ty -let usize_shr #ty = scalar_shr #Usize #ty -let i8_shr #ty = scalar_shr #I8 #ty -let i16_shr #ty = scalar_shr #I16 #ty -let i32_shr #ty = scalar_shr #I32 #ty -let i64_shr #ty = scalar_shr #I64 #ty -let i128_shr #ty = scalar_shr #I128 #ty -let isize_shr #ty = scalar_shr #Isize #ty - -(*** core *) - -/// Trait declaration: [core::clone::Clone] -noeq type core_clone_Clone (self : Type0) = { - clone : self → result self -} - -let core_clone_impls_CloneBool_clone (b : bool) : bool = b - -let core_clone_CloneBool : core_clone_Clone bool = { - clone = fun b -> Ok (core_clone_impls_CloneBool_clone b) -} - -let core_clone_impls_CloneUsize_clone (x : usize) : usize = x -let core_clone_impls_CloneU8_clone (x : u8) : u8 = x -let core_clone_impls_CloneU16_clone (x : u16) : u16 = x -let core_clone_impls_CloneU32_clone (x : u32) : u32 = x -let core_clone_impls_CloneU64_clone (x : u64) : u64 = x -let core_clone_impls_CloneU128_clone (x : u128) : u128 = x - -let core_clone_impls_CloneIsize_clone (x : isize) : isize = x -let core_clone_impls_CloneI8_clone (x : i8) : i8 = x -let core_clone_impls_CloneI16_clone (x : i16) : i16 = x -let core_clone_impls_CloneI32_clone (x : i32) : i32 = x -let core_clone_impls_CloneI64_clone (x : i64) : i64 = x -let core_clone_impls_CloneI128_clone (x : i128) : i128 = x - -let core_clone_CloneUsize : core_clone_Clone usize = { - clone = fun x -> Ok (core_clone_impls_CloneUsize_clone x) -} - -let core_clone_CloneU8 : core_clone_Clone u8 = { - clone = fun x -> Ok (core_clone_impls_CloneU8_clone x) -} - -let core_clone_CloneU16 : core_clone_Clone u16 = { - clone = fun x -> Ok (core_clone_impls_CloneU16_clone x) -} - -let core_clone_CloneU32 : core_clone_Clone u32 = { - clone = fun x -> Ok (core_clone_impls_CloneU32_clone x) -} - -let core_clone_CloneU64 : core_clone_Clone u64 = { - clone = fun x -> Ok (core_clone_impls_CloneU64_clone x) -} - -let core_clone_CloneU128 : core_clone_Clone u128 = { - clone = fun x -> Ok (core_clone_impls_CloneU128_clone x) -} - -let core_clone_CloneIsize : core_clone_Clone isize = { - clone = fun x -> Ok (core_clone_impls_CloneIsize_clone x) -} - -let core_clone_CloneI8 : core_clone_Clone i8 = { - clone = fun x -> Ok (core_clone_impls_CloneI8_clone x) -} - -let core_clone_CloneI16 : core_clone_Clone i16 = { - clone = fun x -> Ok (core_clone_impls_CloneI16_clone x) -} - -let core_clone_CloneI32 : core_clone_Clone i32 = { - clone = fun x -> Ok (core_clone_impls_CloneI32_clone x) -} - -let core_clone_CloneI64 : core_clone_Clone i64 = { - clone = fun x -> Ok (core_clone_impls_CloneI64_clone x) -} - -let core_clone_CloneI128 : core_clone_Clone i128 = { - clone = fun x -> Ok (core_clone_impls_CloneI128_clone x) -} - -(** [core::option::{core::option::Option}::unwrap] *) -let core_option_Option_unwrap (t : Type0) (x : option t) : result t = - match x with - | None -> Fail Failure - | Some x -> Ok x - -(*** core::ops *) - -// Trait declaration: [core::ops::index::Index] -noeq type core_ops_index_Index (self idx : Type0) = { - output : Type0; - index : self → idx → result output -} - -// Trait declaration: [core::ops::index::IndexMut] -noeq type core_ops_index_IndexMut (self idx : Type0) = { - indexInst : core_ops_index_Index self idx; - index_mut : self → idx → result (indexInst.output & (indexInst.output → result self)); -} - -// Trait declaration [core::ops::deref::Deref] -noeq type core_ops_deref_Deref (self : Type0) = { - target : Type0; - deref : self → result target; -} - -// Trait declaration [core::ops::deref::DerefMut] -noeq type core_ops_deref_DerefMut (self : Type0) = { - derefInst : core_ops_deref_Deref self; - deref_mut : self → result (derefInst.target & (derefInst.target → result self)); -} - -type core_ops_range_Range (a : Type0) = { - start : a; - end_ : a; -} - -(*** [alloc] *) - -let alloc_boxed_Box_deref (t : Type0) (x : t) : result t = Ok x -let alloc_boxed_Box_deref_mut (t : Type0) (x : t) : result (t & (t -> result t)) = - Ok (x, (fun x -> Ok x)) - -// Trait instance -let alloc_boxed_Box_coreopsDerefInst (self : Type0) : core_ops_deref_Deref self = { - target = self; - deref = alloc_boxed_Box_deref self; -} - -// Trait instance -let alloc_boxed_Box_coreopsDerefMutInst (self : Type0) : core_ops_deref_DerefMut self = { - derefInst = alloc_boxed_Box_coreopsDerefInst self; - deref_mut = alloc_boxed_Box_deref_mut self; -} - -(*** Array *) -type array (a : Type0) (n : usize) = s:list a{length s = n} - -// We tried putting the normalize_term condition as a refinement on the list -// but it didn't work. It works with the requires clause. -let mk_array (a : Type0) (n : usize) - (l : list a) : - Pure (array a n) - (requires (normalize_term(FStar.List.Tot.length l) = n)) - (ensures (fun _ -> True)) = - normalize_term_spec (FStar.List.Tot.length l); - l - -let array_index_usize (a : Type0) (n : usize) (x : array a n) (i : usize) : result a = - if i < length x then Ok (index x i) - else Fail Failure - -let array_update_usize (a : Type0) (n : usize) (x : array a n) (i : usize) (nx : a) : - result (array a n) = - if i < length x then Ok (list_update x i nx) - else Fail Failure - -let array_index_mut_usize (a : Type0) (n : usize) (x : array a n) (i : usize) : - result (a & (a -> result (array a n))) = - match array_index_usize a n x i with - | Fail e -> Fail e - | Ok v -> - Ok (v, array_update_usize a n x i) - -(*** Slice *) -type slice (a : Type0) = s:list a{length s <= usize_max} - -let slice_len (a : Type0) (s : slice a) : usize = length s - -let slice_index_usize (a : Type0) (x : slice a) (i : usize) : result a = - if i < length x then Ok (index x i) - else Fail Failure - -let slice_update_usize (a : Type0) (x : slice a) (i : usize) (nx : a) : result (slice a) = - if i < length x then Ok (list_update x i nx) - else Fail Failure - -let slice_index_mut_usize (a : Type0) (s : slice a) (i : usize) : - result (a & (a -> result (slice a))) = - match slice_index_usize a s i with - | Fail e -> Fail e - | Ok x -> - Ok (x, slice_update_usize a s i) - -(*** Subslices *) - -let array_to_slice (a : Type0) (n : usize) (x : array a n) : result (slice a) = Ok x -let array_from_slice (a : Type0) (n : usize) (x : array a n) (s : slice a) : result (array a n) = - if length s = n then Ok s - else Fail Failure - -let array_to_slice_mut (a : Type0) (n : usize) (x : array a n) : - result (slice a & (slice a -> result (array a n))) = - Ok (x, array_from_slice a n x) - -// TODO: finish the definitions below (there lacks [List.drop] and [List.take] in the standard library *) -let array_subslice (a : Type0) (n : usize) (x : array a n) (r : core_ops_range_Range usize) : result (slice a) = - admit() - -let array_update_subslice (a : Type0) (n : usize) (x : array a n) (r : core_ops_range_Range usize) (ns : slice a) : result (array a n) = - admit() - -let array_repeat (a : Type0) (n : usize) (x : a) : array a n = - admit() - -let slice_subslice (a : Type0) (x : slice a) (r : core_ops_range_Range usize) : result (slice a) = - admit() - -let slice_update_subslice (a : Type0) (x : slice a) (r : core_ops_range_Range usize) (ns : slice a) : result (slice a) = - admit() - -(*** Vector *) -type alloc_vec_Vec (a : Type0) = v:list a{length v <= usize_max} - -let alloc_vec_Vec_new (a : Type0) : alloc_vec_Vec a = assert_norm(length #a [] == 0); [] -let alloc_vec_Vec_len (a : Type0) (v : alloc_vec_Vec a) : usize = length v - -// Helper -let alloc_vec_Vec_index_usize (#a : Type0) (v : alloc_vec_Vec a) (i : usize) : result a = - if i < length v then Ok (index v i) else Fail Failure -// Helper -let alloc_vec_Vec_update_usize (#a : Type0) (v : alloc_vec_Vec a) (i : usize) (x : a) : result (alloc_vec_Vec a) = - if i < length v then Ok (list_update v i x) else Fail Failure - -let alloc_vec_Vec_index_mut_usize (#a : Type0) (v: alloc_vec_Vec a) (i: usize) : - result (a & (a → result (alloc_vec_Vec a))) = - match alloc_vec_Vec_index_usize v i with - | Ok x -> - Ok (x, alloc_vec_Vec_update_usize v i) - | Fail e -> Fail e - -let alloc_vec_Vec_push (a : Type0) (v : alloc_vec_Vec a) (x : a) : - Pure (result (alloc_vec_Vec a)) - (requires True) - (ensures (fun res -> - match res with - | Fail e -> e == Failure - | Ok v' -> length v' = length v + 1)) = - if length v < usize_max then begin - (**) assert_norm(length [x] == 1); - (**) append_length v [x]; - (**) assert(length (append v [x]) = length v + 1); - Ok (append v [x]) - end - else Fail Failure - -let alloc_vec_Vec_insert (a : Type0) (v : alloc_vec_Vec a) (i : usize) (x : a) : result (alloc_vec_Vec a) = - if i < length v then Ok (list_update v i x) else Fail Failure - -// Trait declaration: [core::slice::index::private_slice_index::Sealed] -type core_slice_index_private_slice_index_Sealed (self : Type0) = unit - -// Trait declaration: [core::slice::index::SliceIndex] -noeq type core_slice_index_SliceIndex (self t : Type0) = { - sealedInst : core_slice_index_private_slice_index_Sealed self; - output : Type0; - get : self → t → result (option output); - get_mut : self → t → result (option output & (option output -> result t)); - get_unchecked : self → const_raw_ptr t → result (const_raw_ptr output); - get_unchecked_mut : self → mut_raw_ptr t → result (mut_raw_ptr output); - index : self → t → result output; - index_mut : self → t → result (output & (output -> result t)); -} - -// [core::slice::index::[T]::index]: forward function -let core_slice_index_Slice_index - (t idx : Type0) (inst : core_slice_index_SliceIndex idx (slice t)) - (s : slice t) (i : idx) : result inst.output = - let* x = inst.get i s in - match x with - | None -> Fail Failure - | Some x -> Ok x - -// [core::slice::index::Range:::get]: forward function -let core_slice_index_RangeUsize_get (t : Type0) (i : core_ops_range_Range usize) (s : slice t) : - result (option (slice t)) = - admit () // TODO - -// [core::slice::index::Range::get_mut]: forward function -let core_slice_index_RangeUsize_get_mut (t : Type0) : - core_ops_range_Range usize → slice t → result (option (slice t) & (option (slice t) -> result (slice t))) = - admit () // TODO - -// [core::slice::index::Range::get_unchecked]: forward function -let core_slice_index_RangeUsize_get_unchecked - (t : Type0) : - core_ops_range_Range usize → const_raw_ptr (slice t) → result (const_raw_ptr (slice t)) = - // Don't know what the model should be - for now we always fail to make - // sure code which uses it fails - fun _ _ -> Fail Failure - -// [core::slice::index::Range::get_unchecked_mut]: forward function -let core_slice_index_RangeUsize_get_unchecked_mut - (t : Type0) : - core_ops_range_Range usize → mut_raw_ptr (slice t) → result (mut_raw_ptr (slice t)) = - // Don't know what the model should be - for now we always fail to make - // sure code which uses it fails - fun _ _ -> Fail Failure - -// [core::slice::index::Range::index]: forward function -let core_slice_index_RangeUsize_index - (t : Type0) : core_ops_range_Range usize → slice t → result (slice t) = - admit () // TODO - -// [core::slice::index::Range::index_mut]: forward function -let core_slice_index_RangeUsize_index_mut (t : Type0) : - core_ops_range_Range usize → slice t → result (slice t & (slice t -> result (slice t))) = - admit () // TODO - -// [core::slice::index::[T]::index_mut]: forward function -let core_slice_index_Slice_index_mut - (t idx : Type0) (inst : core_slice_index_SliceIndex idx (slice t)) : - slice t → idx → result (inst.output & (inst.output -> result (slice t))) = - admit () // - -// [core::array::[T; N]::index]: forward function -let core_array_Array_index - (t idx : Type0) (n : usize) (inst : core_ops_index_Index (slice t) idx) - (a : array t n) (i : idx) : result inst.output = - admit () // TODO - -// [core::array::[T; N]::index_mut]: forward function -let core_array_Array_index_mut - (t idx : Type0) (n : usize) (inst : core_ops_index_IndexMut (slice t) idx) - (a : array t n) (i : idx) : - result (inst.indexInst.output & (inst.indexInst.output -> result (array t n))) = - admit () // TODO - -// Trait implementation: [core::slice::index::private_slice_index::Range] -let core_slice_index_private_slice_index_SealedRangeUsizeInst - : core_slice_index_private_slice_index_Sealed (core_ops_range_Range usize) = () - -// Trait implementation: [core::slice::index::Range] -let core_slice_index_SliceIndexRangeUsizeSliceTInst (t : Type0) : - core_slice_index_SliceIndex (core_ops_range_Range usize) (slice t) = { - sealedInst = core_slice_index_private_slice_index_SealedRangeUsizeInst; - output = slice t; - get = core_slice_index_RangeUsize_get t; - get_mut = core_slice_index_RangeUsize_get_mut t; - get_unchecked = core_slice_index_RangeUsize_get_unchecked t; - get_unchecked_mut = core_slice_index_RangeUsize_get_unchecked_mut t; - index = core_slice_index_RangeUsize_index t; - index_mut = core_slice_index_RangeUsize_index_mut t; -} - -// Trait implementation: [core::slice::index::[T]] -let core_ops_index_IndexSliceTIInst (t idx : Type0) - (inst : core_slice_index_SliceIndex idx (slice t)) : - core_ops_index_Index (slice t) idx = { - output = inst.output; - index = core_slice_index_Slice_index t idx inst; -} - -// Trait implementation: [core::slice::index::[T]] -let core_ops_index_IndexMutSliceTIInst (t idx : Type0) - (inst : core_slice_index_SliceIndex idx (slice t)) : - core_ops_index_IndexMut (slice t) idx = { - indexInst = core_ops_index_IndexSliceTIInst t idx inst; - index_mut = core_slice_index_Slice_index_mut t idx inst; -} - -// Trait implementation: [core::array::[T; N]] -let core_ops_index_IndexArrayInst (t idx : Type0) (n : usize) - (inst : core_ops_index_Index (slice t) idx) : - core_ops_index_Index (array t n) idx = { - output = inst.output; - index = core_array_Array_index t idx n inst; -} - -// Trait implementation: [core::array::[T; N]] -let core_ops_index_IndexMutArrayIInst (t idx : Type0) (n : usize) - (inst : core_ops_index_IndexMut (slice t) idx) : - core_ops_index_IndexMut (array t n) idx = { - indexInst = core_ops_index_IndexArrayInst t idx n inst.indexInst; - index_mut = core_array_Array_index_mut t idx n inst; -} - -// [core::slice::index::usize::get]: forward function -let core_slice_index_usize_get - (t : Type0) : usize → slice t → result (option t) = - admit () // TODO - -// [core::slice::index::usize::get_mut]: forward function -let core_slice_index_usize_get_mut (t : Type0) : - usize → slice t → result (option t & (option t -> result (slice t))) = - admit () // TODO - -// [core::slice::index::usize::get_unchecked]: forward function -let core_slice_index_usize_get_unchecked - (t : Type0) : usize → const_raw_ptr (slice t) → result (const_raw_ptr t) = - admit () // TODO - -// [core::slice::index::usize::get_unchecked_mut]: forward function -let core_slice_index_usize_get_unchecked_mut - (t : Type0) : usize → mut_raw_ptr (slice t) → result (mut_raw_ptr t) = - admit () // TODO - -// [core::slice::index::usize::index]: forward function -let core_slice_index_usize_index (t : Type0) : usize → slice t → result t = - admit () // TODO - -// [core::slice::index::usize::index_mut]: forward function -let core_slice_index_usize_index_mut (t : Type0) : - usize → slice t → result (t & (t -> result (slice t))) = - admit () // TODO - -// Trait implementation: [core::slice::index::private_slice_index::usize] -let core_slice_index_private_slice_index_SealedUsizeInst - : core_slice_index_private_slice_index_Sealed usize = () - -// Trait implementation: [core::slice::index::usize] -let core_slice_index_SliceIndexUsizeSliceTInst (t : Type0) : - core_slice_index_SliceIndex usize (slice t) = { - sealedInst = core_slice_index_private_slice_index_SealedUsizeInst; - output = t; - get = core_slice_index_usize_get t; - get_mut = core_slice_index_usize_get_mut t; - get_unchecked = core_slice_index_usize_get_unchecked t; - get_unchecked_mut = core_slice_index_usize_get_unchecked_mut t; - index = core_slice_index_usize_index t; - index_mut = core_slice_index_usize_index_mut t; -} - -// [alloc::vec::Vec::index]: forward function -let alloc_vec_Vec_index (t idx : Type0) (inst : core_slice_index_SliceIndex idx (slice t)) - (self : alloc_vec_Vec t) (i : idx) : result inst.output = - admit () // TODO - -// [alloc::vec::Vec::index_mut]: forward function -let alloc_vec_Vec_index_mut (t idx : Type0) (inst : core_slice_index_SliceIndex idx (slice t)) - (self : alloc_vec_Vec t) (i : idx) : - result (inst.output & (inst.output -> result (alloc_vec_Vec t))) = - admit () // TODO - -// Trait implementation: [alloc::vec::Vec] -let alloc_vec_Vec_coreopsindexIndexInst (t idx : Type0) - (inst : core_slice_index_SliceIndex idx (slice t)) : - core_ops_index_Index (alloc_vec_Vec t) idx = { - output = inst.output; - index = alloc_vec_Vec_index t idx inst; -} - -// Trait implementation: [alloc::vec::Vec] -let alloc_vec_Vec_coreopsindexIndexMutInst (t idx : Type0) - (inst : core_slice_index_SliceIndex idx (slice t)) : - core_ops_index_IndexMut (alloc_vec_Vec t) idx = { - indexInst = alloc_vec_Vec_coreopsindexIndexInst t idx inst; - index_mut = alloc_vec_Vec_index_mut t idx inst; -} - -(*** Theorems *) - -let alloc_vec_Vec_index_eq (#a : Type0) (v : alloc_vec_Vec a) (i : usize) : - Lemma ( - alloc_vec_Vec_index a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i == - alloc_vec_Vec_index_usize v i) - [SMTPat (alloc_vec_Vec_index a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i)] - = - admit() - -let alloc_vec_Vec_index_mut_eq (#a : Type0) (v : alloc_vec_Vec a) (i : usize) : - Lemma ( - alloc_vec_Vec_index_mut a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i == - alloc_vec_Vec_index_mut_usize v i) - [SMTPat (alloc_vec_Vec_index_mut a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i)] - = - admit() diff --git a/tests/lean/Hashmap/Funs.lean b/tests/lean/Hashmap/Funs.lean index 4f05fbc8..e76912cc 100644 --- a/tests/lean/Hashmap/Funs.lean +++ b/tests/lean/Hashmap/Funs.lean @@ -2,17 +2,18 @@ -- [hashmap]: function definitions import Base import Hashmap.Types +import Hashmap.FunsExternal open Primitives namespace hashmap /- [hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 35:0-35:32 -/ + Source: 'tests/src/hashmap.rs', lines 37:0-37:32 -/ def hash_key (k : Usize) : Result Usize := Result.ok k /- [hashmap::{hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 58:4-64:5 -/ + Source: 'tests/src/hashmap.rs', lines 60:4-66:5 -/ divergent def HashMap.allocate_slots_loop (T : Type) (slots : alloc.vec.Vec (List T)) (n : Usize) : Result (alloc.vec.Vec (List T)) @@ -26,7 +27,7 @@ divergent def HashMap.allocate_slots_loop else Result.ok slots /- [hashmap::{hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 58:4-58:76 -/ + Source: 'tests/src/hashmap.rs', lines 60:4-60:76 -/ def HashMap.allocate_slots (T : Type) (slots : alloc.vec.Vec (List T)) (n : Usize) : Result (alloc.vec.Vec (List T)) @@ -34,7 +35,7 @@ def HashMap.allocate_slots HashMap.allocate_slots_loop T slots n /- [hashmap::{hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 67:4-71:13 -/ + Source: 'tests/src/hashmap.rs', lines 69:4-73:13 -/ def HashMap.new_with_capacity (T : Type) (capacity : Usize) (max_load_dividend : Usize) (max_load_divisor : Usize) : @@ -53,12 +54,12 @@ def HashMap.new_with_capacity } /- [hashmap::{hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 83:4-83:24 -/ + Source: 'tests/src/hashmap.rs', lines 85:4-85:24 -/ def HashMap.new (T : Type) : Result (HashMap T) := HashMap.new_with_capacity T 32#usize 4#usize 5#usize /- [hashmap::{hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 88:4-96:5 -/ + Source: 'tests/src/hashmap.rs', lines 90:4-98:5 -/ divergent def HashMap.clear_loop (T : Type) (slots : alloc.vec.Vec (List T)) (i : Usize) : Result (alloc.vec.Vec (List T)) @@ -76,19 +77,19 @@ divergent def HashMap.clear_loop else Result.ok slots /- [hashmap::{hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 88:4-88:27 -/ + Source: 'tests/src/hashmap.rs', lines 90:4-90: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}::len]: - Source: 'tests/src/hashmap.rs', lines 98:4-98:30 -/ + Source: 'tests/src/hashmap.rs', lines 100:4-100:30 -/ def HashMap.len (T : Type) (self : HashMap T) : Result Usize := Result.ok self.num_entries /- [hashmap::{hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 105:4-122:5 -/ + Source: 'tests/src/hashmap.rs', lines 107:4-124:5 -/ divergent def HashMap.insert_in_list_loop (T : Type) (key : Usize) (value : T) (ls : List T) : Result (Bool × (List T)) @@ -104,7 +105,7 @@ divergent def HashMap.insert_in_list_loop | List.Nil => Result.ok (true, List.Cons key value List.Nil) /- [hashmap::{hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 105:4-105:71 -/ + Source: 'tests/src/hashmap.rs', lines 107:4-107:71 -/ def HashMap.insert_in_list (T : Type) (key : Usize) (value : T) (ls : List T) : Result (Bool × (List T)) @@ -112,7 +113,7 @@ def HashMap.insert_in_list HashMap.insert_in_list_loop T key value ls /- [hashmap::{hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 125:4-125:54 -/ + Source: 'tests/src/hashmap.rs', lines 127:4-127:54 -/ def HashMap.insert_no_resize (T : Type) (self : HashMap T) (key : Usize) (value : T) : Result (HashMap T) @@ -136,7 +137,7 @@ def HashMap.insert_no_resize Result.ok { self with slots := v } /- [hashmap::{hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 191:4-204:5 -/ + Source: 'tests/src/hashmap.rs', lines 193:4-206:5 -/ divergent def HashMap.move_elements_from_list_loop (T : Type) (ntable : HashMap T) (ls : List T) : Result (HashMap T) := match ls with @@ -147,13 +148,13 @@ divergent def HashMap.move_elements_from_list_loop | List.Nil => Result.ok ntable /- [hashmap::{hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 191:4-191:72 -/ + Source: 'tests/src/hashmap.rs', lines 193:4-193:72 -/ def HashMap.move_elements_from_list (T : Type) (ntable : HashMap T) (ls : List T) : Result (HashMap T) := HashMap.move_elements_from_list_loop T ntable ls /- [hashmap::{hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 179:4-188:5 -/ + Source: 'tests/src/hashmap.rs', lines 181:4-190:5 -/ divergent def HashMap.move_elements_loop (T : Type) (ntable : HashMap T) (slots : alloc.vec.Vec (List T)) (i : Usize) : @@ -174,7 +175,7 @@ divergent def HashMap.move_elements_loop else Result.ok (ntable, slots) /- [hashmap::{hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 179:4-179:95 -/ + Source: 'tests/src/hashmap.rs', lines 181:4-181:95 -/ def HashMap.move_elements (T : Type) (ntable : HashMap T) (slots : alloc.vec.Vec (List T)) (i : Usize) : @@ -183,7 +184,7 @@ def HashMap.move_elements HashMap.move_elements_loop T ntable slots i /- [hashmap::{hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 148:4-148:28 -/ + Source: 'tests/src/hashmap.rs', lines 150:4-150:28 -/ def HashMap.try_resize (T : Type) (self : HashMap T) : Result (HashMap T) := do let max_usize ← Scalar.cast .Usize core_u32_max @@ -207,7 +208,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}::insert]: - Source: 'tests/src/hashmap.rs', lines 137:4-137:48 -/ + Source: 'tests/src/hashmap.rs', lines 139:4-139:48 -/ def HashMap.insert (T : Type) (self : HashMap T) (key : Usize) (value : T) : Result (HashMap T) @@ -220,7 +221,7 @@ def HashMap.insert else Result.ok self1 /- [hashmap::{hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 214:4-227:5 -/ + Source: 'tests/src/hashmap.rs', lines 216:4-229:5 -/ divergent def HashMap.contains_key_in_list_loop (T : Type) (key : Usize) (ls : List T) : Result Bool := match ls with @@ -231,13 +232,13 @@ divergent def HashMap.contains_key_in_list_loop | List.Nil => Result.ok false /- [hashmap::{hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 214:4-214:68 -/ + Source: 'tests/src/hashmap.rs', lines 216:4-216:68 -/ def HashMap.contains_key_in_list (T : Type) (key : Usize) (ls : List T) : Result Bool := HashMap.contains_key_in_list_loop T key ls /- [hashmap::{hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 207:4-207:49 -/ + Source: 'tests/src/hashmap.rs', lines 209:4-209:49 -/ def HashMap.contains_key (T : Type) (self : HashMap T) (key : Usize) : Result Bool := do @@ -250,7 +251,7 @@ def HashMap.contains_key HashMap.contains_key_in_list T key l /- [hashmap::{hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 232:4-245:5 -/ + Source: 'tests/src/hashmap.rs', lines 234:4-247:5 -/ divergent def HashMap.get_in_list_loop (T : Type) (key : Usize) (ls : List T) : Result T := match ls with @@ -261,12 +262,12 @@ divergent def HashMap.get_in_list_loop | List.Nil => Result.fail .panic /- [hashmap::{hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 232:4-232:70 -/ + Source: 'tests/src/hashmap.rs', lines 234:4-234:70 -/ def HashMap.get_in_list (T : Type) (key : Usize) (ls : List T) : Result T := HashMap.get_in_list_loop T key ls /- [hashmap::{hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 247:4-247:55 -/ + Source: 'tests/src/hashmap.rs', lines 249:4-249:55 -/ def HashMap.get (T : Type) (self : HashMap T) (key : Usize) : Result T := do let hash ← hash_key key @@ -278,7 +279,7 @@ def HashMap.get (T : Type) (self : HashMap T) (key : Usize) : Result T := HashMap.get_in_list T key l /- [hashmap::{hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 253:4-262:5 -/ + Source: 'tests/src/hashmap.rs', lines 255:4-264:5 -/ divergent def HashMap.get_mut_in_list_loop (T : Type) (ls : List T) (key : Usize) : Result (T × (T → Result (List T))) @@ -301,7 +302,7 @@ divergent def HashMap.get_mut_in_list_loop | List.Nil => Result.fail .panic /- [hashmap::{hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 253:4-253:86 -/ + Source: 'tests/src/hashmap.rs', lines 255:4-255:86 -/ def HashMap.get_mut_in_list (T : Type) (ls : List T) (key : Usize) : Result (T × (T → Result (List T))) @@ -309,7 +310,7 @@ def HashMap.get_mut_in_list HashMap.get_mut_in_list_loop T ls key /- [hashmap::{hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:67 -/ + Source: 'tests/src/hashmap.rs', lines 267:4-267:67 -/ def HashMap.get_mut (T : Type) (self : HashMap T) (key : Usize) : Result (T × (T → Result (HashMap T))) @@ -331,7 +332,7 @@ def HashMap.get_mut Result.ok (t, back) /- [hashmap::{hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 273:4-299:5 -/ + Source: 'tests/src/hashmap.rs', lines 275:4-301:5 -/ divergent def HashMap.remove_from_list_loop (T : Type) (key : Usize) (ls : List T) : Result ((Option T) × (List T)) := match ls with @@ -350,13 +351,13 @@ divergent def HashMap.remove_from_list_loop | List.Nil => Result.ok (none, List.Nil) /- [hashmap::{hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 273:4-273:69 -/ + Source: 'tests/src/hashmap.rs', lines 275:4-275:69 -/ def HashMap.remove_from_list (T : Type) (key : Usize) (ls : List T) : Result ((Option T) × (List T)) := HashMap.remove_from_list_loop T key ls /- [hashmap::{hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 302:4-302:52 -/ + Source: 'tests/src/hashmap.rs', lines 304:4-304:52 -/ def HashMap.remove (T : Type) (self : HashMap T) (key : Usize) : Result ((Option T) × (HashMap T)) @@ -380,8 +381,17 @@ def HashMap.remove let v ← index_mut_back l1 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 -/ +def insert_on_disk + (key : Usize) (value : U64) (st : State) : Result (State × Unit) := + do + let (st1, hm) ← hashmap_utils.deserialize st + let hm1 ← HashMap.insert U64 hm key value + hashmap_utils.serialize hm1 st1 + /- [hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 323:0-323:10 -/ + Source: 'tests/src/hashmap.rs', lines 350:0-350:10 -/ def test1 : Result Unit := do let hm ← HashMap.new U64 diff --git a/tests/lean/Hashmap/FunsExternal.lean b/tests/lean/Hashmap/FunsExternal.lean new file mode 100644 index 00000000..7454696b --- /dev/null +++ b/tests/lean/Hashmap/FunsExternal.lean @@ -0,0 +1,19 @@ +-- [hashmap]: external functions. +import Base +import Hashmap.Types +open Primitives +open hashmap + +-- TODO: fill those bodies + +/- [hashmap::hashmap_utils::deserialize]: + Source: 'tests/src/hashmap.rs', lines 330:4-330:47 -/ +def hashmap_utils.deserialize + : State → Result (State × (HashMap U64)) := + fun _ => .fail .panic + +/- [hashmap::hashmap_utils::serialize]: + Source: 'tests/src/hashmap.rs', lines 325:4-325:46 -/ +def hashmap_utils.serialize + : HashMap U64 → State → Result (State × Unit) := + fun _ _ => .fail .panic diff --git a/tests/lean/Hashmap/FunsExternal_Template.lean b/tests/lean/Hashmap/FunsExternal_Template.lean new file mode 100644 index 00000000..16b4b6af --- /dev/null +++ b/tests/lean/Hashmap/FunsExternal_Template.lean @@ -0,0 +1,17 @@ +-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS +-- [hashmap]: external functions. +-- This is a template file: rename it to "FunsExternal.lean" and fill the holes. +import Base +import Hashmap.Types +open Primitives +open hashmap + +/- [hashmap::hashmap_utils::deserialize]: + Source: 'tests/src/hashmap.rs', lines 330:4-330:47 -/ +axiom hashmap_utils.deserialize : State → Result (State × (HashMap U64)) + +/- [hashmap::hashmap_utils::serialize]: + Source: 'tests/src/hashmap.rs', lines 325:4-325:46 -/ +axiom hashmap_utils.serialize + : HashMap U64 → State → Result (State × Unit) + diff --git a/tests/lean/Hashmap/Types.lean b/tests/lean/Hashmap/Types.lean index a98b972f..b4301106 100644 --- a/tests/lean/Hashmap/Types.lean +++ b/tests/lean/Hashmap/Types.lean @@ -1,18 +1,19 @@ -- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS -- [hashmap]: type definitions import Base +import Hashmap.TypesExternal open Primitives namespace hashmap /- [hashmap::List] - Source: 'tests/src/hashmap.rs', lines 27:0-27:16 -/ + 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::HashMap] - Source: 'tests/src/hashmap.rs', lines 43:0-43:21 -/ + Source: 'tests/src/hashmap.rs', lines 45:0-45:21 -/ structure HashMap (T : Type) where num_entries : Usize max_load_factor : (Usize × Usize) diff --git a/tests/lean/Hashmap/TypesExternal.lean b/tests/lean/Hashmap/TypesExternal.lean new file mode 100644 index 00000000..b32c4c42 --- /dev/null +++ b/tests/lean/Hashmap/TypesExternal.lean @@ -0,0 +1,7 @@ +-- [hashmap]: external types. +import Base +open Primitives + +/- The state type used in the state-error monad -/ +axiom State : Type + diff --git a/tests/lean/Hashmap/TypesExternal_Template.lean b/tests/lean/Hashmap/TypesExternal_Template.lean new file mode 100644 index 00000000..03c3d157 --- /dev/null +++ b/tests/lean/Hashmap/TypesExternal_Template.lean @@ -0,0 +1,9 @@ +-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS +-- [hashmap]: external types. +-- This is a template file: rename it to "TypesExternal.lean" and fill the holes. +import Base +open Primitives + +/- The state type used in the state-error monad -/ +axiom State : Type + diff --git a/tests/lean/HashmapMain.lean b/tests/lean/HashmapMain.lean deleted file mode 100644 index 1a4e7f82..00000000 --- a/tests/lean/HashmapMain.lean +++ /dev/null @@ -1 +0,0 @@ -import HashmapMain.Funs diff --git a/tests/lean/HashmapMain/Funs.lean b/tests/lean/HashmapMain/Funs.lean deleted file mode 100644 index 45d6b058..00000000 --- a/tests/lean/HashmapMain/Funs.lean +++ /dev/null @@ -1,463 +0,0 @@ --- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS --- [hashmap_main]: function definitions -import Base -import HashmapMain.Types -import HashmapMain.FunsExternal -open Primitives - -namespace hashmap_main - -/- [hashmap_main::hashmap::hash_key]: - Source: 'tests/src/hashmap.rs', lines 35:0-35:32 -/ -def hashmap.hash_key (k : Usize) : Result Usize := - Result.ok k - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: loop 0: - Source: 'tests/src/hashmap.rs', lines 58:4-64:5 -/ -divergent def hashmap.HashMap.allocate_slots_loop - (T : Type) (slots : alloc.vec.Vec (hashmap.List T)) (n : Usize) : - Result (alloc.vec.Vec (hashmap.List T)) - := - if n > 0#usize - then - do - let slots1 ← alloc.vec.Vec.push (hashmap.List T) slots hashmap.List.Nil - let n1 ← n - 1#usize - hashmap.HashMap.allocate_slots_loop T slots1 n1 - else Result.ok slots - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::allocate_slots]: - Source: 'tests/src/hashmap.rs', lines 58:4-58:76 -/ -def hashmap.HashMap.allocate_slots - (T : Type) (slots : alloc.vec.Vec (hashmap.List T)) (n : Usize) : - Result (alloc.vec.Vec (hashmap.List T)) - := - hashmap.HashMap.allocate_slots_loop T slots n - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new_with_capacity]: - Source: 'tests/src/hashmap.rs', lines 67:4-71:13 -/ -def hashmap.HashMap.new_with_capacity - (T : Type) (capacity : Usize) (max_load_dividend : Usize) - (max_load_divisor : Usize) : - Result (hashmap.HashMap T) - := - do - let slots ← - hashmap.HashMap.allocate_slots T (alloc.vec.Vec.new (hashmap.List T)) - capacity - let i ← capacity * max_load_dividend - let i1 ← i / max_load_divisor - Result.ok - { - num_entries := 0#usize, - max_load_factor := (max_load_dividend, max_load_divisor), - max_load := i1, - slots := slots - } - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::new]: - Source: 'tests/src/hashmap.rs', lines 83:4-83:24 -/ -def hashmap.HashMap.new (T : Type) : Result (hashmap.HashMap T) := - hashmap.HashMap.new_with_capacity T 32#usize 4#usize 5#usize - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: loop 0: - Source: 'tests/src/hashmap.rs', lines 88:4-96:5 -/ -divergent def hashmap.HashMap.clear_loop - (T : Type) (slots : alloc.vec.Vec (hashmap.List T)) (i : Usize) : - Result (alloc.vec.Vec (hashmap.List T)) - := - let i1 := alloc.vec.Vec.len (hashmap.List T) slots - if i < i1 - then - do - let (_, index_mut_back) ← - alloc.vec.Vec.index_mut (hashmap.List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (hashmap.List T)) slots i - let i2 ← i + 1#usize - let slots1 ← index_mut_back hashmap.List.Nil - hashmap.HashMap.clear_loop T slots1 i2 - else Result.ok slots - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::clear]: - Source: 'tests/src/hashmap.rs', lines 88:4-88:27 -/ -def hashmap.HashMap.clear - (T : Type) (self : hashmap.HashMap T) : Result (hashmap.HashMap T) := - do - let hm ← hashmap.HashMap.clear_loop T self.slots 0#usize - Result.ok { self with num_entries := 0#usize, slots := hm } - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::len]: - Source: 'tests/src/hashmap.rs', lines 98:4-98:30 -/ -def hashmap.HashMap.len (T : Type) (self : hashmap.HashMap T) : Result Usize := - Result.ok self.num_entries - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 105:4-122:5 -/ -divergent def hashmap.HashMap.insert_in_list_loop - (T : Type) (key : Usize) (value : T) (ls : hashmap.List T) : - Result (Bool × (hashmap.List T)) - := - match ls with - | hashmap.List.Cons ckey cvalue tl => - if ckey = key - then Result.ok (false, hashmap.List.Cons ckey value tl) - else - do - let (b, tl1) ← hashmap.HashMap.insert_in_list_loop T key value tl - Result.ok (b, hashmap.List.Cons ckey cvalue tl1) - | hashmap.List.Nil => - Result.ok (true, hashmap.List.Cons key value hashmap.List.Nil) - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_in_list]: - Source: 'tests/src/hashmap.rs', lines 105:4-105:71 -/ -def hashmap.HashMap.insert_in_list - (T : Type) (key : Usize) (value : T) (ls : hashmap.List T) : - Result (Bool × (hashmap.List T)) - := - hashmap.HashMap.insert_in_list_loop T key value ls - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert_no_resize]: - Source: 'tests/src/hashmap.rs', lines 125:4-125:54 -/ -def hashmap.HashMap.insert_no_resize - (T : Type) (self : hashmap.HashMap T) (key : Usize) (value : T) : - Result (hashmap.HashMap T) - := - do - let hash ← hashmap.hash_key key - let i := alloc.vec.Vec.len (hashmap.List T) self.slots - let hash_mod ← hash % i - let (l, index_mut_back) ← - alloc.vec.Vec.index_mut (hashmap.List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (hashmap.List T)) self.slots - hash_mod - let (inserted, l1) ← hashmap.HashMap.insert_in_list T key value l - if inserted - then - do - let i1 ← self.num_entries + 1#usize - let v ← index_mut_back l1 - Result.ok { self with num_entries := i1, slots := v } - else do - let v ← index_mut_back l1 - Result.ok { self with slots := v } - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 191:4-204:5 -/ -divergent def hashmap.HashMap.move_elements_from_list_loop - (T : Type) (ntable : hashmap.HashMap T) (ls : hashmap.List T) : - Result (hashmap.HashMap T) - := - match ls with - | hashmap.List.Cons k v tl => - do - let ntable1 ← hashmap.HashMap.insert_no_resize T ntable k v - hashmap.HashMap.move_elements_from_list_loop T ntable1 tl - | hashmap.List.Nil => Result.ok ntable - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements_from_list]: - Source: 'tests/src/hashmap.rs', lines 191:4-191:72 -/ -def hashmap.HashMap.move_elements_from_list - (T : Type) (ntable : hashmap.HashMap T) (ls : hashmap.List T) : - Result (hashmap.HashMap T) - := - hashmap.HashMap.move_elements_from_list_loop T ntable ls - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: loop 0: - Source: 'tests/src/hashmap.rs', lines 179:4-188:5 -/ -divergent def hashmap.HashMap.move_elements_loop - (T : Type) (ntable : hashmap.HashMap T) - (slots : alloc.vec.Vec (hashmap.List T)) (i : Usize) : - Result ((hashmap.HashMap T) × (alloc.vec.Vec (hashmap.List T))) - := - let i1 := alloc.vec.Vec.len (hashmap.List T) slots - if i < i1 - then - do - let (l, index_mut_back) ← - alloc.vec.Vec.index_mut (hashmap.List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (hashmap.List T)) slots i - let (ls, l1) := core.mem.replace (hashmap.List T) l hashmap.List.Nil - let ntable1 ← hashmap.HashMap.move_elements_from_list T ntable ls - let i2 ← i + 1#usize - let slots1 ← index_mut_back l1 - hashmap.HashMap.move_elements_loop T ntable1 slots1 i2 - else Result.ok (ntable, slots) - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::move_elements]: - Source: 'tests/src/hashmap.rs', lines 179:4-179:95 -/ -def hashmap.HashMap.move_elements - (T : Type) (ntable : hashmap.HashMap T) - (slots : alloc.vec.Vec (hashmap.List T)) (i : Usize) : - Result ((hashmap.HashMap T) × (alloc.vec.Vec (hashmap.List T))) - := - hashmap.HashMap.move_elements_loop T ntable slots i - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::try_resize]: - Source: 'tests/src/hashmap.rs', lines 148:4-148:28 -/ -def hashmap.HashMap.try_resize - (T : Type) (self : hashmap.HashMap T) : Result (hashmap.HashMap T) := - do - let max_usize ← Scalar.cast .Usize core_u32_max - let capacity := alloc.vec.Vec.len (hashmap.List T) self.slots - let n1 ← max_usize / 2#usize - let (i, i1) := self.max_load_factor - let i2 ← n1 / i - if capacity <= i2 - then - do - let i3 ← capacity * 2#usize - let ntable ← hashmap.HashMap.new_with_capacity T i3 i i1 - let p ← hashmap.HashMap.move_elements T ntable self.slots 0#usize - let (ntable1, _) := p - Result.ok - { - ntable1 - with - num_entries := self.num_entries, max_load_factor := (i, i1) - } - else Result.ok { self with max_load_factor := (i, i1) } - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::insert]: - Source: 'tests/src/hashmap.rs', lines 137:4-137:48 -/ -def hashmap.HashMap.insert - (T : Type) (self : hashmap.HashMap T) (key : Usize) (value : T) : - Result (hashmap.HashMap T) - := - do - let self1 ← hashmap.HashMap.insert_no_resize T self key value - let i ← hashmap.HashMap.len T self1 - if i > self1.max_load - then hashmap.HashMap.try_resize T self1 - else Result.ok self1 - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 214:4-227:5 -/ -divergent def hashmap.HashMap.contains_key_in_list_loop - (T : Type) (key : Usize) (ls : hashmap.List T) : Result Bool := - match ls with - | hashmap.List.Cons ckey _ tl => - if ckey = key - then Result.ok true - else hashmap.HashMap.contains_key_in_list_loop T key tl - | hashmap.List.Nil => Result.ok false - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key_in_list]: - Source: 'tests/src/hashmap.rs', lines 214:4-214:68 -/ -def hashmap.HashMap.contains_key_in_list - (T : Type) (key : Usize) (ls : hashmap.List T) : Result Bool := - hashmap.HashMap.contains_key_in_list_loop T key ls - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::contains_key]: - Source: 'tests/src/hashmap.rs', lines 207:4-207:49 -/ -def hashmap.HashMap.contains_key - (T : Type) (self : hashmap.HashMap T) (key : Usize) : Result Bool := - do - let hash ← hashmap.hash_key key - let i := alloc.vec.Vec.len (hashmap.List T) self.slots - let hash_mod ← hash % i - let l ← - alloc.vec.Vec.index (hashmap.List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (hashmap.List T)) self.slots - hash_mod - hashmap.HashMap.contains_key_in_list T key l - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 232:4-245:5 -/ -divergent def hashmap.HashMap.get_in_list_loop - (T : Type) (key : Usize) (ls : hashmap.List T) : Result T := - match ls with - | hashmap.List.Cons ckey cvalue tl => - if ckey = key - then Result.ok cvalue - else hashmap.HashMap.get_in_list_loop T key tl - | hashmap.List.Nil => Result.fail .panic - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_in_list]: - Source: 'tests/src/hashmap.rs', lines 232:4-232:70 -/ -def hashmap.HashMap.get_in_list - (T : Type) (key : Usize) (ls : hashmap.List T) : Result T := - hashmap.HashMap.get_in_list_loop T key ls - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get]: - Source: 'tests/src/hashmap.rs', lines 247:4-247:55 -/ -def hashmap.HashMap.get - (T : Type) (self : hashmap.HashMap T) (key : Usize) : Result T := - do - let hash ← hashmap.hash_key key - let i := alloc.vec.Vec.len (hashmap.List T) self.slots - let hash_mod ← hash % i - let l ← - alloc.vec.Vec.index (hashmap.List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (hashmap.List T)) self.slots - hash_mod - hashmap.HashMap.get_in_list T key l - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 253:4-262:5 -/ -divergent def hashmap.HashMap.get_mut_in_list_loop - (T : Type) (ls : hashmap.List T) (key : Usize) : - Result (T × (T → Result (hashmap.List T))) - := - match ls with - | hashmap.List.Cons ckey cvalue tl => - if ckey = key - then - let back := fun ret => Result.ok (hashmap.List.Cons ckey ret tl) - Result.ok (cvalue, back) - else - do - let (t, back) ← hashmap.HashMap.get_mut_in_list_loop T tl key - let back1 := - fun ret => - do - let tl1 ← back ret - Result.ok (hashmap.List.Cons ckey cvalue tl1) - Result.ok (t, back1) - | hashmap.List.Nil => Result.fail .panic - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut_in_list]: - Source: 'tests/src/hashmap.rs', lines 253:4-253:86 -/ -def hashmap.HashMap.get_mut_in_list - (T : Type) (ls : hashmap.List T) (key : Usize) : - Result (T × (T → Result (hashmap.List T))) - := - hashmap.HashMap.get_mut_in_list_loop T ls key - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::get_mut]: - Source: 'tests/src/hashmap.rs', lines 265:4-265:67 -/ -def hashmap.HashMap.get_mut - (T : Type) (self : hashmap.HashMap T) (key : Usize) : - Result (T × (T → Result (hashmap.HashMap T))) - := - do - let hash ← hashmap.hash_key key - let i := alloc.vec.Vec.len (hashmap.List T) self.slots - let hash_mod ← hash % i - let (l, index_mut_back) ← - alloc.vec.Vec.index_mut (hashmap.List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (hashmap.List T)) self.slots - hash_mod - let (t, get_mut_in_list_back) ← hashmap.HashMap.get_mut_in_list T l key - let back := - fun ret => - do - let l1 ← get_mut_in_list_back ret - let v ← index_mut_back l1 - Result.ok { self with slots := v } - Result.ok (t, back) - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: loop 0: - Source: 'tests/src/hashmap.rs', lines 273:4-299:5 -/ -divergent def hashmap.HashMap.remove_from_list_loop - (T : Type) (key : Usize) (ls : hashmap.List T) : - Result ((Option T) × (hashmap.List T)) - := - match ls with - | hashmap.List.Cons ckey t tl => - if ckey = key - then - let (mv_ls, _) := - core.mem.replace (hashmap.List T) (hashmap.List.Cons ckey t tl) - hashmap.List.Nil - match mv_ls with - | hashmap.List.Cons _ cvalue tl1 => Result.ok (some cvalue, tl1) - | hashmap.List.Nil => Result.fail .panic - else - do - let (o, tl1) ← hashmap.HashMap.remove_from_list_loop T key tl - Result.ok (o, hashmap.List.Cons ckey t tl1) - | hashmap.List.Nil => Result.ok (none, hashmap.List.Nil) - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove_from_list]: - Source: 'tests/src/hashmap.rs', lines 273:4-273:69 -/ -def hashmap.HashMap.remove_from_list - (T : Type) (key : Usize) (ls : hashmap.List T) : - Result ((Option T) × (hashmap.List T)) - := - hashmap.HashMap.remove_from_list_loop T key ls - -/- [hashmap_main::hashmap::{hashmap_main::hashmap::HashMap}::remove]: - Source: 'tests/src/hashmap.rs', lines 302:4-302:52 -/ -def hashmap.HashMap.remove - (T : Type) (self : hashmap.HashMap T) (key : Usize) : - Result ((Option T) × (hashmap.HashMap T)) - := - do - let hash ← hashmap.hash_key key - let i := alloc.vec.Vec.len (hashmap.List T) self.slots - let hash_mod ← hash % i - let (l, index_mut_back) ← - alloc.vec.Vec.index_mut (hashmap.List T) Usize - (core.slice.index.SliceIndexUsizeSliceTInst (hashmap.List T)) self.slots - hash_mod - let (x, l1) ← hashmap.HashMap.remove_from_list T key l - match x with - | none => - do - let v ← index_mut_back l1 - Result.ok (none, { self with slots := v }) - | some x1 => - do - let i1 ← self.num_entries - 1#usize - let v ← index_mut_back l1 - Result.ok (some x1, { self with num_entries := i1, slots := v }) - -/- [hashmap_main::hashmap::test1]: - Source: 'tests/src/hashmap.rs', lines 323:0-323:10 -/ -def hashmap.test1 : Result Unit := - do - let hm ← hashmap.HashMap.new U64 - let hm1 ← hashmap.HashMap.insert U64 hm 0#usize 42#u64 - let hm2 ← hashmap.HashMap.insert U64 hm1 128#usize 18#u64 - let hm3 ← hashmap.HashMap.insert U64 hm2 1024#usize 138#u64 - let hm4 ← hashmap.HashMap.insert U64 hm3 1056#usize 256#u64 - let i ← hashmap.HashMap.get U64 hm4 128#usize - if ¬ (i = 18#u64) - then Result.fail .panic - else - do - let (_, get_mut_back) ← hashmap.HashMap.get_mut U64 hm4 1024#usize - let hm5 ← get_mut_back 56#u64 - let i1 ← hashmap.HashMap.get U64 hm5 1024#usize - if ¬ (i1 = 56#u64) - then Result.fail .panic - else - do - let (x, hm6) ← hashmap.HashMap.remove U64 hm5 1024#usize - match x with - | none => Result.fail .panic - | some x1 => - if ¬ (x1 = 56#u64) - then Result.fail .panic - else - do - let i2 ← hashmap.HashMap.get U64 hm6 0#usize - if ¬ (i2 = 42#u64) - then Result.fail .panic - else - do - let i3 ← hashmap.HashMap.get U64 hm6 128#usize - if ¬ (i3 = 18#u64) - then Result.fail .panic - else - do - let i4 ← hashmap.HashMap.get U64 hm6 1056#usize - if ¬ (i4 = 256#u64) - then Result.fail .panic - else Result.ok () - -/- [hashmap_main::insert_on_disk]: - Source: 'tests/src/hashmap_main.rs', lines 13:0-13:43 -/ -def insert_on_disk - (key : Usize) (value : U64) (st : State) : Result (State × Unit) := - do - let (st1, hm) ← hashmap_utils.deserialize st - let hm1 ← hashmap.HashMap.insert U64 hm key value - hashmap_utils.serialize hm1 st1 - -/- [hashmap_main::main]: - Source: 'tests/src/hashmap_main.rs', lines 22:0-22:13 -/ -def main : Result Unit := - Result.ok () - -end hashmap_main diff --git a/tests/lean/HashmapMain/FunsExternal.lean b/tests/lean/HashmapMain/FunsExternal.lean deleted file mode 100644 index b394b32b..00000000 --- a/tests/lean/HashmapMain/FunsExternal.lean +++ /dev/null @@ -1,17 +0,0 @@ --- [hashmap_main]: templates for the external functions. -import Base -import HashmapMain.Types -open Primitives -open hashmap_main - --- TODO: fill those bodies - -/- [hashmap_main::hashmap_utils::deserialize] -/ -def hashmap_utils.deserialize - : State → Result (State × (hashmap.HashMap U64)) := - fun _ => .fail .panic - -/- [hashmap_main::hashmap_utils::serialize] -/ -def hashmap_utils.serialize - : hashmap.HashMap U64 → State → Result (State × Unit) := - fun _ _ => .fail .panic diff --git a/tests/lean/HashmapMain/FunsExternal_Template.lean b/tests/lean/HashmapMain/FunsExternal_Template.lean deleted file mode 100644 index 1a6c40d5..00000000 --- a/tests/lean/HashmapMain/FunsExternal_Template.lean +++ /dev/null @@ -1,18 +0,0 @@ --- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS --- [hashmap_main]: external functions. --- This is a template file: rename it to "FunsExternal.lean" and fill the holes. -import Base -import HashmapMain.Types -open Primitives -open hashmap_main - -/- [hashmap_main::hashmap_utils::deserialize]: - Source: 'tests/src/hashmap_utils.rs', lines 11:0-11:43 -/ -axiom hashmap_utils.deserialize - : State → Result (State × (hashmap.HashMap U64)) - -/- [hashmap_main::hashmap_utils::serialize]: - Source: 'tests/src/hashmap_utils.rs', lines 6:0-6:42 -/ -axiom hashmap_utils.serialize - : hashmap.HashMap U64 → State → Result (State × Unit) - diff --git a/tests/lean/HashmapMain/Opaque.lean b/tests/lean/HashmapMain/Opaque.lean deleted file mode 100644 index a410b253..00000000 --- a/tests/lean/HashmapMain/Opaque.lean +++ /dev/null @@ -1,18 +0,0 @@ --- [hashmap_main]: opaque function definitions -import Base -import HashmapMain.Types -open Primitives - -namespace hashmap_main - -structure OpaqueDefs where - - /- [hashmap_main::hashmap_utils::deserialize] -/ - hashmap_utils.deserialize_fwd - : State → Result (State × (hashmap_hash_map_t U64)) - - /- [hashmap_main::hashmap_utils::serialize] -/ - hashmap_utils.serialize_fwd - : hashmap_hash_map_t U64 → State → Result (State × Unit) - -end hashmap_main diff --git a/tests/lean/HashmapMain/Types.lean b/tests/lean/HashmapMain/Types.lean deleted file mode 100644 index 54f11e1e..00000000 --- a/tests/lean/HashmapMain/Types.lean +++ /dev/null @@ -1,23 +0,0 @@ --- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS --- [hashmap_main]: type definitions -import Base -import HashmapMain.TypesExternal -open Primitives - -namespace hashmap_main - -/- [hashmap_main::hashmap::List] - Source: 'tests/src/hashmap.rs', lines 27:0-27:16 -/ -inductive hashmap.List (T : Type) := -| Cons : Usize → T → hashmap.List T → hashmap.List T -| Nil : hashmap.List T - -/- [hashmap_main::hashmap::HashMap] - Source: 'tests/src/hashmap.rs', lines 43:0-43:21 -/ -structure hashmap.HashMap (T : Type) where - num_entries : Usize - max_load_factor : (Usize × Usize) - max_load : Usize - slots : alloc.vec.Vec (hashmap.List T) - -end hashmap_main diff --git a/tests/lean/HashmapMain/TypesExternal.lean b/tests/lean/HashmapMain/TypesExternal.lean deleted file mode 100644 index 4e1cdbe9..00000000 --- a/tests/lean/HashmapMain/TypesExternal.lean +++ /dev/null @@ -1,7 +0,0 @@ --- [hashmap_main]: external types. -import Base -open Primitives - -/- The state type used in the state-error monad -/ -axiom State : Type - diff --git a/tests/lean/HashmapMain/TypesExternal_Template.lean b/tests/lean/HashmapMain/TypesExternal_Template.lean deleted file mode 100644 index cfa8bbb1..00000000 --- a/tests/lean/HashmapMain/TypesExternal_Template.lean +++ /dev/null @@ -1,9 +0,0 @@ --- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS --- [hashmap_main]: external types. --- This is a template file: rename it to "TypesExternal.lean" and fill the holes. -import Base -open Primitives - -/- The state type used in the state-error monad -/ -axiom State : Type - diff --git a/tests/lean/lakefile.lean b/tests/lean/lakefile.lean index 3a777824..bb13ddf2 100644 --- a/tests/lean/lakefile.lean +++ b/tests/lean/lakefile.lean @@ -13,7 +13,6 @@ package «tests» {} @[default_target] lean_lib Constants @[default_target] lean_lib External @[default_target] lean_lib Hashmap -@[default_target] lean_lib HashmapMain @[default_target] lean_lib Loops @[default_target] lean_lib NoNestedBorrows @[default_target] lean_lib Paper diff --git a/tests/src/hashmap.rs b/tests/src/hashmap.rs index 4552e4f2..ccb96e1e 100644 --- a/tests/src/hashmap.rs +++ b/tests/src/hashmap.rs @@ -1,9 +1,11 @@ +//@ charon-args=--opaque=hashmap_utils +//@ aeneas-args=-state -split-files //@ [coq] aeneas-args=-use-fuel -//@ aeneas-args=-split-files //@ [fstar] aeneas-args=-decreases-clauses -template-clauses //@ [lean] aeneas-args=-no-gen-lib-entry // ^ the `-no-gen-lib-entry` is because we add a custom import in the Hashmap.lean file: we do not // want to overwrite it. +// Possible to add `--no-code-duplication` if we use the optimized MIR // TODO: reactivate -test-trans-units //! A hashmap implementation. @@ -314,6 +316,31 @@ impl HashMap { } } +// This is a module so we can tell charon to leave it opaque +mod hashmap_utils { + use crate::*; + + /// Serialize a hash map - we don't have traits, so we fix the type of the + /// values (this is not the interesting part anyway) + pub(crate) fn serialize(_hm: HashMap) { + unimplemented!(); + } + /// Deserialize a hash map - we don't have traits, so we fix the type of the + /// values (this is not the interesting part anyway) + pub(crate) fn deserialize() -> HashMap { + unimplemented!(); + } +} + +pub fn insert_on_disk(key: Key, value: u64) { + // Deserialize + let mut hm = hashmap_utils::deserialize(); + // Update + hm.insert(key, value); + // Serialize + hashmap_utils::serialize(hm); +} + /// I currently can't retrieve functions marked with the attribute #[test], /// while I want to extract the unit tests and use the normalize on them, /// so I have to define the test functions somewhere and call them from diff --git a/tests/src/hashmap_main.rs b/tests/src/hashmap_main.rs deleted file mode 100644 index 0c827844..00000000 --- a/tests/src/hashmap_main.rs +++ /dev/null @@ -1,22 +0,0 @@ -//@ charon-args=--opaque=hashmap_utils -//@ aeneas-args=-state -split-files -//@ [coq] aeneas-args=-use-fuel -//@ [fstar] aeneas-args=-decreases-clauses -template-clauses -// Possible to add `--no-code-duplication` if we use the optimized MIR -// TODO: reactivate -test-trans-units -mod hashmap; -mod hashmap_utils; - -use crate::hashmap::*; -use crate::hashmap_utils::*; - -pub fn insert_on_disk(key: Key, value: u64) { - // Deserialize - let mut hm = deserialize(); - // Update - hm.insert(key, value); - // Serialize - serialize(hm); -} - -pub fn main() {} diff --git a/tests/src/hashmap_utils.rs b/tests/src/hashmap_utils.rs deleted file mode 100644 index 33de49e1..00000000 --- a/tests/src/hashmap_utils.rs +++ /dev/null @@ -1,13 +0,0 @@ -//@ skip -use crate::hashmap::*; - -/// Serialize a hash map - we don't have traits, so we fix the type of the -/// values (this is not the interesting part anyway) -pub(crate) fn serialize(_hm: HashMap) { - unimplemented!(); -} -/// Deserialize a hash map - we don't have traits, so we fix the type of the -/// values (this is not the interesting part anyway) -pub(crate) fn deserialize() -> HashMap { - unimplemented!(); -} -- cgit v1.2.3 From f77e3f9cbe2ea7abeb4be815bdbf33d0c98076c2 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Mon, 27 May 2024 13:28:28 +0200 Subject: tests: Rename betree_main -> betree --- flake.nix | 4 +- tests/coq/betree/BetreeMain_Funs.v | 837 ---------- tests/coq/betree/BetreeMain_FunsExternal.v | 39 - .../coq/betree/BetreeMain_FunsExternal_Template.v | 40 - tests/coq/betree/BetreeMain_Types.v | 118 -- tests/coq/betree/BetreeMain_TypesExternal.v | 14 - .../coq/betree/BetreeMain_TypesExternal_Template.v | 15 - tests/coq/betree/Betree_Funs.v | 837 ++++++++++ tests/coq/betree/Betree_FunsExternal.v | 39 + tests/coq/betree/Betree_FunsExternal_Template.v | 40 + tests/coq/betree/Betree_Types.v | 118 ++ tests/coq/betree/Betree_TypesExternal.v | 14 + tests/coq/betree/Betree_TypesExternal_Template.v | 15 + tests/coq/betree/_CoqProject | 12 +- tests/fstar/betree/Betree.Clauses.Template.fst | 117 ++ tests/fstar/betree/Betree.Clauses.fst | 210 +++ tests/fstar/betree/Betree.Funs.fst | 676 ++++++++ tests/fstar/betree/Betree.FunsExternal.fsti | 30 + tests/fstar/betree/Betree.Types.fst | 61 + tests/fstar/betree/Betree.TypesExternal.fsti | 10 + tests/fstar/betree/BetreeMain.Clauses.Template.fst | 117 -- tests/fstar/betree/BetreeMain.Clauses.fst | 210 --- tests/fstar/betree/BetreeMain.Funs.fst | 676 -------- tests/fstar/betree/BetreeMain.FunsExternal.fsti | 30 - tests/fstar/betree/BetreeMain.Types.fst | 61 - tests/fstar/betree/BetreeMain.TypesExternal.fsti | 10 - tests/hol4/betree/betreeMain_FunsScript.sml | 1212 -------------- tests/hol4/betree/betreeMain_FunsTheory.sig | 1230 -------------- tests/hol4/betree/betreeMain_OpaqueScript.sml | 26 - tests/hol4/betree/betreeMain_OpaqueTheory.sig | 11 - tests/hol4/betree/betreeMain_TypesScript.sml | 76 - tests/hol4/betree/betreeMain_TypesTheory.sig | 1751 -------------------- tests/hol4/betree/betree_FunsScript.sml | 1212 ++++++++++++++ tests/hol4/betree/betree_FunsTheory.sig | 1230 ++++++++++++++ tests/hol4/betree/betree_OpaqueScript.sml | 26 + tests/hol4/betree/betree_OpaqueTheory.sig | 11 + tests/hol4/betree/betree_TypesScript.sml | 76 + tests/hol4/betree/betree_TypesTheory.sig | 1751 ++++++++++++++++++++ tests/lean/Betree/Funs.lean | 699 ++++++++ tests/lean/Betree/FunsExternal.lean | 30 + tests/lean/Betree/FunsExternal_Template.lean | 30 + tests/lean/Betree/Types.lean | 83 + tests/lean/Betree/TypesExternal.lean | 7 + tests/lean/Betree/TypesExternal_Template.lean | 9 + tests/lean/BetreeMain.lean | 1 - tests/lean/BetreeMain/Funs.lean | 699 -------- tests/lean/BetreeMain/FunsExternal.lean | 30 - tests/lean/BetreeMain/FunsExternal_Template.lean | 30 - tests/lean/BetreeMain/Types.lean | 83 - tests/lean/BetreeMain/TypesExternal.lean | 7 - tests/lean/BetreeMain/TypesExternal_Template.lean | 9 - tests/lean/lakefile.lean | 2 +- tests/test_runner/run_test.ml | 7 +- 53 files changed, 7342 insertions(+), 7346 deletions(-) delete mode 100644 tests/coq/betree/BetreeMain_Funs.v delete mode 100644 tests/coq/betree/BetreeMain_FunsExternal.v delete mode 100644 tests/coq/betree/BetreeMain_FunsExternal_Template.v delete mode 100644 tests/coq/betree/BetreeMain_Types.v delete mode 100644 tests/coq/betree/BetreeMain_TypesExternal.v delete mode 100644 tests/coq/betree/BetreeMain_TypesExternal_Template.v create mode 100644 tests/coq/betree/Betree_Funs.v create mode 100644 tests/coq/betree/Betree_FunsExternal.v create mode 100644 tests/coq/betree/Betree_FunsExternal_Template.v create mode 100644 tests/coq/betree/Betree_Types.v create mode 100644 tests/coq/betree/Betree_TypesExternal.v create mode 100644 tests/coq/betree/Betree_TypesExternal_Template.v create mode 100644 tests/fstar/betree/Betree.Clauses.Template.fst create mode 100644 tests/fstar/betree/Betree.Clauses.fst create mode 100644 tests/fstar/betree/Betree.Funs.fst create mode 100644 tests/fstar/betree/Betree.FunsExternal.fsti create mode 100644 tests/fstar/betree/Betree.Types.fst create mode 100644 tests/fstar/betree/Betree.TypesExternal.fsti delete mode 100644 tests/fstar/betree/BetreeMain.Clauses.Template.fst delete mode 100644 tests/fstar/betree/BetreeMain.Clauses.fst delete mode 100644 tests/fstar/betree/BetreeMain.Funs.fst delete mode 100644 tests/fstar/betree/BetreeMain.FunsExternal.fsti delete mode 100644 tests/fstar/betree/BetreeMain.Types.fst delete mode 100644 tests/fstar/betree/BetreeMain.TypesExternal.fsti delete mode 100644 tests/hol4/betree/betreeMain_FunsScript.sml delete mode 100644 tests/hol4/betree/betreeMain_FunsTheory.sig delete mode 100644 tests/hol4/betree/betreeMain_OpaqueScript.sml delete mode 100644 tests/hol4/betree/betreeMain_OpaqueTheory.sig delete mode 100644 tests/hol4/betree/betreeMain_TypesScript.sml delete mode 100644 tests/hol4/betree/betreeMain_TypesTheory.sig create mode 100644 tests/hol4/betree/betree_FunsScript.sml create mode 100644 tests/hol4/betree/betree_FunsTheory.sig create mode 100644 tests/hol4/betree/betree_OpaqueScript.sml create mode 100644 tests/hol4/betree/betree_OpaqueTheory.sig create mode 100644 tests/hol4/betree/betree_TypesScript.sml create mode 100644 tests/hol4/betree/betree_TypesTheory.sig create mode 100644 tests/lean/Betree/Funs.lean create mode 100644 tests/lean/Betree/FunsExternal.lean create mode 100644 tests/lean/Betree/FunsExternal_Template.lean create mode 100644 tests/lean/Betree/Types.lean create mode 100644 tests/lean/Betree/TypesExternal.lean create mode 100644 tests/lean/Betree/TypesExternal_Template.lean delete mode 100644 tests/lean/BetreeMain.lean delete mode 100644 tests/lean/BetreeMain/Funs.lean delete mode 100644 tests/lean/BetreeMain/FunsExternal.lean delete mode 100644 tests/lean/BetreeMain/FunsExternal_Template.lean delete mode 100644 tests/lean/BetreeMain/Types.lean delete mode 100644 tests/lean/BetreeMain/TypesExternal.lean delete mode 100644 tests/lean/BetreeMain/TypesExternal_Template.lean diff --git a/flake.nix b/flake.nix index f1a2258a..654c0006 100644 --- a/flake.nix +++ b/flake.nix @@ -90,7 +90,7 @@ betree-llbc = charon.extractCrateWithCharon.${system} { name = "betree"; src = ./tests/src/betree; - charonFlags = "--polonius --opaque=betree_utils --crate betree_main"; + charonFlags = "--polonius --opaque=betree_utils"; craneExtraArgs.checkPhaseCargoCommand = '' cargo rustc -- --test -Zpolonius ./target/debug/betree @@ -125,7 +125,7 @@ mkdir tests/llbc export LLBC_DIR=tests/llbc # Copy over the llbc file we can't generate ourselves. - cp ${betree-llbc}/llbc/betree_main.llbc $LLBC_DIR + cp ${betree-llbc}/llbc/betree.llbc $LLBC_DIR # Run the tests with extra sanity checks enabled IN_CI=1 make test-all -j $NIX_BUILD_CORES diff --git a/tests/coq/betree/BetreeMain_Funs.v b/tests/coq/betree/BetreeMain_Funs.v deleted file mode 100644 index e0a1d8a2..00000000 --- a/tests/coq/betree/BetreeMain_Funs.v +++ /dev/null @@ -1,837 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [betree_main]: function definitions *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Require Import BetreeMain_Types. -Include BetreeMain_Types. -Require Import BetreeMain_FunsExternal. -Include BetreeMain_FunsExternal. -Module BetreeMain_Funs. - -(** [betree_main::betree::load_internal_node]: - Source: 'src/betree.rs', lines 36:0-36:52 *) -Definition betree_load_internal_node - (id : u64) (st : state) : - result (state * (betree_List_t (u64 * betree_Message_t))) - := - betree_utils_load_internal_node id st -. - -(** [betree_main::betree::store_internal_node]: - Source: 'src/betree.rs', lines 41:0-41:60 *) -Definition betree_store_internal_node - (id : u64) (content : betree_List_t (u64 * betree_Message_t)) (st : state) : - result (state * unit) - := - betree_utils_store_internal_node id content st -. - -(** [betree_main::betree::load_leaf_node]: - Source: 'src/betree.rs', lines 46:0-46:44 *) -Definition betree_load_leaf_node - (id : u64) (st : state) : result (state * (betree_List_t (u64 * u64))) := - betree_utils_load_leaf_node id st -. - -(** [betree_main::betree::store_leaf_node]: - Source: 'src/betree.rs', lines 51:0-51:52 *) -Definition betree_store_leaf_node - (id : u64) (content : betree_List_t (u64 * u64)) (st : state) : - result (state * unit) - := - betree_utils_store_leaf_node id content st -. - -(** [betree_main::betree::fresh_node_id]: - Source: 'src/betree.rs', lines 55:0-55:48 *) -Definition betree_fresh_node_id (counter : u64) : result (u64 * u64) := - counter1 <- u64_add counter 1%u64; Ok (counter, counter1) -. - -(** [betree_main::betree::{betree_main::betree::NodeIdCounter}::new]: - Source: 'src/betree.rs', lines 206:4-206:20 *) -Definition betree_NodeIdCounter_new : result betree_NodeIdCounter_t := - Ok {| betree_NodeIdCounter_next_node_id := 0%u64 |} -. - -(** [betree_main::betree::{betree_main::betree::NodeIdCounter}::fresh_id]: - Source: 'src/betree.rs', lines 210:4-210:36 *) -Definition betree_NodeIdCounter_fresh_id - (self : betree_NodeIdCounter_t) : result (u64 * betree_NodeIdCounter_t) := - i <- u64_add self.(betree_NodeIdCounter_next_node_id) 1%u64; - Ok (self.(betree_NodeIdCounter_next_node_id), - {| betree_NodeIdCounter_next_node_id := i |}) -. - -(** [betree_main::betree::upsert_update]: - Source: 'src/betree.rs', lines 234:0-234:70 *) -Definition betree_upsert_update - (prev : option u64) (st : betree_UpsertFunState_t) : result u64 := - match prev with - | None => - match st with - | Betree_UpsertFunState_Add v => Ok v - | Betree_UpsertFunState_Sub _ => Ok 0%u64 - end - | Some prev1 => - match st with - | Betree_UpsertFunState_Add v => - margin <- u64_sub core_u64_max prev1; - if margin s>= v then u64_add prev1 v else Ok core_u64_max - | Betree_UpsertFunState_Sub v => - if prev1 s>= v then u64_sub prev1 v else Ok 0%u64 - end - end -. - -(** [betree_main::betree::{betree_main::betree::List#1}::len]: - Source: 'src/betree.rs', lines 276:4-276:24 *) -Fixpoint betree_List_len - (T : Type) (n : nat) (self : betree_List_t T) : result u64 := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match self with - | Betree_List_Cons _ tl => i <- betree_List_len T n1 tl; u64_add 1%u64 i - | Betree_List_Nil => Ok 0%u64 - end - end -. - -(** [betree_main::betree::{betree_main::betree::List#1}::split_at]: - Source: 'src/betree.rs', lines 284:4-284:51 *) -Fixpoint betree_List_split_at - (T : Type) (n : nat) (self : betree_List_t T) (n1 : u64) : - result ((betree_List_t T) * (betree_List_t T)) - := - match n with - | O => Fail_ OutOfFuel - | S n2 => - if n1 s= 0%u64 - then Ok (Betree_List_Nil, self) - else - match self with - | Betree_List_Cons hd tl => - i <- u64_sub n1 1%u64; - p <- betree_List_split_at T n2 tl i; - let (ls0, ls1) := p in - Ok (Betree_List_Cons hd ls0, ls1) - | Betree_List_Nil => Fail_ Failure - end - end -. - -(** [betree_main::betree::{betree_main::betree::List#1}::push_front]: - Source: 'src/betree.rs', lines 299:4-299:34 *) -Definition betree_List_push_front - (T : Type) (self : betree_List_t T) (x : T) : result (betree_List_t T) := - let (tl, _) := core_mem_replace (betree_List_t T) self Betree_List_Nil in - Ok (Betree_List_Cons x tl) -. - -(** [betree_main::betree::{betree_main::betree::List#1}::pop_front]: - Source: 'src/betree.rs', lines 306:4-306:32 *) -Definition betree_List_pop_front - (T : Type) (self : betree_List_t T) : result (T * (betree_List_t T)) := - let (ls, _) := core_mem_replace (betree_List_t T) self Betree_List_Nil in - match ls with - | Betree_List_Cons x tl => Ok (x, tl) - | Betree_List_Nil => Fail_ Failure - end -. - -(** [betree_main::betree::{betree_main::betree::List#1}::hd]: - Source: 'src/betree.rs', lines 318:4-318:22 *) -Definition betree_List_hd (T : Type) (self : betree_List_t T) : result T := - match self with - | Betree_List_Cons hd _ => Ok hd - | Betree_List_Nil => Fail_ Failure - end -. - -(** [betree_main::betree::{betree_main::betree::List<(u64, T)>#2}::head_has_key]: - Source: 'src/betree.rs', lines 327:4-327:44 *) -Definition betree_ListPairU64T_head_has_key - (T : Type) (self : betree_List_t (u64 * T)) (key : u64) : result bool := - match self with - | Betree_List_Cons hd _ => let (i, _) := hd in Ok (i s= key) - | Betree_List_Nil => Ok false - end -. - -(** [betree_main::betree::{betree_main::betree::List<(u64, T)>#2}::partition_at_pivot]: - Source: 'src/betree.rs', lines 339:4-339:73 *) -Fixpoint betree_ListPairU64T_partition_at_pivot - (T : Type) (n : nat) (self : betree_List_t (u64 * T)) (pivot : u64) : - result ((betree_List_t (u64 * T)) * (betree_List_t (u64 * T))) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match self with - | Betree_List_Cons hd tl => - let (i, t) := hd in - if i s>= pivot - then Ok (Betree_List_Nil, Betree_List_Cons (i, t) tl) - else ( - p <- betree_ListPairU64T_partition_at_pivot T n1 tl pivot; - let (ls0, ls1) := p in - Ok (Betree_List_Cons (i, t) ls0, ls1)) - | Betree_List_Nil => Ok (Betree_List_Nil, Betree_List_Nil) - end - end -. - -(** [betree_main::betree::{betree_main::betree::Leaf#3}::split]: - Source: 'src/betree.rs', lines 359:4-364:17 *) -Definition betree_Leaf_split - (n : nat) (self : betree_Leaf_t) (content : betree_List_t (u64 * u64)) - (params : betree_Params_t) (node_id_cnt : betree_NodeIdCounter_t) - (st : state) : - result (state * (betree_Internal_t * betree_NodeIdCounter_t)) - := - p <- - betree_List_split_at (u64 * u64) n content - params.(betree_Params_split_size); - let (content0, content1) := p in - p1 <- betree_List_hd (u64 * u64) content1; - let (pivot, _) := p1 in - p2 <- betree_NodeIdCounter_fresh_id node_id_cnt; - let (id0, node_id_cnt1) := p2 in - p3 <- betree_NodeIdCounter_fresh_id node_id_cnt1; - let (id1, node_id_cnt2) := p3 in - p4 <- betree_store_leaf_node id0 content0 st; - let (st1, _) := p4 in - p5 <- betree_store_leaf_node id1 content1 st1; - let (st2, _) := p5 in - let n1 := Betree_Node_Leaf - {| - betree_Leaf_id := id0; - betree_Leaf_size := params.(betree_Params_split_size) - |} in - let n2 := Betree_Node_Leaf - {| - betree_Leaf_id := id1; - betree_Leaf_size := params.(betree_Params_split_size) - |} in - Ok (st2, (mkbetree_Internal_t self.(betree_Leaf_id) pivot n1 n2, - node_id_cnt2)) -. - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup_first_message_for_key]: - Source: 'src/betree.rs', lines 789:4-792:34 *) -Fixpoint betree_Node_lookup_first_message_for_key - (n : nat) (key : u64) (msgs : betree_List_t (u64 * betree_Message_t)) : - result ((betree_List_t (u64 * betree_Message_t)) * (betree_List_t (u64 * - betree_Message_t) -> result (betree_List_t (u64 * betree_Message_t)))) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match msgs with - | Betree_List_Cons x next_msgs => - let (i, m) := x in - if i s>= key - then Ok (Betree_List_Cons (i, m) next_msgs, Ok) - else ( - p <- betree_Node_lookup_first_message_for_key n1 key next_msgs; - let (l, lookup_first_message_for_key_back) := p in - let back := - fun (ret : betree_List_t (u64 * betree_Message_t)) => - next_msgs1 <- lookup_first_message_for_key_back ret; - Ok (Betree_List_Cons (i, m) next_msgs1) in - Ok (l, back)) - | Betree_List_Nil => Ok (Betree_List_Nil, Ok) - end - end -. - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup_in_bindings]: - Source: 'src/betree.rs', lines 636:4-636:80 *) -Fixpoint betree_Node_lookup_in_bindings - (n : nat) (key : u64) (bindings : betree_List_t (u64 * u64)) : - result (option u64) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match bindings with - | Betree_List_Cons hd tl => - let (i, i1) := hd in - if i s= key - then Ok (Some i1) - else - if i s> key then Ok None else betree_Node_lookup_in_bindings n1 key tl - | Betree_List_Nil => Ok None - end - end -. - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_upserts]: - Source: 'src/betree.rs', lines 819:4-819:90 *) -Fixpoint betree_Node_apply_upserts - (n : nat) (msgs : betree_List_t (u64 * betree_Message_t)) (prev : option u64) - (key : u64) : - result (u64 * (betree_List_t (u64 * betree_Message_t))) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - b <- betree_ListPairU64T_head_has_key betree_Message_t msgs key; - if b - then ( - p <- betree_List_pop_front (u64 * betree_Message_t) msgs; - let (msg, msgs1) := p in - let (_, m) := msg in - match m with - | Betree_Message_Insert _ => Fail_ Failure - | Betree_Message_Delete => Fail_ Failure - | Betree_Message_Upsert s => - v <- betree_upsert_update prev s; - betree_Node_apply_upserts n1 msgs1 (Some v) key - end) - else ( - v <- core_option_Option_unwrap u64 prev; - msgs1 <- - betree_List_push_front (u64 * betree_Message_t) msgs (key, - Betree_Message_Insert v); - Ok (v, msgs1)) - end -. - -(** [betree_main::betree::{betree_main::betree::Internal#4}::lookup_in_children]: - Source: 'src/betree.rs', lines 395:4-395:63 *) -Fixpoint betree_Internal_lookup_in_children - (n : nat) (self : betree_Internal_t) (key : u64) (st : state) : - result (state * ((option u64) * betree_Internal_t)) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - if key s< self.(betree_Internal_pivot) - then ( - p <- betree_Node_lookup n1 self.(betree_Internal_left) key st; - let (st1, p1) := p in - let (o, n2) := p1 in - Ok (st1, (o, mkbetree_Internal_t self.(betree_Internal_id) - self.(betree_Internal_pivot) n2 self.(betree_Internal_right)))) - else ( - p <- betree_Node_lookup n1 self.(betree_Internal_right) key st; - let (st1, p1) := p in - let (o, n2) := p1 in - Ok (st1, (o, mkbetree_Internal_t self.(betree_Internal_id) - self.(betree_Internal_pivot) self.(betree_Internal_left) n2))) - end - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup]: - Source: 'src/betree.rs', lines 709:4-709:58 *) -with betree_Node_lookup - (n : nat) (self : betree_Node_t) (key : u64) (st : state) : - result (state * ((option u64) * betree_Node_t)) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match self with - | Betree_Node_Internal node => - p <- betree_load_internal_node node.(betree_Internal_id) st; - let (st1, msgs) := p in - p1 <- betree_Node_lookup_first_message_for_key n1 key msgs; - let (pending, lookup_first_message_for_key_back) := p1 in - match pending with - | Betree_List_Cons p2 l => - let (k, msg) := p2 in - if k s<> key - then ( - p3 <- betree_Internal_lookup_in_children n1 node key st1; - let (st2, p4) := p3 in - let (o, node1) := p4 in - _ <- lookup_first_message_for_key_back (Betree_List_Cons (k, msg) l); - Ok (st2, (o, Betree_Node_Internal node1))) - else - match msg with - | Betree_Message_Insert v => - _ <- - lookup_first_message_for_key_back (Betree_List_Cons (k, - Betree_Message_Insert v) l); - Ok (st1, (Some v, Betree_Node_Internal node)) - | Betree_Message_Delete => - _ <- - lookup_first_message_for_key_back (Betree_List_Cons (k, - Betree_Message_Delete) l); - Ok (st1, (None, Betree_Node_Internal node)) - | Betree_Message_Upsert ufs => - p3 <- betree_Internal_lookup_in_children n1 node key st1; - let (st2, p4) := p3 in - let (v, node1) := p4 in - p5 <- - betree_Node_apply_upserts n1 (Betree_List_Cons (k, - Betree_Message_Upsert ufs) l) v key; - let (v1, pending1) := p5 in - msgs1 <- lookup_first_message_for_key_back pending1; - p6 <- - betree_store_internal_node node1.(betree_Internal_id) msgs1 st2; - let (st3, _) := p6 in - Ok (st3, (Some v1, Betree_Node_Internal node1)) - end - | Betree_List_Nil => - p2 <- betree_Internal_lookup_in_children n1 node key st1; - let (st2, p3) := p2 in - let (o, node1) := p3 in - _ <- lookup_first_message_for_key_back Betree_List_Nil; - Ok (st2, (o, Betree_Node_Internal node1)) - end - | Betree_Node_Leaf node => - p <- betree_load_leaf_node node.(betree_Leaf_id) st; - let (st1, bindings) := p in - o <- betree_Node_lookup_in_bindings n1 key bindings; - Ok (st1, (o, Betree_Node_Leaf node)) - end - end -. - -(** [betree_main::betree::{betree_main::betree::Node#5}::filter_messages_for_key]: - Source: 'src/betree.rs', lines 674:4-674:77 *) -Fixpoint betree_Node_filter_messages_for_key - (n : nat) (key : u64) (msgs : betree_List_t (u64 * betree_Message_t)) : - result (betree_List_t (u64 * betree_Message_t)) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match msgs with - | Betree_List_Cons p l => - let (k, m) := p in - if k s= key - then ( - p1 <- - betree_List_pop_front (u64 * betree_Message_t) (Betree_List_Cons (k, - m) l); - let (_, msgs1) := p1 in - betree_Node_filter_messages_for_key n1 key msgs1) - else Ok (Betree_List_Cons (k, m) l) - | Betree_List_Nil => Ok Betree_List_Nil - end - end -. - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup_first_message_after_key]: - Source: 'src/betree.rs', lines 689:4-692:34 *) -Fixpoint betree_Node_lookup_first_message_after_key - (n : nat) (key : u64) (msgs : betree_List_t (u64 * betree_Message_t)) : - result ((betree_List_t (u64 * betree_Message_t)) * (betree_List_t (u64 * - betree_Message_t) -> result (betree_List_t (u64 * betree_Message_t)))) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match msgs with - | Betree_List_Cons p next_msgs => - let (k, m) := p in - if k s= key - then ( - p1 <- betree_Node_lookup_first_message_after_key n1 key next_msgs; - let (l, lookup_first_message_after_key_back) := p1 in - let back := - fun (ret : betree_List_t (u64 * betree_Message_t)) => - next_msgs1 <- lookup_first_message_after_key_back ret; - Ok (Betree_List_Cons (k, m) next_msgs1) in - Ok (l, back)) - else Ok (Betree_List_Cons (k, m) next_msgs, Ok) - | Betree_List_Nil => Ok (Betree_List_Nil, Ok) - end - end -. - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_to_internal]: - Source: 'src/betree.rs', lines 521:4-521:89 *) -Definition betree_Node_apply_to_internal - (n : nat) (msgs : betree_List_t (u64 * betree_Message_t)) (key : u64) - (new_msg : betree_Message_t) : - result (betree_List_t (u64 * betree_Message_t)) - := - p <- betree_Node_lookup_first_message_for_key n key msgs; - let (msgs1, lookup_first_message_for_key_back) := p in - b <- betree_ListPairU64T_head_has_key betree_Message_t msgs1 key; - if b - then - match new_msg with - | Betree_Message_Insert i => - msgs2 <- betree_Node_filter_messages_for_key n key msgs1; - msgs3 <- - betree_List_push_front (u64 * betree_Message_t) msgs2 (key, - Betree_Message_Insert i); - lookup_first_message_for_key_back msgs3 - | Betree_Message_Delete => - msgs2 <- betree_Node_filter_messages_for_key n key msgs1; - msgs3 <- - betree_List_push_front (u64 * betree_Message_t) msgs2 (key, - Betree_Message_Delete); - lookup_first_message_for_key_back msgs3 - | Betree_Message_Upsert s => - p1 <- betree_List_hd (u64 * betree_Message_t) msgs1; - let (_, m) := p1 in - match m with - | Betree_Message_Insert prev => - v <- betree_upsert_update (Some prev) s; - p2 <- betree_List_pop_front (u64 * betree_Message_t) msgs1; - let (_, msgs2) := p2 in - msgs3 <- - betree_List_push_front (u64 * betree_Message_t) msgs2 (key, - Betree_Message_Insert v); - lookup_first_message_for_key_back msgs3 - | Betree_Message_Delete => - p2 <- betree_List_pop_front (u64 * betree_Message_t) msgs1; - let (_, msgs2) := p2 in - v <- betree_upsert_update None s; - msgs3 <- - betree_List_push_front (u64 * betree_Message_t) msgs2 (key, - Betree_Message_Insert v); - lookup_first_message_for_key_back msgs3 - | Betree_Message_Upsert _ => - p2 <- betree_Node_lookup_first_message_after_key n key msgs1; - let (msgs2, lookup_first_message_after_key_back) := p2 in - msgs3 <- - betree_List_push_front (u64 * betree_Message_t) msgs2 (key, - Betree_Message_Upsert s); - msgs4 <- lookup_first_message_after_key_back msgs3; - lookup_first_message_for_key_back msgs4 - end - end - else ( - msgs2 <- - betree_List_push_front (u64 * betree_Message_t) msgs1 (key, new_msg); - lookup_first_message_for_key_back msgs2) -. - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_messages_to_internal]: - Source: 'src/betree.rs', lines 502:4-505:5 *) -Fixpoint betree_Node_apply_messages_to_internal - (n : nat) (msgs : betree_List_t (u64 * betree_Message_t)) - (new_msgs : betree_List_t (u64 * betree_Message_t)) : - result (betree_List_t (u64 * betree_Message_t)) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match new_msgs with - | Betree_List_Cons new_msg new_msgs_tl => - let (i, m) := new_msg in - msgs1 <- betree_Node_apply_to_internal n1 msgs i m; - betree_Node_apply_messages_to_internal n1 msgs1 new_msgs_tl - | Betree_List_Nil => Ok msgs - end - end -. - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup_mut_in_bindings]: - Source: 'src/betree.rs', lines 653:4-656:32 *) -Fixpoint betree_Node_lookup_mut_in_bindings - (n : nat) (key : u64) (bindings : betree_List_t (u64 * u64)) : - result ((betree_List_t (u64 * u64)) * (betree_List_t (u64 * u64) -> result - (betree_List_t (u64 * u64)))) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match bindings with - | Betree_List_Cons hd tl => - let (i, i1) := hd in - if i s>= key - then Ok (Betree_List_Cons (i, i1) tl, Ok) - else ( - p <- betree_Node_lookup_mut_in_bindings n1 key tl; - let (l, lookup_mut_in_bindings_back) := p in - let back := - fun (ret : betree_List_t (u64 * u64)) => - tl1 <- lookup_mut_in_bindings_back ret; - Ok (Betree_List_Cons (i, i1) tl1) in - Ok (l, back)) - | Betree_List_Nil => Ok (Betree_List_Nil, Ok) - end - end -. - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_to_leaf]: - Source: 'src/betree.rs', lines 460:4-460:87 *) -Definition betree_Node_apply_to_leaf - (n : nat) (bindings : betree_List_t (u64 * u64)) (key : u64) - (new_msg : betree_Message_t) : - result (betree_List_t (u64 * u64)) - := - p <- betree_Node_lookup_mut_in_bindings n key bindings; - let (bindings1, lookup_mut_in_bindings_back) := p in - b <- betree_ListPairU64T_head_has_key u64 bindings1 key; - if b - then ( - p1 <- betree_List_pop_front (u64 * u64) bindings1; - let (hd, bindings2) := p1 in - match new_msg with - | Betree_Message_Insert v => - bindings3 <- betree_List_push_front (u64 * u64) bindings2 (key, v); - lookup_mut_in_bindings_back bindings3 - | Betree_Message_Delete => lookup_mut_in_bindings_back bindings2 - | Betree_Message_Upsert s => - let (_, i) := hd in - v <- betree_upsert_update (Some i) s; - bindings3 <- betree_List_push_front (u64 * u64) bindings2 (key, v); - lookup_mut_in_bindings_back bindings3 - end) - else - match new_msg with - | Betree_Message_Insert v => - bindings2 <- betree_List_push_front (u64 * u64) bindings1 (key, v); - lookup_mut_in_bindings_back bindings2 - | Betree_Message_Delete => lookup_mut_in_bindings_back bindings1 - | Betree_Message_Upsert s => - v <- betree_upsert_update None s; - bindings2 <- betree_List_push_front (u64 * u64) bindings1 (key, v); - lookup_mut_in_bindings_back bindings2 - end -. - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_messages_to_leaf]: - Source: 'src/betree.rs', lines 444:4-447:5 *) -Fixpoint betree_Node_apply_messages_to_leaf - (n : nat) (bindings : betree_List_t (u64 * u64)) - (new_msgs : betree_List_t (u64 * betree_Message_t)) : - result (betree_List_t (u64 * u64)) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match new_msgs with - | Betree_List_Cons new_msg new_msgs_tl => - let (i, m) := new_msg in - bindings1 <- betree_Node_apply_to_leaf n1 bindings i m; - betree_Node_apply_messages_to_leaf n1 bindings1 new_msgs_tl - | Betree_List_Nil => Ok bindings - end - end -. - -(** [betree_main::betree::{betree_main::betree::Internal#4}::flush]: - Source: 'src/betree.rs', lines 410:4-415:26 *) -Fixpoint betree_Internal_flush - (n : nat) (self : betree_Internal_t) (params : betree_Params_t) - (node_id_cnt : betree_NodeIdCounter_t) - (content : betree_List_t (u64 * betree_Message_t)) (st : state) : - result (state * ((betree_List_t (u64 * betree_Message_t)) * - (betree_Internal_t * betree_NodeIdCounter_t))) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - p <- - betree_ListPairU64T_partition_at_pivot betree_Message_t n1 content - self.(betree_Internal_pivot); - let (msgs_left, msgs_right) := p in - len_left <- betree_List_len (u64 * betree_Message_t) n1 msgs_left; - if len_left s>= params.(betree_Params_min_flush_size) - then ( - p1 <- - betree_Node_apply_messages n1 self.(betree_Internal_left) params - node_id_cnt msgs_left st; - let (st1, p2) := p1 in - let (n2, node_id_cnt1) := p2 in - len_right <- betree_List_len (u64 * betree_Message_t) n1 msgs_right; - if len_right s>= params.(betree_Params_min_flush_size) - then ( - p3 <- - betree_Node_apply_messages n1 self.(betree_Internal_right) params - node_id_cnt1 msgs_right st1; - let (st2, p4) := p3 in - let (n3, node_id_cnt2) := p4 in - Ok (st2, (Betree_List_Nil, (mkbetree_Internal_t - self.(betree_Internal_id) self.(betree_Internal_pivot) n2 n3, - node_id_cnt2)))) - else - Ok (st1, (msgs_right, (mkbetree_Internal_t self.(betree_Internal_id) - self.(betree_Internal_pivot) n2 self.(betree_Internal_right), - node_id_cnt1)))) - else ( - p1 <- - betree_Node_apply_messages n1 self.(betree_Internal_right) params - node_id_cnt msgs_right st; - let (st1, p2) := p1 in - let (n2, node_id_cnt1) := p2 in - Ok (st1, (msgs_left, (mkbetree_Internal_t self.(betree_Internal_id) - self.(betree_Internal_pivot) self.(betree_Internal_left) n2, - node_id_cnt1)))) - end - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_messages]: - Source: 'src/betree.rs', lines 588:4-593:5 *) -with betree_Node_apply_messages - (n : nat) (self : betree_Node_t) (params : betree_Params_t) - (node_id_cnt : betree_NodeIdCounter_t) - (msgs : betree_List_t (u64 * betree_Message_t)) (st : state) : - result (state * (betree_Node_t * betree_NodeIdCounter_t)) - := - match n with - | O => Fail_ OutOfFuel - | S n1 => - match self with - | Betree_Node_Internal node => - p <- betree_load_internal_node node.(betree_Internal_id) st; - let (st1, content) := p in - content1 <- betree_Node_apply_messages_to_internal n1 content msgs; - num_msgs <- betree_List_len (u64 * betree_Message_t) n1 content1; - if num_msgs s>= params.(betree_Params_min_flush_size) - then ( - p1 <- betree_Internal_flush n1 node params node_id_cnt content1 st1; - let (st2, p2) := p1 in - let (content2, p3) := p2 in - let (node1, node_id_cnt1) := p3 in - p4 <- - betree_store_internal_node node1.(betree_Internal_id) content2 st2; - let (st3, _) := p4 in - Ok (st3, (Betree_Node_Internal node1, node_id_cnt1))) - else ( - p1 <- - betree_store_internal_node node.(betree_Internal_id) content1 st1; - let (st2, _) := p1 in - Ok (st2, (Betree_Node_Internal node, node_id_cnt))) - | Betree_Node_Leaf node => - p <- betree_load_leaf_node node.(betree_Leaf_id) st; - let (st1, content) := p in - content1 <- betree_Node_apply_messages_to_leaf n1 content msgs; - len <- betree_List_len (u64 * u64) n1 content1; - i <- u64_mul 2%u64 params.(betree_Params_split_size); - if len s>= i - then ( - p1 <- betree_Leaf_split n1 node content1 params node_id_cnt st1; - let (st2, p2) := p1 in - let (new_node, node_id_cnt1) := p2 in - p3 <- betree_store_leaf_node node.(betree_Leaf_id) Betree_List_Nil st2; - let (st3, _) := p3 in - Ok (st3, (Betree_Node_Internal new_node, node_id_cnt1))) - else ( - p1 <- betree_store_leaf_node node.(betree_Leaf_id) content1 st1; - let (st2, _) := p1 in - Ok (st2, (Betree_Node_Leaf - {| betree_Leaf_id := node.(betree_Leaf_id); betree_Leaf_size := len - |}, node_id_cnt))) - end - end -. - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply]: - Source: 'src/betree.rs', lines 576:4-582:5 *) -Definition betree_Node_apply - (n : nat) (self : betree_Node_t) (params : betree_Params_t) - (node_id_cnt : betree_NodeIdCounter_t) (key : u64) - (new_msg : betree_Message_t) (st : state) : - result (state * (betree_Node_t * betree_NodeIdCounter_t)) - := - p <- - betree_Node_apply_messages n self params node_id_cnt (Betree_List_Cons - (key, new_msg) Betree_List_Nil) st; - let (st1, p1) := p in - let (self1, node_id_cnt1) := p1 in - Ok (st1, (self1, node_id_cnt1)) -. - -(** [betree_main::betree::{betree_main::betree::BeTree#6}::new]: - Source: 'src/betree.rs', lines 849:4-849:60 *) -Definition betree_BeTree_new - (min_flush_size : u64) (split_size : u64) (st : state) : - result (state * betree_BeTree_t) - := - node_id_cnt <- betree_NodeIdCounter_new; - p <- betree_NodeIdCounter_fresh_id node_id_cnt; - let (id, node_id_cnt1) := p in - p1 <- betree_store_leaf_node id Betree_List_Nil st; - let (st1, _) := p1 in - Ok (st1, - {| - betree_BeTree_params := - {| - betree_Params_min_flush_size := min_flush_size; - betree_Params_split_size := split_size - |}; - betree_BeTree_node_id_cnt := node_id_cnt1; - betree_BeTree_root := - (Betree_Node_Leaf - {| betree_Leaf_id := id; betree_Leaf_size := 0%u64 |}) - |}) -. - -(** [betree_main::betree::{betree_main::betree::BeTree#6}::apply]: - Source: 'src/betree.rs', lines 868:4-868:47 *) -Definition betree_BeTree_apply - (n : nat) (self : betree_BeTree_t) (key : u64) (msg : betree_Message_t) - (st : state) : - result (state * betree_BeTree_t) - := - p <- - betree_Node_apply n self.(betree_BeTree_root) self.(betree_BeTree_params) - self.(betree_BeTree_node_id_cnt) key msg st; - let (st1, p1) := p in - let (n1, nic) := p1 in - Ok (st1, - {| - betree_BeTree_params := self.(betree_BeTree_params); - betree_BeTree_node_id_cnt := nic; - betree_BeTree_root := n1 - |}) -. - -(** [betree_main::betree::{betree_main::betree::BeTree#6}::insert]: - Source: 'src/betree.rs', lines 874:4-874:52 *) -Definition betree_BeTree_insert - (n : nat) (self : betree_BeTree_t) (key : u64) (value : u64) (st : state) : - result (state * betree_BeTree_t) - := - betree_BeTree_apply n self key (Betree_Message_Insert value) st -. - -(** [betree_main::betree::{betree_main::betree::BeTree#6}::delete]: - Source: 'src/betree.rs', lines 880:4-880:38 *) -Definition betree_BeTree_delete - (n : nat) (self : betree_BeTree_t) (key : u64) (st : state) : - result (state * betree_BeTree_t) - := - betree_BeTree_apply n self key Betree_Message_Delete st -. - -(** [betree_main::betree::{betree_main::betree::BeTree#6}::upsert]: - Source: 'src/betree.rs', lines 886:4-886:59 *) -Definition betree_BeTree_upsert - (n : nat) (self : betree_BeTree_t) (key : u64) - (upd : betree_UpsertFunState_t) (st : state) : - result (state * betree_BeTree_t) - := - betree_BeTree_apply n self key (Betree_Message_Upsert upd) st -. - -(** [betree_main::betree::{betree_main::betree::BeTree#6}::lookup]: - Source: 'src/betree.rs', lines 895:4-895:62 *) -Definition betree_BeTree_lookup - (n : nat) (self : betree_BeTree_t) (key : u64) (st : state) : - result (state * ((option u64) * betree_BeTree_t)) - := - p <- betree_Node_lookup n self.(betree_BeTree_root) key st; - let (st1, p1) := p in - let (o, n1) := p1 in - Ok (st1, (o, - {| - betree_BeTree_params := self.(betree_BeTree_params); - betree_BeTree_node_id_cnt := self.(betree_BeTree_node_id_cnt); - betree_BeTree_root := n1 - |})) -. - -(** [betree_main::main]: - Source: 'src/main.rs', lines 4:0-4:9 *) -Definition main : result unit := - Ok tt. - -(** Unit test for [betree_main::main] *) -Check (main )%return. - -End BetreeMain_Funs. diff --git a/tests/coq/betree/BetreeMain_FunsExternal.v b/tests/coq/betree/BetreeMain_FunsExternal.v deleted file mode 100644 index 099a2e8f..00000000 --- a/tests/coq/betree/BetreeMain_FunsExternal.v +++ /dev/null @@ -1,39 +0,0 @@ -(** [betree_main]: external functions. --- This is a template file: rename it to "FunsExternal.lean" and fill the holes. *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Require Export BetreeMain_Types. -Import BetreeMain_Types. -Module BetreeMain_FunsExternal. - -(** [betree_main::betree_utils::load_internal_node]: forward function - Source: 'src/betree_utils.rs', lines 98:0-98:63 *) -Axiom betree_utils_load_internal_node - : u64 -> state -> result (state * (betree_List_t (u64 * betree_Message_t))) -. - -(** [betree_main::betree_utils::store_internal_node]: forward function - Source: 'src/betree_utils.rs', lines 115:0-115:71 *) -Axiom betree_utils_store_internal_node - : - u64 -> betree_List_t (u64 * betree_Message_t) -> state -> result (state * - unit) -. - -(** [betree_main::betree_utils::load_leaf_node]: forward function - Source: 'src/betree_utils.rs', lines 132:0-132:55 *) -Axiom betree_utils_load_leaf_node - : u64 -> state -> result (state * (betree_List_t (u64 * u64))) -. - -(** [betree_main::betree_utils::store_leaf_node]: forward function - Source: 'src/betree_utils.rs', lines 145:0-145:63 *) -Axiom betree_utils_store_leaf_node - : u64 -> betree_List_t (u64 * u64) -> state -> result (state * unit) -. - -End BetreeMain_FunsExternal. diff --git a/tests/coq/betree/BetreeMain_FunsExternal_Template.v b/tests/coq/betree/BetreeMain_FunsExternal_Template.v deleted file mode 100644 index 58be2733..00000000 --- a/tests/coq/betree/BetreeMain_FunsExternal_Template.v +++ /dev/null @@ -1,40 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [betree_main]: external functions. --- This is a template file: rename it to "FunsExternal.lean" and fill the holes. *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Require Import BetreeMain_Types. -Include BetreeMain_Types. -Module BetreeMain_FunsExternal_Template. - -(** [betree_main::betree_utils::load_internal_node]: - Source: 'src/betree_utils.rs', lines 98:0-98:63 *) -Axiom betree_utils_load_internal_node - : u64 -> state -> result (state * (betree_List_t (u64 * betree_Message_t))) -. - -(** [betree_main::betree_utils::store_internal_node]: - Source: 'src/betree_utils.rs', lines 115:0-115:71 *) -Axiom betree_utils_store_internal_node - : - u64 -> betree_List_t (u64 * betree_Message_t) -> state -> result (state * - unit) -. - -(** [betree_main::betree_utils::load_leaf_node]: - Source: 'src/betree_utils.rs', lines 132:0-132:55 *) -Axiom betree_utils_load_leaf_node - : u64 -> state -> result (state * (betree_List_t (u64 * u64))) -. - -(** [betree_main::betree_utils::store_leaf_node]: - Source: 'src/betree_utils.rs', lines 145:0-145:63 *) -Axiom betree_utils_store_leaf_node - : u64 -> betree_List_t (u64 * u64) -> state -> result (state * unit) -. - -End BetreeMain_FunsExternal_Template. diff --git a/tests/coq/betree/BetreeMain_Types.v b/tests/coq/betree/BetreeMain_Types.v deleted file mode 100644 index acbc2085..00000000 --- a/tests/coq/betree/BetreeMain_Types.v +++ /dev/null @@ -1,118 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [betree_main]: type definitions *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Require Import BetreeMain_TypesExternal. -Include BetreeMain_TypesExternal. -Module BetreeMain_Types. - -(** [betree_main::betree::List] - Source: 'src/betree.rs', lines 17:0-17:23 *) -Inductive betree_List_t (T : Type) := -| Betree_List_Cons : T -> betree_List_t T -> betree_List_t T -| Betree_List_Nil : betree_List_t T -. - -Arguments Betree_List_Cons { _ }. -Arguments Betree_List_Nil { _ }. - -(** [betree_main::betree::UpsertFunState] - Source: 'src/betree.rs', lines 63:0-63:23 *) -Inductive betree_UpsertFunState_t := -| Betree_UpsertFunState_Add : u64 -> betree_UpsertFunState_t -| Betree_UpsertFunState_Sub : u64 -> betree_UpsertFunState_t -. - -(** [betree_main::betree::Message] - Source: 'src/betree.rs', lines 69:0-69:23 *) -Inductive betree_Message_t := -| Betree_Message_Insert : u64 -> betree_Message_t -| Betree_Message_Delete : betree_Message_t -| Betree_Message_Upsert : betree_UpsertFunState_t -> betree_Message_t -. - -(** [betree_main::betree::Leaf] - Source: 'src/betree.rs', lines 167:0-167:11 *) -Record betree_Leaf_t := -mkbetree_Leaf_t { - betree_Leaf_id : u64; betree_Leaf_size : u64; -} -. - -(** [betree_main::betree::Internal] - Source: 'src/betree.rs', lines 156:0-156:15 *) -Inductive betree_Internal_t := -| mkbetree_Internal_t : - u64 -> - u64 -> - betree_Node_t -> - betree_Node_t -> - betree_Internal_t - -(** [betree_main::betree::Node] - Source: 'src/betree.rs', lines 179:0-179:9 *) -with betree_Node_t := -| Betree_Node_Internal : betree_Internal_t -> betree_Node_t -| Betree_Node_Leaf : betree_Leaf_t -> betree_Node_t -. - -Definition betree_Internal_id (x : betree_Internal_t) := - match x with | mkbetree_Internal_t x1 _ _ _ => x1 end -. - -Notation "x2 .(betree_Internal_id)" := (betree_Internal_id x2) (at level 9). - -Definition betree_Internal_pivot (x : betree_Internal_t) := - match x with | mkbetree_Internal_t _ x1 _ _ => x1 end -. - -Notation "x2 .(betree_Internal_pivot)" := (betree_Internal_pivot x2) - (at level 9) -. - -Definition betree_Internal_left (x : betree_Internal_t) := - match x with | mkbetree_Internal_t _ _ x1 _ => x1 end -. - -Notation "x2 .(betree_Internal_left)" := (betree_Internal_left x2) (at level 9) -. - -Definition betree_Internal_right (x : betree_Internal_t) := - match x with | mkbetree_Internal_t _ _ _ x1 => x1 end -. - -Notation "x2 .(betree_Internal_right)" := (betree_Internal_right x2) - (at level 9) -. - -(** [betree_main::betree::Params] - Source: 'src/betree.rs', lines 187:0-187:13 *) -Record betree_Params_t := -mkbetree_Params_t { - betree_Params_min_flush_size : u64; betree_Params_split_size : u64; -} -. - -(** [betree_main::betree::NodeIdCounter] - Source: 'src/betree.rs', lines 201:0-201:20 *) -Record betree_NodeIdCounter_t := -mkbetree_NodeIdCounter_t { - betree_NodeIdCounter_next_node_id : u64; -} -. - -(** [betree_main::betree::BeTree] - Source: 'src/betree.rs', lines 218:0-218:17 *) -Record betree_BeTree_t := -mkbetree_BeTree_t { - betree_BeTree_params : betree_Params_t; - betree_BeTree_node_id_cnt : betree_NodeIdCounter_t; - betree_BeTree_root : betree_Node_t; -} -. - -End BetreeMain_Types. diff --git a/tests/coq/betree/BetreeMain_TypesExternal.v b/tests/coq/betree/BetreeMain_TypesExternal.v deleted file mode 100644 index 870d2601..00000000 --- a/tests/coq/betree/BetreeMain_TypesExternal.v +++ /dev/null @@ -1,14 +0,0 @@ -(** [betree_main]: external types. --- This is a template file: rename it to "TypesExternal.lean" and fill the holes. *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Module BetreeMain_TypesExternal. - -(** The state type used in the state-error monad *) -Axiom state : Type. - -End BetreeMain_TypesExternal. diff --git a/tests/coq/betree/BetreeMain_TypesExternal_Template.v b/tests/coq/betree/BetreeMain_TypesExternal_Template.v deleted file mode 100644 index 651de2b7..00000000 --- a/tests/coq/betree/BetreeMain_TypesExternal_Template.v +++ /dev/null @@ -1,15 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [betree_main]: external types. --- This is a template file: rename it to "TypesExternal.lean" and fill the holes. *) -Require Import Primitives. -Import Primitives. -Require Import Coq.ZArith.ZArith. -Require Import List. -Import ListNotations. -Local Open Scope Primitives_scope. -Module BetreeMain_TypesExternal_Template. - -(** The state type used in the state-error monad *) -Axiom state : Type. - -End BetreeMain_TypesExternal_Template. diff --git a/tests/coq/betree/Betree_Funs.v b/tests/coq/betree/Betree_Funs.v new file mode 100644 index 00000000..b965d423 --- /dev/null +++ b/tests/coq/betree/Betree_Funs.v @@ -0,0 +1,837 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [betree]: function definitions *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Require Import Betree_Types. +Include Betree_Types. +Require Import Betree_FunsExternal. +Include Betree_FunsExternal. +Module Betree_Funs. + +(** [betree::betree::load_internal_node]: + Source: 'src/betree.rs', lines 36:0-36:52 *) +Definition betree_load_internal_node + (id : u64) (st : state) : + result (state * (betree_List_t (u64 * betree_Message_t))) + := + betree_utils_load_internal_node id st +. + +(** [betree::betree::store_internal_node]: + Source: 'src/betree.rs', lines 41:0-41:60 *) +Definition betree_store_internal_node + (id : u64) (content : betree_List_t (u64 * betree_Message_t)) (st : state) : + result (state * unit) + := + betree_utils_store_internal_node id content st +. + +(** [betree::betree::load_leaf_node]: + Source: 'src/betree.rs', lines 46:0-46:44 *) +Definition betree_load_leaf_node + (id : u64) (st : state) : result (state * (betree_List_t (u64 * u64))) := + betree_utils_load_leaf_node id st +. + +(** [betree::betree::store_leaf_node]: + Source: 'src/betree.rs', lines 51:0-51:52 *) +Definition betree_store_leaf_node + (id : u64) (content : betree_List_t (u64 * u64)) (st : state) : + result (state * unit) + := + betree_utils_store_leaf_node id content st +. + +(** [betree::betree::fresh_node_id]: + Source: 'src/betree.rs', lines 55:0-55:48 *) +Definition betree_fresh_node_id (counter : u64) : result (u64 * u64) := + counter1 <- u64_add counter 1%u64; Ok (counter, counter1) +. + +(** [betree::betree::{betree::betree::NodeIdCounter}::new]: + Source: 'src/betree.rs', lines 206:4-206:20 *) +Definition betree_NodeIdCounter_new : result betree_NodeIdCounter_t := + Ok {| betree_NodeIdCounter_next_node_id := 0%u64 |} +. + +(** [betree::betree::{betree::betree::NodeIdCounter}::fresh_id]: + Source: 'src/betree.rs', lines 210:4-210:36 *) +Definition betree_NodeIdCounter_fresh_id + (self : betree_NodeIdCounter_t) : result (u64 * betree_NodeIdCounter_t) := + i <- u64_add self.(betree_NodeIdCounter_next_node_id) 1%u64; + Ok (self.(betree_NodeIdCounter_next_node_id), + {| betree_NodeIdCounter_next_node_id := i |}) +. + +(** [betree::betree::upsert_update]: + Source: 'src/betree.rs', lines 234:0-234:70 *) +Definition betree_upsert_update + (prev : option u64) (st : betree_UpsertFunState_t) : result u64 := + match prev with + | None => + match st with + | Betree_UpsertFunState_Add v => Ok v + | Betree_UpsertFunState_Sub _ => Ok 0%u64 + end + | Some prev1 => + match st with + | Betree_UpsertFunState_Add v => + margin <- u64_sub core_u64_max prev1; + if margin s>= v then u64_add prev1 v else Ok core_u64_max + | Betree_UpsertFunState_Sub v => + if prev1 s>= v then u64_sub prev1 v else Ok 0%u64 + end + end +. + +(** [betree::betree::{betree::betree::List#1}::len]: + Source: 'src/betree.rs', lines 276:4-276:24 *) +Fixpoint betree_List_len + (T : Type) (n : nat) (self : betree_List_t T) : result u64 := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match self with + | Betree_List_Cons _ tl => i <- betree_List_len T n1 tl; u64_add 1%u64 i + | Betree_List_Nil => Ok 0%u64 + end + end +. + +(** [betree::betree::{betree::betree::List#1}::split_at]: + Source: 'src/betree.rs', lines 284:4-284:51 *) +Fixpoint betree_List_split_at + (T : Type) (n : nat) (self : betree_List_t T) (n1 : u64) : + result ((betree_List_t T) * (betree_List_t T)) + := + match n with + | O => Fail_ OutOfFuel + | S n2 => + if n1 s= 0%u64 + then Ok (Betree_List_Nil, self) + else + match self with + | Betree_List_Cons hd tl => + i <- u64_sub n1 1%u64; + p <- betree_List_split_at T n2 tl i; + let (ls0, ls1) := p in + Ok (Betree_List_Cons hd ls0, ls1) + | Betree_List_Nil => Fail_ Failure + end + end +. + +(** [betree::betree::{betree::betree::List#1}::push_front]: + Source: 'src/betree.rs', lines 299:4-299:34 *) +Definition betree_List_push_front + (T : Type) (self : betree_List_t T) (x : T) : result (betree_List_t T) := + let (tl, _) := core_mem_replace (betree_List_t T) self Betree_List_Nil in + Ok (Betree_List_Cons x tl) +. + +(** [betree::betree::{betree::betree::List#1}::pop_front]: + Source: 'src/betree.rs', lines 306:4-306:32 *) +Definition betree_List_pop_front + (T : Type) (self : betree_List_t T) : result (T * (betree_List_t T)) := + let (ls, _) := core_mem_replace (betree_List_t T) self Betree_List_Nil in + match ls with + | Betree_List_Cons x tl => Ok (x, tl) + | Betree_List_Nil => Fail_ Failure + end +. + +(** [betree::betree::{betree::betree::List#1}::hd]: + Source: 'src/betree.rs', lines 318:4-318:22 *) +Definition betree_List_hd (T : Type) (self : betree_List_t T) : result T := + match self with + | Betree_List_Cons hd _ => Ok hd + | Betree_List_Nil => Fail_ Failure + end +. + +(** [betree::betree::{betree::betree::List<(u64, T)>#2}::head_has_key]: + Source: 'src/betree.rs', lines 327:4-327:44 *) +Definition betree_ListPairU64T_head_has_key + (T : Type) (self : betree_List_t (u64 * T)) (key : u64) : result bool := + match self with + | Betree_List_Cons hd _ => let (i, _) := hd in Ok (i s= key) + | Betree_List_Nil => Ok false + end +. + +(** [betree::betree::{betree::betree::List<(u64, T)>#2}::partition_at_pivot]: + Source: 'src/betree.rs', lines 339:4-339:73 *) +Fixpoint betree_ListPairU64T_partition_at_pivot + (T : Type) (n : nat) (self : betree_List_t (u64 * T)) (pivot : u64) : + result ((betree_List_t (u64 * T)) * (betree_List_t (u64 * T))) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match self with + | Betree_List_Cons hd tl => + let (i, t) := hd in + if i s>= pivot + then Ok (Betree_List_Nil, Betree_List_Cons (i, t) tl) + else ( + p <- betree_ListPairU64T_partition_at_pivot T n1 tl pivot; + let (ls0, ls1) := p in + Ok (Betree_List_Cons (i, t) ls0, ls1)) + | Betree_List_Nil => Ok (Betree_List_Nil, Betree_List_Nil) + end + end +. + +(** [betree::betree::{betree::betree::Leaf#3}::split]: + Source: 'src/betree.rs', lines 359:4-364:17 *) +Definition betree_Leaf_split + (n : nat) (self : betree_Leaf_t) (content : betree_List_t (u64 * u64)) + (params : betree_Params_t) (node_id_cnt : betree_NodeIdCounter_t) + (st : state) : + result (state * (betree_Internal_t * betree_NodeIdCounter_t)) + := + p <- + betree_List_split_at (u64 * u64) n content + params.(betree_Params_split_size); + let (content0, content1) := p in + p1 <- betree_List_hd (u64 * u64) content1; + let (pivot, _) := p1 in + p2 <- betree_NodeIdCounter_fresh_id node_id_cnt; + let (id0, node_id_cnt1) := p2 in + p3 <- betree_NodeIdCounter_fresh_id node_id_cnt1; + let (id1, node_id_cnt2) := p3 in + p4 <- betree_store_leaf_node id0 content0 st; + let (st1, _) := p4 in + p5 <- betree_store_leaf_node id1 content1 st1; + let (st2, _) := p5 in + let n1 := Betree_Node_Leaf + {| + betree_Leaf_id := id0; + betree_Leaf_size := params.(betree_Params_split_size) + |} in + let n2 := Betree_Node_Leaf + {| + betree_Leaf_id := id1; + betree_Leaf_size := params.(betree_Params_split_size) + |} in + Ok (st2, (mkbetree_Internal_t self.(betree_Leaf_id) pivot n1 n2, + node_id_cnt2)) +. + +(** [betree::betree::{betree::betree::Node#5}::lookup_first_message_for_key]: + Source: 'src/betree.rs', lines 789:4-792:34 *) +Fixpoint betree_Node_lookup_first_message_for_key + (n : nat) (key : u64) (msgs : betree_List_t (u64 * betree_Message_t)) : + result ((betree_List_t (u64 * betree_Message_t)) * (betree_List_t (u64 * + betree_Message_t) -> result (betree_List_t (u64 * betree_Message_t)))) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match msgs with + | Betree_List_Cons x next_msgs => + let (i, m) := x in + if i s>= key + then Ok (Betree_List_Cons (i, m) next_msgs, Ok) + else ( + p <- betree_Node_lookup_first_message_for_key n1 key next_msgs; + let (l, lookup_first_message_for_key_back) := p in + let back := + fun (ret : betree_List_t (u64 * betree_Message_t)) => + next_msgs1 <- lookup_first_message_for_key_back ret; + Ok (Betree_List_Cons (i, m) next_msgs1) in + Ok (l, back)) + | Betree_List_Nil => Ok (Betree_List_Nil, Ok) + end + end +. + +(** [betree::betree::{betree::betree::Node#5}::lookup_in_bindings]: + Source: 'src/betree.rs', lines 636:4-636:80 *) +Fixpoint betree_Node_lookup_in_bindings + (n : nat) (key : u64) (bindings : betree_List_t (u64 * u64)) : + result (option u64) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match bindings with + | Betree_List_Cons hd tl => + let (i, i1) := hd in + if i s= key + then Ok (Some i1) + else + if i s> key then Ok None else betree_Node_lookup_in_bindings n1 key tl + | Betree_List_Nil => Ok None + end + end +. + +(** [betree::betree::{betree::betree::Node#5}::apply_upserts]: + Source: 'src/betree.rs', lines 819:4-819:90 *) +Fixpoint betree_Node_apply_upserts + (n : nat) (msgs : betree_List_t (u64 * betree_Message_t)) (prev : option u64) + (key : u64) : + result (u64 * (betree_List_t (u64 * betree_Message_t))) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + b <- betree_ListPairU64T_head_has_key betree_Message_t msgs key; + if b + then ( + p <- betree_List_pop_front (u64 * betree_Message_t) msgs; + let (msg, msgs1) := p in + let (_, m) := msg in + match m with + | Betree_Message_Insert _ => Fail_ Failure + | Betree_Message_Delete => Fail_ Failure + | Betree_Message_Upsert s => + v <- betree_upsert_update prev s; + betree_Node_apply_upserts n1 msgs1 (Some v) key + end) + else ( + v <- core_option_Option_unwrap u64 prev; + msgs1 <- + betree_List_push_front (u64 * betree_Message_t) msgs (key, + Betree_Message_Insert v); + Ok (v, msgs1)) + end +. + +(** [betree::betree::{betree::betree::Internal#4}::lookup_in_children]: + Source: 'src/betree.rs', lines 395:4-395:63 *) +Fixpoint betree_Internal_lookup_in_children + (n : nat) (self : betree_Internal_t) (key : u64) (st : state) : + result (state * ((option u64) * betree_Internal_t)) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + if key s< self.(betree_Internal_pivot) + then ( + p <- betree_Node_lookup n1 self.(betree_Internal_left) key st; + let (st1, p1) := p in + let (o, n2) := p1 in + Ok (st1, (o, mkbetree_Internal_t self.(betree_Internal_id) + self.(betree_Internal_pivot) n2 self.(betree_Internal_right)))) + else ( + p <- betree_Node_lookup n1 self.(betree_Internal_right) key st; + let (st1, p1) := p in + let (o, n2) := p1 in + Ok (st1, (o, mkbetree_Internal_t self.(betree_Internal_id) + self.(betree_Internal_pivot) self.(betree_Internal_left) n2))) + end + +(** [betree::betree::{betree::betree::Node#5}::lookup]: + Source: 'src/betree.rs', lines 709:4-709:58 *) +with betree_Node_lookup + (n : nat) (self : betree_Node_t) (key : u64) (st : state) : + result (state * ((option u64) * betree_Node_t)) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match self with + | Betree_Node_Internal node => + p <- betree_load_internal_node node.(betree_Internal_id) st; + let (st1, msgs) := p in + p1 <- betree_Node_lookup_first_message_for_key n1 key msgs; + let (pending, lookup_first_message_for_key_back) := p1 in + match pending with + | Betree_List_Cons p2 l => + let (k, msg) := p2 in + if k s<> key + then ( + p3 <- betree_Internal_lookup_in_children n1 node key st1; + let (st2, p4) := p3 in + let (o, node1) := p4 in + _ <- lookup_first_message_for_key_back (Betree_List_Cons (k, msg) l); + Ok (st2, (o, Betree_Node_Internal node1))) + else + match msg with + | Betree_Message_Insert v => + _ <- + lookup_first_message_for_key_back (Betree_List_Cons (k, + Betree_Message_Insert v) l); + Ok (st1, (Some v, Betree_Node_Internal node)) + | Betree_Message_Delete => + _ <- + lookup_first_message_for_key_back (Betree_List_Cons (k, + Betree_Message_Delete) l); + Ok (st1, (None, Betree_Node_Internal node)) + | Betree_Message_Upsert ufs => + p3 <- betree_Internal_lookup_in_children n1 node key st1; + let (st2, p4) := p3 in + let (v, node1) := p4 in + p5 <- + betree_Node_apply_upserts n1 (Betree_List_Cons (k, + Betree_Message_Upsert ufs) l) v key; + let (v1, pending1) := p5 in + msgs1 <- lookup_first_message_for_key_back pending1; + p6 <- + betree_store_internal_node node1.(betree_Internal_id) msgs1 st2; + let (st3, _) := p6 in + Ok (st3, (Some v1, Betree_Node_Internal node1)) + end + | Betree_List_Nil => + p2 <- betree_Internal_lookup_in_children n1 node key st1; + let (st2, p3) := p2 in + let (o, node1) := p3 in + _ <- lookup_first_message_for_key_back Betree_List_Nil; + Ok (st2, (o, Betree_Node_Internal node1)) + end + | Betree_Node_Leaf node => + p <- betree_load_leaf_node node.(betree_Leaf_id) st; + let (st1, bindings) := p in + o <- betree_Node_lookup_in_bindings n1 key bindings; + Ok (st1, (o, Betree_Node_Leaf node)) + end + end +. + +(** [betree::betree::{betree::betree::Node#5}::filter_messages_for_key]: + Source: 'src/betree.rs', lines 674:4-674:77 *) +Fixpoint betree_Node_filter_messages_for_key + (n : nat) (key : u64) (msgs : betree_List_t (u64 * betree_Message_t)) : + result (betree_List_t (u64 * betree_Message_t)) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match msgs with + | Betree_List_Cons p l => + let (k, m) := p in + if k s= key + then ( + p1 <- + betree_List_pop_front (u64 * betree_Message_t) (Betree_List_Cons (k, + m) l); + let (_, msgs1) := p1 in + betree_Node_filter_messages_for_key n1 key msgs1) + else Ok (Betree_List_Cons (k, m) l) + | Betree_List_Nil => Ok Betree_List_Nil + end + end +. + +(** [betree::betree::{betree::betree::Node#5}::lookup_first_message_after_key]: + Source: 'src/betree.rs', lines 689:4-692:34 *) +Fixpoint betree_Node_lookup_first_message_after_key + (n : nat) (key : u64) (msgs : betree_List_t (u64 * betree_Message_t)) : + result ((betree_List_t (u64 * betree_Message_t)) * (betree_List_t (u64 * + betree_Message_t) -> result (betree_List_t (u64 * betree_Message_t)))) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match msgs with + | Betree_List_Cons p next_msgs => + let (k, m) := p in + if k s= key + then ( + p1 <- betree_Node_lookup_first_message_after_key n1 key next_msgs; + let (l, lookup_first_message_after_key_back) := p1 in + let back := + fun (ret : betree_List_t (u64 * betree_Message_t)) => + next_msgs1 <- lookup_first_message_after_key_back ret; + Ok (Betree_List_Cons (k, m) next_msgs1) in + Ok (l, back)) + else Ok (Betree_List_Cons (k, m) next_msgs, Ok) + | Betree_List_Nil => Ok (Betree_List_Nil, Ok) + end + end +. + +(** [betree::betree::{betree::betree::Node#5}::apply_to_internal]: + Source: 'src/betree.rs', lines 521:4-521:89 *) +Definition betree_Node_apply_to_internal + (n : nat) (msgs : betree_List_t (u64 * betree_Message_t)) (key : u64) + (new_msg : betree_Message_t) : + result (betree_List_t (u64 * betree_Message_t)) + := + p <- betree_Node_lookup_first_message_for_key n key msgs; + let (msgs1, lookup_first_message_for_key_back) := p in + b <- betree_ListPairU64T_head_has_key betree_Message_t msgs1 key; + if b + then + match new_msg with + | Betree_Message_Insert i => + msgs2 <- betree_Node_filter_messages_for_key n key msgs1; + msgs3 <- + betree_List_push_front (u64 * betree_Message_t) msgs2 (key, + Betree_Message_Insert i); + lookup_first_message_for_key_back msgs3 + | Betree_Message_Delete => + msgs2 <- betree_Node_filter_messages_for_key n key msgs1; + msgs3 <- + betree_List_push_front (u64 * betree_Message_t) msgs2 (key, + Betree_Message_Delete); + lookup_first_message_for_key_back msgs3 + | Betree_Message_Upsert s => + p1 <- betree_List_hd (u64 * betree_Message_t) msgs1; + let (_, m) := p1 in + match m with + | Betree_Message_Insert prev => + v <- betree_upsert_update (Some prev) s; + p2 <- betree_List_pop_front (u64 * betree_Message_t) msgs1; + let (_, msgs2) := p2 in + msgs3 <- + betree_List_push_front (u64 * betree_Message_t) msgs2 (key, + Betree_Message_Insert v); + lookup_first_message_for_key_back msgs3 + | Betree_Message_Delete => + p2 <- betree_List_pop_front (u64 * betree_Message_t) msgs1; + let (_, msgs2) := p2 in + v <- betree_upsert_update None s; + msgs3 <- + betree_List_push_front (u64 * betree_Message_t) msgs2 (key, + Betree_Message_Insert v); + lookup_first_message_for_key_back msgs3 + | Betree_Message_Upsert _ => + p2 <- betree_Node_lookup_first_message_after_key n key msgs1; + let (msgs2, lookup_first_message_after_key_back) := p2 in + msgs3 <- + betree_List_push_front (u64 * betree_Message_t) msgs2 (key, + Betree_Message_Upsert s); + msgs4 <- lookup_first_message_after_key_back msgs3; + lookup_first_message_for_key_back msgs4 + end + end + else ( + msgs2 <- + betree_List_push_front (u64 * betree_Message_t) msgs1 (key, new_msg); + lookup_first_message_for_key_back msgs2) +. + +(** [betree::betree::{betree::betree::Node#5}::apply_messages_to_internal]: + Source: 'src/betree.rs', lines 502:4-505:5 *) +Fixpoint betree_Node_apply_messages_to_internal + (n : nat) (msgs : betree_List_t (u64 * betree_Message_t)) + (new_msgs : betree_List_t (u64 * betree_Message_t)) : + result (betree_List_t (u64 * betree_Message_t)) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match new_msgs with + | Betree_List_Cons new_msg new_msgs_tl => + let (i, m) := new_msg in + msgs1 <- betree_Node_apply_to_internal n1 msgs i m; + betree_Node_apply_messages_to_internal n1 msgs1 new_msgs_tl + | Betree_List_Nil => Ok msgs + end + end +. + +(** [betree::betree::{betree::betree::Node#5}::lookup_mut_in_bindings]: + Source: 'src/betree.rs', lines 653:4-656:32 *) +Fixpoint betree_Node_lookup_mut_in_bindings + (n : nat) (key : u64) (bindings : betree_List_t (u64 * u64)) : + result ((betree_List_t (u64 * u64)) * (betree_List_t (u64 * u64) -> result + (betree_List_t (u64 * u64)))) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match bindings with + | Betree_List_Cons hd tl => + let (i, i1) := hd in + if i s>= key + then Ok (Betree_List_Cons (i, i1) tl, Ok) + else ( + p <- betree_Node_lookup_mut_in_bindings n1 key tl; + let (l, lookup_mut_in_bindings_back) := p in + let back := + fun (ret : betree_List_t (u64 * u64)) => + tl1 <- lookup_mut_in_bindings_back ret; + Ok (Betree_List_Cons (i, i1) tl1) in + Ok (l, back)) + | Betree_List_Nil => Ok (Betree_List_Nil, Ok) + end + end +. + +(** [betree::betree::{betree::betree::Node#5}::apply_to_leaf]: + Source: 'src/betree.rs', lines 460:4-460:87 *) +Definition betree_Node_apply_to_leaf + (n : nat) (bindings : betree_List_t (u64 * u64)) (key : u64) + (new_msg : betree_Message_t) : + result (betree_List_t (u64 * u64)) + := + p <- betree_Node_lookup_mut_in_bindings n key bindings; + let (bindings1, lookup_mut_in_bindings_back) := p in + b <- betree_ListPairU64T_head_has_key u64 bindings1 key; + if b + then ( + p1 <- betree_List_pop_front (u64 * u64) bindings1; + let (hd, bindings2) := p1 in + match new_msg with + | Betree_Message_Insert v => + bindings3 <- betree_List_push_front (u64 * u64) bindings2 (key, v); + lookup_mut_in_bindings_back bindings3 + | Betree_Message_Delete => lookup_mut_in_bindings_back bindings2 + | Betree_Message_Upsert s => + let (_, i) := hd in + v <- betree_upsert_update (Some i) s; + bindings3 <- betree_List_push_front (u64 * u64) bindings2 (key, v); + lookup_mut_in_bindings_back bindings3 + end) + else + match new_msg with + | Betree_Message_Insert v => + bindings2 <- betree_List_push_front (u64 * u64) bindings1 (key, v); + lookup_mut_in_bindings_back bindings2 + | Betree_Message_Delete => lookup_mut_in_bindings_back bindings1 + | Betree_Message_Upsert s => + v <- betree_upsert_update None s; + bindings2 <- betree_List_push_front (u64 * u64) bindings1 (key, v); + lookup_mut_in_bindings_back bindings2 + end +. + +(** [betree::betree::{betree::betree::Node#5}::apply_messages_to_leaf]: + Source: 'src/betree.rs', lines 444:4-447:5 *) +Fixpoint betree_Node_apply_messages_to_leaf + (n : nat) (bindings : betree_List_t (u64 * u64)) + (new_msgs : betree_List_t (u64 * betree_Message_t)) : + result (betree_List_t (u64 * u64)) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match new_msgs with + | Betree_List_Cons new_msg new_msgs_tl => + let (i, m) := new_msg in + bindings1 <- betree_Node_apply_to_leaf n1 bindings i m; + betree_Node_apply_messages_to_leaf n1 bindings1 new_msgs_tl + | Betree_List_Nil => Ok bindings + end + end +. + +(** [betree::betree::{betree::betree::Internal#4}::flush]: + Source: 'src/betree.rs', lines 410:4-415:26 *) +Fixpoint betree_Internal_flush + (n : nat) (self : betree_Internal_t) (params : betree_Params_t) + (node_id_cnt : betree_NodeIdCounter_t) + (content : betree_List_t (u64 * betree_Message_t)) (st : state) : + result (state * ((betree_List_t (u64 * betree_Message_t)) * + (betree_Internal_t * betree_NodeIdCounter_t))) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + p <- + betree_ListPairU64T_partition_at_pivot betree_Message_t n1 content + self.(betree_Internal_pivot); + let (msgs_left, msgs_right) := p in + len_left <- betree_List_len (u64 * betree_Message_t) n1 msgs_left; + if len_left s>= params.(betree_Params_min_flush_size) + then ( + p1 <- + betree_Node_apply_messages n1 self.(betree_Internal_left) params + node_id_cnt msgs_left st; + let (st1, p2) := p1 in + let (n2, node_id_cnt1) := p2 in + len_right <- betree_List_len (u64 * betree_Message_t) n1 msgs_right; + if len_right s>= params.(betree_Params_min_flush_size) + then ( + p3 <- + betree_Node_apply_messages n1 self.(betree_Internal_right) params + node_id_cnt1 msgs_right st1; + let (st2, p4) := p3 in + let (n3, node_id_cnt2) := p4 in + Ok (st2, (Betree_List_Nil, (mkbetree_Internal_t + self.(betree_Internal_id) self.(betree_Internal_pivot) n2 n3, + node_id_cnt2)))) + else + Ok (st1, (msgs_right, (mkbetree_Internal_t self.(betree_Internal_id) + self.(betree_Internal_pivot) n2 self.(betree_Internal_right), + node_id_cnt1)))) + else ( + p1 <- + betree_Node_apply_messages n1 self.(betree_Internal_right) params + node_id_cnt msgs_right st; + let (st1, p2) := p1 in + let (n2, node_id_cnt1) := p2 in + Ok (st1, (msgs_left, (mkbetree_Internal_t self.(betree_Internal_id) + self.(betree_Internal_pivot) self.(betree_Internal_left) n2, + node_id_cnt1)))) + end + +(** [betree::betree::{betree::betree::Node#5}::apply_messages]: + Source: 'src/betree.rs', lines 588:4-593:5 *) +with betree_Node_apply_messages + (n : nat) (self : betree_Node_t) (params : betree_Params_t) + (node_id_cnt : betree_NodeIdCounter_t) + (msgs : betree_List_t (u64 * betree_Message_t)) (st : state) : + result (state * (betree_Node_t * betree_NodeIdCounter_t)) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + match self with + | Betree_Node_Internal node => + p <- betree_load_internal_node node.(betree_Internal_id) st; + let (st1, content) := p in + content1 <- betree_Node_apply_messages_to_internal n1 content msgs; + num_msgs <- betree_List_len (u64 * betree_Message_t) n1 content1; + if num_msgs s>= params.(betree_Params_min_flush_size) + then ( + p1 <- betree_Internal_flush n1 node params node_id_cnt content1 st1; + let (st2, p2) := p1 in + let (content2, p3) := p2 in + let (node1, node_id_cnt1) := p3 in + p4 <- + betree_store_internal_node node1.(betree_Internal_id) content2 st2; + let (st3, _) := p4 in + Ok (st3, (Betree_Node_Internal node1, node_id_cnt1))) + else ( + p1 <- + betree_store_internal_node node.(betree_Internal_id) content1 st1; + let (st2, _) := p1 in + Ok (st2, (Betree_Node_Internal node, node_id_cnt))) + | Betree_Node_Leaf node => + p <- betree_load_leaf_node node.(betree_Leaf_id) st; + let (st1, content) := p in + content1 <- betree_Node_apply_messages_to_leaf n1 content msgs; + len <- betree_List_len (u64 * u64) n1 content1; + i <- u64_mul 2%u64 params.(betree_Params_split_size); + if len s>= i + then ( + p1 <- betree_Leaf_split n1 node content1 params node_id_cnt st1; + let (st2, p2) := p1 in + let (new_node, node_id_cnt1) := p2 in + p3 <- betree_store_leaf_node node.(betree_Leaf_id) Betree_List_Nil st2; + let (st3, _) := p3 in + Ok (st3, (Betree_Node_Internal new_node, node_id_cnt1))) + else ( + p1 <- betree_store_leaf_node node.(betree_Leaf_id) content1 st1; + let (st2, _) := p1 in + Ok (st2, (Betree_Node_Leaf + {| betree_Leaf_id := node.(betree_Leaf_id); betree_Leaf_size := len + |}, node_id_cnt))) + end + end +. + +(** [betree::betree::{betree::betree::Node#5}::apply]: + Source: 'src/betree.rs', lines 576:4-582:5 *) +Definition betree_Node_apply + (n : nat) (self : betree_Node_t) (params : betree_Params_t) + (node_id_cnt : betree_NodeIdCounter_t) (key : u64) + (new_msg : betree_Message_t) (st : state) : + result (state * (betree_Node_t * betree_NodeIdCounter_t)) + := + p <- + betree_Node_apply_messages n self params node_id_cnt (Betree_List_Cons + (key, new_msg) Betree_List_Nil) st; + let (st1, p1) := p in + let (self1, node_id_cnt1) := p1 in + Ok (st1, (self1, node_id_cnt1)) +. + +(** [betree::betree::{betree::betree::BeTree#6}::new]: + Source: 'src/betree.rs', lines 849:4-849:60 *) +Definition betree_BeTree_new + (min_flush_size : u64) (split_size : u64) (st : state) : + result (state * betree_BeTree_t) + := + node_id_cnt <- betree_NodeIdCounter_new; + p <- betree_NodeIdCounter_fresh_id node_id_cnt; + let (id, node_id_cnt1) := p in + p1 <- betree_store_leaf_node id Betree_List_Nil st; + let (st1, _) := p1 in + Ok (st1, + {| + betree_BeTree_params := + {| + betree_Params_min_flush_size := min_flush_size; + betree_Params_split_size := split_size + |}; + betree_BeTree_node_id_cnt := node_id_cnt1; + betree_BeTree_root := + (Betree_Node_Leaf + {| betree_Leaf_id := id; betree_Leaf_size := 0%u64 |}) + |}) +. + +(** [betree::betree::{betree::betree::BeTree#6}::apply]: + Source: 'src/betree.rs', lines 868:4-868:47 *) +Definition betree_BeTree_apply + (n : nat) (self : betree_BeTree_t) (key : u64) (msg : betree_Message_t) + (st : state) : + result (state * betree_BeTree_t) + := + p <- + betree_Node_apply n self.(betree_BeTree_root) self.(betree_BeTree_params) + self.(betree_BeTree_node_id_cnt) key msg st; + let (st1, p1) := p in + let (n1, nic) := p1 in + Ok (st1, + {| + betree_BeTree_params := self.(betree_BeTree_params); + betree_BeTree_node_id_cnt := nic; + betree_BeTree_root := n1 + |}) +. + +(** [betree::betree::{betree::betree::BeTree#6}::insert]: + Source: 'src/betree.rs', lines 874:4-874:52 *) +Definition betree_BeTree_insert + (n : nat) (self : betree_BeTree_t) (key : u64) (value : u64) (st : state) : + result (state * betree_BeTree_t) + := + betree_BeTree_apply n self key (Betree_Message_Insert value) st +. + +(** [betree::betree::{betree::betree::BeTree#6}::delete]: + Source: 'src/betree.rs', lines 880:4-880:38 *) +Definition betree_BeTree_delete + (n : nat) (self : betree_BeTree_t) (key : u64) (st : state) : + result (state * betree_BeTree_t) + := + betree_BeTree_apply n self key Betree_Message_Delete st +. + +(** [betree::betree::{betree::betree::BeTree#6}::upsert]: + Source: 'src/betree.rs', lines 886:4-886:59 *) +Definition betree_BeTree_upsert + (n : nat) (self : betree_BeTree_t) (key : u64) + (upd : betree_UpsertFunState_t) (st : state) : + result (state * betree_BeTree_t) + := + betree_BeTree_apply n self key (Betree_Message_Upsert upd) st +. + +(** [betree::betree::{betree::betree::BeTree#6}::lookup]: + Source: 'src/betree.rs', lines 895:4-895:62 *) +Definition betree_BeTree_lookup + (n : nat) (self : betree_BeTree_t) (key : u64) (st : state) : + result (state * ((option u64) * betree_BeTree_t)) + := + p <- betree_Node_lookup n self.(betree_BeTree_root) key st; + let (st1, p1) := p in + let (o, n1) := p1 in + Ok (st1, (o, + {| + betree_BeTree_params := self.(betree_BeTree_params); + betree_BeTree_node_id_cnt := self.(betree_BeTree_node_id_cnt); + betree_BeTree_root := n1 + |})) +. + +(** [betree::main]: + Source: 'src/main.rs', lines 4:0-4:9 *) +Definition main : result unit := + Ok tt. + +(** Unit test for [betree::main] *) +Check (main )%return. + +End Betree_Funs. diff --git a/tests/coq/betree/Betree_FunsExternal.v b/tests/coq/betree/Betree_FunsExternal.v new file mode 100644 index 00000000..0263c272 --- /dev/null +++ b/tests/coq/betree/Betree_FunsExternal.v @@ -0,0 +1,39 @@ +(** [betree_main]: external functions. +-- This is a template file: rename it to "FunsExternal.lean" and fill the holes. *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Require Export Betree_Types. +Import Betree_Types. +Module Betree_FunsExternal. + +(** [betree_main::betree_utils::load_internal_node]: forward function + Source: 'src/betree_utils.rs', lines 98:0-98:63 *) +Axiom betree_utils_load_internal_node + : u64 -> state -> result (state * (betree_List_t (u64 * betree_Message_t))) +. + +(** [betree_main::betree_utils::store_internal_node]: forward function + Source: 'src/betree_utils.rs', lines 115:0-115:71 *) +Axiom betree_utils_store_internal_node + : + u64 -> betree_List_t (u64 * betree_Message_t) -> state -> result (state * + unit) +. + +(** [betree_main::betree_utils::load_leaf_node]: forward function + Source: 'src/betree_utils.rs', lines 132:0-132:55 *) +Axiom betree_utils_load_leaf_node + : u64 -> state -> result (state * (betree_List_t (u64 * u64))) +. + +(** [betree_main::betree_utils::store_leaf_node]: forward function + Source: 'src/betree_utils.rs', lines 145:0-145:63 *) +Axiom betree_utils_store_leaf_node + : u64 -> betree_List_t (u64 * u64) -> state -> result (state * unit) +. + +End Betree_FunsExternal. diff --git a/tests/coq/betree/Betree_FunsExternal_Template.v b/tests/coq/betree/Betree_FunsExternal_Template.v new file mode 100644 index 00000000..2a2fddfc --- /dev/null +++ b/tests/coq/betree/Betree_FunsExternal_Template.v @@ -0,0 +1,40 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [betree]: external functions. +-- This is a template file: rename it to "FunsExternal.lean" and fill the holes. *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Require Import Betree_Types. +Include Betree_Types. +Module Betree_FunsExternal_Template. + +(** [betree::betree_utils::load_internal_node]: + Source: 'src/betree_utils.rs', lines 98:0-98:63 *) +Axiom betree_utils_load_internal_node + : u64 -> state -> result (state * (betree_List_t (u64 * betree_Message_t))) +. + +(** [betree::betree_utils::store_internal_node]: + Source: 'src/betree_utils.rs', lines 115:0-115:71 *) +Axiom betree_utils_store_internal_node + : + u64 -> betree_List_t (u64 * betree_Message_t) -> state -> result (state * + unit) +. + +(** [betree::betree_utils::load_leaf_node]: + Source: 'src/betree_utils.rs', lines 132:0-132:55 *) +Axiom betree_utils_load_leaf_node + : u64 -> state -> result (state * (betree_List_t (u64 * u64))) +. + +(** [betree::betree_utils::store_leaf_node]: + Source: 'src/betree_utils.rs', lines 145:0-145:63 *) +Axiom betree_utils_store_leaf_node + : u64 -> betree_List_t (u64 * u64) -> state -> result (state * unit) +. + +End Betree_FunsExternal_Template. diff --git a/tests/coq/betree/Betree_Types.v b/tests/coq/betree/Betree_Types.v new file mode 100644 index 00000000..aa39d8e4 --- /dev/null +++ b/tests/coq/betree/Betree_Types.v @@ -0,0 +1,118 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [betree]: type definitions *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Require Import Betree_TypesExternal. +Include Betree_TypesExternal. +Module Betree_Types. + +(** [betree::betree::List] + Source: 'src/betree.rs', lines 17:0-17:23 *) +Inductive betree_List_t (T : Type) := +| Betree_List_Cons : T -> betree_List_t T -> betree_List_t T +| Betree_List_Nil : betree_List_t T +. + +Arguments Betree_List_Cons { _ }. +Arguments Betree_List_Nil { _ }. + +(** [betree::betree::UpsertFunState] + Source: 'src/betree.rs', lines 63:0-63:23 *) +Inductive betree_UpsertFunState_t := +| Betree_UpsertFunState_Add : u64 -> betree_UpsertFunState_t +| Betree_UpsertFunState_Sub : u64 -> betree_UpsertFunState_t +. + +(** [betree::betree::Message] + Source: 'src/betree.rs', lines 69:0-69:23 *) +Inductive betree_Message_t := +| Betree_Message_Insert : u64 -> betree_Message_t +| Betree_Message_Delete : betree_Message_t +| Betree_Message_Upsert : betree_UpsertFunState_t -> betree_Message_t +. + +(** [betree::betree::Leaf] + Source: 'src/betree.rs', lines 167:0-167:11 *) +Record betree_Leaf_t := +mkbetree_Leaf_t { + betree_Leaf_id : u64; betree_Leaf_size : u64; +} +. + +(** [betree::betree::Internal] + Source: 'src/betree.rs', lines 156:0-156:15 *) +Inductive betree_Internal_t := +| mkbetree_Internal_t : + u64 -> + u64 -> + betree_Node_t -> + betree_Node_t -> + betree_Internal_t + +(** [betree::betree::Node] + Source: 'src/betree.rs', lines 179:0-179:9 *) +with betree_Node_t := +| Betree_Node_Internal : betree_Internal_t -> betree_Node_t +| Betree_Node_Leaf : betree_Leaf_t -> betree_Node_t +. + +Definition betree_Internal_id (x : betree_Internal_t) := + match x with | mkbetree_Internal_t x1 _ _ _ => x1 end +. + +Notation "x2 .(betree_Internal_id)" := (betree_Internal_id x2) (at level 9). + +Definition betree_Internal_pivot (x : betree_Internal_t) := + match x with | mkbetree_Internal_t _ x1 _ _ => x1 end +. + +Notation "x2 .(betree_Internal_pivot)" := (betree_Internal_pivot x2) + (at level 9) +. + +Definition betree_Internal_left (x : betree_Internal_t) := + match x with | mkbetree_Internal_t _ _ x1 _ => x1 end +. + +Notation "x2 .(betree_Internal_left)" := (betree_Internal_left x2) (at level 9) +. + +Definition betree_Internal_right (x : betree_Internal_t) := + match x with | mkbetree_Internal_t _ _ _ x1 => x1 end +. + +Notation "x2 .(betree_Internal_right)" := (betree_Internal_right x2) + (at level 9) +. + +(** [betree::betree::Params] + Source: 'src/betree.rs', lines 187:0-187:13 *) +Record betree_Params_t := +mkbetree_Params_t { + betree_Params_min_flush_size : u64; betree_Params_split_size : u64; +} +. + +(** [betree::betree::NodeIdCounter] + Source: 'src/betree.rs', lines 201:0-201:20 *) +Record betree_NodeIdCounter_t := +mkbetree_NodeIdCounter_t { + betree_NodeIdCounter_next_node_id : u64; +} +. + +(** [betree::betree::BeTree] + Source: 'src/betree.rs', lines 218:0-218:17 *) +Record betree_BeTree_t := +mkbetree_BeTree_t { + betree_BeTree_params : betree_Params_t; + betree_BeTree_node_id_cnt : betree_NodeIdCounter_t; + betree_BeTree_root : betree_Node_t; +} +. + +End Betree_Types. diff --git a/tests/coq/betree/Betree_TypesExternal.v b/tests/coq/betree/Betree_TypesExternal.v new file mode 100644 index 00000000..37487d6d --- /dev/null +++ b/tests/coq/betree/Betree_TypesExternal.v @@ -0,0 +1,14 @@ +(** [betree_main]: external types. +-- This is a template file: rename it to "TypesExternal.lean" and fill the holes. *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Module Betree_TypesExternal. + +(** The state type used in the state-error monad *) +Axiom state : Type. + +End Betree_TypesExternal. diff --git a/tests/coq/betree/Betree_TypesExternal_Template.v b/tests/coq/betree/Betree_TypesExternal_Template.v new file mode 100644 index 00000000..a40811cb --- /dev/null +++ b/tests/coq/betree/Betree_TypesExternal_Template.v @@ -0,0 +1,15 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [betree]: external types. +-- This is a template file: rename it to "TypesExternal.lean" and fill the holes. *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Module Betree_TypesExternal_Template. + +(** The state type used in the state-error monad *) +Axiom state : Type. + +End Betree_TypesExternal_Template. diff --git a/tests/coq/betree/_CoqProject b/tests/coq/betree/_CoqProject index b14bed66..9fb7fe64 100644 --- a/tests/coq/betree/_CoqProject +++ b/tests/coq/betree/_CoqProject @@ -3,10 +3,10 @@ -arg -w -arg all -BetreeMain_TypesExternal_Template.v -BetreeMain_Types.v +Betree_TypesExternal_Template.v +Betree_Types.v Primitives.v -BetreeMain_FunsExternal_Template.v -BetreeMain_Funs.v -BetreeMain_TypesExternal.v -BetreeMain_FunsExternal.v +Betree_FunsExternal_Template.v +Betree_Funs.v +Betree_TypesExternal.v +Betree_FunsExternal.v diff --git a/tests/fstar/betree/Betree.Clauses.Template.fst b/tests/fstar/betree/Betree.Clauses.Template.fst new file mode 100644 index 00000000..fad8c5ba --- /dev/null +++ b/tests/fstar/betree/Betree.Clauses.Template.fst @@ -0,0 +1,117 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [betree]: templates for the decreases clauses *) +module Betree.Clauses.Template +open Primitives +open Betree.Types + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** [betree::betree::{betree::betree::List#1}::len]: decreases clause + Source: 'src/betree.rs', lines 276:4-276:24 *) +unfold +let betree_List_len_decreases (t : Type0) (self : betree_List_t t) : nat = + admit () + +(** [betree::betree::{betree::betree::List#1}::split_at]: decreases clause + Source: 'src/betree.rs', lines 284:4-284:51 *) +unfold +let betree_List_split_at_decreases (t : Type0) (self : betree_List_t t) + (n : u64) : nat = + admit () + +(** [betree::betree::{betree::betree::List<(u64, T)>#2}::partition_at_pivot]: decreases clause + Source: 'src/betree.rs', lines 339:4-339:73 *) +unfold +let betree_ListPairU64T_partition_at_pivot_decreases (t : Type0) + (self : betree_List_t (u64 & t)) (pivot : u64) : nat = + admit () + +(** [betree::betree::{betree::betree::Node#5}::lookup_first_message_for_key]: decreases clause + Source: 'src/betree.rs', lines 789:4-792:34 *) +unfold +let betree_Node_lookup_first_message_for_key_decreases (key : u64) + (msgs : betree_List_t (u64 & betree_Message_t)) : nat = + admit () + +(** [betree::betree::{betree::betree::Node#5}::lookup_in_bindings]: decreases clause + Source: 'src/betree.rs', lines 636:4-636:80 *) +unfold +let betree_Node_lookup_in_bindings_decreases (key : u64) + (bindings : betree_List_t (u64 & u64)) : nat = + admit () + +(** [betree::betree::{betree::betree::Node#5}::apply_upserts]: decreases clause + Source: 'src/betree.rs', lines 819:4-819:90 *) +unfold +let betree_Node_apply_upserts_decreases + (msgs : betree_List_t (u64 & betree_Message_t)) (prev : option u64) + (key : u64) : nat = + admit () + +(** [betree::betree::{betree::betree::Internal#4}::lookup_in_children]: decreases clause + Source: 'src/betree.rs', lines 395:4-395:63 *) +unfold +let betree_Internal_lookup_in_children_decreases (self : betree_Internal_t) + (key : u64) (st : state) : nat = + admit () + +(** [betree::betree::{betree::betree::Node#5}::lookup]: decreases clause + Source: 'src/betree.rs', lines 709:4-709:58 *) +unfold +let betree_Node_lookup_decreases (self : betree_Node_t) (key : u64) + (st : state) : nat = + admit () + +(** [betree::betree::{betree::betree::Node#5}::filter_messages_for_key]: decreases clause + Source: 'src/betree.rs', lines 674:4-674:77 *) +unfold +let betree_Node_filter_messages_for_key_decreases (key : u64) + (msgs : betree_List_t (u64 & betree_Message_t)) : nat = + admit () + +(** [betree::betree::{betree::betree::Node#5}::lookup_first_message_after_key]: decreases clause + Source: 'src/betree.rs', lines 689:4-692:34 *) +unfold +let betree_Node_lookup_first_message_after_key_decreases (key : u64) + (msgs : betree_List_t (u64 & betree_Message_t)) : nat = + admit () + +(** [betree::betree::{betree::betree::Node#5}::apply_messages_to_internal]: decreases clause + Source: 'src/betree.rs', lines 502:4-505:5 *) +unfold +let betree_Node_apply_messages_to_internal_decreases + (msgs : betree_List_t (u64 & betree_Message_t)) + (new_msgs : betree_List_t (u64 & betree_Message_t)) : nat = + admit () + +(** [betree::betree::{betree::betree::Node#5}::lookup_mut_in_bindings]: decreases clause + Source: 'src/betree.rs', lines 653:4-656:32 *) +unfold +let betree_Node_lookup_mut_in_bindings_decreases (key : u64) + (bindings : betree_List_t (u64 & u64)) : nat = + admit () + +(** [betree::betree::{betree::betree::Node#5}::apply_messages_to_leaf]: decreases clause + Source: 'src/betree.rs', lines 444:4-447:5 *) +unfold +let betree_Node_apply_messages_to_leaf_decreases + (bindings : betree_List_t (u64 & u64)) + (new_msgs : betree_List_t (u64 & betree_Message_t)) : nat = + admit () + +(** [betree::betree::{betree::betree::Internal#4}::flush]: decreases clause + Source: 'src/betree.rs', lines 410:4-415:26 *) +unfold +let betree_Internal_flush_decreases (self : betree_Internal_t) + (params : betree_Params_t) (node_id_cnt : betree_NodeIdCounter_t) + (content : betree_List_t (u64 & betree_Message_t)) (st : state) : nat = + admit () + +(** [betree::betree::{betree::betree::Node#5}::apply_messages]: decreases clause + Source: 'src/betree.rs', lines 588:4-593:5 *) +unfold +let betree_Node_apply_messages_decreases (self : betree_Node_t) + (params : betree_Params_t) (node_id_cnt : betree_NodeIdCounter_t) + (msgs : betree_List_t (u64 & betree_Message_t)) (st : state) : nat = + admit () + diff --git a/tests/fstar/betree/Betree.Clauses.fst b/tests/fstar/betree/Betree.Clauses.fst new file mode 100644 index 00000000..ae201cee --- /dev/null +++ b/tests/fstar/betree/Betree.Clauses.fst @@ -0,0 +1,210 @@ +(** [betree_main]: templates for the decreases clauses *) +module Betree.Clauses +open Primitives +open Betree.Types + +#set-options "--z3rlimit 50 --fuel 0 --ifuel 1" + +(*** Well-founded relations *) + +(* We had a few issues when proving termination of the mutually recursive functions: + * - betree_Internal_flush + * - betree_Node_apply_messages + * + * The quantity which effectively decreases is: + * (betree_size, messages_length) + * where messages_length is 0 when there are no messages + * (and where we use the lexicographic ordering, of course) + * + * However, the `%[...]` and `{:well-founded ...} notations are not available outside + * of `decrease` clauses. + * + * We thus resorted to writing and proving correct a well-founded relation over + * pairs of natural numbers. The trick is that `<<` can be used outside of decrease + * clauses, and can be used to trigger SMT patterns. + * + * What follows is adapted from: + * https://www.fstar-lang.org/tutorial/book/part2/part2_well_founded.html + * + * Also, the following PR might make things easier: + * https://github.com/FStarLang/FStar/pull/2561 + *) + +module P = FStar.Preorder +module W = FStar.WellFounded +module L = FStar.LexicographicOrdering + +let lt_nat (x y:nat) : Type = x < y == true +let rec wf_lt_nat (x:nat) + : W.acc lt_nat x + = W.AccIntro (fun y _ -> wf_lt_nat y) + +// A type abbreviation for a pair of nats +let nat_pair = (x:nat & nat) + +// Making a lexicographic ordering from a pair of nat ordering +let lex_order_nat_pair : P.relation nat_pair = + L.lex_t lt_nat (fun _ -> lt_nat) + +// The lex order on nat pairs is well-founded, using our general proof +// of lexicographic composition of well-founded orders +let lex_order_nat_pair_wf : W.well_founded lex_order_nat_pair = + L.lex_t_wf wf_lt_nat (fun _ -> wf_lt_nat) + +// A utility to introduce lt_nat +let mk_lt_nat (x:nat) (y:nat { x < y }) : lt_nat x y = + let _ : equals (x < y) true = Refl in + () + +// A utility to make a lex ordering of nat pairs +let mk_lex_order_nat_pair (xy0:nat_pair) + (xy1:nat_pair { + let (|x0, y0|) = xy0 in + let (|x1, y1|) = xy1 in + x0 < x1 \/ (x0 == x1 /\ y0 < y1) + }) : lex_order_nat_pair xy0 xy1 = + let (|x0, y0|) = xy0 in + let (|x1, y1|) = xy1 in + if x0 < x1 then L.Left_lex x0 x1 y0 y1 (mk_lt_nat x0 x1) + else L.Right_lex x0 y0 y1 (mk_lt_nat y0 y1) + +let rec coerce #a #r #x (p:W.acc #a r x) : Tot (W.acc r x) (decreases p) = + W.AccIntro (fun y r -> coerce (p.access_smaller y r)) + +let coerce_wf #a #r (p: (x:a -> W.acc r x)) : x:a -> W.acc r x = + fun x -> coerce (p x) + +(* We need this axiom, which comes from the following discussion: + * https://github.com/FStarLang/FStar/issues/1916 + * An issue here is that the `{well-founded ... }` notation + *) +assume +val axiom_well_founded (a : Type) (rel : a -> a -> Type0) + (rwf : W.well_founded #a rel) (x y : a) : + Lemma (requires (rel x y)) (ensures (x << y)) + +(* This lemma has a pattern (which makes it work) *) +let wf_nat_pair_lem (p0 p1 : nat_pair) : + Lemma + (requires ( + let (|x0, y0|) = p0 in + let (|x1, y1|) = p1 in + x0 < x1 || (x0 = x1 && y0 < y1))) + (ensures (p0 << p1)) + [SMTPat (p0 << p1)] = + let rel = lex_order_nat_pair in + let rel_wf = lex_order_nat_pair_wf in + let _ = mk_lex_order_nat_pair p0 p1 in + assert(rel p0 p1); + axiom_well_founded nat_pair rel rel_wf p0 p1 + +(*** Decrease clauses *) +/// "Standard" decrease clauses + +(** [betree_main::betree::List::{1}::len]: decreases clause *) +unfold +let betree_List_len_decreases (t : Type0) (self : betree_List_t t) : betree_List_t t = + self + +(** [betree_main::betree::List::{1}::split_at]: decreases clause *) +unfold +let betree_List_split_at_decreases (t : Type0) (self : betree_List_t t) + (n : u64) : nat = + n + +(** [betree_main::betree::List::{2}::partition_at_pivot]: decreases clause *) +unfold +let betree_ListPairU64T_partition_at_pivot_decreases (t : Type0) + (self : betree_List_t (u64 & t)) (pivot : u64) : betree_List_t (u64 & t) = + self + +(** [betree_main::betree::Node::{5}::lookup_in_bindings]: decreases clause *) +unfold +let betree_Node_lookup_in_bindings_decreases (key : u64) + (bindings : betree_List_t (u64 & u64)) : betree_List_t (u64 & u64) = + bindings + +(** [betree_main::betree::Node::{5}::lookup_first_message_for_key]: decreases clause *) +unfold +let betree_Node_lookup_first_message_for_key_decreases (key : u64) + (msgs : betree_List_t (u64 & betree_Message_t)) : betree_List_t (u64 & betree_Message_t) = + msgs + +(** [betree_main::betree::Node::{5}::apply_upserts]: decreases clause *) +unfold +let betree_Node_apply_upserts_decreases + (msgs : betree_List_t (u64 & betree_Message_t)) (prev : option u64) + (key : u64) : betree_List_t (u64 & betree_Message_t) = + msgs + +(** [betree_main::betree::Internal::{4}::lookup_in_children]: decreases clause *) +unfold +let betree_Internal_lookup_in_children_decreases (self : betree_Internal_t) + (key : u64) (st : state) : betree_Internal_t = + self + +(** [betree_main::betree::Node::{5}::lookup]: decreases clause *) +unfold +let betree_Node_lookup_decreases (self : betree_Node_t) (key : u64) + (st : state) : betree_Node_t = + self + +(** [betree_main::betree::Node::{5}::lookup_mut_in_bindings]: decreases clause *) +unfold +let betree_Node_lookup_mut_in_bindings_decreases (key : u64) + (bindings : betree_List_t (u64 & u64)) : betree_List_t (u64 & u64) = + bindings + +unfold +let betree_Node_apply_messages_to_leaf_decreases + (bindings : betree_List_t (u64 & u64)) + (new_msgs : betree_List_t (u64 & betree_Message_t)) : betree_List_t (u64 & betree_Message_t) = + new_msgs + +(** [betree_main::betree::Node::{5}::filter_messages_for_key]: decreases clause *) +unfold +let betree_Node_filter_messages_for_key_decreases (key : u64) + (msgs : betree_List_t (u64 & betree_Message_t)) : betree_List_t (u64 & betree_Message_t) = + msgs + +(** [betree_main::betree::Node::{5}::lookup_first_message_after_key]: decreases clause *) +unfold +let betree_Node_lookup_first_message_after_key_decreases (key : u64) + (msgs : betree_List_t (u64 & betree_Message_t)) : betree_List_t (u64 & betree_Message_t) = + msgs + +let betree_Node_apply_messages_to_internal_decreases + (msgs : betree_List_t (u64 & betree_Message_t)) + (new_msgs : betree_List_t (u64 & betree_Message_t)) : betree_List_t (u64 & betree_Message_t) = + new_msgs + +(*** Decrease clauses - nat_pair *) +/// The following decrease clauses use the [nat_pair] definition and the well-founded +/// relation proven above. + +let rec betree_size (bt : betree_Node_t) : nat = + match bt with + | Betree_Node_Internal node -> 1 + betree_Internal_size node + | Betree_Node_Leaf _ -> 1 + +and betree_Internal_size (node : betree_Internal_t) : nat = + 1 + betree_size node.left + betree_size node.right + +let rec betree_List_len (#a : Type0) (ls : betree_List_t a) : nat = + match ls with + | Betree_List_Cons _ tl -> 1 + betree_List_len tl + | Betree_List_Nil -> 0 + +(** [betree_main::betree::Internal::{4}::flush]: decreases clause *) +unfold +let betree_Internal_flush_decreases (self : betree_Internal_t) + (params : betree_Params_t) (node_id_cnt : betree_NodeIdCounter_t) + (content : betree_List_t (u64 & betree_Message_t)) (st : state) : nat_pair = + (|betree_Internal_size self, 0|) + +(** [betree_main::betree::Node::{5}::apply_messages]: decreases clause *) +unfold +let betree_Node_apply_messages_decreases (self : betree_Node_t) + (params : betree_Params_t) (node_id_cnt : betree_NodeIdCounter_t) + (msgs : betree_List_t (u64 & betree_Message_t)) (st : state) : nat_pair = + (|betree_size self, betree_List_len msgs|) diff --git a/tests/fstar/betree/Betree.Funs.fst b/tests/fstar/betree/Betree.Funs.fst new file mode 100644 index 00000000..07f561f5 --- /dev/null +++ b/tests/fstar/betree/Betree.Funs.fst @@ -0,0 +1,676 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [betree]: function definitions *) +module Betree.Funs +open Primitives +include Betree.Types +include Betree.FunsExternal +include Betree.Clauses + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** [betree::betree::load_internal_node]: + Source: 'src/betree.rs', lines 36:0-36:52 *) +let betree_load_internal_node + (id : u64) (st : state) : + result (state & (betree_List_t (u64 & betree_Message_t))) + = + betree_utils_load_internal_node id st + +(** [betree::betree::store_internal_node]: + Source: 'src/betree.rs', lines 41:0-41:60 *) +let betree_store_internal_node + (id : u64) (content : betree_List_t (u64 & betree_Message_t)) (st : state) : + result (state & unit) + = + betree_utils_store_internal_node id content st + +(** [betree::betree::load_leaf_node]: + Source: 'src/betree.rs', lines 46:0-46:44 *) +let betree_load_leaf_node + (id : u64) (st : state) : result (state & (betree_List_t (u64 & u64))) = + betree_utils_load_leaf_node id st + +(** [betree::betree::store_leaf_node]: + Source: 'src/betree.rs', lines 51:0-51:52 *) +let betree_store_leaf_node + (id : u64) (content : betree_List_t (u64 & u64)) (st : state) : + result (state & unit) + = + betree_utils_store_leaf_node id content st + +(** [betree::betree::fresh_node_id]: + Source: 'src/betree.rs', lines 55:0-55:48 *) +let betree_fresh_node_id (counter : u64) : result (u64 & u64) = + let* counter1 = u64_add counter 1 in Ok (counter, counter1) + +(** [betree::betree::{betree::betree::NodeIdCounter}::new]: + Source: 'src/betree.rs', lines 206:4-206:20 *) +let betree_NodeIdCounter_new : result betree_NodeIdCounter_t = + Ok { next_node_id = 0 } + +(** [betree::betree::{betree::betree::NodeIdCounter}::fresh_id]: + Source: 'src/betree.rs', lines 210:4-210:36 *) +let betree_NodeIdCounter_fresh_id + (self : betree_NodeIdCounter_t) : result (u64 & betree_NodeIdCounter_t) = + let* i = u64_add self.next_node_id 1 in + Ok (self.next_node_id, { next_node_id = i }) + +(** [betree::betree::upsert_update]: + Source: 'src/betree.rs', lines 234:0-234:70 *) +let betree_upsert_update + (prev : option u64) (st : betree_UpsertFunState_t) : result u64 = + begin match prev with + | None -> + begin match st with + | Betree_UpsertFunState_Add v -> Ok v + | Betree_UpsertFunState_Sub _ -> Ok 0 + end + | Some prev1 -> + begin match st with + | Betree_UpsertFunState_Add v -> + let* margin = u64_sub core_u64_max prev1 in + if margin >= v then u64_add prev1 v else Ok core_u64_max + | Betree_UpsertFunState_Sub v -> + if prev1 >= v then u64_sub prev1 v else Ok 0 + end + end + +(** [betree::betree::{betree::betree::List#1}::len]: + Source: 'src/betree.rs', lines 276:4-276:24 *) +let rec betree_List_len + (t : Type0) (self : betree_List_t t) : + Tot (result u64) (decreases (betree_List_len_decreases t self)) + = + begin match self with + | Betree_List_Cons _ tl -> let* i = betree_List_len t tl in u64_add 1 i + | Betree_List_Nil -> Ok 0 + end + +(** [betree::betree::{betree::betree::List#1}::split_at]: + Source: 'src/betree.rs', lines 284:4-284:51 *) +let rec betree_List_split_at + (t : Type0) (self : betree_List_t t) (n : u64) : + Tot (result ((betree_List_t t) & (betree_List_t t))) + (decreases (betree_List_split_at_decreases t self n)) + = + if n = 0 + then Ok (Betree_List_Nil, self) + else + begin match self with + | Betree_List_Cons hd tl -> + let* i = u64_sub n 1 in + let* p = betree_List_split_at t tl i in + let (ls0, ls1) = p in + Ok (Betree_List_Cons hd ls0, ls1) + | Betree_List_Nil -> Fail Failure + end + +(** [betree::betree::{betree::betree::List#1}::push_front]: + Source: 'src/betree.rs', lines 299:4-299:34 *) +let betree_List_push_front + (t : Type0) (self : betree_List_t t) (x : t) : result (betree_List_t t) = + let (tl, _) = core_mem_replace (betree_List_t t) self Betree_List_Nil in + Ok (Betree_List_Cons x tl) + +(** [betree::betree::{betree::betree::List#1}::pop_front]: + Source: 'src/betree.rs', lines 306:4-306:32 *) +let betree_List_pop_front + (t : Type0) (self : betree_List_t t) : result (t & (betree_List_t t)) = + let (ls, _) = core_mem_replace (betree_List_t t) self Betree_List_Nil in + begin match ls with + | Betree_List_Cons x tl -> Ok (x, tl) + | Betree_List_Nil -> Fail Failure + end + +(** [betree::betree::{betree::betree::List#1}::hd]: + Source: 'src/betree.rs', lines 318:4-318:22 *) +let betree_List_hd (t : Type0) (self : betree_List_t t) : result t = + begin match self with + | Betree_List_Cons hd _ -> Ok hd + | Betree_List_Nil -> Fail Failure + end + +(** [betree::betree::{betree::betree::List<(u64, T)>#2}::head_has_key]: + Source: 'src/betree.rs', lines 327:4-327:44 *) +let betree_ListPairU64T_head_has_key + (t : Type0) (self : betree_List_t (u64 & t)) (key : u64) : result bool = + begin match self with + | Betree_List_Cons hd _ -> let (i, _) = hd in Ok (i = key) + | Betree_List_Nil -> Ok false + end + +(** [betree::betree::{betree::betree::List<(u64, T)>#2}::partition_at_pivot]: + Source: 'src/betree.rs', lines 339:4-339:73 *) +let rec betree_ListPairU64T_partition_at_pivot + (t : Type0) (self : betree_List_t (u64 & t)) (pivot : u64) : + Tot (result ((betree_List_t (u64 & t)) & (betree_List_t (u64 & t)))) + (decreases (betree_ListPairU64T_partition_at_pivot_decreases t self pivot)) + = + begin match self with + | Betree_List_Cons hd tl -> + let (i, x) = hd in + if i >= pivot + then Ok (Betree_List_Nil, Betree_List_Cons (i, x) tl) + else + let* p = betree_ListPairU64T_partition_at_pivot t tl pivot in + let (ls0, ls1) = p in + Ok (Betree_List_Cons (i, x) ls0, ls1) + | Betree_List_Nil -> Ok (Betree_List_Nil, Betree_List_Nil) + end + +(** [betree::betree::{betree::betree::Leaf#3}::split]: + Source: 'src/betree.rs', lines 359:4-364:17 *) +let betree_Leaf_split + (self : betree_Leaf_t) (content : betree_List_t (u64 & u64)) + (params : betree_Params_t) (node_id_cnt : betree_NodeIdCounter_t) + (st : state) : + result (state & (betree_Internal_t & betree_NodeIdCounter_t)) + = + let* p = betree_List_split_at (u64 & u64) content params.split_size in + let (content0, content1) = p in + let* p1 = betree_List_hd (u64 & u64) content1 in + let (pivot, _) = p1 in + let* (id0, node_id_cnt1) = betree_NodeIdCounter_fresh_id node_id_cnt in + let* (id1, node_id_cnt2) = betree_NodeIdCounter_fresh_id node_id_cnt1 in + let* (st1, _) = betree_store_leaf_node id0 content0 st in + let* (st2, _) = betree_store_leaf_node id1 content1 st1 in + let n = Betree_Node_Leaf { id = id0; size = params.split_size } in + let n1 = Betree_Node_Leaf { id = id1; size = params.split_size } in + Ok (st2, ({ id = self.id; pivot = pivot; left = n; right = n1 }, + node_id_cnt2)) + +(** [betree::betree::{betree::betree::Node#5}::lookup_first_message_for_key]: + Source: 'src/betree.rs', lines 789:4-792:34 *) +let rec betree_Node_lookup_first_message_for_key + (key : u64) (msgs : betree_List_t (u64 & betree_Message_t)) : + Tot (result ((betree_List_t (u64 & betree_Message_t)) & (betree_List_t (u64 & + betree_Message_t) -> result (betree_List_t (u64 & betree_Message_t))))) + (decreases (betree_Node_lookup_first_message_for_key_decreases key msgs)) + = + begin match msgs with + | Betree_List_Cons x next_msgs -> + let (i, m) = x in + if i >= key + then Ok (Betree_List_Cons (i, m) next_msgs, Ok) + else + let* (l, lookup_first_message_for_key_back) = + betree_Node_lookup_first_message_for_key key next_msgs in + let back = + fun ret -> + let* next_msgs1 = lookup_first_message_for_key_back ret in + Ok (Betree_List_Cons (i, m) next_msgs1) in + Ok (l, back) + | Betree_List_Nil -> Ok (Betree_List_Nil, Ok) + end + +(** [betree::betree::{betree::betree::Node#5}::lookup_in_bindings]: + Source: 'src/betree.rs', lines 636:4-636:80 *) +let rec betree_Node_lookup_in_bindings + (key : u64) (bindings : betree_List_t (u64 & u64)) : + Tot (result (option u64)) + (decreases (betree_Node_lookup_in_bindings_decreases key bindings)) + = + begin match bindings with + | Betree_List_Cons hd tl -> + let (i, i1) = hd in + if i = key + then Ok (Some i1) + else if i > key then Ok None else betree_Node_lookup_in_bindings key tl + | Betree_List_Nil -> Ok None + end + +(** [betree::betree::{betree::betree::Node#5}::apply_upserts]: + Source: 'src/betree.rs', lines 819:4-819:90 *) +let rec betree_Node_apply_upserts + (msgs : betree_List_t (u64 & betree_Message_t)) (prev : option u64) + (key : u64) : + Tot (result (u64 & (betree_List_t (u64 & betree_Message_t)))) + (decreases (betree_Node_apply_upserts_decreases msgs prev key)) + = + let* b = betree_ListPairU64T_head_has_key betree_Message_t msgs key in + if b + then + let* (msg, msgs1) = betree_List_pop_front (u64 & betree_Message_t) msgs in + let (_, m) = msg in + begin match m with + | Betree_Message_Insert _ -> Fail Failure + | Betree_Message_Delete -> Fail Failure + | Betree_Message_Upsert s -> + let* v = betree_upsert_update prev s in + betree_Node_apply_upserts msgs1 (Some v) key + end + else + let* v = core_option_Option_unwrap u64 prev in + let* msgs1 = + betree_List_push_front (u64 & betree_Message_t) msgs (key, + Betree_Message_Insert v) in + Ok (v, msgs1) + +(** [betree::betree::{betree::betree::Internal#4}::lookup_in_children]: + Source: 'src/betree.rs', lines 395:4-395:63 *) +let rec betree_Internal_lookup_in_children + (self : betree_Internal_t) (key : u64) (st : state) : + Tot (result (state & ((option u64) & betree_Internal_t))) + (decreases (betree_Internal_lookup_in_children_decreases self key st)) + = + if key < self.pivot + then + let* (st1, (o, n)) = betree_Node_lookup self.left key st in + Ok (st1, (o, { self with left = n })) + else + let* (st1, (o, n)) = betree_Node_lookup self.right key st in + Ok (st1, (o, { self with right = n })) + +(** [betree::betree::{betree::betree::Node#5}::lookup]: + Source: 'src/betree.rs', lines 709:4-709:58 *) +and betree_Node_lookup + (self : betree_Node_t) (key : u64) (st : state) : + Tot (result (state & ((option u64) & betree_Node_t))) + (decreases (betree_Node_lookup_decreases self key st)) + = + begin match self with + | Betree_Node_Internal node -> + let* (st1, msgs) = betree_load_internal_node node.id st in + let* (pending, lookup_first_message_for_key_back) = + betree_Node_lookup_first_message_for_key key msgs in + begin match pending with + | Betree_List_Cons p l -> + let (k, msg) = p in + if k <> key + then + let* (st2, (o, node1)) = + betree_Internal_lookup_in_children node key st1 in + let* _ = + lookup_first_message_for_key_back (Betree_List_Cons (k, msg) l) in + Ok (st2, (o, Betree_Node_Internal node1)) + else + begin match msg with + | Betree_Message_Insert v -> + let* _ = + lookup_first_message_for_key_back (Betree_List_Cons (k, + Betree_Message_Insert v) l) in + Ok (st1, (Some v, Betree_Node_Internal node)) + | Betree_Message_Delete -> + let* _ = + lookup_first_message_for_key_back (Betree_List_Cons (k, + Betree_Message_Delete) l) in + Ok (st1, (None, Betree_Node_Internal node)) + | Betree_Message_Upsert ufs -> + let* (st2, (v, node1)) = + betree_Internal_lookup_in_children node key st1 in + let* (v1, pending1) = + betree_Node_apply_upserts (Betree_List_Cons (k, + Betree_Message_Upsert ufs) l) v key in + let* msgs1 = lookup_first_message_for_key_back pending1 in + let* (st3, _) = betree_store_internal_node node1.id msgs1 st2 in + Ok (st3, (Some v1, Betree_Node_Internal node1)) + end + | Betree_List_Nil -> + let* (st2, (o, node1)) = betree_Internal_lookup_in_children node key st1 + in + let* _ = lookup_first_message_for_key_back Betree_List_Nil in + Ok (st2, (o, Betree_Node_Internal node1)) + end + | Betree_Node_Leaf node -> + let* (st1, bindings) = betree_load_leaf_node node.id st in + let* o = betree_Node_lookup_in_bindings key bindings in + Ok (st1, (o, Betree_Node_Leaf node)) + end + +(** [betree::betree::{betree::betree::Node#5}::filter_messages_for_key]: + Source: 'src/betree.rs', lines 674:4-674:77 *) +let rec betree_Node_filter_messages_for_key + (key : u64) (msgs : betree_List_t (u64 & betree_Message_t)) : + Tot (result (betree_List_t (u64 & betree_Message_t))) + (decreases (betree_Node_filter_messages_for_key_decreases key msgs)) + = + begin match msgs with + | Betree_List_Cons p l -> + let (k, m) = p in + if k = key + then + let* (_, msgs1) = + betree_List_pop_front (u64 & betree_Message_t) (Betree_List_Cons (k, m) + l) in + betree_Node_filter_messages_for_key key msgs1 + else Ok (Betree_List_Cons (k, m) l) + | Betree_List_Nil -> Ok Betree_List_Nil + end + +(** [betree::betree::{betree::betree::Node#5}::lookup_first_message_after_key]: + Source: 'src/betree.rs', lines 689:4-692:34 *) +let rec betree_Node_lookup_first_message_after_key + (key : u64) (msgs : betree_List_t (u64 & betree_Message_t)) : + Tot (result ((betree_List_t (u64 & betree_Message_t)) & (betree_List_t (u64 & + betree_Message_t) -> result (betree_List_t (u64 & betree_Message_t))))) + (decreases (betree_Node_lookup_first_message_after_key_decreases key msgs)) + = + begin match msgs with + | Betree_List_Cons p next_msgs -> + let (k, m) = p in + if k = key + then + let* (l, lookup_first_message_after_key_back) = + betree_Node_lookup_first_message_after_key key next_msgs in + let back = + fun ret -> + let* next_msgs1 = lookup_first_message_after_key_back ret in + Ok (Betree_List_Cons (k, m) next_msgs1) in + Ok (l, back) + else Ok (Betree_List_Cons (k, m) next_msgs, Ok) + | Betree_List_Nil -> Ok (Betree_List_Nil, Ok) + end + +(** [betree::betree::{betree::betree::Node#5}::apply_to_internal]: + Source: 'src/betree.rs', lines 521:4-521:89 *) +let betree_Node_apply_to_internal + (msgs : betree_List_t (u64 & betree_Message_t)) (key : u64) + (new_msg : betree_Message_t) : + result (betree_List_t (u64 & betree_Message_t)) + = + let* (msgs1, lookup_first_message_for_key_back) = + betree_Node_lookup_first_message_for_key key msgs in + let* b = betree_ListPairU64T_head_has_key betree_Message_t msgs1 key in + if b + then + begin match new_msg with + | Betree_Message_Insert i -> + let* msgs2 = betree_Node_filter_messages_for_key key msgs1 in + let* msgs3 = + betree_List_push_front (u64 & betree_Message_t) msgs2 (key, + Betree_Message_Insert i) in + lookup_first_message_for_key_back msgs3 + | Betree_Message_Delete -> + let* msgs2 = betree_Node_filter_messages_for_key key msgs1 in + let* msgs3 = + betree_List_push_front (u64 & betree_Message_t) msgs2 (key, + Betree_Message_Delete) in + lookup_first_message_for_key_back msgs3 + | Betree_Message_Upsert s -> + let* p = betree_List_hd (u64 & betree_Message_t) msgs1 in + let (_, m) = p in + begin match m with + | Betree_Message_Insert prev -> + let* v = betree_upsert_update (Some prev) s in + let* (_, msgs2) = betree_List_pop_front (u64 & betree_Message_t) msgs1 + in + let* msgs3 = + betree_List_push_front (u64 & betree_Message_t) msgs2 (key, + Betree_Message_Insert v) in + lookup_first_message_for_key_back msgs3 + | Betree_Message_Delete -> + let* (_, msgs2) = betree_List_pop_front (u64 & betree_Message_t) msgs1 + in + let* v = betree_upsert_update None s in + let* msgs3 = + betree_List_push_front (u64 & betree_Message_t) msgs2 (key, + Betree_Message_Insert v) in + lookup_first_message_for_key_back msgs3 + | Betree_Message_Upsert _ -> + let* (msgs2, lookup_first_message_after_key_back) = + betree_Node_lookup_first_message_after_key key msgs1 in + let* msgs3 = + betree_List_push_front (u64 & betree_Message_t) msgs2 (key, + Betree_Message_Upsert s) in + let* msgs4 = lookup_first_message_after_key_back msgs3 in + lookup_first_message_for_key_back msgs4 + end + end + else + let* msgs2 = + betree_List_push_front (u64 & betree_Message_t) msgs1 (key, new_msg) in + lookup_first_message_for_key_back msgs2 + +(** [betree::betree::{betree::betree::Node#5}::apply_messages_to_internal]: + Source: 'src/betree.rs', lines 502:4-505:5 *) +let rec betree_Node_apply_messages_to_internal + (msgs : betree_List_t (u64 & betree_Message_t)) + (new_msgs : betree_List_t (u64 & betree_Message_t)) : + Tot (result (betree_List_t (u64 & betree_Message_t))) + (decreases (betree_Node_apply_messages_to_internal_decreases msgs new_msgs)) + = + begin match new_msgs with + | Betree_List_Cons new_msg new_msgs_tl -> + let (i, m) = new_msg in + let* msgs1 = betree_Node_apply_to_internal msgs i m in + betree_Node_apply_messages_to_internal msgs1 new_msgs_tl + | Betree_List_Nil -> Ok msgs + end + +(** [betree::betree::{betree::betree::Node#5}::lookup_mut_in_bindings]: + Source: 'src/betree.rs', lines 653:4-656:32 *) +let rec betree_Node_lookup_mut_in_bindings + (key : u64) (bindings : betree_List_t (u64 & u64)) : + Tot (result ((betree_List_t (u64 & u64)) & (betree_List_t (u64 & u64) -> + result (betree_List_t (u64 & u64))))) + (decreases (betree_Node_lookup_mut_in_bindings_decreases key bindings)) + = + begin match bindings with + | Betree_List_Cons hd tl -> + let (i, i1) = hd in + if i >= key + then Ok (Betree_List_Cons (i, i1) tl, Ok) + else + let* (l, lookup_mut_in_bindings_back) = + betree_Node_lookup_mut_in_bindings key tl in + let back = + fun ret -> + let* tl1 = lookup_mut_in_bindings_back ret in + Ok (Betree_List_Cons (i, i1) tl1) in + Ok (l, back) + | Betree_List_Nil -> Ok (Betree_List_Nil, Ok) + end + +(** [betree::betree::{betree::betree::Node#5}::apply_to_leaf]: + Source: 'src/betree.rs', lines 460:4-460:87 *) +let betree_Node_apply_to_leaf + (bindings : betree_List_t (u64 & u64)) (key : u64) + (new_msg : betree_Message_t) : + result (betree_List_t (u64 & u64)) + = + let* (bindings1, lookup_mut_in_bindings_back) = + betree_Node_lookup_mut_in_bindings key bindings in + let* b = betree_ListPairU64T_head_has_key u64 bindings1 key in + if b + then + let* (hd, bindings2) = betree_List_pop_front (u64 & u64) bindings1 in + begin match new_msg with + | Betree_Message_Insert v -> + let* bindings3 = betree_List_push_front (u64 & u64) bindings2 (key, v) in + lookup_mut_in_bindings_back bindings3 + | Betree_Message_Delete -> lookup_mut_in_bindings_back bindings2 + | Betree_Message_Upsert s -> + let (_, i) = hd in + let* v = betree_upsert_update (Some i) s in + let* bindings3 = betree_List_push_front (u64 & u64) bindings2 (key, v) in + lookup_mut_in_bindings_back bindings3 + end + else + begin match new_msg with + | Betree_Message_Insert v -> + let* bindings2 = betree_List_push_front (u64 & u64) bindings1 (key, v) in + lookup_mut_in_bindings_back bindings2 + | Betree_Message_Delete -> lookup_mut_in_bindings_back bindings1 + | Betree_Message_Upsert s -> + let* v = betree_upsert_update None s in + let* bindings2 = betree_List_push_front (u64 & u64) bindings1 (key, v) in + lookup_mut_in_bindings_back bindings2 + end + +(** [betree::betree::{betree::betree::Node#5}::apply_messages_to_leaf]: + Source: 'src/betree.rs', lines 444:4-447:5 *) +let rec betree_Node_apply_messages_to_leaf + (bindings : betree_List_t (u64 & u64)) + (new_msgs : betree_List_t (u64 & betree_Message_t)) : + Tot (result (betree_List_t (u64 & u64))) + (decreases (betree_Node_apply_messages_to_leaf_decreases bindings new_msgs)) + = + begin match new_msgs with + | Betree_List_Cons new_msg new_msgs_tl -> + let (i, m) = new_msg in + let* bindings1 = betree_Node_apply_to_leaf bindings i m in + betree_Node_apply_messages_to_leaf bindings1 new_msgs_tl + | Betree_List_Nil -> Ok bindings + end + +(** [betree::betree::{betree::betree::Internal#4}::flush]: + Source: 'src/betree.rs', lines 410:4-415:26 *) +let rec betree_Internal_flush + (self : betree_Internal_t) (params : betree_Params_t) + (node_id_cnt : betree_NodeIdCounter_t) + (content : betree_List_t (u64 & betree_Message_t)) (st : state) : + Tot (result (state & ((betree_List_t (u64 & betree_Message_t)) & + (betree_Internal_t & betree_NodeIdCounter_t)))) + (decreases ( + betree_Internal_flush_decreases self params node_id_cnt content st)) + = + let* p = + betree_ListPairU64T_partition_at_pivot betree_Message_t content self.pivot + in + let (msgs_left, msgs_right) = p in + let* len_left = betree_List_len (u64 & betree_Message_t) msgs_left in + if len_left >= params.min_flush_size + then + let* (st1, p1) = + betree_Node_apply_messages self.left params node_id_cnt msgs_left st in + let (n, node_id_cnt1) = p1 in + let* len_right = betree_List_len (u64 & betree_Message_t) msgs_right in + if len_right >= params.min_flush_size + then + let* (st2, p2) = + betree_Node_apply_messages self.right params node_id_cnt1 msgs_right + st1 in + let (n1, node_id_cnt2) = p2 in + Ok (st2, (Betree_List_Nil, ({ self with left = n; right = n1 }, + node_id_cnt2))) + else Ok (st1, (msgs_right, ({ self with left = n }, node_id_cnt1))) + else + let* (st1, p1) = + betree_Node_apply_messages self.right params node_id_cnt msgs_right st in + let (n, node_id_cnt1) = p1 in + Ok (st1, (msgs_left, ({ self with right = n }, node_id_cnt1))) + +(** [betree::betree::{betree::betree::Node#5}::apply_messages]: + Source: 'src/betree.rs', lines 588:4-593:5 *) +and betree_Node_apply_messages + (self : betree_Node_t) (params : betree_Params_t) + (node_id_cnt : betree_NodeIdCounter_t) + (msgs : betree_List_t (u64 & betree_Message_t)) (st : state) : + Tot (result (state & (betree_Node_t & betree_NodeIdCounter_t))) + (decreases ( + betree_Node_apply_messages_decreases self params node_id_cnt msgs st)) + = + begin match self with + | Betree_Node_Internal node -> + let* (st1, content) = betree_load_internal_node node.id st in + let* content1 = betree_Node_apply_messages_to_internal content msgs in + let* num_msgs = betree_List_len (u64 & betree_Message_t) content1 in + if num_msgs >= params.min_flush_size + then + let* (st2, (content2, p)) = + betree_Internal_flush node params node_id_cnt content1 st1 in + let (node1, node_id_cnt1) = p in + let* (st3, _) = betree_store_internal_node node1.id content2 st2 in + Ok (st3, (Betree_Node_Internal node1, node_id_cnt1)) + else + let* (st2, _) = betree_store_internal_node node.id content1 st1 in + Ok (st2, (Betree_Node_Internal node, node_id_cnt)) + | Betree_Node_Leaf node -> + let* (st1, content) = betree_load_leaf_node node.id st in + let* content1 = betree_Node_apply_messages_to_leaf content msgs in + let* len = betree_List_len (u64 & u64) content1 in + let* i = u64_mul 2 params.split_size in + if len >= i + then + let* (st2, (new_node, node_id_cnt1)) = + betree_Leaf_split node content1 params node_id_cnt st1 in + let* (st3, _) = betree_store_leaf_node node.id Betree_List_Nil st2 in + Ok (st3, (Betree_Node_Internal new_node, node_id_cnt1)) + else + let* (st2, _) = betree_store_leaf_node node.id content1 st1 in + Ok (st2, (Betree_Node_Leaf { node with size = len }, node_id_cnt)) + end + +(** [betree::betree::{betree::betree::Node#5}::apply]: + Source: 'src/betree.rs', lines 576:4-582:5 *) +let betree_Node_apply + (self : betree_Node_t) (params : betree_Params_t) + (node_id_cnt : betree_NodeIdCounter_t) (key : u64) + (new_msg : betree_Message_t) (st : state) : + result (state & (betree_Node_t & betree_NodeIdCounter_t)) + = + let* (st1, p) = + betree_Node_apply_messages self params node_id_cnt (Betree_List_Cons (key, + new_msg) Betree_List_Nil) st in + let (self1, node_id_cnt1) = p in + Ok (st1, (self1, node_id_cnt1)) + +(** [betree::betree::{betree::betree::BeTree#6}::new]: + Source: 'src/betree.rs', lines 849:4-849:60 *) +let betree_BeTree_new + (min_flush_size : u64) (split_size : u64) (st : state) : + result (state & betree_BeTree_t) + = + let* node_id_cnt = betree_NodeIdCounter_new in + let* (id, node_id_cnt1) = betree_NodeIdCounter_fresh_id node_id_cnt in + let* (st1, _) = betree_store_leaf_node id Betree_List_Nil st in + Ok (st1, + { + params = { min_flush_size = min_flush_size; split_size = split_size }; + node_id_cnt = node_id_cnt1; + root = (Betree_Node_Leaf { id = id; size = 0 }) + }) + +(** [betree::betree::{betree::betree::BeTree#6}::apply]: + Source: 'src/betree.rs', lines 868:4-868:47 *) +let betree_BeTree_apply + (self : betree_BeTree_t) (key : u64) (msg : betree_Message_t) (st : state) : + result (state & betree_BeTree_t) + = + let* (st1, p) = + betree_Node_apply self.root self.params self.node_id_cnt key msg st in + let (n, nic) = p in + Ok (st1, { self with node_id_cnt = nic; root = n }) + +(** [betree::betree::{betree::betree::BeTree#6}::insert]: + Source: 'src/betree.rs', lines 874:4-874:52 *) +let betree_BeTree_insert + (self : betree_BeTree_t) (key : u64) (value : u64) (st : state) : + result (state & betree_BeTree_t) + = + betree_BeTree_apply self key (Betree_Message_Insert value) st + +(** [betree::betree::{betree::betree::BeTree#6}::delete]: + Source: 'src/betree.rs', lines 880:4-880:38 *) +let betree_BeTree_delete + (self : betree_BeTree_t) (key : u64) (st : state) : + result (state & betree_BeTree_t) + = + betree_BeTree_apply self key Betree_Message_Delete st + +(** [betree::betree::{betree::betree::BeTree#6}::upsert]: + Source: 'src/betree.rs', lines 886:4-886:59 *) +let betree_BeTree_upsert + (self : betree_BeTree_t) (key : u64) (upd : betree_UpsertFunState_t) + (st : state) : + result (state & betree_BeTree_t) + = + betree_BeTree_apply self key (Betree_Message_Upsert upd) st + +(** [betree::betree::{betree::betree::BeTree#6}::lookup]: + Source: 'src/betree.rs', lines 895:4-895:62 *) +let betree_BeTree_lookup + (self : betree_BeTree_t) (key : u64) (st : state) : + result (state & ((option u64) & betree_BeTree_t)) + = + let* (st1, (o, n)) = betree_Node_lookup self.root key st in + Ok (st1, (o, { self with root = n })) + +(** [betree::main]: + Source: 'src/main.rs', lines 4:0-4:9 *) +let main : result unit = + Ok () + +(** Unit test for [betree::main] *) +let _ = assert_norm (main = Ok ()) + diff --git a/tests/fstar/betree/Betree.FunsExternal.fsti b/tests/fstar/betree/Betree.FunsExternal.fsti new file mode 100644 index 00000000..db96eead --- /dev/null +++ b/tests/fstar/betree/Betree.FunsExternal.fsti @@ -0,0 +1,30 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [betree]: external function declarations *) +module Betree.FunsExternal +open Primitives +include Betree.Types + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** [betree::betree_utils::load_internal_node]: + Source: 'src/betree_utils.rs', lines 98:0-98:63 *) +val betree_utils_load_internal_node + : u64 -> state -> result (state & (betree_List_t (u64 & betree_Message_t))) + +(** [betree::betree_utils::store_internal_node]: + Source: 'src/betree_utils.rs', lines 115:0-115:71 *) +val betree_utils_store_internal_node + : + u64 -> betree_List_t (u64 & betree_Message_t) -> state -> result (state & + unit) + +(** [betree::betree_utils::load_leaf_node]: + Source: 'src/betree_utils.rs', lines 132:0-132:55 *) +val betree_utils_load_leaf_node + : u64 -> state -> result (state & (betree_List_t (u64 & u64))) + +(** [betree::betree_utils::store_leaf_node]: + Source: 'src/betree_utils.rs', lines 145:0-145:63 *) +val betree_utils_store_leaf_node + : u64 -> betree_List_t (u64 & u64) -> state -> result (state & unit) + diff --git a/tests/fstar/betree/Betree.Types.fst b/tests/fstar/betree/Betree.Types.fst new file mode 100644 index 00000000..6b8e063b --- /dev/null +++ b/tests/fstar/betree/Betree.Types.fst @@ -0,0 +1,61 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [betree]: type definitions *) +module Betree.Types +open Primitives +include Betree.TypesExternal + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** [betree::betree::List] + Source: 'src/betree.rs', lines 17:0-17:23 *) +type betree_List_t (t : Type0) = +| Betree_List_Cons : t -> betree_List_t t -> betree_List_t t +| Betree_List_Nil : betree_List_t t + +(** [betree::betree::UpsertFunState] + Source: 'src/betree.rs', lines 63:0-63:23 *) +type betree_UpsertFunState_t = +| Betree_UpsertFunState_Add : u64 -> betree_UpsertFunState_t +| Betree_UpsertFunState_Sub : u64 -> betree_UpsertFunState_t + +(** [betree::betree::Message] + Source: 'src/betree.rs', lines 69:0-69:23 *) +type betree_Message_t = +| Betree_Message_Insert : u64 -> betree_Message_t +| Betree_Message_Delete : betree_Message_t +| Betree_Message_Upsert : betree_UpsertFunState_t -> betree_Message_t + +(** [betree::betree::Leaf] + Source: 'src/betree.rs', lines 167:0-167:11 *) +type betree_Leaf_t = { id : u64; size : u64; } + +(** [betree::betree::Internal] + Source: 'src/betree.rs', lines 156:0-156:15 *) +type betree_Internal_t = +{ + id : u64; pivot : u64; left : betree_Node_t; right : betree_Node_t; +} + +(** [betree::betree::Node] + Source: 'src/betree.rs', lines 179:0-179:9 *) +and betree_Node_t = +| Betree_Node_Internal : betree_Internal_t -> betree_Node_t +| Betree_Node_Leaf : betree_Leaf_t -> betree_Node_t + +(** [betree::betree::Params] + Source: 'src/betree.rs', lines 187:0-187:13 *) +type betree_Params_t = { min_flush_size : u64; split_size : u64; } + +(** [betree::betree::NodeIdCounter] + Source: 'src/betree.rs', lines 201:0-201:20 *) +type betree_NodeIdCounter_t = { next_node_id : u64; } + +(** [betree::betree::BeTree] + Source: 'src/betree.rs', lines 218:0-218:17 *) +type betree_BeTree_t = +{ + params : betree_Params_t; + node_id_cnt : betree_NodeIdCounter_t; + root : betree_Node_t; +} + diff --git a/tests/fstar/betree/Betree.TypesExternal.fsti b/tests/fstar/betree/Betree.TypesExternal.fsti new file mode 100644 index 00000000..39e53ec9 --- /dev/null +++ b/tests/fstar/betree/Betree.TypesExternal.fsti @@ -0,0 +1,10 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [betree]: external type declarations *) +module Betree.TypesExternal +open Primitives + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** The state type used in the state-error monad *) +val state : Type0 + diff --git a/tests/fstar/betree/BetreeMain.Clauses.Template.fst b/tests/fstar/betree/BetreeMain.Clauses.Template.fst deleted file mode 100644 index b317dca4..00000000 --- a/tests/fstar/betree/BetreeMain.Clauses.Template.fst +++ /dev/null @@ -1,117 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [betree_main]: templates for the decreases clauses *) -module BetreeMain.Clauses.Template -open Primitives -open BetreeMain.Types - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** [betree_main::betree::{betree_main::betree::List#1}::len]: decreases clause - Source: 'src/betree.rs', lines 276:4-276:24 *) -unfold -let betree_List_len_decreases (t : Type0) (self : betree_List_t t) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::List#1}::split_at]: decreases clause - Source: 'src/betree.rs', lines 284:4-284:51 *) -unfold -let betree_List_split_at_decreases (t : Type0) (self : betree_List_t t) - (n : u64) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::List<(u64, T)>#2}::partition_at_pivot]: decreases clause - Source: 'src/betree.rs', lines 339:4-339:73 *) -unfold -let betree_ListPairU64T_partition_at_pivot_decreases (t : Type0) - (self : betree_List_t (u64 & t)) (pivot : u64) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup_first_message_for_key]: decreases clause - Source: 'src/betree.rs', lines 789:4-792:34 *) -unfold -let betree_Node_lookup_first_message_for_key_decreases (key : u64) - (msgs : betree_List_t (u64 & betree_Message_t)) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup_in_bindings]: decreases clause - Source: 'src/betree.rs', lines 636:4-636:80 *) -unfold -let betree_Node_lookup_in_bindings_decreases (key : u64) - (bindings : betree_List_t (u64 & u64)) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_upserts]: decreases clause - Source: 'src/betree.rs', lines 819:4-819:90 *) -unfold -let betree_Node_apply_upserts_decreases - (msgs : betree_List_t (u64 & betree_Message_t)) (prev : option u64) - (key : u64) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::Internal#4}::lookup_in_children]: decreases clause - Source: 'src/betree.rs', lines 395:4-395:63 *) -unfold -let betree_Internal_lookup_in_children_decreases (self : betree_Internal_t) - (key : u64) (st : state) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup]: decreases clause - Source: 'src/betree.rs', lines 709:4-709:58 *) -unfold -let betree_Node_lookup_decreases (self : betree_Node_t) (key : u64) - (st : state) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::Node#5}::filter_messages_for_key]: decreases clause - Source: 'src/betree.rs', lines 674:4-674:77 *) -unfold -let betree_Node_filter_messages_for_key_decreases (key : u64) - (msgs : betree_List_t (u64 & betree_Message_t)) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup_first_message_after_key]: decreases clause - Source: 'src/betree.rs', lines 689:4-692:34 *) -unfold -let betree_Node_lookup_first_message_after_key_decreases (key : u64) - (msgs : betree_List_t (u64 & betree_Message_t)) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_messages_to_internal]: decreases clause - Source: 'src/betree.rs', lines 502:4-505:5 *) -unfold -let betree_Node_apply_messages_to_internal_decreases - (msgs : betree_List_t (u64 & betree_Message_t)) - (new_msgs : betree_List_t (u64 & betree_Message_t)) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup_mut_in_bindings]: decreases clause - Source: 'src/betree.rs', lines 653:4-656:32 *) -unfold -let betree_Node_lookup_mut_in_bindings_decreases (key : u64) - (bindings : betree_List_t (u64 & u64)) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_messages_to_leaf]: decreases clause - Source: 'src/betree.rs', lines 444:4-447:5 *) -unfold -let betree_Node_apply_messages_to_leaf_decreases - (bindings : betree_List_t (u64 & u64)) - (new_msgs : betree_List_t (u64 & betree_Message_t)) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::Internal#4}::flush]: decreases clause - Source: 'src/betree.rs', lines 410:4-415:26 *) -unfold -let betree_Internal_flush_decreases (self : betree_Internal_t) - (params : betree_Params_t) (node_id_cnt : betree_NodeIdCounter_t) - (content : betree_List_t (u64 & betree_Message_t)) (st : state) : nat = - admit () - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_messages]: decreases clause - Source: 'src/betree.rs', lines 588:4-593:5 *) -unfold -let betree_Node_apply_messages_decreases (self : betree_Node_t) - (params : betree_Params_t) (node_id_cnt : betree_NodeIdCounter_t) - (msgs : betree_List_t (u64 & betree_Message_t)) (st : state) : nat = - admit () - diff --git a/tests/fstar/betree/BetreeMain.Clauses.fst b/tests/fstar/betree/BetreeMain.Clauses.fst deleted file mode 100644 index b95d4c7e..00000000 --- a/tests/fstar/betree/BetreeMain.Clauses.fst +++ /dev/null @@ -1,210 +0,0 @@ -(** [betree_main]: templates for the decreases clauses *) -module BetreeMain.Clauses -open Primitives -open BetreeMain.Types - -#set-options "--z3rlimit 50 --fuel 0 --ifuel 1" - -(*** Well-founded relations *) - -(* We had a few issues when proving termination of the mutually recursive functions: - * - betree_Internal_flush - * - betree_Node_apply_messages - * - * The quantity which effectively decreases is: - * (betree_size, messages_length) - * where messages_length is 0 when there are no messages - * (and where we use the lexicographic ordering, of course) - * - * However, the `%[...]` and `{:well-founded ...} notations are not available outside - * of `decrease` clauses. - * - * We thus resorted to writing and proving correct a well-founded relation over - * pairs of natural numbers. The trick is that `<<` can be used outside of decrease - * clauses, and can be used to trigger SMT patterns. - * - * What follows is adapted from: - * https://www.fstar-lang.org/tutorial/book/part2/part2_well_founded.html - * - * Also, the following PR might make things easier: - * https://github.com/FStarLang/FStar/pull/2561 - *) - -module P = FStar.Preorder -module W = FStar.WellFounded -module L = FStar.LexicographicOrdering - -let lt_nat (x y:nat) : Type = x < y == true -let rec wf_lt_nat (x:nat) - : W.acc lt_nat x - = W.AccIntro (fun y _ -> wf_lt_nat y) - -// A type abbreviation for a pair of nats -let nat_pair = (x:nat & nat) - -// Making a lexicographic ordering from a pair of nat ordering -let lex_order_nat_pair : P.relation nat_pair = - L.lex_t lt_nat (fun _ -> lt_nat) - -// The lex order on nat pairs is well-founded, using our general proof -// of lexicographic composition of well-founded orders -let lex_order_nat_pair_wf : W.well_founded lex_order_nat_pair = - L.lex_t_wf wf_lt_nat (fun _ -> wf_lt_nat) - -// A utility to introduce lt_nat -let mk_lt_nat (x:nat) (y:nat { x < y }) : lt_nat x y = - let _ : equals (x < y) true = Refl in - () - -// A utility to make a lex ordering of nat pairs -let mk_lex_order_nat_pair (xy0:nat_pair) - (xy1:nat_pair { - let (|x0, y0|) = xy0 in - let (|x1, y1|) = xy1 in - x0 < x1 \/ (x0 == x1 /\ y0 < y1) - }) : lex_order_nat_pair xy0 xy1 = - let (|x0, y0|) = xy0 in - let (|x1, y1|) = xy1 in - if x0 < x1 then L.Left_lex x0 x1 y0 y1 (mk_lt_nat x0 x1) - else L.Right_lex x0 y0 y1 (mk_lt_nat y0 y1) - -let rec coerce #a #r #x (p:W.acc #a r x) : Tot (W.acc r x) (decreases p) = - W.AccIntro (fun y r -> coerce (p.access_smaller y r)) - -let coerce_wf #a #r (p: (x:a -> W.acc r x)) : x:a -> W.acc r x = - fun x -> coerce (p x) - -(* We need this axiom, which comes from the following discussion: - * https://github.com/FStarLang/FStar/issues/1916 - * An issue here is that the `{well-founded ... }` notation - *) -assume -val axiom_well_founded (a : Type) (rel : a -> a -> Type0) - (rwf : W.well_founded #a rel) (x y : a) : - Lemma (requires (rel x y)) (ensures (x << y)) - -(* This lemma has a pattern (which makes it work) *) -let wf_nat_pair_lem (p0 p1 : nat_pair) : - Lemma - (requires ( - let (|x0, y0|) = p0 in - let (|x1, y1|) = p1 in - x0 < x1 || (x0 = x1 && y0 < y1))) - (ensures (p0 << p1)) - [SMTPat (p0 << p1)] = - let rel = lex_order_nat_pair in - let rel_wf = lex_order_nat_pair_wf in - let _ = mk_lex_order_nat_pair p0 p1 in - assert(rel p0 p1); - axiom_well_founded nat_pair rel rel_wf p0 p1 - -(*** Decrease clauses *) -/// "Standard" decrease clauses - -(** [betree_main::betree::List::{1}::len]: decreases clause *) -unfold -let betree_List_len_decreases (t : Type0) (self : betree_List_t t) : betree_List_t t = - self - -(** [betree_main::betree::List::{1}::split_at]: decreases clause *) -unfold -let betree_List_split_at_decreases (t : Type0) (self : betree_List_t t) - (n : u64) : nat = - n - -(** [betree_main::betree::List::{2}::partition_at_pivot]: decreases clause *) -unfold -let betree_ListPairU64T_partition_at_pivot_decreases (t : Type0) - (self : betree_List_t (u64 & t)) (pivot : u64) : betree_List_t (u64 & t) = - self - -(** [betree_main::betree::Node::{5}::lookup_in_bindings]: decreases clause *) -unfold -let betree_Node_lookup_in_bindings_decreases (key : u64) - (bindings : betree_List_t (u64 & u64)) : betree_List_t (u64 & u64) = - bindings - -(** [betree_main::betree::Node::{5}::lookup_first_message_for_key]: decreases clause *) -unfold -let betree_Node_lookup_first_message_for_key_decreases (key : u64) - (msgs : betree_List_t (u64 & betree_Message_t)) : betree_List_t (u64 & betree_Message_t) = - msgs - -(** [betree_main::betree::Node::{5}::apply_upserts]: decreases clause *) -unfold -let betree_Node_apply_upserts_decreases - (msgs : betree_List_t (u64 & betree_Message_t)) (prev : option u64) - (key : u64) : betree_List_t (u64 & betree_Message_t) = - msgs - -(** [betree_main::betree::Internal::{4}::lookup_in_children]: decreases clause *) -unfold -let betree_Internal_lookup_in_children_decreases (self : betree_Internal_t) - (key : u64) (st : state) : betree_Internal_t = - self - -(** [betree_main::betree::Node::{5}::lookup]: decreases clause *) -unfold -let betree_Node_lookup_decreases (self : betree_Node_t) (key : u64) - (st : state) : betree_Node_t = - self - -(** [betree_main::betree::Node::{5}::lookup_mut_in_bindings]: decreases clause *) -unfold -let betree_Node_lookup_mut_in_bindings_decreases (key : u64) - (bindings : betree_List_t (u64 & u64)) : betree_List_t (u64 & u64) = - bindings - -unfold -let betree_Node_apply_messages_to_leaf_decreases - (bindings : betree_List_t (u64 & u64)) - (new_msgs : betree_List_t (u64 & betree_Message_t)) : betree_List_t (u64 & betree_Message_t) = - new_msgs - -(** [betree_main::betree::Node::{5}::filter_messages_for_key]: decreases clause *) -unfold -let betree_Node_filter_messages_for_key_decreases (key : u64) - (msgs : betree_List_t (u64 & betree_Message_t)) : betree_List_t (u64 & betree_Message_t) = - msgs - -(** [betree_main::betree::Node::{5}::lookup_first_message_after_key]: decreases clause *) -unfold -let betree_Node_lookup_first_message_after_key_decreases (key : u64) - (msgs : betree_List_t (u64 & betree_Message_t)) : betree_List_t (u64 & betree_Message_t) = - msgs - -let betree_Node_apply_messages_to_internal_decreases - (msgs : betree_List_t (u64 & betree_Message_t)) - (new_msgs : betree_List_t (u64 & betree_Message_t)) : betree_List_t (u64 & betree_Message_t) = - new_msgs - -(*** Decrease clauses - nat_pair *) -/// The following decrease clauses use the [nat_pair] definition and the well-founded -/// relation proven above. - -let rec betree_size (bt : betree_Node_t) : nat = - match bt with - | Betree_Node_Internal node -> 1 + betree_Internal_size node - | Betree_Node_Leaf _ -> 1 - -and betree_Internal_size (node : betree_Internal_t) : nat = - 1 + betree_size node.left + betree_size node.right - -let rec betree_List_len (#a : Type0) (ls : betree_List_t a) : nat = - match ls with - | Betree_List_Cons _ tl -> 1 + betree_List_len tl - | Betree_List_Nil -> 0 - -(** [betree_main::betree::Internal::{4}::flush]: decreases clause *) -unfold -let betree_Internal_flush_decreases (self : betree_Internal_t) - (params : betree_Params_t) (node_id_cnt : betree_NodeIdCounter_t) - (content : betree_List_t (u64 & betree_Message_t)) (st : state) : nat_pair = - (|betree_Internal_size self, 0|) - -(** [betree_main::betree::Node::{5}::apply_messages]: decreases clause *) -unfold -let betree_Node_apply_messages_decreases (self : betree_Node_t) - (params : betree_Params_t) (node_id_cnt : betree_NodeIdCounter_t) - (msgs : betree_List_t (u64 & betree_Message_t)) (st : state) : nat_pair = - (|betree_size self, betree_List_len msgs|) diff --git a/tests/fstar/betree/BetreeMain.Funs.fst b/tests/fstar/betree/BetreeMain.Funs.fst deleted file mode 100644 index 9942ef68..00000000 --- a/tests/fstar/betree/BetreeMain.Funs.fst +++ /dev/null @@ -1,676 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [betree_main]: function definitions *) -module BetreeMain.Funs -open Primitives -include BetreeMain.Types -include BetreeMain.FunsExternal -include BetreeMain.Clauses - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** [betree_main::betree::load_internal_node]: - Source: 'src/betree.rs', lines 36:0-36:52 *) -let betree_load_internal_node - (id : u64) (st : state) : - result (state & (betree_List_t (u64 & betree_Message_t))) - = - betree_utils_load_internal_node id st - -(** [betree_main::betree::store_internal_node]: - Source: 'src/betree.rs', lines 41:0-41:60 *) -let betree_store_internal_node - (id : u64) (content : betree_List_t (u64 & betree_Message_t)) (st : state) : - result (state & unit) - = - betree_utils_store_internal_node id content st - -(** [betree_main::betree::load_leaf_node]: - Source: 'src/betree.rs', lines 46:0-46:44 *) -let betree_load_leaf_node - (id : u64) (st : state) : result (state & (betree_List_t (u64 & u64))) = - betree_utils_load_leaf_node id st - -(** [betree_main::betree::store_leaf_node]: - Source: 'src/betree.rs', lines 51:0-51:52 *) -let betree_store_leaf_node - (id : u64) (content : betree_List_t (u64 & u64)) (st : state) : - result (state & unit) - = - betree_utils_store_leaf_node id content st - -(** [betree_main::betree::fresh_node_id]: - Source: 'src/betree.rs', lines 55:0-55:48 *) -let betree_fresh_node_id (counter : u64) : result (u64 & u64) = - let* counter1 = u64_add counter 1 in Ok (counter, counter1) - -(** [betree_main::betree::{betree_main::betree::NodeIdCounter}::new]: - Source: 'src/betree.rs', lines 206:4-206:20 *) -let betree_NodeIdCounter_new : result betree_NodeIdCounter_t = - Ok { next_node_id = 0 } - -(** [betree_main::betree::{betree_main::betree::NodeIdCounter}::fresh_id]: - Source: 'src/betree.rs', lines 210:4-210:36 *) -let betree_NodeIdCounter_fresh_id - (self : betree_NodeIdCounter_t) : result (u64 & betree_NodeIdCounter_t) = - let* i = u64_add self.next_node_id 1 in - Ok (self.next_node_id, { next_node_id = i }) - -(** [betree_main::betree::upsert_update]: - Source: 'src/betree.rs', lines 234:0-234:70 *) -let betree_upsert_update - (prev : option u64) (st : betree_UpsertFunState_t) : result u64 = - begin match prev with - | None -> - begin match st with - | Betree_UpsertFunState_Add v -> Ok v - | Betree_UpsertFunState_Sub _ -> Ok 0 - end - | Some prev1 -> - begin match st with - | Betree_UpsertFunState_Add v -> - let* margin = u64_sub core_u64_max prev1 in - if margin >= v then u64_add prev1 v else Ok core_u64_max - | Betree_UpsertFunState_Sub v -> - if prev1 >= v then u64_sub prev1 v else Ok 0 - end - end - -(** [betree_main::betree::{betree_main::betree::List#1}::len]: - Source: 'src/betree.rs', lines 276:4-276:24 *) -let rec betree_List_len - (t : Type0) (self : betree_List_t t) : - Tot (result u64) (decreases (betree_List_len_decreases t self)) - = - begin match self with - | Betree_List_Cons _ tl -> let* i = betree_List_len t tl in u64_add 1 i - | Betree_List_Nil -> Ok 0 - end - -(** [betree_main::betree::{betree_main::betree::List#1}::split_at]: - Source: 'src/betree.rs', lines 284:4-284:51 *) -let rec betree_List_split_at - (t : Type0) (self : betree_List_t t) (n : u64) : - Tot (result ((betree_List_t t) & (betree_List_t t))) - (decreases (betree_List_split_at_decreases t self n)) - = - if n = 0 - then Ok (Betree_List_Nil, self) - else - begin match self with - | Betree_List_Cons hd tl -> - let* i = u64_sub n 1 in - let* p = betree_List_split_at t tl i in - let (ls0, ls1) = p in - Ok (Betree_List_Cons hd ls0, ls1) - | Betree_List_Nil -> Fail Failure - end - -(** [betree_main::betree::{betree_main::betree::List#1}::push_front]: - Source: 'src/betree.rs', lines 299:4-299:34 *) -let betree_List_push_front - (t : Type0) (self : betree_List_t t) (x : t) : result (betree_List_t t) = - let (tl, _) = core_mem_replace (betree_List_t t) self Betree_List_Nil in - Ok (Betree_List_Cons x tl) - -(** [betree_main::betree::{betree_main::betree::List#1}::pop_front]: - Source: 'src/betree.rs', lines 306:4-306:32 *) -let betree_List_pop_front - (t : Type0) (self : betree_List_t t) : result (t & (betree_List_t t)) = - let (ls, _) = core_mem_replace (betree_List_t t) self Betree_List_Nil in - begin match ls with - | Betree_List_Cons x tl -> Ok (x, tl) - | Betree_List_Nil -> Fail Failure - end - -(** [betree_main::betree::{betree_main::betree::List#1}::hd]: - Source: 'src/betree.rs', lines 318:4-318:22 *) -let betree_List_hd (t : Type0) (self : betree_List_t t) : result t = - begin match self with - | Betree_List_Cons hd _ -> Ok hd - | Betree_List_Nil -> Fail Failure - end - -(** [betree_main::betree::{betree_main::betree::List<(u64, T)>#2}::head_has_key]: - Source: 'src/betree.rs', lines 327:4-327:44 *) -let betree_ListPairU64T_head_has_key - (t : Type0) (self : betree_List_t (u64 & t)) (key : u64) : result bool = - begin match self with - | Betree_List_Cons hd _ -> let (i, _) = hd in Ok (i = key) - | Betree_List_Nil -> Ok false - end - -(** [betree_main::betree::{betree_main::betree::List<(u64, T)>#2}::partition_at_pivot]: - Source: 'src/betree.rs', lines 339:4-339:73 *) -let rec betree_ListPairU64T_partition_at_pivot - (t : Type0) (self : betree_List_t (u64 & t)) (pivot : u64) : - Tot (result ((betree_List_t (u64 & t)) & (betree_List_t (u64 & t)))) - (decreases (betree_ListPairU64T_partition_at_pivot_decreases t self pivot)) - = - begin match self with - | Betree_List_Cons hd tl -> - let (i, x) = hd in - if i >= pivot - then Ok (Betree_List_Nil, Betree_List_Cons (i, x) tl) - else - let* p = betree_ListPairU64T_partition_at_pivot t tl pivot in - let (ls0, ls1) = p in - Ok (Betree_List_Cons (i, x) ls0, ls1) - | Betree_List_Nil -> Ok (Betree_List_Nil, Betree_List_Nil) - end - -(** [betree_main::betree::{betree_main::betree::Leaf#3}::split]: - Source: 'src/betree.rs', lines 359:4-364:17 *) -let betree_Leaf_split - (self : betree_Leaf_t) (content : betree_List_t (u64 & u64)) - (params : betree_Params_t) (node_id_cnt : betree_NodeIdCounter_t) - (st : state) : - result (state & (betree_Internal_t & betree_NodeIdCounter_t)) - = - let* p = betree_List_split_at (u64 & u64) content params.split_size in - let (content0, content1) = p in - let* p1 = betree_List_hd (u64 & u64) content1 in - let (pivot, _) = p1 in - let* (id0, node_id_cnt1) = betree_NodeIdCounter_fresh_id node_id_cnt in - let* (id1, node_id_cnt2) = betree_NodeIdCounter_fresh_id node_id_cnt1 in - let* (st1, _) = betree_store_leaf_node id0 content0 st in - let* (st2, _) = betree_store_leaf_node id1 content1 st1 in - let n = Betree_Node_Leaf { id = id0; size = params.split_size } in - let n1 = Betree_Node_Leaf { id = id1; size = params.split_size } in - Ok (st2, ({ id = self.id; pivot = pivot; left = n; right = n1 }, - node_id_cnt2)) - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup_first_message_for_key]: - Source: 'src/betree.rs', lines 789:4-792:34 *) -let rec betree_Node_lookup_first_message_for_key - (key : u64) (msgs : betree_List_t (u64 & betree_Message_t)) : - Tot (result ((betree_List_t (u64 & betree_Message_t)) & (betree_List_t (u64 & - betree_Message_t) -> result (betree_List_t (u64 & betree_Message_t))))) - (decreases (betree_Node_lookup_first_message_for_key_decreases key msgs)) - = - begin match msgs with - | Betree_List_Cons x next_msgs -> - let (i, m) = x in - if i >= key - then Ok (Betree_List_Cons (i, m) next_msgs, Ok) - else - let* (l, lookup_first_message_for_key_back) = - betree_Node_lookup_first_message_for_key key next_msgs in - let back = - fun ret -> - let* next_msgs1 = lookup_first_message_for_key_back ret in - Ok (Betree_List_Cons (i, m) next_msgs1) in - Ok (l, back) - | Betree_List_Nil -> Ok (Betree_List_Nil, Ok) - end - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup_in_bindings]: - Source: 'src/betree.rs', lines 636:4-636:80 *) -let rec betree_Node_lookup_in_bindings - (key : u64) (bindings : betree_List_t (u64 & u64)) : - Tot (result (option u64)) - (decreases (betree_Node_lookup_in_bindings_decreases key bindings)) - = - begin match bindings with - | Betree_List_Cons hd tl -> - let (i, i1) = hd in - if i = key - then Ok (Some i1) - else if i > key then Ok None else betree_Node_lookup_in_bindings key tl - | Betree_List_Nil -> Ok None - end - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_upserts]: - Source: 'src/betree.rs', lines 819:4-819:90 *) -let rec betree_Node_apply_upserts - (msgs : betree_List_t (u64 & betree_Message_t)) (prev : option u64) - (key : u64) : - Tot (result (u64 & (betree_List_t (u64 & betree_Message_t)))) - (decreases (betree_Node_apply_upserts_decreases msgs prev key)) - = - let* b = betree_ListPairU64T_head_has_key betree_Message_t msgs key in - if b - then - let* (msg, msgs1) = betree_List_pop_front (u64 & betree_Message_t) msgs in - let (_, m) = msg in - begin match m with - | Betree_Message_Insert _ -> Fail Failure - | Betree_Message_Delete -> Fail Failure - | Betree_Message_Upsert s -> - let* v = betree_upsert_update prev s in - betree_Node_apply_upserts msgs1 (Some v) key - end - else - let* v = core_option_Option_unwrap u64 prev in - let* msgs1 = - betree_List_push_front (u64 & betree_Message_t) msgs (key, - Betree_Message_Insert v) in - Ok (v, msgs1) - -(** [betree_main::betree::{betree_main::betree::Internal#4}::lookup_in_children]: - Source: 'src/betree.rs', lines 395:4-395:63 *) -let rec betree_Internal_lookup_in_children - (self : betree_Internal_t) (key : u64) (st : state) : - Tot (result (state & ((option u64) & betree_Internal_t))) - (decreases (betree_Internal_lookup_in_children_decreases self key st)) - = - if key < self.pivot - then - let* (st1, (o, n)) = betree_Node_lookup self.left key st in - Ok (st1, (o, { self with left = n })) - else - let* (st1, (o, n)) = betree_Node_lookup self.right key st in - Ok (st1, (o, { self with right = n })) - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup]: - Source: 'src/betree.rs', lines 709:4-709:58 *) -and betree_Node_lookup - (self : betree_Node_t) (key : u64) (st : state) : - Tot (result (state & ((option u64) & betree_Node_t))) - (decreases (betree_Node_lookup_decreases self key st)) - = - begin match self with - | Betree_Node_Internal node -> - let* (st1, msgs) = betree_load_internal_node node.id st in - let* (pending, lookup_first_message_for_key_back) = - betree_Node_lookup_first_message_for_key key msgs in - begin match pending with - | Betree_List_Cons p l -> - let (k, msg) = p in - if k <> key - then - let* (st2, (o, node1)) = - betree_Internal_lookup_in_children node key st1 in - let* _ = - lookup_first_message_for_key_back (Betree_List_Cons (k, msg) l) in - Ok (st2, (o, Betree_Node_Internal node1)) - else - begin match msg with - | Betree_Message_Insert v -> - let* _ = - lookup_first_message_for_key_back (Betree_List_Cons (k, - Betree_Message_Insert v) l) in - Ok (st1, (Some v, Betree_Node_Internal node)) - | Betree_Message_Delete -> - let* _ = - lookup_first_message_for_key_back (Betree_List_Cons (k, - Betree_Message_Delete) l) in - Ok (st1, (None, Betree_Node_Internal node)) - | Betree_Message_Upsert ufs -> - let* (st2, (v, node1)) = - betree_Internal_lookup_in_children node key st1 in - let* (v1, pending1) = - betree_Node_apply_upserts (Betree_List_Cons (k, - Betree_Message_Upsert ufs) l) v key in - let* msgs1 = lookup_first_message_for_key_back pending1 in - let* (st3, _) = betree_store_internal_node node1.id msgs1 st2 in - Ok (st3, (Some v1, Betree_Node_Internal node1)) - end - | Betree_List_Nil -> - let* (st2, (o, node1)) = betree_Internal_lookup_in_children node key st1 - in - let* _ = lookup_first_message_for_key_back Betree_List_Nil in - Ok (st2, (o, Betree_Node_Internal node1)) - end - | Betree_Node_Leaf node -> - let* (st1, bindings) = betree_load_leaf_node node.id st in - let* o = betree_Node_lookup_in_bindings key bindings in - Ok (st1, (o, Betree_Node_Leaf node)) - end - -(** [betree_main::betree::{betree_main::betree::Node#5}::filter_messages_for_key]: - Source: 'src/betree.rs', lines 674:4-674:77 *) -let rec betree_Node_filter_messages_for_key - (key : u64) (msgs : betree_List_t (u64 & betree_Message_t)) : - Tot (result (betree_List_t (u64 & betree_Message_t))) - (decreases (betree_Node_filter_messages_for_key_decreases key msgs)) - = - begin match msgs with - | Betree_List_Cons p l -> - let (k, m) = p in - if k = key - then - let* (_, msgs1) = - betree_List_pop_front (u64 & betree_Message_t) (Betree_List_Cons (k, m) - l) in - betree_Node_filter_messages_for_key key msgs1 - else Ok (Betree_List_Cons (k, m) l) - | Betree_List_Nil -> Ok Betree_List_Nil - end - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup_first_message_after_key]: - Source: 'src/betree.rs', lines 689:4-692:34 *) -let rec betree_Node_lookup_first_message_after_key - (key : u64) (msgs : betree_List_t (u64 & betree_Message_t)) : - Tot (result ((betree_List_t (u64 & betree_Message_t)) & (betree_List_t (u64 & - betree_Message_t) -> result (betree_List_t (u64 & betree_Message_t))))) - (decreases (betree_Node_lookup_first_message_after_key_decreases key msgs)) - = - begin match msgs with - | Betree_List_Cons p next_msgs -> - let (k, m) = p in - if k = key - then - let* (l, lookup_first_message_after_key_back) = - betree_Node_lookup_first_message_after_key key next_msgs in - let back = - fun ret -> - let* next_msgs1 = lookup_first_message_after_key_back ret in - Ok (Betree_List_Cons (k, m) next_msgs1) in - Ok (l, back) - else Ok (Betree_List_Cons (k, m) next_msgs, Ok) - | Betree_List_Nil -> Ok (Betree_List_Nil, Ok) - end - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_to_internal]: - Source: 'src/betree.rs', lines 521:4-521:89 *) -let betree_Node_apply_to_internal - (msgs : betree_List_t (u64 & betree_Message_t)) (key : u64) - (new_msg : betree_Message_t) : - result (betree_List_t (u64 & betree_Message_t)) - = - let* (msgs1, lookup_first_message_for_key_back) = - betree_Node_lookup_first_message_for_key key msgs in - let* b = betree_ListPairU64T_head_has_key betree_Message_t msgs1 key in - if b - then - begin match new_msg with - | Betree_Message_Insert i -> - let* msgs2 = betree_Node_filter_messages_for_key key msgs1 in - let* msgs3 = - betree_List_push_front (u64 & betree_Message_t) msgs2 (key, - Betree_Message_Insert i) in - lookup_first_message_for_key_back msgs3 - | Betree_Message_Delete -> - let* msgs2 = betree_Node_filter_messages_for_key key msgs1 in - let* msgs3 = - betree_List_push_front (u64 & betree_Message_t) msgs2 (key, - Betree_Message_Delete) in - lookup_first_message_for_key_back msgs3 - | Betree_Message_Upsert s -> - let* p = betree_List_hd (u64 & betree_Message_t) msgs1 in - let (_, m) = p in - begin match m with - | Betree_Message_Insert prev -> - let* v = betree_upsert_update (Some prev) s in - let* (_, msgs2) = betree_List_pop_front (u64 & betree_Message_t) msgs1 - in - let* msgs3 = - betree_List_push_front (u64 & betree_Message_t) msgs2 (key, - Betree_Message_Insert v) in - lookup_first_message_for_key_back msgs3 - | Betree_Message_Delete -> - let* (_, msgs2) = betree_List_pop_front (u64 & betree_Message_t) msgs1 - in - let* v = betree_upsert_update None s in - let* msgs3 = - betree_List_push_front (u64 & betree_Message_t) msgs2 (key, - Betree_Message_Insert v) in - lookup_first_message_for_key_back msgs3 - | Betree_Message_Upsert _ -> - let* (msgs2, lookup_first_message_after_key_back) = - betree_Node_lookup_first_message_after_key key msgs1 in - let* msgs3 = - betree_List_push_front (u64 & betree_Message_t) msgs2 (key, - Betree_Message_Upsert s) in - let* msgs4 = lookup_first_message_after_key_back msgs3 in - lookup_first_message_for_key_back msgs4 - end - end - else - let* msgs2 = - betree_List_push_front (u64 & betree_Message_t) msgs1 (key, new_msg) in - lookup_first_message_for_key_back msgs2 - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_messages_to_internal]: - Source: 'src/betree.rs', lines 502:4-505:5 *) -let rec betree_Node_apply_messages_to_internal - (msgs : betree_List_t (u64 & betree_Message_t)) - (new_msgs : betree_List_t (u64 & betree_Message_t)) : - Tot (result (betree_List_t (u64 & betree_Message_t))) - (decreases (betree_Node_apply_messages_to_internal_decreases msgs new_msgs)) - = - begin match new_msgs with - | Betree_List_Cons new_msg new_msgs_tl -> - let (i, m) = new_msg in - let* msgs1 = betree_Node_apply_to_internal msgs i m in - betree_Node_apply_messages_to_internal msgs1 new_msgs_tl - | Betree_List_Nil -> Ok msgs - end - -(** [betree_main::betree::{betree_main::betree::Node#5}::lookup_mut_in_bindings]: - Source: 'src/betree.rs', lines 653:4-656:32 *) -let rec betree_Node_lookup_mut_in_bindings - (key : u64) (bindings : betree_List_t (u64 & u64)) : - Tot (result ((betree_List_t (u64 & u64)) & (betree_List_t (u64 & u64) -> - result (betree_List_t (u64 & u64))))) - (decreases (betree_Node_lookup_mut_in_bindings_decreases key bindings)) - = - begin match bindings with - | Betree_List_Cons hd tl -> - let (i, i1) = hd in - if i >= key - then Ok (Betree_List_Cons (i, i1) tl, Ok) - else - let* (l, lookup_mut_in_bindings_back) = - betree_Node_lookup_mut_in_bindings key tl in - let back = - fun ret -> - let* tl1 = lookup_mut_in_bindings_back ret in - Ok (Betree_List_Cons (i, i1) tl1) in - Ok (l, back) - | Betree_List_Nil -> Ok (Betree_List_Nil, Ok) - end - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_to_leaf]: - Source: 'src/betree.rs', lines 460:4-460:87 *) -let betree_Node_apply_to_leaf - (bindings : betree_List_t (u64 & u64)) (key : u64) - (new_msg : betree_Message_t) : - result (betree_List_t (u64 & u64)) - = - let* (bindings1, lookup_mut_in_bindings_back) = - betree_Node_lookup_mut_in_bindings key bindings in - let* b = betree_ListPairU64T_head_has_key u64 bindings1 key in - if b - then - let* (hd, bindings2) = betree_List_pop_front (u64 & u64) bindings1 in - begin match new_msg with - | Betree_Message_Insert v -> - let* bindings3 = betree_List_push_front (u64 & u64) bindings2 (key, v) in - lookup_mut_in_bindings_back bindings3 - | Betree_Message_Delete -> lookup_mut_in_bindings_back bindings2 - | Betree_Message_Upsert s -> - let (_, i) = hd in - let* v = betree_upsert_update (Some i) s in - let* bindings3 = betree_List_push_front (u64 & u64) bindings2 (key, v) in - lookup_mut_in_bindings_back bindings3 - end - else - begin match new_msg with - | Betree_Message_Insert v -> - let* bindings2 = betree_List_push_front (u64 & u64) bindings1 (key, v) in - lookup_mut_in_bindings_back bindings2 - | Betree_Message_Delete -> lookup_mut_in_bindings_back bindings1 - | Betree_Message_Upsert s -> - let* v = betree_upsert_update None s in - let* bindings2 = betree_List_push_front (u64 & u64) bindings1 (key, v) in - lookup_mut_in_bindings_back bindings2 - end - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_messages_to_leaf]: - Source: 'src/betree.rs', lines 444:4-447:5 *) -let rec betree_Node_apply_messages_to_leaf - (bindings : betree_List_t (u64 & u64)) - (new_msgs : betree_List_t (u64 & betree_Message_t)) : - Tot (result (betree_List_t (u64 & u64))) - (decreases (betree_Node_apply_messages_to_leaf_decreases bindings new_msgs)) - = - begin match new_msgs with - | Betree_List_Cons new_msg new_msgs_tl -> - let (i, m) = new_msg in - let* bindings1 = betree_Node_apply_to_leaf bindings i m in - betree_Node_apply_messages_to_leaf bindings1 new_msgs_tl - | Betree_List_Nil -> Ok bindings - end - -(** [betree_main::betree::{betree_main::betree::Internal#4}::flush]: - Source: 'src/betree.rs', lines 410:4-415:26 *) -let rec betree_Internal_flush - (self : betree_Internal_t) (params : betree_Params_t) - (node_id_cnt : betree_NodeIdCounter_t) - (content : betree_List_t (u64 & betree_Message_t)) (st : state) : - Tot (result (state & ((betree_List_t (u64 & betree_Message_t)) & - (betree_Internal_t & betree_NodeIdCounter_t)))) - (decreases ( - betree_Internal_flush_decreases self params node_id_cnt content st)) - = - let* p = - betree_ListPairU64T_partition_at_pivot betree_Message_t content self.pivot - in - let (msgs_left, msgs_right) = p in - let* len_left = betree_List_len (u64 & betree_Message_t) msgs_left in - if len_left >= params.min_flush_size - then - let* (st1, p1) = - betree_Node_apply_messages self.left params node_id_cnt msgs_left st in - let (n, node_id_cnt1) = p1 in - let* len_right = betree_List_len (u64 & betree_Message_t) msgs_right in - if len_right >= params.min_flush_size - then - let* (st2, p2) = - betree_Node_apply_messages self.right params node_id_cnt1 msgs_right - st1 in - let (n1, node_id_cnt2) = p2 in - Ok (st2, (Betree_List_Nil, ({ self with left = n; right = n1 }, - node_id_cnt2))) - else Ok (st1, (msgs_right, ({ self with left = n }, node_id_cnt1))) - else - let* (st1, p1) = - betree_Node_apply_messages self.right params node_id_cnt msgs_right st in - let (n, node_id_cnt1) = p1 in - Ok (st1, (msgs_left, ({ self with right = n }, node_id_cnt1))) - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply_messages]: - Source: 'src/betree.rs', lines 588:4-593:5 *) -and betree_Node_apply_messages - (self : betree_Node_t) (params : betree_Params_t) - (node_id_cnt : betree_NodeIdCounter_t) - (msgs : betree_List_t (u64 & betree_Message_t)) (st : state) : - Tot (result (state & (betree_Node_t & betree_NodeIdCounter_t))) - (decreases ( - betree_Node_apply_messages_decreases self params node_id_cnt msgs st)) - = - begin match self with - | Betree_Node_Internal node -> - let* (st1, content) = betree_load_internal_node node.id st in - let* content1 = betree_Node_apply_messages_to_internal content msgs in - let* num_msgs = betree_List_len (u64 & betree_Message_t) content1 in - if num_msgs >= params.min_flush_size - then - let* (st2, (content2, p)) = - betree_Internal_flush node params node_id_cnt content1 st1 in - let (node1, node_id_cnt1) = p in - let* (st3, _) = betree_store_internal_node node1.id content2 st2 in - Ok (st3, (Betree_Node_Internal node1, node_id_cnt1)) - else - let* (st2, _) = betree_store_internal_node node.id content1 st1 in - Ok (st2, (Betree_Node_Internal node, node_id_cnt)) - | Betree_Node_Leaf node -> - let* (st1, content) = betree_load_leaf_node node.id st in - let* content1 = betree_Node_apply_messages_to_leaf content msgs in - let* len = betree_List_len (u64 & u64) content1 in - let* i = u64_mul 2 params.split_size in - if len >= i - then - let* (st2, (new_node, node_id_cnt1)) = - betree_Leaf_split node content1 params node_id_cnt st1 in - let* (st3, _) = betree_store_leaf_node node.id Betree_List_Nil st2 in - Ok (st3, (Betree_Node_Internal new_node, node_id_cnt1)) - else - let* (st2, _) = betree_store_leaf_node node.id content1 st1 in - Ok (st2, (Betree_Node_Leaf { node with size = len }, node_id_cnt)) - end - -(** [betree_main::betree::{betree_main::betree::Node#5}::apply]: - Source: 'src/betree.rs', lines 576:4-582:5 *) -let betree_Node_apply - (self : betree_Node_t) (params : betree_Params_t) - (node_id_cnt : betree_NodeIdCounter_t) (key : u64) - (new_msg : betree_Message_t) (st : state) : - result (state & (betree_Node_t & betree_NodeIdCounter_t)) - = - let* (st1, p) = - betree_Node_apply_messages self params node_id_cnt (Betree_List_Cons (key, - new_msg) Betree_List_Nil) st in - let (self1, node_id_cnt1) = p in - Ok (st1, (self1, node_id_cnt1)) - -(** [betree_main::betree::{betree_main::betree::BeTree#6}::new]: - Source: 'src/betree.rs', lines 849:4-849:60 *) -let betree_BeTree_new - (min_flush_size : u64) (split_size : u64) (st : state) : - result (state & betree_BeTree_t) - = - let* node_id_cnt = betree_NodeIdCounter_new in - let* (id, node_id_cnt1) = betree_NodeIdCounter_fresh_id node_id_cnt in - let* (st1, _) = betree_store_leaf_node id Betree_List_Nil st in - Ok (st1, - { - params = { min_flush_size = min_flush_size; split_size = split_size }; - node_id_cnt = node_id_cnt1; - root = (Betree_Node_Leaf { id = id; size = 0 }) - }) - -(** [betree_main::betree::{betree_main::betree::BeTree#6}::apply]: - Source: 'src/betree.rs', lines 868:4-868:47 *) -let betree_BeTree_apply - (self : betree_BeTree_t) (key : u64) (msg : betree_Message_t) (st : state) : - result (state & betree_BeTree_t) - = - let* (st1, p) = - betree_Node_apply self.root self.params self.node_id_cnt key msg st in - let (n, nic) = p in - Ok (st1, { self with node_id_cnt = nic; root = n }) - -(** [betree_main::betree::{betree_main::betree::BeTree#6}::insert]: - Source: 'src/betree.rs', lines 874:4-874:52 *) -let betree_BeTree_insert - (self : betree_BeTree_t) (key : u64) (value : u64) (st : state) : - result (state & betree_BeTree_t) - = - betree_BeTree_apply self key (Betree_Message_Insert value) st - -(** [betree_main::betree::{betree_main::betree::BeTree#6}::delete]: - Source: 'src/betree.rs', lines 880:4-880:38 *) -let betree_BeTree_delete - (self : betree_BeTree_t) (key : u64) (st : state) : - result (state & betree_BeTree_t) - = - betree_BeTree_apply self key Betree_Message_Delete st - -(** [betree_main::betree::{betree_main::betree::BeTree#6}::upsert]: - Source: 'src/betree.rs', lines 886:4-886:59 *) -let betree_BeTree_upsert - (self : betree_BeTree_t) (key : u64) (upd : betree_UpsertFunState_t) - (st : state) : - result (state & betree_BeTree_t) - = - betree_BeTree_apply self key (Betree_Message_Upsert upd) st - -(** [betree_main::betree::{betree_main::betree::BeTree#6}::lookup]: - Source: 'src/betree.rs', lines 895:4-895:62 *) -let betree_BeTree_lookup - (self : betree_BeTree_t) (key : u64) (st : state) : - result (state & ((option u64) & betree_BeTree_t)) - = - let* (st1, (o, n)) = betree_Node_lookup self.root key st in - Ok (st1, (o, { self with root = n })) - -(** [betree_main::main]: - Source: 'src/main.rs', lines 4:0-4:9 *) -let main : result unit = - Ok () - -(** Unit test for [betree_main::main] *) -let _ = assert_norm (main = Ok ()) - diff --git a/tests/fstar/betree/BetreeMain.FunsExternal.fsti b/tests/fstar/betree/BetreeMain.FunsExternal.fsti deleted file mode 100644 index 8be98acf..00000000 --- a/tests/fstar/betree/BetreeMain.FunsExternal.fsti +++ /dev/null @@ -1,30 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [betree_main]: external function declarations *) -module BetreeMain.FunsExternal -open Primitives -include BetreeMain.Types - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** [betree_main::betree_utils::load_internal_node]: - Source: 'src/betree_utils.rs', lines 98:0-98:63 *) -val betree_utils_load_internal_node - : u64 -> state -> result (state & (betree_List_t (u64 & betree_Message_t))) - -(** [betree_main::betree_utils::store_internal_node]: - Source: 'src/betree_utils.rs', lines 115:0-115:71 *) -val betree_utils_store_internal_node - : - u64 -> betree_List_t (u64 & betree_Message_t) -> state -> result (state & - unit) - -(** [betree_main::betree_utils::load_leaf_node]: - Source: 'src/betree_utils.rs', lines 132:0-132:55 *) -val betree_utils_load_leaf_node - : u64 -> state -> result (state & (betree_List_t (u64 & u64))) - -(** [betree_main::betree_utils::store_leaf_node]: - Source: 'src/betree_utils.rs', lines 145:0-145:63 *) -val betree_utils_store_leaf_node - : u64 -> betree_List_t (u64 & u64) -> state -> result (state & unit) - diff --git a/tests/fstar/betree/BetreeMain.Types.fst b/tests/fstar/betree/BetreeMain.Types.fst deleted file mode 100644 index b87219b2..00000000 --- a/tests/fstar/betree/BetreeMain.Types.fst +++ /dev/null @@ -1,61 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [betree_main]: type definitions *) -module BetreeMain.Types -open Primitives -include BetreeMain.TypesExternal - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** [betree_main::betree::List] - Source: 'src/betree.rs', lines 17:0-17:23 *) -type betree_List_t (t : Type0) = -| Betree_List_Cons : t -> betree_List_t t -> betree_List_t t -| Betree_List_Nil : betree_List_t t - -(** [betree_main::betree::UpsertFunState] - Source: 'src/betree.rs', lines 63:0-63:23 *) -type betree_UpsertFunState_t = -| Betree_UpsertFunState_Add : u64 -> betree_UpsertFunState_t -| Betree_UpsertFunState_Sub : u64 -> betree_UpsertFunState_t - -(** [betree_main::betree::Message] - Source: 'src/betree.rs', lines 69:0-69:23 *) -type betree_Message_t = -| Betree_Message_Insert : u64 -> betree_Message_t -| Betree_Message_Delete : betree_Message_t -| Betree_Message_Upsert : betree_UpsertFunState_t -> betree_Message_t - -(** [betree_main::betree::Leaf] - Source: 'src/betree.rs', lines 167:0-167:11 *) -type betree_Leaf_t = { id : u64; size : u64; } - -(** [betree_main::betree::Internal] - Source: 'src/betree.rs', lines 156:0-156:15 *) -type betree_Internal_t = -{ - id : u64; pivot : u64; left : betree_Node_t; right : betree_Node_t; -} - -(** [betree_main::betree::Node] - Source: 'src/betree.rs', lines 179:0-179:9 *) -and betree_Node_t = -| Betree_Node_Internal : betree_Internal_t -> betree_Node_t -| Betree_Node_Leaf : betree_Leaf_t -> betree_Node_t - -(** [betree_main::betree::Params] - Source: 'src/betree.rs', lines 187:0-187:13 *) -type betree_Params_t = { min_flush_size : u64; split_size : u64; } - -(** [betree_main::betree::NodeIdCounter] - Source: 'src/betree.rs', lines 201:0-201:20 *) -type betree_NodeIdCounter_t = { next_node_id : u64; } - -(** [betree_main::betree::BeTree] - Source: 'src/betree.rs', lines 218:0-218:17 *) -type betree_BeTree_t = -{ - params : betree_Params_t; - node_id_cnt : betree_NodeIdCounter_t; - root : betree_Node_t; -} - diff --git a/tests/fstar/betree/BetreeMain.TypesExternal.fsti b/tests/fstar/betree/BetreeMain.TypesExternal.fsti deleted file mode 100644 index 1b2c53a6..00000000 --- a/tests/fstar/betree/BetreeMain.TypesExternal.fsti +++ /dev/null @@ -1,10 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [betree_main]: external type declarations *) -module BetreeMain.TypesExternal -open Primitives - -#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" - -(** The state type used in the state-error monad *) -val state : Type0 - diff --git a/tests/hol4/betree/betreeMain_FunsScript.sml b/tests/hol4/betree/betreeMain_FunsScript.sml deleted file mode 100644 index bd16c16c..00000000 --- a/tests/hol4/betree/betreeMain_FunsScript.sml +++ /dev/null @@ -1,1212 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [betree_main]: function definitions *) -open primitivesLib divDefLib -open betreeMain_TypesTheory betreeMain_OpaqueTheory - -val _ = new_theory "betreeMain_Funs" - - -val betree_load_internal_node_fwd_def = Define ‘ - (** [betree_main::betree::load_internal_node]: forward function *) - betree_load_internal_node_fwd - (id : u64) (st : state) : - (state # (u64 # betree_message_t) betree_list_t) result - = - betree_utils_load_internal_node_fwd id st -’ - -val betree_store_internal_node_fwd_def = Define ‘ - (** [betree_main::betree::store_internal_node]: forward function *) - betree_store_internal_node_fwd - (id : u64) (content : (u64 # betree_message_t) betree_list_t) (st : state) - : - (state # unit) result - = - do - (st0, _) <- betree_utils_store_internal_node_fwd id content st; - Return (st0, ()) - od -’ - -val betree_load_leaf_node_fwd_def = Define ‘ - (** [betree_main::betree::load_leaf_node]: forward function *) - betree_load_leaf_node_fwd - (id : u64) (st : state) : (state # (u64 # u64) betree_list_t) result = - betree_utils_load_leaf_node_fwd id st -’ - -val betree_store_leaf_node_fwd_def = Define ‘ - (** [betree_main::betree::store_leaf_node]: forward function *) - betree_store_leaf_node_fwd - (id : u64) (content : (u64 # u64) betree_list_t) (st : state) : - (state # unit) result - = - do - (st0, _) <- betree_utils_store_leaf_node_fwd id content st; - Return (st0, ()) - od -’ - -val betree_fresh_node_id_fwd_def = Define ‘ - (** [betree_main::betree::fresh_node_id]: forward function *) - betree_fresh_node_id_fwd (counter : u64) : u64 result = - do - _ <- u64_add counter (int_to_u64 1); - Return counter - od -’ - -val betree_fresh_node_id_back_def = Define ‘ - (** [betree_main::betree::fresh_node_id]: backward function 0 *) - betree_fresh_node_id_back (counter : u64) : u64 result = - u64_add counter (int_to_u64 1) -’ - -val betree_node_id_counter_new_fwd_def = Define ‘ - (** [betree_main::betree::NodeIdCounter::{0}::new]: forward function *) - betree_node_id_counter_new_fwd : betree_node_id_counter_t result = - Return (<| betree_node_id_counter_next_node_id := (int_to_u64 0) |>) -’ - -val betree_node_id_counter_fresh_id_fwd_def = Define ‘ - (** [betree_main::betree::NodeIdCounter::{0}::fresh_id]: forward function *) - betree_node_id_counter_fresh_id_fwd - (self : betree_node_id_counter_t) : u64 result = - do - _ <- u64_add self.betree_node_id_counter_next_node_id (int_to_u64 1); - Return self.betree_node_id_counter_next_node_id - od -’ - -val betree_node_id_counter_fresh_id_back_def = Define ‘ - (** [betree_main::betree::NodeIdCounter::{0}::fresh_id]: backward function 0 *) - betree_node_id_counter_fresh_id_back - (self : betree_node_id_counter_t) : betree_node_id_counter_t result = - do - i <- u64_add self.betree_node_id_counter_next_node_id (int_to_u64 1); - Return (<| betree_node_id_counter_next_node_id := i |>) - od -’ - -val betree_upsert_update_fwd_def = Define ‘ - (** [betree_main::betree::upsert_update]: forward function *) - betree_upsert_update_fwd - (prev : u64 option) (st : betree_upsert_fun_state_t) : u64 result = - (case prev of - | NONE => - (case st of - | BetreeUpsertFunStateAdd v => Return v - | BetreeUpsertFunStateSub i => Return (int_to_u64 0)) - | SOME prev0 => - (case st of - | BetreeUpsertFunStateAdd v => - do - margin <- u64_sub core_u64_max prev0; - if u64_ge margin v then u64_add prev0 v else Return core_u64_max - od - | BetreeUpsertFunStateSub v => - if u64_ge prev0 v then u64_sub prev0 v else Return (int_to_u64 0))) -’ - -val [betree_list_len_fwd_def] = DefineDiv ‘ - (** [betree_main::betree::List::{1}::len]: forward function *) - betree_list_len_fwd (self : 't betree_list_t) : u64 result = - (case self of - | BetreeListCons t tl => - do - i <- betree_list_len_fwd tl; - u64_add (int_to_u64 1) i - od - | BetreeListNil => Return (int_to_u64 0)) -’ - -val [betree_list_split_at_fwd_def] = DefineDiv ‘ - (** [betree_main::betree::List::{1}::split_at]: forward function *) - betree_list_split_at_fwd - (self : 't betree_list_t) (n : u64) : - ('t betree_list_t # 't betree_list_t) result - = - if n = int_to_u64 0 - then Return (BetreeListNil, self) - else - (case self of - | BetreeListCons hd tl => - do - i <- u64_sub n (int_to_u64 1); - p <- betree_list_split_at_fwd tl i; - let (ls0, ls1) = p in let l = ls0 in Return (BetreeListCons hd l, ls1) - od - | BetreeListNil => Fail Failure) -’ - -val betree_list_push_front_fwd_back_def = Define ‘ - (** [betree_main::betree::List::{1}::push_front]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - betree_list_push_front_fwd_back - (self : 't betree_list_t) (x : 't) : 't betree_list_t result = - let tl = mem_replace_fwd self BetreeListNil in - let l = tl in - Return (BetreeListCons x l) -’ - -val betree_list_pop_front_fwd_def = Define ‘ - (** [betree_main::betree::List::{1}::pop_front]: forward function *) - betree_list_pop_front_fwd (self : 't betree_list_t) : 't result = - let ls = mem_replace_fwd self BetreeListNil in - (case ls of - | BetreeListCons x tl => Return x - | BetreeListNil => Fail Failure) -’ - -val betree_list_pop_front_back_def = Define ‘ - (** [betree_main::betree::List::{1}::pop_front]: backward function 0 *) - betree_list_pop_front_back - (self : 't betree_list_t) : 't betree_list_t result = - let ls = mem_replace_fwd self BetreeListNil in - (case ls of - | BetreeListCons x tl => Return tl - | BetreeListNil => Fail Failure) -’ - -val betree_list_hd_fwd_def = Define ‘ - (** [betree_main::betree::List::{1}::hd]: forward function *) - betree_list_hd_fwd (self : 't betree_list_t) : 't result = - (case self of - | BetreeListCons hd l => Return hd - | BetreeListNil => Fail Failure) -’ - -val betree_list_head_has_key_fwd_def = Define ‘ - (** [betree_main::betree::List::{2}::head_has_key]: forward function *) - betree_list_head_has_key_fwd - (self : (u64 # 't) betree_list_t) (key : u64) : bool result = - (case self of - | BetreeListCons hd l => let (i, _) = hd in Return (i = key) - | BetreeListNil => Return F) -’ - -val [betree_list_partition_at_pivot_fwd_def] = DefineDiv ‘ - (** [betree_main::betree::List::{2}::partition_at_pivot]: forward function *) - betree_list_partition_at_pivot_fwd - (self : (u64 # 't) betree_list_t) (pivot : u64) : - ((u64 # 't) betree_list_t # (u64 # 't) betree_list_t) result - = - (case self of - | BetreeListCons hd tl => - let (i, t) = hd in - if u64_ge i pivot - then Return (BetreeListNil, BetreeListCons (i, t) tl) - else ( - do - p <- betree_list_partition_at_pivot_fwd tl pivot; - let (ls0, ls1) = p in - let l = ls0 in - Return (BetreeListCons (i, t) l, ls1) - od) - | BetreeListNil => Return (BetreeListNil, BetreeListNil)) -’ - -val betree_leaf_split_fwd_def = Define ‘ - (** [betree_main::betree::Leaf::{3}::split]: forward function *) - betree_leaf_split_fwd - (self : betree_leaf_t) (content : (u64 # u64) betree_list_t) - (params : betree_params_t) (node_id_cnt : betree_node_id_counter_t) - (st : state) : - (state # betree_internal_t) result - = - do - p <- betree_list_split_at_fwd content params.betree_params_split_size; - let (content0, content1) = p in - do - p0 <- betree_list_hd_fwd content1; - let (pivot, _) = p0 in - do - id0 <- betree_node_id_counter_fresh_id_fwd node_id_cnt; - node_id_cnt0 <- betree_node_id_counter_fresh_id_back node_id_cnt; - id1 <- betree_node_id_counter_fresh_id_fwd node_id_cnt0; - (st0, _) <- betree_store_leaf_node_fwd id0 content0 st; - (st1, _) <- betree_store_leaf_node_fwd id1 content1 st0; - let n = BetreeNodeLeaf - (<| - betree_leaf_id := id0; - betree_leaf_size := params.betree_params_split_size - |>) in - let n0 = BetreeNodeLeaf - (<| - betree_leaf_id := id1; - betree_leaf_size := params.betree_params_split_size - |>) in - Return (st1, - <| - betree_internal_id := self.betree_leaf_id; - betree_internal_pivot := pivot; - betree_internal_left := n; - betree_internal_right := n0 - |>) - od - od - od -’ - -val betree_leaf_split_back_def = Define ‘ - (** [betree_main::betree::Leaf::{3}::split]: backward function 2 *) - betree_leaf_split_back - (self : betree_leaf_t) (content : (u64 # u64) betree_list_t) - (params : betree_params_t) (node_id_cnt : betree_node_id_counter_t) - (st : state) : - betree_node_id_counter_t result - = - do - p <- betree_list_split_at_fwd content params.betree_params_split_size; - let (content0, content1) = p in - do - _ <- betree_list_hd_fwd content1; - id0 <- betree_node_id_counter_fresh_id_fwd node_id_cnt; - node_id_cnt0 <- betree_node_id_counter_fresh_id_back node_id_cnt; - id1 <- betree_node_id_counter_fresh_id_fwd node_id_cnt0; - (st0, _) <- betree_store_leaf_node_fwd id0 content0 st; - _ <- betree_store_leaf_node_fwd id1 content1 st0; - betree_node_id_counter_fresh_id_back node_id_cnt0 - od - od -’ - -val [betree_node_lookup_first_message_for_key_fwd_def] = DefineDiv ‘ - (** [betree_main::betree::Node::{5}::lookup_first_message_for_key]: forward function *) - betree_node_lookup_first_message_for_key_fwd - (key : u64) (msgs : (u64 # betree_message_t) betree_list_t) : - (u64 # betree_message_t) betree_list_t result - = - (case msgs of - | BetreeListCons x next_msgs => - let (i, m) = x in - if u64_ge i key - then Return (BetreeListCons (i, m) next_msgs) - else betree_node_lookup_first_message_for_key_fwd key next_msgs - | BetreeListNil => Return BetreeListNil) -’ - -val [betree_node_lookup_first_message_for_key_back_def] = DefineDiv ‘ - (** [betree_main::betree::Node::{5}::lookup_first_message_for_key]: backward function 0 *) - betree_node_lookup_first_message_for_key_back - (key : u64) (msgs : (u64 # betree_message_t) betree_list_t) - (ret : (u64 # betree_message_t) betree_list_t) : - (u64 # betree_message_t) betree_list_t result - = - (case msgs of - | BetreeListCons x next_msgs => - let (i, m) = x in - if u64_ge i key - then Return ret - else ( - do - next_msgs0 <- - betree_node_lookup_first_message_for_key_back key next_msgs ret; - Return (BetreeListCons (i, m) next_msgs0) - od) - | BetreeListNil => Return ret) -’ - -val [betree_node_apply_upserts_fwd_def] = DefineDiv ‘ - (** [betree_main::betree::Node::{5}::apply_upserts]: forward function *) - betree_node_apply_upserts_fwd - (msgs : (u64 # betree_message_t) betree_list_t) (prev : u64 option) - (key : u64) (st : state) : - (state # u64) result - = - do - b <- betree_list_head_has_key_fwd msgs key; - if b - then ( - do - msg <- betree_list_pop_front_fwd msgs; - let (_, m) = msg in - (case m of - | BetreeMessageInsert i => Fail Failure - | BetreeMessageDelete => Fail Failure - | BetreeMessageUpsert s => - do - v <- betree_upsert_update_fwd prev s; - msgs0 <- betree_list_pop_front_back msgs; - betree_node_apply_upserts_fwd msgs0 (SOME v) key st - od) - od) - else ( - do - (st0, v) <- core_option_option_unwrap_fwd prev st; - _ <- betree_list_push_front_fwd_back msgs (key, BetreeMessageInsert v); - Return (st0, v) - od) - od -’ - -val [betree_node_apply_upserts_back_def] = DefineDiv ‘ - (** [betree_main::betree::Node::{5}::apply_upserts]: backward function 0 *) - betree_node_apply_upserts_back - (msgs : (u64 # betree_message_t) betree_list_t) (prev : u64 option) - (key : u64) (st : state) : - (u64 # betree_message_t) betree_list_t result - = - do - b <- betree_list_head_has_key_fwd msgs key; - if b - then ( - do - msg <- betree_list_pop_front_fwd msgs; - let (_, m) = msg in - (case m of - | BetreeMessageInsert i => Fail Failure - | BetreeMessageDelete => Fail Failure - | BetreeMessageUpsert s => - do - v <- betree_upsert_update_fwd prev s; - msgs0 <- betree_list_pop_front_back msgs; - betree_node_apply_upserts_back msgs0 (SOME v) key st - od) - od) - else ( - do - (_, v) <- core_option_option_unwrap_fwd prev st; - betree_list_push_front_fwd_back msgs (key, BetreeMessageInsert v) - od) - od -’ - -val [betree_node_lookup_in_bindings_fwd_def] = DefineDiv ‘ - (** [betree_main::betree::Node::{5}::lookup_in_bindings]: forward function *) - betree_node_lookup_in_bindings_fwd - (key : u64) (bindings : (u64 # u64) betree_list_t) : u64 option result = - (case bindings of - | BetreeListCons hd tl => - let (i, i0) = hd in - if i = key - then Return (SOME i0) - else - if u64_gt i key - then Return NONE - else betree_node_lookup_in_bindings_fwd key tl - | BetreeListNil => Return NONE) -’ - -val [betree_internal_lookup_in_children_fwd_def, betree_internal_lookup_in_children_back_def, betree_node_lookup_fwd_def, betree_node_lookup_back_def] = DefineDiv ‘ - (** [betree_main::betree::Internal::{4}::lookup_in_children]: forward function *) - (betree_internal_lookup_in_children_fwd - (self : betree_internal_t) (key : u64) (st : state) : - (state # u64 option) result - = - if u64_lt key self.betree_internal_pivot - then betree_node_lookup_fwd self.betree_internal_left key st - else betree_node_lookup_fwd self.betree_internal_right key st) - /\ - - (** [betree_main::betree::Internal::{4}::lookup_in_children]: backward function 0 *) - (betree_internal_lookup_in_children_back - (self : betree_internal_t) (key : u64) (st : state) : - betree_internal_t result - = - if u64_lt key self.betree_internal_pivot - then ( - do - n <- betree_node_lookup_back self.betree_internal_left key st; - Return (self with <| betree_internal_left := n |>) - od) - else ( - do - n <- betree_node_lookup_back self.betree_internal_right key st; - Return (self with <| betree_internal_right := n |>) - od)) - /\ - - (** [betree_main::betree::Node::{5}::lookup]: forward function *) - (betree_node_lookup_fwd - (self : betree_node_t) (key : u64) (st : state) : - (state # u64 option) result - = - (case self of - | BetreeNodeInternal node => - do - (st0, msgs) <- betree_load_internal_node_fwd node.betree_internal_id st; - pending <- betree_node_lookup_first_message_for_key_fwd key msgs; - (case pending of - | BetreeListCons p l => - let (k, msg) = p in - if k <> key - then ( - do - (st1, opt) <- betree_internal_lookup_in_children_fwd node key st0; - _ <- - betree_node_lookup_first_message_for_key_back key msgs - (BetreeListCons (k, msg) l); - Return (st1, opt) - od) - else - (case msg of - | BetreeMessageInsert v => - do - _ <- - betree_node_lookup_first_message_for_key_back key msgs - (BetreeListCons (k, BetreeMessageInsert v) l); - Return (st0, SOME v) - od - | BetreeMessageDelete => - do - _ <- - betree_node_lookup_first_message_for_key_back key msgs - (BetreeListCons (k, BetreeMessageDelete) l); - Return (st0, NONE) - od - | BetreeMessageUpsert ufs => - do - (st1, v) <- betree_internal_lookup_in_children_fwd node key st0; - (st2, v0) <- - betree_node_apply_upserts_fwd (BetreeListCons (k, - BetreeMessageUpsert ufs) l) v key st1; - node0 <- betree_internal_lookup_in_children_back node key st0; - pending0 <- - betree_node_apply_upserts_back (BetreeListCons (k, - BetreeMessageUpsert ufs) l) v key st1; - msgs0 <- - betree_node_lookup_first_message_for_key_back key msgs pending0; - (st3, _) <- - betree_store_internal_node_fwd node0.betree_internal_id msgs0 st2; - Return (st3, SOME v0) - od) - | BetreeListNil => - do - (st1, opt) <- betree_internal_lookup_in_children_fwd node key st0; - _ <- - betree_node_lookup_first_message_for_key_back key msgs BetreeListNil; - Return (st1, opt) - od) - od - | BetreeNodeLeaf node => - do - (st0, bindings) <- betree_load_leaf_node_fwd node.betree_leaf_id st; - opt <- betree_node_lookup_in_bindings_fwd key bindings; - Return (st0, opt) - od)) - /\ - - (** [betree_main::betree::Node::{5}::lookup]: backward function 0 *) - betree_node_lookup_back - (self : betree_node_t) (key : u64) (st : state) : betree_node_t result = - (case self of - | BetreeNodeInternal node => - do - (st0, msgs) <- betree_load_internal_node_fwd node.betree_internal_id st; - pending <- betree_node_lookup_first_message_for_key_fwd key msgs; - (case pending of - | BetreeListCons p l => - let (k, msg) = p in - if k <> key - then ( - do - _ <- - betree_node_lookup_first_message_for_key_back key msgs - (BetreeListCons (k, msg) l); - node0 <- betree_internal_lookup_in_children_back node key st0; - Return (BetreeNodeInternal node0) - od) - else - (case msg of - | BetreeMessageInsert v => - do - _ <- - betree_node_lookup_first_message_for_key_back key msgs - (BetreeListCons (k, BetreeMessageInsert v) l); - Return (BetreeNodeInternal node) - od - | BetreeMessageDelete => - do - _ <- - betree_node_lookup_first_message_for_key_back key msgs - (BetreeListCons (k, BetreeMessageDelete) l); - Return (BetreeNodeInternal node) - od - | BetreeMessageUpsert ufs => - do - (st1, v) <- betree_internal_lookup_in_children_fwd node key st0; - (st2, _) <- - betree_node_apply_upserts_fwd (BetreeListCons (k, - BetreeMessageUpsert ufs) l) v key st1; - node0 <- betree_internal_lookup_in_children_back node key st0; - pending0 <- - betree_node_apply_upserts_back (BetreeListCons (k, - BetreeMessageUpsert ufs) l) v key st1; - msgs0 <- - betree_node_lookup_first_message_for_key_back key msgs pending0; - _ <- - betree_store_internal_node_fwd node0.betree_internal_id msgs0 st2; - Return (BetreeNodeInternal node0) - od) - | BetreeListNil => - do - _ <- - betree_node_lookup_first_message_for_key_back key msgs BetreeListNil; - node0 <- betree_internal_lookup_in_children_back node key st0; - Return (BetreeNodeInternal node0) - od) - od - | BetreeNodeLeaf node => - do - (_, bindings) <- betree_load_leaf_node_fwd node.betree_leaf_id st; - _ <- betree_node_lookup_in_bindings_fwd key bindings; - Return (BetreeNodeLeaf node) - od) -’ - -val [betree_node_filter_messages_for_key_fwd_back_def] = DefineDiv ‘ - (** [betree_main::betree::Node::{5}::filter_messages_for_key]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - betree_node_filter_messages_for_key_fwd_back - (key : u64) (msgs : (u64 # betree_message_t) betree_list_t) : - (u64 # betree_message_t) betree_list_t result - = - (case msgs of - | BetreeListCons p l => - let (k, m) = p in - if k = key - then ( - do - msgs0 <- betree_list_pop_front_back (BetreeListCons (k, m) l); - betree_node_filter_messages_for_key_fwd_back key msgs0 - od) - else Return (BetreeListCons (k, m) l) - | BetreeListNil => Return BetreeListNil) -’ - -val [betree_node_lookup_first_message_after_key_fwd_def] = DefineDiv ‘ - (** [betree_main::betree::Node::{5}::lookup_first_message_after_key]: forward function *) - betree_node_lookup_first_message_after_key_fwd - (key : u64) (msgs : (u64 # betree_message_t) betree_list_t) : - (u64 # betree_message_t) betree_list_t result - = - (case msgs of - | BetreeListCons p next_msgs => - let (k, m) = p in - if k = key - then betree_node_lookup_first_message_after_key_fwd key next_msgs - else Return (BetreeListCons (k, m) next_msgs) - | BetreeListNil => Return BetreeListNil) -’ - -val [betree_node_lookup_first_message_after_key_back_def] = DefineDiv ‘ - (** [betree_main::betree::Node::{5}::lookup_first_message_after_key]: backward function 0 *) - betree_node_lookup_first_message_after_key_back - (key : u64) (msgs : (u64 # betree_message_t) betree_list_t) - (ret : (u64 # betree_message_t) betree_list_t) : - (u64 # betree_message_t) betree_list_t result - = - (case msgs of - | BetreeListCons p next_msgs => - let (k, m) = p in - if k = key - then ( - do - next_msgs0 <- - betree_node_lookup_first_message_after_key_back key next_msgs ret; - Return (BetreeListCons (k, m) next_msgs0) - od) - else Return ret - | BetreeListNil => Return ret) -’ - -val betree_node_apply_to_internal_fwd_back_def = Define ‘ - (** [betree_main::betree::Node::{5}::apply_to_internal]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - betree_node_apply_to_internal_fwd_back - (msgs : (u64 # betree_message_t) betree_list_t) (key : u64) - (new_msg : betree_message_t) : - (u64 # betree_message_t) betree_list_t result - = - do - msgs0 <- betree_node_lookup_first_message_for_key_fwd key msgs; - b <- betree_list_head_has_key_fwd msgs0 key; - if b - then - (case new_msg of - | BetreeMessageInsert i => - do - msgs1 <- betree_node_filter_messages_for_key_fwd_back key msgs0; - msgs2 <- - betree_list_push_front_fwd_back msgs1 (key, BetreeMessageInsert i); - betree_node_lookup_first_message_for_key_back key msgs msgs2 - od - | BetreeMessageDelete => - do - msgs1 <- betree_node_filter_messages_for_key_fwd_back key msgs0; - msgs2 <- - betree_list_push_front_fwd_back msgs1 (key, BetreeMessageDelete); - betree_node_lookup_first_message_for_key_back key msgs msgs2 - od - | BetreeMessageUpsert s => - do - p <- betree_list_hd_fwd msgs0; - let (_, m) = p in - (case m of - | BetreeMessageInsert prev => - do - v <- betree_upsert_update_fwd (SOME prev) s; - msgs1 <- betree_list_pop_front_back msgs0; - msgs2 <- - betree_list_push_front_fwd_back msgs1 (key, BetreeMessageInsert v); - betree_node_lookup_first_message_for_key_back key msgs msgs2 - od - | BetreeMessageDelete => - do - v <- betree_upsert_update_fwd NONE s; - msgs1 <- betree_list_pop_front_back msgs0; - msgs2 <- - betree_list_push_front_fwd_back msgs1 (key, BetreeMessageInsert v); - betree_node_lookup_first_message_for_key_back key msgs msgs2 - od - | BetreeMessageUpsert ufs => - do - msgs1 <- betree_node_lookup_first_message_after_key_fwd key msgs0; - msgs2 <- - betree_list_push_front_fwd_back msgs1 (key, BetreeMessageUpsert s); - msgs3 <- - betree_node_lookup_first_message_after_key_back key msgs0 msgs2; - betree_node_lookup_first_message_for_key_back key msgs msgs3 - od) - od) - else ( - do - msgs1 <- betree_list_push_front_fwd_back msgs0 (key, new_msg); - betree_node_lookup_first_message_for_key_back key msgs msgs1 - od) - od -’ - -val [betree_node_apply_messages_to_internal_fwd_back_def] = DefineDiv ‘ - (** [betree_main::betree::Node::{5}::apply_messages_to_internal]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - betree_node_apply_messages_to_internal_fwd_back - (msgs : (u64 # betree_message_t) betree_list_t) - (new_msgs : (u64 # betree_message_t) betree_list_t) : - (u64 # betree_message_t) betree_list_t result - = - (case new_msgs of - | BetreeListCons new_msg new_msgs_tl => - let (i, m) = new_msg in - do - msgs0 <- betree_node_apply_to_internal_fwd_back msgs i m; - betree_node_apply_messages_to_internal_fwd_back msgs0 new_msgs_tl - od - | BetreeListNil => Return msgs) -’ - -val [betree_node_lookup_mut_in_bindings_fwd_def] = DefineDiv ‘ - (** [betree_main::betree::Node::{5}::lookup_mut_in_bindings]: forward function *) - betree_node_lookup_mut_in_bindings_fwd - (key : u64) (bindings : (u64 # u64) betree_list_t) : - (u64 # u64) betree_list_t result - = - (case bindings of - | BetreeListCons hd tl => - let (i, i0) = hd in - if u64_ge i key - then Return (BetreeListCons (i, i0) tl) - else betree_node_lookup_mut_in_bindings_fwd key tl - | BetreeListNil => Return BetreeListNil) -’ - -val [betree_node_lookup_mut_in_bindings_back_def] = DefineDiv ‘ - (** [betree_main::betree::Node::{5}::lookup_mut_in_bindings]: backward function 0 *) - betree_node_lookup_mut_in_bindings_back - (key : u64) (bindings : (u64 # u64) betree_list_t) - (ret : (u64 # u64) betree_list_t) : - (u64 # u64) betree_list_t result - = - (case bindings of - | BetreeListCons hd tl => - let (i, i0) = hd in - if u64_ge i key - then Return ret - else ( - do - tl0 <- betree_node_lookup_mut_in_bindings_back key tl ret; - Return (BetreeListCons (i, i0) tl0) - od) - | BetreeListNil => Return ret) -’ - -val betree_node_apply_to_leaf_fwd_back_def = Define ‘ - (** [betree_main::betree::Node::{5}::apply_to_leaf]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - betree_node_apply_to_leaf_fwd_back - (bindings : (u64 # u64) betree_list_t) (key : u64) - (new_msg : betree_message_t) : - (u64 # u64) betree_list_t result - = - do - bindings0 <- betree_node_lookup_mut_in_bindings_fwd key bindings; - b <- betree_list_head_has_key_fwd bindings0 key; - if b - then ( - do - hd <- betree_list_pop_front_fwd bindings0; - (case new_msg of - | BetreeMessageInsert v => - do - bindings1 <- betree_list_pop_front_back bindings0; - bindings2 <- betree_list_push_front_fwd_back bindings1 (key, v); - betree_node_lookup_mut_in_bindings_back key bindings bindings2 - od - | BetreeMessageDelete => - do - bindings1 <- betree_list_pop_front_back bindings0; - betree_node_lookup_mut_in_bindings_back key bindings bindings1 - od - | BetreeMessageUpsert s => - let (_, i) = hd in - do - v <- betree_upsert_update_fwd (SOME i) s; - bindings1 <- betree_list_pop_front_back bindings0; - bindings2 <- betree_list_push_front_fwd_back bindings1 (key, v); - betree_node_lookup_mut_in_bindings_back key bindings bindings2 - od) - od) - else - (case new_msg of - | BetreeMessageInsert v => - do - bindings1 <- betree_list_push_front_fwd_back bindings0 (key, v); - betree_node_lookup_mut_in_bindings_back key bindings bindings1 - od - | BetreeMessageDelete => - betree_node_lookup_mut_in_bindings_back key bindings bindings0 - | BetreeMessageUpsert s => - do - v <- betree_upsert_update_fwd NONE s; - bindings1 <- betree_list_push_front_fwd_back bindings0 (key, v); - betree_node_lookup_mut_in_bindings_back key bindings bindings1 - od) - od -’ - -val [betree_node_apply_messages_to_leaf_fwd_back_def] = DefineDiv ‘ - (** [betree_main::betree::Node::{5}::apply_messages_to_leaf]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) *) - betree_node_apply_messages_to_leaf_fwd_back - (bindings : (u64 # u64) betree_list_t) - (new_msgs : (u64 # betree_message_t) betree_list_t) : - (u64 # u64) betree_list_t result - = - (case new_msgs of - | BetreeListCons new_msg new_msgs_tl => - let (i, m) = new_msg in - do - bindings0 <- betree_node_apply_to_leaf_fwd_back bindings i m; - betree_node_apply_messages_to_leaf_fwd_back bindings0 new_msgs_tl - od - | BetreeListNil => Return bindings) -’ - -val [betree_internal_flush_fwd_def, betree_internal_flush_back_def, betree_node_apply_messages_fwd_def, betree_node_apply_messages_back_def] = DefineDiv ‘ - (** [betree_main::betree::Internal::{4}::flush]: forward function *) - (betree_internal_flush_fwd - (self : betree_internal_t) (params : betree_params_t) - (node_id_cnt : betree_node_id_counter_t) - (content : (u64 # betree_message_t) betree_list_t) (st : state) : - (state # (u64 # betree_message_t) betree_list_t) result - = - do - p <- betree_list_partition_at_pivot_fwd content self.betree_internal_pivot; - let (msgs_left, msgs_right) = p in - do - len_left <- betree_list_len_fwd msgs_left; - if u64_ge len_left params.betree_params_min_flush_size - then ( - do - (st0, _) <- - betree_node_apply_messages_fwd self.betree_internal_left params - node_id_cnt msgs_left st; - (_, node_id_cnt0) <- - betree_node_apply_messages_back self.betree_internal_left params - node_id_cnt msgs_left st; - len_right <- betree_list_len_fwd msgs_right; - if u64_ge len_right params.betree_params_min_flush_size - then ( - do - (st1, _) <- - betree_node_apply_messages_fwd self.betree_internal_right params - node_id_cnt0 msgs_right st0; - _ <- - betree_node_apply_messages_back self.betree_internal_right params - node_id_cnt0 msgs_right st0; - Return (st1, BetreeListNil) - od) - else Return (st0, msgs_right) - od) - else ( - do - (st0, _) <- - betree_node_apply_messages_fwd self.betree_internal_right params - node_id_cnt msgs_right st; - _ <- - betree_node_apply_messages_back self.betree_internal_right params - node_id_cnt msgs_right st; - Return (st0, msgs_left) - od) - od - od) - /\ - - (** [betree_main::betree::Internal::{4}::flush]: backward function 0 *) - (betree_internal_flush_back - (self : betree_internal_t) (params : betree_params_t) - (node_id_cnt : betree_node_id_counter_t) - (content : (u64 # betree_message_t) betree_list_t) (st : state) : - (betree_internal_t # betree_node_id_counter_t) result - = - do - p <- betree_list_partition_at_pivot_fwd content self.betree_internal_pivot; - let (msgs_left, msgs_right) = p in - do - len_left <- betree_list_len_fwd msgs_left; - if u64_ge len_left params.betree_params_min_flush_size - then ( - do - (st0, _) <- - betree_node_apply_messages_fwd self.betree_internal_left params - node_id_cnt msgs_left st; - (n, node_id_cnt0) <- - betree_node_apply_messages_back self.betree_internal_left params - node_id_cnt msgs_left st; - len_right <- betree_list_len_fwd msgs_right; - if u64_ge len_right params.betree_params_min_flush_size - then ( - do - (n0, node_id_cnt1) <- - betree_node_apply_messages_back self.betree_internal_right params - node_id_cnt0 msgs_right st0; - Return - (self - with - <| - betree_internal_left := n; betree_internal_right := n0 - |>, node_id_cnt1) - od) - else Return (self with <| betree_internal_left := n |>, node_id_cnt0) - od) - else ( - do - (n, node_id_cnt0) <- - betree_node_apply_messages_back self.betree_internal_right params - node_id_cnt msgs_right st; - Return (self with <| betree_internal_right := n |>, node_id_cnt0) - od) - od - od) - /\ - - (** [betree_main::betree::Node::{5}::apply_messages]: forward function *) - (betree_node_apply_messages_fwd - (self : betree_node_t) (params : betree_params_t) - (node_id_cnt : betree_node_id_counter_t) - (msgs : (u64 # betree_message_t) betree_list_t) (st : state) : - (state # unit) result - = - (case self of - | BetreeNodeInternal node => - do - (st0, content) <- - betree_load_internal_node_fwd node.betree_internal_id st; - content0 <- betree_node_apply_messages_to_internal_fwd_back content msgs; - num_msgs <- betree_list_len_fwd content0; - if u64_ge num_msgs params.betree_params_min_flush_size - then ( - do - (st1, content1) <- - betree_internal_flush_fwd node params node_id_cnt content0 st0; - (node0, _) <- - betree_internal_flush_back node params node_id_cnt content0 st0; - (st2, _) <- - betree_store_internal_node_fwd node0.betree_internal_id content1 st1; - Return (st2, ()) - od) - else ( - do - (st1, _) <- - betree_store_internal_node_fwd node.betree_internal_id content0 st0; - Return (st1, ()) - od) - od - | BetreeNodeLeaf node => - do - (st0, content) <- betree_load_leaf_node_fwd node.betree_leaf_id st; - content0 <- betree_node_apply_messages_to_leaf_fwd_back content msgs; - len <- betree_list_len_fwd content0; - i <- u64_mul (int_to_u64 2) params.betree_params_split_size; - if u64_ge len i - then ( - do - (st1, _) <- betree_leaf_split_fwd node content0 params node_id_cnt st0; - (st2, _) <- - betree_store_leaf_node_fwd node.betree_leaf_id BetreeListNil st1; - Return (st2, ()) - od) - else ( - do - (st1, _) <- - betree_store_leaf_node_fwd node.betree_leaf_id content0 st0; - Return (st1, ()) - od) - od)) - /\ - - (** [betree_main::betree::Node::{5}::apply_messages]: backward function 0 *) - betree_node_apply_messages_back - (self : betree_node_t) (params : betree_params_t) - (node_id_cnt : betree_node_id_counter_t) - (msgs : (u64 # betree_message_t) betree_list_t) (st : state) : - (betree_node_t # betree_node_id_counter_t) result - = - (case self of - | BetreeNodeInternal node => - do - (st0, content) <- - betree_load_internal_node_fwd node.betree_internal_id st; - content0 <- betree_node_apply_messages_to_internal_fwd_back content msgs; - num_msgs <- betree_list_len_fwd content0; - if u64_ge num_msgs params.betree_params_min_flush_size - then ( - do - (st1, content1) <- - betree_internal_flush_fwd node params node_id_cnt content0 st0; - (node0, node_id_cnt0) <- - betree_internal_flush_back node params node_id_cnt content0 st0; - _ <- - betree_store_internal_node_fwd node0.betree_internal_id content1 st1; - Return (BetreeNodeInternal node0, node_id_cnt0) - od) - else ( - do - _ <- - betree_store_internal_node_fwd node.betree_internal_id content0 st0; - Return (BetreeNodeInternal node, node_id_cnt) - od) - od - | BetreeNodeLeaf node => - do - (st0, content) <- betree_load_leaf_node_fwd node.betree_leaf_id st; - content0 <- betree_node_apply_messages_to_leaf_fwd_back content msgs; - len <- betree_list_len_fwd content0; - i <- u64_mul (int_to_u64 2) params.betree_params_split_size; - if u64_ge len i - then ( - do - (st1, new_node) <- - betree_leaf_split_fwd node content0 params node_id_cnt st0; - _ <- betree_store_leaf_node_fwd node.betree_leaf_id BetreeListNil st1; - node_id_cnt0 <- - betree_leaf_split_back node content0 params node_id_cnt st0; - Return (BetreeNodeInternal new_node, node_id_cnt0) - od) - else ( - do - _ <- betree_store_leaf_node_fwd node.betree_leaf_id content0 st0; - Return (BetreeNodeLeaf (node with <| betree_leaf_size := len |>), - node_id_cnt) - od) - od) -’ - -val betree_node_apply_fwd_def = Define ‘ - (** [betree_main::betree::Node::{5}::apply]: forward function *) - betree_node_apply_fwd - (self : betree_node_t) (params : betree_params_t) - (node_id_cnt : betree_node_id_counter_t) (key : u64) - (new_msg : betree_message_t) (st : state) : - (state # unit) result - = - let l = BetreeListNil in - do - (st0, _) <- - betree_node_apply_messages_fwd self params node_id_cnt (BetreeListCons - (key, new_msg) l) st; - _ <- - betree_node_apply_messages_back self params node_id_cnt (BetreeListCons - (key, new_msg) l) st; - Return (st0, ()) - od -’ - -val betree_node_apply_back_def = Define ‘ - (** [betree_main::betree::Node::{5}::apply]: backward function 0 *) - betree_node_apply_back - (self : betree_node_t) (params : betree_params_t) - (node_id_cnt : betree_node_id_counter_t) (key : u64) - (new_msg : betree_message_t) (st : state) : - (betree_node_t # betree_node_id_counter_t) result - = - let l = BetreeListNil in - betree_node_apply_messages_back self params node_id_cnt (BetreeListCons - (key, new_msg) l) st -’ - -val betree_be_tree_new_fwd_def = Define ‘ - (** [betree_main::betree::BeTree::{6}::new]: forward function *) - betree_be_tree_new_fwd - (min_flush_size : u64) (split_size : u64) (st : state) : - (state # betree_be_tree_t) result - = - do - node_id_cnt <- betree_node_id_counter_new_fwd; - id <- betree_node_id_counter_fresh_id_fwd node_id_cnt; - (st0, _) <- betree_store_leaf_node_fwd id BetreeListNil st; - node_id_cnt0 <- betree_node_id_counter_fresh_id_back node_id_cnt; - Return (st0, - <| - betree_be_tree_params := - (<| - betree_params_min_flush_size := min_flush_size; - betree_params_split_size := split_size - |>); - betree_be_tree_node_id_cnt := node_id_cnt0; - betree_be_tree_root := - (BetreeNodeLeaf - (<| betree_leaf_id := id; betree_leaf_size := (int_to_u64 0) |>)) - |>) - od -’ - -val betree_be_tree_apply_fwd_def = Define ‘ - (** [betree_main::betree::BeTree::{6}::apply]: forward function *) - betree_be_tree_apply_fwd - (self : betree_be_tree_t) (key : u64) (msg : betree_message_t) (st : state) - : - (state # unit) result - = - do - (st0, _) <- - betree_node_apply_fwd self.betree_be_tree_root self.betree_be_tree_params - self.betree_be_tree_node_id_cnt key msg st; - _ <- - betree_node_apply_back self.betree_be_tree_root - self.betree_be_tree_params self.betree_be_tree_node_id_cnt key msg st; - Return (st0, ()) - od -’ - -val betree_be_tree_apply_back_def = Define ‘ - (** [betree_main::betree::BeTree::{6}::apply]: backward function 0 *) - betree_be_tree_apply_back - (self : betree_be_tree_t) (key : u64) (msg : betree_message_t) (st : state) - : - betree_be_tree_t result - = - do - (n, nic) <- - betree_node_apply_back self.betree_be_tree_root - self.betree_be_tree_params self.betree_be_tree_node_id_cnt key msg st; - Return - (self - with - <| - betree_be_tree_node_id_cnt := nic; betree_be_tree_root := n - |>) - od -’ - -val betree_be_tree_insert_fwd_def = Define ‘ - (** [betree_main::betree::BeTree::{6}::insert]: forward function *) - betree_be_tree_insert_fwd - (self : betree_be_tree_t) (key : u64) (value : u64) (st : state) : - (state # unit) result - = - do - (st0, _) <- - betree_be_tree_apply_fwd self key (BetreeMessageInsert value) st; - _ <- betree_be_tree_apply_back self key (BetreeMessageInsert value) st; - Return (st0, ()) - od -’ - -val betree_be_tree_insert_back_def = Define ‘ - (** [betree_main::betree::BeTree::{6}::insert]: backward function 0 *) - betree_be_tree_insert_back - (self : betree_be_tree_t) (key : u64) (value : u64) (st : state) : - betree_be_tree_t result - = - betree_be_tree_apply_back self key (BetreeMessageInsert value) st -’ - -val betree_be_tree_delete_fwd_def = Define ‘ - (** [betree_main::betree::BeTree::{6}::delete]: forward function *) - betree_be_tree_delete_fwd - (self : betree_be_tree_t) (key : u64) (st : state) : - (state # unit) result - = - do - (st0, _) <- betree_be_tree_apply_fwd self key BetreeMessageDelete st; - _ <- betree_be_tree_apply_back self key BetreeMessageDelete st; - Return (st0, ()) - od -’ - -val betree_be_tree_delete_back_def = Define ‘ - (** [betree_main::betree::BeTree::{6}::delete]: backward function 0 *) - betree_be_tree_delete_back - (self : betree_be_tree_t) (key : u64) (st : state) : - betree_be_tree_t result - = - betree_be_tree_apply_back self key BetreeMessageDelete st -’ - -val betree_be_tree_upsert_fwd_def = Define ‘ - (** [betree_main::betree::BeTree::{6}::upsert]: forward function *) - betree_be_tree_upsert_fwd - (self : betree_be_tree_t) (key : u64) (upd : betree_upsert_fun_state_t) - (st : state) : - (state # unit) result - = - do - (st0, _) <- betree_be_tree_apply_fwd self key (BetreeMessageUpsert upd) st; - _ <- betree_be_tree_apply_back self key (BetreeMessageUpsert upd) st; - Return (st0, ()) - od -’ - -val betree_be_tree_upsert_back_def = Define ‘ - (** [betree_main::betree::BeTree::{6}::upsert]: backward function 0 *) - betree_be_tree_upsert_back - (self : betree_be_tree_t) (key : u64) (upd : betree_upsert_fun_state_t) - (st : state) : - betree_be_tree_t result - = - betree_be_tree_apply_back self key (BetreeMessageUpsert upd) st -’ - -val betree_be_tree_lookup_fwd_def = Define ‘ - (** [betree_main::betree::BeTree::{6}::lookup]: forward function *) - betree_be_tree_lookup_fwd - (self : betree_be_tree_t) (key : u64) (st : state) : - (state # u64 option) result - = - betree_node_lookup_fwd self.betree_be_tree_root key st -’ - -val betree_be_tree_lookup_back_def = Define ‘ - (** [betree_main::betree::BeTree::{6}::lookup]: backward function 0 *) - betree_be_tree_lookup_back - (self : betree_be_tree_t) (key : u64) (st : state) : - betree_be_tree_t result - = - do - n <- betree_node_lookup_back self.betree_be_tree_root key st; - Return (self with <| betree_be_tree_root := n |>) - od -’ - -val main_fwd_def = Define ‘ - (** [betree_main::main]: forward function *) - main_fwd : unit result = - Return () -’ - -(** Unit test for [betree_main::main] *) -val _ = assert_return (“main_fwd”) - -val _ = export_theory () diff --git a/tests/hol4/betree/betreeMain_FunsTheory.sig b/tests/hol4/betree/betreeMain_FunsTheory.sig deleted file mode 100644 index c922ca9f..00000000 --- a/tests/hol4/betree/betreeMain_FunsTheory.sig +++ /dev/null @@ -1,1230 +0,0 @@ -signature betreeMain_FunsTheory = -sig - type thm = Thm.thm - - (* Definitions *) - val betree_be_tree_apply_back_def : thm - val betree_be_tree_apply_fwd_def : thm - val betree_be_tree_delete_back_def : thm - val betree_be_tree_delete_fwd_def : thm - val betree_be_tree_insert_back_def : thm - val betree_be_tree_insert_fwd_def : thm - val betree_be_tree_lookup_back_def : thm - val betree_be_tree_lookup_fwd_def : thm - val betree_be_tree_new_fwd_def : thm - val betree_be_tree_upsert_back_def : thm - val betree_be_tree_upsert_fwd_def : thm - val betree_fresh_node_id_back_def : thm - val betree_fresh_node_id_fwd_def : thm - val betree_internal_flush_back_def : thm - val betree_internal_flush_fwd_def : thm - val betree_internal_lookup_in_children_back_def : thm - val betree_internal_lookup_in_children_fwd_def : thm - val betree_leaf_split_back_def : thm - val betree_leaf_split_fwd_def : thm - val betree_list_hd_fwd_def : thm - val betree_list_head_has_key_fwd_def : thm - val betree_list_len_fwd_def : thm - val betree_list_partition_at_pivot_fwd_def : thm - val betree_list_pop_front_back_def : thm - val betree_list_pop_front_fwd_def : thm - val betree_list_push_front_fwd_back_def : thm - val betree_list_split_at_fwd_def : thm - val betree_load_internal_node_fwd_def : thm - val betree_load_leaf_node_fwd_def : thm - val betree_node_apply_back_def : thm - val betree_node_apply_fwd_def : thm - val betree_node_apply_messages_back_def : thm - val betree_node_apply_messages_fwd_def : thm - val betree_node_apply_messages_to_internal_fwd_back_def : thm - val betree_node_apply_messages_to_leaf_fwd_back_def : thm - val betree_node_apply_to_internal_fwd_back_def : thm - val betree_node_apply_to_leaf_fwd_back_def : thm - val betree_node_apply_upserts_back_def : thm - val betree_node_apply_upserts_fwd_def : thm - val betree_node_filter_messages_for_key_fwd_back_def : thm - val betree_node_id_counter_fresh_id_back_def : thm - val betree_node_id_counter_fresh_id_fwd_def : thm - val betree_node_id_counter_new_fwd_def : thm - val betree_node_lookup_back_def : thm - val betree_node_lookup_first_message_after_key_back_def : thm - val betree_node_lookup_first_message_after_key_fwd_def : thm - val betree_node_lookup_first_message_for_key_back_def : thm - val betree_node_lookup_first_message_for_key_fwd_def : thm - val betree_node_lookup_fwd_def : thm - val betree_node_lookup_in_bindings_fwd_def : thm - val betree_node_lookup_mut_in_bindings_back_def : thm - val betree_node_lookup_mut_in_bindings_fwd_def : thm - val betree_store_internal_node_fwd_def : thm - val betree_store_leaf_node_fwd_def : thm - val betree_upsert_update_fwd_def : thm - val main_fwd_def : thm - - val betreeMain_Funs_grammars : type_grammar.grammar * term_grammar.grammar -(* - [betreeMain_Opaque] Parent theory of "betreeMain_Funs" - - [betree_be_tree_apply_back_def] Definition - - ⊢ ∀self key msg st. - betree_be_tree_apply_back self key msg st = - do - (n,nic) <- - betree_node_apply_back self.betree_be_tree_root - self.betree_be_tree_params self.betree_be_tree_node_id_cnt - key msg st; - Return - (self with - <|betree_be_tree_node_id_cnt := nic; - betree_be_tree_root := n|>) - od - - [betree_be_tree_apply_fwd_def] Definition - - ⊢ ∀self key msg st. - betree_be_tree_apply_fwd self key msg st = - do - (st0,_) <- - betree_node_apply_fwd self.betree_be_tree_root - self.betree_be_tree_params self.betree_be_tree_node_id_cnt - key msg st; - monad_ignore_bind - (betree_node_apply_back self.betree_be_tree_root - self.betree_be_tree_params self.betree_be_tree_node_id_cnt - key msg st) (Return (st0,())) - od - - [betree_be_tree_delete_back_def] Definition - - ⊢ ∀self key st. - betree_be_tree_delete_back self key st = - betree_be_tree_apply_back self key BetreeMessageDelete st - - [betree_be_tree_delete_fwd_def] Definition - - ⊢ ∀self key st. - betree_be_tree_delete_fwd self key st = - do - (st0,_) <- - betree_be_tree_apply_fwd self key BetreeMessageDelete st; - monad_ignore_bind - (betree_be_tree_apply_back self key BetreeMessageDelete st) - (Return (st0,())) - od - - [betree_be_tree_insert_back_def] Definition - - ⊢ ∀self key value st. - betree_be_tree_insert_back self key value st = - betree_be_tree_apply_back self key (BetreeMessageInsert value) st - - [betree_be_tree_insert_fwd_def] Definition - - ⊢ ∀self key value st. - betree_be_tree_insert_fwd self key value st = - do - (st0,_) <- - betree_be_tree_apply_fwd self key (BetreeMessageInsert value) - st; - monad_ignore_bind - (betree_be_tree_apply_back self key - (BetreeMessageInsert value) st) (Return (st0,())) - od - - [betree_be_tree_lookup_back_def] Definition - - ⊢ ∀self key st. - betree_be_tree_lookup_back self key st = - do - n <- betree_node_lookup_back self.betree_be_tree_root key st; - Return (self with betree_be_tree_root := n) - od - - [betree_be_tree_lookup_fwd_def] Definition - - ⊢ ∀self key st. - betree_be_tree_lookup_fwd self key st = - betree_node_lookup_fwd self.betree_be_tree_root key st - - [betree_be_tree_new_fwd_def] Definition - - ⊢ ∀min_flush_size split_size st. - betree_be_tree_new_fwd min_flush_size split_size st = - do - node_id_cnt <- betree_node_id_counter_new_fwd; - id <- betree_node_id_counter_fresh_id_fwd node_id_cnt; - (st0,_) <- betree_store_leaf_node_fwd id BetreeListNil st; - node_id_cnt0 <- - betree_node_id_counter_fresh_id_back node_id_cnt; - Return - (st0, - <|betree_be_tree_params := - <|betree_params_min_flush_size := min_flush_size; - betree_params_split_size := split_size|>; - betree_be_tree_node_id_cnt := node_id_cnt0; - betree_be_tree_root := - BetreeNodeLeaf - <|betree_leaf_id := id; - betree_leaf_size := int_to_u64 0|> |>) - od - - [betree_be_tree_upsert_back_def] Definition - - ⊢ ∀self key upd st. - betree_be_tree_upsert_back self key upd st = - betree_be_tree_apply_back self key (BetreeMessageUpsert upd) st - - [betree_be_tree_upsert_fwd_def] Definition - - ⊢ ∀self key upd st. - betree_be_tree_upsert_fwd self key upd st = - do - (st0,_) <- - betree_be_tree_apply_fwd self key (BetreeMessageUpsert upd) - st; - monad_ignore_bind - (betree_be_tree_apply_back self key (BetreeMessageUpsert upd) - st) (Return (st0,())) - od - - [betree_fresh_node_id_back_def] Definition - - ⊢ ∀counter. - betree_fresh_node_id_back counter = - u64_add counter (int_to_u64 1) - - [betree_fresh_node_id_fwd_def] Definition - - ⊢ ∀counter. - betree_fresh_node_id_fwd counter = - monad_ignore_bind (u64_add counter (int_to_u64 1)) - (Return counter) - - [betree_internal_flush_back_def] Definition - - ⊢ ∀self params node_id_cnt content st. - betree_internal_flush_back self params node_id_cnt content st = - do - p <- - betree_list_partition_at_pivot_fwd content - self.betree_internal_pivot; - (msgs_left,msgs_right) <<- p; - len_left <- betree_list_len_fwd msgs_left; - if u64_ge len_left params.betree_params_min_flush_size then - do - (st0,_) <- - betree_node_apply_messages_fwd self.betree_internal_left - params node_id_cnt msgs_left st; - (n,node_id_cnt0) <- - betree_node_apply_messages_back self.betree_internal_left - params node_id_cnt msgs_left st; - len_right <- betree_list_len_fwd msgs_right; - if - u64_ge len_right params.betree_params_min_flush_size - then - do - (n0,node_id_cnt1) <- - betree_node_apply_messages_back - self.betree_internal_right params node_id_cnt0 - msgs_right st0; - Return - (self with - <|betree_internal_left := n; - betree_internal_right := n0|>,node_id_cnt1) - od - else - Return (self with betree_internal_left := n,node_id_cnt0) - od - else - do - (n,node_id_cnt0) <- - betree_node_apply_messages_back - self.betree_internal_right params node_id_cnt - msgs_right st; - Return (self with betree_internal_right := n,node_id_cnt0) - od - od - - [betree_internal_flush_fwd_def] Definition - - ⊢ ∀self params node_id_cnt content st. - betree_internal_flush_fwd self params node_id_cnt content st = - do - p <- - betree_list_partition_at_pivot_fwd content - self.betree_internal_pivot; - (msgs_left,msgs_right) <<- p; - len_left <- betree_list_len_fwd msgs_left; - if u64_ge len_left params.betree_params_min_flush_size then - do - (st0,_) <- - betree_node_apply_messages_fwd self.betree_internal_left - params node_id_cnt msgs_left st; - (_,node_id_cnt0) <- - betree_node_apply_messages_back self.betree_internal_left - params node_id_cnt msgs_left st; - len_right <- betree_list_len_fwd msgs_right; - if - u64_ge len_right params.betree_params_min_flush_size - then - do - (st1,_) <- - betree_node_apply_messages_fwd - self.betree_internal_right params node_id_cnt0 - msgs_right st0; - monad_ignore_bind - (betree_node_apply_messages_back - self.betree_internal_right params node_id_cnt0 - msgs_right st0) (Return (st1,BetreeListNil)) - od - else Return (st0,msgs_right) - od - else - do - (st0,_) <- - betree_node_apply_messages_fwd self.betree_internal_right - params node_id_cnt msgs_right st; - monad_ignore_bind - (betree_node_apply_messages_back - self.betree_internal_right params node_id_cnt - msgs_right st) (Return (st0,msgs_left)) - od - od - - [betree_internal_lookup_in_children_back_def] Definition - - ⊢ ∀self key st. - betree_internal_lookup_in_children_back self key st = - if u64_lt key self.betree_internal_pivot then - do - n <- betree_node_lookup_back self.betree_internal_left key st; - Return (self with betree_internal_left := n) - od - else - do - n <- - betree_node_lookup_back self.betree_internal_right key st; - Return (self with betree_internal_right := n) - od - - [betree_internal_lookup_in_children_fwd_def] Definition - - ⊢ ∀self key st. - betree_internal_lookup_in_children_fwd self key st = - if u64_lt key self.betree_internal_pivot then - betree_node_lookup_fwd self.betree_internal_left key st - else betree_node_lookup_fwd self.betree_internal_right key st - - [betree_leaf_split_back_def] Definition - - ⊢ ∀self content params node_id_cnt st. - betree_leaf_split_back self content params node_id_cnt st = - do - p <- - betree_list_split_at_fwd content - params.betree_params_split_size; - (content0,content1) <<- p; - monad_ignore_bind (betree_list_hd_fwd content1) - do - id0 <- betree_node_id_counter_fresh_id_fwd node_id_cnt; - node_id_cnt0 <- - betree_node_id_counter_fresh_id_back node_id_cnt; - id1 <- betree_node_id_counter_fresh_id_fwd node_id_cnt0; - (st0,_) <- betree_store_leaf_node_fwd id0 content0 st; - monad_ignore_bind - (betree_store_leaf_node_fwd id1 content1 st0) - (betree_node_id_counter_fresh_id_back node_id_cnt0) - od - od - - [betree_leaf_split_fwd_def] Definition - - ⊢ ∀self content params node_id_cnt st. - betree_leaf_split_fwd self content params node_id_cnt st = - do - p <- - betree_list_split_at_fwd content - params.betree_params_split_size; - (content0,content1) <<- p; - p0 <- betree_list_hd_fwd content1; - (pivot,_) <<- p0; - id0 <- betree_node_id_counter_fresh_id_fwd node_id_cnt; - node_id_cnt0 <- - betree_node_id_counter_fresh_id_back node_id_cnt; - id1 <- betree_node_id_counter_fresh_id_fwd node_id_cnt0; - (st0,_) <- betree_store_leaf_node_fwd id0 content0 st; - (st1,_) <- betree_store_leaf_node_fwd id1 content1 st0; - n <<- - BetreeNodeLeaf - <|betree_leaf_id := id0; - betree_leaf_size := params.betree_params_split_size|>; - n0 <<- - BetreeNodeLeaf - <|betree_leaf_id := id1; - betree_leaf_size := params.betree_params_split_size|>; - Return - (st1, - <|betree_internal_id := self.betree_leaf_id; - betree_internal_pivot := pivot; betree_internal_left := n; - betree_internal_right := n0|>) - od - - [betree_list_hd_fwd_def] Definition - - ⊢ ∀self. - betree_list_hd_fwd self = - case self of - BetreeListCons hd l => Return hd - | BetreeListNil => Fail Failure - - [betree_list_head_has_key_fwd_def] Definition - - ⊢ ∀self key. - betree_list_head_has_key_fwd self key = - case self of - BetreeListCons hd l => (let (i,_) = hd in Return (i = key)) - | BetreeListNil => Return F - - [betree_list_len_fwd_def] Definition - - ⊢ ∀self. - betree_list_len_fwd self = - case self of - BetreeListCons t tl => - do i <- betree_list_len_fwd tl; u64_add (int_to_u64 1) i od - | BetreeListNil => Return (int_to_u64 0) - - [betree_list_partition_at_pivot_fwd_def] Definition - - ⊢ ∀self pivot. - betree_list_partition_at_pivot_fwd self pivot = - case self of - BetreeListCons hd tl => - (let - (i,t) = hd - in - if u64_ge i pivot then - Return (BetreeListNil,BetreeListCons (i,t) tl) - else - do - p <- betree_list_partition_at_pivot_fwd tl pivot; - (ls0,ls1) <<- p; - l <<- ls0; - Return (BetreeListCons (i,t) l,ls1) - od) - | BetreeListNil => Return (BetreeListNil,BetreeListNil) - - [betree_list_pop_front_back_def] Definition - - ⊢ ∀self. - betree_list_pop_front_back self = - (let - ls = mem_replace_fwd self BetreeListNil - in - case ls of - BetreeListCons x tl => Return tl - | BetreeListNil => Fail Failure) - - [betree_list_pop_front_fwd_def] Definition - - ⊢ ∀self. - betree_list_pop_front_fwd self = - (let - ls = mem_replace_fwd self BetreeListNil - in - case ls of - BetreeListCons x tl => Return x - | BetreeListNil => Fail Failure) - - [betree_list_push_front_fwd_back_def] Definition - - ⊢ ∀self x. - betree_list_push_front_fwd_back self x = - (let - tl = mem_replace_fwd self BetreeListNil; - l = tl - in - Return (BetreeListCons x l)) - - [betree_list_split_at_fwd_def] Definition - - ⊢ ∀self n. - betree_list_split_at_fwd self n = - if n = int_to_u64 0 then Return (BetreeListNil,self) - else - case self of - BetreeListCons hd tl => - do - i <- u64_sub n (int_to_u64 1); - p <- betree_list_split_at_fwd tl i; - (ls0,ls1) <<- p; - l <<- ls0; - Return (BetreeListCons hd l,ls1) - od - | BetreeListNil => Fail Failure - - [betree_load_internal_node_fwd_def] Definition - - ⊢ ∀id st. - betree_load_internal_node_fwd id st = - betree_utils_load_internal_node_fwd id st - - [betree_load_leaf_node_fwd_def] Definition - - ⊢ ∀id st. - betree_load_leaf_node_fwd id st = - betree_utils_load_leaf_node_fwd id st - - [betree_node_apply_back_def] Definition - - ⊢ ∀self params node_id_cnt key new_msg st. - betree_node_apply_back self params node_id_cnt key new_msg st = - (let - l = BetreeListNil - in - betree_node_apply_messages_back self params node_id_cnt - (BetreeListCons (key,new_msg) l) st) - - [betree_node_apply_fwd_def] Definition - - ⊢ ∀self params node_id_cnt key new_msg st. - betree_node_apply_fwd self params node_id_cnt key new_msg st = - (let - l = BetreeListNil - in - do - (st0,_) <- - betree_node_apply_messages_fwd self params node_id_cnt - (BetreeListCons (key,new_msg) l) st; - monad_ignore_bind - (betree_node_apply_messages_back self params node_id_cnt - (BetreeListCons (key,new_msg) l) st) (Return (st0,())) - od) - - [betree_node_apply_messages_back_def] Definition - - ⊢ ∀self params node_id_cnt msgs st. - betree_node_apply_messages_back self params node_id_cnt msgs st = - case self of - BetreeNodeInternal node => - do - (st0,content) <- - betree_load_internal_node_fwd node.betree_internal_id st; - content0 <- - betree_node_apply_messages_to_internal_fwd_back content - msgs; - num_msgs <- betree_list_len_fwd content0; - if u64_ge num_msgs params.betree_params_min_flush_size then - do - (st1,content1) <- - betree_internal_flush_fwd node params node_id_cnt - content0 st0; - (node0,node_id_cnt0) <- - betree_internal_flush_back node params node_id_cnt - content0 st0; - monad_ignore_bind - (betree_store_internal_node_fwd - node0.betree_internal_id content1 st1) - (Return (BetreeNodeInternal node0,node_id_cnt0)) - od - else - monad_ignore_bind - (betree_store_internal_node_fwd node.betree_internal_id - content0 st0) - (Return (BetreeNodeInternal node,node_id_cnt)) - od - | BetreeNodeLeaf node' => - do - (st0,content) <- - betree_load_leaf_node_fwd node'.betree_leaf_id st; - content0 <- - betree_node_apply_messages_to_leaf_fwd_back content msgs; - len <- betree_list_len_fwd content0; - i <- u64_mul (int_to_u64 2) params.betree_params_split_size; - if u64_ge len i then - do - (st1,new_node) <- - betree_leaf_split_fwd node' content0 params node_id_cnt - st0; - monad_ignore_bind - (betree_store_leaf_node_fwd node'.betree_leaf_id - BetreeListNil st1) - do - node_id_cnt0 <- - betree_leaf_split_back node' content0 params - node_id_cnt st0; - Return (BetreeNodeInternal new_node,node_id_cnt0) - od - od - else - monad_ignore_bind - (betree_store_leaf_node_fwd node'.betree_leaf_id content0 - st0) - (Return - (BetreeNodeLeaf (node' with betree_leaf_size := len), - node_id_cnt)) - od - - [betree_node_apply_messages_fwd_def] Definition - - ⊢ ∀self params node_id_cnt msgs st. - betree_node_apply_messages_fwd self params node_id_cnt msgs st = - case self of - BetreeNodeInternal node => - do - (st0,content) <- - betree_load_internal_node_fwd node.betree_internal_id st; - content0 <- - betree_node_apply_messages_to_internal_fwd_back content - msgs; - num_msgs <- betree_list_len_fwd content0; - if u64_ge num_msgs params.betree_params_min_flush_size then - do - (st1,content1) <- - betree_internal_flush_fwd node params node_id_cnt - content0 st0; - (node0,_) <- - betree_internal_flush_back node params node_id_cnt - content0 st0; - (st2,_) <- - betree_store_internal_node_fwd - node0.betree_internal_id content1 st1; - Return (st2,()) - od - else - do - (st1,_) <- - betree_store_internal_node_fwd - node.betree_internal_id content0 st0; - Return (st1,()) - od - od - | BetreeNodeLeaf node' => - do - (st0,content) <- - betree_load_leaf_node_fwd node'.betree_leaf_id st; - content0 <- - betree_node_apply_messages_to_leaf_fwd_back content msgs; - len <- betree_list_len_fwd content0; - i <- u64_mul (int_to_u64 2) params.betree_params_split_size; - if u64_ge len i then - do - (st1,_) <- - betree_leaf_split_fwd node' content0 params node_id_cnt - st0; - (st2,_) <- - betree_store_leaf_node_fwd node'.betree_leaf_id - BetreeListNil st1; - Return (st2,()) - od - else - do - (st1,_) <- - betree_store_leaf_node_fwd node'.betree_leaf_id - content0 st0; - Return (st1,()) - od - od - - [betree_node_apply_messages_to_internal_fwd_back_def] Definition - - ⊢ ∀msgs new_msgs. - betree_node_apply_messages_to_internal_fwd_back msgs new_msgs = - case new_msgs of - BetreeListCons new_msg new_msgs_tl => - (let - (i,m) = new_msg - in - do - msgs0 <- betree_node_apply_to_internal_fwd_back msgs i m; - betree_node_apply_messages_to_internal_fwd_back msgs0 - new_msgs_tl - od) - | BetreeListNil => Return msgs - - [betree_node_apply_messages_to_leaf_fwd_back_def] Definition - - ⊢ ∀bindings new_msgs. - betree_node_apply_messages_to_leaf_fwd_back bindings new_msgs = - case new_msgs of - BetreeListCons new_msg new_msgs_tl => - (let - (i,m) = new_msg - in - do - bindings0 <- - betree_node_apply_to_leaf_fwd_back bindings i m; - betree_node_apply_messages_to_leaf_fwd_back bindings0 - new_msgs_tl - od) - | BetreeListNil => Return bindings - - [betree_node_apply_to_internal_fwd_back_def] Definition - - ⊢ ∀msgs key new_msg. - betree_node_apply_to_internal_fwd_back msgs key new_msg = - do - msgs0 <- betree_node_lookup_first_message_for_key_fwd key msgs; - b <- betree_list_head_has_key_fwd msgs0 key; - if b then - case new_msg of - BetreeMessageInsert i => - do - msgs1 <- - betree_node_filter_messages_for_key_fwd_back key - msgs0; - msgs2 <- - betree_list_push_front_fwd_back msgs1 - (key,BetreeMessageInsert i); - betree_node_lookup_first_message_for_key_back key msgs - msgs2 - od - | BetreeMessageDelete => - do - msgs1 <- - betree_node_filter_messages_for_key_fwd_back key msgs0; - msgs2 <- - betree_list_push_front_fwd_back msgs1 - (key,BetreeMessageDelete); - betree_node_lookup_first_message_for_key_back key msgs - msgs2 - od - | BetreeMessageUpsert s => - do - p <- betree_list_hd_fwd msgs0; - (_,m) <<- p; - case m of - BetreeMessageInsert prev => - do - v <- betree_upsert_update_fwd (SOME prev) s; - msgs1 <- betree_list_pop_front_back msgs0; - msgs2 <- - betree_list_push_front_fwd_back msgs1 - (key,BetreeMessageInsert v); - betree_node_lookup_first_message_for_key_back key - msgs msgs2 - od - | BetreeMessageDelete => - do - v <- betree_upsert_update_fwd NONE s; - msgs1 <- betree_list_pop_front_back msgs0; - msgs2 <- - betree_list_push_front_fwd_back msgs1 - (key,BetreeMessageInsert v); - betree_node_lookup_first_message_for_key_back key - msgs msgs2 - od - | BetreeMessageUpsert ufs => - do - msgs1 <- - betree_node_lookup_first_message_after_key_fwd key - msgs0; - msgs2 <- - betree_list_push_front_fwd_back msgs1 - (key,BetreeMessageUpsert s); - msgs3 <- - betree_node_lookup_first_message_after_key_back key - msgs0 msgs2; - betree_node_lookup_first_message_for_key_back key - msgs msgs3 - od - od - else - do - msgs1 <- - betree_list_push_front_fwd_back msgs0 (key,new_msg); - betree_node_lookup_first_message_for_key_back key msgs - msgs1 - od - od - - [betree_node_apply_to_leaf_fwd_back_def] Definition - - ⊢ ∀bindings key new_msg. - betree_node_apply_to_leaf_fwd_back bindings key new_msg = - do - bindings0 <- - betree_node_lookup_mut_in_bindings_fwd key bindings; - b <- betree_list_head_has_key_fwd bindings0 key; - if b then - do - hd <- betree_list_pop_front_fwd bindings0; - case new_msg of - BetreeMessageInsert v => - do - bindings1 <- betree_list_pop_front_back bindings0; - bindings2 <- - betree_list_push_front_fwd_back bindings1 (key,v); - betree_node_lookup_mut_in_bindings_back key bindings - bindings2 - od - | BetreeMessageDelete => - do - bindings1 <- betree_list_pop_front_back bindings0; - betree_node_lookup_mut_in_bindings_back key bindings - bindings1 - od - | BetreeMessageUpsert s => - (let - (_,i) = hd - in - do - v <- betree_upsert_update_fwd (SOME i) s; - bindings1 <- betree_list_pop_front_back bindings0; - bindings2 <- - betree_list_push_front_fwd_back bindings1 (key,v); - betree_node_lookup_mut_in_bindings_back key bindings - bindings2 - od) - od - else - case new_msg of - BetreeMessageInsert v => - do - bindings1 <- - betree_list_push_front_fwd_back bindings0 (key,v); - betree_node_lookup_mut_in_bindings_back key bindings - bindings1 - od - | BetreeMessageDelete => - betree_node_lookup_mut_in_bindings_back key bindings - bindings0 - | BetreeMessageUpsert s => - do - v <- betree_upsert_update_fwd NONE s; - bindings1 <- - betree_list_push_front_fwd_back bindings0 (key,v); - betree_node_lookup_mut_in_bindings_back key bindings - bindings1 - od - od - - [betree_node_apply_upserts_back_def] Definition - - ⊢ ∀msgs prev key st. - betree_node_apply_upserts_back msgs prev key st = - do - b <- betree_list_head_has_key_fwd msgs key; - if b then - do - msg <- betree_list_pop_front_fwd msgs; - (_,m) <<- msg; - case m of - BetreeMessageInsert i => Fail Failure - | BetreeMessageDelete => Fail Failure - | BetreeMessageUpsert s => - do - v <- betree_upsert_update_fwd prev s; - msgs0 <- betree_list_pop_front_back msgs; - betree_node_apply_upserts_back msgs0 (SOME v) key st - od - od - else - do - (_,v) <- core_option_option_unwrap_fwd prev st; - betree_list_push_front_fwd_back msgs - (key,BetreeMessageInsert v) - od - od - - [betree_node_apply_upserts_fwd_def] Definition - - ⊢ ∀msgs prev key st. - betree_node_apply_upserts_fwd msgs prev key st = - do - b <- betree_list_head_has_key_fwd msgs key; - if b then - do - msg <- betree_list_pop_front_fwd msgs; - (_,m) <<- msg; - case m of - BetreeMessageInsert i => Fail Failure - | BetreeMessageDelete => Fail Failure - | BetreeMessageUpsert s => - do - v <- betree_upsert_update_fwd prev s; - msgs0 <- betree_list_pop_front_back msgs; - betree_node_apply_upserts_fwd msgs0 (SOME v) key st - od - od - else - do - (st0,v) <- core_option_option_unwrap_fwd prev st; - monad_ignore_bind - (betree_list_push_front_fwd_back msgs - (key,BetreeMessageInsert v)) (Return (st0,v)) - od - od - - [betree_node_filter_messages_for_key_fwd_back_def] Definition - - ⊢ ∀key msgs. - betree_node_filter_messages_for_key_fwd_back key msgs = - case msgs of - BetreeListCons p l => - (let - (k,m) = p - in - if k = key then - do - msgs0 <- - betree_list_pop_front_back (BetreeListCons (k,m) l); - betree_node_filter_messages_for_key_fwd_back key msgs0 - od - else Return (BetreeListCons (k,m) l)) - | BetreeListNil => Return BetreeListNil - - [betree_node_id_counter_fresh_id_back_def] Definition - - ⊢ ∀self. - betree_node_id_counter_fresh_id_back self = - do - i <- - u64_add self.betree_node_id_counter_next_node_id - (int_to_u64 1); - Return <|betree_node_id_counter_next_node_id := i|> - od - - [betree_node_id_counter_fresh_id_fwd_def] Definition - - ⊢ ∀self. - betree_node_id_counter_fresh_id_fwd self = - monad_ignore_bind - (u64_add self.betree_node_id_counter_next_node_id - (int_to_u64 1)) - (Return self.betree_node_id_counter_next_node_id) - - [betree_node_id_counter_new_fwd_def] Definition - - ⊢ betree_node_id_counter_new_fwd = - Return <|betree_node_id_counter_next_node_id := int_to_u64 0|> - - [betree_node_lookup_back_def] Definition - - ⊢ ∀self key st. - betree_node_lookup_back self key st = - case self of - BetreeNodeInternal node => - do - (st0,msgs) <- - betree_load_internal_node_fwd node.betree_internal_id st; - pending <- - betree_node_lookup_first_message_for_key_fwd key msgs; - case pending of - BetreeListCons p l => - (let - (k,msg) = p - in - if k ≠ key then - monad_ignore_bind - (betree_node_lookup_first_message_for_key_back - key msgs (BetreeListCons (k,msg) l)) - do - node0 <- - betree_internal_lookup_in_children_back node - key st0; - Return (BetreeNodeInternal node0) - od - else - case msg of - BetreeMessageInsert v => - monad_ignore_bind - (betree_node_lookup_first_message_for_key_back - key msgs - (BetreeListCons (k,BetreeMessageInsert v) - l)) (Return (BetreeNodeInternal node)) - | BetreeMessageDelete => - monad_ignore_bind - (betree_node_lookup_first_message_for_key_back - key msgs - (BetreeListCons (k,BetreeMessageDelete) l)) - (Return (BetreeNodeInternal node)) - | BetreeMessageUpsert ufs => - do - (st1,v) <- - betree_internal_lookup_in_children_fwd node - key st0; - (st2,_) <- - betree_node_apply_upserts_fwd - (BetreeListCons - (k,BetreeMessageUpsert ufs) l) v key - st1; - node0 <- - betree_internal_lookup_in_children_back node - key st0; - pending0 <- - betree_node_apply_upserts_back - (BetreeListCons - (k,BetreeMessageUpsert ufs) l) v key - st1; - msgs0 <- - betree_node_lookup_first_message_for_key_back - key msgs pending0; - monad_ignore_bind - (betree_store_internal_node_fwd - node0.betree_internal_id msgs0 st2) - (Return (BetreeNodeInternal node0)) - od) - | BetreeListNil => - monad_ignore_bind - (betree_node_lookup_first_message_for_key_back key msgs - BetreeListNil) - do - node0 <- - betree_internal_lookup_in_children_back node key - st0; - Return (BetreeNodeInternal node0) - od - od - | BetreeNodeLeaf node' => - do - (_,bindings) <- - betree_load_leaf_node_fwd node'.betree_leaf_id st; - monad_ignore_bind - (betree_node_lookup_in_bindings_fwd key bindings) - (Return (BetreeNodeLeaf node')) - od - - [betree_node_lookup_first_message_after_key_back_def] Definition - - ⊢ ∀key msgs ret. - betree_node_lookup_first_message_after_key_back key msgs ret = - case msgs of - BetreeListCons p next_msgs => - (let - (k,m) = p - in - if k = key then - do - next_msgs0 <- - betree_node_lookup_first_message_after_key_back key - next_msgs ret; - Return (BetreeListCons (k,m) next_msgs0) - od - else Return ret) - | BetreeListNil => Return ret - - [betree_node_lookup_first_message_after_key_fwd_def] Definition - - ⊢ ∀key msgs. - betree_node_lookup_first_message_after_key_fwd key msgs = - case msgs of - BetreeListCons p next_msgs => - (let - (k,m) = p - in - if k = key then - betree_node_lookup_first_message_after_key_fwd key - next_msgs - else Return (BetreeListCons (k,m) next_msgs)) - | BetreeListNil => Return BetreeListNil - - [betree_node_lookup_first_message_for_key_back_def] Definition - - ⊢ ∀key msgs ret. - betree_node_lookup_first_message_for_key_back key msgs ret = - case msgs of - BetreeListCons x next_msgs => - (let - (i,m) = x - in - if u64_ge i key then Return ret - else - do - next_msgs0 <- - betree_node_lookup_first_message_for_key_back key - next_msgs ret; - Return (BetreeListCons (i,m) next_msgs0) - od) - | BetreeListNil => Return ret - - [betree_node_lookup_first_message_for_key_fwd_def] Definition - - ⊢ ∀key msgs. - betree_node_lookup_first_message_for_key_fwd key msgs = - case msgs of - BetreeListCons x next_msgs => - (let - (i,m) = x - in - if u64_ge i key then - Return (BetreeListCons (i,m) next_msgs) - else - betree_node_lookup_first_message_for_key_fwd key - next_msgs) - | BetreeListNil => Return BetreeListNil - - [betree_node_lookup_fwd_def] Definition - - ⊢ ∀self key st. - betree_node_lookup_fwd self key st = - case self of - BetreeNodeInternal node => - do - (st0,msgs) <- - betree_load_internal_node_fwd node.betree_internal_id st; - pending <- - betree_node_lookup_first_message_for_key_fwd key msgs; - case pending of - BetreeListCons p l => - (let - (k,msg) = p - in - if k ≠ key then - do - (st1,opt) <- - betree_internal_lookup_in_children_fwd node - key st0; - monad_ignore_bind - (betree_node_lookup_first_message_for_key_back - key msgs (BetreeListCons (k,msg) l)) - (Return (st1,opt)) - od - else - case msg of - BetreeMessageInsert v => - monad_ignore_bind - (betree_node_lookup_first_message_for_key_back - key msgs - (BetreeListCons (k,BetreeMessageInsert v) - l)) (Return (st0,SOME v)) - | BetreeMessageDelete => - monad_ignore_bind - (betree_node_lookup_first_message_for_key_back - key msgs - (BetreeListCons (k,BetreeMessageDelete) l)) - (Return (st0,NONE)) - | BetreeMessageUpsert ufs => - do - (st1,v) <- - betree_internal_lookup_in_children_fwd node - key st0; - (st2,v0) <- - betree_node_apply_upserts_fwd - (BetreeListCons - (k,BetreeMessageUpsert ufs) l) v key - st1; - node0 <- - betree_internal_lookup_in_children_back node - key st0; - pending0 <- - betree_node_apply_upserts_back - (BetreeListCons - (k,BetreeMessageUpsert ufs) l) v key - st1; - msgs0 <- - betree_node_lookup_first_message_for_key_back - key msgs pending0; - (st3,_) <- - betree_store_internal_node_fwd - node0.betree_internal_id msgs0 st2; - Return (st3,SOME v0) - od) - | BetreeListNil => - do - (st1,opt) <- - betree_internal_lookup_in_children_fwd node key st0; - monad_ignore_bind - (betree_node_lookup_first_message_for_key_back key - msgs BetreeListNil) (Return (st1,opt)) - od - od - | BetreeNodeLeaf node' => - do - (st0,bindings) <- - betree_load_leaf_node_fwd node'.betree_leaf_id st; - opt <- betree_node_lookup_in_bindings_fwd key bindings; - Return (st0,opt) - od - - [betree_node_lookup_in_bindings_fwd_def] Definition - - ⊢ ∀key bindings. - betree_node_lookup_in_bindings_fwd key bindings = - case bindings of - BetreeListCons hd tl => - (let - (i,i0) = hd - in - if i = key then Return (SOME i0) - else if u64_gt i key then Return NONE - else betree_node_lookup_in_bindings_fwd key tl) - | BetreeListNil => Return NONE - - [betree_node_lookup_mut_in_bindings_back_def] Definition - - ⊢ ∀key bindings ret. - betree_node_lookup_mut_in_bindings_back key bindings ret = - case bindings of - BetreeListCons hd tl => - (let - (i,i0) = hd - in - if u64_ge i key then Return ret - else - do - tl0 <- - betree_node_lookup_mut_in_bindings_back key tl ret; - Return (BetreeListCons (i,i0) tl0) - od) - | BetreeListNil => Return ret - - [betree_node_lookup_mut_in_bindings_fwd_def] Definition - - ⊢ ∀key bindings. - betree_node_lookup_mut_in_bindings_fwd key bindings = - case bindings of - BetreeListCons hd tl => - (let - (i,i0) = hd - in - if u64_ge i key then Return (BetreeListCons (i,i0) tl) - else betree_node_lookup_mut_in_bindings_fwd key tl) - | BetreeListNil => Return BetreeListNil - - [betree_store_internal_node_fwd_def] Definition - - ⊢ ∀id content st. - betree_store_internal_node_fwd id content st = - do - (st0,_) <- betree_utils_store_internal_node_fwd id content st; - Return (st0,()) - od - - [betree_store_leaf_node_fwd_def] Definition - - ⊢ ∀id content st. - betree_store_leaf_node_fwd id content st = - do - (st0,_) <- betree_utils_store_leaf_node_fwd id content st; - Return (st0,()) - od - - [betree_upsert_update_fwd_def] Definition - - ⊢ ∀prev st. - betree_upsert_update_fwd prev st = - case prev of - NONE => - (case st of - BetreeUpsertFunStateAdd v => Return v - | BetreeUpsertFunStateSub i => Return (int_to_u64 0)) - | SOME prev0 => - case st of - BetreeUpsertFunStateAdd v => - do - margin <- u64_sub core_u64_max prev0; - if u64_ge margin v then u64_add prev0 v - else Return core_u64_max - od - | BetreeUpsertFunStateSub v' => - if u64_ge prev0 v' then u64_sub prev0 v' - else Return (int_to_u64 0) - - [main_fwd_def] Definition - - ⊢ main_fwd = Return () - - -*) -end diff --git a/tests/hol4/betree/betreeMain_OpaqueScript.sml b/tests/hol4/betree/betreeMain_OpaqueScript.sml deleted file mode 100644 index 1d16db4c..00000000 --- a/tests/hol4/betree/betreeMain_OpaqueScript.sml +++ /dev/null @@ -1,26 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [betree_main]: external function declarations *) -open primitivesLib divDefLib -open betreeMain_TypesTheory - -val _ = new_theory "betreeMain_Opaque" - - -(** [betree_main::betree_utils::load_internal_node]: forward function *)val _ = new_constant ("betree_utils_load_internal_node_fwd", - “:u64 -> state -> (state # (u64 # betree_message_t) betree_list_t) - result”) - -(** [betree_main::betree_utils::store_internal_node]: forward function *)val _ = new_constant ("betree_utils_store_internal_node_fwd", - “:u64 -> (u64 # betree_message_t) betree_list_t -> state -> (state # unit) - result”) - -(** [betree_main::betree_utils::load_leaf_node]: forward function *)val _ = new_constant ("betree_utils_load_leaf_node_fwd", - “:u64 -> state -> (state # (u64 # u64) betree_list_t) result”) - -(** [betree_main::betree_utils::store_leaf_node]: forward function *)val _ = new_constant ("betree_utils_store_leaf_node_fwd", - “:u64 -> (u64 # u64) betree_list_t -> state -> (state # unit) result”) - -(** [core::option::Option::{0}::unwrap]: forward function *)val _ = new_constant ("core_option_option_unwrap_fwd", - “:'t option -> state -> (state # 't) result”) - -val _ = export_theory () diff --git a/tests/hol4/betree/betreeMain_OpaqueTheory.sig b/tests/hol4/betree/betreeMain_OpaqueTheory.sig deleted file mode 100644 index da7559a0..00000000 --- a/tests/hol4/betree/betreeMain_OpaqueTheory.sig +++ /dev/null @@ -1,11 +0,0 @@ -signature betreeMain_OpaqueTheory = -sig - type thm = Thm.thm - - val betreeMain_Opaque_grammars : type_grammar.grammar * term_grammar.grammar -(* - [betreeMain_Types] Parent theory of "betreeMain_Opaque" - - -*) -end diff --git a/tests/hol4/betree/betreeMain_TypesScript.sml b/tests/hol4/betree/betreeMain_TypesScript.sml deleted file mode 100644 index 779f6abb..00000000 --- a/tests/hol4/betree/betreeMain_TypesScript.sml +++ /dev/null @@ -1,76 +0,0 @@ -(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) -(** [betree_main]: type definitions *) -open primitivesLib divDefLib - -val _ = new_theory "betreeMain_Types" - - -Datatype: - (** [betree_main::betree::List] *) - betree_list_t = | BetreeListCons 't betree_list_t | BetreeListNil -End - -Datatype: - (** [betree_main::betree::UpsertFunState] *) - betree_upsert_fun_state_t = - | BetreeUpsertFunStateAdd u64 - | BetreeUpsertFunStateSub u64 -End - -Datatype: - (** [betree_main::betree::Message] *) - betree_message_t = - | BetreeMessageInsert u64 - | BetreeMessageDelete - | BetreeMessageUpsert betree_upsert_fun_state_t -End - -Datatype: - (** [betree_main::betree::Leaf] *) - betree_leaf_t = <| betree_leaf_id : u64; betree_leaf_size : u64; |> -End - -Datatype: - (** [betree_main::betree::Internal] *) - betree_internal_t = - <| - betree_internal_id : u64; - betree_internal_pivot : u64; - betree_internal_left : betree_node_t; - betree_internal_right : betree_node_t; - |> - ; - - (** [betree_main::betree::Node] *) - betree_node_t = - | BetreeNodeInternal betree_internal_t - | BetreeNodeLeaf betree_leaf_t -End - -Datatype: - (** [betree_main::betree::Params] *) - betree_params_t = - <| - betree_params_min_flush_size : u64; betree_params_split_size : u64; - |> -End - -Datatype: - (** [betree_main::betree::NodeIdCounter] *) - betree_node_id_counter_t = <| betree_node_id_counter_next_node_id : u64; |> -End - -Datatype: - (** [betree_main::betree::BeTree] *) - betree_be_tree_t = - <| - betree_be_tree_params : betree_params_t; - betree_be_tree_node_id_cnt : betree_node_id_counter_t; - betree_be_tree_root : betree_node_t; - |> -End - -(** The state type used in the state-error monad *) -val _ = new_type ("state", 0) - -val _ = export_theory () diff --git a/tests/hol4/betree/betreeMain_TypesTheory.sig b/tests/hol4/betree/betreeMain_TypesTheory.sig deleted file mode 100644 index cffe6afb..00000000 --- a/tests/hol4/betree/betreeMain_TypesTheory.sig +++ /dev/null @@ -1,1751 +0,0 @@ -signature betreeMain_TypesTheory = -sig - type thm = Thm.thm - - (* Definitions *) - val betree_be_tree_t_TY_DEF : thm - val betree_be_tree_t_betree_be_tree_node_id_cnt : thm - val betree_be_tree_t_betree_be_tree_node_id_cnt_fupd : thm - val betree_be_tree_t_betree_be_tree_params : thm - val betree_be_tree_t_betree_be_tree_params_fupd : thm - val betree_be_tree_t_betree_be_tree_root : thm - val betree_be_tree_t_betree_be_tree_root_fupd : thm - val betree_be_tree_t_case_def : thm - val betree_be_tree_t_size_def : thm - val betree_internal_t_TY_DEF : thm - val betree_internal_t_betree_internal_id : thm - val betree_internal_t_betree_internal_id_fupd : thm - val betree_internal_t_betree_internal_left : thm - val betree_internal_t_betree_internal_left_fupd : thm - val betree_internal_t_betree_internal_pivot : thm - val betree_internal_t_betree_internal_pivot_fupd : thm - val betree_internal_t_betree_internal_right : thm - val betree_internal_t_betree_internal_right_fupd : thm - val betree_internal_t_case_def : thm - val betree_internal_t_size_def : thm - val betree_leaf_t_TY_DEF : thm - val betree_leaf_t_betree_leaf_id : thm - val betree_leaf_t_betree_leaf_id_fupd : thm - val betree_leaf_t_betree_leaf_size : thm - val betree_leaf_t_betree_leaf_size_fupd : thm - val betree_leaf_t_case_def : thm - val betree_leaf_t_size_def : thm - val betree_list_t_TY_DEF : thm - val betree_list_t_case_def : thm - val betree_list_t_size_def : thm - val betree_message_t_TY_DEF : thm - val betree_message_t_case_def : thm - val betree_message_t_size_def : thm - val betree_node_id_counter_t_TY_DEF : thm - val betree_node_id_counter_t_betree_node_id_counter_next_node_id : thm - val betree_node_id_counter_t_betree_node_id_counter_next_node_id_fupd : thm - val betree_node_id_counter_t_case_def : thm - val betree_node_id_counter_t_size_def : thm - val betree_node_t_TY_DEF : thm - val betree_node_t_case_def : thm - val betree_params_t_TY_DEF : thm - val betree_params_t_betree_params_min_flush_size : thm - val betree_params_t_betree_params_min_flush_size_fupd : thm - val betree_params_t_betree_params_split_size : thm - val betree_params_t_betree_params_split_size_fupd : thm - val betree_params_t_case_def : thm - val betree_params_t_size_def : thm - val betree_upsert_fun_state_t_TY_DEF : thm - val betree_upsert_fun_state_t_case_def : thm - val betree_upsert_fun_state_t_size_def : thm - - (* Theorems *) - val EXISTS_betree_be_tree_t : thm - val EXISTS_betree_internal_t : thm - val EXISTS_betree_leaf_t : thm - val EXISTS_betree_node_id_counter_t : thm - val EXISTS_betree_params_t : thm - val FORALL_betree_be_tree_t : thm - val FORALL_betree_internal_t : thm - val FORALL_betree_leaf_t : thm - val FORALL_betree_node_id_counter_t : thm - val FORALL_betree_params_t : thm - val betree_be_tree_t_11 : thm - val betree_be_tree_t_Axiom : thm - val betree_be_tree_t_accessors : thm - val betree_be_tree_t_accfupds : thm - val betree_be_tree_t_case_cong : thm - val betree_be_tree_t_case_eq : thm - val betree_be_tree_t_component_equality : thm - val betree_be_tree_t_fn_updates : thm - val betree_be_tree_t_fupdcanon : thm - val betree_be_tree_t_fupdcanon_comp : thm - val betree_be_tree_t_fupdfupds : thm - val betree_be_tree_t_fupdfupds_comp : thm - val betree_be_tree_t_induction : thm - val betree_be_tree_t_literal_11 : thm - val betree_be_tree_t_literal_nchotomy : thm - val betree_be_tree_t_nchotomy : thm - val betree_be_tree_t_updates_eq_literal : thm - val betree_internal_t_11 : thm - val betree_internal_t_Axiom : thm - val betree_internal_t_accessors : thm - val betree_internal_t_accfupds : thm - val betree_internal_t_case_cong : thm - val betree_internal_t_case_eq : thm - val betree_internal_t_component_equality : thm - val betree_internal_t_fn_updates : thm - val betree_internal_t_fupdcanon : thm - val betree_internal_t_fupdcanon_comp : thm - val betree_internal_t_fupdfupds : thm - val betree_internal_t_fupdfupds_comp : thm - val betree_internal_t_induction : thm - val betree_internal_t_literal_11 : thm - val betree_internal_t_literal_nchotomy : thm - val betree_internal_t_nchotomy : thm - val betree_internal_t_updates_eq_literal : thm - val betree_leaf_t_11 : thm - val betree_leaf_t_Axiom : thm - val betree_leaf_t_accessors : thm - val betree_leaf_t_accfupds : thm - val betree_leaf_t_case_cong : thm - val betree_leaf_t_case_eq : thm - val betree_leaf_t_component_equality : thm - val betree_leaf_t_fn_updates : thm - val betree_leaf_t_fupdcanon : thm - val betree_leaf_t_fupdcanon_comp : thm - val betree_leaf_t_fupdfupds : thm - val betree_leaf_t_fupdfupds_comp : thm - val betree_leaf_t_induction : thm - val betree_leaf_t_literal_11 : thm - val betree_leaf_t_literal_nchotomy : thm - val betree_leaf_t_nchotomy : thm - val betree_leaf_t_updates_eq_literal : thm - val betree_list_t_11 : thm - val betree_list_t_Axiom : thm - val betree_list_t_case_cong : thm - val betree_list_t_case_eq : thm - val betree_list_t_distinct : thm - val betree_list_t_induction : thm - val betree_list_t_nchotomy : thm - val betree_message_t_11 : thm - val betree_message_t_Axiom : thm - val betree_message_t_case_cong : thm - val betree_message_t_case_eq : thm - val betree_message_t_distinct : thm - val betree_message_t_induction : thm - val betree_message_t_nchotomy : thm - val betree_node_id_counter_t_11 : thm - val betree_node_id_counter_t_Axiom : thm - val betree_node_id_counter_t_accessors : thm - val betree_node_id_counter_t_accfupds : thm - val betree_node_id_counter_t_case_cong : thm - val betree_node_id_counter_t_case_eq : thm - val betree_node_id_counter_t_component_equality : thm - val betree_node_id_counter_t_fn_updates : thm - val betree_node_id_counter_t_fupdfupds : thm - val betree_node_id_counter_t_fupdfupds_comp : thm - val betree_node_id_counter_t_induction : thm - val betree_node_id_counter_t_literal_11 : thm - val betree_node_id_counter_t_literal_nchotomy : thm - val betree_node_id_counter_t_nchotomy : thm - val betree_node_id_counter_t_updates_eq_literal : thm - val betree_node_t_11 : thm - val betree_node_t_Axiom : thm - val betree_node_t_case_cong : thm - val betree_node_t_case_eq : thm - val betree_node_t_distinct : thm - val betree_node_t_induction : thm - val betree_node_t_nchotomy : thm - val betree_params_t_11 : thm - val betree_params_t_Axiom : thm - val betree_params_t_accessors : thm - val betree_params_t_accfupds : thm - val betree_params_t_case_cong : thm - val betree_params_t_case_eq : thm - val betree_params_t_component_equality : thm - val betree_params_t_fn_updates : thm - val betree_params_t_fupdcanon : thm - val betree_params_t_fupdcanon_comp : thm - val betree_params_t_fupdfupds : thm - val betree_params_t_fupdfupds_comp : thm - val betree_params_t_induction : thm - val betree_params_t_literal_11 : thm - val betree_params_t_literal_nchotomy : thm - val betree_params_t_nchotomy : thm - val betree_params_t_updates_eq_literal : thm - val betree_upsert_fun_state_t_11 : thm - val betree_upsert_fun_state_t_Axiom : thm - val betree_upsert_fun_state_t_case_cong : thm - val betree_upsert_fun_state_t_case_eq : thm - val betree_upsert_fun_state_t_distinct : thm - val betree_upsert_fun_state_t_induction : thm - val betree_upsert_fun_state_t_nchotomy : thm - val datatype_betree_be_tree_t : thm - val datatype_betree_internal_t : thm - val datatype_betree_leaf_t : thm - val datatype_betree_list_t : thm - val datatype_betree_message_t : thm - val datatype_betree_node_id_counter_t : thm - val datatype_betree_params_t : thm - val datatype_betree_upsert_fun_state_t : thm - - val betreeMain_Types_grammars : type_grammar.grammar * term_grammar.grammar -(* - [divDef] Parent theory of "betreeMain_Types" - - [betree_be_tree_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('betree_be_tree_t'). - (∀a0'. - (∃a0 a1 a2. - a0' = - (λa0 a1 a2. - ind_type$CONSTR 0 (a0,a1,a2) - (λn. ind_type$BOTTOM)) a0 a1 a2) ⇒ - $var$('betree_be_tree_t') a0') ⇒ - $var$('betree_be_tree_t') a0') rep - - [betree_be_tree_t_betree_be_tree_node_id_cnt] Definition - - ⊢ ∀b b0 b1. - (betree_be_tree_t b b0 b1).betree_be_tree_node_id_cnt = b0 - - [betree_be_tree_t_betree_be_tree_node_id_cnt_fupd] Definition - - ⊢ ∀f b b0 b1. - betree_be_tree_t b b0 b1 with - betree_be_tree_node_id_cnt updated_by f = - betree_be_tree_t b (f b0) b1 - - [betree_be_tree_t_betree_be_tree_params] Definition - - ⊢ ∀b b0 b1. (betree_be_tree_t b b0 b1).betree_be_tree_params = b - - [betree_be_tree_t_betree_be_tree_params_fupd] Definition - - ⊢ ∀f b b0 b1. - betree_be_tree_t b b0 b1 with betree_be_tree_params updated_by f = - betree_be_tree_t (f b) b0 b1 - - [betree_be_tree_t_betree_be_tree_root] Definition - - ⊢ ∀b b0 b1. (betree_be_tree_t b b0 b1).betree_be_tree_root = b1 - - [betree_be_tree_t_betree_be_tree_root_fupd] Definition - - ⊢ ∀f b b0 b1. - betree_be_tree_t b b0 b1 with betree_be_tree_root updated_by f = - betree_be_tree_t b b0 (f b1) - - [betree_be_tree_t_case_def] Definition - - ⊢ ∀a0 a1 a2 f. - betree_be_tree_t_CASE (betree_be_tree_t a0 a1 a2) f = f a0 a1 a2 - - [betree_be_tree_t_size_def] Definition - - ⊢ ∀a0 a1 a2. - betree_be_tree_t_size (betree_be_tree_t a0 a1 a2) = - 1 + - (betree_params_t_size a0 + - (betree_node_id_counter_t_size a1 + betree_node_t_size a2)) - - [betree_internal_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('betree_internal_t') $var$('betree_node_t'). - (∀a0'. - (∃a0 a1 a2 a3. - a0' = - (λa0 a1 a2 a3. - ind_type$CONSTR 0 (a0,a1,ARB) - (ind_type$FCONS a2 - (ind_type$FCONS a3 (λn. ind_type$BOTTOM)))) - a0 a1 a2 a3 ∧ $var$('betree_node_t') a2 ∧ - $var$('betree_node_t') a3) ⇒ - $var$('betree_internal_t') a0') ∧ - (∀a1'. - (∃a. a1' = - (λa. - ind_type$CONSTR (SUC 0) (ARB,ARB,ARB) - (ind_type$FCONS a (λn. ind_type$BOTTOM))) - a ∧ $var$('betree_internal_t') a) ∨ - (∃a. a1' = - (λa. - ind_type$CONSTR (SUC (SUC 0)) (ARB,ARB,a) - (λn. ind_type$BOTTOM)) a) ⇒ - $var$('betree_node_t') a1') ⇒ - $var$('betree_internal_t') a0') rep - - [betree_internal_t_betree_internal_id] Definition - - ⊢ ∀u u0 b b0. (betree_internal_t u u0 b b0).betree_internal_id = u - - [betree_internal_t_betree_internal_id_fupd] Definition - - ⊢ ∀f u u0 b b0. - betree_internal_t u u0 b b0 with betree_internal_id updated_by f = - betree_internal_t (f u) u0 b b0 - - [betree_internal_t_betree_internal_left] Definition - - ⊢ ∀u u0 b b0. (betree_internal_t u u0 b b0).betree_internal_left = b - - [betree_internal_t_betree_internal_left_fupd] Definition - - ⊢ ∀f u u0 b b0. - betree_internal_t u u0 b b0 with - betree_internal_left updated_by f = - betree_internal_t u u0 (f b) b0 - - [betree_internal_t_betree_internal_pivot] Definition - - ⊢ ∀u u0 b b0. - (betree_internal_t u u0 b b0).betree_internal_pivot = u0 - - [betree_internal_t_betree_internal_pivot_fupd] Definition - - ⊢ ∀f u u0 b b0. - betree_internal_t u u0 b b0 with - betree_internal_pivot updated_by f = - betree_internal_t u (f u0) b b0 - - [betree_internal_t_betree_internal_right] Definition - - ⊢ ∀u u0 b b0. - (betree_internal_t u u0 b b0).betree_internal_right = b0 - - [betree_internal_t_betree_internal_right_fupd] Definition - - ⊢ ∀f u u0 b b0. - betree_internal_t u u0 b b0 with - betree_internal_right updated_by f = - betree_internal_t u u0 b (f b0) - - [betree_internal_t_case_def] Definition - - ⊢ ∀a0 a1 a2 a3 f. - betree_internal_t_CASE (betree_internal_t a0 a1 a2 a3) f = - f a0 a1 a2 a3 - - [betree_internal_t_size_def] Definition - - ⊢ (∀a0 a1 a2 a3. - betree_internal_t_size (betree_internal_t a0 a1 a2 a3) = - 1 + (betree_node_t_size a2 + betree_node_t_size a3)) ∧ - (∀a. betree_node_t_size (BetreeNodeInternal a) = - 1 + betree_internal_t_size a) ∧ - ∀a. betree_node_t_size (BetreeNodeLeaf a) = - 1 + betree_leaf_t_size a - - [betree_leaf_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('betree_leaf_t'). - (∀a0'. - (∃a0 a1. - a0' = - (λa0 a1. - ind_type$CONSTR 0 (a0,a1) - (λn. ind_type$BOTTOM)) a0 a1) ⇒ - $var$('betree_leaf_t') a0') ⇒ - $var$('betree_leaf_t') a0') rep - - [betree_leaf_t_betree_leaf_id] Definition - - ⊢ ∀u u0. (betree_leaf_t u u0).betree_leaf_id = u - - [betree_leaf_t_betree_leaf_id_fupd] Definition - - ⊢ ∀f u u0. - betree_leaf_t u u0 with betree_leaf_id updated_by f = - betree_leaf_t (f u) u0 - - [betree_leaf_t_betree_leaf_size] Definition - - ⊢ ∀u u0. (betree_leaf_t u u0).betree_leaf_size = u0 - - [betree_leaf_t_betree_leaf_size_fupd] Definition - - ⊢ ∀f u u0. - betree_leaf_t u u0 with betree_leaf_size updated_by f = - betree_leaf_t u (f u0) - - [betree_leaf_t_case_def] Definition - - ⊢ ∀a0 a1 f. betree_leaf_t_CASE (betree_leaf_t a0 a1) f = f a0 a1 - - [betree_leaf_t_size_def] Definition - - ⊢ ∀a0 a1. betree_leaf_t_size (betree_leaf_t a0 a1) = 1 - - [betree_list_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('betree_list_t'). - (∀a0'. - (∃a0 a1. - a0' = - (λa0 a1. - ind_type$CONSTR 0 a0 - (ind_type$FCONS a1 (λn. ind_type$BOTTOM))) - a0 a1 ∧ $var$('betree_list_t') a1) ∨ - a0' = - ind_type$CONSTR (SUC 0) ARB (λn. ind_type$BOTTOM) ⇒ - $var$('betree_list_t') a0') ⇒ - $var$('betree_list_t') a0') rep - - [betree_list_t_case_def] Definition - - ⊢ (∀a0 a1 f v. - betree_list_t_CASE (BetreeListCons a0 a1) f v = f a0 a1) ∧ - ∀f v. betree_list_t_CASE BetreeListNil f v = v - - [betree_list_t_size_def] Definition - - ⊢ (∀f a0 a1. - betree_list_t_size f (BetreeListCons a0 a1) = - 1 + (f a0 + betree_list_t_size f a1)) ∧ - ∀f. betree_list_t_size f BetreeListNil = 0 - - [betree_message_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0. - ∀ $var$('betree_message_t'). - (∀a0. - (∃a. a0 = - (λa. - ind_type$CONSTR 0 (a,ARB) - (λn. ind_type$BOTTOM)) a) ∨ - a0 = - ind_type$CONSTR (SUC 0) (ARB,ARB) - (λn. ind_type$BOTTOM) ∨ - (∃a. a0 = - (λa. - ind_type$CONSTR (SUC (SUC 0)) (ARB,a) - (λn. ind_type$BOTTOM)) a) ⇒ - $var$('betree_message_t') a0) ⇒ - $var$('betree_message_t') a0) rep - - [betree_message_t_case_def] Definition - - ⊢ (∀a f v f1. - betree_message_t_CASE (BetreeMessageInsert a) f v f1 = f a) ∧ - (∀f v f1. betree_message_t_CASE BetreeMessageDelete f v f1 = v) ∧ - ∀a f v f1. - betree_message_t_CASE (BetreeMessageUpsert a) f v f1 = f1 a - - [betree_message_t_size_def] Definition - - ⊢ (∀a. betree_message_t_size (BetreeMessageInsert a) = 1) ∧ - betree_message_t_size BetreeMessageDelete = 0 ∧ - ∀a. betree_message_t_size (BetreeMessageUpsert a) = - 1 + betree_upsert_fun_state_t_size a - - [betree_node_id_counter_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0. - ∀ $var$('betree_node_id_counter_t'). - (∀a0. - (∃a. a0 = - (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) - a) ⇒ - $var$('betree_node_id_counter_t') a0) ⇒ - $var$('betree_node_id_counter_t') a0) rep - - [betree_node_id_counter_t_betree_node_id_counter_next_node_id] Definition - - ⊢ ∀u. (betree_node_id_counter_t u). - betree_node_id_counter_next_node_id = - u - - [betree_node_id_counter_t_betree_node_id_counter_next_node_id_fupd] Definition - - ⊢ ∀f u. - betree_node_id_counter_t u with - betree_node_id_counter_next_node_id updated_by f = - betree_node_id_counter_t (f u) - - [betree_node_id_counter_t_case_def] Definition - - ⊢ ∀a f. - betree_node_id_counter_t_CASE (betree_node_id_counter_t a) f = - f a - - [betree_node_id_counter_t_size_def] Definition - - ⊢ ∀a. betree_node_id_counter_t_size (betree_node_id_counter_t a) = 1 - - [betree_node_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa1'. - ∀ $var$('betree_internal_t') $var$('betree_node_t'). - (∀a0'. - (∃a0 a1 a2 a3. - a0' = - (λa0 a1 a2 a3. - ind_type$CONSTR 0 (a0,a1,ARB) - (ind_type$FCONS a2 - (ind_type$FCONS a3 (λn. ind_type$BOTTOM)))) - a0 a1 a2 a3 ∧ $var$('betree_node_t') a2 ∧ - $var$('betree_node_t') a3) ⇒ - $var$('betree_internal_t') a0') ∧ - (∀a1'. - (∃a. a1' = - (λa. - ind_type$CONSTR (SUC 0) (ARB,ARB,ARB) - (ind_type$FCONS a (λn. ind_type$BOTTOM))) - a ∧ $var$('betree_internal_t') a) ∨ - (∃a. a1' = - (λa. - ind_type$CONSTR (SUC (SUC 0)) (ARB,ARB,a) - (λn. ind_type$BOTTOM)) a) ⇒ - $var$('betree_node_t') a1') ⇒ - $var$('betree_node_t') a1') rep - - [betree_node_t_case_def] Definition - - ⊢ (∀a f f1. betree_node_t_CASE (BetreeNodeInternal a) f f1 = f a) ∧ - ∀a f f1. betree_node_t_CASE (BetreeNodeLeaf a) f f1 = f1 a - - [betree_params_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0'. - ∀ $var$('betree_params_t'). - (∀a0'. - (∃a0 a1. - a0' = - (λa0 a1. - ind_type$CONSTR 0 (a0,a1) - (λn. ind_type$BOTTOM)) a0 a1) ⇒ - $var$('betree_params_t') a0') ⇒ - $var$('betree_params_t') a0') rep - - [betree_params_t_betree_params_min_flush_size] Definition - - ⊢ ∀u u0. (betree_params_t u u0).betree_params_min_flush_size = u - - [betree_params_t_betree_params_min_flush_size_fupd] Definition - - ⊢ ∀f u u0. - betree_params_t u u0 with - betree_params_min_flush_size updated_by f = - betree_params_t (f u) u0 - - [betree_params_t_betree_params_split_size] Definition - - ⊢ ∀u u0. (betree_params_t u u0).betree_params_split_size = u0 - - [betree_params_t_betree_params_split_size_fupd] Definition - - ⊢ ∀f u u0. - betree_params_t u u0 with betree_params_split_size updated_by f = - betree_params_t u (f u0) - - [betree_params_t_case_def] Definition - - ⊢ ∀a0 a1 f. betree_params_t_CASE (betree_params_t a0 a1) f = f a0 a1 - - [betree_params_t_size_def] Definition - - ⊢ ∀a0 a1. betree_params_t_size (betree_params_t a0 a1) = 1 - - [betree_upsert_fun_state_t_TY_DEF] Definition - - ⊢ ∃rep. - TYPE_DEFINITION - (λa0. - ∀ $var$('betree_upsert_fun_state_t'). - (∀a0. - (∃a. a0 = - (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) - a) ∨ - (∃a. a0 = - (λa. - ind_type$CONSTR (SUC 0) a - (λn. ind_type$BOTTOM)) a) ⇒ - $var$('betree_upsert_fun_state_t') a0) ⇒ - $var$('betree_upsert_fun_state_t') a0) rep - - [betree_upsert_fun_state_t_case_def] Definition - - ⊢ (∀a f f1. - betree_upsert_fun_state_t_CASE (BetreeUpsertFunStateAdd a) f f1 = - f a) ∧ - ∀a f f1. - betree_upsert_fun_state_t_CASE (BetreeUpsertFunStateSub a) f f1 = - f1 a - - [betree_upsert_fun_state_t_size_def] Definition - - ⊢ (∀a. betree_upsert_fun_state_t_size (BetreeUpsertFunStateAdd a) = 1) ∧ - ∀a. betree_upsert_fun_state_t_size (BetreeUpsertFunStateSub a) = 1 - - [EXISTS_betree_be_tree_t] Theorem - - ⊢ ∀P. (∃b. P b) ⇔ - ∃b2 b1 b0. - P - <|betree_be_tree_params := b2; - betree_be_tree_node_id_cnt := b1; - betree_be_tree_root := b0|> - - [EXISTS_betree_internal_t] Theorem - - ⊢ ∀P. (∃b. P b) ⇔ - ∃u0 u b1 b0. - P - <|betree_internal_id := u0; betree_internal_pivot := u; - betree_internal_left := b1; betree_internal_right := b0|> - - [EXISTS_betree_leaf_t] Theorem - - ⊢ ∀P. (∃b. P b) ⇔ - ∃u0 u. P <|betree_leaf_id := u0; betree_leaf_size := u|> - - [EXISTS_betree_node_id_counter_t] Theorem - - ⊢ ∀P. (∃b. P b) ⇔ ∃u. P <|betree_node_id_counter_next_node_id := u|> - - [EXISTS_betree_params_t] Theorem - - ⊢ ∀P. (∃b. P b) ⇔ - ∃u0 u. - P - <|betree_params_min_flush_size := u0; - betree_params_split_size := u|> - - [FORALL_betree_be_tree_t] Theorem - - ⊢ ∀P. (∀b. P b) ⇔ - ∀b2 b1 b0. - P - <|betree_be_tree_params := b2; - betree_be_tree_node_id_cnt := b1; - betree_be_tree_root := b0|> - - [FORALL_betree_internal_t] Theorem - - ⊢ ∀P. (∀b. P b) ⇔ - ∀u0 u b1 b0. - P - <|betree_internal_id := u0; betree_internal_pivot := u; - betree_internal_left := b1; betree_internal_right := b0|> - - [FORALL_betree_leaf_t] Theorem - - ⊢ ∀P. (∀b. P b) ⇔ - ∀u0 u. P <|betree_leaf_id := u0; betree_leaf_size := u|> - - [FORALL_betree_node_id_counter_t] Theorem - - ⊢ ∀P. (∀b. P b) ⇔ ∀u. P <|betree_node_id_counter_next_node_id := u|> - - [FORALL_betree_params_t] Theorem - - ⊢ ∀P. (∀b. P b) ⇔ - ∀u0 u. - P - <|betree_params_min_flush_size := u0; - betree_params_split_size := u|> - - [betree_be_tree_t_11] Theorem - - ⊢ ∀a0 a1 a2 a0' a1' a2'. - betree_be_tree_t a0 a1 a2 = betree_be_tree_t a0' a1' a2' ⇔ - a0 = a0' ∧ a1 = a1' ∧ a2 = a2' - - [betree_be_tree_t_Axiom] Theorem - - ⊢ ∀f. ∃fn. ∀a0 a1 a2. fn (betree_be_tree_t a0 a1 a2) = f a0 a1 a2 - - [betree_be_tree_t_accessors] Theorem - - ⊢ (∀b b0 b1. (betree_be_tree_t b b0 b1).betree_be_tree_params = b) ∧ - (∀b b0 b1. - (betree_be_tree_t b b0 b1).betree_be_tree_node_id_cnt = b0) ∧ - ∀b b0 b1. (betree_be_tree_t b b0 b1).betree_be_tree_root = b1 - - [betree_be_tree_t_accfupds] Theorem - - ⊢ (∀f b. - (b with betree_be_tree_node_id_cnt updated_by f). - betree_be_tree_params = - b.betree_be_tree_params) ∧ - (∀f b. - (b with betree_be_tree_root updated_by f).betree_be_tree_params = - b.betree_be_tree_params) ∧ - (∀f b. - (b with betree_be_tree_params updated_by f). - betree_be_tree_node_id_cnt = - b.betree_be_tree_node_id_cnt) ∧ - (∀f b. - (b with betree_be_tree_root updated_by f). - betree_be_tree_node_id_cnt = - b.betree_be_tree_node_id_cnt) ∧ - (∀f b. - (b with betree_be_tree_params updated_by f).betree_be_tree_root = - b.betree_be_tree_root) ∧ - (∀f b. - (b with betree_be_tree_node_id_cnt updated_by f). - betree_be_tree_root = - b.betree_be_tree_root) ∧ - (∀f b. - (b with betree_be_tree_params updated_by f). - betree_be_tree_params = - f b.betree_be_tree_params) ∧ - (∀f b. - (b with betree_be_tree_node_id_cnt updated_by f). - betree_be_tree_node_id_cnt = - f b.betree_be_tree_node_id_cnt) ∧ - ∀f b. - (b with betree_be_tree_root updated_by f).betree_be_tree_root = - f b.betree_be_tree_root - - [betree_be_tree_t_case_cong] Theorem - - ⊢ ∀M M' f. - M = M' ∧ - (∀a0 a1 a2. - M' = betree_be_tree_t a0 a1 a2 ⇒ f a0 a1 a2 = f' a0 a1 a2) ⇒ - betree_be_tree_t_CASE M f = betree_be_tree_t_CASE M' f' - - [betree_be_tree_t_case_eq] Theorem - - ⊢ betree_be_tree_t_CASE x f = v ⇔ - ∃b b0 b1. x = betree_be_tree_t b b0 b1 ∧ f b b0 b1 = v - - [betree_be_tree_t_component_equality] Theorem - - ⊢ ∀b1 b2. - b1 = b2 ⇔ - b1.betree_be_tree_params = b2.betree_be_tree_params ∧ - b1.betree_be_tree_node_id_cnt = b2.betree_be_tree_node_id_cnt ∧ - b1.betree_be_tree_root = b2.betree_be_tree_root - - [betree_be_tree_t_fn_updates] Theorem - - ⊢ (∀f b b0 b1. - betree_be_tree_t b b0 b1 with betree_be_tree_params updated_by f = - betree_be_tree_t (f b) b0 b1) ∧ - (∀f b b0 b1. - betree_be_tree_t b b0 b1 with - betree_be_tree_node_id_cnt updated_by f = - betree_be_tree_t b (f b0) b1) ∧ - ∀f b b0 b1. - betree_be_tree_t b b0 b1 with betree_be_tree_root updated_by f = - betree_be_tree_t b b0 (f b1) - - [betree_be_tree_t_fupdcanon] Theorem - - ⊢ (∀g f b. - b with - <|betree_be_tree_node_id_cnt updated_by f; - betree_be_tree_params updated_by g|> = - b with - <|betree_be_tree_params updated_by g; - betree_be_tree_node_id_cnt updated_by f|>) ∧ - (∀g f b. - b with - <|betree_be_tree_root updated_by f; - betree_be_tree_params updated_by g|> = - b with - <|betree_be_tree_params updated_by g; - betree_be_tree_root updated_by f|>) ∧ - ∀g f b. - b with - <|betree_be_tree_root updated_by f; - betree_be_tree_node_id_cnt updated_by g|> = - b with - <|betree_be_tree_node_id_cnt updated_by g; - betree_be_tree_root updated_by f|> - - [betree_be_tree_t_fupdcanon_comp] Theorem - - ⊢ ((∀g f. - betree_be_tree_node_id_cnt_fupd f ∘ - betree_be_tree_params_fupd g = - betree_be_tree_params_fupd g ∘ - betree_be_tree_node_id_cnt_fupd f) ∧ - ∀h g f. - betree_be_tree_node_id_cnt_fupd f ∘ - betree_be_tree_params_fupd g ∘ h = - betree_be_tree_params_fupd g ∘ - betree_be_tree_node_id_cnt_fupd f ∘ h) ∧ - ((∀g f. - betree_be_tree_root_fupd f ∘ betree_be_tree_params_fupd g = - betree_be_tree_params_fupd g ∘ betree_be_tree_root_fupd f) ∧ - ∀h g f. - betree_be_tree_root_fupd f ∘ betree_be_tree_params_fupd g ∘ h = - betree_be_tree_params_fupd g ∘ betree_be_tree_root_fupd f ∘ h) ∧ - (∀g f. - betree_be_tree_root_fupd f ∘ betree_be_tree_node_id_cnt_fupd g = - betree_be_tree_node_id_cnt_fupd g ∘ betree_be_tree_root_fupd f) ∧ - ∀h g f. - betree_be_tree_root_fupd f ∘ betree_be_tree_node_id_cnt_fupd g ∘ - h = - betree_be_tree_node_id_cnt_fupd g ∘ betree_be_tree_root_fupd f ∘ - h - - [betree_be_tree_t_fupdfupds] Theorem - - ⊢ (∀g f b. - b with - <|betree_be_tree_params updated_by f; - betree_be_tree_params updated_by g|> = - b with betree_be_tree_params updated_by f ∘ g) ∧ - (∀g f b. - b with - <|betree_be_tree_node_id_cnt updated_by f; - betree_be_tree_node_id_cnt updated_by g|> = - b with betree_be_tree_node_id_cnt updated_by f ∘ g) ∧ - ∀g f b. - b with - <|betree_be_tree_root updated_by f; - betree_be_tree_root updated_by g|> = - b with betree_be_tree_root updated_by f ∘ g - - [betree_be_tree_t_fupdfupds_comp] Theorem - - ⊢ ((∀g f. - betree_be_tree_params_fupd f ∘ betree_be_tree_params_fupd g = - betree_be_tree_params_fupd (f ∘ g)) ∧ - ∀h g f. - betree_be_tree_params_fupd f ∘ betree_be_tree_params_fupd g ∘ h = - betree_be_tree_params_fupd (f ∘ g) ∘ h) ∧ - ((∀g f. - betree_be_tree_node_id_cnt_fupd f ∘ - betree_be_tree_node_id_cnt_fupd g = - betree_be_tree_node_id_cnt_fupd (f ∘ g)) ∧ - ∀h g f. - betree_be_tree_node_id_cnt_fupd f ∘ - betree_be_tree_node_id_cnt_fupd g ∘ h = - betree_be_tree_node_id_cnt_fupd (f ∘ g) ∘ h) ∧ - (∀g f. - betree_be_tree_root_fupd f ∘ betree_be_tree_root_fupd g = - betree_be_tree_root_fupd (f ∘ g)) ∧ - ∀h g f. - betree_be_tree_root_fupd f ∘ betree_be_tree_root_fupd g ∘ h = - betree_be_tree_root_fupd (f ∘ g) ∘ h - - [betree_be_tree_t_induction] Theorem - - ⊢ ∀P. (∀b b0 b1. P (betree_be_tree_t b b0 b1)) ⇒ ∀b. P b - - [betree_be_tree_t_literal_11] Theorem - - ⊢ ∀b21 b11 b01 b22 b12 b02. - <|betree_be_tree_params := b21; - betree_be_tree_node_id_cnt := b11; betree_be_tree_root := b01|> = - <|betree_be_tree_params := b22; - betree_be_tree_node_id_cnt := b12; betree_be_tree_root := b02|> ⇔ - b21 = b22 ∧ b11 = b12 ∧ b01 = b02 - - [betree_be_tree_t_literal_nchotomy] Theorem - - ⊢ ∀b. ∃b2 b1 b0. - b = - <|betree_be_tree_params := b2; betree_be_tree_node_id_cnt := b1; - betree_be_tree_root := b0|> - - [betree_be_tree_t_nchotomy] Theorem - - ⊢ ∀bb. ∃b b0 b1. bb = betree_be_tree_t b b0 b1 - - [betree_be_tree_t_updates_eq_literal] Theorem - - ⊢ ∀b b2 b1 b0. - b with - <|betree_be_tree_params := b2; betree_be_tree_node_id_cnt := b1; - betree_be_tree_root := b0|> = - <|betree_be_tree_params := b2; betree_be_tree_node_id_cnt := b1; - betree_be_tree_root := b0|> - - [betree_internal_t_11] Theorem - - ⊢ ∀a0 a1 a2 a3 a0' a1' a2' a3'. - betree_internal_t a0 a1 a2 a3 = betree_internal_t a0' a1' a2' a3' ⇔ - a0 = a0' ∧ a1 = a1' ∧ a2 = a2' ∧ a3 = a3' - - [betree_internal_t_Axiom] Theorem - - ⊢ ∀f0 f1 f2. ∃fn0 fn1. - (∀a0 a1 a2 a3. - fn0 (betree_internal_t a0 a1 a2 a3) = - f0 a0 a1 a2 a3 (fn1 a2) (fn1 a3)) ∧ - (∀a. fn1 (BetreeNodeInternal a) = f1 a (fn0 a)) ∧ - ∀a. fn1 (BetreeNodeLeaf a) = f2 a - - [betree_internal_t_accessors] Theorem - - ⊢ (∀u u0 b b0. (betree_internal_t u u0 b b0).betree_internal_id = u) ∧ - (∀u u0 b b0. - (betree_internal_t u u0 b b0).betree_internal_pivot = u0) ∧ - (∀u u0 b b0. (betree_internal_t u u0 b b0).betree_internal_left = b) ∧ - ∀u u0 b b0. - (betree_internal_t u u0 b b0).betree_internal_right = b0 - - [betree_internal_t_accfupds] Theorem - - ⊢ (∀f b. - (b with betree_internal_pivot updated_by f).betree_internal_id = - b.betree_internal_id) ∧ - (∀f b. - (b with betree_internal_left updated_by f).betree_internal_id = - b.betree_internal_id) ∧ - (∀f b. - (b with betree_internal_right updated_by f).betree_internal_id = - b.betree_internal_id) ∧ - (∀f b. - (b with betree_internal_id updated_by f).betree_internal_pivot = - b.betree_internal_pivot) ∧ - (∀f b. - (b with betree_internal_left updated_by f).betree_internal_pivot = - b.betree_internal_pivot) ∧ - (∀f b. - (b with betree_internal_right updated_by f). - betree_internal_pivot = - b.betree_internal_pivot) ∧ - (∀f b. - (b with betree_internal_id updated_by f).betree_internal_left = - b.betree_internal_left) ∧ - (∀f b. - (b with betree_internal_pivot updated_by f).betree_internal_left = - b.betree_internal_left) ∧ - (∀f b. - (b with betree_internal_right updated_by f).betree_internal_left = - b.betree_internal_left) ∧ - (∀f b. - (b with betree_internal_id updated_by f).betree_internal_right = - b.betree_internal_right) ∧ - (∀f b. - (b with betree_internal_pivot updated_by f). - betree_internal_right = - b.betree_internal_right) ∧ - (∀f b. - (b with betree_internal_left updated_by f).betree_internal_right = - b.betree_internal_right) ∧ - (∀f b. - (b with betree_internal_id updated_by f).betree_internal_id = - f b.betree_internal_id) ∧ - (∀f b. - (b with betree_internal_pivot updated_by f). - betree_internal_pivot = - f b.betree_internal_pivot) ∧ - (∀f b. - (b with betree_internal_left updated_by f).betree_internal_left = - f b.betree_internal_left) ∧ - ∀f b. - (b with betree_internal_right updated_by f).betree_internal_right = - f b.betree_internal_right - - [betree_internal_t_case_cong] Theorem - - ⊢ ∀M M' f. - M = M' ∧ - (∀a0 a1 a2 a3. - M' = betree_internal_t a0 a1 a2 a3 ⇒ - f a0 a1 a2 a3 = f' a0 a1 a2 a3) ⇒ - betree_internal_t_CASE M f = betree_internal_t_CASE M' f' - - [betree_internal_t_case_eq] Theorem - - ⊢ betree_internal_t_CASE x f = v ⇔ - ∃u0 u b b0. x = betree_internal_t u0 u b b0 ∧ f u0 u b b0 = v - - [betree_internal_t_component_equality] Theorem - - ⊢ ∀b1 b2. - b1 = b2 ⇔ - b1.betree_internal_id = b2.betree_internal_id ∧ - b1.betree_internal_pivot = b2.betree_internal_pivot ∧ - b1.betree_internal_left = b2.betree_internal_left ∧ - b1.betree_internal_right = b2.betree_internal_right - - [betree_internal_t_fn_updates] Theorem - - ⊢ (∀f u u0 b b0. - betree_internal_t u u0 b b0 with betree_internal_id updated_by f = - betree_internal_t (f u) u0 b b0) ∧ - (∀f u u0 b b0. - betree_internal_t u u0 b b0 with - betree_internal_pivot updated_by f = - betree_internal_t u (f u0) b b0) ∧ - (∀f u u0 b b0. - betree_internal_t u u0 b b0 with - betree_internal_left updated_by f = - betree_internal_t u u0 (f b) b0) ∧ - ∀f u u0 b b0. - betree_internal_t u u0 b b0 with - betree_internal_right updated_by f = - betree_internal_t u u0 b (f b0) - - [betree_internal_t_fupdcanon] Theorem - - ⊢ (∀g f b. - b with - <|betree_internal_pivot updated_by f; - betree_internal_id updated_by g|> = - b with - <|betree_internal_id updated_by g; - betree_internal_pivot updated_by f|>) ∧ - (∀g f b. - b with - <|betree_internal_left updated_by f; - betree_internal_id updated_by g|> = - b with - <|betree_internal_id updated_by g; - betree_internal_left updated_by f|>) ∧ - (∀g f b. - b with - <|betree_internal_left updated_by f; - betree_internal_pivot updated_by g|> = - b with - <|betree_internal_pivot updated_by g; - betree_internal_left updated_by f|>) ∧ - (∀g f b. - b with - <|betree_internal_right updated_by f; - betree_internal_id updated_by g|> = - b with - <|betree_internal_id updated_by g; - betree_internal_right updated_by f|>) ∧ - (∀g f b. - b with - <|betree_internal_right updated_by f; - betree_internal_pivot updated_by g|> = - b with - <|betree_internal_pivot updated_by g; - betree_internal_right updated_by f|>) ∧ - ∀g f b. - b with - <|betree_internal_right updated_by f; - betree_internal_left updated_by g|> = - b with - <|betree_internal_left updated_by g; - betree_internal_right updated_by f|> - - [betree_internal_t_fupdcanon_comp] Theorem - - ⊢ ((∀g f. - betree_internal_pivot_fupd f ∘ betree_internal_id_fupd g = - betree_internal_id_fupd g ∘ betree_internal_pivot_fupd f) ∧ - ∀h g f. - betree_internal_pivot_fupd f ∘ betree_internal_id_fupd g ∘ h = - betree_internal_id_fupd g ∘ betree_internal_pivot_fupd f ∘ h) ∧ - ((∀g f. - betree_internal_left_fupd f ∘ betree_internal_id_fupd g = - betree_internal_id_fupd g ∘ betree_internal_left_fupd f) ∧ - ∀h g f. - betree_internal_left_fupd f ∘ betree_internal_id_fupd g ∘ h = - betree_internal_id_fupd g ∘ betree_internal_left_fupd f ∘ h) ∧ - ((∀g f. - betree_internal_left_fupd f ∘ betree_internal_pivot_fupd g = - betree_internal_pivot_fupd g ∘ betree_internal_left_fupd f) ∧ - ∀h g f. - betree_internal_left_fupd f ∘ betree_internal_pivot_fupd g ∘ h = - betree_internal_pivot_fupd g ∘ betree_internal_left_fupd f ∘ h) ∧ - ((∀g f. - betree_internal_right_fupd f ∘ betree_internal_id_fupd g = - betree_internal_id_fupd g ∘ betree_internal_right_fupd f) ∧ - ∀h g f. - betree_internal_right_fupd f ∘ betree_internal_id_fupd g ∘ h = - betree_internal_id_fupd g ∘ betree_internal_right_fupd f ∘ h) ∧ - ((∀g f. - betree_internal_right_fupd f ∘ betree_internal_pivot_fupd g = - betree_internal_pivot_fupd g ∘ betree_internal_right_fupd f) ∧ - ∀h g f. - betree_internal_right_fupd f ∘ betree_internal_pivot_fupd g ∘ h = - betree_internal_pivot_fupd g ∘ betree_internal_right_fupd f ∘ h) ∧ - (∀g f. - betree_internal_right_fupd f ∘ betree_internal_left_fupd g = - betree_internal_left_fupd g ∘ betree_internal_right_fupd f) ∧ - ∀h g f. - betree_internal_right_fupd f ∘ betree_internal_left_fupd g ∘ h = - betree_internal_left_fupd g ∘ betree_internal_right_fupd f ∘ h - - [betree_internal_t_fupdfupds] Theorem - - ⊢ (∀g f b. - b with - <|betree_internal_id updated_by f; - betree_internal_id updated_by g|> = - b with betree_internal_id updated_by f ∘ g) ∧ - (∀g f b. - b with - <|betree_internal_pivot updated_by f; - betree_internal_pivot updated_by g|> = - b with betree_internal_pivot updated_by f ∘ g) ∧ - (∀g f b. - b with - <|betree_internal_left updated_by f; - betree_internal_left updated_by g|> = - b with betree_internal_left updated_by f ∘ g) ∧ - ∀g f b. - b with - <|betree_internal_right updated_by f; - betree_internal_right updated_by g|> = - b with betree_internal_right updated_by f ∘ g - - [betree_internal_t_fupdfupds_comp] Theorem - - ⊢ ((∀g f. - betree_internal_id_fupd f ∘ betree_internal_id_fupd g = - betree_internal_id_fupd (f ∘ g)) ∧ - ∀h g f. - betree_internal_id_fupd f ∘ betree_internal_id_fupd g ∘ h = - betree_internal_id_fupd (f ∘ g) ∘ h) ∧ - ((∀g f. - betree_internal_pivot_fupd f ∘ betree_internal_pivot_fupd g = - betree_internal_pivot_fupd (f ∘ g)) ∧ - ∀h g f. - betree_internal_pivot_fupd f ∘ betree_internal_pivot_fupd g ∘ h = - betree_internal_pivot_fupd (f ∘ g) ∘ h) ∧ - ((∀g f. - betree_internal_left_fupd f ∘ betree_internal_left_fupd g = - betree_internal_left_fupd (f ∘ g)) ∧ - ∀h g f. - betree_internal_left_fupd f ∘ betree_internal_left_fupd g ∘ h = - betree_internal_left_fupd (f ∘ g) ∘ h) ∧ - (∀g f. - betree_internal_right_fupd f ∘ betree_internal_right_fupd g = - betree_internal_right_fupd (f ∘ g)) ∧ - ∀h g f. - betree_internal_right_fupd f ∘ betree_internal_right_fupd g ∘ h = - betree_internal_right_fupd (f ∘ g) ∘ h - - [betree_internal_t_induction] Theorem - - ⊢ ∀P0 P1. - (∀b b0. P1 b ∧ P1 b0 ⇒ ∀u u0. P0 (betree_internal_t u0 u b b0)) ∧ - (∀b. P0 b ⇒ P1 (BetreeNodeInternal b)) ∧ - (∀b. P1 (BetreeNodeLeaf b)) ⇒ - (∀b. P0 b) ∧ ∀b. P1 b - - [betree_internal_t_literal_11] Theorem - - ⊢ ∀u01 u1 b11 b01 u02 u2 b12 b02. - <|betree_internal_id := u01; betree_internal_pivot := u1; - betree_internal_left := b11; betree_internal_right := b01|> = - <|betree_internal_id := u02; betree_internal_pivot := u2; - betree_internal_left := b12; betree_internal_right := b02|> ⇔ - u01 = u02 ∧ u1 = u2 ∧ b11 = b12 ∧ b01 = b02 - - [betree_internal_t_literal_nchotomy] Theorem - - ⊢ ∀b. ∃u0 u b1 b0. - b = - <|betree_internal_id := u0; betree_internal_pivot := u; - betree_internal_left := b1; betree_internal_right := b0|> - - [betree_internal_t_nchotomy] Theorem - - ⊢ ∀bb. ∃u0 u b b0. bb = betree_internal_t u0 u b b0 - - [betree_internal_t_updates_eq_literal] Theorem - - ⊢ ∀b u0 u b1 b0. - b with - <|betree_internal_id := u0; betree_internal_pivot := u; - betree_internal_left := b1; betree_internal_right := b0|> = - <|betree_internal_id := u0; betree_internal_pivot := u; - betree_internal_left := b1; betree_internal_right := b0|> - - [betree_leaf_t_11] Theorem - - ⊢ ∀a0 a1 a0' a1'. - betree_leaf_t a0 a1 = betree_leaf_t a0' a1' ⇔ a0 = a0' ∧ a1 = a1' - - [betree_leaf_t_Axiom] Theorem - - ⊢ ∀f. ∃fn. ∀a0 a1. fn (betree_leaf_t a0 a1) = f a0 a1 - - [betree_leaf_t_accessors] Theorem - - ⊢ (∀u u0. (betree_leaf_t u u0).betree_leaf_id = u) ∧ - ∀u u0. (betree_leaf_t u u0).betree_leaf_size = u0 - - [betree_leaf_t_accfupds] Theorem - - ⊢ (∀f b. - (b with betree_leaf_size updated_by f).betree_leaf_id = - b.betree_leaf_id) ∧ - (∀f b. - (b with betree_leaf_id updated_by f).betree_leaf_size = - b.betree_leaf_size) ∧ - (∀f b. - (b with betree_leaf_id updated_by f).betree_leaf_id = - f b.betree_leaf_id) ∧ - ∀f b. - (b with betree_leaf_size updated_by f).betree_leaf_size = - f b.betree_leaf_size - - [betree_leaf_t_case_cong] Theorem - - ⊢ ∀M M' f. - M = M' ∧ (∀a0 a1. M' = betree_leaf_t a0 a1 ⇒ f a0 a1 = f' a0 a1) ⇒ - betree_leaf_t_CASE M f = betree_leaf_t_CASE M' f' - - [betree_leaf_t_case_eq] Theorem - - ⊢ betree_leaf_t_CASE x f = v ⇔ - ∃u u0. x = betree_leaf_t u u0 ∧ f u u0 = v - - [betree_leaf_t_component_equality] Theorem - - ⊢ ∀b1 b2. - b1 = b2 ⇔ - b1.betree_leaf_id = b2.betree_leaf_id ∧ - b1.betree_leaf_size = b2.betree_leaf_size - - [betree_leaf_t_fn_updates] Theorem - - ⊢ (∀f u u0. - betree_leaf_t u u0 with betree_leaf_id updated_by f = - betree_leaf_t (f u) u0) ∧ - ∀f u u0. - betree_leaf_t u u0 with betree_leaf_size updated_by f = - betree_leaf_t u (f u0) - - [betree_leaf_t_fupdcanon] Theorem - - ⊢ ∀g f b. - b with - <|betree_leaf_size updated_by f; betree_leaf_id updated_by g|> = - b with - <|betree_leaf_id updated_by g; betree_leaf_size updated_by f|> - - [betree_leaf_t_fupdcanon_comp] Theorem - - ⊢ (∀g f. - betree_leaf_size_fupd f ∘ betree_leaf_id_fupd g = - betree_leaf_id_fupd g ∘ betree_leaf_size_fupd f) ∧ - ∀h g f. - betree_leaf_size_fupd f ∘ betree_leaf_id_fupd g ∘ h = - betree_leaf_id_fupd g ∘ betree_leaf_size_fupd f ∘ h - - [betree_leaf_t_fupdfupds] Theorem - - ⊢ (∀g f b. - b with - <|betree_leaf_id updated_by f; betree_leaf_id updated_by g|> = - b with betree_leaf_id updated_by f ∘ g) ∧ - ∀g f b. - b with - <|betree_leaf_size updated_by f; betree_leaf_size updated_by g|> = - b with betree_leaf_size updated_by f ∘ g - - [betree_leaf_t_fupdfupds_comp] Theorem - - ⊢ ((∀g f. - betree_leaf_id_fupd f ∘ betree_leaf_id_fupd g = - betree_leaf_id_fupd (f ∘ g)) ∧ - ∀h g f. - betree_leaf_id_fupd f ∘ betree_leaf_id_fupd g ∘ h = - betree_leaf_id_fupd (f ∘ g) ∘ h) ∧ - (∀g f. - betree_leaf_size_fupd f ∘ betree_leaf_size_fupd g = - betree_leaf_size_fupd (f ∘ g)) ∧ - ∀h g f. - betree_leaf_size_fupd f ∘ betree_leaf_size_fupd g ∘ h = - betree_leaf_size_fupd (f ∘ g) ∘ h - - [betree_leaf_t_induction] Theorem - - ⊢ ∀P. (∀u u0. P (betree_leaf_t u u0)) ⇒ ∀b. P b - - [betree_leaf_t_literal_11] Theorem - - ⊢ ∀u01 u1 u02 u2. - <|betree_leaf_id := u01; betree_leaf_size := u1|> = - <|betree_leaf_id := u02; betree_leaf_size := u2|> ⇔ - u01 = u02 ∧ u1 = u2 - - [betree_leaf_t_literal_nchotomy] Theorem - - ⊢ ∀b. ∃u0 u. b = <|betree_leaf_id := u0; betree_leaf_size := u|> - - [betree_leaf_t_nchotomy] Theorem - - ⊢ ∀bb. ∃u u0. bb = betree_leaf_t u u0 - - [betree_leaf_t_updates_eq_literal] Theorem - - ⊢ ∀b u0 u. - b with <|betree_leaf_id := u0; betree_leaf_size := u|> = - <|betree_leaf_id := u0; betree_leaf_size := u|> - - [betree_list_t_11] Theorem - - ⊢ ∀a0 a1 a0' a1'. - BetreeListCons a0 a1 = BetreeListCons a0' a1' ⇔ - a0 = a0' ∧ a1 = a1' - - [betree_list_t_Axiom] Theorem - - ⊢ ∀f0 f1. ∃fn. - (∀a0 a1. fn (BetreeListCons a0 a1) = f0 a0 a1 (fn a1)) ∧ - fn BetreeListNil = f1 - - [betree_list_t_case_cong] Theorem - - ⊢ ∀M M' f v. - M = M' ∧ - (∀a0 a1. M' = BetreeListCons a0 a1 ⇒ f a0 a1 = f' a0 a1) ∧ - (M' = BetreeListNil ⇒ v = v') ⇒ - betree_list_t_CASE M f v = betree_list_t_CASE M' f' v' - - [betree_list_t_case_eq] Theorem - - ⊢ betree_list_t_CASE x f v = v' ⇔ - (∃t b. x = BetreeListCons t b ∧ f t b = v') ∨ - x = BetreeListNil ∧ v = v' - - [betree_list_t_distinct] Theorem - - ⊢ ∀a1 a0. BetreeListCons a0 a1 ≠ BetreeListNil - - [betree_list_t_induction] Theorem - - ⊢ ∀P. (∀b. P b ⇒ ∀t. P (BetreeListCons t b)) ∧ P BetreeListNil ⇒ - ∀b. P b - - [betree_list_t_nchotomy] Theorem - - ⊢ ∀bb. (∃t b. bb = BetreeListCons t b) ∨ bb = BetreeListNil - - [betree_message_t_11] Theorem - - ⊢ (∀a a'. BetreeMessageInsert a = BetreeMessageInsert a' ⇔ a = a') ∧ - ∀a a'. BetreeMessageUpsert a = BetreeMessageUpsert a' ⇔ a = a' - - [betree_message_t_Axiom] Theorem - - ⊢ ∀f0 f1 f2. ∃fn. - (∀a. fn (BetreeMessageInsert a) = f0 a) ∧ - fn BetreeMessageDelete = f1 ∧ - ∀a. fn (BetreeMessageUpsert a) = f2 a - - [betree_message_t_case_cong] Theorem - - ⊢ ∀M M' f v f1. - M = M' ∧ (∀a. M' = BetreeMessageInsert a ⇒ f a = f' a) ∧ - (M' = BetreeMessageDelete ⇒ v = v') ∧ - (∀a. M' = BetreeMessageUpsert a ⇒ f1 a = f1' a) ⇒ - betree_message_t_CASE M f v f1 = - betree_message_t_CASE M' f' v' f1' - - [betree_message_t_case_eq] Theorem - - ⊢ betree_message_t_CASE x f v f1 = v' ⇔ - (∃u. x = BetreeMessageInsert u ∧ f u = v') ∨ - x = BetreeMessageDelete ∧ v = v' ∨ - ∃b. x = BetreeMessageUpsert b ∧ f1 b = v' - - [betree_message_t_distinct] Theorem - - ⊢ (∀a. BetreeMessageInsert a ≠ BetreeMessageDelete) ∧ - (∀a' a. BetreeMessageInsert a ≠ BetreeMessageUpsert a') ∧ - ∀a. BetreeMessageDelete ≠ BetreeMessageUpsert a - - [betree_message_t_induction] Theorem - - ⊢ ∀P. (∀u. P (BetreeMessageInsert u)) ∧ P BetreeMessageDelete ∧ - (∀b. P (BetreeMessageUpsert b)) ⇒ - ∀b. P b - - [betree_message_t_nchotomy] Theorem - - ⊢ ∀bb. - (∃u. bb = BetreeMessageInsert u) ∨ bb = BetreeMessageDelete ∨ - ∃b. bb = BetreeMessageUpsert b - - [betree_node_id_counter_t_11] Theorem - - ⊢ ∀a a'. - betree_node_id_counter_t a = betree_node_id_counter_t a' ⇔ a = a' - - [betree_node_id_counter_t_Axiom] Theorem - - ⊢ ∀f. ∃fn. ∀a. fn (betree_node_id_counter_t a) = f a - - [betree_node_id_counter_t_accessors] Theorem - - ⊢ ∀u. (betree_node_id_counter_t u). - betree_node_id_counter_next_node_id = - u - - [betree_node_id_counter_t_accfupds] Theorem - - ⊢ ∀f b. - (b with betree_node_id_counter_next_node_id updated_by f). - betree_node_id_counter_next_node_id = - f b.betree_node_id_counter_next_node_id - - [betree_node_id_counter_t_case_cong] Theorem - - ⊢ ∀M M' f. - M = M' ∧ (∀a. M' = betree_node_id_counter_t a ⇒ f a = f' a) ⇒ - betree_node_id_counter_t_CASE M f = - betree_node_id_counter_t_CASE M' f' - - [betree_node_id_counter_t_case_eq] Theorem - - ⊢ betree_node_id_counter_t_CASE x f = v ⇔ - ∃u. x = betree_node_id_counter_t u ∧ f u = v - - [betree_node_id_counter_t_component_equality] Theorem - - ⊢ ∀b1 b2. - b1 = b2 ⇔ - b1.betree_node_id_counter_next_node_id = - b2.betree_node_id_counter_next_node_id - - [betree_node_id_counter_t_fn_updates] Theorem - - ⊢ ∀f u. - betree_node_id_counter_t u with - betree_node_id_counter_next_node_id updated_by f = - betree_node_id_counter_t (f u) - - [betree_node_id_counter_t_fupdfupds] Theorem - - ⊢ ∀g f b. - b with - <|betree_node_id_counter_next_node_id updated_by f; - betree_node_id_counter_next_node_id updated_by g|> = - b with betree_node_id_counter_next_node_id updated_by f ∘ g - - [betree_node_id_counter_t_fupdfupds_comp] Theorem - - ⊢ (∀g f. - betree_node_id_counter_next_node_id_fupd f ∘ - betree_node_id_counter_next_node_id_fupd g = - betree_node_id_counter_next_node_id_fupd (f ∘ g)) ∧ - ∀h g f. - betree_node_id_counter_next_node_id_fupd f ∘ - betree_node_id_counter_next_node_id_fupd g ∘ h = - betree_node_id_counter_next_node_id_fupd (f ∘ g) ∘ h - - [betree_node_id_counter_t_induction] Theorem - - ⊢ ∀P. (∀u. P (betree_node_id_counter_t u)) ⇒ ∀b. P b - - [betree_node_id_counter_t_literal_11] Theorem - - ⊢ ∀u1 u2. - <|betree_node_id_counter_next_node_id := u1|> = - <|betree_node_id_counter_next_node_id := u2|> ⇔ u1 = u2 - - [betree_node_id_counter_t_literal_nchotomy] Theorem - - ⊢ ∀b. ∃u. b = <|betree_node_id_counter_next_node_id := u|> - - [betree_node_id_counter_t_nchotomy] Theorem - - ⊢ ∀bb. ∃u. bb = betree_node_id_counter_t u - - [betree_node_id_counter_t_updates_eq_literal] Theorem - - ⊢ ∀b u. - b with betree_node_id_counter_next_node_id := u = - <|betree_node_id_counter_next_node_id := u|> - - [betree_node_t_11] Theorem - - ⊢ (∀a a'. BetreeNodeInternal a = BetreeNodeInternal a' ⇔ a = a') ∧ - ∀a a'. BetreeNodeLeaf a = BetreeNodeLeaf a' ⇔ a = a' - - [betree_node_t_Axiom] Theorem - - ⊢ ∀f0 f1 f2. ∃fn0 fn1. - (∀a0 a1 a2 a3. - fn0 (betree_internal_t a0 a1 a2 a3) = - f0 a0 a1 a2 a3 (fn1 a2) (fn1 a3)) ∧ - (∀a. fn1 (BetreeNodeInternal a) = f1 a (fn0 a)) ∧ - ∀a. fn1 (BetreeNodeLeaf a) = f2 a - - [betree_node_t_case_cong] Theorem - - ⊢ ∀M M' f f1. - M = M' ∧ (∀a. M' = BetreeNodeInternal a ⇒ f a = f' a) ∧ - (∀a. M' = BetreeNodeLeaf a ⇒ f1 a = f1' a) ⇒ - betree_node_t_CASE M f f1 = betree_node_t_CASE M' f' f1' - - [betree_node_t_case_eq] Theorem - - ⊢ betree_node_t_CASE x f f1 = v ⇔ - (∃b. x = BetreeNodeInternal b ∧ f b = v) ∨ - ∃b. x = BetreeNodeLeaf b ∧ f1 b = v - - [betree_node_t_distinct] Theorem - - ⊢ ∀a' a. BetreeNodeInternal a ≠ BetreeNodeLeaf a' - - [betree_node_t_induction] Theorem - - ⊢ ∀P0 P1. - (∀b b0. P1 b ∧ P1 b0 ⇒ ∀u u0. P0 (betree_internal_t u0 u b b0)) ∧ - (∀b. P0 b ⇒ P1 (BetreeNodeInternal b)) ∧ - (∀b. P1 (BetreeNodeLeaf b)) ⇒ - (∀b. P0 b) ∧ ∀b. P1 b - - [betree_node_t_nchotomy] Theorem - - ⊢ ∀bb. (∃b. bb = BetreeNodeInternal b) ∨ ∃b. bb = BetreeNodeLeaf b - - [betree_params_t_11] Theorem - - ⊢ ∀a0 a1 a0' a1'. - betree_params_t a0 a1 = betree_params_t a0' a1' ⇔ - a0 = a0' ∧ a1 = a1' - - [betree_params_t_Axiom] Theorem - - ⊢ ∀f. ∃fn. ∀a0 a1. fn (betree_params_t a0 a1) = f a0 a1 - - [betree_params_t_accessors] Theorem - - ⊢ (∀u u0. (betree_params_t u u0).betree_params_min_flush_size = u) ∧ - ∀u u0. (betree_params_t u u0).betree_params_split_size = u0 - - [betree_params_t_accfupds] Theorem - - ⊢ (∀f b. - (b with betree_params_split_size updated_by f). - betree_params_min_flush_size = - b.betree_params_min_flush_size) ∧ - (∀f b. - (b with betree_params_min_flush_size updated_by f). - betree_params_split_size = - b.betree_params_split_size) ∧ - (∀f b. - (b with betree_params_min_flush_size updated_by f). - betree_params_min_flush_size = - f b.betree_params_min_flush_size) ∧ - ∀f b. - (b with betree_params_split_size updated_by f). - betree_params_split_size = - f b.betree_params_split_size - - [betree_params_t_case_cong] Theorem - - ⊢ ∀M M' f. - M = M' ∧ - (∀a0 a1. M' = betree_params_t a0 a1 ⇒ f a0 a1 = f' a0 a1) ⇒ - betree_params_t_CASE M f = betree_params_t_CASE M' f' - - [betree_params_t_case_eq] Theorem - - ⊢ betree_params_t_CASE x f = v ⇔ - ∃u u0. x = betree_params_t u u0 ∧ f u u0 = v - - [betree_params_t_component_equality] Theorem - - ⊢ ∀b1 b2. - b1 = b2 ⇔ - b1.betree_params_min_flush_size = b2.betree_params_min_flush_size ∧ - b1.betree_params_split_size = b2.betree_params_split_size - - [betree_params_t_fn_updates] Theorem - - ⊢ (∀f u u0. - betree_params_t u u0 with - betree_params_min_flush_size updated_by f = - betree_params_t (f u) u0) ∧ - ∀f u u0. - betree_params_t u u0 with betree_params_split_size updated_by f = - betree_params_t u (f u0) - - [betree_params_t_fupdcanon] Theorem - - ⊢ ∀g f b. - b with - <|betree_params_split_size updated_by f; - betree_params_min_flush_size updated_by g|> = - b with - <|betree_params_min_flush_size updated_by g; - betree_params_split_size updated_by f|> - - [betree_params_t_fupdcanon_comp] Theorem - - ⊢ (∀g f. - betree_params_split_size_fupd f ∘ - betree_params_min_flush_size_fupd g = - betree_params_min_flush_size_fupd g ∘ - betree_params_split_size_fupd f) ∧ - ∀h g f. - betree_params_split_size_fupd f ∘ - betree_params_min_flush_size_fupd g ∘ h = - betree_params_min_flush_size_fupd g ∘ - betree_params_split_size_fupd f ∘ h - - [betree_params_t_fupdfupds] Theorem - - ⊢ (∀g f b. - b with - <|betree_params_min_flush_size updated_by f; - betree_params_min_flush_size updated_by g|> = - b with betree_params_min_flush_size updated_by f ∘ g) ∧ - ∀g f b. - b with - <|betree_params_split_size updated_by f; - betree_params_split_size updated_by g|> = - b with betree_params_split_size updated_by f ∘ g - - [betree_params_t_fupdfupds_comp] Theorem - - ⊢ ((∀g f. - betree_params_min_flush_size_fupd f ∘ - betree_params_min_flush_size_fupd g = - betree_params_min_flush_size_fupd (f ∘ g)) ∧ - ∀h g f. - betree_params_min_flush_size_fupd f ∘ - betree_params_min_flush_size_fupd g ∘ h = - betree_params_min_flush_size_fupd (f ∘ g) ∘ h) ∧ - (∀g f. - betree_params_split_size_fupd f ∘ - betree_params_split_size_fupd g = - betree_params_split_size_fupd (f ∘ g)) ∧ - ∀h g f. - betree_params_split_size_fupd f ∘ - betree_params_split_size_fupd g ∘ h = - betree_params_split_size_fupd (f ∘ g) ∘ h - - [betree_params_t_induction] Theorem - - ⊢ ∀P. (∀u u0. P (betree_params_t u u0)) ⇒ ∀b. P b - - [betree_params_t_literal_11] Theorem - - ⊢ ∀u01 u1 u02 u2. - <|betree_params_min_flush_size := u01; - betree_params_split_size := u1|> = - <|betree_params_min_flush_size := u02; - betree_params_split_size := u2|> ⇔ u01 = u02 ∧ u1 = u2 - - [betree_params_t_literal_nchotomy] Theorem - - ⊢ ∀b. ∃u0 u. - b = - <|betree_params_min_flush_size := u0; - betree_params_split_size := u|> - - [betree_params_t_nchotomy] Theorem - - ⊢ ∀bb. ∃u u0. bb = betree_params_t u u0 - - [betree_params_t_updates_eq_literal] Theorem - - ⊢ ∀b u0 u. - b with - <|betree_params_min_flush_size := u0; - betree_params_split_size := u|> = - <|betree_params_min_flush_size := u0; - betree_params_split_size := u|> - - [betree_upsert_fun_state_t_11] Theorem - - ⊢ (∀a a'. - BetreeUpsertFunStateAdd a = BetreeUpsertFunStateAdd a' ⇔ a = a') ∧ - ∀a a'. - BetreeUpsertFunStateSub a = BetreeUpsertFunStateSub a' ⇔ a = a' - - [betree_upsert_fun_state_t_Axiom] Theorem - - ⊢ ∀f0 f1. ∃fn. - (∀a. fn (BetreeUpsertFunStateAdd a) = f0 a) ∧ - ∀a. fn (BetreeUpsertFunStateSub a) = f1 a - - [betree_upsert_fun_state_t_case_cong] Theorem - - ⊢ ∀M M' f f1. - M = M' ∧ (∀a. M' = BetreeUpsertFunStateAdd a ⇒ f a = f' a) ∧ - (∀a. M' = BetreeUpsertFunStateSub a ⇒ f1 a = f1' a) ⇒ - betree_upsert_fun_state_t_CASE M f f1 = - betree_upsert_fun_state_t_CASE M' f' f1' - - [betree_upsert_fun_state_t_case_eq] Theorem - - ⊢ betree_upsert_fun_state_t_CASE x f f1 = v ⇔ - (∃u. x = BetreeUpsertFunStateAdd u ∧ f u = v) ∨ - ∃u. x = BetreeUpsertFunStateSub u ∧ f1 u = v - - [betree_upsert_fun_state_t_distinct] Theorem - - ⊢ ∀a' a. BetreeUpsertFunStateAdd a ≠ BetreeUpsertFunStateSub a' - - [betree_upsert_fun_state_t_induction] Theorem - - ⊢ ∀P. (∀u. P (BetreeUpsertFunStateAdd u)) ∧ - (∀u. P (BetreeUpsertFunStateSub u)) ⇒ - ∀b. P b - - [betree_upsert_fun_state_t_nchotomy] Theorem - - ⊢ ∀bb. - (∃u. bb = BetreeUpsertFunStateAdd u) ∨ - ∃u. bb = BetreeUpsertFunStateSub u - - [datatype_betree_be_tree_t] Theorem - - ⊢ DATATYPE - (record betree_be_tree_t betree_be_tree_params - betree_be_tree_node_id_cnt betree_be_tree_root) - - [datatype_betree_internal_t] Theorem - - ⊢ DATATYPE - (record betree_internal_t betree_internal_id - betree_internal_pivot betree_internal_left - betree_internal_right ∧ - betree_node_t BetreeNodeInternal BetreeNodeLeaf) - - [datatype_betree_leaf_t] Theorem - - ⊢ DATATYPE (record betree_leaf_t betree_leaf_id betree_leaf_size) - - [datatype_betree_list_t] Theorem - - ⊢ DATATYPE (betree_list_t BetreeListCons BetreeListNil) - - [datatype_betree_message_t] Theorem - - ⊢ DATATYPE - (betree_message_t BetreeMessageInsert BetreeMessageDelete - BetreeMessageUpsert) - - [datatype_betree_node_id_counter_t] Theorem - - ⊢ DATATYPE - (record betree_node_id_counter_t - betree_node_id_counter_next_node_id) - - [datatype_betree_params_t] Theorem - - ⊢ DATATYPE - (record betree_params_t betree_params_min_flush_size - betree_params_split_size) - - [datatype_betree_upsert_fun_state_t] Theorem - - ⊢ DATATYPE - (betree_upsert_fun_state_t BetreeUpsertFunStateAdd - BetreeUpsertFunStateSub) - - -*) -end diff --git a/tests/hol4/betree/betree_FunsScript.sml b/tests/hol4/betree/betree_FunsScript.sml new file mode 100644 index 00000000..49bcb446 --- /dev/null +++ b/tests/hol4/betree/betree_FunsScript.sml @@ -0,0 +1,1212 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [betree]: function definitions *) +open primitivesLib divDefLib +open betree_TypesTheory betree_OpaqueTheory + +val _ = new_theory "betree_Funs" + + +val betree_load_internal_node_fwd_def = Define ‘ + (** [betree::betree::load_internal_node]: forward function *) + betree_load_internal_node_fwd + (id : u64) (st : state) : + (state # (u64 # betree_message_t) betree_list_t) result + = + betree_utils_load_internal_node_fwd id st +’ + +val betree_store_internal_node_fwd_def = Define ‘ + (** [betree::betree::store_internal_node]: forward function *) + betree_store_internal_node_fwd + (id : u64) (content : (u64 # betree_message_t) betree_list_t) (st : state) + : + (state # unit) result + = + do + (st0, _) <- betree_utils_store_internal_node_fwd id content st; + Return (st0, ()) + od +’ + +val betree_load_leaf_node_fwd_def = Define ‘ + (** [betree::betree::load_leaf_node]: forward function *) + betree_load_leaf_node_fwd + (id : u64) (st : state) : (state # (u64 # u64) betree_list_t) result = + betree_utils_load_leaf_node_fwd id st +’ + +val betree_store_leaf_node_fwd_def = Define ‘ + (** [betree::betree::store_leaf_node]: forward function *) + betree_store_leaf_node_fwd + (id : u64) (content : (u64 # u64) betree_list_t) (st : state) : + (state # unit) result + = + do + (st0, _) <- betree_utils_store_leaf_node_fwd id content st; + Return (st0, ()) + od +’ + +val betree_fresh_node_id_fwd_def = Define ‘ + (** [betree::betree::fresh_node_id]: forward function *) + betree_fresh_node_id_fwd (counter : u64) : u64 result = + do + _ <- u64_add counter (int_to_u64 1); + Return counter + od +’ + +val betree_fresh_node_id_back_def = Define ‘ + (** [betree::betree::fresh_node_id]: backward function 0 *) + betree_fresh_node_id_back (counter : u64) : u64 result = + u64_add counter (int_to_u64 1) +’ + +val betree_node_id_counter_new_fwd_def = Define ‘ + (** [betree::betree::NodeIdCounter::{0}::new]: forward function *) + betree_node_id_counter_new_fwd : betree_node_id_counter_t result = + Return (<| betree_node_id_counter_next_node_id := (int_to_u64 0) |>) +’ + +val betree_node_id_counter_fresh_id_fwd_def = Define ‘ + (** [betree::betree::NodeIdCounter::{0}::fresh_id]: forward function *) + betree_node_id_counter_fresh_id_fwd + (self : betree_node_id_counter_t) : u64 result = + do + _ <- u64_add self.betree_node_id_counter_next_node_id (int_to_u64 1); + Return self.betree_node_id_counter_next_node_id + od +’ + +val betree_node_id_counter_fresh_id_back_def = Define ‘ + (** [betree::betree::NodeIdCounter::{0}::fresh_id]: backward function 0 *) + betree_node_id_counter_fresh_id_back + (self : betree_node_id_counter_t) : betree_node_id_counter_t result = + do + i <- u64_add self.betree_node_id_counter_next_node_id (int_to_u64 1); + Return (<| betree_node_id_counter_next_node_id := i |>) + od +’ + +val betree_upsert_update_fwd_def = Define ‘ + (** [betree::betree::upsert_update]: forward function *) + betree_upsert_update_fwd + (prev : u64 option) (st : betree_upsert_fun_state_t) : u64 result = + (case prev of + | NONE => + (case st of + | BetreeUpsertFunStateAdd v => Return v + | BetreeUpsertFunStateSub i => Return (int_to_u64 0)) + | SOME prev0 => + (case st of + | BetreeUpsertFunStateAdd v => + do + margin <- u64_sub core_u64_max prev0; + if u64_ge margin v then u64_add prev0 v else Return core_u64_max + od + | BetreeUpsertFunStateSub v => + if u64_ge prev0 v then u64_sub prev0 v else Return (int_to_u64 0))) +’ + +val [betree_list_len_fwd_def] = DefineDiv ‘ + (** [betree::betree::List::{1}::len]: forward function *) + betree_list_len_fwd (self : 't betree_list_t) : u64 result = + (case self of + | BetreeListCons t tl => + do + i <- betree_list_len_fwd tl; + u64_add (int_to_u64 1) i + od + | BetreeListNil => Return (int_to_u64 0)) +’ + +val [betree_list_split_at_fwd_def] = DefineDiv ‘ + (** [betree::betree::List::{1}::split_at]: forward function *) + betree_list_split_at_fwd + (self : 't betree_list_t) (n : u64) : + ('t betree_list_t # 't betree_list_t) result + = + if n = int_to_u64 0 + then Return (BetreeListNil, self) + else + (case self of + | BetreeListCons hd tl => + do + i <- u64_sub n (int_to_u64 1); + p <- betree_list_split_at_fwd tl i; + let (ls0, ls1) = p in let l = ls0 in Return (BetreeListCons hd l, ls1) + od + | BetreeListNil => Fail Failure) +’ + +val betree_list_push_front_fwd_back_def = Define ‘ + (** [betree::betree::List::{1}::push_front]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + betree_list_push_front_fwd_back + (self : 't betree_list_t) (x : 't) : 't betree_list_t result = + let tl = mem_replace_fwd self BetreeListNil in + let l = tl in + Return (BetreeListCons x l) +’ + +val betree_list_pop_front_fwd_def = Define ‘ + (** [betree::betree::List::{1}::pop_front]: forward function *) + betree_list_pop_front_fwd (self : 't betree_list_t) : 't result = + let ls = mem_replace_fwd self BetreeListNil in + (case ls of + | BetreeListCons x tl => Return x + | BetreeListNil => Fail Failure) +’ + +val betree_list_pop_front_back_def = Define ‘ + (** [betree::betree::List::{1}::pop_front]: backward function 0 *) + betree_list_pop_front_back + (self : 't betree_list_t) : 't betree_list_t result = + let ls = mem_replace_fwd self BetreeListNil in + (case ls of + | BetreeListCons x tl => Return tl + | BetreeListNil => Fail Failure) +’ + +val betree_list_hd_fwd_def = Define ‘ + (** [betree::betree::List::{1}::hd]: forward function *) + betree_list_hd_fwd (self : 't betree_list_t) : 't result = + (case self of + | BetreeListCons hd l => Return hd + | BetreeListNil => Fail Failure) +’ + +val betree_list_head_has_key_fwd_def = Define ‘ + (** [betree::betree::List::{2}::head_has_key]: forward function *) + betree_list_head_has_key_fwd + (self : (u64 # 't) betree_list_t) (key : u64) : bool result = + (case self of + | BetreeListCons hd l => let (i, _) = hd in Return (i = key) + | BetreeListNil => Return F) +’ + +val [betree_list_partition_at_pivot_fwd_def] = DefineDiv ‘ + (** [betree::betree::List::{2}::partition_at_pivot]: forward function *) + betree_list_partition_at_pivot_fwd + (self : (u64 # 't) betree_list_t) (pivot : u64) : + ((u64 # 't) betree_list_t # (u64 # 't) betree_list_t) result + = + (case self of + | BetreeListCons hd tl => + let (i, t) = hd in + if u64_ge i pivot + then Return (BetreeListNil, BetreeListCons (i, t) tl) + else ( + do + p <- betree_list_partition_at_pivot_fwd tl pivot; + let (ls0, ls1) = p in + let l = ls0 in + Return (BetreeListCons (i, t) l, ls1) + od) + | BetreeListNil => Return (BetreeListNil, BetreeListNil)) +’ + +val betree_leaf_split_fwd_def = Define ‘ + (** [betree::betree::Leaf::{3}::split]: forward function *) + betree_leaf_split_fwd + (self : betree_leaf_t) (content : (u64 # u64) betree_list_t) + (params : betree_params_t) (node_id_cnt : betree_node_id_counter_t) + (st : state) : + (state # betree_internal_t) result + = + do + p <- betree_list_split_at_fwd content params.betree_params_split_size; + let (content0, content1) = p in + do + p0 <- betree_list_hd_fwd content1; + let (pivot, _) = p0 in + do + id0 <- betree_node_id_counter_fresh_id_fwd node_id_cnt; + node_id_cnt0 <- betree_node_id_counter_fresh_id_back node_id_cnt; + id1 <- betree_node_id_counter_fresh_id_fwd node_id_cnt0; + (st0, _) <- betree_store_leaf_node_fwd id0 content0 st; + (st1, _) <- betree_store_leaf_node_fwd id1 content1 st0; + let n = BetreeNodeLeaf + (<| + betree_leaf_id := id0; + betree_leaf_size := params.betree_params_split_size + |>) in + let n0 = BetreeNodeLeaf + (<| + betree_leaf_id := id1; + betree_leaf_size := params.betree_params_split_size + |>) in + Return (st1, + <| + betree_internal_id := self.betree_leaf_id; + betree_internal_pivot := pivot; + betree_internal_left := n; + betree_internal_right := n0 + |>) + od + od + od +’ + +val betree_leaf_split_back_def = Define ‘ + (** [betree::betree::Leaf::{3}::split]: backward function 2 *) + betree_leaf_split_back + (self : betree_leaf_t) (content : (u64 # u64) betree_list_t) + (params : betree_params_t) (node_id_cnt : betree_node_id_counter_t) + (st : state) : + betree_node_id_counter_t result + = + do + p <- betree_list_split_at_fwd content params.betree_params_split_size; + let (content0, content1) = p in + do + _ <- betree_list_hd_fwd content1; + id0 <- betree_node_id_counter_fresh_id_fwd node_id_cnt; + node_id_cnt0 <- betree_node_id_counter_fresh_id_back node_id_cnt; + id1 <- betree_node_id_counter_fresh_id_fwd node_id_cnt0; + (st0, _) <- betree_store_leaf_node_fwd id0 content0 st; + _ <- betree_store_leaf_node_fwd id1 content1 st0; + betree_node_id_counter_fresh_id_back node_id_cnt0 + od + od +’ + +val [betree_node_lookup_first_message_for_key_fwd_def] = DefineDiv ‘ + (** [betree::betree::Node::{5}::lookup_first_message_for_key]: forward function *) + betree_node_lookup_first_message_for_key_fwd + (key : u64) (msgs : (u64 # betree_message_t) betree_list_t) : + (u64 # betree_message_t) betree_list_t result + = + (case msgs of + | BetreeListCons x next_msgs => + let (i, m) = x in + if u64_ge i key + then Return (BetreeListCons (i, m) next_msgs) + else betree_node_lookup_first_message_for_key_fwd key next_msgs + | BetreeListNil => Return BetreeListNil) +’ + +val [betree_node_lookup_first_message_for_key_back_def] = DefineDiv ‘ + (** [betree::betree::Node::{5}::lookup_first_message_for_key]: backward function 0 *) + betree_node_lookup_first_message_for_key_back + (key : u64) (msgs : (u64 # betree_message_t) betree_list_t) + (ret : (u64 # betree_message_t) betree_list_t) : + (u64 # betree_message_t) betree_list_t result + = + (case msgs of + | BetreeListCons x next_msgs => + let (i, m) = x in + if u64_ge i key + then Return ret + else ( + do + next_msgs0 <- + betree_node_lookup_first_message_for_key_back key next_msgs ret; + Return (BetreeListCons (i, m) next_msgs0) + od) + | BetreeListNil => Return ret) +’ + +val [betree_node_apply_upserts_fwd_def] = DefineDiv ‘ + (** [betree::betree::Node::{5}::apply_upserts]: forward function *) + betree_node_apply_upserts_fwd + (msgs : (u64 # betree_message_t) betree_list_t) (prev : u64 option) + (key : u64) (st : state) : + (state # u64) result + = + do + b <- betree_list_head_has_key_fwd msgs key; + if b + then ( + do + msg <- betree_list_pop_front_fwd msgs; + let (_, m) = msg in + (case m of + | BetreeMessageInsert i => Fail Failure + | BetreeMessageDelete => Fail Failure + | BetreeMessageUpsert s => + do + v <- betree_upsert_update_fwd prev s; + msgs0 <- betree_list_pop_front_back msgs; + betree_node_apply_upserts_fwd msgs0 (SOME v) key st + od) + od) + else ( + do + (st0, v) <- core_option_option_unwrap_fwd prev st; + _ <- betree_list_push_front_fwd_back msgs (key, BetreeMessageInsert v); + Return (st0, v) + od) + od +’ + +val [betree_node_apply_upserts_back_def] = DefineDiv ‘ + (** [betree::betree::Node::{5}::apply_upserts]: backward function 0 *) + betree_node_apply_upserts_back + (msgs : (u64 # betree_message_t) betree_list_t) (prev : u64 option) + (key : u64) (st : state) : + (u64 # betree_message_t) betree_list_t result + = + do + b <- betree_list_head_has_key_fwd msgs key; + if b + then ( + do + msg <- betree_list_pop_front_fwd msgs; + let (_, m) = msg in + (case m of + | BetreeMessageInsert i => Fail Failure + | BetreeMessageDelete => Fail Failure + | BetreeMessageUpsert s => + do + v <- betree_upsert_update_fwd prev s; + msgs0 <- betree_list_pop_front_back msgs; + betree_node_apply_upserts_back msgs0 (SOME v) key st + od) + od) + else ( + do + (_, v) <- core_option_option_unwrap_fwd prev st; + betree_list_push_front_fwd_back msgs (key, BetreeMessageInsert v) + od) + od +’ + +val [betree_node_lookup_in_bindings_fwd_def] = DefineDiv ‘ + (** [betree::betree::Node::{5}::lookup_in_bindings]: forward function *) + betree_node_lookup_in_bindings_fwd + (key : u64) (bindings : (u64 # u64) betree_list_t) : u64 option result = + (case bindings of + | BetreeListCons hd tl => + let (i, i0) = hd in + if i = key + then Return (SOME i0) + else + if u64_gt i key + then Return NONE + else betree_node_lookup_in_bindings_fwd key tl + | BetreeListNil => Return NONE) +’ + +val [betree_internal_lookup_in_children_fwd_def, betree_internal_lookup_in_children_back_def, betree_node_lookup_fwd_def, betree_node_lookup_back_def] = DefineDiv ‘ + (** [betree::betree::Internal::{4}::lookup_in_children]: forward function *) + (betree_internal_lookup_in_children_fwd + (self : betree_internal_t) (key : u64) (st : state) : + (state # u64 option) result + = + if u64_lt key self.betree_internal_pivot + then betree_node_lookup_fwd self.betree_internal_left key st + else betree_node_lookup_fwd self.betree_internal_right key st) + /\ + + (** [betree::betree::Internal::{4}::lookup_in_children]: backward function 0 *) + (betree_internal_lookup_in_children_back + (self : betree_internal_t) (key : u64) (st : state) : + betree_internal_t result + = + if u64_lt key self.betree_internal_pivot + then ( + do + n <- betree_node_lookup_back self.betree_internal_left key st; + Return (self with <| betree_internal_left := n |>) + od) + else ( + do + n <- betree_node_lookup_back self.betree_internal_right key st; + Return (self with <| betree_internal_right := n |>) + od)) + /\ + + (** [betree::betree::Node::{5}::lookup]: forward function *) + (betree_node_lookup_fwd + (self : betree_node_t) (key : u64) (st : state) : + (state # u64 option) result + = + (case self of + | BetreeNodeInternal node => + do + (st0, msgs) <- betree_load_internal_node_fwd node.betree_internal_id st; + pending <- betree_node_lookup_first_message_for_key_fwd key msgs; + (case pending of + | BetreeListCons p l => + let (k, msg) = p in + if k <> key + then ( + do + (st1, opt) <- betree_internal_lookup_in_children_fwd node key st0; + _ <- + betree_node_lookup_first_message_for_key_back key msgs + (BetreeListCons (k, msg) l); + Return (st1, opt) + od) + else + (case msg of + | BetreeMessageInsert v => + do + _ <- + betree_node_lookup_first_message_for_key_back key msgs + (BetreeListCons (k, BetreeMessageInsert v) l); + Return (st0, SOME v) + od + | BetreeMessageDelete => + do + _ <- + betree_node_lookup_first_message_for_key_back key msgs + (BetreeListCons (k, BetreeMessageDelete) l); + Return (st0, NONE) + od + | BetreeMessageUpsert ufs => + do + (st1, v) <- betree_internal_lookup_in_children_fwd node key st0; + (st2, v0) <- + betree_node_apply_upserts_fwd (BetreeListCons (k, + BetreeMessageUpsert ufs) l) v key st1; + node0 <- betree_internal_lookup_in_children_back node key st0; + pending0 <- + betree_node_apply_upserts_back (BetreeListCons (k, + BetreeMessageUpsert ufs) l) v key st1; + msgs0 <- + betree_node_lookup_first_message_for_key_back key msgs pending0; + (st3, _) <- + betree_store_internal_node_fwd node0.betree_internal_id msgs0 st2; + Return (st3, SOME v0) + od) + | BetreeListNil => + do + (st1, opt) <- betree_internal_lookup_in_children_fwd node key st0; + _ <- + betree_node_lookup_first_message_for_key_back key msgs BetreeListNil; + Return (st1, opt) + od) + od + | BetreeNodeLeaf node => + do + (st0, bindings) <- betree_load_leaf_node_fwd node.betree_leaf_id st; + opt <- betree_node_lookup_in_bindings_fwd key bindings; + Return (st0, opt) + od)) + /\ + + (** [betree::betree::Node::{5}::lookup]: backward function 0 *) + betree_node_lookup_back + (self : betree_node_t) (key : u64) (st : state) : betree_node_t result = + (case self of + | BetreeNodeInternal node => + do + (st0, msgs) <- betree_load_internal_node_fwd node.betree_internal_id st; + pending <- betree_node_lookup_first_message_for_key_fwd key msgs; + (case pending of + | BetreeListCons p l => + let (k, msg) = p in + if k <> key + then ( + do + _ <- + betree_node_lookup_first_message_for_key_back key msgs + (BetreeListCons (k, msg) l); + node0 <- betree_internal_lookup_in_children_back node key st0; + Return (BetreeNodeInternal node0) + od) + else + (case msg of + | BetreeMessageInsert v => + do + _ <- + betree_node_lookup_first_message_for_key_back key msgs + (BetreeListCons (k, BetreeMessageInsert v) l); + Return (BetreeNodeInternal node) + od + | BetreeMessageDelete => + do + _ <- + betree_node_lookup_first_message_for_key_back key msgs + (BetreeListCons (k, BetreeMessageDelete) l); + Return (BetreeNodeInternal node) + od + | BetreeMessageUpsert ufs => + do + (st1, v) <- betree_internal_lookup_in_children_fwd node key st0; + (st2, _) <- + betree_node_apply_upserts_fwd (BetreeListCons (k, + BetreeMessageUpsert ufs) l) v key st1; + node0 <- betree_internal_lookup_in_children_back node key st0; + pending0 <- + betree_node_apply_upserts_back (BetreeListCons (k, + BetreeMessageUpsert ufs) l) v key st1; + msgs0 <- + betree_node_lookup_first_message_for_key_back key msgs pending0; + _ <- + betree_store_internal_node_fwd node0.betree_internal_id msgs0 st2; + Return (BetreeNodeInternal node0) + od) + | BetreeListNil => + do + _ <- + betree_node_lookup_first_message_for_key_back key msgs BetreeListNil; + node0 <- betree_internal_lookup_in_children_back node key st0; + Return (BetreeNodeInternal node0) + od) + od + | BetreeNodeLeaf node => + do + (_, bindings) <- betree_load_leaf_node_fwd node.betree_leaf_id st; + _ <- betree_node_lookup_in_bindings_fwd key bindings; + Return (BetreeNodeLeaf node) + od) +’ + +val [betree_node_filter_messages_for_key_fwd_back_def] = DefineDiv ‘ + (** [betree::betree::Node::{5}::filter_messages_for_key]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + betree_node_filter_messages_for_key_fwd_back + (key : u64) (msgs : (u64 # betree_message_t) betree_list_t) : + (u64 # betree_message_t) betree_list_t result + = + (case msgs of + | BetreeListCons p l => + let (k, m) = p in + if k = key + then ( + do + msgs0 <- betree_list_pop_front_back (BetreeListCons (k, m) l); + betree_node_filter_messages_for_key_fwd_back key msgs0 + od) + else Return (BetreeListCons (k, m) l) + | BetreeListNil => Return BetreeListNil) +’ + +val [betree_node_lookup_first_message_after_key_fwd_def] = DefineDiv ‘ + (** [betree::betree::Node::{5}::lookup_first_message_after_key]: forward function *) + betree_node_lookup_first_message_after_key_fwd + (key : u64) (msgs : (u64 # betree_message_t) betree_list_t) : + (u64 # betree_message_t) betree_list_t result + = + (case msgs of + | BetreeListCons p next_msgs => + let (k, m) = p in + if k = key + then betree_node_lookup_first_message_after_key_fwd key next_msgs + else Return (BetreeListCons (k, m) next_msgs) + | BetreeListNil => Return BetreeListNil) +’ + +val [betree_node_lookup_first_message_after_key_back_def] = DefineDiv ‘ + (** [betree::betree::Node::{5}::lookup_first_message_after_key]: backward function 0 *) + betree_node_lookup_first_message_after_key_back + (key : u64) (msgs : (u64 # betree_message_t) betree_list_t) + (ret : (u64 # betree_message_t) betree_list_t) : + (u64 # betree_message_t) betree_list_t result + = + (case msgs of + | BetreeListCons p next_msgs => + let (k, m) = p in + if k = key + then ( + do + next_msgs0 <- + betree_node_lookup_first_message_after_key_back key next_msgs ret; + Return (BetreeListCons (k, m) next_msgs0) + od) + else Return ret + | BetreeListNil => Return ret) +’ + +val betree_node_apply_to_internal_fwd_back_def = Define ‘ + (** [betree::betree::Node::{5}::apply_to_internal]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + betree_node_apply_to_internal_fwd_back + (msgs : (u64 # betree_message_t) betree_list_t) (key : u64) + (new_msg : betree_message_t) : + (u64 # betree_message_t) betree_list_t result + = + do + msgs0 <- betree_node_lookup_first_message_for_key_fwd key msgs; + b <- betree_list_head_has_key_fwd msgs0 key; + if b + then + (case new_msg of + | BetreeMessageInsert i => + do + msgs1 <- betree_node_filter_messages_for_key_fwd_back key msgs0; + msgs2 <- + betree_list_push_front_fwd_back msgs1 (key, BetreeMessageInsert i); + betree_node_lookup_first_message_for_key_back key msgs msgs2 + od + | BetreeMessageDelete => + do + msgs1 <- betree_node_filter_messages_for_key_fwd_back key msgs0; + msgs2 <- + betree_list_push_front_fwd_back msgs1 (key, BetreeMessageDelete); + betree_node_lookup_first_message_for_key_back key msgs msgs2 + od + | BetreeMessageUpsert s => + do + p <- betree_list_hd_fwd msgs0; + let (_, m) = p in + (case m of + | BetreeMessageInsert prev => + do + v <- betree_upsert_update_fwd (SOME prev) s; + msgs1 <- betree_list_pop_front_back msgs0; + msgs2 <- + betree_list_push_front_fwd_back msgs1 (key, BetreeMessageInsert v); + betree_node_lookup_first_message_for_key_back key msgs msgs2 + od + | BetreeMessageDelete => + do + v <- betree_upsert_update_fwd NONE s; + msgs1 <- betree_list_pop_front_back msgs0; + msgs2 <- + betree_list_push_front_fwd_back msgs1 (key, BetreeMessageInsert v); + betree_node_lookup_first_message_for_key_back key msgs msgs2 + od + | BetreeMessageUpsert ufs => + do + msgs1 <- betree_node_lookup_first_message_after_key_fwd key msgs0; + msgs2 <- + betree_list_push_front_fwd_back msgs1 (key, BetreeMessageUpsert s); + msgs3 <- + betree_node_lookup_first_message_after_key_back key msgs0 msgs2; + betree_node_lookup_first_message_for_key_back key msgs msgs3 + od) + od) + else ( + do + msgs1 <- betree_list_push_front_fwd_back msgs0 (key, new_msg); + betree_node_lookup_first_message_for_key_back key msgs msgs1 + od) + od +’ + +val [betree_node_apply_messages_to_internal_fwd_back_def] = DefineDiv ‘ + (** [betree::betree::Node::{5}::apply_messages_to_internal]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + betree_node_apply_messages_to_internal_fwd_back + (msgs : (u64 # betree_message_t) betree_list_t) + (new_msgs : (u64 # betree_message_t) betree_list_t) : + (u64 # betree_message_t) betree_list_t result + = + (case new_msgs of + | BetreeListCons new_msg new_msgs_tl => + let (i, m) = new_msg in + do + msgs0 <- betree_node_apply_to_internal_fwd_back msgs i m; + betree_node_apply_messages_to_internal_fwd_back msgs0 new_msgs_tl + od + | BetreeListNil => Return msgs) +’ + +val [betree_node_lookup_mut_in_bindings_fwd_def] = DefineDiv ‘ + (** [betree::betree::Node::{5}::lookup_mut_in_bindings]: forward function *) + betree_node_lookup_mut_in_bindings_fwd + (key : u64) (bindings : (u64 # u64) betree_list_t) : + (u64 # u64) betree_list_t result + = + (case bindings of + | BetreeListCons hd tl => + let (i, i0) = hd in + if u64_ge i key + then Return (BetreeListCons (i, i0) tl) + else betree_node_lookup_mut_in_bindings_fwd key tl + | BetreeListNil => Return BetreeListNil) +’ + +val [betree_node_lookup_mut_in_bindings_back_def] = DefineDiv ‘ + (** [betree::betree::Node::{5}::lookup_mut_in_bindings]: backward function 0 *) + betree_node_lookup_mut_in_bindings_back + (key : u64) (bindings : (u64 # u64) betree_list_t) + (ret : (u64 # u64) betree_list_t) : + (u64 # u64) betree_list_t result + = + (case bindings of + | BetreeListCons hd tl => + let (i, i0) = hd in + if u64_ge i key + then Return ret + else ( + do + tl0 <- betree_node_lookup_mut_in_bindings_back key tl ret; + Return (BetreeListCons (i, i0) tl0) + od) + | BetreeListNil => Return ret) +’ + +val betree_node_apply_to_leaf_fwd_back_def = Define ‘ + (** [betree::betree::Node::{5}::apply_to_leaf]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + betree_node_apply_to_leaf_fwd_back + (bindings : (u64 # u64) betree_list_t) (key : u64) + (new_msg : betree_message_t) : + (u64 # u64) betree_list_t result + = + do + bindings0 <- betree_node_lookup_mut_in_bindings_fwd key bindings; + b <- betree_list_head_has_key_fwd bindings0 key; + if b + then ( + do + hd <- betree_list_pop_front_fwd bindings0; + (case new_msg of + | BetreeMessageInsert v => + do + bindings1 <- betree_list_pop_front_back bindings0; + bindings2 <- betree_list_push_front_fwd_back bindings1 (key, v); + betree_node_lookup_mut_in_bindings_back key bindings bindings2 + od + | BetreeMessageDelete => + do + bindings1 <- betree_list_pop_front_back bindings0; + betree_node_lookup_mut_in_bindings_back key bindings bindings1 + od + | BetreeMessageUpsert s => + let (_, i) = hd in + do + v <- betree_upsert_update_fwd (SOME i) s; + bindings1 <- betree_list_pop_front_back bindings0; + bindings2 <- betree_list_push_front_fwd_back bindings1 (key, v); + betree_node_lookup_mut_in_bindings_back key bindings bindings2 + od) + od) + else + (case new_msg of + | BetreeMessageInsert v => + do + bindings1 <- betree_list_push_front_fwd_back bindings0 (key, v); + betree_node_lookup_mut_in_bindings_back key bindings bindings1 + od + | BetreeMessageDelete => + betree_node_lookup_mut_in_bindings_back key bindings bindings0 + | BetreeMessageUpsert s => + do + v <- betree_upsert_update_fwd NONE s; + bindings1 <- betree_list_push_front_fwd_back bindings0 (key, v); + betree_node_lookup_mut_in_bindings_back key bindings bindings1 + od) + od +’ + +val [betree_node_apply_messages_to_leaf_fwd_back_def] = DefineDiv ‘ + (** [betree::betree::Node::{5}::apply_messages_to_leaf]: merged forward/backward function + (there is a single backward function, and the forward function returns ()) *) + betree_node_apply_messages_to_leaf_fwd_back + (bindings : (u64 # u64) betree_list_t) + (new_msgs : (u64 # betree_message_t) betree_list_t) : + (u64 # u64) betree_list_t result + = + (case new_msgs of + | BetreeListCons new_msg new_msgs_tl => + let (i, m) = new_msg in + do + bindings0 <- betree_node_apply_to_leaf_fwd_back bindings i m; + betree_node_apply_messages_to_leaf_fwd_back bindings0 new_msgs_tl + od + | BetreeListNil => Return bindings) +’ + +val [betree_internal_flush_fwd_def, betree_internal_flush_back_def, betree_node_apply_messages_fwd_def, betree_node_apply_messages_back_def] = DefineDiv ‘ + (** [betree::betree::Internal::{4}::flush]: forward function *) + (betree_internal_flush_fwd + (self : betree_internal_t) (params : betree_params_t) + (node_id_cnt : betree_node_id_counter_t) + (content : (u64 # betree_message_t) betree_list_t) (st : state) : + (state # (u64 # betree_message_t) betree_list_t) result + = + do + p <- betree_list_partition_at_pivot_fwd content self.betree_internal_pivot; + let (msgs_left, msgs_right) = p in + do + len_left <- betree_list_len_fwd msgs_left; + if u64_ge len_left params.betree_params_min_flush_size + then ( + do + (st0, _) <- + betree_node_apply_messages_fwd self.betree_internal_left params + node_id_cnt msgs_left st; + (_, node_id_cnt0) <- + betree_node_apply_messages_back self.betree_internal_left params + node_id_cnt msgs_left st; + len_right <- betree_list_len_fwd msgs_right; + if u64_ge len_right params.betree_params_min_flush_size + then ( + do + (st1, _) <- + betree_node_apply_messages_fwd self.betree_internal_right params + node_id_cnt0 msgs_right st0; + _ <- + betree_node_apply_messages_back self.betree_internal_right params + node_id_cnt0 msgs_right st0; + Return (st1, BetreeListNil) + od) + else Return (st0, msgs_right) + od) + else ( + do + (st0, _) <- + betree_node_apply_messages_fwd self.betree_internal_right params + node_id_cnt msgs_right st; + _ <- + betree_node_apply_messages_back self.betree_internal_right params + node_id_cnt msgs_right st; + Return (st0, msgs_left) + od) + od + od) + /\ + + (** [betree::betree::Internal::{4}::flush]: backward function 0 *) + (betree_internal_flush_back + (self : betree_internal_t) (params : betree_params_t) + (node_id_cnt : betree_node_id_counter_t) + (content : (u64 # betree_message_t) betree_list_t) (st : state) : + (betree_internal_t # betree_node_id_counter_t) result + = + do + p <- betree_list_partition_at_pivot_fwd content self.betree_internal_pivot; + let (msgs_left, msgs_right) = p in + do + len_left <- betree_list_len_fwd msgs_left; + if u64_ge len_left params.betree_params_min_flush_size + then ( + do + (st0, _) <- + betree_node_apply_messages_fwd self.betree_internal_left params + node_id_cnt msgs_left st; + (n, node_id_cnt0) <- + betree_node_apply_messages_back self.betree_internal_left params + node_id_cnt msgs_left st; + len_right <- betree_list_len_fwd msgs_right; + if u64_ge len_right params.betree_params_min_flush_size + then ( + do + (n0, node_id_cnt1) <- + betree_node_apply_messages_back self.betree_internal_right params + node_id_cnt0 msgs_right st0; + Return + (self + with + <| + betree_internal_left := n; betree_internal_right := n0 + |>, node_id_cnt1) + od) + else Return (self with <| betree_internal_left := n |>, node_id_cnt0) + od) + else ( + do + (n, node_id_cnt0) <- + betree_node_apply_messages_back self.betree_internal_right params + node_id_cnt msgs_right st; + Return (self with <| betree_internal_right := n |>, node_id_cnt0) + od) + od + od) + /\ + + (** [betree::betree::Node::{5}::apply_messages]: forward function *) + (betree_node_apply_messages_fwd + (self : betree_node_t) (params : betree_params_t) + (node_id_cnt : betree_node_id_counter_t) + (msgs : (u64 # betree_message_t) betree_list_t) (st : state) : + (state # unit) result + = + (case self of + | BetreeNodeInternal node => + do + (st0, content) <- + betree_load_internal_node_fwd node.betree_internal_id st; + content0 <- betree_node_apply_messages_to_internal_fwd_back content msgs; + num_msgs <- betree_list_len_fwd content0; + if u64_ge num_msgs params.betree_params_min_flush_size + then ( + do + (st1, content1) <- + betree_internal_flush_fwd node params node_id_cnt content0 st0; + (node0, _) <- + betree_internal_flush_back node params node_id_cnt content0 st0; + (st2, _) <- + betree_store_internal_node_fwd node0.betree_internal_id content1 st1; + Return (st2, ()) + od) + else ( + do + (st1, _) <- + betree_store_internal_node_fwd node.betree_internal_id content0 st0; + Return (st1, ()) + od) + od + | BetreeNodeLeaf node => + do + (st0, content) <- betree_load_leaf_node_fwd node.betree_leaf_id st; + content0 <- betree_node_apply_messages_to_leaf_fwd_back content msgs; + len <- betree_list_len_fwd content0; + i <- u64_mul (int_to_u64 2) params.betree_params_split_size; + if u64_ge len i + then ( + do + (st1, _) <- betree_leaf_split_fwd node content0 params node_id_cnt st0; + (st2, _) <- + betree_store_leaf_node_fwd node.betree_leaf_id BetreeListNil st1; + Return (st2, ()) + od) + else ( + do + (st1, _) <- + betree_store_leaf_node_fwd node.betree_leaf_id content0 st0; + Return (st1, ()) + od) + od)) + /\ + + (** [betree::betree::Node::{5}::apply_messages]: backward function 0 *) + betree_node_apply_messages_back + (self : betree_node_t) (params : betree_params_t) + (node_id_cnt : betree_node_id_counter_t) + (msgs : (u64 # betree_message_t) betree_list_t) (st : state) : + (betree_node_t # betree_node_id_counter_t) result + = + (case self of + | BetreeNodeInternal node => + do + (st0, content) <- + betree_load_internal_node_fwd node.betree_internal_id st; + content0 <- betree_node_apply_messages_to_internal_fwd_back content msgs; + num_msgs <- betree_list_len_fwd content0; + if u64_ge num_msgs params.betree_params_min_flush_size + then ( + do + (st1, content1) <- + betree_internal_flush_fwd node params node_id_cnt content0 st0; + (node0, node_id_cnt0) <- + betree_internal_flush_back node params node_id_cnt content0 st0; + _ <- + betree_store_internal_node_fwd node0.betree_internal_id content1 st1; + Return (BetreeNodeInternal node0, node_id_cnt0) + od) + else ( + do + _ <- + betree_store_internal_node_fwd node.betree_internal_id content0 st0; + Return (BetreeNodeInternal node, node_id_cnt) + od) + od + | BetreeNodeLeaf node => + do + (st0, content) <- betree_load_leaf_node_fwd node.betree_leaf_id st; + content0 <- betree_node_apply_messages_to_leaf_fwd_back content msgs; + len <- betree_list_len_fwd content0; + i <- u64_mul (int_to_u64 2) params.betree_params_split_size; + if u64_ge len i + then ( + do + (st1, new_node) <- + betree_leaf_split_fwd node content0 params node_id_cnt st0; + _ <- betree_store_leaf_node_fwd node.betree_leaf_id BetreeListNil st1; + node_id_cnt0 <- + betree_leaf_split_back node content0 params node_id_cnt st0; + Return (BetreeNodeInternal new_node, node_id_cnt0) + od) + else ( + do + _ <- betree_store_leaf_node_fwd node.betree_leaf_id content0 st0; + Return (BetreeNodeLeaf (node with <| betree_leaf_size := len |>), + node_id_cnt) + od) + od) +’ + +val betree_node_apply_fwd_def = Define ‘ + (** [betree::betree::Node::{5}::apply]: forward function *) + betree_node_apply_fwd + (self : betree_node_t) (params : betree_params_t) + (node_id_cnt : betree_node_id_counter_t) (key : u64) + (new_msg : betree_message_t) (st : state) : + (state # unit) result + = + let l = BetreeListNil in + do + (st0, _) <- + betree_node_apply_messages_fwd self params node_id_cnt (BetreeListCons + (key, new_msg) l) st; + _ <- + betree_node_apply_messages_back self params node_id_cnt (BetreeListCons + (key, new_msg) l) st; + Return (st0, ()) + od +’ + +val betree_node_apply_back_def = Define ‘ + (** [betree::betree::Node::{5}::apply]: backward function 0 *) + betree_node_apply_back + (self : betree_node_t) (params : betree_params_t) + (node_id_cnt : betree_node_id_counter_t) (key : u64) + (new_msg : betree_message_t) (st : state) : + (betree_node_t # betree_node_id_counter_t) result + = + let l = BetreeListNil in + betree_node_apply_messages_back self params node_id_cnt (BetreeListCons + (key, new_msg) l) st +’ + +val betree_be_tree_new_fwd_def = Define ‘ + (** [betree::betree::BeTree::{6}::new]: forward function *) + betree_be_tree_new_fwd + (min_flush_size : u64) (split_size : u64) (st : state) : + (state # betree_be_tree_t) result + = + do + node_id_cnt <- betree_node_id_counter_new_fwd; + id <- betree_node_id_counter_fresh_id_fwd node_id_cnt; + (st0, _) <- betree_store_leaf_node_fwd id BetreeListNil st; + node_id_cnt0 <- betree_node_id_counter_fresh_id_back node_id_cnt; + Return (st0, + <| + betree_be_tree_params := + (<| + betree_params_min_flush_size := min_flush_size; + betree_params_split_size := split_size + |>); + betree_be_tree_node_id_cnt := node_id_cnt0; + betree_be_tree_root := + (BetreeNodeLeaf + (<| betree_leaf_id := id; betree_leaf_size := (int_to_u64 0) |>)) + |>) + od +’ + +val betree_be_tree_apply_fwd_def = Define ‘ + (** [betree::betree::BeTree::{6}::apply]: forward function *) + betree_be_tree_apply_fwd + (self : betree_be_tree_t) (key : u64) (msg : betree_message_t) (st : state) + : + (state # unit) result + = + do + (st0, _) <- + betree_node_apply_fwd self.betree_be_tree_root self.betree_be_tree_params + self.betree_be_tree_node_id_cnt key msg st; + _ <- + betree_node_apply_back self.betree_be_tree_root + self.betree_be_tree_params self.betree_be_tree_node_id_cnt key msg st; + Return (st0, ()) + od +’ + +val betree_be_tree_apply_back_def = Define ‘ + (** [betree::betree::BeTree::{6}::apply]: backward function 0 *) + betree_be_tree_apply_back + (self : betree_be_tree_t) (key : u64) (msg : betree_message_t) (st : state) + : + betree_be_tree_t result + = + do + (n, nic) <- + betree_node_apply_back self.betree_be_tree_root + self.betree_be_tree_params self.betree_be_tree_node_id_cnt key msg st; + Return + (self + with + <| + betree_be_tree_node_id_cnt := nic; betree_be_tree_root := n + |>) + od +’ + +val betree_be_tree_insert_fwd_def = Define ‘ + (** [betree::betree::BeTree::{6}::insert]: forward function *) + betree_be_tree_insert_fwd + (self : betree_be_tree_t) (key : u64) (value : u64) (st : state) : + (state # unit) result + = + do + (st0, _) <- + betree_be_tree_apply_fwd self key (BetreeMessageInsert value) st; + _ <- betree_be_tree_apply_back self key (BetreeMessageInsert value) st; + Return (st0, ()) + od +’ + +val betree_be_tree_insert_back_def = Define ‘ + (** [betree::betree::BeTree::{6}::insert]: backward function 0 *) + betree_be_tree_insert_back + (self : betree_be_tree_t) (key : u64) (value : u64) (st : state) : + betree_be_tree_t result + = + betree_be_tree_apply_back self key (BetreeMessageInsert value) st +’ + +val betree_be_tree_delete_fwd_def = Define ‘ + (** [betree::betree::BeTree::{6}::delete]: forward function *) + betree_be_tree_delete_fwd + (self : betree_be_tree_t) (key : u64) (st : state) : + (state # unit) result + = + do + (st0, _) <- betree_be_tree_apply_fwd self key BetreeMessageDelete st; + _ <- betree_be_tree_apply_back self key BetreeMessageDelete st; + Return (st0, ()) + od +’ + +val betree_be_tree_delete_back_def = Define ‘ + (** [betree::betree::BeTree::{6}::delete]: backward function 0 *) + betree_be_tree_delete_back + (self : betree_be_tree_t) (key : u64) (st : state) : + betree_be_tree_t result + = + betree_be_tree_apply_back self key BetreeMessageDelete st +’ + +val betree_be_tree_upsert_fwd_def = Define ‘ + (** [betree::betree::BeTree::{6}::upsert]: forward function *) + betree_be_tree_upsert_fwd + (self : betree_be_tree_t) (key : u64) (upd : betree_upsert_fun_state_t) + (st : state) : + (state # unit) result + = + do + (st0, _) <- betree_be_tree_apply_fwd self key (BetreeMessageUpsert upd) st; + _ <- betree_be_tree_apply_back self key (BetreeMessageUpsert upd) st; + Return (st0, ()) + od +’ + +val betree_be_tree_upsert_back_def = Define ‘ + (** [betree::betree::BeTree::{6}::upsert]: backward function 0 *) + betree_be_tree_upsert_back + (self : betree_be_tree_t) (key : u64) (upd : betree_upsert_fun_state_t) + (st : state) : + betree_be_tree_t result + = + betree_be_tree_apply_back self key (BetreeMessageUpsert upd) st +’ + +val betree_be_tree_lookup_fwd_def = Define ‘ + (** [betree::betree::BeTree::{6}::lookup]: forward function *) + betree_be_tree_lookup_fwd + (self : betree_be_tree_t) (key : u64) (st : state) : + (state # u64 option) result + = + betree_node_lookup_fwd self.betree_be_tree_root key st +’ + +val betree_be_tree_lookup_back_def = Define ‘ + (** [betree::betree::BeTree::{6}::lookup]: backward function 0 *) + betree_be_tree_lookup_back + (self : betree_be_tree_t) (key : u64) (st : state) : + betree_be_tree_t result + = + do + n <- betree_node_lookup_back self.betree_be_tree_root key st; + Return (self with <| betree_be_tree_root := n |>) + od +’ + +val main_fwd_def = Define ‘ + (** [betree::main]: forward function *) + main_fwd : unit result = + Return () +’ + +(** Unit test for [betree::main] *) +val _ = assert_return (“main_fwd”) + +val _ = export_theory () diff --git a/tests/hol4/betree/betree_FunsTheory.sig b/tests/hol4/betree/betree_FunsTheory.sig new file mode 100644 index 00000000..4127cba1 --- /dev/null +++ b/tests/hol4/betree/betree_FunsTheory.sig @@ -0,0 +1,1230 @@ +signature betree_FunsTheory = +sig + type thm = Thm.thm + + (* Definitions *) + val betree_be_tree_apply_back_def : thm + val betree_be_tree_apply_fwd_def : thm + val betree_be_tree_delete_back_def : thm + val betree_be_tree_delete_fwd_def : thm + val betree_be_tree_insert_back_def : thm + val betree_be_tree_insert_fwd_def : thm + val betree_be_tree_lookup_back_def : thm + val betree_be_tree_lookup_fwd_def : thm + val betree_be_tree_new_fwd_def : thm + val betree_be_tree_upsert_back_def : thm + val betree_be_tree_upsert_fwd_def : thm + val betree_fresh_node_id_back_def : thm + val betree_fresh_node_id_fwd_def : thm + val betree_internal_flush_back_def : thm + val betree_internal_flush_fwd_def : thm + val betree_internal_lookup_in_children_back_def : thm + val betree_internal_lookup_in_children_fwd_def : thm + val betree_leaf_split_back_def : thm + val betree_leaf_split_fwd_def : thm + val betree_list_hd_fwd_def : thm + val betree_list_head_has_key_fwd_def : thm + val betree_list_len_fwd_def : thm + val betree_list_partition_at_pivot_fwd_def : thm + val betree_list_pop_front_back_def : thm + val betree_list_pop_front_fwd_def : thm + val betree_list_push_front_fwd_back_def : thm + val betree_list_split_at_fwd_def : thm + val betree_load_internal_node_fwd_def : thm + val betree_load_leaf_node_fwd_def : thm + val betree_node_apply_back_def : thm + val betree_node_apply_fwd_def : thm + val betree_node_apply_messages_back_def : thm + val betree_node_apply_messages_fwd_def : thm + val betree_node_apply_messages_to_internal_fwd_back_def : thm + val betree_node_apply_messages_to_leaf_fwd_back_def : thm + val betree_node_apply_to_internal_fwd_back_def : thm + val betree_node_apply_to_leaf_fwd_back_def : thm + val betree_node_apply_upserts_back_def : thm + val betree_node_apply_upserts_fwd_def : thm + val betree_node_filter_messages_for_key_fwd_back_def : thm + val betree_node_id_counter_fresh_id_back_def : thm + val betree_node_id_counter_fresh_id_fwd_def : thm + val betree_node_id_counter_new_fwd_def : thm + val betree_node_lookup_back_def : thm + val betree_node_lookup_first_message_after_key_back_def : thm + val betree_node_lookup_first_message_after_key_fwd_def : thm + val betree_node_lookup_first_message_for_key_back_def : thm + val betree_node_lookup_first_message_for_key_fwd_def : thm + val betree_node_lookup_fwd_def : thm + val betree_node_lookup_in_bindings_fwd_def : thm + val betree_node_lookup_mut_in_bindings_back_def : thm + val betree_node_lookup_mut_in_bindings_fwd_def : thm + val betree_store_internal_node_fwd_def : thm + val betree_store_leaf_node_fwd_def : thm + val betree_upsert_update_fwd_def : thm + val main_fwd_def : thm + + val betree_Funs_grammars : type_grammar.grammar * term_grammar.grammar +(* + [betree_Opaque] Parent theory of "betree_Funs" + + [betree_be_tree_apply_back_def] Definition + + ⊢ ∀self key msg st. + betree_be_tree_apply_back self key msg st = + do + (n,nic) <- + betree_node_apply_back self.betree_be_tree_root + self.betree_be_tree_params self.betree_be_tree_node_id_cnt + key msg st; + Return + (self with + <|betree_be_tree_node_id_cnt := nic; + betree_be_tree_root := n|>) + od + + [betree_be_tree_apply_fwd_def] Definition + + ⊢ ∀self key msg st. + betree_be_tree_apply_fwd self key msg st = + do + (st0,_) <- + betree_node_apply_fwd self.betree_be_tree_root + self.betree_be_tree_params self.betree_be_tree_node_id_cnt + key msg st; + monad_ignore_bind + (betree_node_apply_back self.betree_be_tree_root + self.betree_be_tree_params self.betree_be_tree_node_id_cnt + key msg st) (Return (st0,())) + od + + [betree_be_tree_delete_back_def] Definition + + ⊢ ∀self key st. + betree_be_tree_delete_back self key st = + betree_be_tree_apply_back self key BetreeMessageDelete st + + [betree_be_tree_delete_fwd_def] Definition + + ⊢ ∀self key st. + betree_be_tree_delete_fwd self key st = + do + (st0,_) <- + betree_be_tree_apply_fwd self key BetreeMessageDelete st; + monad_ignore_bind + (betree_be_tree_apply_back self key BetreeMessageDelete st) + (Return (st0,())) + od + + [betree_be_tree_insert_back_def] Definition + + ⊢ ∀self key value st. + betree_be_tree_insert_back self key value st = + betree_be_tree_apply_back self key (BetreeMessageInsert value) st + + [betree_be_tree_insert_fwd_def] Definition + + ⊢ ∀self key value st. + betree_be_tree_insert_fwd self key value st = + do + (st0,_) <- + betree_be_tree_apply_fwd self key (BetreeMessageInsert value) + st; + monad_ignore_bind + (betree_be_tree_apply_back self key + (BetreeMessageInsert value) st) (Return (st0,())) + od + + [betree_be_tree_lookup_back_def] Definition + + ⊢ ∀self key st. + betree_be_tree_lookup_back self key st = + do + n <- betree_node_lookup_back self.betree_be_tree_root key st; + Return (self with betree_be_tree_root := n) + od + + [betree_be_tree_lookup_fwd_def] Definition + + ⊢ ∀self key st. + betree_be_tree_lookup_fwd self key st = + betree_node_lookup_fwd self.betree_be_tree_root key st + + [betree_be_tree_new_fwd_def] Definition + + ⊢ ∀min_flush_size split_size st. + betree_be_tree_new_fwd min_flush_size split_size st = + do + node_id_cnt <- betree_node_id_counter_new_fwd; + id <- betree_node_id_counter_fresh_id_fwd node_id_cnt; + (st0,_) <- betree_store_leaf_node_fwd id BetreeListNil st; + node_id_cnt0 <- + betree_node_id_counter_fresh_id_back node_id_cnt; + Return + (st0, + <|betree_be_tree_params := + <|betree_params_min_flush_size := min_flush_size; + betree_params_split_size := split_size|>; + betree_be_tree_node_id_cnt := node_id_cnt0; + betree_be_tree_root := + BetreeNodeLeaf + <|betree_leaf_id := id; + betree_leaf_size := int_to_u64 0|> |>) + od + + [betree_be_tree_upsert_back_def] Definition + + ⊢ ∀self key upd st. + betree_be_tree_upsert_back self key upd st = + betree_be_tree_apply_back self key (BetreeMessageUpsert upd) st + + [betree_be_tree_upsert_fwd_def] Definition + + ⊢ ∀self key upd st. + betree_be_tree_upsert_fwd self key upd st = + do + (st0,_) <- + betree_be_tree_apply_fwd self key (BetreeMessageUpsert upd) + st; + monad_ignore_bind + (betree_be_tree_apply_back self key (BetreeMessageUpsert upd) + st) (Return (st0,())) + od + + [betree_fresh_node_id_back_def] Definition + + ⊢ ∀counter. + betree_fresh_node_id_back counter = + u64_add counter (int_to_u64 1) + + [betree_fresh_node_id_fwd_def] Definition + + ⊢ ∀counter. + betree_fresh_node_id_fwd counter = + monad_ignore_bind (u64_add counter (int_to_u64 1)) + (Return counter) + + [betree_internal_flush_back_def] Definition + + ⊢ ∀self params node_id_cnt content st. + betree_internal_flush_back self params node_id_cnt content st = + do + p <- + betree_list_partition_at_pivot_fwd content + self.betree_internal_pivot; + (msgs_left,msgs_right) <<- p; + len_left <- betree_list_len_fwd msgs_left; + if u64_ge len_left params.betree_params_min_flush_size then + do + (st0,_) <- + betree_node_apply_messages_fwd self.betree_internal_left + params node_id_cnt msgs_left st; + (n,node_id_cnt0) <- + betree_node_apply_messages_back self.betree_internal_left + params node_id_cnt msgs_left st; + len_right <- betree_list_len_fwd msgs_right; + if + u64_ge len_right params.betree_params_min_flush_size + then + do + (n0,node_id_cnt1) <- + betree_node_apply_messages_back + self.betree_internal_right params node_id_cnt0 + msgs_right st0; + Return + (self with + <|betree_internal_left := n; + betree_internal_right := n0|>,node_id_cnt1) + od + else + Return (self with betree_internal_left := n,node_id_cnt0) + od + else + do + (n,node_id_cnt0) <- + betree_node_apply_messages_back + self.betree_internal_right params node_id_cnt + msgs_right st; + Return (self with betree_internal_right := n,node_id_cnt0) + od + od + + [betree_internal_flush_fwd_def] Definition + + ⊢ ∀self params node_id_cnt content st. + betree_internal_flush_fwd self params node_id_cnt content st = + do + p <- + betree_list_partition_at_pivot_fwd content + self.betree_internal_pivot; + (msgs_left,msgs_right) <<- p; + len_left <- betree_list_len_fwd msgs_left; + if u64_ge len_left params.betree_params_min_flush_size then + do + (st0,_) <- + betree_node_apply_messages_fwd self.betree_internal_left + params node_id_cnt msgs_left st; + (_,node_id_cnt0) <- + betree_node_apply_messages_back self.betree_internal_left + params node_id_cnt msgs_left st; + len_right <- betree_list_len_fwd msgs_right; + if + u64_ge len_right params.betree_params_min_flush_size + then + do + (st1,_) <- + betree_node_apply_messages_fwd + self.betree_internal_right params node_id_cnt0 + msgs_right st0; + monad_ignore_bind + (betree_node_apply_messages_back + self.betree_internal_right params node_id_cnt0 + msgs_right st0) (Return (st1,BetreeListNil)) + od + else Return (st0,msgs_right) + od + else + do + (st0,_) <- + betree_node_apply_messages_fwd self.betree_internal_right + params node_id_cnt msgs_right st; + monad_ignore_bind + (betree_node_apply_messages_back + self.betree_internal_right params node_id_cnt + msgs_right st) (Return (st0,msgs_left)) + od + od + + [betree_internal_lookup_in_children_back_def] Definition + + ⊢ ∀self key st. + betree_internal_lookup_in_children_back self key st = + if u64_lt key self.betree_internal_pivot then + do + n <- betree_node_lookup_back self.betree_internal_left key st; + Return (self with betree_internal_left := n) + od + else + do + n <- + betree_node_lookup_back self.betree_internal_right key st; + Return (self with betree_internal_right := n) + od + + [betree_internal_lookup_in_children_fwd_def] Definition + + ⊢ ∀self key st. + betree_internal_lookup_in_children_fwd self key st = + if u64_lt key self.betree_internal_pivot then + betree_node_lookup_fwd self.betree_internal_left key st + else betree_node_lookup_fwd self.betree_internal_right key st + + [betree_leaf_split_back_def] Definition + + ⊢ ∀self content params node_id_cnt st. + betree_leaf_split_back self content params node_id_cnt st = + do + p <- + betree_list_split_at_fwd content + params.betree_params_split_size; + (content0,content1) <<- p; + monad_ignore_bind (betree_list_hd_fwd content1) + do + id0 <- betree_node_id_counter_fresh_id_fwd node_id_cnt; + node_id_cnt0 <- + betree_node_id_counter_fresh_id_back node_id_cnt; + id1 <- betree_node_id_counter_fresh_id_fwd node_id_cnt0; + (st0,_) <- betree_store_leaf_node_fwd id0 content0 st; + monad_ignore_bind + (betree_store_leaf_node_fwd id1 content1 st0) + (betree_node_id_counter_fresh_id_back node_id_cnt0) + od + od + + [betree_leaf_split_fwd_def] Definition + + ⊢ ∀self content params node_id_cnt st. + betree_leaf_split_fwd self content params node_id_cnt st = + do + p <- + betree_list_split_at_fwd content + params.betree_params_split_size; + (content0,content1) <<- p; + p0 <- betree_list_hd_fwd content1; + (pivot,_) <<- p0; + id0 <- betree_node_id_counter_fresh_id_fwd node_id_cnt; + node_id_cnt0 <- + betree_node_id_counter_fresh_id_back node_id_cnt; + id1 <- betree_node_id_counter_fresh_id_fwd node_id_cnt0; + (st0,_) <- betree_store_leaf_node_fwd id0 content0 st; + (st1,_) <- betree_store_leaf_node_fwd id1 content1 st0; + n <<- + BetreeNodeLeaf + <|betree_leaf_id := id0; + betree_leaf_size := params.betree_params_split_size|>; + n0 <<- + BetreeNodeLeaf + <|betree_leaf_id := id1; + betree_leaf_size := params.betree_params_split_size|>; + Return + (st1, + <|betree_internal_id := self.betree_leaf_id; + betree_internal_pivot := pivot; betree_internal_left := n; + betree_internal_right := n0|>) + od + + [betree_list_hd_fwd_def] Definition + + ⊢ ∀self. + betree_list_hd_fwd self = + case self of + BetreeListCons hd l => Return hd + | BetreeListNil => Fail Failure + + [betree_list_head_has_key_fwd_def] Definition + + ⊢ ∀self key. + betree_list_head_has_key_fwd self key = + case self of + BetreeListCons hd l => (let (i,_) = hd in Return (i = key)) + | BetreeListNil => Return F + + [betree_list_len_fwd_def] Definition + + ⊢ ∀self. + betree_list_len_fwd self = + case self of + BetreeListCons t tl => + do i <- betree_list_len_fwd tl; u64_add (int_to_u64 1) i od + | BetreeListNil => Return (int_to_u64 0) + + [betree_list_partition_at_pivot_fwd_def] Definition + + ⊢ ∀self pivot. + betree_list_partition_at_pivot_fwd self pivot = + case self of + BetreeListCons hd tl => + (let + (i,t) = hd + in + if u64_ge i pivot then + Return (BetreeListNil,BetreeListCons (i,t) tl) + else + do + p <- betree_list_partition_at_pivot_fwd tl pivot; + (ls0,ls1) <<- p; + l <<- ls0; + Return (BetreeListCons (i,t) l,ls1) + od) + | BetreeListNil => Return (BetreeListNil,BetreeListNil) + + [betree_list_pop_front_back_def] Definition + + ⊢ ∀self. + betree_list_pop_front_back self = + (let + ls = mem_replace_fwd self BetreeListNil + in + case ls of + BetreeListCons x tl => Return tl + | BetreeListNil => Fail Failure) + + [betree_list_pop_front_fwd_def] Definition + + ⊢ ∀self. + betree_list_pop_front_fwd self = + (let + ls = mem_replace_fwd self BetreeListNil + in + case ls of + BetreeListCons x tl => Return x + | BetreeListNil => Fail Failure) + + [betree_list_push_front_fwd_back_def] Definition + + ⊢ ∀self x. + betree_list_push_front_fwd_back self x = + (let + tl = mem_replace_fwd self BetreeListNil; + l = tl + in + Return (BetreeListCons x l)) + + [betree_list_split_at_fwd_def] Definition + + ⊢ ∀self n. + betree_list_split_at_fwd self n = + if n = int_to_u64 0 then Return (BetreeListNil,self) + else + case self of + BetreeListCons hd tl => + do + i <- u64_sub n (int_to_u64 1); + p <- betree_list_split_at_fwd tl i; + (ls0,ls1) <<- p; + l <<- ls0; + Return (BetreeListCons hd l,ls1) + od + | BetreeListNil => Fail Failure + + [betree_load_internal_node_fwd_def] Definition + + ⊢ ∀id st. + betree_load_internal_node_fwd id st = + betree_utils_load_internal_node_fwd id st + + [betree_load_leaf_node_fwd_def] Definition + + ⊢ ∀id st. + betree_load_leaf_node_fwd id st = + betree_utils_load_leaf_node_fwd id st + + [betree_node_apply_back_def] Definition + + ⊢ ∀self params node_id_cnt key new_msg st. + betree_node_apply_back self params node_id_cnt key new_msg st = + (let + l = BetreeListNil + in + betree_node_apply_messages_back self params node_id_cnt + (BetreeListCons (key,new_msg) l) st) + + [betree_node_apply_fwd_def] Definition + + ⊢ ∀self params node_id_cnt key new_msg st. + betree_node_apply_fwd self params node_id_cnt key new_msg st = + (let + l = BetreeListNil + in + do + (st0,_) <- + betree_node_apply_messages_fwd self params node_id_cnt + (BetreeListCons (key,new_msg) l) st; + monad_ignore_bind + (betree_node_apply_messages_back self params node_id_cnt + (BetreeListCons (key,new_msg) l) st) (Return (st0,())) + od) + + [betree_node_apply_messages_back_def] Definition + + ⊢ ∀self params node_id_cnt msgs st. + betree_node_apply_messages_back self params node_id_cnt msgs st = + case self of + BetreeNodeInternal node => + do + (st0,content) <- + betree_load_internal_node_fwd node.betree_internal_id st; + content0 <- + betree_node_apply_messages_to_internal_fwd_back content + msgs; + num_msgs <- betree_list_len_fwd content0; + if u64_ge num_msgs params.betree_params_min_flush_size then + do + (st1,content1) <- + betree_internal_flush_fwd node params node_id_cnt + content0 st0; + (node0,node_id_cnt0) <- + betree_internal_flush_back node params node_id_cnt + content0 st0; + monad_ignore_bind + (betree_store_internal_node_fwd + node0.betree_internal_id content1 st1) + (Return (BetreeNodeInternal node0,node_id_cnt0)) + od + else + monad_ignore_bind + (betree_store_internal_node_fwd node.betree_internal_id + content0 st0) + (Return (BetreeNodeInternal node,node_id_cnt)) + od + | BetreeNodeLeaf node' => + do + (st0,content) <- + betree_load_leaf_node_fwd node'.betree_leaf_id st; + content0 <- + betree_node_apply_messages_to_leaf_fwd_back content msgs; + len <- betree_list_len_fwd content0; + i <- u64_mul (int_to_u64 2) params.betree_params_split_size; + if u64_ge len i then + do + (st1,new_node) <- + betree_leaf_split_fwd node' content0 params node_id_cnt + st0; + monad_ignore_bind + (betree_store_leaf_node_fwd node'.betree_leaf_id + BetreeListNil st1) + do + node_id_cnt0 <- + betree_leaf_split_back node' content0 params + node_id_cnt st0; + Return (BetreeNodeInternal new_node,node_id_cnt0) + od + od + else + monad_ignore_bind + (betree_store_leaf_node_fwd node'.betree_leaf_id content0 + st0) + (Return + (BetreeNodeLeaf (node' with betree_leaf_size := len), + node_id_cnt)) + od + + [betree_node_apply_messages_fwd_def] Definition + + ⊢ ∀self params node_id_cnt msgs st. + betree_node_apply_messages_fwd self params node_id_cnt msgs st = + case self of + BetreeNodeInternal node => + do + (st0,content) <- + betree_load_internal_node_fwd node.betree_internal_id st; + content0 <- + betree_node_apply_messages_to_internal_fwd_back content + msgs; + num_msgs <- betree_list_len_fwd content0; + if u64_ge num_msgs params.betree_params_min_flush_size then + do + (st1,content1) <- + betree_internal_flush_fwd node params node_id_cnt + content0 st0; + (node0,_) <- + betree_internal_flush_back node params node_id_cnt + content0 st0; + (st2,_) <- + betree_store_internal_node_fwd + node0.betree_internal_id content1 st1; + Return (st2,()) + od + else + do + (st1,_) <- + betree_store_internal_node_fwd + node.betree_internal_id content0 st0; + Return (st1,()) + od + od + | BetreeNodeLeaf node' => + do + (st0,content) <- + betree_load_leaf_node_fwd node'.betree_leaf_id st; + content0 <- + betree_node_apply_messages_to_leaf_fwd_back content msgs; + len <- betree_list_len_fwd content0; + i <- u64_mul (int_to_u64 2) params.betree_params_split_size; + if u64_ge len i then + do + (st1,_) <- + betree_leaf_split_fwd node' content0 params node_id_cnt + st0; + (st2,_) <- + betree_store_leaf_node_fwd node'.betree_leaf_id + BetreeListNil st1; + Return (st2,()) + od + else + do + (st1,_) <- + betree_store_leaf_node_fwd node'.betree_leaf_id + content0 st0; + Return (st1,()) + od + od + + [betree_node_apply_messages_to_internal_fwd_back_def] Definition + + ⊢ ∀msgs new_msgs. + betree_node_apply_messages_to_internal_fwd_back msgs new_msgs = + case new_msgs of + BetreeListCons new_msg new_msgs_tl => + (let + (i,m) = new_msg + in + do + msgs0 <- betree_node_apply_to_internal_fwd_back msgs i m; + betree_node_apply_messages_to_internal_fwd_back msgs0 + new_msgs_tl + od) + | BetreeListNil => Return msgs + + [betree_node_apply_messages_to_leaf_fwd_back_def] Definition + + ⊢ ∀bindings new_msgs. + betree_node_apply_messages_to_leaf_fwd_back bindings new_msgs = + case new_msgs of + BetreeListCons new_msg new_msgs_tl => + (let + (i,m) = new_msg + in + do + bindings0 <- + betree_node_apply_to_leaf_fwd_back bindings i m; + betree_node_apply_messages_to_leaf_fwd_back bindings0 + new_msgs_tl + od) + | BetreeListNil => Return bindings + + [betree_node_apply_to_internal_fwd_back_def] Definition + + ⊢ ∀msgs key new_msg. + betree_node_apply_to_internal_fwd_back msgs key new_msg = + do + msgs0 <- betree_node_lookup_first_message_for_key_fwd key msgs; + b <- betree_list_head_has_key_fwd msgs0 key; + if b then + case new_msg of + BetreeMessageInsert i => + do + msgs1 <- + betree_node_filter_messages_for_key_fwd_back key + msgs0; + msgs2 <- + betree_list_push_front_fwd_back msgs1 + (key,BetreeMessageInsert i); + betree_node_lookup_first_message_for_key_back key msgs + msgs2 + od + | BetreeMessageDelete => + do + msgs1 <- + betree_node_filter_messages_for_key_fwd_back key msgs0; + msgs2 <- + betree_list_push_front_fwd_back msgs1 + (key,BetreeMessageDelete); + betree_node_lookup_first_message_for_key_back key msgs + msgs2 + od + | BetreeMessageUpsert s => + do + p <- betree_list_hd_fwd msgs0; + (_,m) <<- p; + case m of + BetreeMessageInsert prev => + do + v <- betree_upsert_update_fwd (SOME prev) s; + msgs1 <- betree_list_pop_front_back msgs0; + msgs2 <- + betree_list_push_front_fwd_back msgs1 + (key,BetreeMessageInsert v); + betree_node_lookup_first_message_for_key_back key + msgs msgs2 + od + | BetreeMessageDelete => + do + v <- betree_upsert_update_fwd NONE s; + msgs1 <- betree_list_pop_front_back msgs0; + msgs2 <- + betree_list_push_front_fwd_back msgs1 + (key,BetreeMessageInsert v); + betree_node_lookup_first_message_for_key_back key + msgs msgs2 + od + | BetreeMessageUpsert ufs => + do + msgs1 <- + betree_node_lookup_first_message_after_key_fwd key + msgs0; + msgs2 <- + betree_list_push_front_fwd_back msgs1 + (key,BetreeMessageUpsert s); + msgs3 <- + betree_node_lookup_first_message_after_key_back key + msgs0 msgs2; + betree_node_lookup_first_message_for_key_back key + msgs msgs3 + od + od + else + do + msgs1 <- + betree_list_push_front_fwd_back msgs0 (key,new_msg); + betree_node_lookup_first_message_for_key_back key msgs + msgs1 + od + od + + [betree_node_apply_to_leaf_fwd_back_def] Definition + + ⊢ ∀bindings key new_msg. + betree_node_apply_to_leaf_fwd_back bindings key new_msg = + do + bindings0 <- + betree_node_lookup_mut_in_bindings_fwd key bindings; + b <- betree_list_head_has_key_fwd bindings0 key; + if b then + do + hd <- betree_list_pop_front_fwd bindings0; + case new_msg of + BetreeMessageInsert v => + do + bindings1 <- betree_list_pop_front_back bindings0; + bindings2 <- + betree_list_push_front_fwd_back bindings1 (key,v); + betree_node_lookup_mut_in_bindings_back key bindings + bindings2 + od + | BetreeMessageDelete => + do + bindings1 <- betree_list_pop_front_back bindings0; + betree_node_lookup_mut_in_bindings_back key bindings + bindings1 + od + | BetreeMessageUpsert s => + (let + (_,i) = hd + in + do + v <- betree_upsert_update_fwd (SOME i) s; + bindings1 <- betree_list_pop_front_back bindings0; + bindings2 <- + betree_list_push_front_fwd_back bindings1 (key,v); + betree_node_lookup_mut_in_bindings_back key bindings + bindings2 + od) + od + else + case new_msg of + BetreeMessageInsert v => + do + bindings1 <- + betree_list_push_front_fwd_back bindings0 (key,v); + betree_node_lookup_mut_in_bindings_back key bindings + bindings1 + od + | BetreeMessageDelete => + betree_node_lookup_mut_in_bindings_back key bindings + bindings0 + | BetreeMessageUpsert s => + do + v <- betree_upsert_update_fwd NONE s; + bindings1 <- + betree_list_push_front_fwd_back bindings0 (key,v); + betree_node_lookup_mut_in_bindings_back key bindings + bindings1 + od + od + + [betree_node_apply_upserts_back_def] Definition + + ⊢ ∀msgs prev key st. + betree_node_apply_upserts_back msgs prev key st = + do + b <- betree_list_head_has_key_fwd msgs key; + if b then + do + msg <- betree_list_pop_front_fwd msgs; + (_,m) <<- msg; + case m of + BetreeMessageInsert i => Fail Failure + | BetreeMessageDelete => Fail Failure + | BetreeMessageUpsert s => + do + v <- betree_upsert_update_fwd prev s; + msgs0 <- betree_list_pop_front_back msgs; + betree_node_apply_upserts_back msgs0 (SOME v) key st + od + od + else + do + (_,v) <- core_option_option_unwrap_fwd prev st; + betree_list_push_front_fwd_back msgs + (key,BetreeMessageInsert v) + od + od + + [betree_node_apply_upserts_fwd_def] Definition + + ⊢ ∀msgs prev key st. + betree_node_apply_upserts_fwd msgs prev key st = + do + b <- betree_list_head_has_key_fwd msgs key; + if b then + do + msg <- betree_list_pop_front_fwd msgs; + (_,m) <<- msg; + case m of + BetreeMessageInsert i => Fail Failure + | BetreeMessageDelete => Fail Failure + | BetreeMessageUpsert s => + do + v <- betree_upsert_update_fwd prev s; + msgs0 <- betree_list_pop_front_back msgs; + betree_node_apply_upserts_fwd msgs0 (SOME v) key st + od + od + else + do + (st0,v) <- core_option_option_unwrap_fwd prev st; + monad_ignore_bind + (betree_list_push_front_fwd_back msgs + (key,BetreeMessageInsert v)) (Return (st0,v)) + od + od + + [betree_node_filter_messages_for_key_fwd_back_def] Definition + + ⊢ ∀key msgs. + betree_node_filter_messages_for_key_fwd_back key msgs = + case msgs of + BetreeListCons p l => + (let + (k,m) = p + in + if k = key then + do + msgs0 <- + betree_list_pop_front_back (BetreeListCons (k,m) l); + betree_node_filter_messages_for_key_fwd_back key msgs0 + od + else Return (BetreeListCons (k,m) l)) + | BetreeListNil => Return BetreeListNil + + [betree_node_id_counter_fresh_id_back_def] Definition + + ⊢ ∀self. + betree_node_id_counter_fresh_id_back self = + do + i <- + u64_add self.betree_node_id_counter_next_node_id + (int_to_u64 1); + Return <|betree_node_id_counter_next_node_id := i|> + od + + [betree_node_id_counter_fresh_id_fwd_def] Definition + + ⊢ ∀self. + betree_node_id_counter_fresh_id_fwd self = + monad_ignore_bind + (u64_add self.betree_node_id_counter_next_node_id + (int_to_u64 1)) + (Return self.betree_node_id_counter_next_node_id) + + [betree_node_id_counter_new_fwd_def] Definition + + ⊢ betree_node_id_counter_new_fwd = + Return <|betree_node_id_counter_next_node_id := int_to_u64 0|> + + [betree_node_lookup_back_def] Definition + + ⊢ ∀self key st. + betree_node_lookup_back self key st = + case self of + BetreeNodeInternal node => + do + (st0,msgs) <- + betree_load_internal_node_fwd node.betree_internal_id st; + pending <- + betree_node_lookup_first_message_for_key_fwd key msgs; + case pending of + BetreeListCons p l => + (let + (k,msg) = p + in + if k ≠ key then + monad_ignore_bind + (betree_node_lookup_first_message_for_key_back + key msgs (BetreeListCons (k,msg) l)) + do + node0 <- + betree_internal_lookup_in_children_back node + key st0; + Return (BetreeNodeInternal node0) + od + else + case msg of + BetreeMessageInsert v => + monad_ignore_bind + (betree_node_lookup_first_message_for_key_back + key msgs + (BetreeListCons (k,BetreeMessageInsert v) + l)) (Return (BetreeNodeInternal node)) + | BetreeMessageDelete => + monad_ignore_bind + (betree_node_lookup_first_message_for_key_back + key msgs + (BetreeListCons (k,BetreeMessageDelete) l)) + (Return (BetreeNodeInternal node)) + | BetreeMessageUpsert ufs => + do + (st1,v) <- + betree_internal_lookup_in_children_fwd node + key st0; + (st2,_) <- + betree_node_apply_upserts_fwd + (BetreeListCons + (k,BetreeMessageUpsert ufs) l) v key + st1; + node0 <- + betree_internal_lookup_in_children_back node + key st0; + pending0 <- + betree_node_apply_upserts_back + (BetreeListCons + (k,BetreeMessageUpsert ufs) l) v key + st1; + msgs0 <- + betree_node_lookup_first_message_for_key_back + key msgs pending0; + monad_ignore_bind + (betree_store_internal_node_fwd + node0.betree_internal_id msgs0 st2) + (Return (BetreeNodeInternal node0)) + od) + | BetreeListNil => + monad_ignore_bind + (betree_node_lookup_first_message_for_key_back key msgs + BetreeListNil) + do + node0 <- + betree_internal_lookup_in_children_back node key + st0; + Return (BetreeNodeInternal node0) + od + od + | BetreeNodeLeaf node' => + do + (_,bindings) <- + betree_load_leaf_node_fwd node'.betree_leaf_id st; + monad_ignore_bind + (betree_node_lookup_in_bindings_fwd key bindings) + (Return (BetreeNodeLeaf node')) + od + + [betree_node_lookup_first_message_after_key_back_def] Definition + + ⊢ ∀key msgs ret. + betree_node_lookup_first_message_after_key_back key msgs ret = + case msgs of + BetreeListCons p next_msgs => + (let + (k,m) = p + in + if k = key then + do + next_msgs0 <- + betree_node_lookup_first_message_after_key_back key + next_msgs ret; + Return (BetreeListCons (k,m) next_msgs0) + od + else Return ret) + | BetreeListNil => Return ret + + [betree_node_lookup_first_message_after_key_fwd_def] Definition + + ⊢ ∀key msgs. + betree_node_lookup_first_message_after_key_fwd key msgs = + case msgs of + BetreeListCons p next_msgs => + (let + (k,m) = p + in + if k = key then + betree_node_lookup_first_message_after_key_fwd key + next_msgs + else Return (BetreeListCons (k,m) next_msgs)) + | BetreeListNil => Return BetreeListNil + + [betree_node_lookup_first_message_for_key_back_def] Definition + + ⊢ ∀key msgs ret. + betree_node_lookup_first_message_for_key_back key msgs ret = + case msgs of + BetreeListCons x next_msgs => + (let + (i,m) = x + in + if u64_ge i key then Return ret + else + do + next_msgs0 <- + betree_node_lookup_first_message_for_key_back key + next_msgs ret; + Return (BetreeListCons (i,m) next_msgs0) + od) + | BetreeListNil => Return ret + + [betree_node_lookup_first_message_for_key_fwd_def] Definition + + ⊢ ∀key msgs. + betree_node_lookup_first_message_for_key_fwd key msgs = + case msgs of + BetreeListCons x next_msgs => + (let + (i,m) = x + in + if u64_ge i key then + Return (BetreeListCons (i,m) next_msgs) + else + betree_node_lookup_first_message_for_key_fwd key + next_msgs) + | BetreeListNil => Return BetreeListNil + + [betree_node_lookup_fwd_def] Definition + + ⊢ ∀self key st. + betree_node_lookup_fwd self key st = + case self of + BetreeNodeInternal node => + do + (st0,msgs) <- + betree_load_internal_node_fwd node.betree_internal_id st; + pending <- + betree_node_lookup_first_message_for_key_fwd key msgs; + case pending of + BetreeListCons p l => + (let + (k,msg) = p + in + if k ≠ key then + do + (st1,opt) <- + betree_internal_lookup_in_children_fwd node + key st0; + monad_ignore_bind + (betree_node_lookup_first_message_for_key_back + key msgs (BetreeListCons (k,msg) l)) + (Return (st1,opt)) + od + else + case msg of + BetreeMessageInsert v => + monad_ignore_bind + (betree_node_lookup_first_message_for_key_back + key msgs + (BetreeListCons (k,BetreeMessageInsert v) + l)) (Return (st0,SOME v)) + | BetreeMessageDelete => + monad_ignore_bind + (betree_node_lookup_first_message_for_key_back + key msgs + (BetreeListCons (k,BetreeMessageDelete) l)) + (Return (st0,NONE)) + | BetreeMessageUpsert ufs => + do + (st1,v) <- + betree_internal_lookup_in_children_fwd node + key st0; + (st2,v0) <- + betree_node_apply_upserts_fwd + (BetreeListCons + (k,BetreeMessageUpsert ufs) l) v key + st1; + node0 <- + betree_internal_lookup_in_children_back node + key st0; + pending0 <- + betree_node_apply_upserts_back + (BetreeListCons + (k,BetreeMessageUpsert ufs) l) v key + st1; + msgs0 <- + betree_node_lookup_first_message_for_key_back + key msgs pending0; + (st3,_) <- + betree_store_internal_node_fwd + node0.betree_internal_id msgs0 st2; + Return (st3,SOME v0) + od) + | BetreeListNil => + do + (st1,opt) <- + betree_internal_lookup_in_children_fwd node key st0; + monad_ignore_bind + (betree_node_lookup_first_message_for_key_back key + msgs BetreeListNil) (Return (st1,opt)) + od + od + | BetreeNodeLeaf node' => + do + (st0,bindings) <- + betree_load_leaf_node_fwd node'.betree_leaf_id st; + opt <- betree_node_lookup_in_bindings_fwd key bindings; + Return (st0,opt) + od + + [betree_node_lookup_in_bindings_fwd_def] Definition + + ⊢ ∀key bindings. + betree_node_lookup_in_bindings_fwd key bindings = + case bindings of + BetreeListCons hd tl => + (let + (i,i0) = hd + in + if i = key then Return (SOME i0) + else if u64_gt i key then Return NONE + else betree_node_lookup_in_bindings_fwd key tl) + | BetreeListNil => Return NONE + + [betree_node_lookup_mut_in_bindings_back_def] Definition + + ⊢ ∀key bindings ret. + betree_node_lookup_mut_in_bindings_back key bindings ret = + case bindings of + BetreeListCons hd tl => + (let + (i,i0) = hd + in + if u64_ge i key then Return ret + else + do + tl0 <- + betree_node_lookup_mut_in_bindings_back key tl ret; + Return (BetreeListCons (i,i0) tl0) + od) + | BetreeListNil => Return ret + + [betree_node_lookup_mut_in_bindings_fwd_def] Definition + + ⊢ ∀key bindings. + betree_node_lookup_mut_in_bindings_fwd key bindings = + case bindings of + BetreeListCons hd tl => + (let + (i,i0) = hd + in + if u64_ge i key then Return (BetreeListCons (i,i0) tl) + else betree_node_lookup_mut_in_bindings_fwd key tl) + | BetreeListNil => Return BetreeListNil + + [betree_store_internal_node_fwd_def] Definition + + ⊢ ∀id content st. + betree_store_internal_node_fwd id content st = + do + (st0,_) <- betree_utils_store_internal_node_fwd id content st; + Return (st0,()) + od + + [betree_store_leaf_node_fwd_def] Definition + + ⊢ ∀id content st. + betree_store_leaf_node_fwd id content st = + do + (st0,_) <- betree_utils_store_leaf_node_fwd id content st; + Return (st0,()) + od + + [betree_upsert_update_fwd_def] Definition + + ⊢ ∀prev st. + betree_upsert_update_fwd prev st = + case prev of + NONE => + (case st of + BetreeUpsertFunStateAdd v => Return v + | BetreeUpsertFunStateSub i => Return (int_to_u64 0)) + | SOME prev0 => + case st of + BetreeUpsertFunStateAdd v => + do + margin <- u64_sub core_u64_max prev0; + if u64_ge margin v then u64_add prev0 v + else Return core_u64_max + od + | BetreeUpsertFunStateSub v' => + if u64_ge prev0 v' then u64_sub prev0 v' + else Return (int_to_u64 0) + + [main_fwd_def] Definition + + ⊢ main_fwd = Return () + + +*) +end diff --git a/tests/hol4/betree/betree_OpaqueScript.sml b/tests/hol4/betree/betree_OpaqueScript.sml new file mode 100644 index 00000000..d8fdf6b5 --- /dev/null +++ b/tests/hol4/betree/betree_OpaqueScript.sml @@ -0,0 +1,26 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [betree]: external function declarations *) +open primitivesLib divDefLib +open betree_TypesTheory + +val _ = new_theory "betree_Opaque" + + +(** [betree::betree_utils::load_internal_node]: forward function *)val _ = new_constant ("betree_utils_load_internal_node_fwd", + “:u64 -> state -> (state # (u64 # betree_message_t) betree_list_t) + result”) + +(** [betree::betree_utils::store_internal_node]: forward function *)val _ = new_constant ("betree_utils_store_internal_node_fwd", + “:u64 -> (u64 # betree_message_t) betree_list_t -> state -> (state # unit) + result”) + +(** [betree::betree_utils::load_leaf_node]: forward function *)val _ = new_constant ("betree_utils_load_leaf_node_fwd", + “:u64 -> state -> (state # (u64 # u64) betree_list_t) result”) + +(** [betree::betree_utils::store_leaf_node]: forward function *)val _ = new_constant ("betree_utils_store_leaf_node_fwd", + “:u64 -> (u64 # u64) betree_list_t -> state -> (state # unit) result”) + +(** [core::option::Option::{0}::unwrap]: forward function *)val _ = new_constant ("core_option_option_unwrap_fwd", + “:'t option -> state -> (state # 't) result”) + +val _ = export_theory () diff --git a/tests/hol4/betree/betree_OpaqueTheory.sig b/tests/hol4/betree/betree_OpaqueTheory.sig new file mode 100644 index 00000000..9452ff0f --- /dev/null +++ b/tests/hol4/betree/betree_OpaqueTheory.sig @@ -0,0 +1,11 @@ +signature betree_OpaqueTheory = +sig + type thm = Thm.thm + + val betree_Opaque_grammars : type_grammar.grammar * term_grammar.grammar +(* + [betree_Types] Parent theory of "betree_Opaque" + + +*) +end diff --git a/tests/hol4/betree/betree_TypesScript.sml b/tests/hol4/betree/betree_TypesScript.sml new file mode 100644 index 00000000..93be2f46 --- /dev/null +++ b/tests/hol4/betree/betree_TypesScript.sml @@ -0,0 +1,76 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [betree]: type definitions *) +open primitivesLib divDefLib + +val _ = new_theory "betree_Types" + + +Datatype: + (** [betree::betree::List] *) + betree_list_t = | BetreeListCons 't betree_list_t | BetreeListNil +End + +Datatype: + (** [betree::betree::UpsertFunState] *) + betree_upsert_fun_state_t = + | BetreeUpsertFunStateAdd u64 + | BetreeUpsertFunStateSub u64 +End + +Datatype: + (** [betree::betree::Message] *) + betree_message_t = + | BetreeMessageInsert u64 + | BetreeMessageDelete + | BetreeMessageUpsert betree_upsert_fun_state_t +End + +Datatype: + (** [betree::betree::Leaf] *) + betree_leaf_t = <| betree_leaf_id : u64; betree_leaf_size : u64; |> +End + +Datatype: + (** [betree::betree::Internal] *) + betree_internal_t = + <| + betree_internal_id : u64; + betree_internal_pivot : u64; + betree_internal_left : betree_node_t; + betree_internal_right : betree_node_t; + |> + ; + + (** [betree::betree::Node] *) + betree_node_t = + | BetreeNodeInternal betree_internal_t + | BetreeNodeLeaf betree_leaf_t +End + +Datatype: + (** [betree::betree::Params] *) + betree_params_t = + <| + betree_params_min_flush_size : u64; betree_params_split_size : u64; + |> +End + +Datatype: + (** [betree::betree::NodeIdCounter] *) + betree_node_id_counter_t = <| betree_node_id_counter_next_node_id : u64; |> +End + +Datatype: + (** [betree::betree::BeTree] *) + betree_be_tree_t = + <| + betree_be_tree_params : betree_params_t; + betree_be_tree_node_id_cnt : betree_node_id_counter_t; + betree_be_tree_root : betree_node_t; + |> +End + +(** The state type used in the state-error monad *) +val _ = new_type ("state", 0) + +val _ = export_theory () diff --git a/tests/hol4/betree/betree_TypesTheory.sig b/tests/hol4/betree/betree_TypesTheory.sig new file mode 100644 index 00000000..ee2ec4b5 --- /dev/null +++ b/tests/hol4/betree/betree_TypesTheory.sig @@ -0,0 +1,1751 @@ +signature betree_TypesTheory = +sig + type thm = Thm.thm + + (* Definitions *) + val betree_be_tree_t_TY_DEF : thm + val betree_be_tree_t_betree_be_tree_node_id_cnt : thm + val betree_be_tree_t_betree_be_tree_node_id_cnt_fupd : thm + val betree_be_tree_t_betree_be_tree_params : thm + val betree_be_tree_t_betree_be_tree_params_fupd : thm + val betree_be_tree_t_betree_be_tree_root : thm + val betree_be_tree_t_betree_be_tree_root_fupd : thm + val betree_be_tree_t_case_def : thm + val betree_be_tree_t_size_def : thm + val betree_internal_t_TY_DEF : thm + val betree_internal_t_betree_internal_id : thm + val betree_internal_t_betree_internal_id_fupd : thm + val betree_internal_t_betree_internal_left : thm + val betree_internal_t_betree_internal_left_fupd : thm + val betree_internal_t_betree_internal_pivot : thm + val betree_internal_t_betree_internal_pivot_fupd : thm + val betree_internal_t_betree_internal_right : thm + val betree_internal_t_betree_internal_right_fupd : thm + val betree_internal_t_case_def : thm + val betree_internal_t_size_def : thm + val betree_leaf_t_TY_DEF : thm + val betree_leaf_t_betree_leaf_id : thm + val betree_leaf_t_betree_leaf_id_fupd : thm + val betree_leaf_t_betree_leaf_size : thm + val betree_leaf_t_betree_leaf_size_fupd : thm + val betree_leaf_t_case_def : thm + val betree_leaf_t_size_def : thm + val betree_list_t_TY_DEF : thm + val betree_list_t_case_def : thm + val betree_list_t_size_def : thm + val betree_message_t_TY_DEF : thm + val betree_message_t_case_def : thm + val betree_message_t_size_def : thm + val betree_node_id_counter_t_TY_DEF : thm + val betree_node_id_counter_t_betree_node_id_counter_next_node_id : thm + val betree_node_id_counter_t_betree_node_id_counter_next_node_id_fupd : thm + val betree_node_id_counter_t_case_def : thm + val betree_node_id_counter_t_size_def : thm + val betree_node_t_TY_DEF : thm + val betree_node_t_case_def : thm + val betree_params_t_TY_DEF : thm + val betree_params_t_betree_params_min_flush_size : thm + val betree_params_t_betree_params_min_flush_size_fupd : thm + val betree_params_t_betree_params_split_size : thm + val betree_params_t_betree_params_split_size_fupd : thm + val betree_params_t_case_def : thm + val betree_params_t_size_def : thm + val betree_upsert_fun_state_t_TY_DEF : thm + val betree_upsert_fun_state_t_case_def : thm + val betree_upsert_fun_state_t_size_def : thm + + (* Theorems *) + val EXISTS_betree_be_tree_t : thm + val EXISTS_betree_internal_t : thm + val EXISTS_betree_leaf_t : thm + val EXISTS_betree_node_id_counter_t : thm + val EXISTS_betree_params_t : thm + val FORALL_betree_be_tree_t : thm + val FORALL_betree_internal_t : thm + val FORALL_betree_leaf_t : thm + val FORALL_betree_node_id_counter_t : thm + val FORALL_betree_params_t : thm + val betree_be_tree_t_11 : thm + val betree_be_tree_t_Axiom : thm + val betree_be_tree_t_accessors : thm + val betree_be_tree_t_accfupds : thm + val betree_be_tree_t_case_cong : thm + val betree_be_tree_t_case_eq : thm + val betree_be_tree_t_component_equality : thm + val betree_be_tree_t_fn_updates : thm + val betree_be_tree_t_fupdcanon : thm + val betree_be_tree_t_fupdcanon_comp : thm + val betree_be_tree_t_fupdfupds : thm + val betree_be_tree_t_fupdfupds_comp : thm + val betree_be_tree_t_induction : thm + val betree_be_tree_t_literal_11 : thm + val betree_be_tree_t_literal_nchotomy : thm + val betree_be_tree_t_nchotomy : thm + val betree_be_tree_t_updates_eq_literal : thm + val betree_internal_t_11 : thm + val betree_internal_t_Axiom : thm + val betree_internal_t_accessors : thm + val betree_internal_t_accfupds : thm + val betree_internal_t_case_cong : thm + val betree_internal_t_case_eq : thm + val betree_internal_t_component_equality : thm + val betree_internal_t_fn_updates : thm + val betree_internal_t_fupdcanon : thm + val betree_internal_t_fupdcanon_comp : thm + val betree_internal_t_fupdfupds : thm + val betree_internal_t_fupdfupds_comp : thm + val betree_internal_t_induction : thm + val betree_internal_t_literal_11 : thm + val betree_internal_t_literal_nchotomy : thm + val betree_internal_t_nchotomy : thm + val betree_internal_t_updates_eq_literal : thm + val betree_leaf_t_11 : thm + val betree_leaf_t_Axiom : thm + val betree_leaf_t_accessors : thm + val betree_leaf_t_accfupds : thm + val betree_leaf_t_case_cong : thm + val betree_leaf_t_case_eq : thm + val betree_leaf_t_component_equality : thm + val betree_leaf_t_fn_updates : thm + val betree_leaf_t_fupdcanon : thm + val betree_leaf_t_fupdcanon_comp : thm + val betree_leaf_t_fupdfupds : thm + val betree_leaf_t_fupdfupds_comp : thm + val betree_leaf_t_induction : thm + val betree_leaf_t_literal_11 : thm + val betree_leaf_t_literal_nchotomy : thm + val betree_leaf_t_nchotomy : thm + val betree_leaf_t_updates_eq_literal : thm + val betree_list_t_11 : thm + val betree_list_t_Axiom : thm + val betree_list_t_case_cong : thm + val betree_list_t_case_eq : thm + val betree_list_t_distinct : thm + val betree_list_t_induction : thm + val betree_list_t_nchotomy : thm + val betree_message_t_11 : thm + val betree_message_t_Axiom : thm + val betree_message_t_case_cong : thm + val betree_message_t_case_eq : thm + val betree_message_t_distinct : thm + val betree_message_t_induction : thm + val betree_message_t_nchotomy : thm + val betree_node_id_counter_t_11 : thm + val betree_node_id_counter_t_Axiom : thm + val betree_node_id_counter_t_accessors : thm + val betree_node_id_counter_t_accfupds : thm + val betree_node_id_counter_t_case_cong : thm + val betree_node_id_counter_t_case_eq : thm + val betree_node_id_counter_t_component_equality : thm + val betree_node_id_counter_t_fn_updates : thm + val betree_node_id_counter_t_fupdfupds : thm + val betree_node_id_counter_t_fupdfupds_comp : thm + val betree_node_id_counter_t_induction : thm + val betree_node_id_counter_t_literal_11 : thm + val betree_node_id_counter_t_literal_nchotomy : thm + val betree_node_id_counter_t_nchotomy : thm + val betree_node_id_counter_t_updates_eq_literal : thm + val betree_node_t_11 : thm + val betree_node_t_Axiom : thm + val betree_node_t_case_cong : thm + val betree_node_t_case_eq : thm + val betree_node_t_distinct : thm + val betree_node_t_induction : thm + val betree_node_t_nchotomy : thm + val betree_params_t_11 : thm + val betree_params_t_Axiom : thm + val betree_params_t_accessors : thm + val betree_params_t_accfupds : thm + val betree_params_t_case_cong : thm + val betree_params_t_case_eq : thm + val betree_params_t_component_equality : thm + val betree_params_t_fn_updates : thm + val betree_params_t_fupdcanon : thm + val betree_params_t_fupdcanon_comp : thm + val betree_params_t_fupdfupds : thm + val betree_params_t_fupdfupds_comp : thm + val betree_params_t_induction : thm + val betree_params_t_literal_11 : thm + val betree_params_t_literal_nchotomy : thm + val betree_params_t_nchotomy : thm + val betree_params_t_updates_eq_literal : thm + val betree_upsert_fun_state_t_11 : thm + val betree_upsert_fun_state_t_Axiom : thm + val betree_upsert_fun_state_t_case_cong : thm + val betree_upsert_fun_state_t_case_eq : thm + val betree_upsert_fun_state_t_distinct : thm + val betree_upsert_fun_state_t_induction : thm + val betree_upsert_fun_state_t_nchotomy : thm + val datatype_betree_be_tree_t : thm + val datatype_betree_internal_t : thm + val datatype_betree_leaf_t : thm + val datatype_betree_list_t : thm + val datatype_betree_message_t : thm + val datatype_betree_node_id_counter_t : thm + val datatype_betree_params_t : thm + val datatype_betree_upsert_fun_state_t : thm + + val betree_Types_grammars : type_grammar.grammar * term_grammar.grammar +(* + [divDef] Parent theory of "betree_Types" + + [betree_be_tree_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('betree_be_tree_t'). + (∀a0'. + (∃a0 a1 a2. + a0' = + (λa0 a1 a2. + ind_type$CONSTR 0 (a0,a1,a2) + (λn. ind_type$BOTTOM)) a0 a1 a2) ⇒ + $var$('betree_be_tree_t') a0') ⇒ + $var$('betree_be_tree_t') a0') rep + + [betree_be_tree_t_betree_be_tree_node_id_cnt] Definition + + ⊢ ∀b b0 b1. + (betree_be_tree_t b b0 b1).betree_be_tree_node_id_cnt = b0 + + [betree_be_tree_t_betree_be_tree_node_id_cnt_fupd] Definition + + ⊢ ∀f b b0 b1. + betree_be_tree_t b b0 b1 with + betree_be_tree_node_id_cnt updated_by f = + betree_be_tree_t b (f b0) b1 + + [betree_be_tree_t_betree_be_tree_params] Definition + + ⊢ ∀b b0 b1. (betree_be_tree_t b b0 b1).betree_be_tree_params = b + + [betree_be_tree_t_betree_be_tree_params_fupd] Definition + + ⊢ ∀f b b0 b1. + betree_be_tree_t b b0 b1 with betree_be_tree_params updated_by f = + betree_be_tree_t (f b) b0 b1 + + [betree_be_tree_t_betree_be_tree_root] Definition + + ⊢ ∀b b0 b1. (betree_be_tree_t b b0 b1).betree_be_tree_root = b1 + + [betree_be_tree_t_betree_be_tree_root_fupd] Definition + + ⊢ ∀f b b0 b1. + betree_be_tree_t b b0 b1 with betree_be_tree_root updated_by f = + betree_be_tree_t b b0 (f b1) + + [betree_be_tree_t_case_def] Definition + + ⊢ ∀a0 a1 a2 f. + betree_be_tree_t_CASE (betree_be_tree_t a0 a1 a2) f = f a0 a1 a2 + + [betree_be_tree_t_size_def] Definition + + ⊢ ∀a0 a1 a2. + betree_be_tree_t_size (betree_be_tree_t a0 a1 a2) = + 1 + + (betree_params_t_size a0 + + (betree_node_id_counter_t_size a1 + betree_node_t_size a2)) + + [betree_internal_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('betree_internal_t') $var$('betree_node_t'). + (∀a0'. + (∃a0 a1 a2 a3. + a0' = + (λa0 a1 a2 a3. + ind_type$CONSTR 0 (a0,a1,ARB) + (ind_type$FCONS a2 + (ind_type$FCONS a3 (λn. ind_type$BOTTOM)))) + a0 a1 a2 a3 ∧ $var$('betree_node_t') a2 ∧ + $var$('betree_node_t') a3) ⇒ + $var$('betree_internal_t') a0') ∧ + (∀a1'. + (∃a. a1' = + (λa. + ind_type$CONSTR (SUC 0) (ARB,ARB,ARB) + (ind_type$FCONS a (λn. ind_type$BOTTOM))) + a ∧ $var$('betree_internal_t') a) ∨ + (∃a. a1' = + (λa. + ind_type$CONSTR (SUC (SUC 0)) (ARB,ARB,a) + (λn. ind_type$BOTTOM)) a) ⇒ + $var$('betree_node_t') a1') ⇒ + $var$('betree_internal_t') a0') rep + + [betree_internal_t_betree_internal_id] Definition + + ⊢ ∀u u0 b b0. (betree_internal_t u u0 b b0).betree_internal_id = u + + [betree_internal_t_betree_internal_id_fupd] Definition + + ⊢ ∀f u u0 b b0. + betree_internal_t u u0 b b0 with betree_internal_id updated_by f = + betree_internal_t (f u) u0 b b0 + + [betree_internal_t_betree_internal_left] Definition + + ⊢ ∀u u0 b b0. (betree_internal_t u u0 b b0).betree_internal_left = b + + [betree_internal_t_betree_internal_left_fupd] Definition + + ⊢ ∀f u u0 b b0. + betree_internal_t u u0 b b0 with + betree_internal_left updated_by f = + betree_internal_t u u0 (f b) b0 + + [betree_internal_t_betree_internal_pivot] Definition + + ⊢ ∀u u0 b b0. + (betree_internal_t u u0 b b0).betree_internal_pivot = u0 + + [betree_internal_t_betree_internal_pivot_fupd] Definition + + ⊢ ∀f u u0 b b0. + betree_internal_t u u0 b b0 with + betree_internal_pivot updated_by f = + betree_internal_t u (f u0) b b0 + + [betree_internal_t_betree_internal_right] Definition + + ⊢ ∀u u0 b b0. + (betree_internal_t u u0 b b0).betree_internal_right = b0 + + [betree_internal_t_betree_internal_right_fupd] Definition + + ⊢ ∀f u u0 b b0. + betree_internal_t u u0 b b0 with + betree_internal_right updated_by f = + betree_internal_t u u0 b (f b0) + + [betree_internal_t_case_def] Definition + + ⊢ ∀a0 a1 a2 a3 f. + betree_internal_t_CASE (betree_internal_t a0 a1 a2 a3) f = + f a0 a1 a2 a3 + + [betree_internal_t_size_def] Definition + + ⊢ (∀a0 a1 a2 a3. + betree_internal_t_size (betree_internal_t a0 a1 a2 a3) = + 1 + (betree_node_t_size a2 + betree_node_t_size a3)) ∧ + (∀a. betree_node_t_size (BetreeNodeInternal a) = + 1 + betree_internal_t_size a) ∧ + ∀a. betree_node_t_size (BetreeNodeLeaf a) = + 1 + betree_leaf_t_size a + + [betree_leaf_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('betree_leaf_t'). + (∀a0'. + (∃a0 a1. + a0' = + (λa0 a1. + ind_type$CONSTR 0 (a0,a1) + (λn. ind_type$BOTTOM)) a0 a1) ⇒ + $var$('betree_leaf_t') a0') ⇒ + $var$('betree_leaf_t') a0') rep + + [betree_leaf_t_betree_leaf_id] Definition + + ⊢ ∀u u0. (betree_leaf_t u u0).betree_leaf_id = u + + [betree_leaf_t_betree_leaf_id_fupd] Definition + + ⊢ ∀f u u0. + betree_leaf_t u u0 with betree_leaf_id updated_by f = + betree_leaf_t (f u) u0 + + [betree_leaf_t_betree_leaf_size] Definition + + ⊢ ∀u u0. (betree_leaf_t u u0).betree_leaf_size = u0 + + [betree_leaf_t_betree_leaf_size_fupd] Definition + + ⊢ ∀f u u0. + betree_leaf_t u u0 with betree_leaf_size updated_by f = + betree_leaf_t u (f u0) + + [betree_leaf_t_case_def] Definition + + ⊢ ∀a0 a1 f. betree_leaf_t_CASE (betree_leaf_t a0 a1) f = f a0 a1 + + [betree_leaf_t_size_def] Definition + + ⊢ ∀a0 a1. betree_leaf_t_size (betree_leaf_t a0 a1) = 1 + + [betree_list_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('betree_list_t'). + (∀a0'. + (∃a0 a1. + a0' = + (λa0 a1. + ind_type$CONSTR 0 a0 + (ind_type$FCONS a1 (λn. ind_type$BOTTOM))) + a0 a1 ∧ $var$('betree_list_t') a1) ∨ + a0' = + ind_type$CONSTR (SUC 0) ARB (λn. ind_type$BOTTOM) ⇒ + $var$('betree_list_t') a0') ⇒ + $var$('betree_list_t') a0') rep + + [betree_list_t_case_def] Definition + + ⊢ (∀a0 a1 f v. + betree_list_t_CASE (BetreeListCons a0 a1) f v = f a0 a1) ∧ + ∀f v. betree_list_t_CASE BetreeListNil f v = v + + [betree_list_t_size_def] Definition + + ⊢ (∀f a0 a1. + betree_list_t_size f (BetreeListCons a0 a1) = + 1 + (f a0 + betree_list_t_size f a1)) ∧ + ∀f. betree_list_t_size f BetreeListNil = 0 + + [betree_message_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0. + ∀ $var$('betree_message_t'). + (∀a0. + (∃a. a0 = + (λa. + ind_type$CONSTR 0 (a,ARB) + (λn. ind_type$BOTTOM)) a) ∨ + a0 = + ind_type$CONSTR (SUC 0) (ARB,ARB) + (λn. ind_type$BOTTOM) ∨ + (∃a. a0 = + (λa. + ind_type$CONSTR (SUC (SUC 0)) (ARB,a) + (λn. ind_type$BOTTOM)) a) ⇒ + $var$('betree_message_t') a0) ⇒ + $var$('betree_message_t') a0) rep + + [betree_message_t_case_def] Definition + + ⊢ (∀a f v f1. + betree_message_t_CASE (BetreeMessageInsert a) f v f1 = f a) ∧ + (∀f v f1. betree_message_t_CASE BetreeMessageDelete f v f1 = v) ∧ + ∀a f v f1. + betree_message_t_CASE (BetreeMessageUpsert a) f v f1 = f1 a + + [betree_message_t_size_def] Definition + + ⊢ (∀a. betree_message_t_size (BetreeMessageInsert a) = 1) ∧ + betree_message_t_size BetreeMessageDelete = 0 ∧ + ∀a. betree_message_t_size (BetreeMessageUpsert a) = + 1 + betree_upsert_fun_state_t_size a + + [betree_node_id_counter_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0. + ∀ $var$('betree_node_id_counter_t'). + (∀a0. + (∃a. a0 = + (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) + a) ⇒ + $var$('betree_node_id_counter_t') a0) ⇒ + $var$('betree_node_id_counter_t') a0) rep + + [betree_node_id_counter_t_betree_node_id_counter_next_node_id] Definition + + ⊢ ∀u. (betree_node_id_counter_t u). + betree_node_id_counter_next_node_id = + u + + [betree_node_id_counter_t_betree_node_id_counter_next_node_id_fupd] Definition + + ⊢ ∀f u. + betree_node_id_counter_t u with + betree_node_id_counter_next_node_id updated_by f = + betree_node_id_counter_t (f u) + + [betree_node_id_counter_t_case_def] Definition + + ⊢ ∀a f. + betree_node_id_counter_t_CASE (betree_node_id_counter_t a) f = + f a + + [betree_node_id_counter_t_size_def] Definition + + ⊢ ∀a. betree_node_id_counter_t_size (betree_node_id_counter_t a) = 1 + + [betree_node_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa1'. + ∀ $var$('betree_internal_t') $var$('betree_node_t'). + (∀a0'. + (∃a0 a1 a2 a3. + a0' = + (λa0 a1 a2 a3. + ind_type$CONSTR 0 (a0,a1,ARB) + (ind_type$FCONS a2 + (ind_type$FCONS a3 (λn. ind_type$BOTTOM)))) + a0 a1 a2 a3 ∧ $var$('betree_node_t') a2 ∧ + $var$('betree_node_t') a3) ⇒ + $var$('betree_internal_t') a0') ∧ + (∀a1'. + (∃a. a1' = + (λa. + ind_type$CONSTR (SUC 0) (ARB,ARB,ARB) + (ind_type$FCONS a (λn. ind_type$BOTTOM))) + a ∧ $var$('betree_internal_t') a) ∨ + (∃a. a1' = + (λa. + ind_type$CONSTR (SUC (SUC 0)) (ARB,ARB,a) + (λn. ind_type$BOTTOM)) a) ⇒ + $var$('betree_node_t') a1') ⇒ + $var$('betree_node_t') a1') rep + + [betree_node_t_case_def] Definition + + ⊢ (∀a f f1. betree_node_t_CASE (BetreeNodeInternal a) f f1 = f a) ∧ + ∀a f f1. betree_node_t_CASE (BetreeNodeLeaf a) f f1 = f1 a + + [betree_params_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0'. + ∀ $var$('betree_params_t'). + (∀a0'. + (∃a0 a1. + a0' = + (λa0 a1. + ind_type$CONSTR 0 (a0,a1) + (λn. ind_type$BOTTOM)) a0 a1) ⇒ + $var$('betree_params_t') a0') ⇒ + $var$('betree_params_t') a0') rep + + [betree_params_t_betree_params_min_flush_size] Definition + + ⊢ ∀u u0. (betree_params_t u u0).betree_params_min_flush_size = u + + [betree_params_t_betree_params_min_flush_size_fupd] Definition + + ⊢ ∀f u u0. + betree_params_t u u0 with + betree_params_min_flush_size updated_by f = + betree_params_t (f u) u0 + + [betree_params_t_betree_params_split_size] Definition + + ⊢ ∀u u0. (betree_params_t u u0).betree_params_split_size = u0 + + [betree_params_t_betree_params_split_size_fupd] Definition + + ⊢ ∀f u u0. + betree_params_t u u0 with betree_params_split_size updated_by f = + betree_params_t u (f u0) + + [betree_params_t_case_def] Definition + + ⊢ ∀a0 a1 f. betree_params_t_CASE (betree_params_t a0 a1) f = f a0 a1 + + [betree_params_t_size_def] Definition + + ⊢ ∀a0 a1. betree_params_t_size (betree_params_t a0 a1) = 1 + + [betree_upsert_fun_state_t_TY_DEF] Definition + + ⊢ ∃rep. + TYPE_DEFINITION + (λa0. + ∀ $var$('betree_upsert_fun_state_t'). + (∀a0. + (∃a. a0 = + (λa. ind_type$CONSTR 0 a (λn. ind_type$BOTTOM)) + a) ∨ + (∃a. a0 = + (λa. + ind_type$CONSTR (SUC 0) a + (λn. ind_type$BOTTOM)) a) ⇒ + $var$('betree_upsert_fun_state_t') a0) ⇒ + $var$('betree_upsert_fun_state_t') a0) rep + + [betree_upsert_fun_state_t_case_def] Definition + + ⊢ (∀a f f1. + betree_upsert_fun_state_t_CASE (BetreeUpsertFunStateAdd a) f f1 = + f a) ∧ + ∀a f f1. + betree_upsert_fun_state_t_CASE (BetreeUpsertFunStateSub a) f f1 = + f1 a + + [betree_upsert_fun_state_t_size_def] Definition + + ⊢ (∀a. betree_upsert_fun_state_t_size (BetreeUpsertFunStateAdd a) = 1) ∧ + ∀a. betree_upsert_fun_state_t_size (BetreeUpsertFunStateSub a) = 1 + + [EXISTS_betree_be_tree_t] Theorem + + ⊢ ∀P. (∃b. P b) ⇔ + ∃b2 b1 b0. + P + <|betree_be_tree_params := b2; + betree_be_tree_node_id_cnt := b1; + betree_be_tree_root := b0|> + + [EXISTS_betree_internal_t] Theorem + + ⊢ ∀P. (∃b. P b) ⇔ + ∃u0 u b1 b0. + P + <|betree_internal_id := u0; betree_internal_pivot := u; + betree_internal_left := b1; betree_internal_right := b0|> + + [EXISTS_betree_leaf_t] Theorem + + ⊢ ∀P. (∃b. P b) ⇔ + ∃u0 u. P <|betree_leaf_id := u0; betree_leaf_size := u|> + + [EXISTS_betree_node_id_counter_t] Theorem + + ⊢ ∀P. (∃b. P b) ⇔ ∃u. P <|betree_node_id_counter_next_node_id := u|> + + [EXISTS_betree_params_t] Theorem + + ⊢ ∀P. (∃b. P b) ⇔ + ∃u0 u. + P + <|betree_params_min_flush_size := u0; + betree_params_split_size := u|> + + [FORALL_betree_be_tree_t] Theorem + + ⊢ ∀P. (∀b. P b) ⇔ + ∀b2 b1 b0. + P + <|betree_be_tree_params := b2; + betree_be_tree_node_id_cnt := b1; + betree_be_tree_root := b0|> + + [FORALL_betree_internal_t] Theorem + + ⊢ ∀P. (∀b. P b) ⇔ + ∀u0 u b1 b0. + P + <|betree_internal_id := u0; betree_internal_pivot := u; + betree_internal_left := b1; betree_internal_right := b0|> + + [FORALL_betree_leaf_t] Theorem + + ⊢ ∀P. (∀b. P b) ⇔ + ∀u0 u. P <|betree_leaf_id := u0; betree_leaf_size := u|> + + [FORALL_betree_node_id_counter_t] Theorem + + ⊢ ∀P. (∀b. P b) ⇔ ∀u. P <|betree_node_id_counter_next_node_id := u|> + + [FORALL_betree_params_t] Theorem + + ⊢ ∀P. (∀b. P b) ⇔ + ∀u0 u. + P + <|betree_params_min_flush_size := u0; + betree_params_split_size := u|> + + [betree_be_tree_t_11] Theorem + + ⊢ ∀a0 a1 a2 a0' a1' a2'. + betree_be_tree_t a0 a1 a2 = betree_be_tree_t a0' a1' a2' ⇔ + a0 = a0' ∧ a1 = a1' ∧ a2 = a2' + + [betree_be_tree_t_Axiom] Theorem + + ⊢ ∀f. ∃fn. ∀a0 a1 a2. fn (betree_be_tree_t a0 a1 a2) = f a0 a1 a2 + + [betree_be_tree_t_accessors] Theorem + + ⊢ (∀b b0 b1. (betree_be_tree_t b b0 b1).betree_be_tree_params = b) ∧ + (∀b b0 b1. + (betree_be_tree_t b b0 b1).betree_be_tree_node_id_cnt = b0) ∧ + ∀b b0 b1. (betree_be_tree_t b b0 b1).betree_be_tree_root = b1 + + [betree_be_tree_t_accfupds] Theorem + + ⊢ (∀f b. + (b with betree_be_tree_node_id_cnt updated_by f). + betree_be_tree_params = + b.betree_be_tree_params) ∧ + (∀f b. + (b with betree_be_tree_root updated_by f).betree_be_tree_params = + b.betree_be_tree_params) ∧ + (∀f b. + (b with betree_be_tree_params updated_by f). + betree_be_tree_node_id_cnt = + b.betree_be_tree_node_id_cnt) ∧ + (∀f b. + (b with betree_be_tree_root updated_by f). + betree_be_tree_node_id_cnt = + b.betree_be_tree_node_id_cnt) ∧ + (∀f b. + (b with betree_be_tree_params updated_by f).betree_be_tree_root = + b.betree_be_tree_root) ∧ + (∀f b. + (b with betree_be_tree_node_id_cnt updated_by f). + betree_be_tree_root = + b.betree_be_tree_root) ∧ + (∀f b. + (b with betree_be_tree_params updated_by f). + betree_be_tree_params = + f b.betree_be_tree_params) ∧ + (∀f b. + (b with betree_be_tree_node_id_cnt updated_by f). + betree_be_tree_node_id_cnt = + f b.betree_be_tree_node_id_cnt) ∧ + ∀f b. + (b with betree_be_tree_root updated_by f).betree_be_tree_root = + f b.betree_be_tree_root + + [betree_be_tree_t_case_cong] Theorem + + ⊢ ∀M M' f. + M = M' ∧ + (∀a0 a1 a2. + M' = betree_be_tree_t a0 a1 a2 ⇒ f a0 a1 a2 = f' a0 a1 a2) ⇒ + betree_be_tree_t_CASE M f = betree_be_tree_t_CASE M' f' + + [betree_be_tree_t_case_eq] Theorem + + ⊢ betree_be_tree_t_CASE x f = v ⇔ + ∃b b0 b1. x = betree_be_tree_t b b0 b1 ∧ f b b0 b1 = v + + [betree_be_tree_t_component_equality] Theorem + + ⊢ ∀b1 b2. + b1 = b2 ⇔ + b1.betree_be_tree_params = b2.betree_be_tree_params ∧ + b1.betree_be_tree_node_id_cnt = b2.betree_be_tree_node_id_cnt ∧ + b1.betree_be_tree_root = b2.betree_be_tree_root + + [betree_be_tree_t_fn_updates] Theorem + + ⊢ (∀f b b0 b1. + betree_be_tree_t b b0 b1 with betree_be_tree_params updated_by f = + betree_be_tree_t (f b) b0 b1) ∧ + (∀f b b0 b1. + betree_be_tree_t b b0 b1 with + betree_be_tree_node_id_cnt updated_by f = + betree_be_tree_t b (f b0) b1) ∧ + ∀f b b0 b1. + betree_be_tree_t b b0 b1 with betree_be_tree_root updated_by f = + betree_be_tree_t b b0 (f b1) + + [betree_be_tree_t_fupdcanon] Theorem + + ⊢ (∀g f b. + b with + <|betree_be_tree_node_id_cnt updated_by f; + betree_be_tree_params updated_by g|> = + b with + <|betree_be_tree_params updated_by g; + betree_be_tree_node_id_cnt updated_by f|>) ∧ + (∀g f b. + b with + <|betree_be_tree_root updated_by f; + betree_be_tree_params updated_by g|> = + b with + <|betree_be_tree_params updated_by g; + betree_be_tree_root updated_by f|>) ∧ + ∀g f b. + b with + <|betree_be_tree_root updated_by f; + betree_be_tree_node_id_cnt updated_by g|> = + b with + <|betree_be_tree_node_id_cnt updated_by g; + betree_be_tree_root updated_by f|> + + [betree_be_tree_t_fupdcanon_comp] Theorem + + ⊢ ((∀g f. + betree_be_tree_node_id_cnt_fupd f ∘ + betree_be_tree_params_fupd g = + betree_be_tree_params_fupd g ∘ + betree_be_tree_node_id_cnt_fupd f) ∧ + ∀h g f. + betree_be_tree_node_id_cnt_fupd f ∘ + betree_be_tree_params_fupd g ∘ h = + betree_be_tree_params_fupd g ∘ + betree_be_tree_node_id_cnt_fupd f ∘ h) ∧ + ((∀g f. + betree_be_tree_root_fupd f ∘ betree_be_tree_params_fupd g = + betree_be_tree_params_fupd g ∘ betree_be_tree_root_fupd f) ∧ + ∀h g f. + betree_be_tree_root_fupd f ∘ betree_be_tree_params_fupd g ∘ h = + betree_be_tree_params_fupd g ∘ betree_be_tree_root_fupd f ∘ h) ∧ + (∀g f. + betree_be_tree_root_fupd f ∘ betree_be_tree_node_id_cnt_fupd g = + betree_be_tree_node_id_cnt_fupd g ∘ betree_be_tree_root_fupd f) ∧ + ∀h g f. + betree_be_tree_root_fupd f ∘ betree_be_tree_node_id_cnt_fupd g ∘ + h = + betree_be_tree_node_id_cnt_fupd g ∘ betree_be_tree_root_fupd f ∘ + h + + [betree_be_tree_t_fupdfupds] Theorem + + ⊢ (∀g f b. + b with + <|betree_be_tree_params updated_by f; + betree_be_tree_params updated_by g|> = + b with betree_be_tree_params updated_by f ∘ g) ∧ + (∀g f b. + b with + <|betree_be_tree_node_id_cnt updated_by f; + betree_be_tree_node_id_cnt updated_by g|> = + b with betree_be_tree_node_id_cnt updated_by f ∘ g) ∧ + ∀g f b. + b with + <|betree_be_tree_root updated_by f; + betree_be_tree_root updated_by g|> = + b with betree_be_tree_root updated_by f ∘ g + + [betree_be_tree_t_fupdfupds_comp] Theorem + + ⊢ ((∀g f. + betree_be_tree_params_fupd f ∘ betree_be_tree_params_fupd g = + betree_be_tree_params_fupd (f ∘ g)) ∧ + ∀h g f. + betree_be_tree_params_fupd f ∘ betree_be_tree_params_fupd g ∘ h = + betree_be_tree_params_fupd (f ∘ g) ∘ h) ∧ + ((∀g f. + betree_be_tree_node_id_cnt_fupd f ∘ + betree_be_tree_node_id_cnt_fupd g = + betree_be_tree_node_id_cnt_fupd (f ∘ g)) ∧ + ∀h g f. + betree_be_tree_node_id_cnt_fupd f ∘ + betree_be_tree_node_id_cnt_fupd g ∘ h = + betree_be_tree_node_id_cnt_fupd (f ∘ g) ∘ h) ∧ + (∀g f. + betree_be_tree_root_fupd f ∘ betree_be_tree_root_fupd g = + betree_be_tree_root_fupd (f ∘ g)) ∧ + ∀h g f. + betree_be_tree_root_fupd f ∘ betree_be_tree_root_fupd g ∘ h = + betree_be_tree_root_fupd (f ∘ g) ∘ h + + [betree_be_tree_t_induction] Theorem + + ⊢ ∀P. (∀b b0 b1. P (betree_be_tree_t b b0 b1)) ⇒ ∀b. P b + + [betree_be_tree_t_literal_11] Theorem + + ⊢ ∀b21 b11 b01 b22 b12 b02. + <|betree_be_tree_params := b21; + betree_be_tree_node_id_cnt := b11; betree_be_tree_root := b01|> = + <|betree_be_tree_params := b22; + betree_be_tree_node_id_cnt := b12; betree_be_tree_root := b02|> ⇔ + b21 = b22 ∧ b11 = b12 ∧ b01 = b02 + + [betree_be_tree_t_literal_nchotomy] Theorem + + ⊢ ∀b. ∃b2 b1 b0. + b = + <|betree_be_tree_params := b2; betree_be_tree_node_id_cnt := b1; + betree_be_tree_root := b0|> + + [betree_be_tree_t_nchotomy] Theorem + + ⊢ ∀bb. ∃b b0 b1. bb = betree_be_tree_t b b0 b1 + + [betree_be_tree_t_updates_eq_literal] Theorem + + ⊢ ∀b b2 b1 b0. + b with + <|betree_be_tree_params := b2; betree_be_tree_node_id_cnt := b1; + betree_be_tree_root := b0|> = + <|betree_be_tree_params := b2; betree_be_tree_node_id_cnt := b1; + betree_be_tree_root := b0|> + + [betree_internal_t_11] Theorem + + ⊢ ∀a0 a1 a2 a3 a0' a1' a2' a3'. + betree_internal_t a0 a1 a2 a3 = betree_internal_t a0' a1' a2' a3' ⇔ + a0 = a0' ∧ a1 = a1' ∧ a2 = a2' ∧ a3 = a3' + + [betree_internal_t_Axiom] Theorem + + ⊢ ∀f0 f1 f2. ∃fn0 fn1. + (∀a0 a1 a2 a3. + fn0 (betree_internal_t a0 a1 a2 a3) = + f0 a0 a1 a2 a3 (fn1 a2) (fn1 a3)) ∧ + (∀a. fn1 (BetreeNodeInternal a) = f1 a (fn0 a)) ∧ + ∀a. fn1 (BetreeNodeLeaf a) = f2 a + + [betree_internal_t_accessors] Theorem + + ⊢ (∀u u0 b b0. (betree_internal_t u u0 b b0).betree_internal_id = u) ∧ + (∀u u0 b b0. + (betree_internal_t u u0 b b0).betree_internal_pivot = u0) ∧ + (∀u u0 b b0. (betree_internal_t u u0 b b0).betree_internal_left = b) ∧ + ∀u u0 b b0. + (betree_internal_t u u0 b b0).betree_internal_right = b0 + + [betree_internal_t_accfupds] Theorem + + ⊢ (∀f b. + (b with betree_internal_pivot updated_by f).betree_internal_id = + b.betree_internal_id) ∧ + (∀f b. + (b with betree_internal_left updated_by f).betree_internal_id = + b.betree_internal_id) ∧ + (∀f b. + (b with betree_internal_right updated_by f).betree_internal_id = + b.betree_internal_id) ∧ + (∀f b. + (b with betree_internal_id updated_by f).betree_internal_pivot = + b.betree_internal_pivot) ∧ + (∀f b. + (b with betree_internal_left updated_by f).betree_internal_pivot = + b.betree_internal_pivot) ∧ + (∀f b. + (b with betree_internal_right updated_by f). + betree_internal_pivot = + b.betree_internal_pivot) ∧ + (∀f b. + (b with betree_internal_id updated_by f).betree_internal_left = + b.betree_internal_left) ∧ + (∀f b. + (b with betree_internal_pivot updated_by f).betree_internal_left = + b.betree_internal_left) ∧ + (∀f b. + (b with betree_internal_right updated_by f).betree_internal_left = + b.betree_internal_left) ∧ + (∀f b. + (b with betree_internal_id updated_by f).betree_internal_right = + b.betree_internal_right) ∧ + (∀f b. + (b with betree_internal_pivot updated_by f). + betree_internal_right = + b.betree_internal_right) ∧ + (∀f b. + (b with betree_internal_left updated_by f).betree_internal_right = + b.betree_internal_right) ∧ + (∀f b. + (b with betree_internal_id updated_by f).betree_internal_id = + f b.betree_internal_id) ∧ + (∀f b. + (b with betree_internal_pivot updated_by f). + betree_internal_pivot = + f b.betree_internal_pivot) ∧ + (∀f b. + (b with betree_internal_left updated_by f).betree_internal_left = + f b.betree_internal_left) ∧ + ∀f b. + (b with betree_internal_right updated_by f).betree_internal_right = + f b.betree_internal_right + + [betree_internal_t_case_cong] Theorem + + ⊢ ∀M M' f. + M = M' ∧ + (∀a0 a1 a2 a3. + M' = betree_internal_t a0 a1 a2 a3 ⇒ + f a0 a1 a2 a3 = f' a0 a1 a2 a3) ⇒ + betree_internal_t_CASE M f = betree_internal_t_CASE M' f' + + [betree_internal_t_case_eq] Theorem + + ⊢ betree_internal_t_CASE x f = v ⇔ + ∃u0 u b b0. x = betree_internal_t u0 u b b0 ∧ f u0 u b b0 = v + + [betree_internal_t_component_equality] Theorem + + ⊢ ∀b1 b2. + b1 = b2 ⇔ + b1.betree_internal_id = b2.betree_internal_id ∧ + b1.betree_internal_pivot = b2.betree_internal_pivot ∧ + b1.betree_internal_left = b2.betree_internal_left ∧ + b1.betree_internal_right = b2.betree_internal_right + + [betree_internal_t_fn_updates] Theorem + + ⊢ (∀f u u0 b b0. + betree_internal_t u u0 b b0 with betree_internal_id updated_by f = + betree_internal_t (f u) u0 b b0) ∧ + (∀f u u0 b b0. + betree_internal_t u u0 b b0 with + betree_internal_pivot updated_by f = + betree_internal_t u (f u0) b b0) ∧ + (∀f u u0 b b0. + betree_internal_t u u0 b b0 with + betree_internal_left updated_by f = + betree_internal_t u u0 (f b) b0) ∧ + ∀f u u0 b b0. + betree_internal_t u u0 b b0 with + betree_internal_right updated_by f = + betree_internal_t u u0 b (f b0) + + [betree_internal_t_fupdcanon] Theorem + + ⊢ (∀g f b. + b with + <|betree_internal_pivot updated_by f; + betree_internal_id updated_by g|> = + b with + <|betree_internal_id updated_by g; + betree_internal_pivot updated_by f|>) ∧ + (∀g f b. + b with + <|betree_internal_left updated_by f; + betree_internal_id updated_by g|> = + b with + <|betree_internal_id updated_by g; + betree_internal_left updated_by f|>) ∧ + (∀g f b. + b with + <|betree_internal_left updated_by f; + betree_internal_pivot updated_by g|> = + b with + <|betree_internal_pivot updated_by g; + betree_internal_left updated_by f|>) ∧ + (∀g f b. + b with + <|betree_internal_right updated_by f; + betree_internal_id updated_by g|> = + b with + <|betree_internal_id updated_by g; + betree_internal_right updated_by f|>) ∧ + (∀g f b. + b with + <|betree_internal_right updated_by f; + betree_internal_pivot updated_by g|> = + b with + <|betree_internal_pivot updated_by g; + betree_internal_right updated_by f|>) ∧ + ∀g f b. + b with + <|betree_internal_right updated_by f; + betree_internal_left updated_by g|> = + b with + <|betree_internal_left updated_by g; + betree_internal_right updated_by f|> + + [betree_internal_t_fupdcanon_comp] Theorem + + ⊢ ((∀g f. + betree_internal_pivot_fupd f ∘ betree_internal_id_fupd g = + betree_internal_id_fupd g ∘ betree_internal_pivot_fupd f) ∧ + ∀h g f. + betree_internal_pivot_fupd f ∘ betree_internal_id_fupd g ∘ h = + betree_internal_id_fupd g ∘ betree_internal_pivot_fupd f ∘ h) ∧ + ((∀g f. + betree_internal_left_fupd f ∘ betree_internal_id_fupd g = + betree_internal_id_fupd g ∘ betree_internal_left_fupd f) ∧ + ∀h g f. + betree_internal_left_fupd f ∘ betree_internal_id_fupd g ∘ h = + betree_internal_id_fupd g ∘ betree_internal_left_fupd f ∘ h) ∧ + ((∀g f. + betree_internal_left_fupd f ∘ betree_internal_pivot_fupd g = + betree_internal_pivot_fupd g ∘ betree_internal_left_fupd f) ∧ + ∀h g f. + betree_internal_left_fupd f ∘ betree_internal_pivot_fupd g ∘ h = + betree_internal_pivot_fupd g ∘ betree_internal_left_fupd f ∘ h) ∧ + ((∀g f. + betree_internal_right_fupd f ∘ betree_internal_id_fupd g = + betree_internal_id_fupd g ∘ betree_internal_right_fupd f) ∧ + ∀h g f. + betree_internal_right_fupd f ∘ betree_internal_id_fupd g ∘ h = + betree_internal_id_fupd g ∘ betree_internal_right_fupd f ∘ h) ∧ + ((∀g f. + betree_internal_right_fupd f ∘ betree_internal_pivot_fupd g = + betree_internal_pivot_fupd g ∘ betree_internal_right_fupd f) ∧ + ∀h g f. + betree_internal_right_fupd f ∘ betree_internal_pivot_fupd g ∘ h = + betree_internal_pivot_fupd g ∘ betree_internal_right_fupd f ∘ h) ∧ + (∀g f. + betree_internal_right_fupd f ∘ betree_internal_left_fupd g = + betree_internal_left_fupd g ∘ betree_internal_right_fupd f) ∧ + ∀h g f. + betree_internal_right_fupd f ∘ betree_internal_left_fupd g ∘ h = + betree_internal_left_fupd g ∘ betree_internal_right_fupd f ∘ h + + [betree_internal_t_fupdfupds] Theorem + + ⊢ (∀g f b. + b with + <|betree_internal_id updated_by f; + betree_internal_id updated_by g|> = + b with betree_internal_id updated_by f ∘ g) ∧ + (∀g f b. + b with + <|betree_internal_pivot updated_by f; + betree_internal_pivot updated_by g|> = + b with betree_internal_pivot updated_by f ∘ g) ∧ + (∀g f b. + b with + <|betree_internal_left updated_by f; + betree_internal_left updated_by g|> = + b with betree_internal_left updated_by f ∘ g) ∧ + ∀g f b. + b with + <|betree_internal_right updated_by f; + betree_internal_right updated_by g|> = + b with betree_internal_right updated_by f ∘ g + + [betree_internal_t_fupdfupds_comp] Theorem + + ⊢ ((∀g f. + betree_internal_id_fupd f ∘ betree_internal_id_fupd g = + betree_internal_id_fupd (f ∘ g)) ∧ + ∀h g f. + betree_internal_id_fupd f ∘ betree_internal_id_fupd g ∘ h = + betree_internal_id_fupd (f ∘ g) ∘ h) ∧ + ((∀g f. + betree_internal_pivot_fupd f ∘ betree_internal_pivot_fupd g = + betree_internal_pivot_fupd (f ∘ g)) ∧ + ∀h g f. + betree_internal_pivot_fupd f ∘ betree_internal_pivot_fupd g ∘ h = + betree_internal_pivot_fupd (f ∘ g) ∘ h) ∧ + ((∀g f. + betree_internal_left_fupd f ∘ betree_internal_left_fupd g = + betree_internal_left_fupd (f ∘ g)) ∧ + ∀h g f. + betree_internal_left_fupd f ∘ betree_internal_left_fupd g ∘ h = + betree_internal_left_fupd (f ∘ g) ∘ h) ∧ + (∀g f. + betree_internal_right_fupd f ∘ betree_internal_right_fupd g = + betree_internal_right_fupd (f ∘ g)) ∧ + ∀h g f. + betree_internal_right_fupd f ∘ betree_internal_right_fupd g ∘ h = + betree_internal_right_fupd (f ∘ g) ∘ h + + [betree_internal_t_induction] Theorem + + ⊢ ∀P0 P1. + (∀b b0. P1 b ∧ P1 b0 ⇒ ∀u u0. P0 (betree_internal_t u0 u b b0)) ∧ + (∀b. P0 b ⇒ P1 (BetreeNodeInternal b)) ∧ + (∀b. P1 (BetreeNodeLeaf b)) ⇒ + (∀b. P0 b) ∧ ∀b. P1 b + + [betree_internal_t_literal_11] Theorem + + ⊢ ∀u01 u1 b11 b01 u02 u2 b12 b02. + <|betree_internal_id := u01; betree_internal_pivot := u1; + betree_internal_left := b11; betree_internal_right := b01|> = + <|betree_internal_id := u02; betree_internal_pivot := u2; + betree_internal_left := b12; betree_internal_right := b02|> ⇔ + u01 = u02 ∧ u1 = u2 ∧ b11 = b12 ∧ b01 = b02 + + [betree_internal_t_literal_nchotomy] Theorem + + ⊢ ∀b. ∃u0 u b1 b0. + b = + <|betree_internal_id := u0; betree_internal_pivot := u; + betree_internal_left := b1; betree_internal_right := b0|> + + [betree_internal_t_nchotomy] Theorem + + ⊢ ∀bb. ∃u0 u b b0. bb = betree_internal_t u0 u b b0 + + [betree_internal_t_updates_eq_literal] Theorem + + ⊢ ∀b u0 u b1 b0. + b with + <|betree_internal_id := u0; betree_internal_pivot := u; + betree_internal_left := b1; betree_internal_right := b0|> = + <|betree_internal_id := u0; betree_internal_pivot := u; + betree_internal_left := b1; betree_internal_right := b0|> + + [betree_leaf_t_11] Theorem + + ⊢ ∀a0 a1 a0' a1'. + betree_leaf_t a0 a1 = betree_leaf_t a0' a1' ⇔ a0 = a0' ∧ a1 = a1' + + [betree_leaf_t_Axiom] Theorem + + ⊢ ∀f. ∃fn. ∀a0 a1. fn (betree_leaf_t a0 a1) = f a0 a1 + + [betree_leaf_t_accessors] Theorem + + ⊢ (∀u u0. (betree_leaf_t u u0).betree_leaf_id = u) ∧ + ∀u u0. (betree_leaf_t u u0).betree_leaf_size = u0 + + [betree_leaf_t_accfupds] Theorem + + ⊢ (∀f b. + (b with betree_leaf_size updated_by f).betree_leaf_id = + b.betree_leaf_id) ∧ + (∀f b. + (b with betree_leaf_id updated_by f).betree_leaf_size = + b.betree_leaf_size) ∧ + (∀f b. + (b with betree_leaf_id updated_by f).betree_leaf_id = + f b.betree_leaf_id) ∧ + ∀f b. + (b with betree_leaf_size updated_by f).betree_leaf_size = + f b.betree_leaf_size + + [betree_leaf_t_case_cong] Theorem + + ⊢ ∀M M' f. + M = M' ∧ (∀a0 a1. M' = betree_leaf_t a0 a1 ⇒ f a0 a1 = f' a0 a1) ⇒ + betree_leaf_t_CASE M f = betree_leaf_t_CASE M' f' + + [betree_leaf_t_case_eq] Theorem + + ⊢ betree_leaf_t_CASE x f = v ⇔ + ∃u u0. x = betree_leaf_t u u0 ∧ f u u0 = v + + [betree_leaf_t_component_equality] Theorem + + ⊢ ∀b1 b2. + b1 = b2 ⇔ + b1.betree_leaf_id = b2.betree_leaf_id ∧ + b1.betree_leaf_size = b2.betree_leaf_size + + [betree_leaf_t_fn_updates] Theorem + + ⊢ (∀f u u0. + betree_leaf_t u u0 with betree_leaf_id updated_by f = + betree_leaf_t (f u) u0) ∧ + ∀f u u0. + betree_leaf_t u u0 with betree_leaf_size updated_by f = + betree_leaf_t u (f u0) + + [betree_leaf_t_fupdcanon] Theorem + + ⊢ ∀g f b. + b with + <|betree_leaf_size updated_by f; betree_leaf_id updated_by g|> = + b with + <|betree_leaf_id updated_by g; betree_leaf_size updated_by f|> + + [betree_leaf_t_fupdcanon_comp] Theorem + + ⊢ (∀g f. + betree_leaf_size_fupd f ∘ betree_leaf_id_fupd g = + betree_leaf_id_fupd g ∘ betree_leaf_size_fupd f) ∧ + ∀h g f. + betree_leaf_size_fupd f ∘ betree_leaf_id_fupd g ∘ h = + betree_leaf_id_fupd g ∘ betree_leaf_size_fupd f ∘ h + + [betree_leaf_t_fupdfupds] Theorem + + ⊢ (∀g f b. + b with + <|betree_leaf_id updated_by f; betree_leaf_id updated_by g|> = + b with betree_leaf_id updated_by f ∘ g) ∧ + ∀g f b. + b with + <|betree_leaf_size updated_by f; betree_leaf_size updated_by g|> = + b with betree_leaf_size updated_by f ∘ g + + [betree_leaf_t_fupdfupds_comp] Theorem + + ⊢ ((∀g f. + betree_leaf_id_fupd f ∘ betree_leaf_id_fupd g = + betree_leaf_id_fupd (f ∘ g)) ∧ + ∀h g f. + betree_leaf_id_fupd f ∘ betree_leaf_id_fupd g ∘ h = + betree_leaf_id_fupd (f ∘ g) ∘ h) ∧ + (∀g f. + betree_leaf_size_fupd f ∘ betree_leaf_size_fupd g = + betree_leaf_size_fupd (f ∘ g)) ∧ + ∀h g f. + betree_leaf_size_fupd f ∘ betree_leaf_size_fupd g ∘ h = + betree_leaf_size_fupd (f ∘ g) ∘ h + + [betree_leaf_t_induction] Theorem + + ⊢ ∀P. (∀u u0. P (betree_leaf_t u u0)) ⇒ ∀b. P b + + [betree_leaf_t_literal_11] Theorem + + ⊢ ∀u01 u1 u02 u2. + <|betree_leaf_id := u01; betree_leaf_size := u1|> = + <|betree_leaf_id := u02; betree_leaf_size := u2|> ⇔ + u01 = u02 ∧ u1 = u2 + + [betree_leaf_t_literal_nchotomy] Theorem + + ⊢ ∀b. ∃u0 u. b = <|betree_leaf_id := u0; betree_leaf_size := u|> + + [betree_leaf_t_nchotomy] Theorem + + ⊢ ∀bb. ∃u u0. bb = betree_leaf_t u u0 + + [betree_leaf_t_updates_eq_literal] Theorem + + ⊢ ∀b u0 u. + b with <|betree_leaf_id := u0; betree_leaf_size := u|> = + <|betree_leaf_id := u0; betree_leaf_size := u|> + + [betree_list_t_11] Theorem + + ⊢ ∀a0 a1 a0' a1'. + BetreeListCons a0 a1 = BetreeListCons a0' a1' ⇔ + a0 = a0' ∧ a1 = a1' + + [betree_list_t_Axiom] Theorem + + ⊢ ∀f0 f1. ∃fn. + (∀a0 a1. fn (BetreeListCons a0 a1) = f0 a0 a1 (fn a1)) ∧ + fn BetreeListNil = f1 + + [betree_list_t_case_cong] Theorem + + ⊢ ∀M M' f v. + M = M' ∧ + (∀a0 a1. M' = BetreeListCons a0 a1 ⇒ f a0 a1 = f' a0 a1) ∧ + (M' = BetreeListNil ⇒ v = v') ⇒ + betree_list_t_CASE M f v = betree_list_t_CASE M' f' v' + + [betree_list_t_case_eq] Theorem + + ⊢ betree_list_t_CASE x f v = v' ⇔ + (∃t b. x = BetreeListCons t b ∧ f t b = v') ∨ + x = BetreeListNil ∧ v = v' + + [betree_list_t_distinct] Theorem + + ⊢ ∀a1 a0. BetreeListCons a0 a1 ≠ BetreeListNil + + [betree_list_t_induction] Theorem + + ⊢ ∀P. (∀b. P b ⇒ ∀t. P (BetreeListCons t b)) ∧ P BetreeListNil ⇒ + ∀b. P b + + [betree_list_t_nchotomy] Theorem + + ⊢ ∀bb. (∃t b. bb = BetreeListCons t b) ∨ bb = BetreeListNil + + [betree_message_t_11] Theorem + + ⊢ (∀a a'. BetreeMessageInsert a = BetreeMessageInsert a' ⇔ a = a') ∧ + ∀a a'. BetreeMessageUpsert a = BetreeMessageUpsert a' ⇔ a = a' + + [betree_message_t_Axiom] Theorem + + ⊢ ∀f0 f1 f2. ∃fn. + (∀a. fn (BetreeMessageInsert a) = f0 a) ∧ + fn BetreeMessageDelete = f1 ∧ + ∀a. fn (BetreeMessageUpsert a) = f2 a + + [betree_message_t_case_cong] Theorem + + ⊢ ∀M M' f v f1. + M = M' ∧ (∀a. M' = BetreeMessageInsert a ⇒ f a = f' a) ∧ + (M' = BetreeMessageDelete ⇒ v = v') ∧ + (∀a. M' = BetreeMessageUpsert a ⇒ f1 a = f1' a) ⇒ + betree_message_t_CASE M f v f1 = + betree_message_t_CASE M' f' v' f1' + + [betree_message_t_case_eq] Theorem + + ⊢ betree_message_t_CASE x f v f1 = v' ⇔ + (∃u. x = BetreeMessageInsert u ∧ f u = v') ∨ + x = BetreeMessageDelete ∧ v = v' ∨ + ∃b. x = BetreeMessageUpsert b ∧ f1 b = v' + + [betree_message_t_distinct] Theorem + + ⊢ (∀a. BetreeMessageInsert a ≠ BetreeMessageDelete) ∧ + (∀a' a. BetreeMessageInsert a ≠ BetreeMessageUpsert a') ∧ + ∀a. BetreeMessageDelete ≠ BetreeMessageUpsert a + + [betree_message_t_induction] Theorem + + ⊢ ∀P. (∀u. P (BetreeMessageInsert u)) ∧ P BetreeMessageDelete ∧ + (∀b. P (BetreeMessageUpsert b)) ⇒ + ∀b. P b + + [betree_message_t_nchotomy] Theorem + + ⊢ ∀bb. + (∃u. bb = BetreeMessageInsert u) ∨ bb = BetreeMessageDelete ∨ + ∃b. bb = BetreeMessageUpsert b + + [betree_node_id_counter_t_11] Theorem + + ⊢ ∀a a'. + betree_node_id_counter_t a = betree_node_id_counter_t a' ⇔ a = a' + + [betree_node_id_counter_t_Axiom] Theorem + + ⊢ ∀f. ∃fn. ∀a. fn (betree_node_id_counter_t a) = f a + + [betree_node_id_counter_t_accessors] Theorem + + ⊢ ∀u. (betree_node_id_counter_t u). + betree_node_id_counter_next_node_id = + u + + [betree_node_id_counter_t_accfupds] Theorem + + ⊢ ∀f b. + (b with betree_node_id_counter_next_node_id updated_by f). + betree_node_id_counter_next_node_id = + f b.betree_node_id_counter_next_node_id + + [betree_node_id_counter_t_case_cong] Theorem + + ⊢ ∀M M' f. + M = M' ∧ (∀a. M' = betree_node_id_counter_t a ⇒ f a = f' a) ⇒ + betree_node_id_counter_t_CASE M f = + betree_node_id_counter_t_CASE M' f' + + [betree_node_id_counter_t_case_eq] Theorem + + ⊢ betree_node_id_counter_t_CASE x f = v ⇔ + ∃u. x = betree_node_id_counter_t u ∧ f u = v + + [betree_node_id_counter_t_component_equality] Theorem + + ⊢ ∀b1 b2. + b1 = b2 ⇔ + b1.betree_node_id_counter_next_node_id = + b2.betree_node_id_counter_next_node_id + + [betree_node_id_counter_t_fn_updates] Theorem + + ⊢ ∀f u. + betree_node_id_counter_t u with + betree_node_id_counter_next_node_id updated_by f = + betree_node_id_counter_t (f u) + + [betree_node_id_counter_t_fupdfupds] Theorem + + ⊢ ∀g f b. + b with + <|betree_node_id_counter_next_node_id updated_by f; + betree_node_id_counter_next_node_id updated_by g|> = + b with betree_node_id_counter_next_node_id updated_by f ∘ g + + [betree_node_id_counter_t_fupdfupds_comp] Theorem + + ⊢ (∀g f. + betree_node_id_counter_next_node_id_fupd f ∘ + betree_node_id_counter_next_node_id_fupd g = + betree_node_id_counter_next_node_id_fupd (f ∘ g)) ∧ + ∀h g f. + betree_node_id_counter_next_node_id_fupd f ∘ + betree_node_id_counter_next_node_id_fupd g ∘ h = + betree_node_id_counter_next_node_id_fupd (f ∘ g) ∘ h + + [betree_node_id_counter_t_induction] Theorem + + ⊢ ∀P. (∀u. P (betree_node_id_counter_t u)) ⇒ ∀b. P b + + [betree_node_id_counter_t_literal_11] Theorem + + ⊢ ∀u1 u2. + <|betree_node_id_counter_next_node_id := u1|> = + <|betree_node_id_counter_next_node_id := u2|> ⇔ u1 = u2 + + [betree_node_id_counter_t_literal_nchotomy] Theorem + + ⊢ ∀b. ∃u. b = <|betree_node_id_counter_next_node_id := u|> + + [betree_node_id_counter_t_nchotomy] Theorem + + ⊢ ∀bb. ∃u. bb = betree_node_id_counter_t u + + [betree_node_id_counter_t_updates_eq_literal] Theorem + + ⊢ ∀b u. + b with betree_node_id_counter_next_node_id := u = + <|betree_node_id_counter_next_node_id := u|> + + [betree_node_t_11] Theorem + + ⊢ (∀a a'. BetreeNodeInternal a = BetreeNodeInternal a' ⇔ a = a') ∧ + ∀a a'. BetreeNodeLeaf a = BetreeNodeLeaf a' ⇔ a = a' + + [betree_node_t_Axiom] Theorem + + ⊢ ∀f0 f1 f2. ∃fn0 fn1. + (∀a0 a1 a2 a3. + fn0 (betree_internal_t a0 a1 a2 a3) = + f0 a0 a1 a2 a3 (fn1 a2) (fn1 a3)) ∧ + (∀a. fn1 (BetreeNodeInternal a) = f1 a (fn0 a)) ∧ + ∀a. fn1 (BetreeNodeLeaf a) = f2 a + + [betree_node_t_case_cong] Theorem + + ⊢ ∀M M' f f1. + M = M' ∧ (∀a. M' = BetreeNodeInternal a ⇒ f a = f' a) ∧ + (∀a. M' = BetreeNodeLeaf a ⇒ f1 a = f1' a) ⇒ + betree_node_t_CASE M f f1 = betree_node_t_CASE M' f' f1' + + [betree_node_t_case_eq] Theorem + + ⊢ betree_node_t_CASE x f f1 = v ⇔ + (∃b. x = BetreeNodeInternal b ∧ f b = v) ∨ + ∃b. x = BetreeNodeLeaf b ∧ f1 b = v + + [betree_node_t_distinct] Theorem + + ⊢ ∀a' a. BetreeNodeInternal a ≠ BetreeNodeLeaf a' + + [betree_node_t_induction] Theorem + + ⊢ ∀P0 P1. + (∀b b0. P1 b ∧ P1 b0 ⇒ ∀u u0. P0 (betree_internal_t u0 u b b0)) ∧ + (∀b. P0 b ⇒ P1 (BetreeNodeInternal b)) ∧ + (∀b. P1 (BetreeNodeLeaf b)) ⇒ + (∀b. P0 b) ∧ ∀b. P1 b + + [betree_node_t_nchotomy] Theorem + + ⊢ ∀bb. (∃b. bb = BetreeNodeInternal b) ∨ ∃b. bb = BetreeNodeLeaf b + + [betree_params_t_11] Theorem + + ⊢ ∀a0 a1 a0' a1'. + betree_params_t a0 a1 = betree_params_t a0' a1' ⇔ + a0 = a0' ∧ a1 = a1' + + [betree_params_t_Axiom] Theorem + + ⊢ ∀f. ∃fn. ∀a0 a1. fn (betree_params_t a0 a1) = f a0 a1 + + [betree_params_t_accessors] Theorem + + ⊢ (∀u u0. (betree_params_t u u0).betree_params_min_flush_size = u) ∧ + ∀u u0. (betree_params_t u u0).betree_params_split_size = u0 + + [betree_params_t_accfupds] Theorem + + ⊢ (∀f b. + (b with betree_params_split_size updated_by f). + betree_params_min_flush_size = + b.betree_params_min_flush_size) ∧ + (∀f b. + (b with betree_params_min_flush_size updated_by f). + betree_params_split_size = + b.betree_params_split_size) ∧ + (∀f b. + (b with betree_params_min_flush_size updated_by f). + betree_params_min_flush_size = + f b.betree_params_min_flush_size) ∧ + ∀f b. + (b with betree_params_split_size updated_by f). + betree_params_split_size = + f b.betree_params_split_size + + [betree_params_t_case_cong] Theorem + + ⊢ ∀M M' f. + M = M' ∧ + (∀a0 a1. M' = betree_params_t a0 a1 ⇒ f a0 a1 = f' a0 a1) ⇒ + betree_params_t_CASE M f = betree_params_t_CASE M' f' + + [betree_params_t_case_eq] Theorem + + ⊢ betree_params_t_CASE x f = v ⇔ + ∃u u0. x = betree_params_t u u0 ∧ f u u0 = v + + [betree_params_t_component_equality] Theorem + + ⊢ ∀b1 b2. + b1 = b2 ⇔ + b1.betree_params_min_flush_size = b2.betree_params_min_flush_size ∧ + b1.betree_params_split_size = b2.betree_params_split_size + + [betree_params_t_fn_updates] Theorem + + ⊢ (∀f u u0. + betree_params_t u u0 with + betree_params_min_flush_size updated_by f = + betree_params_t (f u) u0) ∧ + ∀f u u0. + betree_params_t u u0 with betree_params_split_size updated_by f = + betree_params_t u (f u0) + + [betree_params_t_fupdcanon] Theorem + + ⊢ ∀g f b. + b with + <|betree_params_split_size updated_by f; + betree_params_min_flush_size updated_by g|> = + b with + <|betree_params_min_flush_size updated_by g; + betree_params_split_size updated_by f|> + + [betree_params_t_fupdcanon_comp] Theorem + + ⊢ (∀g f. + betree_params_split_size_fupd f ∘ + betree_params_min_flush_size_fupd g = + betree_params_min_flush_size_fupd g ∘ + betree_params_split_size_fupd f) ∧ + ∀h g f. + betree_params_split_size_fupd f ∘ + betree_params_min_flush_size_fupd g ∘ h = + betree_params_min_flush_size_fupd g ∘ + betree_params_split_size_fupd f ∘ h + + [betree_params_t_fupdfupds] Theorem + + ⊢ (∀g f b. + b with + <|betree_params_min_flush_size updated_by f; + betree_params_min_flush_size updated_by g|> = + b with betree_params_min_flush_size updated_by f ∘ g) ∧ + ∀g f b. + b with + <|betree_params_split_size updated_by f; + betree_params_split_size updated_by g|> = + b with betree_params_split_size updated_by f ∘ g + + [betree_params_t_fupdfupds_comp] Theorem + + ⊢ ((∀g f. + betree_params_min_flush_size_fupd f ∘ + betree_params_min_flush_size_fupd g = + betree_params_min_flush_size_fupd (f ∘ g)) ∧ + ∀h g f. + betree_params_min_flush_size_fupd f ∘ + betree_params_min_flush_size_fupd g ∘ h = + betree_params_min_flush_size_fupd (f ∘ g) ∘ h) ∧ + (∀g f. + betree_params_split_size_fupd f ∘ + betree_params_split_size_fupd g = + betree_params_split_size_fupd (f ∘ g)) ∧ + ∀h g f. + betree_params_split_size_fupd f ∘ + betree_params_split_size_fupd g ∘ h = + betree_params_split_size_fupd (f ∘ g) ∘ h + + [betree_params_t_induction] Theorem + + ⊢ ∀P. (∀u u0. P (betree_params_t u u0)) ⇒ ∀b. P b + + [betree_params_t_literal_11] Theorem + + ⊢ ∀u01 u1 u02 u2. + <|betree_params_min_flush_size := u01; + betree_params_split_size := u1|> = + <|betree_params_min_flush_size := u02; + betree_params_split_size := u2|> ⇔ u01 = u02 ∧ u1 = u2 + + [betree_params_t_literal_nchotomy] Theorem + + ⊢ ∀b. ∃u0 u. + b = + <|betree_params_min_flush_size := u0; + betree_params_split_size := u|> + + [betree_params_t_nchotomy] Theorem + + ⊢ ∀bb. ∃u u0. bb = betree_params_t u u0 + + [betree_params_t_updates_eq_literal] Theorem + + ⊢ ∀b u0 u. + b with + <|betree_params_min_flush_size := u0; + betree_params_split_size := u|> = + <|betree_params_min_flush_size := u0; + betree_params_split_size := u|> + + [betree_upsert_fun_state_t_11] Theorem + + ⊢ (∀a a'. + BetreeUpsertFunStateAdd a = BetreeUpsertFunStateAdd a' ⇔ a = a') ∧ + ∀a a'. + BetreeUpsertFunStateSub a = BetreeUpsertFunStateSub a' ⇔ a = a' + + [betree_upsert_fun_state_t_Axiom] Theorem + + ⊢ ∀f0 f1. ∃fn. + (∀a. fn (BetreeUpsertFunStateAdd a) = f0 a) ∧ + ∀a. fn (BetreeUpsertFunStateSub a) = f1 a + + [betree_upsert_fun_state_t_case_cong] Theorem + + ⊢ ∀M M' f f1. + M = M' ∧ (∀a. M' = BetreeUpsertFunStateAdd a ⇒ f a = f' a) ∧ + (∀a. M' = BetreeUpsertFunStateSub a ⇒ f1 a = f1' a) ⇒ + betree_upsert_fun_state_t_CASE M f f1 = + betree_upsert_fun_state_t_CASE M' f' f1' + + [betree_upsert_fun_state_t_case_eq] Theorem + + ⊢ betree_upsert_fun_state_t_CASE x f f1 = v ⇔ + (∃u. x = BetreeUpsertFunStateAdd u ∧ f u = v) ∨ + ∃u. x = BetreeUpsertFunStateSub u ∧ f1 u = v + + [betree_upsert_fun_state_t_distinct] Theorem + + ⊢ ∀a' a. BetreeUpsertFunStateAdd a ≠ BetreeUpsertFunStateSub a' + + [betree_upsert_fun_state_t_induction] Theorem + + ⊢ ∀P. (∀u. P (BetreeUpsertFunStateAdd u)) ∧ + (∀u. P (BetreeUpsertFunStateSub u)) ⇒ + ∀b. P b + + [betree_upsert_fun_state_t_nchotomy] Theorem + + ⊢ ∀bb. + (∃u. bb = BetreeUpsertFunStateAdd u) ∨ + ∃u. bb = BetreeUpsertFunStateSub u + + [datatype_betree_be_tree_t] Theorem + + ⊢ DATATYPE + (record betree_be_tree_t betree_be_tree_params + betree_be_tree_node_id_cnt betree_be_tree_root) + + [datatype_betree_internal_t] Theorem + + ⊢ DATATYPE + (record betree_internal_t betree_internal_id + betree_internal_pivot betree_internal_left + betree_internal_right ∧ + betree_node_t BetreeNodeInternal BetreeNodeLeaf) + + [datatype_betree_leaf_t] Theorem + + ⊢ DATATYPE (record betree_leaf_t betree_leaf_id betree_leaf_size) + + [datatype_betree_list_t] Theorem + + ⊢ DATATYPE (betree_list_t BetreeListCons BetreeListNil) + + [datatype_betree_message_t] Theorem + + ⊢ DATATYPE + (betree_message_t BetreeMessageInsert BetreeMessageDelete + BetreeMessageUpsert) + + [datatype_betree_node_id_counter_t] Theorem + + ⊢ DATATYPE + (record betree_node_id_counter_t + betree_node_id_counter_next_node_id) + + [datatype_betree_params_t] Theorem + + ⊢ DATATYPE + (record betree_params_t betree_params_min_flush_size + betree_params_split_size) + + [datatype_betree_upsert_fun_state_t] Theorem + + ⊢ DATATYPE + (betree_upsert_fun_state_t BetreeUpsertFunStateAdd + BetreeUpsertFunStateSub) + + +*) +end diff --git a/tests/lean/Betree/Funs.lean b/tests/lean/Betree/Funs.lean new file mode 100644 index 00000000..8612ccbc --- /dev/null +++ b/tests/lean/Betree/Funs.lean @@ -0,0 +1,699 @@ +-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS +-- [betree]: function definitions +import Base +import Betree.Types +import Betree.FunsExternal +open Primitives + +namespace betree + +/- [betree::betree::load_internal_node]: + Source: 'src/betree.rs', lines 36:0-36:52 -/ +def betree.load_internal_node + (id : U64) (st : State) : + Result (State × (betree.List (U64 × betree.Message))) + := + betree_utils.load_internal_node id st + +/- [betree::betree::store_internal_node]: + Source: 'src/betree.rs', lines 41:0-41:60 -/ +def betree.store_internal_node + (id : U64) (content : betree.List (U64 × betree.Message)) (st : State) : + Result (State × Unit) + := + betree_utils.store_internal_node id content st + +/- [betree::betree::load_leaf_node]: + Source: 'src/betree.rs', lines 46:0-46:44 -/ +def betree.load_leaf_node + (id : U64) (st : State) : Result (State × (betree.List (U64 × U64))) := + betree_utils.load_leaf_node id st + +/- [betree::betree::store_leaf_node]: + Source: 'src/betree.rs', lines 51:0-51:52 -/ +def betree.store_leaf_node + (id : U64) (content : betree.List (U64 × U64)) (st : State) : + Result (State × Unit) + := + betree_utils.store_leaf_node id content st + +/- [betree::betree::fresh_node_id]: + Source: 'src/betree.rs', lines 55:0-55:48 -/ +def betree.fresh_node_id (counter : U64) : Result (U64 × U64) := + do + let counter1 ← counter + 1#u64 + Result.ok (counter, counter1) + +/- [betree::betree::{betree::betree::NodeIdCounter}::new]: + Source: 'src/betree.rs', lines 206:4-206:20 -/ +def betree.NodeIdCounter.new : Result betree.NodeIdCounter := + Result.ok { next_node_id := 0#u64 } + +/- [betree::betree::{betree::betree::NodeIdCounter}::fresh_id]: + Source: 'src/betree.rs', lines 210:4-210:36 -/ +def betree.NodeIdCounter.fresh_id + (self : betree.NodeIdCounter) : Result (U64 × betree.NodeIdCounter) := + do + let i ← self.next_node_id + 1#u64 + Result.ok (self.next_node_id, { next_node_id := i }) + +/- [betree::betree::upsert_update]: + Source: 'src/betree.rs', lines 234:0-234:70 -/ +def betree.upsert_update + (prev : Option U64) (st : betree.UpsertFunState) : Result U64 := + match prev with + | none => + match st with + | betree.UpsertFunState.Add v => Result.ok v + | betree.UpsertFunState.Sub _ => Result.ok 0#u64 + | some prev1 => + match st with + | betree.UpsertFunState.Add v => + do + let margin ← core_u64_max - prev1 + if margin >= v + then prev1 + v + else Result.ok core_u64_max + | betree.UpsertFunState.Sub v => + if prev1 >= v + then prev1 - v + else Result.ok 0#u64 + +/- [betree::betree::{betree::betree::List#1}::len]: + Source: 'src/betree.rs', lines 276:4-276:24 -/ +divergent def betree.List.len (T : Type) (self : betree.List T) : Result U64 := + match self with + | betree.List.Cons _ tl => do + let i ← betree.List.len T tl + 1#u64 + i + | betree.List.Nil => Result.ok 0#u64 + +/- [betree::betree::{betree::betree::List#1}::split_at]: + Source: 'src/betree.rs', lines 284:4-284:51 -/ +divergent def betree.List.split_at + (T : Type) (self : betree.List T) (n : U64) : + Result ((betree.List T) × (betree.List T)) + := + if n = 0#u64 + then Result.ok (betree.List.Nil, self) + else + match self with + | betree.List.Cons hd tl => + do + let i ← n - 1#u64 + let p ← betree.List.split_at T tl i + let (ls0, ls1) := p + Result.ok (betree.List.Cons hd ls0, ls1) + | betree.List.Nil => Result.fail .panic + +/- [betree::betree::{betree::betree::List#1}::push_front]: + Source: 'src/betree.rs', lines 299:4-299:34 -/ +def betree.List.push_front + (T : Type) (self : betree.List T) (x : T) : Result (betree.List T) := + let (tl, _) := core.mem.replace (betree.List T) self betree.List.Nil + Result.ok (betree.List.Cons x tl) + +/- [betree::betree::{betree::betree::List#1}::pop_front]: + Source: 'src/betree.rs', lines 306:4-306:32 -/ +def betree.List.pop_front + (T : Type) (self : betree.List T) : Result (T × (betree.List T)) := + let (ls, _) := core.mem.replace (betree.List T) self betree.List.Nil + match ls with + | betree.List.Cons x tl => Result.ok (x, tl) + | betree.List.Nil => Result.fail .panic + +/- [betree::betree::{betree::betree::List#1}::hd]: + Source: 'src/betree.rs', lines 318:4-318:22 -/ +def betree.List.hd (T : Type) (self : betree.List T) : Result T := + match self with + | betree.List.Cons hd _ => Result.ok hd + | betree.List.Nil => Result.fail .panic + +/- [betree::betree::{betree::betree::List<(u64, T)>#2}::head_has_key]: + Source: 'src/betree.rs', lines 327:4-327:44 -/ +def betree.ListPairU64T.head_has_key + (T : Type) (self : betree.List (U64 × T)) (key : U64) : Result Bool := + match self with + | betree.List.Cons hd _ => let (i, _) := hd + Result.ok (i = key) + | betree.List.Nil => Result.ok false + +/- [betree::betree::{betree::betree::List<(u64, T)>#2}::partition_at_pivot]: + Source: 'src/betree.rs', lines 339:4-339:73 -/ +divergent def betree.ListPairU64T.partition_at_pivot + (T : Type) (self : betree.List (U64 × T)) (pivot : U64) : + Result ((betree.List (U64 × T)) × (betree.List (U64 × T))) + := + match self with + | betree.List.Cons hd tl => + let (i, t) := hd + if i >= pivot + then Result.ok (betree.List.Nil, betree.List.Cons (i, t) tl) + else + do + let p ← betree.ListPairU64T.partition_at_pivot T tl pivot + let (ls0, ls1) := p + Result.ok (betree.List.Cons (i, t) ls0, ls1) + | betree.List.Nil => Result.ok (betree.List.Nil, betree.List.Nil) + +/- [betree::betree::{betree::betree::Leaf#3}::split]: + Source: 'src/betree.rs', lines 359:4-364:17 -/ +def betree.Leaf.split + (self : betree.Leaf) (content : betree.List (U64 × U64)) + (params : betree.Params) (node_id_cnt : betree.NodeIdCounter) (st : State) : + Result (State × (betree.Internal × betree.NodeIdCounter)) + := + do + let p ← betree.List.split_at (U64 × U64) content params.split_size + let (content0, content1) := p + let p1 ← betree.List.hd (U64 × U64) content1 + let (pivot, _) := p1 + let (id0, node_id_cnt1) ← betree.NodeIdCounter.fresh_id node_id_cnt + let (id1, node_id_cnt2) ← betree.NodeIdCounter.fresh_id node_id_cnt1 + let (st1, _) ← betree.store_leaf_node id0 content0 st + let (st2, _) ← betree.store_leaf_node id1 content1 st1 + let n := betree.Node.Leaf { id := id0, size := params.split_size } + let n1 := betree.Node.Leaf { id := id1, size := params.split_size } + Result.ok (st2, (betree.Internal.mk self.id pivot n n1, node_id_cnt2)) + +/- [betree::betree::{betree::betree::Node#5}::lookup_first_message_for_key]: + Source: 'src/betree.rs', lines 789:4-792:34 -/ +divergent 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 × + betree.Message) → Result (betree.List (U64 × betree.Message)))) + := + match msgs with + | betree.List.Cons x next_msgs => + let (i, m) := x + if i >= key + then Result.ok (betree.List.Cons (i, m) next_msgs, Result.ok) + else + do + let (l, lookup_first_message_for_key_back) ← + betree.Node.lookup_first_message_for_key key next_msgs + let back := + fun ret => + do + let next_msgs1 ← lookup_first_message_for_key_back ret + Result.ok (betree.List.Cons (i, m) next_msgs1) + Result.ok (l, back) + | betree.List.Nil => Result.ok (betree.List.Nil, Result.ok) + +/- [betree::betree::{betree::betree::Node#5}::lookup_in_bindings]: + Source: 'src/betree.rs', lines 636:4-636:80 -/ +divergent def betree.Node.lookup_in_bindings + (key : U64) (bindings : betree.List (U64 × U64)) : Result (Option U64) := + match bindings with + | betree.List.Cons hd tl => + let (i, i1) := hd + if i = key + then Result.ok (some i1) + else + if i > key + then Result.ok none + else betree.Node.lookup_in_bindings key tl + | betree.List.Nil => Result.ok none + +/- [betree::betree::{betree::betree::Node#5}::apply_upserts]: + Source: 'src/betree.rs', lines 819:4-819:90 -/ +divergent def betree.Node.apply_upserts + (msgs : betree.List (U64 × betree.Message)) (prev : Option U64) (key : U64) + : + Result (U64 × (betree.List (U64 × betree.Message))) + := + do + let b ← betree.ListPairU64T.head_has_key betree.Message msgs key + if b + then + do + let (msg, msgs1) ← betree.List.pop_front (U64 × betree.Message) msgs + let (_, m) := msg + match m with + | betree.Message.Insert _ => Result.fail .panic + | betree.Message.Delete => Result.fail .panic + | betree.Message.Upsert s => + do + let v ← betree.upsert_update prev s + betree.Node.apply_upserts msgs1 (some v) key + else + do + let v ← core.option.Option.unwrap U64 prev + let msgs1 ← + betree.List.push_front (U64 × betree.Message) msgs (key, + betree.Message.Insert v) + Result.ok (v, msgs1) + +/- [betree::betree::{betree::betree::Internal#4}::lookup_in_children]: + Source: 'src/betree.rs', lines 395:4-395:63 -/ +mutual divergent def betree.Internal.lookup_in_children + (self : betree.Internal) (key : U64) (st : State) : + Result (State × ((Option U64) × betree.Internal)) + := + if key < self.pivot + then + do + let (st1, (o, n)) ← betree.Node.lookup self.left key st + Result.ok (st1, (o, betree.Internal.mk self.id self.pivot n self.right)) + else + do + let (st1, (o, n)) ← betree.Node.lookup self.right key st + Result.ok (st1, (o, betree.Internal.mk self.id self.pivot self.left n)) + +/- [betree::betree::{betree::betree::Node#5}::lookup]: + Source: 'src/betree.rs', lines 709:4-709:58 -/ +divergent def betree.Node.lookup + (self : betree.Node) (key : U64) (st : State) : + Result (State × ((Option U64) × betree.Node)) + := + match self with + | betree.Node.Internal node => + do + let (st1, msgs) ← betree.load_internal_node node.id st + let (pending, lookup_first_message_for_key_back) ← + betree.Node.lookup_first_message_for_key key msgs + match pending with + | betree.List.Cons p l => + let (k, msg) := p + if k != key + then + do + let (st2, (o, node1)) ← + betree.Internal.lookup_in_children node key st1 + let _ ← + lookup_first_message_for_key_back (betree.List.Cons (k, msg) l) + Result.ok (st2, (o, betree.Node.Internal node1)) + else + match msg with + | betree.Message.Insert v => + do + let _ ← + lookup_first_message_for_key_back (betree.List.Cons (k, + betree.Message.Insert v) l) + Result.ok (st1, (some v, betree.Node.Internal node)) + | betree.Message.Delete => + do + let _ ← + lookup_first_message_for_key_back (betree.List.Cons (k, + betree.Message.Delete) l) + Result.ok (st1, (none, betree.Node.Internal node)) + | betree.Message.Upsert ufs => + do + let (st2, (v, node1)) ← + betree.Internal.lookup_in_children node key st1 + let (v1, pending1) ← + betree.Node.apply_upserts (betree.List.Cons (k, + betree.Message.Upsert ufs) l) v key + let msgs1 ← lookup_first_message_for_key_back pending1 + let (st3, _) ← betree.store_internal_node node1.id msgs1 st2 + Result.ok (st3, (some v1, betree.Node.Internal node1)) + | betree.List.Nil => + do + let (st2, (o, node1)) ← betree.Internal.lookup_in_children node key st1 + let _ ← lookup_first_message_for_key_back betree.List.Nil + Result.ok (st2, (o, betree.Node.Internal node1)) + | betree.Node.Leaf node => + do + let (st1, bindings) ← betree.load_leaf_node node.id st + let o ← betree.Node.lookup_in_bindings key bindings + Result.ok (st1, (o, betree.Node.Leaf node)) + +end + +/- [betree::betree::{betree::betree::Node#5}::filter_messages_for_key]: + Source: 'src/betree.rs', lines 674:4-674:77 -/ +divergent def betree.Node.filter_messages_for_key + (key : U64) (msgs : betree.List (U64 × betree.Message)) : + Result (betree.List (U64 × betree.Message)) + := + match msgs with + | betree.List.Cons p l => + let (k, m) := p + if k = key + then + do + let (_, msgs1) ← + betree.List.pop_front (U64 × betree.Message) (betree.List.Cons (k, m) + l) + betree.Node.filter_messages_for_key key msgs1 + else Result.ok (betree.List.Cons (k, m) l) + | betree.List.Nil => Result.ok betree.List.Nil + +/- [betree::betree::{betree::betree::Node#5}::lookup_first_message_after_key]: + Source: 'src/betree.rs', lines 689:4-692:34 -/ +divergent 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 × + betree.Message) → Result (betree.List (U64 × betree.Message)))) + := + match msgs with + | betree.List.Cons p next_msgs => + let (k, m) := p + if k = key + then + do + let (l, lookup_first_message_after_key_back) ← + betree.Node.lookup_first_message_after_key key next_msgs + let back := + fun ret => + do + let next_msgs1 ← lookup_first_message_after_key_back ret + Result.ok (betree.List.Cons (k, m) next_msgs1) + Result.ok (l, back) + else Result.ok (betree.List.Cons (k, m) next_msgs, Result.ok) + | betree.List.Nil => Result.ok (betree.List.Nil, Result.ok) + +/- [betree::betree::{betree::betree::Node#5}::apply_to_internal]: + Source: 'src/betree.rs', lines 521:4-521:89 -/ +def betree.Node.apply_to_internal + (msgs : betree.List (U64 × betree.Message)) (key : U64) + (new_msg : betree.Message) : + Result (betree.List (U64 × betree.Message)) + := + do + let (msgs1, lookup_first_message_for_key_back) ← + betree.Node.lookup_first_message_for_key key msgs + let b ← betree.ListPairU64T.head_has_key betree.Message msgs1 key + if b + then + match new_msg with + | betree.Message.Insert i => + do + let msgs2 ← betree.Node.filter_messages_for_key key msgs1 + let msgs3 ← + betree.List.push_front (U64 × betree.Message) msgs2 (key, + betree.Message.Insert i) + lookup_first_message_for_key_back msgs3 + | betree.Message.Delete => + do + let msgs2 ← betree.Node.filter_messages_for_key key msgs1 + let msgs3 ← + betree.List.push_front (U64 × betree.Message) msgs2 (key, + betree.Message.Delete) + lookup_first_message_for_key_back msgs3 + | betree.Message.Upsert s => + do + let p ← betree.List.hd (U64 × betree.Message) msgs1 + let (_, m) := p + match m with + | betree.Message.Insert prev => + do + let v ← betree.upsert_update (some prev) s + let (_, msgs2) ← betree.List.pop_front (U64 × betree.Message) msgs1 + let msgs3 ← + betree.List.push_front (U64 × betree.Message) msgs2 (key, + betree.Message.Insert v) + lookup_first_message_for_key_back msgs3 + | betree.Message.Delete => + do + let (_, msgs2) ← betree.List.pop_front (U64 × betree.Message) msgs1 + let v ← betree.upsert_update none s + let msgs3 ← + betree.List.push_front (U64 × betree.Message) msgs2 (key, + betree.Message.Insert v) + lookup_first_message_for_key_back msgs3 + | betree.Message.Upsert _ => + do + let (msgs2, lookup_first_message_after_key_back) ← + betree.Node.lookup_first_message_after_key key msgs1 + let msgs3 ← + betree.List.push_front (U64 × betree.Message) msgs2 (key, + betree.Message.Upsert s) + let msgs4 ← lookup_first_message_after_key_back msgs3 + lookup_first_message_for_key_back msgs4 + else + do + let msgs2 ← + betree.List.push_front (U64 × betree.Message) msgs1 (key, new_msg) + lookup_first_message_for_key_back msgs2 + +/- [betree::betree::{betree::betree::Node#5}::apply_messages_to_internal]: + Source: 'src/betree.rs', lines 502:4-505:5 -/ +divergent def betree.Node.apply_messages_to_internal + (msgs : betree.List (U64 × betree.Message)) + (new_msgs : betree.List (U64 × betree.Message)) : + Result (betree.List (U64 × betree.Message)) + := + match new_msgs with + | betree.List.Cons new_msg new_msgs_tl => + do + let (i, m) := new_msg + let msgs1 ← betree.Node.apply_to_internal msgs i m + betree.Node.apply_messages_to_internal msgs1 new_msgs_tl + | betree.List.Nil => Result.ok msgs + +/- [betree::betree::{betree::betree::Node#5}::lookup_mut_in_bindings]: + Source: 'src/betree.rs', lines 653:4-656:32 -/ +divergent def betree.Node.lookup_mut_in_bindings + (key : U64) (bindings : betree.List (U64 × U64)) : + Result ((betree.List (U64 × U64)) × (betree.List (U64 × U64) → Result + (betree.List (U64 × U64)))) + := + match bindings with + | betree.List.Cons hd tl => + let (i, i1) := hd + if i >= key + then Result.ok (betree.List.Cons (i, i1) tl, Result.ok) + else + do + let (l, lookup_mut_in_bindings_back) ← + betree.Node.lookup_mut_in_bindings key tl + let back := + fun ret => + do + let tl1 ← lookup_mut_in_bindings_back ret + Result.ok (betree.List.Cons (i, i1) tl1) + Result.ok (l, back) + | betree.List.Nil => Result.ok (betree.List.Nil, Result.ok) + +/- [betree::betree::{betree::betree::Node#5}::apply_to_leaf]: + Source: 'src/betree.rs', lines 460:4-460:87 -/ +def betree.Node.apply_to_leaf + (bindings : betree.List (U64 × U64)) (key : U64) (new_msg : betree.Message) + : + Result (betree.List (U64 × U64)) + := + do + let (bindings1, lookup_mut_in_bindings_back) ← + betree.Node.lookup_mut_in_bindings key bindings + let b ← betree.ListPairU64T.head_has_key U64 bindings1 key + if b + then + do + let (hd, bindings2) ← betree.List.pop_front (U64 × U64) bindings1 + match new_msg with + | betree.Message.Insert v => + do + let bindings3 ← betree.List.push_front (U64 × U64) bindings2 (key, v) + lookup_mut_in_bindings_back bindings3 + | betree.Message.Delete => lookup_mut_in_bindings_back bindings2 + | betree.Message.Upsert s => + do + let (_, i) := hd + let v ← betree.upsert_update (some i) s + let bindings3 ← betree.List.push_front (U64 × U64) bindings2 (key, v) + lookup_mut_in_bindings_back bindings3 + else + match new_msg with + | betree.Message.Insert v => + do + let bindings2 ← betree.List.push_front (U64 × U64) bindings1 (key, v) + lookup_mut_in_bindings_back bindings2 + | betree.Message.Delete => lookup_mut_in_bindings_back bindings1 + | betree.Message.Upsert s => + do + let v ← betree.upsert_update none s + let bindings2 ← betree.List.push_front (U64 × U64) bindings1 (key, v) + lookup_mut_in_bindings_back bindings2 + +/- [betree::betree::{betree::betree::Node#5}::apply_messages_to_leaf]: + Source: 'src/betree.rs', lines 444:4-447:5 -/ +divergent def betree.Node.apply_messages_to_leaf + (bindings : betree.List (U64 × U64)) + (new_msgs : betree.List (U64 × betree.Message)) : + Result (betree.List (U64 × U64)) + := + match new_msgs with + | betree.List.Cons new_msg new_msgs_tl => + do + let (i, m) := new_msg + let bindings1 ← betree.Node.apply_to_leaf bindings i m + betree.Node.apply_messages_to_leaf bindings1 new_msgs_tl + | betree.List.Nil => Result.ok bindings + +/- [betree::betree::{betree::betree::Internal#4}::flush]: + Source: 'src/betree.rs', lines 410:4-415:26 -/ +mutual divergent def betree.Internal.flush + (self : betree.Internal) (params : betree.Params) + (node_id_cnt : betree.NodeIdCounter) + (content : betree.List (U64 × betree.Message)) (st : State) : + Result (State × ((betree.List (U64 × betree.Message)) × (betree.Internal + × betree.NodeIdCounter))) + := + do + let p ← + betree.ListPairU64T.partition_at_pivot betree.Message content self.pivot + let (msgs_left, msgs_right) := p + let len_left ← betree.List.len (U64 × betree.Message) msgs_left + if len_left >= params.min_flush_size + then + do + let (st1, p1) ← + betree.Node.apply_messages self.left params node_id_cnt msgs_left st + let (n, node_id_cnt1) := p1 + let len_right ← betree.List.len (U64 × betree.Message) msgs_right + if len_right >= params.min_flush_size + then + do + let (st2, p2) ← + betree.Node.apply_messages self.right params node_id_cnt1 msgs_right + st1 + let (n1, node_id_cnt2) := p2 + Result.ok (st2, (betree.List.Nil, (betree.Internal.mk self.id self.pivot + n n1, node_id_cnt2))) + else + Result.ok (st1, (msgs_right, (betree.Internal.mk self.id self.pivot n + self.right, node_id_cnt1))) + else + do + let (st1, p1) ← + betree.Node.apply_messages self.right params node_id_cnt msgs_right st + let (n, node_id_cnt1) := p1 + Result.ok (st1, (msgs_left, (betree.Internal.mk self.id self.pivot + self.left n, node_id_cnt1))) + +/- [betree::betree::{betree::betree::Node#5}::apply_messages]: + Source: 'src/betree.rs', lines 588:4-593:5 -/ +divergent def betree.Node.apply_messages + (self : betree.Node) (params : betree.Params) + (node_id_cnt : betree.NodeIdCounter) + (msgs : betree.List (U64 × betree.Message)) (st : State) : + Result (State × (betree.Node × betree.NodeIdCounter)) + := + match self with + | betree.Node.Internal node => + do + let (st1, content) ← betree.load_internal_node node.id st + let content1 ← betree.Node.apply_messages_to_internal content msgs + let num_msgs ← betree.List.len (U64 × betree.Message) content1 + if num_msgs >= params.min_flush_size + then + do + let (st2, (content2, p)) ← + betree.Internal.flush node params node_id_cnt content1 st1 + let (node1, node_id_cnt1) := p + let (st3, _) ← betree.store_internal_node node1.id content2 st2 + Result.ok (st3, (betree.Node.Internal node1, node_id_cnt1)) + else + do + let (st2, _) ← betree.store_internal_node node.id content1 st1 + Result.ok (st2, (betree.Node.Internal node, node_id_cnt)) + | betree.Node.Leaf node => + do + let (st1, content) ← betree.load_leaf_node node.id st + let content1 ← betree.Node.apply_messages_to_leaf content msgs + let len ← betree.List.len (U64 × U64) content1 + let i ← 2#u64 * params.split_size + if len >= i + then + do + let (st2, (new_node, node_id_cnt1)) ← + betree.Leaf.split node content1 params node_id_cnt st1 + let (st3, _) ← betree.store_leaf_node node.id betree.List.Nil st2 + Result.ok (st3, (betree.Node.Internal new_node, node_id_cnt1)) + else + do + let (st2, _) ← betree.store_leaf_node node.id content1 st1 + Result.ok (st2, (betree.Node.Leaf { node with size := len }, + node_id_cnt)) + +end + +/- [betree::betree::{betree::betree::Node#5}::apply]: + Source: 'src/betree.rs', lines 576:4-582:5 -/ +def betree.Node.apply + (self : betree.Node) (params : betree.Params) + (node_id_cnt : betree.NodeIdCounter) (key : U64) (new_msg : betree.Message) + (st : State) : + Result (State × (betree.Node × betree.NodeIdCounter)) + := + do + let (st1, p) ← + betree.Node.apply_messages self params node_id_cnt (betree.List.Cons (key, + new_msg) betree.List.Nil) st + let (self1, node_id_cnt1) := p + Result.ok (st1, (self1, node_id_cnt1)) + +/- [betree::betree::{betree::betree::BeTree#6}::new]: + Source: 'src/betree.rs', lines 849:4-849:60 -/ +def betree.BeTree.new + (min_flush_size : U64) (split_size : U64) (st : State) : + Result (State × betree.BeTree) + := + do + let node_id_cnt ← betree.NodeIdCounter.new + let (id, node_id_cnt1) ← betree.NodeIdCounter.fresh_id node_id_cnt + let (st1, _) ← betree.store_leaf_node id betree.List.Nil st + Result.ok (st1, + { + params := { min_flush_size := min_flush_size, split_size := split_size }, + node_id_cnt := node_id_cnt1, + root := (betree.Node.Leaf { id := id, size := 0#u64 }) + }) + +/- [betree::betree::{betree::betree::BeTree#6}::apply]: + Source: 'src/betree.rs', lines 868:4-868:47 -/ +def betree.BeTree.apply + (self : betree.BeTree) (key : U64) (msg : betree.Message) (st : State) : + Result (State × betree.BeTree) + := + do + let (st1, p) ← + betree.Node.apply self.root self.params self.node_id_cnt key msg st + let (n, nic) := p + Result.ok (st1, { self with node_id_cnt := nic, root := n }) + +/- [betree::betree::{betree::betree::BeTree#6}::insert]: + Source: 'src/betree.rs', lines 874:4-874:52 -/ +def betree.BeTree.insert + (self : betree.BeTree) (key : U64) (value : U64) (st : State) : + Result (State × betree.BeTree) + := + betree.BeTree.apply self key (betree.Message.Insert value) st + +/- [betree::betree::{betree::betree::BeTree#6}::delete]: + Source: 'src/betree.rs', lines 880:4-880:38 -/ +def betree.BeTree.delete + (self : betree.BeTree) (key : U64) (st : State) : + Result (State × betree.BeTree) + := + betree.BeTree.apply self key betree.Message.Delete st + +/- [betree::betree::{betree::betree::BeTree#6}::upsert]: + Source: 'src/betree.rs', lines 886:4-886:59 -/ +def betree.BeTree.upsert + (self : betree.BeTree) (key : U64) (upd : betree.UpsertFunState) (st : State) + : + Result (State × betree.BeTree) + := + betree.BeTree.apply self key (betree.Message.Upsert upd) st + +/- [betree::betree::{betree::betree::BeTree#6}::lookup]: + Source: 'src/betree.rs', lines 895:4-895:62 -/ +def betree.BeTree.lookup + (self : betree.BeTree) (key : U64) (st : State) : + Result (State × ((Option U64) × betree.BeTree)) + := + do + let (st1, (o, n)) ← betree.Node.lookup self.root key st + Result.ok (st1, (o, { self with root := n })) + +/- [betree::main]: + Source: 'src/main.rs', lines 4:0-4:9 -/ +def main : Result Unit := + Result.ok () + +/- Unit test for [betree::main] -/ +#assert (main == Result.ok ()) + +end betree diff --git a/tests/lean/Betree/FunsExternal.lean b/tests/lean/Betree/FunsExternal.lean new file mode 100644 index 00000000..859cbd68 --- /dev/null +++ b/tests/lean/Betree/FunsExternal.lean @@ -0,0 +1,30 @@ +-- [betree]: external functions. +import Base +import Betree.Types +open Primitives +open betree + +-- TODO: fill those bodies + +/- [betree::betree_utils::load_internal_node] -/ +def betree_utils.load_internal_node + : + U64 → State → Result (State × (betree.List (U64 × betree.Message))) := + fun _ _ => .fail .panic + +/- [betree::betree_utils::store_internal_node] -/ +def betree_utils.store_internal_node + : + U64 → betree.List (U64 × betree.Message) → State → Result (State + × Unit) := + fun _ _ _ => .fail .panic + +/- [betree::betree_utils::load_leaf_node] -/ +def betree_utils.load_leaf_node + : U64 → State → Result (State × (betree.List (U64 × U64))) := + fun _ _ => .fail .panic + +/- [betree::betree_utils::store_leaf_node] -/ +def betree_utils.store_leaf_node + : U64 → betree.List (U64 × U64) → State → Result (State × Unit) := + fun _ _ _ => .fail .panic diff --git a/tests/lean/Betree/FunsExternal_Template.lean b/tests/lean/Betree/FunsExternal_Template.lean new file mode 100644 index 00000000..014f0d83 --- /dev/null +++ b/tests/lean/Betree/FunsExternal_Template.lean @@ -0,0 +1,30 @@ +-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS +-- [betree]: external functions. +-- This is a template file: rename it to "FunsExternal.lean" and fill the holes. +import Base +import Betree.Types +open Primitives +open betree + +/- [betree::betree_utils::load_internal_node]: + Source: 'src/betree_utils.rs', lines 98:0-98:63 -/ +axiom betree_utils.load_internal_node + : U64 → State → Result (State × (betree.List (U64 × betree.Message))) + +/- [betree::betree_utils::store_internal_node]: + Source: 'src/betree_utils.rs', lines 115:0-115:71 -/ +axiom betree_utils.store_internal_node + : + U64 → betree.List (U64 × betree.Message) → State → Result (State × + Unit) + +/- [betree::betree_utils::load_leaf_node]: + Source: 'src/betree_utils.rs', lines 132:0-132:55 -/ +axiom betree_utils.load_leaf_node + : U64 → State → Result (State × (betree.List (U64 × U64))) + +/- [betree::betree_utils::store_leaf_node]: + Source: 'src/betree_utils.rs', lines 145:0-145:63 -/ +axiom betree_utils.store_leaf_node + : U64 → betree.List (U64 × U64) → State → Result (State × Unit) + diff --git a/tests/lean/Betree/Types.lean b/tests/lean/Betree/Types.lean new file mode 100644 index 00000000..3b46c00c --- /dev/null +++ b/tests/lean/Betree/Types.lean @@ -0,0 +1,83 @@ +-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS +-- [betree]: type definitions +import Base +import Betree.TypesExternal +open Primitives + +namespace betree + +/- [betree::betree::List] + Source: 'src/betree.rs', lines 17:0-17:23 -/ +inductive betree.List (T : Type) := +| Cons : T → betree.List T → betree.List T +| Nil : betree.List T + +/- [betree::betree::UpsertFunState] + Source: 'src/betree.rs', lines 63:0-63:23 -/ +inductive betree.UpsertFunState := +| Add : U64 → betree.UpsertFunState +| Sub : U64 → betree.UpsertFunState + +/- [betree::betree::Message] + Source: 'src/betree.rs', lines 69:0-69:23 -/ +inductive betree.Message := +| Insert : U64 → betree.Message +| Delete : betree.Message +| Upsert : betree.UpsertFunState → betree.Message + +/- [betree::betree::Leaf] + Source: 'src/betree.rs', lines 167:0-167:11 -/ +structure betree.Leaf where + id : U64 + size : U64 + +mutual + +/- [betree::betree::Internal] + Source: 'src/betree.rs', lines 156:0-156:15 -/ +inductive betree.Internal := +| mk : U64 → U64 → betree.Node → betree.Node → betree.Internal + +/- [betree::betree::Node] + Source: 'src/betree.rs', lines 179:0-179:9 -/ +inductive betree.Node := +| Internal : betree.Internal → betree.Node +| Leaf : betree.Leaf → betree.Node + +end + +@[simp, reducible] +def betree.Internal.id (x : betree.Internal) := + match x with | betree.Internal.mk x1 _ _ _ => x1 + +@[simp, reducible] +def betree.Internal.pivot (x : betree.Internal) := + match x with | betree.Internal.mk _ x1 _ _ => x1 + +@[simp, reducible] +def betree.Internal.left (x : betree.Internal) := + match x with | betree.Internal.mk _ _ x1 _ => x1 + +@[simp, reducible] +def betree.Internal.right (x : betree.Internal) := + match x with | betree.Internal.mk _ _ _ x1 => x1 + +/- [betree::betree::Params] + Source: 'src/betree.rs', lines 187:0-187:13 -/ +structure betree.Params where + min_flush_size : U64 + split_size : U64 + +/- [betree::betree::NodeIdCounter] + Source: 'src/betree.rs', lines 201:0-201:20 -/ +structure betree.NodeIdCounter where + next_node_id : U64 + +/- [betree::betree::BeTree] + Source: 'src/betree.rs', lines 218:0-218:17 -/ +structure betree.BeTree where + params : betree.Params + node_id_cnt : betree.NodeIdCounter + root : betree.Node + +end betree diff --git a/tests/lean/Betree/TypesExternal.lean b/tests/lean/Betree/TypesExternal.lean new file mode 100644 index 00000000..34170271 --- /dev/null +++ b/tests/lean/Betree/TypesExternal.lean @@ -0,0 +1,7 @@ +-- [betree]: external types. +import Base +open Primitives + +/- The state type used in the state-error monad -/ +axiom State : Type + diff --git a/tests/lean/Betree/TypesExternal_Template.lean b/tests/lean/Betree/TypesExternal_Template.lean new file mode 100644 index 00000000..12fce657 --- /dev/null +++ b/tests/lean/Betree/TypesExternal_Template.lean @@ -0,0 +1,9 @@ +-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS +-- [betree]: external types. +-- This is a template file: rename it to "TypesExternal.lean" and fill the holes. +import Base +open Primitives + +/- The state type used in the state-error monad -/ +axiom State : Type + diff --git a/tests/lean/BetreeMain.lean b/tests/lean/BetreeMain.lean deleted file mode 100644 index 5f307877..00000000 --- a/tests/lean/BetreeMain.lean +++ /dev/null @@ -1 +0,0 @@ -import BetreeMain.Funs diff --git a/tests/lean/BetreeMain/Funs.lean b/tests/lean/BetreeMain/Funs.lean deleted file mode 100644 index f6fda6db..00000000 --- a/tests/lean/BetreeMain/Funs.lean +++ /dev/null @@ -1,699 +0,0 @@ --- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS --- [betree_main]: function definitions -import Base -import BetreeMain.Types -import BetreeMain.FunsExternal -open Primitives - -namespace betree_main - -/- [betree_main::betree::load_internal_node]: - Source: 'src/betree.rs', lines 36:0-36:52 -/ -def betree.load_internal_node - (id : U64) (st : State) : - Result (State × (betree.List (U64 × betree.Message))) - := - betree_utils.load_internal_node id st - -/- [betree_main::betree::store_internal_node]: - Source: 'src/betree.rs', lines 41:0-41:60 -/ -def betree.store_internal_node - (id : U64) (content : betree.List (U64 × betree.Message)) (st : State) : - Result (State × Unit) - := - betree_utils.store_internal_node id content st - -/- [betree_main::betree::load_leaf_node]: - Source: 'src/betree.rs', lines 46:0-46:44 -/ -def betree.load_leaf_node - (id : U64) (st : State) : Result (State × (betree.List (U64 × U64))) := - betree_utils.load_leaf_node id st - -/- [betree_main::betree::store_leaf_node]: - Source: 'src/betree.rs', lines 51:0-51:52 -/ -def betree.store_leaf_node - (id : U64) (content : betree.List (U64 × U64)) (st : State) : - Result (State × Unit) - := - betree_utils.store_leaf_node id content st - -/- [betree_main::betree::fresh_node_id]: - Source: 'src/betree.rs', lines 55:0-55:48 -/ -def betree.fresh_node_id (counter : U64) : Result (U64 × U64) := - do - let counter1 ← counter + 1#u64 - Result.ok (counter, counter1) - -/- [betree_main::betree::{betree_main::betree::NodeIdCounter}::new]: - Source: 'src/betree.rs', lines 206:4-206:20 -/ -def betree.NodeIdCounter.new : Result betree.NodeIdCounter := - Result.ok { next_node_id := 0#u64 } - -/- [betree_main::betree::{betree_main::betree::NodeIdCounter}::fresh_id]: - Source: 'src/betree.rs', lines 210:4-210:36 -/ -def betree.NodeIdCounter.fresh_id - (self : betree.NodeIdCounter) : Result (U64 × betree.NodeIdCounter) := - do - let i ← self.next_node_id + 1#u64 - Result.ok (self.next_node_id, { next_node_id := i }) - -/- [betree_main::betree::upsert_update]: - Source: 'src/betree.rs', lines 234:0-234:70 -/ -def betree.upsert_update - (prev : Option U64) (st : betree.UpsertFunState) : Result U64 := - match prev with - | none => - match st with - | betree.UpsertFunState.Add v => Result.ok v - | betree.UpsertFunState.Sub _ => Result.ok 0#u64 - | some prev1 => - match st with - | betree.UpsertFunState.Add v => - do - let margin ← core_u64_max - prev1 - if margin >= v - then prev1 + v - else Result.ok core_u64_max - | betree.UpsertFunState.Sub v => - if prev1 >= v - then prev1 - v - else Result.ok 0#u64 - -/- [betree_main::betree::{betree_main::betree::List#1}::len]: - Source: 'src/betree.rs', lines 276:4-276:24 -/ -divergent def betree.List.len (T : Type) (self : betree.List T) : Result U64 := - match self with - | betree.List.Cons _ tl => do - let i ← betree.List.len T tl - 1#u64 + i - | betree.List.Nil => Result.ok 0#u64 - -/- [betree_main::betree::{betree_main::betree::List#1}::split_at]: - Source: 'src/betree.rs', lines 284:4-284:51 -/ -divergent def betree.List.split_at - (T : Type) (self : betree.List T) (n : U64) : - Result ((betree.List T) × (betree.List T)) - := - if n = 0#u64 - then Result.ok (betree.List.Nil, self) - else - match self with - | betree.List.Cons hd tl => - do - let i ← n - 1#u64 - let p ← betree.List.split_at T tl i - let (ls0, ls1) := p - Result.ok (betree.List.Cons hd ls0, ls1) - | betree.List.Nil => Result.fail .panic - -/- [betree_main::betree::{betree_main::betree::List#1}::push_front]: - Source: 'src/betree.rs', lines 299:4-299:34 -/ -def betree.List.push_front - (T : Type) (self : betree.List T) (x : T) : Result (betree.List T) := - let (tl, _) := core.mem.replace (betree.List T) self betree.List.Nil - Result.ok (betree.List.Cons x tl) - -/- [betree_main::betree::{betree_main::betree::List#1}::pop_front]: - Source: 'src/betree.rs', lines 306:4-306:32 -/ -def betree.List.pop_front - (T : Type) (self : betree.List T) : Result (T × (betree.List T)) := - let (ls, _) := core.mem.replace (betree.List T) self betree.List.Nil - match ls with - | betree.List.Cons x tl => Result.ok (x, tl) - | betree.List.Nil => Result.fail .panic - -/- [betree_main::betree::{betree_main::betree::List#1}::hd]: - Source: 'src/betree.rs', lines 318:4-318:22 -/ -def betree.List.hd (T : Type) (self : betree.List T) : Result T := - match self with - | betree.List.Cons hd _ => Result.ok hd - | betree.List.Nil => Result.fail .panic - -/- [betree_main::betree::{betree_main::betree::List<(u64, T)>#2}::head_has_key]: - Source: 'src/betree.rs', lines 327:4-327:44 -/ -def betree.ListPairU64T.head_has_key - (T : Type) (self : betree.List (U64 × T)) (key : U64) : Result Bool := - match self with - | betree.List.Cons hd _ => let (i, _) := hd - Result.ok (i = key) - | betree.List.Nil => Result.ok false - -/- [betree_main::betree::{betree_main::betree::List<(u64, T)>#2}::partition_at_pivot]: - Source: 'src/betree.rs', lines 339:4-339:73 -/ -divergent def betree.ListPairU64T.partition_at_pivot - (T : Type) (self : betree.List (U64 × T)) (pivot : U64) : - Result ((betree.List (U64 × T)) × (betree.List (U64 × T))) - := - match self with - | betree.List.Cons hd tl => - let (i, t) := hd - if i >= pivot - then Result.ok (betree.List.Nil, betree.List.Cons (i, t) tl) - else - do - let p ← betree.ListPairU64T.partition_at_pivot T tl pivot - let (ls0, ls1) := p - Result.ok (betree.List.Cons (i, t) ls0, ls1) - | betree.List.Nil => Result.ok (betree.List.Nil, betree.List.Nil) - -/- [betree_main::betree::{betree_main::betree::Leaf#3}::split]: - Source: 'src/betree.rs', lines 359:4-364:17 -/ -def betree.Leaf.split - (self : betree.Leaf) (content : betree.List (U64 × U64)) - (params : betree.Params) (node_id_cnt : betree.NodeIdCounter) (st : State) : - Result (State × (betree.Internal × betree.NodeIdCounter)) - := - do - let p ← betree.List.split_at (U64 × U64) content params.split_size - let (content0, content1) := p - let p1 ← betree.List.hd (U64 × U64) content1 - let (pivot, _) := p1 - let (id0, node_id_cnt1) ← betree.NodeIdCounter.fresh_id node_id_cnt - let (id1, node_id_cnt2) ← betree.NodeIdCounter.fresh_id node_id_cnt1 - let (st1, _) ← betree.store_leaf_node id0 content0 st - let (st2, _) ← betree.store_leaf_node id1 content1 st1 - let n := betree.Node.Leaf { id := id0, size := params.split_size } - let n1 := betree.Node.Leaf { id := id1, size := params.split_size } - Result.ok (st2, (betree.Internal.mk self.id pivot n n1, node_id_cnt2)) - -/- [betree_main::betree::{betree_main::betree::Node#5}::lookup_first_message_for_key]: - Source: 'src/betree.rs', lines 789:4-792:34 -/ -divergent 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 × - betree.Message) → Result (betree.List (U64 × betree.Message)))) - := - match msgs with - | betree.List.Cons x next_msgs => - let (i, m) := x - if i >= key - then Result.ok (betree.List.Cons (i, m) next_msgs, Result.ok) - else - do - let (l, lookup_first_message_for_key_back) ← - betree.Node.lookup_first_message_for_key key next_msgs - let back := - fun ret => - do - let next_msgs1 ← lookup_first_message_for_key_back ret - Result.ok (betree.List.Cons (i, m) next_msgs1) - Result.ok (l, back) - | betree.List.Nil => Result.ok (betree.List.Nil, Result.ok) - -/- [betree_main::betree::{betree_main::betree::Node#5}::lookup_in_bindings]: - Source: 'src/betree.rs', lines 636:4-636:80 -/ -divergent def betree.Node.lookup_in_bindings - (key : U64) (bindings : betree.List (U64 × U64)) : Result (Option U64) := - match bindings with - | betree.List.Cons hd tl => - let (i, i1) := hd - if i = key - then Result.ok (some i1) - else - if i > key - then Result.ok none - else betree.Node.lookup_in_bindings key tl - | betree.List.Nil => Result.ok none - -/- [betree_main::betree::{betree_main::betree::Node#5}::apply_upserts]: - Source: 'src/betree.rs', lines 819:4-819:90 -/ -divergent def betree.Node.apply_upserts - (msgs : betree.List (U64 × betree.Message)) (prev : Option U64) (key : U64) - : - Result (U64 × (betree.List (U64 × betree.Message))) - := - do - let b ← betree.ListPairU64T.head_has_key betree.Message msgs key - if b - then - do - let (msg, msgs1) ← betree.List.pop_front (U64 × betree.Message) msgs - let (_, m) := msg - match m with - | betree.Message.Insert _ => Result.fail .panic - | betree.Message.Delete => Result.fail .panic - | betree.Message.Upsert s => - do - let v ← betree.upsert_update prev s - betree.Node.apply_upserts msgs1 (some v) key - else - do - let v ← core.option.Option.unwrap U64 prev - let msgs1 ← - betree.List.push_front (U64 × betree.Message) msgs (key, - betree.Message.Insert v) - Result.ok (v, msgs1) - -/- [betree_main::betree::{betree_main::betree::Internal#4}::lookup_in_children]: - Source: 'src/betree.rs', lines 395:4-395:63 -/ -mutual divergent def betree.Internal.lookup_in_children - (self : betree.Internal) (key : U64) (st : State) : - Result (State × ((Option U64) × betree.Internal)) - := - if key < self.pivot - then - do - let (st1, (o, n)) ← betree.Node.lookup self.left key st - Result.ok (st1, (o, betree.Internal.mk self.id self.pivot n self.right)) - else - do - let (st1, (o, n)) ← betree.Node.lookup self.right key st - Result.ok (st1, (o, betree.Internal.mk self.id self.pivot self.left n)) - -/- [betree_main::betree::{betree_main::betree::Node#5}::lookup]: - Source: 'src/betree.rs', lines 709:4-709:58 -/ -divergent def betree.Node.lookup - (self : betree.Node) (key : U64) (st : State) : - Result (State × ((Option U64) × betree.Node)) - := - match self with - | betree.Node.Internal node => - do - let (st1, msgs) ← betree.load_internal_node node.id st - let (pending, lookup_first_message_for_key_back) ← - betree.Node.lookup_first_message_for_key key msgs - match pending with - | betree.List.Cons p l => - let (k, msg) := p - if k != key - then - do - let (st2, (o, node1)) ← - betree.Internal.lookup_in_children node key st1 - let _ ← - lookup_first_message_for_key_back (betree.List.Cons (k, msg) l) - Result.ok (st2, (o, betree.Node.Internal node1)) - else - match msg with - | betree.Message.Insert v => - do - let _ ← - lookup_first_message_for_key_back (betree.List.Cons (k, - betree.Message.Insert v) l) - Result.ok (st1, (some v, betree.Node.Internal node)) - | betree.Message.Delete => - do - let _ ← - lookup_first_message_for_key_back (betree.List.Cons (k, - betree.Message.Delete) l) - Result.ok (st1, (none, betree.Node.Internal node)) - | betree.Message.Upsert ufs => - do - let (st2, (v, node1)) ← - betree.Internal.lookup_in_children node key st1 - let (v1, pending1) ← - betree.Node.apply_upserts (betree.List.Cons (k, - betree.Message.Upsert ufs) l) v key - let msgs1 ← lookup_first_message_for_key_back pending1 - let (st3, _) ← betree.store_internal_node node1.id msgs1 st2 - Result.ok (st3, (some v1, betree.Node.Internal node1)) - | betree.List.Nil => - do - let (st2, (o, node1)) ← betree.Internal.lookup_in_children node key st1 - let _ ← lookup_first_message_for_key_back betree.List.Nil - Result.ok (st2, (o, betree.Node.Internal node1)) - | betree.Node.Leaf node => - do - let (st1, bindings) ← betree.load_leaf_node node.id st - let o ← betree.Node.lookup_in_bindings key bindings - Result.ok (st1, (o, betree.Node.Leaf node)) - -end - -/- [betree_main::betree::{betree_main::betree::Node#5}::filter_messages_for_key]: - Source: 'src/betree.rs', lines 674:4-674:77 -/ -divergent def betree.Node.filter_messages_for_key - (key : U64) (msgs : betree.List (U64 × betree.Message)) : - Result (betree.List (U64 × betree.Message)) - := - match msgs with - | betree.List.Cons p l => - let (k, m) := p - if k = key - then - do - let (_, msgs1) ← - betree.List.pop_front (U64 × betree.Message) (betree.List.Cons (k, m) - l) - betree.Node.filter_messages_for_key key msgs1 - else Result.ok (betree.List.Cons (k, m) l) - | betree.List.Nil => Result.ok betree.List.Nil - -/- [betree_main::betree::{betree_main::betree::Node#5}::lookup_first_message_after_key]: - Source: 'src/betree.rs', lines 689:4-692:34 -/ -divergent 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 × - betree.Message) → Result (betree.List (U64 × betree.Message)))) - := - match msgs with - | betree.List.Cons p next_msgs => - let (k, m) := p - if k = key - then - do - let (l, lookup_first_message_after_key_back) ← - betree.Node.lookup_first_message_after_key key next_msgs - let back := - fun ret => - do - let next_msgs1 ← lookup_first_message_after_key_back ret - Result.ok (betree.List.Cons (k, m) next_msgs1) - Result.ok (l, back) - else Result.ok (betree.List.Cons (k, m) next_msgs, Result.ok) - | betree.List.Nil => Result.ok (betree.List.Nil, Result.ok) - -/- [betree_main::betree::{betree_main::betree::Node#5}::apply_to_internal]: - Source: 'src/betree.rs', lines 521:4-521:89 -/ -def betree.Node.apply_to_internal - (msgs : betree.List (U64 × betree.Message)) (key : U64) - (new_msg : betree.Message) : - Result (betree.List (U64 × betree.Message)) - := - do - let (msgs1, lookup_first_message_for_key_back) ← - betree.Node.lookup_first_message_for_key key msgs - let b ← betree.ListPairU64T.head_has_key betree.Message msgs1 key - if b - then - match new_msg with - | betree.Message.Insert i => - do - let msgs2 ← betree.Node.filter_messages_for_key key msgs1 - let msgs3 ← - betree.List.push_front (U64 × betree.Message) msgs2 (key, - betree.Message.Insert i) - lookup_first_message_for_key_back msgs3 - | betree.Message.Delete => - do - let msgs2 ← betree.Node.filter_messages_for_key key msgs1 - let msgs3 ← - betree.List.push_front (U64 × betree.Message) msgs2 (key, - betree.Message.Delete) - lookup_first_message_for_key_back msgs3 - | betree.Message.Upsert s => - do - let p ← betree.List.hd (U64 × betree.Message) msgs1 - let (_, m) := p - match m with - | betree.Message.Insert prev => - do - let v ← betree.upsert_update (some prev) s - let (_, msgs2) ← betree.List.pop_front (U64 × betree.Message) msgs1 - let msgs3 ← - betree.List.push_front (U64 × betree.Message) msgs2 (key, - betree.Message.Insert v) - lookup_first_message_for_key_back msgs3 - | betree.Message.Delete => - do - let (_, msgs2) ← betree.List.pop_front (U64 × betree.Message) msgs1 - let v ← betree.upsert_update none s - let msgs3 ← - betree.List.push_front (U64 × betree.Message) msgs2 (key, - betree.Message.Insert v) - lookup_first_message_for_key_back msgs3 - | betree.Message.Upsert _ => - do - let (msgs2, lookup_first_message_after_key_back) ← - betree.Node.lookup_first_message_after_key key msgs1 - let msgs3 ← - betree.List.push_front (U64 × betree.Message) msgs2 (key, - betree.Message.Upsert s) - let msgs4 ← lookup_first_message_after_key_back msgs3 - lookup_first_message_for_key_back msgs4 - else - do - let msgs2 ← - betree.List.push_front (U64 × betree.Message) msgs1 (key, new_msg) - lookup_first_message_for_key_back msgs2 - -/- [betree_main::betree::{betree_main::betree::Node#5}::apply_messages_to_internal]: - Source: 'src/betree.rs', lines 502:4-505:5 -/ -divergent def betree.Node.apply_messages_to_internal - (msgs : betree.List (U64 × betree.Message)) - (new_msgs : betree.List (U64 × betree.Message)) : - Result (betree.List (U64 × betree.Message)) - := - match new_msgs with - | betree.List.Cons new_msg new_msgs_tl => - do - let (i, m) := new_msg - let msgs1 ← betree.Node.apply_to_internal msgs i m - betree.Node.apply_messages_to_internal msgs1 new_msgs_tl - | betree.List.Nil => Result.ok msgs - -/- [betree_main::betree::{betree_main::betree::Node#5}::lookup_mut_in_bindings]: - Source: 'src/betree.rs', lines 653:4-656:32 -/ -divergent def betree.Node.lookup_mut_in_bindings - (key : U64) (bindings : betree.List (U64 × U64)) : - Result ((betree.List (U64 × U64)) × (betree.List (U64 × U64) → Result - (betree.List (U64 × U64)))) - := - match bindings with - | betree.List.Cons hd tl => - let (i, i1) := hd - if i >= key - then Result.ok (betree.List.Cons (i, i1) tl, Result.ok) - else - do - let (l, lookup_mut_in_bindings_back) ← - betree.Node.lookup_mut_in_bindings key tl - let back := - fun ret => - do - let tl1 ← lookup_mut_in_bindings_back ret - Result.ok (betree.List.Cons (i, i1) tl1) - Result.ok (l, back) - | betree.List.Nil => Result.ok (betree.List.Nil, Result.ok) - -/- [betree_main::betree::{betree_main::betree::Node#5}::apply_to_leaf]: - Source: 'src/betree.rs', lines 460:4-460:87 -/ -def betree.Node.apply_to_leaf - (bindings : betree.List (U64 × U64)) (key : U64) (new_msg : betree.Message) - : - Result (betree.List (U64 × U64)) - := - do - let (bindings1, lookup_mut_in_bindings_back) ← - betree.Node.lookup_mut_in_bindings key bindings - let b ← betree.ListPairU64T.head_has_key U64 bindings1 key - if b - then - do - let (hd, bindings2) ← betree.List.pop_front (U64 × U64) bindings1 - match new_msg with - | betree.Message.Insert v => - do - let bindings3 ← betree.List.push_front (U64 × U64) bindings2 (key, v) - lookup_mut_in_bindings_back bindings3 - | betree.Message.Delete => lookup_mut_in_bindings_back bindings2 - | betree.Message.Upsert s => - do - let (_, i) := hd - let v ← betree.upsert_update (some i) s - let bindings3 ← betree.List.push_front (U64 × U64) bindings2 (key, v) - lookup_mut_in_bindings_back bindings3 - else - match new_msg with - | betree.Message.Insert v => - do - let bindings2 ← betree.List.push_front (U64 × U64) bindings1 (key, v) - lookup_mut_in_bindings_back bindings2 - | betree.Message.Delete => lookup_mut_in_bindings_back bindings1 - | betree.Message.Upsert s => - do - let v ← betree.upsert_update none s - let bindings2 ← betree.List.push_front (U64 × U64) bindings1 (key, v) - lookup_mut_in_bindings_back bindings2 - -/- [betree_main::betree::{betree_main::betree::Node#5}::apply_messages_to_leaf]: - Source: 'src/betree.rs', lines 444:4-447:5 -/ -divergent def betree.Node.apply_messages_to_leaf - (bindings : betree.List (U64 × U64)) - (new_msgs : betree.List (U64 × betree.Message)) : - Result (betree.List (U64 × U64)) - := - match new_msgs with - | betree.List.Cons new_msg new_msgs_tl => - do - let (i, m) := new_msg - let bindings1 ← betree.Node.apply_to_leaf bindings i m - betree.Node.apply_messages_to_leaf bindings1 new_msgs_tl - | betree.List.Nil => Result.ok bindings - -/- [betree_main::betree::{betree_main::betree::Internal#4}::flush]: - Source: 'src/betree.rs', lines 410:4-415:26 -/ -mutual divergent def betree.Internal.flush - (self : betree.Internal) (params : betree.Params) - (node_id_cnt : betree.NodeIdCounter) - (content : betree.List (U64 × betree.Message)) (st : State) : - Result (State × ((betree.List (U64 × betree.Message)) × (betree.Internal - × betree.NodeIdCounter))) - := - do - let p ← - betree.ListPairU64T.partition_at_pivot betree.Message content self.pivot - let (msgs_left, msgs_right) := p - let len_left ← betree.List.len (U64 × betree.Message) msgs_left - if len_left >= params.min_flush_size - then - do - let (st1, p1) ← - betree.Node.apply_messages self.left params node_id_cnt msgs_left st - let (n, node_id_cnt1) := p1 - let len_right ← betree.List.len (U64 × betree.Message) msgs_right - if len_right >= params.min_flush_size - then - do - let (st2, p2) ← - betree.Node.apply_messages self.right params node_id_cnt1 msgs_right - st1 - let (n1, node_id_cnt2) := p2 - Result.ok (st2, (betree.List.Nil, (betree.Internal.mk self.id self.pivot - n n1, node_id_cnt2))) - else - Result.ok (st1, (msgs_right, (betree.Internal.mk self.id self.pivot n - self.right, node_id_cnt1))) - else - do - let (st1, p1) ← - betree.Node.apply_messages self.right params node_id_cnt msgs_right st - let (n, node_id_cnt1) := p1 - Result.ok (st1, (msgs_left, (betree.Internal.mk self.id self.pivot - self.left n, node_id_cnt1))) - -/- [betree_main::betree::{betree_main::betree::Node#5}::apply_messages]: - Source: 'src/betree.rs', lines 588:4-593:5 -/ -divergent def betree.Node.apply_messages - (self : betree.Node) (params : betree.Params) - (node_id_cnt : betree.NodeIdCounter) - (msgs : betree.List (U64 × betree.Message)) (st : State) : - Result (State × (betree.Node × betree.NodeIdCounter)) - := - match self with - | betree.Node.Internal node => - do - let (st1, content) ← betree.load_internal_node node.id st - let content1 ← betree.Node.apply_messages_to_internal content msgs - let num_msgs ← betree.List.len (U64 × betree.Message) content1 - if num_msgs >= params.min_flush_size - then - do - let (st2, (content2, p)) ← - betree.Internal.flush node params node_id_cnt content1 st1 - let (node1, node_id_cnt1) := p - let (st3, _) ← betree.store_internal_node node1.id content2 st2 - Result.ok (st3, (betree.Node.Internal node1, node_id_cnt1)) - else - do - let (st2, _) ← betree.store_internal_node node.id content1 st1 - Result.ok (st2, (betree.Node.Internal node, node_id_cnt)) - | betree.Node.Leaf node => - do - let (st1, content) ← betree.load_leaf_node node.id st - let content1 ← betree.Node.apply_messages_to_leaf content msgs - let len ← betree.List.len (U64 × U64) content1 - let i ← 2#u64 * params.split_size - if len >= i - then - do - let (st2, (new_node, node_id_cnt1)) ← - betree.Leaf.split node content1 params node_id_cnt st1 - let (st3, _) ← betree.store_leaf_node node.id betree.List.Nil st2 - Result.ok (st3, (betree.Node.Internal new_node, node_id_cnt1)) - else - do - let (st2, _) ← betree.store_leaf_node node.id content1 st1 - Result.ok (st2, (betree.Node.Leaf { node with size := len }, - node_id_cnt)) - -end - -/- [betree_main::betree::{betree_main::betree::Node#5}::apply]: - Source: 'src/betree.rs', lines 576:4-582:5 -/ -def betree.Node.apply - (self : betree.Node) (params : betree.Params) - (node_id_cnt : betree.NodeIdCounter) (key : U64) (new_msg : betree.Message) - (st : State) : - Result (State × (betree.Node × betree.NodeIdCounter)) - := - do - let (st1, p) ← - betree.Node.apply_messages self params node_id_cnt (betree.List.Cons (key, - new_msg) betree.List.Nil) st - let (self1, node_id_cnt1) := p - Result.ok (st1, (self1, node_id_cnt1)) - -/- [betree_main::betree::{betree_main::betree::BeTree#6}::new]: - Source: 'src/betree.rs', lines 849:4-849:60 -/ -def betree.BeTree.new - (min_flush_size : U64) (split_size : U64) (st : State) : - Result (State × betree.BeTree) - := - do - let node_id_cnt ← betree.NodeIdCounter.new - let (id, node_id_cnt1) ← betree.NodeIdCounter.fresh_id node_id_cnt - let (st1, _) ← betree.store_leaf_node id betree.List.Nil st - Result.ok (st1, - { - params := { min_flush_size := min_flush_size, split_size := split_size }, - node_id_cnt := node_id_cnt1, - root := (betree.Node.Leaf { id := id, size := 0#u64 }) - }) - -/- [betree_main::betree::{betree_main::betree::BeTree#6}::apply]: - Source: 'src/betree.rs', lines 868:4-868:47 -/ -def betree.BeTree.apply - (self : betree.BeTree) (key : U64) (msg : betree.Message) (st : State) : - Result (State × betree.BeTree) - := - do - let (st1, p) ← - betree.Node.apply self.root self.params self.node_id_cnt key msg st - let (n, nic) := p - Result.ok (st1, { self with node_id_cnt := nic, root := n }) - -/- [betree_main::betree::{betree_main::betree::BeTree#6}::insert]: - Source: 'src/betree.rs', lines 874:4-874:52 -/ -def betree.BeTree.insert - (self : betree.BeTree) (key : U64) (value : U64) (st : State) : - Result (State × betree.BeTree) - := - betree.BeTree.apply self key (betree.Message.Insert value) st - -/- [betree_main::betree::{betree_main::betree::BeTree#6}::delete]: - Source: 'src/betree.rs', lines 880:4-880:38 -/ -def betree.BeTree.delete - (self : betree.BeTree) (key : U64) (st : State) : - Result (State × betree.BeTree) - := - betree.BeTree.apply self key betree.Message.Delete st - -/- [betree_main::betree::{betree_main::betree::BeTree#6}::upsert]: - Source: 'src/betree.rs', lines 886:4-886:59 -/ -def betree.BeTree.upsert - (self : betree.BeTree) (key : U64) (upd : betree.UpsertFunState) (st : State) - : - Result (State × betree.BeTree) - := - betree.BeTree.apply self key (betree.Message.Upsert upd) st - -/- [betree_main::betree::{betree_main::betree::BeTree#6}::lookup]: - Source: 'src/betree.rs', lines 895:4-895:62 -/ -def betree.BeTree.lookup - (self : betree.BeTree) (key : U64) (st : State) : - Result (State × ((Option U64) × betree.BeTree)) - := - do - let (st1, (o, n)) ← betree.Node.lookup self.root key st - Result.ok (st1, (o, { self with root := n })) - -/- [betree_main::main]: - Source: 'src/main.rs', lines 4:0-4:9 -/ -def main : Result Unit := - Result.ok () - -/- Unit test for [betree_main::main] -/ -#assert (main == Result.ok ()) - -end betree_main diff --git a/tests/lean/BetreeMain/FunsExternal.lean b/tests/lean/BetreeMain/FunsExternal.lean deleted file mode 100644 index d26177fb..00000000 --- a/tests/lean/BetreeMain/FunsExternal.lean +++ /dev/null @@ -1,30 +0,0 @@ --- [betree_main]: external functions. -import Base -import BetreeMain.Types -open Primitives -open betree_main - --- TODO: fill those bodies - -/- [betree_main::betree_utils::load_internal_node] -/ -def betree_utils.load_internal_node - : - U64 → State → Result (State × (betree.List (U64 × betree.Message))) := - fun _ _ => .fail .panic - -/- [betree_main::betree_utils::store_internal_node] -/ -def betree_utils.store_internal_node - : - U64 → betree.List (U64 × betree.Message) → State → Result (State - × Unit) := - fun _ _ _ => .fail .panic - -/- [betree_main::betree_utils::load_leaf_node] -/ -def betree_utils.load_leaf_node - : U64 → State → Result (State × (betree.List (U64 × U64))) := - fun _ _ => .fail .panic - -/- [betree_main::betree_utils::store_leaf_node] -/ -def betree_utils.store_leaf_node - : U64 → betree.List (U64 × U64) → State → Result (State × Unit) := - fun _ _ _ => .fail .panic diff --git a/tests/lean/BetreeMain/FunsExternal_Template.lean b/tests/lean/BetreeMain/FunsExternal_Template.lean deleted file mode 100644 index 0dcce5ca..00000000 --- a/tests/lean/BetreeMain/FunsExternal_Template.lean +++ /dev/null @@ -1,30 +0,0 @@ --- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS --- [betree_main]: external functions. --- This is a template file: rename it to "FunsExternal.lean" and fill the holes. -import Base -import BetreeMain.Types -open Primitives -open betree_main - -/- [betree_main::betree_utils::load_internal_node]: - Source: 'src/betree_utils.rs', lines 98:0-98:63 -/ -axiom betree_utils.load_internal_node - : U64 → State → Result (State × (betree.List (U64 × betree.Message))) - -/- [betree_main::betree_utils::store_internal_node]: - Source: 'src/betree_utils.rs', lines 115:0-115:71 -/ -axiom betree_utils.store_internal_node - : - U64 → betree.List (U64 × betree.Message) → State → Result (State × - Unit) - -/- [betree_main::betree_utils::load_leaf_node]: - Source: 'src/betree_utils.rs', lines 132:0-132:55 -/ -axiom betree_utils.load_leaf_node - : U64 → State → Result (State × (betree.List (U64 × U64))) - -/- [betree_main::betree_utils::store_leaf_node]: - Source: 'src/betree_utils.rs', lines 145:0-145:63 -/ -axiom betree_utils.store_leaf_node - : U64 → betree.List (U64 × U64) → State → Result (State × Unit) - diff --git a/tests/lean/BetreeMain/Types.lean b/tests/lean/BetreeMain/Types.lean deleted file mode 100644 index e79da43f..00000000 --- a/tests/lean/BetreeMain/Types.lean +++ /dev/null @@ -1,83 +0,0 @@ --- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS --- [betree_main]: type definitions -import Base -import BetreeMain.TypesExternal -open Primitives - -namespace betree_main - -/- [betree_main::betree::List] - Source: 'src/betree.rs', lines 17:0-17:23 -/ -inductive betree.List (T : Type) := -| Cons : T → betree.List T → betree.List T -| Nil : betree.List T - -/- [betree_main::betree::UpsertFunState] - Source: 'src/betree.rs', lines 63:0-63:23 -/ -inductive betree.UpsertFunState := -| Add : U64 → betree.UpsertFunState -| Sub : U64 → betree.UpsertFunState - -/- [betree_main::betree::Message] - Source: 'src/betree.rs', lines 69:0-69:23 -/ -inductive betree.Message := -| Insert : U64 → betree.Message -| Delete : betree.Message -| Upsert : betree.UpsertFunState → betree.Message - -/- [betree_main::betree::Leaf] - Source: 'src/betree.rs', lines 167:0-167:11 -/ -structure betree.Leaf where - id : U64 - size : U64 - -mutual - -/- [betree_main::betree::Internal] - Source: 'src/betree.rs', lines 156:0-156:15 -/ -inductive betree.Internal := -| mk : U64 → U64 → betree.Node → betree.Node → betree.Internal - -/- [betree_main::betree::Node] - Source: 'src/betree.rs', lines 179:0-179:9 -/ -inductive betree.Node := -| Internal : betree.Internal → betree.Node -| Leaf : betree.Leaf → betree.Node - -end - -@[simp, reducible] -def betree.Internal.id (x : betree.Internal) := - match x with | betree.Internal.mk x1 _ _ _ => x1 - -@[simp, reducible] -def betree.Internal.pivot (x : betree.Internal) := - match x with | betree.Internal.mk _ x1 _ _ => x1 - -@[simp, reducible] -def betree.Internal.left (x : betree.Internal) := - match x with | betree.Internal.mk _ _ x1 _ => x1 - -@[simp, reducible] -def betree.Internal.right (x : betree.Internal) := - match x with | betree.Internal.mk _ _ _ x1 => x1 - -/- [betree_main::betree::Params] - Source: 'src/betree.rs', lines 187:0-187:13 -/ -structure betree.Params where - min_flush_size : U64 - split_size : U64 - -/- [betree_main::betree::NodeIdCounter] - Source: 'src/betree.rs', lines 201:0-201:20 -/ -structure betree.NodeIdCounter where - next_node_id : U64 - -/- [betree_main::betree::BeTree] - Source: 'src/betree.rs', lines 218:0-218:17 -/ -structure betree.BeTree where - params : betree.Params - node_id_cnt : betree.NodeIdCounter - root : betree.Node - -end betree_main diff --git a/tests/lean/BetreeMain/TypesExternal.lean b/tests/lean/BetreeMain/TypesExternal.lean deleted file mode 100644 index 1701eaaf..00000000 --- a/tests/lean/BetreeMain/TypesExternal.lean +++ /dev/null @@ -1,7 +0,0 @@ --- [betree_main]: external types. -import Base -open Primitives - -/- The state type used in the state-error monad -/ -axiom State : Type - diff --git a/tests/lean/BetreeMain/TypesExternal_Template.lean b/tests/lean/BetreeMain/TypesExternal_Template.lean deleted file mode 100644 index bbac7e99..00000000 --- a/tests/lean/BetreeMain/TypesExternal_Template.lean +++ /dev/null @@ -1,9 +0,0 @@ --- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS --- [betree_main]: external types. --- This is a template file: rename it to "TypesExternal.lean" and fill the holes. -import Base -open Primitives - -/- The state type used in the state-error monad -/ -axiom State : Type - diff --git a/tests/lean/lakefile.lean b/tests/lean/lakefile.lean index bb13ddf2..ba336a4a 100644 --- a/tests/lean/lakefile.lean +++ b/tests/lean/lakefile.lean @@ -9,7 +9,7 @@ require base from "../../backends/lean" package «tests» {} @[default_target] lean_lib Tutorial -@[default_target] lean_lib BetreeMain +@[default_target] lean_lib Betree @[default_target] lean_lib Constants @[default_target] lean_lib External @[default_target] lean_lib Hashmap diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index 1c885d4b..a7b83e88 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -122,8 +122,7 @@ let aeneas_options_for_test backend test_name = (* File-specific options *) let charon_options_for_test test_name = match test_name with - | "betree" -> - [ "--polonius"; "--opaque=betree_utils"; "--crate"; "betree_main" ] + | "betree" -> [ "--polonius"; "--opaque=betree_utils" ] | _ -> [] (* The data for a specific test input *) @@ -257,10 +256,8 @@ end (* Run Aeneas on a specific input with the given options *) let run_aeneas (env : runner_env) (case : Input.t) (backend : Backend.t) = - (* FIXME: remove this special case *) let backend_str = Backend.to_string backend in - let test_name = if case.name = "betree" then "betree_main" else case.name in - let input_file = concat_path [ env.llbc_dir; test_name ] ^ ".llbc" in + let input_file = concat_path [ env.llbc_dir; case.name ] ^ ".llbc" in let output_file = Filename.remove_extension case.path ^ "." ^ backend_str ^ ".out" in -- cgit v1.2.3 From 9cc69020773cc77965a6faa6f0d46f179de3d8b8 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Mon, 27 May 2024 17:22:07 +0200 Subject: runner: Store options for crate tests in a separate file --- tests/src/betree/aeneas-test-options | 4 ++++ tests/test_runner/run_test.ml | 40 +++++++++--------------------------- 2 files changed, 14 insertions(+), 30 deletions(-) create mode 100644 tests/src/betree/aeneas-test-options diff --git a/tests/src/betree/aeneas-test-options b/tests/src/betree/aeneas-test-options new file mode 100644 index 00000000..c71e01df --- /dev/null +++ b/tests/src/betree/aeneas-test-options @@ -0,0 +1,4 @@ +charon-args=--polonius --opaque=betree_utils +aeneas-args=-backward-no-state-update -test-trans-units -state -split-files +[coq] aeneas-args=-use-fuel +[fstar] aeneas-args=-decreases-clauses -template-clauses diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index a7b83e88..c626b4d1 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -100,31 +100,6 @@ module Command = struct | Failure -> () end -let aeneas_options_for_test backend test_name = - if test_name = "betree" then - let options = - [ - "-backward-no-state-update"; - "-test-trans-units"; - "-state"; - "-split-files"; - ] - in - let extra_options = - match backend with - | Backend.Coq -> [ "-use-fuel" ] - | Backend.FStar -> [ "-decreases-clauses"; "-template-clauses" ] - | _ -> [] - in - List.append extra_options options - else [] - -(* File-specific options *) -let charon_options_for_test test_name = - match test_name with - | "betree" -> [ "--polonius"; "--opaque=betree_utils" ] - | _ -> [] - (* The data for a specific test input *) module Input = struct type kind = SingleFile | Crate @@ -218,13 +193,10 @@ module Input = struct else failwith ("`" ^ path ^ "` is not a file or a directory.") in let actions = BackendMap.make (fun _ -> Normal) in - let charon_options = charon_options_for_test name in let subdirs = BackendMap.make (fun backend -> default_subdir backend name) in - let aeneas_options = - BackendMap.make (fun backend -> aeneas_options_for_test backend name) - in + let aeneas_options = BackendMap.make (fun _ -> []) in let check_output = BackendMap.make (fun _ -> true) in let input = { @@ -232,12 +204,14 @@ module Input = struct name; kind; actions; - charon_options; + charon_options = []; subdirs; aeneas_options; check_output; } in + (* For crates, we store the special comments in a separate file. *) + let crate_config_file = Filename.concat path "aeneas-test-options" in match kind with | SingleFile -> let file_lines = Core.In_channel.read_lines path in @@ -251,6 +225,11 @@ module Input = struct List.fold_left (fun input comment -> apply_special_comment comment input) input special_comments + | Crate when Sys.file_exists crate_config_file -> + let special_comments = Core.In_channel.read_lines crate_config_file in + List.fold_left + (fun input comment -> apply_special_comment comment input) + input special_comments | Crate -> input end @@ -300,6 +279,7 @@ let run_aeneas (env : runner_env) (case : Input.t) (backend : Backend.t) = let cmd = { cmd with redirect_out = Some out } in Command.run_command_expecting_failure cmd; Unix.close out + (* Run Charon on a specific input with the given options *) let run_charon (env : runner_env) (case : Input.t) = match case.kind with -- cgit v1.2.3 From c81c96f20b1dbf428a9ed42e83b910e798e1a225 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Mon, 27 May 2024 17:33:56 +0200 Subject: Add some tests --- tests/coq/misc/Issue194RecursiveStructProjector.v | 54 ++++++++++++++++++++++ .../misc/Issue194RecursiveStructProjector.fst | 24 ++++++++++ tests/fstar/misc/Matches.fst | 12 +++++ tests/lean/Issue194RecursiveStructProjector.lean | 35 ++++++++++++++ tests/lean/Matches.lean | 16 +++++++ tests/src/infinite-loop.rs | 11 +++++ tests/src/issue-194-recursive-struct-projector.rs | 16 +++++++ tests/src/matches.rs | 10 ++++ 8 files changed, 178 insertions(+) create mode 100644 tests/coq/misc/Issue194RecursiveStructProjector.v create mode 100644 tests/fstar/misc/Issue194RecursiveStructProjector.fst create mode 100644 tests/fstar/misc/Matches.fst create mode 100644 tests/lean/Issue194RecursiveStructProjector.lean create mode 100644 tests/lean/Matches.lean create mode 100644 tests/src/infinite-loop.rs create mode 100644 tests/src/issue-194-recursive-struct-projector.rs create mode 100644 tests/src/matches.rs diff --git a/tests/coq/misc/Issue194RecursiveStructProjector.v b/tests/coq/misc/Issue194RecursiveStructProjector.v new file mode 100644 index 00000000..d1a07449 --- /dev/null +++ b/tests/coq/misc/Issue194RecursiveStructProjector.v @@ -0,0 +1,54 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [issue_194_recursive_struct_projector] *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Module Issue194RecursiveStructProjector. + +(** [issue_194_recursive_struct_projector::AVLNode] + Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 2:0-2:17 *) +Inductive AVLNode_t (T : Type) := +| mkAVLNode_t : + T -> + option (AVLNode_t T) -> + option (AVLNode_t T) -> + AVLNode_t T +. + +Arguments mkAVLNode_t { _ }. + +Definition aVLNode_value {T : Type} (x : AVLNode_t T) := + match x with | mkAVLNode_t x1 _ _ => x1 end +. + +Notation "x2 .(aVLNode_value)" := (aVLNode_value x2) (at level 9). + +Definition aVLNode_left {T : Type} (x : AVLNode_t T) := + match x with | mkAVLNode_t _ x1 _ => x1 end +. + +Notation "x2 .(aVLNode_left)" := (aVLNode_left x2) (at level 9). + +Definition aVLNode_right {T : Type} (x : AVLNode_t T) := + match x with | mkAVLNode_t _ _ x1 => x1 end +. + +Notation "x2 .(aVLNode_right)" := (aVLNode_right x2) (at level 9). + +(** [issue_194_recursive_struct_projector::get_val]: + Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 10:0-10:33 *) +Definition get_val (T : Type) (x : AVLNode_t T) : result T := + Ok x.(aVLNode_value) +. + +(** [issue_194_recursive_struct_projector::get_left]: + Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 14:0-14:43 *) +Definition get_left + (T : Type) (x : AVLNode_t T) : result (option (AVLNode_t T)) := + Ok x.(aVLNode_left) +. + +End Issue194RecursiveStructProjector. diff --git a/tests/fstar/misc/Issue194RecursiveStructProjector.fst b/tests/fstar/misc/Issue194RecursiveStructProjector.fst new file mode 100644 index 00000000..76368f04 --- /dev/null +++ b/tests/fstar/misc/Issue194RecursiveStructProjector.fst @@ -0,0 +1,24 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [issue_194_recursive_struct_projector] *) +module Issue194RecursiveStructProjector +open Primitives + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** [issue_194_recursive_struct_projector::AVLNode] + Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 2:0-2:17 *) +type aVLNode_t (t : Type0) = +{ + value : t; left : option (aVLNode_t t); right : option (aVLNode_t t); +} + +(** [issue_194_recursive_struct_projector::get_val]: + Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 10:0-10:33 *) +let get_val (t : Type0) (x : aVLNode_t t) : result t = + Ok x.value + +(** [issue_194_recursive_struct_projector::get_left]: + Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 14:0-14:43 *) +let get_left (t : Type0) (x : aVLNode_t t) : result (option (aVLNode_t t)) = + Ok x.left + diff --git a/tests/fstar/misc/Matches.fst b/tests/fstar/misc/Matches.fst new file mode 100644 index 00000000..06a9b6df --- /dev/null +++ b/tests/fstar/misc/Matches.fst @@ -0,0 +1,12 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [matches] *) +module Matches +open Primitives + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** [matches::match_u32]: + Source: 'tests/src/matches.rs', lines 4:0-4:27 *) +let match_u32 (x : u32) : result u32 = + begin match x with | 0 -> Ok 0 | 1 -> Ok 1 | _ -> Ok 2 end + diff --git a/tests/lean/Issue194RecursiveStructProjector.lean b/tests/lean/Issue194RecursiveStructProjector.lean new file mode 100644 index 00000000..4eb23934 --- /dev/null +++ b/tests/lean/Issue194RecursiveStructProjector.lean @@ -0,0 +1,35 @@ +-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS +-- [issue_194_recursive_struct_projector] +import Base +open Primitives + +namespace issue_194_recursive_struct_projector + +/- [issue_194_recursive_struct_projector::AVLNode] + Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 2:0-2:17 -/ +inductive AVLNode (T : Type) := +| mk : T → Option (AVLNode T) → Option (AVLNode T) → AVLNode T + +@[simp, reducible] +def AVLNode.value {T : Type} (x : AVLNode T) := + match x with | AVLNode.mk x1 _ _ => x1 + +@[simp, reducible] +def AVLNode.left {T : Type} (x : AVLNode T) := + match x with | AVLNode.mk _ x1 _ => x1 + +@[simp, reducible] +def AVLNode.right {T : Type} (x : AVLNode T) := + match x with | AVLNode.mk _ _ x1 => x1 + +/- [issue_194_recursive_struct_projector::get_val]: + Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 10:0-10:33 -/ +def get_val (T : Type) (x : AVLNode T) : Result T := + Result.ok x.value + +/- [issue_194_recursive_struct_projector::get_left]: + Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 14:0-14:43 -/ +def get_left (T : Type) (x : AVLNode T) : Result (Option (AVLNode T)) := + Result.ok x.left + +end issue_194_recursive_struct_projector diff --git a/tests/lean/Matches.lean b/tests/lean/Matches.lean new file mode 100644 index 00000000..3e3a558b --- /dev/null +++ b/tests/lean/Matches.lean @@ -0,0 +1,16 @@ +-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS +-- [matches] +import Base +open Primitives + +namespace matches + +/- [matches::match_u32]: + Source: 'tests/src/matches.rs', lines 4:0-4:27 -/ +def match_u32 (x : U32) : Result U32 := + match x with + | 0#u32 => Result.ok 0#u32 + | 1#u32 => Result.ok 1#u32 + | _ => Result.ok 2#u32 + +end matches diff --git a/tests/src/infinite-loop.rs b/tests/src/infinite-loop.rs new file mode 100644 index 00000000..906fc1fa --- /dev/null +++ b/tests/src/infinite-loop.rs @@ -0,0 +1,11 @@ +//@ [coq,fstar,lean] skip +//@ [coq,fstar] aeneas-args=-use-fuel +//@ [coq,fstar] subdir=misc +// FIXME: make it work +fn bar() {} + +fn foo() { + loop { + bar() + } +} diff --git a/tests/src/issue-194-recursive-struct-projector.rs b/tests/src/issue-194-recursive-struct-projector.rs new file mode 100644 index 00000000..9ebdc2bc --- /dev/null +++ b/tests/src/issue-194-recursive-struct-projector.rs @@ -0,0 +1,16 @@ +//@ [coq,fstar] subdir=misc +struct AVLNode { + value: T, + left: AVLTree, + right: AVLTree, +} + +type AVLTree = Option>>; + +fn get_val(x: AVLNode) -> T { + x.value +} + +fn get_left(x: AVLNode) -> AVLTree { + x.left +} diff --git a/tests/src/matches.rs b/tests/src/matches.rs new file mode 100644 index 00000000..5710a604 --- /dev/null +++ b/tests/src/matches.rs @@ -0,0 +1,10 @@ +//@ [coq] skip +//@ [coq,fstar] subdir=misc +//^ note: coq gives "invalid notation for pattern" +fn match_u32(x: u32) -> u32 { + match x { + 0 => 0, + 1 => 1, + _ => 2, + } +} -- cgit v1.2.3 From c4d2af051c18c4c81b1e57a45210c37c89c8330f Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Tue, 28 May 2024 11:18:35 +0200 Subject: tests: Rename hashmap_utils -> utils --- tests/coq/hashmap/Hashmap_Funs.v | 4 ++-- tests/coq/hashmap/Hashmap_FunsExternal.v | 8 ++++---- tests/coq/hashmap/Hashmap_FunsExternal_Template.v | 9 ++++----- tests/fstar/hashmap/Hashmap.Funs.fst | 4 ++-- tests/fstar/hashmap/Hashmap.FunsExternal.fsti | 8 ++++---- tests/fstar/hashmap/Hashmap.Properties.fst | 8 ++++---- tests/hol4/hashmap_main/hashmapMain_FunsScript.sml | 4 ++-- tests/hol4/hashmap_main/hashmapMain_FunsTheory.sig | 4 ++-- tests/hol4/hashmap_main/hashmapMain_OpaqueScript.sml | 4 ++-- tests/lean/Hashmap/Funs.lean | 4 ++-- tests/lean/Hashmap/FunsExternal.lean | 8 ++++---- tests/lean/Hashmap/FunsExternal_Template.lean | 9 ++++----- tests/src/hashmap.rs | 8 ++++---- 13 files changed, 40 insertions(+), 42 deletions(-) diff --git a/tests/coq/hashmap/Hashmap_Funs.v b/tests/coq/hashmap/Hashmap_Funs.v index 89f90210..04df873a 100644 --- a/tests/coq/hashmap/Hashmap_Funs.v +++ b/tests/coq/hashmap/Hashmap_Funs.v @@ -508,10 +508,10 @@ Definition hashMap_remove Source: 'tests/src/hashmap.rs', lines 335:0-335:43 *) Definition insert_on_disk (n : nat) (key : usize) (value : u64) (st : state) : result (state * unit) := - p <- hashmap_utils_deserialize st; + p <- utils_deserialize st; let (st1, hm) := p in hm1 <- hashMap_insert u64 n hm key value; - hashmap_utils_serialize hm1 st1 + utils_serialize hm1 st1 . (** [hashmap::test1]: diff --git a/tests/coq/hashmap/Hashmap_FunsExternal.v b/tests/coq/hashmap/Hashmap_FunsExternal.v index bcf12c17..3bc6c98a 100644 --- a/tests/coq/hashmap/Hashmap_FunsExternal.v +++ b/tests/coq/hashmap/Hashmap_FunsExternal.v @@ -9,13 +9,13 @@ Require Export Hashmap_Types. Import Hashmap_Types. Module Hashmap_FunsExternal. -(** [hashmap::hashmap_utils::deserialize]: +(** [hashmap::utils::deserialize]: Source: 'tests/src/hashmap.rs', lines 330:4-330:47 *) -Axiom hashmap_utils_deserialize : state -> result (state * (HashMap_t u64)). +Axiom utils_deserialize : state -> result (state * (HashMap_t u64)). -(** [hashmap::hashmap_utils::serialize]: +(** [hashmap::utils::serialize]: Source: 'tests/src/hashmap.rs', lines 325:4-325:46 *) -Axiom hashmap_utils_serialize : HashMap_t u64 -> state -> result (state * unit) +Axiom utils_serialize : HashMap_t u64 -> state -> result (state * unit) . End Hashmap_FunsExternal. diff --git a/tests/coq/hashmap/Hashmap_FunsExternal_Template.v b/tests/coq/hashmap/Hashmap_FunsExternal_Template.v index c35b8e14..c58b021d 100644 --- a/tests/coq/hashmap/Hashmap_FunsExternal_Template.v +++ b/tests/coq/hashmap/Hashmap_FunsExternal_Template.v @@ -11,13 +11,12 @@ Require Import Hashmap_Types. Include Hashmap_Types. Module Hashmap_FunsExternal_Template. -(** [hashmap::hashmap_utils::deserialize]: +(** [hashmap::utils::deserialize]: Source: 'tests/src/hashmap.rs', lines 330:4-330:47 *) -Axiom hashmap_utils_deserialize : state -> result (state * (HashMap_t u64)). +Axiom utils_deserialize : state -> result (state * (HashMap_t u64)). -(** [hashmap::hashmap_utils::serialize]: +(** [hashmap::utils::serialize]: Source: 'tests/src/hashmap.rs', lines 325:4-325:46 *) -Axiom hashmap_utils_serialize : HashMap_t u64 -> state -> result (state * unit) -. +Axiom utils_serialize : HashMap_t u64 -> state -> result (state * unit). End Hashmap_FunsExternal_Template. diff --git a/tests/fstar/hashmap/Hashmap.Funs.fst b/tests/fstar/hashmap/Hashmap.Funs.fst index a85be4ed..0e991720 100644 --- a/tests/fstar/hashmap/Hashmap.Funs.fst +++ b/tests/fstar/hashmap/Hashmap.Funs.fst @@ -380,9 +380,9 @@ let hashMap_remove Source: 'tests/src/hashmap.rs', lines 335:0-335:43 *) let insert_on_disk (key : usize) (value : u64) (st : state) : result (state & unit) = - let* (st1, hm) = hashmap_utils_deserialize st in + let* (st1, hm) = utils_deserialize st in let* hm1 = hashMap_insert u64 hm key value in - hashmap_utils_serialize hm1 st1 + utils_serialize hm1 st1 (** [hashmap::test1]: Source: 'tests/src/hashmap.rs', lines 350:0-350:10 *) diff --git a/tests/fstar/hashmap/Hashmap.FunsExternal.fsti b/tests/fstar/hashmap/Hashmap.FunsExternal.fsti index 6e3964c7..f2f305e6 100644 --- a/tests/fstar/hashmap/Hashmap.FunsExternal.fsti +++ b/tests/fstar/hashmap/Hashmap.FunsExternal.fsti @@ -6,11 +6,11 @@ include Hashmap.Types #set-options "--z3rlimit 50 --fuel 1 --ifuel 1" -(** [hashmap::hashmap_utils::deserialize]: +(** [hashmap::utils::deserialize]: Source: 'tests/src/hashmap.rs', lines 330:4-330:47 *) -val hashmap_utils_deserialize : state -> result (state & (hashMap_t u64)) +val utils_deserialize : state -> result (state & (hashMap_t u64)) -(** [hashmap::hashmap_utils::serialize]: +(** [hashmap::utils::serialize]: Source: 'tests/src/hashmap.rs', lines 325:4-325:46 *) -val hashmap_utils_serialize : hashMap_t u64 -> state -> result (state & unit) +val utils_serialize : hashMap_t u64 -> state -> result (state & unit) diff --git a/tests/fstar/hashmap/Hashmap.Properties.fst b/tests/fstar/hashmap/Hashmap.Properties.fst index 0d6372c1..ed118c46 100644 --- a/tests/fstar/hashmap/Hashmap.Properties.fst +++ b/tests/fstar/hashmap/Hashmap.Properties.fst @@ -18,18 +18,18 @@ val state_v : state -> hashMap_t u64 /// [serialize] updates the hash map stored on disk assume val serialize_lem (hm : hashMap_t u64) (st : state) : Lemma ( - match hashmap_utils_serialize hm st with + match utils_serialize hm st with | Fail _ -> True | Ok (st', ()) -> state_v st' == hm) - [SMTPat (hashmap_utils_serialize hm st)] + [SMTPat (utils_serialize hm st)] /// [deserialize] gives us the hash map stored on disk, without updating it assume val deserialize_lem (st : state) : Lemma ( - match hashmap_utils_deserialize st with + match utils_deserialize st with | Fail _ -> True | Ok (st', hm) -> hm == state_v st /\ st' == st) - [SMTPat (hashmap_utils_deserialize st)] + [SMTPat (utils_deserialize st)] (*** Lemmas *) diff --git a/tests/hol4/hashmap_main/hashmapMain_FunsScript.sml b/tests/hol4/hashmap_main/hashmapMain_FunsScript.sml index c1e30aa6..2a17d185 100644 --- a/tests/hol4/hashmap_main/hashmapMain_FunsScript.sml +++ b/tests/hol4/hashmap_main/hashmapMain_FunsScript.sml @@ -628,9 +628,9 @@ val insert_on_disk_fwd_def = Define ‘ insert_on_disk_fwd (key : usize) (value : u64) (st : state) : (state # unit) result = do - (st0, hm) <- hashmap_utils_deserialize_fwd st; + (st0, hm) <- utils_deserialize_fwd st; hm0 <- hashmap_hash_map_insert_fwd_back hm key value; - (st1, _) <- hashmap_utils_serialize_fwd hm0 st0; + (st1, _) <- utils_serialize_fwd hm0 st0; Return (st1, ()) od ’ diff --git a/tests/hol4/hashmap_main/hashmapMain_FunsTheory.sig b/tests/hol4/hashmap_main/hashmapMain_FunsTheory.sig index d4e43d9a..c19673bb 100644 --- a/tests/hol4/hashmap_main/hashmapMain_FunsTheory.sig +++ b/tests/hol4/hashmap_main/hashmapMain_FunsTheory.sig @@ -583,9 +583,9 @@ sig ⊢ ∀key value st. insert_on_disk_fwd key value st = do - (st0,hm) <- hashmap_utils_deserialize_fwd st; + (st0,hm) <- utils_deserialize_fwd st; hm0 <- hashmap_hash_map_insert_fwd_back hm key value; - (st1,_) <- hashmap_utils_serialize_fwd hm0 st0; + (st1,_) <- utils_serialize_fwd hm0 st0; Return (st1,()) od diff --git a/tests/hol4/hashmap_main/hashmapMain_OpaqueScript.sml b/tests/hol4/hashmap_main/hashmapMain_OpaqueScript.sml index f7221d92..9c29c0e0 100644 --- a/tests/hol4/hashmap_main/hashmapMain_OpaqueScript.sml +++ b/tests/hol4/hashmap_main/hashmapMain_OpaqueScript.sml @@ -6,10 +6,10 @@ open hashmapMain_TypesTheory val _ = new_theory "hashmapMain_Opaque" -(** [hashmap_main::hashmap_utils::deserialize]: forward function *)val _ = new_constant ("hashmap_utils_deserialize_fwd", +(** [hashmap_main::utils::deserialize]: forward function *)val _ = new_constant ("utils_deserialize_fwd", “:state -> (state # u64 hashmap_hash_map_t) result”) -(** [hashmap_main::hashmap_utils::serialize]: forward function *)val _ = new_constant ("hashmap_utils_serialize_fwd", +(** [hashmap_main::utils::serialize]: forward function *)val _ = new_constant ("utils_serialize_fwd", “:u64 hashmap_hash_map_t -> state -> (state # unit) result”) val _ = export_theory () diff --git a/tests/lean/Hashmap/Funs.lean b/tests/lean/Hashmap/Funs.lean index e76912cc..d7ac3b05 100644 --- a/tests/lean/Hashmap/Funs.lean +++ b/tests/lean/Hashmap/Funs.lean @@ -386,9 +386,9 @@ def HashMap.remove def insert_on_disk (key : Usize) (value : U64) (st : State) : Result (State × Unit) := do - let (st1, hm) ← hashmap_utils.deserialize st + let (st1, hm) ← utils.deserialize st let hm1 ← HashMap.insert U64 hm key value - hashmap_utils.serialize hm1 st1 + utils.serialize hm1 st1 /- [hashmap::test1]: Source: 'tests/src/hashmap.rs', lines 350:0-350:10 -/ diff --git a/tests/lean/Hashmap/FunsExternal.lean b/tests/lean/Hashmap/FunsExternal.lean index 7454696b..329e5d82 100644 --- a/tests/lean/Hashmap/FunsExternal.lean +++ b/tests/lean/Hashmap/FunsExternal.lean @@ -6,14 +6,14 @@ open hashmap -- TODO: fill those bodies -/- [hashmap::hashmap_utils::deserialize]: +/- [hashmap::utils::deserialize]: Source: 'tests/src/hashmap.rs', lines 330:4-330:47 -/ -def hashmap_utils.deserialize +def utils.deserialize : State → Result (State × (HashMap U64)) := fun _ => .fail .panic -/- [hashmap::hashmap_utils::serialize]: +/- [hashmap::utils::serialize]: Source: 'tests/src/hashmap.rs', lines 325:4-325:46 -/ -def hashmap_utils.serialize +def utils.serialize : HashMap U64 → State → Result (State × Unit) := fun _ _ => .fail .panic diff --git a/tests/lean/Hashmap/FunsExternal_Template.lean b/tests/lean/Hashmap/FunsExternal_Template.lean index 16b4b6af..2af57d10 100644 --- a/tests/lean/Hashmap/FunsExternal_Template.lean +++ b/tests/lean/Hashmap/FunsExternal_Template.lean @@ -6,12 +6,11 @@ import Hashmap.Types open Primitives open hashmap -/- [hashmap::hashmap_utils::deserialize]: +/- [hashmap::utils::deserialize]: Source: 'tests/src/hashmap.rs', lines 330:4-330:47 -/ -axiom hashmap_utils.deserialize : State → Result (State × (HashMap U64)) +axiom utils.deserialize : State → Result (State × (HashMap U64)) -/- [hashmap::hashmap_utils::serialize]: +/- [hashmap::utils::serialize]: Source: 'tests/src/hashmap.rs', lines 325:4-325:46 -/ -axiom hashmap_utils.serialize - : HashMap U64 → State → Result (State × Unit) +axiom utils.serialize : HashMap U64 → State → Result (State × Unit) diff --git a/tests/src/hashmap.rs b/tests/src/hashmap.rs index ccb96e1e..19832a84 100644 --- a/tests/src/hashmap.rs +++ b/tests/src/hashmap.rs @@ -1,4 +1,4 @@ -//@ charon-args=--opaque=hashmap_utils +//@ charon-args=--opaque=utils //@ aeneas-args=-state -split-files //@ [coq] aeneas-args=-use-fuel //@ [fstar] aeneas-args=-decreases-clauses -template-clauses @@ -317,7 +317,7 @@ impl HashMap { } // This is a module so we can tell charon to leave it opaque -mod hashmap_utils { +mod utils { use crate::*; /// Serialize a hash map - we don't have traits, so we fix the type of the @@ -334,11 +334,11 @@ mod hashmap_utils { pub fn insert_on_disk(key: Key, value: u64) { // Deserialize - let mut hm = hashmap_utils::deserialize(); + let mut hm = utils::deserialize(); // Update hm.insert(key, value); // Serialize - hashmap_utils::serialize(hm); + utils::serialize(hm); } /// I currently can't retrieve functions marked with the attribute #[test], -- cgit v1.2.3 From 2b5186fef6932995626790af241eaed4e7cc02ae Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 28 May 2024 18:04:38 +0200 Subject: Fix a bug in SymbolicToPure.translate_loop --- compiler/SymbolicToPure.ml | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/compiler/SymbolicToPure.ml b/compiler/SymbolicToPure.ml index d6d2e018..cc203040 100644 --- a/compiler/SymbolicToPure.ml +++ b/compiler/SymbolicToPure.ml @@ -3529,10 +3529,7 @@ and translate_loop (loop : S.loop) (ctx : bs_ctx) : texpression = if effect_info.stateful then mk_simpl_tuple_ty [ mk_state_ty; output ] else output in - let output = - if effect_info.can_fail && inputs <> [] then mk_result_ty output - else output - in + let output = if effect_info.can_fail then mk_result_ty output else output in (back_info, output) in -- cgit v1.2.3 From ae075db15638ee8878bebe3d31affb1aa320e90f Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 28 May 2024 18:05:01 +0200 Subject: Reactivate the infinite-loop.rs test --- tests/coq/misc/InfiniteLoop.v | 27 +++++++++++++++++++++++++++ tests/fstar/misc/InfiniteLoop.fst | 24 ++++++++++++++++++++++++ tests/lean/InfiniteLoop.lean | 25 +++++++++++++++++++++++++ tests/src/infinite-loop.rs | 1 - 4 files changed, 76 insertions(+), 1 deletion(-) create mode 100644 tests/coq/misc/InfiniteLoop.v create mode 100644 tests/fstar/misc/InfiniteLoop.fst create mode 100644 tests/lean/InfiniteLoop.lean diff --git a/tests/coq/misc/InfiniteLoop.v b/tests/coq/misc/InfiniteLoop.v new file mode 100644 index 00000000..1c0ccbe0 --- /dev/null +++ b/tests/coq/misc/InfiniteLoop.v @@ -0,0 +1,27 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [infinite_loop] *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Require Import List. +Import ListNotations. +Local Open Scope Primitives_scope. +Module InfiniteLoop. + +(** [infinite_loop::bar]: + Source: 'tests/src/infinite-loop.rs', lines 4:0-4:8 *) +Definition bar : result unit := + Ok tt. + +(** [infinite_loop::foo]: loop 0: + Source: 'tests/src/infinite-loop.rs', lines 6:0-10:1 *) +Fixpoint foo_loop (n : nat) : result unit := + match n with | O => Fail_ OutOfFuel | S n1 => _ <- bar; foo_loop n1 end +. + +(** [infinite_loop::foo]: + Source: 'tests/src/infinite-loop.rs', lines 6:0-6:8 *) +Definition foo (n : nat) : result unit := + foo_loop n. + +End InfiniteLoop. diff --git a/tests/fstar/misc/InfiniteLoop.fst b/tests/fstar/misc/InfiniteLoop.fst new file mode 100644 index 00000000..4e7c8c06 --- /dev/null +++ b/tests/fstar/misc/InfiniteLoop.fst @@ -0,0 +1,24 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [infinite_loop] *) +module InfiniteLoop +open Primitives + +#set-options "--z3rlimit 50 --fuel 1 --ifuel 1" + +(** [infinite_loop::bar]: + Source: 'tests/src/infinite-loop.rs', lines 4:0-4:8 *) +let bar : result unit = + Ok () + +(** [infinite_loop::foo]: loop 0: + Source: 'tests/src/infinite-loop.rs', lines 6:0-10:1 *) +let rec foo_loop (n : nat) : result unit = + if is_zero n + then Fail OutOfFuel + else let n1 = decrease n in let* _ = bar in foo_loop n1 + +(** [infinite_loop::foo]: + Source: 'tests/src/infinite-loop.rs', lines 6:0-6:8 *) +let foo (n : nat) : result unit = + foo_loop n + diff --git a/tests/lean/InfiniteLoop.lean b/tests/lean/InfiniteLoop.lean new file mode 100644 index 00000000..9eb8e22c --- /dev/null +++ b/tests/lean/InfiniteLoop.lean @@ -0,0 +1,25 @@ +-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS +-- [infinite_loop] +import Base +open Primitives + +namespace infinite_loop + +/- [infinite_loop::bar]: + Source: 'tests/src/infinite-loop.rs', lines 4:0-4:8 -/ +def bar : Result Unit := + Result.ok () + +/- [infinite_loop::foo]: loop 0: + Source: 'tests/src/infinite-loop.rs', lines 6:0-10:1 -/ +divergent def foo_loop : Result Unit := + do + let _ ← bar + foo_loop + +/- [infinite_loop::foo]: + Source: 'tests/src/infinite-loop.rs', lines 6:0-6:8 -/ +def foo : Result Unit := + foo_loop + +end infinite_loop diff --git a/tests/src/infinite-loop.rs b/tests/src/infinite-loop.rs index 906fc1fa..0e247d20 100644 --- a/tests/src/infinite-loop.rs +++ b/tests/src/infinite-loop.rs @@ -1,4 +1,3 @@ -//@ [coq,fstar,lean] skip //@ [coq,fstar] aeneas-args=-use-fuel //@ [coq,fstar] subdir=misc // FIXME: make it work -- cgit v1.2.3 From b5eac0384818e1f51fbfd900ab580514e851b0ca Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 28 May 2024 18:48:56 +0200 Subject: Fix an issue with some names being ignored when generating unique variable names --- compiler/ExtractBase.ml | 40 ++++++++++++++++++-------------------- tests/coq/arrays/Arrays.v | 4 ++-- tests/fstar/arrays/Arrays.Funs.fst | 4 ++-- 3 files changed, 23 insertions(+), 25 deletions(-) diff --git a/compiler/ExtractBase.ml b/compiler/ExtractBase.ml index ab7eb50c..815e228f 100644 --- a/compiler/ExtractBase.ml +++ b/compiler/ExtractBase.ml @@ -351,13 +351,13 @@ let unsafe_names_map_add (id : id) (name : string) (nm : unsafe_names_map) : [append]: function to append an index to a string *) -let basename_to_unique (names_set : StringSet.t) +let basename_to_unique_aux (collision : string -> bool) (append : string -> int -> string) (basename : string) : string = let rec gen (i : int) : string = let s = append basename i in - if StringSet.mem s names_set then gen (i + 1) else s + if collision s then gen (i + 1) else s in - if StringSet.mem basename names_set then gen 1 else basename + if collision basename then gen 1 else basename type names_maps = { names_map : names_map; @@ -1841,13 +1841,22 @@ let trait_self_clause_basename = "self_clause" let name_append_index (basename : string) (i : int) : string = basename ^ string_of_int i +let basename_to_unique (ctx : extraction_ctx) (name : string) = + let collision s = + (* Note that we ignore the "unsafe" names which contain in particular + field names: we want to allow using field names for variables if + the backend allows such collisions *) + StringSet.mem s ctx.names_maps.names_map.names_set + || StringSet.mem s ctx.names_maps.strict_names_map.names_set + in + + basename_to_unique_aux collision name_append_index name + (** Generate a unique type variable name and add it to the context *) let ctx_add_type_var (span : Meta.span) (basename : string) (id : TypeVarId.id) (ctx : extraction_ctx) : extraction_ctx * string = let name = ctx_compute_type_var_basename ctx basename in - let name = - basename_to_unique ctx.names_maps.names_map.names_set name_append_index name - in + let name = basename_to_unique ctx name in let ctx = ctx_add span (TypeVarId id) name ctx in (ctx, name) @@ -1856,9 +1865,7 @@ let ctx_add_const_generic_var (span : Meta.span) (basename : string) (id : ConstGenericVarId.id) (ctx : extraction_ctx) : extraction_ctx * string = let name = ctx_compute_const_generic_var_basename ctx basename in - let name = - basename_to_unique ctx.names_maps.names_map.names_set name_append_index name - in + let name = basename_to_unique ctx name in let ctx = ctx_add span (ConstGenericVarId id) name ctx in (ctx, name) @@ -1872,10 +1879,7 @@ let ctx_add_type_vars (span : Meta.span) (vars : (string * TypeVarId.id) list) (** Generate a unique variable name and add it to the context *) let ctx_add_var (span : Meta.span) (basename : string) (id : VarId.id) (ctx : extraction_ctx) : extraction_ctx * string = - let name = - basename_to_unique ctx.names_maps.names_map.names_set name_append_index - basename - in + let name = basename_to_unique ctx basename in let ctx = ctx_add span (VarId id) name ctx in (ctx, name) @@ -1883,20 +1887,14 @@ let ctx_add_var (span : Meta.span) (basename : string) (id : VarId.id) let ctx_add_trait_self_clause (span : Meta.span) (ctx : extraction_ctx) : extraction_ctx * string = let basename = trait_self_clause_basename in - let name = - basename_to_unique ctx.names_maps.names_map.names_set name_append_index - basename - in + let name = basename_to_unique ctx basename in let ctx = ctx_add span TraitSelfClauseId name ctx in (ctx, name) (** Generate a unique trait clause name and add it to the context *) let ctx_add_local_trait_clause (span : Meta.span) (basename : string) (id : TraitClauseId.id) (ctx : extraction_ctx) : extraction_ctx * string = - let name = - basename_to_unique ctx.names_maps.names_map.names_set name_append_index - basename - in + let name = basename_to_unique ctx basename in let ctx = ctx_add span (LocalTraitClauseId id) name ctx in (ctx, name) diff --git a/tests/coq/arrays/Arrays.v b/tests/coq/arrays/Arrays.v index 371e4a12..6e4ab188 100644 --- a/tests/coq/arrays/Arrays.v +++ b/tests/coq/arrays/Arrays.v @@ -289,9 +289,9 @@ Definition index_all : result u32 := let (s1, to_slice_mut_back) := p in p1 <- index_mut_slice_u32_0 s1; let (i7, s2) := p1 in - i8 <- u32_add i6 i7; + i9 <- u32_add i6 i7; _ <- to_slice_mut_back s2; - Ok i8 + Ok i9 . (** [arrays::update_array]: diff --git a/tests/fstar/arrays/Arrays.Funs.fst b/tests/fstar/arrays/Arrays.Funs.fst index 289e603d..26a695bb 100644 --- a/tests/fstar/arrays/Arrays.Funs.fst +++ b/tests/fstar/arrays/Arrays.Funs.fst @@ -241,9 +241,9 @@ let index_all : result u32 = let* (s1, to_slice_mut_back) = array_to_slice_mut u32 2 (mk_array u32 2 [ 0; 0 ]) in let* (i7, s2) = index_mut_slice_u32_0 s1 in - let* i8 = u32_add i6 i7 in + let* i9 = u32_add i6 i7 in let* _ = to_slice_mut_back s2 in - Ok i8 + Ok i9 (** [arrays::update_array]: Source: 'tests/src/arrays.rs', lines 187:0-187:36 *) -- cgit v1.2.3 From 5330a6f728add1525854ad4cedd86b8e1d3a1fa7 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 29 May 2024 01:00:38 +0200 Subject: Factor out code in ExtractBuiltin.ml --- compiler/ExtractBuiltin.ml | 284 +++++++++++++++++++-------------------------- compiler/FunsAnalysis.ml | 3 +- 2 files changed, 124 insertions(+), 163 deletions(-) diff --git a/compiler/ExtractBuiltin.ml b/compiler/ExtractBuiltin.ml index a7ab6da0..ff936d2f 100644 --- a/compiler/ExtractBuiltin.ml +++ b/compiler/ExtractBuiltin.ml @@ -233,7 +233,12 @@ let mk_builtin_types_map () = let builtin_types_map = mk_memoized mk_builtin_types_map -type builtin_fun_info = { extract_name : string } [@@deriving show] +type builtin_fun_info = { + extract_name : string; + can_fail : bool; + stateful : bool; +} +[@@deriving show] let int_and_smaller_list : (string * string) list = let uint_names = List.rev [ "u8"; "u16"; "u32"; "u64"; "u128" ] in @@ -257,16 +262,17 @@ let int_and_smaller_list : (string * string) list = ] @ compute_pairs uint_names @ compute_pairs int_names -(** The assumed functions. +(** The builtin functions. The optional list of booleans is filtering information for the type parameters. For instance, in the case of the `Vec` functions, there is a type parameter for the allocator to use, which we want to filter. *) -let builtin_funs () : (pattern * bool list option * builtin_fun_info) list = +let mk_builtin_funs () : (pattern * bool list option * builtin_fun_info) list = (* Small utility *) - let mk_fun (rust_name : string) (extract_name : string option) - (filter : bool list option) : + let mk_fun (rust_name : string) ?(filter : bool list option = None) + ?(can_fail = true) ?(stateful = false) + ?(extract_name : string option = None) () : pattern * bool list option * builtin_fun_info = let rust_name = try parse_pattern rust_name @@ -279,124 +285,151 @@ let builtin_funs () : (pattern * bool list option * builtin_fun_info) list = | Some name -> split_on_separator name in let basename = flatten_name extract_name in - let f = { extract_name = basename } in + let f = { extract_name = basename; can_fail; stateful } in (rust_name, filter, f) in let mk_scalar_fun (rust_name : string -> string) - (extract_name : string -> string) : + (extract_name : string -> string) ?(can_fail = true) () : (pattern * bool list option * builtin_fun_info) list = List.map - (fun ty -> mk_fun (rust_name ty) (Some (extract_name ty)) None) + (fun ty -> + mk_fun (rust_name ty) + ~extract_name:(Some (extract_name ty)) + ~can_fail ()) all_int_names in [ - mk_fun "core::mem::replace" None None; + mk_fun "core::mem::replace" ~can_fail:false (); + mk_fun "core::mem::take" ~can_fail:false (); mk_fun "core::slice::{[@T]}::len" - (Some (backend_choice "slice::len" "Slice::len")) - None; + ~extract_name:(Some (backend_choice "slice::len" "Slice::len")) + ~can_fail:false (); mk_fun "alloc::vec::{alloc::vec::Vec<@T, alloc::alloc::Global>}::new" - (Some "alloc::vec::Vec::new") None; - mk_fun "alloc::vec::{alloc::vec::Vec<@T, @A>}::push" None - (Some [ true; false ]); - mk_fun "alloc::vec::{alloc::vec::Vec<@T, @A>}::insert" None - (Some [ true; false ]); - mk_fun "alloc::vec::{alloc::vec::Vec<@T, @A>}::len" None - (Some [ true; false ]); + ~extract_name:(Some "alloc::vec::Vec::new") ~can_fail:false (); + mk_fun "alloc::vec::{alloc::vec::Vec<@T, @A>}::push" + ~filter:(Some [ true; false ]) + (); + mk_fun "alloc::vec::{alloc::vec::Vec<@T, @A>}::insert" + ~filter:(Some [ true; false ]) + (); + mk_fun "alloc::vec::{alloc::vec::Vec<@T, @A>}::len" + ~filter:(Some [ true; false ]) + ~can_fail:false (); mk_fun "alloc::vec::{core::ops::index::Index, \ @I>}::index" - (Some "alloc.vec.Vec.index") - (Some [ true; true; false ]); + ~extract_name:(Some "alloc.vec.Vec.index") + ~filter:(Some [ true; true; false ]) + (); mk_fun "alloc::vec::{core::ops::index::IndexMut, \ @I>}::index_mut" - (Some "alloc.vec.Vec.index_mut") - (Some [ true; true; false ]); + ~extract_name:(Some "alloc.vec.Vec.index_mut") + ~filter:(Some [ true; true; false ]) + (); mk_fun "alloc::boxed::{core::ops::deref::Deref>}::deref" - (Some "alloc.boxed.Box.deref") - (Some [ true; false ]); + ~extract_name:(Some "alloc.boxed.Box.deref") + ~filter:(Some [ true; false ]) + (); mk_fun "alloc::boxed::{core::ops::deref::DerefMut>}::deref_mut" - (Some "alloc.boxed.Box.deref_mut") - (Some [ true; false ]); + ~extract_name:(Some "alloc.boxed.Box.deref_mut") + ~filter:(Some [ true; false ]) + (); mk_fun "core::slice::index::{core::ops::index::Index<[@T], @I>}::index" - (Some "core.slice.index.Slice.index") None; + ~extract_name:(Some "core.slice.index.Slice.index") (); mk_fun "core::slice::index::{core::ops::index::IndexMut<[@T], @I>}::index_mut" - (Some "core.slice.index.Slice.index_mut") None; + ~extract_name:(Some "core.slice.index.Slice.index_mut") (); mk_fun "core::array::{core::ops::index::Index<[@T; @N], @I>}::index" - (Some "core.array.Array.index") None; + ~extract_name:(Some "core.array.Array.index") (); mk_fun "core::array::{core::ops::index::IndexMut<[@T; @N], @I>}::index_mut" - (Some "core.array.Array.index_mut") None; + ~extract_name:(Some "core.array.Array.index_mut") (); mk_fun "core::slice::index::{core::slice::index::SliceIndex, \ [@T]>}::get" - (Some "core::slice::index::RangeUsize::get") None; + ~extract_name:(Some "core::slice::index::RangeUsize::get") (); mk_fun "core::slice::index::{core::slice::index::SliceIndex, \ [@T]>}::get_mut" - (Some "core::slice::index::RangeUsize::get_mut") None; + ~extract_name:(Some "core::slice::index::RangeUsize::get_mut") (); mk_fun "core::slice::index::{core::slice::index::SliceIndex, \ [@T]>}::index" - (Some "core::slice::index::RangeUsize::index") None; + ~extract_name:(Some "core::slice::index::RangeUsize::index") (); mk_fun "core::slice::index::{core::slice::index::SliceIndex, \ [@T]>}::index_mut" - (Some "core::slice::index::RangeUsize::index_mut") None; + ~extract_name:(Some "core::slice::index::RangeUsize::index_mut") (); mk_fun "core::slice::index::{core::slice::index::SliceIndex, \ [@T]>}::get_unchecked" - (Some "core::slice::index::RangeUsize::get_unchecked") None; + ~extract_name:(Some "core::slice::index::RangeUsize::get_unchecked") (); mk_fun "core::slice::index::{core::slice::index::SliceIndex, \ [@T]>}::get_unchecked_mut" - (Some "core::slice::index::RangeUsize::get_unchecked_mut") None; + ~extract_name:(Some "core::slice::index::RangeUsize::get_unchecked_mut") + (); mk_fun "core::slice::index::{core::slice::index::SliceIndex}::get" - None None; + (); mk_fun "core::slice::index::{core::slice::index::SliceIndex}::get_mut" - None None; + (); mk_fun "core::slice::index::{core::slice::index::SliceIndex}::get_unchecked" - None None; + (); mk_fun "core::slice::index::{core::slice::index::SliceIndex}::get_unchecked_mut" - None None; + (); mk_fun "core::slice::index::{core::slice::index::SliceIndex}::index" - (Some "core_slice_index_Slice_index") None; + ~extract_name:(Some "core_slice_index_Slice_index") (); mk_fun "core::slice::index::{core::slice::index::SliceIndex}::index_mut" - (Some "core_slice_index_Slice_index_mut") None; - mk_fun "alloc::slice::{[@T]}::to_vec" (Some "alloc.slice.Slice.to_vec") None; + ~extract_name:(Some "core_slice_index_Slice_index_mut") (); + mk_fun "alloc::slice::{[@T]}::to_vec" + ~extract_name:(Some "alloc.slice.Slice.to_vec") (); mk_fun "alloc::vec::{alloc::vec::Vec<@T, alloc::alloc::Global>}::with_capacity" - (Some "alloc.vec.Vec.with_capacity") None; - mk_fun "core::slice::{[@T]}::reverse" (Some "core.slice.Slice.reverse") None; + ~extract_name:(Some "alloc.vec.Vec.with_capacity") ~can_fail:false (); + mk_fun "core::slice::{[@T]}::reverse" + ~extract_name:(Some "core.slice.Slice.reverse") ~can_fail:false (); mk_fun "alloc::vec::{core::ops::deref::Deref>}::deref" - (Some "alloc.vec.DerefVec.deref") - (Some [ true; false ]); + ~extract_name:(Some "alloc.vec.DerefVec.deref") + ~filter:(Some [ true; false ]) + ~can_fail:false (); mk_fun "alloc::vec::{core::ops::deref::DerefMut>}::deref_mut" - (Some "alloc.vec.DerefMutVec.deref_mut") - (Some [ true; false ]); + ~extract_name:(Some "alloc.vec.DerefMutVec.deref_mut") + ~filter:(Some [ true; false ]) + (); mk_fun "core::option::{core::option::Option<@T>}::unwrap" - (Some "core.option.Option.unwrap") None; + ~extract_name:(Some "core.option.Option.unwrap") (); ] + @ List.flatten + (List.map + (fun int_name -> + List.map + (fun op -> + mk_fun + ("core::num::" ^ "{" ^ int_name ^ "}::" ^ op) + ~can_fail:false ()) + [ "wrapping_add"; "wrapping_sub"; "rotate_left"; "rotate_right" ]) + all_int_names) @ List.flatten (List.map (fun op -> mk_scalar_fun (fun ty -> "core::num::{" ^ ty ^ "}::checked_" ^ op) (fun ty -> - StringUtils.capitalize_first_letter ty ^ ".checked_" ^ op)) + StringUtils.capitalize_first_letter ty ^ ".checked_" ^ op) + ~can_fail:false ()) [ "add"; "sub"; "mul"; "div"; "rem" ]) (* From *) @ mk_scalar_fun @@ -406,44 +439,50 @@ let builtin_funs () : (pattern * bool list option * builtin_fun_info) list = "core.convert.num.From" ^ StringUtils.capitalize_first_letter ty ^ "Bool.from") + ~can_fail:false () (* From *) @ List.map (fun (big, small) -> mk_fun ("core::convert::num::{core::convert::From<" ^ big ^ ", " ^ small ^ ">}::from") - (Some - ("core.convert.num.From" - ^ StringUtils.capitalize_first_letter big - ^ StringUtils.capitalize_first_letter small - ^ ".from")) - None) + ~extract_name: + (Some + ("core.convert.num.From" + ^ StringUtils.capitalize_first_letter big + ^ StringUtils.capitalize_first_letter small + ^ ".from")) + ~can_fail:false ()) int_and_smaller_list (* Leading zeros *) @ mk_scalar_fun (fun ty -> "core::num::{" ^ ty ^ "}::leading_zeros") (fun ty -> "core.num." ^ StringUtils.capitalize_first_letter ty ^ ".leading_zeros") + ~can_fail:false () (* to_le_bytes *) @ mk_scalar_fun (fun ty -> "core::num::{" ^ ty ^ "}::to_le_bytes") (fun ty -> "core.num." ^ StringUtils.capitalize_first_letter ty ^ ".to_le_bytes") + ~can_fail:false () (* to_be_bytes *) @ mk_scalar_fun (fun ty -> "core::num::{" ^ ty ^ "}::to_be_bytes") (fun ty -> "core.num." ^ StringUtils.capitalize_first_letter ty ^ ".to_be_bytes") + ~can_fail:false () (* Clone *) @ [ mk_fun "core::clone::impls::{core::clone::Clone}::clone" - (Some "core.clone.CloneBool.clone") None; + ~extract_name:(Some "core.clone.CloneBool.clone") ~can_fail:false (); ] (* Clone *) @ mk_scalar_fun (fun ty -> "core::clone::impls::{core::clone::Clone<" ^ ty ^ ">}::clone") (fun ty -> "core.clone.Clone" ^ StringUtils.capitalize_first_letter ty ^ ".clone") + ~can_fail:false () (* Lean-only definitions *) @ mk_lean_only [ @@ -451,15 +490,19 @@ let builtin_funs () : (pattern * bool list option * builtin_fun_info) list = By construction, we cannot write down that parameter in the output in this list *) - mk_fun "core::mem::swap" None None; + mk_fun "core::mem::swap" ~can_fail:false (); mk_fun "core::option::{core::option::Option<@T>}::take" - (Some (backend_choice "" "Option::take")) - None; + ~extract_name:(Some (backend_choice "" "Option::take")) + ~can_fail:false (); mk_fun "core::option::{core::option::Option<@T>}::is_none" - (Some (backend_choice "" "Option::isNone")) - (Some [ false ]); + ~extract_name:(Some (backend_choice "" "Option::isNone")) + ~filter:(Some [ false ]) ~can_fail:false (); ] +let builtin_funs : unit -> (pattern * bool list option * builtin_fun_info) list + = + mk_memoized mk_builtin_funs + let mk_builtin_funs_map () = let m = NameMatcherMap.of_list @@ -475,106 +518,20 @@ let builtin_funs_map = mk_memoized mk_builtin_funs_map type effect_info = { can_fail : bool; stateful : bool } -let builtin_fun_effects = - let int_ops = - [ "wrapping_add"; "wrapping_sub"; "rotate_left"; "rotate_right" ] - in - let int_funs = - List.map - (fun int_name -> - List.map (fun op -> "core::num::" ^ "{" ^ int_name ^ "}::" ^ op) int_ops) - all_int_names - @ List.map - (fun op -> - List.map - (fun ty -> "core::num::{" ^ ty ^ "}::checked_" ^ op) - all_int_names) - [ "add"; "sub"; "mul"; "div"; "rem" ] - (* From *) - @ [ - List.map - (fun int_name -> - "core::convert::num::{core::convert::From<" ^ int_name - ^ ", bool>}::from") - all_int_names; - ] - (* From *) - @ [ - List.map - (fun (big, small) -> - "core::convert::num::{core::convert::From<" ^ big ^ ", " ^ small - ^ ">}::from") - int_and_smaller_list; - ] - (* Leading zeros *) - @ [ - List.map - (fun ty -> "core::num::{" ^ ty ^ "}::leading_zeros") - all_int_names; - ] - (* to_{le,be}_bytes *) - @ List.map - (fun ty -> - let pre = "core::num::{" ^ ty ^ "}::" in - [ pre ^ "to_le_bytes"; pre ^ "to_be_bytes" ]) - all_int_names - (* clone *) - @ [ - List.map - (fun ty -> - "core::clone::impls::{core::clone::Clone<" ^ ty ^ ">}::clone") - all_int_names; - ] +let mk_builtin_fun_effects () : (pattern * effect_info) list = + let builtin_funs : (pattern * bool list option * builtin_fun_info) list = + builtin_funs () in + List.map + (fun ((pattern, _, info) : _ * _ * builtin_fun_info) -> + let info = { can_fail = info.can_fail; stateful = info.stateful } in + (pattern, info)) + builtin_funs - let int_funs = List.concat int_funs in - let no_fail_no_state_funs = - [ - (* TODO: redundancy with the funs information above *) - "core::slice::{[@T]}::len"; - "alloc::vec::{alloc::vec::Vec<@T, alloc::alloc::Global>}::new"; - "alloc::vec::{alloc::vec::Vec<@T, @A>}::len"; - "core::mem::replace"; - "core::mem::take"; - "core::clone::impls::{core::clone::Clone}::clone"; - "alloc::vec::{alloc::vec::Vec<@T, alloc::alloc::Global>}::with_capacity"; - "core::slice::{[@T]}::reverse"; - "alloc::vec::{core::ops::deref::Deref>}::deref"; - ] - @ int_funs - @ mk_lean_only - [ - "core::mem::swap"; - "core::option::{core::option::Option<@T>}::take"; - "core::option::{core::option::Option<@T>}::is_none"; - ] - in - let no_fail_no_state_funs = - List.map - (fun n -> (n, { can_fail = false; stateful = false })) - no_fail_no_state_funs - in - (* TODO: all the functions registered in the [builtin_funs] above should - be considered as not using a state. There is a lot of redundancy - right now. *) - let no_state_funs = - [ - "alloc::vec::{alloc::vec::Vec<@T, @A>}::push"; - "alloc::vec::{core::ops::index::Index, \ - @I>}::index"; - "alloc::vec::{core::ops::index::IndexMut, \ - @I>}::index_mut"; - "core::option::{core::option::Option<@T>}::unwrap"; - ] - in - let no_state_funs = - List.map (fun n -> (n, { can_fail = true; stateful = false })) no_state_funs - in - no_fail_no_state_funs @ no_state_funs +let mk_builtin_fun_effects_map () = + NameMatcherMap.of_list (mk_builtin_fun_effects ()) -let builtin_fun_effects_map = - NameMatcherMap.of_list - (List.map (fun (n, x) -> (parse_pattern n, x)) builtin_fun_effects) +let builtin_fun_effects_map = mk_memoized mk_builtin_fun_effects_map type builtin_trait_decl_info = { rust_name : pattern; @@ -632,13 +589,16 @@ let builtin_trait_decls_info () = if !record_fields_short_names then item_name else extract_name ^ "_" ^ item_name in - let fwd = { extract_name = basename } in + let fwd = + { extract_name = basename; can_fail = true; stateful = false } + in (item_name, fwd) in List.map mk_method methods | Some methods -> List.map - (fun (item_name, extract_name) -> (item_name, { extract_name })) + (fun (item_name, extract_name) -> + (item_name, { extract_name; can_fail = true; stateful = false })) methods in { diff --git a/compiler/FunsAnalysis.ml b/compiler/FunsAnalysis.ml index a11eab87..eadd8f8a 100644 --- a/compiler/FunsAnalysis.ml +++ b/compiler/FunsAnalysis.ml @@ -74,7 +74,8 @@ let analyze_module (m : crate) (funs_map : fun_decl FunDeclId.Map.t) way. *) let get_builtin_info (f : fun_decl) : ExtractBuiltin.effect_info option = let open ExtractBuiltin in - NameMatcherMap.find_opt name_matcher_ctx f.name builtin_fun_effects_map + NameMatcherMap.find_opt name_matcher_ctx f.name + (builtin_fun_effects_map ()) in (* JP: Why not use a reduce visitor here with a tuple of the values to be -- cgit v1.2.3 From 9ba6fc9b83b773ed4aa0e5a90d9103ecd700323d Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 29 May 2024 08:00:37 +0200 Subject: Update the README --- README.md | 13 +++++++------ 1 file changed, 7 insertions(+), 6 deletions(-) diff --git a/README.md b/README.md index 86c1b31e..82f1da55 100644 --- a/README.md +++ b/README.md @@ -94,13 +94,11 @@ design a mechanism to allow using Aeneas in combination with tools targeting uns We have the following limitations, that we plan to address one by one: - **loops**: no nested loops for now. We are working on lifting this limitation. -- **no functions pointers/closures/traits**: ongoing work. We are actively working on this - and plan to have support soon. +- **no functions pointers/closures**: ongoing work. We have support for traits and + will have support for function pointers and closures soon. - **limited type parametricity**: it is not possible for now to instantiate a type parameter with a type containing a borrow. This is mostly an engineering - issue. We intend to quickly address the issue for types (i.e., allow `Option<&mut T>`), - and later address it for functions (i.e., allow `f<&mut T>` - we consider this to - be less urgent). + issue. - **no nested borrows in function signatures**: ongoing work. - **interior mutability**: ongoing work. We are thinking of modeling the effects of interior mutability by using ghost states. @@ -123,4 +121,7 @@ A tutorial for the Lean backend is available [here](./tests/lean/Tutorial.lean). The translation has been formalized and published at ICFP2022: [Aeneas: Rust verification by functional translation](https://dl.acm.org/doi/abs/10.1145/3547647) -([long version](https://arxiv.org/abs/2206.07185)). +([long version](https://arxiv.org/abs/2206.07185)). We also have a proof that +the symbolic execution performed by Aeneas during its translation correctly +implements a borrow checker, and published it in a +[preprint](https://arxiv.org/abs/2404.02680). -- cgit v1.2.3 From b5046454b47aba598a42d3d775d2ec54dc57c75a Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Tue, 28 May 2024 16:05:56 +0200 Subject: ci: Check correctness of the charon pinned commit --- .github/workflows/ci.yml | 6 ++++++ Makefile | 2 +- flake.nix | 1 + scripts/ci-check-charon-pin.sh | 22 ++++++++++++++++++++++ scripts/update-charon-pin.sh | 2 +- 5 files changed, 31 insertions(+), 2 deletions(-) create mode 100755 scripts/ci-check-charon-pin.sh diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 56769cae..5ee28837 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -54,3 +54,9 @@ jobs: - uses: actions/checkout@v4 - uses: cachix/install-nix-action@v26 - run: nix develop --command bash -c "cd tests/lean && make" + + check-charon-pin: + runs-on: [self-hosted, linux, nix] + steps: + - uses: actions/checkout@v4 + - run: nix develop --command ./scripts/ci-check-charon-pin.sh diff --git a/Makefile b/Makefile index 38da9f7e..b27fe0de 100644 --- a/Makefile +++ b/Makefile @@ -90,7 +90,7 @@ update-charon-pin: # Keep the commit revision in `./charon-pin` as well so that non-nix users can # know which commit to use. ./charon-pin: flake.lock - nix-shell -p jq --run './scripts/update-charon-pin.sh' >> ./charon-pin + ./scripts/update-charon-pin.sh >> ./charon-pin # Checks that `./charon` contains a clone of charon at the required commit. # Also checks that `./charon/bin/charon` exists. diff --git a/flake.nix b/flake.nix index 654c0006..3e8b88de 100644 --- a/flake.nix +++ b/flake.nix @@ -226,6 +226,7 @@ pkgs.ocamlPackages.ocamlformat pkgs.ocamlPackages.menhir pkgs.ocamlPackages.odoc + pkgs.jq ]; inputsFrom = [ diff --git a/scripts/ci-check-charon-pin.sh b/scripts/ci-check-charon-pin.sh new file mode 100755 index 00000000..4aacd96e --- /dev/null +++ b/scripts/ci-check-charon-pin.sh @@ -0,0 +1,22 @@ +#!/usr/bin/env bash +# Checks that the charon pin: +# - moves forward from the previous pin, to ensure we don't regress the charon version; +# - is merged into Charon. + +NEW_CHARON_PIN="$(cat flake.lock | jq -r .nodes.charon.locked.rev)" +OLD_CHARON_PIN="$(git show origin/main:flake.lock | jq -r .nodes.charon.locked.rev)" +echo "This PR updates the charon pin from $OLD_CHARON_PIN to $NEW_CHARON_PIN" + +git clone https://github.com/AeneasVerif/charon +cd charon +CHARON_MAIN="$(git rev-parse HEAD)" + +if ! git merge-base --is-ancestor "$OLD_CHARON_PIN" "$NEW_CHARON_PIN"; then + echo "Error: the new charon pin does not have the old one as its ancestor. The pin must only move forward." + exit 1 +fi + +if ! git merge-base --is-ancestor "$NEW_CHARON_PIN" "$CHARON_MAIN"; then + echo "Error: commit $NEW_CHARON_PIN is not merged into Charon." + exit 1 +fi diff --git a/scripts/update-charon-pin.sh b/scripts/update-charon-pin.sh index 418602b8..63dfe3b6 100755 --- a/scripts/update-charon-pin.sh +++ b/scripts/update-charon-pin.sh @@ -1,6 +1,6 @@ #!/usr/bin/env bash if ! which jq 2> /dev/null 1>&2; then - echo 'Error: command `jq` not found.' + echo 'Error: command `jq` not found; please install it.' exit 1 fi echo '# This is the commit from https://github.com/AeneasVerif/charon that should be used with this version of aeneas.' > ./charon-pin -- cgit v1.2.3 From 2d8310261ac4d19bd441de271505a0f0004028b8 Mon Sep 17 00:00:00 2001 From: Ryan Lahfa Date: Thu, 18 Apr 2024 14:26:59 +0200 Subject: chore: explain a Nix-powered workflow To avoid divergence between Charon and Aeneas, we should re-export Charon via our Flake and tell users to use this as a source of truth. Here's an appendix on how I do refresh of my files, which can serve as inspiration for a quick start workflow. Signed-off-by: Ryan Lahfa --- README.md | 19 +++++++++++++++++++ flake.nix | 1 + 2 files changed, 20 insertions(+) diff --git a/README.md b/README.md index 82f1da55..76d1720f 100644 --- a/README.md +++ b/README.md @@ -116,6 +116,25 @@ and tactics specialized for monadic programs (see A tutorial for the Lean backend is available [here](./tests/lean/Tutorial.lean). +## Quick start for Nix users + +Assuming Nix is installed, with a support for Flakes (`*`): + +```console +$ # Run Charon with the exact same version required by Aeneas +$ nix run github:aeneasverif/aeneas#charon -L +$ nix run github:aeneasverif/aeneas -L -- -backend your_preferred_backend your_llbc.file +``` + +To regenerate the extraction, just run step 2 and step 3 again. + +`(*)` : Flakes are not necessary, here is an example of how to do similar steps without it: + +```console +$ nix-shell '' -A packages.x86_64-linux.charon --run "charon" -I aeneas=https://github.com/AeneasVerif/aeneas/archive/main.tar.gz +$ nix-shell '' -A packages.x86_64-linux.default --run "aeneas --backend your_preferred_backend your_llbc.file" -I aeneas=https://github.com/AeneasVerif/aeneas/archive/main.tar.gz +``` + ## Formalization The translation has been formalized and published at ICFP2022: [Aeneas: Rust diff --git a/flake.nix b/flake.nix index 3e8b88de..cb3f84ea 100644 --- a/flake.nix +++ b/flake.nix @@ -217,6 +217,7 @@ { packages = { inherit aeneas; + inherit (charon.packages.${system}) charon-ml charon; default = aeneas; }; devShells.default = pkgs.mkShell { -- cgit v1.2.3 From b669f7c1228efb362cbb56b95090b24c0611ba7b Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Tue, 28 May 2024 12:01:31 +0200 Subject: tests: Add a README --- tests/README.md | 67 +++++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 67 insertions(+) create mode 100644 tests/README.md diff --git a/tests/README.md b/tests/README.md new file mode 100644 index 00000000..4f4f1ff5 --- /dev/null +++ b/tests/README.md @@ -0,0 +1,67 @@ +# Aeneas test suite + +This folder contains the test suite for Aeneas. It is organized as follows: +- `test_runner/`: ocaml script that orchestrates the test suite; +- `src/`: rust files and crates used as test inputs; +- `llbc/`: temporary directory (ignored by git) that stores intermediate files generated during tests; +- `coq/`, `fstar/`, `hol4/`, `lean/`: one folder per backend. Each folder contains: + - Files generated by Aeneas from our test inputs; + - Handwritten proof files. + +# Running the test suite + +Running the test suite has two parts: +- The generation of files from rust to each backend. This is done by running `make test` in the + project root. This will call the test runner on each file or crate in `tests/src`, which will call + `charon` and `aeneas` to generate the backend-specific outputs. +- The verification of proofs for each backend. This is done by running `make verify` in the project + root. It will run each verifier (`fstar`, `lean` etc) on the files in the corresponding folder to + verify the correctness of proofs. + +CI does both of these things; in particular it checks that we correctly committed the generated +files. + +Important note: at the moment we don't regenerate the `hol4` outputs. We do still check them with +hol4. This is tracked in https://github.com/AeneasVerif/aeneas/issues/62. + +# Adding a new test + +Adding a new test is easy: just add a `foo.rs` file or a `foo_crate` folder in `tests/src`. `make +test` will find the new file and call the test runner on it. + +To specify options for how your test should be run, see the next section. + +Ideally, any non-trivial change to Aeneas should have an accompanying test that exercises the new +code. The goal of this test suite is to cover a large portion of what Aeneas can do, so we can work +on it confidently. + +# Passing options to `Aeneas` and `Charon` + +The test runner supports setting several options for each test. +- For single files, it will read comments at the top of the file that start with `//@`; +- For crates, it will read the `crate_dir/aeneas-test-options` file. + +In both cases it supports the same options. Typical options are: +- `charon-args=--polonius --opaque=betree_utils`: pass these arguments to `charon`; +- `aeneas-args=-test-trans-units`: pass these arguments to `aeneas` for all backends; +- `[fstar] aeneas-args=-decreases-clauses -template-clauses`: pass these arguments to `aeneas` for + the `fstar` backend; +- `[coq,lean] skip`: skip the test for the `coq` and `lean` backends; +- `[fstar] known-failure`: see below; +- `[coq,fstar] subdir=misc`: store the output in the `misc` subdirectory for these two backends. + +For an up-to-date list of options, see comments in the test runner code. + +# `known-failure` tests + +There's a special kind of test that doesn't generate backend code: `//@ known-failure` tests. These +are meant to record: +- Cases that produce an error so we don't forget and can fix them later; +- Cases that _should_ error, to ensure we correctly raise an error and that the error output is nice. + +It is useful to add tests there whenever you find an interesting failure case. + +Note however that this is for failures of Aeneas. Tests that cause Charon errors should be submitted +to the Charon project. `known-failure` tests will not pass if the error occurs within Charon. Note +also that there is no option for tests that translate successfully but fail to be verified by the +backend. -- cgit v1.2.3 From 86d0789b5a303f43c0d9bfeff83f37d89750b5d6 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Tue, 28 May 2024 12:14:08 +0200 Subject: runner: make the backend map a submodule of `Backend` --- tests/test_runner/run_test.ml | 84 ++++++++++++++++++++++++------------------- 1 file changed, 48 insertions(+), 36 deletions(-) diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index c626b4d1..5314752c 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -36,27 +36,34 @@ module Backend = struct | Lean -> "lean" | FStar -> "fstar" | HOL4 -> "hol4" -end -module BackendMap = struct - include Map.Make (Backend) + module Map = struct + (* Hack to be able to include things from the parent module with the same names *) + type backend_t = t - (* Make a new map with one entry per backend, given by `f` *) - let make (f : Backend.t -> 'a) : 'a t = - List.fold_left - (fun map backend -> add backend (f backend) map) - empty Backend.all + let backend_compare = compare - (* Set this value for all the backends in `backends` *) - let add_each (backends : Backend.t list) (v : 'a) (map : 'a t) : 'a t = - List.fold_left (fun map backend -> add backend v map) map backends + include Map.Make (struct + type t = backend_t - (* Updates all the backends in `backends` with `f` *) - let update_each (backends : Backend.t list) (f : 'a -> 'a) (map : 'a t) : 'a t - = - List.fold_left - (fun map backend -> update backend (Option.map f) map) - map backends + let compare = backend_compare + end) + + (* Make a new map with one entry per backend, given by `f` *) + let make (f : backend_t -> 'a) : 'a t = + List.fold_left (fun map backend -> add backend (f backend) map) empty all + + (* Set this value for all the backends in `backends` *) + let add_each (backends : backend_t list) (v : 'a) (map : 'a t) : 'a t = + List.fold_left (fun map backend -> add backend v map) map backends + + (* Updates all the backends in `backends` with `f` *) + let update_each (backends : backend_t list) (f : 'a -> 'a) (map : 'a t) : + 'a t = + List.fold_left + (fun map backend -> update backend (Option.map f) map) + map backends + end end let concat_path = List.fold_left Filename.concat "" @@ -109,12 +116,12 @@ module Input = struct name : string; path : string; kind : kind; - actions : action BackendMap.t; + actions : action Backend.Map.t; charon_options : string list; - aeneas_options : string list BackendMap.t; - subdirs : string BackendMap.t; + aeneas_options : string list Backend.Map.t; + subdirs : string Backend.Map.t; (* Whether to store the command output to a file. Only available for known-failure tests. *) - check_output : bool BackendMap.t; + check_output : bool Backend.Map.t; } (* The default subdirectory in which to store the outputs. *) @@ -152,16 +159,16 @@ module Input = struct let subdir = Core.String.chop_prefix comment ~prefix:"subdir=" in if comment = "skip" then - { input with actions = BackendMap.add_each backends Skip input.actions } + { input with actions = Backend.Map.add_each backends Skip input.actions } else if comment = "known-failure" then { input with - actions = BackendMap.add_each backends KnownFailure input.actions; + actions = Backend.Map.add_each backends KnownFailure input.actions; } else if comment = "no-check-output" then { input with - check_output = BackendMap.add_each backends false input.check_output; + check_output = Backend.Map.add_each backends false input.check_output; } else if Option.is_some charon_args then let args = Option.get charon_args in @@ -176,11 +183,14 @@ module Input = struct { input with aeneas_options = - BackendMap.update_each backends add_args input.aeneas_options; + Backend.Map.update_each backends add_args input.aeneas_options; } else if Option.is_some subdir then let subdir = Option.get subdir in - { input with subdirs = BackendMap.add_each backends subdir input.subdirs } + { + input with + subdirs = Backend.Map.add_each backends subdir input.subdirs; + } else failwith ("Unrecognized special comment: `" ^ comment ^ "`") (* Given a path to a rust file or crate, gather the details and options about how to build the test. *) @@ -192,12 +202,12 @@ module Input = struct else if Sys_unix.is_directory_exn path then Crate else failwith ("`" ^ path ^ "` is not a file or a directory.") in - let actions = BackendMap.make (fun _ -> Normal) in + let actions = Backend.Map.make (fun _ -> Normal) in let subdirs = - BackendMap.make (fun backend -> default_subdir backend name) + Backend.Map.make (fun backend -> default_subdir backend name) in - let aeneas_options = BackendMap.make (fun _ -> []) in - let check_output = BackendMap.make (fun _ -> true) in + let aeneas_options = Backend.Map.make (fun _ -> []) in + let check_output = Backend.Map.make (fun _ -> true) in let input = { path; @@ -240,11 +250,11 @@ let run_aeneas (env : runner_env) (case : Input.t) (backend : Backend.t) = let output_file = Filename.remove_extension case.path ^ "." ^ backend_str ^ ".out" in - let subdir = BackendMap.find backend case.subdirs in + let subdir = Backend.Map.find backend case.subdirs in let dest_dir = concat_path [ env.dest_dir; backend_str; subdir ] in - let aeneas_options = BackendMap.find backend case.aeneas_options in - let action = BackendMap.find backend case.actions in - let check_output = BackendMap.find backend case.check_output in + let aeneas_options = Backend.Map.find backend case.aeneas_options in + let action = Backend.Map.find backend case.actions in + let check_output = Backend.Map.find backend case.check_output in (* Build the command *) let args = @@ -332,13 +342,15 @@ let () = { test_case with aeneas_options = - BackendMap.map (List.append aeneas_options) test_case.aeneas_options; + Backend.Map.map + (List.append aeneas_options) + test_case.aeneas_options; } in let skip_all = List.for_all (fun backend -> - BackendMap.find backend test_case.actions = Input.Skip) + Backend.Map.find backend test_case.actions = Input.Skip) Backend.all in if skip_all then () -- cgit v1.2.3 From ec03335a473ffdf9371210e8558c691ea69d212d Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Tue, 28 May 2024 12:18:52 +0200 Subject: runner: Split up into multiple files --- tests/README.md | 2 +- tests/test_runner/Backend.ml | 47 ++++++++++ tests/test_runner/Input.ml | 129 +++++++++++++++++++++++++++ tests/test_runner/Utils.ml | 12 +++ tests/test_runner/dune | 3 +- tests/test_runner/run_test.ml | 196 ------------------------------------------ 6 files changed, 191 insertions(+), 198 deletions(-) create mode 100644 tests/test_runner/Backend.ml create mode 100644 tests/test_runner/Input.ml create mode 100644 tests/test_runner/Utils.ml diff --git a/tests/README.md b/tests/README.md index 4f4f1ff5..f68b509f 100644 --- a/tests/README.md +++ b/tests/README.md @@ -50,7 +50,7 @@ In both cases it supports the same options. Typical options are: - `[fstar] known-failure`: see below; - `[coq,fstar] subdir=misc`: store the output in the `misc` subdirectory for these two backends. -For an up-to-date list of options, see comments in the test runner code. +For an up-to-date list of options, see comments in `tests/test_runner/Input.ml`. # `known-failure` tests diff --git a/tests/test_runner/Backend.ml b/tests/test_runner/Backend.ml new file mode 100644 index 00000000..d21a46fc --- /dev/null +++ b/tests/test_runner/Backend.ml @@ -0,0 +1,47 @@ +(*** Define the backends we support as an enum *) + +type t = Coq | Lean | FStar | HOL4 [@@deriving ord, sexp] + +(* TODO: reactivate HOL4 once traits are parameterized by their associated types *) +let all = [ Coq; Lean; FStar ] + +let of_string = function + | "coq" -> Coq + | "lean" -> Lean + | "fstar" -> FStar + | "hol4" -> HOL4 + | backend -> failwith ("Unknown backend: `" ^ backend ^ "`") + +let to_string = function + | Coq -> "coq" + | Lean -> "lean" + | FStar -> "fstar" + | HOL4 -> "hol4" + +module Map = struct + (* Hack to be able to include things from the parent module with the same names *) + type backend_t = t + + let backend_compare = compare + + include Map.Make (struct + type t = backend_t + + let compare = backend_compare + end) + + (* Make a new map with one entry per backend, given by `f` *) + let make (f : backend_t -> 'a) : 'a t = + List.fold_left (fun map backend -> add backend (f backend) map) empty all + + (* Set this value for all the backends in `backends` *) + let add_each (backends : backend_t list) (v : 'a) (map : 'a t) : 'a t = + List.fold_left (fun map backend -> add backend v map) map backends + + (* Updates all the backends in `backends` with `f` *) + let update_each (backends : backend_t list) (f : 'a -> 'a) (map : 'a t) : 'a t + = + List.fold_left + (fun map backend -> update backend (Option.map f) map) + map backends +end diff --git a/tests/test_runner/Input.ml b/tests/test_runner/Input.ml new file mode 100644 index 00000000..dbb69a47 --- /dev/null +++ b/tests/test_runner/Input.ml @@ -0,0 +1,129 @@ +(*** The data for a specific test input *) + +type kind = SingleFile | Crate +type action = Normal | Skip | KnownFailure + +type t = { + name : string; + path : string; + kind : kind; + actions : action Backend.Map.t; + charon_options : string list; + aeneas_options : string list Backend.Map.t; + subdirs : string Backend.Map.t; + (* Whether to store the command output to a file. Only available for known-failure tests. *) + check_output : bool Backend.Map.t; +} + +(* The default subdirectory in which to store the outputs. *) +let default_subdir backend test_name = + match backend with Backend.Lean -> "." | _ -> test_name + +(* Parse lines that start `//@`. Each of them modifies the options we use for the test. + Supported comments: + - `skip`: don't process the file; + - `known-failure`: TODO; + - `subdir=...: set the subdirectory in which to store the outputs. + Defaults to nothing for lean and to the test name for other backends; + - `charon-args=...`: extra arguments to pass to charon; + - `aeneas-args=...`: extra arguments to pass to aeneas; + - `[backend,..]...`: where each `backend` is the name of a backend supported by + aeneas; this sets options for these backends only. +*) +let apply_special_comment comment input = + let comment = String.trim comment in + (* Parse the backends if any *) + let re = Re.compile (Re.Pcre.re "^\\[([a-zA-Z,]+)\\](.*)$") in + let comment, (backends : Backend.t list) = + match Re.exec_opt re comment with + | Some groups -> + let backends = Re.Group.get groups 1 in + let backends = String.split_on_char ',' backends in + let backends = List.map Backend.of_string backends in + let rest = Re.Group.get groups 2 in + (String.trim rest, backends) + | None -> (comment, Backend.all) + in + (* Parse the other options *) + let charon_args = Core.String.chop_prefix comment ~prefix:"charon-args=" in + let aeneas_args = Core.String.chop_prefix comment ~prefix:"aeneas-args=" in + let subdir = Core.String.chop_prefix comment ~prefix:"subdir=" in + + if comment = "skip" then + { input with actions = Backend.Map.add_each backends Skip input.actions } + else if comment = "known-failure" then + { + input with + actions = Backend.Map.add_each backends KnownFailure input.actions; + } + else if comment = "no-check-output" then + { + input with + check_output = Backend.Map.add_each backends false input.check_output; + } + else if Option.is_some charon_args then + let args = Option.get charon_args in + let args = String.split_on_char ' ' args in + if backends != Backend.all then + failwith "Cannot set per-backend charon-args" + else { input with charon_options = List.append input.charon_options args } + else if Option.is_some aeneas_args then + let args = Option.get aeneas_args in + let args = String.split_on_char ' ' args in + let add_args opts = List.append opts args in + { + input with + aeneas_options = + Backend.Map.update_each backends add_args input.aeneas_options; + } + else if Option.is_some subdir then + let subdir = Option.get subdir in + { input with subdirs = Backend.Map.add_each backends subdir input.subdirs } + else failwith ("Unrecognized special comment: `" ^ comment ^ "`") + +(* Given a path to a rust file or crate, gather the details and options about how to build the test. *) +let build (path : string) : t = + let name = Filename.remove_extension (Filename.basename path) in + let name = Str.global_replace (Str.regexp "-") "_" name in + let kind = + if Sys_unix.is_file_exn path then SingleFile + else if Sys_unix.is_directory_exn path then Crate + else failwith ("`" ^ path ^ "` is not a file or a directory.") + in + let actions = Backend.Map.make (fun _ -> Normal) in + let subdirs = Backend.Map.make (fun backend -> default_subdir backend name) in + let aeneas_options = Backend.Map.make (fun _ -> []) in + let check_output = Backend.Map.make (fun _ -> true) in + let input = + { + path; + name; + kind; + actions; + charon_options = []; + subdirs; + aeneas_options; + check_output; + } + in + (* For crates, we store the special comments in a separate file. *) + let crate_config_file = Filename.concat path "aeneas-test-options" in + match kind with + | SingleFile -> + let file_lines = Core.In_channel.read_lines path in + (* Extract the special lines. Stop at the first non-special line. *) + let special_comments = + Utils.map_while + (fun line -> Core.String.chop_prefix line ~prefix:"//@") + file_lines + in + (* Apply the changes from the special lines to our input. *) + List.fold_left + (fun input comment -> apply_special_comment comment input) + input special_comments + | Crate when Sys.file_exists crate_config_file -> + let special_comments = Core.In_channel.read_lines crate_config_file in + List.fold_left + (fun input comment -> apply_special_comment comment input) + input special_comments + | Crate -> input diff --git a/tests/test_runner/Utils.ml b/tests/test_runner/Utils.ml new file mode 100644 index 00000000..79be617c --- /dev/null +++ b/tests/test_runner/Utils.ml @@ -0,0 +1,12 @@ +(*** Convenience functions *) + +let map_while (f : 'a -> 'b option) (input : 'a list) : 'b list = + let _, result = + List.fold_left + (fun (continue, out) a -> + if continue then + match f a with None -> (false, out) | Some b -> (true, b :: out) + else (continue, out)) + (true, []) input + in + List.rev result diff --git a/tests/test_runner/dune b/tests/test_runner/dune index 6bb3f7b2..52688d14 100644 --- a/tests/test_runner/dune +++ b/tests/test_runner/dune @@ -3,7 +3,8 @@ (libraries core_unix core_unix.filename_unix core_unix.sys_unix re str unix) (preprocess (pps ppx_deriving.show ppx_deriving.ord ppx_sexp_conv)) - (name run_test)) + (name run_test) + (modules Backend Input Utils run_test)) (env (dev diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index 5314752c..e8dd38bb 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -1,15 +1,3 @@ -(* Convenience functions *) -let map_while (f : 'a -> 'b option) (input : 'a list) : 'b list = - let _, result = - List.fold_left - (fun (continue, out) a -> - if continue then - match f a with None -> (false, out) | Some b -> (true, b :: out) - else (continue, out)) - (true, []) input - in - List.rev result - (* Paths to use for tests *) type runner_env = { charon_path : string; @@ -18,54 +6,6 @@ type runner_env = { dest_dir : string; } -module Backend = struct - type t = Coq | Lean | FStar | HOL4 [@@deriving ord, sexp] - - (* TODO: reactivate HOL4 once traits are parameterized by their associated types *) - let all = [ Coq; Lean; FStar ] - - let of_string = function - | "coq" -> Coq - | "lean" -> Lean - | "fstar" -> FStar - | "hol4" -> HOL4 - | backend -> failwith ("Unknown backend: `" ^ backend ^ "`") - - let to_string = function - | Coq -> "coq" - | Lean -> "lean" - | FStar -> "fstar" - | HOL4 -> "hol4" - - module Map = struct - (* Hack to be able to include things from the parent module with the same names *) - type backend_t = t - - let backend_compare = compare - - include Map.Make (struct - type t = backend_t - - let compare = backend_compare - end) - - (* Make a new map with one entry per backend, given by `f` *) - let make (f : backend_t -> 'a) : 'a t = - List.fold_left (fun map backend -> add backend (f backend) map) empty all - - (* Set this value for all the backends in `backends` *) - let add_each (backends : backend_t list) (v : 'a) (map : 'a t) : 'a t = - List.fold_left (fun map backend -> add backend v map) map backends - - (* Updates all the backends in `backends` with `f` *) - let update_each (backends : backend_t list) (f : 'a -> 'a) (map : 'a t) : - 'a t = - List.fold_left - (fun map backend -> update backend (Option.map f) map) - map backends - end -end - let concat_path = List.fold_left Filename.concat "" module Command = struct @@ -107,142 +47,6 @@ module Command = struct | Failure -> () end -(* The data for a specific test input *) -module Input = struct - type kind = SingleFile | Crate - type action = Normal | Skip | KnownFailure - - type t = { - name : string; - path : string; - kind : kind; - actions : action Backend.Map.t; - charon_options : string list; - aeneas_options : string list Backend.Map.t; - subdirs : string Backend.Map.t; - (* Whether to store the command output to a file. Only available for known-failure tests. *) - check_output : bool Backend.Map.t; - } - - (* The default subdirectory in which to store the outputs. *) - let default_subdir backend test_name = - match backend with Backend.Lean -> "." | _ -> test_name - - (* Parse lines that start `//@`. Each of them modifies the options we use for the test. - Supported comments: - - `skip`: don't process the file; - - `known-failure`: TODO; - - `subdir=...: set the subdirectory in which to store the outputs. - Defaults to nothing for lean and to the test name for other backends; - - `charon-args=...`: extra arguments to pass to charon; - - `aeneas-args=...`: extra arguments to pass to aeneas; - - `[backend,..]...`: where each `backend` is the name of a backend supported by - aeneas; this sets options for these backends only. - *) - let apply_special_comment comment input = - let comment = String.trim comment in - (* Parse the backends if any *) - let re = Re.compile (Re.Pcre.re "^\\[([a-zA-Z,]+)\\](.*)$") in - let comment, (backends : Backend.t list) = - match Re.exec_opt re comment with - | Some groups -> - let backends = Re.Group.get groups 1 in - let backends = String.split_on_char ',' backends in - let backends = List.map Backend.of_string backends in - let rest = Re.Group.get groups 2 in - (String.trim rest, backends) - | None -> (comment, Backend.all) - in - (* Parse the other options *) - let charon_args = Core.String.chop_prefix comment ~prefix:"charon-args=" in - let aeneas_args = Core.String.chop_prefix comment ~prefix:"aeneas-args=" in - let subdir = Core.String.chop_prefix comment ~prefix:"subdir=" in - - if comment = "skip" then - { input with actions = Backend.Map.add_each backends Skip input.actions } - else if comment = "known-failure" then - { - input with - actions = Backend.Map.add_each backends KnownFailure input.actions; - } - else if comment = "no-check-output" then - { - input with - check_output = Backend.Map.add_each backends false input.check_output; - } - else if Option.is_some charon_args then - let args = Option.get charon_args in - let args = String.split_on_char ' ' args in - if backends != Backend.all then - failwith "Cannot set per-backend charon-args" - else { input with charon_options = List.append input.charon_options args } - else if Option.is_some aeneas_args then - let args = Option.get aeneas_args in - let args = String.split_on_char ' ' args in - let add_args opts = List.append opts args in - { - input with - aeneas_options = - Backend.Map.update_each backends add_args input.aeneas_options; - } - else if Option.is_some subdir then - let subdir = Option.get subdir in - { - input with - subdirs = Backend.Map.add_each backends subdir input.subdirs; - } - else failwith ("Unrecognized special comment: `" ^ comment ^ "`") - - (* Given a path to a rust file or crate, gather the details and options about how to build the test. *) - let build (path : string) : t = - let name = Filename.remove_extension (Filename.basename path) in - let name = Str.global_replace (Str.regexp "-") "_" name in - let kind = - if Sys_unix.is_file_exn path then SingleFile - else if Sys_unix.is_directory_exn path then Crate - else failwith ("`" ^ path ^ "` is not a file or a directory.") - in - let actions = Backend.Map.make (fun _ -> Normal) in - let subdirs = - Backend.Map.make (fun backend -> default_subdir backend name) - in - let aeneas_options = Backend.Map.make (fun _ -> []) in - let check_output = Backend.Map.make (fun _ -> true) in - let input = - { - path; - name; - kind; - actions; - charon_options = []; - subdirs; - aeneas_options; - check_output; - } - in - (* For crates, we store the special comments in a separate file. *) - let crate_config_file = Filename.concat path "aeneas-test-options" in - match kind with - | SingleFile -> - let file_lines = Core.In_channel.read_lines path in - (* Extract the special lines. Stop at the first non-special line. *) - let special_comments = - map_while - (fun line -> Core.String.chop_prefix line ~prefix:"//@") - file_lines - in - (* Apply the changes from the special lines to our input. *) - List.fold_left - (fun input comment -> apply_special_comment comment input) - input special_comments - | Crate when Sys.file_exists crate_config_file -> - let special_comments = Core.In_channel.read_lines crate_config_file in - List.fold_left - (fun input comment -> apply_special_comment comment input) - input special_comments - | Crate -> input -end - (* Run Aeneas on a specific input with the given options *) let run_aeneas (env : runner_env) (case : Input.t) (backend : Backend.t) = let backend_str = Backend.to_string backend in -- cgit v1.2.3 From 14d9ca2ddf5ccb350d3bd87ca14a7b7468398e9c Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Tue, 28 May 2024 13:52:12 +0200 Subject: runner: Factor out backend-specific options --- tests/test_runner/Input.ml | 83 ++++++++++++++++++++++--------------------- tests/test_runner/run_test.ml | 22 +++++++----- 2 files changed, 56 insertions(+), 49 deletions(-) diff --git a/tests/test_runner/Input.ml b/tests/test_runner/Input.ml index dbb69a47..960ffe8d 100644 --- a/tests/test_runner/Input.ml +++ b/tests/test_runner/Input.ml @@ -3,16 +3,22 @@ type kind = SingleFile | Crate type action = Normal | Skip | KnownFailure +(* Options that can be set for a specific backend *) +type per_backend = { + action : action; + aeneas_options : string list; + subdir : string; + (* Whether to store the command output to a file. Only available for known-failure tests. *) + check_output : bool; +} + +(* The data for a specific test input *) type t = { name : string; path : string; kind : kind; - actions : action Backend.Map.t; charon_options : string list; - aeneas_options : string list Backend.Map.t; - subdirs : string Backend.Map.t; - (* Whether to store the command output to a file. Only available for known-failure tests. *) - check_output : bool Backend.Map.t; + per_backend : per_backend Backend.Map.t; } (* The default subdirectory in which to store the outputs. *) @@ -44,41 +50,42 @@ let apply_special_comment comment input = (String.trim rest, backends) | None -> (comment, Backend.all) in + (* Apply `f` to the selected backends *) + let per_backend f = + { + input with + per_backend = Backend.Map.update_each backends f input.per_backend; + } + in + (* Assert that this option is not available to be set per-backend *) + let assert_no_backend option_name = + if backends != Backend.all then + failwith ("Cannot set the `" ^ option_name ^ "` option per-backend") + in + (* Parse the other options *) let charon_args = Core.String.chop_prefix comment ~prefix:"charon-args=" in let aeneas_args = Core.String.chop_prefix comment ~prefix:"aeneas-args=" in let subdir = Core.String.chop_prefix comment ~prefix:"subdir=" in - if comment = "skip" then - { input with actions = Backend.Map.add_each backends Skip input.actions } + if comment = "skip" then per_backend (fun x -> { x with action = Skip }) else if comment = "known-failure" then - { - input with - actions = Backend.Map.add_each backends KnownFailure input.actions; - } + per_backend (fun x -> { x with action = KnownFailure }) else if comment = "no-check-output" then - { - input with - check_output = Backend.Map.add_each backends false input.check_output; - } - else if Option.is_some charon_args then + per_backend (fun x -> { x with check_output = false }) + else if Option.is_some charon_args then ( let args = Option.get charon_args in let args = String.split_on_char ' ' args in - if backends != Backend.all then - failwith "Cannot set per-backend charon-args" - else { input with charon_options = List.append input.charon_options args } + assert_no_backend "charon-args"; + { input with charon_options = List.append input.charon_options args }) else if Option.is_some aeneas_args then let args = Option.get aeneas_args in let args = String.split_on_char ' ' args in - let add_args opts = List.append opts args in - { - input with - aeneas_options = - Backend.Map.update_each backends add_args input.aeneas_options; - } + per_backend (fun x -> + { x with aeneas_options = List.append x.aeneas_options args }) else if Option.is_some subdir then let subdir = Option.get subdir in - { input with subdirs = Backend.Map.add_each backends subdir input.subdirs } + per_backend (fun x -> { x with subdir }) else failwith ("Unrecognized special comment: `" ^ comment ^ "`") (* Given a path to a rust file or crate, gather the details and options about how to build the test. *) @@ -90,22 +97,16 @@ let build (path : string) : t = else if Sys_unix.is_directory_exn path then Crate else failwith ("`" ^ path ^ "` is not a file or a directory.") in - let actions = Backend.Map.make (fun _ -> Normal) in - let subdirs = Backend.Map.make (fun backend -> default_subdir backend name) in - let aeneas_options = Backend.Map.make (fun _ -> []) in - let check_output = Backend.Map.make (fun _ -> true) in - let input = - { - path; - name; - kind; - actions; - charon_options = []; - subdirs; - aeneas_options; - check_output; - } + let per_backend = + Backend.Map.make (fun backend -> + { + action = Normal; + subdir = default_subdir backend name; + aeneas_options = []; + check_output = true; + }) in + let input = { path; name; kind; charon_options = []; per_backend } in (* For crates, we store the special comments in a separate file. *) let crate_config_file = Filename.concat path "aeneas-test-options" in match kind with diff --git a/tests/test_runner/run_test.ml b/tests/test_runner/run_test.ml index e8dd38bb..c17c17be 100644 --- a/tests/test_runner/run_test.ml +++ b/tests/test_runner/run_test.ml @@ -54,11 +54,12 @@ let run_aeneas (env : runner_env) (case : Input.t) (backend : Backend.t) = let output_file = Filename.remove_extension case.path ^ "." ^ backend_str ^ ".out" in - let subdir = Backend.Map.find backend case.subdirs in + let per_backend = Backend.Map.find backend case.per_backend in + let subdir = per_backend.subdir in + let check_output = per_backend.check_output in + let aeneas_options = per_backend.aeneas_options in + let action = per_backend.action in let dest_dir = concat_path [ env.dest_dir; backend_str; subdir ] in - let aeneas_options = Backend.Map.find backend case.aeneas_options in - let action = Backend.Map.find backend case.actions in - let check_output = Backend.Map.find backend case.check_output in (* Build the command *) let args = @@ -145,16 +146,21 @@ let () = let test_case = { test_case with - aeneas_options = + per_backend = Backend.Map.map - (List.append aeneas_options) - test_case.aeneas_options; + (fun x -> + { + x with + Input.aeneas_options = + List.append aeneas_options x.Input.aeneas_options; + }) + test_case.per_backend; } in let skip_all = List.for_all (fun backend -> - Backend.Map.find backend test_case.actions = Input.Skip) + (Backend.Map.find backend test_case.per_backend).action = Input.Skip) Backend.all in if skip_all then () -- cgit v1.2.3 From 87d5a08f44b46657026a99c154bcec4a733f221d Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Wed, 29 May 2024 10:37:17 +0200 Subject: Improve the tests README --- tests/Makefile | 6 +++--- tests/README.md | 30 +++++++++++++++++++----------- 2 files changed, 22 insertions(+), 14 deletions(-) diff --git a/tests/Makefile b/tests/Makefile index ff4baaba..a911e037 100644 --- a/tests/Makefile +++ b/tests/Makefile @@ -1,6 +1,6 @@ .PHONY: all -all: test-fstar test-coq test-lean test-hol4 +all: verify-fstar verify-coq verify-lean verify-hol4 -.PHONY: test-% -test-%: +.PHONY: verify-% +verify-%: cd $* && $(MAKE) all diff --git a/tests/README.md b/tests/README.md index f68b509f..c9ca0047 100644 --- a/tests/README.md +++ b/tests/README.md @@ -8,23 +8,26 @@ This folder contains the test suite for Aeneas. It is organized as follows: - Files generated by Aeneas from our test inputs; - Handwritten proof files. -# Running the test suite +## Running the test suite Running the test suite has two parts: -- The generation of files from rust to each backend. This is done by running `make test` in the +- We generate files from rust to each backend. This is done by running `make test` in the project root. This will call the test runner on each file or crate in `tests/src`, which will call `charon` and `aeneas` to generate the backend-specific outputs. -- The verification of proofs for each backend. This is done by running `make verify` in the project - root. It will run each verifier (`fstar`, `lean` etc) on the files in the corresponding folder to - verify the correctness of proofs. +- We check that the generated output is valid code by parsing it and type-checking it with the + relevant verifier (Lean, Coq, F*, etc.), and replay the hand-written proofs which use it as + a dependency. This is done by running `make verify` in the project root. -CI does both of these things; in particular it checks that we correctly committed the generated -files. +To generate the output of a single test, run `make test-`, e.g. `make test-loops.rs` or +`make test-betree`. To verify the output for a single backend, run `make verify-` inside the +`tests/` folder, e.g. `make verify-coq`. -Important note: at the moment we don't regenerate the `hol4` outputs. We do still check them with +CI does both of these things; it also checks that we committed all the generated files. + +Important note: at the moment we don't regenerate the HOL4 outputs. We do still check them with hol4. This is tracked in https://github.com/AeneasVerif/aeneas/issues/62. -# Adding a new test +## Adding a new test Adding a new test is easy: just add a `foo.rs` file or a `foo_crate` folder in `tests/src`. `make test` will find the new file and call the test runner on it. @@ -35,7 +38,7 @@ Ideally, any non-trivial change to Aeneas should have an accompanying test that code. The goal of this test suite is to cover a large portion of what Aeneas can do, so we can work on it confidently. -# Passing options to `Aeneas` and `Charon` +## Passing options to `Aeneas` and `Charon` The test runner supports setting several options for each test. - For single files, it will read comments at the top of the file that start with `//@`; @@ -52,13 +55,18 @@ In both cases it supports the same options. Typical options are: For an up-to-date list of options, see comments in `tests/test_runner/Input.ml`. -# `known-failure` tests +## `known-failure` tests There's a special kind of test that doesn't generate backend code: `//@ known-failure` tests. These are meant to record: - Cases that produce an error so we don't forget and can fix them later; - Cases that _should_ error, to ensure we correctly raise an error and that the error output is nice. +We record the output of the `aeneas` call in `test-file.out` alongside the original `test-file.rs`. +This file must be committed like the rest of the test outputs. (Note: sometimes the output will be +different between your local machine and CI, e.g. because release build changed a stacktrace. In +that case we don't record the output, which is done with `//@ no-check-output`). + It is useful to add tests there whenever you find an interesting failure case. Note however that this is for failures of Aeneas. Tests that cause Charon errors should be submitted -- cgit v1.2.3 From 4195b9653b2d00b846edd3b74211f129192a8806 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Thu, 30 May 2024 15:36:08 +0200 Subject: Update charon pin to local commit when developing --- Makefile | 11 +++-------- scripts/update-charon-pin.sh | 11 +++++++++++ 2 files changed, 14 insertions(+), 8 deletions(-) diff --git a/Makefile b/Makefile index b27fe0de..1e832378 100644 --- a/Makefile +++ b/Makefile @@ -81,16 +81,11 @@ build-bin-dir: build-bin build-lib build-runner doc: cd compiler && dune build @doc -# Fetches the latest commit from charon and updates `flake.lock` accordingly. +# Updates `flake.lock` and `charon-pin` with the latest commit from Charon. If +# we're using a symlink, this takes the commit from our local charon. .PHONY: update-charon-pin update-charon-pin: - nix flake lock --update-input charon - $(MAKE) charon-pin - -# Keep the commit revision in `./charon-pin` as well so that non-nix users can -# know which commit to use. -./charon-pin: flake.lock - ./scripts/update-charon-pin.sh >> ./charon-pin + ./scripts/update-charon-pin.sh # Checks that `./charon` contains a clone of charon at the required commit. # Also checks that `./charon/bin/charon` exists. diff --git a/scripts/update-charon-pin.sh b/scripts/update-charon-pin.sh index 63dfe3b6..1fa706ae 100755 --- a/scripts/update-charon-pin.sh +++ b/scripts/update-charon-pin.sh @@ -3,5 +3,16 @@ if ! which jq 2> /dev/null 1>&2; then echo 'Error: command `jq` not found; please install it.' exit 1 fi + +if [ -L charon ]; then + echo '`./charon` is a symlink; we using the commit there for our new pin.' + COMMIT="$(git -C charon rev-parse HEAD)" + nix flake lock --override-input charon "github:aeneasverif/charon/$COMMIT" +else + echo 'Pinning the latest commit from Charon `main`' + nix flake lock --update-input charon +fi + +# Keep the commit revision in `./charon-pin` as well so that non-nix users can know which commit to use. echo '# This is the commit from https://github.com/AeneasVerif/charon that should be used with this version of aeneas.' > ./charon-pin jq -r .nodes.charon.locked.rev flake.lock >> ./charon-pin -- cgit v1.2.3 From 19224a45930e4fb0786d6ace0abcded86d110c65 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Thu, 30 May 2024 16:34:47 +0200 Subject: Remove unused `rust-scripts` --- rust-scripts/Cargo.toml | 7 --- rust-scripts/src/main.rs | 150 ----------------------------------------------- 2 files changed, 157 deletions(-) delete mode 100644 rust-scripts/Cargo.toml delete mode 100644 rust-scripts/src/main.rs diff --git a/rust-scripts/Cargo.toml b/rust-scripts/Cargo.toml deleted file mode 100644 index 31ef5be0..00000000 --- a/rust-scripts/Cargo.toml +++ /dev/null @@ -1,7 +0,0 @@ -[package] -name = "rust-tests" -version = "0.1.0" -authors = ["Son Ho "] -edition = "2018" - -[dependencies] \ No newline at end of file diff --git a/rust-scripts/src/main.rs b/rust-scripts/src/main.rs deleted file mode 100644 index 125b0d76..00000000 --- a/rust-scripts/src/main.rs +++ /dev/null @@ -1,150 +0,0 @@ -/// The following code generates the limits for the scalar types - -fn test_modulo(x: i32, y: i32) { - println!("{} % {} = {}", x, y, x % y); -} - -fn main() { - let ints_lower = [ - "isize", "i8", "i16", "i32", "i64", "i128", "usize", "u8", "u16", "u32", "u64", "u128", - ]; - - let ints_upper = [ - "Isize", "I8", "I16", "I32", "I64", "I128", "Usize", "U8", "U16", "U32", "U64", "U128", - ]; - - let can_fail_binops_lower = ["div", "rem", "add", "sub", "mul"]; - - let mut ints_pairs = vec![]; - for i in 0..ints_lower.len() { - ints_pairs.push((&ints_lower[i], &ints_upper[i])); - } - - // Generate the code to print the scalar ranges - for s in &ints_lower { - println!( - "println!(\"let {}_min = Z.of_string \\\"{{}}\\\"\", {}::MIN);", - s, s - ); - println!( - "println!(\"let {}_max = Z.of_string \\\"{{}}\\\"\", {}::MAX);", - s, s - ); - } - println!("\n"); - - // Generate the OCaml definitions for the ranges - this code is - // generated (comes from the above) - println!("let isize_min = Z.of_string \"{}\"", isize::MIN); - println!("let isize_max = Z.of_string \"{}\"", isize::MAX); - println!("let i8_min = Z.of_string \"{}\"", i8::MIN); - println!("let i8_max = Z.of_string \"{}\"", i8::MAX); - println!("let i16_min = Z.of_string \"{}\"", i16::MIN); - println!("let i16_max = Z.of_string \"{}\"", i16::MAX); - println!("let i32_min = Z.of_string \"{}\"", i32::MIN); - println!("let i32_max = Z.of_string \"{}\"", i32::MAX); - println!("let i64_min = Z.of_string \"{}\"", i64::MIN); - println!("let i64_max = Z.of_string \"{}\"", i64::MAX); - println!("let i128_min = Z.of_string \"{}\"", i128::MIN); - println!("let i128_max = Z.of_string \"{}\"", i128::MAX); - println!("let usize_min = Z.of_string \"{}\"", usize::MIN); - println!("let usize_max = Z.of_string \"{}\"", usize::MAX); - println!("let u8_min = Z.of_string \"{}\"", u8::MIN); - println!("let u8_max = Z.of_string \"{}\"", u8::MAX); - println!("let u16_min = Z.of_string \"{}\"", u16::MIN); - println!("let u16_max = Z.of_string \"{}\"", u16::MAX); - println!("let u32_min = Z.of_string \"{}\"", u32::MIN); - println!("let u32_max = Z.of_string \"{}\"", u32::MAX); - println!("let u64_min = Z.of_string \"{}\"", u64::MIN); - println!("let u64_max = Z.of_string \"{}\"", u64::MAX); - println!("let u128_min = Z.of_string \"{}\"", u128::MIN); - println!("let u128_max = Z.of_string \"{}\"", u128::MAX); - println!("\n"); - - // Generate the check_int_in_range body - for (lo, up) in &ints_pairs { - println!("| {} -> Z.leq {}_min i && Z.leq i {}_max", up, lo, lo); - } - println!("\n"); - - // Generate the scalar_value_get_value_range body - for s in &ints_upper { - println!("| {} i -> i", s); - } - println!("\n"); - - // Generate the mk_scalar body - for s in &ints_upper { - println!("| Types.{} -> Ok ({} i)", s, s); - } - println!("\n"); - - // Generate the code to print the scalar ranges in F* - for s in &ints_lower { - println!("println!(\"let {}_min : int = {{}}\", {}::MIN);", s, s); - println!("println!(\"let {}_max : int = {{}}\", {}::MAX);", s, s); - } - println!("\n"); - - // Generate the F* definitions for the ranges - this code is - // generated (comes from the above) - println!("let isize_min : int = {}", isize::MIN); - println!("let isize_max : int = {}", isize::MAX); - println!("let i8_min : int = {}", i8::MIN); - println!("let i8_max : int = {}", i8::MAX); - println!("let i16_min : int = {}", i16::MIN); - println!("let i16_max : int = {}", i16::MAX); - println!("let i32_min : int = {}", i32::MIN); - println!("let i32_max : int = {}", i32::MAX); - println!("let i64_min : int = {}", i64::MIN); - println!("let i64_max : int = {}", i64::MAX); - println!("let i128_min : int = {}", i128::MIN); - println!("let i128_max : int = {}", i128::MAX); - println!("let usize_min : int = {}", usize::MIN); - println!("let usize_max : int = {}", usize::MAX); - println!("let u8_min : int = {}", u8::MIN); - println!("let u8_max : int = {}", u8::MAX); - println!("let u16_min : int = {}", u16::MIN); - println!("let u16_max : int = {}", u16::MAX); - println!("let u32_min : int = {}", u32::MIN); - println!("let u32_max : int = {}", u32::MAX); - println!("let u64_min : int = {}", u64::MIN); - println!("let u64_max : int = {}", u64::MAX); - println!("let u128_min : int = {}", u128::MIN); - println!("let u128_max : int = {}", u128::MAX); - println!("\n"); - - // Generate the body for the ScalarTy definition - for (_lo, up) in &ints_pairs { - println!("| {}", up); - } - println!("\n"); - - // Generate the body for the max/min F* functions - for (lo, up) in &ints_pairs { - println!("| {} -> {}_min", up, lo); - } - println!("\n"); - - // Generate the scalar types for F* - for (lo, up) in &ints_pairs { - println!("type {} = scalar {}", lo, up); - } - println!("\n"); - - // Generate the unops (rk.: we need to manually filter `-` applied on - // unisgned numbers) - for binop in &can_fail_binops_lower { - for (lo, up) in &ints_pairs { - println!("let {}_{} = scalar_{} #{}", lo, binop, binop, up); - } - println!(""); - } - println!("\n"); - - // Modulo tests - test_modulo(1, 2); - test_modulo(-1, 2); - test_modulo(1, -2); - test_modulo(-1, -2); -} -- cgit v1.2.3 From ec039b63748c2a95f89c0538a843e18d3a51cdf3 Mon Sep 17 00:00:00 2001 From: Lucas Franceschino Date: Mon, 3 Jun 2024 16:53:09 +0200 Subject: feat: basic handling for `RValue::Len`, following AeneasVerif/charon#209 --- compiler/FunsAnalysis.ml | 4 +++- compiler/InterpreterExpressions.ml | 1 + compiler/InterpreterStatements.ml | 1 + compiler/InterpreterUtils.ml | 2 +- 4 files changed, 6 insertions(+), 2 deletions(-) diff --git a/compiler/FunsAnalysis.ml b/compiler/FunsAnalysis.ml index eadd8f8a..815c470f 100644 --- a/compiler/FunsAnalysis.ml +++ b/compiler/FunsAnalysis.ml @@ -103,7 +103,9 @@ let analyze_module (m : crate) (funs_map : fun_decl FunDeclId.Map.t) method! visit_rvalue _env rv = match rv with - | Use _ | RvRef _ | Global _ | Discriminant _ | Aggregate _ -> () + | Use _ | RvRef _ | Global _ | Discriminant _ | Aggregate _ | Len _ + -> + () | UnaryOp (uop, _) -> can_fail := unop_can_fail uop || !can_fail | BinaryOp (bop, _, _) -> can_fail := binop_can_fail bop || !can_fail diff --git a/compiler/InterpreterExpressions.ml b/compiler/InterpreterExpressions.ml index 2223897c..56c0ab5f 100644 --- a/compiler/InterpreterExpressions.ml +++ b/compiler/InterpreterExpressions.ml @@ -816,6 +816,7 @@ let eval_rvalue_not_global (config : config) (span : Meta.span) "Unreachable: discriminant reads should have been eliminated from the \ AST" | Global _ -> craise __FILE__ __LINE__ span "Unreachable" + | Len _ -> craise __FILE__ __LINE__ span "Unhandled Len" let eval_fake_read (config : config) (span : Meta.span) (p : place) : cm_fun = fun ctx -> diff --git a/compiler/InterpreterStatements.ml b/compiler/InterpreterStatements.ml index c6a65757..c60be905 100644 --- a/compiler/InterpreterStatements.ml +++ b/compiler/InterpreterStatements.ml @@ -950,6 +950,7 @@ let rec eval_statement (config : config) (st : statement) : stl_cm_fun = let cc = match rvalue with | Global _ -> craise __FILE__ __LINE__ st.span "Unreachable" + | Len _ -> craise __FILE__ __LINE__ st.span "Unhandled Len" | Use _ | RvRef (_, (BShared | BMut | BTwoPhaseMut | BShallow)) | UnaryOp _ | BinaryOp _ | Discriminant _ | Aggregate _ -> diff --git a/compiler/InterpreterUtils.ml b/compiler/InterpreterUtils.ml index 653a0e24..f3f12906 100644 --- a/compiler/InterpreterUtils.ml +++ b/compiler/InterpreterUtils.ml @@ -291,7 +291,7 @@ let rvalue_get_place (rv : rvalue) : place option = match rv with | Use (Copy p | Move p) -> Some p | Use (Constant _) -> None - | RvRef (p, _) -> Some p + | Len (p, _, _) | RvRef (p, _) -> Some p | UnaryOp _ | BinaryOp _ | Global _ | Discriminant _ | Aggregate _ -> None (** See {!ValuesUtils.symbolic_value_has_borrows} *) -- cgit v1.2.3 From ee11f8135281d34e76c31f4d001854ba11546007 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Tue, 4 Jun 2024 09:24:11 +0200 Subject: Update charon pin --- charon-pin | 2 +- flake.lock | 6 +++--- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/charon-pin b/charon-pin index d18033b9..285b511b 100644 --- a/charon-pin +++ b/charon-pin @@ -1,2 +1,2 @@ # This is the commit from https://github.com/AeneasVerif/charon that should be used with this version of aeneas. -a5fda598f359a2b85e044a884fd977d75f4578b4 +ae610b59b337b191d23f4f1c738ed290b8edd0d2 diff --git a/flake.lock b/flake.lock index 712de213..974a3e25 100644 --- a/flake.lock +++ b/flake.lock @@ -9,11 +9,11 @@ "rust-overlay": "rust-overlay" }, "locked": { - "lastModified": 1716552262, - "narHash": "sha256-anBm+nYYp1IDHb+9Dk5oMv28OK/vjcf+wuFI3IRjKrA=", + "lastModified": 1717443886, + "narHash": "sha256-6tX6AgXQlIIuiKLOij0H2mf00rhU02hbQhs3lTARgMk=", "owner": "aeneasverif", "repo": "charon", - "rev": "a5fda598f359a2b85e044a884fd977d75f4578b4", + "rev": "ae610b59b337b191d23f4f1c738ed290b8edd0d2", "type": "github" }, "original": { -- cgit v1.2.3 From d53dcc06e69b6d607261781b4bfc09c52b737a91 Mon Sep 17 00:00:00 2001 From: Nadrieril Date: Tue, 4 Jun 2024 09:29:41 +0200 Subject: Fix the CI pin check --- .github/workflows/ci.yml | 2 ++ 1 file changed, 2 insertions(+) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 5ee28837..17ed0f26 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -59,4 +59,6 @@ jobs: runs-on: [self-hosted, linux, nix] steps: - uses: actions/checkout@v4 + with: + fetch-depth: 0 # deep clone in order to get access to other commits - run: nix develop --command ./scripts/ci-check-charon-pin.sh -- cgit v1.2.3