Skip to content

Commit

Permalink
chore: remove obsolete porting note about Finset.castSucc (#18929)
Browse files Browse the repository at this point in the history
Removes a comment in `Data.Fin.Tuple.Basic` that says
> Porting note: `i.castSucc` does not work like it did in Lean 3

and rewrites all instances of `(castSucc i)` in that file to `i.castSucc`.

The comment was included when that file was first ported in a43b3f1.
Currently, however, there does not seem to be any issue with writing `i.castSucc`
as was done [in mathlib3](https://github.com/leanprover-community/mathlib3/blob/65a1391a0106c9204fe45bc73a039f056558cb83/src/data/fin/tuple/basic.lean#L382-L383).
  • Loading branch information
dwrensha authored and TobiasLeichtfried committed Nov 21, 2024
1 parent 5ba4267 commit 93651e9
Showing 1 changed file with 15 additions and 17 deletions.
32 changes: 15 additions & 17 deletions Mathlib/Data/Fin/Tuple/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -469,22 +469,20 @@ inductively from `Fin n` starting from the left, not from the right. This implie
more help to realize that elements belong to the right types, i.e., we need to insert casts at
several places. -/

-- Porting note: `i.castSucc` does not work like it did in Lean 3;
-- `(castSucc i)` must be used.
variable {α : Fin (n + 1) → Sort*} (x : α (last n)) (q : ∀ i, α i)
(p : ∀ i : Fin n, α (castSucc i)) (i : Fin n) (y : α (castSucc i)) (z : α (last n))
(p : ∀ i : Fin n, α i.castSucc) (i : Fin n) (y : α i.castSucc) (z : α (last n))

/-- The beginning of an `n+1` tuple, i.e., its first `n` entries -/
def init (q : ∀ i, α i) (i : Fin n) : α (castSucc i) :=
q (castSucc i)
def init (q : ∀ i, α i) (i : Fin n) : α i.castSucc :=
q i.castSucc

theorem init_def {q : ∀ i, α i} :
(init fun k : Fin (n + 1) ↦ q k) = fun k : Fin n ↦ q (castSucc k) :=
(init fun k : Fin (n + 1) ↦ q k) = fun k : Fin n ↦ q k.castSucc :=
rfl

/-- Adding an element at the end of an `n`-tuple, to get an `n+1`-tuple. The name `snoc` comes from
`cons` (i.e., adding an element to the left of a tuple) read in reverse order. -/
def snoc (p : ∀ i : Fin n, α (castSucc i)) (x : α (last n)) (i : Fin (n + 1)) : α i :=
def snoc (p : ∀ i : Fin n, α i.castSucc) (x : α (last n)) (i : Fin (n + 1)) : α i :=
if h : i.val < n then _root_.cast (by rw [Fin.castSucc_castLT i h]) (p (castLT i h))
else _root_.cast (by rw [eq_last_of_not_lt h]) x

Expand All @@ -495,7 +493,7 @@ theorem init_snoc : init (snoc p x) = p := by
convert cast_eq rfl (p i)

@[simp]
theorem snoc_castSucc : snoc p x (castSucc i) = p i := by
theorem snoc_castSucc : snoc p x i.castSucc = p i := by
simp only [snoc, coe_castSucc, is_lt, cast_eq, dite_true]
convert cast_eq rfl (p i)

Expand Down Expand Up @@ -525,7 +523,7 @@ theorem snoc_comp_nat_add {n m : ℕ} {α : Sort*} (f : Fin (m + n) → α) (a :
rw [natAdd_castSucc, snoc_castSucc]

@[simp]
theorem snoc_cast_add {α : Fin (n + m + 1) → Sort*} (f : ∀ i : Fin (n + m), α (castSucc i))
theorem snoc_cast_add {α : Fin (n + m + 1) → Sort*} (f : ∀ i : Fin (n + m), α i.castSucc)
(a : α (last (n + m))) (i : Fin n) : (snoc f a) (castAdd (m + 1) i) = f (castAdd m i) :=
dif_pos _

Expand All @@ -537,20 +535,20 @@ theorem snoc_comp_cast_add {n m : ℕ} {α : Sort*} (f : Fin (n + m) → α) (a

/-- Updating a tuple and adding an element at the end commute. -/
@[simp]
theorem snoc_update : snoc (update p i y) x = update (snoc p x) (castSucc i) y := by
theorem snoc_update : snoc (update p i y) x = update (snoc p x) i.castSucc y := by
ext j
by_cases h : j.val < n
· rw [snoc]
simp only [h]
simp only [dif_pos]
by_cases h' : j = castSucc i
· have C1 : α (castSucc i) = α j := by rw [h']
have E1 : update (snoc p x) (castSucc i) y j = _root_.cast C1 y := by
· have C1 : α i.castSucc = α j := by rw [h']
have E1 : update (snoc p x) i.castSucc y j = _root_.cast C1 y := by
have : update (snoc p x) j (_root_.cast C1 y) j = _root_.cast C1 y := by simp
convert this
· exact h'.symm
· exact heq_of_cast_eq (congr_arg α (Eq.symm h')) rfl
have C2 : α (castSucc i) = α (castSucc (castLT j h)) := by rw [castSucc_castLT, h']
have C2 : α i.castSucc = α (castLT j h).castSucc := by rw [castSucc_castLT, h']
have E2 : update p i y (castLT j h) = _root_.cast C2 y := by
have : update p (castLT j h) (_root_.cast C2 y) (castLT j h) = _root_.cast C2 y := by simp
convert this
Expand Down Expand Up @@ -593,7 +591,7 @@ theorem init_update_last : init (update q (last n) z) = init q := by

/-- Updating an element and taking the beginning commute. -/
@[simp]
theorem init_update_castSucc : init (update q (castSucc i) y) = update (init q) i y := by
theorem init_update_castSucc : init (update q i.castSucc y) = update (init q) i y := by
ext j
by_cases h : j = i
· rw [h]
Expand Down Expand Up @@ -745,11 +743,11 @@ alias forall_iff_succ := forall_fin_succ
alias exists_iff_succ := exists_fin_succ

lemma forall_iff_castSucc {P : Fin (n + 1) → Prop} :
(∀ i, P i) ↔ P (last n) ∧ ∀ i, P (castSucc i) :=
(∀ i, P i) ↔ P (last n) ∧ ∀ i : Fin n, P i.castSucc :=
fun h ↦ ⟨h _, fun _ ↦ h _⟩, fun h ↦ lastCases h.1 h.2

lemma exists_iff_castSucc {P : Fin (n + 1) → Prop} :
(∃ i, P i) ↔ P (last n) ∨ ∃ i, P (castSucc i) where
(∃ i, P i) ↔ P (last n) ∨ ∃ i : Fin n, P i.castSucc where
mp := by
rintro ⟨i, hi⟩
induction' i using lastCases
Expand Down Expand Up @@ -859,7 +857,7 @@ theorem insertNth_last (x : α (last n)) (p : ∀ j : Fin n, α ((last n).succAb
refine insertNth_eq_iff.2by simp, ?_⟩
ext j
apply eq_of_heq
trans snoc (fun j ↦ _root_.cast (congr_arg α (succAbove_last_apply j)) (p j)) x (castSucc j)
trans snoc (fun j ↦ _root_.cast (congr_arg α (succAbove_last_apply j)) (p j)) x j.castSucc
· rw [snoc_castSucc]
exact (cast_heq _ _).symm
· apply congr_arg_heq
Expand Down

0 comments on commit 93651e9

Please sign in to comment.