Skip to content

Commit

Permalink
Merge pull request #68 from AeneasVerif/son/update_lean
Browse files Browse the repository at this point in the history
Update Lean to v4.6.0-rc1
  • Loading branch information
sonmarcho committed Feb 2, 2024
2 parents 0960ad1 + 9cc912e commit eb8bddc
Show file tree
Hide file tree
Showing 45 changed files with 851 additions and 1,246 deletions.
16 changes: 8 additions & 8 deletions Makefile
Expand Up @@ -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
Expand Down Expand Up @@ -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
Expand Down
4 changes: 2 additions & 2 deletions README.md
Expand Up @@ -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
Expand Down
3 changes: 2 additions & 1 deletion backends/lean/.gitignore
@@ -1,2 +1,3 @@
lake-packages/
build/
build/
.lake
37 changes: 22 additions & 15 deletions backends/lean/Base/Diverge/Base.lean
Expand Up @@ -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
Expand Down Expand Up @@ -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
Expand All @@ -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]
Expand All @@ -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
Expand Down Expand Up @@ -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))
Expand Down Expand Up @@ -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))
Expand Down Expand Up @@ -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
Expand Down
11 changes: 4 additions & 7 deletions backends/lean/Base/Diverge/Elab.lean
Expand Up @@ -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.
Expand Down Expand Up @@ -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
Expand Down Expand Up @@ -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
Expand Down
16 changes: 10 additions & 6 deletions backends/lean/Base/Diverge/ElabBase.lean
Expand Up @@ -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"
Expand All @@ -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 => ...)`
Expand All @@ -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"
Expand All @@ -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
Expand Down
10 changes: 5 additions & 5 deletions backends/lean/Base/Extensions.lean
Expand Up @@ -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,
Expand Down
25 changes: 14 additions & 11 deletions backends/lean/Base/IList/IList.lean
Expand Up @@ -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)
Expand All @@ -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
Expand Down Expand Up @@ -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]
Expand All @@ -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 :=
Expand All @@ -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) :
Expand All @@ -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]
Expand All @@ -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
Expand Down
2 changes: 1 addition & 1 deletion backends/lean/Base/Primitives/ArraySlice.lean
Expand Up @@ -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 :=
Expand Down

0 comments on commit eb8bddc

Please sign in to comment.