Documentation

Mathlib.Order.Disjointed

Making a sequence disjoint #

This file defines the way to make a sequence of sets - or, more generally, a map from a partially ordered type ι into a (generalized) Boolean algebra α - into a pairwise disjoint sequence with the same partial sups.

For a sequence f : ℕ → α, this new sequence will be f 0, f 1 \ f 0, f 2 \ (f 0 ⊔ f 1) ⋯. It is actually unique, as disjointed_unique shows.

Main declarations #

We also provide set notation variants of some lemmas.

def disjointed {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [Preorder ι] [LocallyFiniteOrderBot ι] (f : ια) (i : ι) :
α

The function mapping i to f i \ (⨆ j < i, f j). When ι is a partial order, this is the unique function g having the same partialSups as f and such that g i and g j are disjoint whenever i < j.

Equations
Instances For
    theorem disjointed_apply {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [Preorder ι] [LocallyFiniteOrderBot ι] (f : ια) (i : ι) :
    disjointed f i = f i \ (Finset.Iio i).sup f
    theorem disjointed_of_isMin {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [Preorder ι] [LocallyFiniteOrderBot ι] (f : ια) {i : ι} (hn : IsMin i) :
    disjointed f i = f i
    @[simp]
    theorem disjointed_bot {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [Preorder ι] [LocallyFiniteOrderBot ι] [OrderBot ι] (f : ια) :
    theorem disjointed_le {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [Preorder ι] [LocallyFiniteOrderBot ι] (f : ια) :
    theorem disjoint_disjointed_of_lt {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [Preorder ι] [LocallyFiniteOrderBot ι] (f : ια) {i j : ι} (h : i < j) :
    theorem disjointed_eq_self {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [Preorder ι] [LocallyFiniteOrderBot ι] {f : ια} {i : ι} (hf : j < i, Disjoint (f j) (f i)) :
    disjointed f i = f i
    theorem disjointedRec {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [Preorder ι] [LocallyFiniteOrderBot ι] {f : ια} {p : αProp} (hdiff : ∀ ⦃t : α⦄ ⦃i : ι⦄, p tp (t \ f i)) ⦃i : ι :
    p (f i)p (disjointed f i)

    An induction principle for disjointed. To prove something about disjointed f i, it's enough to prove it for f i and being able to extend through diffs.

    @[simp]
    theorem Fintype.sup_disjointed {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [PartialOrder ι] [LocallyFiniteOrderBot ι] [Fintype ι] (f : ια) :
    theorem disjointed_unique {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [PartialOrder ι] [LocallyFiniteOrderBot ι] {f d : ια} (hdisj : ∀ {i j : ι}, i < jDisjoint (d i) (d j)) (hsups : partialSups d = partialSups f) :

    disjointed f is the unique map d : ι → α such that d has the same partial sups as f, and d i and d j are disjoint whenever i < j.

    Linear orders #

    theorem disjointed_unique' {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [LinearOrder ι] [LocallyFiniteOrderBot ι] {f d : ια} (hdisj : Pairwise (Function.onFun Disjoint d)) (hsups : partialSups d = partialSups f) :

    disjointed f is the unique sequence that is pairwise disjoint and has the same partial sups as f.

    theorem disjointed_succ {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [LinearOrder ι] [LocallyFiniteOrderBot ι] [SuccOrder ι] (f : ια) {i : ι} (hi : ¬IsMax i) :
    theorem Monotone.disjointed_succ {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [LinearOrder ι] [LocallyFiniteOrderBot ι] [SuccOrder ι] {f : ια} (hf : Monotone f) {i : ι} (hn : ¬IsMax i) :
    theorem Monotone.disjointed_succ_sup {α : Type u_1} {ι : Type u_2} [GeneralizedBooleanAlgebra α] [LinearOrder ι] [LocallyFiniteOrderBot ι] [SuccOrder ι] {f : ια} (hf : Monotone f) (i : ι) :

    Note this lemma does not require ¬IsMax i, unlike disjointed_succ.

    Functions on an arbitrary fintype #

    theorem Fintype.exists_disjointed_le {α : Type u_1} [GeneralizedBooleanAlgebra α] {ι : Type u_3} [Fintype ι] {f : ια} :

    For any finite family of elements f : ι → α, we can find a pairwise-disjoint family g bounded above by f and having the same supremum. This is non-canonical, depending on an arbitrary choice of ordering of ι.

    Complete Boolean algebras #

    theorem iSup_disjointed {α : Type u_1} {ι : Type u_2} [CompleteBooleanAlgebra α] [PartialOrder ι] [LocallyFiniteOrderBot ι] (f : ια) :
    ⨆ (i : ι), disjointed f i = ⨆ (i : ι), f i
    theorem disjointed_eq_inf_compl {α : Type u_1} {ι : Type u_2} [CompleteBooleanAlgebra α] [Preorder ι] [LocallyFiniteOrderBot ι] (f : ια) (i : ι) :
    disjointed f i = f i ⨅ (j : ι), ⨅ (_ : j < i), (f j)

    Lemmas specific to set-valued functions #

    theorem disjointed_subset {α : Type u_1} {ι : Type u_2} [Preorder ι] [LocallyFiniteOrderBot ι] (f : ιSet α) (i : ι) :
    disjointed f i f i
    theorem iUnion_disjointed {α : Type u_1} {ι : Type u_2} [PartialOrder ι] [LocallyFiniteOrderBot ι] {f : ιSet α} :
    ⋃ (i : ι), disjointed f i = ⋃ (i : ι), f i
    theorem disjointed_eq_inter_compl {α : Type u_1} {ι : Type u_2} [Preorder ι] [LocallyFiniteOrderBot ι] (f : ιSet α) (i : ι) :
    disjointed f i = f i ⋂ (j : ι), ⋂ (_ : j < i), (f j)
    theorem preimage_find_eq_disjointed {α : Type u_1} (s : Set α) (H : ∀ (x : α), ∃ (n : ), x s n) [(x : α) → (n : ) → Decidable (x s n)] (n : ) :
    (fun (x : α) => Nat.find ) ⁻¹' {n} = disjointed s n

    Functions on #

    (See also Mathlib.Algebra.Order.Disjointed for results with more algebra pre-requsisites.)

    @[simp]
    theorem disjointed_zero {α : Type u_1} [GeneralizedBooleanAlgebra α] (f : α) :
    disjointed f 0 = f 0