@@ -4,10 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Aaron Anderson, Kevin Buzzard, Yaël Dillies, Eric Wieser
5
5
6
6
! This file was ported from Lean 3 source module order.sup_indep
7
- ! leanprover-community/mathlib commit 1c857a1f6798cb054be942199463c2cf904cb937
7
+ ! leanprover-community/mathlib commit c4c2ed622f43768eff32608d4a0f8a6cec1c047d
8
8
! Please do not edit these lines, except to modify the commit id
9
9
! if you have ported upstream changes.
10
10
-/
11
+ import Mathlib.Data.Finset.Sigma
11
12
import Mathlib.Data.Finset.Pairwise
12
13
import Mathlib.Data.Finset.Powerset
13
14
import Mathlib.Data.Fintype.Basic
@@ -89,13 +90,43 @@ theorem SupIndep.pairwiseDisjoint (hs : s.SupIndep f) : (s : Set ι).PairwiseDis
89
90
sup_singleton.subst <| hs (singleton_subset_iff.2 hb) ha <| not_mem_singleton.2 hab
90
91
#align finset.sup_indep.pairwise_disjoint Finset.SupIndep.pairwiseDisjoint
91
92
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
+
92
100
/-- The RHS looks like the definition of `CompleteLattice.Independent`. -/
93
101
theorem supIndep_iff_disjoint_erase [DecidableEq ι] :
94
102
s.SupIndep f ↔ ∀ i ∈ s, Disjoint (f i) ((s.erase i).sup f) :=
95
103
⟨fun hs _ hi => hs (erase_subset _ _) hi (not_mem_erase _ _), fun hs _ ht i hi hit =>
96
104
(hs i hi).mono_right (sup_mono fun _ hj => mem_erase.2 ⟨ne_of_mem_of_not_mem hj hit, ht hj⟩)⟩
97
105
#align finset.sup_indep_iff_disjoint_erase Finset.supIndep_iff_disjoint_erase
98
106
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
+
99
130
@[simp]
100
131
theorem supIndep_pair [DecidableEq ι] {i j : ι} (hij : i ≠ j) :
101
132
({i, j} : Finset ι).SupIndep f ↔ Disjoint (f i) (f j) :=
@@ -131,16 +162,39 @@ theorem supIndep_univ_fin_two (f : Fin 2 → α) :
131
162
supIndep_pair this
132
163
#align finset.sup_indep_univ_fin_two Finset.supIndep_univ_fin_two
133
164
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
135
166
intro t _ i _ hi
136
167
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]
138
170
refine' hs (image_subset_iff.2 fun (j : { x // x ∈ s }) _ => j.2 ) i.2 fun hi' => hi _
139
171
rw [mem_image] at hi'
140
172
obtain ⟨j, hj, hji⟩ := hi'
141
173
rwa [Subtype.ext hji] at hj
142
174
#align finset.sup_indep.attach Finset.SupIndep.attach
143
175
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
+
144
198
end Lattice
145
199
146
200
section DistribLattice
@@ -173,6 +227,53 @@ theorem SupIndep.biUnion [DecidableEq ι] {s : Finset ι'} {g : ι' → Finset
173
227
exact hs.sup hg
174
228
#align finset.sup_indep.bUnion Finset.SupIndep.biUnion
175
229
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
+
176
277
end DistribLattice
177
278
178
279
end Finset
0 commit comments