Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - chore: replace Fin.castIso and Fin.revPerm with Fin.cast and Fin.rev for the bump of Std #5847

Closed
wants to merge 15 commits into from
6 changes: 3 additions & 3 deletions Mathlib/Algebra/BigOperators/Fin.lean
Original file line number Diff line number Diff line change
Expand Up @@ -188,11 +188,11 @@ theorem prod_Ioi_succ {M : Type _} [CommMonoid M] {n : ℕ} (i : Fin n) (v : Fin

@[to_additive]
theorem prod_congr' {M : Type _} [CommMonoid M] {a b : ℕ} (f : Fin b → M) (h : a = b) :
(∏ i : Fin a, f (castIso h i)) = ∏ i : Fin b, f i := by
(∏ i : Fin a, f (cast h i)) = ∏ i : Fin b, f i := by
subst h
congr
#align fin.prod_congr' Fin.prod_congr'
#align fin.sum_congr' Fin.sum_congr'
#align fin.prod_congr' Fin.prod_congr'
#align fin.sum_congr' Fin.sum_congr'
Komyyy marked this conversation as resolved.
Show resolved Hide resolved

@[to_additive]
theorem prod_univ_add {M : Type _} [CommMonoid M] {a b : ℕ} (f : Fin (a + b) → M) :
Expand Down
12 changes: 6 additions & 6 deletions Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean
Original file line number Diff line number Diff line change
Expand Up @@ -53,7 +53,7 @@ the $y_i$ are in degree $n$. -/
theorem decomposition_Q (n q : ℕ) :
((Q q).f (n + 1) : X _[n + 1] ⟶ X _[n + 1]) =
∑ i : Fin (n + 1) in Finset.filter (fun i : Fin (n + 1) => (i : ℕ) < q) Finset.univ,
(P i).f (n + 1) ≫ X.δ i.revPerm.succ ≫ X.σ (Fin.revPerm i) := by
(P i).f (n + 1) ≫ X.δ i.rev.succ ≫ X.σ (Fin.rev i) := by
induction' q with q hq
· simp only [Nat.zero_eq, Q_zero, HomologicalComplex.zero_f_apply, Nat.not_lt_zero,
Finset.filter_False, Finset.sum_empty]
Expand All @@ -73,15 +73,15 @@ theorem decomposition_Q (n q : ℕ) :
congr
· have hnaq' : n = a + q := by linarith
simp only [Fin.val_mk, (HigherFacesVanish.of_P q n).comp_Hσ_eq hnaq',
q'.revPerm_eq hnaq', neg_neg]
q'.rev_eq hnaq', neg_neg]
rfl
· ext ⟨i, hi⟩
simp only [Nat.succ_eq_add_one, Nat.lt_succ_iff_lt_or_eq, Finset.mem_univ,
forall_true_left, Finset.mem_filter, lt_self_iff_false, or_true, and_self, not_true,
Finset.mem_erase, ne_eq, Fin.mk.injEq, true_and]
aesop
set_option linter.uppercaseLean3 false in
#align algebraic_topology.dold_kan.decomposition_Q AlgebraicTopology.DoldKan.decomposition_Q
#align algebraic_topology.dold_kan.decomposition_Q AlgebraicTopology.DoldKan.decomposition_Qₓ
Komyyy marked this conversation as resolved.
Show resolved Hide resolved

variable (X)

