summaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--Makefile16
-rw-r--r--README.md4
-rw-r--r--backends/lean/.gitignore3
-rw-r--r--backends/lean/Base/Diverge/Base.lean37
-rw-r--r--backends/lean/Base/Diverge/Elab.lean11
-rw-r--r--backends/lean/Base/Diverge/ElabBase.lean16
-rw-r--r--backends/lean/Base/Extensions.lean10
-rw-r--r--backends/lean/Base/IList/IList.lean25
-rw-r--r--backends/lean/Base/Primitives/ArraySlice.lean2
-rw-r--r--backends/lean/Base/Primitives/Scalar.lean111
-rw-r--r--backends/lean/Base/Primitives/Vec.lean2
-rw-r--r--backends/lean/Base/Progress/Base.lean14
-rw-r--r--backends/lean/Base/Utils.lean6
-rw-r--r--backends/lean/lake-manifest.json117
-rw-r--r--backends/lean/lean-toolchain2
-rw-r--r--compiler/Extract.ml2
-rw-r--r--flake.lock36
-rw-r--r--tests/coq/arrays/Arrays.v (renamed from tests/coq/array/Array.v)234
-rw-r--r--tests/coq/arrays/Makefile (renamed from tests/coq/array/Makefile)0
-rw-r--r--tests/coq/arrays/Primitives.v (renamed from tests/coq/array/Primitives.v)0
-rw-r--r--tests/coq/arrays/_CoqProject (renamed from tests/coq/array/_CoqProject)2
-rw-r--r--tests/coq/misc/_CoqProject2
-rw-r--r--tests/fstar-split/arrays/Arrays.Clauses.Template.fst (renamed from tests/fstar-split/array/Array.Clauses.Template.fst)14
-rw-r--r--tests/fstar-split/arrays/Arrays.Clauses.fst (renamed from tests/fstar-split/array/Array.Clauses.fst)10
-rw-r--r--tests/fstar-split/arrays/Arrays.Funs.fst (renamed from tests/fstar-split/array/Array.Funs.fst)264
-rw-r--r--tests/fstar-split/arrays/Arrays.Types.fst (renamed from tests/fstar/array/Array.Types.fst)8
-rw-r--r--tests/fstar-split/arrays/Makefile (renamed from tests/fstar-split/array/Makefile)0
-rw-r--r--tests/fstar-split/arrays/Primitives.fst (renamed from tests/fstar-split/array/Primitives.fst)0
-rw-r--r--tests/fstar/arrays/Arrays.Clauses.Template.fst (renamed from tests/fstar/array/Array.Clauses.Template.fst)14
-rw-r--r--tests/fstar/arrays/Arrays.Clauses.fst (renamed from tests/fstar/array/Array.Clauses.fst)10
-rw-r--r--tests/fstar/arrays/Arrays.Funs.fst (renamed from tests/fstar/array/Array.Funs.fst)232
-rw-r--r--tests/fstar/arrays/Arrays.Types.fst (renamed from tests/fstar-split/array/Array.Types.fst)8
-rw-r--r--tests/fstar/arrays/Makefile (renamed from tests/fstar/array/Makefile)0
-rw-r--r--tests/fstar/arrays/Primitives.fst (renamed from tests/fstar/array/Primitives.fst)0
-rw-r--r--tests/lean/.gitignore3
-rw-r--r--tests/lean/Array/Funs.lean431
-rw-r--r--tests/lean/Array/Types.lean13
-rw-r--r--tests/lean/Arrays.lean (renamed from tests/lean/Array.lean)236
-rw-r--r--tests/lean/Constants.lean36
-rw-r--r--tests/lean/NoNestedBorrows.lean4
-rw-r--r--tests/lean/Traits.lean4
-rw-r--r--tests/lean/Tutorial.lean2
-rw-r--r--tests/lean/lake-manifest.json128
-rw-r--r--tests/lean/lakefile.lean26
-rw-r--r--tests/lean/lean-toolchain2
45 files changed, 851 insertions, 1246 deletions
diff --git a/Makefile b/Makefile
index 8d49a200..45f191cc 100644
--- a/Makefile
+++ b/Makefile
@@ -93,7 +93,7 @@ tests: test-no_nested_borrows test-paper \
testp-polonius_list testp-betree_main \
ctest-testp-betree_main \
test-loops \
- test-array test-traits test-bitwise
+ test-arrays test-traits test-bitwise
# Verify the F* files generated by the translation
.PHONY: verify
@@ -125,13 +125,13 @@ tlean-paper: SUBDIR :=
thol4-no_nested_borrows: SUBDIR := misc-no_nested_borrows
thol4-paper: SUBDIR := misc-paper
-test-array: OPTIONS +=
-test-array: SUBDIR := array
-tfstar-array: OPTIONS += -decreases-clauses -template-clauses -split-files
-tcoq-array: OPTIONS += -use-fuel
-tlean-array: SUBDIR :=
-tlean-array: OPTIONS +=
-thol4-array: OPTIONS +=
+test-arrays: OPTIONS +=
+test-arrays: SUBDIR := arrays
+tfstar-arrays: OPTIONS += -decreases-clauses -template-clauses -split-files
+tcoq-arrays: OPTIONS += -use-fuel
+tlean-arrays: SUBDIR :=
+tlean-arrays: OPTIONS +=
+thol4-arrays: OPTIONS +=
test-traits: OPTIONS +=
test-traits: SUBDIR := traits
diff --git a/README.md b/README.md
index 0530a0da..82ff3944 100644
--- a/README.md
+++ b/README.md
@@ -83,9 +83,9 @@ to display a detailed documentation.
Files generated by the Lean backend import the `Base` package from Aeneas.
To use those files in Lean, create a new Lean package using `lake new`,
overwrite the `lean-toolchain` with the one inside `./backends/lean`,
-and add `Base` as a dependency in the `lakefile.lean`:
+and add `base` as a dependency in the `lakefile.lean`:
```
-require Base from "PATH_TO_AENEAS_REPO/backends/lean"
+require base from "PATH_TO_AENEAS_REPO/backends/lean"
```
## Targeted Subset And Current Limitations
diff --git a/backends/lean/.gitignore b/backends/lean/.gitignore
index 6aef0860..50d5c125 100644
--- a/backends/lean/.gitignore
+++ b/backends/lean/.gitignore
@@ -1,2 +1,3 @@
lake-packages/
-build/ \ No newline at end of file
+build/
+.lake \ No newline at end of file
diff --git a/backends/lean/Base/Diverge/Base.lean b/backends/lean/Base/Diverge/Base.lean
index 9458c926..e40432bd 100644
--- a/backends/lean/Base/Diverge/Base.lean
+++ b/backends/lean/Base/Diverge/Base.lean
@@ -21,7 +21,7 @@ namespace Lemmas
else
f ⟨ m, by simp_all [Nat.lt_iff_le_and_ne] ⟩ ∧
for_all_fin_aux f (m + 1) (by simp_all [Arith.add_one_le_iff_le_ne])
- termination_by for_all_fin_aux n _ m h => n - m
+ termination_by n - m
decreasing_by
simp_wf
apply Nat.sub_add_lt_sub <;> try simp
@@ -240,8 +240,8 @@ namespace Fix
simp [fix]
-- By property of the least upper bound
revert Hd Hl
- -- TODO: there is no conversion to select the head of a function!
- conv => lhs; apply congr_fun; apply congr_fun; apply congr_fun; simp [fix_fuel_P, div?]
+ conv => lhs; rw [fix_fuel_P]
+ simp [div?]
cases fix_fuel (least (fix_fuel_P f x)) f x <;> simp
have Hmono := fix_fuel_mono Hmono Hineq x
simp [result_rel] at Hmono
@@ -255,7 +255,7 @@ namespace Fix
intros x n Hf
have Hfmono := fix_fuel_fix_mono Hmono n x
-- TODO: there is no conversion to select the head of a function!
- conv => apply congr_fun; simp [fix_fuel_P]
+ rw [fix_fuel_P]
simp [fix_fuel_P] at Hf
revert Hf Hfmono
simp [div?, result_rel, fix]
@@ -268,9 +268,7 @@ namespace Fix
fix f x = f (fix f) x := by
have Hl := fix_fuel_P_least Hmono He
-- TODO: better control of simplification
- conv at Hl =>
- apply congr_fun
- simp [fix_fuel_P]
+ rw [fix_fuel_P] at Hl; simp at Hl
-- The least upper bound is > 0
have ⟨ n, Hsucc ⟩ : ∃ n, least (fix_fuel_P f x) = Nat.succ n := by
revert Hl
@@ -618,12 +616,16 @@ namespace FixI
@[simp] theorem is_valid_p_same
(k : ((i:id) → (x:a i) → Result (b i x)) → (i:id) → (x:a i) → Result (b i x)) (x : Result c) :
is_valid_p k (λ _ => x) := by
- simp [is_valid_p, k_to_gen, e_to_gen]
+ simp [is_valid_p]
+ unfold k_to_gen e_to_gen
+ simp
@[simp] theorem is_valid_p_rec
(k : ((i:id) → (x:a i) → Result (b i x)) → (i:id) → (x:a i) → Result (b i x)) (i : id) (x : a i) :
is_valid_p k (λ k => k i x) := by
- simp [is_valid_p, k_to_gen, e_to_gen, kk_to_gen, kk_of_gen]
+ simp [is_valid_p]
+ unfold k_to_gen e_to_gen kk_to_gen kk_of_gen
+ simp
theorem is_valid_p_ite
(k : ((i:id) → (x:a i) → Result (b i x)) → (i:id) → (x:a i) → Result (b i x))
@@ -826,12 +828,16 @@ namespace FixII
@[simp] theorem is_valid_p_same
(k : ((i:id) → (t:ty i) → a i t → Result (b i t)) → (i:id) → (t:ty i) → a i t → Result (b i t)) (x : Result c) :
is_valid_p k (λ _ => x) := by
- simp [is_valid_p, k_to_gen, e_to_gen]
+ simp [is_valid_p]
+ unfold k_to_gen e_to_gen
+ simp
@[simp] theorem is_valid_p_rec
(k : ((i:id) → (t:ty i) → a i t → Result (b i t)) → (i:id) → (t:ty i) → a i t → Result (b i t)) (i : id) (t : ty i) (x : a i t) :
is_valid_p k (λ k => k i t x) := by
- simp [is_valid_p, k_to_gen, e_to_gen, kk_to_gen, kk_of_gen]
+ simp [is_valid_p]
+ unfold k_to_gen e_to_gen kk_to_gen kk_of_gen
+ simp
theorem is_valid_p_ite
(k : ((i:id) → (t:ty i) → a i t → Result (b i t)) → (i:id) → (t:ty i) → a i t → Result (b i t))
@@ -1531,10 +1537,11 @@ namespace Ex9
intro k a x
simp only [id_body]
split <;> try simp
- apply is_valid_p_bind <;> try simp [*]
- -- We have to show that `map k tl` is valid
- -- Remark: `map_is_valid` doesn't work here, we need the specialized version
- apply map_is_valid_simple
+ . apply is_valid_p_same
+ . apply is_valid_p_bind <;> try simp [*]
+ -- We have to show that `map k tl` is valid
+ -- Remark: `map_is_valid` doesn't work here, we need the specialized version
+ apply map_is_valid_simple
def body (k : (i : Fin 1) → (t : ty i) → (x : input_ty i t) → Result (output_ty i t)) (i: Fin 1) :
(t : ty i) → (x : input_ty i t) → Result (output_ty i t) := get_fun bodies i k
diff --git a/backends/lean/Base/Diverge/Elab.lean b/backends/lean/Base/Diverge/Elab.lean
index 6115b13b..3c2ea877 100644
--- a/backends/lean/Base/Diverge/Elab.lean
+++ b/backends/lean/Base/Diverge/Elab.lean
@@ -383,10 +383,7 @@ def mkFin (n : Nat) : Expr :=
def mkFinVal (n i : Nat) : MetaM Expr := do
let n_lit : Expr := .lit (.natVal (n - 1))
let i_lit : Expr := .lit (.natVal i)
- -- We could use `trySynthInstance`, but as we know the instance that we are
- -- going to use, we can save the lookup
- let ofNat ← mkAppOptM ``Fin.instOfNatFinHAddNatInstHAddInstAddNatOfNat #[n_lit, i_lit]
- mkAppOptM ``OfNat.ofNat #[none, none, ofNat]
+ mkAppOptM ``Fin.ofNat #[.some n_lit, .some i_lit]
/- Information about the type of a function in a declaration group.
@@ -654,8 +651,8 @@ partial def proveExprIsValid (k_var kk_var : Expr) (e : Expr) : MetaM Expr := do
-- Normalize to eliminate the lambdas - TODO: this is slightly dangerous.
let e ← do
if e.isLet ∧ normalize_let_bindings then do
- let updt_config config :=
- { config with transparency := .reducible, zetaNonDep := false }
+ let updt_config (config : Lean.Meta.Config) :=
+ { config with transparency := .reducible }
let e ← withConfig updt_config (whnf e)
trace[Diverge.def.valid] "e (after normalization): {e}"
pure e
@@ -929,7 +926,7 @@ partial def proveAppIsValidApplyThms (k_var kk_var : Expr) (e : Expr)
-- We sometimes need to reduce the term - TODO: this is really dangerous
let e ← do
let updt_config config :=
- { config with transparency := .reducible, zetaNonDep := false }
+ { config with transparency := .reducible }
withConfig updt_config (whnf e)
trace[Diverge.def.valid] "e (after normalization): {e}"
let e_valid ← proveExprIsValid k_var kk_var e
diff --git a/backends/lean/Base/Diverge/ElabBase.lean b/backends/lean/Base/Diverge/ElabBase.lean
index 0d33e9d2..08ef96f7 100644
--- a/backends/lean/Base/Diverge/ElabBase.lean
+++ b/backends/lean/Base/Diverge/ElabBase.lean
@@ -27,12 +27,12 @@ initialize registerTraceClass `Diverge.attr
-- divspec attribute
structure DivSpecAttr where
attr : AttributeImpl
- ext : DiscrTreeExtension Name true
+ ext : DiscrTreeExtension Name
deriving Inhabited
/- The persistent map from expressions to divspec theorems. -/
initialize divspecAttr : DivSpecAttr ← do
- let ext ← mkDiscrTreeExtention `divspecMap true
+ let ext ← mkDiscrTreeExtention `divspecMap
let attrImpl : AttributeImpl := {
name := `divspec
descr := "Marks theorems to use with the `divergent` encoding"
@@ -44,7 +44,7 @@ initialize divspecAttr : DivSpecAttr ← do
-- Lookup the theorem
let env ← getEnv
let thDecl := env.constants.find! thName
- let fKey : Array (DiscrTree.Key true) ← MetaM.run' (do
+ let fKey : Array (DiscrTree.Key) ← MetaM.run' (do
/- The theorem should have the shape:
`∀ ..., is_valid_p k (λ k => ...)`
@@ -59,7 +59,9 @@ initialize divspecAttr : DivSpecAttr ← do
let (_, _, fExpr) ← lambdaMetaTelescope fExpr.consumeMData
trace[Diverge] "Registering divspec theorem for {fExpr}"
-- Convert the function expression to a discrimination tree key
- DiscrTree.mkPath fExpr)
+ -- We use the default configuration
+ let config : WhnfCoreConfig := {}
+ DiscrTree.mkPath fExpr config)
let env := ext.addEntry env (fKey, thName)
setEnv env
trace[Diverge] "Saved the environment"
@@ -69,9 +71,11 @@ initialize divspecAttr : DivSpecAttr ← do
pure { attr := attrImpl, ext := ext }
def DivSpecAttr.find? (s : DivSpecAttr) (e : Expr) : MetaM (Array Name) := do
- (s.ext.getState (← getEnv)).getMatch e
+ -- We use the default configuration
+ let config : WhnfCoreConfig := {}
+ (s.ext.getState (← getEnv)).getMatch e config
-def DivSpecAttr.getState (s : DivSpecAttr) : MetaM (DiscrTree Name true) := do
+def DivSpecAttr.getState (s : DivSpecAttr) : MetaM (DiscrTree Name) := do
pure (s.ext.getState (← getEnv))
def showStoredDivSpec : MetaM Unit := do
diff --git a/backends/lean/Base/Extensions.lean b/backends/lean/Base/Extensions.lean
index b34f41dc..c0e80861 100644
--- a/backends/lean/Base/Extensions.lean
+++ b/backends/lean/Base/Extensions.lean
@@ -31,13 +31,13 @@ def mkMapDeclarationExtension [Inhabited α] (name : Name := by exact decl_name%
store the keys from *after* the transformation (i.e., the `DiscrTreeKey`
below). The transformation itself can be done elsewhere.
-/
-abbrev DiscrTreeKey (simpleReduce : Bool) := Array (DiscrTree.Key simpleReduce)
+abbrev DiscrTreeKey := Array DiscrTree.Key
-abbrev DiscrTreeExtension (α : Type) (simpleReduce : Bool) :=
- SimplePersistentEnvExtension (DiscrTreeKey simpleReduce × α) (DiscrTree α simpleReduce)
+abbrev DiscrTreeExtension (α : Type) :=
+ SimplePersistentEnvExtension (DiscrTreeKey × α) (DiscrTree α)
-def mkDiscrTreeExtention [Inhabited α] [BEq α] (name : Name := by exact decl_name%) (simpleReduce : Bool) :
- IO (DiscrTreeExtension α simpleReduce) :=
+def mkDiscrTreeExtention [Inhabited α] [BEq α] (name : Name := by exact decl_name%) :
+ IO (DiscrTreeExtension α) :=
registerSimplePersistentEnvExtension {
name := name,
addImportedFn := fun a => a.foldl (fun s a => a.foldl (fun s (k, v) => s.insertCore k v) s) DiscrTree.empty,
diff --git a/backends/lean/Base/IList/IList.lean b/backends/lean/Base/IList/IList.lean
index e90d1e0d..51457c20 100644
--- a/backends/lean/Base/IList/IList.lean
+++ b/backends/lean/Base/IList/IList.lean
@@ -66,13 +66,15 @@ theorem indexOpt_eq_index [Inhabited α] (ls : List α) (i : Int) :
i < ls.len →
ls.indexOpt i = some (ls.index i) :=
match ls with
- | [] => by simp; intros; linarith
+ | [] => by simp
| hd :: tl =>
if h: i = 0 then
by simp [*]
- else
+ else by
have hi := indexOpt_eq_index tl (i - 1)
- by simp [*]; intros; apply hi <;> int_tac
+ simp [*]; intros
+ -- TODO: there seems to be syntax errors if we don't put the parentheses below??
+ apply hi <;> (int_tac)
-- Remark: the list is unchanged if the index is not in bounds (in particular
-- if it is < 0)
@@ -83,7 +85,7 @@ def update (ls : List α) (i : Int) (y : α) : List α :=
-- Remark: the whole list is dropped if the index is not in bounds (in particular
-- if it is < 0)
-def idrop (i : Int) (ls : List α) : List α :=
+def idrop {α : Type u} (i : Int) (ls : List α) : List α :=
match ls with
| [] => []
| x :: tl => if i = 0 then x :: tl else idrop (i - 1) tl
@@ -117,7 +119,7 @@ variable {α : Type u}
def ireplicate {α : Type u} (i : ℤ) (x : α) : List α :=
if i ≤ 0 then []
else x :: ireplicate (i - 1) x
-termination_by ireplicate i x => i.toNat
+termination_by i.toNat
decreasing_by int_decr_tac
@[simp] theorem update_nil : update ([] : List α) i y = [] := by simp [update]
@@ -137,7 +139,7 @@ decreasing_by int_decr_tac
@[simp] theorem ireplicate_zero : ireplicate 0 x = [] := by rw [ireplicate]; simp
@[simp] theorem ireplicate_nzero_cons (hne : 0 < i) : ireplicate i x = x :: ireplicate (i - 1) x := by
- rw [ireplicate]; simp [*]; intro; linarith
+ rw [ireplicate]; simp [*]
@[simp]
theorem slice_nzero_cons (i j : Int) (x : α) (tl : List α) (hne : i ≠ 0) : slice i j ((x :: tl) : List α) = slice (i - 1) (j - 1) tl :=
@@ -148,11 +150,12 @@ theorem slice_nzero_cons (i j : Int) (x : α) (tl : List α) (hne : i ≠ 0) : s
have : i = 1 := by int_tac
simp [*, slice]
else
- have := slice_nzero_cons (i - 1) (j - 1) hd tl h
+ have hi := slice_nzero_cons (i - 1) (j - 1) hd tl h
by
conv => lhs; simp [slice, *]
- conv at this => lhs; simp [slice, *]
- simp [*, slice]
+ conv at hi => lhs; simp [slice, *]
+ simp [slice]
+ simp [*]
@[simp]
theorem ireplicate_replicate {α : Type u} (l : ℤ) (x : α) (h : 0 ≤ l) :
@@ -166,7 +169,7 @@ theorem ireplicate_replicate {α : Type u} (l : ℤ) (x : α) (h : 0 ≤ l) :
have hl : l.toNat = .succ (l.toNat - 1) := by
cases hl: l.toNat <;> simp_all
conv => rhs; rw[hl]
-termination_by ireplicate_replicate l x h => l.toNat
+termination_by l.toNat
decreasing_by int_decr_tac
@[simp]
@@ -178,7 +181,7 @@ theorem ireplicate_len {α : Type u} (l : ℤ) (x : α) (h : 0 ≤ l) :
have : 0 < l := by int_tac
have hr := ireplicate_len (l - 1) x (by int_tac)
simp [*]
-termination_by ireplicate_len l x h => l.toNat
+termination_by l.toNat
decreasing_by int_decr_tac
theorem len_eq_length (ls : List α) : ls.len = ls.length := by
diff --git a/backends/lean/Base/Primitives/ArraySlice.lean b/backends/lean/Base/Primitives/ArraySlice.lean
index 5057fb01..c90a85b8 100644
--- a/backends/lean/Base/Primitives/ArraySlice.lean
+++ b/backends/lean/Base/Primitives/ArraySlice.lean
@@ -127,7 +127,7 @@ abbrev Slice.v {α : Type u} (v : Slice α) : List α := v.val
example {a: Type u} (v : Slice a) : v.length ≤ Scalar.max ScalarTy.Usize := by
scalar_tac
-def Slice.new (α : Type u): Slice α := ⟨ [], by apply Scalar.cMax_suffices .Usize; simp ⟩
+def Slice.new (α : Type u): Slice α := ⟨ [], by apply Scalar.cMax_suffices .Usize; simp; decide ⟩
-- TODO: very annoying that the α is an explicit parameter
def Slice.len (α : Type u) (v : Slice α) : Usize :=
diff --git a/backends/lean/Base/Primitives/Scalar.lean b/backends/lean/Base/Primitives/Scalar.lean
index fe8dc8ec..b11bd2a1 100644
--- a/backends/lean/Base/Primitives/Scalar.lean
+++ b/backends/lean/Base/Primitives/Scalar.lean
@@ -98,19 +98,19 @@ def Isize.refined_min : { n:Int // n = I32.min ∨ n = I64.min } :=
⟨ Isize.smin, by
simp [Isize.smin]
cases System.Platform.numBits_eq <;>
- unfold System.Platform.numBits at * <;> simp [*] ⟩
+ unfold System.Platform.numBits at * <;> simp [*] <;> decide ⟩
def Isize.refined_max : { n:Int // n = I32.max ∨ n = I64.max } :=
⟨ Isize.smax, by
simp [Isize.smax]
cases System.Platform.numBits_eq <;>
- unfold System.Platform.numBits at * <;> simp [*] ⟩
+ unfold System.Platform.numBits at * <;> simp [*] <;> decide ⟩
def Usize.refined_max : { n:Int // n = U32.max ∨ n = U64.max } :=
⟨ Usize.smax, by
simp [Usize.smax]
cases System.Platform.numBits_eq <;>
- unfold System.Platform.numBits at * <;> simp [*] ⟩
+ unfold System.Platform.numBits at * <;> simp [*] <;> decide ⟩
def Isize.min := Isize.refined_min.val
def Isize.max := Isize.refined_max.val
@@ -231,30 +231,31 @@ def Scalar.cMax (ty : ScalarTy) : Int :=
| _ => Scalar.max ty
theorem Scalar.min_lt_max (ty : ScalarTy) : Scalar.min ty < Scalar.max ty := by
- cases ty <;> simp [Scalar.min, Scalar.max]
+ cases ty <;> simp [Scalar.min, Scalar.max] <;> try decide
. simp [Isize.min, Isize.max]
have h1 := Isize.refined_min.property
have h2 := Isize.refined_max.property
- cases h1 <;> cases h2 <;> simp [*]
+ cases h1 <;> cases h2 <;> simp [*] <;> decide
. simp [Usize.max]
have h := Usize.refined_max.property
- cases h <;> simp [*]
+ cases h <;> simp [*] <;> decide
theorem Scalar.min_le_max (ty : ScalarTy) : Scalar.min ty ≤ Scalar.max ty := by
have := Scalar.min_lt_max ty
int_tac
theorem Scalar.cMin_bound ty : Scalar.min ty ≤ Scalar.cMin ty := by
- cases ty <;> simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at *
+ cases ty <;> (simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at *; try decide)
have h := Isize.refined_min.property
cases h <;> simp [*, Isize.min]
+ decide
theorem Scalar.cMax_bound ty : Scalar.cMax ty ≤ Scalar.max ty := by
- cases ty <;> simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at *
+ cases ty <;> (simp [Scalar.min, Scalar.max, Scalar.cMin, Scalar.cMax] at *; try decide)
. have h := Isize.refined_max.property
- cases h <;> simp [*, Isize.max]
+ cases h <;> simp [*, Isize.max]; decide
. have h := Usize.refined_max.property
- cases h <;> simp [*, Usize.max]
+ cases h <;> simp [*, Usize.max]; decide
theorem Scalar.cMin_suffices ty (h : Scalar.cMin ty ≤ x) : Scalar.min ty ≤ x := by
have := Scalar.cMin_bound ty
@@ -536,12 +537,11 @@ instance {ty} : HAnd (Scalar ty) (Scalar ty) (Scalar ty) where
theorem Scalar.add_spec {ty} {x y : Scalar ty}
(hmin : Scalar.min ty ≤ x.val + y.val)
(hmax : x.val + y.val ≤ Scalar.max ty) :
- ∃ z, x + y = ret z ∧ z.val = x.val + y.val := by
- simp [HAdd.hAdd, add, Add.add]
- simp [tryMk]
+ (∃ z, x + y = ret z ∧ z.val = x.val + y.val) := by
+ -- Applying the unfoldings only on the left
+ conv => congr; ext; lhs; unfold HAdd.hAdd instHAddScalarResult; simp [add, tryMk]
split
- . simp [pure]
- rfl
+ . simp [pure]; rfl
. tauto
theorem Scalar.add_unsigned_spec {ty} (s: ¬ ty.isSigned) {x y : Scalar ty}
@@ -550,33 +550,33 @@ theorem Scalar.add_unsigned_spec {ty} (s: ¬ ty.isSigned) {x y : Scalar ty}
have hmin : Scalar.min ty ≤ x.val + y.val := by
have hx := x.hmin
have hy := y.hmin
- cases ty <;> simp [min] at * <;> linarith
+ cases ty <;> simp [min, ScalarTy.isSigned] at * <;> linarith
apply add_spec <;> assumption
/- Fine-grained theorems -/
@[pspec] theorem Usize.add_spec {x y : Usize} (hmax : x.val + y.val ≤ Usize.max) :
∃ z, x + y = ret z ∧ z.val = x.val + y.val := by
- apply Scalar.add_unsigned_spec <;> simp only [Scalar.max, *]
+ apply Scalar.add_unsigned_spec <;> simp [ScalarTy.isSigned, Scalar.max, *]
@[pspec] theorem U8.add_spec {x y : U8} (hmax : x.val + y.val ≤ U8.max) :
∃ z, x + y = ret z ∧ z.val = x.val + y.val := by
- apply Scalar.add_unsigned_spec <;> simp only [Scalar.max, *]
+ apply Scalar.add_unsigned_spec <;> simp [ScalarTy.isSigned, Scalar.max, *]
@[pspec] theorem U16.add_spec {x y : U16} (hmax : x.val + y.val ≤ U16.max) :
∃ z, x + y = ret z ∧ z.val = x.val + y.val := by
- apply Scalar.add_unsigned_spec <;> simp only [Scalar.max, *]
+ apply Scalar.add_unsigned_spec <;> simp [ScalarTy.isSigned, Scalar.max, *]
@[pspec] theorem U32.add_spec {x y : U32} (hmax : x.val + y.val ≤ U32.max) :
∃ z, x + y = ret z ∧ z.val = x.val + y.val := by
- apply Scalar.add_unsigned_spec <;> simp only [Scalar.max, *]
+ apply Scalar.add_unsigned_spec <;> simp [ScalarTy.isSigned, Scalar.max, *]
@[pspec] theorem U64.add_spec {x y : U64} (hmax : x.val + y.val ≤ U64.max) :
∃ z, x + y = ret z ∧ z.val = x.val + y.val := by
- apply Scalar.add_unsigned_spec <;> simp only [Scalar.max, *]
+ apply Scalar.add_unsigned_spec <;> simp [ScalarTy.isSigned, Scalar.max, *]
@[pspec] theorem U128.add_spec {x y : U128} (hmax : x.val + y.val ≤ U128.max) :
∃ z, x + y = ret z ∧ z.val = x.val + y.val := by
- apply Scalar.add_unsigned_spec <;> simp only [Scalar.max, *]
+ apply Scalar.add_unsigned_spec <;> simp [ScalarTy.isSigned, Scalar.max, *]
@[pspec] theorem Isize.add_spec {x y : Isize}
(hmin : Isize.min ≤ x.val + y.val) (hmax : x.val + y.val ≤ Isize.max) :
@@ -614,48 +614,47 @@ theorem Scalar.sub_spec {ty} {x y : Scalar ty}
(hmin : Scalar.min ty ≤ x.val - y.val)
(hmax : x.val - y.val ≤ Scalar.max ty) :
∃ z, x - y = ret z ∧ z.val = x.val - y.val := by
- simp [HSub.hSub, sub, Sub.sub]
- simp [tryMk]
+ conv => congr; ext; lhs; simp [HSub.hSub, sub, tryMk, Sub.sub]
split
. simp [pure]
rfl
. tauto
-theorem Scalar.sub_unsigned_spec {ty} (s: ¬ ty.isSigned) {x y : Scalar ty}
- (hmin : Scalar.min ty ≤ x.val - y.val) :
+theorem Scalar.sub_unsigned_spec {ty : ScalarTy} (s : ¬ ty.isSigned)
+ {x y : Scalar ty} (hmin : Scalar.min ty ≤ x.val - y.val) :
∃ z, x - y = ret z ∧ z.val = x.val - y.val := by
have : x.val - y.val ≤ Scalar.max ty := by
have hx := x.hmin
have hxm := x.hmax
have hy := y.hmin
- cases ty <;> simp [min, max] at * <;> linarith
+ cases ty <;> simp [min, max, ScalarTy.isSigned] at * <;> linarith
intros
apply sub_spec <;> assumption
/- Fine-grained theorems -/
@[pspec] theorem Usize.sub_spec {x y : Usize} (hmin : Usize.min ≤ x.val - y.val) :
∃ z, x - y = ret z ∧ z.val = x.val - y.val := by
- apply Scalar.sub_unsigned_spec <;> simp only [Scalar.min, *]
+ apply Scalar.sub_unsigned_spec <;> simp_all [Scalar.min, ScalarTy.isSigned]
@[pspec] theorem U8.sub_spec {x y : U8} (hmin : U8.min ≤ x.val - y.val) :
∃ z, x - y = ret z ∧ z.val = x.val - y.val := by
- apply Scalar.sub_unsigned_spec <;> simp only [Scalar.min, *]
+ apply Scalar.sub_unsigned_spec <;> simp_all [Scalar.min, ScalarTy.isSigned]
@[pspec] theorem U16.sub_spec {x y : U16} (hmin : U16.min ≤ x.val - y.val) :
∃ z, x - y = ret z ∧ z.val = x.val - y.val := by
- apply Scalar.sub_unsigned_spec <;> simp only [Scalar.min, *]
+ apply Scalar.sub_unsigned_spec <;> simp_all [Scalar.min, ScalarTy.isSigned]
@[pspec] theorem U32.sub_spec {x y : U32} (hmin : U32.min ≤ x.val - y.val) :
∃ z, x - y = ret z ∧ z.val = x.val - y.val := by
- apply Scalar.sub_unsigned_spec <;> simp only [Scalar.min, *]
+ apply Scalar.sub_unsigned_spec <;> simp_all [Scalar.min, ScalarTy.isSigned]
@[pspec] theorem U64.sub_spec {x y : U64} (hmin : U64.min ≤ x.val - y.val) :
∃ z, x - y = ret z ∧ z.val = x.val - y.val := by
- apply Scalar.sub_unsigned_spec <;> simp only [Scalar.min, *]
+ apply Scalar.sub_unsigned_spec <;> simp_all [Scalar.min, ScalarTy.isSigned]
@[pspec] theorem U128.sub_spec {x y : U128} (hmin : U128.min ≤ x.val - y.val) :
∃ z, x - y = ret z ∧ z.val = x.val - y.val := by
- apply Scalar.sub_unsigned_spec <;> simp only [Scalar.min, *]
+ apply Scalar.sub_unsigned_spec <;> simp_all [Scalar.min, ScalarTy.isSigned]
@[pspec] theorem Isize.sub_spec {x y : Isize} (hmin : Isize.min ≤ x.val - y.val)
(hmax : x.val - y.val ≤ Isize.max) :
@@ -692,8 +691,8 @@ theorem Scalar.mul_spec {ty} {x y : Scalar ty}
(hmin : Scalar.min ty ≤ x.val * y.val)
(hmax : x.val * y.val ≤ Scalar.max ty) :
∃ z, x * y = ret z ∧ z.val = x.val * y.val := by
- simp [HMul.hMul, mul, Mul.mul]
- simp [tryMk]
+ conv => congr; ext; lhs; simp [HMul.hMul]
+ simp [mul, tryMk]
split
. simp [pure]
rfl
@@ -705,33 +704,33 @@ theorem Scalar.mul_unsigned_spec {ty} (s: ¬ ty.isSigned) {x y : Scalar ty}
have : Scalar.min ty ≤ x.val * y.val := by
have hx := x.hmin
have hy := y.hmin
- cases ty <;> simp at * <;> apply mul_nonneg hx hy
+ cases ty <;> simp [ScalarTy.isSigned] at * <;> apply mul_nonneg hx hy
apply mul_spec <;> assumption
/- Fine-grained theorems -/
@[pspec] theorem Usize.mul_spec {x y : Usize} (hmax : x.val * y.val ≤ Usize.max) :
∃ z, x * y = ret z ∧ z.val = x.val * y.val := by
- apply Scalar.mul_unsigned_spec <;> simp only [Scalar.max, *]
+ apply Scalar.mul_unsigned_spec <;> simp_all [Scalar.max, ScalarTy.isSigned]
@[pspec] theorem U8.mul_spec {x y : U8} (hmax : x.val * y.val ≤ U8.max) :
∃ z, x * y = ret z ∧ z.val = x.val * y.val := by
- apply Scalar.mul_unsigned_spec <;> simp only [Scalar.max, *]
+ apply Scalar.mul_unsigned_spec <;> simp_all [Scalar.max, ScalarTy.isSigned]
@[pspec] theorem U16.mul_spec {x y : U16} (hmax : x.val * y.val ≤ U16.max) :
∃ z, x * y = ret z ∧ z.val = x.val * y.val := by
- apply Scalar.mul_unsigned_spec <;> simp only [Scalar.max, *]
+ apply Scalar.mul_unsigned_spec <;> simp_all [Scalar.max, ScalarTy.isSigned]
@[pspec] theorem U32.mul_spec {x y : U32} (hmax : x.val * y.val ≤ U32.max) :
∃ z, x * y = ret z ∧ z.val = x.val * y.val := by
- apply Scalar.mul_unsigned_spec <;> simp only [Scalar.max, *]
+ apply Scalar.mul_unsigned_spec <;> simp_all [Scalar.max, ScalarTy.isSigned]
@[pspec] theorem U64.mul_spec {x y : U64} (hmax : x.val * y.val ≤ U64.max) :
∃ z, x * y = ret z ∧ z.val = x.val * y.val := by
- apply Scalar.mul_unsigned_spec <;> simp only [Scalar.max, *]
+ apply Scalar.mul_unsigned_spec <;> simp_all [Scalar.max, ScalarTy.isSigned]
@[pspec] theorem U128.mul_spec {x y : U128} (hmax : x.val * y.val ≤ U128.max) :
∃ z, x * y = ret z ∧ z.val = x.val * y.val := by
- apply Scalar.mul_unsigned_spec <;> simp only [Scalar.max, *]
+ apply Scalar.mul_unsigned_spec <;> simp_all [Scalar.max, ScalarTy.isSigned]
@[pspec] theorem Isize.mul_spec {x y : Isize} (hmin : Isize.min ≤ x.val * y.val)
(hmax : x.val * y.val ≤ Isize.max) :
@@ -778,7 +777,7 @@ theorem Scalar.div_spec {ty} {x y : Scalar ty}
theorem Scalar.div_unsigned_spec {ty} (s: ¬ ty.isSigned) (x : Scalar ty) {y : Scalar ty}
(hnz : y.val ≠ 0) :
∃ z, x / y = ret z ∧ z.val = x.val / y.val := by
- have h : Scalar.min ty = 0 := by cases ty <;> simp at *
+ have h : Scalar.min ty = 0 := by cases ty <;> simp [ScalarTy.isSigned, min] at *
have hx := x.hmin
have hy := y.hmin
simp [h] at hx hy
@@ -794,27 +793,27 @@ theorem Scalar.div_unsigned_spec {ty} (s: ¬ ty.isSigned) (x : Scalar ty) {y : S
/- Fine-grained theorems -/
@[pspec] theorem Usize.div_spec (x : Usize) {y : Usize} (hnz : y.val ≠ 0) :
∃ z, x / y = ret z ∧ z.val = x.val / y.val := by
- apply Scalar.div_unsigned_spec <;> simp [*]
+ apply Scalar.div_unsigned_spec <;> simp [ScalarTy.isSigned, *]
@[pspec] theorem U8.div_spec (x : U8) {y : U8} (hnz : y.val ≠ 0) :
∃ z, x / y = ret z ∧ z.val = x.val / y.val := by
- apply Scalar.div_unsigned_spec <;> simp [Scalar.max, *]
+ apply Scalar.div_unsigned_spec <;> simp [ScalarTy.isSigned, *]
@[pspec] theorem U16.div_spec (x : U16) {y : U16} (hnz : y.val ≠ 0) :
∃ z, x / y = ret z ∧ z.val = x.val / y.val := by
- apply Scalar.div_unsigned_spec <;> simp [Scalar.max, *]
+ apply Scalar.div_unsigned_spec <;> simp [ScalarTy.isSigned, *]
@[pspec] theorem U32.div_spec (x : U32) {y : U32} (hnz : y.val ≠ 0) :
∃ z, x / y = ret z ∧ z.val = x.val / y.val := by
- apply Scalar.div_unsigned_spec <;> simp [Scalar.max, *]
+ apply Scalar.div_unsigned_spec <;> simp [ScalarTy.isSigned, *]
@[pspec] theorem U64.div_spec (x : U64) {y : U64} (hnz : y.val ≠ 0) :
∃ z, x / y = ret z ∧ z.val = x.val / y.val := by
- apply Scalar.div_unsigned_spec <;> simp [Scalar.max, *]
+ apply Scalar.div_unsigned_spec <;> simp [ScalarTy.isSigned, *]
@[pspec] theorem U128.div_spec (x : U128) {y : U128} (hnz : y.val ≠ 0) :
∃ z, x / y = ret z ∧ z.val = x.val / y.val := by
- apply Scalar.div_unsigned_spec <;> simp [Scalar.max, *]
+ apply Scalar.div_unsigned_spec <;> simp [ScalarTy.isSigned, *]
@[pspec] theorem Isize.div_spec (x : Isize) {y : Isize}
(hnz : y.val ≠ 0)
@@ -873,7 +872,7 @@ theorem Scalar.rem_spec {ty} {x y : Scalar ty}
theorem Scalar.rem_unsigned_spec {ty} (s: ¬ ty.isSigned) (x : Scalar ty) {y : Scalar ty}
(hnz : y.val ≠ 0) :
∃ z, x % y = ret z ∧ z.val = x.val % y.val := by
- have h : Scalar.min ty = 0 := by cases ty <;> simp at *
+ have h : Scalar.min ty = 0 := by cases ty <;> simp [ScalarTy.isSigned, min] at *
have hx := x.hmin
have hy := y.hmin
simp [h] at hx hy
@@ -889,27 +888,27 @@ theorem Scalar.rem_unsigned_spec {ty} (s: ¬ ty.isSigned) (x : Scalar ty) {y : S
@[pspec] theorem Usize.rem_spec (x : Usize) {y : Usize} (hnz : y.val ≠ 0) :
∃ z, x % y = ret z ∧ z.val = x.val % y.val := by
- apply Scalar.rem_unsigned_spec <;> simp [*]
+ apply Scalar.rem_unsigned_spec <;> simp [ScalarTy.isSigned, *]
@[pspec] theorem U8.rem_spec (x : U8) {y : U8} (hnz : y.val ≠ 0) :
∃ z, x % y = ret z ∧ z.val = x.val % y.val := by
- apply Scalar.rem_unsigned_spec <;> simp [Scalar.max, *]
+ apply Scalar.rem_unsigned_spec <;> simp [ScalarTy.isSigned, *]
@[pspec] theorem U16.rem_spec (x : U16) {y : U16} (hnz : y.val ≠ 0) :
∃ z, x % y = ret z ∧ z.val = x.val % y.val := by
- apply Scalar.rem_unsigned_spec <;> simp [Scalar.max, *]
+ apply Scalar.rem_unsigned_spec <;> simp [ScalarTy.isSigned, *]
@[pspec] theorem U32.rem_spec (x : U32) {y : U32} (hnz : y.val ≠ 0) :
∃ z, x % y = ret z ∧ z.val = x.val % y.val := by
- apply Scalar.rem_unsigned_spec <;> simp [Scalar.max, *]
+ apply Scalar.rem_unsigned_spec <;> simp [ScalarTy.isSigned, *]
@[pspec] theorem U64.rem_spec (x : U64) {y : U64} (hnz : y.val ≠ 0) :
∃ z, x % y = ret z ∧ z.val = x.val % y.val := by
- apply Scalar.rem_unsigned_spec <;> simp [Scalar.max, *]
+ apply Scalar.rem_unsigned_spec <;> simp [ScalarTy.isSigned, *]
@[pspec] theorem U128.rem_spec (x : U128) {y : U128} (hnz : y.val ≠ 0) :
∃ z, x % y = ret z ∧ z.val = x.val % y.val := by
- apply Scalar.rem_unsigned_spec <;> simp [Scalar.max, *]
+ apply Scalar.rem_unsigned_spec <;> simp [ScalarTy.isSigned, *]
@[pspec] theorem I8.rem_spec (x : I8) {y : I8}
(hnz : y.val ≠ 0)
diff --git a/backends/lean/Base/Primitives/Vec.lean b/backends/lean/Base/Primitives/Vec.lean
index 12733a34..b03de15b 100644
--- a/backends/lean/Base/Primitives/Vec.lean
+++ b/backends/lean/Base/Primitives/Vec.lean
@@ -35,7 +35,7 @@ abbrev Vec.v {α : Type u} (v : Vec α) : List α := v.val
example {a: Type u} (v : Vec a) : v.length ≤ Scalar.max ScalarTy.Usize := by
scalar_tac
-def Vec.new (α : Type u): Vec α := ⟨ [], by apply Scalar.cMax_suffices .Usize; simp ⟩
+def Vec.new (α : Type u): Vec α := ⟨ [], by apply Scalar.cMax_suffices .Usize; simp; decide ⟩
instance (α : Type u) : Inhabited (Vec α) := by
constructor
diff --git a/backends/lean/Base/Progress/Base.lean b/backends/lean/Base/Progress/Base.lean
index a64212a5..03c80a42 100644
--- a/backends/lean/Base/Progress/Base.lean
+++ b/backends/lean/Base/Progress/Base.lean
@@ -139,12 +139,12 @@ def getPSpecFunArgsExpr (isGoal : Bool) (th : Expr) : MetaM Expr :=
-- pspec attribute
structure PSpecAttr where
attr : AttributeImpl
- ext : DiscrTreeExtension Name true
+ ext : DiscrTreeExtension Name
deriving Inhabited
/- The persistent map from expressions to pspec theorems. -/
initialize pspecAttr : PSpecAttr ← do
- let ext ← mkDiscrTreeExtention `pspecMap true
+ let ext ← mkDiscrTreeExtention `pspecMap
let attrImpl : AttributeImpl := {
name := `pspec
descr := "Marks theorems to use with the `progress` tactic"
@@ -160,7 +160,9 @@ initialize pspecAttr : PSpecAttr ← do
let fExpr ← getPSpecFunArgsExpr false thDecl.type
trace[Progress] "Registering spec theorem for {fExpr}"
-- Convert the function expression to a discrimination tree key
- DiscrTree.mkPath fExpr)
+ -- We use the default configuration
+ let config : WhnfCoreConfig := {}
+ DiscrTree.mkPath fExpr config)
let env := ext.addEntry env (fKey, thName)
setEnv env
trace[Progress] "Saved the environment"
@@ -170,9 +172,11 @@ initialize pspecAttr : PSpecAttr ← do
pure { attr := attrImpl, ext := ext }
def PSpecAttr.find? (s : PSpecAttr) (e : Expr) : MetaM (Array Name) := do
- (s.ext.getState (← getEnv)).getMatch e
+ -- We use the default configuration
+ let config : WhnfCoreConfig := {}
+ (s.ext.getState (← getEnv)).getMatch e config
-def PSpecAttr.getState (s : PSpecAttr) : MetaM (DiscrTree Name true) := do
+def PSpecAttr.getState (s : PSpecAttr) : MetaM (DiscrTree Name) := do
pure (s.ext.getState (← getEnv))
def showStoredPSpec : MetaM Unit := do
diff --git a/backends/lean/Base/Utils.lean b/backends/lean/Base/Utils.lean
index b0032281..eacfe72b 100644
--- a/backends/lean/Base/Utils.lean
+++ b/backends/lean/Base/Utils.lean
@@ -1,6 +1,5 @@
import Lean
import Mathlib.Tactic.Core
-import Mathlib.Tactic.LeftRight
import Base.UtilsBase
/-
@@ -503,9 +502,8 @@ elab "split_disj " n:ident : tactic => do
example (x y : Int) (h0 : x ≤ y ∨ x ≥ y) : x ≤ y ∨ x ≥ y := by
split_disj h0
- . left; assumption
- . right; assumption
-
+ . apply Or.inl; assumption
+ . apply Or.inr; assumption
-- Tactic to split on an exists.
-- `h` must be an FVar
diff --git a/backends/lean/lake-manifest.json b/backends/lean/lake-manifest.json
index 934ee2d9..3a18466f 100644
--- a/backends/lean/lake-manifest.json
+++ b/backends/lean/lake-manifest.json
@@ -1,51 +1,68 @@
-{"version": 5,
- "packagesDir": "lake-packages",
+{"version": 7,
+ "packagesDir": ".lake/packages",
"packages":
- [{"git":
- {"url": "https://github.com/EdAyers/ProofWidgets4",
- "subDir?": null,
- "rev": "a0c2cd0ac3245a0dade4f925bcfa97e06dd84229",
- "opts": {},
- "name": "proofwidgets",
- "inputRev?": "v0.0.13",
- "inherited": true}},
- {"git":
- {"url": "https://github.com/mhuisi/lean4-cli.git",
- "subDir?": null,
- "rev": "21dac2e9cc7e3cf7da5800814787b833e680b2fd",
- "opts": {},
- "name": "Cli",
- "inputRev?": "nightly",
- "inherited": true}},
- {"git":
- {"url": "https://github.com/leanprover-community/mathlib4.git",
- "subDir?": null,
- "rev": "226948a52f8e19ad95ff6025a96784d7e7ed6ed0",
- "opts": {},
- "name": "mathlib",
- "inputRev?": null,
- "inherited": false}},
- {"git":
- {"url": "https://github.com/gebner/quote4",
- "subDir?": null,
- "rev": "e75daed95ad1c92af4e577fea95e234d7a8401c1",
- "opts": {},
- "name": "Qq",
- "inputRev?": "master",
- "inherited": true}},
- {"git":
- {"url": "https://github.com/JLimperg/aesop",
- "subDir?": null,
- "rev": "1a0cded2be292b5496e659b730d2accc742de098",
- "opts": {},
- "name": "aesop",
- "inputRev?": "master",
- "inherited": true}},
- {"git":
- {"url": "https://github.com/leanprover/std4",
- "subDir?": null,
- "rev": "ba5e5e3af519b4fc5221ad0fa4b2c87276f1d323",
- "opts": {},
- "name": "std",
- "inputRev?": "main",
- "inherited": true}}]}
+ [{"url": "https://github.com/leanprover/std4",
+ "type": "git",
+ "subDir": null,
+ "rev": "276953b13323ca151939eafaaec9129bf7970306",
+ "name": "std",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": "main",
+ "inherited": true,
+ "configFile": "lakefile.lean"},
+ {"url": "https://github.com/leanprover-community/quote4",
+ "type": "git",
+ "subDir": null,
+ "rev": "1c88406514a636d241903e2e288d21dc6d861e01",
+ "name": "Qq",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": "master",
+ "inherited": true,
+ "configFile": "lakefile.lean"},
+ {"url": "https://github.com/leanprover-community/aesop",
+ "type": "git",
+ "subDir": null,
+ "rev": "6beed82dcfbb7731d173cd517675df27d62ad0f4",
+ "name": "aesop",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": "master",
+ "inherited": true,
+ "configFile": "lakefile.lean"},
+ {"url": "https://github.com/leanprover-community/ProofWidgets4",
+ "type": "git",
+ "subDir": null,
+ "rev": "af1e86cf7a37389632a02f4a111e6b501b2b818f",
+ "name": "proofwidgets",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": "v0.0.27",
+ "inherited": true,
+ "configFile": "lakefile.lean"},
+ {"url": "https://github.com/leanprover/lean4-cli",
+ "type": "git",
+ "subDir": null,
+ "rev": "a751d21d4b68c999accb6fc5d960538af26ad5ec",
+ "name": "Cli",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": "main",
+ "inherited": true,
+ "configFile": "lakefile.lean"},
+ {"url": "https://github.com/leanprover-community/import-graph.git",
+ "type": "git",
+ "subDir": null,
+ "rev": "8079d2d1d0e073bde42eab159c24f4c2d0d3a871",
+ "name": "importGraph",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": "main",
+ "inherited": true,
+ "configFile": "lakefile.lean"},
+ {"url": "https://github.com/leanprover-community/mathlib4.git",
+ "type": "git",
+ "subDir": null,
+ "rev": "056cc4b21e25e8d1daaeef3a6e3416872c9fc12c",
+ "name": "mathlib",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": null,
+ "inherited": false,
+ "configFile": "lakefile.lean"}],
+ "name": "base",
+ "lakeDir": ".lake"}
diff --git a/backends/lean/lean-toolchain b/backends/lean/lean-toolchain
index fbca4d37..cfcdd327 100644
--- a/backends/lean/lean-toolchain
+++ b/backends/lean/lean-toolchain
@@ -1 +1 @@
-leanprover/lean4:v4.0.0 \ No newline at end of file
+leanprover/lean4:v4.6.0-rc1
diff --git a/compiler/Extract.ml b/compiler/Extract.ml
index 87dcb1fd..6c523549 100644
--- a/compiler/Extract.ml
+++ b/compiler/Extract.ml
@@ -1864,7 +1864,7 @@ let extract_global_decl (ctx : extraction_ctx) (fmt : F.formatter)
let body =
match !backend with
| FStar -> "eval_global " ^ body_name
- | Lean -> "eval_global " ^ body_name ^ " (by simp)"
+ | Lean -> "eval_global " ^ body_name ^ " (by decide)"
| Coq -> body_name ^ "%global"
| HOL4 -> "get_return_value " ^ body_name
in
diff --git a/flake.lock b/flake.lock
index 2442f6ae..f600cf10 100644
--- a/flake.lock
+++ b/flake.lock
@@ -8,11 +8,11 @@
"rust-overlay": "rust-overlay"
},
"locked": {
- "lastModified": 1706179002,
- "narHash": "sha256-sVAG73/MMnGOFdjUvEyEt3BD2gC6H1VhIQBX3VB7H6A=",
+ "lastModified": 1706913319,
+ "narHash": "sha256-ardrxwhlhzWKpc96Pz2UoJTDOFnk3IpjFxoSRyd/cew=",
"owner": "aeneasverif",
"repo": "charon",
- "rev": "9a4ac0c8c88c6778da31177f69b0f93bac66a88b",
+ "rev": "9aedfc390e7418346afdbb66e1d3c14134be6ddb",
"type": "github"
},
"original": {
@@ -131,11 +131,11 @@
"nixpkgs": "nixpkgs_2"
},
"locked": {
- "lastModified": 1706129113,
- "narHash": "sha256-7YW9RkxDfVQFej2Lw4equuAFb5lEWwsNxW/G+fft768=",
+ "lastModified": 1706647851,
+ "narHash": "sha256-tJgVMcCOEqdrgNUHjHgdc3+Spf3vri5Y3Y6noG1mZgo=",
"owner": "fstarlang",
"repo": "fstar",
- "rev": "1be61a27b7413c4e35c1f9affcc6979e8c9a43d6",
+ "rev": "a48e0aa9935c2a22ea540f9f2734c5f847d96361",
"type": "github"
},
"original": {
@@ -164,11 +164,11 @@
]
},
"locked": {
- "lastModified": 1705864452,
- "narHash": "sha256-vjW9bxQ8gm5c0b316NOfjqXaWDLGDCGCnXd6HcvIV+k=",
+ "lastModified": 1706637944,
+ "narHash": "sha256-Cf3kGqFEsOy5Y+2shxN7BC6ADcmWdqs6XZhPdT8sCZI=",
"owner": "hacl-star",
"repo": "hacl-star",
- "rev": "73e719274a8372122994919ae2722a3b1be2bb32",
+ "rev": "513e026e7096639ee99f0c546c99e2f72f86fd6a",
"type": "github"
},
"original": {
@@ -194,11 +194,11 @@
]
},
"locked": {
- "lastModified": 1706145524,
- "narHash": "sha256-gVS1+zqmQa2ghxbPgHG98QmpTqvTB9JM7G9pbe2rLbM=",
+ "lastModified": 1706663512,
+ "narHash": "sha256-37cicQ3mF8PsZe6Lh48o8n+n6vH53Dn7Vap0HOrIBqc=",
"owner": "hacl-star",
"repo": "hacl-nix",
- "rev": "757aa30e65f111714797d8ba37b38cc0f03aa6b2",
+ "rev": "94fafa6c4fdc4769abbf5f24f170e429e11484bd",
"type": "github"
},
"original": {
@@ -265,11 +265,11 @@
"nixpkgs": "nixpkgs_4"
},
"locked": {
- "lastModified": 1706176580,
- "narHash": "sha256-GPvacgrLp/LGt1YU8P40dNnJBCMs376kB7SZAo6MV88=",
+ "lastModified": 1706830172,
+ "narHash": "sha256-QpLi87ZpYxjvyiCaOpE9bTvLEbOShYtpcSa72s/VO4M=",
"owner": "leanprover",
"repo": "lean4",
- "rev": "1f4359cc80d9942d6ee651017cc17dcd62da6595",
+ "rev": "43bbedca46f890e2d2b29d92f71b1e7b76aa0e93",
"type": "github"
},
"original": {
@@ -318,11 +318,11 @@
"nixpkgs": "nixpkgs_7"
},
"locked": {
- "lastModified": 1706176580,
- "narHash": "sha256-GPvacgrLp/LGt1YU8P40dNnJBCMs376kB7SZAo6MV88=",
+ "lastModified": 1706830172,
+ "narHash": "sha256-QpLi87ZpYxjvyiCaOpE9bTvLEbOShYtpcSa72s/VO4M=",
"owner": "leanprover",
"repo": "lean4",
- "rev": "1f4359cc80d9942d6ee651017cc17dcd62da6595",
+ "rev": "43bbedca46f890e2d2b29d92f71b1e7b76aa0e93",
"type": "github"
},
"original": {
diff --git a/tests/coq/array/Array.v b/tests/coq/arrays/Arrays.v
index 3a30413a..3a6fb02f 100644
--- a/tests/coq/array/Array.v
+++ b/tests/coq/arrays/Arrays.v
@@ -1,31 +1,31 @@
(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *)
-(** [array] *)
+(** [arrays] *)
Require Import Primitives.
Import Primitives.
Require Import Coq.ZArith.ZArith.
Require Import List.
Import ListNotations.
Local Open Scope Primitives_scope.
-Module Array.
+Module Arrays.
-(** [array::AB]
- Source: 'src/array.rs', lines 3:0-3:11 *)
+(** [arrays::AB]
+ Source: 'src/arrays.rs', lines 3:0-3:11 *)
Inductive AB_t := | AB_A : AB_t | AB_B : AB_t.
-(** [array::incr]:
- Source: 'src/array.rs', lines 8:0-8:24 *)
+(** [arrays::incr]:
+ Source: 'src/arrays.rs', lines 8:0-8:24 *)
Definition incr (x : u32) : result u32 :=
u32_add x 1%u32.
-(** [array::array_to_shared_slice_]:
- Source: 'src/array.rs', lines 16:0-16:53 *)
+(** [arrays::array_to_shared_slice_]:
+ Source: 'src/arrays.rs', lines 16:0-16:53 *)
Definition array_to_shared_slice_
(T : Type) (s : array T 32%usize) : result (slice T) :=
array_to_slice T 32%usize s
.
-(** [array::array_to_mut_slice_]:
- Source: 'src/array.rs', lines 21:0-21:58 *)
+(** [arrays::array_to_mut_slice_]:
+ Source: 'src/arrays.rs', lines 21:0-21:58 *)
Definition array_to_mut_slice_
(T : Type) (s : array T 32%usize) :
result ((slice T) * (slice T -> result (array T 32%usize)))
@@ -35,45 +35,45 @@ Definition array_to_mut_slice_
Return (s1, to_slice_mut_back)
.
-(** [array::array_len]:
- Source: 'src/array.rs', lines 25:0-25:40 *)
+(** [arrays::array_len]:
+ Source: 'src/arrays.rs', lines 25:0-25:40 *)
Definition array_len (T : Type) (s : array T 32%usize) : result usize :=
s1 <- array_to_slice T 32%usize s; let i := slice_len T s1 in Return i
.
-(** [array::shared_array_len]:
- Source: 'src/array.rs', lines 29:0-29:48 *)
+(** [arrays::shared_array_len]:
+ Source: 'src/arrays.rs', lines 29:0-29:48 *)
Definition shared_array_len (T : Type) (s : array T 32%usize) : result usize :=
s1 <- array_to_slice T 32%usize s; let i := slice_len T s1 in Return i
.
-(** [array::shared_slice_len]:
- Source: 'src/array.rs', lines 33:0-33:44 *)
+(** [arrays::shared_slice_len]:
+ Source: 'src/arrays.rs', lines 33:0-33:44 *)
Definition shared_slice_len (T : Type) (s : slice T) : result usize :=
let i := slice_len T s in Return i
.
-(** [array::index_array_shared]:
- Source: 'src/array.rs', lines 37:0-37:57 *)
+(** [arrays::index_array_shared]:
+ Source: 'src/arrays.rs', lines 37:0-37:57 *)
Definition index_array_shared
(T : Type) (s : array T 32%usize) (i : usize) : result T :=
array_index_usize T 32%usize s i
.
-(** [array::index_array_u32]:
- Source: 'src/array.rs', lines 44:0-44:53 *)
+(** [arrays::index_array_u32]:
+ Source: 'src/arrays.rs', lines 44:0-44:53 *)
Definition index_array_u32 (s : array u32 32%usize) (i : usize) : result u32 :=
array_index_usize u32 32%usize s i
.
-(** [array::index_array_copy]:
- Source: 'src/array.rs', lines 48:0-48:45 *)
+(** [arrays::index_array_copy]:
+ Source: 'src/arrays.rs', lines 48:0-48:45 *)
Definition index_array_copy (x : array u32 32%usize) : result u32 :=
array_index_usize u32 32%usize x 0%usize
.
-(** [array::index_mut_array]:
- Source: 'src/array.rs', lines 52:0-52:62 *)
+(** [arrays::index_mut_array]:
+ Source: 'src/arrays.rs', lines 52:0-52:62 *)
Definition index_mut_array
(T : Type) (s : array T 32%usize) (i : usize) :
result (T * (T -> result (array T 32%usize)))
@@ -83,14 +83,14 @@ Definition index_mut_array
Return (t, index_mut_back)
.
-(** [array::index_slice]:
- Source: 'src/array.rs', lines 56:0-56:46 *)
+(** [arrays::index_slice]:
+ Source: 'src/arrays.rs', lines 56:0-56:46 *)
Definition index_slice (T : Type) (s : slice T) (i : usize) : result T :=
slice_index_usize T s i
.
-(** [array::index_mut_slice]:
- Source: 'src/array.rs', lines 60:0-60:58 *)
+(** [arrays::index_mut_slice]:
+ Source: 'src/arrays.rs', lines 60:0-60:58 *)
Definition index_mut_slice
(T : Type) (s : slice T) (i : usize) :
result (T * (T -> result (slice T)))
@@ -100,8 +100,8 @@ Definition index_mut_slice
Return (t, index_mut_back)
.
-(** [array::slice_subslice_shared_]:
- Source: 'src/array.rs', lines 64:0-64:70 *)
+(** [arrays::slice_subslice_shared_]:
+ Source: 'src/arrays.rs', lines 64:0-64: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)
@@ -109,8 +109,8 @@ Definition slice_subslice_shared_
{| core_ops_range_Range_start := y; core_ops_range_Range_end_ := z |}
.
-(** [array::slice_subslice_mut_]:
- Source: 'src/array.rs', lines 68:0-68:75 *)
+(** [arrays::slice_subslice_mut_]:
+ Source: 'src/arrays.rs', lines 68:0-68:75 *)
Definition slice_subslice_mut_
(x : slice u32) (y : usize) (z : usize) :
result ((slice u32) * (slice u32 -> result (slice u32)))
@@ -123,15 +123,15 @@ Definition slice_subslice_mut_
Return (s, index_mut_back)
.
-(** [array::array_to_slice_shared_]:
- Source: 'src/array.rs', lines 72:0-72:54 *)
+(** [arrays::array_to_slice_shared_]:
+ Source: 'src/arrays.rs', lines 72:0-72:54 *)
Definition array_to_slice_shared_
(x : array u32 32%usize) : result (slice u32) :=
array_to_slice u32 32%usize x
.
-(** [array::array_to_slice_mut_]:
- Source: 'src/array.rs', lines 76:0-76:59 *)
+(** [arrays::array_to_slice_mut_]:
+ Source: 'src/arrays.rs', lines 76:0-76:59 *)
Definition array_to_slice_mut_
(x : array u32 32%usize) :
result ((slice u32) * (slice u32 -> result (array u32 32%usize)))
@@ -141,8 +141,8 @@ Definition array_to_slice_mut_
Return (s, to_slice_mut_back)
.
-(** [array::array_subslice_shared_]:
- Source: 'src/array.rs', lines 80:0-80:74 *)
+(** [arrays::array_subslice_shared_]:
+ Source: 'src/arrays.rs', lines 80:0-80: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
@@ -151,8 +151,8 @@ Definition array_subslice_shared_
{| core_ops_range_Range_start := y; core_ops_range_Range_end_ := z |}
.
-(** [array::array_subslice_mut_]:
- Source: 'src/array.rs', lines 84:0-84:79 *)
+(** [arrays::array_subslice_mut_]:
+ Source: 'src/arrays.rs', lines 84:0-84:79 *)
Definition array_subslice_mut_
(x : array u32 32%usize) (y : usize) (z : usize) :
result ((slice u32) * (slice u32 -> result (array u32 32%usize)))
@@ -166,20 +166,20 @@ Definition array_subslice_mut_
Return (s, index_mut_back)
.
-(** [array::index_slice_0]:
- Source: 'src/array.rs', lines 88:0-88:38 *)
+(** [arrays::index_slice_0]:
+ Source: 'src/arrays.rs', lines 88:0-88:38 *)
Definition index_slice_0 (T : Type) (s : slice T) : result T :=
slice_index_usize T s 0%usize
.
-(** [array::index_array_0]:
- Source: 'src/array.rs', lines 92:0-92:42 *)
+(** [arrays::index_array_0]:
+ Source: 'src/arrays.rs', lines 92:0-92:42 *)
Definition index_array_0 (T : Type) (s : array T 32%usize) : result T :=
array_index_usize T 32%usize s 0%usize
.
-(** [array::index_index_array]:
- Source: 'src/array.rs', lines 103:0-103:71 *)
+(** [arrays::index_index_array]:
+ Source: 'src/arrays.rs', lines 103:0-103:71 *)
Definition index_index_array
(s : array (array u32 32%usize) 32%usize) (i : usize) (j : usize) :
result u32
@@ -188,8 +188,8 @@ Definition index_index_array
array_index_usize u32 32%usize a j
.
-(** [array::update_update_array]:
- Source: 'src/array.rs', lines 114:0-114:70 *)
+(** [arrays::update_update_array]:
+ Source: 'src/arrays.rs', lines 114:0-114:70 *)
Definition update_update_array
(s : array (array u32 32%usize) 32%usize) (i : usize) (j : usize) :
result unit
@@ -203,48 +203,48 @@ Definition update_update_array
Return tt
.
-(** [array::array_local_deep_copy]:
- Source: 'src/array.rs', lines 118:0-118:43 *)
+(** [arrays::array_local_deep_copy]:
+ Source: 'src/arrays.rs', lines 118:0-118:43 *)
Definition array_local_deep_copy (x : array u32 32%usize) : result unit :=
Return tt
.
-(** [array::take_array]:
- Source: 'src/array.rs', lines 122:0-122:30 *)
+(** [arrays::take_array]:
+ Source: 'src/arrays.rs', lines 122:0-122:30 *)
Definition take_array (a : array u32 2%usize) : result unit :=
Return tt.
-(** [array::take_array_borrow]:
- Source: 'src/array.rs', lines 123:0-123:38 *)
+(** [arrays::take_array_borrow]:
+ Source: 'src/arrays.rs', lines 123:0-123:38 *)
Definition take_array_borrow (a : array u32 2%usize) : result unit :=
Return tt
.
-(** [array::take_slice]:
- Source: 'src/array.rs', lines 124:0-124:28 *)
+(** [arrays::take_slice]:
+ Source: 'src/arrays.rs', lines 124:0-124:28 *)
Definition take_slice (s : slice u32) : result unit :=
Return tt.
-(** [array::take_mut_slice]:
- Source: 'src/array.rs', lines 125:0-125:36 *)
+(** [arrays::take_mut_slice]:
+ Source: 'src/arrays.rs', lines 125:0-125:36 *)
Definition take_mut_slice (s : slice u32) : result (slice u32) :=
Return s.
-(** [array::const_array]:
- Source: 'src/array.rs', lines 127:0-127:32 *)
+(** [arrays::const_array]:
+ Source: 'src/arrays.rs', lines 127:0-127:32 *)
Definition const_array : result (array u32 2%usize) :=
Return (mk_array u32 2%usize [ 0%u32; 0%u32 ])
.
-(** [array::const_slice]:
- Source: 'src/array.rs', lines 131:0-131:20 *)
+(** [arrays::const_slice]:
+ Source: 'src/arrays.rs', lines 131:0-131:20 *)
Definition const_slice : result unit :=
_ <- array_to_slice u32 2%usize (mk_array u32 2%usize [ 0%u32; 0%u32 ]);
Return tt
.
-(** [array::take_all]:
- Source: 'src/array.rs', lines 141:0-141:17 *)
+(** [arrays::take_all]:
+ Source: 'src/arrays.rs', lines 141:0-141: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 ]);
@@ -258,33 +258,33 @@ Definition take_all : result unit :=
Return tt
.
-(** [array::index_array]:
- Source: 'src/array.rs', lines 155:0-155:38 *)
+(** [arrays::index_array]:
+ Source: 'src/arrays.rs', lines 155:0-155:38 *)
Definition index_array (x : array u32 2%usize) : result u32 :=
array_index_usize u32 2%usize x 0%usize
.
-(** [array::index_array_borrow]:
- Source: 'src/array.rs', lines 158:0-158:46 *)
+(** [arrays::index_array_borrow]:
+ Source: 'src/arrays.rs', lines 158:0-158:46 *)
Definition index_array_borrow (x : array u32 2%usize) : result u32 :=
array_index_usize u32 2%usize x 0%usize
.
-(** [array::index_slice_u32_0]:
- Source: 'src/array.rs', lines 162:0-162:42 *)
+(** [arrays::index_slice_u32_0]:
+ Source: 'src/arrays.rs', lines 162:0-162:42 *)
Definition index_slice_u32_0 (x : slice u32) : result u32 :=
slice_index_usize u32 x 0%usize
.
-(** [array::index_mut_slice_u32_0]:
- Source: 'src/array.rs', lines 166:0-166:50 *)
+(** [arrays::index_mut_slice_u32_0]:
+ Source: 'src/arrays.rs', lines 166:0-166:50 *)
Definition index_mut_slice_u32_0
(x : slice u32) : result (u32 * (slice u32)) :=
i <- slice_index_usize u32 x 0%usize; Return (i, x)
.
-(** [array::index_all]:
- Source: 'src/array.rs', lines 170:0-170:25 *)
+(** [arrays::index_all]:
+ Source: 'src/arrays.rs', lines 170:0-170: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 ]);
@@ -303,8 +303,8 @@ Definition index_all : result u32 :=
Return i8
.
-(** [array::update_array]:
- Source: 'src/array.rs', lines 184:0-184:36 *)
+(** [arrays::update_array]:
+ Source: 'src/arrays.rs', lines 184:0-184: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
@@ -312,8 +312,8 @@ Definition update_array (x : array u32 2%usize) : result unit :=
Return tt
.
-(** [array::update_array_mut_borrow]:
- Source: 'src/array.rs', lines 187:0-187:48 *)
+(** [arrays::update_array_mut_borrow]:
+ Source: 'src/arrays.rs', lines 187:0-187: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;
@@ -321,16 +321,16 @@ Definition update_array_mut_borrow
index_mut_back 1%u32
.
-(** [array::update_mut_slice]:
- Source: 'src/array.rs', lines 190:0-190:38 *)
+(** [arrays::update_mut_slice]:
+ Source: 'src/arrays.rs', lines 190:0-190: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
index_mut_back 1%u32
.
-(** [array::update_all]:
- Source: 'src/array.rs', lines 194:0-194:19 *)
+(** [arrays::update_all]:
+ Source: 'src/arrays.rs', lines 194:0-194: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 ]);
@@ -342,8 +342,8 @@ Definition update_all : result unit :=
Return tt
.
-(** [array::range_all]:
- Source: 'src/array.rs', lines 205:0-205:18 *)
+(** [arrays::range_all]:
+ Source: 'src/arrays.rs', lines 205:0-205:18 *)
Definition range_all : result unit :=
p <-
core_array_Array_index_mut u32 (core_ops_range_Range usize) 4%usize
@@ -360,32 +360,32 @@ Definition range_all : result unit :=
Return tt
.
-(** [array::deref_array_borrow]:
- Source: 'src/array.rs', lines 214:0-214:46 *)
+(** [arrays::deref_array_borrow]:
+ Source: 'src/arrays.rs', lines 214:0-214:46 *)
Definition deref_array_borrow (x : array u32 2%usize) : result u32 :=
array_index_usize u32 2%usize x 0%usize
.
-(** [array::deref_array_mut_borrow]:
- Source: 'src/array.rs', lines 219:0-219:54 *)
+(** [arrays::deref_array_mut_borrow]:
+ Source: 'src/arrays.rs', lines 219:0-219: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; Return (i, x)
.
-(** [array::take_array_t]:
- Source: 'src/array.rs', lines 227:0-227:31 *)
+(** [arrays::take_array_t]:
+ Source: 'src/arrays.rs', lines 227:0-227:31 *)
Definition take_array_t (a : array AB_t 2%usize) : result unit :=
Return tt.
-(** [array::non_copyable_array]:
- Source: 'src/array.rs', lines 229:0-229:27 *)
+(** [arrays::non_copyable_array]:
+ Source: 'src/arrays.rs', lines 229:0-229:27 *)
Definition non_copyable_array : result unit :=
_ <- take_array_t (mk_array AB_t 2%usize [ AB_A; AB_B ]); Return tt
.
-(** [array::sum]: loop 0:
- Source: 'src/array.rs', lines 242:0-250:1 *)
+(** [arrays::sum]: loop 0:
+ Source: 'src/arrays.rs', lines 242:0-250:1 *)
Fixpoint sum_loop
(n : nat) (s : slice u32) (sum1 : u32) (i : usize) : result u32 :=
match n with
@@ -402,14 +402,14 @@ Fixpoint sum_loop
end
.
-(** [array::sum]:
- Source: 'src/array.rs', lines 242:0-242:28 *)
+(** [arrays::sum]:
+ Source: 'src/arrays.rs', lines 242:0-242:28 *)
Definition sum (n : nat) (s : slice u32) : result u32 :=
sum_loop n s 0%u32 0%usize
.
-(** [array::sum2]: loop 0:
- Source: 'src/array.rs', lines 252:0-261:1 *)
+(** [arrays::sum2]: loop 0:
+ Source: 'src/arrays.rs', lines 252:0-261:1 *)
Fixpoint sum2_loop
(n : nat) (s : slice u32) (s2 : slice u32) (sum1 : u32) (i : usize) :
result u32
@@ -430,16 +430,16 @@ Fixpoint sum2_loop
end
.
-(** [array::sum2]:
- Source: 'src/array.rs', lines 252:0-252:41 *)
+(** [arrays::sum2]:
+ Source: 'src/arrays.rs', lines 252:0-252: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
if negb (i s= i1) then Fail_ Failure else sum2_loop n s s2 0%u32 0%usize
.
-(** [array::f0]:
- Source: 'src/array.rs', lines 263:0-263:11 *)
+(** [arrays::f0]:
+ Source: 'src/arrays.rs', lines 263:0-263: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
@@ -450,8 +450,8 @@ Definition f0 : result unit :=
Return tt
.
-(** [array::f1]:
- Source: 'src/array.rs', lines 268:0-268:11 *)
+(** [arrays::f1]:
+ Source: 'src/arrays.rs', lines 268:0-268:11 *)
Definition f1 : result unit :=
p <-
array_index_mut_usize u32 2%usize (mk_array u32 2%usize [ 1%u32; 2%u32 ])
@@ -461,13 +461,13 @@ Definition f1 : result unit :=
Return tt
.
-(** [array::f2]:
- Source: 'src/array.rs', lines 273:0-273:17 *)
+(** [arrays::f2]:
+ Source: 'src/arrays.rs', lines 273:0-273:17 *)
Definition f2 (i : u32) : result unit :=
Return tt.
-(** [array::f4]:
- Source: 'src/array.rs', lines 282:0-282:54 *)
+(** [arrays::f4]:
+ Source: 'src/arrays.rs', lines 282:0-282: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
@@ -476,8 +476,8 @@ Definition f4
{| core_ops_range_Range_start := y; core_ops_range_Range_end_ := z |}
.
-(** [array::f3]:
- Source: 'src/array.rs', lines 275:0-275:18 *)
+(** [arrays::f3]:
+ Source: 'src/arrays.rs', lines 275:0-275:18 *)
Definition f3 (n : nat) : result u32 :=
i <-
array_index_usize u32 2%usize (mk_array u32 2%usize [ 1%u32; 2%u32 ])
@@ -489,19 +489,19 @@ Definition f3 (n : nat) : result u32 :=
sum2 n s s1
.
-(** [array::SZ]
- Source: 'src/array.rs', lines 286:0-286:19 *)
+(** [arrays::SZ]
+ Source: 'src/arrays.rs', lines 286:0-286:19 *)
Definition sz_body : result usize := Return 32%usize.
Definition sz_c : usize := sz_body%global.
-(** [array::f5]:
- Source: 'src/array.rs', lines 289:0-289:31 *)
+(** [arrays::f5]:
+ Source: 'src/arrays.rs', lines 289:0-289:31 *)
Definition f5 (x : array u32 32%usize) : result u32 :=
array_index_usize u32 32%usize x 0%usize
.
-(** [array::ite]:
- Source: 'src/array.rs', lines 294:0-294:12 *)
+(** [arrays::ite]:
+ Source: 'src/arrays.rs', lines 294:0-294: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
@@ -516,4 +516,4 @@ Definition ite : result unit :=
Return tt
.
-End Array.
+End Arrays.
diff --git a/tests/coq/array/Makefile b/tests/coq/arrays/Makefile
index 1a5aee4a..1a5aee4a 100644
--- a/tests/coq/array/Makefile
+++ b/tests/coq/arrays/Makefile
diff --git a/tests/coq/array/Primitives.v b/tests/coq/arrays/Primitives.v
index 990e27e4..990e27e4 100644
--- a/tests/coq/array/Primitives.v
+++ b/tests/coq/arrays/Primitives.v
diff --git a/tests/coq/array/_CoqProject b/tests/coq/arrays/_CoqProject
index 87d8fc3d..a4e82408 100644
--- a/tests/coq/array/_CoqProject
+++ b/tests/coq/arrays/_CoqProject
@@ -3,5 +3,5 @@
-arg -w
-arg all
+Arrays.v
Primitives.v
-Array.v
diff --git a/tests/coq/misc/_CoqProject b/tests/coq/misc/_CoqProject
index 64cddedd..869cdb4d 100644
--- a/tests/coq/misc/_CoqProject
+++ b/tests/coq/misc/_CoqProject
@@ -8,9 +8,9 @@ External_Types.v
Primitives.v
External_Funs.v
External_TypesExternal.v
-Paper.v
Constants.v
PoloniusList.v
+Paper.v
NoNestedBorrows.v
External_FunsExternal.v
Bitwise.v
diff --git a/tests/fstar-split/array/Array.Clauses.Template.fst b/tests/fstar-split/arrays/Arrays.Clauses.Template.fst
index b2f2649c..8cc32583 100644
--- a/tests/fstar-split/array/Array.Clauses.Template.fst
+++ b/tests/fstar-split/arrays/Arrays.Clauses.Template.fst
@@ -1,19 +1,19 @@
(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *)
-(** [array]: templates for the decreases clauses *)
-module Array.Clauses.Template
+(** [arrays]: templates for the decreases clauses *)
+module Arrays.Clauses.Template
open Primitives
-open Array.Types
+open Arrays.Types
#set-options "--z3rlimit 50 --fuel 1 --ifuel 1"
-(** [array::sum]: decreases clause
- Source: 'src/array.rs', lines 242:0-250:1 *)
+(** [arrays::sum]: decreases clause
+ Source: 'src/arrays.rs', lines 242:0-250:1 *)
unfold
let sum_loop_decreases (s : slice u32) (sum1 : u32) (i : usize) : nat =
admit ()
-(** [array::sum2]: decreases clause
- Source: 'src/array.rs', lines 252:0-261:1 *)
+(** [arrays::sum2]: decreases clause
+ Source: 'src/arrays.rs', lines 252:0-261:1 *)
unfold
let sum2_loop_decreases (s : slice u32) (s2 : slice u32) (sum1 : u32)
(i : usize) : nat =
diff --git a/tests/fstar-split/array/Array.Clauses.fst b/tests/fstar-split/arrays/Arrays.Clauses.fst
index 68cbf216..aca328c2 100644
--- a/tests/fstar-split/array/Array.Clauses.fst
+++ b/tests/fstar-split/arrays/Arrays.Clauses.fst
@@ -1,17 +1,17 @@
-(** [array]: decreases clauses *)
-module Array.Clauses
+(** [arrays]: decreases clauses *)
+module Arrays.Clauses
open Primitives
-open Array.Types
+open Arrays.Types
open FStar.List.Tot
#set-options "--z3rlimit 50 --fuel 1 --ifuel 1"
-(** [array::sum]: decreases clause *)
+(** [arrays::sum]: decreases clause *)
unfold
let sum_loop_decreases (s : slice u32) (sum : u32) (i : usize) : nat =
if i < length s then length s - i else 0
-(** [array::sum2]: decreases clause *)
+(** [arrays::sum2]: decreases clause *)
unfold
let sum2_loop_decreases (s : slice u32) (s2 : slice u32) (sum : u32)
(i : usize) : nat =
diff --git a/tests/fstar-split/array/Array.Funs.fst b/tests/fstar-split/arrays/Arrays.Funs.fst
index 30b19702..3efe7789 100644
--- a/tests/fstar-split/array/Array.Funs.fst
+++ b/tests/fstar-split/arrays/Arrays.Funs.fst
@@ -1,109 +1,109 @@
(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *)
-(** [array]: function definitions *)
-module Array.Funs
+(** [arrays]: function definitions *)
+module Arrays.Funs
open Primitives
-include Array.Types
-include Array.Clauses
+include Arrays.Types
+include Arrays.Clauses
#set-options "--z3rlimit 50 --fuel 1 --ifuel 1"
-(** [array::incr]: merged forward/backward function
+(** [arrays::incr]: merged forward/backward function
(there is a single backward function, and the forward function returns ())
- Source: 'src/array.rs', lines 8:0-8:24 *)
+ Source: 'src/arrays.rs', lines 8:0-8:24 *)
let incr (x : u32) : result u32 =
u32_add x 1
-(** [array::array_to_shared_slice_]: forward function
- Source: 'src/array.rs', lines 16:0-16:53 *)
+(** [arrays::array_to_shared_slice_]: forward function
+ Source: 'src/arrays.rs', lines 16:0-16:53 *)
let array_to_shared_slice_ (t : Type0) (s : array t 32) : result (slice t) =
array_to_slice t 32 s
-(** [array::array_to_mut_slice_]: forward function
- Source: 'src/array.rs', lines 21:0-21:58 *)
+(** [arrays::array_to_mut_slice_]: forward function
+ Source: 'src/arrays.rs', lines 21:0-21:58 *)
let array_to_mut_slice_ (t : Type0) (s : array t 32) : result (slice t) =
array_to_slice t 32 s
-(** [array::array_to_mut_slice_]: backward function 0
- Source: 'src/array.rs', lines 21:0-21:58 *)
+(** [arrays::array_to_mut_slice_]: backward function 0
+ Source: 'src/arrays.rs', lines 21:0-21:58 *)
let array_to_mut_slice__back
(t : Type0) (s : array t 32) (ret : slice t) : result (array t 32) =
array_from_slice t 32 s ret
-(** [array::array_len]: forward function
- Source: 'src/array.rs', lines 25:0-25:40 *)
+(** [arrays::array_len]: forward function
+ Source: 'src/arrays.rs', lines 25:0-25:40 *)
let array_len (t : Type0) (s : array t 32) : result usize =
let* s1 = array_to_slice t 32 s in let i = slice_len t s1 in Return i
-(** [array::shared_array_len]: forward function
- Source: 'src/array.rs', lines 29:0-29:48 *)
+(** [arrays::shared_array_len]: forward function
+ Source: 'src/arrays.rs', lines 29:0-29:48 *)
let shared_array_len (t : Type0) (s : array t 32) : result usize =
let* s1 = array_to_slice t 32 s in let i = slice_len t s1 in Return i
-(** [array::shared_slice_len]: forward function
- Source: 'src/array.rs', lines 33:0-33:44 *)
+(** [arrays::shared_slice_len]: forward function
+ Source: 'src/arrays.rs', lines 33:0-33:44 *)
let shared_slice_len (t : Type0) (s : slice t) : result usize =
let i = slice_len t s in Return i
-(** [array::index_array_shared]: forward function
- Source: 'src/array.rs', lines 37:0-37:57 *)
+(** [arrays::index_array_shared]: forward function
+ Source: 'src/arrays.rs', lines 37:0-37:57 *)
let index_array_shared (t : Type0) (s : array t 32) (i : usize) : result t =
array_index_usize t 32 s i
-(** [array::index_array_u32]: forward function
- Source: 'src/array.rs', lines 44:0-44:53 *)
+(** [arrays::index_array_u32]: forward function
+ Source: 'src/arrays.rs', lines 44:0-44:53 *)
let index_array_u32 (s : array u32 32) (i : usize) : result u32 =
array_index_usize u32 32 s i
-(** [array::index_array_copy]: forward function
- Source: 'src/array.rs', lines 48:0-48:45 *)
+(** [arrays::index_array_copy]: forward function
+ Source: 'src/arrays.rs', lines 48:0-48:45 *)
let index_array_copy (x : array u32 32) : result u32 =
array_index_usize u32 32 x 0
-(** [array::index_mut_array]: forward function
- Source: 'src/array.rs', lines 52:0-52:62 *)
+(** [arrays::index_mut_array]: forward function
+ Source: 'src/arrays.rs', lines 52:0-52:62 *)
let index_mut_array (t : Type0) (s : array t 32) (i : usize) : result t =
array_index_usize t 32 s i
-(** [array::index_mut_array]: backward function 0
- Source: 'src/array.rs', lines 52:0-52:62 *)
+(** [arrays::index_mut_array]: backward function 0
+ Source: 'src/arrays.rs', lines 52:0-52:62 *)
let index_mut_array_back
(t : Type0) (s : array t 32) (i : usize) (ret : t) : result (array t 32) =
array_update_usize t 32 s i ret
-(** [array::index_slice]: forward function
- Source: 'src/array.rs', lines 56:0-56:46 *)
+(** [arrays::index_slice]: forward function
+ Source: 'src/arrays.rs', lines 56:0-56:46 *)
let index_slice (t : Type0) (s : slice t) (i : usize) : result t =
slice_index_usize t s i
-(** [array::index_mut_slice]: forward function
- Source: 'src/array.rs', lines 60:0-60:58 *)
+(** [arrays::index_mut_slice]: forward function
+ Source: 'src/arrays.rs', lines 60:0-60:58 *)
let index_mut_slice (t : Type0) (s : slice t) (i : usize) : result t =
slice_index_usize t s i
-(** [array::index_mut_slice]: backward function 0
- Source: 'src/array.rs', lines 60:0-60:58 *)
+(** [arrays::index_mut_slice]: backward function 0
+ Source: 'src/arrays.rs', lines 60:0-60:58 *)
let index_mut_slice_back
(t : Type0) (s : slice t) (i : usize) (ret : t) : result (slice t) =
slice_update_usize t s i ret
-(** [array::slice_subslice_shared_]: forward function
- Source: 'src/array.rs', lines 64:0-64:70 *)
+(** [arrays::slice_subslice_shared_]: forward function
+ Source: 'src/arrays.rs', lines 64:0-64: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)
(core_slice_index_SliceIndexRangeUsizeSliceTInst u32) x
{ start = y; end_ = z }
-(** [array::slice_subslice_mut_]: forward function
- Source: 'src/array.rs', lines 68:0-68:75 *)
+(** [arrays::slice_subslice_mut_]: forward function
+ Source: 'src/arrays.rs', lines 68:0-68:75 *)
let slice_subslice_mut_
(x : slice u32) (y : usize) (z : usize) : result (slice u32) =
core_slice_index_Slice_index_mut u32 (core_ops_range_Range usize)
(core_slice_index_SliceIndexRangeUsizeSliceTInst u32) x
{ start = y; end_ = z }
-(** [array::slice_subslice_mut_]: backward function 0
- Source: 'src/array.rs', lines 68:0-68:75 *)
+(** [arrays::slice_subslice_mut_]: backward function 0
+ Source: 'src/arrays.rs', lines 68:0-68:75 *)
let slice_subslice_mut__back
(x : slice u32) (y : usize) (z : usize) (ret : slice u32) :
result (slice u32)
@@ -112,24 +112,24 @@ let slice_subslice_mut__back
(core_slice_index_SliceIndexRangeUsizeSliceTInst u32) x
{ start = y; end_ = z } ret
-(** [array::array_to_slice_shared_]: forward function
- Source: 'src/array.rs', lines 72:0-72:54 *)
+(** [arrays::array_to_slice_shared_]: forward function
+ Source: 'src/arrays.rs', lines 72:0-72:54 *)
let array_to_slice_shared_ (x : array u32 32) : result (slice u32) =
array_to_slice u32 32 x
-(** [array::array_to_slice_mut_]: forward function
- Source: 'src/array.rs', lines 76:0-76:59 *)
+(** [arrays::array_to_slice_mut_]: forward function
+ Source: 'src/arrays.rs', lines 76:0-76:59 *)
let array_to_slice_mut_ (x : array u32 32) : result (slice u32) =
array_to_slice u32 32 x
-(** [array::array_to_slice_mut_]: backward function 0
- Source: 'src/array.rs', lines 76:0-76:59 *)
+(** [arrays::array_to_slice_mut_]: backward function 0
+ Source: 'src/arrays.rs', lines 76:0-76:59 *)
let array_to_slice_mut__back
(x : array u32 32) (ret : slice u32) : result (array u32 32) =
array_from_slice u32 32 x ret
-(** [array::array_subslice_shared_]: forward function
- Source: 'src/array.rs', lines 80:0-80:74 *)
+(** [arrays::array_subslice_shared_]: forward function
+ Source: 'src/arrays.rs', lines 80:0-80: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
@@ -137,8 +137,8 @@ let array_subslice_shared_
(core_slice_index_SliceIndexRangeUsizeSliceTInst u32)) x
{ start = y; end_ = z }
-(** [array::array_subslice_mut_]: forward function
- Source: 'src/array.rs', lines 84:0-84:79 *)
+(** [arrays::array_subslice_mut_]: forward function
+ Source: 'src/arrays.rs', lines 84:0-84:79 *)
let array_subslice_mut_
(x : array u32 32) (y : usize) (z : usize) : result (slice u32) =
core_array_Array_index_mut u32 (core_ops_range_Range usize) 32
@@ -146,8 +146,8 @@ let array_subslice_mut_
(core_slice_index_SliceIndexRangeUsizeSliceTInst u32)) x
{ start = y; end_ = z }
-(** [array::array_subslice_mut_]: backward function 0
- Source: 'src/array.rs', lines 84:0-84:79 *)
+(** [arrays::array_subslice_mut_]: backward function 0
+ Source: 'src/arrays.rs', lines 84:0-84:79 *)
let array_subslice_mut__back
(x : array u32 32) (y : usize) (z : usize) (ret : slice u32) :
result (array u32 32)
@@ -157,25 +157,25 @@ let array_subslice_mut__back
(core_slice_index_SliceIndexRangeUsizeSliceTInst u32)) x
{ start = y; end_ = z } ret
-(** [array::index_slice_0]: forward function
- Source: 'src/array.rs', lines 88:0-88:38 *)
+(** [arrays::index_slice_0]: forward function
+ Source: 'src/arrays.rs', lines 88:0-88:38 *)
let index_slice_0 (t : Type0) (s : slice t) : result t =
slice_index_usize t s 0
-(** [array::index_array_0]: forward function
- Source: 'src/array.rs', lines 92:0-92:42 *)
+(** [arrays::index_array_0]: forward function
+ Source: 'src/arrays.rs', lines 92:0-92:42 *)
let index_array_0 (t : Type0) (s : array t 32) : result t =
array_index_usize t 32 s 0
-(** [array::index_index_array]: forward function
- Source: 'src/array.rs', lines 103:0-103:71 *)
+(** [arrays::index_index_array]: forward function
+ Source: 'src/arrays.rs', lines 103:0-103: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
-(** [array::update_update_array]: forward function
- Source: 'src/array.rs', lines 114:0-114:70 *)
+(** [arrays::update_update_array]: forward function
+ Source: 'src/arrays.rs', lines 114:0-114:70 *)
let update_update_array
(s : array (array u32 32) 32) (i : usize) (j : usize) : result unit =
let* a = array_index_usize (array u32 32) 32 s i in
@@ -183,44 +183,44 @@ let update_update_array
let* _ = array_update_usize (array u32 32) 32 s i a1 in
Return ()
-(** [array::array_local_deep_copy]: forward function
- Source: 'src/array.rs', lines 118:0-118:43 *)
+(** [arrays::array_local_deep_copy]: forward function
+ Source: 'src/arrays.rs', lines 118:0-118:43 *)
let array_local_deep_copy (x : array u32 32) : result unit =
Return ()
-(** [array::take_array]: forward function
- Source: 'src/array.rs', lines 122:0-122:30 *)
+(** [arrays::take_array]: forward function
+ Source: 'src/arrays.rs', lines 122:0-122:30 *)
let take_array (a : array u32 2) : result unit =
Return ()
-(** [array::take_array_borrow]: forward function
- Source: 'src/array.rs', lines 123:0-123:38 *)
+(** [arrays::take_array_borrow]: forward function
+ Source: 'src/arrays.rs', lines 123:0-123:38 *)
let take_array_borrow (a : array u32 2) : result unit =
Return ()
-(** [array::take_slice]: forward function
- Source: 'src/array.rs', lines 124:0-124:28 *)
+(** [arrays::take_slice]: forward function
+ Source: 'src/arrays.rs', lines 124:0-124:28 *)
let take_slice (s : slice u32) : result unit =
Return ()
-(** [array::take_mut_slice]: merged forward/backward function
+(** [arrays::take_mut_slice]: merged forward/backward function
(there is a single backward function, and the forward function returns ())
- Source: 'src/array.rs', lines 125:0-125:36 *)
+ Source: 'src/arrays.rs', lines 125:0-125:36 *)
let take_mut_slice (s : slice u32) : result (slice u32) =
Return s
-(** [array::const_array]: forward function
- Source: 'src/array.rs', lines 127:0-127:32 *)
+(** [arrays::const_array]: forward function
+ Source: 'src/arrays.rs', lines 127:0-127:32 *)
let const_array : result (array u32 2) =
Return (mk_array u32 2 [ 0; 0 ])
-(** [array::const_slice]: forward function
- Source: 'src/array.rs', lines 131:0-131:20 *)
+(** [arrays::const_slice]: forward function
+ Source: 'src/arrays.rs', lines 131:0-131:20 *)
let const_slice : result unit =
let* _ = array_to_slice u32 2 (mk_array u32 2 [ 0; 0 ]) in Return ()
-(** [array::take_all]: forward function
- Source: 'src/array.rs', lines 141:0-141:17 *)
+(** [arrays::take_all]: forward function
+ Source: 'src/arrays.rs', lines 141:0-141:17 *)
let take_all : result unit =
let* _ = take_array (mk_array u32 2 [ 0; 0 ]) in
let* _ = take_array_borrow (mk_array u32 2 [ 0; 0 ]) in
@@ -231,33 +231,33 @@ let take_all : result unit =
let* _ = array_from_slice u32 2 (mk_array u32 2 [ 0; 0 ]) s2 in
Return ()
-(** [array::index_array]: forward function
- Source: 'src/array.rs', lines 155:0-155:38 *)
+(** [arrays::index_array]: forward function
+ Source: 'src/arrays.rs', lines 155:0-155:38 *)
let index_array (x : array u32 2) : result u32 =
array_index_usize u32 2 x 0
-(** [array::index_array_borrow]: forward function
- Source: 'src/array.rs', lines 158:0-158:46 *)
+(** [arrays::index_array_borrow]: forward function
+ Source: 'src/arrays.rs', lines 158:0-158:46 *)
let index_array_borrow (x : array u32 2) : result u32 =
array_index_usize u32 2 x 0
-(** [array::index_slice_u32_0]: forward function
- Source: 'src/array.rs', lines 162:0-162:42 *)
+(** [arrays::index_slice_u32_0]: forward function
+ Source: 'src/arrays.rs', lines 162:0-162:42 *)
let index_slice_u32_0 (x : slice u32) : result u32 =
slice_index_usize u32 x 0
-(** [array::index_mut_slice_u32_0]: forward function
- Source: 'src/array.rs', lines 166:0-166:50 *)
+(** [arrays::index_mut_slice_u32_0]: forward function
+ Source: 'src/arrays.rs', lines 166:0-166:50 *)
let index_mut_slice_u32_0 (x : slice u32) : result u32 =
slice_index_usize u32 x 0
-(** [array::index_mut_slice_u32_0]: backward function 0
- Source: 'src/array.rs', lines 166:0-166:50 *)
+(** [arrays::index_mut_slice_u32_0]: backward function 0
+ Source: 'src/arrays.rs', lines 166:0-166:50 *)
let index_mut_slice_u32_0_back (x : slice u32) : result (slice u32) =
let* _ = slice_index_usize u32 x 0 in Return x
-(** [array::index_all]: forward function
- Source: 'src/array.rs', lines 170:0-170:25 *)
+(** [arrays::index_all]: forward function
+ Source: 'src/arrays.rs', lines 170:0-170: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
@@ -274,25 +274,25 @@ let index_all : result u32 =
let* _ = array_from_slice u32 2 (mk_array u32 2 [ 0; 0 ]) s2 in
Return i8
-(** [array::update_array]: forward function
- Source: 'src/array.rs', lines 184:0-184:36 *)
+(** [arrays::update_array]: forward function
+ Source: 'src/arrays.rs', lines 184:0-184:36 *)
let update_array (x : array u32 2) : result unit =
let* _ = array_update_usize u32 2 x 0 1 in Return ()
-(** [array::update_array_mut_borrow]: merged forward/backward function
+(** [arrays::update_array_mut_borrow]: merged forward/backward function
(there is a single backward function, and the forward function returns ())
- Source: 'src/array.rs', lines 187:0-187:48 *)
+ Source: 'src/arrays.rs', lines 187:0-187:48 *)
let update_array_mut_borrow (x : array u32 2) : result (array u32 2) =
array_update_usize u32 2 x 0 1
-(** [array::update_mut_slice]: merged forward/backward function
+(** [arrays::update_mut_slice]: merged forward/backward function
(there is a single backward function, and the forward function returns ())
- Source: 'src/array.rs', lines 190:0-190:38 *)
+ Source: 'src/arrays.rs', lines 190:0-190:38 *)
let update_mut_slice (x : slice u32) : result (slice u32) =
slice_update_usize u32 x 0 1
-(** [array::update_all]: forward function
- Source: 'src/array.rs', lines 194:0-194:19 *)
+(** [arrays::update_all]: forward function
+ Source: 'src/arrays.rs', lines 194:0-194:19 *)
let update_all : result unit =
let* _ = update_array (mk_array u32 2 [ 0; 0 ]) in
let* x = update_array_mut_borrow (mk_array u32 2 [ 0; 0 ]) in
@@ -301,8 +301,8 @@ let update_all : result unit =
let* _ = array_from_slice u32 2 x s1 in
Return ()
-(** [array::range_all]: forward function
- Source: 'src/array.rs', lines 205:0-205:18 *)
+(** [arrays::range_all]: forward function
+ Source: 'src/arrays.rs', lines 205:0-205:18 *)
let range_all : result unit =
let* s =
core_array_Array_index_mut u32 (core_ops_range_Range usize) 4
@@ -317,33 +317,33 @@ let range_all : result unit =
(mk_array u32 4 [ 0; 0; 0; 0 ]) { start = 1; end_ = 3 } s1 in
Return ()
-(** [array::deref_array_borrow]: forward function
- Source: 'src/array.rs', lines 214:0-214:46 *)
+(** [arrays::deref_array_borrow]: forward function
+ Source: 'src/arrays.rs', lines 214:0-214:46 *)
let deref_array_borrow (x : array u32 2) : result u32 =
array_index_usize u32 2 x 0
-(** [array::deref_array_mut_borrow]: forward function
- Source: 'src/array.rs', lines 219:0-219:54 *)
+(** [arrays::deref_array_mut_borrow]: forward function
+ Source: 'src/arrays.rs', lines 219:0-219:54 *)
let deref_array_mut_borrow (x : array u32 2) : result u32 =
array_index_usize u32 2 x 0
-(** [array::deref_array_mut_borrow]: backward function 0
- Source: 'src/array.rs', lines 219:0-219:54 *)
+(** [arrays::deref_array_mut_borrow]: backward function 0
+ Source: 'src/arrays.rs', lines 219:0-219:54 *)
let deref_array_mut_borrow_back (x : array u32 2) : result (array u32 2) =
let* _ = array_index_usize u32 2 x 0 in Return x
-(** [array::take_array_t]: forward function
- Source: 'src/array.rs', lines 227:0-227:31 *)
+(** [arrays::take_array_t]: forward function
+ Source: 'src/arrays.rs', lines 227:0-227:31 *)
let take_array_t (a : array aB_t 2) : result unit =
Return ()
-(** [array::non_copyable_array]: forward function
- Source: 'src/array.rs', lines 229:0-229:27 *)
+(** [arrays::non_copyable_array]: forward function
+ Source: 'src/arrays.rs', lines 229:0-229:27 *)
let non_copyable_array : result unit =
let* _ = take_array_t (mk_array aB_t 2 [ AB_A; AB_B ]) in Return ()
-(** [array::sum]: loop 0: forward function
- Source: 'src/array.rs', lines 242:0-250:1 *)
+(** [arrays::sum]: loop 0: forward function
+ Source: 'src/arrays.rs', lines 242:0-250:1 *)
let rec sum_loop
(s : slice u32) (sum1 : u32) (i : usize) :
Tot (result u32) (decreases (sum_loop_decreases s sum1 i))
@@ -357,13 +357,13 @@ let rec sum_loop
sum_loop s sum3 i3
else Return sum1
-(** [array::sum]: forward function
- Source: 'src/array.rs', lines 242:0-242:28 *)
+(** [arrays::sum]: forward function
+ Source: 'src/arrays.rs', lines 242:0-242:28 *)
let sum (s : slice u32) : result u32 =
sum_loop s 0 0
-(** [array::sum2]: loop 0: forward function
- Source: 'src/array.rs', lines 252:0-261:1 *)
+(** [arrays::sum2]: loop 0: forward function
+ Source: 'src/arrays.rs', lines 252:0-261: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))
@@ -379,41 +379,41 @@ let rec sum2_loop
sum2_loop s s2 sum3 i5
else Return sum1
-(** [array::sum2]: forward function
- Source: 'src/array.rs', lines 252:0-252:41 *)
+(** [arrays::sum2]: forward function
+ Source: 'src/arrays.rs', lines 252:0-252: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
-(** [array::f0]: forward function
- Source: 'src/array.rs', lines 263:0-263:11 *)
+(** [arrays::f0]: forward function
+ Source: 'src/arrays.rs', lines 263:0-263:11 *)
let f0 : result unit =
let* s = array_to_slice u32 2 (mk_array u32 2 [ 1; 2 ]) in
let* s1 = slice_update_usize u32 s 0 1 in
let* _ = array_from_slice u32 2 (mk_array u32 2 [ 1; 2 ]) s1 in
Return ()
-(** [array::f1]: forward function
- Source: 'src/array.rs', lines 268:0-268:11 *)
+(** [arrays::f1]: forward function
+ Source: 'src/arrays.rs', lines 268:0-268:11 *)
let f1 : result unit =
let* _ = array_update_usize u32 2 (mk_array u32 2 [ 1; 2 ]) 0 1 in Return ()
-(** [array::f2]: forward function
- Source: 'src/array.rs', lines 273:0-273:17 *)
+(** [arrays::f2]: forward function
+ Source: 'src/arrays.rs', lines 273:0-273:17 *)
let f2 (i : u32) : result unit =
Return ()
-(** [array::f4]: forward function
- Source: 'src/array.rs', lines 282:0-282:54 *)
+(** [arrays::f4]: forward function
+ Source: 'src/arrays.rs', lines 282:0-282: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)
(core_slice_index_SliceIndexRangeUsizeSliceTInst u32)) x
{ start = y; end_ = z }
-(** [array::f3]: forward function
- Source: 'src/array.rs', lines 275:0-275:18 *)
+(** [arrays::f3]: forward function
+ Source: 'src/arrays.rs', lines 275:0-275:18 *)
let f3 : result u32 =
let* i = array_index_usize u32 2 (mk_array u32 2 [ 1; 2 ]) 0 in
let* _ = f2 i in
@@ -422,18 +422,18 @@ let f3 : result u32 =
let* s1 = f4 b 16 18 in
sum2 s s1
-(** [array::SZ]
- Source: 'src/array.rs', lines 286:0-286:19 *)
+(** [arrays::SZ]
+ Source: 'src/arrays.rs', lines 286:0-286:19 *)
let sz_body : result usize = Return 32
let sz_c : usize = eval_global sz_body
-(** [array::f5]: forward function
- Source: 'src/array.rs', lines 289:0-289:31 *)
+(** [arrays::f5]: forward function
+ Source: 'src/arrays.rs', lines 289:0-289:31 *)
let f5 (x : array u32 32) : result u32 =
array_index_usize u32 32 x 0
-(** [array::ite]: forward function
- Source: 'src/array.rs', lines 294:0-294:12 *)
+(** [arrays::ite]: forward function
+ Source: 'src/arrays.rs', lines 294:0-294:12 *)
let ite : result unit =
let* s = array_to_slice u32 2 (mk_array u32 2 [ 0; 0 ]) in
let* s1 = array_to_slice u32 2 (mk_array u32 2 [ 0; 0 ]) in
diff --git a/tests/fstar/array/Array.Types.fst b/tests/fstar-split/arrays/Arrays.Types.fst
index 312f6018..d3596e92 100644
--- a/tests/fstar/array/Array.Types.fst
+++ b/tests/fstar-split/arrays/Arrays.Types.fst
@@ -1,11 +1,11 @@
(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *)
-(** [array]: type definitions *)
-module Array.Types
+(** [arrays]: type definitions *)
+module Arrays.Types
open Primitives
#set-options "--z3rlimit 50 --fuel 1 --ifuel 1"
-(** [array::AB]
- Source: 'src/array.rs', lines 3:0-3:11 *)
+(** [arrays::AB]
+ Source: 'src/arrays.rs', lines 3:0-3:11 *)
type aB_t = | AB_A : aB_t | AB_B : aB_t
diff --git a/tests/fstar-split/array/Makefile b/tests/fstar-split/arrays/Makefile
index fa7d1f36..fa7d1f36 100644
--- a/tests/fstar-split/array/Makefile
+++ b/tests/fstar-split/arrays/Makefile
diff --git a/tests/fstar-split/array/Primitives.fst b/tests/fstar-split/arrays/Primitives.fst
index a3ffbde4..a3ffbde4 100644
--- a/tests/fstar-split/array/Primitives.fst
+++ b/tests/fstar-split/arrays/Primitives.fst
diff --git a/tests/fstar/array/Array.Clauses.Template.fst b/tests/fstar/arrays/Arrays.Clauses.Template.fst
index b2f2649c..8cc32583 100644
--- a/tests/fstar/array/Array.Clauses.Template.fst
+++ b/tests/fstar/arrays/Arrays.Clauses.Template.fst
@@ -1,19 +1,19 @@
(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *)
-(** [array]: templates for the decreases clauses *)
-module Array.Clauses.Template
+(** [arrays]: templates for the decreases clauses *)
+module Arrays.Clauses.Template
open Primitives
-open Array.Types
+open Arrays.Types
#set-options "--z3rlimit 50 --fuel 1 --ifuel 1"
-(** [array::sum]: decreases clause
- Source: 'src/array.rs', lines 242:0-250:1 *)
+(** [arrays::sum]: decreases clause
+ Source: 'src/arrays.rs', lines 242:0-250:1 *)
unfold
let sum_loop_decreases (s : slice u32) (sum1 : u32) (i : usize) : nat =
admit ()
-(** [array::sum2]: decreases clause
- Source: 'src/array.rs', lines 252:0-261:1 *)
+(** [arrays::sum2]: decreases clause
+ Source: 'src/arrays.rs', lines 252:0-261:1 *)
unfold
let sum2_loop_decreases (s : slice u32) (s2 : slice u32) (sum1 : u32)
(i : usize) : nat =
diff --git a/tests/fstar/array/Array.Clauses.fst b/tests/fstar/arrays/Arrays.Clauses.fst
index 68cbf216..aca328c2 100644
--- a/tests/fstar/array/Array.Clauses.fst
+++ b/tests/fstar/arrays/Arrays.Clauses.fst
@@ -1,17 +1,17 @@
-(** [array]: decreases clauses *)
-module Array.Clauses
+(** [arrays]: decreases clauses *)
+module Arrays.Clauses
open Primitives
-open Array.Types
+open Arrays.Types
open FStar.List.Tot
#set-options "--z3rlimit 50 --fuel 1 --ifuel 1"
-(** [array::sum]: decreases clause *)
+(** [arrays::sum]: decreases clause *)
unfold
let sum_loop_decreases (s : slice u32) (sum : u32) (i : usize) : nat =
if i < length s then length s - i else 0
-(** [array::sum2]: decreases clause *)
+(** [arrays::sum2]: decreases clause *)
unfold
let sum2_loop_decreases (s : slice u32) (s2 : slice u32) (sum : u32)
(i : usize) : nat =
diff --git a/tests/fstar/array/Array.Funs.fst b/tests/fstar/arrays/Arrays.Funs.fst
index 4193ba7d..b0df7fc2 100644
--- a/tests/fstar/array/Array.Funs.fst
+++ b/tests/fstar/arrays/Arrays.Funs.fst
@@ -1,24 +1,24 @@
(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *)
-(** [array]: function definitions *)
-module Array.Funs
+(** [arrays]: function definitions *)
+module Arrays.Funs
open Primitives
-include Array.Types
-include Array.Clauses
+include Arrays.Types
+include Arrays.Clauses
#set-options "--z3rlimit 50 --fuel 1 --ifuel 1"
-(** [array::incr]:
- Source: 'src/array.rs', lines 8:0-8:24 *)
+(** [arrays::incr]:
+ Source: 'src/arrays.rs', lines 8:0-8:24 *)
let incr (x : u32) : result u32 =
u32_add x 1
-(** [array::array_to_shared_slice_]:
- Source: 'src/array.rs', lines 16:0-16:53 *)
+(** [arrays::array_to_shared_slice_]:
+ Source: 'src/arrays.rs', lines 16:0-16:53 *)
let array_to_shared_slice_ (t : Type0) (s : array t 32) : result (slice t) =
array_to_slice t 32 s
-(** [array::array_to_mut_slice_]:
- Source: 'src/array.rs', lines 21:0-21:58 *)
+(** [arrays::array_to_mut_slice_]:
+ Source: 'src/arrays.rs', lines 21:0-21:58 *)
let array_to_mut_slice_
(t : Type0) (s : array t 32) :
result ((slice t) & (slice t -> result (array t 32)))
@@ -26,38 +26,38 @@ let array_to_mut_slice_
let* (s1, to_slice_mut_back) = array_to_slice_mut t 32 s in
Return (s1, to_slice_mut_back)
-(** [array::array_len]:
- Source: 'src/array.rs', lines 25:0-25:40 *)
+(** [arrays::array_len]:
+ Source: 'src/arrays.rs', lines 25:0-25:40 *)
let array_len (t : Type0) (s : array t 32) : result usize =
let* s1 = array_to_slice t 32 s in let i = slice_len t s1 in Return i
-(** [array::shared_array_len]:
- Source: 'src/array.rs', lines 29:0-29:48 *)
+(** [arrays::shared_array_len]:
+ Source: 'src/arrays.rs', lines 29:0-29:48 *)
let shared_array_len (t : Type0) (s : array t 32) : result usize =
let* s1 = array_to_slice t 32 s in let i = slice_len t s1 in Return i
-(** [array::shared_slice_len]:
- Source: 'src/array.rs', lines 33:0-33:44 *)
+(** [arrays::shared_slice_len]:
+ Source: 'src/arrays.rs', lines 33:0-33:44 *)
let shared_slice_len (t : Type0) (s : slice t) : result usize =
let i = slice_len t s in Return i
-(** [array::index_array_shared]:
- Source: 'src/array.rs', lines 37:0-37:57 *)
+(** [arrays::index_array_shared]:
+ Source: 'src/arrays.rs', lines 37:0-37:57 *)
let index_array_shared (t : Type0) (s : array t 32) (i : usize) : result t =
array_index_usize t 32 s i
-(** [array::index_array_u32]:
- Source: 'src/array.rs', lines 44:0-44:53 *)
+(** [arrays::index_array_u32]:
+ Source: 'src/arrays.rs', lines 44:0-44:53 *)
let index_array_u32 (s : array u32 32) (i : usize) : result u32 =
array_index_usize u32 32 s i
-(** [array::index_array_copy]:
- Source: 'src/array.rs', lines 48:0-48:45 *)
+(** [arrays::index_array_copy]:
+ Source: 'src/arrays.rs', lines 48:0-48:45 *)
let index_array_copy (x : array u32 32) : result u32 =
array_index_usize u32 32 x 0
-(** [array::index_mut_array]:
- Source: 'src/array.rs', lines 52:0-52:62 *)
+(** [arrays::index_mut_array]:
+ Source: 'src/arrays.rs', lines 52:0-52:62 *)
let index_mut_array
(t : Type0) (s : array t 32) (i : usize) :
result (t & (t -> result (array t 32)))
@@ -65,13 +65,13 @@ let index_mut_array
let* (x, index_mut_back) = array_index_mut_usize t 32 s i in
Return (x, index_mut_back)
-(** [array::index_slice]:
- Source: 'src/array.rs', lines 56:0-56:46 *)
+(** [arrays::index_slice]:
+ Source: 'src/arrays.rs', lines 56:0-56:46 *)
let index_slice (t : Type0) (s : slice t) (i : usize) : result t =
slice_index_usize t s i
-(** [array::index_mut_slice]:
- Source: 'src/array.rs', lines 60:0-60:58 *)
+(** [arrays::index_mut_slice]:
+ Source: 'src/arrays.rs', lines 60:0-60:58 *)
let index_mut_slice
(t : Type0) (s : slice t) (i : usize) :
result (t & (t -> result (slice t)))
@@ -79,16 +79,16 @@ let index_mut_slice
let* (x, index_mut_back) = slice_index_mut_usize t s i in
Return (x, index_mut_back)
-(** [array::slice_subslice_shared_]:
- Source: 'src/array.rs', lines 64:0-64:70 *)
+(** [arrays::slice_subslice_shared_]:
+ Source: 'src/arrays.rs', lines 64:0-64: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)
(core_slice_index_SliceIndexRangeUsizeSliceTInst u32) x
{ start = y; end_ = z }
-(** [array::slice_subslice_mut_]:
- Source: 'src/array.rs', lines 68:0-68:75 *)
+(** [arrays::slice_subslice_mut_]:
+ Source: 'src/arrays.rs', lines 68:0-68:75 *)
let slice_subslice_mut_
(x : slice u32) (y : usize) (z : usize) :
result ((slice u32) & (slice u32 -> result (slice u32)))
@@ -99,13 +99,13 @@ let slice_subslice_mut_
{ start = y; end_ = z } in
Return (s, index_mut_back)
-(** [array::array_to_slice_shared_]:
- Source: 'src/array.rs', lines 72:0-72:54 *)
+(** [arrays::array_to_slice_shared_]:
+ Source: 'src/arrays.rs', lines 72:0-72:54 *)
let array_to_slice_shared_ (x : array u32 32) : result (slice u32) =
array_to_slice u32 32 x
-(** [array::array_to_slice_mut_]:
- Source: 'src/array.rs', lines 76:0-76:59 *)
+(** [arrays::array_to_slice_mut_]:
+ Source: 'src/arrays.rs', lines 76:0-76:59 *)
let array_to_slice_mut_
(x : array u32 32) :
result ((slice u32) & (slice u32 -> result (array u32 32)))
@@ -113,8 +113,8 @@ let array_to_slice_mut_
let* (s, to_slice_mut_back) = array_to_slice_mut u32 32 x in
Return (s, to_slice_mut_back)
-(** [array::array_subslice_shared_]:
- Source: 'src/array.rs', lines 80:0-80:74 *)
+(** [arrays::array_subslice_shared_]:
+ Source: 'src/arrays.rs', lines 80:0-80: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
@@ -122,8 +122,8 @@ let array_subslice_shared_
(core_slice_index_SliceIndexRangeUsizeSliceTInst u32)) x
{ start = y; end_ = z }
-(** [array::array_subslice_mut_]:
- Source: 'src/array.rs', lines 84:0-84:79 *)
+(** [arrays::array_subslice_mut_]:
+ Source: 'src/arrays.rs', lines 84:0-84:79 *)
let array_subslice_mut_
(x : array u32 32) (y : usize) (z : usize) :
result ((slice u32) & (slice u32 -> result (array u32 32)))
@@ -135,25 +135,25 @@ let array_subslice_mut_
{ start = y; end_ = z } in
Return (s, index_mut_back)
-(** [array::index_slice_0]:
- Source: 'src/array.rs', lines 88:0-88:38 *)
+(** [arrays::index_slice_0]:
+ Source: 'src/arrays.rs', lines 88:0-88:38 *)
let index_slice_0 (t : Type0) (s : slice t) : result t =
slice_index_usize t s 0
-(** [array::index_array_0]:
- Source: 'src/array.rs', lines 92:0-92:42 *)
+(** [arrays::index_array_0]:
+ Source: 'src/arrays.rs', lines 92:0-92:42 *)
let index_array_0 (t : Type0) (s : array t 32) : result t =
array_index_usize t 32 s 0
-(** [array::index_index_array]:
- Source: 'src/array.rs', lines 103:0-103:71 *)
+(** [arrays::index_index_array]:
+ Source: 'src/arrays.rs', lines 103:0-103: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
-(** [array::update_update_array]:
- Source: 'src/array.rs', lines 114:0-114:70 *)
+(** [arrays::update_update_array]:
+ Source: 'src/arrays.rs', lines 114:0-114: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
@@ -162,43 +162,43 @@ let update_update_array
let* _ = index_mut_back a1 in
Return ()
-(** [array::array_local_deep_copy]:
- Source: 'src/array.rs', lines 118:0-118:43 *)
+(** [arrays::array_local_deep_copy]:
+ Source: 'src/arrays.rs', lines 118:0-118:43 *)
let array_local_deep_copy (x : array u32 32) : result unit =
Return ()
-(** [array::take_array]:
- Source: 'src/array.rs', lines 122:0-122:30 *)
+(** [arrays::take_array]:
+ Source: 'src/arrays.rs', lines 122:0-122:30 *)
let take_array (a : array u32 2) : result unit =
Return ()
-(** [array::take_array_borrow]:
- Source: 'src/array.rs', lines 123:0-123:38 *)
+(** [arrays::take_array_borrow]:
+ Source: 'src/arrays.rs', lines 123:0-123:38 *)
let take_array_borrow (a : array u32 2) : result unit =
Return ()
-(** [array::take_slice]:
- Source: 'src/array.rs', lines 124:0-124:28 *)
+(** [arrays::take_slice]:
+ Source: 'src/arrays.rs', lines 124:0-124:28 *)
let take_slice (s : slice u32) : result unit =
Return ()
-(** [array::take_mut_slice]:
- Source: 'src/array.rs', lines 125:0-125:36 *)
+(** [arrays::take_mut_slice]:
+ Source: 'src/arrays.rs', lines 125:0-125:36 *)
let take_mut_slice (s : slice u32) : result (slice u32) =
Return s
-(** [array::const_array]:
- Source: 'src/array.rs', lines 127:0-127:32 *)
+(** [arrays::const_array]:
+ Source: 'src/arrays.rs', lines 127:0-127:32 *)
let const_array : result (array u32 2) =
Return (mk_array u32 2 [ 0; 0 ])
-(** [array::const_slice]:
- Source: 'src/array.rs', lines 131:0-131:20 *)
+(** [arrays::const_slice]:
+ Source: 'src/arrays.rs', lines 131:0-131:20 *)
let const_slice : result unit =
let* _ = array_to_slice u32 2 (mk_array u32 2 [ 0; 0 ]) in Return ()
-(** [array::take_all]:
- Source: 'src/array.rs', lines 141:0-141:17 *)
+(** [arrays::take_all]:
+ Source: 'src/arrays.rs', lines 141:0-141: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
@@ -211,28 +211,28 @@ let take_all : result unit =
let* _ = to_slice_mut_back s2 in
Return ()
-(** [array::index_array]:
- Source: 'src/array.rs', lines 155:0-155:38 *)
+(** [arrays::index_array]:
+ Source: 'src/arrays.rs', lines 155:0-155:38 *)
let index_array (x : array u32 2) : result u32 =
array_index_usize u32 2 x 0
-(** [array::index_array_borrow]:
- Source: 'src/array.rs', lines 158:0-158:46 *)
+(** [arrays::index_array_borrow]:
+ Source: 'src/arrays.rs', lines 158:0-158:46 *)
let index_array_borrow (x : array u32 2) : result u32 =
array_index_usize u32 2 x 0
-(** [array::index_slice_u32_0]:
- Source: 'src/array.rs', lines 162:0-162:42 *)
+(** [arrays::index_slice_u32_0]:
+ Source: 'src/arrays.rs', lines 162:0-162:42 *)
let index_slice_u32_0 (x : slice u32) : result u32 =
slice_index_usize u32 x 0
-(** [array::index_mut_slice_u32_0]:
- Source: 'src/array.rs', lines 166:0-166:50 *)
+(** [arrays::index_mut_slice_u32_0]:
+ Source: 'src/arrays.rs', lines 166:0-166:50 *)
let index_mut_slice_u32_0 (x : slice u32) : result (u32 & (slice u32)) =
let* i = slice_index_usize u32 x 0 in Return (i, x)
-(** [array::index_all]:
- Source: 'src/array.rs', lines 170:0-170:25 *)
+(** [arrays::index_all]:
+ Source: 'src/arrays.rs', lines 170:0-170: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
@@ -249,26 +249,26 @@ let index_all : result u32 =
let* _ = to_slice_mut_back s2 in
Return i8
-(** [array::update_array]:
- Source: 'src/array.rs', lines 184:0-184:36 *)
+(** [arrays::update_array]:
+ Source: 'src/arrays.rs', lines 184:0-184: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
Return ()
-(** [array::update_array_mut_borrow]:
- Source: 'src/array.rs', lines 187:0-187:48 *)
+(** [arrays::update_array_mut_borrow]:
+ Source: 'src/arrays.rs', lines 187:0-187: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
-(** [array::update_mut_slice]:
- Source: 'src/array.rs', lines 190:0-190:38 *)
+(** [arrays::update_mut_slice]:
+ Source: 'src/arrays.rs', lines 190:0-190: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
-(** [array::update_all]:
- Source: 'src/array.rs', lines 194:0-194:19 *)
+(** [arrays::update_all]:
+ Source: 'src/arrays.rs', lines 194:0-194: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
@@ -278,8 +278,8 @@ let update_all : result unit =
let* _ = to_slice_mut_back s1 in
Return ()
-(** [array::range_all]:
- Source: 'src/array.rs', lines 205:0-205:18 *)
+(** [arrays::range_all]:
+ Source: 'src/arrays.rs', lines 205:0-205:18 *)
let range_all : result unit =
let* (s, index_mut_back) =
core_array_Array_index_mut u32 (core_ops_range_Range usize) 4
@@ -290,28 +290,28 @@ let range_all : result unit =
let* _ = index_mut_back s1 in
Return ()
-(** [array::deref_array_borrow]:
- Source: 'src/array.rs', lines 214:0-214:46 *)
+(** [arrays::deref_array_borrow]:
+ Source: 'src/arrays.rs', lines 214:0-214:46 *)
let deref_array_borrow (x : array u32 2) : result u32 =
array_index_usize u32 2 x 0
-(** [array::deref_array_mut_borrow]:
- Source: 'src/array.rs', lines 219:0-219:54 *)
+(** [arrays::deref_array_mut_borrow]:
+ Source: 'src/arrays.rs', lines 219:0-219: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 Return (i, x)
-(** [array::take_array_t]:
- Source: 'src/array.rs', lines 227:0-227:31 *)
+(** [arrays::take_array_t]:
+ Source: 'src/arrays.rs', lines 227:0-227:31 *)
let take_array_t (a : array aB_t 2) : result unit =
Return ()
-(** [array::non_copyable_array]:
- Source: 'src/array.rs', lines 229:0-229:27 *)
+(** [arrays::non_copyable_array]:
+ Source: 'src/arrays.rs', lines 229:0-229:27 *)
let non_copyable_array : result unit =
let* _ = take_array_t (mk_array aB_t 2 [ AB_A; AB_B ]) in Return ()
-(** [array::sum]: loop 0:
- Source: 'src/array.rs', lines 242:0-250:1 *)
+(** [arrays::sum]: loop 0:
+ Source: 'src/arrays.rs', lines 242:0-250:1 *)
let rec sum_loop
(s : slice u32) (sum1 : u32) (i : usize) :
Tot (result u32) (decreases (sum_loop_decreases s sum1 i))
@@ -325,13 +325,13 @@ let rec sum_loop
sum_loop s sum3 i3
else Return sum1
-(** [array::sum]:
- Source: 'src/array.rs', lines 242:0-242:28 *)
+(** [arrays::sum]:
+ Source: 'src/arrays.rs', lines 242:0-242:28 *)
let sum (s : slice u32) : result u32 =
sum_loop s 0 0
-(** [array::sum2]: loop 0:
- Source: 'src/array.rs', lines 252:0-261:1 *)
+(** [arrays::sum2]: loop 0:
+ Source: 'src/arrays.rs', lines 252:0-261: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))
@@ -347,15 +347,15 @@ let rec sum2_loop
sum2_loop s s2 sum3 i5
else Return sum1
-(** [array::sum2]:
- Source: 'src/array.rs', lines 252:0-252:41 *)
+(** [arrays::sum2]:
+ Source: 'src/arrays.rs', lines 252:0-252: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
-(** [array::f0]:
- Source: 'src/array.rs', lines 263:0-263:11 *)
+(** [arrays::f0]:
+ Source: 'src/arrays.rs', lines 263:0-263:11 *)
let f0 : result unit =
let* (s, to_slice_mut_back) =
array_to_slice_mut u32 2 (mk_array u32 2 [ 1; 2 ]) in
@@ -364,29 +364,29 @@ let f0 : result unit =
let* _ = to_slice_mut_back s1 in
Return ()
-(** [array::f1]:
- Source: 'src/array.rs', lines 268:0-268:11 *)
+(** [arrays::f1]:
+ Source: 'src/arrays.rs', lines 268:0-268:11 *)
let f1 : result unit =
let* (_, index_mut_back) =
array_index_mut_usize u32 2 (mk_array u32 2 [ 1; 2 ]) 0 in
let* _ = index_mut_back 1 in
Return ()
-(** [array::f2]:
- Source: 'src/array.rs', lines 273:0-273:17 *)
+(** [arrays::f2]:
+ Source: 'src/arrays.rs', lines 273:0-273:17 *)
let f2 (i : u32) : result unit =
Return ()
-(** [array::f4]:
- Source: 'src/array.rs', lines 282:0-282:54 *)
+(** [arrays::f4]:
+ Source: 'src/arrays.rs', lines 282:0-282: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)
(core_slice_index_SliceIndexRangeUsizeSliceTInst u32)) x
{ start = y; end_ = z }
-(** [array::f3]:
- Source: 'src/array.rs', lines 275:0-275:18 *)
+(** [arrays::f3]:
+ Source: 'src/arrays.rs', lines 275:0-275:18 *)
let f3 : result u32 =
let* i = array_index_usize u32 2 (mk_array u32 2 [ 1; 2 ]) 0 in
let* _ = f2 i in
@@ -395,18 +395,18 @@ let f3 : result u32 =
let* s1 = f4 b 16 18 in
sum2 s s1
-(** [array::SZ]
- Source: 'src/array.rs', lines 286:0-286:19 *)
+(** [arrays::SZ]
+ Source: 'src/arrays.rs', lines 286:0-286:19 *)
let sz_body : result usize = Return 32
let sz_c : usize = eval_global sz_body
-(** [array::f5]:
- Source: 'src/array.rs', lines 289:0-289:31 *)
+(** [arrays::f5]:
+ Source: 'src/arrays.rs', lines 289:0-289:31 *)
let f5 (x : array u32 32) : result u32 =
array_index_usize u32 32 x 0
-(** [array::ite]:
- Source: 'src/array.rs', lines 294:0-294:12 *)
+(** [arrays::ite]:
+ Source: 'src/arrays.rs', lines 294:0-294:12 *)
let ite : result unit =
let* (s, to_slice_mut_back) =
array_to_slice_mut u32 2 (mk_array u32 2 [ 0; 0 ]) in
diff --git a/tests/fstar-split/array/Array.Types.fst b/tests/fstar/arrays/Arrays.Types.fst
index 312f6018..d3596e92 100644
--- a/tests/fstar-split/array/Array.Types.fst
+++ b/tests/fstar/arrays/Arrays.Types.fst
@@ -1,11 +1,11 @@
(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *)
-(** [array]: type definitions *)
-module Array.Types
+(** [arrays]: type definitions *)
+module Arrays.Types
open Primitives
#set-options "--z3rlimit 50 --fuel 1 --ifuel 1"
-(** [array::AB]
- Source: 'src/array.rs', lines 3:0-3:11 *)
+(** [arrays::AB]
+ Source: 'src/arrays.rs', lines 3:0-3:11 *)
type aB_t = | AB_A : aB_t | AB_B : aB_t
diff --git a/tests/fstar/array/Makefile b/tests/fstar/arrays/Makefile
index fa7d1f36..fa7d1f36 100644
--- a/tests/fstar/array/Makefile
+++ b/tests/fstar/arrays/Makefile
diff --git a/tests/fstar/array/Primitives.fst b/tests/fstar/arrays/Primitives.fst
index fca80829..fca80829 100644
--- a/tests/fstar/array/Primitives.fst
+++ b/tests/fstar/arrays/Primitives.fst
diff --git a/tests/lean/.gitignore b/tests/lean/.gitignore
index 4d1c5853..071df2d0 100644
--- a/tests/lean/.gitignore
+++ b/tests/lean/.gitignore
@@ -1,2 +1,3 @@
lake-packages
-build \ No newline at end of file
+build
+.lake \ No newline at end of file
diff --git a/tests/lean/Array/Funs.lean b/tests/lean/Array/Funs.lean
deleted file mode 100644
index 32ae6248..00000000
--- a/tests/lean/Array/Funs.lean
+++ /dev/null
@@ -1,431 +0,0 @@
--- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS
--- [array]: function definitions
-import Base
-import Array.Types
-open Primitives
-
-namespace array
-
-/- [array::incr]: merged forward/backward function
- (there is a single backward function, and the forward function returns ()) -/
-def incr (x : U32) : Result U32 :=
- x + 1#u32
-
-/- [array::array_to_shared_slice_]: forward function -/
-def array_to_shared_slice_
- (T : Type) (s : Array T 32#usize) : Result (Slice T) :=
- Array.to_slice T 32#usize s
-
-/- [array::array_to_mut_slice_]: forward function -/
-def array_to_mut_slice_ (T : Type) (s : Array T 32#usize) : Result (Slice T) :=
- Array.to_slice T 32#usize s
-
-/- [array::array_to_mut_slice_]: backward function 0 -/
-def array_to_mut_slice__back
- (T : Type) (s : Array T 32#usize) (ret0 : Slice T) :
- Result (Array T 32#usize)
- :=
- Array.from_slice T 32#usize s ret0
-
-/- [array::array_len]: forward function -/
-def array_len (T : Type) (s : Array T 32#usize) : Result Usize :=
- do
- let s0 ← Array.to_slice T 32#usize s
- let i := Slice.len T s0
- Result.ret i
-
-/- [array::shared_array_len]: forward function -/
-def shared_array_len (T : Type) (s : Array T 32#usize) : Result Usize :=
- do
- let s0 ← Array.to_slice T 32#usize s
- let i := Slice.len T s0
- Result.ret i
-
-/- [array::shared_slice_len]: forward function -/
-def shared_slice_len (T : Type) (s : Slice T) : Result Usize :=
- let i := Slice.len T s
- Result.ret i
-
-/- [array::index_array_shared]: forward function -/
-def index_array_shared
- (T : Type) (s : Array T 32#usize) (i : Usize) : Result T :=
- Array.index_usize T 32#usize s i
-
-/- [array::index_array_u32]: forward function -/
-def index_array_u32 (s : Array U32 32#usize) (i : Usize) : Result U32 :=
- Array.index_usize U32 32#usize s i
-
-/- [array::index_array_copy]: forward function -/
-def index_array_copy (x : Array U32 32#usize) : Result U32 :=
- Array.index_usize U32 32#usize x 0#usize
-
-/- [array::index_mut_array]: forward function -/
-def index_mut_array (T : Type) (s : Array T 32#usize) (i : Usize) : Result T :=
- Array.index_usize T 32#usize s i
-
-/- [array::index_mut_array]: backward function 0 -/
-def index_mut_array_back
- (T : Type) (s : Array T 32#usize) (i : Usize) (ret0 : T) :
- Result (Array T 32#usize)
- :=
- Array.update_usize T 32#usize s i ret0
-
-/- [array::index_slice]: forward function -/
-def index_slice (T : Type) (s : Slice T) (i : Usize) : Result T :=
- Slice.index_usize T s i
-
-/- [array::index_mut_slice]: forward function -/
-def index_mut_slice (T : Type) (s : Slice T) (i : Usize) : Result T :=
- Slice.index_usize T s i
-
-/- [array::index_mut_slice]: backward function 0 -/
-def index_mut_slice_back
- (T : Type) (s : Slice T) (i : Usize) (ret0 : T) : Result (Slice T) :=
- Slice.update_usize T s i ret0
-
-/- [array::slice_subslice_shared_]: forward function -/
-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)
- (core.slice.index.Range.coresliceindexSliceIndexInst U32) x
- { start := y, end_ := z }
-
-/- [array::slice_subslice_mut_]: forward function -/
-def slice_subslice_mut_
- (x : Slice U32) (y : Usize) (z : Usize) : Result (Slice U32) :=
- core.slice.index.Slice.index_mut U32 (core.ops.range.Range Usize)
- (core.slice.index.Range.coresliceindexSliceIndexInst U32) x
- { start := y, end_ := z }
-
-/- [array::slice_subslice_mut_]: backward function 0 -/
-def slice_subslice_mut__back
- (x : Slice U32) (y : Usize) (z : Usize) (ret0 : Slice U32) :
- Result (Slice U32)
- :=
- core.slice.index.Slice.index_mut_back U32 (core.ops.range.Range Usize)
- (core.slice.index.Range.coresliceindexSliceIndexInst U32) x
- { start := y, end_ := z } ret0
-
-/- [array::array_to_slice_shared_]: forward function -/
-def array_to_slice_shared_ (x : Array U32 32#usize) : Result (Slice U32) :=
- Array.to_slice U32 32#usize x
-
-/- [array::array_to_slice_mut_]: forward function -/
-def array_to_slice_mut_ (x : Array U32 32#usize) : Result (Slice U32) :=
- Array.to_slice U32 32#usize x
-
-/- [array::array_to_slice_mut_]: backward function 0 -/
-def array_to_slice_mut__back
- (x : Array U32 32#usize) (ret0 : Slice U32) : Result (Array U32 32#usize) :=
- Array.from_slice U32 32#usize x ret0
-
-/- [array::array_subslice_shared_]: forward function -/
-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
- (core.slice.index.Slice.coreopsindexIndexInst U32 (core.ops.range.Range
- Usize) (core.slice.index.Range.coresliceindexSliceIndexInst U32)) x
- { start := y, end_ := z }
-
-/- [array::array_subslice_mut_]: forward function -/
-def array_subslice_mut_
- (x : Array U32 32#usize) (y : Usize) (z : Usize) : Result (Slice U32) :=
- core.array.Array.index_mut U32 (core.ops.range.Range Usize) 32#usize
- (core.slice.index.Slice.coreopsindexIndexMutInst U32 (core.ops.range.Range
- Usize) (core.slice.index.Range.coresliceindexSliceIndexInst U32)) x
- { start := y, end_ := z }
-
-/- [array::array_subslice_mut_]: backward function 0 -/
-def array_subslice_mut__back
- (x : Array U32 32#usize) (y : Usize) (z : Usize) (ret0 : Slice U32) :
- Result (Array U32 32#usize)
- :=
- core.array.Array.index_mut_back U32 (core.ops.range.Range Usize) 32#usize
- (core.slice.index.Slice.coreopsindexIndexMutInst U32 (core.ops.range.Range
- Usize) (core.slice.index.Range.coresliceindexSliceIndexInst U32)) x
- { start := y, end_ := z } ret0
-
-/- [array::index_slice_0]: forward function -/
-def index_slice_0 (T : Type) (s : Slice T) : Result T :=
- Slice.index_usize T s 0#usize
-
-/- [array::index_array_0]: forward function -/
-def index_array_0 (T : Type) (s : Array T 32#usize) : Result T :=
- Array.index_usize T 32#usize s 0#usize
-
-/- [array::index_index_array]: forward function -/
-def index_index_array
- (s : Array (Array U32 32#usize) 32#usize) (i : Usize) (j : Usize) :
- Result U32
- :=
- do
- let a ← Array.index_usize (Array U32 32#usize) 32#usize s i
- Array.index_usize U32 32#usize a j
-
-/- [array::update_update_array]: forward function -/
-def update_update_array
- (s : Array (Array U32 32#usize) 32#usize) (i : Usize) (j : Usize) :
- Result Unit
- :=
- do
- let a ← Array.index_usize (Array U32 32#usize) 32#usize s i
- let a0 ← Array.update_usize U32 32#usize a j 0#u32
- let _ ← Array.update_usize (Array U32 32#usize) 32#usize s i a0
- Result.ret ()
-
-/- [array::array_local_deep_copy]: forward function -/
-def array_local_deep_copy (x : Array U32 32#usize) : Result Unit :=
- Result.ret ()
-
-/- [array::take_array]: forward function -/
-def take_array (a : Array U32 2#usize) : Result Unit :=
- Result.ret ()
-
-/- [array::take_array_borrow]: forward function -/
-def take_array_borrow (a : Array U32 2#usize) : Result Unit :=
- Result.ret ()
-
-/- [array::take_slice]: forward function -/
-def take_slice (s : Slice U32) : Result Unit :=
- Result.ret ()
-
-/- [array::take_mut_slice]: merged forward/backward function
- (there is a single backward function, and the forward function returns ()) -/
-def take_mut_slice (s : Slice U32) : Result (Slice U32) :=
- Result.ret s
-
-/- [array::take_all]: forward function -/
-def take_all : Result Unit :=
- do
- let _ ← take_array (Array.make U32 2#usize [ 0#u32, 0#u32 ])
- let _ ← take_array_borrow (Array.make U32 2#usize [ 0#u32, 0#u32 ])
- let s ←
- Array.to_slice U32 2#usize (Array.make U32 2#usize [ 0#u32, 0#u32 ])
- let _ ← take_slice s
- let s0 ←
- Array.to_slice U32 2#usize (Array.make U32 2#usize [ 0#u32, 0#u32 ])
- let s1 ← take_mut_slice s0
- let _ ←
- Array.from_slice U32 2#usize (Array.make U32 2#usize [ 0#u32, 0#u32 ]) s1
- Result.ret ()
-
-/- [array::index_array]: forward function -/
-def index_array (x : Array U32 2#usize) : Result U32 :=
- Array.index_usize U32 2#usize x 0#usize
-
-/- [array::index_array_borrow]: forward function -/
-def index_array_borrow (x : Array U32 2#usize) : Result U32 :=
- Array.index_usize U32 2#usize x 0#usize
-
-/- [array::index_slice_u32_0]: forward function -/
-def index_slice_u32_0 (x : Slice U32) : Result U32 :=
- Slice.index_usize U32 x 0#usize
-
-/- [array::index_mut_slice_u32_0]: forward function -/
-def index_mut_slice_u32_0 (x : Slice U32) : Result U32 :=
- Slice.index_usize U32 x 0#usize
-
-/- [array::index_mut_slice_u32_0]: backward function 0 -/
-def index_mut_slice_u32_0_back (x : Slice U32) : Result (Slice U32) :=
- do
- let _ ← Slice.index_usize U32 x 0#usize
- Result.ret x
-
-/- [array::index_all]: forward function -/
-def index_all : Result U32 :=
- do
- let i ← index_array (Array.make U32 2#usize [ 0#u32, 0#u32 ])
- let i0 ← index_array (Array.make U32 2#usize [ 0#u32, 0#u32 ])
- let i1 ← i + i0
- let i2 ← index_array_borrow (Array.make U32 2#usize [ 0#u32, 0#u32 ])
- let i3 ← i1 + i2
- let s ←
- Array.to_slice U32 2#usize (Array.make U32 2#usize [ 0#u32, 0#u32 ])
- let i4 ← index_slice_u32_0 s
- let i5 ← i3 + i4
- let s0 ←
- Array.to_slice U32 2#usize (Array.make U32 2#usize [ 0#u32, 0#u32 ])
- let i6 ← index_mut_slice_u32_0 s0
- let i7 ← i5 + i6
- let s1 ← index_mut_slice_u32_0_back s0
- let _ ←
- Array.from_slice U32 2#usize (Array.make U32 2#usize [ 0#u32, 0#u32 ]) s1
- Result.ret i7
-
-/- [array::update_array]: forward function -/
-def update_array (x : Array U32 2#usize) : Result Unit :=
- do
- let _ ← Array.update_usize U32 2#usize x 0#usize 1#u32
- Result.ret ()
-
-/- [array::update_array_mut_borrow]: merged forward/backward function
- (there is a single backward function, and the forward function returns ()) -/
-def update_array_mut_borrow
- (x : Array U32 2#usize) : Result (Array U32 2#usize) :=
- Array.update_usize U32 2#usize x 0#usize 1#u32
-
-/- [array::update_mut_slice]: merged forward/backward function
- (there is a single backward function, and the forward function returns ()) -/
-def update_mut_slice (x : Slice U32) : Result (Slice U32) :=
- Slice.update_usize U32 x 0#usize 1#u32
-
-/- [array::update_all]: forward function -/
-def update_all : Result Unit :=
- do
- let _ ← update_array (Array.make U32 2#usize [ 0#u32, 0#u32 ])
- let x ← update_array_mut_borrow (Array.make U32 2#usize [ 0#u32, 0#u32 ])
- let s ← Array.to_slice U32 2#usize x
- let s0 ← update_mut_slice s
- let _ ← Array.from_slice U32 2#usize x s0
- Result.ret ()
-
-/- [array::range_all]: forward function -/
-def range_all : Result Unit :=
- do
- let s ←
- core.array.Array.index_mut U32 (core.ops.range.Range Usize) 4#usize
- (core.slice.index.Slice.coreopsindexIndexMutInst U32
- (core.ops.range.Range Usize)
- (core.slice.index.Range.coresliceindexSliceIndexInst U32))
- (Array.make U32 4#usize [ 0#u32, 0#u32, 0#u32, 0#u32 ])
- { start := 1#usize, end_ := 3#usize }
- let s0 ← update_mut_slice s
- let _ ←
- core.array.Array.index_mut_back U32 (core.ops.range.Range Usize) 4#usize
- (core.slice.index.Slice.coreopsindexIndexMutInst U32
- (core.ops.range.Range Usize)
- (core.slice.index.Range.coresliceindexSliceIndexInst U32))
- (Array.make U32 4#usize [ 0#u32, 0#u32, 0#u32, 0#u32 ])
- { start := 1#usize, end_ := 3#usize } s0
- Result.ret ()
-
-/- [array::deref_array_borrow]: forward function -/
-def deref_array_borrow (x : Array U32 2#usize) : Result U32 :=
- Array.index_usize U32 2#usize x 0#usize
-
-/- [array::deref_array_mut_borrow]: forward function -/
-def deref_array_mut_borrow (x : Array U32 2#usize) : Result U32 :=
- Array.index_usize U32 2#usize x 0#usize
-
-/- [array::deref_array_mut_borrow]: backward function 0 -/
-def deref_array_mut_borrow_back
- (x : Array U32 2#usize) : Result (Array U32 2#usize) :=
- do
- let _ ← Array.index_usize U32 2#usize x 0#usize
- Result.ret x
-
-/- [array::take_array_t]: forward function -/
-def take_array_t (a : Array AB 2#usize) : Result Unit :=
- Result.ret ()
-
-/- [array::non_copyable_array]: forward function -/
-def non_copyable_array : Result Unit :=
- do
- let _ ← take_array_t (Array.make AB 2#usize [ AB.A, AB.B ])
- Result.ret ()
-
-/- [array::sum]: loop 0: forward function -/
-divergent def sum_loop (s : Slice U32) (sum0 : U32) (i : Usize) : Result U32 :=
- let i0 := Slice.len U32 s
- if i < i0
- then
- do
- let i1 ← Slice.index_usize U32 s i
- let sum1 ← sum0 + i1
- let i2 ← i + 1#usize
- sum_loop s sum1 i2
- else Result.ret sum0
-
-/- [array::sum]: forward function -/
-def sum (s : Slice U32) : Result U32 :=
- sum_loop s 0#u32 0#usize
-
-/- [array::sum2]: loop 0: forward function -/
-divergent def sum2_loop
- (s : Slice U32) (s2 : Slice U32) (sum0 : U32) (i : Usize) : Result U32 :=
- let i0 := Slice.len U32 s
- if i < i0
- then
- do
- let i1 ← Slice.index_usize U32 s i
- let i2 ← Slice.index_usize U32 s2 i
- let i3 ← i1 + i2
- let sum1 ← sum0 + i3
- let i4 ← i + 1#usize
- sum2_loop s s2 sum1 i4
- else Result.ret sum0
-
-/- [array::sum2]: forward function -/
-def sum2 (s : Slice U32) (s2 : Slice U32) : Result U32 :=
- let i := Slice.len U32 s
- let i0 := Slice.len U32 s2
- if not (i = i0)
- then Result.fail Error.panic
- else sum2_loop s s2 0#u32 0#usize
-
-/- [array::f0]: forward function -/
-def f0 : Result Unit :=
- do
- let s ←
- Array.to_slice U32 2#usize (Array.make U32 2#usize [ 1#u32, 2#u32 ])
- let s0 ← Slice.update_usize U32 s 0#usize 1#u32
- let _ ←
- Array.from_slice U32 2#usize (Array.make U32 2#usize [ 1#u32, 2#u32 ]) s0
- Result.ret ()
-
-/- [array::f1]: forward function -/
-def f1 : Result Unit :=
- do
- let _ ←
- Array.update_usize U32 2#usize (Array.make U32 2#usize [ 1#u32, 2#u32 ])
- 0#usize 1#u32
- Result.ret ()
-
-/- [array::f2]: forward function -/
-def f2 (i : U32) : Result Unit :=
- Result.ret ()
-
-/- [array::f4]: forward function -/
-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.slice.index.Slice.coreopsindexIndexInst U32 (core.ops.range.Range
- Usize) (core.slice.index.Range.coresliceindexSliceIndexInst U32)) x
- { start := y, end_ := z }
-
-/- [array::f3]: forward function -/
-def f3 : Result U32 :=
- do
- let i ←
- Array.index_usize U32 2#usize (Array.make U32 2#usize [ 1#u32, 2#u32 ])
- 0#usize
- let _ ← f2 i
- let b := Array.repeat U32 32#usize 0#u32
- let s ←
- Array.to_slice U32 2#usize (Array.make U32 2#usize [ 1#u32, 2#u32 ])
- let s0 ← f4 b 16#usize 18#usize
- sum2 s s0
-
-/- [array::SZ] -/
-def sz_body : Result Usize := Result.ret 32#usize
-def sz_c : Usize := eval_global sz_body (by simp)
-
-/- [array::f5]: forward function -/
-def f5 (x : Array U32 32#usize) : Result U32 :=
- Array.index_usize U32 32#usize x 0#usize
-
-/- [array::ite]: forward function -/
-def ite : Result Unit :=
- do
- let s ←
- Array.to_slice U32 2#usize (Array.make U32 2#usize [ 0#u32, 0#u32 ])
- let s0 ←
- Array.to_slice U32 2#usize (Array.make U32 2#usize [ 0#u32, 0#u32 ])
- let s1 ← index_mut_slice_u32_0_back s0
- let _ ←
- Array.from_slice U32 2#usize (Array.make U32 2#usize [ 0#u32, 0#u32 ]) s1
- let s2 ← index_mut_slice_u32_0_back s
- let _ ←
- Array.from_slice U32 2#usize (Array.make U32 2#usize [ 0#u32, 0#u32 ]) s2
- Result.ret ()
-
-end array
diff --git a/tests/lean/Array/Types.lean b/tests/lean/Array/Types.lean
deleted file mode 100644
index 60fa81ab..00000000
--- a/tests/lean/Array/Types.lean
+++ /dev/null
@@ -1,13 +0,0 @@
--- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS
--- [array]: type definitions
-import Base
-open Primitives
-
-namespace array
-
-/- [array::AB] -/
-inductive AB :=
-| A : AB
-| B : AB
-
-end array
diff --git a/tests/lean/Array.lean b/tests/lean/Arrays.lean
index 7785a208..5158ca28 100644
--- a/tests/lean/Array.lean
+++ b/tests/lean/Arrays.lean
@@ -1,29 +1,29 @@
-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS
--- [array]
+-- [arrays]
import Base
open Primitives
-namespace array
+namespace arrays
-/- [array::AB]
- Source: 'src/array.rs', lines 3:0-3:11 -/
+/- [arrays::AB]
+ Source: 'src/arrays.rs', lines 3:0-3:11 -/
inductive AB :=
| A : AB
| B : AB
-/- [array::incr]:
- Source: 'src/array.rs', lines 8:0-8:24 -/
+/- [arrays::incr]:
+ Source: 'src/arrays.rs', lines 8:0-8:24 -/
def incr (x : U32) : Result U32 :=
x + 1#u32
-/- [array::array_to_shared_slice_]:
- Source: 'src/array.rs', lines 16:0-16:53 -/
+/- [arrays::array_to_shared_slice_]:
+ Source: 'src/arrays.rs', lines 16:0-16:53 -/
def array_to_shared_slice_
(T : Type) (s : Array T 32#usize) : Result (Slice T) :=
Array.to_slice T 32#usize s
-/- [array::array_to_mut_slice_]:
- Source: 'src/array.rs', lines 21:0-21:58 -/
+/- [arrays::array_to_mut_slice_]:
+ Source: 'src/arrays.rs', lines 21:0-21:58 -/
def array_to_mut_slice_
(T : Type) (s : Array T 32#usize) :
Result ((Slice T) × (Slice T → Result (Array T 32#usize)))
@@ -32,46 +32,46 @@ def array_to_mut_slice_
let (s1, to_slice_mut_back) ← Array.to_slice_mut T 32#usize s
Result.ret (s1, to_slice_mut_back)
-/- [array::array_len]:
- Source: 'src/array.rs', lines 25:0-25:40 -/
+/- [arrays::array_len]:
+ Source: 'src/arrays.rs', lines 25:0-25:40 -/
def array_len (T : Type) (s : Array T 32#usize) : Result Usize :=
do
let s1 ← Array.to_slice T 32#usize s
let i := Slice.len T s1
Result.ret i
-/- [array::shared_array_len]:
- Source: 'src/array.rs', lines 29:0-29:48 -/
+/- [arrays::shared_array_len]:
+ Source: 'src/arrays.rs', lines 29:0-29:48 -/
def shared_array_len (T : Type) (s : Array T 32#usize) : Result Usize :=
do
let s1 ← Array.to_slice T 32#usize s
let i := Slice.len T s1
Result.ret i
-/- [array::shared_slice_len]:
- Source: 'src/array.rs', lines 33:0-33:44 -/
+/- [arrays::shared_slice_len]:
+ Source: 'src/arrays.rs', lines 33:0-33:44 -/
def shared_slice_len (T : Type) (s : Slice T) : Result Usize :=
let i := Slice.len T s
Result.ret i
-/- [array::index_array_shared]:
- Source: 'src/array.rs', lines 37:0-37:57 -/
+/- [arrays::index_array_shared]:
+ Source: 'src/arrays.rs', lines 37:0-37:57 -/
def index_array_shared
(T : Type) (s : Array T 32#usize) (i : Usize) : Result T :=
Array.index_usize T 32#usize s i
-/- [array::index_array_u32]:
- Source: 'src/array.rs', lines 44:0-44:53 -/
+/- [arrays::index_array_u32]:
+ Source: 'src/arrays.rs', lines 44:0-44:53 -/
def index_array_u32 (s : Array U32 32#usize) (i : Usize) : Result U32 :=
Array.index_usize U32 32#usize s i
-/- [array::index_array_copy]:
- Source: 'src/array.rs', lines 48:0-48:45 -/
+/- [arrays::index_array_copy]:
+ Source: 'src/arrays.rs', lines 48:0-48:45 -/
def index_array_copy (x : Array U32 32#usize) : Result U32 :=
Array.index_usize U32 32#usize x 0#usize
-/- [array::index_mut_array]:
- Source: 'src/array.rs', lines 52:0-52:62 -/
+/- [arrays::index_mut_array]:
+ Source: 'src/arrays.rs', lines 52:0-52:62 -/
def index_mut_array
(T : Type) (s : Array T 32#usize) (i : Usize) :
Result (T × (T → Result (Array T 32#usize)))
@@ -80,13 +80,13 @@ def index_mut_array
let (t, index_mut_back) ← Array.index_mut_usize T 32#usize s i
Result.ret (t, index_mut_back)
-/- [array::index_slice]:
- Source: 'src/array.rs', lines 56:0-56:46 -/
+/- [arrays::index_slice]:
+ Source: 'src/arrays.rs', lines 56:0-56:46 -/
def index_slice (T : Type) (s : Slice T) (i : Usize) : Result T :=
Slice.index_usize T s i
-/- [array::index_mut_slice]:
- Source: 'src/array.rs', lines 60:0-60:58 -/
+/- [arrays::index_mut_slice]:
+ Source: 'src/arrays.rs', lines 60:0-60:58 -/
def index_mut_slice
(T : Type) (s : Slice T) (i : Usize) :
Result (T × (T → Result (Slice T)))
@@ -95,16 +95,16 @@ def index_mut_slice
let (t, index_mut_back) ← Slice.index_mut_usize T s i
Result.ret (t, index_mut_back)
-/- [array::slice_subslice_shared_]:
- Source: 'src/array.rs', lines 64:0-64:70 -/
+/- [arrays::slice_subslice_shared_]:
+ Source: 'src/arrays.rs', lines 64:0-64: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)
(core.slice.index.SliceIndexRangeUsizeSliceTInst U32) x
{ start := y, end_ := z }
-/- [array::slice_subslice_mut_]:
- Source: 'src/array.rs', lines 68:0-68:75 -/
+/- [arrays::slice_subslice_mut_]:
+ Source: 'src/arrays.rs', lines 68:0-68:75 -/
def slice_subslice_mut_
(x : Slice U32) (y : Usize) (z : Usize) :
Result ((Slice U32) × (Slice U32 → Result (Slice U32)))
@@ -116,13 +116,13 @@ def slice_subslice_mut_
{ start := y, end_ := z }
Result.ret (s, index_mut_back)
-/- [array::array_to_slice_shared_]:
- Source: 'src/array.rs', lines 72:0-72:54 -/
+/- [arrays::array_to_slice_shared_]:
+ Source: 'src/arrays.rs', lines 72:0-72:54 -/
def array_to_slice_shared_ (x : Array U32 32#usize) : Result (Slice U32) :=
Array.to_slice U32 32#usize x
-/- [array::array_to_slice_mut_]:
- Source: 'src/array.rs', lines 76:0-76:59 -/
+/- [arrays::array_to_slice_mut_]:
+ Source: 'src/arrays.rs', lines 76:0-76:59 -/
def array_to_slice_mut_
(x : Array U32 32#usize) :
Result ((Slice U32) × (Slice U32 → Result (Array U32 32#usize)))
@@ -131,8 +131,8 @@ def array_to_slice_mut_
let (s, to_slice_mut_back) ← Array.to_slice_mut U32 32#usize x
Result.ret (s, to_slice_mut_back)
-/- [array::array_subslice_shared_]:
- Source: 'src/array.rs', lines 80:0-80:74 -/
+/- [arrays::array_subslice_shared_]:
+ Source: 'src/arrays.rs', lines 80:0-80: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
@@ -140,8 +140,8 @@ def array_subslice_shared_
(core.slice.index.SliceIndexRangeUsizeSliceTInst U32)) x
{ start := y, end_ := z }
-/- [array::array_subslice_mut_]:
- Source: 'src/array.rs', lines 84:0-84:79 -/
+/- [arrays::array_subslice_mut_]:
+ Source: 'src/arrays.rs', lines 84:0-84:79 -/
def array_subslice_mut_
(x : Array U32 32#usize) (y : Usize) (z : Usize) :
Result ((Slice U32) × (Slice U32 → Result (Array U32 32#usize)))
@@ -154,18 +154,18 @@ def array_subslice_mut_
{ start := y, end_ := z }
Result.ret (s, index_mut_back)
-/- [array::index_slice_0]:
- Source: 'src/array.rs', lines 88:0-88:38 -/
+/- [arrays::index_slice_0]:
+ Source: 'src/arrays.rs', lines 88:0-88:38 -/
def index_slice_0 (T : Type) (s : Slice T) : Result T :=
Slice.index_usize T s 0#usize
-/- [array::index_array_0]:
- Source: 'src/array.rs', lines 92:0-92:42 -/
+/- [arrays::index_array_0]:
+ Source: 'src/arrays.rs', lines 92:0-92:42 -/
def index_array_0 (T : Type) (s : Array T 32#usize) : Result T :=
Array.index_usize T 32#usize s 0#usize
-/- [array::index_index_array]:
- Source: 'src/array.rs', lines 103:0-103:71 -/
+/- [arrays::index_index_array]:
+ Source: 'src/arrays.rs', lines 103:0-103:71 -/
def index_index_array
(s : Array (Array U32 32#usize) 32#usize) (i : Usize) (j : Usize) :
Result U32
@@ -174,8 +174,8 @@ def index_index_array
let a ← Array.index_usize (Array U32 32#usize) 32#usize s i
Array.index_usize U32 32#usize a j
-/- [array::update_update_array]:
- Source: 'src/array.rs', lines 114:0-114:70 -/
+/- [arrays::update_update_array]:
+ Source: 'src/arrays.rs', lines 114:0-114:70 -/
def update_update_array
(s : Array (Array U32 32#usize) 32#usize) (i : Usize) (j : Usize) :
Result Unit
@@ -188,46 +188,46 @@ def update_update_array
let _ ← index_mut_back a1
Result.ret ()
-/- [array::array_local_deep_copy]:
- Source: 'src/array.rs', lines 118:0-118:43 -/
+/- [arrays::array_local_deep_copy]:
+ Source: 'src/arrays.rs', lines 118:0-118:43 -/
def array_local_deep_copy (x : Array U32 32#usize) : Result Unit :=
Result.ret ()
-/- [array::take_array]:
- Source: 'src/array.rs', lines 122:0-122:30 -/
+/- [arrays::take_array]:
+ Source: 'src/arrays.rs', lines 122:0-122:30 -/
def take_array (a : Array U32 2#usize) : Result Unit :=
Result.ret ()
-/- [array::take_array_borrow]:
- Source: 'src/array.rs', lines 123:0-123:38 -/
+/- [arrays::take_array_borrow]:
+ Source: 'src/arrays.rs', lines 123:0-123:38 -/
def take_array_borrow (a : Array U32 2#usize) : Result Unit :=
Result.ret ()
-/- [array::take_slice]:
- Source: 'src/array.rs', lines 124:0-124:28 -/
+/- [arrays::take_slice]:
+ Source: 'src/arrays.rs', lines 124:0-124:28 -/
def take_slice (s : Slice U32) : Result Unit :=
Result.ret ()
-/- [array::take_mut_slice]:
- Source: 'src/array.rs', lines 125:0-125:36 -/
+/- [arrays::take_mut_slice]:
+ Source: 'src/arrays.rs', lines 125:0-125:36 -/
def take_mut_slice (s : Slice U32) : Result (Slice U32) :=
Result.ret s
-/- [array::const_array]:
- Source: 'src/array.rs', lines 127:0-127:32 -/
+/- [arrays::const_array]:
+ Source: 'src/arrays.rs', lines 127:0-127:32 -/
def const_array : Result (Array U32 2#usize) :=
Result.ret (Array.make U32 2#usize [ 0#u32, 0#u32 ])
-/- [array::const_slice]:
- Source: 'src/array.rs', lines 131:0-131:20 -/
+/- [arrays::const_slice]:
+ Source: 'src/arrays.rs', lines 131:0-131:20 -/
def const_slice : Result Unit :=
do
let _ ←
Array.to_slice U32 2#usize (Array.make U32 2#usize [ 0#u32, 0#u32 ])
Result.ret ()
-/- [array::take_all]:
- Source: 'src/array.rs', lines 141:0-141:17 -/
+/- [arrays::take_all]:
+ Source: 'src/arrays.rs', lines 141:0-141:17 -/
def take_all : Result Unit :=
do
let _ ← take_array (Array.make U32 2#usize [ 0#u32, 0#u32 ])
@@ -242,30 +242,30 @@ def take_all : Result Unit :=
let _ ← to_slice_mut_back s2
Result.ret ()
-/- [array::index_array]:
- Source: 'src/array.rs', lines 155:0-155:38 -/
+/- [arrays::index_array]:
+ Source: 'src/arrays.rs', lines 155:0-155:38 -/
def index_array (x : Array U32 2#usize) : Result U32 :=
Array.index_usize U32 2#usize x 0#usize
-/- [array::index_array_borrow]:
- Source: 'src/array.rs', lines 158:0-158:46 -/
+/- [arrays::index_array_borrow]:
+ Source: 'src/arrays.rs', lines 158:0-158:46 -/
def index_array_borrow (x : Array U32 2#usize) : Result U32 :=
Array.index_usize U32 2#usize x 0#usize
-/- [array::index_slice_u32_0]:
- Source: 'src/array.rs', lines 162:0-162:42 -/
+/- [arrays::index_slice_u32_0]:
+ Source: 'src/arrays.rs', lines 162:0-162:42 -/
def index_slice_u32_0 (x : Slice U32) : Result U32 :=
Slice.index_usize U32 x 0#usize
-/- [array::index_mut_slice_u32_0]:
- Source: 'src/array.rs', lines 166:0-166:50 -/
+/- [arrays::index_mut_slice_u32_0]:
+ Source: 'src/arrays.rs', lines 166:0-166: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.ret (i, x)
-/- [array::index_all]:
- Source: 'src/array.rs', lines 170:0-170:25 -/
+/- [arrays::index_all]:
+ Source: 'src/arrays.rs', lines 170:0-170:25 -/
def index_all : Result U32 :=
do
let i ← index_array (Array.make U32 2#usize [ 0#u32, 0#u32 ])
@@ -284,31 +284,31 @@ def index_all : Result U32 :=
let _ ← to_slice_mut_back s2
Result.ret i8
-/- [array::update_array]:
- Source: 'src/array.rs', lines 184:0-184:36 -/
+/- [arrays::update_array]:
+ Source: 'src/arrays.rs', lines 184:0-184: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
let _ ← index_mut_back 1#u32
Result.ret ()
-/- [array::update_array_mut_borrow]:
- Source: 'src/array.rs', lines 187:0-187:48 -/
+/- [arrays::update_array_mut_borrow]:
+ Source: 'src/arrays.rs', lines 187:0-187:48 -/
def update_array_mut_borrow
(x : Array U32 2#usize) : Result (Array U32 2#usize) :=
do
let (_, index_mut_back) ← Array.index_mut_usize U32 2#usize x 0#usize
index_mut_back 1#u32
-/- [array::update_mut_slice]:
- Source: 'src/array.rs', lines 190:0-190:38 -/
+/- [arrays::update_mut_slice]:
+ Source: 'src/arrays.rs', lines 190:0-190: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
-/- [array::update_all]:
- Source: 'src/array.rs', lines 194:0-194:19 -/
+/- [arrays::update_all]:
+ Source: 'src/arrays.rs', lines 194:0-194:19 -/
def update_all : Result Unit :=
do
let _ ← update_array (Array.make U32 2#usize [ 0#u32, 0#u32 ])
@@ -319,8 +319,8 @@ def update_all : Result Unit :=
let _ ← to_slice_mut_back s1
Result.ret ()
-/- [array::range_all]:
- Source: 'src/array.rs', lines 205:0-205:18 -/
+/- [arrays::range_all]:
+ Source: 'src/arrays.rs', lines 205:0-205:18 -/
def range_all : Result Unit :=
do
let (s, index_mut_back) ←
@@ -333,33 +333,33 @@ def range_all : Result Unit :=
let _ ← index_mut_back s1
Result.ret ()
-/- [array::deref_array_borrow]:
- Source: 'src/array.rs', lines 214:0-214:46 -/
+/- [arrays::deref_array_borrow]:
+ Source: 'src/arrays.rs', lines 214:0-214:46 -/
def deref_array_borrow (x : Array U32 2#usize) : Result U32 :=
Array.index_usize U32 2#usize x 0#usize
-/- [array::deref_array_mut_borrow]:
- Source: 'src/array.rs', lines 219:0-219:54 -/
+/- [arrays::deref_array_mut_borrow]:
+ Source: 'src/arrays.rs', lines 219:0-219:54 -/
def deref_array_mut_borrow
(x : Array U32 2#usize) : Result (U32 × (Array U32 2#usize)) :=
do
let i ← Array.index_usize U32 2#usize x 0#usize
Result.ret (i, x)
-/- [array::take_array_t]:
- Source: 'src/array.rs', lines 227:0-227:31 -/
+/- [arrays::take_array_t]:
+ Source: 'src/arrays.rs', lines 227:0-227:31 -/
def take_array_t (a : Array AB 2#usize) : Result Unit :=
Result.ret ()
-/- [array::non_copyable_array]:
- Source: 'src/array.rs', lines 229:0-229:27 -/
+/- [arrays::non_copyable_array]:
+ Source: 'src/arrays.rs', lines 229:0-229:27 -/
def non_copyable_array : Result Unit :=
do
let _ ← take_array_t (Array.make AB 2#usize [ AB.A, AB.B ])
Result.ret ()
-/- [array::sum]: loop 0:
- Source: 'src/array.rs', lines 242:0-250:1 -/
+/- [arrays::sum]: loop 0:
+ Source: 'src/arrays.rs', lines 242:0-250:1 -/
divergent def sum_loop (s : Slice U32) (sum1 : U32) (i : Usize) : Result U32 :=
let i1 := Slice.len U32 s
if i < i1
@@ -371,13 +371,13 @@ divergent def sum_loop (s : Slice U32) (sum1 : U32) (i : Usize) : Result U32 :=
sum_loop s sum3 i3
else Result.ret sum1
-/- [array::sum]:
- Source: 'src/array.rs', lines 242:0-242:28 -/
+/- [arrays::sum]:
+ Source: 'src/arrays.rs', lines 242:0-242:28 -/
def sum (s : Slice U32) : Result U32 :=
sum_loop s 0#u32 0#usize
-/- [array::sum2]: loop 0:
- Source: 'src/array.rs', lines 252:0-261:1 -/
+/- [arrays::sum2]: loop 0:
+ Source: 'src/arrays.rs', lines 252:0-261:1 -/
divergent def sum2_loop
(s : Slice U32) (s2 : Slice U32) (sum1 : U32) (i : Usize) : Result U32 :=
let i1 := Slice.len U32 s
@@ -392,8 +392,8 @@ divergent def sum2_loop
sum2_loop s s2 sum3 i5
else Result.ret sum1
-/- [array::sum2]:
- Source: 'src/array.rs', lines 252:0-252:41 -/
+/- [arrays::sum2]:
+ Source: 'src/arrays.rs', lines 252:0-252:41 -/
def sum2 (s : Slice U32) (s2 : Slice U32) : Result U32 :=
let i := Slice.len U32 s
let i1 := Slice.len U32 s2
@@ -401,8 +401,8 @@ def sum2 (s : Slice U32) (s2 : Slice U32) : Result U32 :=
then Result.fail .panic
else sum2_loop s s2 0#u32 0#usize
-/- [array::f0]:
- Source: 'src/array.rs', lines 263:0-263:11 -/
+/- [arrays::f0]:
+ Source: 'src/arrays.rs', lines 263:0-263:11 -/
def f0 : Result Unit :=
do
let (s, to_slice_mut_back) ←
@@ -412,8 +412,8 @@ def f0 : Result Unit :=
let _ ← to_slice_mut_back s1
Result.ret ()
-/- [array::f1]:
- Source: 'src/array.rs', lines 268:0-268:11 -/
+/- [arrays::f1]:
+ Source: 'src/arrays.rs', lines 268:0-268:11 -/
def f1 : Result Unit :=
do
let (_, index_mut_back) ←
@@ -422,21 +422,21 @@ def f1 : Result Unit :=
let _ ← index_mut_back 1#u32
Result.ret ()
-/- [array::f2]:
- Source: 'src/array.rs', lines 273:0-273:17 -/
+/- [arrays::f2]:
+ Source: 'src/arrays.rs', lines 273:0-273:17 -/
def f2 (i : U32) : Result Unit :=
Result.ret ()
-/- [array::f4]:
- Source: 'src/array.rs', lines 282:0-282:54 -/
+/- [arrays::f4]:
+ Source: 'src/arrays.rs', lines 282:0-282: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)
(core.slice.index.SliceIndexRangeUsizeSliceTInst U32)) x
{ start := y, end_ := z }
-/- [array::f3]:
- Source: 'src/array.rs', lines 275:0-275:18 -/
+/- [arrays::f3]:
+ Source: 'src/arrays.rs', lines 275:0-275:18 -/
def f3 : Result U32 :=
do
let i ←
@@ -449,18 +449,18 @@ def f3 : Result U32 :=
let s1 ← f4 b 16#usize 18#usize
sum2 s s1
-/- [array::SZ]
- Source: 'src/array.rs', lines 286:0-286:19 -/
+/- [arrays::SZ]
+ Source: 'src/arrays.rs', lines 286:0-286:19 -/
def sz_body : Result Usize := Result.ret 32#usize
-def sz_c : Usize := eval_global sz_body (by simp)
+def sz_c : Usize := eval_global sz_body (by decide)
-/- [array::f5]:
- Source: 'src/array.rs', lines 289:0-289:31 -/
+/- [arrays::f5]:
+ Source: 'src/arrays.rs', lines 289:0-289:31 -/
def f5 (x : Array U32 32#usize) : Result U32 :=
Array.index_usize U32 32#usize x 0#usize
-/- [array::ite]:
- Source: 'src/array.rs', lines 294:0-294:12 -/
+/- [arrays::ite]:
+ Source: 'src/arrays.rs', lines 294:0-294:12 -/
def ite : Result Unit :=
do
let (s, to_slice_mut_back) ←
@@ -473,4 +473,4 @@ def ite : Result Unit :=
let _ ← to_slice_mut_back s1
Result.ret ()
-end array
+end arrays
diff --git a/tests/lean/Constants.lean b/tests/lean/Constants.lean
index 2912805f..4c626ab3 100644
--- a/tests/lean/Constants.lean
+++ b/tests/lean/Constants.lean
@@ -8,17 +8,17 @@ namespace constants
/- [constants::X0]
Source: 'src/constants.rs', lines 5:0-5:17 -/
def x0_body : Result U32 := Result.ret 0#u32
-def x0_c : U32 := eval_global x0_body (by simp)
+def x0_c : U32 := eval_global x0_body (by decide)
/- [constants::X1]
Source: 'src/constants.rs', lines 7:0-7:17 -/
def x1_body : Result U32 := Result.ret core_u32_max
-def x1_c : U32 := eval_global x1_body (by simp)
+def x1_c : U32 := eval_global x1_body (by decide)
/- [constants::X2]
Source: 'src/constants.rs', lines 10:0-10:17 -/
def x2_body : Result U32 := Result.ret 3#u32
-def x2_c : U32 := eval_global x2_body (by simp)
+def x2_c : U32 := eval_global x2_body (by decide)
/- [constants::incr]:
Source: 'src/constants.rs', lines 17:0-17:32 -/
@@ -28,7 +28,7 @@ def incr (n : U32) : Result U32 :=
/- [constants::X3]
Source: 'src/constants.rs', lines 15:0-15:17 -/
def x3_body : Result U32 := incr 32#u32
-def x3_c : U32 := eval_global x3_body (by simp)
+def x3_c : U32 := eval_global x3_body (by decide)
/- [constants::mk_pair0]:
Source: 'src/constants.rs', lines 23:0-23:51 -/
@@ -49,22 +49,22 @@ def mk_pair1 (x : U32) (y : U32) : Result (Pair U32 U32) :=
/- [constants::P0]
Source: 'src/constants.rs', lines 31:0-31:24 -/
def p0_body : Result (U32 × U32) := mk_pair0 0#u32 1#u32
-def p0_c : (U32 × U32) := eval_global p0_body (by simp)
+def p0_c : (U32 × U32) := eval_global p0_body (by decide)
/- [constants::P1]
Source: 'src/constants.rs', lines 32:0-32:28 -/
def p1_body : Result (Pair U32 U32) := mk_pair1 0#u32 1#u32
-def p1_c : Pair U32 U32 := eval_global p1_body (by simp)
+def p1_c : Pair U32 U32 := eval_global p1_body (by decide)
/- [constants::P2]
Source: 'src/constants.rs', lines 33:0-33:24 -/
def p2_body : Result (U32 × U32) := Result.ret (0#u32, 1#u32)
-def p2_c : (U32 × U32) := eval_global p2_body (by simp)
+def p2_c : (U32 × U32) := eval_global p2_body (by decide)
/- [constants::P3]
Source: 'src/constants.rs', lines 34:0-34:28 -/
def p3_body : Result (Pair U32 U32) := Result.ret { x := 0#u32, y := 1#u32 }
-def p3_c : Pair U32 U32 := eval_global p3_body (by simp)
+def p3_c : Pair U32 U32 := eval_global p3_body (by decide)
/- [constants::Wrap]
Source: 'src/constants.rs', lines 49:0-49:18 -/
@@ -79,7 +79,7 @@ def Wrap.new (T : Type) (value : T) : Result (Wrap T) :=
/- [constants::Y]
Source: 'src/constants.rs', lines 41:0-41:22 -/
def y_body : Result (Wrap I32) := Wrap.new I32 2#i32
-def y_c : Wrap I32 := eval_global y_body (by simp)
+def y_c : Wrap I32 := eval_global y_body (by decide)
/- [constants::unwrap_y]:
Source: 'src/constants.rs', lines 43:0-43:30 -/
@@ -89,12 +89,12 @@ def unwrap_y : Result I32 :=
/- [constants::YVAL]
Source: 'src/constants.rs', lines 47:0-47:19 -/
def yval_body : Result I32 := unwrap_y
-def yval_c : I32 := eval_global yval_body (by simp)
+def yval_c : I32 := eval_global yval_body (by decide)
/- [constants::get_z1::Z1]
Source: 'src/constants.rs', lines 62:4-62:17 -/
def get_z1_z1_body : Result I32 := Result.ret 3#i32
-def get_z1_z1_c : I32 := eval_global get_z1_z1_body (by simp)
+def get_z1_z1_c : I32 := eval_global get_z1_z1_body (by decide)
/- [constants::get_z1]:
Source: 'src/constants.rs', lines 61:0-61:28 -/
@@ -109,17 +109,17 @@ def add (a : I32) (b : I32) : Result I32 :=
/- [constants::Q1]
Source: 'src/constants.rs', lines 74:0-74:17 -/
def q1_body : Result I32 := Result.ret 5#i32
-def q1_c : I32 := eval_global q1_body (by simp)
+def q1_c : I32 := eval_global q1_body (by decide)
/- [constants::Q2]
Source: 'src/constants.rs', lines 75:0-75:17 -/
def q2_body : Result I32 := Result.ret q1_c
-def q2_c : I32 := eval_global q2_body (by simp)
+def q2_c : I32 := eval_global q2_body (by decide)
/- [constants::Q3]
Source: 'src/constants.rs', lines 76:0-76:17 -/
def q3_body : Result I32 := add q2_c 3#i32
-def q3_c : I32 := eval_global q3_body (by simp)
+def q3_c : I32 := eval_global q3_body (by decide)
/- [constants::get_z2]:
Source: 'src/constants.rs', lines 70:0-70:28 -/
@@ -132,21 +132,21 @@ def get_z2 : Result I32 :=
/- [constants::S1]
Source: 'src/constants.rs', lines 80:0-80:18 -/
def s1_body : Result U32 := Result.ret 6#u32
-def s1_c : U32 := eval_global s1_body (by simp)
+def s1_c : U32 := eval_global s1_body (by decide)
/- [constants::S2]
Source: 'src/constants.rs', lines 81:0-81:18 -/
def s2_body : Result U32 := incr s1_c
-def s2_c : U32 := eval_global s2_body (by simp)
+def s2_c : U32 := eval_global s2_body (by decide)
/- [constants::S3]
Source: 'src/constants.rs', lines 82:0-82:29 -/
def s3_body : Result (Pair U32 U32) := Result.ret p3_c
-def s3_c : Pair U32 U32 := eval_global s3_body (by simp)
+def s3_c : Pair U32 U32 := eval_global s3_body (by decide)
/- [constants::S4]
Source: 'src/constants.rs', lines 83:0-83:29 -/
def s4_body : Result (Pair U32 U32) := mk_pair1 7#u32 8#u32
-def s4_c : Pair U32 U32 := eval_global s4_body (by simp)
+def s4_c : Pair U32 U32 := eval_global s4_body (by decide)
end constants
diff --git a/tests/lean/NoNestedBorrows.lean b/tests/lean/NoNestedBorrows.lean
index 0dd29429..bed71d94 100644
--- a/tests/lean/NoNestedBorrows.lean
+++ b/tests/lean/NoNestedBorrows.lean
@@ -139,12 +139,12 @@ def mix_arith_i32 (x : I32) (y : I32) (z : I32) : Result I32 :=
/- [no_nested_borrows::CONST0]
Source: 'src/no_nested_borrows.rs', lines 125:0-125:23 -/
def const0_body : Result Usize := 1#usize + 1#usize
-def const0_c : Usize := eval_global const0_body (by simp)
+def const0_c : Usize := eval_global const0_body (by decide)
/- [no_nested_borrows::CONST1]
Source: 'src/no_nested_borrows.rs', lines 126:0-126:23 -/
def const1_body : Result Usize := 2#usize * 2#usize
-def const1_c : Usize := eval_global const1_body (by simp)
+def const1_c : Usize := eval_global const1_body (by decide)
/- [no_nested_borrows::cast_u32_to_i32]:
Source: 'src/no_nested_borrows.rs', lines 128:0-128:37 -/
diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean
index d32aba86..3ef4febc 100644
--- a/tests/lean/Traits.lean
+++ b/tests/lean/Traits.lean
@@ -249,7 +249,7 @@ def traits.ToTypetraitsBoolWrapperTInst (T : Type) (ToTypeBoolTInst : ToType
Source: 'src/traits.rs', lines 164:4-164:21 -/
def with_const_ty_len2_body : Result Usize := Result.ret 32#usize
def with_const_ty_len2_c : Usize :=
- eval_global with_const_ty_len2_body (by simp)
+ eval_global with_const_ty_len2_body (by decide)
/- Trait declaration: [traits::WithConstTy]
Source: 'src/traits.rs', lines 161:0-161:39 -/
@@ -264,7 +264,7 @@ structure WithConstTy (Self : Type) (LEN : Usize) where
/- [traits::{bool#8}::LEN1]
Source: 'src/traits.rs', lines 175:4-175:21 -/
def bool_len1_body : Result Usize := Result.ret 12#usize
-def bool_len1_c : Usize := eval_global bool_len1_body (by simp)
+def bool_len1_c : Usize := eval_global bool_len1_body (by decide)
/- [traits::{bool#8}::f]:
Source: 'src/traits.rs', lines 180:4-180:39 -/
diff --git a/tests/lean/Tutorial.lean b/tests/lean/Tutorial.lean
index 840a606e..d92b2dd7 100644
--- a/tests/lean/Tutorial.lean
+++ b/tests/lean/Tutorial.lean
@@ -376,7 +376,7 @@ theorem i32_id_spec (x : I32) (h : 0 ≤ x.val) :
--
-- We first specify a decreasing value. Here, we state that [x], seen as a natural number,
-- decreases at every recursive call.
-termination_by i32_id_spec x _ => x.val.toNat
+termination_by x.val.toNat
-- And we now have to prove that it indeed decreases - you can skip this for now.
decreasing_by
-- We first need to "massage" the goal (in practice, all the proofs of [decreasing_by]
diff --git a/tests/lean/lake-manifest.json b/tests/lean/lake-manifest.json
index 5c20ec3b..e167e841 100644
--- a/tests/lean/lake-manifest.json
+++ b/tests/lean/lake-manifest.json
@@ -1,56 +1,74 @@
-{"version": 5,
- "packagesDir": "lake-packages",
+{"version": 7,
+ "packagesDir": ".lake/packages",
"packages":
- [{"git":
- {"url": "https://github.com/EdAyers/ProofWidgets4",
- "subDir?": null,
- "rev": "a0c2cd0ac3245a0dade4f925bcfa97e06dd84229",
- "opts": {},
- "name": "proofwidgets",
- "inputRev?": "v0.0.13",
- "inherited": true}},
- {"path":
- {"opts": {},
- "name": "Base",
- "inherited": false,
- "dir": "./../../backends/lean"}},
- {"git":
- {"url": "https://github.com/mhuisi/lean4-cli.git",
- "subDir?": null,
- "rev": "21dac2e9cc7e3cf7da5800814787b833e680b2fd",
- "opts": {},
- "name": "Cli",
- "inputRev?": "nightly",
- "inherited": true}},
- {"git":
- {"url": "https://github.com/leanprover-community/mathlib4.git",
- "subDir?": null,
- "rev": "b639e46a19a0328adfb9b1fdf8cbe39dfc1de76b",
- "opts": {},
- "name": "mathlib",
- "inputRev?": null,
- "inherited": false}},
- {"git":
- {"url": "https://github.com/gebner/quote4",
- "subDir?": null,
- "rev": "e75daed95ad1c92af4e577fea95e234d7a8401c1",
- "opts": {},
- "name": "Qq",
- "inputRev?": "master",
- "inherited": true}},
- {"git":
- {"url": "https://github.com/JLimperg/aesop",
- "subDir?": null,
- "rev": "1a0cded2be292b5496e659b730d2accc742de098",
- "opts": {},
- "name": "aesop",
- "inputRev?": "master",
- "inherited": true}},
- {"git":
- {"url": "https://github.com/leanprover/std4",
- "subDir?": null,
- "rev": "ba5e5e3af519b4fc5221ad0fa4b2c87276f1d323",
- "opts": {},
- "name": "std",
- "inputRev?": "main",
- "inherited": true}}]}
+ [{"url": "https://github.com/leanprover/std4",
+ "type": "git",
+ "subDir": null,
+ "rev": "276953b13323ca151939eafaaec9129bf7970306",
+ "name": "std",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": "main",
+ "inherited": true,
+ "configFile": "lakefile.lean"},
+ {"url": "https://github.com/leanprover-community/quote4",
+ "type": "git",
+ "subDir": null,
+ "rev": "1c88406514a636d241903e2e288d21dc6d861e01",
+ "name": "Qq",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": "master",
+ "inherited": true,
+ "configFile": "lakefile.lean"},
+ {"url": "https://github.com/leanprover-community/aesop",
+ "type": "git",
+ "subDir": null,
+ "rev": "6beed82dcfbb7731d173cd517675df27d62ad0f4",
+ "name": "aesop",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": "master",
+ "inherited": true,
+ "configFile": "lakefile.lean"},
+ {"url": "https://github.com/leanprover-community/ProofWidgets4",
+ "type": "git",
+ "subDir": null,
+ "rev": "af1e86cf7a37389632a02f4a111e6b501b2b818f",
+ "name": "proofwidgets",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": "v0.0.27",
+ "inherited": true,
+ "configFile": "lakefile.lean"},
+ {"url": "https://github.com/leanprover/lean4-cli",
+ "type": "git",
+ "subDir": null,
+ "rev": "a751d21d4b68c999accb6fc5d960538af26ad5ec",
+ "name": "Cli",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": "main",
+ "inherited": true,
+ "configFile": "lakefile.lean"},
+ {"url": "https://github.com/leanprover-community/import-graph.git",
+ "type": "git",
+ "subDir": null,
+ "rev": "8079d2d1d0e073bde42eab159c24f4c2d0d3a871",
+ "name": "importGraph",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": "main",
+ "inherited": true,
+ "configFile": "lakefile.lean"},
+ {"url": "https://github.com/leanprover-community/mathlib4.git",
+ "type": "git",
+ "subDir": null,
+ "rev": "d04f8d39c0e47a0d73450b49f6c0665897cdcaf7",
+ "name": "mathlib",
+ "manifestFile": "lake-manifest.json",
+ "inputRev": null,
+ "inherited": false,
+ "configFile": "lakefile.lean"},
+ {"type": "path",
+ "name": "base",
+ "manifestFile": "lake-manifest.json",
+ "inherited": false,
+ "dir": "./../../backends/lean",
+ "configFile": "lakefile.lean"}],
+ "name": "Tests",
+ "lakeDir": ".lake"}
diff --git a/tests/lean/lakefile.lean b/tests/lean/lakefile.lean
index fef94971..781fc8b8 100644
--- a/tests/lean/lakefile.lean
+++ b/tests/lean/lakefile.lean
@@ -4,19 +4,19 @@ open Lake DSL
require mathlib from git
"https://github.com/leanprover-community/mathlib4.git"
-require Base from "../../backends/lean"
+require base from "../../backends/lean"
package «tests» {}
-@[default_target] lean_lib tutorial
-@[default_target] lean_lib betreeMain
-@[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
-@[default_target] lean_lib poloniusList
-@[default_target] lean_lib array
-@[default_target] lean_lib traits
+@[default_target] lean_lib Tutorial
+@[default_target] lean_lib BetreeMain
+@[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
+@[default_target] lean_lib PoloniusList
+@[default_target] lean_lib Arrays
+@[default_target] lean_lib Traits
diff --git a/tests/lean/lean-toolchain b/tests/lean/lean-toolchain
index fbca4d37..cfcdd327 100644
--- a/tests/lean/lean-toolchain
+++ b/tests/lean/lean-toolchain
@@ -1 +1 @@
-leanprover/lean4:v4.0.0 \ No newline at end of file
+leanprover/lean4:v4.6.0-rc1