Skip to content

Commit 600654a

Browse files
committed
feat(Order): Fin.succAboveOrderIso (#21303)
This PR introduces the order isomorphism `Fin.succAboveOrderIso i : Fin (n + 1) ≃o ({i}ᶜ : Finset (Fin (n + 2)))` for `i : Fin (n + 2)`.
1 parent 9d5ba20 commit 600654a

File tree

3 files changed

+49
-0
lines changed

3 files changed

+49
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -4247,6 +4247,7 @@ import Mathlib.Order.Filter.Ultrafilter.Defs
42474247
import Mathlib.Order.Filter.ZeroAndBoundedAtFilter
42484248
import Mathlib.Order.Fin.Basic
42494249
import Mathlib.Order.Fin.Finset
4250+
import Mathlib.Order.Fin.SuccAboveOrderIso
42504251
import Mathlib.Order.Fin.Tuple
42514252
import Mathlib.Order.FixedPoints
42524253
import Mathlib.Order.GaloisConnection.Basic

Mathlib/Data/Fin/Basic.lean

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1265,6 +1265,17 @@ lemma succAbove_predAbove {p : Fin n} {i : Fin (n + 1)} (h : i ≠ castSucc p) :
12651265
· rw [predAbove_of_le_castSucc _ _ (Fin.le_of_lt h), succAbove_castPred_of_lt _ _ h]
12661266
· rw [predAbove_of_castSucc_lt _ _ h, succAbove_pred_of_lt _ _ h]
12671267

1268+
/-- Sending `Fin (n+1)` to `Fin n` by subtracting one from anything above `p`
1269+
then back to `Fin (n+1)` with a gap around `p.succ` is the identity away from `p.succ`. -/
1270+
@[simp]
1271+
lemma succ_succAbove_predAbove {n : ℕ} {p : Fin n} {i : Fin (n + 1)} (h : i ≠ p.succ) :
1272+
p.succ.succAbove (p.predAbove i) = i := by
1273+
obtain h | h := Fin.lt_or_lt_of_ne h
1274+
· rw [predAbove_of_le_castSucc _ _ (le_castSucc_iff.2 h),
1275+
succAbove_castPred_of_lt _ _ h]
1276+
· rw [predAbove_of_castSucc_lt _ _ (Fin.lt_of_le_of_lt (p.castSucc_le_succ) h),
1277+
succAbove_pred_of_lt _ _ h]
1278+
12681279
/-- Sending `Fin n` into `Fin (n + 1)` with a gap at `p`
12691280
then back to `Fin n` by subtracting one from anything above `p` is the identity. -/
12701281
@[simp]
Lines changed: 37 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,37 @@
1+
/-
2+
Copyright (c) 2024 Joël Riou. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Joël Riou
5+
-/
6+
7+
import Mathlib.Order.Fin.Basic
8+
import Mathlib.Data.Fintype.Basic
9+
import Mathlib.Tactic.FinCases
10+
11+
/-!
12+
# The order isomorphism `Fin (n + 1) ≃o {i}ᶜ`
13+
14+
Given `i : Fin (n + 2)`, we show that `Fin.succAboveOrderEmb` induces
15+
an order isomorphism `Fin (n + 1) ≃o ({i}ᶜ : Finset (Fin (n + 2)))`.
16+
17+
-/
18+
19+
open Finset
20+
21+
/-- Given `i : Fin (n + 2)`, this is the order isomorphism
22+
between `Fin (n + 1)` and the finite set `{i}ᶜ`. -/
23+
noncomputable def Fin.succAboveOrderIso {n : ℕ} (i : Fin (n + 2)) :
24+
Fin (n + 1) ≃o ({i}ᶜ : Finset (Fin (n + 2))) where
25+
toEquiv :=
26+
Equiv.ofBijective (f := fun a ↦ ⟨Fin.succAboveOrderEmb i a,
27+
by simpa using Fin.succAbove_ne i a⟩) (by
28+
constructor
29+
· intro a b h
30+
exact (Fin.succAboveOrderEmb i).injective (by simpa using h)
31+
· rintro ⟨j, hj⟩
32+
simp only [mem_compl, mem_singleton] at hj
33+
obtain rfl | ⟨i, rfl⟩ := Fin.eq_zero_or_eq_succ i
34+
· exact ⟨j.pred hj, by simp⟩
35+
· exact ⟨i.predAbove j, by aesop⟩)
36+
map_rel_iff' {a b} := by
37+
simp only [Equiv.ofBijective_apply, Subtype.mk_le_mk, OrderEmbedding.le_iff_le]

0 commit comments

Comments
 (0)