Skip to content

Commit

Permalink
chore: rename fields of Subarray to follow Lean conventions (#3851)
Browse files Browse the repository at this point in the history
Co-authored-by: Eric Wieser <[email protected]>
  • Loading branch information
david-christiansen and eric-wieser authored Apr 13, 2024
1 parent 2e1ef22 commit 864221d
Show file tree
Hide file tree
Showing 4 changed files with 47 additions and 23 deletions.
2 changes: 2 additions & 0 deletions CODEOWNERS
Original file line number Diff line number Diff line change
Expand Up @@ -39,3 +39,5 @@
/src/Lean/Elab/Tactic/Guard.lean @digama0
/src/Init/Guard.lean @digama0
/src/Lean/Server/CodeActions/ @digama0
/src/Init/Data/Array/Subarray.lean @david-christiansen

2 changes: 2 additions & 0 deletions RELEASES.md
Original file line number Diff line number Diff line change
Expand Up @@ -112,6 +112,8 @@ fact.def :

* The coercion from `String` to `Name` was removed. Previously, it was `Name.mkSimple`, which does not separate strings at dots, but experience showed that this is not always the desired coercion. For the previous behavior, manually insert a call to `Name.mkSimple`.

* The `Subarray` fields `as`, `h₁` and `h₂` have been renamed to `array`, `start_le_stop`, and `stop_le_array_size`, respectively. This more closely follows standard Lean conventions. Deprecated aliases for the field projects were added; these will be removed in a future release.

v4.7.0
---------

Expand Down
63 changes: 41 additions & 22 deletions src/Init/Data/Array/Subarray.lean
Original file line number Diff line number Diff line change
Expand Up @@ -9,25 +9,34 @@ import Init.Data.Array.Basic
universe u v w

structure Subarray (α : Type u) where
as : Array α
array : Array α
start : Nat
stop : Nat
h₁ : start ≤ stop
h₂ : stop ≤ as.size
start_le_stop : start ≤ stop
stop_le_array_size : stop ≤ array.size

@[deprecated Subarray.array]
abbrev Subarray.as (s : Subarray α) : Array α := s.array

@[deprecated Subarray.start_le_stop]
theorem Subarray.h₁ (s : Subarray α) : s.start ≤ s.stop := s.start_le_stop

@[deprecated Subarray.stop_le_array_size]
theorem Subarray.h₂ (s : Subarray α) : s.stop ≤ s.as.size := s.stop_le_array_size

namespace Subarray

def size (s : Subarray α) : Nat :=
s.stop - s.start

def get (s : Subarray α) (i : Fin s.size) : α :=
have : s.start + i.val < s.as.size := by
apply Nat.lt_of_lt_of_le _ s.h₂
have : s.start + i.val < s.array.size := by
apply Nat.lt_of_lt_of_le _ s.stop_le_array_size
have := i.isLt
simp [size] at this
rw [Nat.add_comm]
exact Nat.add_lt_of_lt_sub this
s.as[s.start + i.val]
s.array[s.start + i.val]

instance : GetElem (Subarray α) Nat α fun xs i => i < xs.size where
getElem xs i h := xs.get ⟨i, h⟩
Expand All @@ -42,15 +51,15 @@ abbrev get! [Inhabited α] (s : Subarray α) (i : Nat) : α :=

def popFront (s : Subarray α) : Subarray α :=
if h : s.start < s.stop then
{ s with start := s.start + 1, h₁ := Nat.le_of_lt_succ (Nat.add_lt_add_right h 1) }
{ s with start := s.start + 1, start_le_stop := Nat.le_of_lt_succ (Nat.add_lt_add_right h 1) }
else
s

@[inline] unsafe def forInUnsafe {α : Type u} {β : Type v} {m : Type v → Type w} [Monad m] (s : Subarray α) (b : β) (f : α → β → m (ForInStep β)) : m β :=
let sz := USize.ofNat s.stop
let rec @[specialize] loop (i : USize) (b : β) : m β := do
if i < sz then
let a := s.as.uget i lcProof
let a := s.array.uget i lcProof
match (← f a b) with
| ForInStep.done b => pure b
| ForInStep.yield b => loop (i+1) b
Expand All @@ -68,27 +77,27 @@ instance : ForIn m (Subarray α) α where

@[inline]
def foldlM {α : Type u} {β : Type v} {m : Type v → Type w} [Monad m] (f : β → α → m β) (init : β) (as : Subarray α) : m β :=
as.as.foldlM f (init := init) (start := as.start) (stop := as.stop)
as.array.foldlM f (init := init) (start := as.start) (stop := as.stop)

@[inline]
def foldrM {α : Type u} {β : Type v} {m : Type v → Type w} [Monad m] (f : α → β → m β) (init : β) (as : Subarray α) : m β :=
as.as.foldrM f (init := init) (start := as.stop) (stop := as.start)
as.array.foldrM f (init := init) (start := as.stop) (stop := as.start)

@[inline]
def anyM {α : Type u} {m : TypeType w} [Monad m] (p : α → m Bool) (as : Subarray α) : m Bool :=
as.as.anyM p (start := as.start) (stop := as.stop)
as.array.anyM p (start := as.start) (stop := as.stop)

@[inline]
def allM {α : Type u} {m : TypeType w} [Monad m] (p : α → m Bool) (as : Subarray α) : m Bool :=
as.as.allM p (start := as.start) (stop := as.stop)
as.array.allM p (start := as.start) (stop := as.stop)

@[inline]
def forM {α : Type u} {m : Type v → Type w} [Monad m] (f : α → m PUnit) (as : Subarray α) : m PUnit :=
as.as.forM f (start := as.start) (stop := as.stop)
as.array.forM f (start := as.start) (stop := as.stop)

@[inline]
def forRevM {α : Type u} {m : Type v → Type w} [Monad m] (f : α → m PUnit) (as : Subarray α) : m PUnit :=
as.as.forRevM f (start := as.stop) (stop := as.start)
as.array.forRevM f (start := as.stop) (stop := as.start)

@[inline]
def foldl {α : Type u} {β : Type v} (f : β → α → β) (init : β) (as : Subarray α) : β :=
Expand Down Expand Up @@ -135,15 +144,25 @@ variable {α : Type u}

def toSubarray (as : Array α) (start : Nat := 0) (stop : Nat := as.size) : Subarray α :=
if h₂ : stop ≤ as.size then
if h₁ : start ≤ stop then
{ as := as, start := start, stop := stop, h₁ := h₁, h₂ := h₂ }
else
{ as := as, start := stop, stop := stop, h₁ := Nat.le_refl _, h₂ := h₂ }
if h₁ : start ≤ stop then
{ array := as, start := start, stop := stop,
start_le_stop := h₁, stop_le_array_size := h₂ }
else
{ array := as, start := stop, stop := stop,
start_le_stop := Nat.le_refl _, stop_le_array_size := h₂ }
else
if h₁ : start ≤ as.size then
{ as := as, start := start, stop := as.size, h₁ := h₁, h₂ := Nat.le_refl _ }
else
{ as := as, start := as.size, stop := as.size, h₁ := Nat.le_refl _, h₂ := Nat.le_refl _ }
if h₁ : start ≤ as.size then
{ array := as,
start := start,
stop := as.size,
start_le_stop := h₁,
stop_le_array_size := Nat.le_refl _ }
else
{ array := as,
start := as.size,
stop := as.size,
start_le_stop := Nat.le_refl _,
stop_le_array_size := Nat.le_refl _ }

@[coe]
def ofSubarray (s : Subarray α) : Array α := Id.run do
Expand Down
3 changes: 2 additions & 1 deletion src/Init/Data/Stream.lean
Original file line number Diff line number Diff line change
Expand Up @@ -94,7 +94,8 @@ instance : Stream (Subarray α) α where
next? s :=
if h : s.start < s.stop then
have : s.start + 1 ≤ s.stop := Nat.succ_le_of_lt h
some (s.as.get ⟨s.start, Nat.lt_of_lt_of_le h s.h₂⟩, { s with start := s.start + 1, h₁ := this })
some (s.as.get ⟨s.start, Nat.lt_of_lt_of_le h s.stop_le_array_size⟩,
{ s with start := s.start + 1, start_le_stop := this })
else
none

Expand Down

0 comments on commit 864221d

Please sign in to comment.