62
62
exact ⟨λ k b hb, k _ _ hb rfl, λ k a' b hb h, h ▸ k _ hb⟩,
63
63
end
64
64
65
+ @[simp] lemma sup_bUnion [decidable_eq β] (s : finset γ) (t : γ → finset β) :
66
+ (s.bUnion t).sup f = s.sup (λ x, (t x).sup f) :=
67
+ eq_of_forall_ge_iff $ λ c, by simp [@forall_swap _ β]
68
+
65
69
lemma sup_const {s : finset β} (h : s.nonempty) (c : α) : s.sup (λ _, c) = c :=
66
70
eq_of_forall_ge_iff $ λ b, sup_le_iff.trans h.forall_const
67
71
@@ -213,6 +217,13 @@ lemma inf_union [decidable_eq β] : (s₁ ∪ s₂).inf f = s₁.inf f ⊓ s₂.
213
217
theorem inf_congr {f g : β → α} (hs : s₁ = s₂) (hfg : ∀a∈s₂, f a = g a) : s₁.inf f = s₂.inf g :=
214
218
by subst hs; exact finset.fold_congr hfg
215
219
220
+ @[simp] lemma inf_bUnion [decidable_eq β] (s : finset γ) (t : γ → finset β) :
221
+ (s.bUnion t).inf f = s.inf (λ x, (t x).inf f) :=
222
+ @sup_bUnion (order_dual α) _ _ _ _ _ _ _
223
+
224
+ lemma inf_const {s : finset β} (h : s.nonempty) (c : α) : s.inf (λ _, c) = c :=
225
+ @sup_const (order_dual α) _ _ _ h _
226
+
216
227
lemma le_inf_iff {a : α} : a ≤ s.inf f ↔ ∀ b ∈ s, a ≤ f b :=
217
228
@sup_le_iff (order_dual α) _ _ _ _ _
218
229
@@ -337,6 +348,11 @@ begin
337
348
exact bex_congr (λ b hb, with_bot.coe_lt_coe),
338
349
end
339
350
351
+ lemma sup'_bUnion [decidable_eq β] {s : finset γ} (Hs : s.nonempty) {t : γ → finset β}
352
+ (Ht : ∀ b, (t b).nonempty) :
353
+ (s.bUnion t).sup' (Hs.bUnion (λ b _, Ht b)) f = s.sup' Hs (λ b, (t b).sup' (Ht b) f) :=
354
+ eq_of_forall_ge_iff $ λ c, by simp [@forall_swap _ β]
355
+
340
356
lemma comp_sup'_eq_sup'_comp [semilattice_sup γ] {s : finset β} (H : s.nonempty)
341
357
{f : β → α} (g : α → γ) (g_sup : ∀ x y, g (x ⊔ y) = g x ⊔ g y) :
342
358
g (s.sup' H f) = s.sup' H (g ∘ f) :=
@@ -384,6 +400,14 @@ lemma sup'_mem
384
400
t.sup' H p ∈ s :=
385
401
sup'_induction H p w h
386
402
403
+ @[congr] lemma sup'_congr {t : finset β} {f g : β → α} (h₁ : s = t) (h₂ : ∀ x ∈ s, f x = g x) :
404
+ s.sup' H f = t.sup' (h₁ ▸ H) g :=
405
+ begin
406
+ subst s,
407
+ refine eq_of_forall_ge_iff (λ c, _),
408
+ simp only [sup'_le_iff, h₂] { contextual := tt }
409
+ end
410
+
387
411
end sup'
388
412
389
413
section inf'
@@ -436,6 +460,11 @@ lemma inf'_le {b : β} (h : b ∈ s) : s.inf' ⟨b, h⟩ f ≤ f b :=
436
460
@[simp] lemma inf'_lt_iff [is_total α (≤)] {a : α} : s.inf' H f < a ↔ (∃ b ∈ s, f b < a) :=
437
461
@lt_sup'_iff (order_dual α) _ _ _ H f _ _
438
462
463
+ lemma inf'_bUnion [decidable_eq β] {s : finset γ} (Hs : s.nonempty) {t : γ → finset β}
464
+ (Ht : ∀ b, (t b).nonempty) :
465
+ (s.bUnion t).inf' (Hs.bUnion (λ b _, Ht b)) f = s.inf' Hs (λ b, (t b).inf' (Ht b) f) :=
466
+ @sup'_bUnion (order_dual α) _ _ _ _ _ _ Hs _ Ht
467
+
439
468
lemma comp_inf'_eq_inf'_comp [semilattice_inf γ] {s : finset β} (H : s.nonempty)
440
469
{f : β → α} (g : α → γ) (g_inf : ∀ x y, g (x ⊓ y) = g x ⊓ g y) :
441
470
g (s.inf' H f) = s.inf' H (g ∘ f) :=
@@ -453,6 +482,10 @@ lemma inf'_mem (s : set α) (w : ∀ x y ∈ s, x ⊓ y ∈ s)
453
482
t.inf' H p ∈ s :=
454
483
inf'_induction H p w h
455
484
485
+ @[congr] lemma inf'_congr {t : finset β} {f g : β → α} (h₁ : s = t) (h₂ : ∀ x ∈ s, f x = g x) :
486
+ s.inf' H f = t.inf' (h₁ ▸ H) g :=
487
+ @sup'_congr (order_dual α) _ _ _ H _ _ _ h₁ h₂
488
+
456
489
end inf'
457
490
458
491
section sup
0 commit comments