@@ -469,22 +469,20 @@ inductively from `Fin n` starting from the left, not from the right. This implie
469
469
more help to realize that elements belong to the right types, i.e., we need to insert casts at
470
470
several places. -/
471
471
472
- -- Porting note: `i.castSucc` does not work like it did in Lean 3;
473
- -- `(castSucc i)` must be used.
474
472
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))
476
474
477
475
/-- 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
480
478
481
479
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 :=
483
481
rfl
484
482
485
483
/-- Adding an element at the end of an `n`-tuple, to get an `n+1`-tuple. The name `snoc` comes from
486
484
`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 :=
488
486
if h : i.val < n then _root_.cast (by rw [Fin.castSucc_castLT i h]) (p (castLT i h))
489
487
else _root_.cast (by rw [eq_last_of_not_lt h]) x
490
488
@@ -495,7 +493,7 @@ theorem init_snoc : init (snoc p x) = p := by
495
493
convert cast_eq rfl (p i)
496
494
497
495
@[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
499
497
simp only [snoc, coe_castSucc, is_lt, cast_eq, dite_true]
500
498
convert cast_eq rfl (p i)
501
499
@@ -525,7 +523,7 @@ theorem snoc_comp_nat_add {n m : ℕ} {α : Sort*} (f : Fin (m + n) → α) (a :
525
523
rw [natAdd_castSucc, snoc_castSucc]
526
524
527
525
@[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 )
529
527
(a : α (last (n + m))) (i : Fin n) : (snoc f a) (castAdd (m + 1 ) i) = f (castAdd m i) :=
530
528
dif_pos _
531
529
@@ -537,20 +535,20 @@ theorem snoc_comp_cast_add {n m : ℕ} {α : Sort*} (f : Fin (n + m) → α) (a
537
535
538
536
/-- Updating a tuple and adding an element at the end commute. -/
539
537
@[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
541
539
ext j
542
540
by_cases h : j.val < n
543
541
· rw [snoc]
544
542
simp only [h]
545
543
simp only [dif_pos]
546
544
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
549
547
have : update (snoc p x) j (_root_.cast C1 y) j = _root_.cast C1 y := by simp
550
548
convert this
551
549
· exact h'.symm
552
550
· 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']
554
552
have E2 : update p i y (castLT j h) = _root_.cast C2 y := by
555
553
have : update p (castLT j h) (_root_.cast C2 y) (castLT j h) = _root_.cast C2 y := by simp
556
554
convert this
@@ -593,7 +591,7 @@ theorem init_update_last : init (update q (last n) z) = init q := by
593
591
594
592
/-- Updating an element and taking the beginning commute. -/
595
593
@[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
597
595
ext j
598
596
by_cases h : j = i
599
597
· rw [h]
@@ -745,11 +743,11 @@ alias forall_iff_succ := forall_fin_succ
745
743
alias exists_iff_succ := exists_fin_succ
746
744
747
745
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 :=
749
747
⟨fun h ↦ ⟨h _, fun _ ↦ h _⟩, fun h ↦ lastCases h.1 h.2 ⟩
750
748
751
749
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
753
751
mp := by
754
752
rintro ⟨i, hi⟩
755
753
induction' i using lastCases
@@ -859,7 +857,7 @@ theorem insertNth_last (x : α (last n)) (p : ∀ j : Fin n, α ((last n).succAb
859
857
refine insertNth_eq_iff.2 ⟨by simp, ?_⟩
860
858
ext j
861
859
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
863
861
· rw [snoc_castSucc]
864
862
exact (cast_heq _ _).symm
865
863
· apply congr_arg_heq
0 commit comments