Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 271c323

Browse files
committed
feat(order/filter): prod_assoc (#12002)
map (prod_assoc α β γ) ((f ×ᶠ g) ×ᶠ h) = f ×ᶠ (g ×ᶠ h) with two tiny supporting lemmas
1 parent d8d2f54 commit 271c323

File tree

2 files changed

+25
-0
lines changed

2 files changed

+25
-0
lines changed

src/order/filter/bases.lean

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -338,6 +338,13 @@ begin
338338
exact forall_congr (λ s, ⟨λ h, h.1, λ h, ⟨h, λ ⟨t, hl, hP, hts⟩, mem_of_superset hl hts⟩⟩)
339339
end
340340

341+
lemma has_basis.comp_of_surjective (h : l.has_basis p s) {g : ι' → ι} (hg : function.surjective g) :
342+
l.has_basis (p ∘ g) (s ∘ g) :=
343+
⟨λ t, h.mem_iff.trans hg.exists⟩
344+
345+
lemma has_basis.comp_equiv (h : l.has_basis p s) (e : ι' ≃ ι) : l.has_basis (p ∘ e) (s ∘ e) :=
346+
h.comp_of_surjective e.surjective
347+
341348
/-- If `{s i | p i}` is a basis of a filter `l` and each `s i` includes `s j` such that
342349
`p j ∧ q j`, then `{s j | p j ∧ q j}` is a basis of `l`. -/
343350
lemma has_basis.restrict (h : l.has_basis p s) {q : ι → Prop}
@@ -688,6 +695,16 @@ end
688695

689696
end two_types
690697

698+
open equiv
699+
700+
lemma prod_assoc (f : filter α) (g : filter β) (h : filter γ) :
701+
map (prod_assoc α β γ) ((f ×ᶠ g) ×ᶠ h) = f ×ᶠ (g ×ᶠ h) :=
702+
begin
703+
apply ((((basis_sets f).prod $ basis_sets g).prod $ basis_sets h).map _).eq_of_same_basis,
704+
simpa only [prod_assoc_image, function.comp, and_assoc] using
705+
((basis_sets f).prod $ (basis_sets g).prod $ basis_sets h).comp_equiv (prod_assoc _ _ _)
706+
end
707+
691708
end filter
692709

693710
end sort

src/order/filter/basic.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2000,6 +2000,14 @@ le_antisymm
20002000
(λ b (hb : preimage m b ∈ f),
20012001
⟨preimage m b, hb, show preimage (m ∘ n) b ⊆ b, by simp only [h₁]; apply subset.refl⟩)
20022002

2003+
lemma map_equiv_symm (e : α ≃ β) (f : filter β) :
2004+
map e.symm f = comap e f :=
2005+
map_eq_comap_of_inverse e.symm_comp_self e.self_comp_symm
2006+
2007+
lemma comap_equiv_symm (e : α ≃ β) (f : filter α) :
2008+
comap e.symm f = map e f :=
2009+
(map_eq_comap_of_inverse e.self_comp_symm e.symm_comp_self).symm
2010+
20032011
lemma map_swap_eq_comap_swap {f : filter (α × β)} : prod.swap <$> f = comap prod.swap f :=
20042012
map_eq_comap_of_inverse prod.swap_swap_eq prod.swap_swap_eq
20052013

0 commit comments

Comments
 (0)