|
| 1 | +/- |
| 2 | +Copyright (c) 2021 Oliver Nash. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Oliver Nash |
| 5 | +-/ |
| 6 | +import order.well_founded |
| 7 | +import order.order_iso_nat |
| 8 | +import data.set.finite |
| 9 | +import tactic.tfae |
| 10 | + |
| 11 | +/-! |
| 12 | +# Compactness properties for complete lattices |
| 13 | +
|
| 14 | +For complete lattices, there are numerous equivalent ways to express the fact that the relation `>` |
| 15 | +is well-founded. In this file we define three especially-useful characterisations and provide |
| 16 | +proofs that they are indeed equivalent to well-foundedness. |
| 17 | +
|
| 18 | +## Main definitions |
| 19 | + * `complete_lattice.is_sup_closed_compact` |
| 20 | + * `complete_lattice.is_Sup_finite_compact` |
| 21 | + * `complete_lattice.is_compact_element` |
| 22 | + * `complete_lattice.is_compactly_generated` |
| 23 | +
|
| 24 | +## Main results |
| 25 | +The main result is that the following four conditions are equivalent for a complete lattice: |
| 26 | + * `well_founded (>)` |
| 27 | + * `complete_lattice.is_sup_closed_compact` |
| 28 | + * `complete_lattice.is_Sup_finite_compact` |
| 29 | + * `∀ k, complete_lattice.is_compact_element k` |
| 30 | +
|
| 31 | +This is demonstrated by means of the following four lemmas: |
| 32 | + * `complete_lattice.well_founded.is_Sup_finite_compact` |
| 33 | + * `complete_lattice.is_Sup_finite_compact.is_sup_closed_compact` |
| 34 | + * `complete_lattice.is_sup_closed_compact.well_founded` |
| 35 | + * `complete_lattice.is_Sup_finite_compact_iff_all_elements_compact` |
| 36 | +
|
| 37 | + We also show well-founded lattices are compactly generated |
| 38 | + (`complete_lattice.compactly_generated_of_well_founded`). |
| 39 | +
|
| 40 | +## Tags |
| 41 | +
|
| 42 | +complete lattice, well-founded, compact |
| 43 | +-/ |
| 44 | + |
| 45 | +namespace complete_lattice |
| 46 | + |
| 47 | +variables (α : Type*) [complete_lattice α] |
| 48 | + |
| 49 | +/-- A compactness property for a complete lattice is that any `sup`-closed non-empty subset |
| 50 | +contains its `Sup`. -/ |
| 51 | +def is_sup_closed_compact : Prop := |
| 52 | + ∀ (s : set α) (h : s.nonempty), (∀ a b, a ∈ s → b ∈ s → a ⊔ b ∈ s) → (Sup s) ∈ s |
| 53 | + |
| 54 | +/-- A compactness property for a complete lattice is that any subset has a finite subset with the |
| 55 | +same `Sup`. -/ |
| 56 | +def is_Sup_finite_compact : Prop := |
| 57 | +∀ (s : set α), ∃ (t : finset α), ↑t ⊆ s ∧ Sup s = t.sup id |
| 58 | + |
| 59 | +/-- An element `k` of a complete lattice is said to be compact if any set with `Sup` |
| 60 | +above `k` has a finite subset with `Sup` above `k`. Such an element is also called |
| 61 | +"finite" or "S-compact". -/ |
| 62 | +def is_compact_element {α : Type*} [complete_lattice α] (k : α) := |
| 63 | +∀ s : set α, k ≤ Sup s → ∃ t : finset α, ↑t ⊆ s ∧ k ≤ t.sup id |
| 64 | + |
| 65 | +/-- A complete lattice is said to be compactly generated if any |
| 66 | +element is the `Sup` of compact elements. -/ |
| 67 | +def is_compactly_generated : Prop := |
| 68 | +∀ (x : α), ∃ (s : set α), (∀ x ∈ s, is_compact_element x) ∧ Sup s = x |
| 69 | + |
| 70 | +/-- An element `k` is compact if and only if any directed set with `Sup` above |
| 71 | +`k` already got above `k` at some point in the set. -/ |
| 72 | +theorem is_compact_element_iff_le_of_directed_Sup_le (k : α) : |
| 73 | + is_compact_element k ↔ |
| 74 | + ∀ s : set α, s.nonempty → directed_on (≤) s → k ≤ Sup s → ∃ x : α, x ∈ s ∧ k ≤ x := |
| 75 | +begin |
| 76 | + classical, |
| 77 | + split, |
| 78 | + { by_cases hbot : k = ⊥, |
| 79 | + -- Any nonempty directed set certainly has sup above ⊥ |
| 80 | + { rintros _ _ ⟨x, hx⟩ _ _, use x, by simp only [hx, hbot, bot_le, and_self], }, |
| 81 | + { intros hk s hne hdir hsup, |
| 82 | + obtain ⟨t, ht⟩ := hk s hsup, |
| 83 | + -- If t were empty, its sup would be ⊥, which is not above k ≠ ⊥. |
| 84 | + have tne : t.nonempty, |
| 85 | + { by_contradiction n, |
| 86 | + rw [finset.nonempty_iff_ne_empty, not_not] at n, |
| 87 | + simp only [n, true_and, set.empty_subset, finset.coe_empty, |
| 88 | + finset.sup_empty, le_bot_iff] at ht, |
| 89 | + exact absurd ht hbot, }, |
| 90 | + -- certainly every element of t is below something in s, since ↑t ⊆ s. |
| 91 | + have t_below_s : ∀ x ∈ t, ∃ y ∈ s, x ≤ y, from λ x hxt, ⟨x, ht.left hxt, by refl⟩, |
| 92 | + obtain ⟨x, ⟨hxs, hsupx⟩⟩ := finset.sup_le_of_le_directed s hne hdir t t_below_s, |
| 93 | + exact ⟨x, ⟨hxs, le_trans k (t.sup id) x ht.right hsupx⟩⟩, }, }, |
| 94 | + { intros hk s hsup, |
| 95 | + -- Consider the set of finite joins of elements of the (plain) set s. |
| 96 | + let S : set α := { x | ∃ t : finset α, ↑t ⊆ s ∧ x = t.sup id }, |
| 97 | + -- S is directed, nonempty, and still has sup above k. |
| 98 | + have dir_US : directed_on (≤) S, |
| 99 | + { rintros x ⟨c, hc⟩ y ⟨d, hd⟩, |
| 100 | + use x ⊔ y, |
| 101 | + split, |
| 102 | + { use c ∪ d, |
| 103 | + split, |
| 104 | + { simp only [hc.left, hd.left, set.union_subset_iff, finset.coe_union, and_self], }, |
| 105 | + { simp only [hc.right, hd.right, finset.sup_union], }, }, |
| 106 | + simp only [and_self, le_sup_left, le_sup_right], }, |
| 107 | + have sup_S : Sup s ≤ Sup S, |
| 108 | + { apply Sup_le_Sup, |
| 109 | + intros x hx, use {x}, |
| 110 | + simpa only [and_true, id.def, finset.coe_singleton, eq_self_iff_true, finset.sup_singleton, |
| 111 | + set.singleton_subset_iff], }, |
| 112 | + have Sne : S.nonempty, |
| 113 | + { suffices : ⊥ ∈ S, from set.nonempty_of_mem this, |
| 114 | + use ∅, |
| 115 | + simp only [set.empty_subset, finset.coe_empty, finset.sup_empty, |
| 116 | + eq_self_iff_true, and_self], }, |
| 117 | + -- Now apply the defn of compact and finish. |
| 118 | + obtain ⟨j, ⟨hjS, hjk⟩⟩ := hk S Sne dir_US (le_trans k (Sup s) (Sup S) hsup sup_S), |
| 119 | + obtain ⟨t, ⟨htS, htsup⟩⟩ := hjS, |
| 120 | + use t, exact ⟨htS, by rwa ←htsup⟩, }, |
| 121 | +end |
| 122 | + |
| 123 | +lemma well_founded.is_Sup_finite_compact (h : well_founded ((>) : α → α → Prop)) : |
| 124 | + is_Sup_finite_compact α := |
| 125 | +begin |
| 126 | + intros s, |
| 127 | + let p : set α := { x | ∃ (t : finset α), ↑t ⊆ s ∧ t.sup id = x }, |
| 128 | + have hp : p.nonempty, { use [⊥, ∅], simp, }, |
| 129 | + obtain ⟨m, ⟨t, ⟨ht₁, ht₂⟩⟩, hm⟩ := well_founded.well_founded_iff_has_max'.mp h p hp, |
| 130 | + use t, simp only [ht₁, ht₂, true_and], apply le_antisymm, |
| 131 | + { apply Sup_le, intros y hy, classical, |
| 132 | + have hy' : (insert y t).sup id ∈ p, |
| 133 | + { use insert y t, simp, rw set.insert_subset, exact ⟨hy, ht₁⟩, }, |
| 134 | + have hm' : m ≤ (insert y t).sup id, { rw ← ht₂, exact finset.sup_mono (t.subset_insert y), }, |
| 135 | + rw ← hm _ hy' hm', simp, }, |
| 136 | + { rw [← ht₂, finset.sup_eq_Sup], exact Sup_le_Sup ht₁, }, |
| 137 | +end |
| 138 | + |
| 139 | +lemma is_Sup_finite_compact.is_sup_closed_compact (h : is_Sup_finite_compact α) : |
| 140 | + is_sup_closed_compact α := |
| 141 | +begin |
| 142 | + intros s hne hsc, obtain ⟨t, ht₁, ht₂⟩ := h s, clear h, |
| 143 | + cases t.eq_empty_or_nonempty with h h, |
| 144 | + { subst h, rw finset.sup_empty at ht₂, rw ht₂, |
| 145 | + simp [eq_singleton_bot_of_Sup_eq_bot_of_nonempty ht₂ hne], }, |
| 146 | + { rw ht₂, exact t.sup_closed_of_sup_closed h ht₁ hsc, }, |
| 147 | +end |
| 148 | + |
| 149 | +lemma is_sup_closed_compact.well_founded (h : is_sup_closed_compact α) : |
| 150 | + well_founded ((>) : α → α → Prop) := |
| 151 | +begin |
| 152 | + rw rel_embedding.well_founded_iff_no_descending_seq, rintros ⟨a⟩, |
| 153 | + suffices : Sup (set.range a) ∈ set.range a, |
| 154 | + { obtain ⟨n, hn⟩ := set.mem_range.mp this, |
| 155 | + have h' : Sup (set.range a) < a (n+1), { change _ > _, simp [← hn, a.map_rel_iff], }, |
| 156 | + apply lt_irrefl (a (n+1)), apply lt_of_le_of_lt _ h', apply le_Sup, apply set.mem_range_self, }, |
| 157 | + apply h (set.range a), |
| 158 | + { use a 37, apply set.mem_range_self, }, |
| 159 | + { rintros x y ⟨m, hm⟩ ⟨n, hn⟩, use m ⊔ n, rw [← hm, ← hn], apply a.to_rel_hom.map_sup, }, |
| 160 | +end |
| 161 | + |
| 162 | +lemma is_Sup_finite_compact_iff_all_elements_compact : |
| 163 | + is_Sup_finite_compact α ↔ (∀ k : α, is_compact_element k) := |
| 164 | +begin |
| 165 | + split, |
| 166 | + { intros h k s hs, |
| 167 | + obtain ⟨t, ⟨hts, htsup⟩⟩ := h s, |
| 168 | + use [t, hts], |
| 169 | + rwa ←htsup, }, |
| 170 | + { intros h s, |
| 171 | + obtain ⟨t, ⟨hts, htsup⟩⟩ := h (Sup s) s (by refl), |
| 172 | + have : Sup s = t.sup id, |
| 173 | + { suffices : t.sup id ≤ Sup s, by { apply le_antisymm; assumption }, |
| 174 | + simp only [id.def, finset.sup_le_iff], |
| 175 | + intros x hx, |
| 176 | + apply le_Sup, exact hts hx, }, |
| 177 | + use [t, hts], assumption, }, |
| 178 | +end |
| 179 | + |
| 180 | +lemma well_founded_characterisations : |
| 181 | + tfae [well_founded ((>) : α → α → Prop), |
| 182 | + is_Sup_finite_compact α, |
| 183 | + is_sup_closed_compact α, |
| 184 | + ∀ k : α, is_compact_element k] := |
| 185 | +begin |
| 186 | + tfae_have : 1 → 2, by { exact well_founded.is_Sup_finite_compact α, }, |
| 187 | + tfae_have : 2 → 3, by { exact is_Sup_finite_compact.is_sup_closed_compact α, }, |
| 188 | + tfae_have : 3 → 1, by { exact is_sup_closed_compact.well_founded α, }, |
| 189 | + tfae_have : 2 ↔ 4, by { exact is_Sup_finite_compact_iff_all_elements_compact α }, |
| 190 | + tfae_finish, |
| 191 | +end |
| 192 | + |
| 193 | +lemma well_founded_iff_is_Sup_finite_compact : |
| 194 | + well_founded ((>) : α → α → Prop) ↔ is_Sup_finite_compact α := |
| 195 | +(well_founded_characterisations α).out 0 1 |
| 196 | + |
| 197 | +lemma is_Sup_finite_compact_iff_is_sup_closed_compact : |
| 198 | + is_Sup_finite_compact α ↔ is_sup_closed_compact α := |
| 199 | +(well_founded_characterisations α).out 1 2 |
| 200 | + |
| 201 | +lemma is_sup_closed_compact_iff_well_founded : |
| 202 | + is_sup_closed_compact α ↔ well_founded ((>) : α → α → Prop) := |
| 203 | +(well_founded_characterisations α).out 2 0 |
| 204 | + |
| 205 | +alias well_founded_iff_is_Sup_finite_compact ↔ _ is_Sup_finite_compact.well_founded |
| 206 | +alias is_Sup_finite_compact_iff_is_sup_closed_compact ↔ |
| 207 | + _ is_sup_closed_compact.is_Sup_finite_compact |
| 208 | +alias is_sup_closed_compact_iff_well_founded ↔ _ well_founded.is_sup_closed_compact |
| 209 | + |
| 210 | +lemma compactly_generated_of_well_founded (h : well_founded ((>) : α → α → Prop)) : |
| 211 | + is_compactly_generated α := |
| 212 | +begin |
| 213 | + rw [well_founded_iff_is_Sup_finite_compact, is_Sup_finite_compact_iff_all_elements_compact] at h, |
| 214 | + -- x is the join of the set of compact elements {x} |
| 215 | + exact λ x, ⟨{x}, ⟨λ x _, h x, Sup_singleton⟩⟩, |
| 216 | +end |
| 217 | + |
| 218 | +end complete_lattice |
0 commit comments