Skip to content

Commit

Permalink
feat: align {List/Array/Vector}.{attach,attachWith,pmap} lemmas (lean…
Browse files Browse the repository at this point in the history
…prover#6723)

This PR completes the alignment of
{List/Array/Vector}.{attach,attachWith,pmap} lemmas. I had to fill in a
number of gaps in the List API.
  • Loading branch information
kim-em authored and luisacicolini committed Jan 21, 2025
1 parent 8f98fee commit 774244b
Show file tree
Hide file tree
Showing 8 changed files with 738 additions and 27 deletions.
115 changes: 98 additions & 17 deletions src/Init/Data/Array/Attach.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,6 +6,7 @@ Authors: Joachim Breitner, Mario Carneiro
prelude
import Init.Data.Array.Mem
import Init.Data.Array.Lemmas
import Init.Data.Array.Count
import Init.Data.List.Attach

namespace Array
Expand Down Expand Up @@ -142,10 +143,16 @@ theorem pmap_eq_map_attach {p : α → Prop} (f : ∀ a, p a → β) (l H) :
cases l
simp [List.pmap_eq_map_attach]

@[simp]
theorem pmap_eq_attachWith {p q : α → Prop} (f : ∀ a, p a → q a) (l H) :
pmap (fun a h => ⟨a, f a h⟩) l H = l.attachWith q (fun x h => f x (H x h)) := by
cases l
simp [List.pmap_eq_attachWith]

theorem attach_map_coe (l : Array α) (f : α → β) :
(l.attach.map fun (i : {i // i ∈ l}) => f i) = l.map f := by
cases l
simp [List.attach_map_coe]
simp

theorem attach_map_val (l : Array α) (f : α → β) : (l.attach.map fun i => f i.val) = l.map f :=
attach_map_coe _ _
Expand All @@ -172,6 +179,12 @@ theorem mem_attach (l : Array α) : ∀ x, x ∈ l.attach
rcases this with ⟨⟨_, _⟩, m, rfl⟩
exact m

@[simp]
theorem mem_attachWith (l : Array α) {q : α → Prop} (H) (x : {x // q x}) :
x ∈ l.attachWith q H ↔ x.1 ∈ l := by
cases l
simp

@[simp]
theorem mem_pmap {p : α → Prop} {f : ∀ a, p a → β} {l H b} :
b ∈ pmap f l H ↔ ∃ (a : _) (h : a ∈ l), f a (H a h) = b := by
Expand Down Expand Up @@ -223,16 +236,16 @@ theorem attachWith_ne_empty_iff {l : Array α} {P : α → Prop} {H : ∀ a ∈
cases l; simp

@[simp]
theorem getElem?_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : Array α} (h : ∀ a ∈ l, p a) (n : Nat) :
(pmap f l h)[n]? = Option.pmap f l[n]? fun x H => h x (mem_of_getElem? H) := by
theorem getElem?_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : Array α} (h : ∀ a ∈ l, p a) (i : Nat) :
(pmap f l h)[i]? = Option.pmap f l[i]? fun x H => h x (mem_of_getElem? H) := by
cases l; simp

@[simp]
theorem getElem_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : Array α} (h : ∀ a ∈ l, p a) {n : Nat}
(hn : n < (pmap f l h).size) :
(pmap f l h)[n] =
f (l[n]'(@size_pmap _ _ p f l h ▸ hn))
(h _ (getElem_mem (@size_pmap _ _ p f l h ▸ hn))) := by
theorem getElem_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : Array α} (h : ∀ a ∈ l, p a) {i : Nat}
(hi : i < (pmap f l h).size) :
(pmap f l h)[i] =
f (l[i]'(@size_pmap _ _ p f l h ▸ hi))
(h _ (getElem_mem (@size_pmap _ _ p f l h ▸ hi))) := by
cases l; simp

@[simp]
Expand All @@ -256,6 +269,18 @@ theorem getElem_attach {xs : Array α} {i : Nat} (h : i < xs.attach.size) :
xs.attach[i] = ⟨xs[i]'(by simpa using h), getElem_mem (by simpa using h)⟩ :=
getElem_attachWith h

@[simp] theorem pmap_attach (l : Array α) {p : {x // x ∈ l} → Prop} (f : ∀ a, p a → β) (H) :
pmap f l.attach H =
l.pmap (P := fun a => ∃ h : a ∈ l, p ⟨a, h⟩)
(fun a h => f ⟨a, h.1⟩ h.2) (fun a h => ⟨h, H ⟨a, h⟩ (by simp)⟩) := by
ext <;> simp

@[simp] theorem pmap_attachWith (l : Array α) {p : {x // q x} → Prop} (f : ∀ a, p a → β) (H₁ H₂) :
pmap f (l.attachWith q H₁) H₂ =
l.pmap (P := fun a => ∃ h : q a, p ⟨a, h⟩)
(fun a h => f ⟨a, h.1⟩ h.2) (fun a h => ⟨H₁ _ h, H₂ ⟨a, H₁ _ h⟩ (by simpa)⟩) := by
ext <;> simp

theorem foldl_pmap (l : Array α) {P : α → Prop} (f : (a : α) → P a → β)
(H : ∀ (a : α), a ∈ l → P a) (g : γ → β → γ) (x : γ) :
(l.pmap f H).foldl g x = l.attach.foldl (fun acc a => g acc (f a.1 (H _ a.2))) x := by
Expand Down Expand Up @@ -313,11 +338,7 @@ theorem attachWith_map {l : Array α} (f : α → β) {P : β → Prop} {H : ∀
(l.map f).attachWith P H = (l.attachWith (P ∘ f) (fun _ h => H _ (mem_map_of_mem f h))).map
fun ⟨x, h⟩ => ⟨f x, h⟩ := by
cases l
ext
· simp
· simp only [List.map_toArray, List.attachWith_toArray, List.getElem_toArray,
List.getElem_attachWith, List.getElem_map, Function.comp_apply]
erw [List.getElem_attachWith] -- Why is `erw` needed here?
simp [List.attachWith_map]

theorem map_attachWith {l : Array α} {P : α → Prop} {H : ∀ (a : α), a ∈ l → P a}
(f : { x // P x } → β) :
Expand Down Expand Up @@ -347,7 +368,23 @@ theorem attach_filter {l : Array α} (p : α → Bool) :
simp [List.attach_filter, List.map_filterMap, Function.comp_def]

-- We are still missing here `attachWith_filterMap` and `attachWith_filter`.
-- Also missing are `filterMap_attach`, `filter_attach`, `filterMap_attachWith` and `filter_attachWith`.

@[simp]
theorem filterMap_attachWith {q : α → Prop} {l : Array α} {f : {x // q x} → Option β} (H)
(w : stop = (l.attachWith q H).size) :
(l.attachWith q H).filterMap f 0 stop = l.attach.filterMap (fun ⟨x, h⟩ => f ⟨x, H _ h⟩) := by
subst w
cases l
simp [Function.comp_def]

@[simp]
theorem filter_attachWith {q : α → Prop} {l : Array α} {p : {x // q x} → Bool} (H)
(w : stop = (l.attachWith q H).size) :
(l.attachWith q H).filter p 0 stop =
(l.attach.filter (fun ⟨x, h⟩ => p ⟨x, H _ h⟩)).map (fun ⟨x, h⟩ => ⟨x, H _ h⟩) := by
subst w
cases l
simp [Function.comp_def, List.filter_map]

theorem pmap_pmap {p : α → Prop} {q : β → Prop} (g : ∀ a, p a → β) (f : ∀ b, q b → γ) (l H₁ H₂) :
pmap f (pmap g l H₁) H₂ =
Expand Down Expand Up @@ -427,16 +464,48 @@ theorem reverse_attach (xs : Array α) :

@[simp] theorem back?_attachWith {P : α → Prop} {xs : Array α}
{H : ∀ (a : α), a ∈ xs → P a} :
(xs.attachWith P H).back? = xs.back?.pbind (fun a h => some ⟨a, H _ (mem_of_back?_eq_some h)⟩) := by
(xs.attachWith P H).back? = xs.back?.pbind (fun a h => some ⟨a, H _ (mem_of_back? h)⟩) := by
cases xs
simp

@[simp]
theorem back?_attach {xs : Array α} :
xs.attach.back? = xs.back?.pbind fun a h => some ⟨a, mem_of_back?_eq_some h⟩ := by
xs.attach.back? = xs.back?.pbind fun a h => some ⟨a, mem_of_back? h⟩ := by
cases xs
simp

@[simp]
theorem countP_attach (l : Array α) (p : α → Bool) :
l.attach.countP (fun a : {x // x ∈ l} => p a) = l.countP p := by
cases l
simp [Function.comp_def]

@[simp]
theorem countP_attachWith {p : α → Prop} (l : Array α) (H : ∀ a ∈ l, p a) (q : α → Bool) :
(l.attachWith p H).countP (fun a : {x // p x} => q a) = l.countP q := by
cases l
simp

@[simp]
theorem count_attach [DecidableEq α] (l : Array α) (a : {x // x ∈ l}) :
l.attach.count a = l.count ↑a := by
rcases l with ⟨l⟩
simp only [List.attach_toArray, List.attachWith_mem_toArray, List.count_toArray]
rw [List.map_attach, List.count_eq_countP]
simp only [Subtype.beq_iff]
rw [List.countP_pmap, List.countP_attach (p := (fun x => x == a.1)), List.count]

@[simp]
theorem count_attachWith [DecidableEq α] {p : α → Prop} (l : Array α) (H : ∀ a ∈ l, p a) (a : {x // p x}) :
(l.attachWith p H).count a = l.count ↑a := by
cases l
simp

@[simp] theorem countP_pmap {p : α → Prop} (g : ∀ a, p a → β) (f : β → Bool) (l : Array α) (H₁) :
(l.pmap g H₁).countP f =
l.attach.countP (fun ⟨a, m⟩ => f (g a (H₁ a m))) := by
simp [pmap_eq_map_attach, countP_map, Function.comp_def]

/-! ## unattach
`Array.unattach` is the (one-sided) inverse of `Array.attach`. It is a synonym for `Array.map Subtype.val`.
Expand All @@ -455,7 +524,7 @@ and is ideally subsequently simplified away by `unattach_attach`.
If not, usually the right approach is `simp [Array.unattach, -Array.map_subtype]` to unfold.
-/
def unattach {α : Type _} {p : α → Prop} (l : Array { x // p x }) := l.map (·.val)
def unattach {α : Type _} {p : α → Prop} (l : Array { x // p x }) : Array α := l.map (·.val)

@[simp] theorem unattach_nil {p : α → Prop} : (#[] : Array { x // p x }).unattach = #[] := rfl
@[simp] theorem unattach_push {p : α → Prop} {a : { x // p x }} {l : Array { x // p x }} :
Expand Down Expand Up @@ -578,4 +647,16 @@ and simplifies these to the function directly taking the value.
cases l₂
simp

@[simp] theorem unattach_flatten {p : α → Prop} {l : Array (Array { x // p x })} :
l.flatten.unattach = (l.map unattach).flatten := by
unfold unattach
cases l using array₂_induction
simp only [flatten_toArray, List.map_map, Function.comp_def, List.map_id_fun', id_eq,
List.map_toArray, List.map_flatten, map_subtype, map_id_fun', List.unattach_toArray, mk.injEq]
simp only [List.unattach]

@[simp] theorem unattach_mkArray {p : α → Prop} {n : Nat} {x : { x // p x }} :
(Array.mkArray n x).unattach = Array.mkArray n x.1 := by
simp [unattach]

end Array
6 changes: 4 additions & 2 deletions src/Init/Data/Array/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3253,9 +3253,11 @@ theorem back!_eq_back? [Inhabited α] (a : Array α) : a.back! = a.back?.getD de
@[simp] theorem back!_push [Inhabited α] (a : Array α) : (a.push x).back! = x := by
simp [back!_eq_back?]

theorem mem_of_back?_eq_some {xs : Array α} {a : α} (h : xs.back? = some a) : a ∈ xs := by
theorem mem_of_back? {xs : Array α} {a : α} (h : xs.back? = some a) : a ∈ xs := by
cases xs
simpa using List.mem_of_getLast?_eq_some (by simpa using h)
simpa using List.mem_of_getLast? (by simpa using h)

@[deprecated mem_of_back? (since := "2024-10-21")] abbrev mem_of_back?_eq_some := @mem_of_back?

theorem getElem?_push_lt (a : Array α) (x : α) (i : Nat) (h : i < a.size) :
(a.push x)[i]? = some a[i] := by
Expand Down
9 changes: 9 additions & 0 deletions src/Init/Data/Bool.lean
Original file line number Diff line number Diff line change
Expand Up @@ -620,3 +620,12 @@ but may be used locally.
-/
def boolRelToRel : Coe (α → α → Bool) (α → α → Prop) where
coe r := fun a b => Eq (r a b) true

/-! ### subtypes -/

@[simp] theorem Subtype.beq_iff {α : Type u} [DecidableEq α] {p : α → Prop} {x y : {a : α // p a}} :
(x == y) = (x.1 == y.1) := by
cases x
cases y
rw [Bool.eq_iff_iff]
simp [beq_iff_eq]
66 changes: 61 additions & 5 deletions src/Init/Data/List/Attach.lean
Original file line number Diff line number Diff line change
Expand Up @@ -111,6 +111,14 @@ theorem pmap_eq_map_attach {p : α → Prop} (f : ∀ a, p a → β) (l H) :
pmap f l H = l.attach.map fun x => f x.1 (H _ x.2) := by
rw [attach, attachWith, map_pmap]; exact pmap_congr_left l fun _ _ _ _ => rfl

@[simp]
theorem pmap_eq_attachWith {p q : α → Prop} (f : ∀ a, p a → q a) (l H) :
pmap (fun a h => ⟨a, f a h⟩) l H = l.attachWith q (fun x h => f x (H x h)) := by
induction l with
| nil => rfl
| cons a l ih =>
simp [pmap, attachWith, ih]

theorem attach_map_coe (l : List α) (f : α → β) :
(l.attach.map fun (i : {i // i ∈ l}) => f i) = l.map f := by
rw [attach, attachWith, map_pmap]; exact pmap_eq_map _ _ _ _
Expand All @@ -136,10 +144,23 @@ theorem attachWith_map_subtype_val {p : α → Prop} (l : List α) (H : ∀ a
@[simp]
theorem mem_attach (l : List α) : ∀ x, x ∈ l.attach
| ⟨a, h⟩ => by
have := mem_map.1 (by rw [attach_map_subtype_val] <;> exact h)
have := mem_map.1 (by rw [attach_map_subtype_val]; exact h)
rcases this with ⟨⟨_, _⟩, m, rfl⟩
exact m

@[simp]
theorem mem_attachWith (l : List α) {q : α → Prop} (H) (x : {x // q x}) :
x ∈ l.attachWith q H ↔ x.1 ∈ l := by
induction l with
| nil => simp
| cons a l ih =>
simp [ih]
constructor
· rintro (_ | _) <;> simp_all
· rintro (h | h)
· simp [← h]
· simp_all

@[simp]
theorem mem_pmap {p : α → Prop} {f : ∀ a, p a → β} {l H b} :
b ∈ pmap f l H ↔ ∃ (a : _) (h : a ∈ l), f a (H a h) = b := by
Expand Down Expand Up @@ -266,6 +287,18 @@ theorem getElem_attach {xs : List α} {i : Nat} (h : i < xs.attach.length) :
xs.attach[i] = ⟨xs[i]'(by simpa using h), getElem_mem (by simpa using h)⟩ :=
getElem_attachWith h

@[simp] theorem pmap_attach (l : List α) {p : {x // x ∈ l} → Prop} (f : ∀ a, p a → β) (H) :
pmap f l.attach H =
l.pmap (P := fun a => ∃ h : a ∈ l, p ⟨a, h⟩)
(fun a h => f ⟨a, h.1⟩ h.2) (fun a h => ⟨h, H ⟨a, h⟩ (by simp)⟩) := by
apply ext_getElem <;> simp

@[simp] theorem pmap_attachWith (l : List α) {p : {x // q x} → Prop} (f : ∀ a, p a → β) (H₁ H₂) :
pmap f (l.attachWith q H₁) H₂ =
l.pmap (P := fun a => ∃ h : q a, p ⟨a, h⟩)
(fun a h => f ⟨a, h.1⟩ h.2) (fun a h => ⟨H₁ _ h, H₂ ⟨a, H₁ _ h⟩ (by simpa)⟩) := by
apply ext_getElem <;> simp

@[simp] theorem head?_pmap {P : α → Prop} (f : (a : α) → P a → β) (xs : List α)
(H : ∀ (a : α), a ∈ xs → P a) :
(xs.pmap f H).head? = xs.attach.head?.map fun ⟨a, m⟩ => f a (H a m) := by
Expand Down Expand Up @@ -431,7 +464,25 @@ theorem attach_filter {l : List α} (p : α → Bool) :
split <;> simp

-- We are still missing here `attachWith_filterMap` and `attachWith_filter`.
-- Also missing are `filterMap_attach`, `filter_attach`, `filterMap_attachWith` and `filter_attachWith`.

@[simp]
theorem filterMap_attachWith {q : α → Prop} {l : List α} {f : {x // q x} → Option β} (H) :
(l.attachWith q H).filterMap f = l.attach.filterMap (fun ⟨x, h⟩ => f ⟨x, H _ h⟩) := by
induction l with
| nil => rfl
| cons x xs ih =>
simp only [attachWith_cons, filterMap_cons]
split <;> simp_all [Function.comp_def]

@[simp]
theorem filter_attachWith {q : α → Prop} {l : List α} {p : {x // q x} → Bool} (H) :
(l.attachWith q H).filter p =
(l.attach.filter (fun ⟨x, h⟩ => p ⟨x, H _ h⟩)).map (fun ⟨x, h⟩ => ⟨x, H _ h⟩) := by
induction l with
| nil => rfl
| cons x xs ih =>
simp only [attachWith_cons, filter_cons]
split <;> simp_all [Function.comp_def, filter_map]

theorem pmap_pmap {p : α → Prop} {q : β → Prop} (g : ∀ a, p a → β) (f : ∀ b, q b → γ) (l H₁ H₂) :
pmap f (pmap g l H₁) H₂ =
Expand Down Expand Up @@ -520,7 +571,7 @@ theorem reverse_attach (xs : List α) :

@[simp] theorem getLast?_attachWith {P : α → Prop} {xs : List α}
{H : ∀ (a : α), a ∈ xs → P a} :
(xs.attachWith P H).getLast? = xs.getLast?.pbind (fun a h => some ⟨a, H _ (mem_of_getLast?_eq_some h)⟩) := by
(xs.attachWith P H).getLast? = xs.getLast?.pbind (fun a h => some ⟨a, H _ (mem_of_getLast? h)⟩) := by
rw [getLast?_eq_head?_reverse, reverse_attachWith, head?_attachWith]
simp

Expand All @@ -531,7 +582,7 @@ theorem reverse_attach (xs : List α) :

@[simp]
theorem getLast?_attach {xs : List α} :
xs.attach.getLast? = xs.getLast?.pbind fun a h => some ⟨a, mem_of_getLast?_eq_some h⟩ := by
xs.attach.getLast? = xs.getLast?.pbind fun a h => some ⟨a, mem_of_getLast? h⟩ := by
rw [getLast?_eq_head?_reverse, reverse_attach, head?_map, head?_attach]
simp

Expand Down Expand Up @@ -560,6 +611,11 @@ theorem count_attachWith [DecidableEq α] {p : α → Prop} (l : List α) (H :
(l.attachWith p H).count a = l.count ↑a :=
Eq.trans (countP_congr fun _ _ => by simp [Subtype.ext_iff]) <| countP_attachWith _ _ _

@[simp] theorem countP_pmap {p : α → Prop} (g : ∀ a, p a → β) (f : β → Bool) (l : List α) (H₁) :
(l.pmap g H₁).countP f =
l.attach.countP (fun ⟨a, m⟩ => f (g a (H₁ a m))) := by
simp [pmap_eq_map_attach, countP_map, Function.comp_def]

/-! ## unattach
`List.unattach` is the (one-sided) inverse of `List.attach`. It is a synonym for `List.map Subtype.val`.
Expand All @@ -578,7 +634,7 @@ and is ideally subsequently simplified away by `unattach_attach`.
If not, usually the right approach is `simp [List.unattach, -List.map_subtype]` to unfold.
-/
def unattach {α : Type _} {p : α → Prop} (l : List { x // p x }) := l.map (·.val)
def unattach {α : Type _} {p : α → Prop} (l : List { x // p x }) : List α := l.map (·.val)

@[simp] theorem unattach_nil {p : α → Prop} : ([] : List { x // p x }).unattach = [] := rfl
@[simp] theorem unattach_cons {p : α → Prop} {a : { x // p x }} {l : List { x // p x }} :
Expand Down
4 changes: 3 additions & 1 deletion src/Init/Data/List/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2750,10 +2750,12 @@ theorem getLast?_eq_some_iff {xs : List α} {a : α} : xs.getLast? = some a ↔
rw [getLast?_eq_head?_reverse, head?_isSome]
simp

theorem mem_of_getLast?_eq_some {xs : List α} {a : α} (h : xs.getLast? = some a) : a ∈ xs := by
theorem mem_of_getLast? {xs : List α} {a : α} (h : xs.getLast? = some a) : a ∈ xs := by
obtain ⟨ys, rfl⟩ := getLast?_eq_some_iff.1 h
exact mem_concat_self ys a

@[deprecated mem_of_getLast? (since := "2024-10-21")] abbrev mem_of_getLast?_eq_some := @mem_of_getLast?

@[simp] theorem getLast_reverse {l : List α} (h : l.reverse ≠ []) :
l.reverse.getLast h = l.head (by simp_all) := by
simp [getLast_eq_head_reverse]
Expand Down
1 change: 1 addition & 0 deletions src/Init/Data/Subtype.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,6 +5,7 @@ Authors: Johannes Hölzl
-/
prelude
import Init.Ext
import Init.Core

namespace Subtype

Expand Down
Loading

0 comments on commit 774244b

Please sign in to comment.