Skip to content

Commit 66f0c7e

Browse files
committed
chore: remove obsolete porting note about Finset.castSucc (#18929)
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).
1 parent 10702fe commit 66f0c7e

File tree

1 file changed

+15
-17
lines changed

1 file changed

+15
-17
lines changed

Mathlib/Data/Fin/Tuple/Basic.lean

Lines changed: 15 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -469,22 +469,20 @@ inductively from `Fin n` starting from the left, not from the right. This implie
469469
more help to realize that elements belong to the right types, i.e., we need to insert casts at
470470
several places. -/
471471

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

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

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

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

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

497495
@[simp]
498-
theorem snoc_castSucc : snoc p x (castSucc i) = p i := by
496+
theorem snoc_castSucc : snoc p x i.castSucc = p i := by
499497
simp only [snoc, coe_castSucc, is_lt, cast_eq, dite_true]
500498
convert cast_eq rfl (p i)
501499

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

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

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

538536
/-- Updating a tuple and adding an element at the end commute. -/
539537
@[simp]
540-
theorem snoc_update : snoc (update p i y) x = update (snoc p x) (castSucc i) y := by
538+
theorem snoc_update : snoc (update p i y) x = update (snoc p x) i.castSucc y := by
541539
ext j
542540
by_cases h : j.val < n
543541
· rw [snoc]
544542
simp only [h]
545543
simp only [dif_pos]
546544
by_cases h' : j = castSucc i
547-
· have C1 : α (castSucc i) = α j := by rw [h']
548-
have E1 : update (snoc p x) (castSucc i) y j = _root_.cast C1 y := by
545+
· have C1 : α i.castSucc = α j := by rw [h']
546+
have E1 : update (snoc p x) i.castSucc y j = _root_.cast C1 y := by
549547
have : update (snoc p x) j (_root_.cast C1 y) j = _root_.cast C1 y := by simp
550548
convert this
551549
· exact h'.symm
552550
· exact heq_of_cast_eq (congr_arg α (Eq.symm h')) rfl
553-
have C2 : α (castSucc i) = α (castSucc (castLT j h)) := by rw [castSucc_castLT, h']
551+
have C2 : α i.castSucc = α (castLT j h).castSucc := by rw [castSucc_castLT, h']
554552
have E2 : update p i y (castLT j h) = _root_.cast C2 y := by
555553
have : update p (castLT j h) (_root_.cast C2 y) (castLT j h) = _root_.cast C2 y := by simp
556554
convert this
@@ -593,7 +591,7 @@ theorem init_update_last : init (update q (last n) z) = init q := by
593591

594592
/-- Updating an element and taking the beginning commute. -/
595593
@[simp]
596-
theorem init_update_castSucc : init (update q (castSucc i) y) = update (init q) i y := by
594+
theorem init_update_castSucc : init (update q i.castSucc y) = update (init q) i y := by
597595
ext j
598596
by_cases h : j = i
599597
· rw [h]
@@ -745,11 +743,11 @@ alias forall_iff_succ := forall_fin_succ
745743
alias exists_iff_succ := exists_fin_succ
746744

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

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

0 commit comments

Comments
 (0)