Skip to content

Commit e266a22

Browse files
YaelDilliesParcly-Taxeleric-wieser
committed
feat: More sup_indep lemmas (#5196)
Match leanprover-community/mathlib3#11932 Co-authored-by: Jeremy Tan Jie Rui <[email protected]> Co-authored-by: Parcly Taxel <[email protected]> Co-authored-by: Eric Wieser <[email protected]>
1 parent 5f2c0b7 commit e266a22

File tree

5 files changed

+211
-17
lines changed

5 files changed

+211
-17
lines changed

Mathlib/Data/Finset/Pairwise.lean

+15-4
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Yaël Dillies
55
66
! This file was ported from Lean 3 source module data.finset.pairwise
7-
! leanprover-community/mathlib commit 9003f28797c0664a49e4179487267c494477d853
7+
! leanprover-community/mathlib commit c4c2ed622f43768eff32608d4a0f8a6cec1c047d
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -40,13 +40,24 @@ theorem PairwiseDisjoint.elim_finset {s : Set ι} {f : ι → Finset α} (hs : s
4040
hs.elim hi hj (Finset.not_disjoint_iff.2 ⟨a, hai, haj⟩)
4141
#align set.pairwise_disjoint.elim_finset Set.PairwiseDisjoint.elim_finset
4242

43-
theorem PairwiseDisjoint.image_finset_of_le [DecidableEq ι] [SemilatticeInf α] [OrderBot α]
44-
{s : Finset ι} {f : ι → α} (hs : (s : Set ι).PairwiseDisjoint f) {g : ι → ι}
45-
(hf : ∀ a, f (g a) ≤ f a) : (s.image g : Set ι).PairwiseDisjoint f := by
43+
section SemilatticeInf
44+
45+
variable [SemilatticeInf α] [OrderBot α] {s : Finset ι} {f : ι → α}
46+
47+
theorem PairwiseDisjoint.image_finset_of_le [DecidableEq ι] {s : Finset ι} {f : ι → α}
48+
(hs : (s : Set ι).PairwiseDisjoint f) {g : ι → ι} (hf : ∀ a, f (g a) ≤ f a) :
49+
(s.image g : Set ι).PairwiseDisjoint f := by
4650
rw [coe_image]
4751
exact hs.image_of_le hf
4852
#align set.pairwise_disjoint.image_finset_of_le Set.PairwiseDisjoint.image_finset_of_le
4953

54+
theorem PairwiseDisjoint.attach (hs : (s : Set ι).PairwiseDisjoint f) :
55+
(s.attach : Set { x // x ∈ s }).PairwiseDisjoint (f ∘ Subtype.val) := fun i _ j _ hij =>
56+
hs i.2 j.2 <| mt Subtype.ext_val hij
57+
#align set.pairwise_disjoint.attach Set.PairwiseDisjoint.attach
58+
59+
end SemilatticeInf
60+
5061
variable [Lattice α] [OrderBot α]
5162

5263
/-- Bind operation for `Set.PairwiseDisjoint`. In a complete lattice, you can use

Mathlib/Data/Set/Pairwise/Basic.lean

+22-2
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johannes Hölzl
55
66
! This file was ported from Lean 3 source module data.set.pairwise.basic
7-
! leanprover-community/mathlib commit c227d107bbada5d0d9d20287e3282c0a7f1651a0
7+
! leanprover-community/mathlib commit c4c2ed622f43768eff32608d4a0f8a6cec1c047d
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -33,7 +33,7 @@ on `Set.PairwiseDisjoint`, even though the latter unfolds to something nicer.
3333
-/
3434

3535

36-
open Set Function
36+
open Function Order Set
3737

3838
variable {α β γ ι ι' : Type _} {r p q : α → α → Prop}
3939

@@ -345,6 +345,8 @@ end SemilatticeInfBot
345345

346346
/-! ### Pairwise disjoint set of sets -/
347347

348+
variable {s : Set ι} {t : Set ι'}
349+
348350
theorem pairwiseDisjoint_range_singleton :
349351
(range (singleton : ι → Set ι)).PairwiseDisjoint id := by
350352
rintro _ ⟨a, rfl⟩ _ ⟨b, rfl⟩ h
@@ -361,6 +363,24 @@ theorem PairwiseDisjoint.elim_set {s : Set ι} {f : ι → Set α} (hs : s.Pairw
361363
hs.elim hi hj <| not_disjoint_iff.2 ⟨a, hai, haj⟩
362364
#align set.pairwise_disjoint.elim_set Set.PairwiseDisjoint.elim_set
363365

366+
theorem PairwiseDisjoint.prod {f : ι → Set α} {g : ι' → Set β} (hs : s.PairwiseDisjoint f)
367+
(ht : t.PairwiseDisjoint g) :
368+
(s ×ˢ t : Set (ι × ι')).PairwiseDisjoint fun i => f i.1 ×ˢ g i.2 :=
369+
fun ⟨_, _⟩ ⟨hi, hi'⟩ ⟨_, _⟩ ⟨hj, hj'⟩ hij =>
370+
disjoint_left.2 fun ⟨_, _⟩ ⟨hai, hbi⟩ ⟨haj, hbj⟩ =>
371+
hij <| Prod.ext (hs.elim_set hi hj _ hai haj) <| ht.elim_set hi' hj' _ hbi hbj
372+
#align set.pairwise_disjoint.prod Set.PairwiseDisjoint.prod
373+
374+
theorem pairwiseDisjoint_pi {ι' α : ι → Type _} {s : ∀ i, Set (ι' i)} {f : ∀ i, ι' i → Set (α i)}
375+
(hs : ∀ i, (s i).PairwiseDisjoint (f i)) :
376+
((univ : Set ι).pi s).PairwiseDisjoint fun I => (univ : Set ι).pi fun i => f _ (I i) :=
377+
fun _ hI _ hJ hIJ =>
378+
disjoint_left.2 fun a haI haJ =>
379+
hIJ <|
380+
funext fun i =>
381+
(hs i).elim_set (hI i trivial) (hJ i trivial) (a i) (haI i trivial) (haJ i trivial)
382+
#align set.pairwise_disjoint_pi Set.pairwiseDisjoint_pi
383+
364384
/-- The partial images of a binary function `f` whose partial evaluations are injective are pairwise
365385
disjoint iff `f` is injective . -/
366386
theorem pairwiseDisjoint_image_right_iff {f : α → β → γ} {s : Set α} {t : Set β}

Mathlib/Data/Set/Pairwise/Lattice.lean

+40-7
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johannes Hölzl
55
66
! This file was ported from Lean 3 source module data.set.pairwise.lattice
7-
! leanprover-community/mathlib commit c227d107bbada5d0d9d20287e3282c0a7f1651a0
7+
! leanprover-community/mathlib commit c4c2ed622f43768eff32608d4a0f8a6cec1c047d
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -18,17 +18,16 @@ In this file we prove many facts about `Pairwise` and the set lattice.
1818
-/
1919

2020

21-
open Set Function
22-
23-
variable {α β γ ι ι' : Type _} {r p q : α → α → Prop}
21+
open Function Set Order
2422

23+
variable {α β γ ι ι' : Type _} {κ : Sort _} {r p q : α → α → Prop}
2524
section Pairwise
2625

2726
variable {f g : ι → α} {s t u : Set α} {a b : α}
2827

2928
namespace Set
3029

31-
theorem pairwise_iUnion {f : ι → Set α} (h : Directed (· ⊆ ·) f) :
30+
theorem pairwise_iUnion {f : κ → Set α} (h : Directed (· ⊆ ·) f) :
3231
(⋃ n, f n).Pairwise r ↔ ∀ n, (f n).Pairwise r := by
3332
constructor
3433
· intro H n
@@ -69,8 +68,7 @@ end PartialOrderBot
6968

7069
section CompleteLattice
7170

72-
variable [CompleteLattice α]
73-
71+
variable [CompleteLattice α] {s : Set ι} {t : Set ι'}
7472

7573
/-- Bind operation for `Set.PairwiseDisjoint`. If you want to only consider finsets of indices, you
7674
can use `Set.PairwiseDisjoint.biUnion_finset`. -/
@@ -89,8 +87,43 @@ theorem PairwiseDisjoint.biUnion {s : Set ι'} {g : ι' → Set ι} {f : ι →
8987
(le_iSup₂ (f := fun i (_ : i ∈ g d) => f i) b hb)
9088
#align set.pairwise_disjoint.bUnion Set.PairwiseDisjoint.biUnion
9189

90+
/-- If the suprema of columns are pairwise disjoint and suprema of rows as well, then everything is
91+
pairwise disjoint. Not to be confused with `Set.PairwiseDisjoint.prod`. -/
92+
theorem PairwiseDisjoint.prod_left {f : ι × ι' → α}
93+
(hs : s.PairwiseDisjoint fun i => ⨆ i' ∈ t, f (i, i'))
94+
(ht : t.PairwiseDisjoint fun i' => ⨆ i ∈ s, f (i, i')) :
95+
(s ×ˢ t : Set (ι × ι')).PairwiseDisjoint f := by
96+
rintro ⟨i, i'⟩ hi ⟨j, j'⟩ hj h
97+
rw [mem_prod] at hi hj
98+
obtain rfl | hij := eq_or_ne i j
99+
· refine' (ht hi.2 hj.2 <| (Prod.mk.inj_left _).ne_iff.1 h).mono _ _
100+
· convert le_iSup₂ (α := α) i hi.1; rfl
101+
· convert le_iSup₂ (α := α) i hj.1; rfl
102+
· refine' (hs hi.1 hj.1 hij).mono _ _
103+
· convert le_iSup₂ (α := α) i' hi.2; rfl
104+
· convert le_iSup₂ (α := α) j' hj.2; rfl
105+
#align set.pairwise_disjoint.prod_left Set.PairwiseDisjoint.prod_left
106+
92107
end CompleteLattice
93108

109+
section Frame
110+
111+
variable [Frame α]
112+
113+
theorem pairwiseDisjoint_prod_left {s : Set ι} {t : Set ι'} {f : ι × ι' → α} :
114+
(s ×ˢ t : Set (ι × ι')).PairwiseDisjoint f ↔
115+
(s.PairwiseDisjoint fun i => ⨆ i' ∈ t, f (i, i')) ∧
116+
t.PairwiseDisjoint fun i' => ⨆ i ∈ s, f (i, i') := by
117+
refine'
118+
fun h => ⟨fun i hi j hj hij => _, fun i hi j hj hij => _⟩, fun h => h.1.prod_left h.2⟩ <;>
119+
simp_rw [Function.onFun, iSup_disjoint_iff, disjoint_iSup_iff] <;>
120+
intro i' hi' j' hj'
121+
· exact h (mk_mem_prod hi hi') (mk_mem_prod hj hj') (ne_of_apply_ne Prod.fst hij)
122+
· exact h (mk_mem_prod hi' hi) (mk_mem_prod hj' hj) (ne_of_apply_ne Prod.snd hij)
123+
#align set.pairwise_disjoint_prod_left Set.pairwiseDisjoint_prod_left
124+
125+
end Frame
126+
94127
theorem biUnion_diff_biUnion_eq {s t : Set ι} {f : ι → Set α} (h : (s ∪ t).PairwiseDisjoint f) :
95128
((⋃ i ∈ s, f i) \ ⋃ i ∈ t, f i) = ⋃ i ∈ s \ t, f i := by
96129
refine'

Mathlib/Data/Set/Prod.lean

+30-1
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Mario Carneiro, Johannes Hölzl, Patrick Massot
55
66
! This file was ported from Lean 3 source module data.set.prod
7-
! leanprover-community/mathlib commit 27f315c5591c84687852f816d8ef31fe103d03de
7+
! leanprover-community/mathlib commit c4c2ed622f43768eff32608d4a0f8a6cec1c047d
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -173,11 +173,22 @@ theorem prod_inter_prod : s₁ ×ˢ t₁ ∩ s₂ ×ˢ t₂ = (s₁ ∩ s₂) ×
173173
simp [and_assoc, and_left_comm]
174174
#align set.prod_inter_prod Set.prod_inter_prod
175175

176+
@[simp]
176177
theorem disjoint_prod : Disjoint (s₁ ×ˢ t₁) (s₂ ×ˢ t₂) ↔ Disjoint s₁ s₂ ∨ Disjoint t₁ t₂ := by
177178
simp_rw [disjoint_left, mem_prod, not_and_or, Prod.forall, and_imp, ← @forall_or_right α, ←
178179
@forall_or_left β, ← @forall_or_right (_ ∈ s₁), ← @forall_or_left (_ ∈ t₁)]
179180
#align set.disjoint_prod Set.disjoint_prod
180181

182+
theorem Disjoint.set_prod_left (hs : Disjoint s₁ s₂) (t₁ t₂ : Set β) :
183+
Disjoint (s₁ ×ˢ t₁) (s₂ ×ˢ t₂) :=
184+
disjoint_left.2 fun ⟨_a, _b⟩ ⟨ha₁, _⟩ ⟨ha₂, _⟩ => disjoint_left.1 hs ha₁ ha₂
185+
#align set.disjoint.set_prod_left Set.Disjoint.set_prod_left
186+
187+
theorem Disjoint.set_prod_right (ht : Disjoint t₁ t₂) (s₁ s₂ : Set α) :
188+
Disjoint (s₁ ×ˢ t₁) (s₂ ×ˢ t₂) :=
189+
disjoint_left.2 fun ⟨_a, _b⟩ ⟨_, hb₁⟩ ⟨_, hb₂⟩ => disjoint_left.1 ht hb₁ hb₂
190+
#align set.disjoint.set_prod_right Set.Disjoint.set_prod_right
191+
181192
theorem insert_prod : insert a s ×ˢ t = Prod.mk a '' t ∪ s ×ˢ t := by
182193
ext ⟨x, y⟩
183194
simp (config := { contextual := true }) [image, iff_def, or_imp, Imp.swap]
@@ -710,6 +721,24 @@ theorem disjoint_univ_pi : Disjoint (pi univ t₁) (pi univ t₂) ↔ ∃ i, Dis
710721
simp only [disjoint_iff_inter_eq_empty, ← pi_inter_distrib, univ_pi_eq_empty_iff]
711722
#align set.disjoint_univ_pi Set.disjoint_univ_pi
712723

724+
theorem Disjoint.set_pi (hi : i ∈ s) (ht : Disjoint (t₁ i) (t₂ i)) : Disjoint (s.pi t₁) (s.pi t₂) :=
725+
disjoint_left.2 fun _ h₁ h₂ => disjoint_left.1 ht (h₁ _ hi) (h₂ _ hi)
726+
#align set.disjoint.set_pi Set.Disjoint.set_pi
727+
728+
section Nonempty
729+
730+
variable [∀ i, Nonempty (α i)]
731+
732+
theorem pi_eq_empty_iff' : s.pi t = ∅ ↔ ∃ i ∈ s, t i = ∅ := by simp [pi_eq_empty_iff]
733+
#align set.pi_eq_empty_iff' Set.pi_eq_empty_iff'
734+
735+
@[simp]
736+
theorem disjoint_pi : Disjoint (s.pi t₁) (s.pi t₂) ↔ ∃ i ∈ s, Disjoint (t₁ i) (t₂ i) := by
737+
simp only [disjoint_iff_inter_eq_empty, ← pi_inter_distrib, pi_eq_empty_iff']
738+
#align set.disjoint_pi Set.disjoint_pi
739+
740+
end Nonempty
741+
713742
-- Porting note: Removing `simp` - LHS does not simplify
714743
theorem range_dcomp (f : ∀ i, α i → β i) :
715744
(range fun g : ∀ i, α i => fun i => f i (g i)) = pi univ fun i => range (f i) := by

Mathlib/Order/SupIndep.lean

+104-3
Original file line numberDiff line numberDiff line change
@@ -4,10 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Aaron Anderson, Kevin Buzzard, Yaël Dillies, Eric Wieser
55
66
! This file was ported from Lean 3 source module order.sup_indep
7-
! leanprover-community/mathlib commit 1c857a1f6798cb054be942199463c2cf904cb937
7+
! leanprover-community/mathlib commit c4c2ed622f43768eff32608d4a0f8a6cec1c047d
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
11+
import Mathlib.Data.Finset.Sigma
1112
import Mathlib.Data.Finset.Pairwise
1213
import Mathlib.Data.Finset.Powerset
1314
import Mathlib.Data.Fintype.Basic
@@ -89,13 +90,43 @@ theorem SupIndep.pairwiseDisjoint (hs : s.SupIndep f) : (s : Set ι).PairwiseDis
8990
sup_singleton.subst <| hs (singleton_subset_iff.2 hb) ha <| not_mem_singleton.2 hab
9091
#align finset.sup_indep.pairwise_disjoint Finset.SupIndep.pairwiseDisjoint
9192

93+
theorem SupIndep.le_sup_iff (hs : s.SupIndep f) (hts : t ⊆ s) (hi : i ∈ s) (hf : ∀ i, f i ≠ ⊥) :
94+
f i ≤ t.sup f ↔ i ∈ t := by
95+
refine' ⟨fun h => _, le_sup⟩
96+
by_contra hit
97+
exact hf i (disjoint_self.1 <| (hs hts hi hit).mono_right h)
98+
#align finset.sup_indep.le_sup_iff Finset.SupIndep.le_sup_iff
99+
92100
/-- The RHS looks like the definition of `CompleteLattice.Independent`. -/
93101
theorem supIndep_iff_disjoint_erase [DecidableEq ι] :
94102
s.SupIndep f ↔ ∀ i ∈ s, Disjoint (f i) ((s.erase i).sup f) :=
95103
fun hs _ hi => hs (erase_subset _ _) hi (not_mem_erase _ _), fun hs _ ht i hi hit =>
96104
(hs i hi).mono_right (sup_mono fun _ hj => mem_erase.2 ⟨ne_of_mem_of_not_mem hj hit, ht hj⟩)⟩
97105
#align finset.sup_indep_iff_disjoint_erase Finset.supIndep_iff_disjoint_erase
98106

107+
theorem SupIndep.image [DecidableEq ι] {s : Finset ι'} {g : ι' → ι} (hs : s.SupIndep (f ∘ g)) :
108+
(s.image g).SupIndep f := by
109+
intro t ht i hi hit
110+
rw [mem_image] at hi
111+
obtain ⟨i, hi, rfl⟩ := hi
112+
haveI : DecidableEq ι' := Classical.decEq _
113+
suffices hts : t ⊆ (s.erase i).image g
114+
· refine' (supIndep_iff_disjoint_erase.1 hs i hi).mono_right ((sup_mono hts).trans _)
115+
rw [sup_image]
116+
rintro j hjt
117+
obtain ⟨j, hj, rfl⟩ := mem_image.1 (ht hjt)
118+
exact mem_image_of_mem _ (mem_erase.2 ⟨ne_of_apply_ne g (ne_of_mem_of_not_mem hjt hit), hj⟩)
119+
#align finset.sup_indep.image Finset.SupIndep.image
120+
121+
theorem supIndep_map {s : Finset ι'} {g : ι' ↪ ι} : (s.map g).SupIndep f ↔ s.SupIndep (f ∘ g) := by
122+
refine' ⟨fun hs t ht i hi hit => _, fun hs => _⟩
123+
· rw [← sup_map]
124+
exact hs (map_subset_map.2 ht) ((mem_map' _).2 hi) (by rwa [mem_map'])
125+
· classical
126+
rw [map_eq_image]
127+
exact hs.image
128+
#align finset.sup_indep_map Finset.supIndep_map
129+
99130
@[simp]
100131
theorem supIndep_pair [DecidableEq ι] {i j : ι} (hij : i ≠ j) :
101132
({i, j} : Finset ι).SupIndep f ↔ Disjoint (f i) (f j) :=
@@ -131,16 +162,39 @@ theorem supIndep_univ_fin_two (f : Fin 2 → α) :
131162
supIndep_pair this
132163
#align finset.sup_indep_univ_fin_two Finset.supIndep_univ_fin_two
133164

134-
theorem SupIndep.attach (hs : s.SupIndep f) : s.attach.SupIndep (f ∘ Subtype.val) := by
165+
theorem SupIndep.attach (hs : s.SupIndep f) : s.attach.SupIndep fun a => f a := by
135166
intro t _ i _ hi
136167
classical
137-
rw [← Finset.sup_image]
168+
have : (fun (a : { x // x ∈ s }) => f ↑a) = f ∘ (fun a : { x // x ∈ s } => ↑a) := rfl
169+
rw [this, ← Finset.sup_image]
138170
refine' hs (image_subset_iff.2 fun (j : { x // x ∈ s }) _ => j.2) i.2 fun hi' => hi _
139171
rw [mem_image] at hi'
140172
obtain ⟨j, hj, hji⟩ := hi'
141173
rwa [Subtype.ext hji] at hj
142174
#align finset.sup_indep.attach Finset.SupIndep.attach
143175

176+
/-
177+
Porting note: simpNF linter returns
178+
179+
"Left-hand side does not simplify, when using the simp lemma on itself."
180+
181+
However, simp does indeed solve the following. leanprover/std4#71 is related.
182+
183+
example {α ι} [Lattice α] [OrderBot α] (s : Finset ι) (f : ι → α) :
184+
(s.attach.SupIndep fun a => f a) ↔ s.SupIndep f := by simp
185+
-/
186+
@[simp, nolint simpNF]
187+
theorem supIndep_attach : (s.attach.SupIndep fun a => f a) ↔ s.SupIndep f := by
188+
refine' ⟨fun h t ht i his hit => _, SupIndep.attach⟩
189+
classical
190+
convert h (filter_subset (fun (i : { x // x ∈ s }) => (i : ι) ∈ t) _) (mem_attach _ ⟨i, ‹_›⟩)
191+
fun hi => hit <| by simpa using hi using 1
192+
refine' eq_of_forall_ge_iff _
193+
simp only [Finset.sup_le_iff, mem_filter, mem_attach, true_and_iff, Function.comp_apply,
194+
Subtype.forall, Subtype.coe_mk]
195+
exact fun a => forall_congr' fun j => ⟨fun h _ => h, fun h hj => h (ht hj) hj⟩
196+
#align finset.sup_indep_attach Finset.supIndep_attach
197+
144198
end Lattice
145199

146200
section DistribLattice
@@ -173,6 +227,53 @@ theorem SupIndep.biUnion [DecidableEq ι] {s : Finset ι'} {g : ι' → Finset
173227
exact hs.sup hg
174228
#align finset.sup_indep.bUnion Finset.SupIndep.biUnion
175229

230+
/-- Bind operation for `sup_indep`. -/
231+
theorem SupIndep.sigma {β : ι → Type _} {s : Finset ι} {g : ∀ i, Finset (β i)} {f : Sigma β → α}
232+
(hs : s.SupIndep fun i => (g i).sup fun b => f ⟨i, b⟩)
233+
(hg : ∀ i ∈ s, (g i).SupIndep fun b => f ⟨i, b⟩) : (s.sigma g).SupIndep f := by
234+
rintro t ht ⟨i, b⟩ hi hit
235+
rw [Finset.disjoint_sup_right]
236+
rintro ⟨j, c⟩ hj
237+
have hbc := (ne_of_mem_of_not_mem hj hit).symm
238+
replace hj := ht hj
239+
rw [mem_sigma] at hi hj
240+
obtain rfl | hij := eq_or_ne i j
241+
· exact (hg _ hj.1).pairwiseDisjoint hi.2 hj.2 (sigma_mk_injective.ne_iff.1 hbc)
242+
· refine' (hs.pairwiseDisjoint hi.1 hj.1 hij).mono _ _
243+
· convert le_sup (α := α) hi.2; simp
244+
· convert le_sup (α := α) hj.2; simp
245+
#align finset.sup_indep.sigma Finset.SupIndep.sigma
246+
247+
theorem SupIndep.product {s : Finset ι} {t : Finset ι'} {f : ι × ι' → α}
248+
(hs : s.SupIndep fun i => t.sup fun i' => f (i, i'))
249+
(ht : t.SupIndep fun i' => s.sup fun i => f (i, i')) : (s ×ˢ t).SupIndep f := by
250+
rintro u hu ⟨i, i'⟩ hi hiu
251+
rw [Finset.disjoint_sup_right]
252+
rintro ⟨j, j'⟩ hj
253+
have hij := (ne_of_mem_of_not_mem hj hiu).symm
254+
replace hj := hu hj
255+
rw [mem_product] at hi hj
256+
obtain rfl | hij := eq_or_ne i j
257+
· refine' (ht.pairwiseDisjoint hi.2 hj.2 <| (Prod.mk.inj_left _).ne_iff.1 hij).mono _ _
258+
· convert le_sup (α := α) hi.1; simp
259+
· convert le_sup (α := α) hj.1; simp
260+
· refine' (hs.pairwiseDisjoint hi.1 hj.1 hij).mono _ _
261+
· convert le_sup (α := α) hi.2; simp
262+
· convert le_sup (α := α) hj.2; simp
263+
#align finset.sup_indep.product Finset.SupIndep.product
264+
265+
theorem supIndep_product_iff {s : Finset ι} {t : Finset ι'} {f : ι × ι' → α} :
266+
(s.product t).SupIndep f ↔ (s.SupIndep fun i => t.sup fun i' => f (i, i'))
267+
∧ t.SupIndep fun i' => s.sup fun i => f (i, i') := by
268+
refine' ⟨_, fun h => h.1.product h.2
269+
simp_rw [supIndep_iff_pairwiseDisjoint]
270+
refine' fun h => ⟨fun i hi j hj hij => _, fun i hi j hj hij => _⟩ <;>
271+
simp_rw [Function.onFun, Finset.disjoint_sup_left, Finset.disjoint_sup_right] <;>
272+
intro i' hi' j' hj'
273+
· exact h (mk_mem_product hi hi') (mk_mem_product hj hj') (ne_of_apply_ne Prod.fst hij)
274+
· exact h (mk_mem_product hi' hi) (mk_mem_product hj' hj) (ne_of_apply_ne Prod.snd hij)
275+
#align finset.sup_indep_product_iff Finset.supIndep_product_iff
276+
176277
end DistribLattice
177278

178279
end Finset

0 commit comments

Comments
 (0)