Documentation

Mathlib.Data.Set.Pairwise.Chain

Pairwise results for chains #

In this file Pairwise results are applied to chains of sets.

theorem IsChain.pairwise_iUnion₂ {α : Type u_1} {c : Set (Set α)} {r : ααProp} (hc : IsChain (fun (x1 x2 : Set α) => x1 x2) c) :
(⋃ sc, s).Pairwise r sc, s.Pairwise r
theorem IsChain.pairwiseDisjoint_iUnion₂ {α : Type u_1} {β : Type u_2} {c : Set (Set α)} (hc : IsChain (fun (x1 x2 : Set α) => x1 x2) c) [PartialOrder β] [OrderBot β] (f : αβ) :
(⋃ sc, s).PairwiseDisjoint f sc, s.PairwiseDisjoint f
theorem IsChain.pairwise_sUnion {α : Type u_1} {c : Set (Set α)} {r : ααProp} (hc : IsChain (fun (x1 x2 : Set α) => x1 x2) c) :
(⋃₀ c).Pairwise r sc, s.Pairwise r
theorem IsChain.pairwiseDisjoint_sUnion {α : Type u_1} {β : Type u_2} {c : Set (Set α)} (hc : IsChain (fun (x1 x2 : Set α) => x1 x2) c) [PartialOrder β] [OrderBot β] (f : αβ) :