summaryrefslogtreecommitdiff
path: root/backends/lean/Base
diff options
context:
space:
mode:
Diffstat (limited to 'backends/lean/Base')
-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
10 files changed, 123 insertions, 111 deletions
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