Expand All @@ -103,9 +103,9 @@ variable {X} {n : ℕ} {Z Z' : C} (f : MorphComponents X n Z) (g : X' ⟶ X) (h

/-- The morphism `X _[n+1] ⟶ Z ` associated to `f : MorphComponents X n Z`. -/
def φ {Z : C} (f : MorphComponents X n Z) : X _[n + 1] ⟶ Z :=
PInfty.f (n + 1) ≫ f.a + ∑ i : Fin (n + 1), (P i).f (n + 1) ≫ X.δ i.revPerm.succ ≫
f.b (Fin.revPerm i)
#align algebraic_topology.dold_kan.morph_components.φ AlgebraicTopology.DoldKan.MorphComponents.φ
PInfty.f (n + 1) ≫ f.a + ∑ i : Fin (n + 1), (P i).f (n + 1) ≫ X.δ i.rev.succ ≫
f.b (Fin.rev i)
#align algebraic_topology.dold_kan.morph_components.φ AlgebraicTopology.DoldKan.MorphComponents.φₓ
Komyyy marked this conversation as resolved.
Show resolved Hide resolved

variable (X n)

Expand Down
6 changes: 3 additions & 3 deletions Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean
Original file line number Diff line number Diff line change
Expand Up @@ -106,13 +106,13 @@ theorem σ_comp_P_eq_zero (X : SimplicialObject C) {n q : ℕ} (i : Fin (n + 1))
have hi' : i = Fin.castSucc ⟨i, by linarith⟩ := by
ext
simp only [Fin.castSucc_mk, Fin.eta]
have eq := hq j.revPerm.succ (by
simp only [← hk, Fin.revPerm_eq j hk.symm, Nat.succ_eq_add_one, Fin.succ_mk, Fin.val_mk]
have eq := hq j.rev.succ (by
simp only [← hk, Fin.rev_eq j hk.symm, Nat.succ_eq_add_one, Fin.succ_mk, Fin.val_mk]
linarith)
rw [HomologicalComplex.comp_f, assoc, assoc, assoc, hi',
SimplicialObject.σ_comp_σ_assoc, reassoc_of% eq, zero_comp, comp_zero, comp_zero,
comp_zero]
simp only [Fin.revPerm_eq j hk.symm, Fin.le_iff_val_le_val, Fin.val_mk]
simp only [Fin.rev_eq j hk.symm, Fin.le_iff_val_le_val, Fin.val_mk]
linarith
set_option linter.uppercaseLean3 false in
#align algebraic_topology.dold_kan.σ_comp_P_eq_zero AlgebraicTopology.DoldKan.σ_comp_P_eq_zero
Expand Down
14 changes: 7 additions & 7 deletions Mathlib/AlgebraicTopology/DoldKan/Faces.lean
Original file line number Diff line number Diff line change
Expand Up @@ -85,7 +85,7 @@ theorem comp_Hσ_eq {Y : C} {n a q : ℕ} {φ : Y ⟶ X _[n + 1]} (v : HigherFac
swap
· rintro ⟨k, hk⟩
suffices φ ≫ X.δ (⟨a + 2 + k, by linarith⟩ : Fin (n + 2)) = 0 by
simp only [this, Fin.natAdd_mk, Fin.castIso_mk, zero_comp, smul_zero]
simp only [this, Fin.natAdd_mk, Fin.cast_mk, zero_comp, smul_zero]
convert v ⟨a + k + 1, by linarith⟩ (by rw [Fin.val_mk]; linarith)
dsimp
linarith
Expand All @@ -95,21 +95,21 @@ theorem comp_Hσ_eq {Y : C} {n a q : ℕ} {φ : Y ⟶ X _[n + 1]} (v : HigherFac
· rintro ⟨k, hk⟩
rw [assoc, X.δ_comp_σ_of_gt', v.comp_δ_eq_zero_assoc, zero_comp, zsmul_zero]
· simp only [Fin.lt_iff_val_lt_val]
dsimp [Fin.natAdd, Fin.castIso]
dsimp [Fin.natAdd, Fin.cast]
linarith
· intro h
rw [Fin.pred_eq_iff_eq_succ, Fin.ext_iff] at h
dsimp [Fin.castIso] at h
dsimp [Fin.cast] at h
linarith
· dsimp [Fin.castIso, Fin.pred]
· dsimp [Fin.cast, Fin.pred]
rw [Nat.add_right_comm, Nat.add_sub_assoc (by norm_num : 1 ≤ 3)]
linarith
simp only [assoc]
conv_lhs =>
congr
· rw [Fin.sum_univ_castSucc]
· rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc]
dsimp [Fin.castIso, Fin.castLE, Fin.castLT]
dsimp [Fin.cast, Fin.castLE, Fin.castLT]
/- the purpose of the following `simplif` is to create three subgoals in order
to finish the proof -/
have simplif :
Expand Down Expand Up @@ -152,12 +152,12 @@ theorem comp_Hσ_eq_zero {Y : C} {n q : ℕ} {φ : Y ⟶ X _[n + 1]} (v : Higher
AlternatingFaceMapComplex.obj_d_eq]
rw [← Fin.sum_congr' _ (show 2 + (n + 1) = n + 1 + 2 by linarith), Fin.sum_trunc]
· simp only [Fin.sum_univ_castSucc, Fin.sum_univ_zero, zero_add, Fin.last, Fin.castLE_mk,
Fin.castIso_mk, Fin.castSucc_mk]
Fin.cast_mk, Fin.castSucc_mk]
simp only [Fin.mk_zero, Fin.val_zero, pow_zero, one_zsmul, Fin.mk_one, Fin.val_one, pow_one,
neg_smul, comp_neg]
erw [δ_comp_σ_self, δ_comp_σ_succ, add_right_neg]
· intro j
dsimp [Fin.castIso, Fin.castLE, Fin.castLT]
dsimp [Fin.cast, Fin.castLE, Fin.castLT]
rw [comp_zsmul, comp_zsmul, δ_comp_σ_of_gt', v.comp_δ_eq_zero_assoc, zero_comp, zsmul_zero]
· simp only [Fin.lt_iff_val_lt_val]
dsimp [Fin.succ]
Expand Down