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] - feat(Logic/Equiv/Basic): sumSigmaDistrib, finSigmaFinEquiv #19618

Closed
wants to merge 3 commits into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
42 changes: 42 additions & 0 deletions Mathlib/Algebra/BigOperators/Fin.lean
Original file line number Diff line number Diff line change
Expand Up @@ -388,6 +388,48 @@ theorem finPiFinEquiv_single {m : ℕ} {n : Fin m → ℕ} [∀ i, NeZero (n i)]
rintro x hx
rw [Pi.single_eq_of_ne hx, Fin.val_zero', zero_mul]

/-- Equivalence between the Sigma type `(i : Fin m) × Fin (n i)` and `Fin (∑ i : Fin m, n i)`. -/
def finSigmaFinEquiv {m : ℕ} {n : Fin m → ℕ} : (i : Fin m) × Fin (n i) ≃ Fin (∑ i : Fin m, n i) :=
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Could you please provide a few API lemmas that will help working with this defn?

Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I added finSigmaFinEquiv_apply and finSigmaFinEquiv_one. I would add finSigmaFinEquiv_symm_apply but I don't think there's a particularly clean statement of the inverse equivalence (at least, not clean enough to be useful - I think the forward apply should be preferred)

Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Note that there's WIP in Batteries to provide a more cohesive framework for Fin equivalences: leanprover-community/batteries#1007.

My PR also defines auxiliary functions that explicitly identify the inverse equivalence #19013.

Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I see, I guess one can write it in terms of these divSum and modSum you defined. Well, I guess we both have definitions of this equivalence. I'll leave mine like this (without a the _symm_apply lemma), I'm fine with either definition getting merged. #19695 just requires some definition of finSigmaFinEquiv.

Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I'm also fine with either definitions, and wonder if we could have the best of both worlds (explicit formulas from my PR, plus theorems that relate it to your PR).

Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I am merging this PR. @quangvdao could you please merge master into your PR and resolve the conflicts?

match m with
| 0 => @Equiv.equivOfIsEmpty _ _ _ (by simp; exact Fin.isEmpty')
| Nat.succ m =>
calc _ ≃ _ := (@finSumFinEquiv m 1).sigmaCongrLeft.symm
_ ≃ _ := Equiv.sumSigmaDistrib _
_ ≃ _ := finSigmaFinEquiv.sumCongr (Equiv.uniqueSigma _)
_ ≃ _ := finSumFinEquiv
_ ≃ _ := finCongr (Fin.sum_univ_castSucc n).symm

@[simp]
theorem finSigmaFinEquiv_apply {m : ℕ} {n : Fin m → ℕ} (k : (i : Fin m) × Fin (n i)) :
(finSigmaFinEquiv k : ℕ) = ∑ i : Fin k.1, n (Fin.castLE k.1.2.le i) + k.2 := by
induction m
· exact k.fst.elim0
rename_i m ih
rcases k with ⟨⟨iv,hi⟩,j⟩
rw [finSigmaFinEquiv]
unfold finSumFinEquiv
simp only [Equiv.coe_fn_mk, Equiv.sigmaCongrLeft, Equiv.coe_fn_symm_mk, Equiv.instTrans_trans,
Equiv.trans_apply, finCongr_apply, Fin.coe_cast]
conv =>
enter [1,1,3]
apply Equiv.sumCongr_apply
by_cases him : iv < m
· conv in Sigma.mk _ _ =>
equals ⟨Sum.inl ⟨iv, him⟩, j⟩ => simp [Fin.addCases, him]
simpa using ih _
· replace him := Nat.eq_of_lt_succ_of_not_lt hi him
subst him
conv in Sigma.mk _ _ =>
equals ⟨Sum.inr 0, j⟩ => simp [Fin.addCases, Fin.natAdd]
simp
rfl

/-- `finSigmaFinEquiv` on `Fin 1 × f` is just `f`-/
theorem finSigmaFinEquiv_one {n : Fin 1 → ℕ} (ij : (i : Fin 1) × Fin (n i)) :
(finSigmaFinEquiv ij : ℕ) = ij.2 := by
rw [finSigmaFinEquiv_apply, add_left_eq_self]
apply @Finset.sum_of_isEmpty _ _ _ _ (by simpa using Fin.isEmpty')

namespace List

section CommMonoid
Expand Down
32 changes: 31 additions & 1 deletion Mathlib/Logic/Equiv/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -215,6 +215,23 @@ theorem sigmaUnique_symm_apply {α} {β : α → Type*} [∀ a, Unique (β a)] (
(sigmaUnique α β).symm x = ⟨x, default⟩ :=
rfl

/-- Any `Unique` type is a left identity for type sigma up to equivalence. Compare with `uniqueProd`
which is non-dependent. -/
def uniqueSigma {α} (β : α → Type*) [Unique α] : (i:α) × β i ≃ β default :=
⟨fun p ↦ (Unique.eq_default _).rec p.2,
fun b ↦ ⟨default, b⟩,
fun _ ↦ Sigma.ext (Unique.default_eq _) (eqRec_heq _ _),
fun _ ↦ rfl⟩

theorem uniqueSigma_apply {α} {β : α → Type*} [Unique α] (x : (a : α) × β a) :
uniqueSigma β x = (Unique.eq_default _).rec x.2 :=
rfl

@[simp]
theorem uniqueSigma_symm_apply {α} {β : α → Type*} [Unique α] (y : β default) :
(uniqueSigma β).symm y = ⟨default, y⟩ :=
rfl

/-- `Empty` type is a right absorbing element for type product up to an equivalence. -/
def prodEmpty (α) : α × Empty ≃ Empty :=
equivEmpty _
Expand Down Expand Up @@ -919,14 +936,27 @@ theorem prodSumDistrib_symm_apply_right {α β γ} (a : α × γ) :
(prodSumDistrib α β γ).symm (inr a) = (a.1, inr a.2) :=
rfl

/-- An indexed sum of disjoint sums of types is equivalent to the sum of the indexed sums. -/
/-- An indexed sum of disjoint sums of types is equivalent to the sum of the indexed sums. Compare
with `Equiv.sumSigmaDistrib` which is indexed by sums. -/
@[simps]
def sigmaSumDistrib {ι} (α β : ι → Type*) :
(Σ i, α i ⊕ β i) ≃ (Σ i, α i) ⊕ (Σ i, β i) :=
⟨fun p => p.2.map (Sigma.mk p.1) (Sigma.mk p.1),
Sum.elim (Sigma.map id fun _ => Sum.inl) (Sigma.map id fun _ => Sum.inr), fun p => by
rcases p with ⟨i, a | b⟩ <;> rfl, fun p => by rcases p with (⟨i, a⟩ | ⟨i, b⟩) <;> rfl⟩

/-- A type indexed by disjoint sums of types is equivalent to the sum of the sums. Compare with
`Equiv.sigmaSumDistrib` which has the sums as the output type. -/
@[simps]
def sumSigmaDistrib {α β} (t : α ⊕ β → Type*) :
(Σ i, t i) ≃ (Σ i, t (.inl i)) ⊕ (Σ i, t (.inr i)) :=
⟨(match · with
| .mk (.inl x) y => .inl ⟨x, y⟩
| .mk (.inr x) y => .inr ⟨x, y⟩),
Sum.elim (fun a ↦ ⟨.inl a.1, a.2⟩) (fun b ↦ ⟨.inr b.1, b.2⟩),
by rintro ⟨x|x,y⟩ <;> simp,
by rintro (⟨x,y⟩|⟨x,y⟩) <;> simp⟩

/-- The product of an indexed sum of types (formally, a `Sigma`-type `Σ i, α i`) by a type `β` is
equivalent to the sum of products `Σ i, (α i × β)`. -/
@[simps apply symm_apply]
Expand Down
Loading