From fc0b39c4fe48fdbab06d5fd32e0a2b7dcae674e6 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Thu, 4 Apr 2024 15:29:11 +0200 Subject: Fix the coerce notation for scalars and update some lemmas --- backends/lean/Base/Primitives/Scalar.lean | 19 +++++++++++++++---- 1 file changed, 15 insertions(+), 4 deletions(-) (limited to 'backends/lean/Base/Primitives') diff --git a/backends/lean/Base/Primitives/Scalar.lean b/backends/lean/Base/Primitives/Scalar.lean index 3d90f1a5..7668bc59 100644 --- a/backends/lean/Base/Primitives/Scalar.lean +++ b/backends/lean/Base/Primitives/Scalar.lean @@ -265,6 +265,14 @@ theorem Scalar.cMax_suffices ty (h : x ≤ Scalar.cMax ty) : x ≤ Scalar.max ty have := Scalar.cMax_bound ty linarith +/-- The scalar type. + + We could use a subtype, but it using a custom structure type allows us + to have more control over the coercions and the simplifications (we tried + using a subtype and it caused issues especially as we had to make the Scalar + type non-reducible, so that we could have more control, but leading to + some natural equalities not being obvious to the simplifier anymore). + -/ structure Scalar (ty : ScalarTy) where val : Int hmin : Scalar.min ty ≤ val @@ -274,6 +282,9 @@ deriving Repr instance (ty : ScalarTy) : CoeOut (Scalar ty) Int where coe := λ v => v.val +/- Activate the ↑ notation -/ +attribute [coe] Scalar.val + theorem Scalar.bound_suffices (ty : ScalarTy) (x : Int) : Scalar.cMin ty ≤ x ∧ x ≤ Scalar.cMax ty -> Scalar.min ty ≤ x ∧ x ≤ Scalar.max ty @@ -1119,19 +1130,19 @@ theorem Scalar.eq_equiv {ty : ScalarTy} (x y : Scalar ty) : -- This is sometimes useful when rewriting the goal with the local assumptions @[simp] theorem Scalar.eq_imp {ty : ScalarTy} (x y : Scalar ty) : - x = y → (↑x : Int) = ↑y := (eq_equiv x y).mp + (↑x : Int) = ↑y → x = y := (eq_equiv x y).mpr theorem Scalar.lt_equiv {ty : ScalarTy} (x y : Scalar ty) : x < y ↔ (↑x : Int) < ↑y := by simp [LT.lt] @[simp] theorem Scalar.lt_imp {ty : ScalarTy} (x y : Scalar ty) : - x < y → (↑x : Int) < ↑y := (lt_equiv x y).mp + (↑x : Int) < (↑y) → x < y := (lt_equiv x y).mpr theorem Scalar.le_equiv {ty : ScalarTy} (x y : Scalar ty) : x ≤ y ↔ (↑x : Int) ≤ ↑y := by simp [LE.le] @[simp] theorem Scalar.le_imp {ty : ScalarTy} (x y : Scalar ty) : - x ≤ y → (↑x : Int) ≤ ↑y := (le_equiv x y).mp + (↑x : Int) ≤ ↑y → x ≤ y := (le_equiv x y).mpr instance Scalar.decLt {ty} (a b : Scalar ty) : Decidable (LT.lt a b) := Int.decLt .. instance Scalar.decLe {ty} (a b : Scalar ty) : Decidable (LE.le a b) := Int.decLe .. @@ -1152,6 +1163,6 @@ instance (ty : ScalarTy) : DecidableEq (Scalar ty) := | isFalse h => isFalse (Scalar.ne_of_val_ne h) @[simp] theorem Scalar.neq_to_neq_val {ty} : ∀ {i j : Scalar ty}, (¬ i = j) ↔ ¬ i.val = j.val := by - intro i j; cases i; cases j; simp + simp [eq_equiv] end Primitives -- cgit v1.2.3 From 65a77968d0abc2d01da92aa8982256855e7519a6 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Fri, 5 Apr 2024 14:04:25 +0200 Subject: Update the lean toolchain and fix the proofs --- backends/lean/Base/Primitives/ArraySlice.lean | 5 ++--- backends/lean/Base/Primitives/Range.lean | 1 - backends/lean/Base/Primitives/Vec.lean | 1 - 3 files changed, 2 insertions(+), 5 deletions(-) (limited to 'backends/lean/Base/Primitives') diff --git a/backends/lean/Base/Primitives/ArraySlice.lean b/backends/lean/Base/Primitives/ArraySlice.lean index e1a39d40..3bd2aebb 100644 --- a/backends/lean/Base/Primitives/ArraySlice.lean +++ b/backends/lean/Base/Primitives/ArraySlice.lean @@ -2,7 +2,6 @@ import Lean import Lean.Meta.Tactic.Simp import Init.Data.List.Basic -import Mathlib.Tactic.RunCmd import Mathlib.Tactic.Linarith import Base.IList import Base.Primitives.Scalar @@ -269,7 +268,7 @@ def Array.update_subslice (α : Type u) (n : Usize) (a : Array α n) (r : Range . scalar_tac . scalar_tac let na := s_beg.append (s.val.append s_end) - have : na.len = a.val.len := by simp [*] + have : na.len = a.val.len := by simp [na, *] ret ⟨ na, by simp_all [← List.len_eq_length]; scalar_tac ⟩ else fail panic @@ -343,7 +342,7 @@ def Slice.update_subslice (α : Type u) (s : Slice α) (r : Range Usize) (ss : S . scalar_tac . scalar_tac let ns := s_beg.append (ss.val.append s_end) - have : ns.len = s.val.len := by simp [*] + have : ns.len = s.val.len := by simp [ns, *] ret ⟨ ns, by simp_all [← List.len_eq_length]; scalar_tac ⟩ else fail panic diff --git a/backends/lean/Base/Primitives/Range.lean b/backends/lean/Base/Primitives/Range.lean index a268bcba..416cd201 100644 --- a/backends/lean/Base/Primitives/Range.lean +++ b/backends/lean/Base/Primitives/Range.lean @@ -2,7 +2,6 @@ import Lean import Lean.Meta.Tactic.Simp import Init.Data.List.Basic -import Mathlib.Tactic.RunCmd import Mathlib.Tactic.Linarith import Base.IList import Base.Primitives.Scalar diff --git a/backends/lean/Base/Primitives/Vec.lean b/backends/lean/Base/Primitives/Vec.lean index 65249c12..2b8425d8 100644 --- a/backends/lean/Base/Primitives/Vec.lean +++ b/backends/lean/Base/Primitives/Vec.lean @@ -2,7 +2,6 @@ import Lean import Lean.Meta.Tactic.Simp import Init.Data.List.Basic -import Mathlib.Tactic.RunCmd import Mathlib.Tactic.Linarith import Base.IList import Base.Primitives.Scalar -- cgit v1.2.3